From 93d7248fb129e34702bd7d6a472a4eb9ba55072a Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 1 Feb 2022 04:00:15 +0100 Subject: [PATCH 01/73] t --- .github/workflows/ci-ubuntu.yml | 139 ++ .gitignore | 9 + CONTRIBUTING.md | 112 ++ Cubical/Algebra/AbGroup.agda | 4 + Cubical/Algebra/AbGroup/Base.agda | 192 +++ .../Algebra/AbGroup/Instances/DiffInt.agda | 25 + Cubical/Algebra/AbGroup/Instances/Hom.agda | 145 +++ Cubical/Algebra/Algebra.agda | 5 + Cubical/Algebra/Algebra/Base.agda | 269 ++++ Cubical/Algebra/Algebra/Properties.agda | 170 +++ Cubical/Algebra/CommAlgebra.agda | 5 + Cubical/Algebra/CommAlgebra/Base.agda | 250 ++++ Cubical/Algebra/CommAlgebra/FGIdeal.agda | 22 + Cubical/Algebra/CommAlgebra/FPAlgebra.agda | 43 + Cubical/Algebra/CommAlgebra/Ideal.agda | 27 + .../Instances/FreeCommAlgebra.agda | 456 +++++++ .../CommAlgebra/Instances/Initial.agda | 80 ++ .../CommAlgebra/Instances/Pointwise.agda | 26 + Cubical/Algebra/CommAlgebra/Localisation.agda | 225 ++++ Cubical/Algebra/CommAlgebra/Properties.agda | 258 ++++ .../Algebra/CommAlgebra/QuotientAlgebra.agda | 64 + Cubical/Algebra/CommMonoid.agda | 5 + Cubical/Algebra/CommMonoid/Base.agda | 125 ++ Cubical/Algebra/CommMonoid/Properties.agda | 42 + Cubical/Algebra/CommRing.agda | 5 + Cubical/Algebra/CommRing/Base.agda | 153 +++ Cubical/Algebra/CommRing/BinomialThm.agda | 132 ++ Cubical/Algebra/CommRing/FGIdeal.agda | 345 +++++ Cubical/Algebra/CommRing/FiberedProduct.agda | 100 ++ Cubical/Algebra/CommRing/Ideal.agda | 331 +++++ .../Algebra/CommRing/Instances/BiInvInt.agda | 26 + Cubical/Algebra/CommRing/Instances/Int.agda | 26 + .../Algebra/CommRing/Instances/Pointwise.agda | 36 + Cubical/Algebra/CommRing/Instances/Unit.agda | 27 + .../Algebra/CommRing/Localisation/Base.agda | 288 +++++ .../Localisation/InvertingElements.agda | 708 ++++++++++ .../CommRing/Localisation/PullbackSquare.agda | 586 +++++++++ .../Localisation/UniversalProperty.agda | 417 ++++++ Cubical/Algebra/CommRing/Properties.agda | 339 +++++ Cubical/Algebra/CommRing/QuotientRing.agda | 31 + Cubical/Algebra/CommRing/RadicalIdeal.agda | 235 ++++ Cubical/Algebra/DistLattice.agda | 4 + Cubical/Algebra/DistLattice/Base.agda | 276 ++++ Cubical/Algebra/DistLattice/Basis.agda | 79 ++ Cubical/Algebra/DistLattice/BigOps.agda | 162 +++ Cubical/Algebra/Group.agda | 15 + Cubical/Algebra/Group/Base.agda | 186 +++ Cubical/Algebra/Group/DirProd.agda | 33 + Cubical/Algebra/Group/EilenbergMacLane1.agda | 151 +++ Cubical/Algebra/Group/Exact.agda | 54 + Cubical/Algebra/Group/GroupPath.agda | 146 +++ Cubical/Algebra/Group/Instances/Bool.agda | 93 ++ Cubical/Algebra/Group/Instances/DiffInt.agda | 24 + Cubical/Algebra/Group/Instances/Int.agda | 21 + Cubical/Algebra/Group/Instances/IntMod.agda | 69 + Cubical/Algebra/Group/Instances/Unit.agda | 70 + .../Algebra/Group/IsomorphismTheorems.agda | 105 ++ Cubical/Algebra/Group/MorphismProperties.agda | 278 ++++ Cubical/Algebra/Group/Morphisms.agda | 106 ++ Cubical/Algebra/Group/Properties.agda | 97 ++ Cubical/Algebra/Group/QuotientGroup.agda | 105 ++ Cubical/Algebra/Group/Subgroup.agda | 196 +++ Cubical/Algebra/Group/ZAction.agda | 469 +++++++ Cubical/Algebra/Lattice.agda | 5 + Cubical/Algebra/Lattice/Base.agda | 231 ++++ Cubical/Algebra/Lattice/Properties.agda | 95 ++ Cubical/Algebra/Matrix.agda | 336 +++++ Cubical/Algebra/Module.agda | 4 + Cubical/Algebra/Module/Base.agda | 165 +++ Cubical/Algebra/Monoid.agda | 4 + Cubical/Algebra/Monoid/Base.agda | 151 +++ Cubical/Algebra/Monoid/BigOp.agda | 59 + Cubical/Algebra/NatSolver/EvalHom.agda | 196 +++ Cubical/Algebra/NatSolver/Examples.agda | 147 +++ Cubical/Algebra/NatSolver/HornerForms.agda | 100 ++ Cubical/Algebra/NatSolver/NatExpression.agda | 28 + Cubical/Algebra/NatSolver/Reflection.agda | 251 ++++ Cubical/Algebra/NatSolver/Solver.agda | 123 ++ Cubical/Algebra/Ring.agda | 5 + Cubical/Algebra/Ring/Base.agda | 276 ++++ Cubical/Algebra/Ring/BigOps.agda | 119 ++ Cubical/Algebra/Ring/Ideal.agda | 81 ++ Cubical/Algebra/Ring/Kernel.agda | 56 + Cubical/Algebra/Ring/Properties.agda | 312 +++++ Cubical/Algebra/Ring/QuotientRing.agda | 236 ++++ .../Algebra/RingSolver/AlgebraExpression.agda | 38 + Cubical/Algebra/RingSolver/EvalHom.agda | 275 ++++ Cubical/Algebra/RingSolver/Examples.agda | 97 ++ Cubical/Algebra/RingSolver/HornerEval.agda | 133 ++ Cubical/Algebra/RingSolver/HornerForms.agda | 184 +++ Cubical/Algebra/RingSolver/IntAsRawRing.agda | 15 + Cubical/Algebra/RingSolver/README.md | 25 + Cubical/Algebra/RingSolver/RawAlgebra.agda | 177 +++ Cubical/Algebra/RingSolver/RawRing.agda | 29 + Cubical/Algebra/RingSolver/Reflection.agda | 328 +++++ Cubical/Algebra/RingSolver/Solver.agda | 145 +++ Cubical/Algebra/RingSolver/Utility.agda | 37 + Cubical/Algebra/Semigroup.agda | 4 + Cubical/Algebra/Semigroup/Base.agda | 100 ++ Cubical/Algebra/Semilattice.agda | 4 + Cubical/Algebra/Semilattice/Base.agda | 233 ++++ Cubical/Algebra/SymmetricGroup.agda | 26 + Cubical/Algebra/ZariskiLattice/Base.agda | 585 +++++++++ Cubical/Categories/Adjoint.agda | 318 +++++ Cubical/Categories/Category.agda | 21 + Cubical/Categories/Category/Base.agda | 102 ++ Cubical/Categories/Category/Precategory.agda | 54 + Cubical/Categories/Category/Properties.agda | 92 ++ Cubical/Categories/Commutativity.agda | 33 + .../Categories/Constructions/BinProduct.agda | 64 + .../Categories/Constructions/Elements.agda | 138 ++ Cubical/Categories/Constructions/Product.agda | 64 + .../Categories/Constructions/Quotient.agda | 83 ++ Cubical/Categories/Constructions/Slice.agda | 377 ++++++ Cubical/Categories/DistLatticeSheaf.agda | 230 ++++ Cubical/Categories/Equivalence.agda | 7 + Cubical/Categories/Equivalence/Base.agda | 30 + .../Categories/Equivalence/Properties.agda | 92 ++ Cubical/Categories/Functor.agda | 8 + Cubical/Categories/Functor/Base.agda | 79 ++ Cubical/Categories/Functor/BinProduct.agda | 26 + Cubical/Categories/Functor/Compose.agda | 44 + Cubical/Categories/Functor/Product.agda | 26 + Cubical/Categories/Functor/Properties.agda | 116 ++ Cubical/Categories/Instances/AbGroups.agda | 23 + Cubical/Categories/Instances/Algebras.agda | 28 + Cubical/Categories/Instances/Categories.agda | 25 + .../Categories/Instances/CommAlgebras.agda | 62 + Cubical/Categories/Instances/CommRings.agda | 66 + Cubical/Categories/Instances/Cospan.agda | 77 ++ Cubical/Categories/Instances/Discrete.agda | 55 + Cubical/Categories/Instances/DistLattice.agda | 14 + Cubical/Categories/Instances/Functors.agda | 52 + Cubical/Categories/Instances/Lattice.agda | 15 + Cubical/Categories/Instances/Poset.agda | 28 + Cubical/Categories/Instances/Rings.agda | 22 + Cubical/Categories/Instances/Semilattice.agda | 25 + Cubical/Categories/Instances/Sets.agda | 92 ++ .../Categories/Instances/TypePrecategory.agda | 19 + Cubical/Categories/Limits.agda | 9 + Cubical/Categories/Limits/BinCoproduct.agda | 43 + Cubical/Categories/Limits/BinProduct.agda | 43 + Cubical/Categories/Limits/Initial.agda | 61 + Cubical/Categories/Limits/Limits.agda | 139 ++ Cubical/Categories/Limits/Pullback.agda | 153 +++ Cubical/Categories/Limits/Terminal.agda | 61 + Cubical/Categories/Monoidal.agda | 7 + Cubical/Categories/Monoidal/Base.agda | 121 ++ Cubical/Categories/Monoidal/Enriched.agda | 29 + Cubical/Categories/Morphism.agda | 129 ++ Cubical/Categories/NaturalTransformation.agda | 6 + .../NaturalTransformation/Base.agda | 216 ++++ .../NaturalTransformation/Properties.agda | 102 ++ Cubical/Categories/Presheaf.agda | 7 + Cubical/Categories/Presheaf/Base.agda | 13 + Cubical/Categories/Presheaf/KanExtension.agda | 341 +++++ Cubical/Categories/Presheaf/Properties.agda | 389 ++++++ .../TypesOfCategories/TypeCategory.agda | 191 +++ Cubical/Categories/Yoneda.agda | 179 +++ Cubical/Codata/Conat.agda | 7 + Cubical/Codata/Conat/Base.agda | 55 + Cubical/Codata/Conat/Properties.agda | 188 +++ Cubical/Codata/Everything.agda | 32 + Cubical/Codata/EverythingSafe.agda | 2 + Cubical/Codata/M.agda | 95 ++ Cubical/Codata/M/AsLimit/Coalg.agda | 5 + Cubical/Codata/M/AsLimit/Coalg/Base.agda | 39 + Cubical/Codata/M/AsLimit/Container.agda | 88 ++ Cubical/Codata/M/AsLimit/M.agda | 6 + Cubical/Codata/M/AsLimit/M/Base.agda | 194 +++ Cubical/Codata/M/AsLimit/M/Properties.agda | 61 + Cubical/Codata/M/AsLimit/helper.agda | 116 ++ Cubical/Codata/M/AsLimit/itree.agda | 75 ++ Cubical/Codata/M/AsLimit/stream.agda | 30 + Cubical/Codata/M/Bisimilarity.agda | 147 +++ Cubical/Codata/Stream.agda | 7 + Cubical/Codata/Stream/Base.agda | 11 + Cubical/Codata/Stream/Properties.agda | 121 ++ Cubical/Core/Everything.agda | 11 + Cubical/Core/Glue.agda | 141 ++ Cubical/Core/Id.agda | 24 + Cubical/Core/Primitives.agda | 211 +++ Cubical/Core/README.md | 16 + Cubical/Data/BinNat.agda | 4 + Cubical/Data/BinNat/BinNat.agda | 492 +++++++ Cubical/Data/Bool.agda | 6 + Cubical/Data/Bool/Base.agda | 85 ++ Cubical/Data/Bool/Properties.agda | 221 ++++ Cubical/Data/Bool/SwitchStatement.agda | 42 + Cubical/Data/DescendingList.agda | 9 + Cubical/Data/DescendingList/Base.agda | 33 + Cubical/Data/DescendingList/Examples.agda | 121 ++ Cubical/Data/DescendingList/Properties.agda | 279 ++++ Cubical/Data/DescendingList/Strict.agda | 14 + .../DescendingList/Strict/Properties.agda | 227 ++++ Cubical/Data/Empty.agda | 6 + Cubical/Data/Empty/Base.agda | 20 + Cubical/Data/Empty/Properties.agda | 34 + Cubical/Data/Equality.agda | 50 + Cubical/Data/Fin.agda | 7 + Cubical/Data/Fin/Arithmetic.agda | 87 ++ Cubical/Data/Fin/Base.agda | 97 ++ Cubical/Data/Fin/LehmerCode.agda | 222 ++++ Cubical/Data/Fin/Literals.agda | 18 + Cubical/Data/Fin/Properties.agda | 636 +++++++++ Cubical/Data/Fin/Recursive.agda | 6 + Cubical/Data/Fin/Recursive/Base.agda | 34 + Cubical/Data/Fin/Recursive/Properties.agda | 266 ++++ Cubical/Data/FinData.agda | 5 + Cubical/Data/FinData/Base.agda | 80 ++ Cubical/Data/FinData/Properties.agda | 178 +++ Cubical/Data/FinInd.agda | 48 + Cubical/Data/FinSet.agda | 6 + Cubical/Data/FinSet/Base.agda | 34 + Cubical/Data/FinSet/Binary/Large.agda | 147 +++ Cubical/Data/FinSet/Binary/Small.agda | 6 + Cubical/Data/FinSet/Binary/Small/Base.agda | 20 + .../Data/FinSet/Binary/Small/Properties.agda | 115 ++ Cubical/Data/FinSet/Constructors.agda | 196 +++ Cubical/Data/FinSet/FiniteChoice.agda | 72 ++ Cubical/Data/FinSet/Properties.agda | 152 +++ Cubical/Data/Graph.agda | 5 + Cubical/Data/Graph/Base.agda | 55 + Cubical/Data/Graph/Examples.agda | 161 +++ Cubical/Data/HomotopyGroup.agda | 4 + Cubical/Data/HomotopyGroup/Base.agda | 52 + Cubical/Data/InfNat.agda | 5 + Cubical/Data/InfNat/Base.agda | 36 + Cubical/Data/InfNat/Properties.agda | 26 + Cubical/Data/Int.agda | 7 + Cubical/Data/Int/Base.agda | 85 ++ Cubical/Data/Int/MoreInts/BiInvInt.agda | 6 + Cubical/Data/Int/MoreInts/BiInvInt/Base.agda | 310 +++++ .../Int/MoreInts/BiInvInt/Properties.agda | 243 ++++ Cubical/Data/Int/MoreInts/DeltaInt.agda | 6 + Cubical/Data/Int/MoreInts/DeltaInt/Base.agda | 61 + .../Int/MoreInts/DeltaInt/Properties.agda | 86 ++ Cubical/Data/Int/MoreInts/DiffInt.agda | 5 + Cubical/Data/Int/MoreInts/DiffInt/Base.agda | 84 ++ .../Data/Int/MoreInts/DiffInt/Properties.agda | 227 ++++ Cubical/Data/Int/MoreInts/QuoInt.agda | 6 + Cubical/Data/Int/MoreInts/QuoInt/Base.agda | 215 ++++ .../Data/Int/MoreInts/QuoInt/Properties.agda | 235 ++++ Cubical/Data/Int/Properties.agda | 487 +++++++ Cubical/Data/List.agda | 5 + Cubical/Data/List/Base.agda | 52 + Cubical/Data/List/Properties.agda | 176 +++ Cubical/Data/Maybe.agda | 5 + Cubical/Data/Maybe/Base.agda | 25 + Cubical/Data/Maybe/Properties.agda | 158 +++ Cubical/Data/Nat.agda | 6 + Cubical/Data/Nat/Algebra.agda | 299 +++++ Cubical/Data/Nat/Base.agda | 61 + Cubical/Data/Nat/Coprime.agda | 95 ++ Cubical/Data/Nat/Divisibility.agda | 104 ++ Cubical/Data/Nat/GCD.agda | 161 +++ Cubical/Data/Nat/Literals.agda | 22 + Cubical/Data/Nat/Mod.agda | 153 +++ Cubical/Data/Nat/Order.agda | 362 ++++++ Cubical/Data/Nat/Order/Recursive.agda | 104 ++ Cubical/Data/Nat/Properties.agda | 206 +++ Cubical/Data/NatMinusOne.agda | 6 + Cubical/Data/NatMinusOne/Base.agda | 33 + Cubical/Data/NatMinusOne/Properties.agda | 13 + Cubical/Data/NatPlusOne.agda | 6 + Cubical/Data/NatPlusOne/Base.agda | 32 + Cubical/Data/NatPlusOne/Properties.agda | 38 + Cubical/Data/Prod.agda | 5 + Cubical/Data/Prod/Base.agda | 51 + Cubical/Data/Prod/Properties.agda | 111 ++ Cubical/Data/Queue.agda | 7 + Cubical/Data/Queue/1List.agda | 87 ++ Cubical/Data/Queue/Base.agda | 6 + Cubical/Data/Queue/Finite.agda | 58 + Cubical/Data/Queue/Truncated2List.agda | 150 +++ Cubical/Data/Queue/Untruncated2List.agda | 166 +++ .../Data/Queue/Untruncated2ListInvariant.agda | 81 ++ Cubical/Data/Sigma.agda | 6 + Cubical/Data/Sigma/Base.agda | 52 + Cubical/Data/Sigma/Properties.agda | 391 ++++++ Cubical/Data/SubFinSet.agda | 49 + Cubical/Data/Sum.agda | 5 + Cubical/Data/Sum/Base.agda | 37 + Cubical/Data/Sum/Properties.agda | 186 +++ Cubical/Data/SumFin.agda | 5 + Cubical/Data/SumFin/Base.agda | 71 + Cubical/Data/SumFin/Properties.agda | 92 ++ Cubical/Data/Unit.agda | 5 + Cubical/Data/Unit/Base.agda | 19 + Cubical/Data/Unit/Pointed.agda | 14 + Cubical/Data/Unit/Properties.agda | 96 ++ Cubical/Data/Vec.agda | 6 + Cubical/Data/Vec/Base.agda | 67 + Cubical/Data/Vec/NAry.agda | 55 + Cubical/Data/Vec/Properties.agda | 60 + Cubical/Displayed/Auto.agda | 367 ++++++ Cubical/Displayed/Base.agda | 87 ++ Cubical/Displayed/Constant.agda | 39 + Cubical/Displayed/Function.agda | 135 ++ Cubical/Displayed/Generic.agda | 36 + Cubical/Displayed/Morphism.agda | 67 + Cubical/Displayed/Prop.agda | 52 + Cubical/Displayed/Properties.agda | 141 ++ Cubical/Displayed/Record.agda | 231 ++++ Cubical/Displayed/Sigma.agda | 101 ++ Cubical/Displayed/Subst.agda | 60 + Cubical/Displayed/Unit.agda | 27 + Cubical/Displayed/Universe.agda | 32 + Cubical/Experiments/Brunerie.agda | 298 +++++ Cubical/Experiments/CohomologyGroups.agda | 140 ++ Cubical/Experiments/Combinatorics.agda | 59 + Cubical/Experiments/EscardoSIP.agda | 218 ++++ Cubical/Experiments/Everything.agda | 17 + Cubical/Experiments/FunExtFromUA.agda | 116 ++ Cubical/Experiments/Generic.agda | 120 ++ Cubical/Experiments/HAEquivInt.agda | 5 + Cubical/Experiments/HAEquivInt/Base.agda | 23 + Cubical/Experiments/HInt.agda | 98 ++ Cubical/Experiments/HoTT-UF.agda | 104 ++ Cubical/Experiments/IsoInt.agda | 5 + Cubical/Experiments/IsoInt/Base.agda | 112 ++ Cubical/Experiments/List.agda | 285 +++++ Cubical/Experiments/LocalisationDefs.agda | 146 +++ Cubical/Experiments/NatMinusTwo.agda | 21 + Cubical/Experiments/NatMinusTwo/Base.agda | 41 + .../Experiments/NatMinusTwo/Properties.agda | 11 + .../NatMinusTwo/ToNatMinusOne.agda | 24 + Cubical/Experiments/Poset.agda | 326 +++++ Cubical/Experiments/Problem.agda | 150 +++ .../Experiments/ZCohomology/Benchmarks.agda | 388 ++++++ .../Experiments/ZCohomologyExperiments.agda | 79 ++ Cubical/Experiments/ZCohomologyOld/Base.agda | 46 + .../ZCohomologyOld/KcompPrelims.agda | 213 +++ .../MayerVietorisUnreduced.agda | 384 ++++++ .../ZCohomologyOld/Properties.agda | 634 +++++++++ .../Experiments/ZariskiLatticeBasicOpens.agda | 249 ++++ Cubical/Foundations/CartesianKanOps.agda | 149 +++ Cubical/Foundations/Equiv.agda | 307 +++++ Cubical/Foundations/Equiv/Base.agda | 31 + Cubical/Foundations/Equiv/BiInvertible.agda | 82 ++ Cubical/Foundations/Equiv/Fiberwise.agda | 85 ++ Cubical/Foundations/Equiv/HalfAdjoint.agda | 166 +++ Cubical/Foundations/Equiv/PathSplit.agda | 140 ++ Cubical/Foundations/Equiv/Properties.agda | 200 +++ Cubical/Foundations/Everything.agda | 51 + Cubical/Foundations/Function.agda | 154 +++ Cubical/Foundations/GroupoidLaws.agda | 497 +++++++ Cubical/Foundations/HLevels.agda | 658 ++++++++++ Cubical/Foundations/Id.agda | 313 +++++ Cubical/Foundations/Isomorphism.agda | 185 +++ Cubical/Foundations/Path.agda | 362 ++++++ Cubical/Foundations/Pointed.agda | 9 + Cubical/Foundations/Pointed/Base.agda | 128 ++ Cubical/Foundations/Pointed/FunExt.agda | 48 + Cubical/Foundations/Pointed/Homogeneous.agda | 141 ++ Cubical/Foundations/Pointed/Homotopy.agda | 119 ++ Cubical/Foundations/Pointed/Properties.agda | 99 ++ Cubical/Foundations/Powerset.agda | 65 + Cubical/Foundations/Prelude.agda | 510 ++++++++ Cubical/Foundations/RelationalStructure.agda | 268 ++++ Cubical/Foundations/SIP.agda | 124 ++ Cubical/Foundations/Structure.agda | 49 + Cubical/Foundations/Transport.agda | 154 +++ Cubical/Foundations/Univalence.agda | 309 +++++ Cubical/Foundations/Univalence/Universe.agda | 108 ++ Cubical/Functions/Bundle.agda | 50 + Cubical/Functions/Embedding.agda | 427 +++++++ Cubical/Functions/Fibration.agda | 111 ++ Cubical/Functions/Fixpoint.agda | 43 + Cubical/Functions/FunExtEquiv.agda | 193 +++ Cubical/Functions/Implicit.agda | 17 + Cubical/Functions/Involution.agda | 42 + Cubical/Functions/Logic.agda | 290 +++++ Cubical/Functions/Morphism.agda | 84 ++ Cubical/Functions/Surjection.agda | 80 ++ Cubical/HITs/2GroupoidTruncation.agda | 5 + Cubical/HITs/2GroupoidTruncation/Base.agda | 17 + .../HITs/2GroupoidTruncation/Properties.agda | 68 + Cubical/HITs/AssocList.agda | 6 + Cubical/HITs/AssocList/Base.agda | 69 + Cubical/HITs/AssocList/Properties.agda | 162 +++ Cubical/HITs/Colimit.agda | 5 + Cubical/HITs/Colimit/Base.agda | 132 ++ Cubical/HITs/Colimit/Examples.agda | 76 ++ Cubical/HITs/Cost.agda | 4 + Cubical/HITs/Cost/Base.agda | 124 ++ Cubical/HITs/Cylinder.agda | 6 + Cubical/HITs/Cylinder/Base.agda | 265 ++++ Cubical/HITs/Delooping/Two/Base.agda | 55 + Cubical/HITs/Delooping/Two/Properties.agda | 198 +++ Cubical/HITs/DunceCap.agda | 6 + Cubical/HITs/DunceCap/Base.agda | 27 + Cubical/HITs/DunceCap/Properties.agda | 64 + Cubical/HITs/EilenbergMacLane1.agda | 5 + Cubical/HITs/EilenbergMacLane1/Base.agda | 43 + .../HITs/EilenbergMacLane1/Properties.agda | 100 ++ Cubical/HITs/FiniteMultiset.agda | 6 + Cubical/HITs/FiniteMultiset/Base.agda | 50 + .../FiniteMultiset/CountExtensionality.agda | 191 +++ Cubical/HITs/FiniteMultiset/Properties.agda | 219 ++++ Cubical/HITs/GroupoidQuotients.agda | 5 + Cubical/HITs/GroupoidQuotients/Base.agda | 43 + .../HITs/GroupoidQuotients/Properties.agda | 106 ++ Cubical/HITs/GroupoidTruncation.agda | 5 + Cubical/HITs/GroupoidTruncation/Base.agda | 17 + .../HITs/GroupoidTruncation/Properties.agda | 67 + Cubical/HITs/InfNat.agda | 6 + Cubical/HITs/InfNat/Base.agda | 16 + Cubical/HITs/InfNat/Properties.agda | 61 + Cubical/HITs/Interval.agda | 6 + Cubical/HITs/Interval/Base.agda | 39 + Cubical/HITs/Join.agda | 5 + Cubical/HITs/Join/Base.agda | 87 ++ Cubical/HITs/Join/Properties.agda | 476 +++++++ Cubical/HITs/KleinBottle.agda | 5 + Cubical/HITs/KleinBottle/Base.agda | 15 + Cubical/HITs/KleinBottle/Properties.agda | 159 +++ Cubical/HITs/ListedFiniteSet.agda | 5 + Cubical/HITs/ListedFiniteSet/Base.agda | 115 ++ Cubical/HITs/ListedFiniteSet/Properties.agda | 106 ++ Cubical/HITs/Localization.agda | 6 + Cubical/HITs/Localization/Base.agda | 26 + Cubical/HITs/Localization/Properties.agda | 23 + Cubical/HITs/MappingCones.agda | 6 + Cubical/HITs/MappingCones/Base.agda | 26 + Cubical/HITs/MappingCones/Properties.agda | 51 + Cubical/HITs/Modulo.agda | 8 + Cubical/HITs/Modulo/Base.agda | 60 + Cubical/HITs/Modulo/FinEquiv.agda | 77 ++ Cubical/HITs/Modulo/Properties.agda | 80 ++ Cubical/HITs/Nullification.agda | 6 + Cubical/HITs/Nullification/Base.agda | 25 + Cubical/HITs/Nullification/Properties.agda | 140 ++ Cubical/HITs/PropositionalTruncation.agda | 7 + .../HITs/PropositionalTruncation/Base.agda | 17 + .../PropositionalTruncation/MagicTrick.agda | 88 ++ .../HITs/PropositionalTruncation/Monad.agda | 33 + .../PropositionalTruncation/Properties.agda | 499 ++++++++ Cubical/HITs/Pushout.agda | 6 + Cubical/HITs/Pushout/Base.agda | 117 ++ Cubical/HITs/Pushout/Flattening.agda | 91 ++ Cubical/HITs/Pushout/KrausVonRaumer.agda | 198 +++ Cubical/HITs/Pushout/Properties.agda | 543 ++++++++ Cubical/HITs/RPn.agda | 4 + Cubical/HITs/RPn/Base.agda | 320 +++++ Cubical/HITs/Rationals/HITQ.agda | 6 + Cubical/HITs/Rationals/HITQ/Base.agda | 36 + Cubical/HITs/Rationals/QuoQ.agda | 6 + Cubical/HITs/Rationals/QuoQ/Base.agda | 76 ++ Cubical/HITs/Rationals/QuoQ/Properties.agda | 245 ++++ Cubical/HITs/Rationals/SigmaQ.agda | 6 + Cubical/HITs/Rationals/SigmaQ/Base.agda | 54 + Cubical/HITs/Rationals/SigmaQ/Properties.agda | 78 ++ Cubical/HITs/S1.agda | 5 + Cubical/HITs/S1/Base.agda | 493 +++++++ Cubical/HITs/S1/Properties.agda | 34 + Cubical/HITs/S2.agda | 6 + Cubical/HITs/S2/Base.agda | 57 + Cubical/HITs/S3.agda | 6 + Cubical/HITs/S3/Base.agda | 8 + Cubical/HITs/SetQuotients.agda | 5 + Cubical/HITs/SetQuotients/Base.agda | 17 + Cubical/HITs/SetQuotients/Properties.agda | 334 +++++ Cubical/HITs/SetTruncation.agda | 5 + Cubical/HITs/SetTruncation/Base.agda | 23 + Cubical/HITs/SetTruncation/Properties.agda | 328 +++++ Cubical/HITs/SmashProduct.agda | 6 + Cubical/HITs/SmashProduct/Base.agda | 213 +++ Cubical/HITs/Sn.agda | 5 + Cubical/HITs/Sn/Base.agda | 29 + Cubical/HITs/Sn/Properties.agda | 476 +++++++ Cubical/HITs/Susp.agda | 6 + Cubical/HITs/Susp/Base.agda | 195 +++ Cubical/HITs/Susp/LoopAdjunction.agda | 113 ++ Cubical/HITs/Susp/Properties.agda | 193 +++ Cubical/HITs/Torus.agda | 6 + Cubical/HITs/Torus/Base.agda | 95 ++ Cubical/HITs/Truncation.agda | 7 + Cubical/HITs/Truncation/Base.agda | 43 + Cubical/HITs/Truncation/FromNegTwo.agda | 5 + Cubical/HITs/Truncation/FromNegTwo/Base.agda | 29 + .../Truncation/FromNegTwo/Properties.agda | 396 ++++++ Cubical/HITs/Truncation/Properties.agda | 566 ++++++++ Cubical/HITs/TypeQuotients.agda | 5 + Cubical/HITs/TypeQuotients/Base.agda | 17 + Cubical/HITs/TypeQuotients/Properties.agda | 60 + Cubical/HITs/Wedge.agda | 4 + Cubical/HITs/Wedge/Base.agda | 59 + Cubical/Homotopy/Base.agda | 19 + Cubical/Homotopy/BlakersMassey.agda | 809 ++++++++++++ Cubical/Homotopy/Connected.agda | 405 ++++++ Cubical/Homotopy/EilenbergSteenrod.agda | 47 + Cubical/Homotopy/Freudenthal.agda | 135 ++ Cubical/Homotopy/Group/Base.agda | 1046 +++++++++++++++ Cubical/Homotopy/Group/LES.agda | 665 ++++++++++ Cubical/Homotopy/Group/Pi3S2.agda | 154 +++ .../Homotopy/Group/Pi4S3/S3PushoutIso.agda | 682 ++++++++++ .../Homotopy/Group/Pi4S3/S3PushoutIso2.agda | 64 + Cubical/Homotopy/Group/Pi4S3/Summary.agda | 133 ++ Cubical/Homotopy/Group/Pi4S3/Tricky.agda | 772 +++++++++++ Cubical/Homotopy/Group/PinSn.agda | 334 +++++ Cubical/Homotopy/Group/SuspensionMap.agda | 608 +++++++++ Cubical/Homotopy/HSpace.agda | 51 + Cubical/Homotopy/Hopf.agda | 704 ++++++++++ Cubical/Homotopy/HopfInvariant/Base.agda | 313 +++++ .../Homotopy/HopfInvariant/Homomorphism.agda | 884 +++++++++++++ Cubical/Homotopy/HopfInvariant/HopfMap.agda | 559 ++++++++ Cubical/Homotopy/HopfInvariant/Whitehead.agda | 808 ++++++++++++ Cubical/Homotopy/Loopspace.agda | 449 +++++++ Cubical/Homotopy/MayerVietorisCofiber.agda | 175 +++ Cubical/Homotopy/Prespectrum.agda | 54 + Cubical/Homotopy/Spectrum.agda | 22 + Cubical/Homotopy/WedgeConnectivity.agda | 66 + Cubical/Homotopy/Whitehead.agda | 407 ++++++ Cubical/Induction/WellFounded.agda | 47 + Cubical/Modalities/Lex.agda | 271 ++++ Cubical/Modalities/Modality.agda | 155 +++ .../Papers/RepresentationIndependence.agda | 314 +++++ Cubical/Papers/Synthetic.agda | 214 ++++ Cubical/Papers/ZCohomology.agda | 536 ++++++++ Cubical/README.agda | 73 ++ Cubical/Reflection/Base.agda | 59 + Cubical/Reflection/RecordEquiv.agda | 196 +++ Cubical/Reflection/StrictEquiv.agda | 81 ++ Cubical/Relation/Binary.agda | 6 + Cubical/Relation/Binary/Base.agda | 148 +++ Cubical/Relation/Binary/Poset.agda | 138 ++ Cubical/Relation/Binary/Properties.agda | 5 + Cubical/Relation/Nullary.agda | 5 + Cubical/Relation/Nullary/Base.agda | 66 + Cubical/Relation/Nullary/DecidableEq.agda | 5 + Cubical/Relation/Nullary/HLevels.agda | 38 + Cubical/Relation/Nullary/Properties.agda | 168 +++ .../ZigZag/Applications/MultiSet.agda | 358 ++++++ Cubical/Relation/ZigZag/Base.agda | 161 +++ Cubical/Structures/Auto.agda | 250 ++++ Cubical/Structures/Axioms.agda | 69 + Cubical/Structures/Constant.agda | 35 + Cubical/Structures/Function.agda | 82 ++ Cubical/Structures/LeftAction.agda | 17 + Cubical/Structures/Macro.agda | 157 +++ Cubical/Structures/Maybe.agda | 108 ++ Cubical/Structures/MultiSet.agda | 42 + Cubical/Structures/Parameterized.agda | 48 + Cubical/Structures/Pointed.agda | 41 + Cubical/Structures/Product.agda | 50 + Cubical/Structures/Queue.agda | 115 ++ Cubical/Structures/Record.agda | 453 +++++++ Cubical/Structures/Relational/Auto.agda | 239 ++++ Cubical/Structures/Relational/Constant.agda | 57 + Cubical/Structures/Relational/Equalizer.agda | 67 + Cubical/Structures/Relational/Function.agda | 191 +++ Cubical/Structures/Relational/Macro.agda | 161 +++ Cubical/Structures/Relational/Maybe.agda | 127 ++ .../Structures/Relational/Parameterized.agda | 70 + Cubical/Structures/Relational/Pointed.agda | 56 + Cubical/Structures/Relational/Product.agda | 140 ++ Cubical/Structures/Successor.agda | 29 + Cubical/Structures/Transfer.agda | 47 + Cubical/Structures/TypeEqvTo.agda | 41 + .../Syntax/\342\237\250\342\237\251.agda" | 12 + Cubical/Talks/EPA2020.agda | 343 +++++ Cubical/WithK.agda | 48 + Cubical/ZCohomology/Base.agda | 51 + Cubical/ZCohomology/EilenbergSteenrodZ.agda | 379 ++++++ Cubical/ZCohomology/GroupStructure.agda | 930 ++++++++++++++ Cubical/ZCohomology/Groups/CP2.agda | 249 ++++ Cubical/ZCohomology/Groups/Connected.agda | 45 + Cubical/ZCohomology/Groups/KleinBottle.agda | 454 +++++++ Cubical/ZCohomology/Groups/Prelims.agda | 280 ++++ Cubical/ZCohomology/Groups/RP2.agda | 128 ++ Cubical/ZCohomology/Groups/Sn.agda | 330 +++++ Cubical/ZCohomology/Groups/Torus.agda | 309 +++++ Cubical/ZCohomology/Groups/Unit.agda | 85 ++ Cubical/ZCohomology/Groups/Wedge.agda | 243 ++++ .../ZCohomology/Groups/WedgeOfSpheres.agda | 134 ++ Cubical/ZCohomology/Gysin.agda | 1137 +++++++++++++++++ .../ZCohomology/MayerVietorisUnreduced.agda | 250 ++++ Cubical/ZCohomology/Properties.agda | 676 ++++++++++ .../ZCohomology/RingStructure/CupProduct.agda | 87 ++ .../RingStructure/GradedCommutativity.agda | 1105 ++++++++++++++++ .../ZCohomology/RingStructure/RingLaws.agda | 686 ++++++++++ Everythings.hs | 124 ++ GNUmakefile | 67 + INSTALL.md | 281 ++++ LICENSE | 21 + MAKEWINDOWS.md | 32 + README.md | 62 + RELEASE.md | 16 + cubical.agda-lib | 4 + fix-whitespace.yaml | 40 + 591 files changed, 79011 insertions(+) create mode 100644 .github/workflows/ci-ubuntu.yml create mode 100644 .gitignore create mode 100644 CONTRIBUTING.md create mode 100644 Cubical/Algebra/AbGroup.agda create mode 100644 Cubical/Algebra/AbGroup/Base.agda create mode 100644 Cubical/Algebra/AbGroup/Instances/DiffInt.agda create mode 100644 Cubical/Algebra/AbGroup/Instances/Hom.agda create mode 100644 Cubical/Algebra/Algebra.agda create mode 100644 Cubical/Algebra/Algebra/Base.agda create mode 100644 Cubical/Algebra/Algebra/Properties.agda create mode 100644 Cubical/Algebra/CommAlgebra.agda create mode 100644 Cubical/Algebra/CommAlgebra/Base.agda create mode 100644 Cubical/Algebra/CommAlgebra/FGIdeal.agda create mode 100644 Cubical/Algebra/CommAlgebra/FPAlgebra.agda create mode 100644 Cubical/Algebra/CommAlgebra/Ideal.agda create mode 100644 Cubical/Algebra/CommAlgebra/Instances/FreeCommAlgebra.agda create mode 100644 Cubical/Algebra/CommAlgebra/Instances/Initial.agda create mode 100644 Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda create mode 100644 Cubical/Algebra/CommAlgebra/Localisation.agda create mode 100644 Cubical/Algebra/CommAlgebra/Properties.agda create mode 100644 Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda create mode 100644 Cubical/Algebra/CommMonoid.agda create mode 100644 Cubical/Algebra/CommMonoid/Base.agda create mode 100644 Cubical/Algebra/CommMonoid/Properties.agda create mode 100644 Cubical/Algebra/CommRing.agda create mode 100644 Cubical/Algebra/CommRing/Base.agda create mode 100644 Cubical/Algebra/CommRing/BinomialThm.agda create mode 100644 Cubical/Algebra/CommRing/FGIdeal.agda create mode 100644 Cubical/Algebra/CommRing/FiberedProduct.agda create mode 100644 Cubical/Algebra/CommRing/Ideal.agda create mode 100644 Cubical/Algebra/CommRing/Instances/BiInvInt.agda create mode 100644 Cubical/Algebra/CommRing/Instances/Int.agda create mode 100644 Cubical/Algebra/CommRing/Instances/Pointwise.agda create mode 100644 Cubical/Algebra/CommRing/Instances/Unit.agda create mode 100644 Cubical/Algebra/CommRing/Localisation/Base.agda create mode 100644 Cubical/Algebra/CommRing/Localisation/InvertingElements.agda create mode 100644 Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda create mode 100644 Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda create mode 100644 Cubical/Algebra/CommRing/Properties.agda create mode 100644 Cubical/Algebra/CommRing/QuotientRing.agda create mode 100644 Cubical/Algebra/CommRing/RadicalIdeal.agda create mode 100644 Cubical/Algebra/DistLattice.agda create mode 100644 Cubical/Algebra/DistLattice/Base.agda create mode 100644 Cubical/Algebra/DistLattice/Basis.agda create mode 100644 Cubical/Algebra/DistLattice/BigOps.agda create mode 100644 Cubical/Algebra/Group.agda create mode 100644 Cubical/Algebra/Group/Base.agda create mode 100644 Cubical/Algebra/Group/DirProd.agda create mode 100644 Cubical/Algebra/Group/EilenbergMacLane1.agda create mode 100644 Cubical/Algebra/Group/Exact.agda create mode 100644 Cubical/Algebra/Group/GroupPath.agda create mode 100644 Cubical/Algebra/Group/Instances/Bool.agda create mode 100644 Cubical/Algebra/Group/Instances/DiffInt.agda create mode 100644 Cubical/Algebra/Group/Instances/Int.agda create mode 100644 Cubical/Algebra/Group/Instances/IntMod.agda create mode 100644 Cubical/Algebra/Group/Instances/Unit.agda create mode 100644 Cubical/Algebra/Group/IsomorphismTheorems.agda create mode 100644 Cubical/Algebra/Group/MorphismProperties.agda create mode 100644 Cubical/Algebra/Group/Morphisms.agda create mode 100644 Cubical/Algebra/Group/Properties.agda create mode 100644 Cubical/Algebra/Group/QuotientGroup.agda create mode 100644 Cubical/Algebra/Group/Subgroup.agda create mode 100644 Cubical/Algebra/Group/ZAction.agda create mode 100644 Cubical/Algebra/Lattice.agda create mode 100644 Cubical/Algebra/Lattice/Base.agda create mode 100644 Cubical/Algebra/Lattice/Properties.agda create mode 100644 Cubical/Algebra/Matrix.agda create mode 100644 Cubical/Algebra/Module.agda create mode 100644 Cubical/Algebra/Module/Base.agda create mode 100644 Cubical/Algebra/Monoid.agda create mode 100644 Cubical/Algebra/Monoid/Base.agda create mode 100644 Cubical/Algebra/Monoid/BigOp.agda create mode 100644 Cubical/Algebra/NatSolver/EvalHom.agda create mode 100644 Cubical/Algebra/NatSolver/Examples.agda create mode 100644 Cubical/Algebra/NatSolver/HornerForms.agda create mode 100644 Cubical/Algebra/NatSolver/NatExpression.agda create mode 100644 Cubical/Algebra/NatSolver/Reflection.agda create mode 100644 Cubical/Algebra/NatSolver/Solver.agda create mode 100644 Cubical/Algebra/Ring.agda create mode 100644 Cubical/Algebra/Ring/Base.agda create mode 100644 Cubical/Algebra/Ring/BigOps.agda create mode 100644 Cubical/Algebra/Ring/Ideal.agda create mode 100644 Cubical/Algebra/Ring/Kernel.agda create mode 100644 Cubical/Algebra/Ring/Properties.agda create mode 100644 Cubical/Algebra/Ring/QuotientRing.agda create mode 100644 Cubical/Algebra/RingSolver/AlgebraExpression.agda create mode 100644 Cubical/Algebra/RingSolver/EvalHom.agda create mode 100644 Cubical/Algebra/RingSolver/Examples.agda create mode 100644 Cubical/Algebra/RingSolver/HornerEval.agda create mode 100644 Cubical/Algebra/RingSolver/HornerForms.agda create mode 100644 Cubical/Algebra/RingSolver/IntAsRawRing.agda create mode 100644 Cubical/Algebra/RingSolver/README.md create mode 100644 Cubical/Algebra/RingSolver/RawAlgebra.agda create mode 100644 Cubical/Algebra/RingSolver/RawRing.agda create mode 100644 Cubical/Algebra/RingSolver/Reflection.agda create mode 100644 Cubical/Algebra/RingSolver/Solver.agda create mode 100644 Cubical/Algebra/RingSolver/Utility.agda create mode 100644 Cubical/Algebra/Semigroup.agda create mode 100644 Cubical/Algebra/Semigroup/Base.agda create mode 100644 Cubical/Algebra/Semilattice.agda create mode 100644 Cubical/Algebra/Semilattice/Base.agda create mode 100644 Cubical/Algebra/SymmetricGroup.agda create mode 100644 Cubical/Algebra/ZariskiLattice/Base.agda create mode 100644 Cubical/Categories/Adjoint.agda create mode 100644 Cubical/Categories/Category.agda create mode 100644 Cubical/Categories/Category/Base.agda create mode 100644 Cubical/Categories/Category/Precategory.agda create mode 100644 Cubical/Categories/Category/Properties.agda create mode 100644 Cubical/Categories/Commutativity.agda create mode 100644 Cubical/Categories/Constructions/BinProduct.agda create mode 100644 Cubical/Categories/Constructions/Elements.agda create mode 100644 Cubical/Categories/Constructions/Product.agda create mode 100644 Cubical/Categories/Constructions/Quotient.agda create mode 100644 Cubical/Categories/Constructions/Slice.agda create mode 100644 Cubical/Categories/DistLatticeSheaf.agda create mode 100644 Cubical/Categories/Equivalence.agda create mode 100644 Cubical/Categories/Equivalence/Base.agda create mode 100644 Cubical/Categories/Equivalence/Properties.agda create mode 100644 Cubical/Categories/Functor.agda create mode 100644 Cubical/Categories/Functor/Base.agda create mode 100644 Cubical/Categories/Functor/BinProduct.agda create mode 100644 Cubical/Categories/Functor/Compose.agda create mode 100644 Cubical/Categories/Functor/Product.agda create mode 100644 Cubical/Categories/Functor/Properties.agda create mode 100644 Cubical/Categories/Instances/AbGroups.agda create mode 100644 Cubical/Categories/Instances/Algebras.agda create mode 100644 Cubical/Categories/Instances/Categories.agda create mode 100644 Cubical/Categories/Instances/CommAlgebras.agda create mode 100644 Cubical/Categories/Instances/CommRings.agda create mode 100644 Cubical/Categories/Instances/Cospan.agda create mode 100644 Cubical/Categories/Instances/Discrete.agda create mode 100644 Cubical/Categories/Instances/DistLattice.agda create mode 100644 Cubical/Categories/Instances/Functors.agda create mode 100644 Cubical/Categories/Instances/Lattice.agda create mode 100644 Cubical/Categories/Instances/Poset.agda create mode 100644 Cubical/Categories/Instances/Rings.agda create mode 100644 Cubical/Categories/Instances/Semilattice.agda create mode 100644 Cubical/Categories/Instances/Sets.agda create mode 100644 Cubical/Categories/Instances/TypePrecategory.agda create mode 100644 Cubical/Categories/Limits.agda create mode 100644 Cubical/Categories/Limits/BinCoproduct.agda create mode 100644 Cubical/Categories/Limits/BinProduct.agda create mode 100644 Cubical/Categories/Limits/Initial.agda create mode 100644 Cubical/Categories/Limits/Limits.agda create mode 100644 Cubical/Categories/Limits/Pullback.agda create mode 100644 Cubical/Categories/Limits/Terminal.agda create mode 100644 Cubical/Categories/Monoidal.agda create mode 100644 Cubical/Categories/Monoidal/Base.agda create mode 100644 Cubical/Categories/Monoidal/Enriched.agda create mode 100644 Cubical/Categories/Morphism.agda create mode 100644 Cubical/Categories/NaturalTransformation.agda create mode 100644 Cubical/Categories/NaturalTransformation/Base.agda create mode 100644 Cubical/Categories/NaturalTransformation/Properties.agda create mode 100644 Cubical/Categories/Presheaf.agda create mode 100644 Cubical/Categories/Presheaf/Base.agda create mode 100644 Cubical/Categories/Presheaf/KanExtension.agda create mode 100644 Cubical/Categories/Presheaf/Properties.agda create mode 100644 Cubical/Categories/TypesOfCategories/TypeCategory.agda create mode 100644 Cubical/Categories/Yoneda.agda create mode 100644 Cubical/Codata/Conat.agda create mode 100644 Cubical/Codata/Conat/Base.agda create mode 100644 Cubical/Codata/Conat/Properties.agda create mode 100644 Cubical/Codata/Everything.agda create mode 100644 Cubical/Codata/EverythingSafe.agda create mode 100644 Cubical/Codata/M.agda create mode 100644 Cubical/Codata/M/AsLimit/Coalg.agda create mode 100644 Cubical/Codata/M/AsLimit/Coalg/Base.agda create mode 100644 Cubical/Codata/M/AsLimit/Container.agda create mode 100644 Cubical/Codata/M/AsLimit/M.agda create mode 100644 Cubical/Codata/M/AsLimit/M/Base.agda create mode 100644 Cubical/Codata/M/AsLimit/M/Properties.agda create mode 100644 Cubical/Codata/M/AsLimit/helper.agda create mode 100644 Cubical/Codata/M/AsLimit/itree.agda create mode 100644 Cubical/Codata/M/AsLimit/stream.agda create mode 100644 Cubical/Codata/M/Bisimilarity.agda create mode 100644 Cubical/Codata/Stream.agda create mode 100644 Cubical/Codata/Stream/Base.agda create mode 100644 Cubical/Codata/Stream/Properties.agda create mode 100644 Cubical/Core/Everything.agda create mode 100644 Cubical/Core/Glue.agda create mode 100644 Cubical/Core/Id.agda create mode 100644 Cubical/Core/Primitives.agda create mode 100644 Cubical/Core/README.md create mode 100644 Cubical/Data/BinNat.agda create mode 100644 Cubical/Data/BinNat/BinNat.agda create mode 100644 Cubical/Data/Bool.agda create mode 100644 Cubical/Data/Bool/Base.agda create mode 100644 Cubical/Data/Bool/Properties.agda create mode 100644 Cubical/Data/Bool/SwitchStatement.agda create mode 100644 Cubical/Data/DescendingList.agda create mode 100644 Cubical/Data/DescendingList/Base.agda create mode 100644 Cubical/Data/DescendingList/Examples.agda create mode 100644 Cubical/Data/DescendingList/Properties.agda create mode 100644 Cubical/Data/DescendingList/Strict.agda create mode 100644 Cubical/Data/DescendingList/Strict/Properties.agda create mode 100644 Cubical/Data/Empty.agda create mode 100644 Cubical/Data/Empty/Base.agda create mode 100644 Cubical/Data/Empty/Properties.agda create mode 100644 Cubical/Data/Equality.agda create mode 100644 Cubical/Data/Fin.agda create mode 100644 Cubical/Data/Fin/Arithmetic.agda create mode 100644 Cubical/Data/Fin/Base.agda create mode 100644 Cubical/Data/Fin/LehmerCode.agda create mode 100644 Cubical/Data/Fin/Literals.agda create mode 100644 Cubical/Data/Fin/Properties.agda create mode 100644 Cubical/Data/Fin/Recursive.agda create mode 100644 Cubical/Data/Fin/Recursive/Base.agda create mode 100644 Cubical/Data/Fin/Recursive/Properties.agda create mode 100644 Cubical/Data/FinData.agda create mode 100644 Cubical/Data/FinData/Base.agda create mode 100644 Cubical/Data/FinData/Properties.agda create mode 100644 Cubical/Data/FinInd.agda create mode 100644 Cubical/Data/FinSet.agda create mode 100644 Cubical/Data/FinSet/Base.agda create mode 100644 Cubical/Data/FinSet/Binary/Large.agda create mode 100644 Cubical/Data/FinSet/Binary/Small.agda create mode 100644 Cubical/Data/FinSet/Binary/Small/Base.agda create mode 100644 Cubical/Data/FinSet/Binary/Small/Properties.agda create mode 100644 Cubical/Data/FinSet/Constructors.agda create mode 100644 Cubical/Data/FinSet/FiniteChoice.agda create mode 100644 Cubical/Data/FinSet/Properties.agda create mode 100644 Cubical/Data/Graph.agda create mode 100644 Cubical/Data/Graph/Base.agda create mode 100644 Cubical/Data/Graph/Examples.agda create mode 100644 Cubical/Data/HomotopyGroup.agda create mode 100644 Cubical/Data/HomotopyGroup/Base.agda create mode 100644 Cubical/Data/InfNat.agda create mode 100644 Cubical/Data/InfNat/Base.agda create mode 100644 Cubical/Data/InfNat/Properties.agda create mode 100644 Cubical/Data/Int.agda create mode 100644 Cubical/Data/Int/Base.agda create mode 100644 Cubical/Data/Int/MoreInts/BiInvInt.agda create mode 100644 Cubical/Data/Int/MoreInts/BiInvInt/Base.agda create mode 100644 Cubical/Data/Int/MoreInts/BiInvInt/Properties.agda create mode 100644 Cubical/Data/Int/MoreInts/DeltaInt.agda create mode 100644 Cubical/Data/Int/MoreInts/DeltaInt/Base.agda create mode 100644 Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda create mode 100644 Cubical/Data/Int/MoreInts/DiffInt.agda create mode 100644 Cubical/Data/Int/MoreInts/DiffInt/Base.agda create mode 100644 Cubical/Data/Int/MoreInts/DiffInt/Properties.agda create mode 100644 Cubical/Data/Int/MoreInts/QuoInt.agda create mode 100644 Cubical/Data/Int/MoreInts/QuoInt/Base.agda create mode 100644 Cubical/Data/Int/MoreInts/QuoInt/Properties.agda create mode 100644 Cubical/Data/Int/Properties.agda create mode 100644 Cubical/Data/List.agda create mode 100644 Cubical/Data/List/Base.agda create mode 100644 Cubical/Data/List/Properties.agda create mode 100644 Cubical/Data/Maybe.agda create mode 100644 Cubical/Data/Maybe/Base.agda create mode 100644 Cubical/Data/Maybe/Properties.agda create mode 100644 Cubical/Data/Nat.agda create mode 100644 Cubical/Data/Nat/Algebra.agda create mode 100644 Cubical/Data/Nat/Base.agda create mode 100644 Cubical/Data/Nat/Coprime.agda create mode 100644 Cubical/Data/Nat/Divisibility.agda create mode 100644 Cubical/Data/Nat/GCD.agda create mode 100644 Cubical/Data/Nat/Literals.agda create mode 100644 Cubical/Data/Nat/Mod.agda create mode 100644 Cubical/Data/Nat/Order.agda create mode 100644 Cubical/Data/Nat/Order/Recursive.agda create mode 100644 Cubical/Data/Nat/Properties.agda create mode 100644 Cubical/Data/NatMinusOne.agda create mode 100644 Cubical/Data/NatMinusOne/Base.agda create mode 100644 Cubical/Data/NatMinusOne/Properties.agda create mode 100644 Cubical/Data/NatPlusOne.agda create mode 100644 Cubical/Data/NatPlusOne/Base.agda create mode 100644 Cubical/Data/NatPlusOne/Properties.agda create mode 100644 Cubical/Data/Prod.agda create mode 100644 Cubical/Data/Prod/Base.agda create mode 100644 Cubical/Data/Prod/Properties.agda create mode 100644 Cubical/Data/Queue.agda create mode 100644 Cubical/Data/Queue/1List.agda create mode 100644 Cubical/Data/Queue/Base.agda create mode 100644 Cubical/Data/Queue/Finite.agda create mode 100644 Cubical/Data/Queue/Truncated2List.agda create mode 100644 Cubical/Data/Queue/Untruncated2List.agda create mode 100644 Cubical/Data/Queue/Untruncated2ListInvariant.agda create mode 100644 Cubical/Data/Sigma.agda create mode 100644 Cubical/Data/Sigma/Base.agda create mode 100644 Cubical/Data/Sigma/Properties.agda create mode 100644 Cubical/Data/SubFinSet.agda create mode 100644 Cubical/Data/Sum.agda create mode 100644 Cubical/Data/Sum/Base.agda create mode 100644 Cubical/Data/Sum/Properties.agda create mode 100644 Cubical/Data/SumFin.agda create mode 100644 Cubical/Data/SumFin/Base.agda create mode 100644 Cubical/Data/SumFin/Properties.agda create mode 100644 Cubical/Data/Unit.agda create mode 100644 Cubical/Data/Unit/Base.agda create mode 100644 Cubical/Data/Unit/Pointed.agda create mode 100644 Cubical/Data/Unit/Properties.agda create mode 100644 Cubical/Data/Vec.agda create mode 100644 Cubical/Data/Vec/Base.agda create mode 100644 Cubical/Data/Vec/NAry.agda create mode 100644 Cubical/Data/Vec/Properties.agda create mode 100644 Cubical/Displayed/Auto.agda create mode 100644 Cubical/Displayed/Base.agda create mode 100644 Cubical/Displayed/Constant.agda create mode 100644 Cubical/Displayed/Function.agda create mode 100644 Cubical/Displayed/Generic.agda create mode 100644 Cubical/Displayed/Morphism.agda create mode 100644 Cubical/Displayed/Prop.agda create mode 100644 Cubical/Displayed/Properties.agda create mode 100644 Cubical/Displayed/Record.agda create mode 100644 Cubical/Displayed/Sigma.agda create mode 100644 Cubical/Displayed/Subst.agda create mode 100644 Cubical/Displayed/Unit.agda create mode 100644 Cubical/Displayed/Universe.agda create mode 100644 Cubical/Experiments/Brunerie.agda create mode 100644 Cubical/Experiments/CohomologyGroups.agda create mode 100644 Cubical/Experiments/Combinatorics.agda create mode 100644 Cubical/Experiments/EscardoSIP.agda create mode 100644 Cubical/Experiments/Everything.agda create mode 100644 Cubical/Experiments/FunExtFromUA.agda create mode 100644 Cubical/Experiments/Generic.agda create mode 100644 Cubical/Experiments/HAEquivInt.agda create mode 100644 Cubical/Experiments/HAEquivInt/Base.agda create mode 100644 Cubical/Experiments/HInt.agda create mode 100644 Cubical/Experiments/HoTT-UF.agda create mode 100644 Cubical/Experiments/IsoInt.agda create mode 100644 Cubical/Experiments/IsoInt/Base.agda create mode 100644 Cubical/Experiments/List.agda create mode 100644 Cubical/Experiments/LocalisationDefs.agda create mode 100644 Cubical/Experiments/NatMinusTwo.agda create mode 100644 Cubical/Experiments/NatMinusTwo/Base.agda create mode 100644 Cubical/Experiments/NatMinusTwo/Properties.agda create mode 100644 Cubical/Experiments/NatMinusTwo/ToNatMinusOne.agda create mode 100644 Cubical/Experiments/Poset.agda create mode 100644 Cubical/Experiments/Problem.agda create mode 100644 Cubical/Experiments/ZCohomology/Benchmarks.agda create mode 100644 Cubical/Experiments/ZCohomologyExperiments.agda create mode 100644 Cubical/Experiments/ZCohomologyOld/Base.agda create mode 100644 Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda create mode 100644 Cubical/Experiments/ZCohomologyOld/MayerVietorisUnreduced.agda create mode 100644 Cubical/Experiments/ZCohomologyOld/Properties.agda create mode 100644 Cubical/Experiments/ZariskiLatticeBasicOpens.agda create mode 100644 Cubical/Foundations/CartesianKanOps.agda create mode 100644 Cubical/Foundations/Equiv.agda create mode 100644 Cubical/Foundations/Equiv/Base.agda create mode 100644 Cubical/Foundations/Equiv/BiInvertible.agda create mode 100644 Cubical/Foundations/Equiv/Fiberwise.agda create mode 100644 Cubical/Foundations/Equiv/HalfAdjoint.agda create mode 100644 Cubical/Foundations/Equiv/PathSplit.agda create mode 100644 Cubical/Foundations/Equiv/Properties.agda create mode 100644 Cubical/Foundations/Everything.agda create mode 100644 Cubical/Foundations/Function.agda create mode 100644 Cubical/Foundations/GroupoidLaws.agda create mode 100644 Cubical/Foundations/HLevels.agda create mode 100644 Cubical/Foundations/Id.agda create mode 100644 Cubical/Foundations/Isomorphism.agda create mode 100644 Cubical/Foundations/Path.agda create mode 100644 Cubical/Foundations/Pointed.agda create mode 100644 Cubical/Foundations/Pointed/Base.agda create mode 100644 Cubical/Foundations/Pointed/FunExt.agda create mode 100644 Cubical/Foundations/Pointed/Homogeneous.agda create mode 100644 Cubical/Foundations/Pointed/Homotopy.agda create mode 100644 Cubical/Foundations/Pointed/Properties.agda create mode 100644 Cubical/Foundations/Powerset.agda create mode 100644 Cubical/Foundations/Prelude.agda create mode 100644 Cubical/Foundations/RelationalStructure.agda create mode 100644 Cubical/Foundations/SIP.agda create mode 100644 Cubical/Foundations/Structure.agda create mode 100644 Cubical/Foundations/Transport.agda create mode 100644 Cubical/Foundations/Univalence.agda create mode 100644 Cubical/Foundations/Univalence/Universe.agda create mode 100644 Cubical/Functions/Bundle.agda create mode 100644 Cubical/Functions/Embedding.agda create mode 100644 Cubical/Functions/Fibration.agda create mode 100644 Cubical/Functions/Fixpoint.agda create mode 100644 Cubical/Functions/FunExtEquiv.agda create mode 100644 Cubical/Functions/Implicit.agda create mode 100644 Cubical/Functions/Involution.agda create mode 100644 Cubical/Functions/Logic.agda create mode 100644 Cubical/Functions/Morphism.agda create mode 100644 Cubical/Functions/Surjection.agda create mode 100644 Cubical/HITs/2GroupoidTruncation.agda create mode 100644 Cubical/HITs/2GroupoidTruncation/Base.agda create mode 100644 Cubical/HITs/2GroupoidTruncation/Properties.agda create mode 100644 Cubical/HITs/AssocList.agda create mode 100644 Cubical/HITs/AssocList/Base.agda create mode 100644 Cubical/HITs/AssocList/Properties.agda create mode 100644 Cubical/HITs/Colimit.agda create mode 100644 Cubical/HITs/Colimit/Base.agda create mode 100644 Cubical/HITs/Colimit/Examples.agda create mode 100644 Cubical/HITs/Cost.agda create mode 100644 Cubical/HITs/Cost/Base.agda create mode 100644 Cubical/HITs/Cylinder.agda create mode 100644 Cubical/HITs/Cylinder/Base.agda create mode 100644 Cubical/HITs/Delooping/Two/Base.agda create mode 100644 Cubical/HITs/Delooping/Two/Properties.agda create mode 100644 Cubical/HITs/DunceCap.agda create mode 100644 Cubical/HITs/DunceCap/Base.agda create mode 100644 Cubical/HITs/DunceCap/Properties.agda create mode 100644 Cubical/HITs/EilenbergMacLane1.agda create mode 100644 Cubical/HITs/EilenbergMacLane1/Base.agda create mode 100644 Cubical/HITs/EilenbergMacLane1/Properties.agda create mode 100644 Cubical/HITs/FiniteMultiset.agda create mode 100644 Cubical/HITs/FiniteMultiset/Base.agda create mode 100644 Cubical/HITs/FiniteMultiset/CountExtensionality.agda create mode 100644 Cubical/HITs/FiniteMultiset/Properties.agda create mode 100644 Cubical/HITs/GroupoidQuotients.agda create mode 100644 Cubical/HITs/GroupoidQuotients/Base.agda create mode 100644 Cubical/HITs/GroupoidQuotients/Properties.agda create mode 100644 Cubical/HITs/GroupoidTruncation.agda create mode 100644 Cubical/HITs/GroupoidTruncation/Base.agda create mode 100644 Cubical/HITs/GroupoidTruncation/Properties.agda create mode 100644 Cubical/HITs/InfNat.agda create mode 100644 Cubical/HITs/InfNat/Base.agda create mode 100644 Cubical/HITs/InfNat/Properties.agda create mode 100644 Cubical/HITs/Interval.agda create mode 100644 Cubical/HITs/Interval/Base.agda create mode 100644 Cubical/HITs/Join.agda create mode 100644 Cubical/HITs/Join/Base.agda create mode 100644 Cubical/HITs/Join/Properties.agda create mode 100644 Cubical/HITs/KleinBottle.agda create mode 100644 Cubical/HITs/KleinBottle/Base.agda create mode 100644 Cubical/HITs/KleinBottle/Properties.agda create mode 100644 Cubical/HITs/ListedFiniteSet.agda create mode 100644 Cubical/HITs/ListedFiniteSet/Base.agda create mode 100644 Cubical/HITs/ListedFiniteSet/Properties.agda create mode 100644 Cubical/HITs/Localization.agda create mode 100644 Cubical/HITs/Localization/Base.agda create mode 100644 Cubical/HITs/Localization/Properties.agda create mode 100644 Cubical/HITs/MappingCones.agda create mode 100644 Cubical/HITs/MappingCones/Base.agda create mode 100644 Cubical/HITs/MappingCones/Properties.agda create mode 100644 Cubical/HITs/Modulo.agda create mode 100644 Cubical/HITs/Modulo/Base.agda create mode 100644 Cubical/HITs/Modulo/FinEquiv.agda create mode 100644 Cubical/HITs/Modulo/Properties.agda create mode 100644 Cubical/HITs/Nullification.agda create mode 100644 Cubical/HITs/Nullification/Base.agda create mode 100644 Cubical/HITs/Nullification/Properties.agda create mode 100644 Cubical/HITs/PropositionalTruncation.agda create mode 100644 Cubical/HITs/PropositionalTruncation/Base.agda create mode 100644 Cubical/HITs/PropositionalTruncation/MagicTrick.agda create mode 100644 Cubical/HITs/PropositionalTruncation/Monad.agda create mode 100644 Cubical/HITs/PropositionalTruncation/Properties.agda create mode 100644 Cubical/HITs/Pushout.agda create mode 100644 Cubical/HITs/Pushout/Base.agda create mode 100644 Cubical/HITs/Pushout/Flattening.agda create mode 100644 Cubical/HITs/Pushout/KrausVonRaumer.agda create mode 100644 Cubical/HITs/Pushout/Properties.agda create mode 100644 Cubical/HITs/RPn.agda create mode 100644 Cubical/HITs/RPn/Base.agda create mode 100644 Cubical/HITs/Rationals/HITQ.agda create mode 100644 Cubical/HITs/Rationals/HITQ/Base.agda create mode 100644 Cubical/HITs/Rationals/QuoQ.agda create mode 100644 Cubical/HITs/Rationals/QuoQ/Base.agda create mode 100644 Cubical/HITs/Rationals/QuoQ/Properties.agda create mode 100644 Cubical/HITs/Rationals/SigmaQ.agda create mode 100644 Cubical/HITs/Rationals/SigmaQ/Base.agda create mode 100644 Cubical/HITs/Rationals/SigmaQ/Properties.agda create mode 100644 Cubical/HITs/S1.agda create mode 100644 Cubical/HITs/S1/Base.agda create mode 100644 Cubical/HITs/S1/Properties.agda create mode 100644 Cubical/HITs/S2.agda create mode 100644 Cubical/HITs/S2/Base.agda create mode 100644 Cubical/HITs/S3.agda create mode 100644 Cubical/HITs/S3/Base.agda create mode 100644 Cubical/HITs/SetQuotients.agda create mode 100644 Cubical/HITs/SetQuotients/Base.agda create mode 100644 Cubical/HITs/SetQuotients/Properties.agda create mode 100644 Cubical/HITs/SetTruncation.agda create mode 100644 Cubical/HITs/SetTruncation/Base.agda create mode 100644 Cubical/HITs/SetTruncation/Properties.agda create mode 100644 Cubical/HITs/SmashProduct.agda create mode 100644 Cubical/HITs/SmashProduct/Base.agda create mode 100644 Cubical/HITs/Sn.agda create mode 100644 Cubical/HITs/Sn/Base.agda create mode 100644 Cubical/HITs/Sn/Properties.agda create mode 100644 Cubical/HITs/Susp.agda create mode 100644 Cubical/HITs/Susp/Base.agda create mode 100644 Cubical/HITs/Susp/LoopAdjunction.agda create mode 100644 Cubical/HITs/Susp/Properties.agda create mode 100644 Cubical/HITs/Torus.agda create mode 100644 Cubical/HITs/Torus/Base.agda create mode 100644 Cubical/HITs/Truncation.agda create mode 100644 Cubical/HITs/Truncation/Base.agda create mode 100644 Cubical/HITs/Truncation/FromNegTwo.agda create mode 100644 Cubical/HITs/Truncation/FromNegTwo/Base.agda create mode 100644 Cubical/HITs/Truncation/FromNegTwo/Properties.agda create mode 100644 Cubical/HITs/Truncation/Properties.agda create mode 100644 Cubical/HITs/TypeQuotients.agda create mode 100644 Cubical/HITs/TypeQuotients/Base.agda create mode 100644 Cubical/HITs/TypeQuotients/Properties.agda create mode 100644 Cubical/HITs/Wedge.agda create mode 100644 Cubical/HITs/Wedge/Base.agda create mode 100644 Cubical/Homotopy/Base.agda create mode 100644 Cubical/Homotopy/BlakersMassey.agda create mode 100644 Cubical/Homotopy/Connected.agda create mode 100644 Cubical/Homotopy/EilenbergSteenrod.agda create mode 100644 Cubical/Homotopy/Freudenthal.agda create mode 100644 Cubical/Homotopy/Group/Base.agda create mode 100644 Cubical/Homotopy/Group/LES.agda create mode 100644 Cubical/Homotopy/Group/Pi3S2.agda create mode 100644 Cubical/Homotopy/Group/Pi4S3/S3PushoutIso.agda create mode 100644 Cubical/Homotopy/Group/Pi4S3/S3PushoutIso2.agda create mode 100644 Cubical/Homotopy/Group/Pi4S3/Summary.agda create mode 100644 Cubical/Homotopy/Group/Pi4S3/Tricky.agda create mode 100644 Cubical/Homotopy/Group/PinSn.agda create mode 100644 Cubical/Homotopy/Group/SuspensionMap.agda create mode 100644 Cubical/Homotopy/HSpace.agda create mode 100644 Cubical/Homotopy/Hopf.agda create mode 100644 Cubical/Homotopy/HopfInvariant/Base.agda create mode 100644 Cubical/Homotopy/HopfInvariant/Homomorphism.agda create mode 100644 Cubical/Homotopy/HopfInvariant/HopfMap.agda create mode 100644 Cubical/Homotopy/HopfInvariant/Whitehead.agda create mode 100644 Cubical/Homotopy/Loopspace.agda create mode 100644 Cubical/Homotopy/MayerVietorisCofiber.agda create mode 100644 Cubical/Homotopy/Prespectrum.agda create mode 100644 Cubical/Homotopy/Spectrum.agda create mode 100644 Cubical/Homotopy/WedgeConnectivity.agda create mode 100644 Cubical/Homotopy/Whitehead.agda create mode 100644 Cubical/Induction/WellFounded.agda create mode 100644 Cubical/Modalities/Lex.agda create mode 100644 Cubical/Modalities/Modality.agda create mode 100644 Cubical/Papers/RepresentationIndependence.agda create mode 100644 Cubical/Papers/Synthetic.agda create mode 100644 Cubical/Papers/ZCohomology.agda create mode 100644 Cubical/README.agda create mode 100644 Cubical/Reflection/Base.agda create mode 100644 Cubical/Reflection/RecordEquiv.agda create mode 100644 Cubical/Reflection/StrictEquiv.agda create mode 100644 Cubical/Relation/Binary.agda create mode 100644 Cubical/Relation/Binary/Base.agda create mode 100644 Cubical/Relation/Binary/Poset.agda create mode 100644 Cubical/Relation/Binary/Properties.agda create mode 100644 Cubical/Relation/Nullary.agda create mode 100644 Cubical/Relation/Nullary/Base.agda create mode 100644 Cubical/Relation/Nullary/DecidableEq.agda create mode 100644 Cubical/Relation/Nullary/HLevels.agda create mode 100644 Cubical/Relation/Nullary/Properties.agda create mode 100644 Cubical/Relation/ZigZag/Applications/MultiSet.agda create mode 100644 Cubical/Relation/ZigZag/Base.agda create mode 100644 Cubical/Structures/Auto.agda create mode 100644 Cubical/Structures/Axioms.agda create mode 100644 Cubical/Structures/Constant.agda create mode 100644 Cubical/Structures/Function.agda create mode 100644 Cubical/Structures/LeftAction.agda create mode 100644 Cubical/Structures/Macro.agda create mode 100644 Cubical/Structures/Maybe.agda create mode 100644 Cubical/Structures/MultiSet.agda create mode 100644 Cubical/Structures/Parameterized.agda create mode 100644 Cubical/Structures/Pointed.agda create mode 100644 Cubical/Structures/Product.agda create mode 100644 Cubical/Structures/Queue.agda create mode 100644 Cubical/Structures/Record.agda create mode 100644 Cubical/Structures/Relational/Auto.agda create mode 100644 Cubical/Structures/Relational/Constant.agda create mode 100644 Cubical/Structures/Relational/Equalizer.agda create mode 100644 Cubical/Structures/Relational/Function.agda create mode 100644 Cubical/Structures/Relational/Macro.agda create mode 100644 Cubical/Structures/Relational/Maybe.agda create mode 100644 Cubical/Structures/Relational/Parameterized.agda create mode 100644 Cubical/Structures/Relational/Pointed.agda create mode 100644 Cubical/Structures/Relational/Product.agda create mode 100644 Cubical/Structures/Successor.agda create mode 100644 Cubical/Structures/Transfer.agda create mode 100644 Cubical/Structures/TypeEqvTo.agda create mode 100644 "Cubical/Syntax/\342\237\250\342\237\251.agda" create mode 100644 Cubical/Talks/EPA2020.agda create mode 100644 Cubical/WithK.agda create mode 100644 Cubical/ZCohomology/Base.agda create mode 100644 Cubical/ZCohomology/EilenbergSteenrodZ.agda create mode 100644 Cubical/ZCohomology/GroupStructure.agda create mode 100644 Cubical/ZCohomology/Groups/CP2.agda create mode 100644 Cubical/ZCohomology/Groups/Connected.agda create mode 100644 Cubical/ZCohomology/Groups/KleinBottle.agda create mode 100644 Cubical/ZCohomology/Groups/Prelims.agda create mode 100644 Cubical/ZCohomology/Groups/RP2.agda create mode 100644 Cubical/ZCohomology/Groups/Sn.agda create mode 100644 Cubical/ZCohomology/Groups/Torus.agda create mode 100644 Cubical/ZCohomology/Groups/Unit.agda create mode 100644 Cubical/ZCohomology/Groups/Wedge.agda create mode 100644 Cubical/ZCohomology/Groups/WedgeOfSpheres.agda create mode 100644 Cubical/ZCohomology/Gysin.agda create mode 100644 Cubical/ZCohomology/MayerVietorisUnreduced.agda create mode 100644 Cubical/ZCohomology/Properties.agda create mode 100644 Cubical/ZCohomology/RingStructure/CupProduct.agda create mode 100644 Cubical/ZCohomology/RingStructure/GradedCommutativity.agda create mode 100644 Cubical/ZCohomology/RingStructure/RingLaws.agda create mode 100644 Everythings.hs create mode 100644 GNUmakefile create mode 100644 INSTALL.md create mode 100644 LICENSE create mode 100644 MAKEWINDOWS.md create mode 100644 README.md create mode 100644 RELEASE.md create mode 100644 cubical.agda-lib create mode 100644 fix-whitespace.yaml diff --git a/.github/workflows/ci-ubuntu.yml b/.github/workflows/ci-ubuntu.yml new file mode 100644 index 0000000000..5c4230c14c --- /dev/null +++ b/.github/workflows/ci-ubuntu.yml @@ -0,0 +1,139 @@ +## Adapted from the agda/agda-stdlib workflow + +name: Ubuntu build +on: + push: + branches: + - master + pull_request: + branches: + - master + +######################################################################## +## CONFIGURATION +## +## See SETTINGS for the most important configuration variable: AGDA_COMMIT. +## It has to be defined as a build step because it is potentially branch +## dependent. +## +## As for the rest: +## +## Basically do not touch GHC_VERSION and CABAL_VERSION as long as +## they aren't a problem in the build. If you have time to waste, it +## could be worth investigating whether newer versions of ghc produce +## more efficient Agda executable and could cut down the build time. +## Just be aware that actions are flaky and small variations are to be +## expected. +## +## The CABAL_INSTALL variable only passes `-O1` optimisations to ghc +## because github actions cannot currently handle a build using `-O2`. +## To be experimented with again in the future to see if things have +## gotten better. +## +######################################################################## + +env: + GHC_VERSION: 8.6.5 + CABAL_VERSION: 3.2.0.0 + CABAL_INSTALL: cabal install --overwrite-policy=always --ghc-options='-O1 +RTS -M6G -RTS' + +jobs: + test-cubical: + runs-on: ubuntu-latest + steps: + + ######################################################################## + ## SETTINGS + ## + ## AGDA_COMMIT picks the version of Agda to use to build the library. + ## It can either be a hash of a specific commit (to target a bugfix for + ## instance) or a tag e.g. tags/v2.6.1.3 (to target a released version). + ######################################################################## + + - name: Initialise variables + run: | + # Pick Agda version + echo "AGDA_BRANCH=v2.6.2" >> $GITHUB_ENV + + ######################################################################## + ## CACHING + ######################################################################## + + # This caching step allows us to save a lot of building time by only + # downloading ghc and cabal and rebuilding Agda if absolutely necessary + # i.e. if we change either the version of Agda, ghc, or cabal that we want + # to use for the build. + - name: Cache cabal packages + uses: actions/cache@v2 + id: cache-cabal + with: + path: | + ~/.ghcup/bin + ~/.cabal/packages + ~/.cabal/store + ~/.cabal/bin + key: ${{ runner.os }}-${{ env.GHC_VERSION }}-${{ env.CABAL_VERSION }}-${{ env.AGDA_BRANCH }} + + ######################################################################## + ## INSTALLATION STEPS + ######################################################################## + + - name: Install cabal + if: steps.cache-cabal.outputs.cache-hit != 'true' + uses: actions/setup-haskell@v1.1.3 + with: + ghc-version: ${{ env.GHC_VERSION }} + cabal-version: ${{ env.CABAL_VERSION }} + + - name: Put cabal programs in PATH + run: echo "~/.cabal/bin" >> $GITHUB_PATH + + - name: Cabal update + run: cabal update + + - name: Download and install Agda from github + if: steps.cache-cabal.outputs.cache-hit != 'true' + run: | + git clone https://github.com/agda/agda --branch ${{ env.AGDA_BRANCH }} --depth=1 + cd agda + mkdir -p doc + touch doc/user-manual.pdf + ${{ env.CABAL_INSTALL }} + cd .. + + - name: Download and install fix-whitespace + if: steps.cache-cabal.outputs.cache-hit != 'true' + run: | + git clone https://github.com/agda/fix-whitespace --depth=1 + cd fix-whitespace + ${{ env.CABAL_INSTALL }} fix-whitespace.cabal + cd .. + + ######################################################################## + ## TESTING + ######################################################################## + + # By default github actions do not pull the repo + - name: Checkout cubical + uses: actions/checkout@v2 + + - name: Test cubical + run: | + make test \ + AGDA_EXEC='~/.cabal/bin/agda -W error' \ + FIX_WHITESPACE='~/.cabal/bin/fix-whitespace' \ + RUNHASKELL='~/.ghcup/bin/runhaskell' + + - name: Htmlize cubical + run: | + make listings \ + AGDA_EXEC='~/.cabal/bin/agda -W error' \ + FIX_WHITESPACE='~/.cabal/bin/fix-whitespace' \ + RUNHASKELL='~/.ghcup/bin/runhaskell' + + - name: Deploy to GitHub Pages + if: github.event_name == 'push' && github.ref_name == 'master' + uses: JamesIves/github-pages-deploy-action@4.1.8 + with: + branch: gh-pages + folder: html diff --git a/.gitignore b/.gitignore new file mode 100644 index 0000000000..19f6faa6d4 --- /dev/null +++ b/.gitignore @@ -0,0 +1,9 @@ +*.agdai +*~ +\.*#* +\#* +Cubical/*/Everything.agda +!Cubical/Core/Everything.agda +!Cubical/Foundations/Everything.agda +!Cubical/Codata/Everything.agda +!Cubical/Experiments/Everything.agda diff --git a/CONTRIBUTING.md b/CONTRIBUTING.md new file mode 100644 index 0000000000..e92036b479 --- /dev/null +++ b/CONTRIBUTING.md @@ -0,0 +1,112 @@ +CONTRIBUTING +============ + +We very much appreciate contributions in the form of pull requests (PRs). +When preparing a PR here are some general guidelines: + +- To test your changes before submission, run `make` at the top level, + which will generate all required `Everything` files in + `Cubical/README.agda` and then typecheck the latter file. + If you're using Windows, [here](MAKEWINDOWS.md) are some instructions + to get the `make` command working. + +- Please read through and clean your code before making a PR. Clean + code has reasonable line length (<100 characters), good indentation, + appropriate amounts of comments and consistent naming. + +- Local definitions should be put in `where`, `let-in` or `private` so + that it is easy to see which are the main results of a file and + which are just helper definitions. + +- If the code is based on some paper add a reference to the version of + the paper that the file follows in a comment at the top of the + file. Avoid calling top-level definitions `thm123`, `lem321`, etc., + instead have informative names and put pointers to the theorems and + lemmas in comments above the definition. + +- No global naming scheme for the library was enforced in the past, but new + contributions should only use names in camel-case. Names of types should + start with a capital letter and names of terms which are not types should + start with a lower case letter. However, for names of types that encode a + property (e.g. `isSet`), we prefer lowercase names. But this is not a + hard-and-fast rule. + +- Use `private variable` to quantify over universe levels at the top + of the file. All definitions should be maximally universe + polymorphic. + +- Make reasonably many arguments implicit. If you find yourself having + to provide some argument explicitly most of the time then it should + not be implicit. The same applies the other way around, if some argument + most often can be replaced by `_` then it should be made implicit. + +- Use `Type ℓ` for universes (so `Set ℓ` is not allowed in order to + avoid confusion with the type of h-sets). + +- All files should start with + + `{-# OPTIONS --safe #-}` + + unless there is a good reason for it not to. The `--cubical` and + `--no-import-sorts` flags are added in the `cubical.agda-lib` file. + +- It is much easier for us to review and merge smaller and + self-contained PRs. If a PR changes a lot of files all over the + library then they will conflict with other PRs making it harder for + us to merge them. It also takes longer to comment on the code + if the PR is very big. If you have a large PR please consider + splitting it into smaller parts that build on each other. + +- Large refactoring, renaming, etc., should be done in designated PRs + that only do one thing. + +- Draft PRs are very appreciated so that we can discuss the code as it + develops. It also helps with knowing who's working on what and + reducing duplicate efforts. If your code contains postulates then it + should be in a draft PR until the postulates have been filled in. + +- We automatically check for whitespace in the end of lines using + Travis. It is easy to configure emacs so that this is visible while + editing files by adding `(setq-default show-trailing-whitespace t)` + to `~/.emacs`. The command `M-x delete-trailing-whitespace` is also + very useful. It is possible to add a hook that runs this command + automatically when saving Agda files, by adding the following to your + `~/.emacs`: + ``` + ;; delete trailing whitespace before saving in agda-mode + (defun agda-mode-delete-whitespace-before-save () + (when (eq major-mode 'agda2-mode) + (delete-trailing-whitespace))) + + (add-hook 'before-save-hook #'agda-mode-delete-whitespace-before-save) + ``` + +- Use copattern-matching when instantiating records for efficiency. + This seems especially important when constructing Iso's. + +- If typechecking starts becoming slow try to fix the efficiency + problems directly. We don't want to merge files that are very slow + to typecheck so they will have to optimized at some point and it's + often much easier to fix these things directly. If you don't know + what to try make a draft PR and ask for help. + +- It is often useful to give explicit names to the Iso, Equiv and Path + version of a result. Try to avoid switching between these when + constructing something, for instance if you want to construct a Path + out of a series of Iso's then compose the Isos first and then apply + `isoToPath` once instead of converting all of them to paths and + composing them as paths. + +- Some useful lemmas have specialized versions with higher arity to + make code easier to read. See for example `isPropΠ2` and `isSetΠ2` + in [HLevels.agda](https://github.com/agda/cubical/blob/master/Cubical/Foundations/HLevels.agda) + as well as various versions of function extensionality in + [FunExtEquiv.agda](https://github.com/agda/cubical/blob/master/Cubical/Functions/FunExtEquiv.agda). + +- Unless a file is in the `Core`, `Foundations`, `Codata` or + `Experiments` package you don't need to add it manually to the + `Everything` file as it is automatically generated when running + `make`. + +- The `Experiments` folder contains various experiments and nothing in + the library should depend on something from this folder. diff --git a/Cubical/Algebra/AbGroup.agda b/Cubical/Algebra/AbGroup.agda new file mode 100644 index 0000000000..dd0ffe8410 --- /dev/null +++ b/Cubical/Algebra/AbGroup.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.AbGroup where + +open import Cubical.Algebra.AbGroup.Base public diff --git a/Cubical/Algebra/AbGroup/Base.agda b/Cubical/Algebra/AbGroup/Base.agda new file mode 100644 index 0000000000..84dc2c07b7 --- /dev/null +++ b/Cubical/Algebra/AbGroup/Base.agda @@ -0,0 +1,192 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.AbGroup.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Group + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open Iso + +private + variable + ℓ ℓ' : Level + +record IsAbGroup {G : Type ℓ} + (0g : G) (_+_ : G → G → G) (-_ : G → G) : Type ℓ where + + constructor isabgroup + + field + isGroup : IsGroup 0g _+_ -_ + comm : (x y : G) → x + y ≡ y + x + + open IsGroup isGroup public + +record AbGroupStr (A : Type ℓ) : Type (ℓ-suc ℓ) where + + constructor abgroupstr + + field + 0g : A + _+_ : A → A → A + -_ : A → A + isAbGroup : IsAbGroup 0g _+_ -_ + + infix 8 -_ + infixr 7 _+_ + + open IsAbGroup isAbGroup public + +AbGroup : ∀ ℓ → Type (ℓ-suc ℓ) +AbGroup ℓ = TypeWithStr ℓ AbGroupStr + +makeIsAbGroup : {G : Type ℓ} {0g : G} {_+_ : G → G → G} { -_ : G → G} + (is-setG : isSet G) + (assoc : (x y z : G) → x + (y + z) ≡ (x + y) + z) + (rid : (x : G) → x + 0g ≡ x) + (rinv : (x : G) → x + (- x) ≡ 0g) + (comm : (x y : G) → x + y ≡ y + x) + → IsAbGroup 0g _+_ -_ +makeIsAbGroup is-setG assoc rid rinv comm = + isabgroup (makeIsGroup is-setG assoc rid (λ x → comm _ _ ∙ rid x) rinv (λ x → comm _ _ ∙ rinv x)) comm + +makeAbGroup : {G : Type ℓ} (0g : G) (_+_ : G → G → G) (-_ : G → G) + (is-setG : isSet G) + (assoc : (x y z : G) → x + (y + z) ≡ (x + y) + z) + (rid : (x : G) → x + 0g ≡ x) + (rinv : (x : G) → x + (- x) ≡ 0g) + (comm : (x y : G) → x + y ≡ y + x) + → AbGroup ℓ +makeAbGroup 0g _+_ -_ is-setG assoc rid rinv comm = + _ , abgroupstr 0g _+_ -_ (makeIsAbGroup is-setG assoc rid rinv comm) + +open GroupStr +open AbGroupStr +open IsAbGroup + +AbGroupStr→GroupStr : {G : Type ℓ} → AbGroupStr G → GroupStr G +AbGroupStr→GroupStr A .1g = A .0g +AbGroupStr→GroupStr A ._·_ = A ._+_ +AbGroupStr→GroupStr A .inv = A .-_ +AbGroupStr→GroupStr A .isGroup = A .isAbGroup .isGroup + +AbGroup→Group : AbGroup ℓ → Group ℓ +fst (AbGroup→Group A) = fst A +snd (AbGroup→Group A) = AbGroupStr→GroupStr (snd A) + +Group→AbGroup : (G : Group ℓ) → ((x y : fst G) → _·_ (snd G) x y ≡ _·_ (snd G) y x) → AbGroup ℓ +fst (Group→AbGroup G comm) = fst G +AbGroupStr.0g (snd (Group→AbGroup G comm)) = 1g (snd G) +AbGroupStr._+_ (snd (Group→AbGroup G comm)) = _·_ (snd G) +AbGroupStr.- snd (Group→AbGroup G comm) = inv (snd G) +IsAbGroup.isGroup (AbGroupStr.isAbGroup (snd (Group→AbGroup G comm))) = isGroup (snd G) +IsAbGroup.comm (AbGroupStr.isAbGroup (snd (Group→AbGroup G comm))) = comm + +isSetAbGroup : (A : AbGroup ℓ) → isSet ⟨ A ⟩ +isSetAbGroup A = isSetGroup (AbGroup→Group A) + +AbGroupHom : (G : AbGroup ℓ) (H : AbGroup ℓ') → Type (ℓ-max ℓ ℓ') +AbGroupHom G H = GroupHom (AbGroup→Group G) (AbGroup→Group H) + +IsAbGroupEquiv : {A : Type ℓ} {B : Type ℓ'} + (G : AbGroupStr A) (e : A ≃ B) (H : AbGroupStr B) → Type (ℓ-max ℓ ℓ') +IsAbGroupEquiv G e H = IsGroupHom (AbGroupStr→GroupStr G) (e .fst) (AbGroupStr→GroupStr H) + +AbGroupEquiv : (G : AbGroup ℓ) (H : AbGroup ℓ') → Type (ℓ-max ℓ ℓ') +AbGroupEquiv G H = Σ[ e ∈ (G .fst ≃ H .fst) ] IsAbGroupEquiv (G .snd) e (H .snd) + +isPropIsAbGroup : {G : Type ℓ} (0g : G) (_+_ : G → G → G) (- : G → G) + → isProp (IsAbGroup 0g _+_ -) +isPropIsAbGroup 0g _+_ -_ (isabgroup GG GC) (isabgroup HG HC) = + λ i → isabgroup (isPropIsGroup _ _ _ GG HG i) (isPropComm GC HC i) + where + isSetG : isSet _ + isSetG = GG .IsGroup.isMonoid .IsMonoid.isSemigroup .IsSemigroup.is-set + + isPropComm : isProp ((x y : _) → x + y ≡ y + x) + isPropComm = isPropΠ2 λ _ _ → isSetG _ _ + +𝒮ᴰ-AbGroup : DUARel (𝒮-Univ ℓ) AbGroupStr ℓ +𝒮ᴰ-AbGroup = + 𝒮ᴰ-Record (𝒮-Univ _) IsAbGroupEquiv + (fields: + data[ _+_ ∣ autoDUARel _ _ ∣ pres· ] + data[ 0g ∣ autoDUARel _ _ ∣ pres1 ] + data[ -_ ∣ autoDUARel _ _ ∣ presinv ] + prop[ isAbGroup ∣ (λ _ _ → isPropIsAbGroup _ _ _) ]) + where + open AbGroupStr + open IsGroupHom + +-- Extract the characterization of equality of groups +AbGroupPath : (G H : AbGroup ℓ) → (AbGroupEquiv G H) ≃ (G ≡ H) +AbGroupPath = ∫ 𝒮ᴰ-AbGroup .UARel.ua + +-- TODO: Induced structure results are temporarily inconvenient while we transition between algebra +-- representations +module _ (G : AbGroup ℓ) {A : Type ℓ} (m : A → A → A) + (e : ⟨ G ⟩ ≃ A) + (p· : ∀ x y → e .fst (G .snd ._+_ x y) ≡ m (e .fst x) (e .fst y)) + where + + private + module G = AbGroupStr (G .snd) + + FamilyΣ : Σ[ B ∈ Type ℓ ] (B → B → B) → Type ℓ + FamilyΣ (B , n) = + Σ[ e ∈ B ] + Σ[ i ∈ (B → B) ] + IsAbGroup e n i + + inducedΣ : FamilyΣ (A , m) + inducedΣ = + subst FamilyΣ + (UARel.≅→≡ (autoUARel (Σ[ B ∈ Type ℓ ] (B → B → B))) (e , p·)) + (G.0g , G.-_ , G.isAbGroup) + + InducedAbGroup : AbGroup ℓ + InducedAbGroup .fst = A + InducedAbGroup .snd ._+_ = m + InducedAbGroup .snd .0g = inducedΣ .fst + InducedAbGroup .snd .-_ = inducedΣ .snd .fst + InducedAbGroup .snd .isAbGroup = inducedΣ .snd .snd + + InducedAbGroupPath : G ≡ InducedAbGroup + InducedAbGroupPath = AbGroupPath _ _ .fst (e , makeIsGroupHom p·) + +open IsMonoid +open IsSemigroup +open IsGroup + +dirProdAb : AbGroup ℓ → AbGroup ℓ' → AbGroup (ℓ-max ℓ ℓ') +dirProdAb A B = + Group→AbGroup (DirProd (AbGroup→Group A) (AbGroup→Group B)) + λ p q → ΣPathP (comm (isAbGroup (snd A)) _ _ + , comm (isAbGroup (snd B)) _ _) + +trivialAbGroup : ∀ {ℓ} → AbGroup ℓ +fst trivialAbGroup = Unit* +0g (snd trivialAbGroup) = tt* +_+_ (snd trivialAbGroup) _ _ = tt* +(- snd trivialAbGroup) _ = tt* +is-set (isSemigroup (isMonoid (isGroup (isAbGroup (snd trivialAbGroup))))) = + isProp→isSet isPropUnit* +assoc (isSemigroup (isMonoid (isGroup (isAbGroup (snd trivialAbGroup))))) _ _ _ = refl +identity (isMonoid (isGroup (isAbGroup (snd trivialAbGroup)))) _ = refl , refl +inverse (isGroup (isAbGroup (snd trivialAbGroup))) _ = refl , refl +comm (isAbGroup (snd trivialAbGroup)) _ _ = refl diff --git a/Cubical/Algebra/AbGroup/Instances/DiffInt.agda b/Cubical/Algebra/AbGroup/Instances/DiffInt.agda new file mode 100644 index 0000000000..60979c9121 --- /dev/null +++ b/Cubical/Algebra/AbGroup/Instances/DiffInt.agda @@ -0,0 +1,25 @@ +-- It is recommended to use Cubical.Algebra.CommRing.Instances.Int +-- instead of this file. + +{-# OPTIONS --safe #-} +module Cubical.Algebra.AbGroup.Instances.DiffInt where + +open import Cubical.Foundations.Prelude +open import Cubical.HITs.SetQuotients + +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Data.Int.MoreInts.DiffInt + renaming ( + _+_ to _+ℤ_ + ) + +DiffℤasAbGroup : AbGroup ℓ-zero +DiffℤasAbGroup = makeAbGroup {G = ℤ} + [ (0 , 0) ] + _+ℤ_ + -ℤ_ + ℤ-isSet + +ℤ-assoc + (λ x → zero-identityʳ 0 x) + (λ x → -ℤ-invʳ x) + +ℤ-comm diff --git a/Cubical/Algebra/AbGroup/Instances/Hom.agda b/Cubical/Algebra/AbGroup/Instances/Hom.agda new file mode 100644 index 0000000000..fb8a343eb0 --- /dev/null +++ b/Cubical/Algebra/AbGroup/Instances/Hom.agda @@ -0,0 +1,145 @@ +-- Given two abelian groups A, B +-- the set of all group homomorphisms from A to B +-- is itself an abelian group. +-- In other words, Ab is cartesian closed. +-- This is needed to show Ab is an abelian category. + +{-# OPTIONS --safe #-} + +module Cubical.Algebra.AbGroup.Instances.Hom where + +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Properties +open import Cubical.Foundations.Prelude + +private + variable + ℓ ℓ' : Level + +module _ (A : AbGroup ℓ) (B : AbGroup ℓ') where + + -- These names are useful for the proofs + private + open IsGroupHom + open AbGroupStr (A .snd) using () renaming (0g to 0A; _+_ to _⋆_; -_ to inv) + open AbGroupStr (B .snd) using (_+_; -_; comm; assoc; identity; inverse) + renaming (0g to 0B) + open GroupTheory (AbGroup→Group B) using (invDistr) renaming (inv1g to inv0B) + + -- Some lemmas + idrB : (b : B .fst) → b + 0B ≡ b + idrB b = identity b .fst + + invrB : (b : B .fst) → b + (- b) ≡ 0B + invrB b = inverse b .fst + + hom0AB : (f : AbGroupHom A B) → f .fst 0A ≡ 0B + hom0AB f = hom1g (AbGroupStr→GroupStr (A .snd)) (f .fst) + (AbGroupStr→GroupStr (B .snd)) (f .snd .pres·) + + homInvAB : (f : AbGroupHom A B) → (a : A .fst) → f .fst (inv a) ≡ (- f .fst a) + homInvAB f a = homInv (AbGroupStr→GroupStr (A .snd)) (f .fst) + (AbGroupStr→GroupStr (B .snd)) (f .snd .pres·) a + + + -- Zero morphism + zero : AbGroupHom A B + zero .fst a = 0B + zero .snd .pres· a a' = sym (idrB _) + zero .snd .pres1 = refl + zero .snd .presinv a = sym (inv0B) + + + -- Pointwise addition of morphisms + module _ (f* g* : AbGroupHom A B) where + private + f = f* .fst + g = g* .fst + + HomAdd : AbGroupHom A B + HomAdd .fst = λ a → f a + g a + + HomAdd .snd .pres· a a' = + f (a ⋆ a') + g (a ⋆ a') ≡⟨ cong (_+ g(a ⋆ a')) (f* .snd .pres· _ _) ⟩ + (f a + f a') + g (a ⋆ a') ≡⟨ cong ((f a + f a') +_) (g* .snd .pres· _ _) ⟩ + (f a + f a') + (g a + g a') ≡⟨ sym (assoc _ _ _) ⟩ + f a + (f a' + (g a + g a')) ≡⟨ cong (f a +_) (assoc _ _ _) ⟩ + f a + ((f a' + g a) + g a') ≡⟨ cong (λ b → (f a + b + g a')) (comm _ _) ⟩ + f a + ((g a + f a') + g a') ≡⟨ cong (f a +_) (sym (assoc _ _ _)) ⟩ + f a + (g a + (f a' + g a')) ≡⟨ assoc _ _ _ ⟩ + (f a + g a) + (f a' + g a') ∎ + + HomAdd .snd .pres1 = + f 0A + g 0A ≡⟨ cong (_+ g 0A) (hom0AB f*) ⟩ + 0B + g 0A ≡⟨ cong (0B +_) (hom0AB g*) ⟩ + 0B + 0B ≡⟨ idrB _ ⟩ + 0B ∎ + + HomAdd .snd .presinv a = + f (inv a) + g (inv a) ≡⟨ cong (_+ g (inv a)) (homInvAB f* _) ⟩ + (- f a) + g (inv a) ≡⟨ cong ((- f a) +_) (homInvAB g* _) ⟩ + (- f a) + (- g a) ≡⟨ comm _ _ ⟩ + (- g a) + (- f a) ≡⟨ sym (invDistr _ _) ⟩ + - (f a + g a) ∎ + + + -- Pointwise inverse of morphism + module _ (f* : AbGroupHom A B) where + private + f = f* .fst + + HomInv : AbGroupHom A B + HomInv .fst = λ a → - f a + + HomInv .snd .pres· a a' = + - f (a ⋆ a') ≡⟨ cong -_ (f* .snd .pres· _ _) ⟩ + - (f a + f a') ≡⟨ invDistr _ _ ⟩ + (- f a') + (- f a) ≡⟨ comm _ _ ⟩ + (- f a) + (- f a') ∎ + + HomInv .snd .pres1 = + - (f 0A) ≡⟨ cong -_ (f* .snd .pres1) ⟩ + - 0B ≡⟨ inv0B ⟩ + 0B ∎ + + HomInv .snd .presinv a = + - f (inv a) ≡⟨ cong -_ (homInvAB f* _) ⟩ + - (- f a) ∎ + + + -- Group laws for morphisms + private + 0ₕ = zero + _+ₕ_ = HomAdd + -ₕ_ = HomInv + + -- Morphism addition is associative + HomAdd-assoc : (f g h : AbGroupHom A B) → (f +ₕ (g +ₕ h)) ≡ ((f +ₕ g) +ₕ h) + HomAdd-assoc f g h = GroupHom≡ (funExt λ a → assoc _ _ _) + + -- Morphism addition is commutative + HomAdd-comm : (f g : AbGroupHom A B) → (f +ₕ g) ≡ (g +ₕ f) + HomAdd-comm f g = GroupHom≡ (funExt λ a → comm _ _) + + -- zero is right identity + HomAdd-zero : (f : AbGroupHom A B) → (f +ₕ zero) ≡ f + HomAdd-zero f = GroupHom≡ (funExt λ a → idrB _) + + -- -ₕ is right inverse + HomInv-invr : (f : AbGroupHom A B) → (f +ₕ (-ₕ f)) ≡ zero + HomInv-invr f = GroupHom≡ (funExt λ a → invrB _) + + +-- Abelian group structure on AbGroupHom A B +open AbGroupStr +HomAbGroupStr : (A : AbGroup ℓ) → (B : AbGroup ℓ') → AbGroupStr (AbGroupHom A B) +HomAbGroupStr A B .0g = zero A B +HomAbGroupStr A B ._+_ = HomAdd A B +HomAbGroupStr A B .-_ = HomInv A B +HomAbGroupStr A B .isAbGroup = makeIsAbGroup isSetGroupHom + (HomAdd-assoc A B) (HomAdd-zero A B) (HomInv-invr A B) (HomAdd-comm A B) + +HomAbGroup : (A : AbGroup ℓ) → (B : AbGroup ℓ') → AbGroup (ℓ-max ℓ ℓ') +HomAbGroup A B = AbGroupHom A B , HomAbGroupStr A B diff --git a/Cubical/Algebra/Algebra.agda b/Cubical/Algebra/Algebra.agda new file mode 100644 index 0000000000..acdb869da5 --- /dev/null +++ b/Cubical/Algebra/Algebra.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Algebra where + +open import Cubical.Algebra.Algebra.Base public +open import Cubical.Algebra.Algebra.Properties public diff --git a/Cubical/Algebra/Algebra/Base.agda b/Cubical/Algebra/Algebra/Base.agda new file mode 100644 index 0000000000..54e3c80ec5 --- /dev/null +++ b/Cubical/Algebra/Algebra/Base.agda @@ -0,0 +1,269 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Algebra.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Reflection.RecordEquiv + +open import Cubical.Algebra.Module +open import Cubical.Algebra.Ring +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group +open import Cubical.Algebra.Monoid + +open Iso + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + +record IsAlgebra (R : Ring ℓ) {A : Type ℓ'} + (0a 1a : A) (_+_ _·_ : A → A → A) (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where + + constructor isalgebra + + open RingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·r_) + + field + isLeftModule : IsLeftModule R 0a _+_ -_ _⋆_ + ·-isMonoid : IsMonoid 1a _·_ + dist : (x y z : A) → (x · (y + z) ≡ (x · y) + (x · z)) + × ((x + y) · z ≡ (x · z) + (y · z)) + ⋆-lassoc : (r : ⟨ R ⟩) (x y : A) → (r ⋆ x) · y ≡ r ⋆ (x · y) + ⋆-rassoc : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x · y) ≡ x · (r ⋆ y) + + open IsLeftModule isLeftModule public + + isRing : IsRing _ _ _ _ _ + isRing = isring (IsLeftModule.+-isAbGroup isLeftModule) ·-isMonoid dist + open IsRing isRing public hiding (_-_; +Assoc; +Lid; +Linv; +Rid; +Rinv; +Comm) + +unquoteDecl IsAlgebraIsoΣ = declareRecordIsoΣ IsAlgebraIsoΣ (quote IsAlgebra) + +record AlgebraStr (R : Ring ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where + + constructor algebrastr + + field + 0a : A + 1a : A + _+_ : A → A → A + _·_ : A → A → A + -_ : A → A + _⋆_ : ⟨ R ⟩ → A → A + isAlgebra : IsAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + + open IsAlgebra isAlgebra public + +Algebra : (R : Ring ℓ) → ∀ ℓ' → Type (ℓ-max ℓ (ℓ-suc ℓ')) +Algebra R ℓ' = Σ[ A ∈ Type ℓ' ] AlgebraStr R A + +module commonExtractors {R : Ring ℓ} where + + Algebra→Module : (A : Algebra R ℓ') → LeftModule R ℓ' + Algebra→Module (_ , algebrastr A _ _ _ _ _ (isalgebra isLeftModule _ _ _ _)) = + _ , leftmodulestr A _ _ _ isLeftModule + + Algebra→Ring : (A : Algebra R ℓ') → Ring ℓ' + Algebra→Ring (_ , str) = _ , ringstr _ _ _ _ _ (IsAlgebra.isRing (AlgebraStr.isAlgebra str)) + + Algebra→AbGroup : (A : Algebra R ℓ') → AbGroup ℓ' + Algebra→AbGroup A = LeftModule→AbGroup (Algebra→Module A) + + Algebra→Group : (A : Algebra R ℓ') → Group ℓ' + Algebra→Group A = Ring→Group (Algebra→Ring A) + + Algebra→AddMonoid : (A : Algebra R ℓ') → Monoid ℓ' + Algebra→AddMonoid A = Group→Monoid (Algebra→Group A) + + Algebra→MultMonoid : (A : Algebra R ℓ') → Monoid ℓ' + Algebra→MultMonoid A = Ring→MultMonoid (Algebra→Ring A) + + isSetAlgebra : (A : Algebra R ℓ') → isSet ⟨ A ⟩ + isSetAlgebra A = isSetAbGroup (Algebra→AbGroup A) + + open RingStr (snd R) using (1r; ·Ldist+) renaming (_+_ to _+r_; _·_ to _·s_) + + makeIsAlgebra : {A : Type ℓ'} {0a 1a : A} + {_+_ _·_ : A → A → A} { -_ : A → A} {_⋆_ : ⟨ R ⟩ → A → A} + (isSet-A : isSet A) + (+-assoc : (x y z : A) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : A) → x + 0a ≡ x) + (+-rinv : (x : A) → x + (- x) ≡ 0a) + (+-comm : (x y : A) → x + y ≡ y + x) + (·-assoc : (x y z : A) → x · (y · z) ≡ (x · y) · z) + (·-rid : (x : A) → x · 1a ≡ x) + (·-lid : (x : A) → 1a · x ≡ x) + (·-rdist-+ : (x y z : A) → x · (y + z) ≡ (x · y) + (x · z)) + (·-ldist-+ : (x y z : A) → (x + y) · z ≡ (x · z) + (y · z)) + (⋆-assoc : (r s : ⟨ R ⟩) (x : A) → (r ·s s) ⋆ x ≡ r ⋆ (s ⋆ x)) + (⋆-ldist : (r s : ⟨ R ⟩) (x : A) → (r +r s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x)) + (⋆-rdist : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) + (⋆-lid : (x : A) → 1r ⋆ x ≡ x) + (⋆-lassoc : (r : ⟨ R ⟩) (x y : A) → (r ⋆ x) · y ≡ r ⋆ (x · y)) + (⋆-rassoc : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x · y) ≡ x · (r ⋆ y)) + → IsAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + makeIsAlgebra isSet-A + +-assoc +-rid +-rinv +-comm + ·-assoc ·-rid ·-lid ·-rdist-+ ·-ldist-+ + ⋆-assoc ⋆-ldist ⋆-rdist ⋆-lid ⋆-lassoc ⋆-rassoc = + isalgebra + (makeIsLeftModule isSet-A + +-assoc +-rid +-rinv +-comm + ⋆-assoc ⋆-ldist ⋆-rdist ⋆-lid) + (makeIsMonoid isSet-A ·-assoc ·-rid ·-lid) + (λ x y z → ·-rdist-+ x y z , ·-ldist-+ x y z) + ⋆-lassoc ⋆-rassoc + + +open commonExtractors public + +record IsAlgebraHom {R : Ring ℓ} {A : Type ℓ'} {B : Type ℓ''} + (M : AlgebraStr R A) (f : A → B) (N : AlgebraStr R B) + : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) + where + + -- Shorter qualified names + private + module M = AlgebraStr M + module N = AlgebraStr N + + field + pres0 : f M.0a ≡ N.0a + pres1 : f M.1a ≡ N.1a + pres+ : (x y : A) → f (x M.+ y) ≡ f x N.+ f y + pres· : (x y : A) → f (x M.· y) ≡ f x N.· f y + pres- : (x : A) → f (M.- x) ≡ N.- (f x) + pres⋆ : (r : ⟨ R ⟩) (y : A) → f (r M.⋆ y) ≡ r N.⋆ f y + +unquoteDecl IsAlgebraHomIsoΣ = declareRecordIsoΣ IsAlgebraHomIsoΣ (quote IsAlgebraHom) +open IsAlgebraHom + +AlgebraHom : {R : Ring ℓ} (M : Algebra R ℓ') (N : Algebra R ℓ'') → Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) +AlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ → ⟨ N ⟩) ] IsAlgebraHom (M .snd) f (N .snd) + +IsAlgebraEquiv : {R : Ring ℓ} {A : Type ℓ'} {B : Type ℓ''} + (M : AlgebraStr R A) (e : A ≃ B) (N : AlgebraStr R B) + → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') +IsAlgebraEquiv M e N = IsAlgebraHom M (e .fst) N + +AlgebraEquiv : {R : Ring ℓ} (M : Algebra R ℓ') (N : Algebra R ℓ'') → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') +AlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsAlgebraEquiv (M .snd) e (N .snd) + +_$a_ : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} → AlgebraHom A B → ⟨ A ⟩ → ⟨ B ⟩ +f $a x = fst f x + +AlgebraEquiv→AlgebraHom : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} + → AlgebraEquiv A B → AlgebraHom A B +AlgebraEquiv→AlgebraHom (e , eIsHom) = e .fst , eIsHom + +isPropIsAlgebra : (R : Ring ℓ) {A : Type ℓ'} + (0a 1a : A) + (_+_ _·_ : A → A → A) + (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) + → isProp (IsAlgebra R 0a 1a _+_ _·_ -_ _⋆_) +isPropIsAlgebra R _ _ _ _ _ _ = let open IsLeftModule in + isOfHLevelRetractFromIso 1 IsAlgebraIsoΣ + (isPropΣ + (isPropIsLeftModule _ _ _ _ _) + (λ mo → isProp×3 (isPropIsMonoid _ _) + (isPropΠ3 λ _ _ _ → isProp× (mo .is-set _ _) (mo .is-set _ _)) + (isPropΠ3 λ _ _ _ → mo .is-set _ _) + (isPropΠ3 λ _ _ _ → mo .is-set _ _) )) + + +isPropIsAlgebraHom : (R : Ring ℓ) {A : Type ℓ'} {B : Type ℓ''} + (AS : AlgebraStr R A) (f : A → B) (BS : AlgebraStr R B) + → isProp (IsAlgebraHom AS f BS) +isPropIsAlgebraHom R AS f BS = isOfHLevelRetractFromIso 1 IsAlgebraHomIsoΣ + (isProp×5 (isSetAlgebra (_ , BS) _ _) + (isSetAlgebra (_ , BS) _ _) + (isPropΠ2 λ _ _ → isSetAlgebra (_ , BS) _ _) + (isPropΠ2 λ _ _ → isSetAlgebra (_ , BS) _ _) + (isPropΠ λ _ → isSetAlgebra (_ , BS) _ _) + (isPropΠ2 λ _ _ → isSetAlgebra (_ , BS) _ _)) + +isSetAlgebraHom : {R : Ring ℓ} (M : Algebra R ℓ') (N : Algebra R ℓ'') + → isSet (AlgebraHom M N) +isSetAlgebraHom _ N = isSetΣ (isSetΠ (λ _ → isSetAlgebra N)) + λ _ → isProp→isSet (isPropIsAlgebraHom _ _ _ _) + + +isSetAlgebraEquiv : {R : Ring ℓ} (M N : Algebra R ℓ') + → isSet (AlgebraEquiv M N) +isSetAlgebraEquiv M N = isSetΣ (isOfHLevel≃ 2 (isSetAlgebra M) (isSetAlgebra N)) + λ _ → isProp→isSet (isPropIsAlgebraHom _ _ _ _) + +AlgebraHom≡ : {R : Ring ℓ} {A B : Algebra R ℓ'} {φ ψ : AlgebraHom A B} → fst φ ≡ fst ψ → φ ≡ ψ +AlgebraHom≡ = Σ≡Prop λ f → isPropIsAlgebraHom _ _ f _ + +𝒮ᴰ-Algebra : (R : Ring ℓ) → DUARel (𝒮-Univ ℓ') (AlgebraStr R) (ℓ-max ℓ ℓ') +𝒮ᴰ-Algebra R = + 𝒮ᴰ-Record (𝒮-Univ _) (IsAlgebraEquiv {R = R}) + (fields: + data[ 0a ∣ nul ∣ pres0 ] + data[ 1a ∣ nul ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ autoDUARel _ _ ∣ pres- ] + data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] + prop[ isAlgebra ∣ (λ _ _ → isPropIsAlgebra _ _ _ _ _ _ _) ]) + where + open AlgebraStr + + -- faster with some sharing + nul = autoDUARel (𝒮-Univ _) (λ A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +AlgebraPath : {R : Ring ℓ} (A B : Algebra R ℓ') → (AlgebraEquiv A B) ≃ (A ≡ B) +AlgebraPath {R = R} = ∫ (𝒮ᴰ-Algebra R) .UARel.ua + +uaAlgebra : {R : Ring ℓ} {A B : Algebra R ℓ'} → AlgebraEquiv A B → A ≡ B +uaAlgebra {A = A} {B = B} = equivFun (AlgebraPath A B) + +isGroupoidAlgebra : {R : Ring ℓ} → isGroupoid (Algebra R ℓ') +isGroupoidAlgebra _ _ = isOfHLevelRespectEquiv 2 (AlgebraPath _ _) (isSetAlgebraEquiv _ _) + +-- Smart constructor for ring homomorphisms +-- that infers the other equations from pres1, pres+, and pres· + +module _ {R : Ring ℓ} {A : Algebra R ℓ} {B : Algebra R ℓ'} {f : ⟨ A ⟩ → ⟨ B ⟩} where + + private + module A = AlgebraStr (A .snd) + module B = AlgebraStr (B .snd) + + module _ + (p1 : f A.1a ≡ B.1a) + (p+ : (x y : ⟨ A ⟩) → f (x A.+ y) ≡ f x B.+ f y) + (p· : (x y : ⟨ A ⟩) → f (x A.· y) ≡ f x B.· f y) + (p⋆ : (r : ⟨ R ⟩) (x : ⟨ A ⟩) → f (r A.⋆ x) ≡ r B.⋆ f x) + where + + open IsAlgebraHom + private + isGHom : IsGroupHom (Algebra→Group A .snd) f (Algebra→Group B .snd) + isGHom = makeIsGroupHom p+ + + makeIsAlgebraHom : IsAlgebraHom (A .snd) f (B .snd) + makeIsAlgebraHom .pres0 = isGHom .IsGroupHom.pres1 + makeIsAlgebraHom .pres1 = p1 + makeIsAlgebraHom .pres+ = p+ + makeIsAlgebraHom .pres· = p· + makeIsAlgebraHom .pres- = isGHom .IsGroupHom.presinv + makeIsAlgebraHom .pres⋆ = p⋆ diff --git a/Cubical/Algebra/Algebra/Properties.agda b/Cubical/Algebra/Algebra/Properties.agda new file mode 100644 index 0000000000..7cac6d54bf --- /dev/null +++ b/Cubical/Algebra/Algebra/Properties.agda @@ -0,0 +1,170 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Algebra.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Sigma + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Auto +open import Cubical.Structures.Macro +open import Cubical.Algebra.Module +open import Cubical.Algebra.Ring +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Algebra.Base + +open Iso + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + + + + +module AlgebraTheory (R : Ring ℓ) (A : Algebra R ℓ') where + open RingStr (snd R) renaming (_+_ to _+r_ ; _·_ to _·r_) + open AlgebraStr (A .snd) + + 0-actsNullifying : (x : ⟨ A ⟩) → 0r ⋆ x ≡ 0a + 0-actsNullifying x = + let idempotent-+ = 0r ⋆ x ≡⟨ cong (λ u → u ⋆ x) (sym (RingTheory.0Idempotent R)) ⟩ + (0r +r 0r) ⋆ x ≡⟨ ⋆-ldist 0r 0r x ⟩ + (0r ⋆ x) + (0r ⋆ x) ∎ + in RingTheory.+Idempotency→0 (Algebra→Ring A) (0r ⋆ x) idempotent-+ + + ⋆Dist· : (x y : ⟨ R ⟩) (a b : ⟨ A ⟩) → (x ·r y) ⋆ (a · b) ≡ (x ⋆ a) · (y ⋆ b) + ⋆Dist· x y a b = (x ·r y) ⋆ (a · b) ≡⟨ ⋆-rassoc _ _ _ ⟩ + a · ((x ·r y) ⋆ b) ≡⟨ cong (a ·_) (⋆-assoc _ _ _) ⟩ + a · (x ⋆ (y ⋆ b)) ≡⟨ sym (⋆-rassoc _ _ _) ⟩ + x ⋆ (a · (y ⋆ b)) ≡⟨ sym (⋆-lassoc _ _ _) ⟩ + (x ⋆ a) · (y ⋆ b) ∎ + + +module AlgebraHoms {R : Ring ℓ} where + open IsAlgebraHom + + idAlgebraHom : (A : Algebra R ℓ') → AlgebraHom A A + fst (idAlgebraHom A) x = x + pres0 (snd (idAlgebraHom A)) = refl + pres1 (snd (idAlgebraHom A)) = refl + pres+ (snd (idAlgebraHom A)) x y = refl + pres· (snd (idAlgebraHom A)) x y = refl + pres- (snd (idAlgebraHom A)) x = refl + pres⋆ (snd (idAlgebraHom A)) r x = refl + + compIsAlgebraHom : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + {g : ⟨ B ⟩ → ⟨ C ⟩} {f : ⟨ A ⟩ → ⟨ B ⟩} + → IsAlgebraHom (B .snd) g (C .snd) + → IsAlgebraHom (A .snd) f (B .snd) + → IsAlgebraHom (A .snd) (g ∘ f) (C .snd) + compIsAlgebraHom {g = g} {f} gh fh .pres0 = cong g (fh .pres0) ∙ gh .pres0 + compIsAlgebraHom {g = g} {f} gh fh .pres1 = cong g (fh .pres1) ∙ gh .pres1 + compIsAlgebraHom {g = g} {f} gh fh .pres+ x y = cong g (fh .pres+ x y) ∙ gh .pres+ (f x) (f y) + compIsAlgebraHom {g = g} {f} gh fh .pres· x y = cong g (fh .pres· x y) ∙ gh .pres· (f x) (f y) + compIsAlgebraHom {g = g} {f} gh fh .pres- x = cong g (fh .pres- x) ∙ gh .pres- (f x) + compIsAlgebraHom {g = g} {f} gh fh .pres⋆ r x = cong g (fh .pres⋆ r x) ∙ gh .pres⋆ r (f x) + + compAlgebraHom : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + → AlgebraHom A B → AlgebraHom B C → AlgebraHom A C + compAlgebraHom f g .fst = g .fst ∘ f .fst + compAlgebraHom f g .snd = compIsAlgebraHom (g .snd) (f .snd) + + syntax compAlgebraHom f g = g ∘a f + + compIdAlgebraHom : {A B : Algebra R ℓ'} (φ : AlgebraHom A B) → compAlgebraHom (idAlgebraHom A) φ ≡ φ + compIdAlgebraHom φ = AlgebraHom≡ refl + + idCompAlgebraHom : {A B : Algebra R ℓ'} (φ : AlgebraHom A B) → compAlgebraHom φ (idAlgebraHom B) ≡ φ + idCompAlgebraHom φ = AlgebraHom≡ refl + + compAssocAlgebraHom : {A B C D : Algebra R ℓ'} + (φ : AlgebraHom A B) (ψ : AlgebraHom B C) (χ : AlgebraHom C D) + → compAlgebraHom (compAlgebraHom φ ψ) χ ≡ compAlgebraHom φ (compAlgebraHom ψ χ) + compAssocAlgebraHom _ _ _ = AlgebraHom≡ refl + + +module AlgebraEquivs {R : Ring ℓ} where + open IsAlgebraHom + open AlgebraHoms + + compIsAlgebraEquiv : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + {g : ⟨ B ⟩ ≃ ⟨ C ⟩} {f : ⟨ A ⟩ ≃ ⟨ B ⟩} + → IsAlgebraEquiv (B .snd) g (C .snd) + → IsAlgebraEquiv (A .snd) f (B .snd) + → IsAlgebraEquiv (A .snd) (compEquiv f g) (C .snd) + compIsAlgebraEquiv {g = g} {f} gh fh = compIsAlgebraHom {g = g .fst} {f .fst} gh fh + + compAlgebraEquiv : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + → AlgebraEquiv A B → AlgebraEquiv B C → AlgebraEquiv A C + fst (compAlgebraEquiv f g) = compEquiv (f .fst) (g .fst) + snd (compAlgebraEquiv f g) = compIsAlgebraEquiv {g = g .fst} {f = f .fst} (g .snd) (f .snd) + + + +-- the Algebra version of uaCompEquiv +module AlgebraUAFunctoriality {R : Ring ℓ} where + open AlgebraStr + open AlgebraEquivs + + Algebra≡ : (A B : Algebra R ℓ') → ( + Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] + Σ[ q0 ∈ PathP (λ i → p i) (0a (snd A)) (0a (snd B)) ] + Σ[ q1 ∈ PathP (λ i → p i) (1a (snd A)) (1a (snd B)) ] + Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] + Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] + Σ[ s- ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] + Σ[ s⋆ ∈ PathP (λ i → ⟨ R ⟩ → p i → p i) (_⋆_ (snd A)) (_⋆_ (snd B)) ] + PathP (λ i → IsAlgebra R (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i)) (isAlgebra (snd A)) + (isAlgebra (snd B))) + ≃ (A ≡ B) + Algebra≡ A B = isoToEquiv theIso + where + open Iso + theIso : Iso _ _ + fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i = p i + , algebrastr (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i) (t i) + inv theIso x = cong ⟨_⟩ x , cong (0a ∘ snd) x , cong (1a ∘ snd) x + , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (_⋆_ ∘ snd) x + , cong (isAlgebra ∘ snd) x + rightInv theIso _ = refl + leftInv theIso _ = refl + + caracAlgebra≡ : {A B : Algebra R ℓ'} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q + caracAlgebra≡ {A = A} {B = B} p q P = + sym (transportTransport⁻ (ua (Algebra≡ A B)) p) + ∙∙ cong (transport (ua (Algebra≡ A B))) helper + ∙∙ transportTransport⁻ (ua (Algebra≡ A B)) q + where + helper : transport (sym (ua (Algebra≡ A B))) p ≡ transport (sym (ua (Algebra≡ A B))) q + helper = Σ≡Prop + (λ _ → isPropΣ + (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isOfHLevelPathP 1 (isPropIsAlgebra _ _ _ _ _ _ _) _ _) + (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + + uaCompAlgebraEquiv : {A B C : Algebra R ℓ'} (f : AlgebraEquiv A B) (g : AlgebraEquiv B C) + → uaAlgebra (compAlgebraEquiv f g) ≡ uaAlgebra f ∙ uaAlgebra g + uaCompAlgebraEquiv f g = caracAlgebra≡ _ _ ( + cong ⟨_⟩ (uaAlgebra (compAlgebraEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaAlgebra f) ∙ cong ⟨_⟩ (uaAlgebra g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaAlgebra f) (uaAlgebra g)) ⟩ + cong ⟨_⟩ (uaAlgebra f ∙ uaAlgebra g) ∎) diff --git a/Cubical/Algebra/CommAlgebra.agda b/Cubical/Algebra/CommAlgebra.agda new file mode 100644 index 0000000000..d150261beb --- /dev/null +++ b/Cubical/Algebra/CommAlgebra.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra where + +open import Cubical.Algebra.CommAlgebra.Base public +open import Cubical.Algebra.CommAlgebra.Properties public diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda new file mode 100644 index 0000000000..428ff7dfe4 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -0,0 +1,250 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Reflection.RecordEquiv + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + +record IsCommAlgebra (R : CommRing ℓ) {A : Type ℓ'} + (0a : A) (1a : A) + (_+_ : A → A → A) (_·_ : A → A → A) (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where + + constructor iscommalgebra + + field + isAlgebra : IsAlgebra (CommRing→Ring R) 0a 1a _+_ _·_ -_ _⋆_ + ·-comm : (x y : A) → x · y ≡ y · x + + open IsAlgebra isAlgebra public + +unquoteDecl IsCommAlgebraIsoΣ = declareRecordIsoΣ IsCommAlgebraIsoΣ (quote IsCommAlgebra) + +record CommAlgebraStr (R : CommRing ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where + + constructor commalgebrastr + + field + 0a : A + 1a : A + _+_ : A → A → A + _·_ : A → A → A + -_ : A → A + _⋆_ : ⟨ R ⟩ → A → A + isCommAlgebra : IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + + open IsCommAlgebra isCommAlgebra public + + infix 8 -_ + infixl 7 _·_ + infixl 7 _⋆_ + infixl 6 _+_ + +CommAlgebra : (R : CommRing ℓ) → ∀ ℓ' → Type (ℓ-max ℓ (ℓ-suc ℓ')) +CommAlgebra R ℓ' = Σ[ A ∈ Type ℓ' ] CommAlgebraStr R A + +module _ {R : CommRing ℓ} where + open CommRingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·s_) + + CommAlgebraStr→AlgebraStr : {A : Type ℓ'} → CommAlgebraStr R A → AlgebraStr (CommRing→Ring R) A + CommAlgebraStr→AlgebraStr (commalgebrastr _ _ _ _ _ _ (iscommalgebra isAlgebra ·-comm)) = + algebrastr _ _ _ _ _ _ isAlgebra + + CommAlgebra→Algebra : (A : CommAlgebra R ℓ') → Algebra (CommRing→Ring R) ℓ' + CommAlgebra→Algebra (_ , str) = (_ , CommAlgebraStr→AlgebraStr str) + + CommAlgebra→CommRing : (A : CommAlgebra R ℓ') → CommRing ℓ' + CommAlgebra→CommRing (_ , commalgebrastr _ _ _ _ _ _ (iscommalgebra isAlgebra ·-comm)) = + _ , commringstr _ _ _ _ _ (iscommring (IsAlgebra.isRing isAlgebra) ·-comm) + + isSetCommAlgebra : (A : CommAlgebra R ℓ') → isSet ⟨ A ⟩ + isSetCommAlgebra A = isSetAlgebra (CommAlgebra→Algebra A) + + makeIsCommAlgebra : {A : Type ℓ'} {0a 1a : A} + {_+_ _·_ : A → A → A} { -_ : A → A} {_⋆_ : ⟨ R ⟩ → A → A} + (isSet-A : isSet A) + (+-assoc : (x y z : A) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : A) → x + 0a ≡ x) + (+-rinv : (x : A) → x + (- x) ≡ 0a) + (+-comm : (x y : A) → x + y ≡ y + x) + (·-assoc : (x y z : A) → x · (y · z) ≡ (x · y) · z) + (·-lid : (x : A) → 1a · x ≡ x) + (·-ldist-+ : (x y z : A) → (x + y) · z ≡ (x · z) + (y · z)) + (·-comm : (x y : A) → x · y ≡ y · x) + (⋆-assoc : (r s : ⟨ R ⟩) (x : A) → (r ·s s) ⋆ x ≡ r ⋆ (s ⋆ x)) + (⋆-ldist : (r s : ⟨ R ⟩) (x : A) → (r +r s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x)) + (⋆-rdist : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) + (⋆-lid : (x : A) → 1r ⋆ x ≡ x) + (⋆-lassoc : (r : ⟨ R ⟩) (x y : A) → (r ⋆ x) · y ≡ r ⋆ (x · y)) + → IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + makeIsCommAlgebra {A = A} {0a} {1a} {_+_} {_·_} { -_} {_⋆_} isSet-A + +-assoc +-rid +-rinv +-comm + ·-assoc ·-lid ·-ldist-+ ·-comm + ⋆-assoc ⋆-ldist ⋆-rdist ⋆-lid ⋆-lassoc + = iscommalgebra + (makeIsAlgebra + isSet-A + +-assoc +-rid +-rinv +-comm + ·-assoc + (λ x → x · 1a ≡⟨ ·-comm _ _ ⟩ 1a · x ≡⟨ ·-lid _ ⟩ x ∎) + ·-lid + (λ x y z → x · (y + z) ≡⟨ ·-comm _ _ ⟩ + (y + z) · x ≡⟨ ·-ldist-+ _ _ _ ⟩ + (y · x) + (z · x) ≡⟨ cong (λ u → (y · x) + u) (·-comm _ _) ⟩ + (y · x) + (x · z) ≡⟨ cong (λ u → u + (x · z)) (·-comm _ _) ⟩ + (x · y) + (x · z) ∎) + ·-ldist-+ + ⋆-assoc + ⋆-ldist + ⋆-rdist + ⋆-lid + ⋆-lassoc + λ r x y → r ⋆ (x · y) ≡⟨ cong (λ u → r ⋆ u) (·-comm _ _) ⟩ + r ⋆ (y · x) ≡⟨ sym (⋆-lassoc _ _ _) ⟩ + (r ⋆ y) · x ≡⟨ ·-comm _ _ ⟩ + x · (r ⋆ y) ∎) + ·-comm + + module _ (S : CommRing ℓ) where + open CommRingStr (snd S) renaming (1r to 1S) + open CommRingStr (snd R) using () renaming (_·_ to _·R_; _+_ to _+R_; 1r to 1R) + commAlgebraFromCommRing : + (_⋆_ : fst R → fst S → fst S) + → ((r s : fst R) (x : fst S) → (r ·R s) ⋆ x ≡ r ⋆ (s ⋆ x)) + → ((r s : fst R) (x : fst S) → (r +R s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x)) + → ((r : fst R) (x y : fst S) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) + → ((x : fst S) → 1R ⋆ x ≡ x) + → ((r : fst R) (x y : fst S) → (r ⋆ x) · y ≡ r ⋆ (x · y)) + → CommAlgebra R ℓ + commAlgebraFromCommRing _⋆_ ·Assoc⋆ ⋆DistR ⋆DistL ⋆Lid ⋆Assoc· = fst S , + commalgebrastr 0r 1S _+_ _·_ -_ _⋆_ + (makeIsCommAlgebra is-set +Assoc +Rid +Rinv +Comm ·Assoc ·Lid ·Ldist+ ·Comm + ·Assoc⋆ ⋆DistR ⋆DistL ⋆Lid ⋆Assoc·) + + + IsCommAlgebraEquiv : {A : Type ℓ'} {B : Type ℓ''} + (M : CommAlgebraStr R A) (e : A ≃ B) (N : CommAlgebraStr R B) + → Type _ + IsCommAlgebraEquiv M e N = + IsAlgebraHom (CommAlgebraStr→AlgebraStr M) (e .fst) (CommAlgebraStr→AlgebraStr N) + + CommAlgebraEquiv : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') → Type _ + CommAlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsCommAlgebraEquiv (M .snd) e (N .snd) + + IsCommAlgebraHom : {A : Type ℓ'} {B : Type ℓ''} + (M : CommAlgebraStr R A) (f : A → B) (N : CommAlgebraStr R B) + → Type _ + IsCommAlgebraHom M f N = + IsAlgebraHom (CommAlgebraStr→AlgebraStr M) f (CommAlgebraStr→AlgebraStr N) + + CommAlgebraHom : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') → Type _ + CommAlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ → ⟨ N ⟩) ] IsCommAlgebraHom (M .snd) f (N .snd) + + CommAlgebraEquiv→CommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraEquiv A B → CommAlgebraHom A B + CommAlgebraEquiv→CommAlgebraHom (e , eIsHom) = e .fst , eIsHom + + module _ {M N : CommAlgebra R ℓ'} where + open CommAlgebraStr {{...}} + open IsAlgebraHom + private + instance + _ = snd M + _ = snd N + + makeCommAlgebraHom : (f : fst M → fst N) + → (fPres1 : f 1a ≡ 1a) + → (fPres+ : (x y : fst M) → f (x + y) ≡ f x + f y) + → (fPres· : (x y : fst M) → f (x · y) ≡ f x · f y) + → (fPres⋆ : (r : fst R) (x : fst M) → f (r ⋆ x) ≡ r ⋆ f x) + → CommAlgebraHom M N + makeCommAlgebraHom f fPres1 fPres+ fPres· fPres⋆ = f , isHom + where fPres0 = + f 0a ≡⟨ sym (+-rid _) ⟩ + f 0a + 0a ≡⟨ cong (λ u → f 0a + u) (sym (+-rinv (f 0a))) ⟩ + f 0a + (f 0a - f 0a) ≡⟨ +-assoc (f 0a) (f 0a) (- f 0a) ⟩ + (f 0a + f 0a) - f 0a ≡⟨ cong (λ u → u - f 0a) (sym (fPres+ 0a 0a)) ⟩ + f (0a + 0a) - f 0a ≡⟨ cong (λ u → f u - f 0a) (+-lid 0a) ⟩ + f 0a - f 0a ≡⟨ +-rinv (f 0a) ⟩ + 0a ∎ + + isHom : IsCommAlgebraHom (snd M) f (snd N) + pres0 isHom = fPres0 + pres1 isHom = fPres1 + pres+ isHom = fPres+ + pres· isHom = fPres· + pres- isHom = (λ x → + f (- x) ≡⟨ sym (+-rid _) ⟩ + (f (- x) + 0a) ≡⟨ cong (λ u → f (- x) + u) (sym (+-rinv (f x))) ⟩ + (f (- x) + (f x - f x)) ≡⟨ +-assoc _ _ _ ⟩ + ((f (- x) + f x) - f x) ≡⟨ cong (λ u → u - f x) (sym (fPres+ _ _)) ⟩ + (f ((- x) + x) - f x) ≡⟨ cong (λ u → f u - f x) (+-linv x) ⟩ + (f 0a - f x) ≡⟨ cong (λ u → u - f x) fPres0 ⟩ + (0a - f x) ≡⟨ +-lid _ ⟩ (- f x) ∎) + pres⋆ isHom = fPres⋆ + + isPropIsCommAlgebraHom : (f : fst M → fst N) → isProp (IsCommAlgebraHom (snd M) f (snd N)) + isPropIsCommAlgebraHom f = isPropIsAlgebraHom + (CommRing→Ring R) + (snd (CommAlgebra→Algebra M)) + f + (snd (CommAlgebra→Algebra N)) + +isPropIsCommAlgebra : (R : CommRing ℓ) {A : Type ℓ'} + (0a 1a : A) + (_+_ _·_ : A → A → A) + (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) + → isProp (IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_) +isPropIsCommAlgebra R _ _ _ _ _ _ = + isOfHLevelRetractFromIso 1 IsCommAlgebraIsoΣ + (isPropΣ (isPropIsAlgebra _ _ _ _ _ _ _) + (λ alg → isPropΠ2 λ _ _ → alg .IsAlgebra.is-set _ _)) + +𝒮ᴰ-CommAlgebra : (R : CommRing ℓ) → DUARel (𝒮-Univ ℓ') (CommAlgebraStr R) (ℓ-max ℓ ℓ') +𝒮ᴰ-CommAlgebra R = + 𝒮ᴰ-Record (𝒮-Univ _) (IsCommAlgebraEquiv {R = R}) + (fields: + data[ 0a ∣ nul ∣ pres0 ] + data[ 1a ∣ nul ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ autoDUARel _ _ ∣ pres- ] + data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] + prop[ isCommAlgebra ∣ (λ _ _ → isPropIsCommAlgebra _ _ _ _ _ _ _) ]) + where + open CommAlgebraStr + open IsAlgebraHom + + -- faster with some sharing + nul = autoDUARel (𝒮-Univ _) (λ A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) +CommAlgebraPath R = ∫ (𝒮ᴰ-CommAlgebra R) .UARel.ua + +uaCommAlgebra : {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} → CommAlgebraEquiv A B → A ≡ B +uaCommAlgebra {R = R} {A = A} {B = B} = equivFun (CommAlgebraPath R A B) + +isGroupoidCommAlgebra : {R : CommRing ℓ} → isGroupoid (CommAlgebra R ℓ') +isGroupoidCommAlgebra A B = isOfHLevelRespectEquiv 2 (CommAlgebraPath _ _ _) (isSetAlgebraEquiv _ _) diff --git a/Cubical/Algebra/CommAlgebra/FGIdeal.agda b/Cubical/Algebra/CommAlgebra/FGIdeal.agda new file mode 100644 index 0000000000..5c5c546b70 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/FGIdeal.agda @@ -0,0 +1,22 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.FGIdeal where +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Vec + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.FGIdeal renaming (generatedIdeal to generatedIdealCommRing) +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.Ideal + +private + variable + ℓ : Level + +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where + open CommAlgebraStr (snd A) + + generatedIdeal : {n : ℕ} → FinVec (fst A) n → IdealsIn A + generatedIdeal = generatedIdealCommRing (CommAlgebra→CommRing A) diff --git a/Cubical/Algebra/CommAlgebra/FPAlgebra.agda b/Cubical/Algebra/CommAlgebra/FPAlgebra.agda new file mode 100644 index 0000000000..ab3cf0713e --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/FPAlgebra.agda @@ -0,0 +1,43 @@ +{- + Finitely presented algebras. + An R-algebra A is finitely presented, if there merely is an exact sequence + of R-modules: + + (a₁,⋯,aₘ) → R[X₁,⋯,Xₙ] → A → 0 +-} +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.FPAlgebra where +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Vec +open import Cubical.Data.Sigma + +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.Instances.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebra.QuotientAlgebra +open import Cubical.Algebra.CommAlgebra.Ideal +open import Cubical.Algebra.CommAlgebra.FGIdeal + +private + variable + ℓ : Level + +module _ {R : CommRing ℓ} where + freeAlgebra : (n : ℕ) → CommAlgebra R ℓ + freeAlgebra n = R [ Fin n ] + + makeFPAlgebra : {m : ℕ} (n : ℕ) (l : FinVec (fst (freeAlgebra n)) m) + → CommAlgebra R ℓ + makeFPAlgebra n l = freeAlgebra n / generatedIdeal (freeAlgebra n) l + + isFPAlgebra : (A : CommAlgebra R ℓ) → Type _ + isFPAlgebra A = ∃[ ((n , m) , l) ∈ Σ[ (n , m) ∈ ℕ × ℕ ] FinVec (fst (freeAlgebra n)) m ] + makeFPAlgebra n l ≡ A + + isFPAlgebraIsProp : {A : CommAlgebra R ℓ} → isProp (isFPAlgebra A) + isFPAlgebraIsProp = isPropPropTrunc diff --git a/Cubical/Algebra/CommAlgebra/Ideal.agda b/Cubical/Algebra/CommAlgebra/Ideal.agda new file mode 100644 index 0000000000..60ba64ffb6 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Ideal.agda @@ -0,0 +1,27 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.Ideal where +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Ideal renaming (IdealsIn to IdealsInCommRing; + makeIdeal to makeIdealCommRing) +open import Cubical.Algebra.CommAlgebra + +private + variable + ℓ : Level + +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where + IdealsIn : Type _ + IdealsIn = IdealsInCommRing (CommAlgebra→CommRing A) + + open CommAlgebraStr (snd A) + + makeIdeal : (I : fst A → hProp ℓ) + → (+-closed : {x y : fst A} → x ∈ I → y ∈ I → (x + y) ∈ I) + → (0-closed : 0a ∈ I) + → (·-closedLeft : {x : fst A} → (r : fst A) → x ∈ I → r · x ∈ I) + → IdealsIn + makeIdeal = makeIdealCommRing (CommAlgebra→CommRing A) diff --git a/Cubical/Algebra/CommAlgebra/Instances/FreeCommAlgebra.agda b/Cubical/Algebra/CommAlgebra/Instances/FreeCommAlgebra.agda new file mode 100644 index 0000000000..c6f30f99ad --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Instances/FreeCommAlgebra.agda @@ -0,0 +1,456 @@ +{-# OPTIONS --safe #-} + +module Cubical.Algebra.CommAlgebra.Instances.FreeCommAlgebra where +{- + The free commutative algebra over a commutative ring, + or in other words the ring of polynomials with coefficients in a given ring. + Note that this is a constructive definition, which entails that polynomials + cannot be represented by lists of coefficients, where the last one is non-zero. + For rings with decidable equality, that is still possible. + + I learned about this (and other) definition(s) from David Jaz Myers. + You can watch him talk about these things here: + https://www.youtube.com/watch?v=VNp-f_9MnVk + + This file contains + * the definition of the free commutative algebra on a type I over a commutative ring R as a HIT + (let us call that R[I]) + * a prove that the construction is an commutative R-algebra + * definitions of the induced maps appearing in the universal property of R[I], + that is: * for any map I → A, where A is a commutative R-algebra, + the induced algebra homomorphism R[I] → A + ('inducedHom') + * for any hom R[I] → A, the 'restricttion to variables' I → A + ('evaluateAt') + * a proof that the two constructions are inverse to each other + ('homRetrievable' and 'mapRetrievable') + * a proof, that the corresponding pointwise equivalence of functors is natural + ('naturalR', 'naturalL') +-} +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function hiding (const) + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring using () +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.Algebra +open import Cubical.HITs.SetTruncation + +private + variable + ℓ ℓ' ℓ'' : Level + +module Construction (R : CommRing ℓ) where + open CommRingStr (snd R) using (1r; 0r) renaming (_+_ to _+r_; _·_ to _·r_) + + data R[_] (I : Type ℓ') : Type (ℓ-max ℓ ℓ') where + var : I → R[ I ] + const : ⟨ R ⟩ → R[ I ] + _+_ : R[ I ] → R[ I ] → R[ I ] + -_ : R[ I ] → R[ I ] + _·_ : R[ I ] → R[ I ] → R[ I ] -- \cdot + + +-assoc : (x y z : R[ I ]) → x + (y + z) ≡ (x + y) + z + +-rid : (x : R[ I ]) → x + (const 0r) ≡ x + +-rinv : (x : R[ I ]) → x + (- x) ≡ (const 0r) + +-comm : (x y : R[ I ]) → x + y ≡ y + x + + ·-assoc : (x y z : R[ I ]) → x · (y · z) ≡ (x · y) · z + ·-lid : (x : R[ I ]) → (const 1r) · x ≡ x + ·-comm : (x y : R[ I ]) → x · y ≡ y · x + + ldist : (x y z : R[ I ]) → (x + y) · z ≡ (x · z) + (y · z) + + +HomConst : (s t : ⟨ R ⟩) → const (s +r t) ≡ const s + const t + ·HomConst : (s t : ⟨ R ⟩) → const (s ·r t) ≡ (const s) · (const t) + + 0-trunc : (x y : R[ I ]) (p q : x ≡ y) → p ≡ q + + _⋆_ : {I : Type ℓ'} → ⟨ R ⟩ → R[ I ] → R[ I ] + r ⋆ x = const r · x + + ⋆-assoc : {I : Type ℓ'} → (s t : ⟨ R ⟩) (x : R[ I ]) → (s ·r t) ⋆ x ≡ s ⋆ (t ⋆ x) + ⋆-assoc s t x = const (s ·r t) · x ≡⟨ cong (λ u → u · x) (·HomConst _ _) ⟩ + (const s · const t) · x ≡⟨ sym (·-assoc _ _ _) ⟩ + const s · (const t · x) ≡⟨ refl ⟩ + s ⋆ (t ⋆ x) ∎ + + ⋆-ldist-+ : {I : Type ℓ'} → (s t : ⟨ R ⟩) (x : R[ I ]) → (s +r t) ⋆ x ≡ (s ⋆ x) + (t ⋆ x) + ⋆-ldist-+ s t x = (s +r t) ⋆ x ≡⟨ cong (λ u → u · x) (+HomConst _ _) ⟩ + (const s + const t) · x ≡⟨ ldist _ _ _ ⟩ + (s ⋆ x) + (t ⋆ x) ∎ + + ⋆-rdist-+ : {I : Type ℓ'} → (s : ⟨ R ⟩) (x y : R[ I ]) → s ⋆ (x + y) ≡ (s ⋆ x) + (s ⋆ y) + ⋆-rdist-+ s x y = const s · (x + y) ≡⟨ ·-comm _ _ ⟩ + (x + y) · const s ≡⟨ ldist _ _ _ ⟩ + (x · const s) + (y · const s) ≡⟨ cong (λ u → u + (y · const s)) (·-comm _ _) ⟩ + (s ⋆ x) + (y · const s) ≡⟨ cong (λ u → (s ⋆ x) + u) (·-comm _ _) ⟩ + (s ⋆ x) + (s ⋆ y) ∎ + + ⋆-assoc-· : {I : Type ℓ'} → (s : ⟨ R ⟩) (x y : R[ I ]) → (s ⋆ x) · y ≡ s ⋆ (x · y) + ⋆-assoc-· s x y = (s ⋆ x) · y ≡⟨ sym (·-assoc _ _ _) ⟩ + s ⋆ (x · y) ∎ + + 0a : {I : Type ℓ'} → R[ I ] + 0a = (const 0r) + + 1a : {I : Type ℓ'} → R[ I ] + 1a = (const 1r) + + isCommAlgebra : {I : Type ℓ'} → IsCommAlgebra R {A = R[ I ]} 0a 1a _+_ _·_ -_ _⋆_ + isCommAlgebra = makeIsCommAlgebra 0-trunc + +-assoc +-rid +-rinv +-comm + ·-assoc ·-lid ldist ·-comm + ⋆-assoc ⋆-ldist-+ ⋆-rdist-+ ·-lid ⋆-assoc-· + +_[_] : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R (ℓ-max ℓ ℓ') +(R [ I ]) = R[ I ] , commalgebrastr 0a 1a _+_ _·_ -_ _⋆_ isCommAlgebra + where + open Construction R + +module Theory {R : CommRing ℓ} {I : Type ℓ'} where + open CommRingStr (snd R) + using (0r; 1r) + renaming (_·_ to _·r_; _+_ to _+r_; ·Comm to ·r-comm; ·Rid to ·r-rid) + + module _ (A : CommAlgebra R ℓ'') (φ : I → ⟨ A ⟩) where + open CommAlgebraStr (A .snd) + open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) + open Construction using (var; const) renaming (_+_ to _+c_; -_ to -c_; _·_ to _·c_) + + imageOf0Works : 0r ⋆ 1a ≡ 0a + imageOf0Works = 0-actsNullifying 1a + + imageOf1Works : 1r ⋆ 1a ≡ 1a + imageOf1Works = ⋆-lid 1a + + inducedMap : ⟨ R [ I ] ⟩ → ⟨ A ⟩ + inducedMap (var x) = φ x + inducedMap (const r) = r ⋆ 1a + inducedMap (P +c Q) = (inducedMap P) + (inducedMap Q) + inducedMap (-c P) = - inducedMap P + inducedMap (Construction.+-assoc P Q S i) = +-assoc (inducedMap P) (inducedMap Q) (inducedMap S) i + inducedMap (Construction.+-rid P i) = + let + eq : (inducedMap P) + (inducedMap (const 0r)) ≡ (inducedMap P) + eq = (inducedMap P) + (inducedMap (const 0r)) ≡⟨ refl ⟩ + (inducedMap P) + (0r ⋆ 1a) ≡⟨ cong + (λ u → (inducedMap P) + u) + (imageOf0Works) ⟩ + (inducedMap P) + 0a ≡⟨ +-rid _ ⟩ + (inducedMap P) ∎ + in eq i + inducedMap (Construction.+-rinv P i) = + let eq : (inducedMap P - inducedMap P) ≡ (inducedMap (const 0r)) + eq = (inducedMap P - inducedMap P) ≡⟨ +-rinv _ ⟩ + 0a ≡⟨ sym imageOf0Works ⟩ + (inducedMap (const 0r))∎ + in eq i + inducedMap (Construction.+-comm P Q i) = +-comm (inducedMap P) (inducedMap Q) i + inducedMap (P ·c Q) = inducedMap P · inducedMap Q + inducedMap (Construction.·-assoc P Q S i) = ·Assoc (inducedMap P) (inducedMap Q) (inducedMap S) i + inducedMap (Construction.·-lid P i) = + let eq = inducedMap (const 1r) · inducedMap P ≡⟨ cong (λ u → u · inducedMap P) imageOf1Works ⟩ + 1a · inducedMap P ≡⟨ ·Lid (inducedMap P) ⟩ + inducedMap P ∎ + in eq i + inducedMap (Construction.·-comm P Q i) = ·-comm (inducedMap P) (inducedMap Q) i + inducedMap (Construction.ldist P Q S i) = ·Ldist+ (inducedMap P) (inducedMap Q) (inducedMap S) i + inducedMap (Construction.+HomConst s t i) = ⋆-ldist s t 1a i + inducedMap (Construction.·HomConst s t i) = + let eq = (s ·r t) ⋆ 1a ≡⟨ cong (λ u → u ⋆ 1a) (·r-comm _ _) ⟩ + (t ·r s) ⋆ 1a ≡⟨ ⋆-assoc t s 1a ⟩ + t ⋆ (s ⋆ 1a) ≡⟨ cong (λ u → t ⋆ u) (sym (·Rid _)) ⟩ + t ⋆ ((s ⋆ 1a) · 1a) ≡⟨ ⋆-rassoc t (s ⋆ 1a) 1a ⟩ + (s ⋆ 1a) · (t ⋆ 1a) ∎ + in eq i + inducedMap (Construction.0-trunc P Q p q i j) = + isSetAlgebra (CommAlgebra→Algebra A) (inducedMap P) (inducedMap Q) (cong _ p) (cong _ q) i j + + module _ where + open IsAlgebraHom + + inducedHom : AlgebraHom (CommAlgebra→Algebra (R [ I ])) (CommAlgebra→Algebra A) + inducedHom .fst = inducedMap + inducedHom .snd .pres0 = 0-actsNullifying _ + inducedHom .snd .pres1 = imageOf1Works + inducedHom .snd .pres+ x y = refl + inducedHom .snd .pres· x y = refl + inducedHom .snd .pres- x = refl + inducedHom .snd .pres⋆ r x = + (r ⋆ 1a) · inducedMap x ≡⟨ ⋆-lassoc r 1a (inducedMap x) ⟩ + r ⋆ (1a · inducedMap x) ≡⟨ cong (λ u → r ⋆ u) (·Lid (inducedMap x)) ⟩ + r ⋆ inducedMap x ∎ + + module _ (A : CommAlgebra R ℓ'') where + open CommAlgebraStr (A .snd) + open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) + open Construction using (var; const) renaming (_+_ to _+c_; -_ to -c_; _·_ to _·c_) + + Hom = AlgebraHom (CommAlgebra→Algebra (R [ I ])) (CommAlgebra→Algebra A) + open IsAlgebraHom + + evaluateAt : Hom → I → ⟨ A ⟩ + evaluateAt φ x = φ .fst (var x) + + mapRetrievable : ∀ (φ : I → ⟨ A ⟩) + → evaluateAt (inducedHom A φ) ≡ φ + mapRetrievable φ = refl + + proveEq : ∀ {X : Type ℓ''} (isSetX : isSet X) (f g : ⟨ R [ I ] ⟩ → X) + → (var-eq : (x : I) → f (var x) ≡ g (var x)) + → (const-eq : (r : ⟨ R ⟩) → f (const r) ≡ g (const r)) + → (+-eq : (x y : ⟨ R [ I ] ⟩) → (eq-x : f x ≡ g x) → (eq-y : f y ≡ g y) + → f (x +c y) ≡ g (x +c y)) + → (·-eq : (x y : ⟨ R [ I ] ⟩) → (eq-x : f x ≡ g x) → (eq-y : f y ≡ g y) + → f (x ·c y) ≡ g (x ·c y)) + → (-eq : (x : ⟨ R [ I ] ⟩) → (eq-x : f x ≡ g x) + → f (-c x) ≡ g (-c x)) + → f ≡ g + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (var x) = var-eq x i + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (const x) = const-eq x i + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (x +c y) = + +-eq x y + (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i y) + i + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (-c x) = + -eq x ((λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x)) i + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (x ·c y) = + ·-eq x y + (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i y) + i + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.+-assoc x y z j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f (x +c (y +c z)) ≡ g (x +c (y +c z)) + a₀₋ = +-eq _ _ (rec x) (+-eq _ _ (rec y) (rec z)) + a₁₋ : f ((x +c y) +c z) ≡ g ((x +c y) +c z) + a₁₋ = +-eq _ _ (+-eq _ _ (rec x) (rec y)) (rec z) + a₋₀ : f (x +c (y +c z)) ≡ f ((x +c y) +c z) + a₋₀ = cong f (Construction.+-assoc x y z) + a₋₁ : g (x +c (y +c z)) ≡ g ((x +c y) +c z) + a₋₁ = cong g (Construction.+-assoc x y z) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.+-rid x j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f (x +c (const 0r)) ≡ g (x +c (const 0r)) + a₀₋ = +-eq _ _ (rec x) (const-eq 0r) + a₁₋ : f x ≡ g x + a₁₋ = rec x + a₋₀ : f (x +c (const 0r)) ≡ f x + a₋₀ = cong f (Construction.+-rid x) + a₋₁ : g (x +c (const 0r)) ≡ g x + a₋₁ = cong g (Construction.+-rid x) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.+-rinv x j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f (x +c (-c x)) ≡ g (x +c (-c x)) + a₀₋ = +-eq x (-c x) (rec x) (-eq x (rec x)) + a₁₋ : f (const 0r) ≡ g (const 0r) + a₁₋ = const-eq 0r + a₋₀ : f (x +c (-c x)) ≡ f (const 0r) + a₋₀ = cong f (Construction.+-rinv x) + a₋₁ : g (x +c (-c x)) ≡ g (const 0r) + a₋₁ = cong g (Construction.+-rinv x) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.+-comm x y j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f (x +c y) ≡ g (x +c y) + a₀₋ = +-eq x y (rec x) (rec y) + a₁₋ : f (y +c x) ≡ g (y +c x) + a₁₋ = +-eq y x (rec y) (rec x) + a₋₀ : f (x +c y) ≡ f (y +c x) + a₋₀ = cong f (Construction.+-comm x y) + a₋₁ : g (x +c y) ≡ g (y +c x) + a₋₁ = cong g (Construction.+-comm x y) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.·-assoc x y z j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f (x ·c (y ·c z)) ≡ g (x ·c (y ·c z)) + a₀₋ = ·-eq _ _ (rec x) (·-eq _ _ (rec y) (rec z)) + a₁₋ : f ((x ·c y) ·c z) ≡ g ((x ·c y) ·c z) + a₁₋ = ·-eq _ _ (·-eq _ _ (rec x) (rec y)) (rec z) + a₋₀ : f (x ·c (y ·c z)) ≡ f ((x ·c y) ·c z) + a₋₀ = cong f (Construction.·-assoc x y z) + a₋₁ : g (x ·c (y ·c z)) ≡ g ((x ·c y) ·c z) + a₋₁ = cong g (Construction.·-assoc x y z) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.·-lid x j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f ((const 1r) ·c x) ≡ g ((const 1r) ·c x) + a₀₋ = ·-eq _ _ (const-eq 1r) (rec x) + a₁₋ : f x ≡ g x + a₁₋ = rec x + a₋₀ : f ((const 1r) ·c x) ≡ f x + a₋₀ = cong f (Construction.·-lid x) + a₋₁ : g ((const 1r) ·c x) ≡ g x + a₋₁ = cong g (Construction.·-lid x) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.·-comm x y j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f (x ·c y) ≡ g (x ·c y) + a₀₋ = ·-eq _ _ (rec x) (rec y) + a₁₋ : f (y ·c x) ≡ g (y ·c x) + a₁₋ = ·-eq _ _ (rec y) (rec x) + a₋₀ : f (x ·c y) ≡ f (y ·c x) + a₋₀ = cong f (Construction.·-comm x y) + a₋₁ : g (x ·c y) ≡ g (y ·c x) + a₋₁ = cong g (Construction.·-comm x y) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.ldist x y z j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f ((x +c y) ·c z) ≡ g ((x +c y) ·c z) + a₀₋ = ·-eq (x +c y) z + (+-eq _ _ (rec x) (rec y)) + (rec z) + a₁₋ : f ((x ·c z) +c (y ·c z)) ≡ g ((x ·c z) +c (y ·c z)) + a₁₋ = +-eq _ _ (·-eq _ _ (rec x) (rec z)) (·-eq _ _ (rec y) (rec z)) + a₋₀ : f ((x +c y) ·c z) ≡ f ((x ·c z) +c (y ·c z)) + a₋₀ = cong f (Construction.ldist x y z) + a₋₁ : g ((x +c y) ·c z) ≡ g ((x ·c z) +c (y ·c z)) + a₋₁ = cong g (Construction.ldist x y z) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.+HomConst s t j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f (const (s +r t)) ≡ g (const (s +r t)) + a₀₋ = const-eq (s +r t) + a₁₋ : f (const s +c const t) ≡ g (const s +c const t) + a₁₋ = +-eq _ _ (const-eq s) (const-eq t) + a₋₀ : f (const (s +r t)) ≡ f (const s +c const t) + a₋₀ = cong f (Construction.+HomConst s t) + a₋₁ : g (const (s +r t)) ≡ g (const s +c const t) + a₋₁ = cong g (Construction.+HomConst s t) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.·HomConst s t j) = + let + rec : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + rec x = (λ i → proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x) + a₀₋ : f (const (s ·r t)) ≡ g (const (s ·r t)) + a₀₋ = const-eq (s ·r t) + a₁₋ : f (const s ·c const t) ≡ g (const s ·c const t) + a₁₋ = ·-eq _ _ (const-eq s) (const-eq t) + a₋₀ : f (const (s ·r t)) ≡ f (const s ·c const t) + a₋₀ = cong f (Construction.·HomConst s t) + a₋₁ : g (const (s ·r t)) ≡ g (const s ·c const t) + a₋₁ = cong g (Construction.·HomConst s t) + in isSet→isSet' isSetX a₀₋ a₁₋ a₋₀ a₋₁ j i + + proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i (Construction.0-trunc x y p q j k) = + let + P : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + P x i = proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x + Q : (x : ⟨ R [ I ] ⟩) → f x ≡ g x + Q x i = proveEq isSetX f g var-eq const-eq +-eq ·-eq -eq i x + in isOfHLevel→isOfHLevelDep 2 + (λ z → isProp→isSet (isSetX (f z) (g z))) _ _ + (cong P p) + (cong Q q) + (Construction.0-trunc x y p q) j k i + + + homRetrievable : ∀ (f : Hom) + → inducedMap A (evaluateAt f) ≡ fst f + homRetrievable f = + proveEq + (isSetAlgebra (CommAlgebra→Algebra A)) + (inducedMap A (evaluateAt f)) + (λ x → f $a x) + (λ x → refl) + (λ r → r ⋆ 1a ≡⟨ cong (λ u → r ⋆ u) (sym f.pres1) ⟩ + r ⋆ (f $a (const 1r)) ≡⟨ sym (f.pres⋆ r _) ⟩ + f $a (const r ·c const 1r) ≡⟨ cong (λ u → f $a u) (sym (Construction.·HomConst r 1r)) ⟩ + f $a (const (r ·r 1r)) ≡⟨ cong (λ u → f $a (const u)) (·r-rid r) ⟩ + f $a (const r) ∎) + + (λ x y eq-x eq-y → + ι (x +c y) ≡⟨ refl ⟩ + (ι x + ι y) ≡⟨ cong (λ u → u + ι y) eq-x ⟩ + ((f $a x) + ι y) ≡⟨ + cong (λ u → (f $a x) + u) eq-y ⟩ + ((f $a x) + (f $a y)) ≡⟨ sym (f.pres+ _ _) ⟩ (f $a (x +c y)) ∎) + + (λ x y eq-x eq-y → + ι (x ·c y) ≡⟨ refl ⟩ + ι x · ι y ≡⟨ cong (λ u → u · ι y) eq-x ⟩ + (f $a x) · (ι y) ≡⟨ cong (λ u → (f $a x) · u) eq-y ⟩ + (f $a x) · (f $a y) ≡⟨ sym (f.pres· _ _) ⟩ + f $a (x ·c y) ∎) + (λ x eq-x → + ι (-c x) ≡⟨ refl ⟩ + - ι x ≡⟨ cong (λ u → - u) eq-x ⟩ + - (f $a x) ≡⟨ sym (f.pres- x) ⟩ + f $a (-c x) ∎) + where + ι = inducedMap A (evaluateAt f) + module f = IsAlgebraHom (f .snd) + + +evaluateAt : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') + (f : AlgebraHom (CommAlgebra→Algebra (R [ I ])) (CommAlgebra→Algebra A)) + → (I → ⟨ A ⟩) +evaluateAt A f x = f $a (Construction.var x) + +inducedHom : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') + (φ : I → ⟨ A ⟩) + → AlgebraHom (CommAlgebra→Algebra (R [ I ])) (CommAlgebra→Algebra A) +inducedHom A φ = Theory.inducedHom A φ + +module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ''} where + open AlgebraHoms + A′ = CommAlgebra→Algebra A + B′ = CommAlgebra→Algebra B + R′ = (CommRing→Ring R) + ν : AlgebraHom A′ B′ → (⟨ A ⟩ → ⟨ B ⟩) + ν φ = φ .fst + + {- + Hom(R[I],A) → (I → A) + ↓ ↓ + Hom(R[I],B) → (I → B) + -} + naturalR : {I : Type ℓ'} (ψ : AlgebraHom A′ B′) + (f : AlgebraHom (CommAlgebra→Algebra (R [ I ])) A′) + → (ν ψ) ∘ evaluateAt A f ≡ evaluateAt B (ψ ∘a f) + naturalR ψ f = refl + + {- + Hom(R[I],A) → (I → A) + ↓ ↓ + Hom(R[J],A) → (J → A) + -} + naturalL : {I J : Type ℓ'} (φ : J → I) + (f : AlgebraHom (CommAlgebra→Algebra (R [ I ])) A′) + → (evaluateAt A f) ∘ φ + ≡ evaluateAt A (f ∘a (inducedHom (R [ I ]) (λ x → Construction.var (φ x)))) + naturalL φ f = refl diff --git a/Cubical/Algebra/CommAlgebra/Instances/Initial.agda b/Cubical/Algebra/CommAlgebra/Instances/Initial.agda new file mode 100644 index 0000000000..8e8e630997 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Instances/Initial.agda @@ -0,0 +1,80 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.Instances.Initial where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma.Properties using (Σ≡Prop) + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom) +open import Cubical.Algebra.CommAlgebra.Base + +private + variable + ℓ : Level + +module _ ((R , str) : CommRing ℓ) where + + initialCAlg : CommAlgebra (R , str) ℓ + initialCAlg = + let open CommRingStr str + in (R , commalgebrastr _ _ _ _ _ (λ r x → r · x) + (makeIsCommAlgebra (isSetRing (CommRing→Ring (R , str))) + +Assoc +Rid +Rinv +Comm + ·Assoc ·Lid + ·Ldist+ ·Comm + (λ x y z → sym (·Assoc x y z)) ·Ldist+ ·Rdist+ ·Lid + λ x y z → sym (·Assoc x y z))) + + + module _ (A : CommAlgebra (R , str) ℓ) where + open CommAlgebraStr ⦃... ⦄ + private + instance + _ : CommAlgebraStr (R , str) (fst A) + _ = snd A + _ : CommAlgebraStr (R , str) R + _ = snd initialCAlg + + _*_ : R → (fst A) → (fst A) + r * a = CommAlgebraStr._⋆_ (snd A) r a + + initialMap : CommAlgebraHom initialCAlg A + initialMap = + makeCommAlgebraHom {M = initialCAlg} {N = A} + (λ r → r * 1a) + (⋆-lid _) + (λ x y → ⋆-ldist x y 1a) + (λ x y → (x · y) * 1a ≡⟨ ⋆-assoc _ _ _ ⟩ + x * (y * 1a) ≡[ i ]⟨ x * (·Lid (y * 1a) (~ i)) ⟩ + x * (1a · (y * 1a)) ≡⟨ sym (⋆-lassoc _ _ _) ⟩ + (x * 1a) · (y * 1a) ∎) + (λ r x → (r · x) * 1a ≡⟨ ⋆-assoc _ _ _ ⟩ + (r * (x * 1a)) ∎) + + initialMapEq : (f : CommAlgebraHom initialCAlg A) + → f ≡ initialMap + initialMapEq f = + let open IsAlgebraHom (snd f) + in Σ≡Prop + (isPropIsCommAlgebraHom {M = initialCAlg} {N = A}) + λ i x → + ((fst f) x ≡⟨ cong (fst f) (sym (·Rid _)) ⟩ + fst f (x · 1a) ≡⟨ pres⋆ x 1a ⟩ + CommAlgebraStr._⋆_ (snd A) x (fst f 1a) ≡⟨ cong + (λ u → (snd A CommAlgebraStr.⋆ x) u) + pres1 ⟩ + (CommAlgebraStr._⋆_ (snd A) x 1a) ∎) i + + initialityIso : Iso (CommAlgebraHom initialCAlg A) (Unit* {ℓ = ℓ}) + initialityIso = iso (λ _ → tt*) + (λ _ → initialMap) + (λ {tt*x → refl}) + λ f → sym (initialMapEq f) + + initialityPath : CommAlgebraHom initialCAlg A ≡ Unit* + initialityPath = isoToPath initialityIso diff --git a/Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda b/Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda new file mode 100644 index 0000000000..171a9a6eae --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda @@ -0,0 +1,26 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.Instances.Pointwise where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Algebra.CommRing.Base +open import Cubical.Algebra.CommRing.Instances.Pointwise +open import Cubical.Algebra.CommAlgebra.Base + +private + variable + ℓ : Level + +pointwiseAlgebra : {R : CommRing ℓ} (X : Type ℓ) (A : CommAlgebra R ℓ) → CommAlgebra R ℓ +pointwiseAlgebra {R = R} X A = + let open CommAlgebraStr (snd A) + isSetX→A = isOfHLevelΠ 2 (λ (x : X) → isSetCommRing (CommAlgebra→CommRing A)) + in commAlgebraFromCommRing + (pointwiseRing X (CommAlgebra→CommRing A)) + (λ r f → (λ x → r ⋆ (f x))) + (λ r s f i x → ⋆-assoc r s (f x) i) + (λ r s f i x → ⋆-ldist r s (f x) i) + (λ r f g i x → ⋆-rdist r (f x) (g x) i) + (λ f i x → ⋆-lid (f x) i) + λ r f g i x → ⋆-lassoc r (f x) (g x) i diff --git a/Cubical/Algebra/CommAlgebra/Localisation.agda b/Cubical/Algebra/CommAlgebra/Localisation.agda new file mode 100644 index 0000000000..e9baa6bc98 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Localisation.agda @@ -0,0 +1,225 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommAlgebra.Localisation where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Powerset + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Structures.Axioms +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommRing.Base +open import Cubical.Algebra.CommRing.Properties +open import Cubical.Algebra.CommRing.Localisation.Base +open import Cubical.Algebra.CommRing.Localisation.UniversalProperty +open import Cubical.Algebra.CommRing.Localisation.InvertingElements +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Properties + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + + +private + variable + ℓ ℓ′ : Level + + + +module AlgLoc (R' : CommRing ℓ) + (S' : ℙ (fst R')) (SMultClosedSubset : isMultClosedSubset R' S') where + open isMultClosedSubset + private R = fst R' + open CommAlgebraStr + open IsAlgebraHom + open CommRingStr (snd R') renaming (_+_ to _+r_ ; _·_ to _·r_ ; ·Rid to ·rRid) + open RingTheory (CommRing→Ring R') + open CommRingTheory R' + open Loc R' S' SMultClosedSubset + open S⁻¹RUniversalProp R' S' SMultClosedSubset + open CommAlgChar R' + + + S⁻¹RAsCommAlg : CommAlgebra R' ℓ + S⁻¹RAsCommAlg = toCommAlg (S⁻¹RAsCommRing , /1AsCommRingHom) + + + hasLocAlgUniversalProp : (A : CommAlgebra R' ℓ) + → (∀ s → s ∈ S' → _⋆_ (snd A) s (1a (snd A)) ∈ (CommAlgebra→CommRing A) ˣ) + → Type (ℓ-suc ℓ) + hasLocAlgUniversalProp A _ = (B : CommAlgebra R' ℓ) + → (∀ s → s ∈ S' → _⋆_ (snd B) s (1a (snd B)) ∈ (CommAlgebra→CommRing B) ˣ) + → isContr (CommAlgebraHom A B) + + S⋆1⊆S⁻¹Rˣ : ∀ s → s ∈ S' → _⋆_ (snd S⁻¹RAsCommAlg) s (1a (snd S⁻¹RAsCommAlg)) ∈ S⁻¹Rˣ + S⋆1⊆S⁻¹Rˣ s s∈S' = subst-∈ S⁻¹Rˣ + (cong [_] (≡-× (sym (·rRid s)) (Σ≡Prop (λ x → S' x .snd) (sym (·rRid _))))) + (S/1⊆S⁻¹Rˣ s s∈S') + + + S⁻¹RHasAlgUniversalProp : hasLocAlgUniversalProp S⁻¹RAsCommAlg S⋆1⊆S⁻¹Rˣ + S⁻¹RHasAlgUniversalProp B' S⋆1⊆Bˣ = χᴬ , χᴬuniqueness + where + B = fromCommAlg B' .fst + φ = fromCommAlg B' .snd + open CommRingStr (snd B) renaming (_·_ to _·b_ ; 1r to 1b ; ·Lid to ·bLid) + + χ : CommRingHom S⁻¹RAsCommRing B + χ = S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .fst .fst + + χcomp : ∀ r → fst χ (r /1) ≡ fst φ r + χcomp = funExt⁻ (S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .fst .snd) + + χᴬ : CommAlgebraHom S⁻¹RAsCommAlg B' + fst χᴬ = fst χ + pres0 (snd χᴬ) = IsRingHom.pres0 (snd χ) + pres1 (snd χᴬ) = IsRingHom.pres1 (snd χ) + pres+ (snd χᴬ) = IsRingHom.pres+ (snd χ) + pres· (snd χᴬ) = IsRingHom.pres· (snd χ) + pres- (snd χᴬ) = IsRingHom.pres- (snd χ) + pres⋆ (snd χᴬ) r x = path + where + path : fst χ ((r /1) ·ₗ x) ≡ _⋆_ (snd B') r (fst χ x) + path = fst χ ((r /1) ·ₗ x) ≡⟨ IsRingHom.pres· (snd χ) _ _ ⟩ + fst χ (r /1) ·b fst χ x ≡⟨ cong (_·b fst χ x) (χcomp r) ⟩ + fst φ r ·b fst χ x ≡⟨ refl ⟩ + _⋆_ (snd B') r 1b ·b fst χ x ≡⟨ ⋆-lassoc (snd B') _ _ _ ⟩ + _⋆_ (snd B') r (1b ·b fst χ x) ≡⟨ cong (_⋆_ (snd B') r) (·bLid _) ⟩ + _⋆_ (snd B') r (fst χ x) ∎ + + + χᴬuniqueness : (ψ : CommAlgebraHom S⁻¹RAsCommAlg B') → χᴬ ≡ ψ + χᴬuniqueness ψ = Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) + (cong (fst ∘ fst) (χuniqueness (ψ' , funExt ψ'r/1≡φr))) + where + χuniqueness = S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .snd + + ψ' : CommRingHom S⁻¹RAsCommRing B + fst ψ' = fst ψ + IsRingHom.pres0 (snd ψ') = pres0 (snd ψ) + IsRingHom.pres1 (snd ψ') = pres1 (snd ψ) + IsRingHom.pres+ (snd ψ') = pres+ (snd ψ) + IsRingHom.pres· (snd ψ') = pres· (snd ψ) + IsRingHom.pres- (snd ψ') = pres- (snd ψ) + + ψ'r/1≡φr : ∀ r → fst ψ (r /1) ≡ fst φ r + ψ'r/1≡φr r = + fst ψ (r /1) ≡⟨ cong (fst ψ) (sym (·ₗ-rid _)) ⟩ + fst ψ (_⋆_ (snd S⁻¹RAsCommAlg) r (1a (snd S⁻¹RAsCommAlg))) ≡⟨ pres⋆ (snd ψ) _ _ ⟩ + _⋆_ (snd B') r (fst ψ (1a (snd S⁻¹RAsCommAlg))) ≡⟨ cong (_⋆_ (snd B') r) (pres1 (snd ψ)) ⟩ + _⋆_ (snd B') r 1b ∎ + + + -- an immediate corollary: + isContrHomS⁻¹RS⁻¹R : isContr (CommAlgebraHom S⁻¹RAsCommAlg S⁻¹RAsCommAlg) + isContrHomS⁻¹RS⁻¹R = S⁻¹RHasAlgUniversalProp S⁻¹RAsCommAlg S⋆1⊆S⁻¹Rˣ + + +-- the special case of localizing at a single element +R[1/_]AsCommAlgebra : {R : CommRing ℓ} → ⟨ R ⟩ → CommAlgebra R ℓ +R[1/_]AsCommAlgebra {R = R} f = S⁻¹RAsCommAlg [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + where + open AlgLoc R + open InvertingElementsBase R + +module AlgLocTwoSubsets (R' : CommRing ℓ) + (S₁ : ℙ (fst R')) (S₁MultClosedSubset : isMultClosedSubset R' S₁) + (S₂ : ℙ (fst R')) (S₂MultClosedSubset : isMultClosedSubset R' S₂) where + open isMultClosedSubset + open CommRingStr (snd R') hiding (is-set) + open RingTheory (CommRing→Ring R') + open Loc R' S₁ S₁MultClosedSubset renaming (S⁻¹R to S₁⁻¹R ; + S⁻¹RAsCommRing to S₁⁻¹RAsCommRing) + open Loc R' S₂ S₂MultClosedSubset renaming (S⁻¹R to S₂⁻¹R ; + S⁻¹RAsCommRing to S₂⁻¹RAsCommRing) + open AlgLoc R' S₁ S₁MultClosedSubset renaming ( S⁻¹RAsCommAlg to S₁⁻¹RAsCommAlg + ; S⋆1⊆S⁻¹Rˣ to S₁⋆1⊆S₁⁻¹Rˣ + ; S⁻¹RHasAlgUniversalProp to S₁⁻¹RHasAlgUniversalProp + ; isContrHomS⁻¹RS⁻¹R to isContrHomS₁⁻¹RS₁⁻¹R) + open AlgLoc R' S₂ S₂MultClosedSubset renaming ( S⁻¹RAsCommAlg to S₂⁻¹RAsCommAlg + ; S⋆1⊆S⁻¹Rˣ to S₂⋆1⊆S₂⁻¹Rˣ + ; S⁻¹RHasAlgUniversalProp to S₂⁻¹RHasAlgUniversalProp + ; isContrHomS⁻¹RS⁻¹R to isContrHomS₂⁻¹RS₂⁻¹R) + + open IsAlgebraHom + open AlgebraHoms + open CommAlgebraHoms + open CommAlgebraStr ⦃...⦄ + + private + R = fst R' + S₁⁻¹Rˣ = S₁⁻¹RAsCommRing ˣ + S₂⁻¹Rˣ = S₂⁻¹RAsCommRing ˣ + instance + _ = snd S₁⁻¹RAsCommAlg + _ = snd S₂⁻¹RAsCommAlg + + + isContrS₁⁻¹R≡S₂⁻¹R : (∀ s₁ → s₁ ∈ S₁ → s₁ ⋆ 1a ∈ S₂⁻¹Rˣ) + → (∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1a ∈ S₁⁻¹Rˣ) + → isContr (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) + isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ = isOfHLevelRetractFromIso 0 + (equivToIso (invEquiv (CommAlgebraPath _ _ _))) + isContrS₁⁻¹R≅S₂⁻¹R + where + χ₁ : CommAlgebraHom S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg + χ₁ = S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .fst + + χ₂ : CommAlgebraHom S₂⁻¹RAsCommAlg S₁⁻¹RAsCommAlg + χ₂ = S₂⁻¹RHasAlgUniversalProp S₁⁻¹RAsCommAlg S₂⊆S₁⁻¹Rˣ .fst + + χ₁∘χ₂≡id : χ₁ ∘a χ₂ ≡ idCommAlgebraHom _ + χ₁∘χ₂≡id = isContr→isProp isContrHomS₂⁻¹RS₂⁻¹R _ _ + + χ₂∘χ₁≡id : χ₂ ∘a χ₁ ≡ idCommAlgebraHom _ + χ₂∘χ₁≡id = isContr→isProp isContrHomS₁⁻¹RS₁⁻¹R _ _ + + IsoS₁⁻¹RS₂⁻¹R : Iso S₁⁻¹R S₂⁻¹R + Iso.fun IsoS₁⁻¹RS₂⁻¹R = fst χ₁ + Iso.inv IsoS₁⁻¹RS₂⁻¹R = fst χ₂ + Iso.rightInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong fst χ₁∘χ₂≡id) + Iso.leftInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong fst χ₂∘χ₁≡id) + + isContrS₁⁻¹R≅S₂⁻¹R : isContr (CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) + isContrS₁⁻¹R≅S₂⁻¹R = center , uniqueness + where + center : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg + fst center = isoToEquiv IsoS₁⁻¹RS₂⁻¹R + pres0 (snd center) = pres0 (snd χ₁) + pres1 (snd center) = pres1 (snd χ₁) + pres+ (snd center) = pres+ (snd χ₁) + pres· (snd center) = pres· (snd χ₁) + pres- (snd center) = pres- (snd χ₁) + pres⋆ (snd center) = pres⋆ (snd χ₁) + + uniqueness : (φ : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) → center ≡ φ + uniqueness φ = Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) + (equivEq (cong fst + (S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .snd + (AlgebraEquiv→AlgebraHom φ)))) + + + isPropS₁⁻¹R≡S₂⁻¹R : isProp (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) + isPropS₁⁻¹R≡S₂⁻¹R S₁⁻¹R≡S₂⁻¹R = + isContr→isProp (isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ) S₁⁻¹R≡S₂⁻¹R + where + S₁⊆S₂⁻¹Rˣ : ∀ s₁ → s₁ ∈ S₁ → s₁ ⋆ 1a ∈ S₂⁻¹Rˣ + S₁⊆S₂⁻¹Rˣ s₁ s₁∈S₁ = + transport (λ i → _⋆_ ⦃ S₁⁻¹R≡S₂⁻¹R i .snd ⦄ s₁ (1a ⦃ S₁⁻¹R≡S₂⁻¹R i .snd ⦄) + ∈ (CommAlgebra→CommRing (S₁⁻¹R≡S₂⁻¹R i)) ˣ) (S₁⋆1⊆S₁⁻¹Rˣ s₁ s₁∈S₁) + + S₂⊆S₁⁻¹Rˣ : ∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1a ∈ S₁⁻¹Rˣ + S₂⊆S₁⁻¹Rˣ s₂ s₂∈S₂ = + transport (λ i → _⋆_ ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄ s₂ (1a ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄) + ∈ (CommAlgebra→CommRing ((sym S₁⁻¹R≡S₂⁻¹R) i)) ˣ) (S₂⋆1⊆S₂⁻¹Rˣ s₂ s₂∈S₂) diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda new file mode 100644 index 0000000000..f294042d35 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -0,0 +1,258 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommAlgebra.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Structures.Axioms +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommAlgebra.Base + +open import Cubical.HITs.PropositionalTruncation + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + + +-- An R-algebra is the same as a CommRing A with a CommRingHom φ : R → A +module CommAlgChar (R : CommRing ℓ) where + open Iso + open IsRingHom + open CommRingTheory + + + CommRingWithHom : Type (ℓ-suc ℓ) + CommRingWithHom = Σ[ A ∈ CommRing ℓ ] CommRingHom R A + + toCommAlg : CommRingWithHom → CommAlgebra R ℓ + toCommAlg (A , φ , φIsHom) = ⟨ A ⟩ , ACommAlgStr + where + open CommRingStr (snd A) + ACommAlgStr : CommAlgebraStr R ⟨ A ⟩ + CommAlgebraStr.0a ACommAlgStr = 0r + CommAlgebraStr.1a ACommAlgStr = 1r + CommAlgebraStr._+_ ACommAlgStr = _+_ + CommAlgebraStr._·_ ACommAlgStr = _·_ + CommAlgebraStr.- ACommAlgStr = -_ + CommAlgebraStr._⋆_ ACommAlgStr r a = (φ r) · a + CommAlgebraStr.isCommAlgebra ACommAlgStr = makeIsCommAlgebra + is-set +Assoc +Rid +Rinv +Comm ·Assoc ·Lid ·Ldist+ ·Comm + (λ _ _ x → cong (λ y → y · x) (pres· φIsHom _ _) ∙ sym (·Assoc _ _ _)) + (λ _ _ x → cong (λ y → y · x) (pres+ φIsHom _ _) ∙ ·Ldist+ _ _ _) + (λ _ _ _ → ·Rdist+ _ _ _) + (λ x → cong (λ y → y · x) (pres1 φIsHom) ∙ ·Lid x) + (λ _ _ _ → sym (·Assoc _ _ _)) + + + fromCommAlg : CommAlgebra R ℓ → CommRingWithHom + fromCommAlg A = (CommAlgebra→CommRing A) , φ , φIsHom + where + open CommRingStr (snd R) renaming (_·_ to _·r_) hiding (·Lid) + open CommAlgebraStr (snd A) + open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) + φ : ⟨ R ⟩ → ⟨ A ⟩ + φ r = r ⋆ 1a + φIsHom : IsRingHom (CommRing→Ring R .snd) φ (CommRing→Ring (CommAlgebra→CommRing A) .snd) + φIsHom = makeIsRingHom (⋆-lid _) (λ _ _ → ⋆-ldist _ _ _) + λ x y → cong (λ a → (x ·r y) ⋆ a) (sym (·Lid _)) ∙ ⋆Dist· _ _ _ _ + + + CommRingWithHomRoundTrip : (Aφ : CommRingWithHom) → fromCommAlg (toCommAlg Aφ) ≡ Aφ + CommRingWithHomRoundTrip (A , φ) = ΣPathP (APath , φPathP) + where + open CommRingStr + -- note that the proofs of the axioms might differ! + APath : fst (fromCommAlg (toCommAlg (A , φ))) ≡ A + fst (APath i) = ⟨ A ⟩ + 0r (snd (APath i)) = 0r (snd A) + 1r (snd (APath i)) = 1r (snd A) + _+_ (snd (APath i)) = _+_ (snd A) + _·_ (snd (APath i)) = _·_ (snd A) + -_ (snd (APath i)) = -_ (snd A) + isCommRing (snd (APath i)) = isProp→PathP (λ i → isPropIsCommRing _ _ _ _ _ ) + (isCommRing (snd (fst (fromCommAlg (toCommAlg (A , φ)))))) (isCommRing (snd A)) i + + -- this only works because fst (APath i) = fst A definitionally! + φPathP : PathP (λ i → CommRingHom R (APath i)) (snd (fromCommAlg (toCommAlg (A , φ)))) φ + φPathP = RingHomPathP _ _ _ _ _ _ λ i x → ·Rid (snd A) (fst φ x) i + + + CommAlgRoundTrip : (A : CommAlgebra R ℓ) → toCommAlg (fromCommAlg A) ≡ A + CommAlgRoundTrip A = ΣPathP (refl , AlgStrPathP) + where + open CommAlgebraStr ⦃...⦄ + instance + _ = snd A + AlgStrPathP : PathP (λ i → CommAlgebraStr R ⟨ A ⟩) (snd (toCommAlg (fromCommAlg A))) (snd A) + CommAlgebraStr.0a (AlgStrPathP i) = 0a + CommAlgebraStr.1a (AlgStrPathP i) = 1a + CommAlgebraStr._+_ (AlgStrPathP i) = _+_ + CommAlgebraStr._·_ (AlgStrPathP i) = _·_ + CommAlgebraStr.-_ (AlgStrPathP i) = -_ + CommAlgebraStr._⋆_ (AlgStrPathP i) r x = (⋆-lassoc r 1a x ∙ cong (r ⋆_) (·Lid x)) i + CommAlgebraStr.isCommAlgebra (AlgStrPathP i) = isProp→PathP + (λ i → isPropIsCommAlgebra _ _ _ _ _ _ (CommAlgebraStr._⋆_ (AlgStrPathP i))) + (CommAlgebraStr.isCommAlgebra (snd (toCommAlg (fromCommAlg A)))) isCommAlgebra i + + + CommAlgIso : Iso (CommAlgebra R ℓ) CommRingWithHom + fun CommAlgIso = fromCommAlg + inv CommAlgIso = toCommAlg + rightInv CommAlgIso = CommRingWithHomRoundTrip + leftInv CommAlgIso = CommAlgRoundTrip + + +module CommAlgebraHoms {R : CommRing ℓ} where + open AlgebraHoms + + idCommAlgebraHom : (A : CommAlgebra R ℓ') → CommAlgebraHom A A + idCommAlgebraHom A = idAlgebraHom (CommAlgebra→Algebra A) + + compCommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') (C : CommAlgebra R ℓ''') + → CommAlgebraHom A B → CommAlgebraHom B C → CommAlgebraHom A C + compCommAlgebraHom A B C = compAlgebraHom {A = CommAlgebra→Algebra A} + {CommAlgebra→Algebra B} {CommAlgebra→Algebra C} + + _∘ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} + → CommAlgebraHom B C → CommAlgebraHom A B → CommAlgebraHom A C + g ∘ca f = compCommAlgebraHom _ _ _ f g + + compIdCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) + → compCommAlgebraHom _ _ _ (idCommAlgebraHom A) f ≡ f + compIdCommAlgebraHom = compIdAlgebraHom + + idCompCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) + → compCommAlgebraHom _ _ _ f (idCommAlgebraHom B) ≡ f + idCompCommAlgebraHom = idCompAlgebraHom + + compAssocCommAlgebraHom : {A B C D : CommAlgebra R ℓ'} + (f : CommAlgebraHom A B) (g : CommAlgebraHom B C) (h : CommAlgebraHom C D) + → compCommAlgebraHom _ _ _ (compCommAlgebraHom _ _ _ f g) h + ≡ compCommAlgebraHom _ _ _ f (compCommAlgebraHom _ _ _ g h) + compAssocCommAlgebraHom = compAssocAlgebraHom + +module CommAlgebraEquivs {R : CommRing ℓ} where + open AlgebraEquivs + + compCommAlgebraEquiv : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} + → CommAlgebraEquiv A B → CommAlgebraEquiv B C → CommAlgebraEquiv A C + compCommAlgebraEquiv {A = A} {B = B} {C = C} = compAlgebraEquiv {A = CommAlgebra→Algebra A} + {B = CommAlgebra→Algebra B} + {C = CommAlgebra→Algebra C} + + +-- the CommAlgebra version of uaCompEquiv +module CommAlgebraUAFunctoriality {R : CommRing ℓ} where + open CommAlgebraStr + open CommAlgebraEquivs + + CommAlgebra≡ : (A B : CommAlgebra R ℓ') → ( + Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] + Σ[ q0 ∈ PathP (λ i → p i) (0a (snd A)) (0a (snd B)) ] + Σ[ q1 ∈ PathP (λ i → p i) (1a (snd A)) (1a (snd B)) ] + Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] + Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] + Σ[ s- ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] + Σ[ s⋆ ∈ PathP (λ i → ⟨ R ⟩ → p i → p i) (_⋆_ (snd A)) (_⋆_ (snd B)) ] + PathP (λ i → IsCommAlgebra R (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i)) (isCommAlgebra (snd A)) + (isCommAlgebra (snd B))) + ≃ (A ≡ B) + CommAlgebra≡ A B = isoToEquiv theIso + where + open Iso + theIso : Iso _ _ + fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i = p i + , commalgebrastr (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i) (t i) + inv theIso x = cong ⟨_⟩ x , cong (0a ∘ snd) x , cong (1a ∘ snd) x + , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (_⋆_ ∘ snd) x + , cong (isCommAlgebra ∘ snd) x + rightInv theIso _ = refl + leftInv theIso _ = refl + + caracCommAlgebra≡ : {A B : CommAlgebra R ℓ'} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q + caracCommAlgebra≡ {A = A} {B = B} p q P = + sym (transportTransport⁻ (ua (CommAlgebra≡ A B)) p) + ∙∙ cong (transport (ua (CommAlgebra≡ A B))) helper + ∙∙ transportTransport⁻ (ua (CommAlgebra≡ A B)) q + where + helper : transport (sym (ua (CommAlgebra≡ A B))) p ≡ transport (sym (ua (CommAlgebra≡ A B))) q + helper = Σ≡Prop + (λ _ → isPropΣ + (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isOfHLevelPathP 1 (isPropIsCommAlgebra _ _ _ _ _ _ _) _ _) + (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + + uaCompCommAlgebraEquiv : {A B C : CommAlgebra R ℓ'} (f : CommAlgebraEquiv A B) (g : CommAlgebraEquiv B C) + → uaCommAlgebra (compCommAlgebraEquiv f g) ≡ uaCommAlgebra f ∙ uaCommAlgebra g + uaCompCommAlgebraEquiv f g = caracCommAlgebra≡ _ _ ( + cong ⟨_⟩ (uaCommAlgebra (compCommAlgebraEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaCommAlgebra f) ∙ cong ⟨_⟩ (uaCommAlgebra g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaCommAlgebra f) (uaCommAlgebra g)) ⟩ + cong ⟨_⟩ (uaCommAlgebra f ∙ uaCommAlgebra g) ∎) + + +open CommAlgebraHoms +open CommAlgebraEquivs +open CommAlgebraUAFunctoriality +recPT→CommAlgebra : {R : CommRing ℓ} {A : Type ℓ'} (𝓕 : A → CommAlgebra R ℓ'') + → (σ : ∀ x y → CommAlgebraEquiv (𝓕 x) (𝓕 y)) + → (∀ x y z → σ x z ≡ compCommAlgebraEquiv (σ x y) (σ y z)) + ------------------------------------------------------ + → ∥ A ∥ → CommAlgebra R ℓ'' +recPT→CommAlgebra 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommAlgebra 𝓕 + (3-ConstantCompChar 𝓕 (λ x y → uaCommAlgebra (σ x y)) + λ x y z → sym ( cong uaCommAlgebra (compCoh x y z) + ∙ uaCompCommAlgebraEquiv (σ x y) (σ y z))) + + +contrCommAlgebraHom→contrCommAlgebraEquiv : {R : CommRing ℓ} {A : Type ℓ'} + (σ : A → CommAlgebra R ℓ'') + → (∀ x y → isContr (CommAlgebraHom (σ x) (σ y))) + ---------------------------------------------------------------------------- + → ∀ x y → isContr (CommAlgebraEquiv (σ x) (σ y)) +contrCommAlgebraHom→contrCommAlgebraEquiv σ contrHom x y = σEquiv , + λ e → Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) + (Σ≡Prop isPropIsEquiv (cong fst (contrHom _ _ .snd (CommAlgebraEquiv→CommAlgebraHom e)))) + where + open Iso + χ₁ = contrHom x y .fst + χ₂ = contrHom y x .fst + χ₁∘χ₂≡id : χ₁ ∘ca χ₂ ≡ idCommAlgebraHom _ + χ₁∘χ₂≡id = isContr→isProp (contrHom _ _) _ _ + + χ₂∘χ₁≡id : χ₂ ∘ca χ₁ ≡ idCommAlgebraHom _ + χ₂∘χ₁≡id = isContr→isProp (contrHom _ _) _ _ + + σIso : Iso ⟨ σ x ⟩ ⟨ σ y ⟩ + fun σIso = fst χ₁ + inv σIso = fst χ₂ + rightInv σIso = funExt⁻ (cong fst χ₁∘χ₂≡id) + leftInv σIso = funExt⁻ (cong fst χ₂∘χ₁≡id) + + σEquiv : CommAlgebraEquiv (σ x) (σ y) + fst σEquiv = isoToEquiv σIso + snd σEquiv = snd χ₁ diff --git a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda new file mode 100644 index 0000000000..860621aafe --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda @@ -0,0 +1,64 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.QuotientAlgebra where +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset using (_∈_) + +open import Cubical.HITs.SetQuotients hiding (_/_) + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.QuotientRing renaming (_/_ to _/Ring_) hiding ([_]/) +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.Ideal +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Ring.Ideal using (isIdeal) + +private + variable + ℓ : Level + +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where + open CommRingStr {{...}} hiding (_-_; -_; dist; ·Lid; ·Rdist+) renaming (_·_ to _·R_; _+_ to _+R_) + open CommAlgebraStr {{...}} + open RingTheory (CommRing→Ring (CommAlgebra→CommRing A)) using (-DistR·) + instance + _ : CommRingStr _ + _ = snd R + _ : CommAlgebraStr _ _ + _ = snd A + + _/_ : (I : IdealsIn A) → CommAlgebra R ℓ + _/_ I = commAlgebraFromCommRing + ((CommAlgebra→CommRing A) /Ring I) + (λ r → elim (λ _ → squash/) (λ x → [ r ⋆ x ]) (eq r)) + (λ r s → elimProp (λ _ → squash/ _ _) + λ x i → [ ((r ·R s) ⋆ x ≡⟨ ⋆-assoc r s x ⟩ + r ⋆ (s ⋆ x) ∎) i ]) + (λ r s → elimProp (λ _ → squash/ _ _) + λ x i → [ ((r +R s) ⋆ x ≡⟨ ⋆-ldist r s x ⟩ + r ⋆ x + s ⋆ x ∎) i ]) + (λ r → elimProp2 (λ _ _ → squash/ _ _) + λ x y i → [ (r ⋆ (x + y) ≡⟨ ⋆-rdist r x y ⟩ + r ⋆ x + r ⋆ y ∎) i ]) + (elimProp (λ _ → squash/ _ _) + (λ x i → [ (1r ⋆ x ≡⟨ ⋆-lid x ⟩ x ∎) i ])) + λ r → elimProp2 (λ _ _ → squash/ _ _) + λ x y i → [ ((r ⋆ x) · y ≡⟨ ⋆-lassoc r x y ⟩ + r ⋆ (x · y) ∎) i ] + + where + eq : (r : fst R) (x y : fst A) → x - y ∈ (fst I) → [ r ⋆ x ] ≡ [ r ⋆ y ] + eq r x y x-y∈I = eq/ _ _ + (subst (λ u → u ∈ fst I) + ((r ⋆ 1a) · (x - y) ≡⟨ ·Rdist+ (r ⋆ 1a) x (- y) ⟩ + (r ⋆ 1a) · x + (r ⋆ 1a) · (- y) ≡[ i ]⟨ (r ⋆ 1a) · x + -DistR· (r ⋆ 1a) y i ⟩ + (r ⋆ 1a) · x - (r ⋆ 1a) · y ≡[ i ]⟨ ⋆-lassoc r 1a x i + - ⋆-lassoc r 1a y i ⟩ + r ⋆ (1a · x) - r ⋆ (1a · y) ≡[ i ]⟨ r ⋆ (·Lid x i) - r ⋆ (·Lid y i) ⟩ + r ⋆ x - r ⋆ y ∎ ) + (isIdeal.·-closedLeft (snd I) _ x-y∈I)) + + +[_]/ : {R : CommRing ℓ} {A : CommAlgebra R ℓ} {I : IdealsIn A} + → (a : fst A) → fst (A / I) +[ a ]/ = [ a ] diff --git a/Cubical/Algebra/CommMonoid.agda b/Cubical/Algebra/CommMonoid.agda new file mode 100644 index 0000000000..651f1e3cca --- /dev/null +++ b/Cubical/Algebra/CommMonoid.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommMonoid where + +open import Cubical.Algebra.CommMonoid.Base public +open import Cubical.Algebra.CommMonoid.Properties public diff --git a/Cubical/Algebra/CommMonoid/Base.agda b/Cubical/Algebra/CommMonoid/Base.agda new file mode 100644 index 0000000000..7bec68f71e --- /dev/null +++ b/Cubical/Algebra/CommMonoid/Base.agda @@ -0,0 +1,125 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommMonoid.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid.Base + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Reflection.RecordEquiv + +open Iso + +private + variable + ℓ ℓ' : Level + +record IsCommMonoid {A : Type ℓ} (ε : A) (_·_ : A → A → A) : Type ℓ where + constructor iscommmonoid + + field + isMonoid : IsMonoid ε _·_ + comm : (x y : A) → x · y ≡ y · x + + open IsMonoid isMonoid public + +unquoteDecl IsCommMonoidIsoΣ = declareRecordIsoΣ IsCommMonoidIsoΣ (quote IsCommMonoid) + +record CommMonoidStr (A : Type ℓ) : Type ℓ where + constructor commmonoidstr + + field + ε : A + _·_ : A → A → A + isCommMonoid : IsCommMonoid ε _·_ + + infixl 7 _·_ + + open IsCommMonoid isCommMonoid public + +CommMonoid : ∀ ℓ → Type (ℓ-suc ℓ) +CommMonoid ℓ = TypeWithStr ℓ CommMonoidStr + +commmonoid : (A : Type ℓ) (ε : A) (_·_ : A → A → A) (h : IsCommMonoid ε _·_) → CommMonoid ℓ +commmonoid A ε _·_ h = A , commmonoidstr ε _·_ h + +-- Easier to use constructors +makeIsCommMonoid : {M : Type ℓ} {ε : M} {_·_ : M → M → M} + (is-setM : isSet M) + (assoc : (x y z : M) → x · (y · z) ≡ (x · y) · z) + (rid : (x : M) → x · ε ≡ x) + (lid : (x : M) → ε · x ≡ x) + (comm : (x y : M) → x · y ≡ y · x) + → IsCommMonoid ε _·_ +IsCommMonoid.isMonoid (makeIsCommMonoid is-setM assoc rid lid comm) = + makeIsMonoid is-setM assoc rid lid +IsCommMonoid.comm (makeIsCommMonoid is-setM assoc rid lid comm) = comm + +makeCommMonoid : {M : Type ℓ} (ε : M) (_·_ : M → M → M) + (is-setM : isSet M) + (assoc : (x y z : M) → x · (y · z) ≡ (x · y) · z) + (rid : (x : M) → x · ε ≡ x) + (lid : (x : M) → ε · x ≡ x) + (comm : (x y : M) → x · y ≡ y · x) + → CommMonoid ℓ +makeCommMonoid ε _·_ is-setM assoc rid lid comm = + commmonoid _ ε _·_ (makeIsCommMonoid is-setM assoc rid lid comm) + + + +CommMonoidStr→MonoidStr : {A : Type ℓ} → CommMonoidStr A → MonoidStr A +CommMonoidStr→MonoidStr (commmonoidstr _ _ H) = monoidstr _ _ (IsCommMonoid.isMonoid H) + +CommMonoid→Monoid : CommMonoid ℓ → Monoid ℓ +CommMonoid→Monoid (_ , commmonoidstr _ _ H) = _ , monoidstr _ _ (IsCommMonoid.isMonoid H) + + +CommMonoidHom : (L : CommMonoid ℓ) (M : CommMonoid ℓ') → Type (ℓ-max ℓ ℓ') +CommMonoidHom L M = MonoidHom (CommMonoid→Monoid L) (CommMonoid→Monoid M) + +IsCommMonoidEquiv : {A : Type ℓ} {B : Type ℓ'} + (M : CommMonoidStr A) (e : A ≃ B) (N : CommMonoidStr B) → Type (ℓ-max ℓ ℓ') +IsCommMonoidEquiv M e N = IsMonoidHom (CommMonoidStr→MonoidStr M) (e .fst) (CommMonoidStr→MonoidStr N) + +CommMonoidEquiv : (M : CommMonoid ℓ) (N : CommMonoid ℓ') → Type (ℓ-max ℓ ℓ') +CommMonoidEquiv M N = Σ[ e ∈ (M .fst ≃ N .fst) ] IsCommMonoidEquiv (M .snd) e (N .snd) + +isPropIsCommMonoid : {M : Type ℓ} (ε : M) (_·_ : M → M → M) + → isProp (IsCommMonoid ε _·_) +isPropIsCommMonoid ε _·_ (iscommmonoid MM MC) (iscommmonoid SM SC) = + λ i → iscommmonoid (isPropIsMonoid _ _ MM SM i) (isPropComm MC SC i) + where + isSetM : isSet _ + isSetM = MM .IsMonoid.isSemigroup .IsSemigroup.is-set + + isPropComm : isProp ((x y : _) → x · y ≡ y · x) + isPropComm = isPropΠ2 λ _ _ → isSetM _ _ + +𝒮ᴰ-CommMonoid : DUARel (𝒮-Univ ℓ) CommMonoidStr ℓ +𝒮ᴰ-CommMonoid = + 𝒮ᴰ-Record (𝒮-Univ _) IsCommMonoidEquiv + (fields: + data[ ε ∣ autoDUARel _ _ ∣ presε ] + data[ _·_ ∣ autoDUARel _ _ ∣ isHom ] + prop[ isCommMonoid ∣ (λ _ _ → isPropIsCommMonoid _ _) ]) + where + open CommMonoidStr + open IsMonoidHom + +CommMonoidPath : (M N : CommMonoid ℓ) → CommMonoidEquiv M N ≃ (M ≡ N) +CommMonoidPath = ∫ 𝒮ᴰ-CommMonoid .UARel.ua diff --git a/Cubical/Algebra/CommMonoid/Properties.agda b/Cubical/Algebra/CommMonoid/Properties.agda new file mode 100644 index 0000000000..d85d441250 --- /dev/null +++ b/Cubical/Algebra/CommMonoid/Properties.agda @@ -0,0 +1,42 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommMonoid.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Auto +open import Cubical.Structures.Macro +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommMonoid.Base + +private + variable + ℓ : Level + +module CommMonoidTheory (M' : CommMonoid ℓ) where + open CommMonoidStr (snd M') + private M = ⟨ M' ⟩ + + commAssocl : (x y z : M) → x · (y · z) ≡ y · (x · z) + commAssocl x y z = assoc x y z ∙∙ cong (_· z) (comm x y) ∙∙ sym (assoc y x z) + + commAssocr : (x y z : M) → x · y · z ≡ x · z · y + commAssocr x y z = sym (assoc x y z) ∙∙ cong (x ·_) (comm y z) ∙∙ assoc x z y + + + commAssocr2 : (x y z : M) → x · y · z ≡ z · y · x + commAssocr2 x y z = commAssocr _ _ _ ∙∙ cong (_· y) (comm _ _) ∙∙ commAssocr _ _ _ + + commAssocSwap : (x y z w : M) → (x · y) · (z · w) ≡ (x · z) · (y · w) + commAssocSwap x y z w = assoc (x · y) z w ∙∙ cong (_· w) (commAssocr x y z) + ∙∙ sym (assoc (x · z) y w) diff --git a/Cubical/Algebra/CommRing.agda b/Cubical/Algebra/CommRing.agda new file mode 100644 index 0000000000..0f159f20ac --- /dev/null +++ b/Cubical/Algebra/CommRing.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing where + +open import Cubical.Algebra.CommRing.Base public +open import Cubical.Algebra.CommRing.Properties public diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda new file mode 100644 index 0000000000..8b973f9c44 --- /dev/null +++ b/Cubical/Algebra/CommRing/Base.agda @@ -0,0 +1,153 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Ring.Base + +open Iso + +private + variable + ℓ ℓ' : Level + +record IsCommRing {R : Type ℓ} + (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) : Type ℓ where + + constructor iscommring + + field + isRing : IsRing 0r 1r _+_ _·_ -_ + ·Comm : (x y : R) → x · y ≡ y · x + + open IsRing isRing public + +record CommRingStr (A : Type ℓ) : Type (ℓ-suc ℓ) where + + constructor commringstr + + field + 0r : A + 1r : A + _+_ : A → A → A + _·_ : A → A → A + -_ : A → A + isCommRing : IsCommRing 0r 1r _+_ _·_ -_ + + infix 8 -_ + infixl 7 _·_ + infixl 6 _+_ + + open IsCommRing isCommRing public + +CommRing : ∀ ℓ → Type (ℓ-suc ℓ) +CommRing ℓ = TypeWithStr ℓ CommRingStr + + +makeIsCommRing : {R : Type ℓ} {0r 1r : R} {_+_ _·_ : R → R → R} { -_ : R → R} + (is-setR : isSet R) + (+-assoc : (x y z : R) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : R) → x + 0r ≡ x) + (+-rinv : (x : R) → x + (- x) ≡ 0r) + (+-comm : (x y : R) → x + y ≡ y + x) + (·-assoc : (x y z : R) → x · (y · z) ≡ (x · y) · z) + (·-rid : (x : R) → x · 1r ≡ x) + (·-rdist-+ : (x y z : R) → x · (y + z) ≡ (x · y) + (x · z)) + (·-comm : (x y : R) → x · y ≡ y · x) + → IsCommRing 0r 1r _+_ _·_ -_ +makeIsCommRing {_+_ = _+_} is-setR +-assoc +-rid +-rinv +-comm ·-assoc ·-rid ·-rdist-+ ·-comm = + iscommring (makeIsRing is-setR +-assoc +-rid +-rinv +-comm ·-assoc ·-rid + (λ x → ·-comm _ _ ∙ ·-rid x) ·-rdist-+ + (λ x y z → ·-comm _ _ ∙∙ ·-rdist-+ z x y ∙∙ λ i → (·-comm z x i) + (·-comm z y i))) ·-comm + +makeCommRing : {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) + (is-setR : isSet R) + (+-assoc : (x y z : R) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : R) → x + 0r ≡ x) + (+-rinv : (x : R) → x + (- x) ≡ 0r) + (+-comm : (x y : R) → x + y ≡ y + x) + (·-assoc : (x y z : R) → x · (y · z) ≡ (x · y) · z) + (·-rid : (x : R) → x · 1r ≡ x) + (·-rdist-+ : (x y z : R) → x · (y + z) ≡ (x · y) + (x · z)) + (·-comm : (x y : R) → x · y ≡ y · x) + → CommRing ℓ +makeCommRing 0r 1r _+_ _·_ -_ is-setR +-assoc +-rid +-rinv +-comm ·-assoc ·-rid ·-rdist-+ ·-comm = + _ , commringstr _ _ _ _ _ (makeIsCommRing is-setR +-assoc +-rid +-rinv +-comm ·-assoc ·-rid ·-rdist-+ ·-comm) + +CommRingStr→RingStr : {A : Type ℓ} → CommRingStr A → RingStr A +CommRingStr→RingStr (commringstr _ _ _ _ _ H) = ringstr _ _ _ _ _ (IsCommRing.isRing H) + +CommRing→Ring : CommRing ℓ → Ring ℓ +CommRing→Ring (_ , commringstr _ _ _ _ _ H) = _ , ringstr _ _ _ _ _ (IsCommRing.isRing H) + +CommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') +CommRingHom R S = RingHom (CommRing→Ring R) (CommRing→Ring S) + +IsCommRingEquiv : {A : Type ℓ} {B : Type ℓ'} + (R : CommRingStr A) (e : A ≃ B) (S : CommRingStr B) → Type (ℓ-max ℓ ℓ') +IsCommRingEquiv R e S = IsRingHom (CommRingStr→RingStr R) (e .fst) (CommRingStr→RingStr S) + +CommRingEquiv : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') +CommRingEquiv R S = Σ[ e ∈ (R .fst ≃ S .fst) ] IsCommRingEquiv (R .snd) e (S .snd) + +CommRingEquiv→CommRingHom : {A : CommRing ℓ} {B : CommRing ℓ'} → CommRingEquiv A B → CommRingHom A B +CommRingEquiv→CommRingHom (e , eIsHom) = e .fst , eIsHom + +isPropIsCommRing : {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) + → isProp (IsCommRing 0r 1r _+_ _·_ -_) +isPropIsCommRing 0r 1r _+_ _·_ -_ (iscommring RR RC) (iscommring SR SC) = + λ i → iscommring (isPropIsRing _ _ _ _ _ RR SR i) + (isPropComm RC SC i) + where + isSetR : isSet _ + isSetR = RR .IsRing.·IsMonoid .IsMonoid.isSemigroup .IsSemigroup.is-set + + isPropComm : isProp ((x y : _) → x · y ≡ y · x) + isPropComm = isPropΠ2 λ _ _ → isSetR _ _ + +𝒮ᴰ-CommRing : DUARel (𝒮-Univ ℓ) CommRingStr ℓ +𝒮ᴰ-CommRing = + 𝒮ᴰ-Record (𝒮-Univ _) IsCommRingEquiv + (fields: + data[ 0r ∣ null ∣ pres0 ] + data[ 1r ∣ null ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ autoDUARel _ _ ∣ pres- ] + prop[ isCommRing ∣ (λ _ _ → isPropIsCommRing _ _ _ _ _) ]) + where + open CommRingStr + open IsRingHom + + -- faster with some sharing + null = autoDUARel (𝒮-Univ _) (λ A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +CommRingPath : (R S : CommRing ℓ) → CommRingEquiv R S ≃ (R ≡ S) +CommRingPath = ∫ 𝒮ᴰ-CommRing .UARel.ua + +uaCommRing : {A B : CommRing ℓ} → CommRingEquiv A B → A ≡ B +uaCommRing {A = A} {B = B} = equivFun (CommRingPath A B) + +isSetCommRing : ((R , str) : CommRing ℓ) → isSet R +isSetCommRing (R , str) = str .CommRingStr.is-set + +isGroupoidCommRing : isGroupoid (CommRing ℓ) +isGroupoidCommRing _ _ = isOfHLevelRespectEquiv 2 (CommRingPath _ _) (isSetRingEquiv _ _) diff --git a/Cubical/Algebra/CommRing/BinomialThm.agda b/Cubical/Algebra/CommRing/BinomialThm.agda new file mode 100644 index 0000000000..f8e49041d9 --- /dev/null +++ b/Cubical/Algebra/CommRing/BinomialThm.agda @@ -0,0 +1,132 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommRing.BinomialThm where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm + ; _choose_ to _ℕchoose_ ; snotz to ℕsnotz) +open import Cubical.Data.Nat.Order +open import Cubical.Data.FinData +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Algebra.Monoid.BigOp +open import Cubical.Algebra.Ring.BigOps +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.RingSolver.Reflection + +private + variable + ℓ : Level + +module BinomialThm (R' : CommRing ℓ) where + open CommRingStr (snd R') + open Exponentiation R' + open CommRingTheory R' + open Sum (CommRing→Ring R') + private R = fst R' + + _choose_ : ℕ → ℕ → R + n choose zero = 1r + zero choose suc k = 0r + suc n choose suc k = n choose (suc k) + n choose k + + n n≡m + + powersFormMultClosedSubset : (f : R) → isMultClosedSubset R' [ f ⁿ|n≥0] + powersFormMultClosedSubset f .containsOne = PT.∣ zero , refl ∣ + powersFormMultClosedSubset f .multClosed = + PT.map2 λ (m , p) (n , q) → (m +ℕ n) , (λ i → (p i) · (q i)) ∙ ·-of-^-is-^-of-+ f m n + + + R[1/_] : R → Type ℓ + R[1/ f ] = Loc.S⁻¹R R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + + + R[1/_]AsCommRing : R → CommRing ℓ + R[1/ f ]AsCommRing = Loc.S⁻¹RAsCommRing R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + + -- A useful lemma: (gⁿ/1)≡(g/1)ⁿ in R[1/f] + ^-respects-/1 : {f g : R} (n : ℕ) → [ (g ^ n) , 1r , PT.∣ 0 , (λ _ → 1r) ∣ ] ≡ + Exponentiation._^_ R[1/ f ]AsCommRing [ g , 1r , powersFormMultClosedSubset _ .containsOne ] n + ^-respects-/1 zero = refl + ^-respects-/1 {f} {g} (suc n) = eq/ _ _ ( (1r , powersFormMultClosedSubset f .containsOne) + , cong (1r · (g · (g ^ n)) ·_) (·Lid 1r)) + ∙ cong (CommRingStr._·_ (R[1/ f ]AsCommRing .snd) + [ g , 1r , powersFormMultClosedSubset f .containsOne ]) (^-respects-/1 n) + + -- A slight improvement for eliminating into propositions + invElPropElim : {f : R} {P : R[1/ f ] → Type ℓ'} + → (∀ x → isProp (P x)) + → (∀ (r : R) (n : ℕ) → P [ r , (f ^ n) , PT.∣ n , refl ∣ ]) -- ∀ r n → P (r/fⁿ) + ---------------------------------------------------------- + → ∀ x → P x + invElPropElim {f = f} {P = P} PisProp base = elimProp (λ _ → PisProp _) []-case + where + S[f] = Loc.S R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + []-case : (a : R × S[f]) → P [ a ] + []-case (r , s , s∈S[f]) = PT.rec (PisProp _) Σhelper s∈S[f] + where + Σhelper : Σ[ n ∈ ℕ ] s ≡ f ^ n → P [ r , s , s∈S[f] ] + Σhelper (n , p) = subst P (cong [_] (≡-× refl (Σ≡Prop (λ _ → isPropPropTrunc) (sym p)))) (base r n) + + invElPropElim2 : {f g : R} {P : R[1/ f ] → R[1/ g ] → Type ℓ'} + → (∀ x y → isProp (P x y)) + → (∀ (r s : R) (n : ℕ) → P [ r , (f ^ n) , PT.∣ n , refl ∣ ] + [ s , (g ^ n) , PT.∣ n , refl ∣ ]) + ---------------------------------------------------------- + → ∀ x y → P x y + invElPropElim2 {f = f} {g = g} {P = P} PisProp base = + invElPropElim (λ _ → isPropΠ (λ _ → PisProp _ _)) reduce1 + where + reduce1 : ∀ (r : R) (n : ℕ) (y : R[1/ g ]) → P [ r , f ^ n , ∣ n , refl ∣ ] y + reduce1 r n = invElPropElim (λ _ → PisProp _ _) reduce2 + where + reduce2 : (s : R) (m : ℕ) → P [ r , f ^ n , ∣ n , refl ∣ ] [ s , g ^ m , ∣ m , refl ∣ ] + reduce2 s m = subst2 P p q (base _ _ l) + where + l = max m n + x : R[1/ f ] + x = [ r , f ^ n , ∣ n , refl ∣ ] + y : R[1/ g ] + y = [ s , g ^ m , ∣ m , refl ∣ ] + x' : R[1/ f ] + x' = [ r · f ^ (l ∸ n) , f ^ l , ∣ l , refl ∣ ] + y' : R[1/ g ] + y' = [ s · g ^ (l ∸ m) , g ^ l , ∣ l , refl ∣ ] + + p : x' ≡ x + p = eq/ _ _ ((1r , ∣ 0 , refl ∣) , path) + where + useSolver1 : ∀ a b c → 1r · (a · b) · c ≡ a · (b · c) + useSolver1 = solve R' + useSolver2 : ∀ a b → a · b ≡ 1r · a · b + useSolver2 = solve R' + path : 1r · (r · f ^ (l ∸ n)) · f ^ n ≡ 1r · r · f ^ l + path = 1r · (r · f ^ (l ∸ n)) · f ^ n ≡⟨ useSolver1 _ _ _ ⟩ + r · (f ^ (l ∸ n) · f ^ n) ≡⟨ cong (r ·_) (·-of-^-is-^-of-+ _ _ _) ⟩ + r · f ^ (l ∸ n +ℕ n) ≡⟨ cong (λ k → r · f ^ k) (≤-∸-+-cancel right-≤-max) ⟩ + r · f ^ l ≡⟨ useSolver2 _ _ ⟩ + 1r · r · f ^ l ∎ + + q : y' ≡ y + q = eq/ _ _ ((1r , ∣ 0 , refl ∣) , path) + where + useSolver1 : ∀ a b c → 1r · (a · b) · c ≡ a · (b · c) + useSolver1 = solve R' + useSolver2 : ∀ a b → a · b ≡ 1r · a · b + useSolver2 = solve R' + path : 1r · (s · g ^ (l ∸ m)) · g ^ m ≡ 1r · s · g ^ l + path = 1r · (s · g ^ (l ∸ m)) · g ^ m ≡⟨ useSolver1 _ _ _ ⟩ + s · (g ^ (l ∸ m) · g ^ m) ≡⟨ cong (s ·_) (·-of-^-is-^-of-+ _ _ _) ⟩ + s · g ^ (l ∸ m +ℕ m) ≡⟨ cong (λ k → s · g ^ k) (≤-∸-+-cancel left-≤-max) ⟩ + s · g ^ l ≡⟨ useSolver2 _ _ ⟩ + 1r · s · g ^ l ∎ + + + -- For predicates over the set of powers + powersPropElim : {f : R} {P : R → Type ℓ'} + → (∀ x → isProp (P x)) + → (∀ n → P (f ^ n)) + ------------------------------ + → ∀ s → s ∈ [ f ⁿ|n≥0] → P s + powersPropElim {f = f} {P = P} PisProp base s = + PT.rec (PisProp s) λ (n , p) → subst P (sym p) (base n) + + +module RadicalLemma (R' : CommRing ℓ) (f g : (fst R')) where + open IsRingHom + open isMultClosedSubset + open CommRingTheory R' + open RingTheory (CommRing→Ring R') + open CommIdeal R' hiding (subst-∈) renaming (_∈_ to _∈ᵢ_) + open RadicalIdeal R' + open Exponentiation R' + open InvertingElementsBase R' + + open S⁻¹RUniversalProp R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + hiding (S⁻¹RHasUniversalProp) + open S⁻¹RUniversalProp R' [ g ⁿ|n≥0] (powersFormMultClosedSubset g) + hiding (_/1 ; /1AsCommRingHom) + open CommRingStr (R' .snd) + + private + R = R' .fst + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + + unitHelper : f ∈ᵢ √ ⟨ replicateFinVec 1 g ⟩ → (g /1) ∈ R[1/ f ]AsCommRing ˣ + unitHelper = PT.rec isPropGoal (uncurry ℕhelper) + where + isPropGoal = Units.inverseUniqueness _ (g /1) + + ℕhelper : (n : ℕ) → f ^ n ∈ᵢ ⟨ replicateFinVec 1 g ⟩ → (g /1) ∈ R[1/ f ]AsCommRing ˣ + ℕhelper n = PT.rec isPropGoal -- fⁿ≡αg → g⁻¹≡α/fⁿ + λ (α , p) → [ (α zero) , (f ^ n) , ∣ n , refl ∣ ] + , eq/ _ _ ((1r , powersFormMultClosedSubset f .containsOne) + , useSolver1 _ _ ∙ sym p ∙ useSolver2 _) + where + useSolver1 : ∀ x y → 1r · (x · y) · 1r ≡ y · x + 0r + useSolver1 = solve R' + + useSolver2 : ∀ x → x ≡ 1r · 1r · (1r · x) + useSolver2 = solve R' + + toUnit : f ∈ᵢ √ ⟨ replicateFinVec 1 g ⟩ + → ∀ s → s ∈ [ g ⁿ|n≥0] → (s /1) ∈ R[1/ f ]AsCommRing ˣ + toUnit f∈√⟨g⟩ = powersPropElim (λ x → Units.inverseUniqueness _ (x /1)) + λ n → subst-∈ (R[1/ f ]AsCommRing ˣ) (sym (^-respects-/1 n)) + (Exponentiation.^-presUnit _ _ n (unitHelper f∈√⟨g⟩)) + + +module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where + open isMultClosedSubset + open CommRingStr (snd R') + open CommRingTheory R' + open Exponentiation R' + open InvertingElementsBase + open RingTheory (CommRing→Ring R') + open CommRingStr (snd (R[1/_]AsCommRing R' f)) renaming ( _·_ to _·ᶠ_ ; 1r to 1ᶠ + ; _+_ to _+ᶠ_ ; 0r to 0ᶠ + ; ·Lid to ·ᶠ-lid ; ·Rid to ·ᶠ-rid + ; ·Assoc to ·ᶠ-assoc ; ·Comm to ·ᶠ-comm) + open IsRingHom + + private + R = fst R' + R[1/f] = R[1/_] R' f + R[1/f]AsCommRing = R[1/_]AsCommRing R' f + R[1/fg] = R[1/_] R' (f · g) + R[1/fg]AsCommRing = R[1/_]AsCommRing R' (f · g) + R[1/f][1/g] = R[1/_] (R[1/_]AsCommRing R' f) + [ g , 1r , powersFormMultClosedSubset R' f .containsOne ] + R[1/f][1/g]AsCommRing = R[1/_]AsCommRing (R[1/_]AsCommRing R' f) + [ g , 1r , powersFormMultClosedSubset R' f .containsOne ] + R[1/f][1/g]ˣ = R[1/f][1/g]AsCommRing ˣ + + + _/1/1 : R → R[1/f][1/g] + r /1/1 = [ [ r , 1r , PT.∣ 0 , refl ∣ ] , 1ᶠ , PT.∣ 0 , refl ∣ ] + + /1/1AsCommRingHom : CommRingHom R' R[1/f][1/g]AsCommRing + fst /1/1AsCommRingHom = _/1/1 + snd /1/1AsCommRingHom = makeIsRingHom refl lem+ lem· + where + lem+ : _ + lem+ r r' = + cong [_] + (≡-× + (cong [_] + (≡-× + (cong₂ _+_ + (sym (·Rid _) ∙ (λ i → (·Rid r (~ i)) · (·Rid 1r (~ i)))) + (sym (·Rid _) ∙ (λ i → (·Rid r' (~ i)) · (·Rid 1r (~ i))))) + (Σ≡Prop (λ _ → isPropPropTrunc) + (sym (·Lid _) ∙ (λ i → (·Lid 1r (~ i)) · (·Lid 1r (~ i))))))) + (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·ᶠ-lid 1ᶠ)))) + + lem· : _ + lem· r r' = + cong [_] + (≡-× + (cong [_] (≡-× refl (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·Lid _))))) + (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·ᶠ-lid 1ᶠ)))) + + -- this will give us a map R[1/fg] → R[1/f][1/g] by the universal property of localisation + fⁿgⁿ/1/1∈R[1/f][1/g]ˣ : (s : R) → s ∈ ([_ⁿ|n≥0] R' (f · g)) → s /1/1 ∈ R[1/f][1/g]ˣ + fⁿgⁿ/1/1∈R[1/f][1/g]ˣ = powersPropElim R' (λ s → R[1/f][1/g]ˣ (s /1/1) .snd) ℕcase + where + ℕcase : (n : ℕ) → ((f · g) ^ n) /1/1 ∈ R[1/f][1/g]ˣ + ℕcase n = [ [ 1r , (f ^ n) , PT.∣ n , refl ∣ ] + , [ (g ^ n) , 1r , PT.∣ 0 , refl ∣ ] --denominator + , PT.∣ n , ^-respects-/1 _ n ∣ ] + , eq/ _ _ ((1ᶠ , powersFormMultClosedSubset _ _ .containsOne) + , eq/ _ _ ((1r , powersFormMultClosedSubset _ _ .containsOne) , path)) + where + eq1 : ∀ x → 1r · (1r · (x · 1r) · 1r) · (1r · 1r · (1r · 1r)) ≡ x + eq1 = solve R' + + eq2 : ∀ x y → x · y ≡ 1r · (1r · 1r · (1r · y)) · (1r · (1r · x) · 1r) + eq2 = solve R' + + path : 1r · (1r · ((f · g) ^ n · 1r) · 1r) · (1r · 1r · (1r · 1r)) + ≡ 1r · (1r · 1r · (1r · g ^ n)) · (1r · (1r · f ^ n) · 1r) + path = 1r · (1r · ((f · g) ^ n · 1r) · 1r) · (1r · 1r · (1r · 1r)) ≡⟨ eq1 ((f · g) ^ n) ⟩ + (f · g) ^ n ≡⟨ ^-ldist-· _ _ _ ⟩ + f ^ n · g ^ n ≡⟨ eq2 (f ^ n) (g ^ n) ⟩ + 1r · (1r · 1r · (1r · g ^ n)) · (1r · (1r · f ^ n) · 1r) ∎ + + + -- the main result: localising at one element and then at another is + -- the same as localising at the product. + -- takes forever to compute without experimental lossy unification + R[1/fg]≡R[1/f][1/g] : R[1/fg]AsCommRing ≡ R[1/f][1/g]AsCommRing + R[1/fg]≡R[1/f][1/g] = S⁻¹RChar R' ([_ⁿ|n≥0] R' (f · g)) + (powersFormMultClosedSubset R' (f · g)) _ /1/1AsCommRingHom pathtoR[1/fg] + where + open PathToS⁻¹R + pathtoR[1/fg] : PathToS⁻¹R R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + R[1/f][1/g]AsCommRing /1/1AsCommRingHom + φS⊆Aˣ pathtoR[1/fg] = fⁿgⁿ/1/1∈R[1/f][1/g]ˣ + + kerφ⊆annS pathtoR[1/fg] r p = toGoal helperR[1/f] + where + open RingTheory (CommRing→Ring R[1/f]AsCommRing) renaming ( 0RightAnnihilates to 0ᶠRightAnnihilates + ; 0LeftAnnihilates to 0ᶠ-leftNullifies) + open Exponentiation R[1/f]AsCommRing renaming (_^_ to _^ᶠ_) + hiding (·-of-^-is-^-of-+ ; ^-ldist-·) + + S[f] = Loc.S R' ([_ⁿ|n≥0] R' f) (powersFormMultClosedSubset R' f) + S[fg] = Loc.S R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + g/1 : R[1/_] R' f + g/1 = [ g , 1r , powersFormMultClosedSubset R' f .containsOne ] + S[g/1] = Loc.S R[1/f]AsCommRing + ([_ⁿ|n≥0] R[1/f]AsCommRing g/1) + (powersFormMultClosedSubset R[1/f]AsCommRing g/1) + r/1 : R[1/_] R' f + r/1 = [ r , 1r , powersFormMultClosedSubset R' f .containsOne ] + + -- this is the crucial step, modulo truncation we can take p to be generated + -- by the quotienting relation of localisation. Note that we wouldn't be able + -- to prove our goal if kerφ⊆annS was formulated with a Σ instead of a ∃ + ∥r/1,1/1≈0/1,1/1∥ : ∃[ u ∈ S[g/1] ] fst u ·ᶠ r/1 ·ᶠ 1ᶠ ≡ fst u ·ᶠ 0ᶠ ·ᶠ 1ᶠ + ∥r/1,1/1≈0/1,1/1∥ = Iso.fun (SQ.isEquivRel→TruncIso (Loc.locIsEquivRel _ _ _) _ _) p + + helperR[1/f] : ∃[ n ∈ ℕ ] [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + helperR[1/f] = PT.rec isPropPropTrunc + (uncurry (uncurry (powersPropElim R[1/f]AsCommRing + (λ _ → isPropΠ (λ _ → isPropPropTrunc)) baseCase))) + ∥r/1,1/1≈0/1,1/1∥ + where + baseCase : ∀ n → g/1 ^ᶠ n ·ᶠ r/1 ·ᶠ 1ᶠ ≡ g/1 ^ᶠ n ·ᶠ 0ᶠ ·ᶠ 1ᶠ + → ∃[ n ∈ ℕ ] [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + baseCase n q = PT.∣ n , path ∣ + where + path : [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + path = [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] + + ≡⟨ cong [_] (≡-× refl (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·Rid _)))) ⟩ + + [ g ^ n , 1r , PT.∣ 0 , refl ∣ ] ·ᶠ r/1 + + ≡⟨ cong (_·ᶠ r/1) (^-respects-/1 _ n) ⟩ + + g/1 ^ᶠ n ·ᶠ r/1 + + ≡⟨ sym (·ᶠ-rid _) ⟩ + + g/1 ^ᶠ n ·ᶠ r/1 ·ᶠ 1ᶠ + + ≡⟨ q ⟩ + + g/1 ^ᶠ n ·ᶠ 0ᶠ ·ᶠ 1ᶠ + + ≡⟨ cong (_·ᶠ 1ᶠ) (0ᶠRightAnnihilates _) ∙ 0ᶠ-leftNullifies 1ᶠ ⟩ + + 0ᶠ ∎ + + + toGoal : ∃[ n ∈ ℕ ] [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + → ∃[ u ∈ S[fg] ] fst u · r ≡ 0r + toGoal = PT.rec isPropPropTrunc Σhelper + where + Σhelper : Σ[ n ∈ ℕ ] [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + → ∃[ u ∈ S[fg] ] fst u · r ≡ 0r + Σhelper (n , q) = PT.map Σhelper2 helperR + where + -- now, repeat the above strategy with q + ∥gⁿr≈0∥ : ∃[ u ∈ S[f] ] fst u · (g ^ n · r) · 1r ≡ fst u · 0r · 1r + ∥gⁿr≈0∥ = Iso.fun (SQ.isEquivRel→TruncIso (Loc.locIsEquivRel _ _ _) _ _) q + + helperR : ∃[ m ∈ ℕ ] f ^ m · g ^ n · r ≡ 0r + helperR = PT.rec isPropPropTrunc + (uncurry (uncurry (powersPropElim R' + (λ _ → isPropΠ (λ _ → isPropPropTrunc)) baseCase))) + ∥gⁿr≈0∥ + where + baseCase : (m : ℕ) → f ^ m · (g ^ n · r) · 1r ≡ f ^ m · 0r · 1r + → ∃[ m ∈ ℕ ] f ^ m · g ^ n · r ≡ 0r + baseCase m q' = PT.∣ m , path ∣ + where + path : f ^ m · g ^ n · r ≡ 0r + path = (λ i → ·Rid (·Assoc (f ^ m) (g ^ n) r (~ i)) (~ i)) + ∙∙ q' ∙∙ (λ i → ·Rid (0RightAnnihilates (f ^ m) i) i) + + Σhelper2 : Σ[ m ∈ ℕ ] f ^ m · g ^ n · r ≡ 0r + → Σ[ u ∈ S[fg] ] fst u · r ≡ 0r + Σhelper2 (m , q') = (((f · g) ^ l) , PT.∣ l , refl ∣) , path + where + l = max m n + + path : (f · g) ^ l · r ≡ 0r + path = (f · g) ^ l · r + + ≡⟨ cong (_· r) (^-ldist-· _ _ _) ⟩ + + f ^ l · g ^ l · r + + ≡⟨ cong₂ (λ x y → f ^ x · g ^ y · r) (sym (≤-∸-+-cancel {m = m} left-≤-max)) + (sym (≤-∸-+-cancel {m = n} right-≤-max)) ⟩ + + f ^ (l ∸ m +ℕ m) · g ^ (l ∸ n +ℕ n) · r + + ≡⟨ cong₂ (λ x y → x · y · r) (sym (·-of-^-is-^-of-+ _ _ _)) + (sym (·-of-^-is-^-of-+ _ _ _)) ⟩ + + f ^ (l ∸ m) · f ^ m · (g ^ (l ∸ n) · g ^ n) · r + + ≡⟨ cong (_· r) (·CommAssocSwap _ _ _ _) ⟩ + + f ^ (l ∸ m) · g ^ (l ∸ n) · (f ^ m · g ^ n) · r + + ≡⟨ sym (·Assoc _ _ _) ⟩ + + f ^ (l ∸ m) · g ^ (l ∸ n) · (f ^ m · g ^ n · r) + + ≡⟨ cong (f ^ (l ∸ m) · g ^ (l ∸ n) ·_) q' ⟩ + + f ^ (l ∸ m) · g ^ (l ∸ n) · 0r + + ≡⟨ 0RightAnnihilates _ ⟩ + + 0r ∎ + + + surχ pathtoR[1/fg] = invElPropElim _ (λ _ → isPropPropTrunc) toGoal + where + open Exponentiation R[1/f]AsCommRing renaming (_^_ to _^ᶠ_) + hiding (·-of-^-is-^-of-+ ; ^-ldist-·) + open CommRingStr (snd R[1/f][1/g]AsCommRing) renaming (_·_ to _·R[1/f][1/g]_) + hiding (1r ; ·Lid ; ·Rid ; ·Assoc) + open Units R[1/f][1/g]AsCommRing + g/1 : R[1/_] R' f + g/1 = [ g , 1r , powersFormMultClosedSubset R' f .containsOne ] + S[fg] = Loc.S R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + + baseCase : (r : R) (m n : ℕ) → ∃[ x ∈ R × S[fg] ] (x .fst /1/1) + ≡ [ [ r , f ^ m , PT.∣ m , refl ∣ ] + , [ g ^ n , 1r , PT.∣ 0 , refl ∣ ] , PT.∣ n , ^-respects-/1 _ n ∣ ] + ·R[1/f][1/g] (x .snd .fst /1/1) + baseCase r m n = PT.∣ ((r · f ^ (l ∸ m) · g ^ (l ∸ n)) -- x .fst + , (f · g) ^ l , PT.∣ l , refl ∣) -- x .snd + , eq/ _ _ ((1ᶠ , PT.∣ 0 , refl ∣) , eq/ _ _ ((1r , PT.∣ 0 , refl ∣) , path)) ∣ + -- reduce equality of double fractions into equality in R + where + eq1 : ∀ r flm gln gn fm + → 1r · (1r · (r · flm · gln) · (gn · 1r)) · (1r · (fm · 1r) · 1r) + ≡ r · flm · (gln · gn) · fm + eq1 = solve R' + + eq2 : ∀ r flm gl fm → r · flm · gl · fm ≡ r · (flm · fm ) · gl + eq2 = solve R' + + eq3 : ∀ r fgl → r · fgl ≡ 1r · (1r · (r · fgl) · 1r) · (1r · 1r · (1r · 1r)) + eq3 = solve R' + + l = max m n + + path : 1r · (1r · (r · f ^ (l ∸ m) · g ^ (l ∸ n)) · (g ^ n · 1r)) · (1r · (f ^ m · 1r) · 1r) + ≡ 1r · (1r · (r · (f · g) ^ l) · 1r) · (1r · 1r · (1r · 1r)) + path = 1r · (1r · (r · f ^ (l ∸ m) · g ^ (l ∸ n)) · (g ^ n · 1r)) · (1r · (f ^ m · 1r) · 1r) + + ≡⟨ eq1 r (f ^ (l ∸ m)) (g ^ (l ∸ n)) (g ^ n) (f ^ m) ⟩ + + r · f ^ (l ∸ m) · (g ^ (l ∸ n) · g ^ n) · f ^ m + + ≡⟨ cong (λ x → r · f ^ (l ∸ m) · x · f ^ m) (·-of-^-is-^-of-+ _ _ _) ⟩ + + r · f ^ (l ∸ m) · g ^ (l ∸ n +ℕ n) · f ^ m + + ≡⟨ cong (λ x → r · f ^ (l ∸ m) · g ^ x · f ^ m) (≤-∸-+-cancel right-≤-max) ⟩ + + r · f ^ (l ∸ m) · g ^ l · f ^ m + + ≡⟨ eq2 r (f ^ (l ∸ m)) (g ^ l) (f ^ m) ⟩ + + r · (f ^ (l ∸ m) · f ^ m) · g ^ l + + ≡⟨ cong (λ x → r · x · g ^ l) (·-of-^-is-^-of-+ _ _ _) ⟩ + + r · f ^ (l ∸ m +ℕ m) · g ^ l + + ≡⟨ cong (λ x → r · f ^ x · g ^ l) (≤-∸-+-cancel left-≤-max) ⟩ + + r · f ^ l · g ^ l + + ≡⟨ sym (·Assoc _ _ _) ⟩ + + r · (f ^ l · g ^ l) + + ≡⟨ cong (r ·_) (sym (^-ldist-· _ _ _)) ⟩ + + r · (f · g) ^ l + + ≡⟨ eq3 r ((f · g) ^ l) ⟩ + + 1r · (1r · (r · (f · g) ^ l) · 1r) · (1r · 1r · (1r · 1r)) ∎ + + + base-^ᶠ-helper : (r : R) (m n : ℕ) → ∃[ x ∈ R × S[fg] ] (x .fst /1/1) + ≡ [ [ r , f ^ m , PT.∣ m , refl ∣ ] + , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] ·R[1/f][1/g] (x .snd .fst /1/1) + base-^ᶠ-helper r m n = subst (λ y → ∃[ x ∈ R × S[fg] ] (x .fst /1/1) + ≡ [ [ r , f ^ m , PT.∣ m , refl ∣ ] , y ] ·R[1/f][1/g] (x .snd .fst /1/1)) + (Σ≡Prop (λ _ → isPropPropTrunc) (^-respects-/1 _ n)) (baseCase r m n) + + indStep : (r : R[1/_] R' f) (n : ℕ) → ∃[ x ∈ R × S[fg] ] + (x .fst /1/1) ≡ [ r , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] ·R[1/f][1/g] (x .snd .fst /1/1) + indStep = invElPropElim _ (λ _ → isPropΠ λ _ → isPropPropTrunc) base-^ᶠ-helper + + toGoal : (r : R[1/_] R' f) (n : ℕ) → ∃[ x ∈ R × S[fg] ] + (x .fst /1/1) ·R[1/f][1/g] + ((x .snd .fst /1/1) ⁻¹) ⦃ φS⊆Aˣ pathtoR[1/fg] (x .snd .fst) (x .snd .snd) ⦄ + ≡ [ r , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] + toGoal r n = PT.map Σhelper (indStep r n) + where + Σhelper : Σ[ x ∈ R × S[fg] ] + (x .fst /1/1) ≡ [ r , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] ·R[1/f][1/g] (x .snd .fst /1/1) + → Σ[ x ∈ R × S[fg] ] + (x .fst /1/1) ·R[1/f][1/g] ((x .snd .fst /1/1) ⁻¹) + ⦃ φS⊆Aˣ pathtoR[1/fg] (x .snd .fst) (x .snd .snd) ⦄ + ≡ [ r , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] + Σhelper ((r' , s , s∈S[fg]) , p) = (r' , s , s∈S[fg]) + , ⁻¹-eq-elim ⦃ φS⊆Aˣ pathtoR[1/fg] s s∈S[fg] ⦄ p + + + -- In this module we construct the map R[1/fg]→R[1/f][1/g] directly + -- and show that it is equal (although not judgementally) to the map induced + -- by the universal property of localisation, i.e. transporting along the path + -- constructed above. Given that this is the easier direction, we can see that + -- it is pretty cumbersome to prove R[1/fg]≡R[1/f][1/g] directly, + -- which illustrates the usefulness of S⁻¹RChar quite well. + private + module check where + φ : R[1/fg] → R[1/f][1/g] + φ = SQ.rec squash/ ϕ ϕcoh + where + S[fg] = Loc.S R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + + curriedϕΣ : (r s : R) → Σ[ n ∈ ℕ ] s ≡ (f · g) ^ n → R[1/f][1/g] + curriedϕΣ r s (n , s≡fg^n) = [ [ r , (f ^ n) , PT.∣ n , refl ∣ ] + , [ (g ^ n) , 1r , PT.∣ 0 , refl ∣ ] --denominator + , PT.∣ n , ^-respects-/1 R' n ∣ ] + + curriedϕ : (r s : R) → ∃[ n ∈ ℕ ] s ≡ (f · g) ^ n → R[1/f][1/g] + curriedϕ r s = elim→Set (λ _ → squash/) (curriedϕΣ r s) coh + where + coh : (x y : Σ[ n ∈ ℕ ] s ≡ (f · g) ^ n) → curriedϕΣ r s x ≡ curriedϕΣ r s y + coh (n , s≡fg^n) (m , s≡fg^m) = eq/ _ _ ((1ᶠ , PT.∣ 0 , refl ∣) , + eq/ _ _ ( (1r , powersFormMultClosedSubset R' f .containsOne) + , path)) + where + eq1 : ∀ r gm fm → 1r · (1r · r · gm) · (1r · fm · 1r) ≡ r · (gm · fm) + eq1 = solve R' + + path : 1r · (1r · r · (g ^ m)) · (1r · (f ^ m) · 1r) + ≡ 1r · (1r · r · (g ^ n)) · (1r · (f ^ n) · 1r) + path = 1r · (1r · r · (g ^ m)) · (1r · (f ^ m) · 1r) + + ≡⟨ eq1 r (g ^ m) (f ^ m) ⟩ + + r · (g ^ m · f ^ m) + + ≡⟨ cong (r ·_) (sym (^-ldist-· g f m)) ⟩ + + r · ((g · f) ^ m) + + ≡⟨ cong (λ x → r · (x ^ m)) (·Comm _ _) ⟩ + + r · ((f · g) ^ m) + + ≡⟨ cong (r ·_) ((sym s≡fg^m) ∙ s≡fg^n) ⟩ + + r · ((f · g) ^ n) + + ≡⟨ cong (λ x → r · (x ^ n)) (·Comm _ _) ⟩ + + r · ((g · f) ^ n) + + ≡⟨ cong (r ·_) (^-ldist-· g f n) ⟩ + + r · (g ^ n · f ^ n) + + ≡⟨ sym (eq1 r (g ^ n) (f ^ n)) ⟩ + + 1r · (1r · r · (g ^ n)) · (1r · (f ^ n) · 1r) ∎ + + + ϕ : R × S[fg] → R[1/f][1/g] + ϕ = uncurry2 curriedϕ -- λ (r / (fg)ⁿ) → ((r / fⁿ) / gⁿ) + + curriedϕcohΣ : (r s r' s' u : R) → (p : u · r · s' ≡ u · r' · s) + → (α : Σ[ n ∈ ℕ ] s ≡ (f · g) ^ n) + → (β : Σ[ m ∈ ℕ ] s' ≡ (f · g) ^ m) + → (γ : Σ[ l ∈ ℕ ] u ≡ (f · g) ^ l) + → ϕ (r , s , PT.∣ α ∣) ≡ ϕ (r' , s' , PT.∣ β ∣) + curriedϕcohΣ r s r' s' u p (n , s≡fgⁿ) (m , s'≡fgᵐ) (l , u≡fgˡ) = + eq/ _ _ ( ( [ (g ^ l) , 1r , powersFormMultClosedSubset R' f .containsOne ] + , PT.∣ l , ^-respects-/1 R' l ∣) + , eq/ _ _ ((f ^ l , PT.∣ l , refl ∣) , path)) + where + eq1 : ∀ fl gl r gm fm + → fl · (gl · r · gm) · (1r · fm · 1r) ≡ fl · gl · r · (gm · fm) + eq1 = solve R' + + path : f ^ l · (g ^ l · transp (λ i → R) i0 r · transp (λ i → R) i0 (g ^ m)) + · (1r · transp (λ i → R) i0 (f ^ m) · transp (λ i → R) i0 1r) + ≡ f ^ l · (g ^ l · transp (λ i → R) i0 r' · transp (λ i → R) i0 (g ^ n)) + · (1r · transp (λ i → R) i0 (f ^ n) · transp (λ i → R) i0 1r) + path = f ^ l · (g ^ l · transp (λ i → R) i0 r · transp (λ i → R) i0 (g ^ m)) + · (1r · transp (λ i → R) i0 (f ^ m) · transp (λ i → R) i0 1r) + + ≡⟨ (λ i → f ^ l · (g ^ l · transportRefl r i · transportRefl (g ^ m) i) + · (1r · transportRefl (f ^ m) i · transportRefl 1r i)) ⟩ + + f ^ l · (g ^ l · r · g ^ m) · (1r · f ^ m · 1r) + + ≡⟨ eq1 (f ^ l) (g ^ l) r (g ^ m) (f ^ m) ⟩ + + f ^ l · g ^ l · r · (g ^ m · f ^ m) + + ≡⟨ (λ i → ^-ldist-· f g l (~ i) · r · ^-ldist-· g f m (~ i)) ⟩ + + (f · g) ^ l · r · (g · f) ^ m + + ≡⟨ cong (λ x → (f · g) ^ l · r · x ^ m) (·Comm _ _) ⟩ + + (f · g) ^ l · r · (f · g) ^ m + + ≡⟨ (λ i → u≡fgˡ (~ i) · r · s'≡fgᵐ (~ i)) ⟩ + + u · r · s' + + ≡⟨ p ⟩ + + u · r' · s + + ≡⟨ (λ i → u≡fgˡ i · r' · s≡fgⁿ i) ⟩ + + (f · g) ^ l · r' · (f · g) ^ n + + ≡⟨ cong (λ x → (f · g) ^ l · r' · x ^ n) (·Comm _ _) ⟩ + + (f · g) ^ l · r' · (g · f) ^ n + + ≡⟨ (λ i → ^-ldist-· f g l i · r' · ^-ldist-· g f n i) ⟩ + + f ^ l · g ^ l · r' · (g ^ n · f ^ n) + + ≡⟨ sym (eq1 (f ^ l) (g ^ l) r' (g ^ n) (f ^ n)) ⟩ + + f ^ l · (g ^ l · r' · g ^ n) · (1r · f ^ n · 1r) + + ≡⟨ (λ i → f ^ l · (g ^ l · transportRefl r' (~ i) · transportRefl (g ^ n) (~ i)) + · (1r · transportRefl (f ^ n) (~ i) · transportRefl 1r (~ i))) ⟩ + + f ^ l · (g ^ l · transp (λ i → R) i0 r' · transp (λ i → R) i0 (g ^ n)) + · (1r · transp (λ i → R) i0 (f ^ n) · transp (λ i → R) i0 1r) ∎ + + + curriedϕcoh : (r s r' s' u : R) → (p : u · r · s' ≡ u · r' · s) + → (α : ∃[ n ∈ ℕ ] s ≡ (f · g) ^ n) + → (β : ∃[ m ∈ ℕ ] s' ≡ (f · g) ^ m) + → (γ : ∃[ l ∈ ℕ ] u ≡ (f · g) ^ l) + → ϕ (r , s , α) ≡ ϕ (r' , s' , β) + curriedϕcoh r s r' s' u p = PT.elim (λ _ → isPropΠ2 (λ _ _ → squash/ _ _)) + λ α → PT.elim (λ _ → isPropΠ (λ _ → squash/ _ _)) + λ β → PT.rec (squash/ _ _) + λ γ → curriedϕcohΣ r s r' s' u p α β γ + + ϕcoh : (a b : R × S[fg]) + → Loc._≈_ R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) a b + → ϕ a ≡ ϕ b + ϕcoh (r , s , α) (r' , s' , β) ((u , γ) , p) = curriedϕcoh r s r' s' u p α β γ + + + + + -- the map induced by the universal property + open S⁻¹RUniversalProp R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + χ : R[1/fg] → R[1/f][1/g] + χ = S⁻¹RHasUniversalProp R[1/f][1/g]AsCommRing /1/1AsCommRingHom fⁿgⁿ/1/1∈R[1/f][1/g]ˣ .fst .fst .fst + + -- the sanity check: + -- both maps send a fraction r/(fg)ⁿ to a double fraction, + -- where numerator and denominator are almost the same fraction respectively. + -- unfortunately the proofs that the denominators are powers are quite different for + -- the two maps, but of course we can ignore them. + -- The definition of χ introduces a lot of (1r ·_). Perhaps most surprisingly, + -- we have to give the path eq1 for the equality of the numerator of the numerator. + φ≡χ : ∀ r → φ r ≡ χ r + φ≡χ = invElPropElim _ (λ _ → squash/ _ _) ℕcase + where + ℕcase : (r : R) (n : ℕ) + → φ [ r , (f · g) ^ n , PT.∣ n , refl ∣ ] ≡ χ [ r , (f · g) ^ n , PT.∣ n , refl ∣ ] + ℕcase r n = cong [_] (ΣPathP --look into the components of the double-fractions + ( cong [_] (ΣPathP (eq1 , Σ≡Prop (λ x → S'[f] x .snd) (sym (·Lid _)))) + , Σ≡Prop (λ x → S'[f][g] x .snd) --ignore proof that denominator is power of g/1 + ( cong [_] (ΣPathP (sym (·Lid _) , Σ≡Prop (λ x → S'[f] x .snd) (sym (·Lid _))))))) + where + S'[f] = ([_ⁿ|n≥0] R' f) + S'[f][g] = ([_ⁿ|n≥0] R[1/f]AsCommRing [ g , 1r , powersFormMultClosedSubset R' f .containsOne ]) + + eq1 : transp (λ i → fst R') i0 r ≡ r · transp (λ i → fst R') i0 1r + eq1 = transportRefl r ∙∙ sym (·Rid r) ∙∙ cong (r ·_) (sym (transportRefl 1r)) diff --git a/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda b/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda new file mode 100644 index 0000000000..d41f692f03 --- /dev/null +++ b/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda @@ -0,0 +1,586 @@ +{- + + This file contains a proof of the following fact: + for a commutative ring R with elements f and g s.t. + ⟨f,g⟩ = R, we get a get a pullback square + + _/1 + R ----> R[1/f] + ⌟ + _/1 | | χ₁ + v v + + R[1/g] -> R[1/fg] + χ₂ + + where the morphisms χ are induced by the universal property + of localization. + + -} + + +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommRing.Localisation.PullbackSquare where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset renaming (_∈_ to _∈ₚ_) +open import Cubical.Foundations.Transport +open import Cubical.Functions.FunExtEquiv + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Vec +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Algebra.Group +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Localisation.Base +open import Cubical.Algebra.CommRing.Localisation.UniversalProperty +open import Cubical.Algebra.CommRing.Localisation.InvertingElements +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.RadicalIdeal + +open import Cubical.Algebra.RingSolver.Reflection + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Instances.CommRings +open import Cubical.Categories.Limits.Pullback + +open Iso + +private + variable + ℓ ℓ' : Level + A : Type ℓ + + +module _ (R' : CommRing ℓ) (f g : (fst R')) where + + open IsRingHom + open isMultClosedSubset + open CommRingTheory R' + open RingTheory (CommRing→Ring R') + open CommIdeal R' + open RadicalIdeal R' + open Exponentiation R' + open InvertingElementsBase R' + + open CommRingStr ⦃...⦄ + private + R = R' .fst + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + fgVec : FinVec R 2 + fgVec zero = f + fgVec (suc zero) = g + ⟨f,g⟩ = ⟨ fgVec ⟩ + fⁿgⁿVec : (n : ℕ) → FinVec R 2 + fⁿgⁿVec n zero = f ^ n + fⁿgⁿVec n (suc zero) = g ^ n + ⟨fⁿ,gⁿ⟩ : (n : ℕ) → CommIdeal + ⟨fⁿ,gⁿ⟩ n = ⟨ (fⁿgⁿVec n) ⟩ + + instance + _ = R' .snd + _ = R[1/ f ]AsCommRing .snd + _ = R[1/ g ]AsCommRing .snd + _ = R[1/ (R' .snd .CommRingStr._·_ f g) ]AsCommRing .snd + + open Loc R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + renaming (_≈_ to _≈ᶠ_ ; locIsEquivRel to locIsEquivRelᶠ ; S to Sᶠ) + open S⁻¹RUniversalProp R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + renaming ( _/1 to _/1ᶠ ; /1AsCommRingHom to /1ᶠAsCommRingHom + ; S⁻¹RHasUniversalProp to R[1/f]HasUniversalProp) + open Loc R' [ g ⁿ|n≥0] (powersFormMultClosedSubset g) + renaming (_≈_ to _≈ᵍ_ ; locIsEquivRel to locIsEquivRelᵍ ; S to Sᵍ) + open S⁻¹RUniversalProp R' [ g ⁿ|n≥0] (powersFormMultClosedSubset g) + renaming ( _/1 to _/1ᵍ ; /1AsCommRingHom to /1ᵍAsCommRingHom + ; S⁻¹RHasUniversalProp to R[1/g]HasUniversalProp) + open Loc R' [ (f · g) ⁿ|n≥0] (powersFormMultClosedSubset (f · g)) + renaming (_≈_ to _≈ᶠᵍ_ ; locIsEquivRel to locIsEquivRelᶠᵍ ; S to Sᶠᵍ) + open S⁻¹RUniversalProp R' [ (f · g) ⁿ|n≥0] (powersFormMultClosedSubset (f · g)) + renaming ( _/1 to _/1ᶠᵍ ; /1AsCommRingHom to /1ᶠᵍAsCommRingHom) + + + -- the pullback legs + private + -- using RadicalLemma doesn't compute... + χ₁ : CommRingHom R[1/ f ]AsCommRing R[1/ (f · g) ]AsCommRing + χ₁ = R[1/f]HasUniversalProp _ /1ᶠᵍAsCommRingHom unitHelper .fst .fst + where + unitHelper : ∀ s → s ∈ₚ [ f ⁿ|n≥0] → s /1ᶠᵍ ∈ₚ (R[1/ (f · g) ]AsCommRing) ˣ + unitHelper = powersPropElim (λ s → Units.inverseUniqueness _ (s /1ᶠᵍ)) + λ n → [ g ^ n , (f · g) ^ n , ∣ n , refl ∣ ] + , eq/ _ _ ((1r , powersFormMultClosedSubset (f · g) .containsOne) + , path n) + where + useSolver1 : ∀ a b → 1r · (a · b) · 1r ≡ a · b + useSolver1 = solve R' + useSolver2 : ∀ a → a ≡ (1r · 1r) · (1r · a) + useSolver2 = solve R' + + path : (n : ℕ) → 1r · (f ^ n · g ^ n) · 1r ≡ (1r · 1r) · (1r · ((f · g) ^ n)) + path n = useSolver1 _ _ ∙ sym (^-ldist-· f g n) ∙ useSolver2 _ + + χ₂ : CommRingHom R[1/ g ]AsCommRing R[1/ (f · g) ]AsCommRing + χ₂ = R[1/g]HasUniversalProp _ /1ᶠᵍAsCommRingHom unitHelper .fst .fst + where + unitHelper : ∀ s → s ∈ₚ [ g ⁿ|n≥0] → s /1ᶠᵍ ∈ₚ (R[1/ (f · g) ]AsCommRing) ˣ + unitHelper = powersPropElim (λ s → Units.inverseUniqueness _ (s /1ᶠᵍ)) + λ n → [ f ^ n , (f · g) ^ n , ∣ n , refl ∣ ] + , eq/ _ _ ((1r , powersFormMultClosedSubset (f · g) .containsOne) + , path n) + where + useSolver1 : ∀ a b → 1r · (a · b) · 1r ≡ b · a + useSolver1 = solve R' + useSolver2 : ∀ a → a ≡ (1r · 1r) · (1r · a) + useSolver2 = solve R' + + path : (n : ℕ) → 1r · (g ^ n · f ^ n) · 1r ≡ (1r · 1r) · (1r · ((f · g) ^ n)) + path n = useSolver1 _ _ ∙ sym (^-ldist-· f g n) ∙ useSolver2 _ + + + + injectivityLemma : 1r ∈ ⟨f,g⟩ → ∀ (x : R) → x /1ᶠ ≡ 0r → x /1ᵍ ≡ 0r → x ≡ 0r + injectivityLemma 1∈⟨f,g⟩ x x≡0overF x≡0overG = + PT.rec2 (is-set _ _) annihilatorHelper exAnnihilatorF exAnnihilatorG + where + exAnnihilatorF : ∃[ s ∈ Sᶠ ] (fst s · x · 1r ≡ fst s · 0r · 1r) + exAnnihilatorF = isEquivRel→TruncIso locIsEquivRelᶠ _ _ .fun x≡0overF + + exAnnihilatorG : ∃[ s ∈ Sᵍ ] (fst s · x · 1r ≡ fst s · 0r · 1r) + exAnnihilatorG = isEquivRel→TruncIso locIsEquivRelᵍ _ _ .fun x≡0overG + + annihilatorHelper : Σ[ s ∈ Sᶠ ] (fst s · x · 1r ≡ fst s · 0r · 1r) + → Σ[ s ∈ Sᵍ ] (fst s · x · 1r ≡ fst s · 0r · 1r) + → x ≡ 0r + annihilatorHelper ((u , u∈[fⁿ]) , p) ((v , v∈[gⁿ]) , q) = + PT.rec2 (is-set _ _) exponentHelper u∈[fⁿ] v∈[gⁿ] + where + ux≡0 : u · x ≡ 0r + ux≡0 = sym (·Rid _) ∙ p ∙ cong (_· 1r) (0RightAnnihilates _) ∙ (·Rid _) + + vx≡0 : v · x ≡ 0r + vx≡0 = sym (·Rid _) ∙ q ∙ cong (_· 1r) (0RightAnnihilates _) ∙ (·Rid _) + + exponentHelper : Σ[ n ∈ ℕ ] u ≡ f ^ n + → Σ[ n ∈ ℕ ] v ≡ g ^ n + → x ≡ 0r + exponentHelper (n , u≡fⁿ) (m , v≡gᵐ) = + PT.rec (is-set _ _) Σhelper (GeneratingExponents.lemma R' f g l 1∈⟨f,g⟩) + where + l = max n m + + fˡx≡0 : f ^ l · x ≡ 0r + fˡx≡0 = f ^ l · x ≡⟨ cong (λ k → f ^ k · x) (sym (≤-∸-+-cancel left-≤-max)) ⟩ + f ^ ((l ∸ n) +ℕ n) · x ≡⟨ cong (_· x) (sym (·-of-^-is-^-of-+ _ _ _)) ⟩ + f ^ (l ∸ n) · f ^ n · x ≡⟨ cong (λ y → f ^ (l ∸ n) · y · x) (sym u≡fⁿ) ⟩ + f ^ (l ∸ n) · u · x ≡⟨ sym (·Assoc _ _ _) ⟩ + f ^ (l ∸ n) · (u · x) ≡⟨ cong (f ^ (l ∸ n) ·_) ux≡0 ⟩ + f ^ (l ∸ n) · 0r ≡⟨ 0RightAnnihilates _ ⟩ + 0r ∎ + + gˡx≡0 : g ^ l · x ≡ 0r + gˡx≡0 = g ^ l · x ≡⟨ cong (λ k → g ^ k · x) (sym (≤-∸-+-cancel right-≤-max)) ⟩ + g ^ ((l ∸ m) +ℕ m) · x ≡⟨ cong (_· x) (sym (·-of-^-is-^-of-+ _ _ _)) ⟩ + g ^ (l ∸ m) · g ^ m · x ≡⟨ cong (λ y → g ^ (l ∸ m) · y · x) (sym v≡gᵐ) ⟩ + g ^ (l ∸ m) · v · x ≡⟨ sym (·Assoc _ _ _) ⟩ + g ^ (l ∸ m) · (v · x) ≡⟨ cong (g ^ (l ∸ m) ·_) vx≡0 ⟩ + g ^ (l ∸ m) · 0r ≡⟨ 0RightAnnihilates _ ⟩ + 0r ∎ + + Σhelper : Σ[ α ∈ FinVec R 2 ] 1r ≡ α zero · f ^ l + (α (suc zero) · g ^ l + 0r) + → x ≡ 0r + Σhelper (α , 1≡α₀fˡ+α₁gˡ+0) = path + where + α₀ = α zero + α₁ = α (suc zero) + + 1≡α₀fˡ+α₁gˡ : 1r ≡ α₀ · f ^ l + α₁ · g ^ l + 1≡α₀fˡ+α₁gˡ = 1≡α₀fˡ+α₁gˡ+0 ∙ cong (α₀ · f ^ l +_) (+Rid _) + + path : x ≡ 0r + path = x ≡⟨ sym (·Lid _) ⟩ + 1r · x ≡⟨ cong (_· x) 1≡α₀fˡ+α₁gˡ ⟩ + (α₀ · f ^ l + α₁ · g ^ l) · x ≡⟨ ·Ldist+ _ _ _ ⟩ + α₀ · f ^ l · x + α₁ · g ^ l · x ≡⟨ cong₂ _+_ (sym (·Assoc _ _ _)) + (sym (·Assoc _ _ _)) ⟩ + α₀ · (f ^ l · x) + α₁ · (g ^ l · x) ≡⟨ cong₂ (λ y z → α₀ · y + α₁ · z) + fˡx≡0 gˡx≡0 ⟩ + α₀ · 0r + α₁ · 0r ≡⟨ cong₂ _+_ (0RightAnnihilates _) + (0RightAnnihilates _) ∙ +Rid _ ⟩ + 0r ∎ + + + equalizerLemma : 1r ∈ ⟨f,g⟩ + → ∀ (x : R[1/ f ]) (y : R[1/ g ]) + → χ₁ .fst x ≡ χ₂ .fst y + → ∃![ z ∈ R ] (z /1ᶠ ≡ x) × (z /1ᵍ ≡ y) + equalizerLemma 1∈⟨f,g⟩ = invElPropElim2 (λ _ _ → isPropΠ (λ _ → isProp∃!)) baseCase + where + baseCase : ∀ (x y : R) (n : ℕ) + → fst χ₁ ([ x , f ^ n , ∣ n , refl ∣ ]) ≡ fst χ₂ ([ y , g ^ n , ∣ n , refl ∣ ]) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + baseCase x y n χ₁[x/fⁿ]≡χ₂[y/gⁿ] = PT.rec isProp∃! annihilatorHelper exAnnihilator + where + -- doesn't compute that well but at least it computes... + exAnnihilator : ∃[ s ∈ Sᶠᵍ ] -- s.t. + (fst s · (x · transport refl (g ^ n)) · (1r · transport refl ((f · g) ^ n)) + ≡ fst s · (y · transport refl (f ^ n)) · (1r · transport refl ((f · g) ^ n))) + exAnnihilator = isEquivRel→TruncIso locIsEquivRelᶠᵍ _ _ .fun χ₁[x/fⁿ]≡χ₂[y/gⁿ] + + annihilatorHelper : Σ[ s ∈ Sᶠᵍ ] + (fst s · (x · transport refl (g ^ n)) · (1r · transport refl ((f · g) ^ n)) + ≡ fst s · (y · transport refl (f ^ n)) · (1r · transport refl ((f · g) ^ n))) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + annihilatorHelper ((s , s∈[fgⁿ]) , p) = PT.rec isProp∃! exponentHelper s∈[fgⁿ] + where + sxgⁿ[fg]ⁿ≡syfⁿ[fg]ⁿ : s · x · g ^ n · (f · g) ^ n ≡ s · y · f ^ n · (f · g) ^ n + sxgⁿ[fg]ⁿ≡syfⁿ[fg]ⁿ = + s · x · g ^ n · (f · g) ^ n + + ≡⟨ transpHelper _ _ _ _ ⟩ + + s · x · transport refl (g ^ n) · transport refl ((f · g) ^ n) + + ≡⟨ useSolver _ _ _ _ ⟩ + + s · (x · transport refl (g ^ n)) · (1r · transport refl ((f · g) ^ n)) + + ≡⟨ p ⟩ + + s · (y · transport refl (f ^ n)) · (1r · transport refl ((f · g) ^ n)) + + ≡⟨ sym (useSolver _ _ _ _) ⟩ + + s · y · transport refl (f ^ n) · transport refl ((f · g) ^ n) + + ≡⟨ sym (transpHelper _ _ _ _) ⟩ + + s · y · f ^ n · (f · g) ^ n ∎ + + where + transpHelper : ∀ a b c d → a · b · c · d + ≡ a · b · transport refl c · transport refl d + transpHelper a b c d i = a · b · transportRefl c (~ i) · transportRefl d (~ i) + useSolver : ∀ a b c d → a · b · c · d ≡ a · (b · c) · (1r · d) + useSolver = solve R' + + + exponentHelper : Σ[ m ∈ ℕ ] s ≡ (f · g) ^ m + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + exponentHelper (m , s≡[fg]ᵐ) = + PT.rec isProp∃! Σhelper (GeneratingExponents.lemma R' f g 2n+m 1∈⟨f,g⟩) + where + -- the path we'll actually work with + xgⁿ[fg]ⁿ⁺ᵐ≡yfⁿ[fg]ⁿ⁺ᵐ : x · g ^ n · (f · g) ^ (n +ℕ m) ≡ y · f ^ n · (f · g) ^ (n +ℕ m) + xgⁿ[fg]ⁿ⁺ᵐ≡yfⁿ[fg]ⁿ⁺ᵐ = + x · g ^ n · (f · g) ^ (n +ℕ m) + + ≡⟨ cong (x · (g ^ n) ·_) (sym (·-of-^-is-^-of-+ _ _ _)) ⟩ + + x · g ^ n · ((f · g) ^ n · (f · g) ^ m) + + ≡⟨ useSolver _ _ _ _ ⟩ + + (f · g) ^ m · x · g ^ n · (f · g) ^ n + + ≡⟨ cong (λ a → a · x · g ^ n · (f · g) ^ n) (sym s≡[fg]ᵐ) ⟩ + + s · x · g ^ n · (f · g) ^ n + + ≡⟨ sxgⁿ[fg]ⁿ≡syfⁿ[fg]ⁿ ⟩ + + s · y · f ^ n · (f · g) ^ n + + ≡⟨ cong (λ a → a · y · f ^ n · (f · g) ^ n) s≡[fg]ᵐ ⟩ + + (f · g) ^ m · y · f ^ n · (f · g) ^ n + + ≡⟨ sym (useSolver _ _ _ _) ⟩ + + y · f ^ n · ((f · g) ^ n · (f · g) ^ m) + + ≡⟨ cong (y · (f ^ n) ·_) (·-of-^-is-^-of-+ _ _ _) ⟩ + + y · f ^ n · (f · g) ^ (n +ℕ m) ∎ + + where + useSolver : ∀ a b c d → a · b · (c · d) ≡ d · a · b · c + useSolver = solve R' + + -- critical exponent + 2n+m = n +ℕ (n +ℕ m) + -- extracting information from the fact that R=⟨f,g⟩ + Σhelper : Σ[ α ∈ FinVec R 2 ] 1r ≡ linearCombination R' α (fⁿgⁿVec 2n+m) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + Σhelper (α , linCombi) = uniqueExists z (z/1≡x/fⁿ , z/1≡y/gⁿ) + (λ _ → isProp× (is-set _ _) (is-set _ _)) + uniqueness + where + α₀ = α zero + α₁ = α (suc zero) + + 1≡α₀f²ⁿ⁺ᵐ+α₁g²ⁿ⁺ᵐ : 1r ≡ α₀ · f ^ 2n+m + α₁ · g ^ 2n+m + 1≡α₀f²ⁿ⁺ᵐ+α₁g²ⁿ⁺ᵐ = linCombi ∙ cong (α₀ · f ^ 2n+m +_) (+Rid _) + + -- definition of the element + z = α₀ · x · f ^ (n +ℕ m) + α₁ · y · g ^ (n +ℕ m) + + z/1≡x/fⁿ : (z /1ᶠ) ≡ [ x , f ^ n , ∣ n , refl ∣ ] + z/1≡x/fⁿ = eq/ _ _ ((f ^ (n +ℕ m) , ∣ n +ℕ m , refl ∣) , path) + where + useSolver1 : ∀ x y α₀ α₁ fⁿ⁺ᵐ gⁿ⁺ᵐ fⁿ + → fⁿ⁺ᵐ · (α₀ · x · fⁿ⁺ᵐ + α₁ · y · gⁿ⁺ᵐ) · fⁿ + ≡ fⁿ⁺ᵐ · (α₀ · x · (fⁿ · fⁿ⁺ᵐ)) + α₁ · (y · fⁿ · (fⁿ⁺ᵐ · gⁿ⁺ᵐ)) + useSolver1 = solve R' + + useSolver2 : ∀ x α₀ α₁ fⁿ⁺ᵐ g²ⁿ⁺ᵐ f²ⁿ⁺ᵐ + → fⁿ⁺ᵐ · (α₀ · x · f²ⁿ⁺ᵐ) + α₁ · (x · fⁿ⁺ᵐ · g²ⁿ⁺ᵐ) + ≡ fⁿ⁺ᵐ · x · (α₀ · f²ⁿ⁺ᵐ + α₁ · g²ⁿ⁺ᵐ) + useSolver2 = solve R' + + path : f ^ (n +ℕ m) · z · f ^ n ≡ f ^ (n +ℕ m) · x · 1r + path = + f ^ (n +ℕ m) · z · f ^ n + + ≡⟨ useSolver1 _ _ _ _ _ _ _ ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ n · f ^ (n +ℕ m))) + α₁ · (y · f ^ n · (f ^ (n +ℕ m) · g ^ (n +ℕ m))) + + ≡⟨ cong₂ (λ a b → f ^ (n +ℕ m) · (α₀ · x · a) + α₁ · ((y · f ^ n) · b)) + (·-of-^-is-^-of-+ _ _ _) (sym (^-ldist-· _ _ _)) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (y · f ^ n · (f · g) ^ (n +ℕ m)) + + ≡⟨ cong (λ a → f ^ (n +ℕ m) · (α₀ · x · f ^ 2n+m) + α₁ · a) + (sym xgⁿ[fg]ⁿ⁺ᵐ≡yfⁿ[fg]ⁿ⁺ᵐ) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · g ^ n · (f · g) ^ (n +ℕ m)) + + ≡⟨ cong (λ a → f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · g ^ n · a)) (^-ldist-· _ _ _) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · g ^ n · (f ^ (n +ℕ m) · g ^ (n +ℕ m))) + + ≡⟨ cong (λ a → f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · a) (·CommAssocSwap _ _ _ _) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · f ^ (n +ℕ m) · (g ^ n · g ^ (n +ℕ m))) + + ≡⟨ cong (λ a → f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · f ^ (n +ℕ m) · a)) + (·-of-^-is-^-of-+ _ _ _) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · f ^ (n +ℕ m) · g ^ 2n+m) + + ≡⟨ useSolver2 _ _ _ _ _ _ ⟩ + + f ^ (n +ℕ m) · x · (α₀ · f ^ 2n+m + α₁ · g ^ 2n+m) + + ≡⟨ cong (f ^ (n +ℕ m) · x ·_) (sym 1≡α₀f²ⁿ⁺ᵐ+α₁g²ⁿ⁺ᵐ) ⟩ + + f ^ (n +ℕ m) · x · 1r ∎ + + z/1≡y/gⁿ : (z /1ᵍ) ≡ [ y , g ^ n , ∣ n , refl ∣ ] + z/1≡y/gⁿ = eq/ _ _ ((g ^ (n +ℕ m) , ∣ n +ℕ m , refl ∣) , path) + where + useSolver1 : ∀ x y α₀ α₁ fⁿ⁺ᵐ gⁿ⁺ᵐ gⁿ + → gⁿ⁺ᵐ · (α₀ · x · fⁿ⁺ᵐ + α₁ · y · gⁿ⁺ᵐ) · gⁿ + ≡ α₀ · (x · gⁿ · (fⁿ⁺ᵐ · gⁿ⁺ᵐ)) + gⁿ⁺ᵐ · (α₁ · y · (gⁿ · gⁿ⁺ᵐ)) + useSolver1 = solve R' + + useSolver2 : ∀ y α₀ α₁ gⁿ⁺ᵐ g²ⁿ⁺ᵐ f²ⁿ⁺ᵐ + → α₀ · (y · f²ⁿ⁺ᵐ · gⁿ⁺ᵐ) + gⁿ⁺ᵐ · (α₁ · y · g²ⁿ⁺ᵐ) + ≡ gⁿ⁺ᵐ · y · (α₀ · f²ⁿ⁺ᵐ + α₁ · g²ⁿ⁺ᵐ) + useSolver2 = solve R' + + path : g ^ (n +ℕ m) · z · g ^ n ≡ g ^ (n +ℕ m) · y · 1r + path = + g ^ (n +ℕ m) · z · g ^ n + + ≡⟨ useSolver1 _ _ _ _ _ _ _ ⟩ + + α₀ · (x · g ^ n · (f ^ (n +ℕ m) · g ^ (n +ℕ m))) + g ^ (n +ℕ m) · (α₁ · y · (g ^ n · g ^ (n +ℕ m))) + + ≡⟨ cong₂ (λ a b → α₀ · (x · g ^ n · a) + g ^ (n +ℕ m) · (α₁ · y · b)) + (sym (^-ldist-· _ _ _)) (·-of-^-is-^-of-+ _ _ _) ⟩ + + α₀ · (x · g ^ n · (f · g) ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ cong (λ a → α₀ · a + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m)) + xgⁿ[fg]ⁿ⁺ᵐ≡yfⁿ[fg]ⁿ⁺ᵐ ⟩ + + α₀ · (y · f ^ n · (f · g) ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ cong (λ a → α₀ · (y · f ^ n · a) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m)) (^-ldist-· _ _ _) ⟩ + + α₀ · (y · f ^ n · (f ^ (n +ℕ m) · g ^ (n +ℕ m))) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ cong (λ a → α₀ · a + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m)) (·-assoc2 _ _ _ _) ⟩ + + α₀ · (y · (f ^ n · f ^ (n +ℕ m)) · g ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ cong (λ a → α₀ · (y · a · g ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m)) + (·-of-^-is-^-of-+ _ _ _) ⟩ + + α₀ · (y · f ^ 2n+m · g ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ useSolver2 _ _ _ _ _ _ ⟩ + + g ^ (n +ℕ m) · y · (α₀ · f ^ 2n+m + α₁ · g ^ 2n+m) + + ≡⟨ cong (g ^ (n +ℕ m) · y ·_) (sym 1≡α₀f²ⁿ⁺ᵐ+α₁g²ⁿ⁺ᵐ) ⟩ + + g ^ (n +ℕ m) · y · 1r ∎ + + + uniqueness : ∀ a → ((a /1ᶠ) ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × ((a /1ᵍ) ≡ [ y , g ^ n , ∣ n , refl ∣ ]) + → z ≡ a + uniqueness a (a/1≡x/fⁿ , a/1≡y/gⁿ) = equalByDifference _ _ + (injectivityLemma 1∈⟨f,g⟩ (z - a) [z-a]/1≡0overF [z-a]/1≡0overG) + where + [z-a]/1≡0overF : (z - a) /1ᶠ ≡ 0r + [z-a]/1≡0overF = (z - a) /1ᶠ + + ≡⟨ /1ᶠAsCommRingHom .snd .pres+ _ _ ⟩ -- (-a)/1=-(a/1) by refl! + + (z /1ᶠ) - (a /1ᶠ) + + ≡⟨ cong₂ _-_ z/1≡x/fⁿ a/1≡x/fⁿ ⟩ + + [ x , f ^ n , ∣ n , refl ∣ ] - [ x , f ^ n , ∣ n , refl ∣ ] + + ≡⟨ +Rinv ([ x , f ^ n , ∣ n , refl ∣ ]) ⟩ + + 0r ∎ + + [z-a]/1≡0overG : (z - a) /1ᵍ ≡ 0r + [z-a]/1≡0overG = (z - a) /1ᵍ + + ≡⟨ /1ᵍAsCommRingHom .snd .pres+ _ _ ⟩ -- (-a)/1=-(a/1) by refl! + + (z /1ᵍ) - (a /1ᵍ) + + ≡⟨ cong₂ _-_ z/1≡y/gⁿ a/1≡y/gⁿ ⟩ + + [ y , g ^ n , ∣ n , refl ∣ ] - [ y , g ^ n , ∣ n , refl ∣ ] + + ≡⟨ +Rinv ([ y , g ^ n , ∣ n , refl ∣ ]) ⟩ + + 0r ∎ + + + {- + putting everything together with the pullback machinery: + If ⟨f,g⟩ = R then we get a square + + _/1ᶠ + R ----> R[1/f] + ⌟ + _/1ᵍ | | χ₁ + v v + + R[1/g] -> R[1/fg] + χ₂ + -} + + open Category (CommRingsCategory {ℓ}) + open Cospan + + fgCospan : Cospan CommRingsCategory + l fgCospan = R[1/ f ]AsCommRing + m fgCospan = R[1/ (f · g) ]AsCommRing + r fgCospan = R[1/ g ]AsCommRing + s₁ fgCospan = χ₁ + s₂ fgCospan = χ₂ + + -- the commutative square + private + /1χComm : ∀ (x : R) → χ₁ .fst (x /1ᶠ) ≡ χ₂ .fst (x /1ᵍ) + /1χComm x = eq/ _ _ ((1r , powersFormMultClosedSubset (f · g) .containsOne) , refl) + + /1χHomComm : /1ᶠAsCommRingHom ⋆ χ₁ ≡ /1ᵍAsCommRingHom ⋆ χ₂ + /1χHomComm = RingHom≡ (funExt /1χComm) + + fgSquare : 1r ∈ ⟨f,g⟩ + → isPullback _ fgCospan /1ᶠAsCommRingHom /1ᵍAsCommRingHom /1χHomComm + fgSquare 1∈⟨f,g⟩ {d = A} φ ψ φχ₁≡ψχ₂ = (χ , χCoh) , χUniqueness + where + instance + _ = snd A + + applyEqualizerLemma : ∀ a → ∃![ χa ∈ R ] (χa /1ᶠ ≡ fst φ a) × (χa /1ᵍ ≡ fst ψ a) + applyEqualizerLemma a = + equalizerLemma 1∈⟨f,g⟩ (fst φ a) (fst ψ a) (cong (_$ a) φχ₁≡ψχ₂) + + χ : CommRingHom A R' + fst χ a = applyEqualizerLemma a .fst .fst + snd χ = makeIsRingHom + (cong fst (applyEqualizerLemma 1r .snd (1r , 1Coh))) + (λ x y → cong fst (applyEqualizerLemma (x + y) .snd (_ , +Coh x y))) + (λ x y → cong fst (applyEqualizerLemma (x · y) .snd (_ , ·Coh x y))) + where + open IsRingHom + 1Coh : (1r /1ᶠ ≡ fst φ 1r) × (1r /1ᵍ ≡ fst ψ 1r) + 1Coh = (sym (φ .snd .pres1)) , sym (ψ .snd .pres1) + + +Coh : ∀ x y → ((fst χ x + fst χ y) /1ᶠ ≡ fst φ (x + y)) + × ((fst χ x + fst χ y) /1ᵍ ≡ fst ψ (x + y)) + fst (+Coh x y) = /1ᶠAsCommRingHom .snd .pres+ _ _ + ∙∙ cong₂ _+_ (applyEqualizerLemma x .fst .snd .fst) + (applyEqualizerLemma y .fst .snd .fst) + ∙∙ sym (φ .snd .pres+ x y) + snd (+Coh x y) = /1ᵍAsCommRingHom .snd .pres+ _ _ + ∙∙ cong₂ _+_ (applyEqualizerLemma x .fst .snd .snd) + (applyEqualizerLemma y .fst .snd .snd) + ∙∙ sym (ψ .snd .pres+ x y) + + ·Coh : ∀ x y → ((fst χ x · fst χ y) /1ᶠ ≡ fst φ (x · y)) + × ((fst χ x · fst χ y) /1ᵍ ≡ fst ψ (x · y)) + fst (·Coh x y) = /1ᶠAsCommRingHom .snd .pres· _ _ + ∙∙ cong₂ _·_ (applyEqualizerLemma x .fst .snd .fst) + (applyEqualizerLemma y .fst .snd .fst) + ∙∙ sym (φ .snd .pres· x y) + snd (·Coh x y) = /1ᵍAsCommRingHom .snd .pres· _ _ + ∙∙ cong₂ _·_ (applyEqualizerLemma x .fst .snd .snd) + (applyEqualizerLemma y .fst .snd .snd) + ∙∙ sym (ψ .snd .pres· x y) + + + χCoh : (φ ≡ χ ⋆ /1ᶠAsCommRingHom) × (ψ ≡ χ ⋆ /1ᵍAsCommRingHom) + fst χCoh = RingHom≡ (funExt (λ a → sym (applyEqualizerLemma a .fst .snd .fst))) + snd χCoh = RingHom≡ (funExt (λ a → sym (applyEqualizerLemma a .fst .snd .snd))) + + χUniqueness : (y : Σ[ θ ∈ CommRingHom A R' ] + (φ ≡ θ ⋆ /1ᶠAsCommRingHom) × (ψ ≡ θ ⋆ /1ᵍAsCommRingHom)) + → (χ , χCoh) ≡ y + χUniqueness (θ , θCoh) = Σ≡Prop (λ _ → isProp× (isSetRingHom _ _ _ _) + (isSetRingHom _ _ _ _)) + (RingHom≡ (funExt (λ a → cong fst (applyEqualizerLemma a .snd (θtriple a))))) + where + θtriple : ∀ a → Σ[ x ∈ R ] (x /1ᶠ ≡ fst φ a) × (x /1ᵍ ≡ fst ψ a) + θtriple a = fst θ a , sym (cong (_$ a) (θCoh .fst)) + , sym (cong (_$ a) (θCoh .snd)) diff --git a/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda b/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda new file mode 100644 index 0000000000..2d2e00cf81 --- /dev/null +++ b/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda @@ -0,0 +1,417 @@ +-- We define what it means to satisfy the universal property +-- of localisation and show that the localisation in Base satisfies +-- it. We will also show that the localisation is uniquely determined +-- by the universal property, and give sufficient criteria for +-- satisfying the universal property. + +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.Localisation.UniversalProperty where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Transport +open import Cubical.Functions.FunExtEquiv +open import Cubical.Functions.Surjection +open import Cubical.Functions.Embedding + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) +open import Cubical.Data.Vec +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.RingSolver.Reflection +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Localisation.Base + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open Iso + +private + variable + ℓ ℓ' : Level + + +module _ (R' : CommRing ℓ) (S' : ℙ (fst R')) (SMultClosedSubset : isMultClosedSubset R' S') where + open isMultClosedSubset + private R = fst R' + open CommRingStr (snd R') hiding (is-set) + open RingTheory (CommRing→Ring R') + open IsRingHom + + + + hasLocUniversalProp : (A : CommRing ℓ) (φ : CommRingHom R' A) + → (∀ s → s ∈ S' → fst φ s ∈ A ˣ) + → Type (ℓ-suc ℓ) + hasLocUniversalProp A φ _ = (B : CommRing ℓ) (ψ : CommRingHom R' B) + → (∀ s → s ∈ S' → fst ψ s ∈ B ˣ) + → ∃![ χ ∈ CommRingHom A B ] (fst χ) ∘ (fst φ) ≡ (fst ψ) + + isPropUniversalProp : (A : CommRing ℓ) (φ : CommRingHom R' A) + → (φS⊆Aˣ : ∀ s → s ∈ S' → fst φ s ∈ A ˣ) + → isProp (hasLocUniversalProp A φ φS⊆Aˣ) + isPropUniversalProp A φ φS⊆Aˣ = isPropΠ3 (λ _ _ _ → isPropIsContr) + + -- S⁻¹R has the universal property + module S⁻¹RUniversalProp where + open Loc R' S' SMultClosedSubset + _/1 : R → S⁻¹R + r /1 = [ r , 1r , SMultClosedSubset .containsOne ] + + /1AsCommRingHom : CommRingHom R' S⁻¹RAsCommRing + fst /1AsCommRingHom = _/1 + snd /1AsCommRingHom = + makeIsRingHom + refl + (λ r r' → cong [_] (≡-× (cong₂ (_+_) (sym (·Rid r)) (sym (·Rid r'))) + (Σ≡Prop (λ x → S' x .snd) (sym (·Lid 1r))))) + (λ _ _ → cong [_] (≡-× refl (Σ≡Prop (λ x → S' x .snd) (sym (·Lid 1r))))) + + S⁻¹Rˣ = S⁻¹RAsCommRing ˣ + S/1⊆S⁻¹Rˣ : ∀ s → s ∈ S' → (s /1) ∈ S⁻¹Rˣ + S/1⊆S⁻¹Rˣ s s∈S' = [ 1r , s , s∈S' ] , eq/ _ _ ((1r , SMultClosedSubset .containsOne) , path s) + where + path : ∀ s → 1r · (s · 1r) · 1r ≡ 1r · 1r · (1r · s) + path = solve R' + + S⁻¹RHasUniversalProp : hasLocUniversalProp S⁻¹RAsCommRing /1AsCommRingHom S/1⊆S⁻¹Rˣ + S⁻¹RHasUniversalProp B' ψ ψS⊆Bˣ = (χ , funExt χcomp) , χunique + where + B = fst B' + open CommRingStr (snd B') renaming ( is-set to Bset ; _·_ to _·B_ ; 1r to 1b + ; _+_ to _+B_ + ; ·Assoc to ·B-assoc ; ·Comm to ·B-comm + ; ·Lid to ·B-lid ; ·Rid to ·B-rid + ; ·Ldist+ to ·B-ldist-+) + open Units B' renaming (Rˣ to Bˣ ; RˣMultClosed to BˣMultClosed ; RˣContainsOne to BˣContainsOne) + open RingTheory (CommRing→Ring B') renaming (·-assoc2 to ·B-assoc2) + open CommRingTheory B' renaming (·CommAssocl to ·B-commAssocl ; ·CommAssocSwap to ·B-commAssocSwap) + + ψ₀ = fst ψ + module ψ = IsRingHom (snd ψ) + + χ : CommRingHom S⁻¹RAsCommRing B' + fst χ = SQ.rec Bset fχ fχcoh + where + fχ : R × S → B + fχ (r , s , s∈S') = (fst ψ r) ·B ((fst ψ s) ⁻¹) ⦃ ψS⊆Bˣ s s∈S' ⦄ + fχcoh : (a b : R × S) → a ≈ b → fχ a ≡ fχ b + fχcoh (r , s , s∈S') (r' , s' , s'∈S') ((u , u∈S') , p) = instancepath + ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦃ ψS⊆Bˣ (u · s · s') + (SMultClosedSubset .multClosed (SMultClosedSubset .multClosed u∈S' s∈S') s'∈S') ⦄ + ⦃ BˣMultClosed _ _ ⦃ ψS⊆Bˣ (u · s) (SMultClosedSubset .multClosed u∈S' s∈S') ⦄ + ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦄ + ⦃ ψS⊆Bˣ (u · s) (SMultClosedSubset .multClosed u∈S' s∈S') ⦄ + where + -- only a few individual steps can be solved by the ring solver yet + instancepath : ⦃ _ : ψ₀ s ∈ Bˣ ⦄ ⦃ _ : ψ₀ s' ∈ Bˣ ⦄ + ⦃ _ : ψ₀ (u · s · s') ∈ Bˣ ⦄ ⦃ _ : ψ₀ (u · s) ·B ψ₀ s' ∈ Bˣ ⦄ + ⦃ _ : ψ₀ (u · s) ∈ Bˣ ⦄ + → ψ₀ r ·B ψ₀ s ⁻¹ ≡ ψ₀ r' ·B ψ₀ s' ⁻¹ + instancepath = ψ₀ r ·B ψ₀ s ⁻¹ + + ≡⟨ sym (·B-rid _) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B 1b + + ≡⟨ cong (ψ₀ r ·B ψ₀ s ⁻¹ ·B_) (sym (·-rinv _)) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ (u · s · s') ·B ψ₀ (u · s · s') ⁻¹) + + ≡⟨ ·B-assoc _ _ _ ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B ψ₀ (u · s · s') ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → ψ₀ r ·B ψ₀ s ⁻¹ ·B x ·B ψ₀ (u · s · s') ⁻¹) (ψ.pres· _ _) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ (u · s) ·B ψ₀ s') ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (·B-assoc _ _ _) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B ψ₀ (u · s) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → ψ₀ r ·B ψ₀ s ⁻¹ ·B x ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹) + (ψ.pres· _ _) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ u ·B ψ₀ s) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → x ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹) (·B-commAssocSwap _ _ _ _) ⟩ + + ψ₀ r ·B ψ₀ u ·B (ψ₀ s ⁻¹ ·B ψ₀ s) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ (λ i → ·B-comm (ψ₀ r) (ψ₀ u) i ·B (·-linv (ψ₀ s) i) + ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹) ⟩ + + ψ₀ u ·B ψ₀ r ·B 1b ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ (λ i → (·B-rid (sym (ψ.pres· u r) i) i) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹) ⟩ + + ψ₀ (u · r) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (sym (ψ.pres· _ _)) ⟩ + + ψ₀ (u · r · s') ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → ψ₀ x ·B ψ₀ (u · s · s') ⁻¹) p ⟩ + + ψ₀ (u · r' · s) ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (ψ.pres· _ _) ⟩ + + ψ₀ (u · r') ·B ψ₀ s ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → x ·B ψ₀ s ·B ψ₀ (u · s · s') ⁻¹) (ψ.pres· _ _) ⟩ + + ψ₀ u ·B ψ₀ r' ·B ψ₀ s ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (sym (·B-assoc _ _ _)) ⟩ + + ψ₀ u ·B (ψ₀ r' ·B ψ₀ s) ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (·B-commAssocl _ _ _) ⟩ + + ψ₀ r' ·B (ψ₀ u ·B ψ₀ s) ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → ψ₀ r' ·B x ·B ψ₀ (u · s · s') ⁻¹) (sym (ψ.pres· _ _)) ⟩ + + ψ₀ r' ·B ψ₀ (u · s) ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (ψ₀ r' ·B ψ₀ (u · s) ·B_) (unitCong (ψ.pres· _ _)) ⟩ + + ψ₀ r' ·B ψ₀ (u · s) ·B (ψ₀ (u · s) ·B ψ₀ s') ⁻¹ + + ≡⟨ cong (ψ₀ r' ·B ψ₀ (u · s) ·B_) (⁻¹-dist-· _ _) ⟩ + + ψ₀ r' ·B ψ₀ (u · s) ·B (ψ₀ (u · s) ⁻¹ ·B ψ₀ s' ⁻¹) + + ≡⟨ ·B-assoc2 _ _ _ _ ⟩ + + ψ₀ r' ·B (ψ₀ (u · s) ·B ψ₀ (u · s) ⁻¹) ·B ψ₀ s' ⁻¹ + + ≡⟨ cong (λ x → ψ₀ r' ·B x ·B ψ₀ s' ⁻¹) (·-rinv _) ⟩ + + ψ₀ r' ·B 1b ·B ψ₀ s' ⁻¹ + + ≡⟨ cong (_·B ψ₀ s' ⁻¹) (·B-rid _) ⟩ + + ψ₀ r' ·B ψ₀ s' ⁻¹ ∎ + + snd χ = + makeIsRingHom + (instancepres1χ ⦃ ψS⊆Bˣ 1r (SMultClosedSubset .containsOne) ⦄ ⦃ BˣContainsOne ⦄) + (elimProp2 (λ _ _ _ _ → Bset _ _ _ _) pres+[]) + (elimProp2 (λ _ _ _ _ → Bset _ _ _ _) pres·[]) + where + instancepres1χ : ⦃ _ : ψ₀ 1r ∈ Bˣ ⦄ ⦃ _ : 1b ∈ Bˣ ⦄ + → ψ₀ 1r ·B (ψ₀ 1r) ⁻¹ ≡ 1b + instancepres1χ = (λ i → (ψ.pres1 i) ·B (unitCong (ψ.pres1) i)) + ∙ (λ i → ·B-lid (1⁻¹≡1 i) i) + + pres+[] : (a b : R × S) → fst χ ([ a ] +ₗ [ b ]) ≡ (fst χ [ a ]) +B (fst χ [ b ]) + pres+[] (r , s , s∈S') (r' , s' , s'∈S') = instancepath + ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦃ ψS⊆Bˣ (s · s') (SMultClosedSubset .multClosed s∈S' s'∈S') ⦄ + ⦃ BˣMultClosed _ _ ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦄ + where + -- only a few individual steps can be solved by the ring solver yet + instancepath : ⦃ _ : ψ₀ s ∈ Bˣ ⦄ ⦃ _ : ψ₀ s' ∈ Bˣ ⦄ + ⦃ _ : ψ₀ (s · s') ∈ Bˣ ⦄ ⦃ _ : ψ₀ s ·B ψ₀ s' ∈ Bˣ ⦄ + → ψ₀ (r · s' + r' · s) ·B ψ₀ (s · s') ⁻¹ ≡ ψ₀ r ·B ψ₀ s ⁻¹ +B ψ₀ r' ·B ψ₀ s' ⁻¹ + instancepath = + ψ₀ (r · s' + r' · s) ·B ψ₀ (s · s') ⁻¹ + + ≡⟨ (λ i → ψ.pres+ (r · s') (r' · s) i ·B unitCong (ψ.pres· s s') i) ⟩ + + (ψ₀ (r · s') +B ψ₀ (r' · s)) ·B (ψ₀ s ·B ψ₀ s') ⁻¹ + + ≡⟨ (λ i → (ψ.pres· r s' i +B ψ.pres· r' s i) ·B ⁻¹-dist-· (ψ₀ s) (ψ₀ s') i) ⟩ + + (ψ₀ r ·B ψ₀ s' +B ψ₀ r' ·B ψ₀ s) ·B (ψ₀ s ⁻¹ ·B ψ₀ s' ⁻¹) + + ≡⟨ ·B-ldist-+ _ _ _ ⟩ + + ψ₀ r ·B ψ₀ s' ·B (ψ₀ s ⁻¹ ·B ψ₀ s' ⁻¹) +B ψ₀ r' ·B ψ₀ s ·B (ψ₀ s ⁻¹ ·B ψ₀ s' ⁻¹) + + ≡⟨ (λ i → ·B-commAssocSwap (ψ₀ r) (ψ₀ s') (ψ₀ s ⁻¹) (ψ₀ s' ⁻¹) i + +B ·B-assoc2 (ψ₀ r') (ψ₀ s) (ψ₀ s ⁻¹) (ψ₀ s' ⁻¹) i) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ s' ·B ψ₀ s' ⁻¹) +B ψ₀ r' ·B (ψ₀ s ·B ψ₀ s ⁻¹) ·B ψ₀ s' ⁻¹ + + ≡⟨ (λ i → ψ₀ r ·B ψ₀ s ⁻¹ ·B (·-rinv (ψ₀ s') i) + +B ψ₀ r' ·B (·-rinv (ψ₀ s) i) ·B ψ₀ s' ⁻¹) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B 1b +B ψ₀ r' ·B 1b ·B ψ₀ s' ⁻¹ + + ≡⟨ (λ i → ·B-rid (ψ₀ r ·B ψ₀ s ⁻¹) i +B ·B-rid (ψ₀ r') i ·B ψ₀ s' ⁻¹) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ +B ψ₀ r' ·B ψ₀ s' ⁻¹ ∎ + + pres·[] : (a b : R × S) → fst χ ([ a ] ·ₗ [ b ]) ≡ (fst χ [ a ]) ·B (fst χ [ b ]) + pres·[] (r , s , s∈S') (r' , s' , s'∈S') = instancepath + ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦃ ψS⊆Bˣ (s · s') (SMultClosedSubset .multClosed s∈S' s'∈S') ⦄ + ⦃ BˣMultClosed _ _ ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦄ + where + -- only a few individual steps can be solved by the ring solver yet + instancepath : ⦃ _ : ψ₀ s ∈ Bˣ ⦄ ⦃ _ : ψ₀ s' ∈ Bˣ ⦄ + ⦃ _ : ψ₀ (s · s') ∈ Bˣ ⦄ ⦃ _ : ψ₀ s ·B ψ₀ s' ∈ Bˣ ⦄ + → ψ₀ (r · r') ·B ψ₀ (s · s') ⁻¹ ≡ (ψ₀ r ·B ψ₀ s ⁻¹) ·B (ψ₀ r' ·B ψ₀ s' ⁻¹) + instancepath = ψ₀ (r · r') ·B ψ₀ (s · s') ⁻¹ + + ≡⟨ (λ i → ψ.pres· r r' i ·B unitCong (ψ.pres· s s') i) ⟩ + + ψ₀ r ·B ψ₀ r' ·B (ψ₀ s ·B ψ₀ s') ⁻¹ + + ≡⟨ cong (ψ₀ r ·B ψ₀ r' ·B_) (⁻¹-dist-· _ _) ⟩ + + ψ₀ r ·B ψ₀ r' ·B (ψ₀ s ⁻¹ ·B ψ₀ s' ⁻¹) + + ≡⟨ ·B-commAssocSwap _ _ _ _ ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ r' ·B ψ₀ s' ⁻¹) ∎ + + + χcomp : (r : R) → fst χ (r /1) ≡ ψ₀ r + χcomp = instanceχcomp ⦃ ψS⊆Bˣ 1r (SMultClosedSubset .containsOne) ⦄ ⦃ Units.RˣContainsOne B' ⦄ + where + instanceχcomp : ⦃ _ : ψ₀ 1r ∈ Bˣ ⦄ ⦃ _ : 1b ∈ Bˣ ⦄ + (r : R) → ψ₀ r ·B (ψ₀ 1r) ⁻¹ ≡ ψ₀ r + instanceχcomp r = ψ₀ r ·B (ψ₀ 1r) ⁻¹ ≡⟨ cong (ψ₀ r ·B_) (unitCong (ψ.pres1)) ⟩ + ψ₀ r ·B 1b ⁻¹ ≡⟨ cong (ψ₀ r ·B_) 1⁻¹≡1 ⟩ + ψ₀ r ·B 1b ≡⟨ ·B-rid _ ⟩ + ψ₀ r ∎ + + + χunique : (y : Σ[ χ' ∈ CommRingHom S⁻¹RAsCommRing B' ] fst χ' ∘ _/1 ≡ ψ₀) + → (χ , funExt χcomp) ≡ y + χunique (χ' , χ'/1≡ψ) = Σ≡Prop (λ x → isSetΠ (λ _ → Bset) _ _) (RingHom≡ fχ≡fχ') + where + open CommRingHomTheory {A' = S⁻¹RAsCommRing} {B' = B'} χ' + renaming (φ[x⁻¹]≡φ[x]⁻¹ to χ'[x⁻¹]≡χ'[x]⁻¹) + + module χ' = IsRingHom (χ' .snd) + + []-path : (a : R × S) → fst χ [ a ] ≡ fst χ' [ a ] + []-path (r , s , s∈S') = instancepath ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ S/1⊆S⁻¹Rˣ s s∈S' ⦄ + ⦃ RingHomRespInv _ ⦃ S/1⊆S⁻¹Rˣ s s∈S' ⦄ ⦄ + where + open Units S⁻¹RAsCommRing renaming (_⁻¹ to _⁻¹ˡ ; inverseUniqueness to S⁻¹RInverseUniqueness) + hiding (unitCong) + + s-inv : ⦃ s/1∈S⁻¹Rˣ : s /1 ∈ S⁻¹Rˣ ⦄ → s /1 ⁻¹ˡ ≡ [ 1r , s , s∈S' ] + s-inv ⦃ s/1∈S⁻¹Rˣ ⦄ = PathPΣ (S⁻¹RInverseUniqueness (s /1) s/1∈S⁻¹Rˣ + (_ , eq/ _ _ ((1r , SMultClosedSubset .containsOne) , path s))) .fst + where + path : ∀ s → 1r · (s · 1r) · 1r ≡ 1r · 1r · (1r · s) + path = solve R' + + ·ₗ-path : [ r , s , s∈S' ] ≡ [ r , 1r , SMultClosedSubset .containsOne ] + ·ₗ [ 1r , s , s∈S' ] + ·ₗ-path = cong [_] (≡-× (sym (·Rid r)) (Σ≡Prop (λ x → S' x .snd) (sym (·Lid s)))) + + instancepath : ⦃ _ : ψ₀ s ∈ Bˣ ⦄ ⦃ _ : s /1 ∈ S⁻¹Rˣ ⦄ ⦃ _ : fst χ' (s /1) ∈ Bˣ ⦄ + → ψ₀ r ·B ψ₀ s ⁻¹ ≡ fst χ' [ r , s , s∈S' ] + instancepath = ψ₀ r ·B ψ₀ s ⁻¹ + + ≡⟨ cong (ψ₀ r ·B_) (unitCong (cong (λ φ → φ s) (sym χ'/1≡ψ))) ⟩ + + ψ₀ r ·B fst χ' (s /1) ⁻¹ + + ≡⟨ cong (ψ₀ r ·B_) (sym (χ'[x⁻¹]≡χ'[x]⁻¹ _)) ⟩ + + ψ₀ r ·B fst χ' (s /1 ⁻¹ˡ) + + ≡⟨ cong (λ x → ψ₀ r ·B fst χ' x) s-inv ⟩ + + ψ₀ r ·B fst χ' [ 1r , s , s∈S' ] + + ≡⟨ cong (_·B fst χ' [ 1r , s , s∈S' ]) (cong (λ φ → φ r) (sym χ'/1≡ψ)) ⟩ + + fst χ' [ r , 1r , SMultClosedSubset .containsOne ] ·B fst χ' [ 1r , s , s∈S' ] + + ≡⟨ sym (χ'.pres· _ _) ⟩ + + fst χ' ([ r , 1r , SMultClosedSubset .containsOne ] ·ₗ [ 1r , s , s∈S' ]) + + ≡⟨ cong (fst χ') (sym ·ₗ-path) ⟩ + + fst χ' [ r , s , s∈S' ] ∎ + + fχ≡fχ' : fst χ ≡ fst χ' + fχ≡fχ' = funExt (SQ.elimProp (λ _ → Bset _ _) []-path) + + + -- sufficient conditions for having the universal property + -- used as API in the leanprover-community/mathlib + -- Corollary 3.2 in Atiyah-McDonald + open S⁻¹RUniversalProp + open Loc R' S' SMultClosedSubset + + record PathToS⁻¹R (A' : CommRing ℓ) (φ : CommRingHom R' A') : Type ℓ where + constructor + pathtoS⁻¹R + open Units A' renaming (Rˣ to Aˣ) + open CommRingStr (snd A') renaming (is-set to Aset ; 0r to 0a ; _·_ to _·A_) + field + φS⊆Aˣ : ∀ s → s ∈ S' → fst φ s ∈ Aˣ + kerφ⊆annS : ∀ r → fst φ r ≡ 0a → ∃[ s ∈ S ] (s .fst) · r ≡ 0r + surχ : ∀ a → ∃[ x ∈ R × S ] fst φ (x .fst) ·A (fst φ (x .snd .fst) ⁻¹) ⦃ φS⊆Aˣ _ (x .snd .snd) ⦄ ≡ a + + S⁻¹RChar : (A' : CommRing ℓ) (φ : CommRingHom R' A') + → PathToS⁻¹R A' φ + → S⁻¹RAsCommRing ≡ A' + S⁻¹RChar A' φ cond = CommRingPath S⁻¹RAsCommRing A' .fst + (S⁻¹R≃A , χ .snd) + where + open CommRingStr (snd A') renaming ( is-set to Aset ; 0r to 0a ; _·_ to _·A_ ; 1r to 1a + ; ·Rid to ·A-rid) + open Units A' renaming (Rˣ to Aˣ ; RˣInvClosed to AˣInvClosed) + open PathToS⁻¹R ⦃...⦄ + private + A = fst A' + instance + _ = cond + χ = (S⁻¹RHasUniversalProp A' φ φS⊆Aˣ .fst .fst) + open RingHomTheory χ + + S⁻¹R≃A : S⁻¹R ≃ A + S⁻¹R≃A = fst χ , isEmbedding×isSurjection→isEquiv (Embχ , Surχ) + where + Embχ : isEmbedding (fst χ) + Embχ = injEmbedding squash/ Aset (ker≡0→inj λ {x} → kerχ≡0 x) + where + kerχ≡0 : (r/s : S⁻¹R) → fst χ r/s ≡ 0a → r/s ≡ 0ₗ + kerχ≡0 = SQ.elimProp (λ _ → isPropΠ λ _ → squash/ _ _) kerχ≡[] + where + kerχ≡[] : (a : R × S) → fst χ [ a ] ≡ 0a → [ a ] ≡ 0ₗ + kerχ≡[] (r , s , s∈S') p = PT.rec (squash/ _ _) Σhelper + (kerφ⊆annS r (UnitsAreNotZeroDivisors _ _ p)) + where + instance + _ : fst φ s ∈ Aˣ + _ = φS⊆Aˣ s s∈S' + _ : fst φ s ⁻¹ ∈ Aˣ + _ = AˣInvClosed (fst φ s) + + Σhelper : Σ[ s ∈ S ] (s .fst) · r ≡ 0r → [ r , s , s∈S' ] ≡ 0ₗ + Σhelper ((u , u∈S') , q) = eq/ _ _ ((u , u∈S') , path) + where + path : u · r · 1r ≡ u · 0r · s + path = (λ i → ·Rid (q i) i) ∙∙ sym (0LeftAnnihilates _) + ∙∙ cong (_· s) (sym (0RightAnnihilates _)) + + Surχ : isSurjection (fst χ) + Surχ a = PT.rec isPropPropTrunc (λ x → PT.∣ [ x .fst ] , x .snd ∣) (surχ a) diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda new file mode 100644 index 0000000000..5dc51a0150 --- /dev/null +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -0,0 +1,339 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommRing.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.GroupoidLaws hiding (_⁻¹) +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Auto +open import Cubical.Structures.Macro +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing.Base + +open import Cubical.HITs.PropositionalTruncation + +private + variable + ℓ ℓ' ℓ'' : Level + +module Units (R' : CommRing ℓ) where + open CommRingStr (snd R') + open RingTheory (CommRing→Ring R') + private R = fst R' + + inverseUniqueness : (r : R) → isProp (Σ[ r' ∈ R ] r · r' ≡ 1r) + inverseUniqueness r (r' , rr'≡1) (r'' , rr''≡1) = Σ≡Prop (λ _ → is-set _ _) path + where + path : r' ≡ r'' + path = r' ≡⟨ sym (·Rid _) ⟩ + r' · 1r ≡⟨ cong (r' ·_) (sym rr''≡1) ⟩ + r' · (r · r'') ≡⟨ ·Assoc _ _ _ ⟩ + (r' · r) · r'' ≡⟨ cong (_· r'') (·Comm _ _) ⟩ + (r · r') · r'' ≡⟨ cong (_· r'') rr'≡1 ⟩ + 1r · r'' ≡⟨ ·Lid _ ⟩ + r'' ∎ + + + Rˣ : ℙ R + Rˣ r = (Σ[ r' ∈ R ] r · r' ≡ 1r) , inverseUniqueness r + + -- some notation using instance arguments + _⁻¹ : (r : R) → ⦃ r ∈ Rˣ ⦄ → R + _⁻¹ r ⦃ r∈Rˣ ⦄ = r∈Rˣ .fst + + infix 9 _⁻¹ + + -- some results about inverses + ·-rinv : (r : R) ⦃ r∈Rˣ : r ∈ Rˣ ⦄ → r · r ⁻¹ ≡ 1r + ·-rinv r ⦃ r∈Rˣ ⦄ = r∈Rˣ .snd + + ·-linv : (r : R) ⦃ r∈Rˣ : r ∈ Rˣ ⦄ → r ⁻¹ · r ≡ 1r + ·-linv r ⦃ r∈Rˣ ⦄ = ·Comm _ _ ∙ r∈Rˣ .snd + + + RˣMultClosed : (r r' : R) ⦃ r∈Rˣ : r ∈ Rˣ ⦄ ⦃ r'∈Rˣ : r' ∈ Rˣ ⦄ + → (r · r') ∈ Rˣ + RˣMultClosed r r' = (r ⁻¹ · r' ⁻¹) , path + where + path : r · r' · (r ⁻¹ · r' ⁻¹) ≡ 1r + path = r · r' · (r ⁻¹ · r' ⁻¹) ≡⟨ cong (_· (r ⁻¹ · r' ⁻¹)) (·Comm _ _) ⟩ + r' · r · (r ⁻¹ · r' ⁻¹) ≡⟨ ·Assoc _ _ _ ⟩ + r' · r · r ⁻¹ · r' ⁻¹ ≡⟨ cong (_· r' ⁻¹) (sym (·Assoc _ _ _)) ⟩ + r' · (r · r ⁻¹) · r' ⁻¹ ≡⟨ cong (λ x → r' · x · r' ⁻¹) (·-rinv _) ⟩ + r' · 1r · r' ⁻¹ ≡⟨ cong (_· r' ⁻¹) (·Rid _) ⟩ + r' · r' ⁻¹ ≡⟨ ·-rinv _ ⟩ + 1r ∎ + + RˣContainsOne : 1r ∈ Rˣ + RˣContainsOne = 1r , ·Lid _ + + RˣInvClosed : (r : R) ⦃ _ : r ∈ Rˣ ⦄ → r ⁻¹ ∈ Rˣ + RˣInvClosed r = r , ·-linv _ + + UnitsAreNotZeroDivisors : (r : R) ⦃ _ : r ∈ Rˣ ⦄ + → ∀ r' → r' · r ≡ 0r → r' ≡ 0r + UnitsAreNotZeroDivisors r r' p = r' ≡⟨ sym (·Rid _) ⟩ + r' · 1r ≡⟨ cong (r' ·_) (sym (·-rinv _)) ⟩ + r' · (r · r ⁻¹) ≡⟨ ·Assoc _ _ _ ⟩ + r' · r · r ⁻¹ ≡⟨ cong (_· r ⁻¹) p ⟩ + 0r · r ⁻¹ ≡⟨ 0LeftAnnihilates _ ⟩ + 0r ∎ + + + -- laws keeping the instance arguments + 1⁻¹≡1 : ⦃ 1∈Rˣ' : 1r ∈ Rˣ ⦄ → 1r ⁻¹ ≡ 1r + 1⁻¹≡1 ⦃ 1∈Rˣ' ⦄ = (sym (·Lid _)) ∙ 1∈Rˣ' .snd + + ⁻¹-dist-· : (r r' : R) ⦃ r∈Rˣ : r ∈ Rˣ ⦄ ⦃ r'∈Rˣ : r' ∈ Rˣ ⦄ ⦃ rr'∈Rˣ : (r · r') ∈ Rˣ ⦄ + → (r · r') ⁻¹ ≡ r ⁻¹ · r' ⁻¹ + ⁻¹-dist-· r r' ⦃ r∈Rˣ ⦄ ⦃ r'∈Rˣ ⦄ ⦃ rr'∈Rˣ ⦄ = + sym path ∙∙ cong (r ⁻¹ · r' ⁻¹ ·_) (rr'∈Rˣ .snd) ∙∙ (·Rid _) + where + path : r ⁻¹ · r' ⁻¹ · (r · r' · (r · r') ⁻¹) ≡ (r · r') ⁻¹ + path = r ⁻¹ · r' ⁻¹ · (r · r' · (r · r') ⁻¹) + ≡⟨ ·Assoc _ _ _ ⟩ + r ⁻¹ · r' ⁻¹ · (r · r') · (r · r') ⁻¹ + ≡⟨ cong (λ x → r ⁻¹ · r' ⁻¹ · x · (r · r') ⁻¹) (·Comm _ _) ⟩ + r ⁻¹ · r' ⁻¹ · (r' · r) · (r · r') ⁻¹ + ≡⟨ cong (_· (r · r') ⁻¹) (sym (·Assoc _ _ _)) ⟩ + r ⁻¹ · (r' ⁻¹ · (r' · r)) · (r · r') ⁻¹ + ≡⟨ cong (λ x → r ⁻¹ · x · (r · r') ⁻¹) (·Assoc _ _ _) ⟩ + r ⁻¹ · (r' ⁻¹ · r' · r) · (r · r') ⁻¹ + ≡⟨ cong (λ x → r ⁻¹ · (x · r) · (r · r') ⁻¹) (·-linv _) ⟩ + r ⁻¹ · (1r · r) · (r · r') ⁻¹ + ≡⟨ cong (λ x → r ⁻¹ · x · (r · r') ⁻¹) (·Lid _) ⟩ + r ⁻¹ · r · (r · r') ⁻¹ + ≡⟨ cong (_· (r · r') ⁻¹) (·-linv _) ⟩ + 1r · (r · r') ⁻¹ + ≡⟨ ·Lid _ ⟩ + (r · r') ⁻¹ ∎ + + unitCong : {r r' : R} → r ≡ r' → ⦃ r∈Rˣ : r ∈ Rˣ ⦄ ⦃ r'∈Rˣ : r' ∈ Rˣ ⦄ → r ⁻¹ ≡ r' ⁻¹ + unitCong {r = r} {r' = r'} p ⦃ r∈Rˣ ⦄ ⦃ r'∈Rˣ ⦄ = + PathPΣ (inverseUniqueness r' (r ⁻¹ , subst (λ x → x · r ⁻¹ ≡ 1r) p (r∈Rˣ .snd)) r'∈Rˣ) .fst + + ⁻¹-eq-elim : {r r' r'' : R} ⦃ r∈Rˣ : r ∈ Rˣ ⦄ → r' ≡ r'' · r → r' · r ⁻¹ ≡ r'' + ⁻¹-eq-elim {r = r} {r'' = r''} p = cong (_· r ⁻¹) p + ∙ sym (·Assoc _ _ _) + ∙ cong (r'' ·_) (·-rinv _) + ∙ ·Rid _ + + +-- some convenient notation +_ˣ : (R' : CommRing ℓ) → ℙ (R' .fst) +R' ˣ = Units.Rˣ R' + +module CommRingHoms where + open RingHoms + + idCommRingHom : (R : CommRing ℓ) → CommRingHom R R + idCommRingHom R = idRingHom (CommRing→Ring R) + + compCommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') (T : CommRing ℓ'') + → CommRingHom R S → CommRingHom S T → CommRingHom R T + compCommRingHom R S T = compRingHom {R = CommRing→Ring R} {CommRing→Ring S} {CommRing→Ring T} + + _∘cr_ : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} + → CommRingHom S T → CommRingHom R S → CommRingHom R T + g ∘cr f = compCommRingHom _ _ _ f g + + compIdCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) + → compCommRingHom _ _ _ (idCommRingHom R) f ≡ f + compIdCommRingHom = compIdRingHom + + idCompCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) + → compCommRingHom _ _ _ f (idCommRingHom S) ≡ f + idCompCommRingHom = idCompRingHom + + compAssocCommRingHom : {R S T U : CommRing ℓ} + (f : CommRingHom R S) (g : CommRingHom S T) (h : CommRingHom T U) + → compCommRingHom _ _ _ (compCommRingHom _ _ _ f g) h + ≡ compCommRingHom _ _ _ f (compCommRingHom _ _ _ g h) + compAssocCommRingHom = compAssocRingHom + +module CommRingEquivs where + open RingEquivs + + compCommRingEquiv : {A : CommRing ℓ} {B : CommRing ℓ'} {C : CommRing ℓ''} + → CommRingEquiv A B → CommRingEquiv B C → CommRingEquiv A C + compCommRingEquiv {A = A} {B = B} {C = C} = compRingEquiv {A = CommRing→Ring A} + {B = CommRing→Ring B} + {C = CommRing→Ring C} + +module CommRingHomTheory {A' B' : CommRing ℓ} (φ : CommRingHom A' B') where + open Units A' renaming (Rˣ to Aˣ ; _⁻¹ to _⁻¹ᵃ ; ·-rinv to ·A-rinv ; ·-linv to ·A-linv) + private A = fst A' + open CommRingStr (snd A') renaming (_·_ to _·A_ ; 1r to 1a) + open Units B' renaming (Rˣ to Bˣ ; _⁻¹ to _⁻¹ᵇ ; ·-rinv to ·B-rinv) + open CommRingStr (snd B') renaming ( _·_ to _·B_ ; 1r to 1b + ; ·Lid to ·B-lid ; ·Rid to ·B-rid + ; ·Assoc to ·B-assoc) + + private + f = fst φ + open IsRingHom (φ .snd) + + RingHomRespInv : (r : A) ⦃ r∈Aˣ : r ∈ Aˣ ⦄ → f r ∈ Bˣ + RingHomRespInv r = f (r ⁻¹ᵃ) , (sym (pres· r (r ⁻¹ᵃ)) ∙∙ cong (f) (·A-rinv r) ∙∙ pres1) + + φ[x⁻¹]≡φ[x]⁻¹ : (r : A) ⦃ r∈Aˣ : r ∈ Aˣ ⦄ ⦃ φr∈Bˣ : f r ∈ Bˣ ⦄ + → f (r ⁻¹ᵃ) ≡ (f r) ⁻¹ᵇ + φ[x⁻¹]≡φ[x]⁻¹ r ⦃ r∈Aˣ ⦄ ⦃ φr∈Bˣ ⦄ = + f (r ⁻¹ᵃ) ≡⟨ sym (·B-rid _) ⟩ + f (r ⁻¹ᵃ) ·B 1b ≡⟨ cong (f (r ⁻¹ᵃ) ·B_) (sym (·B-rinv _)) ⟩ + f (r ⁻¹ᵃ) ·B ((f r) ·B (f r) ⁻¹ᵇ) ≡⟨ ·B-assoc _ _ _ ⟩ + f (r ⁻¹ᵃ) ·B (f r) ·B (f r) ⁻¹ᵇ ≡⟨ cong (_·B (f r) ⁻¹ᵇ) (sym (pres· _ _)) ⟩ + f (r ⁻¹ᵃ ·A r) ·B (f r) ⁻¹ᵇ ≡⟨ cong (λ x → f x ·B (f r) ⁻¹ᵇ) (·A-linv _) ⟩ + f 1a ·B (f r) ⁻¹ᵇ ≡⟨ cong (_·B (f r) ⁻¹ᵇ) (pres1) ⟩ + 1b ·B (f r) ⁻¹ᵇ ≡⟨ ·B-lid _ ⟩ + (f r) ⁻¹ᵇ ∎ + + +module Exponentiation (R' : CommRing ℓ) where + open CommRingStr (snd R') + private R = fst R' + + -- introduce exponentiation + _^_ : R → ℕ → R + f ^ zero = 1r + f ^ suc n = f · (f ^ n) + + infix 9 _^_ + + -- and prove some laws + 1ⁿ≡1 : (n : ℕ) → 1r ^ n ≡ 1r + 1ⁿ≡1 zero = refl + 1ⁿ≡1 (suc n) = ·Lid _ ∙ 1ⁿ≡1 n + + ·-of-^-is-^-of-+ : (f : R) (m n : ℕ) → (f ^ m) · (f ^ n) ≡ f ^ (m +ℕ n) + ·-of-^-is-^-of-+ f zero n = ·Lid _ + ·-of-^-is-^-of-+ f (suc m) n = sym (·Assoc _ _ _) ∙ cong (f ·_) (·-of-^-is-^-of-+ f m n) + + ^-ldist-· : (f g : R) (n : ℕ) → (f · g) ^ n ≡ (f ^ n) · (g ^ n) + ^-ldist-· f g zero = sym (·Lid 1r) + ^-ldist-· f g (suc n) = path + where + path : f · g · ((f · g) ^ n) ≡ f · (f ^ n) · (g · (g ^ n)) + path = f · g · ((f · g) ^ n) ≡⟨ cong (f · g ·_) (^-ldist-· f g n) ⟩ + f · g · ((f ^ n) · (g ^ n)) ≡⟨ ·Assoc _ _ _ ⟩ + f · g · (f ^ n) · (g ^ n) ≡⟨ cong (_· (g ^ n)) (sym (·Assoc _ _ _)) ⟩ + f · (g · (f ^ n)) · (g ^ n) ≡⟨ cong (λ r → (f · r) · (g ^ n)) (·Comm _ _) ⟩ + f · ((f ^ n) · g) · (g ^ n) ≡⟨ cong (_· (g ^ n)) (·Assoc _ _ _) ⟩ + f · (f ^ n) · g · (g ^ n) ≡⟨ sym (·Assoc _ _ _) ⟩ + f · (f ^ n) · (g · (g ^ n)) ∎ + + ^-rdist-·ℕ : (f : R) (n m : ℕ) → f ^ (n ·ℕ m) ≡ (f ^ n) ^ m + ^-rdist-·ℕ f zero m = sym (1ⁿ≡1 m) + ^-rdist-·ℕ f (suc n) m = sym (·-of-^-is-^-of-+ f m (n ·ℕ m)) + ∙∙ cong (f ^ m ·_) (^-rdist-·ℕ f n m) + ∙∙ sym (^-ldist-· f (f ^ n) m) + + -- interaction of exponentiation and units + open Units R' + + ^-presUnit : ∀ (f : R) (n : ℕ) → f ∈ Rˣ → f ^ n ∈ Rˣ + ^-presUnit f zero f∈Rˣ = RˣContainsOne + ^-presUnit f (suc n) f∈Rˣ = RˣMultClosed f (f ^ n) ⦃ f∈Rˣ ⦄ ⦃ ^-presUnit f n f∈Rˣ ⦄ + + +-- like in Ring.Properties we provide helpful lemmas here +module CommRingTheory (R' : CommRing ℓ) where + open CommRingStr (snd R') + private R = fst R' + + ·CommAssocl : (x y z : R) → x · (y · z) ≡ y · (x · z) + ·CommAssocl x y z = ·Assoc x y z ∙∙ cong (_· z) (·Comm x y) ∙∙ sym (·Assoc y x z) + + ·CommAssocr : (x y z : R) → x · y · z ≡ x · z · y + ·CommAssocr x y z = sym (·Assoc x y z) ∙∙ cong (x ·_) (·Comm y z) ∙∙ ·Assoc x z y + + ·CommAssocr2 : (x y z : R) → x · y · z ≡ z · y · x + ·CommAssocr2 x y z = ·CommAssocr _ _ _ ∙∙ cong (_· y) (·Comm _ _) ∙∙ ·CommAssocr _ _ _ + + ·CommAssocSwap : (x y z w : R) → (x · y) · (z · w) ≡ (x · z) · (y · w) + ·CommAssocSwap x y z w = + ·Assoc (x · y) z w ∙∙ cong (_· w) (·CommAssocr x y z) ∙∙ sym (·Assoc (x · z) y w) + + + +-- the CommRing version of uaCompEquiv +module CommRingUAFunctoriality where + open CommRingStr + open CommRingEquivs + + CommRing≡ : (A B : CommRing ℓ) → ( + Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] + Σ[ q0 ∈ PathP (λ i → p i) (0r (snd A)) (0r (snd B)) ] + Σ[ q1 ∈ PathP (λ i → p i) (1r (snd A)) (1r (snd B)) ] + Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] + Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] + Σ[ s ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] + PathP (λ i → IsCommRing (q0 i) (q1 i) (r+ i) (r· i) (s i)) (isCommRing (snd A)) (isCommRing (snd B))) + ≃ (A ≡ B) + CommRing≡ A B = isoToEquiv theIso + where + open Iso + theIso : Iso _ _ + fun theIso (p , q0 , q1 , r+ , r· , s , t) i = p i + , commringstr (q0 i) (q1 i) (r+ i) (r· i) (s i) (t i) + inv theIso x = cong ⟨_⟩ x , cong (0r ∘ snd) x , cong (1r ∘ snd) x + , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (isCommRing ∘ snd) x + rightInv theIso _ = refl + leftInv theIso _ = refl + + caracCommRing≡ : {A B : CommRing ℓ} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q + caracCommRing≡ {A = A} {B = B} p q P = + sym (transportTransport⁻ (ua (CommRing≡ A B)) p) + ∙∙ cong (transport (ua (CommRing≡ A B))) helper + ∙∙ transportTransport⁻ (ua (CommRing≡ A B)) q + where + helper : transport (sym (ua (CommRing≡ A B))) p ≡ transport (sym (ua (CommRing≡ A B))) q + helper = Σ≡Prop + (λ _ → isPropΣ + (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) + λ _ → isOfHLevelPathP 1 (isPropIsCommRing _ _ _ _ _) _ _) + (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + + uaCompCommRingEquiv : {A B C : CommRing ℓ} (f : CommRingEquiv A B) (g : CommRingEquiv B C) + → uaCommRing (compCommRingEquiv f g) ≡ uaCommRing f ∙ uaCommRing g + uaCompCommRingEquiv f g = caracCommRing≡ _ _ ( + cong ⟨_⟩ (uaCommRing (compCommRingEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaCommRing f) ∙ cong ⟨_⟩ (uaCommRing g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaCommRing f) (uaCommRing g)) ⟩ + cong ⟨_⟩ (uaCommRing f ∙ uaCommRing g) ∎) + + + +open CommRingEquivs +open CommRingUAFunctoriality +recPT→CommRing : {A : Type ℓ'} (𝓕 : A → CommRing ℓ) + → (σ : ∀ x y → CommRingEquiv (𝓕 x) (𝓕 y)) + → (∀ x y z → σ x z ≡ compCommRingEquiv (σ x y) (σ y z)) + ------------------------------------------------------ + → ∥ A ∥ → CommRing ℓ +recPT→CommRing 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommRing 𝓕 + (3-ConstantCompChar 𝓕 (λ x y → uaCommRing (σ x y)) + λ x y z → sym ( cong uaCommRing (compCoh x y z) + ∙ uaCompCommRingEquiv (σ x y) (σ y z))) diff --git a/Cubical/Algebra/CommRing/QuotientRing.agda b/Cubical/Algebra/CommRing/QuotientRing.agda new file mode 100644 index 0000000000..71a133250d --- /dev/null +++ b/Cubical/Algebra/CommRing/QuotientRing.agda @@ -0,0 +1,31 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.QuotientRing where + +open import Cubical.Foundations.Prelude + +open import Cubical.HITs.SetQuotients hiding (_/_) + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Ring.QuotientRing renaming (_/_ to _/Ring_) hiding (asRing) + +private + variable + ℓ : Level + +_/_ : (R : CommRing ℓ) → (I : IdealsIn R) → CommRing ℓ +R / I = + fst asRing , commringstr _ _ _ _ _ + (iscommring (RingStr.isRing (snd asRing)) + (elimProp2 (λ _ _ → squash/ _ _) + commEq)) + where + asRing = (CommRing→Ring R) /Ring I + _·/_ : fst asRing → fst asRing → fst asRing + _·/_ = RingStr._·_ (snd asRing) + commEq : (x y : fst R) → ([ x ] ·/ [ y ]) ≡ ([ y ] ·/ [ x ]) + commEq x y i = [ CommRingStr.·Comm (snd R) x y i ] + +[_]/ : {R : CommRing ℓ} {I : IdealsIn R} → (a : fst R) → fst (R / I) +[ a ]/ = [ a ] diff --git a/Cubical/Algebra/CommRing/RadicalIdeal.agda b/Cubical/Algebra/CommRing/RadicalIdeal.agda new file mode 100644 index 0000000000..6d4740cd3b --- /dev/null +++ b/Cubical/Algebra/CommRing/RadicalIdeal.agda @@ -0,0 +1,235 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommRing.RadicalIdeal where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Powerset using (⊆-isProp) +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma +open import Cubical.Data.Sum hiding (map) +open import Cubical.Data.FinData hiding (elim) +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm + ; _choose_ to _ℕchoose_ ; snotz to ℕsnotz) +open import Cubical.Data.Nat.Order + +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.BinomialThm +open import Cubical.Algebra.Ring.Properties +open import Cubical.Algebra.Ring.BigOps +open import Cubical.Algebra.RingSolver.Reflection + +private + variable + ℓ : Level + +module RadicalIdeal (R' : CommRing ℓ) where + private R = fst R' + open CommRingStr (snd R') + open RingTheory (CommRing→Ring R') + open Sum (CommRing→Ring R') + open CommRingTheory R' + open Exponentiation R' + open BinomialThm R' + open CommIdeal R' + open isCommIdeal + + + √ : CommIdeal → CommIdeal + fst (√ I) x = (∃[ n ∈ ℕ ] x ^ n ∈ I) , isPropPropTrunc + +Closed (snd (√ I)) {x = x} {y = y} = map2 +ClosedΣ + where + +ClosedΣ : Σ[ n ∈ ℕ ] x ^ n ∈ I → Σ[ n ∈ ℕ ] y ^ n ∈ I → Σ[ n ∈ ℕ ] (x + y) ^ n ∈ I + +ClosedΣ (n , xⁿ∈I) (m , yᵐ∈I) = (n +ℕ m) + , subst-∈ I (sym (BinomialThm (n +ℕ m) _ _)) ∑Binomial∈I + where + binomialCoeff∈I : ∀ i → ((n +ℕ m) choose toℕ i) · x ^ toℕ i · y ^ (n +ℕ m ∸ toℕ i) ∈ I + binomialCoeff∈I i with ≤-+-split n m (toℕ i) (pred-≤-pred (toℕ>= normalise + just (varInfos , equation) ← returnTC (getVarsAndEquation hole′) + where + nothing + → typeError (strErr "Something went wrong when getting the variable names in " + ∷ termErr hole′ ∷ []) + just (lhs , rhs) ← returnTC (toAlgebraExpression (getArgs equation)) + where + nothing + → typeError( + strErr "Error while trying to build ASTs for the equation " ∷ + termErr equation ∷ []) + let solution = solverCallWithLambdas (length varInfos) varInfos lhs rhs + unify hole solution + +macro + solve : Term → TC _ + solve = solve-macro diff --git a/Cubical/Algebra/NatSolver/Solver.agda b/Cubical/Algebra/NatSolver/Solver.agda new file mode 100644 index 0000000000..fa15aab9a4 --- /dev/null +++ b/Cubical/Algebra/NatSolver/Solver.agda @@ -0,0 +1,123 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.NatSolver.Solver where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order using (zero-≤) +open import Cubical.Data.Vec.Base +open import Cubical.Algebra.NatSolver.NatExpression +open import Cubical.Algebra.NatSolver.HornerForms +open import Cubical.Algebra.NatSolver.EvalHom + +private + variable + ℓ : Level + +module EqualityToNormalform where + open Eval + open IteratedHornerOperations + open HomomorphismProperties + + normalize : {n : ℕ} → Expr n → IteratedHornerForms n + normalize {n = n} (K r) = Constant n r + normalize {n = n} (∣ k) = Variable n k + normalize (x +' y) = + (normalize x) +ₕ (normalize y) + normalize (x ·' y) = + (normalize x) ·ₕ (normalize y) + + isEqualToNormalform : + {n : ℕ} + (e : Expr n) (xs : Vec ℕ n) + → eval (normalize e) xs ≡ ⟦ e ⟧ xs + isEqualToNormalform (K r) [] = refl + isEqualToNormalform {n = ℕ.suc n} (K r) (x ∷ xs) = + eval (Constant (ℕ.suc n) r) (x ∷ xs) ≡⟨ refl ⟩ + eval (0ₕ ·X+ Constant n r) (x ∷ xs) ≡⟨ refl ⟩ + eval 0ₕ (x ∷ xs) · x + eval (Constant n r) xs + ≡⟨ cong (λ u → u · x + eval (Constant n r) xs) (eval0H (x ∷ xs)) ⟩ + 0 · x + eval (Constant n r) xs + ≡⟨ refl ⟩ + eval (Constant n r) xs + ≡⟨ isEqualToNormalform (K r) xs ⟩ + r ∎ + + isEqualToNormalform (∣ zero) (x ∷ xs) = + eval (1ₕ ·X+ 0ₕ) (x ∷ xs) ≡⟨ refl ⟩ + eval 1ₕ (x ∷ xs) · x + eval 0ₕ xs ≡⟨ cong (λ u → u · x + eval 0ₕ xs) + (eval1ₕ (x ∷ xs)) ⟩ + 1 · x + eval 0ₕ xs ≡⟨ cong (λ u → 1 · x + u ) (eval0H xs) ⟩ + 1 · x + 0 ≡⟨ +-zero _ ⟩ + 1 · x ≡⟨ ·-identityˡ _ ⟩ + x ∎ + isEqualToNormalform {n = ℕ.suc n} (∣ (suc k)) (x ∷ xs) = + eval (0ₕ ·X+ Variable n k) (x ∷ xs) ≡⟨ refl ⟩ + eval 0ₕ (x ∷ xs) · x + eval (Variable n k) xs + ≡⟨ cong (λ u → u · x + eval (Variable n k) xs) (eval0H (x ∷ xs)) ⟩ + 0 · x + eval (Variable n k) xs + ≡⟨ refl ⟩ + eval (Variable n k) xs + ≡⟨ isEqualToNormalform (∣ k) xs ⟩ + ⟦ ∣ (suc k) ⟧ (x ∷ xs) ∎ + + isEqualToNormalform (e +' e₁) [] = + eval (normalize e +ₕ normalize e₁) [] + ≡⟨ +Homeval (normalize e) _ [] ⟩ + eval (normalize e) [] + + eval (normalize e₁) [] + ≡⟨ cong (λ u → u + eval (normalize e₁) []) + (isEqualToNormalform e []) ⟩ + ⟦ e ⟧ [] + + eval (normalize e₁) [] + ≡⟨ cong (λ u → ⟦ e ⟧ [] + u) (isEqualToNormalform e₁ []) ⟩ + ⟦ e ⟧ [] + ⟦ e₁ ⟧ [] ∎ + isEqualToNormalform (e +' e₁) (x ∷ xs) = + eval (normalize e + +ₕ normalize e₁) (x ∷ xs) + ≡⟨ +Homeval (normalize e) _ (x ∷ xs) ⟩ + eval (normalize e) (x ∷ xs) + + eval (normalize e₁) (x ∷ xs) + ≡⟨ cong (λ u → u + eval (normalize e₁) (x ∷ xs)) + (isEqualToNormalform e (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) + + eval (normalize e₁) (x ∷ xs) + ≡⟨ cong (λ u → ⟦ e ⟧ (x ∷ xs) + u) + (isEqualToNormalform e₁ (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) + ⟦ e₁ ⟧ (x ∷ xs) ∎ + + isEqualToNormalform (e ·' e₁) [] = + eval (normalize e ·ₕ normalize e₁) [] + ≡⟨ ·Homeval (normalize e) _ [] ⟩ + eval (normalize e) [] + · eval (normalize e₁) [] + ≡⟨ cong (λ u → u · eval (normalize e₁) []) + (isEqualToNormalform e []) ⟩ + ⟦ e ⟧ [] + · eval (normalize e₁) [] + ≡⟨ cong (λ u → ⟦ e ⟧ [] · u) (isEqualToNormalform e₁ []) ⟩ + ⟦ e ⟧ [] · ⟦ e₁ ⟧ [] ∎ + + isEqualToNormalform (e ·' e₁) (x ∷ xs) = + eval (normalize e ·ₕ normalize e₁) (x ∷ xs) + ≡⟨ ·Homeval (normalize e) _ (x ∷ xs) ⟩ + eval (normalize e) (x ∷ xs) + · eval (normalize e₁) (x ∷ xs) + ≡⟨ cong (λ u → u · eval (normalize e₁) (x ∷ xs)) + (isEqualToNormalform e (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) + · eval (normalize e₁) (x ∷ xs) + ≡⟨ cong (λ u → ⟦ e ⟧ (x ∷ xs) · u) + (isEqualToNormalform e₁ (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) · ⟦ e₁ ⟧ (x ∷ xs) ∎ + + solve : + {n : ℕ} (e₁ e₂ : Expr n) (xs : Vec ℕ n) + (p : eval (normalize e₁) xs ≡ eval (normalize e₂) xs) + → ⟦ e₁ ⟧ xs ≡ ⟦ e₂ ⟧ xs + solve e₁ e₂ xs p = + ⟦ e₁ ⟧ xs ≡⟨ sym (isEqualToNormalform e₁ xs) ⟩ + eval (normalize e₁) xs ≡⟨ p ⟩ + eval (normalize e₂) xs ≡⟨ isEqualToNormalform e₂ xs ⟩ + ⟦ e₂ ⟧ xs ∎ diff --git a/Cubical/Algebra/Ring.agda b/Cubical/Algebra/Ring.agda new file mode 100644 index 0000000000..ca6659831b --- /dev/null +++ b/Cubical/Algebra/Ring.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Ring where + +open import Cubical.Algebra.Ring.Base public +open import Cubical.Algebra.Ring.Properties public diff --git a/Cubical/Algebra/Ring/Base.agda b/Cubical/Algebra/Ring/Base.agda new file mode 100644 index 0000000000..4c2a00f49f --- /dev/null +++ b/Cubical/Algebra/Ring/Base.agda @@ -0,0 +1,276 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Ring.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Group +open import Cubical.Algebra.AbGroup + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Reflection.RecordEquiv + +open Iso + +private + variable + ℓ ℓ' ℓ'' : Level + +record IsRing {R : Type ℓ} + (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) : Type ℓ where + + constructor isring + + field + +IsAbGroup : IsAbGroup 0r _+_ -_ + ·IsMonoid : IsMonoid 1r _·_ + dist : (x y z : R) → (x · (y + z) ≡ (x · y) + (x · z)) + × ((x + y) · z ≡ (x · z) + (y · z)) + -- This is in the Agda stdlib, but it's redundant + -- zero : (x : R) → (x · 0r ≡ 0r) × (0r · x ≡ 0r) + + open IsAbGroup +IsAbGroup public + renaming + ( assoc to +Assoc + ; identity to +Identity + ; lid to +Lid + ; rid to +Rid + ; inverse to +Inv + ; invl to +Linv + ; invr to +Rinv + ; comm to +Comm + ; isSemigroup to +IsSemigroup + ; isMonoid to +IsMonoid + ; isGroup to +IsGroup ) + + open IsMonoid ·IsMonoid public + renaming + ( assoc to ·Assoc + ; identity to ·Identity + ; lid to ·Lid + ; rid to ·Rid + ; isSemigroup to ·IsSemigroup ) + hiding + ( is-set ) -- We only want to export one proof of this + + ·Rdist+ : (x y z : R) → x · (y + z) ≡ (x · y) + (x · z) + ·Rdist+ x y z = dist x y z .fst + + ·Ldist+ : (x y z : R) → (x + y) · z ≡ (x · z) + (y · z) + ·Ldist+ x y z = dist x y z .snd + +record RingStr (A : Type ℓ) : Type (ℓ-suc ℓ) where + + constructor ringstr + + field + 0r : A + 1r : A + _+_ : A → A → A + _·_ : A → A → A + -_ : A → A + isRing : IsRing 0r 1r _+_ _·_ -_ + + infix 8 -_ + infixl 7 _·_ + infixl 6 _+_ + + open IsRing isRing public + +Ring : ∀ ℓ → Type (ℓ-suc ℓ) +Ring ℓ = TypeWithStr ℓ RingStr + +isSetRing : (R : Ring ℓ) → isSet ⟨ R ⟩ +isSetRing R = R .snd .RingStr.isRing .IsRing.·IsMonoid .IsMonoid.isSemigroup .IsSemigroup.is-set + +makeIsRing : {R : Type ℓ} {0r 1r : R} {_+_ _·_ : R → R → R} { -_ : R → R} + (is-setR : isSet R) + (+-assoc : (x y z : R) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : R) → x + 0r ≡ x) + (+-rinv : (x : R) → x + (- x) ≡ 0r) + (+-comm : (x y : R) → x + y ≡ y + x) + (r+-assoc : (x y z : R) → x · (y · z) ≡ (x · y) · z) + (·-rid : (x : R) → x · 1r ≡ x) + (·-lid : (x : R) → 1r · x ≡ x) + (·-rdist-+ : (x y z : R) → x · (y + z) ≡ (x · y) + (x · z)) + (·-ldist-+ : (x y z : R) → (x + y) · z ≡ (x · z) + (y · z)) + → IsRing 0r 1r _+_ _·_ -_ +makeIsRing is-setR assoc +-rid +-rinv +-comm ·-assoc ·-rid ·-lid ·-rdist-+ ·-ldist-+ = + isring (makeIsAbGroup is-setR assoc +-rid +-rinv +-comm) + (makeIsMonoid is-setR ·-assoc ·-rid ·-lid) + λ x y z → ·-rdist-+ x y z , ·-ldist-+ x y z + +makeRing : {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) + (is-setR : isSet R) + (+-assoc : (x y z : R) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : R) → x + 0r ≡ x) + (+-rinv : (x : R) → x + (- x) ≡ 0r) + (+-comm : (x y : R) → x + y ≡ y + x) + (+-assoc : (x y z : R) → x · (y · z) ≡ (x · y) · z) + (·-rid : (x : R) → x · 1r ≡ x) + (·-lid : (x : R) → 1r · x ≡ x) + (·-rdist-+ : (x y z : R) → x · (y + z) ≡ (x · y) + (x · z)) + (·-ldist-+ : (x y z : R) → (x + y) · z ≡ (x · z) + (y · z)) + → Ring ℓ +makeRing 0r 1r _+_ _·_ -_ is-setR assoc +-rid +-rinv +-comm ·-assoc ·-rid ·-lid ·-rdist-+ ·-ldist-+ = + _ , ringstr 0r 1r _+_ _·_ -_ + (makeIsRing is-setR assoc +-rid +-rinv +-comm + ·-assoc ·-rid ·-lid ·-rdist-+ ·-ldist-+ ) + +record IsRingHom {A : Type ℓ} {B : Type ℓ'} (R : RingStr A) (f : A → B) (S : RingStr B) + : Type (ℓ-max ℓ ℓ') + where + + -- Shorter qualified names + private + module R = RingStr R + module S = RingStr S + + field + pres0 : f R.0r ≡ S.0r + pres1 : f R.1r ≡ S.1r + pres+ : (x y : A) → f (x R.+ y) ≡ f x S.+ f y + pres· : (x y : A) → f (x R.· y) ≡ f x S.· f y + pres- : (x : A) → f (R.- x) ≡ S.- (f x) + +unquoteDecl IsRingHomIsoΣ = declareRecordIsoΣ IsRingHomIsoΣ (quote IsRingHom) + +RingHom : (R : Ring ℓ) (S : Ring ℓ') → Type (ℓ-max ℓ ℓ') +RingHom R S = Σ[ f ∈ (⟨ R ⟩ → ⟨ S ⟩) ] IsRingHom (R .snd) f (S .snd) + +IsRingEquiv : {A : Type ℓ} {B : Type ℓ'} (M : RingStr A) (e : A ≃ B) (N : RingStr B) + → Type (ℓ-max ℓ ℓ') +IsRingEquiv M e N = IsRingHom M (e .fst) N + +RingEquiv : (R : Ring ℓ) (S : Ring ℓ') → Type (ℓ-max ℓ ℓ') +RingEquiv R S = Σ[ e ∈ (⟨ R ⟩ ≃ ⟨ S ⟩) ] IsRingEquiv (R .snd) e (S .snd) + +_$_ : {R S : Ring ℓ} → (φ : RingHom R S) → (x : ⟨ R ⟩) → ⟨ S ⟩ +φ $ x = φ .fst x + +RingEquiv→RingHom : {A B : Ring ℓ} → RingEquiv A B → RingHom A B +RingEquiv→RingHom (e , eIsHom) = e .fst , eIsHom + +isPropIsRing : {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) + → isProp (IsRing 0r 1r _+_ _·_ -_) +isPropIsRing 0r 1r _+_ _·_ -_ (isring RG RM RD) (isring SG SM SD) = + λ i → isring (isPropIsAbGroup _ _ _ RG SG i) + (isPropIsMonoid _ _ RM SM i) + (isPropDistr RD SD i) + where + isSetR : isSet _ + isSetR = RM .IsMonoid.isSemigroup .IsSemigroup.is-set + + isPropDistr : isProp ((x y z : _) → ((x · (y + z)) ≡ ((x · y) + (x · z))) + × (((x + y) · z) ≡ ((x · z) + (y · z)))) + isPropDistr = isPropΠ3 λ _ _ _ → isProp× (isSetR _ _) (isSetR _ _) + +isPropIsRingHom : {A : Type ℓ} {B : Type ℓ'} (R : RingStr A) (f : A → B) (S : RingStr B) + → isProp (IsRingHom R f S) +isPropIsRingHom R f S = isOfHLevelRetractFromIso 1 IsRingHomIsoΣ + (isProp×4 (isSetRing (_ , S) _ _) + (isSetRing (_ , S) _ _) + (isPropΠ2 λ _ _ → isSetRing (_ , S) _ _) + (isPropΠ2 λ _ _ → isSetRing (_ , S) _ _) + (isPropΠ λ _ → isSetRing (_ , S) _ _)) + +isSetRingHom : (R : Ring ℓ) (S : Ring ℓ') → isSet (RingHom R S) +isSetRingHom R S = isSetΣSndProp (isSetΠ (λ _ → isSetRing S)) (λ f → isPropIsRingHom (snd R) f (snd S)) + +isSetRingEquiv : (A : Ring ℓ) (B : Ring ℓ') → isSet (RingEquiv A B) +isSetRingEquiv A B = isSetΣSndProp (isOfHLevel≃ 2 (isSetRing A) (isSetRing B)) + (λ e → isPropIsRingHom (snd A) (fst e) (snd B)) + +RingHomPathP : (R S T : Ring ℓ) (p : S ≡ T) (φ : RingHom R S) (ψ : RingHom R T) + → PathP (λ i → R .fst → p i .fst) (φ .fst) (ψ .fst) + → PathP (λ i → RingHom R (p i)) φ ψ +RingHomPathP R S T p φ ψ q = ΣPathP (q , isProp→PathP (λ _ → isPropIsRingHom _ _ _) _ _) + +RingHom≡ : {R S : Ring ℓ} {φ ψ : RingHom R S} → fst φ ≡ fst ψ → φ ≡ ψ +RingHom≡ = Σ≡Prop λ f → isPropIsRingHom _ f _ + +𝒮ᴰ-Ring : DUARel (𝒮-Univ ℓ) RingStr ℓ +𝒮ᴰ-Ring = + 𝒮ᴰ-Record (𝒮-Univ _) IsRingEquiv + (fields: + data[ 0r ∣ null ∣ pres0 ] + data[ 1r ∣ null ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ un ∣ pres- ] + prop[ isRing ∣ (λ _ _ → isPropIsRing _ _ _ _ _) ]) + where + open RingStr + open IsRingHom + + -- faster with some sharing + null = autoDUARel (𝒮-Univ _) (λ A → A) + un = autoDUARel (𝒮-Univ _) (λ A → A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +RingPath : (R S : Ring ℓ) → RingEquiv R S ≃ (R ≡ S) +RingPath = ∫ 𝒮ᴰ-Ring .UARel.ua + +uaRing : {A B : Ring ℓ} → RingEquiv A B → A ≡ B +uaRing {A = A} {B = B} = equivFun (RingPath A B) + +isGroupoidRing : isGroupoid (Ring ℓ) +isGroupoidRing _ _ = isOfHLevelRespectEquiv 2 (RingPath _ _) (isSetRingEquiv _ _) + + +-- Rings have an abelian group and a monoid + +Ring→AbGroup : Ring ℓ → AbGroup ℓ +Ring→AbGroup (A , ringstr _ _ _ _ _ R) = A , abgroupstr _ _ _ (IsRing.+IsAbGroup R) + +Ring→Group : Ring ℓ → Group ℓ +Ring→Group = AbGroup→Group ∘ Ring→AbGroup + +Ring→AddMonoid : Ring ℓ → Monoid ℓ +Ring→AddMonoid = Group→Monoid ∘ Ring→Group + +Ring→MultMonoid : Ring ℓ → Monoid ℓ +Ring→MultMonoid (A , ringstr _ _ _ _ _ R) = monoid _ _ _ (IsRing.·IsMonoid R) + + +-- Smart constructor for ring homomorphisms +-- that infers the other equations from pres1, pres+, and pres· + +module _ {R : Ring ℓ} {S : Ring ℓ'} {f : ⟨ R ⟩ → ⟨ S ⟩} where + + private + module R = RingStr (R .snd) + module S = RingStr (S .snd) + + module _ + (p1 : f R.1r ≡ S.1r) + (p+ : (x y : ⟨ R ⟩) → f (x R.+ y) ≡ f x S.+ f y) + (p· : (x y : ⟨ R ⟩) → f (x R.· y) ≡ f x S.· f y) + where + + open IsRingHom + private + isGHom : IsGroupHom (Ring→Group R .snd) f (Ring→Group S .snd) + isGHom = makeIsGroupHom p+ + + makeIsRingHom : IsRingHom (R .snd) f (S .snd) + makeIsRingHom .pres0 = isGHom .IsGroupHom.pres1 + makeIsRingHom .pres1 = p1 + makeIsRingHom .pres+ = p+ + makeIsRingHom .pres· = p· + makeIsRingHom .pres- = isGHom .IsGroupHom.presinv diff --git a/Cubical/Algebra/Ring/BigOps.agda b/Cubical/Algebra/Ring/BigOps.agda new file mode 100644 index 0000000000..252eff1c3b --- /dev/null +++ b/Cubical/Algebra/Ring/BigOps.agda @@ -0,0 +1,119 @@ +-- define ∑ and ∏ as the bigOps of a Ring when interpreted +-- as an additive/multiplicative monoid + +{-# OPTIONS --safe #-} +module Cubical.Algebra.Ring.BigOps where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ ; zero ; suc) +open import Cubical.Data.FinData +open import Cubical.Data.Bool + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Auto +open import Cubical.Structures.Macro +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Monoid.BigOp +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Ring.Base +open import Cubical.Algebra.Ring.Properties + +private + variable + ℓ : Level + +module KroneckerDelta (R' : Ring ℓ) where + private + R = fst R' + open RingStr (snd R') + + δ : {n : ℕ} (i j : Fin n) → R + δ i j = if i == j then 1r else 0r + + + +module Sum (R' : Ring ℓ) where + private + R = fst R' + open RingStr (snd R') + open MonoidBigOp (Ring→AddMonoid R') + open RingTheory R' + open KroneckerDelta R' + + ∑ = bigOp + ∑Ext = bigOpExt + ∑0r = bigOpε + ∑Last = bigOpLast + + ∑Split : ∀ {n} → (V W : FinVec R n) → ∑ (λ i → V i + W i) ≡ ∑ V + ∑ W + ∑Split = bigOpSplit +Comm + + ∑Split++ : ∀ {n m : ℕ} (V : FinVec R n) (W : FinVec R m) + → ∑ (V ++Fin W) ≡ ∑ V + ∑ W + ∑Split++ = bigOpSplit++ +Comm + + ∑Mulrdist : ∀ {n} → (x : R) → (V : FinVec R n) + → x · ∑ V ≡ ∑ λ i → x · V i + ∑Mulrdist {n = zero} x _ = 0RightAnnihilates x + ∑Mulrdist {n = suc n} x V = + x · (V zero + ∑ (V ∘ suc)) ≡⟨ ·Rdist+ _ _ _ ⟩ + x · V zero + x · ∑ (V ∘ suc) ≡⟨ (λ i → x · V zero + ∑Mulrdist x (V ∘ suc) i) ⟩ + x · V zero + ∑ (λ i → x · V (suc i)) ∎ + + ∑Mulldist : ∀ {n} → (x : R) → (V : FinVec R n) + → (∑ V) · x ≡ ∑ λ i → V i · x + ∑Mulldist {n = zero} x _ = 0LeftAnnihilates x + ∑Mulldist {n = suc n} x V = + (V zero + ∑ (V ∘ suc)) · x ≡⟨ ·Ldist+ _ _ _ ⟩ + V zero · x + (∑ (V ∘ suc)) · x ≡⟨ (λ i → V zero · x + ∑Mulldist x (V ∘ suc) i) ⟩ + V zero · x + ∑ (λ i → V (suc i) · x) ∎ + + ∑Mulr0 : ∀ {n} → (V : FinVec R n) → ∑ (λ i → V i · 0r) ≡ 0r + ∑Mulr0 V = sym (∑Mulldist 0r V) ∙ 0RightAnnihilates _ + + ∑Mul0r : ∀ {n} → (V : FinVec R n) → ∑ (λ i → 0r · V i) ≡ 0r + ∑Mul0r V = sym (∑Mulrdist 0r V) ∙ 0LeftAnnihilates _ + + ∑Mulr1 : (n : ℕ) (V : FinVec R n) → (j : Fin n) → ∑ (λ i → V i · δ i j) ≡ V j + ∑Mulr1 (suc n) V zero = (λ k → ·Rid (V zero) k + ∑Mulr0 (V ∘ suc) k) ∙ +Rid (V zero) + ∑Mulr1 (suc n) V (suc j) = + (λ i → 0RightAnnihilates (V zero) i + ∑ (λ x → V (suc x) · δ x j)) + ∙∙ +Lid _ ∙∙ ∑Mulr1 n (V ∘ suc) j + + ∑Mul1r : (n : ℕ) (V : FinVec R n) → (j : Fin n) → ∑ (λ i → (δ j i) · V i) ≡ V j + ∑Mul1r (suc n) V zero = (λ k → ·Lid (V zero) k + ∑Mul0r (V ∘ suc) k) ∙ +Rid (V zero) + ∑Mul1r (suc n) V (suc j) = + (λ i → 0LeftAnnihilates (V zero) i + ∑ (λ i → (δ j i) · V (suc i))) + ∙∙ +Lid _ ∙∙ ∑Mul1r n (V ∘ suc) j + + ∑Dist- : ∀ {n : ℕ} (V : FinVec R n) → ∑ (λ i → - V i) ≡ - ∑ V + ∑Dist- V = ∑Ext (λ i → -IsMult-1 (V i)) ∙ sym (∑Mulrdist _ V) ∙ sym (-IsMult-1 _) + + +-- anything interesting to prove here? +module Product (R' : Ring ℓ) where + private + R = fst R' + open RingStr (snd R') + open RingTheory R' + open MonoidBigOp (Ring→MultMonoid R') + + ∏ = bigOp + ∏Ext = bigOpExt + ∏0r = bigOpε + ∏Last = bigOpLast + + -- only holds in CommRings! + -- ∏Split : ∀ {n} → (V W : FinVec R n) → ∏ (λ i → V i · W i) ≡ ∏ V · ∏ W + -- ∏Split = bigOpSplit MultR' ·-comm diff --git a/Cubical/Algebra/Ring/Ideal.agda b/Cubical/Algebra/Ring/Ideal.agda new file mode 100644 index 0000000000..bf49c65143 --- /dev/null +++ b/Cubical/Algebra/Ring/Ideal.agda @@ -0,0 +1,81 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Ring.Ideal where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset + +open import Cubical.Algebra.Ring.Base +open import Cubical.Algebra.Ring.Properties + + +private + variable + ℓ : Level + +module _ (R' : Ring ℓ) where + + open RingStr (snd R') + private R = ⟨ R' ⟩ + + {- by default, 'ideal' means two-sided ideal -} + record isIdeal (I : ℙ R) : Type ℓ where + field + +-closed : {x y : R} → x ∈ I → y ∈ I → x + y ∈ I + -closed : {x : R} → x ∈ I → - x ∈ I + 0r-closed : 0r ∈ I + ·-closedLeft : {x : R} → (r : R) → x ∈ I → r · x ∈ I + ·-closedRight : {x : R} → (r : R) → x ∈ I → x · r ∈ I + + Ideal : Type _ + Ideal = Σ[ I ∈ ℙ R ] isIdeal I + + record isLeftIdeal (I : ℙ R) : Type ℓ where + field + +-closed : {x y : R} → x ∈ I → y ∈ I → x + y ∈ I + -closed : {x : R} → x ∈ I → - x ∈ I + 0r-closed : 0r ∈ I + ·-closedLeft : {x : R} → (r : R) → x ∈ I → r · x ∈ I + + + record isRightIdeal (I : ℙ R) : Type ℓ where + field + +-closed : {x y : R} → x ∈ I → y ∈ I → x + y ∈ I + -closed : {x : R} → x ∈ I → - x ∈ I + 0r-closed : 0r ∈ I + ·-closedRight : {x : R} → (r : R) → x ∈ I → x · r ∈ I + + {- Examples of ideals -} + zeroSubset : ℙ R + zeroSubset x = (x ≡ 0r) , isSetRing R' _ _ + + open RingTheory R' + open isIdeal + + isIdealZeroIdeal : isIdeal zeroSubset + +-closed isIdealZeroIdeal x≡0 y≡0 = + _ + _ ≡⟨ cong (λ u → u + _) x≡0 ⟩ + 0r + _ ≡⟨ +Lid _ ⟩ + _ ≡⟨ y≡0 ⟩ + 0r ∎ + -closed isIdealZeroIdeal x≡0 = + - _ ≡⟨ cong (λ u → - u) x≡0 ⟩ + - 0r ≡⟨ 0Selfinverse ⟩ + 0r ∎ + 0r-closed isIdealZeroIdeal = refl + ·-closedLeft isIdealZeroIdeal r x≡0 = + r · _ ≡⟨ cong (λ u → r · u) x≡0 ⟩ + r · 0r ≡⟨ 0RightAnnihilates r ⟩ + 0r ∎ + ·-closedRight isIdealZeroIdeal r x≡0 = + _ · r ≡⟨ cong (λ u → u · r) x≡0 ⟩ + 0r · r ≡⟨ 0LeftAnnihilates r ⟩ + 0r ∎ + + zeroIdeal : Ideal + zeroIdeal = zeroSubset , isIdealZeroIdeal + +IdealsIn : (R : Ring ℓ) → Type _ +IdealsIn {ℓ} R = Σ[ I ∈ ℙ ⟨ R ⟩ ] isIdeal R I diff --git a/Cubical/Algebra/Ring/Kernel.agda b/Cubical/Algebra/Ring/Kernel.agda new file mode 100644 index 0000000000..2c1e81c757 --- /dev/null +++ b/Cubical/Algebra/Ring/Kernel.agda @@ -0,0 +1,56 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Ring.Kernel where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset + +open import Cubical.Algebra.Ring.Base +open import Cubical.Algebra.Ring.Properties +open import Cubical.Algebra.Ring.Ideal + +private + variable + ℓ : Level + +module _ {{R S : Ring ℓ}} (f′ : RingHom R S) where + open IsRingHom (f′ .snd) + open RingStr ⦃...⦄ + open isIdeal + open RingTheory + private + instance + _ = R + _ = S + _ = snd R + _ = snd S + f = fst f′ + + kernel : ⟨ R ⟩ → hProp ℓ + kernel x = (f x ≡ 0r) , isSetRing S _ _ + + kernelIsIdeal : isIdeal R kernel + +-closed kernelIsIdeal = + λ fx≡0 fy≡0 → f (_ + _) ≡⟨ pres+ _ _ ⟩ + f _ + f _ ≡⟨ cong (λ u → u + f _) fx≡0 ⟩ + 0r + f _ ≡⟨ cong (λ u → 0r + u) fy≡0 ⟩ + 0r + 0r ≡⟨ 0Idempotent S ⟩ + 0r ∎ + -closed kernelIsIdeal = + λ fx≡0 → f (- _) ≡⟨ pres- _ ⟩ + - f _ ≡⟨ cong -_ fx≡0 ⟩ + - 0r ≡⟨ 0Selfinverse S ⟩ + 0r ∎ + 0r-closed kernelIsIdeal = f 0r ≡⟨ pres0 ⟩ 0r ∎ + ·-closedLeft kernelIsIdeal = λ r fx≡0 → + f (r · _) ≡⟨ pres· _ _ ⟩ + f r · f (_) ≡⟨ cong (λ u → f r · u) fx≡0 ⟩ + f r · 0r ≡⟨ 0RightAnnihilates S _ ⟩ + 0r ∎ + ·-closedRight kernelIsIdeal = λ r fx≡0 → + f (_ · r) ≡⟨ pres· _ _ ⟩ + f _ · f r ≡⟨ cong (λ u → u · f r) fx≡0 ⟩ + 0r · f r ≡⟨ 0LeftAnnihilates S _ ⟩ + 0r ∎ diff --git a/Cubical/Algebra/Ring/Properties.agda b/Cubical/Algebra/Ring/Properties.agda new file mode 100644 index 0000000000..954d0ea080 --- /dev/null +++ b/Cubical/Algebra/Ring/Properties.agda @@ -0,0 +1,312 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.Ring.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Sigma +open import Cubical.Relation.Binary.Poset + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Auto +open import Cubical.Structures.Macro +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Ring.Base + +open import Cubical.HITs.PropositionalTruncation + +private + variable + ℓ ℓ' ℓ'' ℓ''' ℓ'''' : Level + +{- + some basic calculations (used for example in QuotientRing.agda), + that should become obsolete or subject to change once we have a + ring solver (see https://github.com/agda/cubical/issues/297) +-} +module RingTheory (R' : Ring ℓ) where + + open RingStr (snd R') + private R = ⟨ R' ⟩ + implicitInverse : (x y : R) + → x + y ≡ 0r + → y ≡ - x + implicitInverse x y p = + y ≡⟨ sym (+Lid y) ⟩ + 0r + y ≡⟨ cong (λ u → u + y) (sym (+Linv x)) ⟩ + (- x + x) + y ≡⟨ sym (+Assoc _ _ _) ⟩ + (- x) + (x + y) ≡⟨ cong (λ u → (- x) + u) p ⟩ + (- x) + 0r ≡⟨ +Rid _ ⟩ + - x ∎ + + equalByDifference : (x y : R) + → x - y ≡ 0r + → x ≡ y + equalByDifference x y p = + x ≡⟨ sym (+Rid _) ⟩ + x + 0r ≡⟨ cong (λ u → x + u) (sym (+Linv y)) ⟩ + x + ((- y) + y) ≡⟨ +Assoc _ _ _ ⟩ + (x - y) + y ≡⟨ cong (λ u → u + y) p ⟩ + 0r + y ≡⟨ +Lid _ ⟩ + y ∎ + + 0Selfinverse : - 0r ≡ 0r + 0Selfinverse = sym (implicitInverse _ _ (+Rid 0r)) + + 0Idempotent : 0r + 0r ≡ 0r + 0Idempotent = +Lid 0r + + +Idempotency→0 : (x : R) → x ≡ x + x → x ≡ 0r + +Idempotency→0 x p = + x ≡⟨ sym (+Rid x) ⟩ + x + 0r ≡⟨ cong (λ u → x + u) (sym (+Rinv _)) ⟩ + x + (x + (- x)) ≡⟨ +Assoc _ _ _ ⟩ + (x + x) + (- x) ≡⟨ cong (λ u → u + (- x)) (sym p) ⟩ + x + (- x) ≡⟨ +Rinv _ ⟩ + 0r ∎ + + -Idempotent : (x : R) → -(- x) ≡ x + -Idempotent x = - (- x) ≡⟨ sym (implicitInverse (- x) x (+Linv _)) ⟩ + x ∎ + + 0RightAnnihilates : (x : R) → x · 0r ≡ 0r + 0RightAnnihilates x = + let x·0-is-idempotent : x · 0r ≡ x · 0r + x · 0r + x·0-is-idempotent = + x · 0r ≡⟨ cong (λ u → x · u) (sym 0Idempotent) ⟩ + x · (0r + 0r) ≡⟨ ·Rdist+ _ _ _ ⟩ + (x · 0r) + (x · 0r) ∎ + in (+Idempotency→0 _ x·0-is-idempotent) + + 0LeftAnnihilates : (x : R) → 0r · x ≡ 0r + 0LeftAnnihilates x = + let 0·x-is-idempotent : 0r · x ≡ 0r · x + 0r · x + 0·x-is-idempotent = + 0r · x ≡⟨ cong (λ u → u · x) (sym 0Idempotent) ⟩ + (0r + 0r) · x ≡⟨ ·Ldist+ _ _ _ ⟩ + (0r · x) + (0r · x) ∎ + in +Idempotency→0 _ 0·x-is-idempotent + + -DistR· : (x y : R) → x · (- y) ≡ - (x · y) + -DistR· x y = implicitInverse (x · y) (x · (- y)) + + (x · y + x · (- y) ≡⟨ sym (·Rdist+ _ _ _) ⟩ + x · (y + (- y)) ≡⟨ cong (λ u → x · u) (+Rinv y) ⟩ + x · 0r ≡⟨ 0RightAnnihilates x ⟩ + 0r ∎) + + -DistL· : (x y : R) → (- x) · y ≡ - (x · y) + -DistL· x y = implicitInverse (x · y) ((- x) · y) + + (x · y + (- x) · y ≡⟨ sym (·Ldist+ _ _ _) ⟩ + (x - x) · y ≡⟨ cong (λ u → u · y) (+Rinv x) ⟩ + 0r · y ≡⟨ 0LeftAnnihilates y ⟩ + 0r ∎) + + -Swap· : (x y : R) → (- x) · y ≡ x · (- y) + -Swap· _ _ = -DistL· _ _ ∙ sym (-DistR· _ _) + + -IsMult-1 : (x : R) → - x ≡ (- 1r) · x + -IsMult-1 _ = sym (·Lid _) ∙ sym (-Swap· _ _) + + -Dist : (x y : R) → (- x) + (- y) ≡ - (x + y) + -Dist x y = + implicitInverse _ _ + ((x + y) + ((- x) + (- y)) ≡⟨ sym (+Assoc _ _ _) ⟩ + x + (y + ((- x) + (- y))) ≡⟨ cong + (λ u → x + (y + u)) + (+Comm _ _) ⟩ + x + (y + ((- y) + (- x))) ≡⟨ cong (λ u → x + u) (+Assoc _ _ _) ⟩ + x + ((y + (- y)) + (- x)) ≡⟨ cong (λ u → x + (u + (- x))) + (+Rinv _) ⟩ + x + (0r + (- x)) ≡⟨ cong (λ u → x + u) (+Lid _) ⟩ + x + (- x) ≡⟨ +Rinv _ ⟩ + 0r ∎) + + translatedDifference : (x a b : R) → a - b ≡ (x + a) - (x + b) + translatedDifference x a b = + a - b ≡⟨ cong (λ u → a + u) + (sym (+Lid _)) ⟩ + (a + (0r + (- b))) ≡⟨ cong (λ u → a + (u + (- b))) + (sym (+Rinv _)) ⟩ + (a + ((x + (- x)) + (- b))) ≡⟨ cong (λ u → a + u) + (sym (+Assoc _ _ _)) ⟩ + (a + (x + ((- x) + (- b)))) ≡⟨ (+Assoc _ _ _) ⟩ + ((a + x) + ((- x) + (- b))) ≡⟨ cong (λ u → u + ((- x) + (- b))) + (+Comm _ _) ⟩ + ((x + a) + ((- x) + (- b))) ≡⟨ cong (λ u → (x + a) + u) + (-Dist _ _) ⟩ + ((x + a) - (x + b)) ∎ + + +Assoc-comm1 : (x y z : R) → x + (y + z) ≡ y + (x + z) + +Assoc-comm1 x y z = +Assoc x y z ∙∙ cong (λ x → x + z) (+Comm x y) ∙∙ sym (+Assoc y x z) + + +Assoc-comm2 : (x y z : R) → x + (y + z) ≡ z + (y + x) + +Assoc-comm2 x y z = +Assoc-comm1 x y z ∙∙ cong (λ x → y + x) (+Comm x z) ∙∙ +Assoc-comm1 y z x + + +ShufflePairs : (a b c d : R) + → (a + b) + (c + d) ≡ (a + c) + (b + d) + +ShufflePairs a b c d = + (a + b) + (c + d) ≡⟨ +Assoc _ _ _ ⟩ + ((a + b) + c) + d ≡⟨ cong (λ u → u + d) (sym (+Assoc _ _ _)) ⟩ + (a + (b + c)) + d ≡⟨ cong (λ u → (a + u) + d) (+Comm _ _) ⟩ + (a + (c + b)) + d ≡⟨ cong (λ u → u + d) (+Assoc _ _ _) ⟩ + ((a + c) + b) + d ≡⟨ sym (+Assoc _ _ _) ⟩ + (a + c) + (b + d) ∎ + + ·-assoc2 : (x y z w : R) → (x · y) · (z · w) ≡ x · (y · z) · w + ·-assoc2 x y z w = ·Assoc (x · y) z w ∙ cong (_· w) (sym (·Assoc x y z)) + + +module RingHoms where + open IsRingHom + + idRingHom : (R : Ring ℓ) → RingHom R R + fst (idRingHom R) = idfun (fst R) + snd (idRingHom R) = makeIsRingHom refl (λ _ _ → refl) (λ _ _ → refl) + + compIsRingHom : {A : Ring ℓ} {B : Ring ℓ'} {C : Ring ℓ''} + {g : ⟨ B ⟩ → ⟨ C ⟩} {f : ⟨ A ⟩ → ⟨ B ⟩} + → IsRingHom (B .snd) g (C .snd) + → IsRingHom (A .snd) f (B .snd) + → IsRingHom (A .snd) (g ∘ f) (C .snd) + compIsRingHom {g = g} {f} gh fh .pres0 = cong g (fh .pres0) ∙ gh .pres0 + compIsRingHom {g = g} {f} gh fh .pres1 = cong g (fh .pres1) ∙ gh .pres1 + compIsRingHom {g = g} {f} gh fh .pres+ x y = cong g (fh .pres+ x y) ∙ gh .pres+ (f x) (f y) + compIsRingHom {g = g} {f} gh fh .pres· x y = cong g (fh .pres· x y) ∙ gh .pres· (f x) (f y) + compIsRingHom {g = g} {f} gh fh .pres- x = cong g (fh .pres- x) ∙ gh .pres- (f x) + + compRingHom : {R : Ring ℓ} {S : Ring ℓ'} {T : Ring ℓ''} + → RingHom R S → RingHom S T → RingHom R T + fst (compRingHom f g) x = g .fst (f .fst x) + snd (compRingHom f g) = compIsRingHom (g .snd) (f .snd) + + syntax compRingHom f g = g ∘r f + + compIdRingHom : {R S : Ring ℓ} (φ : RingHom R S) → compRingHom (idRingHom R) φ ≡ φ + compIdRingHom φ = RingHom≡ refl + + idCompRingHom : {R S : Ring ℓ} (φ : RingHom R S) → compRingHom φ (idRingHom S) ≡ φ + idCompRingHom φ = RingHom≡ refl + + compAssocRingHom : {R S T U : Ring ℓ} (φ : RingHom R S) (ψ : RingHom S T) (χ : RingHom T U) → + compRingHom (compRingHom φ ψ) χ ≡ compRingHom φ (compRingHom ψ χ) + compAssocRingHom _ _ _ = RingHom≡ refl + + +module RingEquivs where + open IsRingHom + open RingHoms + + compIsRingEquiv : {A : Ring ℓ} {B : Ring ℓ'} {C : Ring ℓ''} + {g : ⟨ B ⟩ ≃ ⟨ C ⟩} {f : ⟨ A ⟩ ≃ ⟨ B ⟩} + → IsRingEquiv (B .snd) g (C .snd) + → IsRingEquiv (A .snd) f (B .snd) + → IsRingEquiv (A .snd) (compEquiv f g) (C .snd) + compIsRingEquiv {g = g} {f} gh fh = compIsRingHom {g = g .fst} {f .fst} gh fh + + compRingEquiv : {A : Ring ℓ} {B : Ring ℓ'} {C : Ring ℓ''} + → RingEquiv A B → RingEquiv B C → RingEquiv A C + fst (compRingEquiv f g) = compEquiv (f .fst) (g .fst) + snd (compRingEquiv f g) = compIsRingEquiv {g = g .fst} {f = f .fst} (g .snd) (f .snd) + + +module RingHomTheory {R S : Ring ℓ} (φ : RingHom R S) where + open RingTheory ⦃...⦄ + open RingStr ⦃...⦄ + open IsRingHom (φ .snd) + private + instance + _ = R + _ = S + _ = snd R + _ = snd S + f = fst φ + + ker≡0→inj : ({x : ⟨ R ⟩} → f x ≡ 0r → x ≡ 0r) + → ({x y : ⟨ R ⟩} → f x ≡ f y → x ≡ y) + ker≡0→inj ker≡0 {x} {y} p = equalByDifference _ _ (ker≡0 path) + where + path : f (x - y) ≡ 0r + path = f (x - y) ≡⟨ pres+ _ _ ⟩ + f x + f (- y) ≡⟨ cong (f x +_) (pres- _) ⟩ + f x - f y ≡⟨ cong (_- f y) p ⟩ + f y - f y ≡⟨ +Rinv _ ⟩ + 0r ∎ + + +-- the Ring version of uaCompEquiv +module RingUAFunctoriality where + open RingStr + open RingEquivs + + Ring≡ : (A B : Ring ℓ) → ( + Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] + Σ[ q0 ∈ PathP (λ i → p i) (0r (snd A)) (0r (snd B)) ] + Σ[ q1 ∈ PathP (λ i → p i) (1r (snd A)) (1r (snd B)) ] + Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] + Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] + Σ[ s ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] + PathP (λ i → IsRing (q0 i) (q1 i) (r+ i) (r· i) (s i)) (isRing (snd A)) (isRing (snd B))) + ≃ (A ≡ B) + Ring≡ A B = isoToEquiv theIso + where + open Iso + theIso : Iso _ _ + fun theIso (p , q0 , q1 , r+ , r· , s , t) i = p i + , ringstr (q0 i) (q1 i) (r+ i) (r· i) (s i) (t i) + inv theIso x = cong ⟨_⟩ x , cong (0r ∘ snd) x , cong (1r ∘ snd) x + , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (isRing ∘ snd) x + rightInv theIso _ = refl + leftInv theIso _ = refl + + caracRing≡ : {A B : Ring ℓ} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q + caracRing≡ {A = A} {B = B} p q P = + sym (transportTransport⁻ (ua (Ring≡ A B)) p) + ∙∙ cong (transport (ua (Ring≡ A B))) helper + ∙∙ transportTransport⁻ (ua (Ring≡ A B)) q + where + helper : transport (sym (ua (Ring≡ A B))) p ≡ transport (sym (ua (Ring≡ A B))) q + helper = Σ≡Prop + (λ _ → isPropΣ + (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) + λ _ → isOfHLevelPathP 1 (isPropIsRing _ _ _ _ _) _ _) + (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + + uaCompRingEquiv : {A B C : Ring ℓ} (f : RingEquiv A B) (g : RingEquiv B C) + → uaRing (compRingEquiv f g) ≡ uaRing f ∙ uaRing g + uaCompRingEquiv f g = caracRing≡ _ _ ( + cong ⟨_⟩ (uaRing (compRingEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaRing f) ∙ cong ⟨_⟩ (uaRing g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaRing f) (uaRing g)) ⟩ + cong ⟨_⟩ (uaRing f ∙ uaRing g) ∎) + + + +open RingEquivs +open RingUAFunctoriality +recPT→Ring : {A : Type ℓ'} (𝓕 : A → Ring ℓ) + → (σ : ∀ x y → RingEquiv (𝓕 x) (𝓕 y)) + → (∀ x y z → σ x z ≡ compRingEquiv (σ x y) (σ y z)) + ------------------------------------------------------ + → ∥ A ∥ → Ring ℓ +recPT→Ring 𝓕 σ compCoh = rec→Gpd isGroupoidRing 𝓕 + (3-ConstantCompChar 𝓕 (λ x y → uaRing (σ x y)) + λ x y z → sym ( cong uaRing (compCoh x y z) + ∙ uaCompRingEquiv (σ x y) (σ y z))) diff --git a/Cubical/Algebra/Ring/QuotientRing.agda b/Cubical/Algebra/Ring/QuotientRing.agda new file mode 100644 index 0000000000..20a415fa9e --- /dev/null +++ b/Cubical/Algebra/Ring/QuotientRing.agda @@ -0,0 +1,236 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Ring.QuotientRing where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Powerset using (_∈_; _⊆_) -- \in, \sub= + +open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/ₛ_) +open import Cubical.HITs.SetQuotients.Properties + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Ring.Ideal +open import Cubical.Algebra.Ring.Kernel + +private + variable + ℓ : Level + +module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' I) where + open RingStr (snd R') + private R = ⟨ R' ⟩ + open isIdeal I-isIdeal + open RingTheory R' + + R/I : Type ℓ + R/I = R /ₛ (λ x y → x - y ∈ I) + + private + homogeneity : ∀ (x a b : R) + → (a - b ∈ I) + → (x + a) - (x + b) ∈ I + homogeneity x a b p = subst (λ u → u ∈ I) (translatedDifference x a b) p + + isSetR/I : isSet R/I + isSetR/I = squash/ + [_]/I : (a : R) → R/I + [ a ]/I = [ a ] + + lemma : (x y a : R) + → x - y ∈ I + → [ x + a ]/I ≡ [ y + a ]/I + lemma x y a x-y∈I = eq/ (x + a) (y + a) (subst (λ u → u ∈ I) calculate x-y∈I) + where calculate : x - y ≡ (x + a) - (y + a) + calculate = + x - y ≡⟨ translatedDifference a x y ⟩ + ((a + x) - (a + y)) ≡⟨ cong (λ u → u - (a + y)) (+Comm _ _) ⟩ + ((x + a) - (a + y)) ≡⟨ cong (λ u → (x + a) - u) (+Comm _ _) ⟩ + ((x + a) - (y + a)) ∎ + + pre-+/I : R → R/I → R/I + pre-+/I x = elim + (λ _ → squash/) + (λ y → [ x + y ]) + λ y y' diffrenceInIdeal + → eq/ (x + y) (x + y') (homogeneity x y y' diffrenceInIdeal) + + pre-+/I-DescendsToQuotient : (x y : R) → (x - y ∈ I) + → pre-+/I x ≡ pre-+/I y + pre-+/I-DescendsToQuotient x y x-y∈I i r = pointwise-equal r i + where + pointwise-equal : ∀ (u : R/I) + → pre-+/I x u ≡ pre-+/I y u + pointwise-equal = elimProp (λ u → isSetR/I (pre-+/I x u) (pre-+/I y u)) + (λ a → lemma x y a x-y∈I) + + _+/I_ : R/I → R/I → R/I + x +/I y = (elim R/I→R/I-isSet pre-+/I pre-+/I-DescendsToQuotient x) y + where + R/I→R/I-isSet : R/I → isSet (R/I → R/I) + R/I→R/I-isSet _ = isSetΠ (λ _ → squash/) + + +/I-comm : (x y : R/I) → x +/I y ≡ y +/I x + +/I-comm = elimProp2 (λ _ _ → squash/ _ _) eq + where eq : (x y : R) → [ x ] +/I [ y ] ≡ [ y ] +/I [ x ] + eq x y i = [ +Comm x y i ] + + +/I-assoc : (x y z : R/I) → x +/I (y +/I z) ≡ (x +/I y) +/I z + +/I-assoc = elimProp3 (λ _ _ _ → squash/ _ _) eq + where eq : (x y z : R) → [ x ] +/I ([ y ] +/I [ z ]) ≡ ([ x ] +/I [ y ]) +/I [ z ] + eq x y z i = [ +Assoc x y z i ] + + + 0/I : R/I + 0/I = [ 0r ] + + 1/I : R/I + 1/I = [ 1r ] + + -/I : R/I → R/I + -/I = elim (λ _ → squash/) (λ x' → [ - x' ]) eq + where + eq : (x y : R) → (x - y ∈ I) → [ - x ] ≡ [ - y ] + eq x y x-y∈I = eq/ (- x) (- y) (subst (λ u → u ∈ I) eq' (isIdeal.-closed I-isIdeal x-y∈I)) + where + eq' = - (x + (- y)) ≡⟨ sym (-Dist _ _) ⟩ + (- x) - (- y) ∎ + + +/I-rinv : (x : R/I) → x +/I (-/I x) ≡ 0/I + +/I-rinv = elimProp (λ x → squash/ _ _) eq + where + eq : (x : R) → [ x ] +/I (-/I [ x ]) ≡ 0/I + eq x i = [ +Rinv x i ] + + + +/I-rid : (x : R/I) → x +/I 0/I ≡ x + +/I-rid = elimProp (λ x → squash/ _ _) eq + where + eq : (x : R) → [ x ] +/I 0/I ≡ [ x ] + eq x i = [ +Rid x i ] + + _·/I_ : R/I → R/I → R/I + _·/I_ = + elim (λ _ → isSetΠ (λ _ → squash/)) + (λ x → left· x) + eq' + where + eq : (x y y' : R) → (y - y' ∈ I) → [ x · y ] ≡ [ x · y' ] + eq x y y' y-y'∈I = eq/ _ _ + (subst (λ u → u ∈ I) + (x · (y - y') ≡⟨ ·Rdist+ _ _ _ ⟩ + ((x · y) + x · (- y')) ≡⟨ cong (λ u → (x · y) + u) + (-DistR· x y') ⟩ + (x · y) - (x · y') ∎) + (isIdeal.·-closedLeft I-isIdeal x y-y'∈I)) + left· : (x : R) → R/I → R/I + left· x = elim (λ y → squash/) + (λ y → [ x · y ]) + (eq x) + eq' : (x x' : R) → (x - x' ∈ I) → left· x ≡ left· x' + eq' x x' x-x'∈I i y = elimProp (λ y → squash/ (left· x y) (left· x' y)) + (λ y → eq′ y) + y i + where + eq′ : (y : R) → left· x [ y ] ≡ left· x' [ y ] + eq′ y = eq/ (x · y) (x' · y) + (subst (λ u → u ∈ I) + ((x - x') · y ≡⟨ ·Ldist+ x (- x') y ⟩ + x · y + (- x') · y ≡⟨ cong + (λ u → x · y + u) + (-DistL· x' y) ⟩ + x · y - x' · y ∎) + (isIdeal.·-closedRight I-isIdeal y x-x'∈I)) + + + -- more or less copy paste from '+/I' - this is preliminary anyway + ·/I-assoc : (x y z : R/I) → x ·/I (y ·/I z) ≡ (x ·/I y) ·/I z + ·/I-assoc = elimProp3 (λ _ _ _ → squash/ _ _) eq + where eq : (x y z : R) → [ x ] ·/I ([ y ] ·/I [ z ]) ≡ ([ x ] ·/I [ y ]) ·/I [ z ] + eq x y z i = [ ·Assoc x y z i ] + + ·/I-lid : (x : R/I) → 1/I ·/I x ≡ x + ·/I-lid = elimProp (λ x → squash/ _ _) eq + where + eq : (x : R) → 1/I ·/I [ x ] ≡ [ x ] + eq x i = [ ·Lid x i ] + + ·/I-rid : (x : R/I) → x ·/I 1/I ≡ x + ·/I-rid = elimProp (λ x → squash/ _ _) eq + where + eq : (x : R) → [ x ] ·/I 1/I ≡ [ x ] + eq x i = [ ·Rid x i ] + + + /I-ldist : (x y z : R/I) → (x +/I y) ·/I z ≡ (x ·/I z) +/I (y ·/I z) + /I-ldist = elimProp3 (λ _ _ _ → squash/ _ _) eq + where + eq : (x y z : R) → ([ x ] +/I [ y ]) ·/I [ z ] ≡ ([ x ] ·/I [ z ]) +/I ([ y ] ·/I [ z ]) + eq x y z i = [ ·Ldist+ x y z i ] + + /I-rdist : (x y z : R/I) → x ·/I (y +/I z) ≡ (x ·/I y) +/I (x ·/I z) + /I-rdist = elimProp3 (λ _ _ _ → squash/ _ _) eq + where + eq : (x y z : R) → [ x ] ·/I ([ y ] +/I [ z ]) ≡ ([ x ] ·/I [ y ]) +/I ([ x ] ·/I [ z ]) + eq x y z i = [ ·Rdist+ x y z i ] + + asRing : Ring ℓ + asRing = makeRing 0/I 1/I _+/I_ _·/I_ -/I isSetR/I + +/I-assoc +/I-rid +/I-rinv +/I-comm + ·/I-assoc ·/I-rid ·/I-lid /I-rdist /I-ldist + +_/_ : (R : Ring ℓ) → (I : IdealsIn R) → Ring ℓ +R / (I , IisIdeal) = asRing R I IisIdeal + +[_]/I : {R : Ring ℓ} {I : IdealsIn R} → (a : ⟨ R ⟩) → ⟨ R / I ⟩ +[ a ]/I = [ a ] + + +module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where + open RingStr ⦃...⦄ + open RingTheory ⦃...⦄ + Iₛ = fst I + private + instance + _ = R + _ = snd R + + module _ {S : Ring ℓ} (φ : RingHom R S) where + open IsRingHom + open RingHomTheory φ + private + instance + _ = S + _ = snd S + f = fst φ + module φ = IsRingHom (snd φ) + + + inducedHom : Iₛ ⊆ kernel φ → RingHom (R / I) S + fst (inducedHom Iₛ⊆kernel) = + elim + (λ _ → isSetRing S) + f + λ r₁ r₂ r₁-r₂∈I → equalByDifference (f r₁) (f r₂) + (f r₁ - f r₂ ≡⟨ cong (λ u → f r₁ + u) (sym (φ.pres- _)) ⟩ + f r₁ + f (- r₂) ≡⟨ sym (φ.pres+ _ _) ⟩ + f (r₁ - r₂) ≡⟨ Iₛ⊆kernel (r₁ - r₂) r₁-r₂∈I ⟩ + 0r ∎) + pres0 (snd (inducedHom Iₛ⊆kernel)) = φ.pres0 + pres1 (snd (inducedHom Iₛ⊆kernel)) = φ.pres1 + pres+ (snd (inducedHom Iₛ⊆kernel)) = + elimProp2 (λ _ _ → isSetRing S _ _) φ.pres+ + pres· (snd (inducedHom Iₛ⊆kernel)) = + elimProp2 (λ _ _ → isSetRing S _ _) φ.pres· + pres- (snd (inducedHom Iₛ⊆kernel)) = + elimProp (λ _ → isSetRing S _ _) φ.pres- + + solution : (p : Iₛ ⊆ kernel φ) + → (x : ⟨ R ⟩) → inducedHom p $ [ x ] ≡ φ $ x + solution p x = refl + + unique : (p : Iₛ ⊆ kernel φ) + → (ψ : RingHom (R / I) S) → (ψIsSolution : (x : ⟨ R ⟩) → ψ $ [ x ] ≡ φ $ x) + → (x : ⟨ R ⟩) → ψ $ [ x ] ≡ inducedHom p $ [ x ] + unique p ψ ψIsSolution x = ψIsSolution x diff --git a/Cubical/Algebra/RingSolver/AlgebraExpression.agda b/Cubical/Algebra/RingSolver/AlgebraExpression.agda new file mode 100644 index 0000000000..9db5861b25 --- /dev/null +++ b/Cubical/Algebra/RingSolver/AlgebraExpression.agda @@ -0,0 +1,38 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.AlgebraExpression where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData +open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Nat.Order using (zero-≤) +open import Cubical.Data.Vec.Base +open import Cubical.Algebra.RingSolver.RawRing +open import Cubical.Algebra.RingSolver.RawAlgebra renaming (⟨_⟩ to ⟨_⟩ₐ) + +private + variable + ℓ ℓ′ : Level + +infixl 6 _+'_ +infixl 7 -'_ +infixl 8 _·'_ + +-- Expression in an R-Algebra A with n variables +data Expr {ℓ} (R : RawRing ℓ) (A : Type ℓ′) (n : ℕ) : Type ℓ where + K : ⟨ R ⟩ → Expr R A n + ∣ : Fin n → Expr R A n + _+'_ : Expr R A n → Expr R A n → Expr R A n + _·'_ : Expr R A n → Expr R A n → Expr R A n + -'_ : Expr R A n → Expr R A n + +module Eval (R : RawRing ℓ) (A : RawAlgebra R ℓ′) where + open import Cubical.Data.Vec + open RawAlgebra A renaming (scalar to scalarₐ) + + ⟦_⟧ : ∀ {n} → Expr R ⟨ A ⟩ₐ n → Vec ⟨ A ⟩ₐ n → ⟨ A ⟩ₐ + ⟦ K r ⟧ v = scalarₐ r + ⟦ ∣ k ⟧ v = lookup k v + ⟦ x +' y ⟧ v = ⟦ x ⟧ v + ⟦ y ⟧ v + ⟦ x ·' y ⟧ v = ⟦ x ⟧ v · ⟦ y ⟧ v + ⟦ -' x ⟧ v = - ⟦ x ⟧ v diff --git a/Cubical/Algebra/RingSolver/EvalHom.agda b/Cubical/Algebra/RingSolver/EvalHom.agda new file mode 100644 index 0000000000..a86c38171e --- /dev/null +++ b/Cubical/Algebra/RingSolver/EvalHom.agda @@ -0,0 +1,275 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.EvalHom where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Int.Base hiding (_+_ ; _·_ ; -_) +open import Cubical.Data.FinData +open import Cubical.Data.Vec +open import Cubical.Data.Bool +open import Cubical.Relation.Nullary.Base + +open import Cubical.Algebra.RingSolver.Utility +open import Cubical.Algebra.RingSolver.RawAlgebra +open import Cubical.Algebra.RingSolver.IntAsRawRing +open import Cubical.Algebra.RingSolver.HornerForms +open import Cubical.Algebra.RingSolver.HornerEval +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring + +private + variable + ℓ : Level + +module HomomorphismProperties (R : CommRing ℓ) where + private + νR = CommRing→RawℤAlgebra R + open CommRingStr (snd R) + open RingTheory (CommRing→Ring R) + open IteratedHornerOperations νR + + EvalHom+0 : {n : ℕ} (P : IteratedHornerForms νR n) (xs : Vec ⟨ νR ⟩ n) + → eval (0ₕ +ₕ P) xs ≡ eval P xs + EvalHom+0 {n = ℕ.zero} (const x) [] = cong (scalar R) (+Ridℤ x) + EvalHom+0 {n = ℕ.suc _} P xs = refl + + Eval0H : {n : ℕ} (xs : Vec ⟨ νR ⟩ n) + → eval {A = νR} 0ₕ xs ≡ 0r + Eval0H [] = refl + Eval0H (x ∷ xs) = refl + + Eval1ₕ : {n : ℕ} (xs : Vec ⟨ νR ⟩ n) + → eval {A = νR} 1ₕ xs ≡ 1r + Eval1ₕ [] = refl + Eval1ₕ (x ∷ xs) = + eval 1ₕ (x ∷ xs) ≡⟨ refl ⟩ + eval (0H ·X+ 1ₕ) (x ∷ xs) ≡⟨ combineCasesEval R 0H 1ₕ x xs ⟩ + eval {A = νR} 0H (x ∷ xs) · x + eval 1ₕ xs ≡⟨ cong (λ u → u · x + eval 1ₕ xs) + (Eval0H (x ∷ xs)) ⟩ + 0r · x + eval 1ₕ xs ≡⟨ cong (λ u → 0r · x + u) + (Eval1ₕ xs) ⟩ + 0r · x + 1r ≡⟨ cong (λ u → u + 1r) + (0LeftAnnihilates _) ⟩ + 0r + 1r ≡⟨ +Lid _ ⟩ + 1r ∎ + + -EvalDist : + {n : ℕ} (P : IteratedHornerForms νR n) (xs : Vec ⟨ νR ⟩ n) + → eval (-ₕ P) xs ≡ - eval P xs + -EvalDist (const x) [] = -DistScalar R x + -EvalDist 0H xs = + eval (-ₕ 0H) xs ≡⟨ Eval0H xs ⟩ + 0r ≡⟨ sym 0Selfinverse ⟩ + - 0r ≡⟨ cong -_ (sym (Eval0H xs)) ⟩ + - eval 0H xs ∎ + -EvalDist (P ·X+ Q) (x ∷ xs) = + eval (-ₕ (P ·X+ Q)) (x ∷ xs) + ≡⟨ refl ⟩ + eval ((-ₕ P) ·X+ (-ₕ Q)) (x ∷ xs) + ≡⟨ combineCasesEval R (-ₕ P) (-ₕ Q) x xs ⟩ + (eval (-ₕ P) (x ∷ xs)) · x + eval (-ₕ Q) xs + ≡⟨ cong (λ u → u · x + eval (-ₕ Q) xs) (-EvalDist P _) ⟩ + (- eval P (x ∷ xs)) · x + eval (-ₕ Q) xs + ≡⟨ cong (λ u → (- eval P (x ∷ xs)) · x + u) (-EvalDist Q _) ⟩ + (- eval P (x ∷ xs)) · x + - eval Q xs + ≡[ i ]⟨ -DistL· (eval P (x ∷ xs)) x i + - eval Q xs ⟩ + - ((eval P (x ∷ xs)) · x) + (- eval Q xs) + ≡⟨ -Dist _ _ ⟩ + - ((eval P (x ∷ xs)) · x + eval Q xs) + ≡[ i ]⟨ - combineCasesEval R P Q x xs (~ i) ⟩ + - eval (P ·X+ Q) (x ∷ xs) ∎ + + combineCases+ : {n : ℕ} (P Q : IteratedHornerForms νR (ℕ.suc n)) + (r s : IteratedHornerForms νR n) + (x : fst R) (xs : Vec (fst R) n) + → eval ((P ·X+ r) +ₕ (Q ·X+ s)) (x ∷ xs) + ≡ eval ((P +ₕ Q) ·X+ (r +ₕ s)) (x ∷ xs) + combineCases+ {n = n} P Q r s x xs with (isZero νR (P +ₕ Q) and isZero νR (r +ₕ s)) ≟ true + ... | yes p = compute+ₕEvalBothZero R n P Q r s x xs p + ... | no p = compute+ₕEvalNotBothZero R n P Q r s x xs (¬true→false _ p) + + +Homeval : + {n : ℕ} (P Q : IteratedHornerForms νR n) (xs : Vec ⟨ νR ⟩ n) + → eval (P +ₕ Q) xs ≡ (eval P xs) + (eval Q xs) + +Homeval (const x) (const y) [] = +HomScalar R x y + +Homeval 0H Q xs = + eval (0H +ₕ Q) xs ≡⟨ refl ⟩ + eval Q xs ≡⟨ sym (+Lid _) ⟩ + 0r + eval Q xs ≡⟨ cong (λ u → u + eval Q xs) (sym (Eval0H xs)) ⟩ + eval 0H xs + eval Q xs ∎ + +Homeval (P ·X+ Q) 0H xs = + eval ((P ·X+ Q) +ₕ 0H) xs ≡⟨ refl ⟩ + eval (P ·X+ Q) xs ≡⟨ sym (+Rid _) ⟩ + eval (P ·X+ Q) xs + 0r + ≡⟨ cong (λ u → eval (P ·X+ Q) xs + u) (sym (Eval0H xs)) ⟩ + eval (P ·X+ Q) xs + eval 0H xs ∎ + +Homeval (P ·X+ Q) (S ·X+ T) (x ∷ xs) = + eval ((P ·X+ Q) +ₕ (S ·X+ T)) (x ∷ xs) + ≡⟨ combineCases+ P S Q T x xs ⟩ + eval ((P +ₕ S) ·X+ (Q +ₕ T)) (x ∷ xs) + ≡⟨ combineCasesEval R (P +ₕ S) (Q +ₕ T) x xs ⟩ + (eval (P +ₕ S) (x ∷ xs)) · x + eval (Q +ₕ T) xs + ≡⟨ cong (λ u → (eval (P +ₕ S) (x ∷ xs)) · x + u) (+Homeval Q T xs) ⟩ + (eval (P +ₕ S) (x ∷ xs)) · x + (eval Q xs + eval T xs) + ≡⟨ cong (λ u → u · x + (eval Q xs + eval T xs)) (+Homeval P S (x ∷ xs)) ⟩ + (eval P (x ∷ xs) + eval S (x ∷ xs)) · x + + (eval Q xs + eval T xs) + ≡⟨ cong (λ u → u + (eval Q xs + eval T xs)) (·Ldist+ _ _ _) ⟩ + (eval P (x ∷ xs)) · x + (eval S (x ∷ xs)) · x + + (eval Q xs + eval T xs) + ≡⟨ +ShufflePairs _ _ _ _ ⟩ + ((eval P (x ∷ xs)) · x + eval Q xs) + + ((eval S (x ∷ xs)) · x + eval T xs) + ≡[ i ]⟨ combineCasesEval R P Q x xs (~ i) + combineCasesEval R S T x xs (~ i) ⟩ + eval (P ·X+ Q) (x ∷ xs) + + eval (S ·X+ T) (x ∷ xs) ∎ + + ⋆Homeval : {n : ℕ} + (r : IteratedHornerForms νR n) + (P : IteratedHornerForms νR (ℕ.suc n)) (x : ⟨ νR ⟩) (xs : Vec ⟨ νR ⟩ n) + → eval (r ⋆ P) (x ∷ xs) ≡ eval r xs · eval P (x ∷ xs) + + ⋆0LeftAnnihilates : + {n : ℕ} (P : IteratedHornerForms νR (ℕ.suc n)) (xs : Vec ⟨ νR ⟩ (ℕ.suc n)) + → eval (0ₕ ⋆ P) xs ≡ 0r + ⋆0LeftAnnihilates 0H xs = Eval0H xs + ⋆0LeftAnnihilates {n = ℕ.zero} (P ·X+ Q) (x ∷ xs) = refl + ⋆0LeftAnnihilates {n = ℕ.suc _} (P ·X+ Q) (x ∷ xs) = refl + + ⋆isZeroLeftAnnihilates : + {n : ℕ} (r : IteratedHornerForms νR n) + (P : IteratedHornerForms νR (ℕ.suc n)) + (xs : Vec ⟨ νR ⟩ (ℕ.suc n)) + → isZero νR r ≡ true + → eval (r ⋆ P) xs ≡ 0r + ⋆isZeroLeftAnnihilates r P xs isZero-r = evalIsZero R (r ⋆ P) xs (isZeroPresLeft⋆ r P isZero-r) + + ·0LeftAnnihilates : + {n : ℕ} (P : IteratedHornerForms νR n) (xs : Vec ⟨ νR ⟩ n) + → eval (0ₕ ·ₕ P) xs ≡ 0r + ·0LeftAnnihilates (const x) xs = + eval (const _) xs ≡⟨ Eval0H xs ⟩ 0r ∎ + ·0LeftAnnihilates 0H xs = Eval0H xs + ·0LeftAnnihilates (P ·X+ P₁) xs = Eval0H xs + + ·isZeroLeftAnnihilates : + {n : ℕ} (P Q : IteratedHornerForms νR n) + (xs : Vec (fst R) n) + → isZero νR P ≡ true + → eval (P ·ₕ Q) xs ≡ 0r + ·isZeroLeftAnnihilates P Q xs isZeroP = evalIsZero R (P ·ₕ Q) xs (isZeroPresLeft·ₕ P Q isZeroP) + + ·Homeval : {n : ℕ} (P Q : IteratedHornerForms νR n) (xs : Vec ⟨ νR ⟩ n) + → eval (P ·ₕ Q) xs ≡ (eval P xs) · (eval Q xs) + + combineCases⋆ : {n : ℕ} (x : fst R) (xs : Vec (fst R) n) + → (r : IteratedHornerForms νR n) + → (P : IteratedHornerForms νR (ℕ.suc n)) + → (Q : IteratedHornerForms νR n) + → eval (r ⋆ (P ·X+ Q)) (x ∷ xs) ≡ eval ((r ⋆ P) ·X+ (r ·ₕ Q)) (x ∷ xs) + combineCases⋆ x xs r P Q with isZero νR r ≟ true + ... | yes p = + eval (r ⋆ (P ·X+ Q)) (x ∷ xs) ≡⟨ ⋆isZeroLeftAnnihilates r (P ·X+ Q) (x ∷ xs) p ⟩ + 0r ≡⟨ someCalculation R ⟩ + 0r · x + 0r ≡⟨ step1 ⟩ + eval (r ⋆ P) (x ∷ xs) · x + eval (r ·ₕ Q) xs ≡⟨ sym (combineCasesEval R (r ⋆ P) (r ·ₕ Q) x xs) ⟩ + eval ((r ⋆ P) ·X+ (r ·ₕ Q)) (x ∷ xs) ∎ + where + step1 : 0r · x + 0r ≡ eval (r ⋆ P) (x ∷ xs) · x + eval (r ·ₕ Q) xs + step1 i = ⋆isZeroLeftAnnihilates r P (x ∷ xs) p (~ i) · x + ·isZeroLeftAnnihilates r Q xs p (~ i) + ... | no p with isZero νR r + ... | true = byAbsurdity (p refl) + ... | false = refl + + ⋆Homeval r 0H x xs = + eval (r ⋆ 0H) (x ∷ xs) ≡⟨ refl ⟩ + 0r ≡⟨ sym (0RightAnnihilates _) ⟩ + eval r xs · 0r ≡⟨ refl ⟩ + eval r xs · eval {A = νR} 0H (x ∷ xs) ∎ + ⋆Homeval r (P ·X+ Q) x xs = + eval (r ⋆ (P ·X+ Q)) (x ∷ xs) ≡⟨ combineCases⋆ x xs r P Q ⟩ + eval ((r ⋆ P) ·X+ (r ·ₕ Q)) (x ∷ xs) + ≡⟨ combineCasesEval R (r ⋆ P) (r ·ₕ Q) x xs ⟩ + (eval (r ⋆ P) (x ∷ xs)) · x + eval (r ·ₕ Q) xs + ≡⟨ cong (λ u → u · x + eval (r ·ₕ Q) xs) (⋆Homeval r P x xs) ⟩ + (eval r xs · eval P (x ∷ xs)) · x + eval (r ·ₕ Q) xs + ≡⟨ cong (λ u → (eval r xs · eval P (x ∷ xs)) · x + u) (·Homeval r Q xs) ⟩ + (eval r xs · eval P (x ∷ xs)) · x + eval r xs · eval Q xs + ≡⟨ cong (λ u → u + eval r xs · eval Q xs) (sym (·Assoc _ _ _)) ⟩ + eval r xs · (eval P (x ∷ xs) · x) + eval r xs · eval Q xs + ≡⟨ sym (·Rdist+ _ _ _) ⟩ + eval r xs · ((eval P (x ∷ xs) · x) + eval Q xs) + ≡[ i ]⟨ eval r xs · combineCasesEval R P Q x xs (~ i) ⟩ + eval r xs · eval (P ·X+ Q) (x ∷ xs) ∎ + + lemmaForCombineCases· : + {n : ℕ} (Q : IteratedHornerForms νR n) (P S : IteratedHornerForms νR (ℕ.suc n)) + (xs : Vec (fst R) (ℕ.suc n)) + → isZero νR (P ·ₕ S) ≡ true + → eval ((P ·X+ Q) ·ₕ S) xs ≡ eval (Q ⋆ S) xs + lemmaForCombineCases· Q P S xs isZeroProd with isZero νR (P ·ₕ S) + ... | true = refl + ... | false = byBoolAbsurdity isZeroProd + + combineCases· : + {n : ℕ} (Q : IteratedHornerForms νR n) (P S : IteratedHornerForms νR (ℕ.suc n)) + (xs : Vec (fst R) (ℕ.suc n)) + → eval ((P ·X+ Q) ·ₕ S) xs ≡ eval (((P ·ₕ S) ·X+ 0ₕ) +ₕ (Q ⋆ S)) xs + combineCases· Q P S (x ∷ xs) with isZero νR (P ·ₕ S) ≟ true + ... | yes p = + eval ((P ·X+ Q) ·ₕ S) (x ∷ xs) ≡⟨ lemmaForCombineCases· Q P S (x ∷ xs) p ⟩ + eval (Q ⋆ S) (x ∷ xs) ≡⟨ sym (+Lid _) ⟩ + 0r + eval (Q ⋆ S) (x ∷ xs) ≡⟨ step1 ⟩ + eval ((P ·ₕ S) ·X+ 0ₕ) (x ∷ xs) + eval (Q ⋆ S) (x ∷ xs) ≡⟨ step2 ⟩ + eval (((P ·ₕ S) ·X+ 0ₕ) +ₕ (Q ⋆ S)) (x ∷ xs) ∎ + where + lemma = + eval ((P ·ₕ S) ·X+ 0ₕ) (x ∷ xs) ≡⟨ combineCasesEval R (P ·ₕ S) 0ₕ x xs ⟩ + eval (P ·ₕ S) (x ∷ xs) · x + eval 0ₕ xs ≡[ i ]⟨ evalIsZero R (P ·ₕ S) (x ∷ xs) p i · x + Eval0H xs i ⟩ + 0r · x + 0r ≡⟨ sym (someCalculation R) ⟩ + 0r ∎ + step1 : _ ≡ _ + step1 i = lemma (~ i) + eval (Q ⋆ S) (x ∷ xs) + step2 = sym (+Homeval ((P ·ₕ S) ·X+ 0ₕ) (Q ⋆ S) (x ∷ xs)) + ... | no p with isZero νR (P ·ₕ S) + ... | true = byAbsurdity (p refl) + ... | false = refl + + ·Homeval (const x) (const y) [] = ·HomScalar R x y + ·Homeval 0H Q xs = + eval (0H ·ₕ Q) xs ≡⟨ Eval0H xs ⟩ + 0r ≡⟨ sym (0LeftAnnihilates _) ⟩ + 0r · eval Q xs ≡⟨ cong (λ u → u · eval Q xs) (sym (Eval0H xs)) ⟩ + eval 0H xs · eval Q xs ∎ + ·Homeval (P ·X+ Q) S (x ∷ xs) = + eval ((P ·X+ Q) ·ₕ S) (x ∷ xs) + ≡⟨ combineCases· Q P S (x ∷ xs) ⟩ + eval (((P ·ₕ S) ·X+ 0ₕ) +ₕ (Q ⋆ S)) (x ∷ xs) + ≡⟨ +Homeval ((P ·ₕ S) ·X+ 0ₕ) (Q ⋆ S) (x ∷ xs) ⟩ + eval ((P ·ₕ S) ·X+ 0ₕ) (x ∷ xs) + eval (Q ⋆ S) (x ∷ xs) + ≡⟨ cong (λ u → u + eval (Q ⋆ S) (x ∷ xs)) (combineCasesEval R (P ·ₕ S) 0ₕ x xs) ⟩ + (eval (P ·ₕ S) (x ∷ xs) · x + eval 0ₕ xs) + eval (Q ⋆ S) (x ∷ xs) + ≡⟨ cong (λ u → u + eval (Q ⋆ S) (x ∷ xs)) + ((eval (P ·ₕ S) (x ∷ xs) · x + eval 0ₕ xs) + ≡⟨ cong (λ u → eval (P ·ₕ S) (x ∷ xs) · x + u) (Eval0H xs) ⟩ + (eval (P ·ₕ S) (x ∷ xs) · x + 0r) + ≡⟨ +Rid _ ⟩ + (eval (P ·ₕ S) (x ∷ xs) · x) + ≡⟨ cong (λ u → u · x) (·Homeval P S (x ∷ xs)) ⟩ + ((eval P (x ∷ xs) · eval S (x ∷ xs)) · x) + ≡⟨ sym (·Assoc _ _ _) ⟩ + (eval P (x ∷ xs) · (eval S (x ∷ xs) · x)) + ≡⟨ cong (λ u → eval P (x ∷ xs) · u) (·Comm _ _) ⟩ + (eval P (x ∷ xs) · (x · eval S (x ∷ xs))) + ≡⟨ ·Assoc _ _ _ ⟩ + (eval P (x ∷ xs) · x) · eval S (x ∷ xs) + ∎) ⟩ + (eval P (x ∷ xs) · x) · eval S (x ∷ xs) + eval (Q ⋆ S) (x ∷ xs) + ≡⟨ cong (λ u → (eval P (x ∷ xs) · x) · eval S (x ∷ xs) + u) + (⋆Homeval Q S x xs) ⟩ + (eval P (x ∷ xs) · x) · eval S (x ∷ xs) + eval Q xs · eval S (x ∷ xs) + ≡⟨ sym (·Ldist+ _ _ _) ⟩ + ((eval P (x ∷ xs) · x) + eval Q xs) · eval S (x ∷ xs) + ≡⟨ cong (λ u → u · eval S (x ∷ xs)) (sym (combineCasesEval R P Q x xs)) ⟩ + eval (P ·X+ Q) (x ∷ xs) · eval S (x ∷ xs) ∎ diff --git a/Cubical/Algebra/RingSolver/Examples.agda b/Cubical/Algebra/RingSolver/Examples.agda new file mode 100644 index 0000000000..319d97de71 --- /dev/null +++ b/Cubical/Algebra/RingSolver/Examples.agda @@ -0,0 +1,97 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.Examples where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Int.Base hiding (_+_ ; _·_ ; _-_) +open import Cubical.Data.List + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.RingSolver.Reflection + +private + variable + ℓ : Level + + +module Test (R : CommRing ℓ) where + open CommRingStr (snd R) + + _ : 0r ≡ 0r + _ = solve R + + _ : 1r · (1r + 0r) + ≡ (1r · 0r) + 1r + _ = solve R + + _ : 1r · 0r + (1r - 1r) + ≡ 0r - 0r + _ = solve R + + _ : (x : fst R) → x ≡ x + _ = solve R + + _ : (x y : fst R) → x ≡ x + _ = solve R + + _ : (x y : fst R) → x + y ≡ y + x + _ = solve R + + _ : (x y : fst R) → (x + y) · (x - y) ≡ x · x - y · y + _ = solve R + + {- + A bigger example, copied from the other example files: + -} + _ : (x y z : (fst R)) → (x + y) · (x + y) · (x + y) · (x + y) + ≡ x · x · x · x + (fromℤ R 4) · x · x · x · y + (fromℤ R 6) · x · x · y · y + + (fromℤ R 4) · x · y · y · y + y · y · y · y + _ = solve R + + {- + Examples that used to fail (see #513): + -} + + _ : (x : (fst R)) → x · 0r ≡ 0r + _ = solve R + + _ : (x y z : (fst R)) → x · (y - z) ≡ x · y - x · z + _ = solve R + + {- + Keep in mind, that the solver can lead to wrong error locations. + For example, the commented code below tries to solve an equation that does not hold, + with the result of an error at the wrong location. + + _ : (x y : (fst R)) → x ≡ y + _ = solve R + -} + +module TestInPlaceSolving (R : CommRing ℓ) where + open CommRingStr (snd R) + + testWithOneVariabl : (x : fst R) → x + 0r ≡ 0r + x + testWithOneVariabl x = solveInPlace R (x ∷ []) + + testEquationalReasoning : (x : fst R) → x + 0r ≡ 0r + x + testEquationalReasoning x = + x + 0r ≡⟨solveIn R withVars (x ∷ []) ⟩ + 0r + x ∎ + + testWithTwoVariables : (x y : fst R) → x + y ≡ y + x + testWithTwoVariables x y = + x + y ≡⟨solveIn R withVars (x ∷ y ∷ []) ⟩ + y + x ∎ + + {- + So far, solving during equational reasoning has a serious + restriction: + The solver identifies variables by deBruijn indices and the variables + appearing in the equations to solve need to have indices 0,...,n. This + entails that in the following code, the order of 'p' and 'x' cannot be + switched. + -} + testEquationalReasoning' : (p : (y : fst R) → 0r + y ≡ 1r) (x : fst R) → x + 0r ≡ 1r + testEquationalReasoning' p x = + x + 0r ≡⟨solveIn R withVars (x ∷ []) ⟩ + 0r + x ≡⟨ p x ⟩ + 1r ∎ diff --git a/Cubical/Algebra/RingSolver/HornerEval.agda b/Cubical/Algebra/RingSolver/HornerEval.agda new file mode 100644 index 0000000000..9e07e288de --- /dev/null +++ b/Cubical/Algebra/RingSolver/HornerEval.agda @@ -0,0 +1,133 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.HornerEval where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Int hiding (_+_ ; _·_ ; -_) +open import Cubical.Data.Vec +open import Cubical.Data.Bool + +open import Cubical.Relation.Nullary.Base using (¬_; yes; no) + +open import Cubical.Algebra.RingSolver.Utility +open import Cubical.Algebra.RingSolver.RawAlgebra +open import Cubical.Algebra.RingSolver.IntAsRawRing +open import Cubical.Algebra.RingSolver.HornerForms +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring + +private + variable + ℓ ℓ′ : Level + +eval : {A : RawAlgebra ℤAsRawRing ℓ′} + {n : ℕ} (P : IteratedHornerForms A n) + → Vec ⟨ A ⟩ n → ⟨ A ⟩ +eval {A = A} (const r) [] = RawAlgebra.scalar A r +eval {A = A} 0H (_ ∷ _) = RawAlgebra.0r A +eval {A = A} (P ·X+ Q) (x ∷ xs) = + let open RawAlgebra A + P' = (eval P (x ∷ xs)) + Q' = eval Q xs + in if (isZero A P) + then Q' + else P' · x + Q' + +module _ (R : CommRing ℓ) where + private + νR = CommRing→RawℤAlgebra R + open CommRingStr (snd R) + open RingTheory (CommRing→Ring R) + open IteratedHornerOperations νR + + someCalculation : {x : fst R} → _ ≡ _ + someCalculation {x = x} = + 0r ≡⟨ sym (+Rid 0r) ⟩ + 0r + 0r ≡[ i ]⟨ 0LeftAnnihilates x (~ i) + 0r ⟩ + 0r · x + 0r ∎ + + + evalIsZero : {n : ℕ} (P : IteratedHornerForms νR n) + → (l : Vec (fst R) n) + → isZero νR P ≡ true + → eval P l ≡ 0r + evalIsZero (const (pos ℕ.zero)) [] isZeroP = refl + evalIsZero (const (pos (ℕ.suc n))) [] isZeroP = byBoolAbsurdity isZeroP + evalIsZero (const (negsuc _)) [] isZeroP = byBoolAbsurdity isZeroP + evalIsZero 0H (x ∷ xs) _ = refl + evalIsZero {n = ℕ.suc n} (P ·X+ Q) (x ∷ xs) isZeroPandQ with isZero νR P + ... | true = eval Q xs ≡⟨ evalIsZero Q xs isZeroQ ⟩ + 0r ∎ + where isZeroQ = snd (extractFromAnd _ _ isZeroPandQ) + ... | false = byBoolAbsurdity isZeroP + where isZeroP = fst (extractFromAnd _ _ isZeroPandQ) + + computeEvalSummandIsZero : + {n : ℕ} + (P : IteratedHornerForms νR (ℕ.suc n)) + (Q : IteratedHornerForms νR n) + → (xs : Vec (fst R) n) + → (x : (fst R)) + → isZero νR P ≡ true + → eval (P ·X+ Q) (x ∷ xs) ≡ eval Q xs + computeEvalSummandIsZero P Q xs x isZeroP with isZero νR P + ... | true = refl + ... | false = byBoolAbsurdity isZeroP + + computeEvalNotZero : + {n : ℕ} + (P : IteratedHornerForms νR (ℕ.suc n)) + (Q : IteratedHornerForms νR n) + → (xs : Vec (fst R) n) + → (x : (fst R)) + → ¬ (isZero νR P ≡ true) + → eval (P ·X+ Q) (x ∷ xs) ≡ (eval P (x ∷ xs)) · x + eval Q xs + computeEvalNotZero P Q xs x notZeroP with isZero νR P + ... | true = byBoolAbsurdity (sym (¬true→false true notZeroP)) + ... | false = refl + + combineCasesEval : + {n : ℕ} (P : IteratedHornerForms νR (ℕ.suc n)) (Q : IteratedHornerForms νR n) + (x : (fst R)) (xs : Vec (fst R) n) + → eval (P ·X+ Q) (x ∷ xs) + ≡ (eval P (x ∷ xs)) · x + eval Q xs + combineCasesEval P Q x xs with isZero νR P ≟ true + ... | yes p = + eval (P ·X+ Q) (x ∷ xs) ≡⟨ computeEvalSummandIsZero P Q xs x p ⟩ + eval Q xs ≡⟨ sym (+Lid _) ⟩ + 0r + eval Q xs ≡[ i ]⟨ 0LeftAnnihilates x (~ i) + eval Q xs ⟩ + 0r · x + eval Q xs ≡[ i ]⟨ (evalIsZero P (x ∷ xs) p (~ i)) · x + eval Q xs ⟩ + (eval P (x ∷ xs)) · x + eval Q xs ∎ + ... | no p = computeEvalNotZero P Q xs x p + + + compute+ₕEvalBothZero : + (n : ℕ) (P Q : IteratedHornerForms νR (ℕ.suc n)) + (r s : IteratedHornerForms νR n) + (x : (fst R)) (xs : Vec (fst R) n) + → (isZero νR (P +ₕ Q) and isZero νR (r +ₕ s)) ≡ true + → eval ((P ·X+ r) +ₕ (Q ·X+ s)) (x ∷ xs) ≡ eval ((P +ₕ Q) ·X+ (r +ₕ s)) (x ∷ xs) + compute+ₕEvalBothZero n P Q r s x xs bothZero with isZero νR (P +ₕ Q) and isZero νR (r +ₕ s) | bothZero + ... | true | p = + eval {A = νR} 0H (x ∷ xs) ≡⟨ refl ⟩ + 0r ≡⟨ someCalculation ⟩ + 0r · x + 0r ≡⟨ step1 ⟩ + (eval (P +ₕ Q) (x ∷ xs)) · x + eval (r +ₕ s) xs ≡⟨ step2 ⟩ + eval ((P +ₕ Q) ·X+ (r +ₕ s)) (x ∷ xs) ∎ + where step1 : 0r · x + 0r ≡ (eval (P +ₕ Q) (x ∷ xs)) · x + eval (r +ₕ s) xs + step1 i = (evalIsZero (P +ₕ Q) (x ∷ xs) (fst (extractFromAnd _ _ (bothZero))) (~ i)) · x + + (evalIsZero (r +ₕ s) xs (snd (extractFromAnd _ _ (bothZero))) (~ i)) + step2 = sym (combineCasesEval (P +ₕ Q) (r +ₕ s) x xs) + ... | false | p = byBoolAbsurdity p + + compute+ₕEvalNotBothZero : + (n : ℕ) (P Q : IteratedHornerForms νR (ℕ.suc n)) + (r s : IteratedHornerForms νR n) + (x : (fst R)) (xs : Vec (fst R) n) + → (isZero νR (P +ₕ Q) and isZero νR (r +ₕ s)) ≡ false + → eval ((P ·X+ r) +ₕ (Q ·X+ s)) (x ∷ xs) ≡ eval ((P +ₕ Q) ·X+ (r +ₕ s)) (x ∷ xs) + compute+ₕEvalNotBothZero n P Q r s _ _ notBothZero + with isZero νR (P +ₕ Q) and isZero νR (r +ₕ s) | notBothZero + ... | true | p = byBoolAbsurdity (sym p) + ... | false | p = refl diff --git a/Cubical/Algebra/RingSolver/HornerForms.agda b/Cubical/Algebra/RingSolver/HornerForms.agda new file mode 100644 index 0000000000..9c3679f7a5 --- /dev/null +++ b/Cubical/Algebra/RingSolver/HornerForms.agda @@ -0,0 +1,184 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.HornerForms where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Int hiding (_+_ ; _·_ ; -_) +open import Cubical.Data.FinData +open import Cubical.Data.Vec +open import Cubical.Data.Bool + +open import Cubical.Relation.Nullary.Base using (yes; no) + +open import Cubical.Algebra.RingSolver.Utility + +open import Cubical.Algebra.RingSolver.RawRing +open import Cubical.Algebra.RingSolver.IntAsRawRing +open import Cubical.Algebra.RingSolver.RawAlgebra renaming (⟨_⟩ to ⟨_⟩ₐ) +open import Cubical.Algebra.RingSolver.AlgebraExpression public + +private + variable + ℓ ℓ′ : Level + +{- + This defines the type of multivariate Polynomials over the RawRing R. + The construction is based on the algebraic fact + + R[X₀][X₁]⋯[Xₙ] ≅ R[X₀,⋯,Xₙ] + + BUT: Contrary to algebraic convetions, we will give 'Xₙ' the lowest index + in the definition of 'Variable' below. So if 'Variable n R k' is identified + with 'Xₖ', then the RawRing we construct should rather be denoted with + + R[Xₙ][Xₙ₋₁]⋯[X₀] + + or, to be precise about the evaluation order: + + (⋯((R[Xₙ])[Xₙ₋₁])⋯)[X₀] + +-} + +data IteratedHornerForms (A : RawAlgebra ℤAsRawRing ℓ) : ℕ → Type ℓ where + const : ℤ → IteratedHornerForms A ℕ.zero + 0H : {n : ℕ} → IteratedHornerForms A (ℕ.suc n) + _·X+_ : {n : ℕ} → IteratedHornerForms A (ℕ.suc n) → IteratedHornerForms A n + → IteratedHornerForms A (ℕ.suc n) + + +{- + The following function returns true, if there is some + obvious reason that the Horner-Expression should be zero. + Since Equality is undecidable in a general RawAlgebra, we cannot + have a function that fully lives up to the name 'isZero'. +-} +module _ (A : RawAlgebra ℤAsRawRing ℓ) where + open RawRing ℤAsRawRing + isZero : {n : ℕ} → IteratedHornerForms A n + → Bool + isZero (const (pos ℕ.zero)) = true + isZero (const (pos (ℕ.suc _))) = false + isZero (const (negsuc _)) = false + isZero 0H = true + isZero (P ·X+ Q) = (isZero P) and (isZero Q) + + leftIsZero : {n : ℕ} + (P : IteratedHornerForms A (ℕ.suc n)) + (Q : IteratedHornerForms A n) + → isZero (P ·X+ Q) ≡ true + → isZero P ≡ true + leftIsZero P Q isZeroSum with isZero P + ... | true = refl + ... | false = byBoolAbsurdity (fst (extractFromAnd _ _ isZeroSum)) + + rightIsZero : {n : ℕ} + (P : IteratedHornerForms A (ℕ.suc n)) + (Q : IteratedHornerForms A n) + → isZero (P ·X+ Q) ≡ true + → isZero Q ≡ true + rightIsZero P Q isZeroSum with isZero Q + ... | true = refl + ... | false = byBoolAbsurdity (snd (extractFromAnd _ _ isZeroSum)) + +module IteratedHornerOperations (A : RawAlgebra ℤAsRawRing ℓ) where + open RawRing ℤAsRawRing + + private + 1H' : (n : ℕ) → IteratedHornerForms A n + 1H' ℕ.zero = const 1r + 1H' (ℕ.suc n) = 0H ·X+ 1H' n + + 0H' : (n : ℕ) → IteratedHornerForms A n + 0H' ℕ.zero = const 0r + 0H' (ℕ.suc n) = 0H + + 1ₕ : {n : ℕ} → IteratedHornerForms A n + 1ₕ {n = n} = 1H' n + + 0ₕ : {n : ℕ} → IteratedHornerForms A n + 0ₕ {n = n} = 0H' n + + X : (n : ℕ) (k : Fin n) → IteratedHornerForms A n + X (ℕ.suc m) zero = 1ₕ ·X+ 0ₕ + X (ℕ.suc m) (suc k) = 0ₕ ·X+ X m k + + _+ₕ_ : {n : ℕ} → IteratedHornerForms A n → IteratedHornerForms A n + → IteratedHornerForms A n + (const r) +ₕ (const s) = const (r + s) + 0H +ₕ Q = Q + (P ·X+ r) +ₕ 0H = P ·X+ r + (P ·X+ r) +ₕ (Q ·X+ s) = + let left = (P +ₕ Q) + right = (r +ₕ s) + in if ((isZero A left) and (isZero A right)) + then 0ₕ + else left ·X+ right + + -ₕ : {n : ℕ} → IteratedHornerForms A n → IteratedHornerForms A n + -ₕ (const x) = const (- x) + -ₕ 0H = 0H + -ₕ (P ·X+ Q) = (-ₕ P) ·X+ (-ₕ Q) + + _⋆_ : {n : ℕ} → IteratedHornerForms A n → IteratedHornerForms A (ℕ.suc n) + → IteratedHornerForms A (ℕ.suc n) + _·ₕ_ : {n : ℕ} → IteratedHornerForms A n → IteratedHornerForms A n + → IteratedHornerForms A n + r ⋆ 0H = 0H + r ⋆ (P ·X+ Q) = + if (isZero A r) + then 0ₕ + else (r ⋆ P) ·X+ (r ·ₕ Q) + + const x ·ₕ const y = const (x · y) + 0H ·ₕ Q = 0H + (P ·X+ Q) ·ₕ S = + let + z = (P ·ₕ S) + in if (isZero A z) + then (Q ⋆ S) + else (z ·X+ 0ₕ) +ₕ (Q ⋆ S) + + isZeroPresLeft⋆ : + {n : ℕ} + (r : IteratedHornerForms A n) + (P : IteratedHornerForms A (ℕ.suc n)) + → isZero A r ≡ true + → isZero A (r ⋆ P) ≡ true + isZeroPresLeft⋆ r 0H isZero-r = refl + isZeroPresLeft⋆ r (P ·X+ Q) isZero-r with isZero A r + ... | true = refl + ... | false = byBoolAbsurdity isZero-r + + isZeroPresLeft·ₕ : + {n : ℕ} (P Q : IteratedHornerForms A n) + → isZero A P ≡ true + → isZero A (P ·ₕ Q) ≡ true + isZeroPresLeft·ₕ (const (pos ℕ.zero)) (const _) isZeroP = refl + isZeroPresLeft·ₕ (const (pos (ℕ.suc n))) (const _) isZeroP = byBoolAbsurdity isZeroP + isZeroPresLeft·ₕ (const (negsuc n)) (const _) isZeroP = byBoolAbsurdity isZeroP + isZeroPresLeft·ₕ 0H Q isZeroP = refl + isZeroPresLeft·ₕ (P ·X+ Q) S isZeroSum with isZero A (P ·ₕ S) ≟ true + ... | no p = byBoolAbsurdity (sym notZeroProd ∙ isZeroProd) + where notZeroProd = ¬true→false _ p + isZeroP = extractFromAndLeft isZeroSum + isZeroProd = isZeroPresLeft·ₕ P S isZeroP + ... | yes p with isZero A (P ·ₕ S) + ... | true = isZeroPresLeft⋆ Q S isZeroQ + where isZeroQ = extractFromAndRight isZeroSum + ... | false = byBoolAbsurdity p + + asRawRing : (n : ℕ) → RawRing ℓ + RawRing.Carrier (asRawRing n) = IteratedHornerForms A n + RawRing.0r (asRawRing n) = 0ₕ + RawRing.1r (asRawRing n) = 1ₕ + RawRing._+_ (asRawRing n) = _+ₕ_ + RawRing._·_ (asRawRing n) = _·ₕ_ + RawRing.- (asRawRing n) = -ₕ + +Variable : (n : ℕ) (R : RawAlgebra ℤAsRawRing ℓ′) (k : Fin n) → IteratedHornerForms R n +Variable n R k = IteratedHornerOperations.X R n k + +Constant : (n : ℕ) (R : RawAlgebra ℤAsRawRing ℓ′) (r : ℤ) → IteratedHornerForms R n +Constant ℕ.zero R r = const r +Constant (ℕ.suc n) R r = IteratedHornerOperations.0ₕ R ·X+ Constant n R r diff --git a/Cubical/Algebra/RingSolver/IntAsRawRing.agda b/Cubical/Algebra/RingSolver/IntAsRawRing.agda new file mode 100644 index 0000000000..5055a15468 --- /dev/null +++ b/Cubical/Algebra/RingSolver/IntAsRawRing.agda @@ -0,0 +1,15 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.IntAsRawRing where + +open import Cubical.Data.Nat hiding (_+_; _·_) +open import Cubical.Data.Int +open import Cubical.Data.Int.Base + +open import Cubical.Foundations.Prelude +open import Cubical.Algebra.RingSolver.RawRing + +ℤAsRawRing : RawRing ℓ-zero +ℤAsRawRing = rawring ℤ (pos zero) (pos (suc zero)) _+_ _·_ (λ k → - k) + ++Ridℤ : (k : ℤ) → (pos zero) + k ≡ k ++Ridℤ k = sym (pos0+ k) diff --git a/Cubical/Algebra/RingSolver/README.md b/Cubical/Algebra/RingSolver/README.md new file mode 100644 index 0000000000..c87573f3dc --- /dev/null +++ b/Cubical/Algebra/RingSolver/README.md @@ -0,0 +1,25 @@ +Ring solving +======================================== + +This is a ring solver similar to this one: + +https://github.com/oisdk/agda-ring-solver + +The latter is explained in this thesis: + +https://github.com/oisdk/agda-ring-solver-report/blob/master/report.pdf + +The thesis refers to this paper: + +"Proving Equalities in a Commutative Ring Done Right in Coq" +https://link.springer.com/content/pdf/10.1007%2F11541868_7.pdf + +There are three parts of the appraoch to prove x=y: +* turn x and y into Expressions (i.e. syntax trees) using reflection +* map the expressions to polynomials in horner form (normalize) +* let agda compare the results with unification + +There are two versions of the solver, one which works for natural numbers and one which works for commutatitive rings (CommRing). Only the CommRingSolver has a reflection interface and the NatSolver has some additional problems, but it could still be useful. + +To see how the ring solver might be used, check out 'Examples.agda'. To understand how it works, it is probably good to have a look at 'NatExamples.agda' and 'CommRingExamples.agda'. + diff --git a/Cubical/Algebra/RingSolver/RawAlgebra.agda b/Cubical/Algebra/RingSolver/RawAlgebra.agda new file mode 100644 index 0000000000..f7f1b9476b --- /dev/null +++ b/Cubical/Algebra/RingSolver/RawAlgebra.agda @@ -0,0 +1,177 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.RawAlgebra where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Int renaming (_+_ to _+ℤ_ ; _·_ to _·ℤ_ ; -_ to -ℤ_ ; _-_ to _-ℤ_ ; +Assoc to +ℤAssoc ; +Comm to +ℤComm ; -DistL· to -ℤDistL·ℤ) + +open import Cubical.Algebra.RingSolver.RawRing renaming (⟨_⟩ to ⟨_⟩ᵣ) +open import Cubical.Algebra.RingSolver.IntAsRawRing +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring + +private + variable + ℓ ℓ′ : Level + +record RawAlgebra (R : RawRing ℓ) (ℓ′ : Level) : Type (ℓ-suc (ℓ-max ℓ ℓ′)) where + + constructor rawalgebra + + field + Carrier : Type ℓ′ + scalar : ⟨ R ⟩ᵣ → Carrier + 0r : Carrier + 1r : Carrier + _+_ : Carrier → Carrier → Carrier + _·_ : Carrier → Carrier → Carrier + -_ : Carrier → Carrier + + infixl 8 _·_ + infixl 7 -_ + infixl 6 _+_ + +⟨_⟩ : {R : RawRing ℓ} → RawAlgebra R ℓ′ → Type ℓ′ +⟨_⟩ = RawAlgebra.Carrier + +{- + Mapping to integer scalars and its (homorphism) properties. +-} +module _ (R : CommRing ℓ) where + open CommRingStr (snd R) + open Cubical.Algebra.Ring.RingTheory (CommRing→Ring R) + + scalarℕ : ℕ → (fst R) + scalarℕ ℕ.zero = 0r + scalarℕ (ℕ.suc ℕ.zero) = 1r + scalarℕ (ℕ.suc (ℕ.suc n)) = 1r + scalarℕ (ℕ.suc n) + + scalar : ℤ → (fst R) + scalar (pos k) = scalarℕ k + scalar (negsuc k) = - scalarℕ (ℕ.suc k) + + -DistScalar : (k : ℤ) + → scalar (-ℤ k) ≡ - (scalar k) + -DistScalar (pos ℕ.zero) = sym 0Selfinverse + -DistScalar (pos (ℕ.suc n)) = refl + -DistScalar (negsuc n) = sym (-Idempotent _) + + lemmaSuc : (k : ℤ) + → scalar (sucℤ k) ≡ 1r + scalar k + lemmaSuc (pos ℕ.zero) = sym (+Rid _) + lemmaSuc (pos (ℕ.suc ℕ.zero)) = refl + lemmaSuc (pos (ℕ.suc (ℕ.suc n))) = refl + lemmaSuc (negsuc ℕ.zero) = sym (+Rinv _) + lemmaSuc (negsuc (ℕ.suc n)) = + scalar (negsuc n) ≡⟨ sym (+Lid (scalar (negsuc n))) ⟩ + 0r + scalar (negsuc n) ≡[ i ]⟨ +Rinv 1r (~ i) + scalar (negsuc n) ⟩ + (1r - 1r) + scalar (negsuc n) ≡⟨ sym (+Assoc _ _ _) ⟩ + 1r + (- 1r + - scalar (pos (ℕ.suc n))) ≡[ i ]⟨ 1r + -Dist 1r (scalar (pos (ℕ.suc n))) i ⟩ + 1r + -(1r + scalar (pos (ℕ.suc n))) ≡⟨ refl ⟩ + 1r + -(scalar (pos (ℕ.suc (ℕ.suc n)))) ≡⟨ refl ⟩ + 1r + scalar (negsuc (ℕ.suc n)) ∎ + + lemmaPred : (k : ℤ) + → scalar (predℤ k) ≡ - 1r + scalar k + lemmaPred k = sym( + - 1r + scalar k ≡[ i ]⟨ - 1r + scalar (sucPred k (~ i)) ⟩ + - 1r + scalar (sucℤ (predℤ k)) ≡[ i ]⟨ - 1r + lemmaSuc (predℤ k) i ⟩ + - 1r + (1r + scalar (predℤ k)) ≡⟨ +Assoc _ _ _ ⟩ + (- 1r + 1r) + scalar (predℤ k) ≡[ i ]⟨ +Linv 1r i + scalar (predℤ k) ⟩ + 0r + scalar (predℤ k) ≡⟨ +Lid _ ⟩ + scalar (predℤ k) ∎) + + +HomScalar : (k l : ℤ) + → scalar (k +ℤ l) ≡ (scalar k) + (scalar l) + +HomScalar (pos ℕ.zero) l = + scalar (0 +ℤ l) ≡[ i ]⟨ scalar (sym (pos0+ l) i) ⟩ + scalar l ≡⟨ sym (+Lid _) ⟩ + 0r + scalar l ≡⟨ refl ⟩ + scalar 0 + scalar l ∎ + + +HomScalar (pos (ℕ.suc ℕ.zero)) l = + scalar (1 +ℤ l) ≡[ i ]⟨ scalar (+ℤComm 1 l i) ⟩ + scalar (l +ℤ 1) ≡⟨ refl ⟩ + scalar (sucℤ l) ≡⟨ lemmaSuc l ⟩ + 1r + scalar l ≡⟨ refl ⟩ + scalar (pos (ℕ.suc ℕ.zero)) + scalar l ∎ + + +HomScalar (pos (ℕ.suc (ℕ.suc n))) l = + scalar (pos (ℕ.suc (ℕ.suc n)) +ℤ l) ≡⟨ refl ⟩ + scalar ((pos (ℕ.suc n) +ℤ 1) +ℤ l) ≡[ i ]⟨ scalar ((+ℤComm (pos (ℕ.suc n)) 1 i) +ℤ l) ⟩ + scalar ((1 +ℤ (pos (ℕ.suc n))) +ℤ l) ≡[ i ]⟨ scalar (+ℤAssoc 1 (pos (ℕ.suc n)) l (~ i)) ⟩ + scalar (1 +ℤ (pos (ℕ.suc n) +ℤ l)) ≡⟨ +HomScalar (pos (ℕ.suc ℕ.zero)) (pos (ℕ.suc n) +ℤ l) ⟩ + scalar 1 + scalar (pos (ℕ.suc n) +ℤ l) ≡⟨ refl ⟩ + 1r + (scalar (pos (ℕ.suc n) +ℤ l)) ≡[ i ]⟨ 1r + +HomScalar (pos (ℕ.suc n)) l i ⟩ + 1r + (scalar (pos (ℕ.suc n)) + scalar l) ≡⟨ +Assoc _ _ _ ⟩ + (1r + scalar (pos (ℕ.suc n))) + scalar l ≡⟨ refl ⟩ + scalar (pos (ℕ.suc (ℕ.suc n))) + scalar l ∎ + + +HomScalar (negsuc ℕ.zero) l = + scalar (-1 +ℤ l) ≡[ i ]⟨ scalar (+ℤComm -1 l i) ⟩ + scalar (l +ℤ -1) ≡⟨ refl ⟩ + scalar (predℤ l) ≡⟨ lemmaPred l ⟩ + - 1r + scalar l ≡⟨ refl ⟩ + scalar -1 + scalar l ∎ + + +HomScalar (negsuc (ℕ.suc n)) l = + scalar (negsuc (ℕ.suc n) +ℤ l) ≡⟨ refl ⟩ + scalar ((negsuc n +ℤ -1) +ℤ l) ≡[ i ]⟨ scalar (+ℤComm (negsuc n) -1 i +ℤ l) ⟩ + scalar ((-1 +ℤ negsuc n) +ℤ l) ≡[ i ]⟨ scalar (+ℤAssoc -1 (negsuc n) l (~ i)) ⟩ + scalar (-1 +ℤ (negsuc n +ℤ l)) ≡⟨ +HomScalar -1 (negsuc n +ℤ l) ⟩ + - 1r + scalar (negsuc n +ℤ l) ≡[ i ]⟨ - 1r + +HomScalar (negsuc n) l i ⟩ + - 1r + (scalar (negsuc n) + scalar l) ≡⟨ +Assoc (- 1r) _ _ ⟩ + (- 1r + (scalar (negsuc n))) + scalar l ≡⟨ refl ⟩ + (- 1r + - scalar (pos (ℕ.suc n))) + scalar l ≡[ i ]⟨ -Dist 1r (scalar (pos (ℕ.suc n))) i + scalar l ⟩ + (- (1r + scalar (pos (ℕ.suc n)))) + scalar l ≡⟨ refl ⟩ + scalar (negsuc (ℕ.suc n)) + scalar l ∎ + + + lemma1 : (n : ℕ) + → 1r + scalar (pos n) ≡ scalar (pos (ℕ.suc n)) + lemma1 ℕ.zero = +Rid _ + lemma1 (ℕ.suc k) = refl + + lemma-1 : (n : ℕ) + → - 1r + scalar (negsuc n) ≡ scalar (negsuc (ℕ.suc n)) + lemma-1 ℕ.zero = -Dist _ _ + lemma-1 (ℕ.suc k) = + - 1r + scalar (negsuc (ℕ.suc k)) ≡⟨ refl ⟩ + - 1r + - scalar (pos (ℕ.suc (ℕ.suc k))) ≡⟨ -Dist _ _ ⟩ + - (1r + scalar (pos (ℕ.suc (ℕ.suc k)))) ≡⟨ refl ⟩ + scalar (negsuc (ℕ.suc (ℕ.suc k))) ∎ + + ·HomScalar : (k l : ℤ) + → scalar (k ·ℤ l) ≡ scalar k · scalar l + ·HomScalar (pos ℕ.zero) l = 0r ≡⟨ sym (0LeftAnnihilates (scalar l)) ⟩ 0r · scalar l ∎ + ·HomScalar (pos (ℕ.suc n)) l = + scalar (l +ℤ (pos n ·ℤ l)) ≡⟨ +HomScalar l (pos n ·ℤ l) ⟩ + scalar l + scalar (pos n ·ℤ l) ≡[ i ]⟨ scalar l + ·HomScalar (pos n) l i ⟩ + scalar l + (scalar (pos n) · scalar l) ≡[ i ]⟨ ·Lid (scalar l) (~ i) + (scalar (pos n) · scalar l) ⟩ + 1r · scalar l + (scalar (pos n) · scalar l) ≡⟨ sym (·Ldist+ 1r _ _) ⟩ + (1r + scalar (pos n)) · scalar l ≡[ i ]⟨ lemma1 n i · scalar l ⟩ + scalar (pos (ℕ.suc n)) · scalar l ∎ + + ·HomScalar (negsuc ℕ.zero) l = + scalar (-ℤ l) ≡⟨ -DistScalar l ⟩ + - scalar l ≡[ i ]⟨ - (·Lid (scalar l) (~ i)) ⟩ + - (1r · scalar l) ≡⟨ sym (-DistL· _ _) ⟩ + - 1r · scalar l ≡⟨ refl ⟩ + scalar (negsuc ℕ.zero) · scalar l ∎ + + ·HomScalar (negsuc (ℕ.suc n)) l = + scalar ((-ℤ l) +ℤ (negsuc n ·ℤ l)) ≡⟨ +HomScalar (-ℤ l) (negsuc n ·ℤ l) ⟩ + scalar (-ℤ l) + scalar (negsuc n ·ℤ l) ≡[ i ]⟨ -DistScalar l i + scalar (negsuc n ·ℤ l) ⟩ + - scalar l + scalar (negsuc n ·ℤ l) ≡[ i ]⟨ - scalar l + ·HomScalar (negsuc n) l i ⟩ + - scalar l + scalar (negsuc n) · scalar l ≡[ i ]⟨ (- ·Lid (scalar l) (~ i)) + + scalar (negsuc n) · scalar l ⟩ + - (1r · scalar l) + scalar (negsuc n) · scalar l ≡[ i ]⟨ -DistL· 1r (scalar l) (~ i) + + scalar (negsuc n) · scalar l ⟩ + - 1r · scalar l + scalar (negsuc n) · scalar l ≡⟨ sym (·Ldist+ _ _ _) ⟩ + (- 1r + scalar (negsuc n)) · scalar l ≡[ i ]⟨ lemma-1 n i · scalar l ⟩ + scalar (negsuc (ℕ.suc n)) · scalar l ∎ + +CommRing→RawℤAlgebra : CommRing ℓ → RawAlgebra ℤAsRawRing ℓ +CommRing→RawℤAlgebra (R , commringstr 0r 1r _+_ _·_ -_ isCommRing) = + rawalgebra R (scalar ((R , commringstr 0r 1r _+_ _·_ -_ isCommRing))) 0r 1r _+_ _·_ -_ diff --git a/Cubical/Algebra/RingSolver/RawRing.agda b/Cubical/Algebra/RingSolver/RawRing.agda new file mode 100644 index 0000000000..605f42d54c --- /dev/null +++ b/Cubical/Algebra/RingSolver/RawRing.agda @@ -0,0 +1,29 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.RawRing where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ) + +private + variable + ℓ : Level + +record RawRing ℓ : Type (ℓ-suc ℓ) where + + constructor rawring + + field + Carrier : Type ℓ + 0r : Carrier + 1r : Carrier + _+_ : Carrier → Carrier → Carrier + _·_ : Carrier → Carrier → Carrier + -_ : Carrier → Carrier + + infixl 8 _·_ + infixl 7 -_ + infixl 6 _+_ + +⟨_⟩ : RawRing ℓ → Type ℓ +⟨_⟩ = RawRing.Carrier diff --git a/Cubical/Algebra/RingSolver/Reflection.agda b/Cubical/Algebra/RingSolver/Reflection.agda new file mode 100644 index 0000000000..afe886a50c --- /dev/null +++ b/Cubical/Algebra/RingSolver/Reflection.agda @@ -0,0 +1,328 @@ +{-# OPTIONS --safe #-} +{- + This is inspired by/copied from: + https://github.com/agda/agda-stdlib/blob/master/src/Tactic/MonoidSolver.agda + + Boilerplate code for calling the ring solver is constructed automatically + with agda's reflection features. +-} +module Cubical.Algebra.RingSolver.Reflection where + +open import Cubical.Foundations.Prelude hiding (Type) +open import Cubical.Functions.Logic + +open import Agda.Builtin.Reflection hiding (Type) +open import Agda.Builtin.String + +open import Cubical.Reflection.Base + +open import Cubical.Data.Maybe +open import Cubical.Data.Sigma +open import Cubical.Data.List +open import Cubical.Data.Nat.Literals +open import Cubical.Data.Int.Base hiding (abs) +open import Cubical.Data.Int using (fromNegℤ; fromNatℤ) +open import Cubical.Data.Nat using (ℕ) renaming (_+_ to _+ℕ_) +open import Cubical.Data.FinData using () renaming (zero to fzero; suc to fsuc) +open import Cubical.Data.Bool +open import Cubical.Data.Bool.SwitchStatement +open import Cubical.Data.Vec using (Vec) renaming ([] to emptyVec; _∷_ to _∷vec_) public + +open import Cubical.Algebra.RingSolver.AlgebraExpression +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.RingSolver.RawAlgebra +open import Cubical.Algebra.RingSolver.IntAsRawRing +open import Cubical.Algebra.RingSolver.Solver renaming (solve to ringSolve) + +private + variable + ℓ : Level + + _==_ = primQNameEquality + {-# INLINE _==_ #-} + + record VarInfo : Type ℓ-zero where + field + varName : String + varType : Arg Term + index : ℕ + + {- + `getLastTwoArgsOf` maps a term 'def n (z₁ ∷ … ∷ zₙ ∷ x ∷ y ∷ [])' to the pair '(x,y)' + non-visible arguments are ignored. + -} + getLastTwoArgsOf : Name → Term → Maybe (Term × Term) + getLastTwoArgsOf n' (def n xs) = + if n == n' + then go xs + else nothing + where + go : List (Arg Term) → Maybe (Term × Term) + go (varg x ∷ varg y ∷ []) = just (x , y) + go (x ∷ xs) = go xs + go _ = nothing + getLastTwoArgsOf n' _ = nothing + + {- + `getArgs` maps a term 'x ≡ y' to the pair '(x,y)' + -} + getArgs : Term → Maybe (Term × Term) + getArgs = getLastTwoArgsOf (quote PathP) + + + firstVisibleArg : List (Arg Term) → Maybe Term + firstVisibleArg [] = nothing + firstVisibleArg (varg x ∷ l) = just x + firstVisibleArg (x ∷ l) = firstVisibleArg l + + {- + If the solver needs to be applied during equational reasoning, + the right hand side of the equation to solve cannot be given to + the solver directly. The folllowing function extracts this term y + from a more complex expression as in: + x ≡⟨ solve ... ⟩ (y ≡⟨ ... ⟩ z ∎) + -} + getRhs : Term → Maybe Term + getRhs reasoningToTheRight@(def n xs) = + if n == (quote _∎) + then firstVisibleArg xs + else (if n == (quote _≡⟨_⟩_) + then firstVisibleArg xs + else nothing) + getRhs _ = nothing + + + private + solverCallAsTerm : Term → Arg Term → Term → Term → Term + solverCallAsTerm R varList lhs rhs = + def + (quote ringSolve) + (varg R ∷ varg lhs ∷ varg rhs + ∷ varList + ∷ varg (def (quote refl) []) ∷ []) + + solverCallWithLambdas : ℕ → List VarInfo → Term → Term → Term → Term + solverCallWithLambdas n varInfos R lhs rhs = + encloseWithIteratedLambda + (map VarInfo.varName varInfos) + (solverCallAsTerm R (variableList (rev varInfos)) lhs rhs) + where + encloseWithIteratedLambda : List String → Term → Term + encloseWithIteratedLambda (varName ∷ xs) t = lam visible (abs varName (encloseWithIteratedLambda xs t)) + encloseWithIteratedLambda [] t = t + + variableList : List VarInfo → Arg Term + variableList [] = varg (con (quote emptyVec) []) + variableList (varInfo ∷ varInfos) + = varg (con (quote _∷vec_) (varg (var (VarInfo.index varInfo) []) ∷ (variableList varInfos) ∷ [])) + + solverCallByVarIndices : ℕ → List ℕ → Term → Term → Term → Term + solverCallByVarIndices n varIndices R lhs rhs = + solverCallAsTerm R (variableList (rev varIndices)) lhs rhs + where + variableList : List ℕ → Arg Term + variableList [] = varg (con (quote emptyVec) []) + variableList (varIndex ∷ varIndices) + = varg (con (quote _∷vec_) (varg (var (varIndex) []) ∷ (variableList varIndices) ∷ [])) + + + +module pr (R : CommRing ℓ) {n : ℕ} where + open CommRingStr (snd R) + + 0' : Expr ℤAsRawRing (fst R) n + 0' = K 0 + + 1' : Expr ℤAsRawRing (fst R) n + 1' = K 1 + +module _ (cring : Term) where + open pr + + `0` : List (Arg Term) → Term + `0` [] = def (quote 0') (varg cring ∷ []) + `0` (varg fstcring ∷ xs) = `0` xs + `0` (harg _ ∷ xs) = `0` xs + `0` _ = unknown + + `1` : List (Arg Term) → Term + `1` [] = def (quote 1') (varg cring ∷ []) + `1` (varg fstcring ∷ xs) = `1` xs + `1` (harg _ ∷ xs) = `1` xs + `1` _ = unknown + + mutual + + `_·_` : List (Arg Term) → Term + `_·_` (harg _ ∷ xs) = `_·_` xs + `_·_` (varg _ ∷ varg x ∷ varg y ∷ []) = + con + (quote _·'_) (varg (buildExpression x) ∷ varg (buildExpression y) ∷ []) + `_·_` _ = unknown + + `_+_` : List (Arg Term) → Term + `_+_` (harg _ ∷ xs) = `_+_` xs + `_+_` (varg _ ∷ varg x ∷ varg y ∷ []) = + con + (quote _+'_) (varg (buildExpression x) ∷ varg (buildExpression y) ∷ []) + `_+_` _ = unknown + + `-_` : List (Arg Term) → Term + `-_` (harg _ ∷ xs) = `-_` xs + `-_` (varg _ ∷ varg x ∷ []) = + con + (quote -'_) (varg (buildExpression x) ∷ []) + `-_` _ = unknown + + K' : List (Arg Term) → Term + K' xs = con (quote K) xs + + finiteNumberAsTerm : ℕ → Term + finiteNumberAsTerm ℕ.zero = con (quote fzero) [] + finiteNumberAsTerm (ℕ.suc n) = con (quote fsuc) (varg (finiteNumberAsTerm n) ∷ []) + + buildExpression : Term → Term + buildExpression (var index _) = con (quote ∣) (varg (finiteNumberAsTerm index) ∷ []) + buildExpression t@(def n xs) = + switch (n ==_) cases + case (quote CommRingStr.0r) ⇒ `0` xs break + case (quote CommRingStr.1r) ⇒ `1` xs break + case (quote CommRingStr._·_) ⇒ `_·_` xs break + case (quote CommRingStr._+_) ⇒ `_+_` xs break + case (quote (CommRingStr.-_)) ⇒ `-_` xs break + default⇒ (K' xs) + buildExpression t@(con n xs) = + switch (n ==_) cases + case (quote CommRingStr.0r) ⇒ `0` xs break + case (quote CommRingStr.1r) ⇒ `1` xs break + case (quote CommRingStr._·_) ⇒ `_·_` xs break + case (quote CommRingStr._+_) ⇒ `_+_` xs break + case (quote (CommRingStr.-_)) ⇒ `-_` xs break + default⇒ (K' xs) + buildExpression t = unknown + + toAlgebraExpression : Maybe (Term × Term) → Maybe (Term × Term) + toAlgebraExpression nothing = nothing + toAlgebraExpression (just (lhs , rhs)) = just (buildExpression lhs , buildExpression rhs) + +private + adjustDeBruijnIndex : (n : ℕ) → Term → Term + adjustDeBruijnIndex n (var k args) = var (k +ℕ n) args + adjustDeBruijnIndex n _ = unknown + + extractVarIndices : Maybe (List Term) → Maybe (List ℕ) + extractVarIndices (just ((var index _) ∷ l)) with extractVarIndices (just l) + ... | just indices = just (index ∷ indices) + ... | nothing = nothing + extractVarIndices (just []) = just [] + extractVarIndices _ = nothing + + getVarsAndEquation : Term → Maybe (List VarInfo × Term) + getVarsAndEquation t = + let + (rawVars , equationTerm) = extractVars t + maybeVars = addIndices (length rawVars) rawVars + in map-Maybe (_, equationTerm) maybeVars + where + extractVars : Term → List (String × Arg Term) × Term + extractVars (pi argType (abs varName t)) with extractVars t + ... | xs , equation + = (varName , argType) ∷ xs , equation + extractVars equation = [] , equation + + addIndices : ℕ → List (String × Arg Term) → Maybe (List VarInfo) + addIndices ℕ.zero [] = just [] + addIndices (ℕ.suc countVar) ((varName , argType) ∷ list) = + map-Maybe (λ varList → record { varName = varName ; varType = argType ; index = countVar } + ∷ varList) + (addIndices countVar list) + addIndices _ _ = nothing + + toListOfTerms : Term → Maybe (List Term) + toListOfTerms (con c []) = if (c == (quote [])) then just [] else nothing + toListOfTerms (con c (varg t ∷ varg s ∷ args)) with toListOfTerms s + ... | just terms = if (c == (quote _∷_)) then just (t ∷ terms) else nothing + ... | nothing = nothing + toListOfTerms (con c (harg t ∷ args)) = toListOfTerms (con c args) + toListOfTerms _ = nothing + + solve-macro : Term → Term → TC Unit + solve-macro cring hole = + do + hole′ ← inferType hole >>= normalise + just (varInfos , equation) ← returnTC (getVarsAndEquation hole′) + where + nothing + → typeError (strErr "Something went wrong when getting the variable names in " + ∷ termErr hole′ ∷ []) + {- + The call to the ring solver will be inside a lamba-expression. + That means, that we have to adjust the deBruijn-indices of the variables in 'cring' + -} + adjustedCring ← returnTC (adjustDeBruijnIndex (length varInfos) cring) + just (lhs , rhs) ← returnTC (toAlgebraExpression adjustedCring (getArgs equation)) + where + nothing + → typeError( + strErr "Error while trying to build ASTs for the equation " ∷ + termErr equation ∷ []) + let solution = solverCallWithLambdas (length varInfos) varInfos adjustedCring lhs rhs + unify hole solution + + solveInPlace-macro : Term → Term → Term → TC Unit + solveInPlace-macro cring varsToSolve hole = + do + equation ← inferType hole >>= normalise + just varIndices ← returnTC (extractVarIndices (toListOfTerms varsToSolve)) + where + nothing + → typeError( + strErr "Error reading variables to solve " ∷ + termErr varsToSolve ∷ []) + just (lhs , rhs) ← returnTC (toAlgebraExpression cring (getArgs equation)) + where + nothing + → typeError( + strErr "Error while trying to build ASTs for the equation " ∷ + termErr equation ∷ []) + let solution = solverCallByVarIndices (length varIndices) varIndices cring lhs rhs + unify hole solution + + solveEqReasoning-macro : Term → Term → Term → Term → Term → TC Unit + solveEqReasoning-macro lhs cring varsToSolve reasoningToTheRight hole = + do + just varIndices ← returnTC (extractVarIndices (toListOfTerms varsToSolve)) + where + nothing + → typeError( + strErr "Error reading variables to solve " ∷ + termErr varsToSolve ∷ []) + just rhs ← returnTC (getRhs reasoningToTheRight) + where + nothing + → typeError( + strErr "Failed to extract right hand side of equation to solve from " ∷ + termErr reasoningToTheRight ∷ []) + just (lhsAST , rhsAST) ← returnTC (toAlgebraExpression cring (just (lhs , rhs))) + where + nothing + → typeError( + strErr "Error while trying to build ASTs from " ∷ + termErr lhs ∷ strErr " and " ∷ termErr rhs ∷ []) + let solverCall = solverCallByVarIndices (length varIndices) varIndices cring lhsAST rhsAST + unify hole (def (quote _≡⟨_⟩_) (varg lhs ∷ varg solverCall ∷ varg reasoningToTheRight ∷ [])) + +macro + solve : Term → Term → TC _ + solve = solve-macro + + solveInPlace : Term → Term → Term → TC _ + solveInPlace = solveInPlace-macro + + infixr 2 _≡⟨solveIn_withVars_⟩_ + _≡⟨solveIn_withVars_⟩_ : Term → Term → Term → Term → Term → TC Unit + _≡⟨solveIn_withVars_⟩_ = solveEqReasoning-macro + + +fromℤ : (R : CommRing ℓ) → ℤ → fst R +fromℤ = scalar diff --git a/Cubical/Algebra/RingSolver/Solver.agda b/Cubical/Algebra/RingSolver/Solver.agda new file mode 100644 index 0000000000..13442126c5 --- /dev/null +++ b/Cubical/Algebra/RingSolver/Solver.agda @@ -0,0 +1,145 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.Solver where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData +open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Nat.Order using (zero-≤) +open import Cubical.Data.Vec.Base +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring +open import Cubical.Algebra.RingSolver.RawAlgebra renaming (⟨_⟩ to ⟨_⟩ᵣ) +open import Cubical.Algebra.RingSolver.AlgebraExpression +open import Cubical.Algebra.RingSolver.IntAsRawRing +open import Cubical.Algebra.RingSolver.HornerForms +open import Cubical.Algebra.RingSolver.EvalHom +open import Cubical.Algebra.RingSolver.HornerEval + +private + variable + ℓ : Level + +module EqualityToNormalform (R : CommRing ℓ) where + νR = CommRing→RawℤAlgebra R + open CommRingStr (snd R) + open RingTheory (CommRing→Ring R) + open Eval ℤAsRawRing νR + open IteratedHornerOperations νR + open HomomorphismProperties R + + ℤExpr : (n : ℕ) → Type _ + ℤExpr = Expr ℤAsRawRing (fst R) + + normalize : {n : ℕ} → ℤExpr n → IteratedHornerForms νR n + normalize {n = n} (K r) = Constant n νR r + normalize {n = n} (∣ k) = Variable n νR k + normalize (x +' y) = + (normalize x) +ₕ (normalize y) + normalize (x ·' y) = + (normalize x) ·ₕ (normalize y) + normalize (-' x) = -ₕ (normalize x) + + isEqualToNormalform : + {n : ℕ} (e : ℤExpr n) (xs : Vec (fst R) n) + → eval (normalize e) xs ≡ ⟦ e ⟧ xs + isEqualToNormalform (K r) [] = refl + isEqualToNormalform {n = ℕ.suc n} (K r) (x ∷ xs) = + eval (Constant (ℕ.suc n) νR r) (x ∷ xs) ≡⟨ refl ⟩ + eval (0ₕ ·X+ Constant n νR r) (x ∷ xs) ≡⟨ combineCasesEval R 0ₕ (Constant n νR r) x xs ⟩ + eval 0ₕ (x ∷ xs) · x + eval (Constant n νR r) xs ≡⟨ cong (λ u → u · x + eval (Constant n νR r) xs) + (Eval0H (x ∷ xs)) ⟩ + 0r · x + eval (Constant n νR r) xs ≡⟨ cong + (λ u → u + eval (Constant n νR r) xs) + (0LeftAnnihilates _) ⟩ + 0r + eval (Constant n νR r) xs ≡⟨ +Lid _ ⟩ + eval (Constant n νR r) xs ≡⟨ isEqualToNormalform (K r) xs ⟩ + _ ∎ + + isEqualToNormalform (∣ zero) (x ∷ xs) = + eval (1ₕ ·X+ 0ₕ) (x ∷ xs) ≡⟨ combineCasesEval R 1ₕ 0ₕ x xs ⟩ + eval 1ₕ (x ∷ xs) · x + eval 0ₕ xs ≡⟨ cong (λ u → u · x + eval 0ₕ xs) + (Eval1ₕ (x ∷ xs)) ⟩ + 1r · x + eval 0ₕ xs ≡⟨ cong (λ u → 1r · x + u ) (Eval0H xs) ⟩ + 1r · x + 0r ≡⟨ +Rid _ ⟩ + 1r · x ≡⟨ ·Lid _ ⟩ + x ∎ + isEqualToNormalform {n = ℕ.suc n} (∣ (suc k)) (x ∷ xs) = + eval (0ₕ ·X+ Variable n νR k) (x ∷ xs) ≡⟨ combineCasesEval R 0ₕ (Variable n νR k) x xs ⟩ + eval 0ₕ (x ∷ xs) · x + eval (Variable n νR k) xs ≡⟨ cong (λ u → u · x + eval (Variable n νR k) xs) + (Eval0H (x ∷ xs)) ⟩ + 0r · x + eval (Variable n νR k) xs ≡⟨ cong (λ u → u + eval (Variable n νR k) xs) + (0LeftAnnihilates _) ⟩ + 0r + eval (Variable n νR k) xs ≡⟨ +Lid _ ⟩ + eval (Variable n νR k) xs ≡⟨ isEqualToNormalform (∣ k) xs ⟩ + ⟦ ∣ (suc k) ⟧ (x ∷ xs) ∎ + + isEqualToNormalform (-' e) [] = + eval (-ₕ (normalize e)) [] ≡⟨ -EvalDist (normalize e) [] ⟩ + - eval (normalize e) [] ≡⟨ cong -_ (isEqualToNormalform e [] ) ⟩ + - ⟦ e ⟧ [] ∎ + isEqualToNormalform (-' e) (x ∷ xs) = + eval (-ₕ (normalize e)) (x ∷ xs) ≡⟨ -EvalDist (normalize e) (x ∷ xs) ⟩ + - eval (normalize e) (x ∷ xs) ≡⟨ cong -_ (isEqualToNormalform e (x ∷ xs) ) ⟩ + - ⟦ e ⟧ (x ∷ xs) ∎ + + isEqualToNormalform (e +' e₁) [] = + eval (normalize e +ₕ normalize e₁) [] + ≡⟨ +Homeval (normalize e) _ [] ⟩ + eval (normalize e) [] + + eval (normalize e₁) [] + ≡⟨ cong (λ u → u + eval (normalize e₁) []) + (isEqualToNormalform e []) ⟩ + ⟦ e ⟧ [] + + eval (normalize e₁) [] + ≡⟨ cong (λ u → ⟦ e ⟧ [] + u) (isEqualToNormalform e₁ []) ⟩ + ⟦ e ⟧ [] + ⟦ e₁ ⟧ [] ∎ + isEqualToNormalform (e +' e₁) (x ∷ xs) = + eval (normalize e +ₕ normalize e₁) (x ∷ xs) + ≡⟨ +Homeval (normalize e) _ (x ∷ xs) ⟩ + eval (normalize e) (x ∷ xs) + eval (normalize e₁) (x ∷ xs) + ≡⟨ cong (λ u → u + eval (normalize e₁) (x ∷ xs)) + (isEqualToNormalform e (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) + eval (normalize e₁) (x ∷ xs) + ≡⟨ cong (λ u → ⟦ e ⟧ (x ∷ xs) + u) (isEqualToNormalform e₁ (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) + ⟦ e₁ ⟧ (x ∷ xs) ∎ + + isEqualToNormalform (e ·' e₁) [] = + eval (normalize e ·ₕ normalize e₁) [] + ≡⟨ ·Homeval (normalize e) _ [] ⟩ + eval (normalize e) [] · eval (normalize e₁) [] + ≡⟨ cong (λ u → u · eval (normalize e₁) []) + (isEqualToNormalform e []) ⟩ + ⟦ e ⟧ [] · eval (normalize e₁) [] + ≡⟨ cong (λ u → ⟦ e ⟧ [] · u) (isEqualToNormalform e₁ []) ⟩ + ⟦ e ⟧ [] · ⟦ e₁ ⟧ [] ∎ + + isEqualToNormalform (e ·' e₁) (x ∷ xs) = + eval (normalize e ·ₕ normalize e₁) (x ∷ xs) + ≡⟨ ·Homeval (normalize e) _ (x ∷ xs) ⟩ + eval (normalize e) (x ∷ xs) · eval (normalize e₁) (x ∷ xs) + ≡⟨ cong (λ u → u · eval (normalize e₁) (x ∷ xs)) + (isEqualToNormalform e (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) · eval (normalize e₁) (x ∷ xs) + ≡⟨ cong (λ u → ⟦ e ⟧ (x ∷ xs) · u) (isEqualToNormalform e₁ (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) · ⟦ e₁ ⟧ (x ∷ xs) ∎ + + solve : + {n : ℕ} (e₁ e₂ : ℤExpr n) (xs : Vec (fst R) n) + (p : eval (normalize e₁) xs ≡ eval (normalize e₂) xs) + → ⟦ e₁ ⟧ xs ≡ ⟦ e₂ ⟧ xs + solve e₁ e₂ xs p = + ⟦ e₁ ⟧ xs ≡⟨ sym (isEqualToNormalform e₁ xs) ⟩ + eval (normalize e₁) xs ≡⟨ p ⟩ + eval (normalize e₂) xs ≡⟨ isEqualToNormalform e₂ xs ⟩ + ⟦ e₂ ⟧ xs ∎ + +ℤExpr : (R : CommRing ℓ) (n : ℕ) + → _ +ℤExpr R n = EqualityToNormalform.ℤExpr R n + +solve : (R : CommRing ℓ) + {n : ℕ} (e₁ e₂ : ℤExpr R n) (xs : Vec (fst R) n) + (p : eval (EqualityToNormalform.normalize R e₁) xs ≡ eval (EqualityToNormalform.normalize R e₂) xs) + → _ +solve R = EqualityToNormalform.solve R diff --git a/Cubical/Algebra/RingSolver/Utility.agda b/Cubical/Algebra/RingSolver/Utility.agda new file mode 100644 index 0000000000..3042d0d318 --- /dev/null +++ b/Cubical/Algebra/RingSolver/Utility.agda @@ -0,0 +1,37 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.Utility where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Bool +open import Cubical.Data.Empty using (⊥) renaming (rec to recEmpty) +open import Cubical.Data.Sigma + +open import Cubical.Relation.Nullary.Base + +private + variable + ℓ ℓ′ : Level +byBoolAbsurdity : {Anything : Type ℓ} → false ≡ true → Anything +byBoolAbsurdity p = recEmpty (false≢true p) + +byAbsurdity : {Anything : Type ℓ} → ⊥ → Anything +byAbsurdity x = recEmpty x + +extractFromAnd : (P Q : Bool) + → P and Q ≡ true + → (P ≡ true) × (Q ≡ true) +extractFromAnd false false eq = byBoolAbsurdity eq +extractFromAnd false true eq = byBoolAbsurdity eq +extractFromAnd true false eq = byBoolAbsurdity eq +extractFromAnd true true eq = eq , eq + +extractFromAndLeft : {P Q : Bool} + → P and Q ≡ true + → P ≡ true +extractFromAndLeft eq = fst (extractFromAnd _ _ eq) + +extractFromAndRight : {P Q : Bool} + → P and Q ≡ true + → Q ≡ true +extractFromAndRight eq = snd (extractFromAnd _ _ eq) diff --git a/Cubical/Algebra/Semigroup.agda b/Cubical/Algebra/Semigroup.agda new file mode 100644 index 0000000000..53e0c0385c --- /dev/null +++ b/Cubical/Algebra/Semigroup.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Semigroup where + +open import Cubical.Algebra.Semigroup.Base public diff --git a/Cubical/Algebra/Semigroup/Base.agda b/Cubical/Algebra/Semigroup/Base.agda new file mode 100644 index 0000000000..4388d929f5 --- /dev/null +++ b/Cubical/Algebra/Semigroup/Base.agda @@ -0,0 +1,100 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Semigroup.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.RecordEquiv +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open Iso + + +private + variable + ℓ : Level + +-- Semigroups as a record, inspired by the Agda standard library: +-- +-- https://github.com/agda/agda-stdlib/blob/master/src/Algebra/Bundles.agda#L48 +-- https://github.com/agda/agda-stdlib/blob/master/src/Algebra/Structures.agda#L50 +-- +-- Note that as we are using Path for all equations the IsMagma record +-- would only contain isSet A if we had it. +record IsSemigroup {A : Type ℓ} (_·_ : A → A → A) : Type ℓ where + no-eta-equality + constructor issemigroup + + field + is-set : isSet A + assoc : (x y z : A) → x · (y · z) ≡ (x · y) · z + +unquoteDecl IsSemigroupIsoΣ = declareRecordIsoΣ IsSemigroupIsoΣ (quote IsSemigroup) + +record SemigroupStr (A : Type ℓ) : Type ℓ where + + constructor semigroupstr + + field + _·_ : A → A → A + isSemigroup : IsSemigroup _·_ + + infixl 7 _·_ + + open IsSemigroup isSemigroup public + +Semigroup : ∀ ℓ → Type (ℓ-suc ℓ) +Semigroup ℓ = TypeWithStr ℓ SemigroupStr + +semigroup : (A : Type ℓ) (_·_ : A → A → A) (h : IsSemigroup _·_) → Semigroup ℓ +semigroup A _·_ h = A , semigroupstr _·_ h + +record IsSemigroupEquiv {A : Type ℓ} {B : Type ℓ} + (M : SemigroupStr A) (e : A ≃ B) (N : SemigroupStr B) + : Type ℓ + where + + -- Shorter qualified names + private + module M = SemigroupStr M + module N = SemigroupStr N + + field + isHom : (x y : A) → equivFun e (x M.· y) ≡ equivFun e x N.· equivFun e y + +open SemigroupStr +open IsSemigroup +open IsSemigroupEquiv + +SemigroupEquiv : (M N : Semigroup ℓ) → Type ℓ +SemigroupEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsSemigroupEquiv (M .snd) e (N .snd) + +isPropIsSemigroup : {A : Type ℓ} (_·_ : A → A → A) → isProp (IsSemigroup _·_) +isPropIsSemigroup _·_ = + isOfHLevelRetractFromIso 1 IsSemigroupIsoΣ + (isPropΣ + isPropIsSet + (λ isSetA → isPropΠ3 λ _ _ _ → isSetA _ _)) + +𝒮ᴰ-Semigroup : DUARel (𝒮-Univ ℓ) SemigroupStr ℓ +𝒮ᴰ-Semigroup = + 𝒮ᴰ-Record (𝒮-Univ _) IsSemigroupEquiv + (fields: + data[ _·_ ∣ autoDUARel _ _ ∣ isHom ] + prop[ isSemigroup ∣ (λ _ _ → isPropIsSemigroup _) ]) + +SemigroupPath : (M N : Semigroup ℓ) → SemigroupEquiv M N ≃ (M ≡ N) +SemigroupPath = ∫ 𝒮ᴰ-Semigroup .UARel.ua diff --git a/Cubical/Algebra/Semilattice.agda b/Cubical/Algebra/Semilattice.agda new file mode 100644 index 0000000000..240f180946 --- /dev/null +++ b/Cubical/Algebra/Semilattice.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Semilattice where + +open import Cubical.Algebra.Semilattice.Base public diff --git a/Cubical/Algebra/Semilattice/Base.agda b/Cubical/Algebra/Semilattice/Base.agda new file mode 100644 index 0000000000..381d5a96e0 --- /dev/null +++ b/Cubical/Algebra/Semilattice/Base.agda @@ -0,0 +1,233 @@ +{- + following Johnstone's book "Stone Spaces" we define semilattices + to be commutative monoids such that every element is idempotent. + In particular, we take every semilattice to have a neutral element + that is either the maximal or minimal element depending on whether + we have a join or meet semilattice. +-} + +{-# OPTIONS --safe #-} +module Cubical.Algebra.Semilattice.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommMonoid + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Relation.Binary +open import Cubical.Relation.Binary.Poset + +open import Cubical.Reflection.RecordEquiv + +open Iso + +private + variable + ℓ ℓ' : Level + +record IsSemilattice {A : Type ℓ} (ε : A) (_·_ : A → A → A) : Type ℓ where + constructor issemilattice + + field + isCommMonoid : IsCommMonoid ε _·_ + idem : (x : A) → x · x ≡ x + + open IsCommMonoid isCommMonoid public + +unquoteDecl IsSemilatticeIsoΣ = declareRecordIsoΣ IsSemilatticeIsoΣ (quote IsSemilattice) + +record SemilatticeStr (A : Type ℓ) : Type ℓ where + constructor semilatticestr + + field + ε : A + _·_ : A → A → A + isSemilattice : IsSemilattice ε _·_ + + infixl 7 _·_ + + open IsSemilattice isSemilattice public + +Semilattice : ∀ ℓ → Type (ℓ-suc ℓ) +Semilattice ℓ = TypeWithStr ℓ SemilatticeStr + +semilattice : (A : Type ℓ) (ε : A) (_·_ : A → A → A) (h : IsSemilattice ε _·_) → Semilattice ℓ +semilattice A ε _·_ h = A , semilatticestr ε _·_ h + +-- Easier to use constructors +makeIsSemilattice : {L : Type ℓ} {ε : L} {_·_ : L → L → L} + (is-setL : isSet L) + (assoc : (x y z : L) → x · (y · z) ≡ (x · y) · z) + (rid : (x : L) → x · ε ≡ x) + (lid : (x : L) → ε · x ≡ x) + (comm : (x y : L) → x · y ≡ y · x) + (idem : (x : L) → x · x ≡ x) + → IsSemilattice ε _·_ +IsSemilattice.isCommMonoid (makeIsSemilattice is-setL assoc rid lid comm idem) = + makeIsCommMonoid is-setL assoc rid lid comm +IsSemilattice.idem (makeIsSemilattice is-setL assoc rid lid comm idem) = idem + +makeSemilattice : {L : Type ℓ} (ε : L) (_·_ : L → L → L) + (is-setL : isSet L) + (assoc : (x y z : L) → x · (y · z) ≡ (x · y) · z) + (rid : (x : L) → x · ε ≡ x) + (lid : (x : L) → ε · x ≡ x) + (comm : (x y : L) → x · y ≡ y · x) + (idem : (x : L) → x · x ≡ x) + → Semilattice ℓ +makeSemilattice ε _·_ is-setL assoc rid lid comm idem = + semilattice _ ε _·_ (makeIsSemilattice is-setL assoc rid lid comm idem) + + +SemilatticeStr→MonoidStr : {A : Type ℓ} → SemilatticeStr A → MonoidStr A +SemilatticeStr→MonoidStr (semilatticestr _ _ H) = + monoidstr _ _ (H .IsSemilattice.isCommMonoid .IsCommMonoid.isMonoid) + +Semilattice→Monoid : Semilattice ℓ → Monoid ℓ +Semilattice→Monoid (_ , semilatticestr _ _ H) = + _ , monoidstr _ _ (H .IsSemilattice.isCommMonoid .IsCommMonoid.isMonoid) + +Semilattice→CommMonoid : Semilattice ℓ → CommMonoid ℓ +Semilattice→CommMonoid (_ , semilatticestr _ _ H) = + _ , commmonoidstr _ _ (H .IsSemilattice.isCommMonoid) + +SemilatticeHom : (L : Semilattice ℓ) (M : Semilattice ℓ') → Type (ℓ-max ℓ ℓ') +SemilatticeHom L M = MonoidHom (Semilattice→Monoid L) (Semilattice→Monoid M) + +IsSemilatticeEquiv : {A : Type ℓ} {B : Type ℓ'} + (M : SemilatticeStr A) (e : A ≃ B) (N : SemilatticeStr B) → Type (ℓ-max ℓ ℓ') +IsSemilatticeEquiv M e N = + IsMonoidHom (SemilatticeStr→MonoidStr M) (e .fst) (SemilatticeStr→MonoidStr N) + +SemilatticeEquiv : (M : Semilattice ℓ) (N : Semilattice ℓ') → Type (ℓ-max ℓ ℓ') +SemilatticeEquiv M N = Σ[ e ∈ (M .fst ≃ N .fst) ] IsSemilatticeEquiv (M .snd) e (N .snd) + +isPropIsSemilattice : {L : Type ℓ} (ε : L) (_·_ : L → L → L) + → isProp (IsSemilattice ε _·_) +isPropIsSemilattice ε _·_ (issemilattice LL LC) (issemilattice SL SC) = + λ i → issemilattice (isPropIsCommMonoid _ _ LL SL i) (isPropIdem LC SC i) + where + isSetL : isSet _ + isSetL = LL .IsCommMonoid.isMonoid .IsMonoid.isSemigroup .IsSemigroup.is-set + + isPropIdem : isProp ((x : _) → x · x ≡ x) + isPropIdem = isPropΠ λ _ → isSetL _ _ + +𝒮ᴰ-Semilattice : DUARel (𝒮-Univ ℓ) SemilatticeStr ℓ +𝒮ᴰ-Semilattice = + 𝒮ᴰ-Record (𝒮-Univ _) IsSemilatticeEquiv + (fields: + data[ ε ∣ autoDUARel _ _ ∣ presε ] + data[ _·_ ∣ autoDUARel _ _ ∣ isHom ] + prop[ isSemilattice ∣ (λ _ _ → isPropIsSemilattice _ _) ]) + where + open SemilatticeStr + open IsMonoidHom + +SemilatticePath : (L K : Semilattice ℓ) → SemilatticeEquiv L K ≃ (L ≡ K) +SemilatticePath = ∫ 𝒮ᴰ-Semilattice .UARel.ua + + +-- TODO: decide if that's the right approach +module JoinSemilattice (L' : Semilattice ℓ) where + private L = fst L' + open SemilatticeStr (snd L') renaming (_·_ to _∨l_ ; ε to 1l) + open CommMonoidTheory (Semilattice→CommMonoid L') + + + _≤_ : L → L → Type ℓ + x ≤ y = x ∨l y ≡ y + + infix 4 _≤_ + + IndPoset : Poset ℓ ℓ + fst IndPoset = L + PosetStr._≤_ (snd IndPoset) = _≤_ + IsPoset.is-set (PosetStr.isPoset (snd IndPoset)) = is-set + IsPoset.is-prop-valued (PosetStr.isPoset (snd IndPoset)) = λ _ _ → is-set _ _ + IsPoset.is-refl (PosetStr.isPoset (snd IndPoset)) = idem + IsPoset.is-trans (PosetStr.isPoset (snd IndPoset)) = path + where + path : (a b c : L) → a ∨l b ≡ b → b ∨l c ≡ c → a ∨l c ≡ c + path a b c a∨b≡b b∨c≡c = a ∨l c ≡⟨ cong (a ∨l_) (sym b∨c≡c) ⟩ + a ∨l (b ∨l c) ≡⟨ assoc _ _ _ ⟩ + (a ∨l b) ∨l c ≡⟨ cong (_∨l c) a∨b≡b ⟩ + b ∨l c ≡⟨ b∨c≡c ⟩ + c ∎ + IsPoset.is-antisym (PosetStr.isPoset (snd IndPoset)) = + λ _ _ a∨b≡b b∨a≡a → sym b∨a≡a ∙∙ comm _ _ ∙∙ a∨b≡b + + ∨lIsMax : ∀ x y z → x ≤ z → y ≤ z → x ∨l y ≤ z + ∨lIsMax x y z x≤z y≤z = cong ((x ∨l y) ∨l_) (sym (idem z)) ∙ commAssocSwap x y z z + ∙ cong₂ (_∨l_) x≤z y≤z + ∙ idem z + + ∨≤LCancel : ∀ x y → y ≤ x ∨l y + ∨≤LCancel x y = commAssocl y x y ∙ cong (x ∨l_) (idem y) + + ∨≤RCancel : ∀ x y → x ≤ x ∨l y + ∨≤RCancel x y = assoc _ _ _ ∙ cong (_∨l y) (idem x) + + ≤-∨Pres : ∀ x y u w → x ≤ y → u ≤ w → x ∨l u ≤ y ∨l w + ≤-∨Pres x y u w x≤y u≤w = commAssocSwap x u y w ∙ cong₂ (_∨l_) x≤y u≤w + + ≤-∨LPres : ∀ x y z → x ≤ y → z ∨l x ≤ z ∨l y + ≤-∨LPres x y z x≤y = ≤-∨Pres _ _ _ _ (idem z) x≤y + + +module MeetSemilattice (L' : Semilattice ℓ) where + private L = fst L' + open SemilatticeStr (snd L') renaming (_·_ to _∧l_ ; ε to 0l) + open CommMonoidTheory (Semilattice→CommMonoid L') + + _≤_ : L → L → Type ℓ + x ≤ y = x ∧l y ≡ x + + infix 4 _≤_ + + IndPoset : Poset ℓ ℓ + fst IndPoset = L + PosetStr._≤_ (snd IndPoset) = _≤_ + IsPoset.is-set (PosetStr.isPoset (snd IndPoset)) = is-set + IsPoset.is-prop-valued (PosetStr.isPoset (snd IndPoset)) = λ _ _ → is-set _ _ + IsPoset.is-refl (PosetStr.isPoset (snd IndPoset)) = idem + IsPoset.is-trans (PosetStr.isPoset (snd IndPoset)) = path + where + path : (a b c : L) → a ∧l b ≡ a → b ∧l c ≡ b → a ∧l c ≡ a + path a b c a∧b≡a b∧c≡b = a ∧l c ≡⟨ cong (_∧l c) (sym a∧b≡a) ⟩ + (a ∧l b) ∧l c ≡⟨ sym (assoc _ _ _) ⟩ + a ∧l (b ∧l c) ≡⟨ cong (a ∧l_) b∧c≡b ⟩ + a ∧l b ≡⟨ a∧b≡a ⟩ + a ∎ + IsPoset.is-antisym (PosetStr.isPoset (snd IndPoset)) = + λ _ _ a∧b≡a b∧a≡b → sym a∧b≡a ∙∙ comm _ _ ∙∙ b∧a≡b + + ≤-∧LPres : ∀ x y z → x ≤ y → z ∧l x ≤ z ∧l y + ≤-∧LPres x y z x≤y = commAssocSwap z x z y ∙∙ cong (_∧l (x ∧l y)) (idem z) ∙∙ cong (z ∧l_) x≤y + + ∧≤LCancel : ∀ x y → x ∧l y ≤ y + ∧≤LCancel x y = sym (assoc _ _ _) ∙ cong (x ∧l_) (idem y) + + ∧≤RCancel : ∀ x y → x ∧l y ≤ x + ∧≤RCancel x y = commAssocr x y x ∙ cong (_∧l y) (idem x) + + ∧lIsMin : ∀ x y z → z ≤ x → z ≤ y → z ≤ x ∧l y + ∧lIsMin x y z z≤x z≤y = cong (_∧l (x ∧l y)) (sym (idem z)) ∙ commAssocSwap z z x y + ∙ cong₂ (_∧l_) z≤x z≤y + ∙ idem z diff --git a/Cubical/Algebra/SymmetricGroup.agda b/Cubical/Algebra/SymmetricGroup.agda new file mode 100644 index 0000000000..aa4395f7d4 --- /dev/null +++ b/Cubical/Algebra/SymmetricGroup.agda @@ -0,0 +1,26 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.SymmetricGroup where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ ; suc ; zero) +open import Cubical.Data.Fin using (Fin ; isSetFin) +open import Cubical.Data.Empty +open import Cubical.Relation.Nullary using (¬_) + +open import Cubical.Algebra.Group + +private + variable + ℓ : Level + +Symmetric-Group : (X : Type ℓ) → isSet X → Group ℓ +Symmetric-Group X isSetX = makeGroup (idEquiv X) compEquiv invEquiv (isOfHLevel≃ 2 isSetX isSetX) + compEquiv-assoc compEquivEquivId compEquivIdEquiv invEquiv-is-rinv invEquiv-is-linv + +-- Finite symmetrics groups + +Sym : ℕ → Group _ +Sym n = Symmetric-Group (Fin n) isSetFin diff --git a/Cubical/Algebra/ZariskiLattice/Base.agda b/Cubical/Algebra/ZariskiLattice/Base.agda new file mode 100644 index 0000000000..341fca72bd --- /dev/null +++ b/Cubical/Algebra/ZariskiLattice/Base.agda @@ -0,0 +1,585 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.ZariskiLattice.Base where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Powerset using (ℙ ; ⊆-refl-consequence) + renaming (_∈_ to _∈ₚ_ ; subst-∈ to subst-∈ₚ) + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm + ; ·-identityʳ to ·ℕ-rid) +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary +open import Cubical.Relation.Binary.Poset + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Ring.Properties +open import Cubical.Algebra.Ring.BigOps +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.BinomialThm +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.RadicalIdeal +open import Cubical.Algebra.CommRing.Localisation.Base +open import Cubical.Algebra.CommRing.Localisation.UniversalProperty +open import Cubical.Algebra.CommRing.Localisation.InvertingElements +open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Properties +open import Cubical.Algebra.CommAlgebra.Localisation +open import Cubical.Algebra.RingSolver.Reflection +open import Cubical.Algebra.Semilattice +open import Cubical.Algebra.Lattice +open import Cubical.Algebra.DistLattice +open import Cubical.Algebra.DistLattice.Basis +open import Cubical.Algebra.DistLattice.BigOps +open import Cubical.Algebra.Matrix + +open import Cubical.Categories.Category.Base hiding (_[_,_]) +open import Cubical.Categories.Functor +open import Cubical.Categories.Instances.CommAlgebras +open import Cubical.Categories.Instances.DistLattice +open import Cubical.Categories.Instances.Semilattice + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open Iso +open BinaryRelation +open isEquivRel + +private + variable + ℓ ℓ' : Level + + +module ZarLat (R' : CommRing ℓ) where + open CommRingStr (snd R') + open RingTheory (CommRing→Ring R') + open Sum (CommRing→Ring R') + open CommRingTheory R' + open Exponentiation R' + open BinomialThm R' + open CommIdeal R' + open RadicalIdeal R' + open isCommIdeal + open ProdFin R' + + private + R = fst R' + A = Σ[ n ∈ ℕ ] (FinVec R n) + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + + _∼_ : A → A → Type (ℓ-suc ℓ) + (_ , α) ∼ (_ , β) = √ ⟨ α ⟩ ≡ √ ⟨ β ⟩ + + ∼EquivRel : isEquivRel (_∼_) + reflexive ∼EquivRel _ = refl + symmetric ∼EquivRel _ _ = sym + transitive ∼EquivRel _ _ _ = _∙_ + + ZL : Type (ℓ-suc ℓ) + ZL = A / _∼_ + + 0z : ZL + 0z = [ 0 , (λ ()) ] + + 1z : ZL + 1z = [ 1 , (replicateFinVec 1 1r) ] + + _∨z_ : ZL → ZL → ZL + _∨z_ = setQuotSymmBinOp (reflexive ∼EquivRel) (transitive ∼EquivRel) + (λ (_ , α) (_ , β) → (_ , α ++Fin β)) + (λ (_ , α) (_ , β) → cong √ (FGIdealAddLemma _ α β ∙∙ +iComm _ _ ∙∙ sym (FGIdealAddLemma _ β α))) + λ (_ , α) (_ , β) (_ , γ) α∼β → --need to show α∨γ ∼ β∨γ + √ ⟨ α ++Fin γ ⟩ ≡⟨ cong √ (FGIdealAddLemma _ α γ) ⟩ + √ (⟨ α ⟩ +i ⟨ γ ⟩) ≡⟨ sym (√+LContr _ _) ⟩ + √ (√ ⟨ α ⟩ +i ⟨ γ ⟩) ≡⟨ cong (λ I → √ (I +i ⟨ γ ⟩)) α∼β ⟩ + √ (√ ⟨ β ⟩ +i ⟨ γ ⟩) ≡⟨ √+LContr _ _ ⟩ + √ (⟨ β ⟩ +i ⟨ γ ⟩) ≡⟨ cong √ (sym (FGIdealAddLemma _ β γ)) ⟩ + √ ⟨ β ++Fin γ ⟩ ∎ + + _∧z_ : ZL → ZL → ZL + _∧z_ = setQuotSymmBinOp (reflexive ∼EquivRel) (transitive ∼EquivRel) + (λ (_ , α) (_ , β) → (_ , α ··Fin β)) + (λ (_ , α) (_ , β) → cong √ (FGIdealMultLemma _ α β ∙∙ ·iComm _ _ ∙∙ sym (FGIdealMultLemma _ β α))) + λ (_ , α) (_ , β) (_ , γ) α∼β → --need to show α∧γ ∼ β∧γ + √ ⟨ α ··Fin γ ⟩ ≡⟨ cong √ (FGIdealMultLemma _ α γ) ⟩ + √ (⟨ α ⟩ ·i ⟨ γ ⟩) ≡⟨ sym (√·LContr _ _) ⟩ + √ (√ ⟨ α ⟩ ·i ⟨ γ ⟩) ≡⟨ cong (λ I → √ (I ·i ⟨ γ ⟩)) α∼β ⟩ + √ (√ ⟨ β ⟩ ·i ⟨ γ ⟩) ≡⟨ √·LContr _ _ ⟩ + √ (⟨ β ⟩ ·i ⟨ γ ⟩) ≡⟨ cong √ (sym (FGIdealMultLemma _ β γ)) ⟩ + √ ⟨ β ··Fin γ ⟩ ∎ + + -- join axioms + ∨zAssoc : ∀ (𝔞 𝔟 𝔠 : ZL) → 𝔞 ∨z (𝔟 ∨z 𝔠) ≡ (𝔞 ∨z 𝔟) ∨z 𝔠 + ∨zAssoc = SQ.elimProp3 (λ _ _ _ → squash/ _ _) + λ (_ , α) (_ , β) (_ , γ) → eq/ _ _ (cong √ (IdealAddAssoc _ _ _ _)) + + ∨zComm : ∀ (𝔞 𝔟 : ZL) → 𝔞 ∨z 𝔟 ≡ 𝔟 ∨z 𝔞 + ∨zComm = SQ.elimProp2 (λ _ _ → squash/ _ _) + λ (_ , α) (_ , β) → eq/ _ _ + (cong √ (FGIdealAddLemma _ α β ∙∙ +iComm _ _ ∙∙ sym (FGIdealAddLemma _ β α))) + + ∨zLid : ∀ (𝔞 : ZL) → 0z ∨z 𝔞 ≡ 𝔞 + ∨zLid = SQ.elimProp (λ _ → squash/ _ _) λ _ → eq/ _ _ refl + + ∨zRid : ∀ (𝔞 : ZL) → 𝔞 ∨z 0z ≡ 𝔞 + ∨zRid _ = ∨zComm _ _ ∙ ∨zLid _ + + + -- -- meet axioms + ∧zAssoc : ∀ (𝔞 𝔟 𝔠 : ZL) → 𝔞 ∧z (𝔟 ∧z 𝔠) ≡ (𝔞 ∧z 𝔟) ∧z 𝔠 + ∧zAssoc = SQ.elimProp3 (λ _ _ _ → squash/ _ _) + λ (_ , α) (_ , β) (_ , γ) → eq/ _ _ + (√ ⟨ α ··Fin (β ··Fin γ) ⟩ ≡⟨ cong √ (FGIdealMultLemma _ _ _) ⟩ + √ (⟨ α ⟩ ·i ⟨ β ··Fin γ ⟩) ≡⟨ cong (λ x → √ (⟨ α ⟩ ·i x)) (FGIdealMultLemma _ _ _) ⟩ + √ (⟨ α ⟩ ·i (⟨ β ⟩ ·i ⟨ γ ⟩)) ≡⟨ cong √ (·iAssoc _ _ _) ⟩ + √ ((⟨ α ⟩ ·i ⟨ β ⟩) ·i ⟨ γ ⟩) ≡⟨ cong (λ x → √ (x ·i ⟨ γ ⟩)) (sym (FGIdealMultLemma _ _ _)) ⟩ + √ (⟨ α ··Fin β ⟩ ·i ⟨ γ ⟩) ≡⟨ cong √ (sym (FGIdealMultLemma _ _ _)) ⟩ + √ ⟨ (α ··Fin β) ··Fin γ ⟩ ∎) + + ∧zComm : ∀ (𝔞 𝔟 : ZL) → 𝔞 ∧z 𝔟 ≡ 𝔟 ∧z 𝔞 + ∧zComm = SQ.elimProp2 (λ _ _ → squash/ _ _) + λ (_ , α) (_ , β) → eq/ _ _ + (cong √ (FGIdealMultLemma _ α β ∙∙ ·iComm _ _ ∙∙ sym (FGIdealMultLemma _ β α))) + + ∧zRid : ∀ (𝔞 : ZL) → 𝔞 ∧z 1z ≡ 𝔞 + ∧zRid = SQ.elimProp (λ _ → squash/ _ _) + λ (_ , α) → eq/ _ _ (cong √ + (⟨ α ··Fin (replicateFinVec 1 1r) ⟩ ≡⟨ FGIdealMultLemma _ _ _ ⟩ + ⟨ α ⟩ ·i ⟨ (replicateFinVec 1 1r) ⟩ ≡⟨ cong (⟨ α ⟩ ·i_) (contains1Is1 _ (indInIdeal _ _ zero)) ⟩ + ⟨ α ⟩ ·i 1Ideal ≡⟨ ·iRid _ ⟩ + ⟨ α ⟩ ∎)) + + + -- absorption and distributivity + ∧zAbsorb∨z : ∀ (𝔞 𝔟 : ZL) → 𝔞 ∧z (𝔞 ∨z 𝔟) ≡ 𝔞 + ∧zAbsorb∨z = SQ.elimProp2 (λ _ _ → squash/ _ _) + λ (_ , α) (_ , β) → eq/ _ _ + (√ ⟨ α ··Fin (α ++Fin β) ⟩ ≡⟨ cong √ (FGIdealMultLemma _ α (α ++Fin β)) ⟩ + √ (⟨ α ⟩ ·i ⟨ α ++Fin β ⟩) ≡⟨ cong (λ x → √ (⟨ α ⟩ ·i x)) (FGIdealAddLemma _ α β) ⟩ + √ (⟨ α ⟩ ·i (⟨ α ⟩ +i ⟨ β ⟩)) ≡⟨ √·Absorb+ _ _ ⟩ + √ ⟨ α ⟩ ∎) + + ∧zLDist∨z : ∀ (𝔞 𝔟 𝔠 : ZL) → 𝔞 ∧z (𝔟 ∨z 𝔠) ≡ (𝔞 ∧z 𝔟) ∨z (𝔞 ∧z 𝔠) + ∧zLDist∨z = SQ.elimProp3 (λ _ _ _ → squash/ _ _) + λ (_ , α) (_ , β) (_ , γ) → eq/ _ _ + (√ ⟨ α ··Fin (β ++Fin γ) ⟩ ≡⟨ cong √ (FGIdealMultLemma _ _ _) ⟩ + √ (⟨ α ⟩ ·i ⟨ β ++Fin γ ⟩) ≡⟨ cong (λ x → √ (⟨ α ⟩ ·i x)) (FGIdealAddLemma _ _ _) ⟩ + √ (⟨ α ⟩ ·i (⟨ β ⟩ +i ⟨ γ ⟩)) ≡⟨ cong √ (·iRdist+i _ _ _) ⟩ + -- L/R-dist are swapped + -- in Lattices vs Rings + √ (⟨ α ⟩ ·i ⟨ β ⟩ +i ⟨ α ⟩ ·i ⟨ γ ⟩) ≡⟨ cong₂ (λ x y → √ (x +i y)) + (sym (FGIdealMultLemma _ _ _)) + (sym (FGIdealMultLemma _ _ _)) ⟩ + √ (⟨ α ··Fin β ⟩ +i ⟨ α ··Fin γ ⟩) ≡⟨ cong √ (sym (FGIdealAddLemma _ _ _)) ⟩ + √ ⟨ (α ··Fin β) ++Fin (α ··Fin γ) ⟩ ∎) + + + ZariskiLattice : DistLattice (ℓ-suc ℓ) + fst ZariskiLattice = ZL + DistLatticeStr.0l (snd ZariskiLattice) = 0z + DistLatticeStr.1l (snd ZariskiLattice) = 1z + DistLatticeStr._∨l_ (snd ZariskiLattice) = _∨z_ + DistLatticeStr._∧l_ (snd ZariskiLattice) = _∧z_ + DistLatticeStr.isDistLattice (snd ZariskiLattice) = + makeIsDistLattice∧lOver∨l squash/ ∨zAssoc ∨zRid ∨zComm + ∧zAssoc ∧zRid ∧zComm ∧zAbsorb∨z ∧zLDist∨z + + + +module _ (R' : CommRing ℓ) (L' : DistLattice ℓ') where + + open CommRingStr (R' .snd) + open RingTheory (CommRing→Ring R') + open Sum (CommRing→Ring R') + open CommRingTheory R' + open Exponentiation R' + + open DistLatticeStr (L' .snd) renaming (is-set to isSetL) + open Join L' + open LatticeTheory (DistLattice→Lattice L') + open Order (DistLattice→Lattice L') + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice L')) + open PosetReasoning IndPoset + open PosetStr (IndPoset .snd) hiding (_≤_) + private + R = fst R' + L = fst L' + + record IsZarMap (d : R → L) : Type (ℓ-max ℓ ℓ') where + constructor iszarmap + + field + pres0 : d 0r ≡ 0l + pres1 : d 1r ≡ 1l + ·≡∧ : ∀ x y → d (x · y) ≡ d x ∧l d y + +≤∨ : ∀ x y → d (x + y) ≤ d x ∨l d y + + ∑≤⋁ : {n : ℕ} (U : FinVec R n) → d (∑ U) ≤ ⋁ λ i → d (U i) + ∑≤⋁ {n = zero} U = ∨lRid _ ∙ pres0 + ∑≤⋁ {n = suc n} U = d (∑ U) ≤⟨ ∨lIdem _ ⟩ + d (U zero + ∑ (U ∘ suc)) ≤⟨ +≤∨ _ _ ⟩ + d (U zero) ∨l d (∑ (U ∘ suc)) ≤⟨ ≤-∨LPres _ _ _ (∑≤⋁ (U ∘ suc)) ⟩ + d (U zero) ∨l ⋁ (d ∘ U ∘ suc) ≤⟨ ∨lIdem _ ⟩ + ⋁ (d ∘ U) ◾ + + d·LCancel : ∀ x y → d (x · y) ≤ d y + d·LCancel x y = subst (λ a → a ≤ d y) (sym (·≡∧ x y)) (∧≤LCancelJoin _ _) + + linearCombination≤LCancel : {n : ℕ} (α β : FinVec R n) + → d (linearCombination R' α β) ≤ ⋁ (d ∘ β) + linearCombination≤LCancel α β = is-trans _ _ _ (∑≤⋁ (λ i → α i · β i)) + (≤-⋁Ext _ _ λ i → d·LCancel (α i) (β i)) + + ZarMapIdem : ∀ (n : ℕ) (x : R) → d (x ^ (suc n)) ≡ d x + ZarMapIdem zero x = ·≡∧ _ _ ∙∙ cong (d x ∧l_) pres1 ∙∙ ∧lRid _ + ZarMapIdem (suc n) x = ·≡∧ _ _ ∙∙ cong (d x ∧l_) (ZarMapIdem n x) ∙∙ ∧lIdem _ + + ZarMapExpIneq : ∀ (n : ℕ) (x : R) → d x ≤ d (x ^ n) + ZarMapExpIneq zero x = cong (d x ∨l_) pres1 ∙∙ 1lRightAnnihilates∨l _ ∙∙ sym pres1 + ZarMapExpIneq (suc n) x = subst (λ y → d x ≤ y) (sym (ZarMapIdem _ x)) (∨lIdem _) + + -- the crucial lemma about "Zariski maps" + open CommIdeal R' + open RadicalIdeal R' + open isCommIdeal + private + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + + ZarMapRadicalIneq : ∀ {n : ℕ} (α : FinVec R n) (x : R) + → x ∈ √ ⟨ α ⟩ → d x ≤ ⋁ (d ∘ α) + ZarMapRadicalIneq α x = PT.elim (λ _ → isSetL _ _) + (uncurry (λ n → (PT.elim (λ _ → isSetL _ _) (uncurry (curriedHelper n))))) + where + curriedHelper : (n : ℕ) (β : FinVec R _) + → x ^ n ≡ linearCombination R' β α → d x ≤ ⋁ (d ∘ α) + curriedHelper n β xⁿ≡∑βα = d x ≤⟨ ZarMapExpIneq n x ⟩ + d (x ^ n) + ≤⟨ subst (λ y → y ≤ ⋁ (d ∘ α)) (sym (cong d xⁿ≡∑βα)) (linearCombination≤LCancel β α) ⟩ + ⋁ (d ∘ α) ◾ + +module ZarLatUniversalProp (R' : CommRing ℓ) where + open CommRingStr (snd R') + open RingTheory (CommRing→Ring R') + open Sum (CommRing→Ring R') + open CommRingTheory R' + open Exponentiation R' + open BinomialThm R' + open CommIdeal R' + open RadicalIdeal R' + open isCommIdeal + open ProdFin R' + + open ZarLat R' + open IsZarMap + + private + R = fst R' + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + + + D : R → ZL + D x = [ 1 , replicateFinVec 1 x ] -- λ x → √⟨x⟩ + + isZarMapD : IsZarMap R' ZariskiLattice D + pres0 isZarMapD = eq/ _ _ (cong √ (0FGIdeal _ ∙ sym (emptyFGIdeal _ _))) + pres1 isZarMapD = refl + ·≡∧ isZarMapD x y = cong {B = λ _ → ZL} (λ U → [ 1 , U ]) (Length1··Fin x y) + +≤∨ isZarMapD x y = eq/ _ _ (cong √ (CommIdeal≡Char (inclOfFGIdeal _ 3Vec ⟨ 2Vec ⟩ 3Vec⊆2Vec) + (inclOfFGIdeal _ 2Vec ⟨ 3Vec ⟩ 2Vec⊆3Vec))) + where + 2Vec = replicateFinVec 1 x ++Fin replicateFinVec 1 y + 3Vec = replicateFinVec 1 (x + y) ++Fin (replicateFinVec 1 x ++Fin replicateFinVec 1 y) + + 3Vec⊆2Vec : ∀ (i : Fin 3) → 3Vec i ∈ ⟨ 2Vec ⟩ + 3Vec⊆2Vec zero = ⟨ 2Vec ⟩ .snd .+Closed (indInIdeal _ _ zero) (indInIdeal _ _ (suc zero)) + 3Vec⊆2Vec (suc zero) = indInIdeal _ _ zero + 3Vec⊆2Vec (suc (suc zero)) = indInIdeal _ _ (suc zero) + + 2Vec⊆3Vec : ∀ (i : Fin 2) → 2Vec i ∈ ⟨ 3Vec ⟩ + 2Vec⊆3Vec zero = indInIdeal _ _ (suc zero) + 2Vec⊆3Vec (suc zero) = indInIdeal _ _ (suc (suc zero)) + + + -- defintion of the universal property + hasZarLatUniversalProp : (L : DistLattice ℓ') (D : R → fst L) + → IsZarMap R' L D + → Type _ + hasZarLatUniversalProp {ℓ' = ℓ'} L D _ = ∀ (L' : DistLattice ℓ') (d : R → fst L') + → IsZarMap R' L' d + → ∃![ χ ∈ DistLatticeHom L L' ] (fst χ) ∘ D ≡ d + + isPropZarLatUniversalProp : (L : DistLattice ℓ') (D : R → fst L) (isZarMapD : IsZarMap R' L D) + → isProp (hasZarLatUniversalProp L D isZarMapD) + isPropZarLatUniversalProp L D isZarMapD = isPropΠ3 (λ _ _ _ → isPropIsContr) + + ZLHasUniversalProp : hasZarLatUniversalProp ZariskiLattice D isZarMapD + ZLHasUniversalProp L' d isZarMapd = (χ , funExt χcomp) , χunique + where + open DistLatticeStr (snd L') renaming (is-set to isSetL) + open LatticeTheory (DistLattice→Lattice L') + open Join L' + open IsLatticeHom + L = fst L' + + χ : DistLatticeHom ZariskiLattice L' + fst χ = SQ.rec isSetL (λ (_ , α) → ⋁ (d ∘ α)) + λ (_ , α) (_ , β) → curriedHelper α β + where + curriedHelper : {n m : ℕ} (α : FinVec R n) (β : FinVec R m) + → √ ⟨ α ⟩ ≡ √ ⟨ β ⟩ → ⋁ (d ∘ α) ≡ ⋁ (d ∘ β) + curriedHelper α β √⟨α⟩≡√⟨β⟩ = is-antisym _ _ ineq1 ineq2 + where + open Order (DistLattice→Lattice L') + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice L')) + open PosetReasoning IndPoset + open PosetStr (IndPoset .snd) hiding (_≤_) + + incl1 : √ ⟨ α ⟩ ⊆ √ ⟨ β ⟩ + incl1 = ⊆-refl-consequence _ _ (cong fst √⟨α⟩≡√⟨β⟩) .fst + + ineq1 : ⋁ (d ∘ α) ≤ ⋁ (d ∘ β) + ineq1 = ⋁IsMax (d ∘ α) (⋁ (d ∘ β)) + λ i → ZarMapRadicalIneq isZarMapd β (α i) (√FGIdealCharLImpl α ⟨ β ⟩ incl1 i) + + incl2 : √ ⟨ β ⟩ ⊆ √ ⟨ α ⟩ + incl2 = ⊆-refl-consequence _ _ (cong fst √⟨α⟩≡√⟨β⟩) .snd + + ineq2 : ⋁ (d ∘ β) ≤ ⋁ (d ∘ α) + ineq2 = ⋁IsMax (d ∘ β) (⋁ (d ∘ α)) + λ i → ZarMapRadicalIneq isZarMapd α (β i) (√FGIdealCharLImpl β ⟨ α ⟩ incl2 i) + + + pres0 (snd χ) = refl + pres1 (snd χ) = ∨lRid _ ∙ isZarMapd .pres1 + pres∨l (snd χ) = elimProp2 (λ _ _ → isSetL _ _) (uncurry (λ n α → uncurry (curriedHelper n α))) + where + curriedHelper : (n : ℕ) (α : FinVec R n) (m : ℕ) (β : FinVec R m) + → ⋁ (d ∘ (α ++Fin β)) ≡ ⋁ (d ∘ α) ∨l ⋁ (d ∘ β) + curriedHelper zero α _ β = sym (∨lLid _) + curriedHelper (suc n) α _ β = + ⋁ (d ∘ (α ++Fin β)) ≡⟨ refl ⟩ + d (α zero) ∨l ⋁ (d ∘ ((α ∘ suc) ++Fin β)) + + ≡⟨ cong (d (α zero) ∨l_) (curriedHelper _ (α ∘ suc) _ β) ⟩ + + d (α zero) ∨l (⋁ (d ∘ α ∘ suc) ∨l ⋁ (d ∘ β)) + ≡⟨ ∨lAssoc _ _ _ ⟩ + + ⋁ (d ∘ α) ∨l ⋁ (d ∘ β) ∎ + + pres∧l (snd χ) = elimProp2 (λ _ _ → isSetL _ _) (uncurry (λ n α → uncurry (curriedHelper n α))) + where + -- have to repeat this one here so the termination checker won't complain + oldHelper : (n : ℕ) (α : FinVec R n) (m : ℕ) (β : FinVec R m) + → ⋁ (d ∘ (α ++Fin β)) ≡ ⋁ (d ∘ α) ∨l ⋁ (d ∘ β) + oldHelper zero α _ β = sym (∨lLid _) + oldHelper (suc n) α _ β = cong (d (α zero) ∨l_) (oldHelper _ (α ∘ suc) _ β) ∙ ∨lAssoc _ _ _ + + curriedHelper : (n : ℕ) (α : FinVec R n) (m : ℕ) (β : FinVec R m) + → ⋁ (d ∘ (α ··Fin β)) ≡ ⋁ (d ∘ α) ∧l ⋁ (d ∘ β) + curriedHelper zero α _ β = sym (0lLeftAnnihilates∧l _) + curriedHelper (suc n) α _ β = + ⋁ (d ∘ (α ··Fin β)) ≡⟨ refl ⟩ + ⋁ (d ∘ ((λ j → α zero · β j) ++Fin ((α ∘ suc) ··Fin β))) + + ≡⟨ oldHelper _ (λ j → α zero · β j) _ ((α ∘ suc) ··Fin β) ⟩ + + ⋁ (d ∘ (λ j → α zero · β j)) ∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β)) + + ≡⟨ cong (_∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β))) (⋁Ext (λ j → isZarMapd .·≡∧ (α zero) (β j))) ⟩ + + ⋁ (λ j → d (α zero) ∧l d (β j)) ∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β)) + + ≡⟨ cong (_∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β))) (sym (⋁Meetrdist _ _)) ⟩ + + (d (α zero) ∧l ⋁ (d ∘ β)) ∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β)) + + ≡⟨ cong ((d (α zero) ∧l ⋁ (d ∘ β)) ∨l_) (curriedHelper _ (α ∘ suc) _ β) ⟩ + + (d (α zero) ∧l ⋁ (d ∘ β)) ∨l (⋁ (d ∘ α ∘ suc) ∧l ⋁ (d ∘ β)) + + ≡⟨ sym (∧lRdist∨l _ _ _) ⟩ + + ⋁ (d ∘ α) ∧l ⋁ (d ∘ β) ∎ + + + χcomp : ∀ (f : R) → χ .fst (D f) ≡ d f + χcomp f = ∨lRid (d f) + + χunique : (y : Σ[ χ' ∈ DistLatticeHom ZariskiLattice L' ] fst χ' ∘ D ≡ d) + → (χ , funExt χcomp) ≡ y + χunique (χ' , χ'∘D≡d) = Σ≡Prop (λ _ → isSetΠ (λ _ → isSetL) _ _) (LatticeHom≡f _ _ + (funExt (elimProp (λ _ → isSetL _ _) (uncurry uniqHelper)))) + where + uniqHelper : (n : ℕ) (α : FinVec R n) → fst χ [ n , α ] ≡ fst χ' [ n , α ] + uniqHelper zero _ = sym (cong (λ α → fst χ' [ 0 , α ]) (funExt (λ ())) ∙ χ' .snd .pres0) + uniqHelper (suc n) α = + ⋁ (d ∘ α) ≡⟨ refl ⟩ + d (α zero) ∨l ⋁ (d ∘ α ∘ suc) + + ≡⟨ cong (d (α zero) ∨l_) (uniqHelper n (α ∘ suc)) ⟩ -- the inductive step + + d (α zero) ∨l fst χ' [ n , α ∘ suc ] + + ≡⟨ cong (_∨l fst χ' [ n , α ∘ suc ]) (sym (funExt⁻ χ'∘D≡d (α zero))) ⟩ + + fst χ' (D (α zero)) ∨l fst χ' [ n , α ∘ suc ] + + ≡⟨ sym (χ' .snd .pres∨l _ _) ⟩ + + fst χ' (D (α zero) ∨z [ n , α ∘ suc ]) + + ≡⟨ cong (λ β → fst χ' [ suc n , β ]) (funExt (λ { zero → refl ; (suc i) → refl })) ⟩ + + fst χ' [ suc n , α ] ∎ + + + -- the map induced by applying the universal property to the Zariski lattice + -- itself is the identity hom + ZLUniversalPropCorollary : ZLHasUniversalProp ZariskiLattice D isZarMapD .fst .fst + ≡ idDistLatticeHom ZariskiLattice + ZLUniversalPropCorollary = cong fst + (ZLHasUniversalProp ZariskiLattice D isZarMapD .snd + (idDistLatticeHom ZariskiLattice , refl)) + + +-- An equivalent definition that doesn't bump up the unviverse level +module SmallZarLat (R' : CommRing ℓ) where + open CommRingStr (snd R') + open CommIdeal R' + open RadicalIdeal R' + open ZarLat R' + + open Iso + + private + R = fst R' + A = Σ[ n ∈ ℕ ] (FinVec R n) + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + -- This is small! + _≼_ : A → A → Type ℓ + (_ , α) ≼ (_ , β) = ∀ i → α i ∈ √ ⟨ β ⟩ + + _∼'_ : A → A → Type ℓ + α ∼' β = (α ≼ β) × (β ≼ α) + + -- lives in the same universe as R + ZL' : Type ℓ + ZL' = A / (_∼'_) + + + IsoLarLatSmall : Iso ZL ZL' + IsoLarLatSmall = relBiimpl→TruncIso ~→∼' ~'→∼ + where + ~→∼' : ∀ {a b : A} → a ∼ b → a ∼' b + ~→∼' r = √FGIdealCharLImpl _ ⟨ _ ⟩ (λ x h → subst (λ p → x ∈ p) r h) + , √FGIdealCharLImpl _ ⟨ _ ⟩ (λ x h → subst (λ p → x ∈ p) (sym r) h) + + ~'→∼ : ∀ {a b : A} → a ∼' b → a ∼ b + ~'→∼ r = CommIdeal≡Char (√FGIdealCharRImpl _ ⟨ _ ⟩ (fst r)) + (√FGIdealCharRImpl _ ⟨ _ ⟩ (snd r)) + + ZL≃ZL' : ZL ≃ ZL' + ZL≃ZL' = isoToEquiv IsoLarLatSmall + + +module BasicOpens (R' : CommRing ℓ) where + open CommRingStr ⦃...⦄ + open RingTheory (CommRing→Ring R') + open CommIdeal R' + open isCommIdeal + + open ZarLat R' + open ZarLatUniversalProp R' + open IsZarMap + + open Join ZariskiLattice + open IsBasis + + private + R = fst R' + instance + _ = snd R' + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + + BasicOpens : ℙ ZL + BasicOpens 𝔞 = (∃[ f ∈ R ] (D f ≡ 𝔞)) , isPropPropTrunc + + BO : Type (ℓ-suc ℓ) + BO = Σ[ 𝔞 ∈ ZL ] (𝔞 ∈ₚ BasicOpens) + + basicOpensAreBasis : IsBasis ZariskiLattice BasicOpens + contains1 basicOpensAreBasis = ∣ 1r , isZarMapD .pres1 ∣ + ∧lClosed basicOpensAreBasis 𝔞 𝔟 = map2 + λ (f , Df≡𝔞) (g , Dg≡𝔟) → (f · g) , isZarMapD .·≡∧ f g ∙ cong₂ (_∧z_) Df≡𝔞 Dg≡𝔟 + ⋁Basis basicOpensAreBasis = elimProp (λ _ → isPropPropTrunc) Σhelper + where + Σhelper : (a : Σ[ n ∈ ℕ ] FinVec R n) + → ∃[ n ∈ ℕ ] Σ[ α ∈ FinVec ZL n ] (∀ i → α i ∈ₚ BasicOpens) × (⋁ α ≡ [ a ]) + Σhelper (n , α) = ∣ n , (D ∘ α) , (λ i → ∣ α i , refl ∣) , path ∣ + where + path : ⋁ (D ∘ α) ≡ [ n , α ] + path = funExt⁻ (cong fst ZLUniversalPropCorollary) _ + + + -- The structure presheaf on BO + BOCat : Category (ℓ-suc ℓ) (ℓ-suc ℓ) + BOCat = ΣPropCat (DistLatticeCategory ZariskiLattice) BasicOpens + + BasisStructurePShf : Functor (BOCat ^op) (CommAlgebrasCategory R') + BasisStructurePShf = universalPShf (DistLatticeCategory ZariskiLattice) + (λ 𝔞 → Σ[ f ∈ R ] (D f ≡ 𝔞)) -- the untruncated defining prop + (λ (_ , f , _) → R[1/ f ]AsCommAlgebra) -- D(f) ↦ R[1/f] + λ (𝔞 , f , p) (𝔟 , g , q) → contrHoms 𝔞 𝔟 f g p q + where + open PreSheafFromUniversalProp + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice ZariskiLattice)) + open InvertingElementsBase R' + + contrHoms : (𝔞 𝔟 : ZL) (f g : R) (p : D f ≡ 𝔞) (q : D g ≡ 𝔟) + → 𝔞 ≤ 𝔟 → isContr (CommAlgebraHom R[1/ g ]AsCommAlgebra R[1/ f ]AsCommAlgebra) + contrHoms 𝔞 𝔟 f g p q 𝔞≤𝔟 = R[1/g]HasAlgUniversalProp R[1/ f ]AsCommAlgebra + λ s s∈[gⁿ|n≥0] → subst-∈ₚ (R[1/ f ]AsCommRing ˣ) + (sym (·Rid (s /1))) --can't apply the lemma directly as we get mult with 1 somewhere + (RadicalLemma.toUnit R' f g f∈√⟨g⟩ s s∈[gⁿ|n≥0]) + where + open AlgLoc R' [ g ⁿ|n≥0] (powersFormMultClosedSubset g) + renaming (S⁻¹RHasAlgUniversalProp to R[1/g]HasAlgUniversalProp) + open S⁻¹RUniversalProp R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) using (_/1) + open RadicalIdeal R' + + private + instance + _ = snd R[1/ f ]AsCommRing + + Df≤Dg : D f ≤ D g + Df≤Dg = subst2 _≤_ (sym p) (sym q) 𝔞≤𝔟 + + radicalHelper : √ ⟨ replicateFinVec 1 f ++Fin replicateFinVec 1 g ⟩ + ≡ √ ⟨ replicateFinVec 1 g ⟩ + radicalHelper = + isEquivRel→effectiveIso (λ _ _ → isSetCommIdeal _ _) ∼EquivRel _ _ .fun Df≤Dg + + f∈√⟨g⟩ : f ∈ √ ⟨ replicateFinVec 1 g ⟩ + f∈√⟨g⟩ = subst (f ∈_) radicalHelper (∈→∈√ _ _ (indInIdeal _ _ zero)) diff --git a/Cubical/Categories/Adjoint.agda b/Cubical/Categories/Adjoint.agda new file mode 100644 index 0000000000..fa6bbd0ef1 --- /dev/null +++ b/Cubical/Categories/Adjoint.agda @@ -0,0 +1,318 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Adjoint where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation +open import Cubical.Foundations.Isomorphism + +open Functor + +open Iso +open Category + +{- +============================================== + Overview +============================================== + +This module contains two definitions for adjoint +functors, and functions witnessing their +logical (and maybe eventually actual?) +equivalence. +-} + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +{- +============================================== + Adjoint definitions +============================================== + +We provide two alternative definitions for +adjoint functors: the unit-counit +definition, followed by the natural bijection +definition. +-} + +module UnitCounit where + + -- Adjoint def 1: unit-counit + record _⊣_ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) (G : Functor D C) + : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + -- unit + η : 𝟙⟨ C ⟩ ⇒ (funcComp G F) + -- counit + ε : (funcComp F G) ⇒ 𝟙⟨ D ⟩ + -- triangle identities + Δ₁ : PathP (λ i → NatTrans (F-lUnit {F = F} i) (F-rUnit {F = F} i)) + (seqTransP F-assoc (F ∘ʳ η) (ε ∘ˡ F)) + (1[ F ]) + Δ₂ : PathP (λ i → NatTrans (F-rUnit {F = G} i) (F-lUnit {F = G} i)) + (seqTransP (sym F-assoc) (η ∘ˡ G) (G ∘ʳ ε)) + (1[ G ]) + + {- + Helper function for building unit-counit adjunctions between categories, + using that equality of natural transformations in a category is equality on objects + -} + + module _ {ℓC ℓC' ℓD ℓD'} + {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} {F : Functor C D} {G : Functor D C} + (η : 𝟙⟨ C ⟩ ⇒ (funcComp G F)) + (ε : (funcComp F G) ⇒ 𝟙⟨ D ⟩) + (Δ₁ : ∀ c → F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε ⟦ F ⟅ c ⟆ ⟧ ≡ D .id) + (Δ₂ : ∀ d → η ⟦ G ⟅ d ⟆ ⟧ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫ ≡ C .id) + where + + make⊣ : F ⊣ G + make⊣ ._⊣_.η = η + make⊣ ._⊣_.ε = ε + make⊣ ._⊣_.Δ₁ = + makeNatTransPathP F-lUnit F-rUnit + (funExt λ c → cong (D ._⋆_ (F ⟪ η ⟦ c ⟧ ⟫)) (transportRefl _) ∙ Δ₁ c) + make⊣ ._⊣_.Δ₂ = + makeNatTransPathP F-rUnit F-lUnit + (funExt λ d → cong (C ._⋆_ (η ⟦ G ⟅ d ⟆ ⟧)) (transportRefl _) ∙ Δ₂ d) + +module NaturalBijection where + -- Adjoint def 2: natural bijection + record _⊣_ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) (G : Functor D C) : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + adjIso : ∀ {c d} → Iso (D [ F ⟅ c ⟆ , d ]) (C [ c , G ⟅ d ⟆ ]) + + infix 40 _♭ + infix 40 _♯ + _♭ : ∀ {c d} → (D [ F ⟅ c ⟆ , d ]) → (C [ c , G ⟅ d ⟆ ]) + (_♭) {_} {_} = adjIso .fun + + _♯ : ∀ {c d} → (C [ c , G ⟅ d ⟆ ]) → (D [ F ⟅ c ⟆ , d ]) + (_♯) {_} {_} = adjIso .inv + + field + adjNatInD : ∀ {c : C .ob} {d d'} (f : D [ F ⟅ c ⟆ , d ]) (k : D [ d , d' ]) + → (f ⋆⟨ D ⟩ k) ♭ ≡ f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ + + adjNatInC : ∀ {c' c d} (g : C [ c , G ⟅ d ⟆ ]) (h : C [ c' , c ]) + → (h ⋆⟨ C ⟩ g) ♯ ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯ + +{- +============================================== + Proofs of equivalence +============================================== + +This first unnamed module provides a function +adj'→adj which takes you from the second +definition to the first. + +The second unnamed module does the reverse. +-} + +module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) (G : Functor D C) where + open UnitCounit + open NaturalBijection renaming (_⊣_ to _⊣²_) + module _ (adj : F ⊣² G) where + open _⊣²_ adj + open _⊣_ + + -- Naturality condition implies that a commutative square in C + -- appears iff the transpose in D is commutative as well + -- Used in adj'→adj + adjNat' : ∀ {c c' d d'} {f : D [ F ⟅ c ⟆ , d ]} {k : D [ d , d' ]} + → {h : C [ c , c' ]} {g : C [ c' , G ⟅ d' ⟆ ]} + -- commutativity of squares is iff + → ((f ⋆⟨ D ⟩ k ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯) → (f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ ≡ h ⋆⟨ C ⟩ g)) + × ((f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ ≡ h ⋆⟨ C ⟩ g) → (f ⋆⟨ D ⟩ k ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯)) + adjNat' {c} {c'} {d} {d'} {f} {k} {h} {g} = D→C , C→D + where + D→C : (f ⋆⟨ D ⟩ k ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯) → (f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ ≡ h ⋆⟨ C ⟩ g) + D→C eq = f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ + ≡⟨ sym (adjNatInD _ _) ⟩ + ((f ⋆⟨ D ⟩ k) ♭) + ≡⟨ cong _♭ eq ⟩ + (F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯) ♭ + ≡⟨ sym (cong _♭ (adjNatInC _ _)) ⟩ + (h ⋆⟨ C ⟩ g) ♯ ♭ + ≡⟨ adjIso .rightInv _ ⟩ + h ⋆⟨ C ⟩ g + ∎ + C→D : (f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ ≡ h ⋆⟨ C ⟩ g) → (f ⋆⟨ D ⟩ k ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯) + C→D eq = f ⋆⟨ D ⟩ k + ≡⟨ sym (adjIso .leftInv _) ⟩ + (f ⋆⟨ D ⟩ k) ♭ ♯ + ≡⟨ cong _♯ (adjNatInD _ _) ⟩ + (f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫) ♯ + ≡⟨ cong _♯ eq ⟩ + (h ⋆⟨ C ⟩ g) ♯ + ≡⟨ adjNatInC _ _ ⟩ + F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯ + ∎ + + open NatTrans + + -- note : had to make this record syntax because termination checker was complaining + -- due to referencing η and ε from the definitions of Δs + adj'→adj : F ⊣ G + adj'→adj = record + { η = η' + ; ε = ε' + ; Δ₁ = Δ₁' + ; Δ₂ = Δ₂' } + + where + -- ETA + + -- trivial commutative diagram between identities in D + commInD : ∀ {x y} (f : C [ x , y ]) → D .id ⋆⟨ D ⟩ F ⟪ f ⟫ ≡ F ⟪ f ⟫ ⋆⟨ D ⟩ D .id + commInD f = (D .⋆IdL _) ∙ sym (D .⋆IdR _) + + sharpen1 : ∀ {x y} (f : C [ x , y ]) → F ⟪ f ⟫ ⋆⟨ D ⟩ D .id ≡ F ⟪ f ⟫ ⋆⟨ D ⟩ D .id ♭ ♯ + sharpen1 f = cong (λ v → F ⟪ f ⟫ ⋆⟨ D ⟩ v) (sym (adjIso .leftInv _)) + + η' : 𝟙⟨ C ⟩ ⇒ G ∘F F + η' .N-ob x = D .id ♭ + η' .N-hom f = sym (fst (adjNat') (commInD f ∙ sharpen1 f)) + + -- EPSILON + + -- trivial commutative diagram between identities in C + commInC : ∀ {x y} (g : D [ x , y ]) → C .id ⋆⟨ C ⟩ G ⟪ g ⟫ ≡ G ⟪ g ⟫ ⋆⟨ C ⟩ C .id + commInC g = (C .⋆IdL _) ∙ sym (C .⋆IdR _) + + sharpen2 : ∀ {x y} (g : D [ x , y ]) → C .id ♯ ♭ ⋆⟨ C ⟩ G ⟪ g ⟫ ≡ C .id ⋆⟨ C ⟩ G ⟪ g ⟫ + sharpen2 g = cong (λ v → v ⋆⟨ C ⟩ G ⟪ g ⟫) (adjIso .rightInv _) + + ε' : F ∘F G ⇒ 𝟙⟨ D ⟩ + ε' .N-ob x = C .id ♯ + ε' .N-hom g = sym (snd adjNat' (sharpen2 g ∙ commInC g)) + + -- DELTA 1 + + expL : ∀ (c) + → (seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F) .N-ob c) + ≡ F ⟪ η' ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε' ⟦ F ⟅ c ⟆ ⟧ + expL c = seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F) .N-ob c + ≡⟨ refl ⟩ + seqP {C = D} {p = refl} (F ⟪ η' ⟦ c ⟧ ⟫) (ε' ⟦ F ⟅ c ⟆ ⟧) + ≡⟨ seqP≡seq {C = D} _ _ ⟩ + F ⟪ η' ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε' ⟦ F ⟅ c ⟆ ⟧ + ∎ + + body : ∀ (c) + → (idTrans F) ⟦ c ⟧ ≡ (seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F) .N-ob c) + body c = (idTrans F) ⟦ c ⟧ + ≡⟨ refl ⟩ + D .id + ≡⟨ sym (D .⋆IdL _) ⟩ + D .id ⋆⟨ D ⟩ D .id + ≡⟨ snd adjNat' (cong (λ v → (η' ⟦ c ⟧) ⋆⟨ C ⟩ v) (G .F-id)) ⟩ + F ⟪ η' ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε' ⟦ F ⟅ c ⟆ ⟧ + ≡⟨ sym (expL c) ⟩ + seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F) .N-ob c + ∎ + + Δ₁' : PathP (λ i → NatTrans (F-lUnit {F = F} i) (F-rUnit {F = F} i)) + (seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F)) + (1[ F ]) + Δ₁' = makeNatTransPathP F-lUnit F-rUnit (sym (funExt body)) + + -- DELTA 2 + + body2 : ∀ (d) + → seqP {C = C} {p = refl} ((η' ∘ˡ G) ⟦ d ⟧) ((G ∘ʳ ε') ⟦ d ⟧) ≡ C .id + body2 d = seqP {C = C} {p = refl} ((η' ∘ˡ G) ⟦ d ⟧) ((G ∘ʳ ε') ⟦ d ⟧) + ≡⟨ seqP≡seq {C = C} _ _ ⟩ + ((η' ∘ˡ G) ⟦ d ⟧) ⋆⟨ C ⟩ ((G ∘ʳ ε') ⟦ d ⟧) + ≡⟨ refl ⟩ + (η' ⟦ G ⟅ d ⟆ ⟧) ⋆⟨ C ⟩ (G ⟪ ε' ⟦ d ⟧ ⟫) + ≡⟨ fst adjNat' (cong (λ v → v ⋆⟨ D ⟩ (ε' ⟦ d ⟧)) (sym (F .F-id))) ⟩ + C .id ⋆⟨ C ⟩ C .id + ≡⟨ C .⋆IdL _ ⟩ + C .id + ∎ + + Δ₂' : PathP (λ i → NatTrans (F-rUnit {F = G} i) (F-lUnit {F = G} i)) + (seqTransP (sym F-assoc) (η' ∘ˡ G) (G ∘ʳ ε')) + (1[ G ]) + Δ₂' = makeNatTransPathP F-rUnit F-lUnit (funExt body2) + + + module _ (adj : F ⊣ G) where + open _⊣_ adj + open _⊣²_ + open NatTrans + + -- helper functions for working with this Adjoint definition + + δ₁ : ∀ {c} → (F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε ⟦ F ⟅ c ⟆ ⟧) ≡ D .id + δ₁ {c} = (F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε ⟦ F ⟅ c ⟆ ⟧) + ≡⟨ sym (seqP≡seq {C = D} _ _) ⟩ + seqP {C = D} {p = refl} (F ⟪ η ⟦ c ⟧ ⟫) (ε ⟦ F ⟅ c ⟆ ⟧) + ≡⟨ (λ j → (Δ₁ j) .N-ob c) ⟩ + D .id + ∎ + + δ₂ : ∀ {d} → (η ⟦ G ⟅ d ⟆ ⟧ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫) ≡ C .id + δ₂ {d} = (η ⟦ G ⟅ d ⟆ ⟧ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫) + ≡⟨ sym (seqP≡seq {C = C} _ _) ⟩ + seqP {C = C} {p = refl} (η ⟦ G ⟅ d ⟆ ⟧) (G ⟪ ε ⟦ d ⟧ ⟫) + ≡⟨ (λ j → (Δ₂ j) .N-ob d) ⟩ + C .id + ∎ + + + adj→adj' : F ⊣² G + -- ∀ {c d} → Iso (D [ F ⟅ c ⟆ , d ]) (C [ c , G ⟅ d ⟆ ]) + -- takes f to Gf precomposed with the unit + adj→adj' .adjIso {c = c} .fun f = η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ f ⟫ + -- takes g to Fg postcomposed with the counit + adj→adj' .adjIso {d = d} .inv g = F ⟪ g ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ + -- invertibility follows from the triangle identities + adj→adj' .adjIso {c = c} {d} .rightInv g + = η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ F ⟪ g ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ ⟫ -- step0 ∙ step1 ∙ step2 ∙ (C .⋆IdR _) + ≡⟨ cong (λ v → η ⟦ c ⟧ ⋆⟨ C ⟩ v) (G .F-seq _ _) ⟩ + η ⟦ c ⟧ ⋆⟨ C ⟩ (G ⟪ F ⟪ g ⟫ ⟫ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫) + ≡⟨ sym (C .⋆Assoc _ _ _) ⟩ + η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ F ⟪ g ⟫ ⟫ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫ + -- apply naturality + ≡⟨ rCatWhisker {C = C} _ _ _ natu ⟩ + (g ⋆⟨ C ⟩ η ⟦ G ⟅ d ⟆ ⟧) ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫ + ≡⟨ C .⋆Assoc _ _ _ ⟩ + g ⋆⟨ C ⟩ (η ⟦ G ⟅ d ⟆ ⟧ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫) + ≡⟨ lCatWhisker {C = C} _ _ _ δ₂ ⟩ + g ⋆⟨ C ⟩ C .id + ≡⟨ C .⋆IdR _ ⟩ + g + ∎ + where + natu : η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ F ⟪ g ⟫ ⟫ ≡ g ⋆⟨ C ⟩ η ⟦ G ⟅ d ⟆ ⟧ + natu = sym (η .N-hom _) + adj→adj' .adjIso {c = c} {d} .leftInv f + = F ⟪ η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ f ⟫ ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ + ≡⟨ cong (λ v → v ⋆⟨ D ⟩ ε ⟦ d ⟧) (F .F-seq _ _) ⟩ + F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ F ⟪ G ⟪ f ⟫ ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ + ≡⟨ D .⋆Assoc _ _ _ ⟩ + F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ (F ⟪ G ⟪ f ⟫ ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧) + -- apply naturality + ≡⟨ lCatWhisker {C = D} _ _ _ natu ⟩ + F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ (ε ⟦ F ⟅ c ⟆ ⟧ ⋆⟨ D ⟩ f) + ≡⟨ sym (D .⋆Assoc _ _ _) ⟩ + F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε ⟦ F ⟅ c ⟆ ⟧ ⋆⟨ D ⟩ f + -- apply triangle identity + ≡⟨ rCatWhisker {C = D} _ _ _ δ₁ ⟩ + D .id ⋆⟨ D ⟩ f + ≡⟨ D .⋆IdL _ ⟩ + f + ∎ + where + natu : F ⟪ G ⟪ f ⟫ ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ ≡ ε ⟦ F ⟅ c ⟆ ⟧ ⋆⟨ D ⟩ f + natu = ε .N-hom _ + -- follows directly from functoriality + adj→adj' .adjNatInD {c = c} f k = cong (λ v → η ⟦ c ⟧ ⋆⟨ C ⟩ v) (G .F-seq _ _) ∙ (sym (C .⋆Assoc _ _ _)) + adj→adj' .adjNatInC {d = d} g h = cong (λ v → v ⋆⟨ D ⟩ ε ⟦ d ⟧) (F .F-seq _ _) ∙ D .⋆Assoc _ _ _ diff --git a/Cubical/Categories/Category.agda b/Cubical/Categories/Category.agda new file mode 100644 index 0000000000..f92b528338 --- /dev/null +++ b/Cubical/Categories/Category.agda @@ -0,0 +1,21 @@ +{- + Definition of various kinds of categories. + + This library follows the UniMath terminology, that is: + + Concept Ob C Hom C Univalence + + Precategory Type Type No + Category Type Set No + Univalent Category Type Set Yes + + The most useful notion is Category and the library is hence based on + them. If one needs precategories then they can be found in + Cubical.Categories.Category.Precategory + +-} +{-# OPTIONS --safe #-} +module Cubical.Categories.Category where + +open import Cubical.Categories.Category.Base public +open import Cubical.Categories.Category.Properties public diff --git a/Cubical/Categories/Category/Base.agda b/Cubical/Categories/Category/Base.agda new file mode 100644 index 0000000000..6f31d4ad82 --- /dev/null +++ b/Cubical/Categories/Category/Base.agda @@ -0,0 +1,102 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Category.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset + +private + variable + ℓ ℓ' : Level + +-- Categories with hom-sets +record Category ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + -- no-eta-equality ; NOTE: need eta equality for `opop` + field + ob : Type ℓ + Hom[_,_] : ob → ob → Type ℓ' + id : ∀ {x} → Hom[ x , x ] + _⋆_ : ∀ {x y z} (f : Hom[ x , y ]) (g : Hom[ y , z ]) → Hom[ x , z ] + ⋆IdL : ∀ {x y} (f : Hom[ x , y ]) → id ⋆ f ≡ f + ⋆IdR : ∀ {x y} (f : Hom[ x , y ]) → f ⋆ id ≡ f + ⋆Assoc : ∀ {x y z w} (f : Hom[ x , y ]) (g : Hom[ y , z ]) (h : Hom[ z , w ]) + → (f ⋆ g) ⋆ h ≡ f ⋆ (g ⋆ h) + isSetHom : ∀ {x y} → isSet Hom[ x , y ] + + -- composition: alternative to diagramatic order + _∘_ : ∀ {x y z} (g : Hom[ y , z ]) (f : Hom[ x , y ]) → Hom[ x , z ] + g ∘ f = f ⋆ g + + infixr 9 _⋆_ + infixr 9 _∘_ + +open Category + +-- Helpful syntax/notation +_[_,_] : (C : Category ℓ ℓ') → (x y : C .ob) → Type ℓ' +_[_,_] = Hom[_,_] + +-- Needed to define this in order to be able to make the subsequence syntax declaration +seq' : ∀ (C : Category ℓ ℓ') {x y z} (f : C [ x , y ]) (g : C [ y , z ]) → C [ x , z ] +seq' = _⋆_ + +infixl 15 seq' +syntax seq' C f g = f ⋆⟨ C ⟩ g + +-- composition +comp' : ∀ (C : Category ℓ ℓ') {x y z} (g : C [ y , z ]) (f : C [ x , y ]) → C [ x , z ] +comp' = _∘_ + +infixr 16 comp' +syntax comp' C g f = g ∘⟨ C ⟩ f + +-- Isomorphisms and paths in categories +record CatIso (C : Category ℓ ℓ') (x y : C .ob) : Type ℓ' where + constructor catiso + field + mor : C [ x , y ] + inv : C [ y , x ] + sec : inv ⋆⟨ C ⟩ mor ≡ C .id + ret : mor ⋆⟨ C ⟩ inv ≡ C .id + +pathToIso : {C : Category ℓ ℓ'} {x y : C .ob} (p : x ≡ y) → CatIso C x y +pathToIso {C = C} p = J (λ z _ → CatIso _ _ z) (catiso idx idx (C .⋆IdL idx) (C .⋆IdL idx)) p + where + idx = C .id + +-- Univalent Categories +record isUnivalent (C : Category ℓ ℓ') : Type (ℓ-max ℓ ℓ') where + field + univ : (x y : C .ob) → isEquiv (pathToIso {C = C} {x = x} {y = y}) + + -- package up the univalence equivalence + univEquiv : ∀ (x y : C .ob) → (x ≡ y) ≃ (CatIso _ x y) + univEquiv x y = pathToIso , univ x y + + -- The function extracting paths from category-theoretic isomorphisms. + CatIsoToPath : {x y : C .ob} (p : CatIso _ x y) → x ≡ y + CatIsoToPath {x = x} {y = y} p = + equivFun (invEquiv (univEquiv x y)) p + +-- Opposite category +_^op : Category ℓ ℓ' → Category ℓ ℓ' +ob (C ^op) = ob C +Hom[_,_] (C ^op) x y = C [ y , x ] +id (C ^op) = id C +_⋆_ (C ^op) f g = g ⋆⟨ C ⟩ f +⋆IdL (C ^op) = C .⋆IdR +⋆IdR (C ^op) = C .⋆IdL +⋆Assoc (C ^op) f g h = sym (C .⋆Assoc _ _ _) +isSetHom (C ^op) = C .isSetHom + + +ΣPropCat : (C : Category ℓ ℓ') (P : ℙ (ob C)) → Category ℓ ℓ' +ob (ΣPropCat C P) = Σ[ x ∈ ob C ] x ∈ P +Hom[_,_] (ΣPropCat C P) x y = C [ fst x , fst y ] +id (ΣPropCat C P) = id C +_⋆_ (ΣPropCat C P) = _⋆_ C +⋆IdL (ΣPropCat C P) = ⋆IdL C +⋆IdR (ΣPropCat C P) = ⋆IdR C +⋆Assoc (ΣPropCat C P) = ⋆Assoc C +isSetHom (ΣPropCat C P) = isSetHom C diff --git a/Cubical/Categories/Category/Precategory.agda b/Cubical/Categories/Category/Precategory.agda new file mode 100644 index 0000000000..274403bc41 --- /dev/null +++ b/Cubical/Categories/Category/Precategory.agda @@ -0,0 +1,54 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Category.Precategory where + +open import Cubical.Foundations.Prelude + +private + variable + ℓ ℓ' : Level + +-- Precategories +record Precategory ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + -- no-eta-equality ; NOTE: need eta equality for `opop` + field + ob : Type ℓ + Hom[_,_] : ob → ob → Type ℓ' + id : ∀ {x} → Hom[ x , x ] + _⋆_ : ∀ {x y z} (f : Hom[ x , y ]) (g : Hom[ y , z ]) → Hom[ x , z ] + ⋆IdL : ∀ {x y} (f : Hom[ x , y ]) → id ⋆ f ≡ f + ⋆IdR : ∀ {x y} (f : Hom[ x , y ]) → f ⋆ id ≡ f + ⋆Assoc : ∀ {u v w x} (f : Hom[ u , v ]) (g : Hom[ v , w ]) (h : Hom[ w , x ]) → (f ⋆ g) ⋆ h ≡ f ⋆ (g ⋆ h) + + -- composition: alternative to diagramatic order + _∘_ : ∀ {x y z} (g : Hom[ y , z ]) (f : Hom[ x , y ]) → Hom[ x , z ] + g ∘ f = f ⋆ g + +open Precategory + +-- Helpful syntax/notation +_[_,_] : (C : Precategory ℓ ℓ') → (x y : C .ob) → Type ℓ' +_[_,_] = Hom[_,_] + +-- Needed to define this in order to be able to make the subsequence syntax declaration +seq' : ∀ (C : Precategory ℓ ℓ') {x y z} (f : C [ x , y ]) (g : C [ y , z ]) → C [ x , z ] +seq' = _⋆_ + +infixl 15 seq' +syntax seq' C f g = f ⋆⟨ C ⟩ g + +-- composition +comp' : ∀ (C : Precategory ℓ ℓ') {x y z} (g : C [ y , z ]) (f : C [ x , y ]) → C [ x , z ] +comp' = _∘_ + +infixr 16 comp' +syntax comp' C g f = g ∘⟨ C ⟩ f + +-- Opposite precategory +_^op : Precategory ℓ ℓ' → Precategory ℓ ℓ' +(C ^op) .ob = C .ob +(C ^op) .Hom[_,_] x y = C .Hom[_,_] y x +(C ^op) .id = C .id +(C ^op) ._⋆_ f g = C ._⋆_ g f +(C ^op) .⋆IdL = C .⋆IdR +(C ^op) .⋆IdR = C .⋆IdL +(C ^op) .⋆Assoc f g h = sym (C .⋆Assoc _ _ _) diff --git a/Cubical/Categories/Category/Properties.agda b/Cubical/Categories/Category/Properties.agda new file mode 100644 index 0000000000..ad158fa1f0 --- /dev/null +++ b/Cubical/Categories/Category/Properties.agda @@ -0,0 +1,92 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Category.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Categories.Category.Base + +open Category + +private + variable + ℓ ℓ' : Level + +module _ {C : Category ℓ ℓ'} where + + -- isSet where your allowed to compare paths where one side is only + -- equal up to path. Is there a generalization of this? + isSetHomP1 : ∀ {x y : C .ob} {n : C [ x , y ]} + → isOfHLevelDep 1 (λ m → m ≡ n) + isSetHomP1 {x = x} {y} {n} = isOfHLevel→isOfHLevelDep 1 (λ m → isSetHom C m n) + + -- isSet where the arrows can be between non-definitionally equal obs + isSetHomP2l : ∀ {y : C .ob} + → isOfHLevelDep 2 (λ x → C [ x , y ]) + isSetHomP2l = isOfHLevel→isOfHLevelDep 2 (λ a → isSetHom C {x = a}) + + isSetHomP2r : ∀ {x : C .ob} + → isOfHLevelDep 2 (λ y → C [ x , y ]) + isSetHomP2r = isOfHLevel→isOfHLevelDep 2 (λ a → isSetHom C {y = a}) + + +-- opposite of opposite is definitionally equal to itself +involutiveOp : ∀ {C : Category ℓ ℓ'} → C ^op ^op ≡ C +involutiveOp = refl + +module _ {C : Category ℓ ℓ'} where + -- Other useful operations on categories + + -- whisker the parallel morphisms g and g' with f + lCatWhisker : {x y z : C .ob} (f : C [ x , y ]) (g g' : C [ y , z ]) (p : g ≡ g') + → f ⋆⟨ C ⟩ g ≡ f ⋆⟨ C ⟩ g' + lCatWhisker f _ _ p = cong (_⋆_ C f) p + + rCatWhisker : {x y z : C .ob} (f f' : C [ x , y ]) (g : C [ y , z ]) (p : f ≡ f') + → f ⋆⟨ C ⟩ g ≡ f' ⋆⟨ C ⟩ g + rCatWhisker _ _ g p = cong (λ v → v ⋆⟨ C ⟩ g) p + + -- working with equal objects + idP : ∀ {x x'} {p : x ≡ x'} → C [ x , x' ] + idP {x} {x'} {p} = subst (λ v → C [ x , v ]) p (C .id) + + -- heterogeneous seq + seqP : ∀ {x y y' z} {p : y ≡ y'} + → (f : C [ x , y ]) (g : C [ y' , z ]) + → C [ x , z ] + seqP {x} {_} {_} {z} {p} f g = f ⋆⟨ C ⟩ (subst (λ a → C [ a , z ]) (sym p) g) + + -- also heterogeneous seq, but substituting on the left + seqP' : ∀ {x y y' z} {p : y ≡ y'} + → (f : C [ x , y ]) (g : C [ y' , z ]) + → C [ x , z ] + seqP' {x} {_} {_} {z} {p} f g = subst (λ a → C [ x , a ]) p f ⋆⟨ C ⟩ g + + -- show that they're equal + seqP≡seqP' : ∀ {x y y' z} {p : y ≡ y'} + → (f : C [ x , y ]) (g : C [ y' , z ]) + → seqP {p = p} f g ≡ seqP' {p = p} f g + seqP≡seqP' {x = x} {z = z} {p = p} f g i = + (toPathP {A = λ i' → C [ x , p i' ]} {f} refl i) + ⋆⟨ C ⟩ + (toPathP {A = λ i' → C [ p (~ i') , z ]} {x = g} (sym refl) (~ i)) + + -- seqP is equal to normal seq when y ≡ y' + seqP≡seq : ∀ {x y z} + → (f : C [ x , y ]) (g : C [ y , z ]) + → seqP {p = refl} f g ≡ f ⋆⟨ C ⟩ g + seqP≡seq {y = y} {z} f g i = f ⋆⟨ C ⟩ toPathP {A = λ _ → C [ y , z ]} {x = g} refl (~ i) + + + -- whiskering with heterogenous seq -- (maybe should let z be heterogeneous too) + lCatWhiskerP : {x y z y' : C .ob} {p : y ≡ y'} + → (f : C [ x , y ]) (g : C [ y , z ]) (g' : C [ y' , z ]) + → (r : PathP (λ i → C [ p i , z ]) g g') + → f ⋆⟨ C ⟩ g ≡ seqP {p = p} f g' + lCatWhiskerP f g g' r = cong (λ v → f ⋆⟨ C ⟩ v) (sym (fromPathP (symP r))) + + rCatWhiskerP : {x y' y z : C .ob} {p : y' ≡ y} + → (f' : C [ x , y' ]) (f : C [ x , y ]) (g : C [ y , z ]) + → (r : PathP (λ i → C [ x , p i ]) f' f) + → f ⋆⟨ C ⟩ g ≡ seqP' {p = p} f' g + rCatWhiskerP f' f g r = cong (λ v → v ⋆⟨ C ⟩ g) (sym (fromPathP r)) diff --git a/Cubical/Categories/Commutativity.agda b/Cubical/Categories/Commutativity.agda new file mode 100644 index 0000000000..ae902eb084 --- /dev/null +++ b/Cubical/Categories/Commutativity.agda @@ -0,0 +1,33 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Commutativity where + +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Category + +private + variable + ℓ ℓ' : Level + +module _ {C : Category ℓ ℓ'} where + open Category C + + compSq : ∀ {x y z w u v} {f : C [ x , y ]} {g h} {k : C [ z , w ]} {l} {m} {n : C [ u , v ]} + -- square 1 + → f ⋆ g ≡ h ⋆ k + -- square 2 (sharing g) + → k ⋆ l ≡ m ⋆ n + → f ⋆ (g ⋆ l) ≡ (h ⋆ m) ⋆ n + compSq {f = f} {g} {h} {k} {l} {m} {n} p q + = f ⋆ (g ⋆ l) + ≡⟨ sym (⋆Assoc _ _ _) ⟩ + (f ⋆ g) ⋆ l + ≡⟨ cong (_⋆ l) p ⟩ + (h ⋆ k) ⋆ l + ≡⟨ ⋆Assoc _ _ _ ⟩ + h ⋆ (k ⋆ l) + ≡⟨ cong (h ⋆_) q ⟩ + h ⋆ (m ⋆ n) + ≡⟨ sym (⋆Assoc _ _ _) ⟩ + (h ⋆ m) ⋆ n + ∎ diff --git a/Cubical/Categories/Constructions/BinProduct.agda b/Cubical/Categories/Constructions/BinProduct.agda new file mode 100644 index 0000000000..c6498dc2d3 --- /dev/null +++ b/Cubical/Categories/Constructions/BinProduct.agda @@ -0,0 +1,64 @@ +-- Product of two categories +{-# OPTIONS --safe #-} + +module Cubical.Categories.Constructions.BinProduct where + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor.Base +open import Cubical.Data.Sigma renaming (_×_ to _×'_) +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude + +private + variable + ℓC ℓC' ℓD ℓD' ℓE ℓE' : Level + + +open Category + +_×_ : (C : Category ℓC ℓC') → (D : Category ℓD ℓD') + → Category (ℓ-max ℓC ℓD) (ℓ-max ℓC' ℓD') + +(C × D) .ob = (ob C) ×' (ob D) +(C × D) .Hom[_,_] (c , d) (c' , d') = (C [ c , c' ]) ×' (D [ d , d' ]) +(C × D) .id = (id C , id D) +(C × D) ._⋆_ _ _ = (_ ⋆⟨ C ⟩ _ , _ ⋆⟨ D ⟩ _) +(C × D) .⋆IdL _ = ≡-× (⋆IdL C _) (⋆IdL D _) +(C × D) .⋆IdR _ = ≡-× (⋆IdR C _) (⋆IdR D _) +(C × D) .⋆Assoc _ _ _ = ≡-× (⋆Assoc C _ _ _) (⋆Assoc D _ _ _) +(C × D) .isSetHom = isSet× (isSetHom C) (isSetHom D) + +infixr 5 _×_ + + +-- Some useful functors +module _ (C : Category ℓC ℓC') + (D : Category ℓD ℓD') where + open Functor + + module _ (E : Category ℓE ℓE') where + -- Associativity of product + ×C-assoc : Functor (C × (D × E)) ((C × D) × E) + ×C-assoc .F-ob (c , (d , e)) = ((c , d), e) + ×C-assoc .F-hom (f , (g , h)) = ((f , g), h) + ×C-assoc .F-id = refl + ×C-assoc .F-seq _ _ = refl + + -- Left/right injections into product + linj : (d : ob D) → Functor C (C × D) + linj d .F-ob c = (c , d) + linj d .F-hom f = (f , id D) + linj d .F-id = refl + linj d .F-seq f g = ≡-× refl (sym (⋆IdL D _)) + + rinj : (c : ob C) → Functor D (C × D) + rinj c .F-ob d = (c , d) + rinj c .F-hom f = (id C , f) + rinj c .F-id = refl + rinj c .F-seq f g = ≡-× (sym (⋆IdL C _)) refl + +{- + TODO: + - define inverse to `assoc`, prove isomorphism + - prove product is commutative up to isomorphism +-} diff --git a/Cubical/Categories/Constructions/Elements.agda b/Cubical/Categories/Constructions/Elements.agda new file mode 100644 index 0000000000..5f33b969ae --- /dev/null +++ b/Cubical/Categories/Constructions/Elements.agda @@ -0,0 +1,138 @@ +{-# OPTIONS --safe #-} + +-- The Category of Elements + +open import Cubical.Categories.Category + +module Cubical.Categories.Constructions.Elements {ℓ ℓ'} {C : Category ℓ ℓ'} where + +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Functor +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma + +import Cubical.Categories.Morphism as Morphism +import Cubical.Categories.Constructions.Slice as Slice + +-- some issues +-- * always need to specify objects during composition because can't infer isSet +open Category +open Functor + + +getIsSet : ∀ {ℓS} {C : Category ℓ ℓ'} (F : Functor C (SET ℓS)) → (c : C .ob) → isSet (fst (F ⟅ c ⟆)) +getIsSet F c = snd (F ⟅ c ⟆) + + +infix 50 ∫_ +∫_ : ∀ {ℓS} → Functor C (SET ℓS) → Category (ℓ-max ℓ ℓS) (ℓ-max ℓ' ℓS) +-- objects are (c , x) pairs where c ∈ C and x ∈ F c +(∫ F) .ob = Σ[ c ∈ C .ob ] fst (F ⟅ c ⟆) +-- morphisms are f : c → c' which take x to x' +(∫ F) .Hom[_,_] (c , x) (c' , x') = Σ[ f ∈ C [ c , c' ] ] x' ≡ (F ⟪ f ⟫) x +(∫ F) .id {x = (c , x)} = C .id , sym (funExt⁻ (F .F-id) x ∙ refl) +(∫ F) ._⋆_ {c , x} {c₁ , x₁} {c₂ , x₂} (f , p) (g , q) + = (f ⋆⟨ C ⟩ g) , (x₂ + ≡⟨ q ⟩ + (F ⟪ g ⟫) x₁ -- basically expanding out function composition + ≡⟨ cong (F ⟪ g ⟫) p ⟩ + (F ⟪ g ⟫) ((F ⟪ f ⟫) x) + ≡⟨ funExt⁻ (sym (F .F-seq _ _)) _ ⟩ + (F ⟪ f ⋆⟨ C ⟩ g ⟫) x + ∎) +(∫ F) .⋆IdL o@{c , x} o1@{c' , x'} f'@(f , p) i + = (cIdL i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS' x' ((F ⟪ a ⟫) x)) p' p cIdL i + where + isS = getIsSet F c + isS' = getIsSet F c' + cIdL = C .⋆IdL f + + -- proof from composition with id + p' : x' ≡ (F ⟪ C .id ⋆⟨ C ⟩ f ⟫) x + p' = snd ((∫ F) ._⋆_ ((∫ F) .id) f') +(∫ F) .⋆IdR o@{c , x} o1@{c' , x'} f'@(f , p) i + = (cIdR i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS' x' ((F ⟪ a ⟫) x)) p' p cIdR i + where + cIdR = C .⋆IdR f + isS' = getIsSet F c' + + p' : x' ≡ (F ⟪ f ⋆⟨ C ⟩ C .id ⟫) x + p' = snd ((∫ F) ._⋆_ f' ((∫ F) .id)) +(∫ F) .⋆Assoc o@{c , x} o1@{c₁ , x₁} o2@{c₂ , x₂} o3@{c₃ , x₃} f'@(f , p) g'@(g , q) h'@(h , r) i + = (cAssoc i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS₃ x₃ ((F ⟪ a ⟫) x)) p1 p2 cAssoc i + where + cAssoc = C .⋆Assoc f g h + isS₃ = getIsSet F c₃ + + p1 : x₃ ≡ (F ⟪ (f ⋆⟨ C ⟩ g) ⋆⟨ C ⟩ h ⟫) x + p1 = snd ((∫ F) ._⋆_ ((∫ F) ._⋆_ {o} {o1} {o2} f' g') h') + + p2 : x₃ ≡ (F ⟪ f ⋆⟨ C ⟩ (g ⋆⟨ C ⟩ h) ⟫) x + p2 = snd ((∫ F) ._⋆_ f' ((∫ F) ._⋆_ {o1} {o2} {o3} g' h')) +(∫ F) .isSetHom {x} {y} = isSetΣSndProp (C .isSetHom) λ _ → (F ⟅ fst y ⟆) .snd _ _ + +-- same thing but for presheaves +∫ᴾ_ : ∀ {ℓS} → Functor (C ^op) (SET ℓS) → Category (ℓ-max ℓ ℓS) (ℓ-max ℓ' ℓS) +-- objects are (c , x) pairs where c ∈ C and x ∈ F c +(∫ᴾ F) .ob = Σ[ c ∈ C .ob ] fst (F ⟅ c ⟆) +-- morphisms are f : c → c' which take x to x' +(∫ᴾ F) .Hom[_,_] (c , x) (c' , x') = Σ[ f ∈ C [ c , c' ] ] x ≡ (F ⟪ f ⟫) x' +(∫ᴾ F) .id {x = (c , x)} = C .id , sym (funExt⁻ (F .F-id) x ∙ refl) +(∫ᴾ F) ._⋆_ {c , x} {c₁ , x₁} {c₂ , x₂} (f , p) (g , q) + = (f ⋆⟨ C ⟩ g) , (x + ≡⟨ p ⟩ + (F ⟪ f ⟫) x₁ -- basically expanding out function composition + ≡⟨ cong (F ⟪ f ⟫) q ⟩ + (F ⟪ f ⟫) ((F ⟪ g ⟫) x₂) + ≡⟨ funExt⁻ (sym (F .F-seq _ _)) _ ⟩ + (F ⟪ f ⋆⟨ C ⟩ g ⟫) x₂ + ∎) +(∫ᴾ F) .⋆IdL o@{c , x} o1@{c' , x'} f'@(f , p) i + = (cIdL i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS x ((F ⟪ a ⟫) x')) p' p cIdL i + where + isS = getIsSet F c + isS' = getIsSet F c' + cIdL = C .⋆IdL f + + -- proof from composition with id + p' : x ≡ (F ⟪ C .id ⋆⟨ C ⟩ f ⟫) x' + p' = snd ((∫ᴾ F) ._⋆_ ((∫ᴾ F) .id) f') +(∫ᴾ F) .⋆IdR o@{c , x} o1@{c' , x'} f'@(f , p) i + = (cIdR i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS x ((F ⟪ a ⟫) x')) p' p cIdR i + where + cIdR = C .⋆IdR f + isS = getIsSet F c + + p' : x ≡ (F ⟪ f ⋆⟨ C ⟩ C .id ⟫) x' + p' = snd ((∫ᴾ F) ._⋆_ f' ((∫ᴾ F) .id)) +(∫ᴾ F) .⋆Assoc o@{c , x} o1@{c₁ , x₁} o2@{c₂ , x₂} o3@{c₃ , x₃} f'@(f , p) g'@(g , q) h'@(h , r) i + = (cAssoc i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS x ((F ⟪ a ⟫) x₃)) p1 p2 cAssoc i + where + cAssoc = C .⋆Assoc f g h + isS = getIsSet F c + + p1 : x ≡ (F ⟪ (f ⋆⟨ C ⟩ g) ⋆⟨ C ⟩ h ⟫) x₃ + p1 = snd ((∫ᴾ F) ._⋆_ ((∫ᴾ F) ._⋆_ {o} {o1} {o2} f' g') h') + + p2 : x ≡ (F ⟪ f ⋆⟨ C ⟩ (g ⋆⟨ C ⟩ h) ⟫) x₃ + p2 = snd ((∫ᴾ F) ._⋆_ f' ((∫ᴾ F) ._⋆_ {o1} {o2} {o3} g' h')) +(∫ᴾ F) .isSetHom {x} = isSetΣSndProp (C .isSetHom) (λ _ → (F ⟅ fst x ⟆) .snd _ _) + +-- helpful results + +module _ {ℓS} {F : Functor (C ^op) (SET ℓS)} where + + -- morphisms are equal as long as the morphisms in C are equals + ∫ᴾhomEq : ∀ {o1 o1' o2 o2'} (f : (∫ᴾ F) [ o1 , o2 ]) (g : (∫ᴾ F) [ o1' , o2' ]) + → (p : o1 ≡ o1') (q : o2 ≡ o2') + → (eqInC : PathP (λ i → C [ fst (p i) , fst (q i) ]) (fst f) (fst g)) + → PathP (λ i → (∫ᴾ F) [ p i , q i ]) f g + ∫ᴾhomEq (f , eqf) (g , eqg) p q eqInC + = ΣPathP (eqInC + , isOfHLevel→isOfHLevelDep 1 {A = Σ[ (o1 , o2) ∈ (∫ᴾ F) .ob × (∫ᴾ F) .ob ] (C [ fst o1 , fst o2 ])} + {B = λ ((o1 , o2) , f) → snd o1 ≡ (F ⟪ f ⟫) (snd o2)} + (λ ((o1 , o2) , f) → snd (F ⟅ (fst o1) ⟆) (snd o1) ((F ⟪ f ⟫) (snd o2))) + eqf + eqg + λ i → ((p i , q i) , eqInC i)) diff --git a/Cubical/Categories/Constructions/Product.agda b/Cubical/Categories/Constructions/Product.agda new file mode 100644 index 0000000000..a013278ea8 --- /dev/null +++ b/Cubical/Categories/Constructions/Product.agda @@ -0,0 +1,64 @@ +-- Product of two categories +{-# OPTIONS --safe #-} + +module Cubical.Categories.Constructions.Product where + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor.Base +open import Cubical.Data.Sigma renaming (_×_ to _×'_) +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude + +private + variable + ℓC ℓC' ℓD ℓD' ℓE ℓE' : Level + + +open Category + +_×_ : (C : Category ℓC ℓC') → (D : Category ℓD ℓD') + → Category (ℓ-max ℓC ℓD) (ℓ-max ℓC' ℓD') + +(C × D) .ob = (ob C) ×' (ob D) +(C × D) .Hom[_,_] (c , d) (c' , d') = (C [ c , c' ]) ×' (D [ d , d' ]) +(C × D) .id = (id C , id D) +(C × D) ._⋆_ _ _ = (_ ⋆⟨ C ⟩ _ , _ ⋆⟨ D ⟩ _) +(C × D) .⋆IdL _ = ≡-× (⋆IdL C _) (⋆IdL D _) +(C × D) .⋆IdR _ = ≡-× (⋆IdR C _) (⋆IdR D _) +(C × D) .⋆Assoc _ _ _ = ≡-× (⋆Assoc C _ _ _) (⋆Assoc D _ _ _) +(C × D) .isSetHom = isSet× (isSetHom C) (isSetHom D) + +infixr 5 _×_ + + +-- Some useful functors +module _ (C : Category ℓC ℓC') + (D : Category ℓD ℓD') where + open Functor + + module _ (E : Category ℓE ℓE') where + -- Associativity of product + ×C-assoc : Functor (C × (D × E)) ((C × D) × E) + ×C-assoc .F-ob (c , (d , e)) = ((c , d), e) + ×C-assoc .F-hom (f , (g , h)) = ((f , g), h) + ×C-assoc .F-id = refl + ×C-assoc .F-seq _ _ = refl + + -- Left/right injections into product + linj : (d : ob D) → Functor C (C × D) + linj d .F-ob c = (c , d) + linj d .F-hom f = (f , id D) + linj d .F-id = refl + linj d .F-seq f g = ≡-× refl (sym (⋆IdL D _)) + + rinj : (c : ob C) → Functor D (C × D) + rinj c .F-ob d = (c , d) + rinj c .F-hom f = (id C , f) + rinj c .F-id = refl + rinj c .F-seq f g = ≡-× (sym (⋆IdL C _)) refl + +{- + TODO: + - define inverse to `assoc`, prove isomorphism + - prove product is commutative up to isomorphism +-} diff --git a/Cubical/Categories/Constructions/Quotient.agda b/Cubical/Categories/Constructions/Quotient.agda new file mode 100644 index 0000000000..7082bd9c0d --- /dev/null +++ b/Cubical/Categories/Constructions/Quotient.agda @@ -0,0 +1,83 @@ +-- Quotient category +{-# OPTIONS --safe #-} + +module Cubical.Categories.Constructions.Quotient where + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Limits.Terminal +open import Cubical.Categories.Limits.Initial +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.HITs.SetQuotients renaming ([_] to ⟦_⟧) + +private + variable + ℓ ℓ' ℓq : Level + +module _ (C : Category ℓ ℓ') where + open Category C + + module _ (_~_ : {x y : ob} (f g : Hom[ x , y ] ) → Type ℓq) + (~refl : {x y : ob} (f : Hom[ x , y ] ) → f ~ f) + (~cong : {x y z : ob} + (f f' : Hom[ x , y ]) → f ~ f' + → (g g' : Hom[ y , z ]) → g ~ g' + → (f ⋆ g) ~ (f' ⋆ g')) where + + private + Hom[_,_]/~ = λ (x y : ob) → Hom[ x , y ] / _~_ + + module _ {x y z : ob} where + _⋆/~_ : Hom[ x , y ]/~ → Hom[ y , z ]/~ → Hom[ x , z ]/~ + _⋆/~_ = rec2 squash/ (λ f g → ⟦ f ⋆ g ⟧) + (λ f f' g f~f' → eq/ _ _ (~cong _ _ f~f' _ _ (~refl _))) + (λ f g g' g~g' → eq/ _ _ (~cong _ _ (~refl _) _ _ g~g')) + + module _ {x y : ob} where + ⋆/~IdL : (f : Hom[ x , y ]/~) → (⟦ id ⟧ ⋆/~ f) ≡ f + ⋆/~IdL = elimProp (λ _ → squash/ _ _) (λ _ → cong ⟦_⟧ (⋆IdL _)) + + ⋆/~IdR : (f : Hom[ x , y ]/~) → (f ⋆/~ ⟦ id ⟧) ≡ f + ⋆/~IdR = elimProp (λ _ → squash/ _ _) (λ _ → cong ⟦_⟧ (⋆IdR _)) + + module _ {x y z w : ob} where + ⋆/~Assoc : (f : Hom[ x , y ]/~) + (g : Hom[ y , z ]/~) + (h : Hom[ z , w ]/~) + → ((f ⋆/~ g) ⋆/~ h) ≡ (f ⋆/~ (g ⋆/~ h)) + + ⋆/~Assoc = elimProp3 (λ _ _ _ → squash/ _ _) (λ _ _ _ → cong ⟦_⟧ (⋆Assoc _ _ _)) + + + open Category + QuotientCategory : Category ℓ (ℓ-max ℓ' ℓq) + QuotientCategory .ob = ob C + QuotientCategory .Hom[_,_] x y = (C [ x , y ]) / _~_ + QuotientCategory .id = ⟦ id C ⟧ + QuotientCategory ._⋆_ = _⋆/~_ + QuotientCategory .⋆IdL = ⋆/~IdL + QuotientCategory .⋆IdR = ⋆/~IdR + QuotientCategory .⋆Assoc = ⋆/~Assoc + QuotientCategory .isSetHom = squash/ + + + private + C/~ = QuotientCategory + + -- Quotient map + open Functor + QuoFunctor : Functor C C/~ + QuoFunctor .F-ob x = x + QuoFunctor .F-hom = ⟦_⟧ + QuoFunctor .F-id = refl + QuoFunctor .F-seq f g = refl + + -- Quotients preserve initial / terminal objects + isInitial/~ : {z : ob C} → isInitial C z → isInitial C/~ z + isInitial/~ zInit x = ⟦ zInit x .fst ⟧ , elimProp (λ _ → squash/ _ _) + λ f → eq/ _ _ (subst (_~ f) (sym (zInit x .snd f)) (~refl _)) + + isTerminal/~ : {z : ob C} → isTerminal C z → isTerminal C/~ z + isTerminal/~ zTerminal x = ⟦ zTerminal x .fst ⟧ , elimProp (λ _ → squash/ _ _) + λ f → eq/ _ _ (subst (_~ f) (sym (zTerminal x .snd f)) (~refl _)) diff --git a/Cubical/Categories/Constructions/Slice.agda b/Cubical/Categories/Constructions/Slice.agda new file mode 100644 index 0000000000..faede1e471 --- /dev/null +++ b/Cubical/Categories/Constructions/Slice.agda @@ -0,0 +1,377 @@ +{-# OPTIONS --safe #-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport using (transpFill) + +open import Cubical.Categories.Category +open import Cubical.Categories.Morphism renaming (isIso to isIsoC) + +open import Cubical.Data.Sigma + +open Category +open isUnivalent +open Iso + +module Cubical.Categories.Constructions.Slice {ℓ ℓ' : Level} (C : Category ℓ ℓ') (c : C .ob) where + +-- just a helper to prevent redundency +TypeC : Type (ℓ-suc (ℓ-max ℓ ℓ')) +TypeC = Type (ℓ-max ℓ ℓ') + +-- Components of a slice category + +record SliceOb : TypeC where + constructor sliceob + field + {S-ob} : C .ob + S-arr : C [ S-ob , c ] + +open SliceOb public + +record SliceHom (a b : SliceOb) : Type ℓ' where + constructor slicehom + field + S-hom : C [ S-ob a , S-ob b ] + -- commutative diagram + S-comm : S-hom ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a + +open SliceHom public + +-- Helpers for working with equality +-- can probably replace these by showing that SliceOb is isomorphic to Sigma and +-- that paths are isomorphic to Sigma? But sounds like that would need a lot of transp + +SliceOb-≡-intro : ∀ {a b} {f g} + → (p : a ≡ b) + → PathP (λ i → C [ p i , c ]) f g + → sliceob {a} f ≡ sliceob {b} g +SliceOb-≡-intro p q = λ i → sliceob {p i} (q i) + +module _ {xf yg : SliceOb} where + private + x = xf .S-ob + f = xf .S-arr + y = yg .S-ob + g = yg .S-arr + + -- a path between slice objects is the "same" as a pair of paths between C obs and C arrows + SOPathIsoPathΣ : Iso (xf ≡ yg) (Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g) + SOPathIsoPathΣ .fun p = (λ i → (p i) .S-ob) , (λ i → (p i) .S-arr) + SOPathIsoPathΣ .inv (p , q) i = sliceob {p i} (q i) + SOPathIsoPathΣ .rightInv _ = refl + SOPathIsoPathΣ .leftInv _ = refl + + SOPath≃PathΣ = isoToEquiv SOPathIsoPathΣ + + SOPath≡PathΣ = ua (isoToEquiv SOPathIsoPathΣ) + +-- intro and elim for working with SliceHom equalities (is there a better way to do this?) +SliceHom-≡-intro : ∀ {a b} {f g} {c₁} {c₂} + → (p : f ≡ g) + → PathP (λ i → (p i) ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) c₁ c₂ + → slicehom f c₁ ≡ slicehom g c₂ +SliceHom-≡-intro p q = λ i → slicehom (p i) (q i) + +SliceHom-≡-elim : ∀ {a b} {f g} {c₁} {c₂} + → slicehom f c₁ ≡ slicehom g c₂ + → Σ[ p ∈ f ≡ g ] PathP (λ i → (p i) ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) c₁ c₂ +SliceHom-≡-elim r = (λ i → S-hom (r i)) , λ i → S-comm (r i) + + +SliceHom-≡-intro' : ∀ {a b} {f g : C [ a .S-ob , b .S-ob ]} {c₁} {c₂} + → (p : f ≡ g) + → slicehom f c₁ ≡ slicehom g c₂ +SliceHom-≡-intro' {a} {b} {f} {g} {c₁} {c₂} p i = slicehom (p i) (c₁≡c₂ i) + where + c₁≡c₂ : PathP (λ i → (p i) ⋆⟨ C ⟩ (b .S-arr) ≡ a .S-arr) c₁ c₂ + c₁≡c₂ = isOfHLevel→isOfHLevelDep 1 (λ _ → C .isSetHom _ _) c₁ c₂ p + +-- SliceHom is isomorphic to the Sigma type with the same components +SliceHom-Σ-Iso : ∀ {a b} + → Iso (SliceHom a b) (Σ[ h ∈ C [ S-ob a , S-ob b ] ] h ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) +SliceHom-Σ-Iso .fun (slicehom h c) = h , c +SliceHom-Σ-Iso .inv (h , c) = slicehom h c +SliceHom-Σ-Iso .rightInv = λ x → refl +SliceHom-Σ-Iso .leftInv = λ x → refl + + +-- Category definition + +SliceCat : Category (ℓ-max ℓ ℓ') ℓ' +ob SliceCat = SliceOb +Hom[_,_] SliceCat = SliceHom +id SliceCat = slicehom (C .id) (C .⋆IdL _) +_⋆_ SliceCat {sliceob j} {sliceob k} {sliceob l} (slicehom f p) (slicehom g p') = + slicehom + (f ⋆⟨ C ⟩ g) + ( f ⋆⟨ C ⟩ g ⋆⟨ C ⟩ l + ≡⟨ C .⋆Assoc _ _ _ ⟩ + f ⋆⟨ C ⟩ (g ⋆⟨ C ⟩ l) + ≡⟨ cong (λ v → f ⋆⟨ C ⟩ v) p' ⟩ + f ⋆⟨ C ⟩ k + ≡⟨ p ⟩ + j + ∎) +⋆IdL SliceCat (slicehom S-hom S-comm) = + SliceHom-≡-intro (⋆IdL C _) (toPathP (C .isSetHom _ _ _ _)) +⋆IdR SliceCat (slicehom S-hom S-comm) = + SliceHom-≡-intro (⋆IdR C _) (toPathP (C .isSetHom _ _ _ _)) +⋆Assoc SliceCat f g h = + SliceHom-≡-intro (⋆Assoc C _ _ _) (toPathP (C .isSetHom _ _ _ _)) +isSetHom SliceCat {a} {b} (slicehom f c₁) (slicehom g c₂) p q = cong isoP p'≡q' + where + -- paths between SliceHoms are equivalent to the projection paths + p' : Σ[ p ∈ f ≡ g ] PathP (λ i → (p i) ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) c₁ c₂ + p' = SliceHom-≡-elim p + q' : Σ[ p ∈ f ≡ g ] PathP (λ i → (p i) ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) c₁ c₂ + q' = SliceHom-≡-elim q + + -- we want all paths between (dependent) paths of this type to be equal + B = λ v → v ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a + + -- need the groupoidness for dependent paths + isGroupoidDepHom : isOfHLevelDep 2 B + isGroupoidDepHom = isOfHLevel→isOfHLevelDep 2 (λ v x y → isSet→isGroupoid (C .isSetHom) _ _ x y) + + -- we first prove that the projected paths are equal + p'≡q' : p' ≡ q' + p'≡q' = ΣPathP (C .isSetHom _ _ _ _ , toPathP (isGroupoidDepHom _ _ _ _ _)) + + -- and then we can use equivalence to lift these paths up + -- to actual SliceHom paths + isoP = λ g → cong (inv SliceHom-Σ-Iso) (fun (ΣPathIsoPathΣ) g) + +-- SliceCat is univalent if C is univalent + +module _ ⦃ isU : isUnivalent C ⦄ where + open CatIso + open Iso + + module _ { xf yg : SliceOb } where + private + x = xf .S-ob + y = yg .S-ob + + -- names for the equivalences/isos + + pathIsoEquiv : (x ≡ y) ≃ (CatIso _ x y) + pathIsoEquiv = univEquiv isU x y + + isoPathEquiv : (CatIso _ x y) ≃ (x ≡ y) + isoPathEquiv = invEquiv pathIsoEquiv + + pToIIso' : Iso (x ≡ y) (CatIso _ x y) + pToIIso' = equivToIso pathIsoEquiv + + -- the iso in SliceCat we're given induces an iso in C between x and y + module _ ( cIso@(catiso kc lc s r) : CatIso SliceCat xf yg ) where + extractIso' : CatIso C x y + extractIso' .mor = kc .S-hom + extractIso' .inv = lc .S-hom + extractIso' .sec i = (s i) .S-hom + extractIso' .ret i = (r i) .S-hom + + instance + preservesUnivalenceSlice : isUnivalent SliceCat + -- we prove the equivalence by going through Iso + preservesUnivalenceSlice .univ xf@(sliceob {x} f) yg@(sliceob {y} g) = isoToIsEquiv sIso + where + -- this is just here because the type checker can't seem to infer xf and yg + pToIIso : Iso (x ≡ y) (CatIso _ x y) + pToIIso = pToIIso' {xf = xf} {yg} + + -- the meat of the proof + sIso : Iso (xf ≡ yg) (CatIso _ xf yg) + sIso .fun p = pathToIso p -- we use the normal pathToIso via path induction to get an isomorphism + sIso .inv is@(catiso kc lc s r) = SliceOb-≡-intro x≡y (symP (sym (lc .S-comm) ◁ lf≡f)) + where + -- we get a path between xf and yg by combining paths between + -- x and y, and f and g + -- 1. x≡y follows from univalence of C + -- 2. f≡g is more tricky; by commutativity, we know that g ≡ l ⋆ f + -- so we want l to be id; we get this by showing: id ≡ pathToIso x y x≡y ≡ l + -- where the first step follows from path induction, and the second from univalence of C + + -- morphisms in C from kc and lc + k = kc .S-hom + l = lc .S-hom + + -- extract out the iso between x and y + extractIso : CatIso C x y + extractIso = extractIso' is + + -- and we can use univalence of C to get x ≡ y + x≡y : x ≡ y + x≡y = pToIIso .inv extractIso + + -- to show that f ≡ g, we show that l ≡ id + -- by using C's isomorphism + pToI≡id : PathP (λ i → C [ x≡y (~ i) , x ]) (pathToIso {C = C} x≡y .inv) (C .id) + pToI≡id = J (λ y p → PathP (λ i → C [ p (~ i) , x ]) (pathToIso {C = C} p .inv) (C .id)) + (λ j → JRefl pToIFam pToIBase j .inv) + x≡y + where + idx = C .id + pToIFam = (λ z _ → CatIso C x z) + pToIBase = catiso (C .id) idx (C .⋆IdL idx) (C .⋆IdL idx) + + l≡pToI : l ≡ pathToIso {C = C} x≡y .inv + l≡pToI i = pToIIso .rightInv extractIso (~ i) .inv + + l≡id : PathP (λ i → C [ x≡y (~ i) , x ]) l (C .id) + l≡id = l≡pToI ◁ pToI≡id + + lf≡f : PathP (λ i → C [ x≡y (~ i) , c ]) (l ⋆⟨ C ⟩ f) f + lf≡f = (λ i → (l≡id i) ⋆⟨ C ⟩ f) ▷ C .⋆IdL _ + + sIso .rightInv is@(catiso kc lc s r) i = catiso (kc'≡kc i) (lc'≡lc i) (s'≡s i) (r'≡r i) + -- we prove rightInv using a combination of univalence and the fact that homs are an h-set + where + kc' = (sIso .fun) (sIso .inv is) .mor + lc' = (sIso .fun) (sIso .inv is) .inv + k' = kc' .S-hom + l' = lc' .S-hom + k = kc .S-hom + l = lc .S-hom + + extractIso : CatIso C x y + extractIso = extractIso' is + + -- we do the equality component wise + + -- mor + + k'≡k : k' ≡ k + k'≡k i = (pToIIso .rightInv extractIso) i .mor + + kcom'≡kcom : PathP (λ j → (k'≡k j) ⋆⟨ C ⟩ g ≡ f) (kc' .S-comm) (kc .S-comm) + kcom'≡kcom = isSetHomP1 {C = C} _ _ λ i → (k'≡k i) ⋆⟨ C ⟩ g + kc'≡kc : kc' ≡ kc + kc'≡kc i = slicehom (k'≡k i) (kcom'≡kcom i) + + -- inv + + l'≡l : l' ≡ l + l'≡l i = (pToIIso .rightInv extractIso) i .inv + + lcom'≡lcom : PathP (λ j → (l'≡l j) ⋆⟨ C ⟩ f ≡ g) (lc' .S-comm) (lc .S-comm) + lcom'≡lcom = isSetHomP1 {C = C} _ _ λ i → (l'≡l i) ⋆⟨ C ⟩ f + + lc'≡lc : lc' ≡ lc + lc'≡lc i = slicehom (l'≡l i) (lcom'≡lcom i) + + -- sec + + s' = (sIso .fun) (sIso .inv is) .sec + s'≡s : PathP (λ i → lc'≡lc i ⋆⟨ SliceCat ⟩ kc'≡kc i ≡ SliceCat .id) s' s + s'≡s = isSetHomP1 {C = SliceCat} _ _ λ i → lc'≡lc i ⋆⟨ SliceCat ⟩ kc'≡kc i + + -- ret + + r' = (sIso .fun) (sIso .inv is) .ret + r'≡r : PathP (λ i → kc'≡kc i ⋆⟨ SliceCat ⟩ lc'≡lc i ≡ SliceCat .id) r' r + r'≡r = isSetHomP1 {C = SliceCat} _ _ λ i → kc'≡kc i ⋆⟨ SliceCat ⟩ lc'≡lc i + + sIso .leftInv p = p'≡p + -- to show that the round trip is equivalent to the identity + -- we show that this is true for each component (S-ob, S-arr) + -- and then combine + -- specifically, we show that p'Ob≡pOb and p'Mor≡pMor + -- and it follows that p'≡p + where + p' = (sIso .inv) (sIso .fun p) + + pOb : x ≡ y + pOb i = (p i) .S-ob + + p'Ob : x ≡ y + p'Ob i = (p' i) .S-ob + + + pMor : PathP (λ i → C [ pOb i , c ]) f g + pMor i = (p i) .S-arr + + p'Mor : PathP (λ i → C [ p'Ob i , c ]) f g + p'Mor i = (p' i) .S-arr + + -- we first show that it's equivalent to use sIso first then extract, or to extract first than use pToIIso + extractCom : extractIso' (sIso .fun p) ≡ pToIIso .fun pOb + extractCom = J (λ yg' p̃ → extractIso' (pathToIso p̃) ≡ pToIIso' {xf = xf} {yg'} .fun (λ i → (p̃ i) .S-ob)) + (cong extractIso' (JRefl pToIFam' pToIBase') ∙ sym (JRefl pToIFam pToIBase)) + p + where + idx = C .id + pToIFam = (λ z _ → CatIso C x z) + pToIBase = catiso (C .id) idx (C .⋆IdL idx) (C .⋆IdL idx) + + idxf = SliceCat .id + pToIFam' = (λ z _ → CatIso SliceCat xf z) + pToIBase' = catiso (SliceCat .id) idxf (SliceCat .⋆IdL idxf) (SliceCat .⋆IdL idxf) + + -- why does this not follow definitionally? + -- from extractCom, we get that performing the roundtrip on pOb gives us back p'Ob + ppp : p'Ob ≡ (pToIIso .inv) (pToIIso .fun pOb) + ppp = cong (pToIIso .inv) extractCom + + -- apply univalence of C + -- this gives us the first component that we want + p'Ob≡pOb : p'Ob ≡ pOb + p'Ob≡pOb = ppp ∙ pToIIso .leftInv pOb + + -- isSetHom gives us the second component, path between morphisms + p'Mor≡pMor : PathP (λ j → PathP (λ i → C [ (p'Ob≡pOb j) i , c ]) f g) p'Mor pMor + p'Mor≡pMor = isSetHomP2l {C = C} _ _ p'Mor pMor p'Ob≡pOb + + -- we can use the above paths to show that p' ≡ p + p'≡p : p' ≡ p + p'≡p i = comp (λ i' → SOPath≡PathΣ {xf = xf} {yg} (~ i')) + (λ j → λ { (i = i0) → left (~ j) ; (i = i1) → right (~ j) }) + (p'Σ≡pΣ i) + where + -- we break up p' and p into their constituent paths + -- first via transport and then via our component definitions from before + -- we show that p'ΣT ≡ p'Σ (and same for p) via univalence + -- and p'Σ≡pΣ follows from our work from above + p'ΣT : Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g + p'ΣT = transport SOPath≡PathΣ p' + p'Σ : Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g + p'Σ = (p'Ob , p'Mor) + + pΣT : Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g + pΣT = transport SOPath≡PathΣ p + pΣ : Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g + pΣ = (pOb , pMor)-- transport SOPathP≡PathPSO p + + -- using the computation rule to ua + p'ΣT≡p'Σ : p'ΣT ≡ p'Σ + p'ΣT≡p'Σ = uaβ SOPath≃PathΣ p' + + pΣT≡pΣ : pΣT ≡ pΣ + pΣT≡pΣ = uaβ SOPath≃PathΣ p + + p'Σ≡pΣ : p'Σ ≡ pΣ + p'Σ≡pΣ = ΣPathP (p'Ob≡pOb , p'Mor≡pMor) + + -- two sides of the square we're connecting + left : PathP (λ i → SOPath≡PathΣ {xf = xf} {yg} i) p' p'Σ + left = transport-filler SOPath≡PathΣ p' ▷ p'ΣT≡p'Σ + + right : PathP (λ i → SOPath≡PathΣ {xf = xf} {yg} i) p pΣ + right = transport-filler SOPath≡PathΣ p ▷ pΣT≡pΣ + +-- properties +-- TODO: move to own file + +open isIsoC renaming (inv to invC) + +-- make a slice isomorphism from just the hom +sliceIso : ∀ {a b} (f : C [ a .S-ob , b .S-ob ]) (c : (f ⋆⟨ C ⟩ b .S-arr) ≡ a .S-arr) + → isIsoC C f + → isIsoC SliceCat (slicehom f c) +sliceIso f c isof .invC = slicehom (isof .invC) (sym (invMoveL (isIso→areInv isof) c)) +sliceIso f c isof .sec = SliceHom-≡-intro' (isof .sec) +sliceIso f c isof .ret = SliceHom-≡-intro' (isof .ret) diff --git a/Cubical/Categories/DistLatticeSheaf.agda b/Cubical/Categories/DistLatticeSheaf.agda new file mode 100644 index 0000000000..a02392aabe --- /dev/null +++ b/Cubical/Categories/DistLatticeSheaf.agda @@ -0,0 +1,230 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.DistLatticeSheaf where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Powerset +open import Cubical.Data.Sigma + +open import Cubical.Relation.Binary.Poset + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Semilattice +open import Cubical.Algebra.Lattice +open import Cubical.Algebra.DistLattice +open import Cubical.Algebra.DistLattice.Basis + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Limits.Terminal +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.Instances.CommRings +open import Cubical.Categories.Instances.Poset +open import Cubical.Categories.Instances.Semilattice +open import Cubical.Categories.Instances.Lattice +open import Cubical.Categories.Instances.DistLattice + +private + variable + ℓ ℓ' ℓ'' : Level + +module _ (L : DistLattice ℓ) (C : Category ℓ' ℓ'') (T : Terminal C) where + open Category hiding (_⋆_) + open Functor + open Order (DistLattice→Lattice L) + open DistLatticeStr (snd L) + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice L)) + open MeetSemilattice (Lattice→MeetSemilattice (DistLattice→Lattice L)) + using (∧≤RCancel ; ∧≤LCancel) + open PosetStr (IndPoset .snd) hiding (_≤_) + + 𝟙 : ob C + 𝟙 = terminalOb C T + + DLCat : Category ℓ ℓ + DLCat = DistLatticeCategory L + + open Category DLCat + + -- C-valued presheaves on a distributive lattice + DLPreSheaf : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + DLPreSheaf = Functor (DLCat ^op) C + + hom-∨₁ : (x y : L .fst) → DLCat [ x , x ∨l y ] + hom-∨₁ = ∨≤RCancel + -- TODO: isn't the fixity of the operators a bit weird? + + hom-∨₂ : (x y : L .fst) → DLCat [ y , x ∨l y ] + hom-∨₂ = ∨≤LCancel + + hom-∧₁ : (x y : L .fst) → DLCat [ x ∧l y , x ] + hom-∧₁ _ _ = (≤m→≤j _ _ (∧≤RCancel _ _)) + + hom-∧₂ : (x y : L .fst) → DLCat [ x ∧l y , y ] + hom-∧₂ _ _ = (≤m→≤j _ _ (∧≤LCancel _ _)) + + + {- + x ∧ y ----→ y + | | + | sq | + V V + x ----→ x ∨ y + -} + sq : (x y : L .fst) → hom-∧₂ x y ⋆ hom-∨₂ x y ≡ hom-∧₁ x y ⋆ hom-∨₁ x y + sq x y = is-prop-valued (x ∧l y) (x ∨l y) (hom-∧₂ x y ⋆ hom-∨₂ x y) (hom-∧₁ x y ⋆ hom-∨₁ x y) + + {- + F(x ∨ y) ----→ F(y) + | | + | Fsq | + V V + F(x) ------→ F(x ∧ y) + -} + Fsq : (F : DLPreSheaf) (x y : L .fst) + → F .F-hom (hom-∨₂ x y) ⋆⟨ C ⟩ F .F-hom (hom-∧₂ x y) ≡ + F .F-hom (hom-∨₁ x y) ⋆⟨ C ⟩ F .F-hom (hom-∧₁ x y) + Fsq F x y = sym (F-seq F (hom-∨₂ x y) (hom-∧₂ x y)) + ∙∙ cong (F .F-hom) (sq x y) + ∙∙ F-seq F (hom-∨₁ x y) (hom-∧₁ x y) + + isDLSheaf : (F : DLPreSheaf) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isDLSheaf F = (F-ob F 0l ≡ 𝟙) + × ((x y : L .fst) → isPullback C _ _ _ (Fsq F x y)) + + -- TODO: might be better to define this as a record + DLSheaf : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + DLSheaf = Σ[ F ∈ DLPreSheaf ] isDLSheaf F + + +module SheafOnBasis (L : DistLattice ℓ) (C : Category ℓ' ℓ'') (T : Terminal C) + (L' : ℙ (fst L)) (hB : IsBasis L L') where + + open Category + open Functor + + open DistLatticeStr ⦃...⦄ + open SemilatticeStr ⦃...⦄ + open PosetStr ⦃...⦄ hiding (_≤_) + open IsBasis hB + + private + BasisCat = MeetSemilatticeCategory (Basis→MeetSemilattice L L' hB) + DLBasisPreSheaf = Functor (BasisCat ^op) C + + -- to avoid writing 𝟙 L C T + 1c : ob C + 1c = terminalOb C T + + instance + _ = snd L + _ = snd (Basis→MeetSemilattice L L' hB) + + + module condSquare (x y : ob BasisCat) (x∨y∈L' : fst x ∨l fst y ∈ L') where + + open MeetSemilattice (Lattice→MeetSemilattice (DistLattice→Lattice L)) + using (∧≤RCancel ; ∧≤LCancel) + open MeetSemilattice (Basis→MeetSemilattice L L' hB) + using (IndPoset) + + instance + _ = snd IndPoset + + x∨y : ob BasisCat -- = Σ[ x ∈ L ] (x ∈ L') + x∨y = fst x ∨l fst y , x∨y∈L' + + Bhom-∨₁ : BasisCat [ x , x∨y ] + Bhom-∨₁ = Σ≡Prop (λ _ → L' _ .snd) (∧lAbsorb∨l _ _) + + Bhom-∨₂ : BasisCat [ y , x∨y ] + Bhom-∨₂ = Σ≡Prop (λ _ → L' _ .snd) (cong (fst y ∧l_) (∨lComm _ _) ∙ ∧lAbsorb∨l _ _) + + Bhom-∧₁ : BasisCat [ x · y , x ] + Bhom-∧₁ = Σ≡Prop (λ _ → L' _ .snd) (∧≤RCancel _ _) + + Bhom-∧₂ : BasisCat [ x · y , y ] + Bhom-∧₂ = Σ≡Prop (λ _ → L' _ .snd) (∧≤LCancel _ _) + + {- + x ∧ y ----→ y + | | + | sq | + V V + x ----→ x ∨ y + -} + Bsq : Bhom-∧₂ ⋆⟨ BasisCat ⟩ Bhom-∨₂ ≡ Bhom-∧₁ ⋆⟨ BasisCat ⟩ Bhom-∨₁ + Bsq = is-prop-valued (x · y) x∨y (Bhom-∧₂ ⋆⟨ BasisCat ⟩ Bhom-∨₂) (Bhom-∧₁ ⋆⟨ BasisCat ⟩ Bhom-∨₁) + + {- + F(x ∨ y) ----→ F(y) + | | + | Fsq | + V V + F(x) ------→ F(x ∧ y) + -} + BFsq : (F : DLBasisPreSheaf) + → F .F-hom Bhom-∨₂ ⋆⟨ C ⟩ F .F-hom Bhom-∧₂ ≡ + F .F-hom Bhom-∨₁ ⋆⟨ C ⟩ F .F-hom Bhom-∧₁ + BFsq F = sym (F-seq F Bhom-∨₂ Bhom-∧₂) + ∙∙ cong (F .F-hom) Bsq + ∙∙ F-seq F Bhom-∨₁ Bhom-∧₁ + + + -- TODO: check that this is equivalent to the functor + -- preserving terminal objects and pullbacks + isDLBasisSheaf : DLBasisPreSheaf → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isDLBasisSheaf F = ((0∈L' : 0l ∈ L') → F .F-ob (0l , 0∈L') ≡ 1c) + × ((x y : ob BasisCat) (x∨y∈L' : fst x ∨l fst y ∈ L') + → isPullback C _ _ _ (BFsq x y x∨y∈L' F)) + where + open condSquare + + DLBasisSheaf : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + DLBasisSheaf = Σ[ F ∈ DLBasisPreSheaf ] isDLBasisSheaf F + + -- To prove the statement we probably need that C is: + -- 1. univalent + -- 2. has finite limits (or pullbacks and a terminal object) + -- 3. isGroupoid (C .ob) + -- The last point is not strictly necessary, but we have to have some + -- control over the hLevel as we want to write F(x) in terms of its + -- basis cover which is information hidden under a prop truncation... + -- Alternatively we just prove the statement for C = CommRingsCategory + + -- TODO: is unique existence expressed like this what we want? + -- statement : (F' : DLBasisSheaf) + -- → ∃![ F ∈ DLSheaf L C T ] ((x : fst L) → (x ∈ L') → CatIso C (F-ob (fst F) x) (F-ob (fst F') x)) -- TODO: if C is univalent the CatIso could be ≡? + -- statement (F' , h1 , hPb) = ? + + -- It might be easier to prove all of these if we use the definition + -- in terms of particular limits instead + + + + + + -- Scrap zone: + + -- -- Sublattices: upstream later + -- record isSublattice (L' : ℙ (fst L)) : Type ℓ where + -- field + -- 1l-closed : 1l ∈ L' + -- 0l-closed : 0l ∈ L' + -- ∧l-closed : {x y : fst L} → x ∈ L' → y ∈ L' → x ∧l y ∈ L' + -- ∨l-closed : {x y : fst L} → x ∈ L' → y ∈ L' → x ∨l y ∈ L' + + -- open isSublattice + + -- Sublattice : Type (ℓ-suc ℓ) + -- Sublattice = Σ[ L' ∈ ℙ (fst L) ] isSublattice L' + + -- restrictDLSheaf : DLSheaf → Sublattice → DLSheaf + -- F-ob (fst (restrictDLSheaf F (L' , HL'))) x = {!F-ob (fst F) x!} -- Hmm, not nice... + -- F-hom (fst (restrictDLSheaf F L')) = {!!} + -- F-id (fst (restrictDLSheaf F L')) = {!!} + -- F-seq (fst (restrictDLSheaf F L')) = {!!} + -- snd (restrictDLSheaf F L') = {!!} diff --git a/Cubical/Categories/Equivalence.agda b/Cubical/Categories/Equivalence.agda new file mode 100644 index 0000000000..cea02fe388 --- /dev/null +++ b/Cubical/Categories/Equivalence.agda @@ -0,0 +1,7 @@ + +{-# OPTIONS --safe #-} + +module Cubical.Categories.Equivalence where + +open import Cubical.Categories.Equivalence.Base public +open import Cubical.Categories.Equivalence.Properties public diff --git a/Cubical/Categories/Equivalence/Base.agda b/Cubical/Categories/Equivalence/Base.agda new file mode 100644 index 0000000000..e5d8255aa8 --- /dev/null +++ b/Cubical/Categories/Equivalence/Base.agda @@ -0,0 +1,30 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Equivalence.Base where + +open import Cubical.Foundations.Prelude + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation + +open Category +open Functor + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +record isEquivalence {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} + (func : Functor C D) : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + invFunc : Functor D C + + η : 𝟙⟨ C ⟩ ≅ᶜ invFunc ∘F func + ε : func ∘F invFunc ≅ᶜ 𝟙⟨ D ⟩ + +record _≃ᶜ_ (C : Category ℓC ℓC') (D : Category ℓD ℓD') : + Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + func : Functor C D + isEquiv : isEquivalence func + diff --git a/Cubical/Categories/Equivalence/Properties.agda b/Cubical/Categories/Equivalence/Properties.agda new file mode 100644 index 0000000000..6e05df1276 --- /dev/null +++ b/Cubical/Categories/Equivalence/Properties.agda @@ -0,0 +1,92 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Equivalence.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Morphism +open import Cubical.Categories.Equivalence.Base +open import Cubical.HITs.PropositionalTruncation.Base + +open Category +open Functor +open NatIso +open CatIso +open NatTrans +open isEquivalence + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +-- Equivalence implies Full, Faithul, and Essentially Surjective + +module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where + symEquiv : ∀ {F : Functor C D} → (e : isEquivalence F) → isEquivalence (e .invFunc) + symEquiv {F} record { invFunc = G ; η = η ; ε = ε } = record { invFunc = F ; η = symNatIso ε ; ε = symNatIso η } + + isEquiv→Faithful : ∀ {F : Functor C D} → isEquivalence F → isFaithful F + isEquiv→Faithful {F} record { invFunc = G + ; η = η + ; ε = _ } + c c' f g p = f + ≡⟨ sqRL η ⟩ + cIso .mor ⋆⟨ C ⟩ G ⟪ F ⟪ f ⟫ ⟫ ⋆⟨ C ⟩ c'Iso .inv + ≡⟨ cong (λ v → cIso .mor ⋆⟨ C ⟩ (G ⟪ v ⟫) ⋆⟨ C ⟩ c'Iso .inv) p ⟩ + cIso .mor ⋆⟨ C ⟩ G ⟪ F ⟪ g ⟫ ⟫ ⋆⟨ C ⟩ c'Iso .inv + ≡⟨ sym (sqRL η) ⟩ + g + ∎ + where + + -- isomorphism between c and GFc + cIso = isIso→CatIso (η .nIso c) + -- isomorphism between c' and GFc' + c'Iso = isIso→CatIso (η .nIso c') + +module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where + isEquiv→Full : ∀ {F : Functor C D} → isEquivalence F → isFull F + isEquiv→Full {F} eq@record { invFunc = G + ; η = η + ; ε = _ } + c c' g = ∣ h , isEquiv→Faithful (symEquiv eq) _ _ _ _ GFh≡Gg ∣ -- apply faithfulness of G + where + -- isomorphism between c and GFc + cIso = isIso→CatIso (η .nIso c) + -- isomorphism between c' and GFc' + c'Iso = isIso→CatIso (η .nIso c') + + -- reverses + cIso⁻ = symCatIso cIso + c'Iso⁻ = symCatIso c'Iso + + h = cIso .mor ⋆⟨ C ⟩ G ⟪ g ⟫ ⋆⟨ C ⟩ c'Iso .inv + + -- we show that both `G ⟪ g ⟫` and `G ⟪ F ⟪ h ⟫ ⟫` + -- are equal to the same thing + -- namely : cIso .inv ⋆⟨ C ⟩ h ⋆⟨ C ⟩ c'Iso .mor + Gg≡ηhη : G ⟪ g ⟫ ≡ cIso .inv ⋆⟨ C ⟩ h ⋆⟨ C ⟩ c'Iso .mor + Gg≡ηhη = invMoveL cAreInv move-c' ∙ sym (C .⋆Assoc _ _ _) + where + cAreInv : areInv _ (cIso .mor) (cIso .inv) + cAreInv = CatIso→areInv cIso + + c'AreInv : areInv _ (c'Iso .mor) (c'Iso .inv) + c'AreInv = CatIso→areInv c'Iso + + move-c' : cIso .mor ⋆⟨ C ⟩ G ⟪ g ⟫ ≡ h ⋆⟨ C ⟩ c'Iso .mor + move-c' = invMoveR (symAreInv c'AreInv) refl + + GFh≡Gg : G ⟪ F ⟪ h ⟫ ⟫ ≡ G ⟪ g ⟫ + GFh≡Gg = G ⟪ F ⟪ h ⟫ ⟫ + ≡⟨ sqLR η ⟩ + cIso .inv ⋆⟨ C ⟩ h ⋆⟨ C ⟩ c'Iso .mor + ≡⟨ sym Gg≡ηhη ⟩ + G ⟪ g ⟫ + ∎ + + isEquiv→Surj : ∀ {F : Functor C D} → isEquivalence F → isEssentiallySurj F + isEquiv→Surj isE d = (isE .invFunc ⟅ d ⟆) , isIso→CatIso ((isE .ε .nIso) d) diff --git a/Cubical/Categories/Functor.agda b/Cubical/Categories/Functor.agda new file mode 100644 index 0000000000..9b1e7a30bd --- /dev/null +++ b/Cubical/Categories/Functor.agda @@ -0,0 +1,8 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Functor where + +open import Cubical.Categories.Functor.Base public +open import Cubical.Categories.Functor.BinProduct public +open import Cubical.Categories.Functor.Compose public +open import Cubical.Categories.Functor.Properties public diff --git a/Cubical/Categories/Functor/Base.agda b/Cubical/Categories/Functor/Base.agda new file mode 100644 index 0000000000..871dd6d72e --- /dev/null +++ b/Cubical/Categories/Functor/Base.agda @@ -0,0 +1,79 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Functor.Base where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Sigma + +open import Cubical.Categories.Category + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +record Functor (C : Category ℓC ℓC') (D : Category ℓD ℓD') : + Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + no-eta-equality + + open Category + + field + F-ob : C .ob → D .ob + F-hom : {x y : C .ob} → C [ x , y ] → D [ F-ob x , F-ob y ] + F-id : {x : C .ob} → F-hom (C .id) ≡ D .id {x = F-ob x} + F-seq : {x y z : C .ob} (f : C [ x , y ]) (g : C [ y , z ]) + → F-hom (f ⋆⟨ C ⟩ g) ≡ (F-hom f) ⋆⟨ D ⟩ (F-hom g) + + isFull = (x y : _) (F[f] : D [ F-ob x , F-ob y ]) → ∃[ f ∈ C [ x , y ] ] F-hom f ≡ F[f] + isFaithful = (x y : _) (f g : C [ x , y ]) → F-hom f ≡ F-hom g → f ≡ g + isEssentiallySurj = (d : D .ob) → Σ[ c ∈ C .ob ] CatIso D (F-ob c) d + +private + variable + ℓ ℓ' : Level + C D E : Category ℓ ℓ' + +open Category +open Functor + +-- Helpful notation + +-- action on objects +infix 30 _⟅_⟆ +_⟅_⟆ : (F : Functor C D) + → C .ob + → D .ob +_⟅_⟆ = F-ob + +-- action on morphisms +infix 30 _⟪_⟫ -- same infix level as on objects since these will never be used in the same context +_⟪_⟫ : (F : Functor C D) + → ∀ {x y} + → C [ x , y ] + → D [(F ⟅ x ⟆) , (F ⟅ y ⟆)] +_⟪_⟫ = F-hom + + +-- Functor constructions + +𝟙⟨_⟩ : ∀ (C : Category ℓ ℓ') → Functor C C +𝟙⟨ C ⟩ .F-ob x = x +𝟙⟨ C ⟩ .F-hom f = f +𝟙⟨ C ⟩ .F-id = refl +𝟙⟨ C ⟩ .F-seq _ _ = refl + +-- functor composition +funcComp : ∀ (G : Functor D E) (F : Functor C D) → Functor C E +(funcComp G F) .F-ob c = G ⟅ F ⟅ c ⟆ ⟆ +(funcComp G F) .F-hom f = G ⟪ F ⟪ f ⟫ ⟫ +(funcComp G F) .F-id = cong (G ⟪_⟫) (F .F-id) ∙ G .F-id +(funcComp G F) .F-seq f g = cong (G ⟪_⟫) (F .F-seq _ _) ∙ G .F-seq _ _ + +infixr 30 funcComp +syntax funcComp G F = G ∘F F + +_^opF : Functor C D → Functor (C ^op) (D ^op) +(F ^opF) .F-ob = F .F-ob +(F ^opF) .F-hom = F .F-hom +(F ^opF) .F-id = F .F-id +(F ^opF) .F-seq f g = F .F-seq g f diff --git a/Cubical/Categories/Functor/BinProduct.agda b/Cubical/Categories/Functor/BinProduct.agda new file mode 100644 index 0000000000..8b15074431 --- /dev/null +++ b/Cubical/Categories/Functor/BinProduct.agda @@ -0,0 +1,26 @@ +-- Product of two functors +{-# OPTIONS --safe #-} + +module Cubical.Categories.Functor.BinProduct where + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Constructions.BinProduct +open import Cubical.Categories.Functor.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Foundations.Prelude + +private + variable + ℓA ℓA' ℓB ℓB' ℓC ℓC' ℓD ℓD' : Level + A : Category ℓA ℓA' + B : Category ℓB ℓB' + C : Category ℓC ℓC' + D : Category ℓD ℓD' + +open Functor + +_×F_ : Functor A C → Functor B D → Functor (A × B) (C × D) +(G ×F H) .F-ob (a , b) = (G ⟅ a ⟆ , H ⟅ b ⟆) +(G ×F H) .F-hom (g , h) = (G ⟪ g ⟫ , H ⟪ h ⟫) +(G ×F H) .F-id = ≡-× (G .F-id) (H .F-id) +(G ×F H) .F-seq _ _ = ≡-× (G .F-seq _ _) (H .F-seq _ _) diff --git a/Cubical/Categories/Functor/Compose.agda b/Cubical/Categories/Functor/Compose.agda new file mode 100644 index 0000000000..d413a8f575 --- /dev/null +++ b/Cubical/Categories/Functor/Compose.agda @@ -0,0 +1,44 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Functor.Compose where + +open import Cubical.Foundations.Prelude + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.NaturalTransformation.Base + +open import Cubical.Categories.Instances.Functors + +module _ {ℓC ℓC' ℓD ℓD' ℓE ℓE'} + {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (E : Category ℓE ℓE') + (F : Functor C D) + where + + open Functor + open NatTrans + + precomposeF : Functor (FUNCTOR D E) (FUNCTOR C E) + precomposeF .F-ob G = funcComp G F + precomposeF .F-hom α .N-ob c = α .N-ob (F .F-ob c) + precomposeF .F-hom α .N-hom f = α .N-hom (F .F-hom f) + precomposeF .F-id = refl + precomposeF .F-seq f g = refl + +module _ {ℓC ℓC' ℓD ℓD' ℓE ℓE'} + (C : Category ℓC ℓC') {D : Category ℓD ℓD'} {E : Category ℓE ℓE'} + (G : Functor D E) + where + + open Functor + open NatTrans + + postcomposeF : Functor (FUNCTOR C D) (FUNCTOR C E) + postcomposeF .F-ob F = funcComp G F + postcomposeF .F-hom α .N-ob c = G. F-hom (α .N-ob c) + postcomposeF .F-hom {F₀} {F₁} α .N-hom {x} {y} f = + sym (G .F-seq (F₀ ⟪ f ⟫) (α ⟦ y ⟧)) + ∙∙ cong (G ⟪_⟫) (α .N-hom f) + ∙∙ G .F-seq (α ⟦ x ⟧) (F₁ ⟪ f ⟫) + postcomposeF .F-id = makeNatTransPath (funExt λ _ → G .F-id) + postcomposeF .F-seq f g = makeNatTransPath (funExt λ _ → G .F-seq _ _) diff --git a/Cubical/Categories/Functor/Product.agda b/Cubical/Categories/Functor/Product.agda new file mode 100644 index 0000000000..85c2272ee7 --- /dev/null +++ b/Cubical/Categories/Functor/Product.agda @@ -0,0 +1,26 @@ +-- Product of two functors +{-# OPTIONS --safe #-} + +module Cubical.Categories.Functor.Product where + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Constructions.Product +open import Cubical.Categories.Functor.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Foundations.Prelude + +private + variable + ℓA ℓA' ℓB ℓB' ℓC ℓC' ℓD ℓD' : Level + A : Category ℓA ℓA' + B : Category ℓB ℓB' + C : Category ℓC ℓC' + D : Category ℓD ℓD' + +open Functor + +_×F_ : Functor A C → Functor B D → Functor (A × B) (C × D) +(G ×F H) .F-ob (a , b) = (G ⟅ a ⟆ , H ⟅ b ⟆) +(G ×F H) .F-hom (g , h) = (G ⟪ g ⟫ , H ⟪ h ⟫) +(G ×F H) .F-id = ≡-× (G .F-id) (H .F-id) +(G ×F H) .F-seq _ _ = ≡-× (G .F-seq _ _) (H .F-seq _ _) diff --git a/Cubical/Categories/Functor/Properties.agda b/Cubical/Categories/Functor/Properties.agda new file mode 100644 index 0000000000..18799c47c9 --- /dev/null +++ b/Cubical/Categories/Functor/Properties.agda @@ -0,0 +1,116 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Functor.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function renaming (_∘_ to _◍_) +open import Cubical.Foundations.GroupoidLaws using (lUnit; rUnit; assoc; cong-∙) +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base + +private + variable + ℓ ℓ' ℓ'' : Level + B C D E : Category ℓ ℓ' + +open Category +open Functor + +{- +x ---p--- x' + ⇓ᵍ + g x' ---q--- y + ⇓ʰ + h y ---r--- z + +The path from `h (g x) ≡ z` obtained by + 1. first applying cong to p and composing with q; then applying cong again and composing with r + 2. first applying cong to q and composing with r; then applying a double cong to p and precomposing +are path equal. +-} +congAssoc : ∀ {X : Type ℓ} {Y : Type ℓ'} {Z : Type ℓ''} (g : X → Y) (h : Y → Z) {x x' : X} {y : Y} {z : Z} + → (p : x ≡ x') (q : g x' ≡ y) (r : h y ≡ z) + → cong (h ◍ g) p ∙ (cong h q ∙ r) ≡ cong h (cong g p ∙ q) ∙ r +congAssoc g h p q r + = cong (h ◍ g) p ∙ (cong h q ∙ r) + ≡⟨ assoc _ _ _ ⟩ + ((cong (h ◍ g) p) ∙ (cong h q)) ∙ r + ≡⟨ refl ⟩ + (cong h (cong g p) ∙ (cong h q)) ∙ r + ≡⟨ cong (_∙ r) (sym (cong-∙ h _ _)) ⟩ + cong h (cong g p ∙ q) ∙ r + ∎ + +-- composition is associative +F-assoc : {F : Functor B C} {G : Functor C D} {H : Functor D E} + → H ∘F (G ∘F F) ≡ (H ∘F G) ∘F F +F-assoc {F = F} {G} {H} i .F-ob x = H ⟅ G ⟅ F ⟅ x ⟆ ⟆ ⟆ +F-assoc {F = F} {G} {H} i .F-hom f = H ⟪ G ⟪ F ⟪ f ⟫ ⟫ ⟫ +F-assoc {F = F} {G} {H} i .F-id {x} = congAssoc (G ⟪_⟫) (H ⟪_⟫) (F .F-id {x}) (G .F-id {F ⟅ x ⟆}) (H .F-id) (~ i) +F-assoc {F = F} {G} {H} i .F-seq f g = congAssoc (G ⟪_⟫) (H ⟪_⟫) (F .F-seq f g) (G .F-seq _ _) (H .F-seq _ _) (~ i) + +-- Results about functors + +module _ {F : Functor C D} where + + -- the identity is the identity + F-lUnit : F ∘F 𝟙⟨ C ⟩ ≡ F + F-lUnit i .F-ob x = F ⟅ x ⟆ + F-lUnit i .F-hom f = F ⟪ f ⟫ + F-lUnit i .F-id {x} = lUnit (F .F-id) (~ i) + F-lUnit i .F-seq f g = lUnit (F .F-seq f g) (~ i) + + F-rUnit : 𝟙⟨ D ⟩ ∘F F ≡ F + F-rUnit i .F-ob x = F ⟅ x ⟆ + F-rUnit i .F-hom f = F ⟪ f ⟫ + F-rUnit i .F-id {x} = rUnit (F .F-id) (~ i) + F-rUnit i .F-seq f g = rUnit (F .F-seq f g) (~ i) + + -- functors preserve commutative diagrams (specificallysqures here) + preserveCommF : ∀ {x y z w} {f : C [ x , y ]} {g : C [ y , w ]} {h : C [ x , z ]} {k : C [ z , w ]} + → f ⋆⟨ C ⟩ g ≡ h ⋆⟨ C ⟩ k + → (F ⟪ f ⟫) ⋆⟨ D ⟩ (F ⟪ g ⟫) ≡ (F ⟪ h ⟫) ⋆⟨ D ⟩ (F ⟪ k ⟫) + preserveCommF {f = f} {g = g} {h = h} {k = k} eq + = (F ⟪ f ⟫) ⋆⟨ D ⟩ (F ⟪ g ⟫) + ≡⟨ sym (F .F-seq _ _) ⟩ + F ⟪ f ⋆⟨ C ⟩ g ⟫ + ≡⟨ cong (F ⟪_⟫) eq ⟩ + F ⟪ h ⋆⟨ C ⟩ k ⟫ + ≡⟨ F .F-seq _ _ ⟩ + (F ⟪ h ⟫) ⋆⟨ D ⟩ (F ⟪ k ⟫) + ∎ + + -- functors preserve isomorphisms + preserveIsosF : ∀ {x y} → CatIso C x y → CatIso D (F ⟅ x ⟆) (F ⟅ y ⟆) + preserveIsosF {x} {y} (catiso f f⁻¹ sec' ret') = + catiso + g g⁻¹ + -- sec + ( (g⁻¹ ⋆⟨ D ⟩ g) + ≡⟨ sym (F .F-seq f⁻¹ f) ⟩ + F ⟪ f⁻¹ ⋆⟨ C ⟩ f ⟫ + ≡⟨ cong (F .F-hom) sec' ⟩ + F ⟪ C .id ⟫ + ≡⟨ F .F-id ⟩ + D .id + ∎ ) + -- ret + ( (g ⋆⟨ D ⟩ g⁻¹) + ≡⟨ sym (F .F-seq f f⁻¹) ⟩ + F ⟪ f ⋆⟨ C ⟩ f⁻¹ ⟫ + ≡⟨ cong (F .F-hom) ret' ⟩ + F ⟪ C .id ⟫ + ≡⟨ F .F-id ⟩ + D .id + ∎ ) + + where + x' : D .ob + x' = F ⟅ x ⟆ + y' : D .ob + y' = F ⟅ y ⟆ + + g : D [ x' , y' ] + g = F ⟪ f ⟫ + g⁻¹ : D [ y' , x' ] + g⁻¹ = F ⟪ f⁻¹ ⟫ diff --git a/Cubical/Categories/Instances/AbGroups.agda b/Cubical/Categories/Instances/AbGroups.agda new file mode 100644 index 0000000000..59205e2963 --- /dev/null +++ b/Cubical/Categories/Instances/AbGroups.agda @@ -0,0 +1,23 @@ +-- The category Ab of abelian groups +{-# OPTIONS --safe #-} + +module Cubical.Categories.Instances.AbGroups where + +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Categories.Category.Base +open import Cubical.Foundations.Prelude + + +module _ {ℓ : Level} where + open Category + + AbGroupCategory : Category (ℓ-suc ℓ) ℓ + AbGroupCategory .ob = AbGroup ℓ + AbGroupCategory .Hom[_,_] = AbGroupHom + AbGroupCategory .id = idGroupHom + AbGroupCategory ._⋆_ = compGroupHom + AbGroupCategory .⋆IdL f = GroupHom≡ refl + AbGroupCategory .⋆IdR f = GroupHom≡ refl + AbGroupCategory .⋆Assoc f g h = GroupHom≡ refl + AbGroupCategory .isSetHom = isSetGroupHom diff --git a/Cubical/Categories/Instances/Algebras.agda b/Cubical/Categories/Instances/Algebras.agda new file mode 100644 index 0000000000..4998c436fd --- /dev/null +++ b/Cubical/Categories/Instances/Algebras.agda @@ -0,0 +1,28 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.Algebras where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra + +open import Cubical.Categories.Category + +open Category +open AlgebraHoms + +private + variable + ℓ ℓ' : Level + +module _ (R : Ring ℓ) where + AlgebrasCategory : Category (ℓ-suc (ℓ-max ℓ ℓ')) (ℓ-max ℓ ℓ') + ob AlgebrasCategory = Algebra R _ + Hom[_,_] AlgebrasCategory = AlgebraHom + id AlgebrasCategory {A} = idAlgebraHom A + _⋆_ AlgebrasCategory = compAlgebraHom + ⋆IdL AlgebrasCategory = compIdAlgebraHom + ⋆IdR AlgebrasCategory = idCompAlgebraHom + ⋆Assoc AlgebrasCategory = compAssocAlgebraHom + isSetHom AlgebrasCategory = isSetAlgebraHom _ _ diff --git a/Cubical/Categories/Instances/Categories.agda b/Cubical/Categories/Instances/Categories.agda new file mode 100644 index 0000000000..7ce6b89a2f --- /dev/null +++ b/Cubical/Categories/Instances/Categories.agda @@ -0,0 +1,25 @@ +-- The (pre)category of (small) categories +{-# OPTIONS --safe #-} + +module Cubical.Categories.Instances.Categories where + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Category.Precategory +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Functor.Properties +open import Cubical.Foundations.Prelude + + +module _ (ℓ ℓ' : Level) where + open Precategory + + CatPrecategory : Precategory (ℓ-suc (ℓ-max ℓ ℓ')) (ℓ-max ℓ ℓ') + CatPrecategory .ob = Category ℓ ℓ' + CatPrecategory .Hom[_,_] = Functor + CatPrecategory .id = 𝟙⟨ _ ⟩ + CatPrecategory ._⋆_ G H = H ∘F G + CatPrecategory .⋆IdL _ = F-lUnit + CatPrecategory .⋆IdR _ = F-rUnit + CatPrecategory .⋆Assoc _ _ _ = F-assoc + +-- TODO: what is required for Functor C D to be a set? diff --git a/Cubical/Categories/Instances/CommAlgebras.agda b/Cubical/Categories/Instances/CommAlgebras.agda new file mode 100644 index 0000000000..29e3a8eb9c --- /dev/null +++ b/Cubical/Categories/Instances/CommAlgebras.agda @@ -0,0 +1,62 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.CommAlgebras where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Powerset + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommAlgebra + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base + +open import Cubical.HITs.PropositionalTruncation + +open Category +open CommAlgebraHoms + +private + variable + ℓ ℓ' ℓ'' : Level + +module _ (R : CommRing ℓ) where + CommAlgebrasCategory : Category (ℓ-suc (ℓ-max ℓ ℓ')) (ℓ-max ℓ ℓ') + ob CommAlgebrasCategory = CommAlgebra R _ + Hom[_,_] CommAlgebrasCategory = CommAlgebraHom + id CommAlgebrasCategory {A} = idCommAlgebraHom A + _⋆_ CommAlgebrasCategory {A} {B} {C} = compCommAlgebraHom A B C + ⋆IdL CommAlgebrasCategory {A} {B} = compIdCommAlgebraHom {A = A} {B} + ⋆IdR CommAlgebrasCategory {A} {B} = idCompCommAlgebraHom {A = A} {B} + ⋆Assoc CommAlgebrasCategory {A} {B} {C} {D} = compAssocCommAlgebraHom {A = A} {B} {C} {D} + isSetHom CommAlgebrasCategory = isSetAlgebraHom _ _ + + +module PreSheafFromUniversalProp (C : Category ℓ ℓ') (P : ob C → Type ℓ) + {R : CommRing ℓ''} (𝓕 : Σ (ob C) P → CommAlgebra R ℓ'') + (uniqueHom : ∀ x y → C [ fst x , fst y ] → isContr (CommAlgebraHom (𝓕 y) (𝓕 x))) + where + + private + ∥P∥ : ℙ (ob C) + ∥P∥ x = ∥ P x ∥ , isPropPropTrunc + ΣC∥P∥Cat = ΣPropCat C ∥P∥ + + 𝓕UniqueEquiv : (x : ob C) (p q : P x) → isContr (CommAlgebraEquiv (𝓕 (x , p)) (𝓕 (x , q))) + 𝓕UniqueEquiv x = contrCommAlgebraHom→contrCommAlgebraEquiv (curry 𝓕 x) λ p q → uniqueHom _ _ (id C) + + theMap : (x : ob C) → ∥ P x ∥ → CommAlgebra R ℓ'' + theMap x = recPT→CommAlgebra (curry 𝓕 x) (λ p q → 𝓕UniqueEquiv x p q .fst) + λ p q r → 𝓕UniqueEquiv x p r .snd _ + + theAction : (x y : ob C) → C [ x , y ] + → (p : ∥ P x ∥) (q : ∥ P y ∥) → isContr (CommAlgebraHom (theMap y q) (theMap x p)) + theAction _ _ f = elim2 (λ _ _ → isPropIsContr) λ _ _ → uniqueHom _ _ f + + open Functor + universalPShf : Functor (ΣC∥P∥Cat ^op) (CommAlgebrasCategory {ℓ = ℓ''} R {ℓ' = ℓ''}) + F-ob universalPShf = uncurry theMap + F-hom universalPShf {x = x} {y = y} f = theAction _ _ f (y .snd) (x. snd) .fst + F-id universalPShf {x = x} = theAction (x .fst) (x .fst) (id C) (x .snd) (x .snd) .snd _ + F-seq universalPShf {x = x} {z = z} f g = theAction _ _ (g ⋆⟨ C ⟩ f) (z .snd) (x .snd) .snd _ diff --git a/Cubical/Categories/Instances/CommRings.agda b/Cubical/Categories/Instances/CommRings.agda new file mode 100644 index 0000000000..5b6d39ef25 --- /dev/null +++ b/Cubical/Categories/Instances/CommRings.agda @@ -0,0 +1,66 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.CommRings where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Powerset + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.FiberedProduct +open import Cubical.Algebra.CommRing.Instances.Unit + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.Limits.Terminal +open import Cubical.Categories.Limits.Pullback + +open import Cubical.HITs.PropositionalTruncation + +open Category hiding (_∘_) +open CommRingHoms + +private + variable + ℓ ℓ' ℓ'' : Level + +CommRingsCategory : Category (ℓ-suc ℓ) ℓ +ob CommRingsCategory = CommRing _ +Hom[_,_] CommRingsCategory = CommRingHom +id CommRingsCategory {R} = idCommRingHom R +_⋆_ CommRingsCategory {R} {S} {T} = compCommRingHom R S T +⋆IdL CommRingsCategory {R} {S} = compIdCommRingHom {R = R} {S} +⋆IdR CommRingsCategory {R} {S} = idCompCommRingHom {R = R} {S} +⋆Assoc CommRingsCategory {R} {S} {T} {U} = compAssocCommRingHom {R = R} {S} {T} {U} +isSetHom CommRingsCategory = isSetRingHom _ _ + +TerminalCommRing : Terminal {ℓ-suc ℓ-zero} CommRingsCategory +fst TerminalCommRing = UnitCommRing +fst (fst (snd TerminalCommRing y)) _ = tt +snd (fst (snd TerminalCommRing y)) = makeIsRingHom refl (λ _ _ → refl) (λ _ _ → refl) +snd (snd TerminalCommRing y) f = RingHom≡ (funExt (λ _ → refl)) + + +open Pullback + +{- + A x_C B -----> A + | | + | | α + | | + V V + B --------> C + β +-} +PullbackCommRing : Pullbacks {ℓ-suc ℓ} CommRingsCategory +pbOb (PullbackCommRing (cospan A C B α β)) = fiberedProduct A B C α β +pbPr₁ (PullbackCommRing (cospan A C B α β)) = fiberedProductPr₁ A B C α β +pbPr₂ (PullbackCommRing (cospan A C B α β)) = fiberedProductPr₂ A B C α β +pbCommutes (PullbackCommRing (cospan A C B α β)) = fiberedProductPr₁₂Commutes A B C α β +univProp (PullbackCommRing (cospan A C B α β)) {d = D} = fiberedProductUnivProp A B C α β D diff --git a/Cubical/Categories/Instances/Cospan.agda b/Cubical/Categories/Instances/Cospan.agda new file mode 100644 index 0000000000..d108b1f692 --- /dev/null +++ b/Cubical/Categories/Instances/Cospan.agda @@ -0,0 +1,77 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.Cospan where + +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Category +open import Cubical.Data.Unit +open import Cubical.Data.Empty + +open Category + +data 𝟛 : Type ℓ-zero where + ⓪ : 𝟛 + ① : 𝟛 + ② : 𝟛 + +CospanCat : Category ℓ-zero ℓ-zero +CospanCat .ob = 𝟛 + +CospanCat .Hom[_,_] ⓪ ① = Unit +CospanCat .Hom[_,_] ② ① = Unit +CospanCat .Hom[_,_] ⓪ ⓪ = Unit +CospanCat .Hom[_,_] ① ① = Unit +CospanCat .Hom[_,_] ② ② = Unit +CospanCat .Hom[_,_] _ _ = ⊥ + +CospanCat ._⋆_ {x = ⓪} {⓪} {⓪} f g = tt +CospanCat ._⋆_ {x = ①} {①} {①} f g = tt +CospanCat ._⋆_ {x = ②} {②} {②} f g = tt +CospanCat ._⋆_ {x = ⓪} {①} {①} f g = tt +CospanCat ._⋆_ {x = ②} {①} {①} f g = tt +CospanCat ._⋆_ {x = ⓪} {⓪} {①} f g = tt +CospanCat ._⋆_ {x = ②} {②} {①} f g = tt + +CospanCat .id {⓪} = tt +CospanCat .id {①} = tt +CospanCat .id {②} = tt + +CospanCat .⋆IdL {⓪} {①} _ = refl +CospanCat .⋆IdL {②} {①} _ = refl +CospanCat .⋆IdL {⓪} {⓪} _ = refl +CospanCat .⋆IdL {①} {①} _ = refl +CospanCat .⋆IdL {②} {②} _ = refl + +CospanCat .⋆IdR {⓪} {①} _ = refl +CospanCat .⋆IdR {②} {①} _ = refl +CospanCat .⋆IdR {⓪} {⓪} _ = refl +CospanCat .⋆IdR {①} {①} _ = refl +CospanCat .⋆IdR {②} {②} _ = refl + +CospanCat .⋆Assoc {⓪} {⓪} {⓪} {⓪} _ _ _ = refl +CospanCat .⋆Assoc {⓪} {⓪} {⓪} {①} _ _ _ = refl +CospanCat .⋆Assoc {⓪} {⓪} {①} {①} _ _ _ = refl +CospanCat .⋆Assoc {⓪} {①} {①} {①} _ _ _ = refl +CospanCat .⋆Assoc {①} {①} {①} {①} _ _ _ = refl +CospanCat .⋆Assoc {②} {②} {②} {②} _ _ _ = refl +CospanCat .⋆Assoc {②} {②} {②} {①} _ _ _ = refl +CospanCat .⋆Assoc {②} {②} {①} {①} _ _ _ = refl +CospanCat .⋆Assoc {②} {①} {①} {①} _ _ _ = refl + +CospanCat .isSetHom {⓪} {⓪} = isSetUnit +CospanCat .isSetHom {⓪} {①} = isSetUnit +CospanCat .isSetHom {①} {①} = isSetUnit +CospanCat .isSetHom {②} {①} = isSetUnit +CospanCat .isSetHom {②} {②} = isSetUnit + + +-- makes it easier to write functors into CospanCat +isPropHomCospanCat : (x y : ob CospanCat) → isProp (CospanCat [ x , y ]) +isPropHomCospanCat ⓪ ⓪ = isPropUnit +isPropHomCospanCat ⓪ ① = isPropUnit +isPropHomCospanCat ⓪ ② = isProp⊥ +isPropHomCospanCat ① ⓪ = isProp⊥ +isPropHomCospanCat ① ① = isPropUnit +isPropHomCospanCat ① ② = isProp⊥ +isPropHomCospanCat ② ⓪ = isProp⊥ +isPropHomCospanCat ② ① = isPropUnit +isPropHomCospanCat ② ② = isPropUnit diff --git a/Cubical/Categories/Instances/Discrete.agda b/Cubical/Categories/Instances/Discrete.agda new file mode 100644 index 0000000000..8fff2527cb --- /dev/null +++ b/Cubical/Categories/Instances/Discrete.agda @@ -0,0 +1,55 @@ +-- Discrete category over a type A +-- A must be an h-groupoid for the homs to be sets +{-# OPTIONS --safe #-} + +module Cubical.Categories.Instances.Discrete where + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor.Base +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Transport + +private + variable + ℓ ℓC ℓC' : Level + +open Category + +-- Discrete category +DiscreteCategory : hGroupoid ℓ → Category ℓ ℓ +DiscreteCategory A .ob = A .fst +DiscreteCategory A .Hom[_,_] a a' = a ≡ a' +DiscreteCategory A .id = refl +DiscreteCategory A ._⋆_ = _∙_ +DiscreteCategory A .⋆IdL f = sym (lUnit f) +DiscreteCategory A .⋆IdR f = sym (rUnit f) +DiscreteCategory A .⋆Assoc f g h = sym (assoc f g h) +DiscreteCategory A .isSetHom {x} {y} = A .snd x y + + +module _ {A : hGroupoid ℓ} + {C : Category ℓC ℓC'} where + open Functor + + -- Functions f: A → ob C give functors F: DiscreteCategory A → C + DiscFunc : (fst A → ob C) → Functor (DiscreteCategory A) C + DiscFunc f .F-ob = f + DiscFunc f .F-hom {x} p = subst (λ z → C [ f x , f z ]) p (id C) + DiscFunc f .F-id {x} = substRefl {B = λ z → C [ f x , f z ]} (id C) + + -- Preserves composition + DiscFunc f .F-seq {x} {y} p q = + let open Category C using () renaming (_⋆_ to _●_) in + + let Hom[fx,f—] = (λ (w : A .fst) → C [ f x , f w ]) in + let Hom[fy,f—] = (λ (w : A .fst) → C [ f y , f w ]) in + let id-fx = id C {f x} in + let id-fy = id C {f y} in + let Fp = (subst Hom[fx,f—] (p) id-fx) in + + subst Hom[fx,f—] (p ∙ q) id-fx ≡⟨ substComposite Hom[fx,f—] _ _ _ ⟩ + subst Hom[fx,f—] (q) (Fp) ≡⟨ cong (subst _ q) (sym (⋆IdR C _)) ⟩ + subst Hom[fx,f—] (q) (Fp ● id-fy) ≡⟨ substCommSlice _ _ (λ _ → Fp ●_) q _ ⟩ + Fp ● (subst Hom[fy,f—] (q) id-fy) ∎ diff --git a/Cubical/Categories/Instances/DistLattice.agda b/Cubical/Categories/Instances/DistLattice.agda new file mode 100644 index 0000000000..b78c38afb5 --- /dev/null +++ b/Cubical/Categories/Instances/DistLattice.agda @@ -0,0 +1,14 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.DistLattice where + +open import Cubical.Foundations.Prelude + +open import Cubical.Algebra.DistLattice + +open import Cubical.Categories.Category +open import Cubical.Categories.Instances.Lattice + +open Category + +DistLatticeCategory : ∀ {ℓ} (L : DistLattice ℓ) → Category ℓ ℓ +DistLatticeCategory L = LatticeCategory (DistLattice→Lattice L) diff --git a/Cubical/Categories/Instances/Functors.agda b/Cubical/Categories/Instances/Functors.agda new file mode 100644 index 0000000000..dde890a76b --- /dev/null +++ b/Cubical/Categories/Instances/Functors.agda @@ -0,0 +1,52 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Instances.Functors where + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.NaturalTransformation.Base +open import Cubical.Categories.NaturalTransformation.Properties +open import Cubical.Categories.Morphism renaming (isIso to isIsoC) +open import Cubical.Foundations.Prelude + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +module _ (C : Category ℓC ℓC') (D : Category ℓD ℓD') where + open Category + open NatTrans + open Functor + + FUNCTOR : Category (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) (ℓ-max (ℓ-max ℓC ℓC') ℓD') + ob FUNCTOR = Functor C D + Hom[_,_] FUNCTOR = NatTrans + id FUNCTOR {F} = idTrans F + _⋆_ FUNCTOR = seqTrans + ⋆IdL FUNCTOR α = makeNatTransPath λ i x → D .⋆IdL (α .N-ob x) i + ⋆IdR FUNCTOR α = makeNatTransPath λ i x → D .⋆IdR (α .N-ob x) i + ⋆Assoc FUNCTOR α β γ = makeNatTransPath λ i x → D .⋆Assoc (α .N-ob x) (β .N-ob x) (γ .N-ob x) i + isSetHom FUNCTOR = isSetNatTrans + + open isIsoC renaming (inv to invC) + -- componentwise iso is an iso in Functor + FUNCTORIso : ∀ {F G : Functor C D} (α : F ⇒ G) + → (∀ (c : C .ob) → isIsoC D (α ⟦ c ⟧)) + → isIsoC FUNCTOR α + FUNCTORIso α is .invC .N-ob c = (is c) .invC + FUNCTORIso {F} {G} α is .invC .N-hom {c} {d} f + = invMoveL areInv-αc + ( α ⟦ c ⟧ ⋆⟨ D ⟩ (G ⟪ f ⟫ ⋆⟨ D ⟩ is d .invC) + ≡⟨ sym (D .⋆Assoc _ _ _) ⟩ + (α ⟦ c ⟧ ⋆⟨ D ⟩ G ⟪ f ⟫) ⋆⟨ D ⟩ is d .invC + ≡⟨ sym (invMoveR areInv-αd (α .N-hom f)) ⟩ + F ⟪ f ⟫ + ∎ ) + where + areInv-αc : areInv _ (α ⟦ c ⟧) ((is c) .invC) + areInv-αc = isIso→areInv (is c) + + areInv-αd : areInv _ (α ⟦ d ⟧) ((is d) .invC) + areInv-αd = isIso→areInv (is d) + FUNCTORIso α is .sec = makeNatTransPath (funExt (λ c → (is c) .sec)) + FUNCTORIso α is .ret = makeNatTransPath (funExt (λ c → (is c) .ret)) diff --git a/Cubical/Categories/Instances/Lattice.agda b/Cubical/Categories/Instances/Lattice.agda new file mode 100644 index 0000000000..3139dcc649 --- /dev/null +++ b/Cubical/Categories/Instances/Lattice.agda @@ -0,0 +1,15 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.Lattice where + +open import Cubical.Foundations.Prelude + +open import Cubical.Algebra.Lattice + +open import Cubical.Categories.Category +open import Cubical.Categories.Instances.Semilattice + +open Category + + +LatticeCategory : ∀ {ℓ} (L : Lattice ℓ) → Category ℓ ℓ +LatticeCategory L = JoinSemilatticeCategory (Lattice→JoinSemilattice L) diff --git a/Cubical/Categories/Instances/Poset.agda b/Cubical/Categories/Instances/Poset.agda new file mode 100644 index 0000000000..720326e9b8 --- /dev/null +++ b/Cubical/Categories/Instances/Poset.agda @@ -0,0 +1,28 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.Poset where + +open import Cubical.Foundations.Prelude + +open import Cubical.Relation.Binary.Poset + +open import Cubical.Categories.Category + +open Category + +private + variable + ℓ ℓ' : Level + +module _ (P : Poset ℓ ℓ') where + + open PosetStr (snd P) + + PosetCategory : Category ℓ ℓ' + ob PosetCategory = fst P + Hom[_,_] PosetCategory = _≤_ + id PosetCategory = is-refl _ + _⋆_ PosetCategory = is-trans _ _ _ + ⋆IdL PosetCategory _ = is-prop-valued _ _ _ _ + ⋆IdR PosetCategory _ = is-prop-valued _ _ _ _ + ⋆Assoc PosetCategory _ _ _ = is-prop-valued _ _ _ _ + isSetHom PosetCategory = isProp→isSet (is-prop-valued _ _) diff --git a/Cubical/Categories/Instances/Rings.agda b/Cubical/Categories/Instances/Rings.agda new file mode 100644 index 0000000000..f36e336bec --- /dev/null +++ b/Cubical/Categories/Instances/Rings.agda @@ -0,0 +1,22 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.Rings where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Algebra.Ring + +open import Cubical.Categories.Category + +open Category +open RingHoms + +RingsCategory : ∀ {ℓ} → Category (ℓ-suc ℓ) ℓ +ob RingsCategory = Ring _ +Hom[_,_] RingsCategory = RingHom +id RingsCategory {R} = idRingHom R +_⋆_ RingsCategory = compRingHom +⋆IdL RingsCategory = compIdRingHom +⋆IdR RingsCategory = idCompRingHom +⋆Assoc RingsCategory = compAssocRingHom +isSetHom RingsCategory = isSetRingHom _ _ diff --git a/Cubical/Categories/Instances/Semilattice.agda b/Cubical/Categories/Instances/Semilattice.agda new file mode 100644 index 0000000000..e698f39ae2 --- /dev/null +++ b/Cubical/Categories/Instances/Semilattice.agda @@ -0,0 +1,25 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.Semilattice where + +open import Cubical.Foundations.Prelude + +open import Cubical.Algebra.Semilattice + +open import Cubical.Categories.Category +open import Cubical.Categories.Instances.Poset + +open Category + + +module _ {ℓ} (L : Semilattice ℓ) where + open JoinSemilattice L + + JoinSemilatticeCategory : Category ℓ ℓ + JoinSemilatticeCategory = PosetCategory IndPoset + + +module _ {ℓ} (L : Semilattice ℓ) where + open MeetSemilattice L + + MeetSemilatticeCategory : Category ℓ ℓ + MeetSemilatticeCategory = PosetCategory IndPoset diff --git a/Cubical/Categories/Instances/Sets.agda b/Cubical/Categories/Instances/Sets.agda new file mode 100644 index 0000000000..962a0e06f1 --- /dev/null +++ b/Cubical/Categories/Instances/Sets.agda @@ -0,0 +1,92 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Instances.Sets where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation + +open import Cubical.Categories.Limits + +open Category + +module _ ℓ where + SET : Category (ℓ-suc ℓ) ℓ + ob SET = hSet ℓ + Hom[_,_] SET (A , _) (B , _) = A → B + id SET x = x + _⋆_ SET f g x = g (f x) + ⋆IdL SET f = refl + ⋆IdR SET f = refl + ⋆Assoc SET f g h = refl + isSetHom SET {A} {B} = isSetΠ (λ _ → snd B) + +private + variable + ℓ ℓ' : Level + +open Functor + +-- Hom functors +_[-,_] : (C : Category ℓ ℓ') → (c : C .ob) → Functor (C ^op) (SET ℓ') +(C [-, c ]) .F-ob x = (C [ x , c ]) , C .isSetHom +(C [-, c ]) .F-hom f k = f ⋆⟨ C ⟩ k +(C [-, c ]) .F-id = funExt λ _ → C .⋆IdL _ +(C [-, c ]) .F-seq _ _ = funExt λ _ → C .⋆Assoc _ _ _ + +_[_,-] : (C : Category ℓ ℓ') → (c : C .ob)→ Functor C (SET ℓ') +(C [ c ,-]) .F-ob x = (C [ c , x ]) , C .isSetHom +(C [ c ,-]) .F-hom f k = k ⋆⟨ C ⟩ f +(C [ c ,-]) .F-id = funExt λ _ → C .⋆IdR _ +(C [ c ,-]) .F-seq _ _ = funExt λ _ → sym (C .⋆Assoc _ _ _) + +module _ {C : Category ℓ ℓ'} {F : Functor C (SET ℓ')} where + open NatTrans + + -- natural transformations by pre/post composition + preComp : {x y : C .ob} + → (f : C [ x , y ]) + → C [ x ,-] ⇒ F + → C [ y ,-] ⇒ F + preComp f α .N-ob c k = (α ⟦ c ⟧) (f ⋆⟨ C ⟩ k) + preComp f α .N-hom {x = c} {d} k + = (λ l → (α ⟦ d ⟧) (f ⋆⟨ C ⟩ (l ⋆⟨ C ⟩ k))) + ≡[ i ]⟨ (λ l → (α ⟦ d ⟧) (⋆Assoc C f l k (~ i))) ⟩ + (λ l → (α ⟦ d ⟧) (f ⋆⟨ C ⟩ l ⋆⟨ C ⟩ k)) + ≡[ i ]⟨ (λ l → (α .N-hom k) i (f ⋆⟨ C ⟩ l)) ⟩ + (λ l → (F ⟪ k ⟫) ((α ⟦ c ⟧) (f ⋆⟨ C ⟩ l))) + ∎ + +-- properties +-- TODO: move to own file +open CatIso renaming (inv to cInv) +open Iso + +Iso→CatIso : ∀ {A B : (SET ℓ) .ob} + → Iso (fst A) (fst B) + → CatIso (SET ℓ) A B +Iso→CatIso is .mor = is .fun +Iso→CatIso is .cInv = is .inv +Iso→CatIso is .sec = funExt λ b → is .rightInv b -- is .rightInv +Iso→CatIso is .ret = funExt λ b → is .leftInv b -- is .rightInv + +-- SET is complete + +open LimCone +open Cone + +completeSET : ∀ {ℓJ ℓJ'} → Limits {ℓJ} {ℓJ'} (SET (ℓ-max ℓJ ℓJ')) +lim (completeSET J D) = Cone D (Unit* , isOfHLevelLift 2 isSetUnit) , isSetCone D _ +coneOut (limCone (completeSET J D)) j e = coneOut e j tt* +coneOutCommutes (limCone (completeSET J D)) j i e = coneOutCommutes e j i tt* +univProp (completeSET J D) c cc = + uniqueExists + (λ x → cone (λ v _ → coneOut cc v x) (λ e i _ → coneOutCommutes cc e i x)) + (λ _ → funExt (λ _ → refl)) + (λ x → isPropIsConeMor cc (limCone (completeSET J D)) x) + (λ x hx → funExt (λ d → cone≡ λ u → funExt (λ _ → sym (funExt⁻ (hx u) d)))) diff --git a/Cubical/Categories/Instances/TypePrecategory.agda b/Cubical/Categories/Instances/TypePrecategory.agda new file mode 100644 index 0000000000..41f62ef709 --- /dev/null +++ b/Cubical/Categories/Instances/TypePrecategory.agda @@ -0,0 +1,19 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.TypePrecategory where + +open import Cubical.Foundations.Prelude + +open import Cubical.Categories.Category.Precategory + +open Precategory + +-- TYPE precategory has types as objects +module _ ℓ where + TYPE : Precategory (ℓ-suc ℓ) ℓ + TYPE .ob = Type ℓ + TYPE .Hom[_,_] A B = A → B + TYPE .id = λ x → x + TYPE ._⋆_ f g = λ x → g (f x) + TYPE .⋆IdL f = refl + TYPE .⋆IdR f = refl + TYPE .⋆Assoc f g h = refl diff --git a/Cubical/Categories/Limits.agda b/Cubical/Categories/Limits.agda new file mode 100644 index 0000000000..63083fac59 --- /dev/null +++ b/Cubical/Categories/Limits.agda @@ -0,0 +1,9 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Limits where + +open import Cubical.Categories.Limits.Limits public +open import Cubical.Categories.Limits.BinProduct public +open import Cubical.Categories.Limits.BinCoproduct public +open import Cubical.Categories.Limits.Initial public +open import Cubical.Categories.Limits.Terminal public +open import Cubical.Categories.Limits.Pullback public diff --git a/Cubical/Categories/Limits/BinCoproduct.agda b/Cubical/Categories/Limits/BinCoproduct.agda new file mode 100644 index 0000000000..352f0b57a9 --- /dev/null +++ b/Cubical/Categories/Limits/BinCoproduct.agda @@ -0,0 +1,43 @@ +-- Binary coproducts +{-# OPTIONS --safe #-} + +module Cubical.Categories.Limits.BinCoproduct where + +open import Cubical.Categories.Category.Base +open import Cubical.Data.Sigma.Base +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓ ℓ' : Level + +module _ (C : Category ℓ ℓ') where + open Category C + + module _ {x y x+y : ob} + (i₁ : Hom[ x , x+y ]) + (i₂ : Hom[ y , x+y ]) where + + isBinCoproduct : Type (ℓ-max ℓ ℓ') + isBinCoproduct = ∀ {z : ob} (f₁ : Hom[ x , z ]) (f₂ : Hom[ y , z ]) → + ∃![ f ∈ Hom[ x+y , z ] ] (i₁ ⋆ f ≡ f₁) × (i₂ ⋆ f ≡ f₂) + + isPropIsBinCoproduct : isProp (isBinCoproduct) + isPropIsBinCoproduct = isPropImplicitΠ (λ _ → isPropΠ2 (λ _ _ → isPropIsContr)) + + + record BinCoproduct (x y : ob) : Type (ℓ-max ℓ ℓ') where + field + binCoprodOb : ob + binCoprodInj₁ : Hom[ x , binCoprodOb ] + binCoprodInj₂ : Hom[ y , binCoprodOb ] + univProp : isBinCoproduct binCoprodInj₁ binCoprodInj₂ + + + BinCoproducts : Type (ℓ-max ℓ ℓ') + BinCoproducts = (x y : ob) → BinCoproduct x y + + hasBinCoproducts : Type (ℓ-max ℓ ℓ') + hasBinCoproducts = ∥ BinCoproducts ∥ diff --git a/Cubical/Categories/Limits/BinProduct.agda b/Cubical/Categories/Limits/BinProduct.agda new file mode 100644 index 0000000000..02cc35ba44 --- /dev/null +++ b/Cubical/Categories/Limits/BinProduct.agda @@ -0,0 +1,43 @@ +-- Binary products +{-# OPTIONS --safe #-} + +module Cubical.Categories.Limits.BinProduct where + +open import Cubical.Categories.Category.Base +open import Cubical.Data.Sigma.Base +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓ ℓ' : Level + +module _ (C : Category ℓ ℓ') where + open Category C + + module _ {x y x×y : ob} + (π₁ : Hom[ x×y , x ]) + (π₂ : Hom[ x×y , y ]) where + + isBinProduct : Type (ℓ-max ℓ ℓ') + isBinProduct = ∀ {z : ob} (f₁ : Hom[ z , x ]) (f₂ : Hom[ z , y ]) → + ∃![ f ∈ Hom[ z , x×y ] ] (f ⋆ π₁ ≡ f₁) × (f ⋆ π₂ ≡ f₂) + + isPropIsBinProduct : isProp (isBinProduct) + isPropIsBinProduct = isPropImplicitΠ (λ _ → isPropΠ2 (λ _ _ → isPropIsContr)) + + + record BinProduct (x y : ob) : Type (ℓ-max ℓ ℓ') where + field + binProdOb : ob + binProdPr₁ : Hom[ binProdOb , x ] + binProdPr₂ : Hom[ binProdOb , y ] + univProp : isBinProduct binProdPr₁ binProdPr₂ + + + BinProducts : Type (ℓ-max ℓ ℓ') + BinProducts = (x y : ob) → BinProduct x y + + hasBinProducts : Type (ℓ-max ℓ ℓ') + hasBinProducts = ∥ BinProducts ∥ diff --git a/Cubical/Categories/Limits/Initial.agda b/Cubical/Categories/Limits/Initial.agda new file mode 100644 index 0000000000..ea34d15da5 --- /dev/null +++ b/Cubical/Categories/Limits/Initial.agda @@ -0,0 +1,61 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Limits.Initial where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.HITs.PropositionalTruncation.Base + +open import Cubical.Data.Sigma + +open import Cubical.Categories.Category + +private + variable + ℓ ℓ' : Level + +module _ (C : Category ℓ ℓ') where + open Category C + + isInitial : (x : ob) → Type (ℓ-max ℓ ℓ') + isInitial x = ∀ (y : ob) → isContr (C [ x , y ]) + + Initial : Type (ℓ-max ℓ ℓ') + Initial = Σ[ x ∈ ob ] isInitial x + + initialOb : Initial → ob + initialOb = fst + + initialArrow : (T : Initial) (y : ob) → C [ initialOb T , y ] + initialArrow T y = T .snd y .fst + + initialArrowUnique : {T : Initial} {y : ob} (f : C [ initialOb T , y ]) + → initialArrow T y ≡ f + initialArrowUnique {T} {y} f = T .snd y .snd f + + initialEndoIsId : (T : Initial) (f : C [ initialOb T , initialOb T ]) + → f ≡ id + initialEndoIsId T f = isContr→isProp (T .snd (initialOb T)) f id + + hasInitial : Type (ℓ-max ℓ ℓ') + hasInitial = ∥ Initial ∥ + + -- Initiality of an object is a proposition. + isPropIsInitial : (x : ob) → isProp (isInitial x) + isPropIsInitial _ = isPropΠ λ _ → isPropIsContr + + open CatIso + + -- Objects that are initial are isomorphic. + initialToIso : (x y : Initial) → CatIso C (initialOb x) (initialOb y) + mor (initialToIso x y) = initialArrow x (initialOb y) + inv (initialToIso x y) = initialArrow y (initialOb x) + sec (initialToIso x y) = initialEndoIsId y _ + ret (initialToIso x y) = initialEndoIsId x _ + + open isUnivalent + + -- The type of initial objects of a univalent category is a proposition, + -- i.e. all initial objects are equal. + isPropInitial : (hC : isUnivalent C) → isProp Initial + isPropInitial hC x y = + Σ≡Prop isPropIsInitial (CatIsoToPath hC (initialToIso x y)) diff --git a/Cubical/Categories/Limits/Limits.agda b/Cubical/Categories/Limits/Limits.agda new file mode 100644 index 0000000000..5d2b54845d --- /dev/null +++ b/Cubical/Categories/Limits/Limits.agda @@ -0,0 +1,139 @@ +-- Limits. +-- Heavily inspired by https://github.com/UniMath/UniMath/blob/master/UniMath/CategoryTheory/limits/graphs/limits.v +-- (except we're using categories instead of graphs to index limits) + +{-# OPTIONS --safe #-} +module Cubical.Categories.Limits.Limits where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma.Base + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor + +open import Cubical.HITs.PropositionalTruncation.Base + +module _ {ℓJ ℓJ' ℓC ℓC' : Level} {J : Category ℓJ ℓJ'} {C : Category ℓC ℓC'} where + open Category + open Functor + + private + ℓ = ℓ-max (ℓ-max (ℓ-max ℓJ ℓJ') ℓC) ℓC' + + record Cone (D : Functor J C) (c : ob C) : Type (ℓ-max (ℓ-max ℓJ ℓJ') ℓC') where + constructor cone + + field + coneOut : (v : ob J) → C [ c , F-ob D v ] + coneOutCommutes : {u v : ob J} (e : J [ u , v ]) → + coneOut u ⋆⟨ C ⟩ D .F-hom e ≡ coneOut v + + open Cone + + cone≡ : {D : Functor J C} {c : ob C} → {c1 c2 : Cone D c} + → ((v : ob J) → coneOut c1 v ≡ coneOut c2 v) → c1 ≡ c2 + coneOut (cone≡ h i) v = h v i + coneOutCommutes (cone≡ {D} {_} {c1} {c2} h i) {u} {v} e = + isProp→PathP (λ j → isSetHom C (h u j ⋆⟨ C ⟩ D .F-hom e) (h v j)) + (coneOutCommutes c1 e) (coneOutCommutes c2 e) i + + -- TODO: can we automate this a bit? + isSetCone : (D : Functor J C) (c : ob C) → isSet (Cone D c) + isSetCone D c = isSetRetract toConeΣ fromConeΣ (λ _ → refl) + (isSetΣSndProp (isSetΠ (λ _ → isSetHom C)) + (λ _ → isPropImplicitΠ2 (λ _ _ → isPropΠ (λ _ → isSetHom C _ _)))) + where + ConeΣ = Σ[ f ∈ ((v : ob J) → C [ c , F-ob D v ]) ] + ({u v : ob J} (e : J [ u , v ]) → f u ⋆⟨ C ⟩ D .F-hom e ≡ f v) + + toConeΣ : Cone D c → ConeΣ + fst (toConeΣ x) = coneOut x + snd (toConeΣ x) = coneOutCommutes x + + fromConeΣ : ConeΣ → Cone D c + coneOut (fromConeΣ x) = fst x + coneOutCommutes (fromConeΣ x) = snd x + + isConeMor : {c1 c2 : ob C} {D : Functor J C} + (cc1 : Cone D c1) (cc2 : Cone D c2) + → C [ c1 , c2 ] → Type (ℓ-max ℓJ ℓC') + isConeMor cc1 cc2 f = (v : ob J) → f ⋆⟨ C ⟩ coneOut cc2 v ≡ coneOut cc1 v + + isPropIsConeMor : {c1 c2 : ob C} {D : Functor J C} + (cc1 : Cone D c1) (cc2 : Cone D c2) (f : C [ c1 , c2 ]) + → isProp (isConeMor cc1 cc2 f) + isPropIsConeMor cc1 cc2 f = isPropΠ (λ _ → isSetHom C _ _) + + isLimCone : (D : Functor J C) (c0 : ob C) → Cone D c0 → Type ℓ + isLimCone D c0 cc0 = (c : ob C) → (cc : Cone D c) + → ∃![ f ∈ C [ c , c0 ] ] isConeMor cc cc0 f + + isPropIsLimCone : (D : Functor J C) (c0 : ob C) (cc0 : Cone D c0) + → isProp (isLimCone D c0 cc0) + isPropIsLimCone D c0 cc0 = isPropΠ2 λ _ _ → isProp∃! + + record LimCone (D : Functor J C) : Type ℓ where + constructor climCone + + field + lim : ob C + limCone : Cone D lim + univProp : isLimCone D lim limCone + + limOut : (v : ob J) → C [ lim , D .F-ob v ] + limOut = coneOut limCone + + limOutCommutes : {u v : ob J} (e : J [ u , v ]) + → limOut u ⋆⟨ C ⟩ D .F-hom e ≡ limOut v + limOutCommutes = coneOutCommutes limCone + + limArrow : (c : ob C) → (cc : Cone D c) → C [ c , lim ] + limArrow c cc = univProp c cc .fst .fst + + limArrowCommutes : (c : ob C) → (cc : Cone D c) (u : ob J) + → limArrow c cc ⋆⟨ C ⟩ limOut u ≡ coneOut cc u + limArrowCommutes c cc = univProp c cc .fst .snd + + limArrowUnique : (c : ob C) → (cc : Cone D c) (k : C [ c , lim ]) + → isConeMor cc limCone k → limArrow c cc ≡ k + limArrowUnique c cc k hk = cong fst (univProp c cc .snd (k , hk)) + + -- TODO: define limOfArrows + +-- A category is complete if it has all limits +Limits : {ℓJ ℓJ' ℓC ℓC' : Level} → Category ℓC ℓC' → Type _ +Limits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') → (D : Functor J C) → LimCone D + +hasLimits : {ℓJ ℓJ' ℓC ℓC' : Level} → Category ℓC ℓC' → Type _ +hasLimits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') → (D : Functor J C) → ∥ LimCone D ∥ + +-- Limits of a specific shape J in a category C +LimitsOfShape : {ℓJ ℓJ' ℓC ℓC' : Level} → Category ℓJ ℓJ' → Category ℓC ℓC' → Type _ +LimitsOfShape J C = (D : Functor J C) → LimCone D + + +-- Preservation of limits + +module _ {ℓJ ℓJ' ℓC1 ℓC1' ℓC2 ℓC2' : Level} + {C1 : Category ℓC1 ℓC1'} {C2 : Category ℓC2 ℓC2'} + (F : Functor C1 C2) where + open Category + open Functor + open Cone + + private + ℓ = ℓ-max (ℓ-max (ℓ-max (ℓ-max (ℓ-max ℓJ ℓJ') ℓC1) ℓC1') ℓC2) ℓC2' + + mapCone : {J : Category ℓJ ℓJ'} (D : Functor J C1) {x : ob C1} (ccx : Cone D x) → Cone (funcComp F D) (F .F-ob x) + coneOut (mapCone D ccx) v = F .F-hom (coneOut ccx v) + coneOutCommutes (mapCone D ccx) e = + sym (F-seq F (coneOut ccx _) (D ⟪ e ⟫)) ∙ cong (F .F-hom) (coneOutCommutes ccx e) + + preservesLimit : {J : Category ℓJ ℓJ'} (D : Functor J C1) + → (L : ob C1) → Cone D L → Type ℓ + preservesLimit D L ccL = + isLimCone D L ccL → isLimCone (funcComp F D) (F .F-ob L) (mapCone D ccL) + + -- TODO: prove that right adjoints preserve limits diff --git a/Cubical/Categories/Limits/Pullback.agda b/Cubical/Categories/Limits/Pullback.agda new file mode 100644 index 0000000000..ae1e7a333f --- /dev/null +++ b/Cubical/Categories/Limits/Pullback.agda @@ -0,0 +1,153 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Limits.Pullback where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.HITs.PropositionalTruncation.Base + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.Instances.Cospan +open import Cubical.Categories.Limits.Limits + +private + variable + ℓ ℓ' : Level + +module _ (C : Category ℓ ℓ') where + + open Category C + open Functor + + record Cospan : Type (ℓ-max ℓ ℓ') where + constructor cospan + field + l m r : ob + s₁ : C [ l , m ] + s₂ : C [ r , m ] + + open Cospan + + isPullback : (cspn : Cospan) → + {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ]) + (H : p₁ ⋆ cspn .s₁ ≡ p₂ ⋆ cspn .s₂) → Type (ℓ-max ℓ ℓ') + isPullback cspn {c} p₁ p₂ H = + ∀ {d} (h : C [ d , cspn .l ]) (k : C [ d , cspn .r ]) + (H' : h ⋆ cspn .s₁ ≡ k ⋆ cspn .s₂) + → ∃![ hk ∈ C [ d , c ] ] (h ≡ hk ⋆ p₁) × (k ≡ hk ⋆ p₂) + + isPropIsPullback : (cspn : Cospan) → + {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ]) + (H : p₁ ⋆ cspn .s₁ ≡ p₂ ⋆ cspn .s₂) → isProp (isPullback cspn p₁ p₂ H) + isPropIsPullback cspn p₁ p₂ H = + isPropImplicitΠ (λ x → isPropΠ3 λ h k H' → isPropIsContr) + + record Pullback (cspn : Cospan) : Type (ℓ-max ℓ ℓ') where + field + pbOb : ob + pbPr₁ : C [ pbOb , cspn .l ] + pbPr₂ : C [ pbOb , cspn .r ] + pbCommutes : pbPr₁ ⋆ cspn .s₁ ≡ pbPr₂ ⋆ cspn .s₂ + univProp : isPullback cspn pbPr₁ pbPr₂ pbCommutes + + open Pullback + + pullbackArrow : + {cspn : Cospan} (pb : Pullback cspn) + {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ]) + (H : p₁ ⋆ cspn .s₁ ≡ p₂ ⋆ cspn .s₂) → C [ c , pb . pbOb ] + pullbackArrow pb p₁ p₂ H = pb .univProp p₁ p₂ H .fst .fst + + pullbackArrowPr₁ : + {cspn : Cospan} (pb : Pullback cspn) + {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ]) + (H : p₁ ⋆ cspn .s₁ ≡ p₂ ⋆ cspn .s₂) → + p₁ ≡ pullbackArrow pb p₁ p₂ H ⋆ pbPr₁ pb + pullbackArrowPr₁ pb p₁ p₂ H = pb .univProp p₁ p₂ H .fst .snd .fst + + pullbackArrowPr₂ : + {cspn : Cospan} (pb : Pullback cspn) + {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ]) + (H : p₁ ⋆ cspn .s₁ ≡ p₂ ⋆ cspn .s₂) → + p₂ ≡ pullbackArrow pb p₁ p₂ H ⋆ pbPr₂ pb + pullbackArrowPr₂ pb p₁ p₂ H = pb .univProp p₁ p₂ H .fst .snd .snd + + pullbackArrowUnique : + {cspn : Cospan} (pb : Pullback cspn) + {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ]) + (H : p₁ ⋆ cspn .s₁ ≡ p₂ ⋆ cspn .s₂) + (pbArrow' : C [ c , pb .pbOb ]) + (H₁ : p₁ ≡ pbArrow' ⋆ pbPr₁ pb) (H₂ : p₂ ≡ pbArrow' ⋆ pbPr₂ pb) + → pullbackArrow pb p₁ p₂ H ≡ pbArrow' + pullbackArrowUnique pb p₁ p₂ H pbArrow' H₁ H₂ = + cong fst (pb .univProp p₁ p₂ H .snd (pbArrow' , (H₁ , H₂))) + + Pullbacks : Type (ℓ-max ℓ ℓ') + Pullbacks = (cspn : Cospan) → Pullback cspn + + hasPullbacks : Type (ℓ-max ℓ ℓ') + hasPullbacks = ∥ Pullbacks ∥ + + +-- Pullbacks from limits +module _ {C : Category ℓ ℓ'} where + open Category C + open Functor + open Pullback + open LimCone + open Cone + open Cospan + + Cospan→Func : Cospan C → Functor CospanCat C + Cospan→Func (cospan l m r f g) .F-ob ⓪ = l + Cospan→Func (cospan l m r f g) .F-ob ① = m + Cospan→Func (cospan l m r f g) .F-ob ② = r + Cospan→Func (cospan l m r f g) .F-hom {⓪} {①} k = f + Cospan→Func (cospan l m r f g) .F-hom {②} {①} k = g + Cospan→Func (cospan l m r f g) .F-hom {⓪} {⓪} k = id + Cospan→Func (cospan l m r f g) .F-hom {①} {①} k = id + Cospan→Func (cospan l m r f g) .F-hom {②} {②} k = id + Cospan→Func (cospan l m r f g) .F-id {⓪} = refl + Cospan→Func (cospan l m r f g) .F-id {①} = refl + Cospan→Func (cospan l m r f g) .F-id {②} = refl + Cospan→Func (cospan l m r f g) .F-seq {⓪} {⓪} {⓪} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {⓪} {⓪} {①} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {⓪} {①} {①} φ ψ = sym (⋆IdR _) + Cospan→Func (cospan l m r f g) .F-seq {①} {①} {①} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {②} {②} {②} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {②} {②} {①} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {②} {①} {①} φ ψ = sym (⋆IdR _) + + LimitsOfShapeCospanCat→Pullbacks : LimitsOfShape CospanCat C → Pullbacks C + pbOb (LimitsOfShapeCospanCat→Pullbacks H cspn) = lim (H (Cospan→Func cspn)) + pbPr₁ (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOut (H (Cospan→Func cspn)) ⓪ + pbPr₂ (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOut (H (Cospan→Func cspn)) ② + pbCommutes (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOutCommutes (H (Cospan→Func cspn)) {v = ①} tt + ∙ sym (limOutCommutes (H (Cospan→Func cspn)) tt) + univProp (LimitsOfShapeCospanCat→Pullbacks H cspn) {d = d} h k H' = + uniqueExists (limArrow (H (Cospan→Func cspn)) d cc) + ( sym (limArrowCommutes (H (Cospan→Func cspn)) d cc ⓪) + , sym (limArrowCommutes (H (Cospan→Func cspn)) d cc ②)) + (λ _ → isProp× (isSetHom _ _) (isSetHom _ _)) + λ a' ha' → limArrowUnique (H (Cospan→Func cspn)) d cc a' + (λ { ⓪ → sym (ha' .fst) + ; ① → cong (a' ⋆_) (sym (limOutCommutes (H (Cospan→Func cspn)) {⓪} {①} tt)) + ∙∙ sym (⋆Assoc _ _ _) + ∙∙ cong (_⋆ cspn .s₁) (sym (ha' .fst)) + ; ② → sym (ha' .snd) }) + where + cc : Cone (Cospan→Func cspn) d + coneOut cc ⓪ = h + coneOut cc ① = h ⋆ cspn .s₁ + coneOut cc ② = k + coneOutCommutes cc {⓪} {⓪} e = ⋆IdR h + coneOutCommutes cc {⓪} {①} e = refl + coneOutCommutes cc {①} {①} e = ⋆IdR _ + coneOutCommutes cc {②} {①} e = sym H' + coneOutCommutes cc {②} {②} e = ⋆IdR k + + Limits→Pullbacks : Limits C → Pullbacks C + Limits→Pullbacks H = LimitsOfShapeCospanCat→Pullbacks (H CospanCat) diff --git a/Cubical/Categories/Limits/Terminal.agda b/Cubical/Categories/Limits/Terminal.agda new file mode 100644 index 0000000000..e81d9b6d04 --- /dev/null +++ b/Cubical/Categories/Limits/Terminal.agda @@ -0,0 +1,61 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Limits.Terminal where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.HITs.PropositionalTruncation.Base + +open import Cubical.Data.Sigma + +open import Cubical.Categories.Category + +private + variable + ℓ ℓ' : Level + +module _ (C : Category ℓ ℓ') where + open Category C + + isTerminal : (x : ob) → Type (ℓ-max ℓ ℓ') + isTerminal x = ∀ (y : ob) → isContr (C [ y , x ]) + + Terminal : Type (ℓ-max ℓ ℓ') + Terminal = Σ[ x ∈ ob ] isTerminal x + + terminalOb : Terminal → ob + terminalOb = fst + + terminalArrow : (T : Terminal) (y : ob) → C [ y , terminalOb T ] + terminalArrow T y = T .snd y .fst + + terminalArrowUnique : {T : Terminal} {y : ob} (f : C [ y , terminalOb T ]) + → terminalArrow T y ≡ f + terminalArrowUnique {T} {y} f = T .snd y .snd f + + terminalEndoIsId : (T : Terminal) (f : C [ terminalOb T , terminalOb T ]) + → f ≡ id + terminalEndoIsId T f = isContr→isProp (T .snd (terminalOb T)) f id + + hasTerminal : Type (ℓ-max ℓ ℓ') + hasTerminal = ∥ Terminal ∥ + + -- Terminality of an object is a proposition. + isPropIsTerminal : (x : ob) → isProp (isTerminal x) + isPropIsTerminal _ = isPropΠ λ _ → isPropIsContr + + open CatIso + + -- Objects that are initial are isomorphic. + terminalToIso : (x y : Terminal) → CatIso C (terminalOb x) (terminalOb y) + mor (terminalToIso x y) = terminalArrow y (terminalOb x) + inv (terminalToIso x y) = terminalArrow x (terminalOb y) + sec (terminalToIso x y) = terminalEndoIsId y _ + ret (terminalToIso x y) = terminalEndoIsId x _ + + open isUnivalent + + -- The type of terminal objects of a univalent category is a proposition, + -- i.e. all terminal objects are equal. + isPropTerminal : (hC : isUnivalent C) → isProp Terminal + isPropTerminal hC x y = + Σ≡Prop isPropIsTerminal (CatIsoToPath hC (terminalToIso x y)) diff --git a/Cubical/Categories/Monoidal.agda b/Cubical/Categories/Monoidal.agda new file mode 100644 index 0000000000..bc2a153a14 --- /dev/null +++ b/Cubical/Categories/Monoidal.agda @@ -0,0 +1,7 @@ +-- Monoidal categories +{-# OPTIONS --safe #-} + +module Cubical.Categories.Monoidal where + +open import Cubical.Categories.Monoidal.Base public +open import Cubical.Categories.Monoidal.Enriched public diff --git a/Cubical/Categories/Monoidal/Base.agda b/Cubical/Categories/Monoidal/Base.agda new file mode 100644 index 0000000000..c9f0e7546f --- /dev/null +++ b/Cubical/Categories/Monoidal/Base.agda @@ -0,0 +1,121 @@ +-- Monoidal categories +{-# OPTIONS --safe #-} + +module Cubical.Categories.Monoidal.Base where + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Constructions.Product +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Functor.Product +open import Cubical.Categories.Morphism +open import Cubical.Categories.NaturalTransformation.Base +open import Cubical.Foundations.Prelude + + +module _ {ℓ ℓ' : Level} (C : Category ℓ ℓ') where + open Category C + + private + record TensorStr : Type (ℓ-max ℓ ℓ') where + field + ─⊗─ : Functor (C × C) C + unit : ob + + open Functor + + -- Useful tensor product notation + _⊗_ : ob → ob → ob + x ⊗ y = ─⊗─ .F-ob (x , y) + + _⊗ₕ_ : ∀ {x y z w} → Hom[ x , y ] → Hom[ z , w ] + → Hom[ x ⊗ z , y ⊗ w ] + f ⊗ₕ g = ─⊗─ .F-hom (f , g) + + + record StrictMonStr : Type (ℓ-max ℓ ℓ') where + field + tenstr : TensorStr + + open TensorStr tenstr public + + field + -- Axioms - strict + assoc : ∀ x y z → x ⊗ (y ⊗ z) ≡ (x ⊗ y) ⊗ z + idl : ∀ x → unit ⊗ x ≡ x + idr : ∀ x → x ⊗ unit ≡ x + + + record MonoidalStr : Type (ℓ-max ℓ ℓ') where + field + tenstr : TensorStr + + open TensorStr tenstr public + + private + -- Private names to make the axioms below look nice + x⊗[y⊗z] : Functor (C × C × C) C + x⊗[y⊗z] = ─⊗─ ∘F (𝟙⟨ C ⟩ ×F ─⊗─) + + [x⊗y]⊗z : Functor (C × C × C) C + [x⊗y]⊗z = ─⊗─ ∘F (─⊗─ ×F 𝟙⟨ C ⟩) ∘F (×C-assoc C C C) + + x = 𝟙⟨ C ⟩ + 1⊗x = ─⊗─ ∘F (rinj C C unit) + x⊗1 = ─⊗─ ∘F (linj C C unit) + + field + -- "Axioms" - up to natural isomorphism + α : x⊗[y⊗z] ≅ᶜ [x⊗y]⊗z + η : 1⊗x ≅ᶜ x + ρ : x⊗1 ≅ᶜ x + + open NatIso + + -- More nice notations + α⟨_,_,_⟩ : (x y z : ob) → Hom[ x ⊗ (y ⊗ z) , (x ⊗ y) ⊗ z ] + α⟨ x , y , z ⟩ = α .trans ⟦ ( x , y , z ) ⟧ + + η⟨_⟩ : (x : ob) → Hom[ unit ⊗ x , x ] + η⟨ x ⟩ = η .trans ⟦ x ⟧ + + ρ⟨_⟩ : (x : ob) → Hom[ x ⊗ unit , x ] + ρ⟨ x ⟩ = ρ .trans ⟦ x ⟧ + + field + -- Coherence conditions + pentagon : ∀ w x y z → + id ⊗ₕ α⟨ x , y , z ⟩ ⋆ α⟨ w , x ⊗ y , z ⟩ ⋆ α⟨ w , x , y ⟩ ⊗ₕ id + ≡ α⟨ w , x , y ⊗ z ⟩ ⋆ α⟨ w ⊗ x , y , z ⟩ + + triangle : ∀ x y → + α⟨ x , unit , y ⟩ ⋆ ρ⟨ x ⟩ ⊗ₕ id ≡ id ⊗ₕ η⟨ y ⟩ + + open isIso + + -- Inverses of α, η, ρ for convenience + α⁻¹⟨_,_,_⟩ : (x y z : ob) → Hom[ (x ⊗ y) ⊗ z , x ⊗ (y ⊗ z) ] + α⁻¹⟨ x , y , z ⟩ = α .nIso (x , y , z) .inv + + η⁻¹⟨_⟩ : (x : ob) → Hom[ x , unit ⊗ x ] + η⁻¹⟨ x ⟩ = η .nIso (x) .inv + + ρ⁻¹⟨_⟩ : (x : ob) → Hom[ x , x ⊗ unit ] + ρ⁻¹⟨ x ⟩ = ρ .nIso (x) .inv + + +record StrictMonCategory ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + field + C : Category ℓ ℓ' + sms : StrictMonStr C + + open Category C public + open StrictMonStr sms public + + +record MonoidalCategory ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + field + C : Category ℓ ℓ' + monstr : MonoidalStr C + + open Category C public + open MonoidalStr monstr public diff --git a/Cubical/Categories/Monoidal/Enriched.agda b/Cubical/Categories/Monoidal/Enriched.agda new file mode 100644 index 0000000000..c7ea5d60ad --- /dev/null +++ b/Cubical/Categories/Monoidal/Enriched.agda @@ -0,0 +1,29 @@ +-- Enriched categories +{-# OPTIONS --safe #-} + +module Cubical.Categories.Monoidal.Enriched where + +open import Cubical.Categories.Monoidal.Base +open import Cubical.Foundations.Prelude + +module _ {ℓV ℓV' : Level} (V : MonoidalCategory ℓV ℓV') (ℓE : Level) where + + open MonoidalCategory V + renaming (ob to obV; Hom[_,_] to V[_,_]; id to idV; _⋆_ to _⋆V_) + + record EnrichedCategory : Type (ℓ-max (ℓ-max ℓV ℓV') (ℓ-suc ℓE)) where + field + ob : Type ℓE + Hom[_,_] : ob → ob → obV + id : ∀ {x} → V[ unit , Hom[ x , x ] ] + seq : ∀ x y z → V[ Hom[ x , y ] ⊗ Hom[ y , z ] , Hom[ x , z ] ] + + -- Axioms + ⋆IdL : ∀ x y → η⟨ _ ⟩ ≡ (id {x} ⊗ₕ idV) ⋆V (seq x x y) + ⋆IdR : ∀ x y → ρ⟨ _ ⟩ ≡ (idV ⊗ₕ id {y}) ⋆V (seq x y y) + ⋆Assoc : ∀ x y z w → + α⟨ _ , _ , _ ⟩ ⋆V ((seq x y z) ⊗ₕ idV) ⋆V (seq x z w) + ≡ (idV ⊗ₕ (seq y z w)) ⋆V (seq x y w) + + +-- TODO: define underlying category using Hom[ x , y ] := V[ unit , Hom[ x , y ] ] diff --git a/Cubical/Categories/Morphism.agda b/Cubical/Categories/Morphism.agda new file mode 100644 index 0000000000..0cf33a1c93 --- /dev/null +++ b/Cubical/Categories/Morphism.agda @@ -0,0 +1,129 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Morphism where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Categories.Category + + +private + variable + ℓ ℓ' : Level + +-- C needs to be explicit for these definitions as Agda can't infer it +module _ (C : Category ℓ ℓ') where + open Category C + + private + variable + x y z w : ob + + isMonic : Hom[ x , y ] → Type (ℓ-max ℓ ℓ') + isMonic {x} {y} f = + ∀ {z} {a a' : Hom[ z , x ]} → f ∘ a ≡ f ∘ a' → a ≡ a' + + isEpic : (Hom[ x , y ]) → Type (ℓ-max ℓ ℓ') + isEpic {x} {y} g = + ∀ {z} {b b' : Hom[ y , z ]} → b ∘ g ≡ b' ∘ g → b ≡ b' + + -- A morphism is split monic if it has a *retraction* + isSplitMon : (Hom[ x , y ]) → Type ℓ' + isSplitMon {x} {y} f = ∃[ r ∈ Hom[ y , x ] ] r ∘ f ≡ id + + -- A morphism is split epic if it has a *section* + isSplitEpi : (Hom[ x , y ]) → Type ℓ' + isSplitEpi {x} {y} g = ∃[ s ∈ Hom[ y , x ] ] g ∘ s ≡ id + + record areInv (f : Hom[ x , y ]) (g : Hom[ y , x ]) : Type ℓ' where + field + sec : g ⋆ f ≡ id + ret : f ⋆ g ≡ id + + record isIso (f : Hom[ x , y ]) : Type ℓ' where + field + inv : Hom[ y , x ] + sec : inv ⋆ f ≡ id + ret : f ⋆ inv ≡ id + + +-- C can be implicit here +module _ {C : Category ℓ ℓ'} where + open Category C + + private + variable + x y z w : ob + + open areInv + + symAreInv : {f : Hom[ x , y ]} {g : Hom[ y , x ]} → areInv C f g → areInv C g f + sec (symAreInv x) = ret x + ret (symAreInv x) = sec x + + -- equational reasoning with inverses + invMoveR : ∀ {f : Hom[ x , y ]} {g : Hom[ y , x ]} {h : Hom[ z , x ]} {k : Hom[ z , y ]} + → areInv C f g + → h ⋆ f ≡ k + → h ≡ k ⋆ g + invMoveR {f = f} {g} {h} {k} ai p + = h + ≡⟨ sym (⋆IdR _) ⟩ + (h ⋆ id) + ≡⟨ cong (h ⋆_) (sym (ai .ret)) ⟩ + (h ⋆ (f ⋆ g)) + ≡⟨ sym (⋆Assoc _ _ _) ⟩ + ((h ⋆ f) ⋆ g) + ≡⟨ cong (_⋆ g) p ⟩ + k ⋆ g + ∎ + + invMoveL : ∀ {f : Hom[ x , y ]} {g : Hom[ y , x ]} {h : Hom[ y , z ]} {k : Hom[ x , z ]} + → areInv C f g + → f ⋆ h ≡ k + → h ≡ g ⋆ k + invMoveL {f = f} {g} {h} {k} ai p + = h + ≡⟨ sym (⋆IdL _) ⟩ + id ⋆ h + ≡⟨ cong (_⋆ h) (sym (ai .sec)) ⟩ + (g ⋆ f) ⋆ h + ≡⟨ ⋆Assoc _ _ _ ⟩ + g ⋆ (f ⋆ h) + ≡⟨ cong (g ⋆_) p ⟩ + (g ⋆ k) + ∎ + + open isIso + + isIso→areInv : ∀ {f : Hom[ x , y ]} + → (isI : isIso C f) + → areInv C f (isI .inv) + sec (isIso→areInv isI) = sec isI + ret (isIso→areInv isI) = ret isI + + open CatIso + + -- isIso agrees with CatIso + isIso→CatIso : ∀ {f : C [ x , y ]} + → isIso C f + → CatIso C x y + mor (isIso→CatIso {f = f} x) = f + inv (isIso→CatIso x) = inv x + sec (isIso→CatIso x) = sec x + ret (isIso→CatIso x) = ret x + + CatIso→isIso : (cIso : CatIso C x y) + → isIso C (cIso .mor) + inv (CatIso→isIso f) = inv f + sec (CatIso→isIso f) = sec f + ret (CatIso→isIso f) = ret f + + CatIso→areInv : (cIso : CatIso C x y) + → areInv C (cIso .mor) (cIso .inv) + CatIso→areInv cIso = isIso→areInv (CatIso→isIso cIso) + + -- reverse of an iso is also an iso + symCatIso : ∀ {x y} + → CatIso C x y + → CatIso C y x + symCatIso (catiso mor inv sec ret) = catiso inv mor ret sec diff --git a/Cubical/Categories/NaturalTransformation.agda b/Cubical/Categories/NaturalTransformation.agda new file mode 100644 index 0000000000..ccde33ef6f --- /dev/null +++ b/Cubical/Categories/NaturalTransformation.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.NaturalTransformation where + +open import Cubical.Categories.NaturalTransformation.Base public +open import Cubical.Categories.NaturalTransformation.Properties public diff --git a/Cubical/Categories/NaturalTransformation/Base.agda b/Cubical/Categories/NaturalTransformation/Base.agda new file mode 100644 index 0000000000..52aaba0a71 --- /dev/null +++ b/Cubical/Categories/NaturalTransformation/Base.agda @@ -0,0 +1,216 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.NaturalTransformation.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism renaming (iso to iIso) +open import Cubical.Data.Sigma +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Functor.Properties +open import Cubical.Categories.Commutativity +open import Cubical.Categories.Morphism renaming (isIso to isIsoC) + +private + variable + ℓA ℓA' ℓB ℓB' ℓC ℓC' ℓD ℓD' : Level + +module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where + -- syntax for sequencing in category D + infixl 15 _⋆ᴰ_ + private + _⋆ᴰ_ : ∀ {x y z} (f : D [ x , y ]) (g : D [ y , z ]) → D [ x , z ] + f ⋆ᴰ g = f ⋆⟨ D ⟩ g + + open Category + open Functor + + -- type aliases because it gets tedious typing it out all the time + N-ob-Type : (F G : Functor C D) → Type _ + N-ob-Type F G = (x : C .ob) → D [(F .F-ob x) , (G .F-ob x)] + + N-hom-Type : (F G : Functor C D) → N-ob-Type F G → Type _ + N-hom-Type F G ϕ = {x y : C .ob} (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (ϕ y) ≡ (ϕ x) ⋆ᴰ (G .F-hom f) + + record NatTrans (F G : Functor C D) : Type (ℓ-max (ℓ-max ℓC ℓC') ℓD') where + constructor natTrans + field + -- components of the natural transformation + N-ob : N-ob-Type F G + -- naturality condition + N-hom : N-hom-Type F G N-ob + + record NatIso (F G : Functor C D): Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + trans : NatTrans F G + open NatTrans trans + + field + nIso : ∀ (x : C .ob) → isIsoC D (N-ob x) + + open isIsoC + + -- the three other commuting squares + sqRL : ∀ {x y : C .ob} {f : C [ x , y ]} + → F ⟪ f ⟫ ≡ (N-ob x) ⋆ᴰ G ⟪ f ⟫ ⋆ᴰ (nIso y) .inv + sqRL {x} {y} {f} = invMoveR (isIso→areInv (nIso y)) (N-hom f) + + sqLL : ∀ {x y : C .ob} {f : C [ x , y ]} + → G ⟪ f ⟫ ⋆ᴰ (nIso y) .inv ≡ (nIso x) .inv ⋆ᴰ F ⟪ f ⟫ + sqLL {x} {y} {f} = invMoveL (isIso→areInv (nIso x)) (sym sqRL') + where + sqRL' : F ⟪ f ⟫ ≡ (N-ob x) ⋆ᴰ ( G ⟪ f ⟫ ⋆ᴰ (nIso y) .inv ) + sqRL' = sqRL ∙ (D .⋆Assoc _ _ _) + + sqLR : ∀ {x y : C .ob} {f : C [ x , y ]} + → G ⟪ f ⟫ ≡ (nIso x) .inv ⋆ᴰ F ⟪ f ⟫ ⋆ᴰ (N-ob y) + sqLR {x} {y} {f} = invMoveR (symAreInv (isIso→areInv (nIso y))) sqLL + + open NatTrans + open NatIso + + infix 10 NatTrans + syntax NatTrans F G = F ⇒ G + + infix 9 NatIso + syntax NatIso F G = F ≅ᶜ G -- c superscript to indicate that this is in the context of categories + + -- component of a natural transformation + infix 30 _⟦_⟧ + _⟦_⟧ : ∀ {F G : Functor C D} → F ⇒ G → (x : C .ob) → D [ F .F-ob x , G .F-ob x ] + _⟦_⟧ = N-ob + + idTrans : (F : Functor C D) → NatTrans F F + idTrans F .N-ob x = D .id + idTrans F .N-hom f = + (F .F-hom f) ⋆ᴰ (idTrans F .N-ob _) + ≡⟨ D .⋆IdR _ ⟩ + F .F-hom f + ≡⟨ sym (D .⋆IdL _) ⟩ + (D .id) ⋆ᴰ (F .F-hom f) + ∎ + + syntax idTrans F = 1[ F ] + + + -- vertical sequencing + seqTrans : {F G H : Functor C D} (α : NatTrans F G) (β : NatTrans G H) → NatTrans F H + seqTrans α β .N-ob x = (α .N-ob x) ⋆ᴰ (β .N-ob x) + seqTrans {F} {G} {H} α β .N-hom f = + (F .F-hom f) ⋆ᴰ ((α .N-ob _) ⋆ᴰ (β .N-ob _)) + ≡⟨ sym (D .⋆Assoc _ _ _) ⟩ + ((F .F-hom f) ⋆ᴰ (α .N-ob _)) ⋆ᴰ (β .N-ob _) + ≡[ i ]⟨ (α .N-hom f i) ⋆ᴰ (β .N-ob _) ⟩ + ((α .N-ob _) ⋆ᴰ (G .F-hom f)) ⋆ᴰ (β .N-ob _) + ≡⟨ D .⋆Assoc _ _ _ ⟩ + (α .N-ob _) ⋆ᴰ ((G .F-hom f) ⋆ᴰ (β .N-ob _)) + ≡[ i ]⟨ (α .N-ob _) ⋆ᴰ (β .N-hom f i) ⟩ + (α .N-ob _) ⋆ᴰ ((β .N-ob _) ⋆ᴰ (H .F-hom f)) + ≡⟨ sym (D .⋆Assoc _ _ _) ⟩ + ((α .N-ob _) ⋆ᴰ (β .N-ob _)) ⋆ᴰ (H .F-hom f) + ∎ + + compTrans : {F G H : Functor C D} (β : NatTrans G H) (α : NatTrans F G) → NatTrans F H + compTrans β α = seqTrans α β + + infixl 8 seqTrans + syntax seqTrans α β = α ●ᵛ β + + + -- vertically sequence natural transformations whose + -- common functor is not definitional equal + seqTransP : {F G G' H : Functor C D} (p : G ≡ G') + → (α : NatTrans F G) (β : NatTrans G' H) + → NatTrans F H + seqTransP {F} {G} {G'} {H} p α β .N-ob x + -- sequence morphisms with non-judgementally equal (co)domain + = seqP {C = D} {p = Gx≡G'x} (α ⟦ x ⟧) (β ⟦ x ⟧) + where + Gx≡G'x : ∀ {x} → G ⟅ x ⟆ ≡ G' ⟅ x ⟆ + Gx≡G'x {x} i = F-ob (p i) x + seqTransP {F} {G} {G'} {H} p α β .N-hom {x = x} {y} f + -- compose the two commuting squares + -- 1. α's commuting square + -- 2. β's commuting square, but extended to G since β is only G' ≡> H + = compSq {C = D} (α .N-hom f) βSq + where + -- functor equality implies equality of actions on objects and morphisms + Gx≡G'x : G ⟅ x ⟆ ≡ G' ⟅ x ⟆ + Gx≡G'x i = F-ob (p i) x + + Gy≡G'y : G ⟅ y ⟆ ≡ G' ⟅ y ⟆ + Gy≡G'y i = F-ob (p i) y + + Gf≡G'f : PathP (λ i → D [ Gx≡G'x i , Gy≡G'y i ]) (G ⟪ f ⟫) (G' ⟪ f ⟫) + Gf≡G'f i = p i ⟪ f ⟫ + + -- components of β extended out to Gx and Gy respectively + βx' = subst (λ a → D [ a , H ⟅ x ⟆ ]) (sym Gx≡G'x) (β ⟦ x ⟧) + βy' = subst (λ a → D [ a , H ⟅ y ⟆ ]) (sym Gy≡G'y) (β ⟦ y ⟧) + + -- extensions are equal to originals + βy'≡βy : PathP (λ i → D [ Gy≡G'y i , H ⟅ y ⟆ ]) βy' (β ⟦ y ⟧) + βy'≡βy = symP (toPathP {A = λ i → D [ Gy≡G'y (~ i) , H ⟅ y ⟆ ]} refl) + + βx≡βx' : PathP (λ i → D [ Gx≡G'x (~ i) , H ⟅ x ⟆ ]) (β ⟦ x ⟧) βx' + βx≡βx' = toPathP refl + + -- left wall of square + left : PathP (λ i → D [ Gx≡G'x i , H ⟅ y ⟆ ]) (G ⟪ f ⟫ ⋆⟨ D ⟩ βy') (G' ⟪ f ⟫ ⋆⟨ D ⟩ β ⟦ y ⟧) + left i = Gf≡G'f i ⋆⟨ D ⟩ βy'≡βy i + + -- right wall of square + right : PathP (λ i → D [ Gx≡G'x (~ i) , H ⟅ y ⟆ ]) (β ⟦ x ⟧ ⋆⟨ D ⟩ H ⟪ f ⟫) (βx' ⋆⟨ D ⟩ H ⟪ f ⟫) + right i = βx≡βx' i ⋆⟨ D ⟩ refl {x = H ⟪ f ⟫} i + + -- putting it all together + βSq : G ⟪ f ⟫ ⋆⟨ D ⟩ βy' ≡ βx' ⋆⟨ D ⟩ H ⟪ f ⟫ + βSq i = comp (λ k → D [ Gx≡G'x (~ k) , H ⟅ y ⟆ ]) + (λ j → λ { (i = i0) → left (~ j) + ; (i = i1) → right j }) + (β .N-hom f i) + + + module _ {F G : Functor C D} {α β : NatTrans F G} where + open Category + open Functor + open NatTrans + + makeNatTransPath : α .N-ob ≡ β .N-ob → α ≡ β + makeNatTransPath p i .N-ob = p i + makeNatTransPath p i .N-hom f = rem i + where + rem : PathP (λ i → (F .F-hom f) ⋆ᴰ (p i _) ≡ (p i _) ⋆ᴰ (G .F-hom f)) + (α .N-hom f) (β .N-hom f) + rem = toPathP (D .isSetHom _ _ _ _) + + module _ {F F' G G' : Functor C D} {α : NatTrans F G} {β : NatTrans F' G'} where + open Category + open Functor + open NatTrans + makeNatTransPathP : ∀ (p : F ≡ F') (q : G ≡ G') + → PathP (λ i → (x : C .ob) → D [ (p i) .F-ob x , (q i) .F-ob x ]) + (α .N-ob) (β .N-ob) + → PathP (λ i → NatTrans (p i) (q i)) α β + makeNatTransPathP p q P i .N-ob = P i + makeNatTransPathP p q P i .N-hom f = rem i + where + rem : PathP (λ i → ((p i) .F-hom f) ⋆ᴰ (P i _) ≡ (P i _) ⋆ᴰ ((q i) .F-hom f)) + (α .N-hom f) (β .N-hom f) + rem = toPathP (D .isSetHom _ _ _ _) + +module _ {B : Category ℓB ℓB'} {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where + open NatTrans + -- whiskering + -- αF + _∘ˡ_ : ∀ {G H : Functor C D} (α : NatTrans G H) → (F : Functor B C) + → NatTrans (G ∘F F) (H ∘F F) + (_∘ˡ_ {G} {H} α F) .N-ob x = α ⟦ F ⟅ x ⟆ ⟧ + (_∘ˡ_ {G} {H} α F) .N-hom f = (α .N-hom) _ + + -- Kβ + _∘ʳ_ : ∀ (K : Functor C D) → {G H : Functor B C} (β : NatTrans G H) + → NatTrans (K ∘F G) (K ∘F H) + (_∘ʳ_ K {G} {H} β) .N-ob x = K ⟪ β ⟦ x ⟧ ⟫ + (_∘ʳ_ K {G} {H} β) .N-hom f = preserveCommF {C = C} {D = D} {K} (β .N-hom f) diff --git a/Cubical/Categories/NaturalTransformation/Properties.agda b/Cubical/Categories/NaturalTransformation/Properties.agda new file mode 100644 index 0000000000..a79e1cc619 --- /dev/null +++ b/Cubical/Categories/NaturalTransformation/Properties.agda @@ -0,0 +1,102 @@ + +{-# OPTIONS --safe #-} + +module Cubical.Categories.NaturalTransformation.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism renaming (iso to iIso) +open import Cubical.Data.Sigma +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Morphism renaming (isIso to isIsoC) +open import Cubical.Categories.NaturalTransformation.Base + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +open isIsoC +open NatIso +open NatTrans +open Category +open Functor +open Iso + +module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where + private + _⋆ᴰ_ : ∀ {x y z} (f : D [ x , y ]) (g : D [ y , z ]) → D [ x , z ] + f ⋆ᴰ g = f ⋆⟨ D ⟩ g + + -- natural isomorphism is symmetric + symNatIso : ∀ {F G : Functor C D} + → F ≅ᶜ G + → G ≅ᶜ F + symNatIso η .trans .N-ob x = η .nIso x .inv + symNatIso η .trans .N-hom _ = sqLL η + symNatIso η .nIso x .inv = η .trans .N-ob x + symNatIso η .nIso x .sec = η .nIso x .ret + symNatIso η .nIso x .ret = η .nIso x .sec + + -- Properties + + -- path helpers + module NatTransP where + + module _ {F G : Functor C D} where + + -- same as Sigma version + NatTransΣ : Type (ℓ-max (ℓ-max ℓC ℓC') ℓD') + NatTransΣ = Σ[ ob ∈ ((x : C .ob) → D [(F .F-ob x) , (G .F-ob x)]) ] + ({x y : _ } (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (ob y) ≡ (ob x) ⋆ᴰ (G .F-hom f)) + + NatTransIsoΣ : Iso (NatTrans F G) NatTransΣ + NatTransIsoΣ .fun (natTrans N-ob N-hom) = N-ob , N-hom + NatTransIsoΣ .inv (N-ob , N-hom) = (natTrans N-ob N-hom) + NatTransIsoΣ .rightInv _ = refl + NatTransIsoΣ .leftInv _ = refl + + NatTrans≡Σ : NatTrans F G ≡ NatTransΣ + NatTrans≡Σ = isoToPath NatTransIsoΣ + + -- introducing paths + NatTrans-≡-intro : ∀ {αo βo : N-ob-Type F G} + {αh : N-hom-Type F G αo} + {βh : N-hom-Type F G βo} + → (p : αo ≡ βo) + → PathP (λ i → ({x y : C .ob} (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (p i y) ≡ (p i x) ⋆ᴰ (G .F-hom f))) αh βh + → natTrans {F = F} {G} αo αh ≡ natTrans βo βh + NatTrans-≡-intro p q i = natTrans (p i) (q i) + + module _ {F G : Functor C D} {α β : NatTrans F G} where + open Iso + private + αOb = α .N-ob + βOb = β .N-ob + αHom = α .N-hom + βHom = β .N-hom + -- path between natural transformations is the same as a pair of paths (between ob and hom) + NTPathIsoPathΣ : Iso (α ≡ β) + (Σ[ p ∈ (αOb ≡ βOb) ] + (PathP (λ i → ({x y : _} (f : _) → F ⟪ f ⟫ ⋆ᴰ (p i y) ≡ (p i x) ⋆ᴰ G ⟪ f ⟫)) + αHom + βHom)) + NTPathIsoPathΣ .fun p = (λ i → p i .N-ob) , (λ i → p i .N-hom) + NTPathIsoPathΣ .inv (po , ph) i = record { N-ob = po i ; N-hom = ph i } + NTPathIsoPathΣ .rightInv pσ = refl + NTPathIsoPathΣ .leftInv p = refl + + NTPath≃PathΣ = isoToEquiv NTPathIsoPathΣ + + NTPath≡PathΣ = ua NTPath≃PathΣ + + module _ where + open NatTransP + + -- if the target category has hom Sets, then any natural transformation is a set + isSetNatTrans : {F G : Functor C D} → isSet (NatTrans F G) + isSetNatTrans = + isSetRetract (fun NatTransIsoΣ) (inv NatTransIsoΣ) (leftInv NatTransIsoΣ) + (isSetΣSndProp (isSetΠ (λ _ → isSetHom D)) + (λ _ → isPropImplicitΠ2 (λ _ _ → isPropΠ (λ _ → isSetHom D _ _)))) diff --git a/Cubical/Categories/Presheaf.agda b/Cubical/Categories/Presheaf.agda new file mode 100644 index 0000000000..4fe1dc15c2 --- /dev/null +++ b/Cubical/Categories/Presheaf.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --postfix-projections --safe #-} + +module Cubical.Categories.Presheaf where + +open import Cubical.Categories.Presheaf.Base public +open import Cubical.Categories.Presheaf.KanExtension public +open import Cubical.Categories.Presheaf.Properties public diff --git a/Cubical/Categories/Presheaf/Base.agda b/Cubical/Categories/Presheaf/Base.agda new file mode 100644 index 0000000000..c1b4e2d7f4 --- /dev/null +++ b/Cubical/Categories/Presheaf/Base.agda @@ -0,0 +1,13 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Presheaf.Base where + +open import Cubical.Foundations.Prelude + +open import Cubical.Categories.Category +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Instances.Functors + +PreShv : ∀ {ℓ ℓ'} → Category ℓ ℓ' → (ℓS : Level) + → Category (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓS)) + (ℓ-max (ℓ-max ℓ ℓ') ℓS) +PreShv C ℓS = FUNCTOR (C ^op) (SET ℓS) diff --git a/Cubical/Categories/Presheaf/KanExtension.agda b/Cubical/Categories/Presheaf/KanExtension.agda new file mode 100644 index 0000000000..2ec7392ae7 --- /dev/null +++ b/Cubical/Categories/Presheaf/KanExtension.agda @@ -0,0 +1,341 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} + +{- + Kan extension of a functor C → D to a functor PreShv C ℓ → PreShv D ℓ left or right adjoint to + precomposition. +-} + +module Cubical.Categories.Presheaf.KanExtension where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Functions.FunExtEquiv + +open import Cubical.HITs.SetQuotients + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Adjoint +open import Cubical.Categories.Presheaf.Base +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.Instances.Sets + +{- + Left Kan extension of a functor C → D to a functor PreShv C ℓ → PreShv D ℓ left adjoint to precomposition. +-} + +module Lan {ℓC ℓC' ℓD ℓD'} ℓS + {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} + (F : Functor C D) + where + + open Functor + open NatTrans + + private + module C = Category C + module D = Category D + + {- + We want the category SET ℓ we're mapping into to be large enough that the coend will take presheaves + Cᵒᵖ → Set ℓ to presheaves Dᵒᵖ → Set ℓ, otherwise we get no adjunction with precomposition. + So we must have ℓC,ℓC',ℓD' ≤ ℓ; the parameter ℓS allows ℓ to be larger than their maximum. + -} + ℓ = ℓ-max (ℓ-max (ℓ-max ℓC ℓC') ℓD') ℓS + + module _ (G : Functor (C ^op) (SET ℓ)) where + + -- Definition of the coend + + module _ (d : D.ob) where + + Raw : Type ℓ + Raw = Σ[ c ∈ C.ob ] Σ[ g ∈ D.Hom[ d , F ⟅ c ⟆ ] ] G .F-ob c .fst + + data _≈_ : (u v : Raw) → Type ℓ where + shift : {c c' : C.ob} (g : D.Hom[ d , F ⟅ c ⟆ ]) (f : C.Hom[ c , c' ]) (a : (G ⟅ c' ⟆) .fst) + → (c' , (g D.⋆ F ⟪ f ⟫) , a) ≈ (c , g , (G ⟪ f ⟫) a) + + Quo = Raw / _≈_ + + pattern shift/ g f a i = eq/ _ _ (shift g f a) i + + -- Action of Quo on arrows in D + + mapR : {d d' : D.ob} (h : D.Hom[ d' , d ]) → Quo d → Quo d' + mapR h [ c , g , a ] = [ c , h D.⋆ g , a ] + mapR h (shift/ g f a i) = + hcomp + (λ j → λ + { (i = i0) → [ _ , D.⋆Assoc h g (F ⟪ f ⟫) j , a ] + ; (i = i1) → [ _ , h D.⋆ g , (G ⟪ f ⟫) a ] + }) + (shift/ (h D.⋆ g) f a i) + mapR h (squash/ t u p q i j) = + squash/ (mapR h t) (mapR h u) (cong (mapR h) p) (cong (mapR h) q) i j + + mapRId : (d : D.ob) → mapR (D.id {x = d}) ≡ idfun (Quo d) + mapRId d = + funExt (elimProp (λ _ → squash/ _ _) (λ (c , g , a) i → [ c , D.⋆IdL g i , a ])) + + mapR∘ : {d d' d'' : D.ob} + (h' : D.Hom[ d'' , d' ]) (h : D.Hom[ d' , d ]) + → mapR (h' D.⋆ h) ≡ mapR h' ∘ mapR h + mapR∘ h' h = + funExt (elimProp (λ _ → squash/ _ _) (λ (c , g , a) i → [ c , D.⋆Assoc h' h g i , a ])) + + LanOb : Functor (C ^op) (SET ℓ) → Functor (D ^op) (SET _) + LanOb G .F-ob d .fst = Quo G d + LanOb G .F-ob d .snd = squash/ + LanOb G .F-hom = mapR G + LanOb G .F-id {d} = mapRId G d + LanOb G .F-seq h h' = mapR∘ G h' h + + -- Action of Quo on arrows in Cᵒᵖ → Set + + module _ {G G' : Functor (C ^op) (SET ℓ)} (α : NatTrans G G') where + + mapL : (d : D.ob) → Quo G d → Quo G' d + mapL d [ c , g , a ] = [ c , g , α .N-ob c a ] + mapL d (shift/ g f a i) = + hcomp + (λ j → λ + { (i = i0) → [ _ , (g D.⋆ F ⟪ f ⟫) , α .N-ob _ a ] + ; (i = i1) → [ _ , g , funExt⁻ (α .N-hom f) a (~ j) ] + }) + (shift/ g f ((α ⟦ _ ⟧) a) i) + mapL d (squash/ t u p q i j) = + squash/ (mapL d t) (mapL d u) (cong (mapL d) p) (cong (mapL d) q) i j + + mapLR : {d d' : D.ob} (h : D.Hom[ d' , d ]) + → mapL d' ∘ mapR G h ≡ mapR G' h ∘ mapL d + mapLR h = funExt (elimProp (λ _ → squash/ _ _) (λ _ → refl)) + + mapLId : (G : Functor (C ^op) (SET ℓ)) + (d : D.ob) → mapL (idTrans G) d ≡ idfun (Quo G d) + mapLId G d = funExt (elimProp (λ _ → squash/ _ _) (λ _ → refl)) + + mapL∘ : {G G' G'' : Functor (C ^op) (SET ℓ)} + (β : NatTrans G' G'') (α : NatTrans G G') + (d : D.ob) → mapL (seqTrans α β) d ≡ mapL β d ∘ mapL α d + mapL∘ β α d = funExt (elimProp (λ _ → squash/ _ _) (λ _ → refl)) + + LanHom : {G G' : Functor (C ^op) (SET ℓ)} + → NatTrans G G' → NatTrans (LanOb G) (LanOb G') + LanHom α .N-ob = mapL α + LanHom α .N-hom = mapLR α + + -- Definition of the left Kan extension functor + + Lan : Functor (FUNCTOR (C ^op) (SET ℓ)) (FUNCTOR (D ^op) (SET ℓ)) + Lan .F-ob = LanOb + Lan .F-hom = LanHom + Lan .F-id {G} = makeNatTransPath (funExt (mapLId G)) + Lan .F-seq α β = makeNatTransPath (funExt (mapL∘ β α)) + + -- Adjunction between the left Kan extension and precomposition + + private + F* = precomposeF (SET ℓ) (F ^opF) + + open UnitCounit + + η : 𝟙⟨ FUNCTOR (C ^op) (SET ℓ) ⟩ ⇒ funcComp F* Lan + η .N-ob G .N-ob c a = [ c , D.id , a ] + η .N-ob G .N-hom {c'} {c} f = + funExt λ a → + [ c , D.id , (G ⟪ f ⟫) a ] + ≡⟨ sym (shift/ D.id f a) ⟩ + [ c' , (D.id D.⋆ F ⟪ f ⟫) , a ] + ≡[ i ]⟨ [ c' , lem i , a ] ⟩ + [ c' , (F ⟪ f ⟫ D.⋆ D.id) , a ] + ∎ + where + lem : D.id D.⋆ F ⟪ f ⟫ ≡ F ⟪ f ⟫ D.⋆ D.id + lem = D.⋆IdL (F ⟪ f ⟫) ∙ sym (D.⋆IdR (F ⟪ f ⟫)) + η .N-hom f = makeNatTransPath refl + + ε : funcComp Lan F* ⇒ 𝟙⟨ FUNCTOR (D ^op) (SET ℓ) ⟩ + ε .N-ob H .N-ob d = + elim + (λ _ → (H ⟅ d ⟆) .snd) + (λ (c , g , a) → (H ⟪ g ⟫) a) + (λ {_ _ (shift g f a) i → H .F-seq (F ⟪ f ⟫) g i a}) + ε .N-ob H .N-hom g' = + funExt (elimProp (λ _ → (H ⟅ _ ⟆) .snd _ _) (λ (c , g , a) → funExt⁻ (H .F-seq g g') a)) + ε .N-hom {H} {H'} α = + makeNatTransPath + (funExt₂ λ d → + elimProp (λ _ → (H' ⟅ _ ⟆) .snd _ _) + (λ (c , g , a) → sym (funExt⁻ (α .N-hom g) a))) + + Δ₁ : ∀ G → seqTrans (Lan ⟪ η ⟦ G ⟧ ⟫) (ε ⟦ Lan ⟅ G ⟆ ⟧) ≡ idTrans _ + Δ₁ G = + makeNatTransPath + (funExt₂ λ d → + elimProp (λ _ → squash/ _ _) + (λ (c , g , a) → + [ c , g D.⋆ D.id , a ] + ≡[ i ]⟨ [ c , (g D.⋆ F .F-id (~ i)) , a ] ⟩ + [ c , g D.⋆ (F ⟪ C.id ⟫) , a ] + ≡⟨ shift/ g C.id a ⟩ + [ c , g , (G ⟪ C.id ⟫) a ] + ≡[ i ]⟨ [ c , g , G .F-id i a ] ⟩ + [ c , g , a ] + ∎)) + + Δ₂ : ∀ H → seqTrans (η ⟦ F* ⟅ H ⟆ ⟧) (F* ⟪ ε ⟦ H ⟧ ⟫) ≡ idTrans _ + Δ₂ H = makeNatTransPath (funExt λ c → H .F-id) + + adj : Lan ⊣ F* + adj = make⊣ η ε Δ₁ Δ₂ + +{- + Right Kan extension of a functor C → D to a functor PreShv C ℓ → PreShv D ℓ right adjoint to precomposition. +-} + +module Ran {ℓC ℓC' ℓD ℓD'} ℓS + {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} + (F : Functor C D) + where + + open Functor + open NatTrans + + private + module C = Category C + module D = Category D + + {- + We want the category SET ℓ we're mapping into to be large enough that the coend will take presheaves + Cᵒᵖ → Set ℓ to presheaves Dᵒᵖ → Set ℓ, otherwise we get no adjunction with precomposition. + So we must have ℓC,ℓC',ℓD' ≤ ℓ; the parameter ℓS allows ℓ to be larger than their maximum. + -} + ℓ = ℓ-max (ℓ-max (ℓ-max ℓC ℓC') ℓD') ℓS + + module _ (G : Functor (C ^op) (SET ℓ)) where + + -- Definition of the end + + record End (d : D.ob) : Type ℓ where + field + fun : (c : C.ob) (g : D.Hom[ F ⟅ c ⟆ , d ]) → G .F-ob c .fst + coh : {c c' : C.ob} (f : C.Hom[ c , c' ]) (g : D.Hom[ F ⟅ c' ⟆ , d ]) + → fun c (F ⟪ f ⟫ ⋆⟨ D ⟩ g) ≡ (G ⟪ f ⟫) (fun c' g) + + open End + + end≡ : {d : D.ob} {x x' : End d} → (∀ c g → x .fun c g ≡ x' .fun c g) → x ≡ x' + end≡ h i .fun c g = h c g i + end≡ {_} {x} {x'} h i .coh f g = + isSet→isSet' (G .F-ob _ .snd) + (x .coh f g) + (x' .coh f g) + (h _ (F ⟪ f ⟫ ⋆⟨ D ⟩ g)) + (cong (G ⟪ f ⟫) (h _ g)) + i + + -- Action of End on arrows in D + + mapR : {d d' : D.ob} (h : D.Hom[ d' , d ]) → End d → End d' + mapR h x .fun c g = x .fun c (g ⋆⟨ D ⟩ h) + mapR h x .coh f g = cong (x .fun _) (D.⋆Assoc (F ⟪ f ⟫) g h) ∙ x .coh f (g ⋆⟨ D ⟩ h) + + mapRId : (d : D.ob) → mapR (D.id {x = d}) ≡ idfun (End d) + mapRId h = funExt λ x → end≡ λ c g → cong (x .fun c) (D.⋆IdR g) + + mapR∘ : {d d' d'' : D.ob} + (h' : D.Hom[ d'' , d' ]) (h : D.Hom[ d' , d ]) + → mapR (h' D.⋆ h) ≡ mapR h' ∘ mapR h + mapR∘ h' h = funExt λ x → end≡ λ c g → cong (x .fun c) (sym (D.⋆Assoc g h' h)) + + open End + + RanOb : Functor (C ^op) (SET ℓ) → Functor (D ^op) (SET _) + RanOb G .F-ob d .fst = End G d + RanOb G .F-ob d .snd = + -- We use that End is equivalent to a Σ-type to prove its HLevel more easily + isOfHLevelRetract 2 + {B = + Σ[ z ∈ ((c : C.ob) (g : D.Hom[ F ⟅ c ⟆ , d ]) → G .F-ob c .fst) ] + ({c c' : C.ob} (f : C.Hom[ c , c' ]) (g : D.Hom[ F ⟅ c' ⟆ , d ]) + → z c (F ⟪ f ⟫ ⋆⟨ D ⟩ g) ≡ (G ⟪ f ⟫) (z c' g)) + } + (λ x → λ where .fst → x .fun; .snd → x .coh) + (λ σ → λ where .fun → σ .fst; .coh → σ .snd) + (λ _ → refl) + (isSetΣ + (isSetΠ2 λ _ _ → G .F-ob _ .snd) + (λ _ → isProp→isSet + (isPropImplicitΠ λ _ → isPropImplicitΠ λ _ → isPropΠ2 λ _ _ → G .F-ob _ .snd _ _))) + RanOb G .F-hom = mapR G + RanOb G .F-id {d} = mapRId G d + RanOb G .F-seq h h' = mapR∘ G h' h + + -- Action of End on arrows in Cᵒᵖ → Set + + module _ {G G' : Functor (C ^op) (SET ℓ)} (α : NatTrans G G') where + + mapL : (d : D.ob) → End G d → End G' d + mapL d x .fun c g = (α ⟦ c ⟧) (x .fun c g) + mapL d x .coh f g = + cong (α ⟦ _ ⟧) (x .coh f g) + ∙ funExt⁻ (α .N-hom f) (x .fun _ g) + + mapLR : {d d' : D.ob} (h : D.Hom[ d' , d ]) + → mapL d' ∘ mapR G h ≡ mapR G' h ∘ mapL d + mapLR h = funExt λ _ → end≡ _ λ _ _ → refl + + mapLId : (G : Functor (C ^op) (SET ℓ)) + (d : D.ob) → mapL (idTrans G) d ≡ idfun (End G d) + mapLId G d = funExt λ _ → end≡ _ λ _ _ → refl + + mapL∘ : {G G' G'' : Functor (C ^op) (SET ℓ)} + (β : NatTrans G' G'') (α : NatTrans G G') + (d : D.ob) → mapL (seqTrans α β) d ≡ mapL β d ∘ mapL α d + mapL∘ β α d = funExt λ _ → end≡ _ λ _ _ → refl + + RanHom : {G G' : Functor (C ^op) (SET ℓ)} + → NatTrans G G' → NatTrans (RanOb G) (RanOb G') + RanHom α .N-ob = mapL α + RanHom α .N-hom = mapLR α + + -- Definition of the right Kan extension functor + + Ran : Functor (FUNCTOR (C ^op) (SET ℓ)) (FUNCTOR (D ^op) (SET ℓ)) + Ran .F-ob = RanOb + Ran .F-hom = RanHom + Ran .F-id {G} = makeNatTransPath (funExt (mapLId G)) + Ran .F-seq α β = makeNatTransPath (funExt (mapL∘ β α)) + + -- Adjunction between precomposition and right Kan extension + + private + F* = precomposeF (SET ℓ) (F ^opF) + + open UnitCounit + + η : 𝟙⟨ FUNCTOR (D ^op) (SET ℓ) ⟩ ⇒ (funcComp Ran F*) + η .N-ob G .N-ob d a .fun c g = (G ⟪ g ⟫) a + η .N-ob G .N-ob d a .coh f g = funExt⁻ (G .F-seq g (F ⟪ f ⟫)) a + η .N-ob G .N-hom h = funExt λ a → end≡ _ λ c g → sym (funExt⁻ (G .F-seq h g) a) + η .N-hom {G} {G'} α = + makeNatTransPath (funExt₂ λ d a → end≡ _ λ c g → sym (funExt⁻ (α .N-hom g) a)) + + ε : funcComp F* Ran ⇒ 𝟙⟨ FUNCTOR (C ^op) (SET ℓ) ⟩ + ε .N-ob H .N-ob c x = x .fun c D.id + ε .N-ob H .N-hom {c} {c'} g = + funExt λ x → + cong (x .fun c') (D.⋆IdL _ ∙ sym (D.⋆IdR _)) ∙ x .coh g D.id + ε .N-hom {H} {H'} α = makeNatTransPath refl + + Δ₁ : ∀ G → seqTrans (F* ⟪ η ⟦ G ⟧ ⟫) (ε ⟦ F* ⟅ G ⟆ ⟧) ≡ idTrans _ + Δ₁ G = makeNatTransPath (funExt₂ λ c a → funExt⁻ (G .F-id) a) + + Δ₂ : ∀ H → seqTrans (η ⟦ Ran ⟅ H ⟆ ⟧) (Ran ⟪ ε ⟦ H ⟧ ⟫) ≡ idTrans _ + Δ₂ H = makeNatTransPath (funExt₂ λ c x → end≡ _ λ c' g → cong (x .fun c') (D.⋆IdL g)) + + adj : F* ⊣ Ran + adj = make⊣ η ε Δ₁ Δ₂ diff --git a/Cubical/Categories/Presheaf/Properties.agda b/Cubical/Categories/Presheaf/Properties.agda new file mode 100644 index 0000000000..5c0833a1e6 --- /dev/null +++ b/Cubical/Categories/Presheaf/Properties.agda @@ -0,0 +1,389 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Presheaf.Properties where + +open import Cubical.Categories.Category +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.Functor +open import Cubical.Categories.Presheaf.Base +open import Cubical.Categories.Equivalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv using (fiber) +open import Cubical.Data.Sigma + +import Cubical.Categories.Morphism as Morphism +import Cubical.Categories.Constructions.Slice as Slice +import Cubical.Categories.Constructions.Elements as Elements +import Cubical.Functions.Fibration as Fibration + +private + variable + ℓ ℓ' : Level + e e' : Level + + +-- (PreShv C) / F ≃ᶜ PreShv (∫ᴾ F) +module _ {ℓS : Level} (C : Category ℓ ℓ') (F : Functor (C ^op) (SET ℓS)) where + open Category + open Functor + open _≃ᶜ_ + open isEquivalence + open NatTrans + open NatIso + open Slice (PreShv C ℓS) F + open Elements {C = C} + + open Fibration.ForSets + + -- specific case of fiber under natural transformation + fibersEqIfRepsEqNatTrans : ∀ {A} (ϕ : A ⇒ F) {c x x'} {px : x ≡ x'} {a' : fiber (ϕ ⟦ c ⟧) x} {b' : fiber (ϕ ⟦ c ⟧) x'} + → fst a' ≡ fst b' + → PathP (λ i → fiber (ϕ ⟦ c ⟧) (px i)) a' b' + fibersEqIfRepsEqNatTrans ϕ {c} {x} {x'} {px} {a , fiba} {b , fibb} p + = fibersEqIfRepsEq {isSetB = snd (F ⟅ c ⟆)} (ϕ ⟦ c ⟧) p + + -- ======================================== + -- K : Slice → PreShv + -- ======================================== + + -- action on (slice) objects + K-ob : (s : SliceCat .ob) → (PreShv (∫ᴾ F) ℓS .ob) + -- we take (c , x) to the fiber in A of ϕ over x + K-ob (sliceob {A} ϕ) .F-ob (c , x) + = (fiber (ϕ ⟦ c ⟧) x) + , isOfHLevelΣ 2 (snd (A ⟅ c ⟆)) λ _ → isSet→isGroupoid (snd (F ⟅ c ⟆)) _ _ + -- for morphisms, we just apply A ⟪ h ⟫ (plus equality proof) + K-ob (sliceob {A} ϕ) .F-hom {d , y} {c , x} (h , com) (b , eq) + = ((A ⟪ h ⟫) b) + , ((ϕ ⟦ c ⟧) ((A ⟪ h ⟫) b) + ≡[ i ]⟨ (ϕ .N-hom h) i b ⟩ + (F ⟪ h ⟫) ((ϕ ⟦ d ⟧) b) + ≡[ i ]⟨ (F ⟪ h ⟫) (eq i) ⟩ + (F ⟪ h ⟫) y + ≡⟨ sym com ⟩ + x + ∎) + -- functoriality follows from functoriality of A + K-ob (sliceob {A} ϕ) .F-id {x = (c , x)} + = funExt λ { (a , fibp) + → fibersEqIfRepsEqNatTrans ϕ (λ i → A .F-id i a) } + K-ob (sliceob {A} ϕ) .F-seq {x = (c , x)} {(d , y)} {(e , z)} (f' , eq1) (g' , eq2) + = funExt λ { ( a , fibp ) + → fibersEqIfRepsEqNatTrans ϕ (λ i → (A .F-seq f' g') i a) } + + + -- action on morphisms (in this case, natural transformation) + K-hom : {sA sB : SliceCat .ob} + → (ε : SliceCat [ sA , sB ]) + → (K-ob sA) ⇒ (K-ob sB) + K-hom {sA = s1@(sliceob {A} ϕ)} {s2@(sliceob {B} ψ)} (slicehom ε com) = natTrans η-ob (λ h → funExt (η-hom h)) + where + P = K-ob s1 + Q = K-ob s2 + + -- just apply the natural transformation (ε) we're given + -- this ensures that we stay in the fiber over x due to the commutativity given by slicenesss + η-ob : (el : (∫ᴾ F) .ob) → (fst (P ⟅ el ⟆) → fst (Q ⟅ el ⟆) ) + η-ob (c , x) (a , ϕa≡x) = ((ε ⟦ c ⟧) a) , εψ≡ϕ ∙ ϕa≡x + where + εψ≡ϕ : (ψ ⟦ c ⟧) ((ε ⟦ c ⟧) a) ≡ (ϕ ⟦ c ⟧) a + εψ≡ϕ i = ((com i) ⟦ c ⟧) a + + η-hom : ∀ {el1 el2} (h : (∫ᴾ F) [ el1 , el2 ]) (ae : fst (P ⟅ el2 ⟆)) → η-ob el1 ((P ⟪ h ⟫) ae) ≡ (Q ⟪ h ⟫) (η-ob el2 ae) + η-hom {el1 = (c , x)} {d , y} (h , eqh) (a , eqa) + = fibersEqIfRepsEqNatTrans ψ (λ i → ε .N-hom h i a) + + + K : Functor SliceCat (PreShv (∫ᴾ F) ℓS) + K .F-ob = K-ob + K .F-hom = K-hom + K .F-id = makeNatTransPath + (funExt λ cx@(c , x) + → funExt λ aeq@(a , eq) + → fibersEqIfRepsEq {isSetB = snd (F ⟅ c ⟆)} _ refl) + K .F-seq (slicehom α eqa) (slicehom β eqb) + = makeNatTransPath + (funExt λ cx@(c , x) + → funExt λ aeq@(a , eq) + → fibersEqIfRepsEq {isSetB = snd (F ⟅ c ⟆)} _ refl) + + + -- ======================================== + -- L : PreShv → Slice + -- ======================================== + + -- action on objects (presheaves) + L-ob : (P : PreShv (∫ᴾ F) ℓS .ob) + → SliceCat .ob + L-ob P = sliceob {S-ob = L-ob-ob} L-ob-hom + where + -- sends c to the disjoint union of all the images under P + LF-ob : (c : C .ob) → (SET _) .ob + LF-ob c = (Σ[ x ∈ fst (F ⟅ c ⟆) ] fst (P ⟅ c , x ⟆)) , isSetΣ (snd (F ⟅ c ⟆)) (λ x → snd (P ⟅ c , x ⟆)) + + -- defines a function piecewise over the fibers by applying P + LF-hom : ∀ {x y} + → (f : C [ y , x ]) + → (SET _) [ LF-ob x , LF-ob y ] + LF-hom {x = c} {d} f (x , a) = ((F ⟪ f ⟫) x) , (P ⟪ f , refl ⟫) a + + L-ob-ob : Functor (C ^op) (SET _) + L-ob-ob .F-ob = LF-ob + L-ob-ob .F-hom = LF-hom + L-ob-ob .F-id {x = c} + = funExt idFunExt + where + idFunExt : ∀ (un : fst (LF-ob c)) + → (LF-hom (C .id) un) ≡ un + idFunExt (x , X) = ΣPathP (leftEq , rightEq) + where + leftEq : (F ⟪ C .id ⟫) x ≡ x + leftEq i = F .F-id i x + + rightEq : PathP (λ i → fst (P ⟅ c , leftEq i ⟆)) + ((P ⟪ C .id , refl ⟫) X) X + rightEq = left ▷ right + where + -- the id morphism in (∫ᴾ F) + ∫id = C .id , sym (funExt⁻ (F .F-id) x ∙ refl) + + -- functoriality of P gives us close to what we want + right : (P ⟪ ∫id ⟫) X ≡ X + right i = P .F-id i X + + -- but need to do more work to show that (C .id , refl) ≡ ∫id + left : PathP (λ i → fst (P ⟅ c , leftEq i ⟆)) + ((P ⟪ C .id , refl ⟫) X) + ((P ⟪ ∫id ⟫) X) + left i = (P ⟪ ∫ᴾhomEq {F = F} (C .id , refl) ∫id (λ i → (c , leftEq i)) refl refl i ⟫) X + L-ob-ob .F-seq {x = c} {d} {e} f g + = funExt seqFunEq + where + seqFunEq : ∀ (un : fst (LF-ob c)) + → (LF-hom (g ⋆⟨ C ⟩ f) un) ≡ (LF-hom g) (LF-hom f un) + seqFunEq un@(x , X) = ΣPathP (leftEq , rightEq) + where + -- the left component is comparing the action of F on x + -- equality follows from functoriality of F + -- leftEq : fst (LF-hom (g ⋆⟨ C ⟩ f) un) ≡ fst ((LF-hom g) (LF-hom f un)) + leftEq : (F ⟪ g ⋆⟨ C ⟩ f ⟫) x ≡ (F ⟪ g ⟫) ((F ⟪ f ⟫) x) + leftEq i = F .F-seq f g i x + + -- on the right, equality also follows from functoriality of P + -- but it's more complicated because of heterogeneity + -- since leftEq is not a definitional equality + rightEq : PathP (λ i → fst (P ⟅ e , leftEq i ⟆)) + ((P ⟪ g ⋆⟨ C ⟩ f , refl ⟫) X) + ((P ⟪ g , refl ⟫) ((P ⟪ f , refl ⟫) X)) + rightEq = left ▷ right + where + -- functoriality of P only gets us to this weird composition on the left + right : (P ⟪ (g , refl) ⋆⟨ (∫ᴾ F) ⟩ (f , refl) ⟫) X ≡ (P ⟪ g , refl ⟫) ((P ⟪ f , refl ⟫) X) + right i = P .F-seq (f , refl) (g , refl) i X + + -- so we need to show that this composition is actually equal to the one we want + left : PathP (λ i → fst (P ⟅ e , leftEq i ⟆)) + ((P ⟪ g ⋆⟨ C ⟩ f , refl ⟫) X) + ((P ⟪ (g , refl) ⋆⟨ (∫ᴾ F) ⟩ (f , refl) ⟫) X) + left i = (P ⟪ ∫ᴾhomEq {F = F} (g ⋆⟨ C ⟩ f , refl) ((g , refl) ⋆⟨ (∫ᴾ F) ⟩ (f , refl)) (λ i → (e , leftEq i)) refl refl i ⟫) X + L-ob-hom : L-ob-ob ⇒ F + L-ob-hom .N-ob c (x , _) = x + L-ob-hom .N-hom f = funExt λ (x , _) → refl + + -- action on morphisms (aka natural transformations between presheaves) + -- is essentially the identity (plus equality proofs for naturality and slice commutativity) + L-hom : ∀ {P Q} → PreShv (∫ᴾ F) ℓS [ P , Q ] → + SliceCat [ L-ob P , L-ob Q ] + L-hom {P} {Q} η = slicehom arr com + where + A = S-ob (L-ob P) + ϕ = S-arr (L-ob P) + B = S-ob (L-ob Q) + ψ = S-arr (L-ob Q) + arr : A ⇒ B + arr .N-ob c (x , X) = x , ((η ⟦ c , x ⟧) X) + arr .N-hom {c} {d} f = funExt natu + where + natuType : fst (A ⟅ c ⟆) → Type _ + natuType xX@(x , X) = ((F ⟪ f ⟫) x , (η ⟦ d , (F ⟪ f ⟫) x ⟧) ((P ⟪ f , refl ⟫) X)) ≡ ((F ⟪ f ⟫) x , (Q ⟪ f , refl ⟫) ((η ⟦ c , x ⟧) X)) + natu : ∀ (xX : fst (A ⟅ c ⟆)) → natuType xX + natu (x , X) = ΣPathP (refl , λ i → (η .N-hom (f , refl) i) X) + + com : arr ⋆⟨ PreShv C ℓS ⟩ ψ ≡ ϕ + com = makeNatTransPath (funExt comFunExt) + where + comFunExt : ∀ (c : C .ob) + → (arr ●ᵛ ψ) ⟦ c ⟧ ≡ ϕ ⟦ c ⟧ + comFunExt c = funExt λ x → refl + + L : Functor (PreShv (∫ᴾ F) ℓS) SliceCat + L .F-ob = L-ob + L .F-hom = L-hom + L .F-id {cx} = SliceHom-≡-intro' (makeNatTransPath (funExt λ c → refl)) + L .F-seq {cx} {dy} P Q = SliceHom-≡-intro' (makeNatTransPath (funExt λ c → refl)) + + -- ======================================== + -- η : 𝟙 ≅ LK + -- ======================================== + + module _ where + open Iso + open Morphism renaming (isIso to isIsoC) + -- the iso we need + -- a type is isomorphic to the disjoint union of all its fibers + typeSectionIso : ∀ {A B : Type ℓS} {isSetB : isSet B} → (ϕ : A → B) + → Iso A (Σ[ b ∈ B ] fiber ϕ b) + typeSectionIso ϕ .fun a = (ϕ a) , (a , refl) + typeSectionIso ϕ .inv (b , (a , eq)) = a + typeSectionIso {isSetB = isSetB} ϕ .rightInv (b , (a , eq)) + = ΣPathP (eq + , ΣPathP (refl + , isOfHLevel→isOfHLevelDep 1 (λ b' → isSetB _ _) refl eq eq)) + typeSectionIso ϕ .leftInv a = refl + + -- the natural transformation + -- just applies typeSectionIso + ηTrans : 𝟙⟨ SliceCat ⟩ ⇒ (L ∘F K) + ηTrans .N-ob sob@(sliceob {A} ϕ) = slicehom A⇒LK comm + where + LKA = S-ob (L ⟅ K ⟅ sob ⟆ ⟆) + ψ = S-arr (L ⟅ K ⟅ sob ⟆ ⟆) + + A⇒LK : A ⇒ LKA + A⇒LK .N-ob c = typeSectionIso {isSetB = snd (F ⟅ c ⟆)} (ϕ ⟦ c ⟧) .fun + A⇒LK .N-hom {c} {d} f = funExt homFunExt + where + homFunExt : (x : fst (A ⟅ c ⟆)) + → (((ϕ ⟦ d ⟧) ((A ⟪ f ⟫) x)) , ((A ⟪ f ⟫) x , refl)) ≡ ((F ⟪ f ⟫) ((ϕ ⟦ c ⟧) x) , (A ⟪ f ⟫) x , _) + homFunExt x = ΣPathP ((λ i → (ϕ .N-hom f i) x) , fibersEqIfRepsEqNatTrans ϕ refl) + + comm : (A⇒LK) ●ᵛ ψ ≡ ϕ + comm = makeNatTransPath (funExt λ x → refl) + ηTrans .N-hom {sliceob {A} α} {sliceob {B} β} (slicehom ϕ eq) + = SliceHom-≡-intro' (makeNatTransPath (funExt (λ c → funExt λ a → natFunExt c a))) + where + natFunExt : ∀ (c : C .ob) (a : fst (A ⟅ c ⟆)) + → ((β ⟦ c ⟧) ((ϕ ⟦ c ⟧) a) , (ϕ ⟦ c ⟧) a , _) ≡ ((α ⟦ c ⟧) a , (ϕ ⟦ c ⟧) a , _) + natFunExt c a = ΣPathP ((λ i → ((eq i) ⟦ c ⟧) a) , fibersEqIfRepsEqNatTrans β refl) + + -- isomorphism follows from typeSectionIso + ηIso : ∀ (sob : SliceCat .ob) + → isIsoC SliceCat (ηTrans ⟦ sob ⟧) + ηIso sob@(sliceob ϕ) = sliceIso _ _ (FUNCTORIso _ _ _ isIsoCf) + where + isIsoCf : ∀ (c : C .ob) + → isIsoC _ (ηTrans .N-ob sob .S-hom ⟦ c ⟧) + isIsoCf c = CatIso→isIso (Iso→CatIso (typeSectionIso {isSetB = snd (F ⟅ c ⟆)} (ϕ ⟦ c ⟧))) + + + -- ======================================== + -- ε : KL ≅ 𝟙 + -- ======================================== + + module _ where + open Iso + open Morphism renaming (isIso to isIsoC) + -- the iso we deserve + -- says that a type family at x is isomorphic to the fiber over x of that type family packaged up + typeFiberIso : ∀ {ℓ ℓ'} {A : Type ℓ} {isSetA : isSet A} {x} (B : A → Type ℓ') + → Iso (B x) (fiber {A = Σ[ a ∈ A ] B a} (λ (x , _) → x) x) + typeFiberIso {x = x} _ .fun b = (x , b) , refl + typeFiberIso _ .inv ((a , b) , eq) = subst _ eq b + typeFiberIso {isSetA = isSetA} {x = x} B .rightInv ((a , b) , eq) + = fibersEqIfRepsEq {isSetB = isSetA} (λ (x , _) → x) (ΣPathP (sym eq , symP (transport-filler (λ i → B (eq i)) b))) + typeFiberIso {x = x} _ .leftInv b = sym (transport-filler refl b) + + -- the natural isomorphism + -- applies typeFiberIso (inv) + εTrans : (K ∘F L) ⇒ 𝟙⟨ PreShv (∫ᴾ F) ℓS ⟩ + εTrans .N-ob P = natTrans γ-ob (λ f → funExt (λ a → γ-homFunExt f a)) + where + KLP = K ⟅ L ⟅ P ⟆ ⟆ + + γ-ob : (el : (∫ᴾ F) .ob) + → (fst (KLP ⟅ el ⟆) → fst (P ⟅ el ⟆) ) + γ-ob el@(c , _) = typeFiberIso {isSetA = snd (F ⟅ c ⟆)} (λ x → fst (P ⟅ c , x ⟆)) .inv + + -- naturality + -- the annoying part is all the substs + γ-homFunExt : ∀ {el2 el1} → (f' : (∫ᴾ F) [ el2 , el1 ]) + → (∀ (a : fst (KLP ⟅ el1 ⟆)) → γ-ob el2 ((KLP ⟪ f' ⟫) a) ≡ (P ⟪ f' ⟫) (γ-ob el1 a)) + γ-homFunExt {d , y} {c , x} f'@(f , comm) a@((x' , X') , eq) i + = comp (λ j → fst (P ⟅ d , eq' j ⟆)) (λ j → λ { (i = i0) → left j + ; (i = i1) → right j }) ((P ⟪ f , refl ⟫) X') + where + -- fiber equality proof that we get from an application of KLP + eq' = snd ((KLP ⟪ f' ⟫) a) + + -- top right of the commuting diagram + -- "remove" the subst from the inside + right : PathP (λ i → fst (P ⟅ d , eq' i ⟆)) ((P ⟪ f , refl ⟫) X') ((P ⟪ f , comm ⟫) (subst _ eq X')) + right i = (P ⟪ f , refl≡comm i ⟫) (X'≡subst i) + where + refl≡comm : PathP (λ i → (eq' i) ≡ (F ⟪ f ⟫) (eq i)) refl comm + refl≡comm = isOfHLevel→isOfHLevelDep 1 (λ (v , w) → snd (F ⟅ d ⟆) v ((F ⟪ f ⟫) w)) refl comm λ i → (eq' i , eq i) + + X'≡subst : PathP (λ i → fst (P ⟅ c , eq i ⟆)) X' (subst _ eq X') + X'≡subst = transport-filler (λ i → fst (P ⟅ c , eq i ⟆)) X' + + -- bottom left of the commuting diagram + -- "remove" the subst from the outside + left : PathP (λ i → fst (P ⟅ d , eq' i ⟆)) ((P ⟪ f , refl ⟫) X') (subst (λ v → fst (P ⟅ d , v ⟆)) eq' ((P ⟪ f , refl ⟫) X')) + left = transport-filler (λ i → fst (P ⟅ d , eq' i ⟆)) ((P ⟪ f , refl ⟫) X') + εTrans .N-hom {P} {Q} α = makeNatTransPath (funExt λ cx → funExt λ xX' → ε-homFunExt cx xX') + where + KLP = K ⟅ L ⟅ P ⟆ ⟆ + + -- naturality of the above construction applies a similar argument as in `γ-homFunExt` + ε-homFunExt : ∀ (cx@(c , x) : (∫ᴾ F) .ob) (xX'@((x' , X') , eq) : fst (KLP ⟅ cx ⟆)) + → subst (λ v → fst (Q ⟅ c , v ⟆)) (snd ((K ⟪ L ⟪ α ⟫ ⟫ ⟦ cx ⟧) xX')) ((α ⟦ c , x' ⟧) X') + ≡ (α ⟦ c , x ⟧) (subst _ eq X') + ε-homFunExt cx@(c , x) xX'@((x' , X') , eq) i + = comp (λ j → fst (Q ⟅ c , eq j ⟆)) (λ j → λ { (i = i0) → left j + ; (i = i1) → right j }) ((α ⟦ c , x' ⟧) X') + where + eq' : x' ≡ x + eq' = snd ((K ⟪ L ⟪ α ⟫ ⟫ ⟦ cx ⟧) xX') + + right : PathP (λ i → fst (Q ⟅ c , eq i ⟆)) ((α ⟦ c , x' ⟧) X') ((α ⟦ c , x ⟧) (subst _ eq X')) + right i = (α ⟦ c , eq i ⟧) (X'≡subst i) + where + -- this is exactly the same as the one from before, can refactor? + X'≡subst : PathP (λ i → fst (P ⟅ c , eq i ⟆)) X' (subst _ eq X') + X'≡subst = transport-filler _ _ + + -- extracted out type since need to use in in 'left' body as well + leftTy : (x' ≡ x) → Type _ + leftTy eq* = PathP (λ i → fst (Q ⟅ c , eq* i ⟆)) ((α ⟦ c , x' ⟧) X') (subst (λ v → fst (Q ⟅ c , v ⟆)) eq' ((α ⟦ c , x' ⟧) X')) + + left : leftTy eq + left = subst + (λ eq* → leftTy eq*) + eq'≡eq + (transport-filler _ _) + where + eq'≡eq : eq' ≡ eq + eq'≡eq = snd (F ⟅ c ⟆) _ _ eq' eq + + εIso : ∀ (P : PreShv (∫ᴾ F) ℓS .ob) + → isIsoC (PreShv (∫ᴾ F) ℓS) (εTrans ⟦ P ⟧) + εIso P = FUNCTORIso _ _ _ isIsoC' + where + isIsoC' : ∀ (cx : (∫ᴾ F) .ob) + → isIsoC (SET _) ((εTrans ⟦ P ⟧) ⟦ cx ⟧) + isIsoC' cx@(c , _) = CatIso→isIso (Iso→CatIso (invIso (typeFiberIso {isSetA = snd (F ⟅ c ⟆)} _))) + + + -- putting it all together + + preshvSlice≃preshvElem : SliceCat ≃ᶜ PreShv (∫ᴾ F) ℓS + preshvSlice≃preshvElem .func = K + preshvSlice≃preshvElem .isEquiv .invFunc = L + preshvSlice≃preshvElem .isEquiv .η .trans = ηTrans + preshvSlice≃preshvElem .isEquiv .η .nIso = ηIso + preshvSlice≃preshvElem .isEquiv .ε .trans = εTrans + preshvSlice≃preshvElem .isEquiv .ε .nIso = εIso diff --git a/Cubical/Categories/TypesOfCategories/TypeCategory.agda b/Cubical/Categories/TypesOfCategories/TypeCategory.agda new file mode 100644 index 0000000000..99305c6b3d --- /dev/null +++ b/Cubical/Categories/TypesOfCategories/TypeCategory.agda @@ -0,0 +1,191 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.TypesOfCategories.TypeCategory where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Data.Sigma +import Cubical.Functions.Fibration as Fibration + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.Presheaf +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Instances.Sets + + +open Fibration.ForSets + +record isTypeCategory {ℓ ℓ' ℓ''} (C : Category ℓ ℓ') + : Type (ℓ-max ℓ (ℓ-max ℓ' (ℓ-suc ℓ''))) where + open Category C + open Cospan + + field + -- a Type of types over a context + Ty[_] : ob → Type ℓ'' + -- extend a context with a type + cext : ∀ (Γ : _) → (A : Ty[ Γ ]) → Σ[ ΓA ∈ ob ] (C [ ΓA , Γ ]) + + -- the new object from a context extension + _⍮_ : (Γ : _) → (A : Ty[ Γ ]) → ob + Γ ⍮ A = fst (cext Γ A) + + -- the projection from the extended context to the original + π : (Γ : _) → (A : Ty[ Γ ]) → C [ Γ ⍮ A , Γ ] + π Γ A = snd (cext Γ A) + + field + -- pullback over context extentions + reindex : ∀ {Γ' Γ} + → C [ Γ' , Γ ] + → (Ty[ Γ ] → Ty[ Γ' ]) + + q⟨_,_⟩ : ∀ {Γ' Γ} + → (f : C [ Γ' , Γ ]) + → (A : Ty[ Γ ]) + → C [ Γ' ⍮ (reindex f A) , Γ ⍮ A ] + + sq : ∀ {Γ' Γ : ob} (f : C [ Γ' , Γ ]) (A : Ty[ Γ ]) + → π Γ' (reindex f A) ⋆ f ≡ q⟨ f , A ⟩ ⋆ π Γ A + + isPB : ∀ {Γ' Γ : ob} (f : C [ Γ' , Γ ]) (A : Ty[ Γ ]) + → isPullback C (cospan Γ' Γ (Γ ⍮ A) f (π Γ A)) + (π Γ' (reindex f A)) (q⟨ f , A ⟩) (sq f A) + +-- presheaves are type contexts +module _ {ℓ ℓ' ℓ'' : Level} (C : Category ℓ ℓ') where + open isTypeCategory + open Category + open Functor + open NatTrans + + private + isSurjSET : ∀ {ℓ} {A B : SET ℓ .ob} → (f : SET ℓ [ A , B ]) → Type _ + isSurjSET {A = A} {B} f = ∀ (b : fst B) → Σ[ a ∈ fst A ] f a ≡ b + + -- types over Γ are types with a "projection" (aka surjection) to Γ + PSTy[_] : PreShv C ℓ'' .ob → Type _ + PSTy[ Γ ] = Σ[ ΓA ∈ PreShv C ℓ'' .ob ] + Σ[ π ∈ ΓA ⇒ Γ ] + (∀ (c : C .ob) → isSurjSET {A = ΓA ⟅ c ⟆} {Γ ⟅ c ⟆} (π ⟦ c ⟧)) + + -- just directly use types from above as context extensions + PSCext : (Γ : _) → PSTy[ Γ ] → Σ[ ΓA ∈ PreShv C ℓ'' .ob ] ΓA ⇒ Γ + PSCext Γ (ΓA , π , _) = ΓA , π + + -- the pullback or reindexed set is the disjoint union of the fibers + -- from the projection + module _ {Δ Γ : PreShv C ℓ'' .ob} (γ : Δ ⇒ Γ) + (A'@(ΓA , π , isSurjπ) : PSTy[ Γ ]) where + ΔA : PreShv C ℓ'' .ob + ΔA .F-ob c = ΔATy , isSetΔA + where + ΔATy = (Σ[ x ∈ fst (Δ ⟅ c ⟆) ] fiber (π ⟦ c ⟧) ((γ ⟦ c ⟧) x)) + + isSetΔA : isSet ΔATy + isSetΔA = isOfHLevelΣ 2 (snd (Δ ⟅ c ⟆)) λ Γc → isOfHLevelΣ 2 (snd (ΓA ⟅ c ⟆)) λ ΓAc → isProp→isSet (snd (Γ ⟅ c ⟆) _ _) + -- for morphisms, we apply Δ ⟪ f ⟫ to the first component + -- and ΓA ⟪ f ⟫ to the second + -- the fiber rule + ΔA .F-hom {c} {d} f (δax , γax , eq) + = ((Δ ⟪ f ⟫) δax) + , (((ΓA ⟪ f ⟫) γax) + , ((π ⟦ d ⟧) ((ΓA ⟪ f ⟫) γax) + ≡[ i ]⟨ π .N-hom f i γax ⟩ + (Γ ⟪ f ⟫) ((π ⟦ c ⟧) γax) + ≡[ i ]⟨ (Γ ⟪ f ⟫) (eq i) ⟩ + (Γ ⟪ f ⟫) ((γ ⟦ c ⟧) δax) + ≡[ i ]⟨ γ .N-hom f (~ i) δax ⟩ + (γ ⟦ d ⟧) ((Δ ⟪ f ⟫) δax) + ∎)) + ΔA .F-id {x = c} + = funExt λ (δax , γax , eq) + → ΣPathP ((λ i → Δ .F-id i δax) + , fibersEqIfRepsEq {isSetB = snd (Γ ⟅ c ⟆)} _ + (λ i → ΓA .F-id i γax)) + ΔA .F-seq {a} {b} {c} f g + = funExt λ (δax , γax , eq) + → ΣPathP ((λ i → Δ .F-seq f g i δax) + , fibersEqIfRepsEq {isSetB = snd (Γ ⟅ c ⟆)} _ + λ i → ΓA .F-seq f g i γax) + π' : ΔA ⇒ Δ + π' .N-ob c (x , snd) = x + π' .N-hom f = refl + PSReindex : PSTy[ Δ ] + PSReindex = ΔA , (π' , isSurj) + where + + isSurj : ∀ (c : C .ob) → isSurjSET {A = ΔA ⟅ c ⟆} {B = Δ ⟅ c ⟆} (π' ⟦ c ⟧) + isSurj c δx = (δx , isSurjπ c ((γ ⟦ c ⟧) δx)) , refl + + PSq : ΔA ⇒ ΓA + PSq .N-ob c (δax , γax , eq) = γax + PSq .N-hom {c} {d} f = funExt λ (δax , γax , eq) → refl + + PSSq : (PreShv C ℓ'' ⋆ snd (PSCext Δ (PSReindex))) γ ≡ + (PreShv C ℓ'' ⋆ PSq) (snd (PSCext Γ A')) + PSSq = makeNatTransPath (funExt sqExt) + where + sqExt : ∀ (c : C .ob) → _ + sqExt c = funExt λ (δax , γax , eq) → sym eq + + PSIsPB : isPullback (PreShv C ℓ'') + (cospan Δ Γ (fst (PSCext Γ A')) γ (snd (PSCext Γ A'))) + (snd (PSCext Δ PSReindex)) PSq PSSq + PSIsPB {Θ} p₁ p₂ sq = (α , eq) , unique + where + α : Θ ⇒ ΔA + α .N-ob c t = ((p₁ ⟦ c ⟧) t) + , (((p₂ ⟦ c ⟧) t) + , (λ i → (sq (~ i) ⟦ c ⟧) t)) + α .N-hom {d} {c} f = funExt αHomExt + where + αHomExt : ∀ (t : fst (Θ ⟅ d ⟆)) + → ((p₁ ⟦ c ⟧) ((Θ ⟪ f ⟫) t) , (p₂ ⟦ c ⟧) ((Θ ⟪ f ⟫) t), _) + ≡ ((Δ ⟪ f ⟫) ((p₁ ⟦ d ⟧) t) , (ΓA ⟪ f ⟫) ((p₂ ⟦ d ⟧) t) , _) + αHomExt t = ΣPathP ((λ i → p₁ .N-hom f i t) + , fibersEqIfRepsEq {isSetB = snd (Γ ⟅ c ⟆)} _ + (λ i → p₂ .N-hom f i t)) + + eq : _ + eq = makeNatTransPath (funExt (λ _ → funExt λ _ → refl)) + , makeNatTransPath (funExt (λ _ → funExt λ _ → refl)) + + unique : ∀ (βeq : Σ[ β ∈ Θ ⇒ ΔA ] _) + → (α , eq) ≡ βeq + unique (β , eqβ) = ΣPathP (α≡β , eq≡eqβ) + where + α≡β : α ≡ β + α≡β = makeNatTransPath (funExt λ c → funExt λ t → eqExt c t) + where + eqβ1 = eqβ .fst + eqβ2 = eqβ .snd + eqExt : ∀ (c : C .ob) + → (t : fst (Θ ⟅ c ⟆)) + → (α ⟦ c ⟧) t ≡ (β ⟦ c ⟧) t + eqExt c t = ΣPathP ((λ i → (eqβ1 i ⟦ c ⟧) t) + , fibersEqIfRepsEq {isSetB = snd (Γ ⟅ c ⟆)} _ + (λ i → (eqβ2 i ⟦ c ⟧) t)) + eq≡eqβ : PathP (λ i + → (p₁ ≡ (α≡β i) ●ᵛ π') + × (p₂ ≡ (α≡β i) ●ᵛ PSq)) eq eqβ + eq≡eqβ = ΣPathP ( isPropNatP1 (eq .fst) (eqβ .fst) α≡β + , isPropNatP2 (eq .snd) (eqβ .snd) α≡β) + where + isPropNatP1 : isOfHLevelDep 1 (λ γ → p₁ ≡ γ ●ᵛ π') + isPropNatP1 = isOfHLevel→isOfHLevelDep 1 (λ _ → isSetNatTrans _ _) + + isPropNatP2 : isOfHLevelDep 1 (λ γ → p₂ ≡ γ ●ᵛ PSq) + isPropNatP2 = isOfHLevel→isOfHLevelDep 1 (λ _ → isSetNatTrans _ _) + + -- putting everything together + isTypeCategoryPresheaf : isTypeCategory (PreShv C ℓ'') + isTypeCategoryPresheaf .Ty[_] Γ = PSTy[ Γ ] + isTypeCategoryPresheaf .cext = PSCext + isTypeCategoryPresheaf .reindex = PSReindex + isTypeCategoryPresheaf .q⟨_,_⟩ = PSq + isTypeCategoryPresheaf .sq = PSSq + isTypeCategoryPresheaf .isPB = PSIsPB diff --git a/Cubical/Categories/Yoneda.agda b/Cubical/Categories/Yoneda.agda new file mode 100644 index 0000000000..83f03f0419 --- /dev/null +++ b/Cubical/Categories/Yoneda.agda @@ -0,0 +1,179 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Yoneda where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence using (ua) +open import Cubical.Foundations.Function renaming (_∘_ to _◍_) +open import Cubical.Data.Sigma +open import Cubical.Foundations.Equiv +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Categories.Category +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Functor +open import Cubical.Categories.Presheaf + +private + variable + ℓ ℓ' ℓ'' : Level + +-- THE YONEDA LEMMA + +open NatTrans +open NatTransP +open Functor +open Iso + +module _ {C : Category ℓ ℓ'} where + open Category + + yoneda : (F : Functor C (SET ℓ')) + → (c : C .ob) + → Iso ((FUNCTOR C (SET ℓ')) [ C [ c ,-] , F ]) (fst (F ⟅ c ⟆)) + yoneda F c = theIso + where + natType = (FUNCTOR C (SET ℓ')) [ C [ c ,-] , F ] + setType = fst (F ⟅ c ⟆) + + -- takes a natural transformation to what it does on id + ϕ : natType → setType + ϕ α = (α ⟦ _ ⟧) (C .id) + + -- takes an element x of F c and sends it to the (only) natural transformation + -- which takes the identity to x + Ψ : setType → natType + Ψ x .N-ob c = λ f → (F ⟪ f ⟫) x + Ψ x .N-hom g + = funExt (λ f → (F ⟪ f ⋆⟨ C ⟩ g ⟫) x + ≡[ i ]⟨ (F .F-seq f g) i x ⟩ + (F ⟪ g ⟫) ((F ⟪ f ⟫) x) + ∎) + + theIso : Iso natType setType + theIso .fun = ϕ + theIso .inv = Ψ + theIso .rightInv x i = F .F-id i x + theIso .leftInv α@(natTrans αo αh) = NatTrans-≡-intro (sym αo≡βo) (symP αh≡βh) + where + β = Ψ (ϕ α) + βo = β .N-ob + βh = β .N-hom + + -- equivalence of action on objects follows + -- from simple equational reasoning using naturality + αo≡βo : αo ≡ βo + αo≡βo = funExt λ x → funExt λ f + → αo x f + ≡[ i ]⟨ αo x (C .⋆IdL f (~ i)) ⟩ -- expand into the bottom left of the naturality diagram + αo x (C .id ⋆⟨ C ⟩ f) + ≡[ i ]⟨ αh f i (C .id) ⟩ -- apply naturality + (F ⟪ f ⟫) ((αo _) (C .id)) + ∎ + + -- type aliases for natural transformation + NOType = N-ob-Type (C [ c ,-]) F + NHType = N-hom-Type (C [ c ,-]) F + + -- equivalence of commutative squares follows from SET being a Category + αh≡βh : PathP (λ i → NHType (αo≡βo i)) αh βh -- αh βh + αh≡βh = isPropHomP αh βh αo≡βo + where + isProp-hom : (ϕ : NOType) → isProp (NHType ϕ) + isProp-hom γ = isPropImplicitΠ2 λ x y → isPropΠ λ f → isSetHom (SET _) {x = (C [ c , x ]) , C .isSetHom } {F ⟅ y ⟆} _ _ + + isPropHomP : isOfHLevelDep 1 (λ ηo → NHType ηo) + isPropHomP = isOfHLevel→isOfHLevelDep 1 λ a → isProp-hom a + + -- Naturality of the bijection + + -- in the functor + -- it's equivalent to apply ϕ to α then do β ⟦ c ⟧ + -- or apply ϕ the the composite nat trans α ◍ β + -- where ϕ takes a natural transformation to its representing element + yonedaIsNaturalInFunctor : ∀ {F G : Functor C (SET ℓ')} (c : C .ob) + → (β : F ⇒ G) + → (fun (yoneda G c) ◍ compTrans β) ≡ (β ⟦ c ⟧ ◍ fun (yoneda F c)) + yonedaIsNaturalInFunctor {F = F} {G} c β = funExt λ α → refl + + -- in the object + -- it's equivalent to apply ϕ and then F ⟪ f ⟫ + -- or to apply ϕ to the natural transformation obtained by precomposing with f + yonedaIsNaturalInOb : ∀ {F : Functor C (SET ℓ')} + → (c c' : C .ob) + → (f : C [ c , c' ]) + → yoneda F c' .fun ◍ preComp f ≡ F ⟪ f ⟫ ◍ yoneda F c .fun + yonedaIsNaturalInOb {F = F} c c' f = funExt (λ α + → (yoneda F c' .fun ◍ preComp f) α + ≡⟨ refl ⟩ + (α ⟦ c' ⟧) (f ⋆⟨ C ⟩ C .id) + ≡[ i ]⟨ (α ⟦ c' ⟧) (C .⋆IdR f i) ⟩ + (α ⟦ c' ⟧) f + ≡[ i ]⟨ (α ⟦ c' ⟧) (C .⋆IdL f (~ i)) ⟩ + (α ⟦ c' ⟧) (C .id ⋆⟨ C ⟩ f) + ≡[ i ]⟨ (α .N-hom f i) (C .id) ⟩ + (F ⟪ f ⟫) ((α ⟦ c ⟧) (C .id)) + ≡⟨ refl ⟩ + ((F ⟪ f ⟫) ◍ yoneda F c .fun) α + ∎) + +-- Yoneda embedding +-- TODO: probably want to rename/refactor +module _ {C : Category ℓ ℓ} where + open Functor + open NatTrans + open Category C + + yo : ob → Functor (C ^op) (SET ℓ) + yo x .F-ob y .fst = C [ y , x ] + yo x .F-ob y .snd = isSetHom + yo x .F-hom f g = f ⋆⟨ C ⟩ g + yo x .F-id i f = ⋆IdL f i + yo x .F-seq f g i h = ⋆Assoc g f h i + + YO : Functor C (PreShv C ℓ) + YO .F-ob = yo + YO .F-hom f .N-ob z g = g ⋆⟨ C ⟩ f + YO .F-hom f .N-hom g i h = ⋆Assoc g h f i + YO .F-id = makeNatTransPath λ i _ → λ f → ⋆IdR f i + YO .F-seq f g = makeNatTransPath λ i _ → λ h → ⋆Assoc h f g (~ i) + + + module _ {x} (F : Functor (C ^op) (SET ℓ)) where + yo-yo-yo : NatTrans (yo x) F → F .F-ob x .fst + yo-yo-yo α = α .N-ob _ id + + no-no-no : F .F-ob x .fst → NatTrans (yo x) F + no-no-no a .N-ob y f = F .F-hom f a + no-no-no a .N-hom f = funExt λ g i → F .F-seq g f i a + + yoIso : Iso (NatTrans (yo x) F) (F .F-ob x .fst) + yoIso .Iso.fun = yo-yo-yo + yoIso .Iso.inv = no-no-no + yoIso .Iso.rightInv b i = F .F-id i b + yoIso .Iso.leftInv a = makeNatTransPath (funExt λ _ → funExt rem) + where + rem : ∀ {z} (x₁ : C [ z , x ]) → F .F-hom x₁ (yo-yo-yo a) ≡ (a .N-ob z) x₁ + rem g = + F .F-hom g (yo-yo-yo a) + ≡[ i ]⟨ a .N-hom g (~ i) id ⟩ + a .N-hom g i0 id + ≡[ i ]⟨ a .N-ob _ (⋆IdR g i) ⟩ + (a .N-ob _) g + ∎ + + yoEquiv : NatTrans (yo x) F ≃ F .F-ob x .fst + yoEquiv = isoToEquiv yoIso + + + isFullYO : isFull YO + isFullYO x y F[f] = ∣ yo-yo-yo _ F[f] , yoIso {x} (yo y) .Iso.leftInv F[f] ∣ + + isFaithfulYO : isFaithful YO + isFaithfulYO x y f g p i = + hcomp + (λ j → λ{ (i = i0) → ⋆IdL f j; (i = i1) → ⋆IdL g j}) + (yo-yo-yo _ (p i)) diff --git a/Cubical/Codata/Conat.agda b/Cubical/Codata/Conat.agda new file mode 100644 index 0000000000..f35a316c92 --- /dev/null +++ b/Cubical/Codata/Conat.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --guardedness #-} + +module Cubical.Codata.Conat where + +open import Cubical.Codata.Conat.Base public + +open import Cubical.Codata.Conat.Properties public diff --git a/Cubical/Codata/Conat/Base.agda b/Cubical/Codata/Conat/Base.agda new file mode 100644 index 0000000000..93974fb4e0 --- /dev/null +++ b/Cubical/Codata/Conat/Base.agda @@ -0,0 +1,55 @@ +{- Conatural numbers (Tesla Ice Zhang, Feb. 2019) + +This file defines: + +- A coinductive natural number representation which is dual to + the inductive version (zero | suc Nat → Nat) of natural numbers. + +- Trivial operations (succ, pred) and the pattern synonyms on conaturals. + +While this definition can be seen as a coinductive wrapper of an inductive +datatype, another way of definition is to define an inductive datatype that +wraps a coinductive thunk of Nat. +The standard library uses the second approach: + +https://github.com/agda/agda-stdlib/blob/master/src/Codata/Conat.agda + +The first approach is chosen to exploit guarded recursion and to avoid the use +of Sized Types. +-} + +{-# OPTIONS --safe --guardedness #-} +module Cubical.Codata.Conat.Base where + +open import Cubical.Data.Unit +open import Cubical.Data.Sum + +open import Cubical.Core.Everything + +record Conat : Type₀ +Conat′ = Unit ⊎ Conat +record Conat where + coinductive + constructor conat′ + field force : Conat′ +open Conat public + +pattern zero = inl tt +pattern suc n = inr n + +conat : Conat′ → Conat +force (conat a) = a + +succ : Conat → Conat +force (succ a) = suc a + +succ′ : Conat′ → Conat′ +succ′ n = suc λ where .force → n + +pred′ : Conat′ → Conat′ +pred′ zero = zero +pred′ (suc x) = force x + +pred′′ : Conat′ → Conat +force (pred′′ zero) = zero +pred′′ (suc x) = x diff --git a/Cubical/Codata/Conat/Properties.agda b/Cubical/Codata/Conat/Properties.agda new file mode 100644 index 0000000000..b4469ccf57 --- /dev/null +++ b/Cubical/Codata/Conat/Properties.agda @@ -0,0 +1,188 @@ +{- Conatural number properties (Tesla Ice Zhang et al., Feb. 2019) + +This file defines operations and properties on conatural numbers: + +- Infinity (∞). + +- Proof that ∞ + 1 is equivalent to ∞. + +- Proof that conatural is an hSet. + +- Bisimulation on conatural + +- Proof that bisimulation is equivalent to equivalence (Coinductive Proof + Principle). + +- Proof that this bisimulation is prop valued + +The standard library also defines bisimulation on conaturals: + +https://github.com/agda/agda-stdlib/blob/master/src/Codata/Conat/Bisimilarity.agda +-} + +{-# OPTIONS --safe --guardedness #-} +module Cubical.Codata.Conat.Properties where + +open import Cubical.Data.Unit +open import Cubical.Data.Sum +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Relation.Nullary +open import Cubical.Codata.Conat.Base + +Unwrap-prev : Conat′ → Type₀ +Unwrap-prev zero = Unit +Unwrap-prev (suc _) = Conat + +unwrap-prev : (n : Conat′) -> Unwrap-prev n +unwrap-prev zero = _ +unwrap-prev (suc x) = x + +private -- tests + 𝟘 = conat zero + 𝟙 = succ 𝟘 + 𝟚 = succ 𝟙 + + succ𝟙≡𝟚 : succ 𝟙 ≡ 𝟚 + succ𝟙≡𝟚 i = 𝟚 + + pred𝟚≡𝟙 : unwrap-prev (force 𝟚) ≡ 𝟙 + pred𝟚≡𝟙 i = 𝟙 + +∞ : Conat +force ∞ = suc ∞ + +∞+1≡∞ : succ ∞ ≡ ∞ +force (∞+1≡∞ _) = suc ∞ + +∞+2≡∞ : succ (succ ∞) ≡ ∞ +∞+2≡∞ = (cong succ ∞+1≡∞) ∙ ∞+1≡∞ + +_+_ : Conat → Conat → Conat +_+′_ : Conat′ → Conat → Conat′ + +force (x + y) = force x +′ y +zero +′ y = force y +suc x +′ y = suc (x + y) + +n+∞≡∞ : ∀ n → n + ∞ ≡ ∞ +n+′∞≡∞′ : ∀ n → n +′ ∞ ≡ suc ∞ + +force (n+∞≡∞ n i) = n+′∞≡∞′ (force n) i +n+′∞≡∞′ zero = refl +n+′∞≡∞′ (suc n) = λ i → suc (n+∞≡∞ n i) + +∞+∞≡∞ : ∞ + ∞ ≡ ∞ +force (∞+∞≡∞ i) = suc (∞+∞≡∞ i) + ++-zeroˡ : ∀ n → 𝟘 + n ≡ n +force (+-zeroˡ n _) = force n + ++-zeroʳ : ∀ n → n + 𝟘 ≡ n ++′-zeroʳ : ∀ n → n +′ 𝟘 ≡ n + +force (+-zeroʳ n i) = +′-zeroʳ (force n) i ++′-zeroʳ zero _ = zero ++′-zeroʳ (suc n) i = suc (+-zeroʳ n i) + ++-assoc : ∀ m n p → (m + n) + p ≡ m + (n + p) ++′-assoc : ∀ m n p → (m +′ n) +′ p ≡ m +′ (n + p) + +force (+-assoc m n p i) = +′-assoc (force m) n p i ++′-assoc zero _ _ = refl ++′-assoc (suc m) n p i = suc (+-assoc m n p i) + + +conat-absurd : ∀ {y : Conat} {ℓ} {Whatever : Type ℓ} → zero ≡ suc y → Whatever +conat-absurd eq = ⊥.rec (transport (cong diag eq) tt) + where + diag : Conat′ → Type₀ + diag zero = Unit + diag (suc _) = ⊥ + +module IsSet where + ≡-stable : {x y : Conat} → Stable (x ≡ y) + ≡′-stable : {x y : Conat′} → Stable (x ≡ y) + + force (≡-stable ¬¬p i) = ≡′-stable (λ ¬p → ¬¬p (λ p → ¬p (cong force p))) i + ≡′-stable {zero} {zero} ¬¬p′ = refl + ≡′-stable {suc x} {suc y} ¬¬p′ = + cong′ suc (≡-stable λ ¬p → ¬¬p′ λ p → ¬p (cong pred′′ p)) + ≡′-stable {zero} {suc y} ¬¬p′ = ⊥.rec (¬¬p′ conat-absurd) + ≡′-stable {suc x} {zero} ¬¬p′ = ⊥.rec (¬¬p′ λ p → conat-absurd (sym p)) + + isSetConat : isSet Conat + isSetConat _ _ = Separated→isSet (λ _ _ → ≡-stable) _ _ + + isSetConat′ : isSet Conat′ + isSetConat′ m n p′ q′ = cong (cong force) (isSetConat (conat m) (conat n) p q) + where p = λ where i .force → p′ i + q = λ where i .force → q′ i + +module Bisimulation where + open IsSet using (isSetConat) + + record _≈_ (x y : Conat) : Type₀ + data _≈′_ (x y : Conat′) : Type₀ + _≈′′_ : Conat′ → Conat′ → Type₀ + zero ≈′′ zero = Unit + suc x ≈′′ suc y = x ≈ y + -- So impossible proofs are preserved + x ≈′′ y = ⊥ + + record _≈_ x y where + coinductive + field prove : force x ≈′ force y + + data _≈′_ x y where + con : x ≈′′ y → x ≈′ y + + open _≈_ public + + bisim : ∀ {x y} → x ≈ y → x ≡ y + bisim′ : ∀ {x y} → x ≈′ y → x ≡ y + + bisim′ {zero} {zero} (con tt) = refl + bisim′ {zero} {suc x} (con ()) + bisim′ {suc x} {zero} (con ()) + bisim′ {suc x} {suc y} (con eq) i = suc (bisim eq i) + force (bisim eq i) = bisim′ (prove eq) i + + misib : ∀ {x y} → x ≡ y → x ≈ y + misib′ : ∀ {x y} → x ≡ y → x ≈′ y + + misib′ {zero} {zero} _ = con tt + misib′ {zero} {suc x} = conat-absurd + misib′ {suc x} {zero} p = conat-absurd (sym p) + -- misib′ {suc x} {suc y} p = con λ where .prove → misib′ (cong pred′ p) + misib′ {suc x} {suc y} p = con (misib (cong pred′′ p)) + prove (misib x≡y) = misib′ (cong force x≡y) + + iso″ : ∀ {x y} → (p : x ≈ y) → misib (bisim p) ≡ p + iso′ : ∀ {x y} → (p : x ≈′ y) → misib′ (bisim′ p) ≡ p + + iso′ {zero} {zero} (con p) = refl + iso′ {zero} {suc x} (con ()) + iso′ {suc x} {zero} (con ()) + iso′ {suc x} {suc y} (con p) = cong con (iso″ p) + prove (iso″ p i) = iso′ (prove p) i + + osi : ∀ {x y} → (p : x ≡ y) → bisim (misib p) ≡ p + osi p = isSetConat _ _ _ p + + path≃bisim : ∀ {x y} → (x ≡ y) ≃ (x ≈ y) + path≃bisim = isoToEquiv (iso misib bisim iso″ osi) + + path≡bisim : ∀ {x y} → (x ≡ y) ≡ (x ≈ y) + path≡bisim = ua path≃bisim + + isProp≈ : ∀ {x y} → isProp (x ≈ y) + isProp≈ = subst isProp path≡bisim (isSetConat _ _) diff --git a/Cubical/Codata/Everything.agda b/Cubical/Codata/Everything.agda new file mode 100644 index 0000000000..3031b6a8a8 --- /dev/null +++ b/Cubical/Codata/Everything.agda @@ -0,0 +1,32 @@ +{-# OPTIONS --guardedness #-} + +module Cubical.Codata.Everything where + +open import Cubical.Codata.EverythingSafe public + +--- Modules making assumptions that might be incompatible with other +-- flags or make use of potentially unsafe features. + +-- Assumes --guardedness +open import Cubical.Codata.Stream public + +open import Cubical.Codata.Conat public + +open import Cubical.Codata.M public + +-- Also uses {-# TERMINATING #-}. +open import Cubical.Codata.M.Bisimilarity public + +{- +-- Alternative M type implemetation, based on +-- https://arxiv.org/pdf/1504.02949.pdf +-- "Non-wellfounded trees in Homotopy Type Theory" +-- Benedikt Ahrens, Paolo Capriotti, Régis Spadotti +-} + +open import Cubical.Codata.M.AsLimit.M +open import Cubical.Codata.M.AsLimit.Coalg +open import Cubical.Codata.M.AsLimit.helper +open import Cubical.Codata.M.AsLimit.Container +open import Cubical.Codata.M.AsLimit.itree +open import Cubical.Codata.M.AsLimit.stream diff --git a/Cubical/Codata/EverythingSafe.agda b/Cubical/Codata/EverythingSafe.agda new file mode 100644 index 0000000000..0f7a364688 --- /dev/null +++ b/Cubical/Codata/EverythingSafe.agda @@ -0,0 +1,2 @@ +{-# OPTIONS --safe #-} +module Cubical.Codata.EverythingSafe where diff --git a/Cubical/Codata/M.agda b/Cubical/Codata/M.agda new file mode 100644 index 0000000000..4399a07c62 --- /dev/null +++ b/Cubical/Codata/M.agda @@ -0,0 +1,95 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Codata.M where + +open import Cubical.Foundations.Prelude + +-- TODO move +module Helpers where + module _ {ℓ ℓ'} {A : Type ℓ} {x : A} (P : ∀ y → x ≡ y → Type ℓ') (d : P x refl) where + -- A version with y as an explicit argument can be used to make Agda + -- infer the family P. + J' : ∀ y (p : x ≡ y) → P y p + J' y p = J P d p + + lem-transp : ∀ {ℓ} {A : Type ℓ} (i : I) → (B : Type ℓ) (P : A ≡ B) → (p : P i) → PathP (\ j → P j) (transp (\ k → P (~ k ∧ i)) (~ i) p) + (transp (\ k → P (k ∨ i)) i p) + lem-transp {A = A} i = J' _ (\ p j → transp (\ _ → A) ((~ j ∧ ~ i) ∨ (j ∧ i)) p ) + + transp-over : ∀ {ℓ} (A : I → Type ℓ) (i j : I) → A i → A j + transp-over A i k p = transp (\ j → A ((~ j ∧ i) ∨ (j ∧ k))) (~ i ∧ ~ k) p + + transp-over-1 : ∀ {ℓ} (A : I → Type ℓ) (i j : I) → A i → A j + transp-over-1 A i k p = transp (\ j → A ((j ∨ i) ∧ (~ j ∨ k))) (i ∧ k) p + + compPathD : {ℓ ℓ' : _} {X : Type ℓ} (F : X → Type ℓ') {A B C : X} (P : A ≡ B) (Q : B ≡ C) + → ∀ {x y z} → (\ i → F (P i)) [ x ≡ y ] → (\ i → F (Q i)) [ y ≡ z ] → (\ i → F ((P ∙ Q) i)) [ x ≡ z ] + compPathD F {A = A} P Q {x} p q i = + comp (\ j → F (hfill (λ j → \ { (i = i0) → A ; (i = i1) → Q j }) + (inS (P i)) + j)) + (λ j → \ { (i = i0) → x; (i = i1) → q j }) + (p i) + +open Helpers + +IxCont : ∀ {ℓ} -> Type ℓ → Type (ℓ-suc ℓ) +IxCont {ℓ} X = Σ (X → Type ℓ) λ S → ∀ x → S x → X → Type ℓ + + +⟦_⟧ : ∀ {ℓ} {X : Type ℓ} → IxCont X → (X → Type ℓ) → (X → Type ℓ) +⟦ (S , P) ⟧ X x = Σ (S x) λ s → ∀ y → P x s y → X y + + +record M {ℓ} {X : Type ℓ} (C : IxCont X) (x : X) : Type ℓ where -- Type₀ + coinductive + field + head : C .fst x + tails : ∀ y → C .snd x head y → M C y + +open M public + +module _ {ℓ} {X : Type ℓ} {C : IxCont X} where + + private + F = ⟦ C ⟧ + + inM : ∀ x → F (M C) x → M C x + inM x (head , tail) = record { head = head ; tails = tail } + + out : ∀ x → M C x → F (M C) x + out x a = (a .head) , (a .tails) + + mapF : ∀ {A B} → (∀ x → A x → B x) → ∀ x → F A x → F B x + mapF f x (s , t) = s , \ y p → f _ (t y p) + + unfold : ∀ {A} (α : ∀ x → A x → F A x) → ∀ x → A x → M C x + unfold α x a .head = α x a .fst + unfold α x a .tails y p = unfold α y (α x a .snd y p) + + + -- We generalize the type to avoid upsetting --guardedness by + -- transporting after the corecursive call. + -- Recognizing hcomp/transp as guardedness-preserving could be a better solution. + unfold-η' : ∀ {A} (α : ∀ x → A x → F A x) → (h : ∀ x → A x → M C x) + → (∀ (x : X) (a : A x) → out x (h x a) ≡ mapF h x (α x a)) + → ∀ (x : X) (a : A x) m → h x a ≡ m → m ≡ unfold α x a + unfold-η' α h eq x a m eq' = let heq = cong head (sym eq') ∙ cong fst (eq x a) + in \ where + i .head → heq i + i .tails y p → + let p0 = (transp-over-1 (\ k → C .snd x (heq k) y) i i1 p) + p1 = (transp-over (\ k → C .snd x (heq k) y) i i0 p) + pe = lem-transp i _ (\ k → C .snd x (heq k) y) p + tl = compPathD (λ p → C .snd x p y → M C y) + (cong head (sym eq')) (cong fst (eq x a)) + (cong (\ f → f .tails y) (sym eq')) + (cong (\ f → f .snd y) (eq x a)) + in unfold-η' α h eq y (α x a .snd y p0) + (m .tails y p1) + (sym (\ k → tl k (pe k))) + i + + unfold-η : ∀ {A} (α : ∀ x → A x → F A x) → (h : ∀ x → A x → M C x) + → (∀ (x : X) (a : A x) → out x (h x a) ≡ mapF h x (α x a)) + → ∀ (x : X) (a : A x) → h x a ≡ unfold α x a + unfold-η α h eq x a = unfold-η' α h eq x a _ refl diff --git a/Cubical/Codata/M/AsLimit/Coalg.agda b/Cubical/Codata/M/AsLimit/Coalg.agda new file mode 100644 index 0000000000..4118df13c7 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/Coalg.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.Coalg where + +open import Cubical.Codata.M.AsLimit.Coalg.Base public diff --git a/Cubical/Codata/M/AsLimit/Coalg/Base.agda b/Cubical/Codata/M/AsLimit/Coalg/Base.agda new file mode 100644 index 0000000000..d4c7105c23 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/Coalg/Base.agda @@ -0,0 +1,39 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.Coalg.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function using ( _∘_ ) +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.Data.Sigma + +open import Cubical.Codata.M.AsLimit.Container +open import Cubical.Codata.M.AsLimit.helper + +------------------------------- +-- Definition of a Coalgebra -- +------------------------------- + +Coalg₀ : ∀ {ℓ} {S : Container ℓ} -> Type (ℓ-suc ℓ) +Coalg₀ {ℓ} {S = S} = Σ[ C ∈ Type ℓ ] (C → P₀ S C) + +-------------------------- +-- Definition of a Cone -- +-------------------------- + +Cone₀ : ∀ {ℓ} {S : Container ℓ} {C,γ : Coalg₀ {S = S}} -> Type ℓ +Cone₀ {S = S} {C , _} = (n : ℕ) → C → X (sequence S) n + +Cone₁ : ∀ {ℓ} {S : Container ℓ} {C,γ : Coalg₀ {S = S}} -> (f : Cone₀ {C,γ = C,γ}) -> Type ℓ +Cone₁ {S = S} {C , _} f = (n : ℕ) → π (sequence S) ∘ (f (suc n)) ≡ f n + +Cone : ∀ {ℓ} {S : Container ℓ} (C,γ : Coalg₀ {S = S}) -> Type ℓ +Cone {S = S} C,γ = Σ[ Cone ∈ (Cone₀ {C,γ = C,γ}) ] (Cone₁{C,γ = C,γ} Cone) diff --git a/Cubical/Codata/M/AsLimit/Container.agda b/Cubical/Codata/M/AsLimit/Container.agda new file mode 100644 index 0000000000..f0c9cb6e8f --- /dev/null +++ b/Cubical/Codata/M/AsLimit/Container.agda @@ -0,0 +1,88 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.Container where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv using (_≃_) +open import Cubical.Foundations.Function using (_∘_) + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) + +open import Cubical.Foundations.Transport + +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open import Cubical.Data.Sum + +open import Cubical.Foundations.Structure + +open import Cubical.Codata.M.AsLimit.helper + +------------------------------------- +-- Container and Container Functor -- +------------------------------------- + +-- Σ[ A ∈ (Type ℓ) ] (A → Type ℓ) +Container : ∀ ℓ -> Type (ℓ-suc ℓ) +Container ℓ = TypeWithStr ℓ (λ x → x → Type ℓ) + +-- Polynomial functor (P₀ , P₁) defined over a container +-- https://ncatlab.org/nlab/show/polynomial+functor + +-- P₀ object part of functor +P₀ : ∀ {ℓ} (S : Container ℓ) -> Type ℓ -> Type ℓ +P₀ (A , B) X = Σ[ a ∈ A ] (B a -> X) + +-- P₁ morphism part of functor +P₁ : ∀ {ℓ} {S : Container ℓ} {X Y} (f : X -> Y) -> P₀ S X -> P₀ S Y +P₁ {S = S} f = λ { (a , g) -> a , f ∘ g } + +----------------------- +-- Chains and Limits -- +----------------------- + +record Chain ℓ : Type (ℓ-suc ℓ) where + constructor _,,_ + field + X : ℕ -> Type ℓ + π : {n : ℕ} -> X (suc n) -> X n + +open Chain public + +limit-of-chain : ∀ {ℓ} -> Chain ℓ → Type ℓ +limit-of-chain (x ,, pi) = Σ[ x ∈ ((n : ℕ) → x n) ] ((n : ℕ) → pi {n = n} (x (suc n)) ≡ x n) + +shift-chain : ∀ {ℓ} -> Chain ℓ -> Chain ℓ +shift-chain = λ X,π -> ((λ x → X X,π (suc x)) ,, λ {n} → π X,π {suc n}) + +------------------------------------------------------ +-- Limit type of a Container , and Shift of a Limit -- +------------------------------------------------------ + +Wₙ : ∀ {ℓ} -> Container ℓ -> ℕ -> Type ℓ +Wₙ S 0 = Lift Unit +Wₙ S (suc n) = P₀ S (Wₙ S n) + +πₙ : ∀ {ℓ} -> (S : Container ℓ) -> {n : ℕ} -> Wₙ S (suc n) -> Wₙ S n +πₙ {ℓ} S {0} _ = lift tt +πₙ S {suc n} = P₁ (πₙ S {n}) + +sequence : ∀ {ℓ} -> Container ℓ -> Chain ℓ +X (sequence S) n = Wₙ S n +π (sequence S) {n} = πₙ S {n} + +PX,Pπ : ∀ {ℓ} (S : Container ℓ) -> Chain ℓ +PX,Pπ {ℓ} S = + (λ n → P₀ S (X (sequence S) n)) ,, + (λ {n : ℕ} x → P₁ (λ z → z) (π (sequence S) {n = suc n} x )) + +----------------------------------- +-- M type is limit of a sequence -- +----------------------------------- + +M : ∀ {ℓ} -> Container ℓ → Type ℓ +M = limit-of-chain ∘ sequence diff --git a/Cubical/Codata/M/AsLimit/M.agda b/Cubical/Codata/M/AsLimit/M.agda new file mode 100644 index 0000000000..22c3b9e654 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/M.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.M where + +open import Cubical.Codata.M.AsLimit.M.Base public +open import Cubical.Codata.M.AsLimit.M.Properties public diff --git a/Cubical/Codata/M/AsLimit/M/Base.agda b/Cubical/Codata/M/AsLimit/M/Base.agda new file mode 100644 index 0000000000..a07b2d8116 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/M/Base.agda @@ -0,0 +1,194 @@ +{-# OPTIONS --guardedness --safe #-} + +-- Construction of M-types from +-- https://arxiv.org/pdf/1504.02949.pdf +-- "Non-wellfounded trees in Homotopy Type Theory" +-- Benedikt Ahrens, Paolo Capriotti, Régis Spadotti + +module Cubical.Codata.M.AsLimit.M.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv using (_≃_) +open import Cubical.Foundations.Function using (_∘_) + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.Data.Nat.Algebra + +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Sum + +open import Cubical.Codata.M.AsLimit.helper +open import Cubical.Codata.M.AsLimit.Container + +open import Cubical.Codata.M.AsLimit.Coalg.Base + +open Iso + +private + limit-collapse : ∀ {ℓ} {S : Container ℓ} (X : ℕ → Type ℓ) (l : (n : ℕ) → X n → X (suc n)) → (x₀ : X 0) → ∀ (n : ℕ) → X n + limit-collapse X l x₀ 0 = x₀ + limit-collapse {S = S} X l x₀ (suc n) = l n (limit-collapse {S = S} X l x₀ n) + +lemma11-Iso : + ∀ {ℓ} {S : Container ℓ} (X : ℕ → Type ℓ) (l : (n : ℕ) → X n → X (suc n)) + → Iso (Σ[ x ∈ ((n : ℕ) → X n)] ((n : ℕ) → x (suc n) ≡ l n (x n))) + (X 0) +fun (lemma11-Iso X l) (x , y) = x 0 +inv (lemma11-Iso {S = S} X l) x₀ = limit-collapse {S = S} X l x₀ , (λ n → refl {x = limit-collapse {S = S} X l x₀ (suc n)}) +rightInv (lemma11-Iso X l) _ = refl +leftInv (lemma11-Iso {ℓ = ℓ} {S = S} X l) (x , y) i = + let temp = χ-prop (x 0) (fst (inv (lemma11-Iso {S = S} X l) (fun (lemma11-Iso {S = S} X l) (x , y))) , refl , (λ n → refl {x = limit-collapse {S = S} X l (x 0) (suc n)})) (x , refl , y) + in temp i .fst , proj₂ (temp i .snd) + where + open AlgebraPropositionality + open NatSection + + X-fiber-over-ℕ : (x₀ : X 0) -> NatFiber NatAlgebraℕ ℓ + X-fiber-over-ℕ x₀ = record { Fiber = X ; fib-zero = x₀ ; fib-suc = λ {n : ℕ} xₙ → l n xₙ } + + X-section : (x₀ : X 0) → (z : Σ[ x ∈ ((n : ℕ) → X n)] (x 0 ≡ x₀) × (∀ n → (x (suc n)) ≡ l n (x n))) -> NatSection (X-fiber-over-ℕ x₀) + X-section = λ x₀ z → record { section = fst z ; sec-comm-zero = proj₁ (snd z) ; sec-comm-suc = proj₂ (snd z) } + + Z-is-Section : (x₀ : X 0) → + Iso (Σ[ x ∈ ((n : ℕ) → X n)] (x 0 ≡ x₀) × (∀ n → (x (suc n)) ≡ l n (x n))) + (NatSection (X-fiber-over-ℕ x₀)) + fun (Z-is-Section x₀) (x , (z , y)) = record { section = x ; sec-comm-zero = z ; sec-comm-suc = y } + inv (Z-is-Section x₀) x = NatSection.section x , (sec-comm-zero x , sec-comm-suc x) + rightInv (Z-is-Section x₀) _ = refl + leftInv (Z-is-Section x₀) (x , (z , y)) = refl + + -- S≡T + χ-prop' : (x₀ : X 0) → isProp (NatSection (X-fiber-over-ℕ x₀)) + χ-prop' x₀ a b = SectionProp.S≡T isNatInductiveℕ (X-section x₀ (inv (Z-is-Section x₀) a)) (X-section x₀ (inv (Z-is-Section x₀) b)) + + χ-prop : (x₀ : X 0) → isProp (Σ[ x ∈ ((n : ℕ) → X n) ] (x 0 ≡ x₀) × (∀ n → (x (suc n)) ≡ l n (x n))) + χ-prop x₀ = subst isProp (sym (isoToPath (Z-is-Section x₀))) (χ-prop' x₀) + +----------------------------------------------------- +-- Shifting the limit of a chain is an equivalence -- +----------------------------------------------------- + +-- Shift is equivalence (12) and (13) in the proof of Theorem 7 +-- https://arxiv.org/pdf/1504.02949.pdf +-- "Non-wellfounded trees in Homotopy Type Theory" +-- Benedikt Ahrens, Paolo Capriotti, Régis Spadotti + +-- TODO: This definition is inefficient, it should be updated to use some cubical features! +shift-iso : ∀ {ℓ} (S : Container ℓ) -> Iso (P₀ S (M S)) (M S) +shift-iso S@(A , B) = + P₀ S (M S) + Iso⟨ Σ-cong-iso-snd + (λ x → iso (λ f → (λ n z → f z .fst n) , λ n i a → f a .snd n i) + (λ (u , q) z → (λ n → u n z) , λ n i → q n i z) + (λ _ → refl) + (λ _ → refl)) ⟩ + (Σ[ a ∈ A ] (Σ[ u ∈ ((n : ℕ) → B a → X (sequence S) n) ] ((n : ℕ) → π (sequence S) ∘ (u (suc n)) ≡ u n))) + Iso⟨ invIso α-iso-step-5-Iso ⟩ + (Σ[ a ∈ (Σ[ a ∈ ((n : ℕ) → A) ] ((n : ℕ) → a (suc n) ≡ a n)) ] + Σ[ u ∈ ((n : ℕ) → B (a .fst n) → X (sequence S) n) ] + ((n : ℕ) → PathP (λ x → B (a .snd n x) → X (sequence S) n) + (π (sequence S) ∘ u (suc n)) + (u n))) + Iso⟨ α-iso-step-1-4-Iso-lem-12 ⟩ + M S ∎Iso + where + α-iso-step-5-Iso-helper0 : + ∀ (a : (ℕ -> A)) + → (p : (n : ℕ) → a (suc n) ≡ a n) + → (n : ℕ) + → a n ≡ a 0 + α-iso-step-5-Iso-helper0 a p 0 = refl + α-iso-step-5-Iso-helper0 a p (suc n) = p n ∙ α-iso-step-5-Iso-helper0 a p n + + α-iso-step-5-Iso-helper1-Iso : + ∀ (a : ℕ -> A) + → (p : (n : ℕ) → a (suc n) ≡ a n) + → (u : (n : ℕ) → B (a n) → Wₙ S n) + → (n : ℕ) + → (PathP (λ x → B (p n x) → Wₙ S n) (πₙ S ∘ u (suc n)) (u n)) ≡ + (πₙ S ∘ (subst (\k -> B k → Wₙ S (suc n)) (α-iso-step-5-Iso-helper0 a p (suc n))) (u (suc n)) + ≡ subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n)) + α-iso-step-5-Iso-helper1-Iso a p u n = + PathP (λ x → B (p n x) → Wₙ S n) (πₙ S ∘ u (suc n)) (u n) + ≡⟨ PathP≡Path (λ x → B (p n x) → Wₙ S n) (πₙ S ∘ u (suc n)) (u n) ⟩ + subst (λ k → B k → Wₙ S n) (p n) (πₙ S ∘ u (suc n)) ≡ (u n) + ≡⟨ (λ i → transp (λ j → B (α-iso-step-5-Iso-helper0 a p n (i ∧ j)) → Wₙ S n) (~ i) (subst (λ k → B k → Wₙ S n) (p n) (πₙ S ∘ u (suc n))) + ≡ transp (λ j → B (α-iso-step-5-Iso-helper0 a p n (i ∧ j)) → Wₙ S n) (~ i) (u n)) ⟩ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (subst (λ k → B k → Wₙ S n) (p n) (πₙ S ∘ u (suc n))) ≡ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n) + ≡⟨ (λ i → sym (substComposite (λ k → B k → Wₙ S n) (p n) (α-iso-step-5-Iso-helper0 a p n) (πₙ S ∘ u (suc n))) i + ≡ subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n)) ⟩ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p (suc n)) (πₙ S ∘ u (suc n)) ≡ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n) + ≡⟨ (λ i → substCommSlice (λ k → B k → Wₙ S (suc n)) (λ k → B k → Wₙ S n) + (λ a x x₁ → (πₙ S) (x x₁)) + (α-iso-step-5-Iso-helper0 a p (suc n)) (u (suc n)) i + ≡ subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n)) ⟩ + πₙ S ∘ subst (λ k → B k → Wₙ S (suc n)) (α-iso-step-5-Iso-helper0 a p (suc n)) (u (suc n)) + ≡ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n) ∎ + + α-iso-step-5-Iso : + Iso + (Σ[ a ∈ (Σ[ a ∈ ((n : ℕ) → A) ] ((n : ℕ) → a (suc n) ≡ a n)) ] + Σ[ u ∈ ((n : ℕ) → B (a .fst n) → X (sequence S) n) ] + ((n : ℕ) → PathP (λ x → B (a .snd n x) → X (sequence S) n) + (π (sequence S) ∘ u (suc n)) + (u n))) + (Σ[ a ∈ A ] (Σ[ u ∈ ((n : ℕ) → B a → X (sequence S) n) ] ((n : ℕ) → π (sequence S) ∘ (u (suc n)) ≡ u n))) + α-iso-step-5-Iso = + Σ-cong-iso (lemma11-Iso {S = S} (λ _ → A) (λ _ x → x)) (λ a,p → + Σ-cong-iso (pathToIso (cong (λ k → (n : ℕ) → k n) (funExt λ n → cong (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 (a,p .fst) (a,p .snd) n)))) λ u → + pathToIso (cong (λ k → (n : ℕ) → k n) (funExt λ n → α-iso-step-5-Iso-helper1-Iso (a,p .fst) (a,p .snd) u n))) + + α-iso-step-1-4-Iso-lem-12 : + Iso (Σ[ a ∈ (Σ[ a ∈ ((n : ℕ) → A)] ((n : ℕ) → a (suc n) ≡ a n)) ] + (Σ[ u ∈ ((n : ℕ) → B (a .fst n) → X (sequence S) n) ] + ((n : ℕ) → PathP (λ x → B (a .snd n x) → X (sequence S) n) + (π (sequence S) ∘ u (suc n)) + (u n)))) + (limit-of-chain (sequence S)) + fun α-iso-step-1-4-Iso-lem-12 (a , b) = (λ { 0 → lift tt ; (suc n) → (a .fst n) , (b .fst n)}) , λ { 0 → refl {x = lift tt} ; (suc m) i → a .snd m i , b .snd m i } + inv α-iso-step-1-4-Iso-lem-12 x = ((λ n → (x .fst) (suc n) .fst) , λ n i → (x .snd) (suc n) i .fst) , (λ n → (x .fst) (suc n) .snd) , λ n i → (x .snd) (suc n) i .snd + fst (rightInv α-iso-step-1-4-Iso-lem-12 (b , c) i) 0 = lift tt + fst (rightInv α-iso-step-1-4-Iso-lem-12 (b , c) i) (suc n) = refl i + snd (rightInv α-iso-step-1-4-Iso-lem-12 (b , c) i) 0 = refl + snd (rightInv α-iso-step-1-4-Iso-lem-12 (b , c) i) (suc n) = c (suc n) + leftInv α-iso-step-1-4-Iso-lem-12 (a , b) = refl + +shift : ∀ {ℓ} (S : Container ℓ) -> P₀ S (M S) ≡ M S +shift S = isoToPath (shift-iso S) -- lemma 13 & lemma 12 + +-- Transporting along shift + +in-fun : ∀ {ℓ} {S : Container ℓ} -> P₀ S (M S) -> M S +in-fun {S = S} = fun (shift-iso S) + +out-fun : ∀ {ℓ} {S : Container ℓ} -> M S -> P₀ S (M S) +out-fun {S = S} = inv (shift-iso S) + +-- Property of functions into M-types + +lift-to-M : ∀ {ℓ} {A : Type ℓ} {S : Container ℓ} + → (x : (n : ℕ) -> A → X (sequence S) n) + → ((n : ℕ) → (a : A) → π (sequence S) (x (suc n) a) ≡ x n a) + --------------- + → (A → M S) +lift-to-M x p a = (λ n → x n a) , λ n i → p n a i + +lift-direct-M : ∀ {ℓ} {S : Container ℓ} + → (x : (n : ℕ) → X (sequence S) n) + → ((n : ℕ) → π (sequence S) (x (suc n)) ≡ x n) + --------------- + → M S +lift-direct-M x p = x , p diff --git a/Cubical/Codata/M/AsLimit/M/Properties.agda b/Cubical/Codata/M/AsLimit/M/Properties.agda new file mode 100644 index 0000000000..adabd33a1e --- /dev/null +++ b/Cubical/Codata/M/AsLimit/M/Properties.agda @@ -0,0 +1,61 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.M.Properties where + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) +open import Cubical.Data.Sum +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv using (_≃_) +open import Cubical.Foundations.Function using (_∘_) +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv + +open import Cubical.Functions.Embedding + +open import Cubical.Codata.M.AsLimit.helper +open import Cubical.Codata.M.AsLimit.M.Base +open import Cubical.Codata.M.AsLimit.Container + +-- in-fun and out-fun are inverse + +open Iso + +in-inverse-out : ∀ {ℓ} {S : Container ℓ} → (in-fun {S = S} ∘ out-fun {S = S}) ≡ idfun (M S) +in-inverse-out {S = S} = subst (λ inv → in-fun {S = S} ∘ inv ≡ idfun (M S)) idpath def where + -- substituting refl makes type-checking work a lot faster, but introduces a transport + -- TODO (2020-05-23): revisit this at some point to see if it's still needed in future versions of agda + def : (in-fun {S = S} ∘ inv (shift-iso S)) ≡ idfun (M S) + def = funExt (rightInv (shift-iso S)) + idpath : inv (shift-iso S) ≡ out-fun {S = S} + idpath = refl + +out-inverse-in : ∀ {ℓ} {S : Container ℓ} → (out-fun {S = S} ∘ in-fun {S = S}) ≡ idfun (P₀ S (M S)) +out-inverse-in {S = S} = subst (λ fun → out-fun {S = S} ∘ fun ≡ idfun (P₀ S (M S))) idpath def where + def : (out-fun {S = S} ∘ fun (shift-iso S)) ≡ idfun (P₀ S (M S)) + def = funExt (leftInv (shift-iso S)) + idpath : fun (shift-iso S) ≡ in-fun {S = S} + idpath = refl + +in-out-id : ∀ {ℓ} {S : Container ℓ} -> ∀ {x y : M S} → (in-fun (out-fun x) ≡ in-fun (out-fun y)) ≡ (x ≡ y) +in-out-id {x = x} {y} i = (in-inverse-out i x) ≡ (in-inverse-out i y) + +-- constructor properties + +in-inj : ∀ {ℓ} {S : Container ℓ} {Z : Type ℓ} → ∀ {f g : Z → P₀ S (M S)} → (in-fun ∘ f ≡ in-fun ∘ g) ≡ (f ≡ g) +in-inj {ℓ} {S = S} {Z = Z} {f = f} {g = g} = iso→fun-Injection-Path {ℓ = ℓ} {A = P₀ S (M S)} {B = M S} {C = Z} (shift-iso S) {f = f} {g = g} + +out-inj : ∀ {ℓ} {S : Container ℓ} {Z : Type ℓ} → ∀ {f g : Z → M S} → (out-fun ∘ f ≡ out-fun ∘ g) ≡ (f ≡ g) +out-inj {ℓ} {S = S} {Z = Z} {f = f} {g = g} = iso→inv-Injection-Path {ℓ = ℓ} {A = P₀ S (M S)} {B = M S} {C = Z} (shift-iso S) {f = f} {g = g} + +in-inj-x : ∀ {ℓ} {S : Container ℓ} → ∀ {x y : P₀ S (M S)} → (in-fun x ≡ in-fun y) ≡ (x ≡ y) +in-inj-x {ℓ} {S = S} {x = x} {y} = iso→fun-Injection-Path-x (shift-iso S) {x} {y} + +out-inj-x : ∀ {ℓ} {S : Container ℓ} → ∀ {x y : M S} → (out-fun x ≡ out-fun y) ≡ (x ≡ y) +out-inj-x {ℓ} {S = S} {x = x} {y} = iso→inv-Injection-Path-x (shift-iso S) {x} {y} diff --git a/Cubical/Codata/M/AsLimit/helper.agda b/Cubical/Codata/M/AsLimit/helper.agda new file mode 100644 index 0000000000..33be276fd7 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/helper.agda @@ -0,0 +1,116 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.helper where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv using (_≃_) +open import Cubical.Foundations.Function using (_∘_) + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) +open import Cubical.Data.Sigma hiding (_×_) + +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Functions.Embedding +open import Cubical.Functions.FunExtEquiv + +open Iso + +-- General + +iso→fun-Injection : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) + → ∀ {f g : C -> A} + → (x : C) → (Iso.fun isom (f x) ≡ Iso.fun isom (g x)) ≡ (f x ≡ g x) +iso→fun-Injection {A = A} {B} {C} isom {f = f} {g} = + isEmbedding→Injection {A = A} {B} {C} (Iso.fun isom) (iso→isEmbedding {A = A} {B} isom) {f = f} {g = g} + +abstract + iso→Pi-fun-Injection : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) + → ∀ {f g : C -> A} + → Iso (∀ x → (fun isom) (f x) ≡ (fun isom) (g x)) (∀ x → f x ≡ g x) + iso→Pi-fun-Injection {A = A} {B} {C} isom {f = f} {g} = + pathToIso (cong (λ k → ∀ x → k x) (funExt (iso→fun-Injection isom {f = f} {g = g}))) + +iso→fun-Injection-Iso : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) + → ∀ {f g : C -> A} + → Iso (fun isom ∘ f ≡ fun isom ∘ g) (f ≡ g) +iso→fun-Injection-Iso {A = A} {B} {C} isom {f = f} {g} = + (fun isom) ∘ f ≡ (fun isom) ∘ g + Iso⟨ invIso funExtIso ⟩ + (∀ x → (fun isom) (f x) ≡ (fun isom) (g x)) + Iso⟨ iso→Pi-fun-Injection isom ⟩ + (∀ x → f x ≡ g x) + Iso⟨ funExtIso ⟩ + f ≡ g ∎Iso + +iso→fun-Injection-Path : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) + → ∀ {f g : C -> A} + → (fun isom ∘ f ≡ fun isom ∘ g) ≡ (f ≡ g) +iso→fun-Injection-Path {A = A} {B} {C} isom {f = f} {g} = + isoToPath (iso→fun-Injection-Iso isom) + +iso→inv-Injection-Path : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) → + ∀ {f g : C -> B} → + ----------------------- + ((inv isom) ∘ f ≡ (inv isom) ∘ g) ≡ (f ≡ g) +iso→inv-Injection-Path {A = A} {B} {C} isom {f = f} {g} = iso→fun-Injection-Path (invIso isom) + +iso→fun-Injection-Iso-x : + ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + → ∀ {x y : A} + → Iso ((fun isom) x ≡ (fun isom) y) (x ≡ y) +iso→fun-Injection-Iso-x isom {x} {y} = + let tempx = λ {(lift tt) → x} + tempy = λ {(lift tt) → y} in + fun isom x ≡ fun isom y + Iso⟨ iso (λ x₁ t → x₁) + (λ x₁ → x₁ (lift tt)) + (λ x → refl) + (λ x → refl) ⟩ + (∀ (t : Lift Unit) -> (((fun isom) ∘ tempx) t ≡ ((fun isom) ∘ tempy) t)) + Iso⟨ iso→Pi-fun-Injection isom ⟩ + (∀ (t : Lift Unit) -> tempx t ≡ tempy t) + Iso⟨ iso (λ x₁ → x₁ (lift tt)) + (λ x₁ t → x₁) + (λ x → refl) + (λ x → refl) ⟩ + x ≡ y ∎Iso + +iso→inv-Injection-Iso-x : + ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + → ∀ {x y : B} + → Iso ((inv isom) x ≡ (inv isom) y) (x ≡ y) +iso→inv-Injection-Iso-x {A = A} {B = B} isom = + iso→fun-Injection-Iso-x {A = B} {B = A} (invIso isom) + +iso→fun-Injection-Path-x : + ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + → ∀ {x y : A} + → ((fun isom) x ≡ (fun isom) y) ≡ (x ≡ y) +iso→fun-Injection-Path-x isom {x} {y} = + isoToPath (iso→fun-Injection-Iso-x isom) + +iso→inv-Injection-Path-x : + ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + → ∀ {x y : B} + → ((inv isom) x ≡ (inv isom) y) ≡ (x ≡ y) +iso→inv-Injection-Path-x isom = + isoToPath (iso→inv-Injection-Iso-x isom) + diff --git a/Cubical/Codata/M/AsLimit/itree.agda b/Cubical/Codata/M/AsLimit/itree.agda new file mode 100644 index 0000000000..50d8e10757 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/itree.agda @@ -0,0 +1,75 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.itree where + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) +open import Cubical.Data.Sum +open import Cubical.Data.Empty +open import Cubical.Data.Bool + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude + +open import Cubical.Codata.M.AsLimit.Container +open import Cubical.Codata.M.AsLimit.M +open import Cubical.Codata.M.AsLimit.Coalg + +-- Delay monad defined as an M-type +delay-S : (R : Type₀) -> Container ℓ-zero +delay-S R = (Unit ⊎ R) , λ { (inr _) -> ⊥ ; (inl tt) -> Unit } + +delay : (R : Type₀) -> Type₀ +delay R = M (delay-S R) + +delay-ret : {R : Type₀} -> R -> delay R +delay-ret r = in-fun (inr r , λ ()) + +delay-tau : {R : Type₀} -> delay R -> delay R +delay-tau S = in-fun (inl tt , λ x → S) + +-- Bottom element raised +data ⊥₁ : Type₁ where + +-- TREES +tree-S : (E : Type₀ -> Type₁) (R : Type₀) -> Container (ℓ-suc ℓ-zero) +tree-S E R = (R ⊎ (Σ[ A ∈ Type₀ ] (E A))) , (λ { (inl _) -> ⊥₁ ; (inr (A , e)) -> Lift A } ) + +tree : (E : Type₀ -> Type₁) (R : Type₀) -> Type₁ +tree E R = M (tree-S E R) + +tree-ret : ∀ {E} {R} -> R -> tree E R +tree-ret {E} {R} r = in-fun (inl r , λ ()) + +tree-vis : ∀ {E} {R} -> ∀ {A} -> E A -> (A -> tree E R) -> tree E R +tree-vis {A = A} e k = in-fun (inr (A , e) , λ { (lift x) -> k x } ) + +-- ITrees (and buildup examples) +-- https://arxiv.org/pdf/1906.00046.pdf +-- Interaction Trees: Representing Recursive and Impure Programs in Coq +-- Li-yao Xia, Yannick Zakowski, Paul He, Chung-Kil Hur, Gregory Malecha, Benjamin C. Pierce, Steve Zdancewic + +itree-S : ∀ (E : Type₀ -> Type₁) (R : Type₀) -> Container (ℓ-suc ℓ-zero) +itree-S E R = ((Unit ⊎ R) ⊎ (Σ[ A ∈ Type₀ ] (E A))) , (λ { (inl (inl _)) -> Lift Unit ; (inl (inr _)) -> ⊥₁ ; (inr (A , e)) -> Lift A } ) + +itree : ∀ (E : Type₀ -> Type₁) (R : Type₀) -> Type₁ +itree E R = M (itree-S E R) + +ret' : ∀ {E} {R} -> R -> P₀ (itree-S E R) (itree E R) +ret' {E} {R} r = inl (inr r) , λ () + +tau' : {E : Type₀ -> Type₁} -> {R : Type₀} -> itree E R -> P₀ (itree-S E R) (itree E R) +tau' t = inl (inl tt) , λ x → t + +vis' : ∀ {E} {R} -> ∀ {A : Type₀} -> E A -> (A -> itree E R) -> P₀ (itree-S E R) (itree E R) +vis' {A = A} e k = inr (A , e) , λ { (lift x) -> k x } + +ret : ∀ {E} {R} -> R -> itree E R +ret = in-fun ∘ ret' + +tau : {E : Type₀ -> Type₁} -> {R : Type₀} -> itree E R -> itree E R +tau = in-fun ∘ tau' + +vis : ∀ {E} {R} -> ∀ {A : Type₀} -> E A -> (A -> itree E R) -> itree E R +vis {A = A} e = in-fun ∘ (vis' {A = A} e) diff --git a/Cubical/Codata/M/AsLimit/stream.agda b/Cubical/Codata/M/AsLimit/stream.agda new file mode 100644 index 0000000000..f6fc3c19eb --- /dev/null +++ b/Cubical/Codata/M/AsLimit/stream.agda @@ -0,0 +1,30 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.stream where + +open import Cubical.Data.Unit + +open import Cubical.Foundations.Prelude + +open import Cubical.Codata.M.AsLimit.M +open import Cubical.Codata.M.AsLimit.helper +open import Cubical.Codata.M.AsLimit.Container + +-------------------------------------- +-- Stream definitions using M.AsLimit -- +-------------------------------------- + +stream-S : ∀ A -> Container ℓ-zero +stream-S A = (A , (λ _ → Unit)) + +stream : ∀ (A : Type₀) -> Type₀ +stream A = M (stream-S A) + +cons : ∀ {A} -> A -> stream A -> stream A +cons x xs = in-fun (x , λ { tt -> xs } ) + +hd : ∀ {A} -> stream A -> A +hd {A} S = out-fun S .fst + +tl : ∀ {A} -> stream A -> stream A +tl {A} S = out-fun S .snd tt diff --git a/Cubical/Codata/M/Bisimilarity.agda b/Cubical/Codata/M/Bisimilarity.agda new file mode 100644 index 0000000000..b88086d6e1 --- /dev/null +++ b/Cubical/Codata/M/Bisimilarity.agda @@ -0,0 +1,147 @@ +{-# OPTIONS --postfix-projections --guardedness #-} +module Cubical.Codata.M.Bisimilarity where + +open import Cubical.Core.Everything +open import Cubical.Codata.M +open import Cubical.Foundations.Equiv.Fiberwise +open import Cubical.Foundations.Everything +open Helpers using (J') + +-- Bisimilarity as a coinductive record type. +record _≈_ {X : Type₀} {C : IxCont X} {x : X} (a b : M C x) : Type₀ where + coinductive + constructor _,_ + field + head≈ : a .head ≡ b .head + tails≈ : ∀ y → (pa : C .snd x (a .head) y) (pb : C .snd x (b .head) y) + → (\ i → C .snd x (head≈ i) y) [ pa ≡ pb ] + → a .tails y pa ≈ b .tails y pb + +open _≈_ public + + + +module _ {X : Type₀} {C : IxCont X} where + + -- Here we show that `a ≡ b` and `a ≈ b` are equivalent. + -- + -- A direct construction of an isomorphism, like we do for streams, + -- would be complicated by the type dependencies between the fields + -- of `M C x` and even more so between the fields of the bisimilarity relation itself. + -- + -- Instead we rely on theorem 4.7.7 of the HoTT book (fiberwise equivalences) to show that `misib` is an equivalence. + + misib : {x : X} (a b : M C x) → a ≡ b → a ≈ b + misib a b eq .head≈ i = eq i .head + misib a b eq .tails≈ y pa pb q = misib (a .tails y pa) (b .tails y pb) (\ i → eq i .tails y (q i)) + + + -- With `a` fixed, `misib` is a fiberwise transformation between (a ≡_) and (a ≈_). + -- + -- We show that the induced map on the total spaces is an + -- equivalence because it is a map of contractible types. + -- + -- The domain is the HoTT singleton type, so contractible, while the + -- codomain is shown to be contractible by `contr-T` below. + + T : ∀ {x} → M C x → Type _ + T a = Σ (M C _) \ b → a ≈ b + + private + lemma : ∀ {A} (B : Type₀) (P : A ≡ B) (pa : P i0) (pb : P i1) (peq : PathP (\ i → P i) pa pb) + → PathP (\ i → PathP (\ j → P j) (transp (\ k → P (~ k ∧ i)) (~ i) (peq i)) pb) + peq + (\ j → transp (\ k → P (~ k ∨ j)) j pb) + lemma {A} = J' _ \ pa → J' _ \ { i j → transp (\ _ → A) (~ i ∨ j) pa } + + + + -- We predefine `u'` so that Agda will agree that `contr-T-fst` is productive. + private + module Tails x a φ (u : Partial φ (T {x} a)) y (p : C .snd x (hcomp (λ i .o → u o .snd .head≈ i) (a .head)) y) where + q = transp (\ i → C .snd x (hfill (\ i o → u o .snd .head≈ i) (inS (a .head)) (~ i)) y) i0 p + a' = a .tails y q + u' : Partial φ (T a') + u' (φ = i1) = u 1=1 .fst .tails y p + , u 1=1 .snd .tails≈ y q p \ j → transp (\ i → C .snd x (u 1=1 .snd .head≈ (~ i ∨ j)) y) j p + + + contr-T-fst : ∀ x a φ → Partial φ (T {x} a) → M C x + contr-T-fst x a φ u .head = hcomp (\ i o → u o .snd .head≈ i) (a .head) + contr-T-fst x a φ u .tails y p = contr-T-fst y a' φ u' + where + open Tails x a φ u y p + + -- `contr-T-snd` is productive as the corecursive call appears as + -- the main argument of transport, which is guardedness-preserving. + {-# TERMINATING #-} + contr-T-snd : ∀ x a φ → (u : Partial φ (T {x} a)) → a ≈ contr-T-fst x a φ u + contr-T-snd x a φ u .head≈ i = hfill (λ { i (φ = i1) → u 1=1 .snd .head≈ i }) (inS (a .head)) i + contr-T-snd x a φ u .tails≈ y pa pb peq = + let r = contr-T-snd y (a .tails y pa) φ (\ { (φ = i1) → u 1=1 .fst .tails y pb , u 1=1 .snd .tails≈ y pa pb peq }) in + transport (\ i → a .tails y pa + ≈ contr-T-fst y (a .tails y (sym (fromPathP (\ i → peq (~ i))) i)) φ + (\ { (φ = i1) → u 1=1 .fst .tails y pb , u 1=1 .snd .tails≈ y + ((fromPathP (\ i → peq (~ i))) (~ i)) pb + \ j → lemma _ (λ h → C .snd x (u _ .snd .head≈ h) y) pa pb peq i j })) r + + contr-T : ∀ x a φ → Partial φ (T {x} a) → T a + contr-T x a φ u .fst = contr-T-fst x a φ u + contr-T x a φ u .snd = contr-T-snd x a φ u + + + contr-T-φ-fst : ∀ x a → (u : Partial i1 (T {x} a)) → contr-T x a i1 u .fst ≡ u 1=1 .fst + contr-T-φ-fst x a u i .head = u 1=1 .fst .head + contr-T-φ-fst x a u i .tails y p + = let + q = (transp (\ i → C .snd x (hfill (\ i o → u o .snd .head≈ i) (inS (a .head)) (~ i)) y) i0 p) + in contr-T-φ-fst y (a .tails y q) + (\ o → u o .fst .tails y p + , u o .snd .tails≈ y q p \ j → transp (\ i → C .snd x (u 1=1 .snd .head≈ (~ i ∨ j)) y) j p) + i + + -- `contr-T-φ-snd` is productive as the corecursive call appears as + -- the main argument of transport, which is guardedness-preserving (even for paths of a coinductive type). + {-# TERMINATING #-} + contr-T-φ-snd : ∀ x a → (u : Partial i1 (T {x} a)) → (\ i → a ≈ contr-T-φ-fst x a u i) [ contr-T x a i1 u .snd ≡ u 1=1 .snd ] + contr-T-φ-snd x a u i .head≈ = u _ .snd .head≈ + contr-T-φ-snd x a u i .tails≈ y pa pb peq = let + eqh = u 1=1 .snd .head≈ + r = contr-T-φ-snd y (a .tails y pa) (\ o → u o .fst .tails y pb , u 1=1 .snd .tails≈ y pa pb peq) + F : I → Type _ + F k = a .tails y pa + ≈ contr-T-fst y + (a .tails y (transp (λ j → C .snd x (eqh (k ∧ ~ j)) y) (~ k) (peq k))) + i1 + (λ _ → u _ .fst .tails y pb + , u _ .snd .tails≈ y + (transp (λ j → C .snd x (eqh (k ∧ ~ j)) y) (~ k) (peq k)) + pb + (λ j → lemma (C .snd x (u 1=1 .fst .head) y) (λ h → C .snd x (eqh h) y) pa pb peq k j) + ) + u0 = contr-T-snd y (a .tails y pa) i1 (λ o → u o .fst .tails y pb , u o .snd .tails≈ y pa pb peq) + in transport + (λ l → PathP + (λ z → a .tails y pa + ≈ contr-T-φ-fst y + (a .tails y (transp (λ k → C .snd x (u 1=1 .snd .head≈ (~ k ∧ l)) y) (~ l) (peq l))) + (λ _ → u _ .fst .tails y pb + , u _ .snd .tails≈ y (transp (λ k → C .snd x (u _ .snd .head≈ (~ k ∧ l)) y) (~ l) (peq l)) pb + \ j → lemma (C .snd x (u 1=1 .fst .head) y) (λ h → C .snd x (eqh h) y) pa pb peq l j) + z) + (transpFill {A = F i0} i0 (\ i → inS (F i)) u0 l) + (u _ .snd .tails≈ y pa pb peq)) + r + i + + contr-T-φ : ∀ x a → (u : Partial i1 (T {x} a)) → contr-T x a i1 u ≡ u 1=1 + contr-T-φ x a u i .fst = contr-T-φ-fst x a u i + contr-T-φ x a u i .snd = contr-T-φ-snd x a u i + + + contr-T' : ∀ {x} a → isContr (T {x} a) + contr-T' a = isContrPartial→isContr (contr-T _ a) \ u → sym (contr-T-φ _ a (\ _ → u)) + + + bisimEquiv : ∀ {x} {a b : M C x} → isEquiv (misib a b) + bisimEquiv = isContrToUniv _≈_ (misib _ _) contr-T' diff --git a/Cubical/Codata/Stream.agda b/Cubical/Codata/Stream.agda new file mode 100644 index 0000000000..45011fd312 --- /dev/null +++ b/Cubical/Codata/Stream.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --guardedness #-} + +module Cubical.Codata.Stream where + +open import Cubical.Codata.Stream.Base public + +open import Cubical.Codata.Stream.Properties public diff --git a/Cubical/Codata/Stream/Base.agda b/Cubical/Codata/Stream/Base.agda new file mode 100644 index 0000000000..d4a62600a3 --- /dev/null +++ b/Cubical/Codata/Stream/Base.agda @@ -0,0 +1,11 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Codata.Stream.Base where + +open import Cubical.Core.Everything + +record Stream (A : Type₀) : Type₀ where + coinductive + constructor _,_ + field + head : A + tail : Stream A diff --git a/Cubical/Codata/Stream/Properties.agda b/Cubical/Codata/Stream/Properties.agda new file mode 100644 index 0000000000..85927262c1 --- /dev/null +++ b/Cubical/Codata/Stream/Properties.agda @@ -0,0 +1,121 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Codata.Stream.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Data.Nat + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Codata.Stream.Base + +open Stream + +mapS : ∀ {A B} → (A → B) → Stream A → Stream B +head (mapS f xs) = f (head xs) +tail (mapS f xs) = mapS f (tail xs) + +even : ∀ {A} → Stream A → Stream A +head (even a) = head a +tail (even a) = even (tail (tail a)) + +odd : ∀ {A} → Stream A → Stream A +head (odd a) = head (tail a) +tail (odd a) = odd (tail (tail a)) + +merge : ∀ {A} → Stream A → Stream A → Stream A +head (merge a _) = head a +head (tail (merge _ b)) = head b +tail (tail (merge a b)) = merge (tail a) (tail b) + +mapS-id : ∀ {A} {xs : Stream A} → mapS (λ x → x) xs ≡ xs +head (mapS-id {xs = xs} i) = head xs +tail (mapS-id {xs = xs} i) = mapS-id {xs = tail xs} i + +Stream-η : ∀ {A} {xs : Stream A} → xs ≡ (head xs , tail xs) +head (Stream-η {A} {xs} i) = head xs +tail (Stream-η {A} {xs} i) = tail xs + +elimS : ∀ {A} (P : Stream A → Type₀) (c : ∀ x xs → P (x , xs)) (xs : Stream A) → P xs +elimS P c xs = transp (λ i → P (Stream-η {xs = xs} (~ i))) i0 + (c (head xs) (tail xs)) + +odd≡even∘tail : ∀ {A} → (a : Stream A) → odd a ≡ even (tail a) +head (odd≡even∘tail a i) = head (tail a) +tail (odd≡even∘tail a i) = odd≡even∘tail (tail (tail a)) i + +mergeEvenOdd≡id : ∀ {A} → (a : Stream A) → merge (even a) (odd a) ≡ a +head (mergeEvenOdd≡id a i) = head a +head (tail (mergeEvenOdd≡id a i)) = head (tail a) +tail (tail (mergeEvenOdd≡id a i)) = mergeEvenOdd≡id (tail (tail a)) i + +module Equality≅Bisimulation where + +-- Bisimulation + record _≈_ {A : Type₀} (x y : Stream A) : Type₀ where + coinductive + field + ≈head : head x ≡ head y + ≈tail : tail x ≈ tail y + + open _≈_ + + bisim : {A : Type₀} → {x y : Stream A} → x ≈ y → x ≡ y + head (bisim x≈y i) = ≈head x≈y i + tail (bisim x≈y i) = bisim (≈tail x≈y) i + + misib : {A : Type₀} → {x y : Stream A} → x ≡ y → x ≈ y + ≈head (misib p) = λ i → head (p i) + ≈tail (misib p) = misib (λ i → tail (p i)) + + iso1 : {A : Type₀} → {x y : Stream A} → (p : x ≡ y) → bisim (misib p) ≡ p + head (iso1 p i j) = head (p j) + tail (iso1 p i j) = iso1 (λ i → tail (p i)) i j + + iso2 : {A : Type₀} → {x y : Stream A} → (p : x ≈ y) → misib (bisim p) ≡ p + ≈head (iso2 p i) = ≈head p + ≈tail (iso2 p i) = iso2 (≈tail p) i + + path≃bisim : {A : Type₀} → {x y : Stream A} → (x ≡ y) ≃ (x ≈ y) + path≃bisim = isoToEquiv (iso misib bisim iso2 iso1) + + path≡bisim : {A : Type₀} → {x y : Stream A} → (x ≡ y) ≡ (x ≈ y) + path≡bisim = ua path≃bisim + + -- misib can be implemented by transport as well. + refl≈ : {A : Type₀} {x : Stream A} → x ≈ x + ≈head refl≈ = refl + ≈tail refl≈ = refl≈ + + cast : ∀ {A : Type₀} {x y : Stream A} (p : x ≡ y) → x ≈ y + cast {x = x} p = transport (λ i → x ≈ p i) refl≈ + + misib-refl : ∀ {A : Type₀} {x : Stream A} → misib {x = x} refl ≡ refl≈ + ≈head (misib-refl i) = refl + ≈tail (misib-refl i) = misib-refl i + + misibTransp : ∀ {A : Type₀} {x y : Stream A} (p : x ≡ y) → cast p ≡ misib p + misibTransp p = J (λ _ p → cast p ≡ misib p) ((transportRefl refl≈) ∙ (sym misib-refl)) p + +module Stream≅Nat→ {A : Type₀} where + lookup : Stream A → ℕ → A + lookup xs zero = head xs + lookup xs (suc n) = lookup (tail xs) n + + tabulate : (ℕ → A) → Stream A + head (tabulate f) = f zero + tail (tabulate f) = tabulate (λ n → f (suc n)) + + lookup∘tabulate : (λ (x : _ → A) → lookup (tabulate x)) ≡ (λ x → x) + lookup∘tabulate i f zero = f zero + lookup∘tabulate i f (suc n) = lookup∘tabulate i (λ n → f (suc n)) n + + tabulate∘lookup : (λ (x : Stream A) → tabulate (lookup x)) ≡ (λ x → x) + head (tabulate∘lookup i xs) = head xs + tail (tabulate∘lookup i xs) = tabulate∘lookup i (tail xs) + + Stream≡Nat→ : Stream A ≡ (ℕ → A) + Stream≡Nat→ = isoToPath (iso lookup tabulate (λ f i → lookup∘tabulate i f) (λ xs i → tabulate∘lookup i xs)) diff --git a/Cubical/Core/Everything.agda b/Cubical/Core/Everything.agda new file mode 100644 index 0000000000..9ce877e584 --- /dev/null +++ b/Cubical/Core/Everything.agda @@ -0,0 +1,11 @@ +{-# OPTIONS --safe #-} +module Cubical.Core.Everything where + +-- Basic primitives (some are from Agda.Primitive) +open import Cubical.Core.Primitives public + +-- Definition of equivalences and Glue types +open import Cubical.Core.Glue public + +-- Definition of cubical Identity types +open import Cubical.Core.Id diff --git a/Cubical/Core/Glue.agda b/Cubical/Core/Glue.agda new file mode 100644 index 0000000000..f214a683ba --- /dev/null +++ b/Cubical/Core/Glue.agda @@ -0,0 +1,141 @@ +{- + +This file contains: + +- Definitions equivalences + +- Glue types + +-} +{-# OPTIONS --safe #-} +module Cubical.Core.Glue where + +open import Cubical.Core.Primitives + +open import Agda.Builtin.Cubical.Glue public + using ( isEquiv -- ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ ⊔ ℓ') + + ; equiv-proof -- ∀ (y : B) → isContr (fiber f y) + + ; _≃_ -- ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Type (ℓ ⊔ ℓ') + + ; equivFun -- ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ≃ B → A → B + + ; equivProof -- ∀ {ℓ ℓ'} (T : Type ℓ) (A : Type ℓ') (w : T ≃ A) (a : A) φ → + -- Partial φ (fiber (equivFun w) a) → fiber (equivFun w) a + + ; primGlue -- ∀ {ℓ ℓ'} (A : Type ℓ) {φ : I} (T : Partial φ (Type ℓ')) + -- → (e : PartialP φ (λ o → T o ≃ A)) → Type ℓ' + + ; prim^unglue -- ∀ {ℓ ℓ'} {A : Type ℓ} {φ : I} {T : Partial φ (Type ℓ')} + -- → {e : PartialP φ (λ o → T o ≃ A)} → primGlue A T e → A + + -- The ∀ operation on I. This is commented out as it is not currently used for anything + -- ; primFaceForall -- (I → I) → I + ) + renaming ( prim^glue to glue -- ∀ {ℓ ℓ'} {A : Type ℓ} {φ : I} {T : Partial φ (Type ℓ')} + -- → {e : PartialP φ (λ o → T o ≃ A)} + -- → PartialP φ T → A → primGlue A T e + + ; pathToEquiv to lineToEquiv -- ∀ {ℓ : I → Level} (P : (i : I) → Type (ℓ i)) → P i0 ≃ P i1 + ) + +private + variable + ℓ ℓ' : Level + +-- Uncurry Glue to make it more pleasant to use +Glue : (A : Type ℓ) {φ : I} + → (Te : Partial φ (Σ[ T ∈ Type ℓ' ] T ≃ A)) + → Type ℓ' +Glue A Te = primGlue A (λ x → Te x .fst) (λ x → Te x .snd) + +-- Make the φ argument of prim^unglue explicit +unglue : {A : Type ℓ} (φ : I) {T : Partial φ (Type ℓ')} + {e : PartialP φ (λ o → T o ≃ A)} → primGlue A T e → A +unglue φ = prim^unglue {φ = φ} + +-- People unfamiliar with [Glue], [glue] and [uglue] can find the types below more +-- informative as they demonstrate the computational behavior. +-- +-- Full inference rules can be found in Section 6 of CCHM: +-- https://arxiv.org/pdf/1611.02108.pdf +-- Cubical Type Theory: a constructive interpretation of the univalence axiom +-- Cyril Cohen, Thierry Coquand, Simon Huber, Anders Mörtberg +private + + module GluePrims (A : Type ℓ) {φ : I} (Te : Partial φ (Σ[ T ∈ Type ℓ' ] T ≃ A)) where + T : Partial φ (Type ℓ') + T φ1 = Te φ1 .fst + e : PartialP φ (λ φ → T φ ≃ A) + e φ1 = Te φ1 .snd + + -- Glue can be seen as a subtype of Type that, at φ, is definitionally equal to the left type + -- of the given equivalences. + Glue-S : Type ℓ' [ φ ↦ T ] + Glue-S = inS (Glue A Te) + + -- Which means partial elements of T are partial elements of Glue + coeT→G : + ∀ (t : PartialP φ T) + → Partial φ (Glue A Te) + coeT→G t (φ = i1) = t 1=1 + + -- ... and elements of Glue can be seen as partial elements of T + coeG→T : + ∀ (g : Glue A Te) + → PartialP φ T + coeG→T g (φ = i1) = g + + -- What about elements that are applied to the equivalences? + trans-e : + ∀ (t : PartialP φ T) + → Partial φ A + trans-e t ϕ1 = equivFun (e ϕ1) (t ϕ1) + + -- glue gives a partial element of Glue given an element of A. Note that it "undoes" + -- the application of the equivalences! + glue-S : + ∀ (t : PartialP φ T) + → A [ φ ↦ trans-e t ] + → Glue A Te [ φ ↦ coeT→G t ] + glue-S t s = inS (glue t (outS s)) + + -- typechecking glue enforces this, e.g. you can not simply write + -- glue-bad : (t : PartialP φ T) → A → Glue A Te + -- glue-bad t s = glue t s + + -- unglue does the inverse: + unglue-S : + ∀ (b : Glue A Te) + → A [ φ ↦ trans-e (coeG→T b) ] + unglue-S b = inS (unglue φ b) + + module GlueTransp (A : I → Type ℓ) (Te : (i : I) → Partial (i ∨ ~ i) (Σ[ T ∈ Type ℓ' ] T ≃ A i)) where + A0 A1 : Type ℓ + A0 = A i0 + A1 = A i1 + T : (i : I) → Partial (i ∨ ~ i) (Type ℓ') + T i φ = Te i φ .fst + e : (i : I) → PartialP (i ∨ ~ i) (λ φ → T i φ ≃ A i) + e i φ = Te i φ .snd + T0 T1 : Type ℓ' + T0 = T i0 1=1 + T1 = T i1 1=1 + e0 : T0 ≃ A0 + e0 = e i0 1=1 + e1 : T1 ≃ A1 + e1 = e i1 1=1 + + open import Cubical.Foundations.Prelude using (transport) + transportA : A0 → A1 + transportA = transport (λ i → A i) + + -- copied over from Foundations/Equiv for readability, can't directly import due to cyclic dependency + invEq : ∀ {X : Type ℓ'} {ℓ''} {Y : Type ℓ''} (w : X ≃ Y) → Y → X + invEq w y = w .snd .equiv-proof y .fst .fst + + -- transport in Glue reduces to transport in A + the application of the equivalences in forward and backward + -- direction. + transp-S : (t0 : T0) → T1 [ i1 ↦ (λ _ → invEq e1 (transportA (equivFun e0 t0))) ] + transp-S t0 = inS (transport (λ i → Glue (A i) (Te i)) t0) diff --git a/Cubical/Core/Id.agda b/Cubical/Core/Id.agda new file mode 100644 index 0000000000..8229e5a75d --- /dev/null +++ b/Cubical/Core/Id.agda @@ -0,0 +1,24 @@ +{- This file exports the primitives of cubical Id types -} +{-# OPTIONS --safe #-} +module Cubical.Core.Id where + +open import Cubical.Core.Primitives hiding ( _≡_ ) + +open import Agda.Builtin.Cubical.Id public + renaming ( conid to ⟨_,_⟩ + -- TODO: should the user really be able to access these two? + ; primIdFace to faceId -- ∀ {ℓ} {A : Type ℓ} {x y : A} → Id x y → I + ; primIdPath to pathId -- ∀ {ℓ} {A : Type ℓ} {x y : A} → Id x y → Path A x y + + ; primIdElim to elimId -- ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} + -- (P : ∀ (y : A) → x ≡ y → Type ℓ') + -- (h : ∀ (φ : I) (y : A [ φ ↦ (λ _ → x) ]) + -- (w : (Path _ x (outS y)) [ φ ↦ (λ { (φ = i1) → λ _ → x}) ] ) → + -- P (outS y) ⟨ φ , outS w ⟩) → + -- {y : A} (w' : x ≡ y) → P y w' + ) + hiding ( primIdJ ) -- this should not be used as it is using compCCHM + +{- BUILTIN ID Id -} +_≡_ : ∀ {ℓ} {A : Type ℓ} → A → A → Type ℓ +_≡_ = Id diff --git a/Cubical/Core/Primitives.agda b/Cubical/Core/Primitives.agda new file mode 100644 index 0000000000..da206a9843 --- /dev/null +++ b/Cubical/Core/Primitives.agda @@ -0,0 +1,211 @@ +{- + +This file document and export the main primitives of Cubical Agda. It +also defines some basic derived operations (composition and filling). + +-} +{-# OPTIONS --safe #-} +module Cubical.Core.Primitives where + +open import Agda.Builtin.Cubical.Path public +open import Agda.Builtin.Cubical.Sub public + renaming ( inc to inS + ; primSubOut to outS + ) +open import Agda.Primitive.Cubical public + renaming ( primIMin to _∧_ -- I → I → I + ; primIMax to _∨_ -- I → I → I + ; primINeg to ~_ -- I → I + ; isOneEmpty to empty + ; primComp to comp + ; primHComp to hcomp + ; primTransp to transp + ; itIsOne to 1=1 ) + +-- These two are to make sure all the primitives are loaded and ready +-- to compute hcomp/transp for the universe. +import Agda.Builtin.Cubical.Glue +-- HCompU is already imported from Glue, and older Agda versions do +-- not have it. So we comment it out for now. +-- import Agda.Builtin.Cubical.HCompU + +open import Agda.Primitive public + using ( Level ) + renaming ( lzero to ℓ-zero + ; lsuc to ℓ-suc + ; _⊔_ to ℓ-max + ; Set to Type + ; Setω to Typeω ) +open import Agda.Builtin.Sigma public + +-- This file document the Cubical Agda primitives. The primitives +-- themselves are bound by the Agda files imported above. + +-- * The Interval +-- I : Typeω + +-- Endpoints, Connections, Reversal +-- i0 i1 : I +-- _∧_ _∨_ : I → I → I +-- ~_ : I → I + + +-- * Dependent path type. (Path over Path) + +-- Introduced with lambda abstraction and eliminated with application, +-- just like function types. + +-- PathP : ∀ {ℓ} (A : I → Type ℓ) → A i0 → A i1 → Type ℓ + +infix 4 _[_≡_] + +_[_≡_] : ∀ {ℓ} (A : I → Type ℓ) → A i0 → A i1 → Type ℓ +_[_≡_] = PathP + + +-- Non dependent path types + +Path : ∀ {ℓ} (A : Type ℓ) → A → A → Type ℓ +Path A a b = PathP (λ _ → A) a b + +-- PathP (λ i → A) x y gets printed as x ≡ y when A does not mention i. +-- _≡_ : ∀ {ℓ} {A : Type ℓ} → A → A → Type ℓ +-- _≡_ {A = A} = PathP (λ _ → A) + + +-- * @IsOne r@ represents the constraint "r = i1". +-- Often we will use "φ" for elements of I, when we intend to use them +-- with IsOne (or Partial[P]). +-- IsOne : I → Typeω + +-- i1 is indeed equal to i1. +-- 1=1 : IsOne i1 + + +-- * Types of partial elements, and their dependent version. + +-- "Partial φ A" is a special version of "IsOne φ → A" with a more +-- extensional judgmental equality. +-- "PartialP φ A" allows "A" to be defined only on "φ". + +-- Partial : ∀ {ℓ} → I → Type ℓ → Typeω +-- PartialP : ∀ {ℓ} → (φ : I) → Partial φ (Type ℓ) → Typeω + +-- Partial elements are introduced by pattern matching with (r = i0) +-- or (r = i1) constraints, like so: + +private + sys : ∀ i → Partial (i ∨ ~ i) Type₁ + sys i (i = i0) = Type₀ + sys i (i = i1) = Type₀ → Type₀ + + -- It also works with pattern matching lambdas: + -- http://wiki.portal.chalmers.se/agda/pmwiki.php?n=ReferenceManual.PatternMatchingLambdas + sys' : ∀ i → Partial (i ∨ ~ i) Type₁ + sys' i = λ { (i = i0) → Type₀ + ; (i = i1) → Type₀ → Type₀ + } + + -- When the cases overlap they must agree. + sys2 : ∀ i j → Partial (i ∨ (i ∧ j)) Type₁ + sys2 i j = λ { (i = i1) → Type₀ + ; (i = i1) (j = i1) → Type₀ + } + + -- (i0 = i1) is actually absurd. + sys3 : Partial i0 Type₁ + sys3 = λ { () } + + +-- * There are cubical subtypes as in CCHM. Note that these are not +-- fibrant (hence in Typeω): + +_[_↦_] : ∀ {ℓ} (A : Type ℓ) (φ : I) (u : Partial φ A) → _ +A [ φ ↦ u ] = Sub A φ u + +infix 4 _[_↦_] + +-- Any element u : A can be seen as an element of A [ φ ↦ u ] which +-- agrees with u on φ: + +-- inS : ∀ {ℓ} {A : Type ℓ} {φ} (u : A) → A [ φ ↦ (λ _ → u) ] + +-- One can also forget that an element agrees with u on φ: + +-- outS : ∀ {ℓ} {A : Type ℓ} {φ : I} {u : Partial φ A} → A [ φ ↦ u ] → A + + +-- * Composition operation according to [CCHM 18]. +-- When calling "comp A φ u a" Agda makes sure that "a" agrees with "u i0" on "φ". +-- compCCHM : ∀ {ℓ} (A : (i : I) → Type ℓ) (φ : I) (u : ∀ i → Partial φ (A i)) (a : A i0) → A i1 + +-- Note: this is not recommended to use, instead use the CHM +-- primitives! The reason is that these work with HITs and produce +-- fewer empty systems. + + +-- * Generalized transport and homogeneous composition [CHM 18]. + +-- When calling "transp A φ a" Agda makes sure that "A" is constant on "φ" (see below). +-- transp : ∀ {ℓ} (A : I → Type ℓ) (φ : I) (a : A i0) → A i1 + +-- "A" being constant on "φ" means that "A" should be a constant function whenever the +-- constraint "φ = i1" is satisfied. For example: +-- - If "φ" is "i0" then "A" can be anything, since this condition is vacuously true. +-- - If "φ" is "i1" then "A" must be a constant function. +-- - If "φ" is some in-scope variable "i" then "A" only needs to be a constant function +-- when substituting "i1" for "i". + +-- When calling "hcomp A φ u a" Agda makes sure that "a" agrees with "u i0" on "φ". +-- hcomp : ∀ {ℓ} {A : Type ℓ} {φ : I} (u : I → Partial φ A) (a : A) → A + +private + variable + ℓ : Level + ℓ′ : I → Level + +-- Homogeneous filling +hfill : {A : Type ℓ} + {φ : I} + (u : ∀ i → Partial φ A) + (u0 : A [ φ ↦ u i0 ]) + ----------------------- + (i : I) → A +hfill {φ = φ} u u0 i = + hcomp (λ j → λ { (φ = i1) → u (i ∧ j) 1=1 + ; (i = i0) → outS u0 }) + (outS u0) + +-- Heterogeneous composition can defined as in CHM, however we use the +-- builtin one as it doesn't require u0 to be a cubical subtype. This +-- reduces the number of inS's a lot. +-- comp : (A : ∀ i → Type (ℓ′ i)) +-- {φ : I} +-- (u : ∀ i → Partial φ (A i)) +-- (u0 : A i0 [ φ ↦ u i0 ]) +-- → --------------------------- +-- A i1 +-- comp A {φ = φ} u u0 = +-- hcomp (λ i → λ { (φ = i1) → transp (λ j → A (i ∨ j)) i (u _ 1=1) }) +-- (transp A i0 (outS u0)) + +-- Heterogeneous filling defined using comp +fill : (A : ∀ i → Type (ℓ′ i)) + {φ : I} + (u : ∀ i → Partial φ (A i)) + (u0 : A i0 [ φ ↦ u i0 ]) + --------------------------- + (i : I) → A i +fill A {φ = φ} u u0 i = + comp (λ j → A (i ∧ j)) + (λ j → λ { (φ = i1) → u (i ∧ j) 1=1 + ; (i = i0) → outS u0 }) + (outS u0) + +-- Σ-types +infix 2 Σ-syntax + +Σ-syntax : ∀ {ℓ ℓ'} (A : Type ℓ) (B : A → Type ℓ') → Type (ℓ-max ℓ ℓ') +Σ-syntax = Σ + +syntax Σ-syntax A (λ x → B) = Σ[ x ∈ A ] B diff --git a/Cubical/Core/README.md b/Cubical/Core/README.md new file mode 100644 index 0000000000..da6b3180d1 --- /dev/null +++ b/Cubical/Core/README.md @@ -0,0 +1,16 @@ +Core library for Cubical Agda +============================= + +This folder contains the core library for Cubical Agda. It contains +the following things: + +* **Primitives**: exposes cubical agda primitives. + +* **Glue**: definition of equivalences, Glue types and the univalence + theorem. + +* **Id**: identity types and definitions of J, funExt, univalence and + propositional truncation using Id instead of Path. + +This library is intentionally kept as minimal as possible and does not +depend on the Agda standard library. diff --git a/Cubical/Data/BinNat.agda b/Cubical/Data/BinNat.agda new file mode 100644 index 0000000000..0832a1481c --- /dev/null +++ b/Cubical/Data/BinNat.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.BinNat where + +open import Cubical.Data.BinNat.BinNat public diff --git a/Cubical/Data/BinNat/BinNat.agda b/Cubical/Data/BinNat/BinNat.agda new file mode 100644 index 0000000000..1d2eca79b8 --- /dev/null +++ b/Cubical/Data/BinNat/BinNat.agda @@ -0,0 +1,492 @@ +{- Binary natural numbers (Anders Mörtberg, Jan. 2019) + +This file defines two representations of binary numbers. We prove that +they are equivalent to unary numbers and univalence is then used to +transport both programs and properties between the representations. +This is an example of how having computational univalence can be +useful for practical programming. + +The first definition is [Binℕ] and the numbers are essentially lists +of 0/1 with no trailing zeroes (in little-endian format). The main +definitions and examples are: + +- Equivalence between Binℕ and ℕ ([Binℕ≃ℕ]) with an equality obtained + using univalence ([Binℕ≡ℕ]). + +- Addition on Binℕ defined by transporting addition on ℕ to Binℕ + ([_+Binℕ_]) along Binℕ≡ℕ together with a proof that addition on Binℕ + is associative obtained by transporting the proof for ℕ ([+Binℕ-assoc]). + +- Functions testing whether a binary number is odd or even in O(1) + ([oddBinℕ], [evenBinℕ]) and the corresponding functions for ℕ + obtained by transport. Proof that odd numbers are not even + transported from Binℕ to ℕ ([oddℕnotEvenℕ]). + +- An example of the structure identity principle for natural number + structures ([NatImpl]). We first prove that Binℕ≡ℕ lifts to natural + number structures ([NatImplℕ≡Binℕ]) and we then use this to + transport "+-suc : m + suc n ≡ suc (m + n)" from ℕ to Binℕ ([+Binℕ-suc]). + +- An example of program/data refinement using the structure identity + principle where we transport a property that is infeasible to prove + by computation for ℕ ([propDoubleℕ]): + + 2^20 · 2^10 = 2^5 · (2^15 · 2^10) + + from the corresponding result on Binℕ which is proved instantly by + refl ([propDoubleBinℕ]). + + +These examples are inspired from an old cubicaltt formalization: + +https://github.com/mortberg/cubicaltt/blob/master/examples/binnat.ctt + +which itself is based on an even older cubical formalization (from 2014): + +https://github.com/simhu/cubical/blob/master/examples/binnat.cub + + + +The second representation is more non-standard and inspired by: + +https://github.com/RedPRL/redtt/blob/master/library/cool/nats.red + +Only some of the experiments have been done for this representation, +but it has the virtue of being a bit simpler to prove equivalent to +ℕ. The same representation can be found in: + +http://www.cs.bham.ac.uk/~mhe/agda-new/BinaryNaturals.html + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Data.BinNat.BinNat where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Empty + +open import Cubical.Relation.Nullary + +data Binℕ : Type₀ +data Pos : Type₀ + +-- Binary natural numbers +data Binℕ where + binℕ0 : Binℕ + binℕpos : Pos → Binℕ + +-- Positive binary numbers +data Pos where + x0 : Pos → Pos + x1 : Binℕ → Pos + +pattern pos1 = x1 binℕ0 +pattern x1-pos n = x1 (binℕpos n) + +-- Note on notation: +-- We use "⇒" for functions that are equivalences (and therefore +-- they don't preserve the numerical value where the ranges don't +-- match, as with Binℕ⇒Pos). +-- +-- We use "→" for the opposite situation (numerical value is preserved, +-- but the function is not necessarily an equivalence) +Binℕ⇒Pos : Binℕ → Pos +sucPos : Pos → Pos +Binℕ⇒Pos binℕ0 = pos1 +Binℕ⇒Pos (binℕpos n) = sucPos n +sucPos (x0 ps) = x1-pos ps +sucPos (x1 ps) = x0 (Binℕ⇒Pos ps) + +Binℕ→ℕ : Binℕ → ℕ +Pos⇒ℕ : Pos → ℕ +Pos→ℕ : Pos → ℕ +Binℕ→ℕ binℕ0 = zero +Binℕ→ℕ (binℕpos x) = Pos→ℕ x +Pos→ℕ ps = suc (Pos⇒ℕ ps) + +Pos⇒ℕ (x0 ps) = suc (doubleℕ (Pos⇒ℕ ps)) +Pos⇒ℕ (x1 ps) = doubleℕ (Binℕ→ℕ ps) + +posInd : {P : Pos → Type₀} → P pos1 → ((p : Pos) → P p → P (sucPos p)) → (p : Pos) → P p +posInd {P} h1 hs ps = f ps + where + H : (p : Pos) → P (x0 p) → P (x0 (sucPos p)) + H p hx0p = hs (x1-pos p) (hs (x0 p) hx0p) + + f : (ps : Pos) → P ps + f pos1 = h1 + f (x0 ps) = posInd (hs pos1 h1) H ps + f (x1-pos ps) = hs (x0 ps) (posInd (hs pos1 h1) H ps) + +Binℕ⇒Pos⇒ℕ : (p : Binℕ) → Pos⇒ℕ (Binℕ⇒Pos p) ≡ Binℕ→ℕ p +Binℕ⇒Pos⇒ℕ binℕ0 = refl +Binℕ⇒Pos⇒ℕ (binℕpos (x0 p)) = refl +Binℕ⇒Pos⇒ℕ (binℕpos (x1 x)) = λ i → suc (doubleℕ (Binℕ⇒Pos⇒ℕ x i)) + +Pos⇒ℕsucPos : (p : Pos) → Pos⇒ℕ (sucPos p) ≡ suc (Pos⇒ℕ p) +Pos⇒ℕsucPos p = Binℕ⇒Pos⇒ℕ (binℕpos p) + +Pos→ℕsucPos : (p : Pos) → Pos→ℕ (sucPos p) ≡ suc (Pos→ℕ p) +Pos→ℕsucPos p = cong suc (Binℕ⇒Pos⇒ℕ (binℕpos p)) + +ℕ⇒Pos : ℕ → Pos +ℕ⇒Pos zero = pos1 +ℕ⇒Pos (suc n) = sucPos (ℕ⇒Pos n) + +ℕ→Pos : ℕ → Pos +ℕ→Pos zero = pos1 +ℕ→Pos (suc n) = ℕ⇒Pos n + +Pos⇒ℕ⇒Pos : (p : Pos) → ℕ⇒Pos (Pos⇒ℕ p) ≡ p +Pos⇒ℕ⇒Pos p = posInd refl hs p + where + hs : (p : Pos) → ℕ⇒Pos (Pos⇒ℕ p) ≡ p → ℕ⇒Pos (Pos⇒ℕ (sucPos p)) ≡ sucPos p + hs p hp = + ℕ⇒Pos (Pos⇒ℕ (sucPos p)) ≡⟨ cong ℕ⇒Pos (Pos⇒ℕsucPos p) ⟩ + sucPos (ℕ⇒Pos (Pos⇒ℕ p)) ≡⟨ cong sucPos hp ⟩ + sucPos p ∎ + +ℕ⇒Pos⇒ℕ : (n : ℕ) → Pos⇒ℕ (ℕ⇒Pos n) ≡ n +ℕ⇒Pos⇒ℕ zero = refl +ℕ⇒Pos⇒ℕ (suc n) = + Pos⇒ℕ (ℕ⇒Pos (suc n)) ≡⟨ Pos⇒ℕsucPos (ℕ⇒Pos n) ⟩ + suc (Pos⇒ℕ (ℕ⇒Pos n)) ≡⟨ cong suc (ℕ⇒Pos⇒ℕ n) ⟩ + suc n ∎ + +ℕ→Binℕ : ℕ → Binℕ +ℕ→Binℕ zero = binℕ0 +ℕ→Binℕ (suc n) = binℕpos (ℕ⇒Pos n) + +ℕ→Binℕ→ℕ : (n : ℕ) → Binℕ→ℕ (ℕ→Binℕ n) ≡ n +ℕ→Binℕ→ℕ zero = refl +ℕ→Binℕ→ℕ (suc n) = cong suc (ℕ⇒Pos⇒ℕ n) + +Binℕ→ℕ→Binℕ : (n : Binℕ) → ℕ→Binℕ (Binℕ→ℕ n) ≡ n +Binℕ→ℕ→Binℕ binℕ0 = refl +Binℕ→ℕ→Binℕ (binℕpos p) = cong binℕpos (Pos⇒ℕ⇒Pos p) + +Binℕ≃ℕ : Binℕ ≃ ℕ +Binℕ≃ℕ = isoToEquiv (iso Binℕ→ℕ ℕ→Binℕ ℕ→Binℕ→ℕ Binℕ→ℕ→Binℕ) + +-- Use univalence (in fact only "ua") to get an equality from the +-- above equivalence +Binℕ≡ℕ : Binℕ ≡ ℕ +Binℕ≡ℕ = ua Binℕ≃ℕ + +sucBinℕ : Binℕ → Binℕ +sucBinℕ x = binℕpos (Binℕ⇒Pos x) + +Binℕ→ℕsuc : (x : Binℕ) → suc (Binℕ→ℕ x) ≡ Binℕ→ℕ (sucBinℕ x) +Binℕ→ℕsuc binℕ0 = refl +Binℕ→ℕsuc (binℕpos x) = sym (Pos→ℕsucPos x) + +-- We can transport addition on ℕ to Binℕ +_+Binℕ_ : Binℕ → Binℕ → Binℕ +_+Binℕ_ = transport (λ i → Binℕ≡ℕ (~ i) → Binℕ≡ℕ (~ i) → Binℕ≡ℕ (~ i)) _+_ + +-- Test: 4 + 1 = 5 +private + _ : binℕpos (x0 (x0 pos1)) +Binℕ binℕpos pos1 ≡ binℕpos (x1-pos (x0 pos1)) + _ = refl + +-- It is easy to test if binary numbers are odd +oddBinℕ : Binℕ → Bool +oddBinℕ binℕ0 = false +oddBinℕ (binℕpos (x0 _)) = false +oddBinℕ (binℕpos (x1 _)) = true + +evenBinℕ : Binℕ → Bool +evenBinℕ n = oddBinℕ (sucBinℕ n) + +-- And prove the following property (without induction) +oddBinℕnotEvenBinℕ : (n : Binℕ) → oddBinℕ n ≡ not (evenBinℕ n) +oddBinℕnotEvenBinℕ binℕ0 = refl +oddBinℕnotEvenBinℕ (binℕpos (x0 x)) = refl +oddBinℕnotEvenBinℕ (binℕpos (x1 x)) = refl + +-- It is also easy to define and prove the property for unary numbers, +-- however the definition uses recursion and the proof induction +private + oddn : ℕ → Bool + oddn zero = true + oddn (suc x) = not (oddn x) + + evenn : ℕ → Bool + evenn n = not (oddn n) + + oddnSuc : (n : ℕ) → oddn n ≡ not (evenn n) + oddnSuc zero = refl + oddnSuc (suc n) = cong not (oddnSuc n) + + +-- So what we can do instead is to transport the odd test from Binℕ to +-- ℕ along the equality +oddℕ : ℕ → Bool +oddℕ = transport (λ i → Binℕ≡ℕ i → Bool) oddBinℕ + +evenℕ : ℕ → Bool +evenℕ = transport (λ i → Binℕ≡ℕ i → Bool) evenBinℕ + +-- We can then also transport the property +oddℕnotEvenℕ : (n : ℕ) → oddℕ n ≡ not (evenℕ n) +oddℕnotEvenℕ = + let -- We first build a path from oddBinℕ to oddℕ. When i=1 this is + -- "transp (λ j → Binℕ≡ℕ j → Bool) i0 oddBinℕ" (i.e. oddℕ) + oddp : PathP (λ i → Binℕ≡ℕ i → Bool) oddBinℕ oddℕ + oddp i = transp (λ j → Binℕ≡ℕ (i ∧ j) → Bool) (~ i) oddBinℕ + + -- We then build a path from evenBinℕ to evenℕ + evenp : PathP (λ i → Binℕ≡ℕ i → Bool) evenBinℕ evenℕ + evenp i = transp (λ j → Binℕ≡ℕ (i ∧ j) → Bool) (~ i) evenBinℕ + in -- Then transport oddBinℕnotEvenBinℕ in a suitable equality type + -- When i=0 this is "(n : Binℕ) → oddBinℕ n ≡ not (evenBinℕ n)" + -- When i=1 this is "(n : ℕ) → oddℕ n ≡ not (evenℕ n)" + transport (λ i → (n : Binℕ≡ℕ i) → oddp i n ≡ not (evenp i n)) oddBinℕnotEvenBinℕ + + +-- We can do the same for natural numbers: + +-- First construct the path +addp : PathP (λ i → Binℕ≡ℕ (~ i) → Binℕ≡ℕ (~ i) → Binℕ≡ℕ (~ i)) _+_ _+Binℕ_ +addp i = transp (λ j → Binℕ≡ℕ (~ i ∨ ~ j) → Binℕ≡ℕ (~ i ∨ ~ j) → Binℕ≡ℕ (~ i ∨ ~ j)) (~ i) _+_ + +-- Then transport associativity: ++Binℕ-assoc : ∀ m n o → m +Binℕ (n +Binℕ o) ≡ (m +Binℕ n) +Binℕ o ++Binℕ-assoc = + transport (λ i → (m n o : Binℕ≡ℕ (~ i)) + → addp i m (addp i n o) ≡ addp i (addp i m n) o) +-assoc + + +-- We can also define what it means to be an implementation of natural +-- numbers and use this to transport properties between different +-- implementation of natural numbers. This can be seen as a special +-- case of the structure identity principle: any property that holds +-- for one structure also holds for an equivalent one. + +-- An implementation of natural numbers (i.e. a "natural number +-- structure") has a zero and successor. +record NatImpl (A : Type₀) : Type₀ where + field + z : A + s : A → A +open NatImpl + +NatImplℕ : NatImpl ℕ +z NatImplℕ = zero +s NatImplℕ = suc + +NatImplBinℕ : NatImpl Binℕ +z NatImplBinℕ = binℕ0 +s NatImplBinℕ = sucBinℕ + +-- Using the equality between binary and unary numbers we can get an +-- equality between the two implementations of the NatImpl interface +NatImplℕ≡Binℕ : PathP (λ i → NatImpl (Binℕ≡ℕ (~ i))) NatImplℕ NatImplBinℕ +z (NatImplℕ≡Binℕ i) = transp (λ j → Binℕ≡ℕ (~ i ∨ ~ j)) (~ i) zero +s (NatImplℕ≡Binℕ i) = + λ x → glue (λ { (i = i0) → suc x + ; (i = i1) → sucBinℕ x }) + -- We need to do use and hcomp to do and endpoint + -- correction as "suc (unglue x)" connects "suc x" + -- with "suc (Binℕ→ℕ x)" along i (which makes sense as + -- x varies from ℕ to Binℕ along i), but we need + -- something from "suc x" to "Binℕ→ℕ (sucBinℕ x)" for + -- the glue to be well-formed + (hcomp (λ j → λ { (i = i0) → suc x + ; (i = i1) → Binℕ→ℕsuc x j }) + (suc (unglue (i ∨ ~ i) x))) + +-- We then use this to transport +-suc from unary to binary numbers ++Binℕ-suc : ∀ m n → m +Binℕ sucBinℕ n ≡ sucBinℕ (m +Binℕ n) ++Binℕ-suc = + transport (λ i → (m n : Binℕ≡ℕ (~ i)) + → addp i m (NatImplℕ≡Binℕ i .s n) ≡ NatImplℕ≡Binℕ i .s (addp i m n)) +-suc + + + +-- Doubling experiment: we define a notion of "doubling structure" and +-- transport a proof that is proved directly using refl for binary +-- numbers to unary numbers. This is an example of program/data +-- refinement: we can use univalence to prove properties about +-- inefficient data-structures using efficient ones. + +-- Doubling structures +record Double {ℓ} (A : Type ℓ) : Type (ℓ-suc ℓ) where + field + -- doubling function computing 2 · x + double : A → A + -- element to double + elt : A +open Double + +-- Compute: 2^n · x +doubles : ∀ {ℓ} {A : Type ℓ} (D : Double A) → ℕ → A → A +doubles D n x = iter n (double D) x + +Doubleℕ : Double ℕ +double Doubleℕ = doubleℕ +elt Doubleℕ = n1024 + where + -- 1024 = 2^8 · 2^2 = 2^10 + n1024 : ℕ + n1024 = doublesℕ 8 4 + +-- The doubling operation on binary numbers is O(1), while for unary +-- numbers it is O(n). What is of course even more problematic is that +-- we cannot handle very big unary natural numbers, but with binary +-- there is no problem to represent very big numbers +doubleBinℕ : Binℕ → Binℕ +doubleBinℕ binℕ0 = binℕ0 +doubleBinℕ (binℕpos x) = binℕpos (x0 x) + +DoubleBinℕ : Double Binℕ +double DoubleBinℕ = doubleBinℕ +elt DoubleBinℕ = bin1024 + where + -- 1024 = 2^10 = 10000000000₂ + bin1024 : Binℕ + bin1024 = binℕpos (x0 (x0 (x0 (x0 (x0 (x0 (x0 (x0 (x0 (x0 pos1)))))))))) + +-- As these function don't commute strictly we have to prove it +-- separately and insert it in the proof of DoubleBinℕ≡Doubleℕ below +-- (just like we had to in NatImplℕ≡NatImplBinℕ +Binℕ→ℕdouble : (x : Binℕ) → doubleℕ (Binℕ→ℕ x) ≡ Binℕ→ℕ (doubleBinℕ x) +Binℕ→ℕdouble binℕ0 = refl +Binℕ→ℕdouble (binℕpos x) = refl + +-- We use the equality between Binℕ and ℕ to get an equality of +-- doubling structures +DoubleBinℕ≡Doubleℕ : PathP (λ i → Double (Binℕ≡ℕ i)) DoubleBinℕ Doubleℕ +double (DoubleBinℕ≡Doubleℕ i) = + λ x → glue (λ { (i = i0) → doubleBinℕ x + ; (i = i1) → doubleℕ x }) + (hcomp (λ j → λ { (i = i0) → Binℕ→ℕdouble x j + ; (i = i1) → doubleℕ x }) + (doubleℕ (unglue (i ∨ ~ i) x))) +elt (DoubleBinℕ≡Doubleℕ i) = transp (λ j → Binℕ≡ℕ (i ∨ ~ j)) i (Doubleℕ .elt) + +-- We can now use transport to prove a property that is too slow to +-- check with unary numbers. We define the property we want to check +-- as a record so that Agda does not try to unfold it eagerly. +record propDouble {ℓ} {A : Type ℓ} (D : Double A) : Type ℓ where + field + -- 2^20 · e = 2^5 · (2^15 · e) + proof : doubles D 20 (elt D) ≡ doubles D 5 (doubles D 15 (elt D)) +open propDouble + +-- The property we want to prove takes too long to typecheck for ℕ: +-- propDoubleℕ : propDouble Doubleℕ +-- propDoubleℕ = refl + +-- With binary numbers it is instant +propDoubleBinℕ : propDouble DoubleBinℕ +proof propDoubleBinℕ = refl + +-- By transporting the proof along the equality we then get it for +-- unary numbers +propDoubleℕ : propDouble Doubleℕ +propDoubleℕ = transport (λ i → propDouble (DoubleBinℕ≡Doubleℕ i)) propDoubleBinℕ + + +-------------------------------------------------------------------------------- +-- +-- Alternative encoding of binary natural numbers inspired by: +-- https://github.com/RedPRL/redtt/blob/master/library/cool/nats.red +-- +-- This representation makes the equivalence with ℕ a bit easier to +-- prove, but the doubling example wouldn't work as nicely as we +-- cannot define it as an O(1) operation + +data binnat : Type₀ where + zero : binnat -- 0 + consOdd : binnat → binnat -- 2·n + 1 + consEven : binnat → binnat -- 2·{n+1} + +binnat→ℕ : binnat → ℕ +binnat→ℕ zero = 0 +binnat→ℕ (consOdd n) = suc (doubleℕ (binnat→ℕ n)) +binnat→ℕ (consEven n) = suc (suc (doubleℕ (binnat→ℕ n))) + +suc-binnat : binnat → binnat +suc-binnat zero = consOdd zero +suc-binnat (consOdd n) = consEven n +suc-binnat (consEven n) = consOdd (suc-binnat n) + +ℕ→binnat : ℕ → binnat +ℕ→binnat zero = zero +ℕ→binnat (suc n) = suc-binnat (ℕ→binnat n) + +binnat→ℕ-suc : (n : binnat) → binnat→ℕ (suc-binnat n) ≡ suc (binnat→ℕ n) +binnat→ℕ-suc zero = refl +binnat→ℕ-suc (consOdd n) = refl +binnat→ℕ-suc (consEven n) = λ i → suc (doubleℕ (binnat→ℕ-suc n i)) + +ℕ→binnat→ℕ : (n : ℕ) → binnat→ℕ (ℕ→binnat n) ≡ n +ℕ→binnat→ℕ zero = refl +ℕ→binnat→ℕ (suc n) = (binnat→ℕ-suc (ℕ→binnat n)) ∙ (cong suc (ℕ→binnat→ℕ n)) + +suc-ℕ→binnat-double : (n : ℕ) → suc-binnat (ℕ→binnat (doubleℕ n)) ≡ consOdd (ℕ→binnat n) +suc-ℕ→binnat-double zero = refl +suc-ℕ→binnat-double (suc n) = λ i → suc-binnat (suc-binnat (suc-ℕ→binnat-double n i)) + +binnat→ℕ→binnat : (n : binnat) → ℕ→binnat (binnat→ℕ n) ≡ n +binnat→ℕ→binnat zero = refl +binnat→ℕ→binnat (consOdd n) = + (suc-ℕ→binnat-double (binnat→ℕ n)) ∙ (cong consOdd (binnat→ℕ→binnat n)) +binnat→ℕ→binnat (consEven n) = + (λ i → suc-binnat (suc-ℕ→binnat-double (binnat→ℕ n) i)) ∙ (cong consEven (binnat→ℕ→binnat n)) + +ℕ≃binnat : ℕ ≃ binnat +ℕ≃binnat = isoToEquiv (iso ℕ→binnat binnat→ℕ binnat→ℕ→binnat ℕ→binnat→ℕ) + +ℕ≡binnat : ℕ ≡ binnat +ℕ≡binnat = ua ℕ≃binnat + +-- We can transport addition on ℕ to binnat +_+binnat_ : binnat → binnat → binnat +_+binnat_ = transport (λ i → ℕ≡binnat i → ℕ≡binnat i → ℕ≡binnat i) _+_ + +-- Test: 4 + 1 = 5 +_ : consEven (consOdd zero) +binnat consOdd zero ≡ consOdd (consEven zero) +_ = refl + +oddbinnat : binnat → Bool +oddbinnat zero = false +oddbinnat (consOdd _) = true +oddbinnat (consEven _) = false + +oddℕ' : ℕ → Bool +oddℕ' = transport (λ i → ℕ≡binnat (~ i) → Bool) oddbinnat + +-- The NatImpl example for this representation of binary numbers +private + NatImplbinnat : NatImpl binnat + z NatImplbinnat = zero + s NatImplbinnat = suc-binnat + + -- Note that the s case is a bit simpler as no end-point correction + -- is necessary (things commute strictly) + NatImplℕ≡NatImplbinnat : PathP (λ i → NatImpl (ℕ≡binnat i)) NatImplℕ NatImplbinnat + z (NatImplℕ≡NatImplbinnat i) = transp (λ j → ℕ≡binnat (i ∨ ~ j)) i zero + s (NatImplℕ≡NatImplbinnat i) = + λ x → glue (λ { (i = i0) → suc x + ; (i = i1) → suc-binnat x }) + (suc-binnat (unglue (i ∨ ~ i) x)) + + oddSuc : (n : binnat) → oddbinnat n ≡ not (oddbinnat (suc-binnat n)) + oddSuc zero = refl + oddSuc (consOdd _) = refl + oddSuc (consEven _) = refl + + oddℕSuc' : (n : ℕ) → oddℕ' n ≡ not (oddℕ' (suc n)) + oddℕSuc' = + let eq : (i : I) → ℕ≡binnat (~ i) → Bool + eq i = transp (λ j → ℕ≡binnat (~ i ∨ ~ j) → Bool) (~ i) oddbinnat + in transport + (λ i → (n : ℕ≡binnat (~ i)) → eq i n ≡ not (eq i (NatImplℕ≡NatImplbinnat (~ i) .NatImpl.s n))) + oddSuc diff --git a/Cubical/Data/Bool.agda b/Cubical/Data/Bool.agda new file mode 100644 index 0000000000..5a084c5f9a --- /dev/null +++ b/Cubical/Data/Bool.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Bool where + +open import Cubical.Data.Bool.Base public + +open import Cubical.Data.Bool.Properties public diff --git a/Cubical/Data/Bool/Base.agda b/Cubical/Data/Bool/Base.agda new file mode 100644 index 0000000000..f2676e7354 --- /dev/null +++ b/Cubical/Data/Bool/Base.agda @@ -0,0 +1,85 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Bool.Base where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Empty +open import Cubical.Data.Sum.Base +open import Cubical.Data.Unit.Base + +open import Cubical.Relation.Nullary.Base +open import Cubical.Relation.Nullary.DecidableEq + +-- Obtain the booleans +open import Agda.Builtin.Bool public + +private + variable + ℓ : Level + A : Type ℓ + +infixr 6 _and_ +infixr 5 _or_ +infix 0 if_then_else_ + +not : Bool → Bool +not true = false +not false = true + +_or_ : Bool → Bool → Bool +false or false = false +false or true = true +true or false = true +true or true = true + +_and_ : Bool → Bool → Bool +false and false = false +false and true = false +true and false = false +true and true = true + +-- xor / mod-2 addition +_⊕_ : Bool → Bool → Bool +false ⊕ x = x +true ⊕ x = not x + +if_then_else_ : Bool → A → A → A +if true then x else y = x +if false then x else y = y + +_≟_ : Discrete Bool +false ≟ false = yes refl +false ≟ true = no λ p → subst (λ b → if b then ⊥ else Bool) p true +true ≟ false = no λ p → subst (λ b → if b then Bool else ⊥) p true +true ≟ true = yes refl + +Dec→Bool : Dec A → Bool +Dec→Bool (yes p) = true +Dec→Bool (no ¬p) = false + +-- Helpers for automatic proof +Bool→Type : Bool → Type₀ +Bool→Type true = Unit +Bool→Type false = ⊥ + +True : Dec A → Type₀ +True Q = Bool→Type (Dec→Bool Q) + +False : Dec A → Type₀ +False Q = Bool→Type (not (Dec→Bool Q)) + +toWitness : {Q : Dec A} → True Q → A +toWitness {Q = yes p} _ = p + +toWitnessFalse : {Q : Dec A} → False Q → ¬ A +toWitnessFalse {Q = no ¬p} _ = ¬p + +dichotomyBool : (x : Bool) → (x ≡ true) ⊎ (x ≡ false) +dichotomyBool true = inl refl +dichotomyBool false = inr refl + +-- TODO: this should be uncommented and implemented using instance arguments +-- _==_ : {dA : Discrete A} → A → A → Bool +-- _==_ {dA = dA} x y = Dec→Bool (dA x y) diff --git a/Cubical/Data/Bool/Properties.agda b/Cubical/Data/Bool/Properties.agda new file mode 100644 index 0000000000..c899f997d7 --- /dev/null +++ b/Cubical/Data/Bool/Properties.agda @@ -0,0 +1,221 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Bool.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Functions.Involution + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed + +open import Cubical.Data.Bool.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidableEq + +notnot : ∀ x → not (not x) ≡ x +notnot true = refl +notnot false = refl + +notIso : Iso Bool Bool +Iso.fun notIso = not +Iso.inv notIso = not +Iso.rightInv notIso = notnot +Iso.leftInv notIso = notnot + +notIsEquiv : isEquiv not +notIsEquiv = involIsEquiv {f = not} notnot + +notEquiv : Bool ≃ Bool +notEquiv = involEquiv {f = not} notnot + +notEq : Bool ≡ Bool +notEq = involPath {f = not} notnot + +private + variable + ℓ : Level + + -- This computes to false as expected + nfalse : Bool + nfalse = transp (λ i → notEq i) i0 true + + -- Sanity check + nfalsepath : nfalse ≡ false + nfalsepath = refl + +K-Bool + : (P : {b : Bool} → b ≡ b → Type ℓ) + → (∀{b} → P {b} refl) + → ∀{b} → (q : b ≡ b) → P q +K-Bool P Pr {false} = J (λ{ false q → P q ; true _ → Lift ⊥ }) Pr +K-Bool P Pr {true} = J (λ{ true q → P q ; false _ → Lift ⊥ }) Pr + +isSetBool : isSet Bool +isSetBool a b = J (λ _ p → ∀ q → p ≡ q) (K-Bool (refl ≡_) refl) + +true≢false : ¬ true ≡ false +true≢false p = subst (λ b → if b then Bool else ⊥) p true + +false≢true : ¬ false ≡ true +false≢true p = subst (λ b → if b then ⊥ else Bool) p true + +¬true→false : (x : Bool) → ¬ x ≡ true → x ≡ false +¬true→false false _ = refl +¬true→false true p = rec (p refl) + +¬false→true : (x : Bool) → ¬ x ≡ false → x ≡ true +¬false→true false p = rec (p refl) +¬false→true true _ = refl + +not≢const : ∀ x → ¬ not x ≡ x +not≢const false = true≢false +not≢const true = false≢true + +zeroˡ : ∀ x → true or x ≡ true +zeroˡ false = refl +zeroˡ true = refl + +zeroʳ : ∀ x → x or true ≡ true +zeroʳ false = refl +zeroʳ true = refl + +or-identityˡ : ∀ x → false or x ≡ x +or-identityˡ false = refl +or-identityˡ true = refl + +or-identityʳ : ∀ x → x or false ≡ x +or-identityʳ false = refl +or-identityʳ true = refl + +or-comm : ∀ x y → x or y ≡ y or x +or-comm false y = + false or y ≡⟨ or-identityˡ y ⟩ + y ≡⟨ sym (or-identityʳ y) ⟩ + y or false ∎ +or-comm true y = + true or y ≡⟨ zeroˡ y ⟩ + true ≡⟨ sym (zeroʳ y) ⟩ + y or true ∎ + +or-assoc : ∀ x y z → x or (y or z) ≡ (x or y) or z +or-assoc false y z = + false or (y or z) ≡⟨ or-identityˡ _ ⟩ + y or z ≡[ i ]⟨ or-identityˡ y (~ i) or z ⟩ + ((false or y) or z) ∎ +or-assoc true y z = + true or (y or z) ≡⟨ zeroˡ _ ⟩ + true ≡⟨ sym (zeroˡ _) ⟩ + true or z ≡[ i ]⟨ zeroˡ y (~ i) or z ⟩ + (true or y) or z ∎ + +or-idem : ∀ x → x or x ≡ x +or-idem false = refl +or-idem true = refl + +⊕-identityʳ : ∀ x → x ⊕ false ≡ x +⊕-identityʳ false = refl +⊕-identityʳ true = refl + +⊕-comm : ∀ x y → x ⊕ y ≡ y ⊕ x +⊕-comm false false = refl +⊕-comm false true = refl +⊕-comm true false = refl +⊕-comm true true = refl + +⊕-assoc : ∀ x y z → x ⊕ (y ⊕ z) ≡ (x ⊕ y) ⊕ z +⊕-assoc false y z = refl +⊕-assoc true false z = refl +⊕-assoc true true z = notnot z + +not-⊕ˡ : ∀ x y → not (x ⊕ y) ≡ not x ⊕ y +not-⊕ˡ false y = refl +not-⊕ˡ true y = notnot y + +⊕-invol : ∀ x y → x ⊕ (x ⊕ y) ≡ y +⊕-invol false x = refl +⊕-invol true x = notnot x + +isEquiv-⊕ : ∀ x → isEquiv (x ⊕_) +isEquiv-⊕ x = involIsEquiv (⊕-invol x) + +⊕-Path : ∀ x → Bool ≡ Bool +⊕-Path x = involPath {f = x ⊕_} (⊕-invol x) + +⊕-Path-refl : ⊕-Path false ≡ refl +⊕-Path-refl = isInjectiveTransport refl + +¬transportNot : ∀(P : Bool ≡ Bool) b → ¬ (transport P (not b) ≡ transport P b) +¬transportNot P b eq = not≢const b sub + where + sub : not b ≡ b + sub = subst {A = Bool → Bool} (λ f → f (not b) ≡ f b) + (λ i c → transport⁻Transport P c i) (cong (transport⁻ P) eq) + +module BoolReflection where + data Table (A : Type₀) (P : Bool ≡ A) : Type₀ where + inspect : (b c : A) + → transport P false ≡ b + → transport P true ≡ c + → Table A P + + table : ∀ P → Table Bool P + table = J Table (inspect false true refl refl) + + reflLemma : (P : Bool ≡ Bool) + → transport P false ≡ false + → transport P true ≡ true + → transport P ≡ transport (⊕-Path false) + reflLemma P ff tt i false = ff i + reflLemma P ff tt i true = tt i + + notLemma : (P : Bool ≡ Bool) + → transport P false ≡ true + → transport P true ≡ false + → transport P ≡ transport (⊕-Path true) + notLemma P ft tf i false = ft i + notLemma P ft tf i true = tf i + + categorize : ∀ P → transport P ≡ transport (⊕-Path (transport P false)) + categorize P with table P + categorize P | inspect false true p q + = subst (λ b → transport P ≡ transport (⊕-Path b)) (sym p) (reflLemma P p q) + categorize P | inspect true false p q + = subst (λ b → transport P ≡ transport (⊕-Path b)) (sym p) (notLemma P p q) + categorize P | inspect false false p q + = rec (¬transportNot P false (q ∙ sym p)) + categorize P | inspect true true p q + = rec (¬transportNot P false (q ∙ sym p)) + + ⊕-complete : ∀ P → P ≡ ⊕-Path (transport P false) + ⊕-complete P = isInjectiveTransport (categorize P) + + ⊕-comp : ∀ p q → ⊕-Path p ∙ ⊕-Path q ≡ ⊕-Path (q ⊕ p) + ⊕-comp p q = isInjectiveTransport (λ i x → ⊕-assoc q p x i) + + open Iso + + reflectIso : Iso Bool (Bool ≡ Bool) + reflectIso .fun = ⊕-Path + reflectIso .inv P = transport P false + reflectIso .leftInv = ⊕-identityʳ + reflectIso .rightInv P = sym (⊕-complete P) + + reflectEquiv : Bool ≃ (Bool ≡ Bool) + reflectEquiv = isoToEquiv reflectIso + +IsoBool→∙ : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) +Iso.fun IsoBool→∙ f = fst f false +fst (Iso.inv IsoBool→∙ a) false = a +fst (Iso.inv (IsoBool→∙ {A = A}) a) true = pt A +snd (Iso.inv IsoBool→∙ a) = refl +Iso.rightInv IsoBool→∙ a = refl +Iso.leftInv IsoBool→∙ (f , p) = + ΣPathP ((funExt (λ { false → refl ; true → sym p})) + , λ i j → p (~ i ∨ j)) diff --git a/Cubical/Data/Bool/SwitchStatement.agda b/Cubical/Data/Bool/SwitchStatement.agda new file mode 100644 index 0000000000..f86bc476dd --- /dev/null +++ b/Cubical/Data/Bool/SwitchStatement.agda @@ -0,0 +1,42 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Bool.SwitchStatement where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Bool.Base +open import Cubical.Data.Nat + +{- + Switch-case: + + _==_ : A → A → Bool + + _ : B + _ = switch (λ x → x == fixedValue) cases + case value1 ⇒ result1 break + case value2 ⇒ result2 break + ... + case valueN ⇒ resultN break + default⇒ defaultResult +-} + + +private + variable + ℓ ℓ′ : Level + + +infixr 6 default⇒_ +infixr 5 case_⇒_break_ +infixr 4 switch_cases_ + +switch_cases_ : {A : Type ℓ} {B : Type ℓ′} → (A → Bool) → ((A → Bool) → B) → B +switch caseIndicator cases caseData = caseData caseIndicator + +case_⇒_break_ : {A : Type ℓ} {B : Type ℓ′} → A → B → (otherCases : (A → Bool) → B) → (A → Bool) → B +case forValue ⇒ result break otherCases = λ caseIndicator → if (caseIndicator forValue) then result else (otherCases caseIndicator) + +default⇒_ : {A : Type ℓ} {B : Type ℓ′} → B → (A → Bool) → B +default⇒_ value caseIndicator = value diff --git a/Cubical/Data/DescendingList.agda b/Cubical/Data/DescendingList.agda new file mode 100644 index 0000000000..f98f3975e8 --- /dev/null +++ b/Cubical/Data/DescendingList.agda @@ -0,0 +1,9 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.DescendingList where + +open import Cubical.Data.DescendingList.Base public +open import Cubical.Data.DescendingList.Properties public +open import Cubical.Data.DescendingList.Examples public +open import Cubical.Data.DescendingList.Strict public + +open import Cubical.Data.DescendingList.Strict.Properties public diff --git a/Cubical/Data/DescendingList/Base.agda b/Cubical/Data/DescendingList/Base.agda new file mode 100644 index 0000000000..4ef8e9220d --- /dev/null +++ b/Cubical/Data/DescendingList/Base.agda @@ -0,0 +1,33 @@ +------------------------------------------------------------------------ +-- Descending lists +------------------------------------------------------------------------ + +{-# OPTIONS --safe #-} + +open import Cubical.Foundations.Everything + + +module Cubical.Data.DescendingList.Base + (A : Type₀) + (_≥_ : A → A → Type₀) + where + +------------------------------------------------------------------------ +-- Descending lists +-- +-- Defined simultaneously with the relation "x ≥ the HEAD of u" + +data DL : Type₀ +data _≥ᴴ_ (x : A) : DL → Type₀ + +data DL where + [] : DL + cons : (x : A) (u : DL) → x ≥ᴴ u → DL + +data _≥ᴴ_ x where + ≥ᴴ[] : x ≥ᴴ [] + ≥ᴴcons : {y : A} {u : DL} {r : y ≥ᴴ u} + → x ≥ y → x ≥ᴴ (cons y u r) + +[_] : A → DL +[ x ] = cons x [] ≥ᴴ[] diff --git a/Cubical/Data/DescendingList/Examples.agda b/Cubical/Data/DescendingList/Examples.agda new file mode 100644 index 0000000000..bc5894d725 --- /dev/null +++ b/Cubical/Data/DescendingList/Examples.agda @@ -0,0 +1,121 @@ +------------------------------------------------------------------------ +-- Descending lists +-- +-- We present some simple examples to demonstrate +-- +-- 1. the concatenation operation on sorted lists obtained by +-- transporting the one on finite multisets, and +-- +-- 2. "sorting" finite multisets by converting into sorted lists. +------------------------------------------------------------------------ + +{-# OPTIONS --safe --experimental-lossy-unification #-} + +module Cubical.Data.DescendingList.Examples where + +open import Cubical.Foundations.Everything + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidableEq + +open import Cubical.HITs.FiniteMultiset + + +------------------------------------------------------------------------ +-- The ordering relation on natural numbers +-- +-- we work with the definition from the standard library. + +infix 4 _≤_ _≥_ + +data _≤_ : ℕ → ℕ → Type where + z≤n : ∀ {n} → zero ≤ n + s≤s : ∀ {m n} (m≤n : m ≤ n) → suc m ≤ suc n + +_≥_ : ℕ → ℕ → Type +m ≥ n = n ≤ m + +≤pred : {n m : ℕ} → suc n ≤ suc m → n ≤ m +≤pred (s≤s r) = r + +≥isPropValued : {n m : ℕ} → isProp (n ≥ m) +≥isPropValued z≤n z≤n = refl +≥isPropValued (s≤s p) (s≤s q) = cong s≤s (≥isPropValued p q) + +≥dec : (x y : ℕ) → Dec (x ≥ y) +≥dec zero zero = yes z≤n +≥dec zero (suc y) = no (λ ()) +≥dec (suc x) zero = yes z≤n +≥dec (suc x) (suc y) with ≥dec x y +≥dec (suc x) (suc y) | yes p = yes (s≤s p) +≥dec (suc x) (suc y) | no ¬p = no (λ r → ¬p (≤pred r)) + +≰→≥ : {x y : ℕ} → ¬ (x ≥ y) → y ≥ x +≰→≥ {zero} {y} f = z≤n +≰→≥ {suc x} {zero} f = ⊥.rec (f z≤n) +≰→≥ {suc x} {suc y} f = s≤s (≰→≥ λ g → f (s≤s g)) + +≥trans : {x y z : ℕ} → x ≥ y → y ≥ z → x ≥ z +≥trans z≤n z≤n = z≤n +≥trans (s≤s r) z≤n = z≤n +≥trans (s≤s r) (s≤s s) = s≤s (≥trans r s) + +≤≥→≡ : {x y : ℕ} → x ≥ y → y ≥ x → x ≡ y +≤≥→≡ z≤n z≤n = refl +≤≥→≡ (s≤s r) (s≤s s) = cong suc (≤≥→≡ r s) + + +open import Cubical.Data.DescendingList.Base ℕ _≥_ +open import Cubical.Data.DescendingList.Properties ℕ _≥_ +open isSetDL ≥isPropValued discreteℕ +open IsoToFMSet discreteℕ ≥dec ≥isPropValued ≥trans ≰→≥ ≤≥→≡ + +------------------------------------------------------------------------ +-- A simple example of concatenating sorted lists + +l1 : DL +l1 = cons 4 (cons 3 (cons 2 [] ≥ᴴ[]) (≥ᴴcons (s≤s (s≤s z≤n)))) (≥ᴴcons (s≤s (s≤s (s≤s z≤n)))) + +l2 : DL +l2 = cons 2 (cons 0 [] ≥ᴴ[]) (≥ᴴcons z≤n) + +l3 : DL +l3 = l1 ++ᴰᴸ l2 + +ValueOfl3 : l3 ≡ cons 4 (cons 3 (cons 2 (cons 2 (cons 0 _ _) _) _) _) _ +ValueOfl3 = refl + +l3=l2++l1 : l3 ≡ l2 ++ᴰᴸ l1 +l3=l2++l1 = refl + +-- Commented as it was the slowest definition in the whole library :-) +-- LongerExample : l1 ++ᴰᴸ l2 ++ᴰᴸ l1 ++ᴰᴸ l1 ++ᴰᴸ l2 +-- ≡ l2 ++ᴰᴸ l1 ++ᴰᴸ l2 ++ᴰᴸ l1 ++ᴰᴸ l1 +-- LongerExample = refl + + +-- ------------------------------------------------------------------------ +-- -- A simple example of sorting finite multisets + +m1 : FMSet ℕ +m1 = 13 ∷ 9 ∷ 78 ∷ 31 ∷ 86 ∷ 3 ∷ 0 ∷ 99 ∷ [] + +m2 : FMSet ℕ +m2 = 3 ∷ 1 ∷ 4 ∷ 1 ∷ 5 ∷ 9 ∷ [] + +m3 : FMSet ℕ +m3 = toFMSet (toDL (m1 ++ m2)) + +ValueOfm3 : + m3 ≡ 99 ∷ 86 ∷ 78 ∷ 31 ∷ 13 ∷ 9 ∷ 9 ∷ 5 ∷ 4 ∷ 3 ∷ 3 ∷ 1 ∷ 1 ∷ 0 ∷ [] +ValueOfm3 = refl + +ValueOfm1++m2 : + m1 ++ m2 ≡ 13 ∷ 9 ∷ 78 ∷ 31 ∷ 86 ∷ 3 ∷ 0 ∷ 99 ∷ 3 ∷ 1 ∷ 4 ∷ 1 ∷ 5 ∷ 9 ∷ [] +ValueOfm1++m2 = refl + +m3=m1++m2 : m3 ≡ m1 ++ m2 +m3=m1++m2 = toFMSet∘toDL≡id (m1 ++ m2) diff --git a/Cubical/Data/DescendingList/Properties.agda b/Cubical/Data/DescendingList/Properties.agda new file mode 100644 index 0000000000..1992c02cc3 --- /dev/null +++ b/Cubical/Data/DescendingList/Properties.agda @@ -0,0 +1,279 @@ +------------------------------------------------------------------------ +-- Descending lists +-- +-- Anders Mörtberg and Chuangjie Xu, October 2019 +-- +-- We define descending lists via simultaneous definitions and show +-- that they are isomorphic to finite multisets. The conversion from +-- finite multisets to descending lists is exactly insertion sort. We +-- obtain the concatenation operation on descending lists and its +-- properties by transporting those on finite multisets. +------------------------------------------------------------------------ + +{-# OPTIONS --safe #-} + +open import Cubical.Foundations.Everything + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit +open import Cubical.Data.List using (List ; [] ; _∷_) + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidableEq + +open import Cubical.HITs.FiniteMultiset as FMSet hiding ([_]) + + +module Cubical.Data.DescendingList.Properties + (A : Type₀) + (_≥_ : A → A → Type₀) + where + +open import Cubical.Data.DescendingList.Base A _≥_ + +toFMSet : DL → FMSet A +toFMSet [] = [] +toFMSet (cons x u _) = x ∷ toFMSet u + +toList : DL → List A +toList [] = [] +toList (cons x u _) = x ∷ toList u + +-- "x ≥ ALL the elements of u" +data _≥ᴬ_ (x : A) : List A → Type₀ where + ≥ᴬ[] : x ≥ᴬ [] + ≥ᴬcons : {y : A} {u : List A} → x ≥ y → x ≥ᴬ u → x ≥ᴬ (y ∷ u) + +data descending : List A → Type₀ where + []-descending : descending [] + cons-descending : {x : A} {u : List A} → x ≥ᴬ u → descending u → descending (x ∷ u) + +module DescendingDL + (≥trans : {x y z : A} → x ≥ y → y ≥ z → x ≥ z) + where + + ≥≥ᴴtrans : {x y : A} {u : DL} + → x ≥ y → y ≥ᴴ u → x ≥ᴴ u + ≥≥ᴴtrans _ ≥ᴴ[] = ≥ᴴ[] + ≥≥ᴴtrans x≥y (≥ᴴcons y≥z) = ≥ᴴcons (≥trans x≥y y≥z) + + ≥ᴴ→≥ᴬ : {x : A} {u : DL} → x ≥ᴴ u → x ≥ᴬ toList u + ≥ᴴ→≥ᴬ ≥ᴴ[] = ≥ᴬ[] + ≥ᴴ→≥ᴬ (≥ᴴcons {r = y≥u} x≥y) = ≥ᴬcons x≥y (≥ᴴ→≥ᴬ (≥≥ᴴtrans x≥y y≥u)) + + descendingDL : (u : DL) → descending (toList u) + descendingDL [] = []-descending + descendingDL (cons x u x≥u) = cons-descending (≥ᴴ→≥ᴬ x≥u) (descendingDL u) + + +------------------------------------------------------------------------ +-- Descending lists are a set. + +head≡ : {x y : A} {u v : DL} {r : x ≥ᴴ u} {s : y ≥ᴴ v} + → cons x u r ≡ cons y v s → x ≡ y +head≡ {x} = cong head + where + head : DL → A + head [] = x + head (cons z _ _) = z + +tail≡ : {x y : A} {u v : DL} {r : x ≥ᴴ u} {s : y ≥ᴴ v} + → cons x u r ≡ cons y v s → u ≡ v +tail≡ = cong tail + where + tail : DL → DL + tail [] = [] + tail (cons _ u _) = u + +caseDL : ∀ {ℓ} → {X : Type ℓ} → (x y : X) → DL → X +caseDL x y [] = x +caseDL x y (cons _ _ _) = y + +[]≢cons : {x : A} {u : DL} {r : x ≥ᴴ u} → ¬ ([] ≡ cons x u r) +[]≢cons p = subst (caseDL Unit ⊥) p tt + +cons≢[] : {x : A} {u : DL} {r : x ≥ᴴ u} → ¬ (cons x u r ≡ []) +cons≢[] p = subst (caseDL ⊥ Unit) p tt + +module isSetDL + (≥isPropValued : {x y : A} → isProp (x ≥ y)) + (discreteA : Discrete A) + where + + ≥ᴴisPropValued : {x : A} {u : DL} → isProp (x ≥ᴴ u) + ≥ᴴisPropValued ≥ᴴ[] ≥ᴴ[] = refl + ≥ᴴisPropValued (≥ᴴcons x≥y) (≥ᴴcons x≥y') = cong ≥ᴴcons (≥isPropValued x≥y x≥y') + + cons≡ : {x y : A} {u v : DL} {r : x ≥ᴴ u} {s : y ≥ᴴ v} + → x ≡ y → u ≡ v → cons x u r ≡ cons y v s + cons≡ {x} p = subst P p d + where + P : A → Type₀ + P y = {u v : DL} {r : x ≥ᴴ u} {s : y ≥ᴴ v} → u ≡ v → cons x u r ≡ cons y v s + d : P x + d {u} q = subst Q q c + where + Q : (v : DL) → Type₀ + Q v = {r : x ≥ᴴ u} {s : x ≥ᴴ v} → cons x u r ≡ cons x v s + c : Q u + c = cong (cons x u) (≥ᴴisPropValued _ _) + + discreteDL : Discrete DL + discreteDL [] [] = yes refl + discreteDL [] (cons y v s) = no []≢cons + discreteDL (cons x u r) [] = no cons≢[] + discreteDL (cons x u r) (cons y v s) with discreteA x y + discreteDL (cons x u r) (cons y v s) | yes x≡y with discreteDL u v + discreteDL (cons x u r) (cons y v s) | yes x≡y | yes u≡v = yes (cons≡ x≡y u≡v) + discreteDL (cons x u r) (cons y v s) | yes x≡y | no u≢v = no (λ e → u≢v (tail≡ e)) + discreteDL (cons x u r) (cons y v s) | no x≢y = no (λ e → x≢y (head≡ e)) + + isSetDL : isSet DL + isSetDL = Discrete→isSet discreteDL + + +------------------------------------------------------------------------ +-- Descending lists are isomorphic to finite multisets. + +module IsoToFMSet + (discreteA : Discrete A) + (≥dec : (x y : A) → Dec (x ≥ y)) + (≥isPropValued : {x y : A} → isProp (x ≥ y)) + (≥trans : {x y z : A} → x ≥ y → y ≥ z → x ≥ z) + (≰→≥ : {x y : A} → ¬ (x ≥ y) → y ≥ x) + (≤≥→≡ : {x y : A} → x ≥ y → y ≥ x → x ≡ y) + where + + ------------------------------------------------------------------------ + -- The insert function + -- + -- The type DL is defined simultaneously with the relation _≥ᴴ_. + -- Hence the insert function has to be defined by simultaneously + -- proving a property of _≥ᴴ_. + + insert : A → DL → DL + ≥ᴴinsert : {x y : A} {u : DL} + → y ≥ᴴ u → ¬ (x ≥ y) → y ≥ᴴ insert x u + + insert x [] = [ x ] + insert x (cons y u r) with ≥dec x y + insert x (cons y u r) | yes x≥y = cons x (cons y u r) (≥ᴴcons x≥y) + insert x (cons y u r) | no x≱y = cons y (insert x u) (≥ᴴinsert r x≱y) + + ≥ᴴinsert ≥ᴴ[] x≱y = ≥ᴴcons (≰→≥ x≱y) + ≥ᴴinsert {x} {y} {cons z u z≥u} (≥ᴴcons y≥z) x≱y with ≥dec x z + ≥ᴴinsert {x} {y} {cons z u z≥u} (≥ᴴcons y≥z) x≱y | yes x≥z = ≥ᴴcons (≰→≥ x≱y) + ≥ᴴinsert {x} {y} {cons z u z≥u} (≥ᴴcons y≥z) x≱y | no x≱z = ≥ᴴcons y≥z + + open isSetDL ≥isPropValued discreteA + + insert-swap : (x y : A) (u : DL) + → insert x (insert y u) ≡ insert y (insert x u) + insert-swap x y [] with ≥dec x y + insert-swap x y [] | yes x≥y with ≥dec y x + insert-swap x y [] | yes x≥y | yes y≥x = cons≡ (≤≥→≡ x≥y y≥x) (cons≡ (≤≥→≡ y≥x x≥y) refl) + insert-swap x y [] | yes x≥y | no y≱x = cons≡ refl (cons≡ refl refl) + insert-swap x y [] | no x≱y with ≥dec y x + insert-swap x y [] | no x≱y | yes y≥x = cons≡ refl (cons≡ refl refl) + insert-swap x y [] | no x≱y | no y≱x = ⊥.rec (x≱y (≰→≥ y≱x)) + insert-swap x y (cons z u z≥u) with ≥dec y z + insert-swap x y (cons z u z≥u) | yes y≥z with ≥dec x y + insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y with ≥dec x z + insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z with ≥dec y x + insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z | yes y≥x = cons≡ (≤≥→≡ x≥y y≥x) (cons≡ (≤≥→≡ y≥x x≥y) refl) + insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z | no y≱x with ≥dec y z + insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z | no y≱x | yes y≥z' = cons≡ refl (cons≡ refl refl) + insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z | no y≱x | no y≱z = ⊥.rec (y≱z y≥z) + insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | no x≱z = ⊥.rec (x≱z (≥trans x≥y y≥z)) + insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y with ≥dec x z + insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | yes x≥z with ≥dec y x + insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | yes x≥z | yes y≥x = cons≡ refl (cons≡ refl refl) + insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | yes x≥z | no y≱x = ⊥.rec (x≱y (≰→≥ y≱x)) + insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | no x≱z with ≥dec y z + insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | no x≱z | yes y≥z' = cons≡ refl (cons≡ refl refl) + insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | no x≱z | no y≱z = ⊥.rec (y≱z y≥z) + insert-swap x y (cons z u z≥u) | no y≱z with ≥dec x z + insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z with ≥dec y x + insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | yes y≥x = ⊥.rec (y≱z (≥trans y≥x x≥z)) + insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | no y≱x with ≥dec y z + insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | no y≱x | yes y≥z = ⊥.rec (y≱z y≥z) + insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | no y≱x | no y≱z' = cons≡ refl (cons≡ refl refl) + insert-swap x y (cons z u z≥u) | no y≱z | no x≱z with ≥dec y z + insert-swap x y (cons z u z≥u) | no y≱z | no x≱z | yes y≥z = ⊥.rec (y≱z y≥z) + insert-swap x y (cons z u z≥u) | no y≱z | no x≱z | no y≱z' = cons≡ refl (insert-swap x y u) + + -- Insertion sort + toDL : FMSet A → DL + toDL = FMSet.Rec.f isSetDL [] insert insert-swap + {- + toDL [] = [] + toDL (x ∷ u) = insert x (toDL u) + toDL (comm x y u i) = insert-swap x y (toDL u) i + toDL (trunc x y p q i j) = isSetDL (toDL x) (toDL y) (cong toDL p) (cong toDL q) i j + -} + + insert-cons : (x : A) (u : DL) (r : x ≥ᴴ u) + → insert x u ≡ cons x u r + insert-cons x [] _ = cons≡ refl refl + insert-cons x (cons y u _) _ with ≥dec x y + insert-cons x (cons y u _) _ | yes x≥y = cons≡ refl refl + insert-cons x (cons y u _) (≥ᴴcons x≥y) | no x≱y = ⊥.rec (x≱y x≥y) + + toDL∘toFMSet≡id : (u : DL) → toDL (toFMSet u) ≡ u + toDL∘toFMSet≡id [] = refl + toDL∘toFMSet≡id (cons x u r) i = + hcomp (λ j → λ { (i = i0) → insert x (toDL∘toFMSet≡id u (~ j)) + ; (i = i1) → cons x u r }) + (insert-cons x u r i) + + insert-∷ : (x : A) (u : DL) + → toFMSet (insert x u) ≡ x ∷ toFMSet u + insert-∷ x [] = refl + insert-∷ x (cons y u _) with ≥dec x y + insert-∷ x (cons y u _) | yes x≥y = refl + insert-∷ x (cons y u _) | no x≱y = cong (λ z → y ∷ z) (insert-∷ x u) ∙ comm y x (toFMSet u) + + toFMSet∘toDL≡id : (u : FMSet A) → toFMSet (toDL u) ≡ u + toFMSet∘toDL≡id = FMSet.ElimProp.f (trunc _ _) + refl + (λ x {u} p → insert-∷ x (toDL u) ∙ cong (λ z → x ∷ z) p) + + FMSet≡DL : FMSet A ≡ DL + FMSet≡DL = isoToPath (iso toDL toFMSet toDL∘toFMSet≡id toFMSet∘toDL≡id) + + ------------------------------------------------------------------------ + -- Concatenation of sorted lists + -- + -- Defined by transporting the one on finite multisets + + infixr 30 _++ᴰᴸ_ + + _++ᴰᴸ_ : DL → DL → DL + _++ᴰᴸ_ = transport (λ i → FMSet≡DL i → FMSet≡DL i → FMSet≡DL i) _++_ + + []Path : PathP (λ i → FMSet≡DL i) [] [] + []Path i = transp (λ j → FMSet≡DL (i ∧ j)) (~ i) [] + + ++Path : PathP (λ i → FMSet≡DL i → FMSet≡DL i → FMSet≡DL i) _++_ _++ᴰᴸ_ + ++Path i = transp (λ j → FMSet≡DL (i ∧ j) → FMSet≡DL (i ∧ j) → FMSet≡DL (i ∧ j)) (~ i) _++_ + + unitl-++ᴰᴸ : ∀ u → [] ++ᴰᴸ u ≡ u + unitl-++ᴰᴸ = transport (λ i → (u : FMSet≡DL i) → ++Path i ([]Path i) u ≡ u) unitl-++ + + unitr-++ᴰᴸ : ∀ u → u ++ᴰᴸ [] ≡ u + unitr-++ᴰᴸ = transport (λ i → (u : FMSet≡DL i) → ++Path i u ([]Path i) ≡ u) unitr-++ + + assoc-++ᴰᴸ : ∀ u v w → u ++ᴰᴸ (v ++ᴰᴸ w) ≡ (u ++ᴰᴸ v) ++ᴰᴸ w + assoc-++ᴰᴸ = transport (λ i → (u v w : FMSet≡DL i) + → ++Path i u (++Path i v w) ≡ ++Path i (++Path i u v) w) + assoc-++ + + comm-++ᴰᴸ : ∀ u v → u ++ᴰᴸ v ≡ v ++ᴰᴸ u + comm-++ᴰᴸ = transport (λ i → (u v : FMSet≡DL i) → ++Path i u v ≡ ++Path i v u) comm-++ + + + ------------------------------------------------------------------------ + -- Converting multisets to (descending) lists + + FMSet→List : FMSet A → List A + FMSet→List u = toList (toDL u) diff --git a/Cubical/Data/DescendingList/Strict.agda b/Cubical/Data/DescendingList/Strict.agda new file mode 100644 index 0000000000..4141c6d76f --- /dev/null +++ b/Cubical/Data/DescendingList/Strict.agda @@ -0,0 +1,14 @@ +------------------------------------------------------------------------ +-- Strictly descending lists +------------------------------------------------------------------------ + +{-# OPTIONS --safe #-} + +open import Cubical.Core.Everything + +module Cubical.Data.DescendingList.Strict + (A : Type₀) + (_>_ : A → A → Type₀) + where + +open import Cubical.Data.DescendingList.Base A _>_ renaming (_≥ᴴ_ to _>ᴴ_; ≥ᴴ[] to >ᴴ[]; ≥ᴴcons to >ᴴcons; DL to SDL) using ([]; cons) public diff --git a/Cubical/Data/DescendingList/Strict/Properties.agda b/Cubical/Data/DescendingList/Strict/Properties.agda new file mode 100644 index 0000000000..ed23dc0009 --- /dev/null +++ b/Cubical/Data/DescendingList/Strict/Properties.agda @@ -0,0 +1,227 @@ +{-# OPTIONS --safe #-} + +open import Cubical.Foundations.Prelude + +module Cubical.Data.DescendingList.Strict.Properties + (A : Type₀) + (_>_ : A → A → Type₀) + where + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.DescendingList.Strict A _>_ +open import Cubical.HITs.ListedFiniteSet as LFSet renaming (_∈_ to _∈ʰ_) + +import Cubical.Data.Empty as ⊥ +open import Cubical.Relation.Nullary.DecidableEq + +open import Cubical.Relation.Nullary using (Dec; Discrete) renaming (¬_ to Type¬_) + +unsort : SDL → LFSet A +unsort [] = [] +unsort (cons x xs x>xs) = x ∷ unsort xs + +module _ where + open import Cubical.Relation.Nullary + data Tri (A B C : Type) : Type where + tri-< : A → ¬ B → ¬ C → Tri A B C + tri-≡ : ¬ A → B → ¬ C → Tri A B C + tri-> : ¬ A → ¬ B → C → Tri A B C + +module IsoToLFSet + (A-discrete : Discrete A) + (>-isProp : ∀ {x y} → isProp (x > y)) + (tri : ∀ x y → Tri (y > x) (x ≡ y) (x > y)) + (>-trans : ∀ {x y z} → x > y → y > z → x > z) + (>-irreflexive : ∀ {x} → Type¬ x > x) + where + + Tri' : A → A → Type + Tri' x y = Tri (y > x) (x ≡ y) (x > y) + + open import Cubical.HITs.PropositionalTruncation as PropTrunc + open import Cubical.Functions.Logic + + -- Membership is defined via `LFSet`. + -- This computes just as well as a direct inductive definition, + -- and additionally lets us use the extra `comm` and `dup` paths to prove + -- things about membership. + _∈ˡ_ : A → SDL → hProp ℓ-zero + a ∈ˡ l = a ∈ʰ unsort l + + Memˡ : SDL → A → hProp ℓ-zero + Memˡ l a = a ∈ˡ l + + Memʰ : LFSet A → A → hProp ℓ-zero + Memʰ l a = a ∈ʰ l + + >ᴴ-trans : ∀ x y zs → x > y → y >ᴴ zs → x >ᴴ zs + >ᴴ-trans x y [] x>y y>zs = >ᴴ[] + >ᴴ-trans x y (cons z zs _) x>y (>ᴴcons y>z) = >ᴴcons (>-trans x>y y>z) + + ≡ₚ-sym : ∀ {A : Type} {x y : A} → ⟨ x ≡ₚ y ⟩ → ⟨ y ≡ₚ x ⟩ + ≡ₚ-sym p = PropTrunc.rec squash (λ p → ∣ sym p ∣) p + + >-all : ∀ x l → x >ᴴ l → ∀ a → ⟨ a ∈ˡ l ⟩ → x > a + >-all x (cons y zs y>zs) (>ᴴcons x>y) a a∈l = + ⊔-elim (a ≡ₚ y) (a ∈ˡ zs) (λ _ → (x > a) , >-isProp {x} {a}) + (λ a≡ₚy → substₚ (λ q → x > q , >-isProp) (≡ₚ-sym a≡ₚy) x>y) + (λ a∈zs → >-all x zs (>ᴴ-trans x y zs x>y y>zs) a a∈zs) + a∈l + + >-absent : ∀ x l → x >ᴴ l → ⟨ ¬ (x ∈ˡ l) ⟩ + >-absent x l x>l x∈l = ⊥.rec (>-irreflexive (>-all x l x>l x x∈l)) + + >ᴴ-isProp : ∀ x xs → isProp (x >ᴴ xs) + >ᴴ-isProp x _ >ᴴ[] >ᴴ[] = refl + >ᴴ-isProp x _ (>ᴴcons p) (>ᴴcons q) = cong >ᴴcons (>-isProp p q) + + SDL-isSet : isSet SDL + SDL-isSet = isSetDL.isSetDL A _>_ >-isProp A-discrete where + open import Cubical.Data.DescendingList.Properties + + insert : A → SDL → SDL + >ᴴinsert : {x y : A} {u : SDL} → y >ᴴ u → (y > x) → y >ᴴ insert x u + insert x [] = cons x [] >ᴴ[] + insert x (cons y zs good) with tri x y + insert x (cons y zs good) | tri-< xᴴinsert good x _ _ x>y = cons x (cons y zs good) (>ᴴcons x>y) + >ᴴinsert >ᴴ[] y>x = >ᴴcons y>x + >ᴴinsert {x} (>ᴴcons {y} y>ys) y>x with tri x y + >ᴴinsert {x} {y} (>ᴴcons {z} z>zs) y>x | tri-< _ _ e = >ᴴcons z>zs + >ᴴinsert {x} (>ᴴcons {y} y>ys) y>x | tri-≡ _ _ e = >ᴴcons y>ys + >ᴴinsert {x} (>ᴴcons {y} y>ys) y>x | tri-> _ _ _ = >ᴴcons y>x + + insert-correct : ∀ x ys → unsort (insert x ys) ≡ (x ∷ unsort ys) + insert-correct x [] = refl + insert-correct x (cons y zs y>zs) with tri x y + ... | tri-< _ _ _ = + y ∷ unsort (insert x zs) ≡⟨ (λ i → y ∷ (insert-correct x zs i)) ⟩ + y ∷ x ∷ unsort zs ≡⟨ comm _ _ _ ⟩ + x ∷ y ∷ unsort zs ∎ + ... | tri-≡ _ x≡y _ = sym (dup y (unsort zs)) ∙ (λ i → (x≡y (~ i)) ∷ y ∷ unsort zs) + ... | tri-> _ _ _ = refl + + insert-correct₂ : ∀ x y zs → unsort (insert x (insert y zs)) ≡ (x ∷ y ∷ unsort zs) + insert-correct₂ x y zs = + insert-correct x (insert y zs) + ∙ cong (λ q → x ∷ q) (insert-correct y zs) + + abstract + -- for some reason, making [exclude] non-abstract makes + -- typechecking noticeably slower + exclude : A → (A → hProp ℓ-zero) → (A → hProp ℓ-zero) + exclude x h a = ¬ a ≡ₚ x ⊓ h a + + >-excluded : ∀ x xs → x >ᴴ xs → exclude x (Memʰ (x ∷ unsort xs)) ≡ Memˡ xs + >-excluded x xs x>xs = funExt (λ a → ⇔toPath (to a) (from a)) where + + import Cubical.Data.Sigma as D + import Cubical.Data.Sum as D + + from : ∀ a → ⟨ a ∈ˡ xs ⟩ → ⟨ ¬ a ≡ₚ x ⊓ (a ≡ₚ x ⊔ a ∈ˡ xs) ⟩ + from a a∈xs = (PropTrunc.rec (snd ⊥) a≢x) D., inr a∈xs where + a≢x : Type¬ (a ≡ x) + a≢x = λ a≡x → (>-absent x xs x>xs (transport (λ i → ⟨ a≡x i ∈ˡ xs ⟩) a∈xs )) + + to : ∀ a → ⟨ ¬ a ≡ₚ x ⊓ (a ≡ₚ x ⊔ a ∈ˡ xs) ⟩ → ⟨ a ∈ˡ xs ⟩ + to a (a≢x D., x) = PropTrunc.rec (snd (a ∈ˡ xs)) (λ { + (D.inl a≡x) → ⊥.rec (a≢x a≡x); + (D.inr x) → x }) x + + cons-eq : ∀ x xs x>xs y ys y>ys → x ≡ y → xs ≡ ys → cons x xs x>xs ≡ cons y ys y>ys + cons-eq x xs x>xs y ys y>ys x≡y xs≡ys i = cons (x≡y i) (xs≡ys i) + (>ᴴ-isProp (x≡y i) (xs≡ys i) + (transp (λ j → (x≡y (i ∧ j)) >ᴴ (xs≡ys) (i ∧ j)) (~ i) x>xs) + (transp (λ j → (x≡y (i ∨ ~ j)) >ᴴ (xs≡ys) (i ∨ ~ j)) i y>ys) + i) + + Memˡ-inj-cons : ∀ x xs x>xs y ys y>ys → x ≡ y → Memˡ (cons x xs x>xs) ≡ Memˡ (cons y ys y>ys) → Memˡ xs ≡ Memˡ ys + Memˡ-inj-cons x xs x>xs y ys y>ys x≡y e = + Memˡ xs ≡⟨ + sym (>-excluded x xs x>xs) + ⟩ exclude x (Memʰ (x ∷ unsort xs)) ≡⟨ + (λ i → exclude (x≡y i) (e i)) + ⟩ exclude y (Memʰ (y ∷ unsort ys)) ≡⟨ + (>-excluded y ys y>ys) + ⟩ Memˡ ys ∎ + + Memˡ-inj : ∀ l₁ l₂ → Memˡ l₁ ≡ Memˡ l₂ → l₁ ≡ l₂ + Memˡ-inj [] [] eq = refl + Memˡ-inj [] (cons y ys y>ys) eq = ⊥.rec (lower (transport (λ i → ⟨ eq (~ i) y ⟩) (inl ∣ refl ∣))) + Memˡ-inj (cons y ys y>ys) [] eq = ⊥.rec (lower (transport (λ i → ⟨ eq i y ⟩) (inl ∣ refl ∣))) + Memˡ-inj (cons x xs x>xs) (cons y ys y>ys) e = + ⊔-elim (x ≡ₚ y) (x ∈ʰ unsort ys) + (λ _ → ((cons x xs x>xs) ≡ (cons y ys y>ys)) , SDL-isSet _ _) + (PropTrunc.rec (SDL-isSet _ _) with-x≡y) + (⊥.rec ∘ x∉ys) + (transport (λ i → ⟨ e i x ⟩) (inl ∣ refl ∣)) where + + xxs = cons x xs x>xs + + x∉ys : ⟨ ¬ x ∈ˡ ys ⟩ + x∉ys x∈ys = ⊥.rec (>-irreflexive y>y) where + y>x : y > x + y>x = (>-all y ys y>ys x x∈ys) + + y∈xxs : ⟨ y ∈ˡ (cons x xs x>xs) ⟩ + y∈xxs = (transport (λ i → ⟨ e (~ i) y ⟩) (inl ∣ refl ∣)) + + y>y : y > y + y>y = >-all y xxs (>ᴴcons y>x) y y∈xxs + + with-x≡y : x ≡ y → (cons x xs x>xs) ≡ (cons y ys y>ys) + with-x≡y x≡y = cons-eq x xs x>xs y ys y>ys x≡y r where + + r : xs ≡ ys + r = Memˡ-inj _ _ (Memˡ-inj-cons x xs x>xs y ys y>ys x≡y e) + + unsort-inj : ∀ x y → unsort x ≡ unsort y → x ≡ y + unsort-inj x y e = Memˡ-inj x y λ i a → a ∈ʰ (e i) + + insert-swap : (x y : A) (zs : SDL) → insert x (insert y zs) ≡ insert y (insert x zs) + insert-swap x y zs = + unsort-inj (insert x (insert y zs)) (insert y (insert x zs)) + (unsort (insert x (insert y zs)) + ≡⟨ (λ i → insert-correct₂ x y zs i) ⟩ + x ∷ y ∷ unsort zs + ≡⟨ (λ i → comm x y (unsort zs) i) ⟩ + y ∷ x ∷ unsort zs + ≡⟨ (λ i → insert-correct₂ y x zs (~ i)) ⟩ + unsort (insert y (insert x zs)) ∎) + + insert-dup : (x : A) (ys : SDL) → insert x (insert x ys) ≡ insert x ys + insert-dup x ys = unsort-inj (insert x (insert x ys)) (insert x ys) + ( + unsort (insert x (insert x ys)) + ≡⟨ (λ i → insert-correct₂ x x ys i) ⟩ + x ∷ x ∷ unsort ys + ≡⟨ dup x (unsort ys) ⟩ + x ∷ unsort ys + ≡⟨ (λ i → insert-correct x ys (~ i)) ⟩ + unsort (insert x ys) ∎ + ) + + sort : LFSet A → SDL + sort = LFSet.Rec.f [] insert insert-swap insert-dup SDL-isSet + + unsort∘sort : ∀ x → unsort (sort x) ≡ x + unsort∘sort = + LFSet.PropElim.f + refl + (λ x {ys} ys-hyp → insert-correct x (sort ys) ∙ cong (λ q → x ∷ q) ys-hyp) + (λ xs → trunc (unsort (sort xs)) xs) + + sort∘unsort : ∀ x → sort (unsort x) ≡ x + sort∘unsort x = unsort-inj (sort (unsort x)) x (unsort∘sort (unsort x)) + + SDL-LFSet-iso : Iso SDL (LFSet A) + SDL-LFSet-iso = (iso unsort sort unsort∘sort sort∘unsort) + + SDL≡LFSet : SDL ≡ LFSet A + SDL≡LFSet = ua (isoToEquiv SDL-LFSet-iso) diff --git a/Cubical/Data/Empty.agda b/Cubical/Data/Empty.agda new file mode 100644 index 0000000000..d3e44609ce --- /dev/null +++ b/Cubical/Data/Empty.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Empty where + +open import Cubical.Data.Empty.Base public + +open import Cubical.Data.Empty.Properties public diff --git a/Cubical/Data/Empty/Base.agda b/Cubical/Data/Empty/Base.agda new file mode 100644 index 0000000000..c6384d3b09 --- /dev/null +++ b/Cubical/Data/Empty/Base.agda @@ -0,0 +1,20 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Empty.Base where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude + +private + variable + ℓ : Level + +data ⊥ : Type₀ where + +⊥* : Type ℓ +⊥* = Lift ⊥ + +rec : {A : Type ℓ} → ⊥ → A +rec () + +elim : {A : ⊥ → Type ℓ} → (x : ⊥) → A x +elim () diff --git a/Cubical/Data/Empty/Properties.agda b/Cubical/Data/Empty/Properties.agda new file mode 100644 index 0000000000..f0cc830651 --- /dev/null +++ b/Cubical/Data/Empty/Properties.agda @@ -0,0 +1,34 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Empty.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Empty.Base + +isProp⊥ : isProp ⊥ +isProp⊥ () + +isProp⊥* : ∀ {ℓ} → isProp {ℓ} ⊥* +isProp⊥* _ () + +isContr⊥→A : ∀ {ℓ} {A : Type ℓ} → isContr (⊥ → A) +fst isContr⊥→A () +snd isContr⊥→A f i () + +isContrΠ⊥ : ∀ {ℓ} {A : ⊥ → Type ℓ} → isContr ((x : ⊥) → A x) +fst isContrΠ⊥ () +snd isContrΠ⊥ f i () + +uninhabEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (A → ⊥) → (B → ⊥) → A ≃ B +uninhabEquiv ¬a ¬b = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun a = rec (¬a a) + isom .inv b = rec (¬b b) + isom .rightInv b = rec (¬b b) + isom .leftInv a = rec (¬a a) diff --git a/Cubical/Data/Equality.agda b/Cubical/Data/Equality.agda new file mode 100644 index 0000000000..044216caac --- /dev/null +++ b/Cubical/Data/Equality.agda @@ -0,0 +1,50 @@ +{- + +This module converts between the path equality +and the inductively define equality types. + +- _≡c_ stands for "c"ubical equality. +- _≡p_ stands for "p"ropositional equality. + +TODO: reconsider naming scheme. +-} +{-# OPTIONS --safe #-} + +module Cubical.Data.Equality where + +open import Cubical.Foundations.Prelude + renaming ( _≡_ to _≡c_ ; refl to reflc ) + public +open import Agda.Builtin.Equality + renaming ( _≡_ to _≡p_ ; refl to reflp ) + public + +open import Cubical.Foundations.Isomorphism + +private + variable + ℓ : Level + A : Type ℓ + x y z : A + +ptoc : x ≡p y → x ≡c y +ptoc reflp = reflc + +ctop : x ≡c y → x ≡p y +ctop p = transport (cong (λ u → _ ≡p u) p) reflp + +ptoc-ctop : (p : x ≡c y) → ptoc (ctop p) ≡c p +ptoc-ctop p = + J (λ _ h → ptoc (ctop h) ≡c h) (cong ptoc (transportRefl reflp)) p + +ctop-ptoc : (p : x ≡p y) → ctop (ptoc p) ≡c p +ctop-ptoc {x = x} reflp = transportRefl reflp + +p≅c : {x y : A} → Iso (x ≡c y) (x ≡p y) +p≅c = iso ctop ptoc ctop-ptoc ptoc-ctop + +p-c : {x y : A} → (x ≡c y) ≡c (x ≡p y) +p-c = isoToPath p≅c + +p=c : {x y : A} → (x ≡c y) ≡p (x ≡p y) +p=c = ctop p-c diff --git a/Cubical/Data/Fin.agda b/Cubical/Data/Fin.agda new file mode 100644 index 0000000000..458cb01e01 --- /dev/null +++ b/Cubical/Data/Fin.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --safe #-} + +module Cubical.Data.Fin where + +open import Cubical.Data.Fin.Base public +open import Cubical.Data.Fin.Properties public +open import Cubical.Data.Fin.Literals public diff --git a/Cubical/Data/Fin/Arithmetic.agda b/Cubical/Data/Fin/Arithmetic.agda new file mode 100644 index 0000000000..318b6174a4 --- /dev/null +++ b/Cubical/Data/Fin/Arithmetic.agda @@ -0,0 +1,87 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.Data.Fin.Arithmetic where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Mod +open import Cubical.Data.Nat.Order +open import Cubical.Data.Fin +open import Cubical.Data.Sigma + +infixl 6 _+ₘ_ _-ₘ_ _·ₘ_ +infix 7 -ₘ_ + +-- Addition, subtraction and multiplication +_+ₘ_ : {n : ℕ} → Fin (suc n) → Fin (suc n) → Fin (suc n) +fst (_+ₘ_ {n = n} x y) = ((fst x) + (fst y)) mod (suc n) +snd (_+ₘ_ {n = n} x y) = mod< _ ((fst x) + (fst y)) + +-ₘ_ : {n : ℕ} → (x : Fin (suc n)) → Fin (suc n) +fst (-ₘ_ {n = n} x) = + (+induction n _ (λ x _ → ((suc n) ∸ x) mod (suc n)) λ _ x → x) (fst x) +snd (-ₘ_ {n = n} x) = lem (fst x) + where + ≡<-trans : {x y z : ℕ} → x < y → x ≡ z → z < y + ≡<-trans (k , p) q = k , cong (λ x → k + suc x) (sym q) ∙ p + + lem : {n : ℕ} (x : ℕ) + → (+induction n _ _ _) x < suc n + lem {n = n} = + +induction n _ + (λ x p → ≡<-trans (mod< n (suc n ∸ x)) + (sym (+inductionBase n _ _ _ x p))) + λ x p → ≡<-trans p + (sym (+inductionStep n _ _ _ x)) + +_-ₘ_ : {n : ℕ} → (x y : Fin (suc n)) → Fin (suc n) +_-ₘ_ x y = x +ₘ (-ₘ y) + +_·ₘ_ : {n : ℕ} → (x y : Fin (suc n)) → Fin (suc n) +fst (_·ₘ_ {n = n} x y) = (fst x · fst y) mod (suc n) +snd (_·ₘ_ {n = n} x y) = mod< n (fst x · fst y) + +-- Group laws ++ₘ-assoc : {n : ℕ} (x y z : Fin (suc n)) + → (x +ₘ y) +ₘ z ≡ (x +ₘ (y +ₘ z)) ++ₘ-assoc {n = n} x y z = + Σ≡Prop (λ _ → m≤n-isProp) + ((mod-rCancel (suc n) ((fst x + fst y) mod (suc n)) (fst z)) + ∙∙ sym (mod+mod≡mod (suc n) (fst x + fst y) (fst z)) + ∙∙ cong (_mod suc n) (sym (+-assoc (fst x) (fst y) (fst z))) + ∙∙ mod+mod≡mod (suc n) (fst x) (fst y + fst z) + ∙∙ sym (mod-lCancel (suc n) (fst x) ((fst y + fst z) mod suc n))) + ++ₘ-comm : {n : ℕ} (x y : Fin (suc n)) → (x +ₘ y) ≡ (y +ₘ x) ++ₘ-comm {n = n} x y = + Σ≡Prop (λ _ → m≤n-isProp) + (cong (_mod suc n) (+-comm (fst x) (fst y))) + ++ₘ-lUnit : {n : ℕ} (x : Fin (suc n)) → 0 +ₘ x ≡ x ++ₘ-lUnit {n = n} (x , p) = + Σ≡Prop (λ _ → m≤n-isProp) + (+inductionBase n _ _ _ x p) + ++ₘ-rUnit : {n : ℕ} (x : Fin (suc n)) → x +ₘ 0 ≡ x ++ₘ-rUnit x = +ₘ-comm x 0 ∙ (+ₘ-lUnit x) + ++ₘ-rCancel : {n : ℕ} (x : Fin (suc n)) → x -ₘ x ≡ 0 ++ₘ-rCancel {n = n} x = + Σ≡Prop (λ _ → m≤n-isProp) + (cong (λ z → (fst x + z) mod (suc n)) + (+inductionBase n _ _ _ (fst x) (snd x)) + ∙∙ sym (mod-rCancel (suc n) (fst x) ((suc n) ∸ (fst x))) + ∙∙ cong (_mod (suc n)) (+-comm (fst x) ((suc n) ∸ (fst x))) + ∙∙ cong (_mod (suc n)) (≤-∸-+-cancel (<-weaken (snd x))) + ∙∙ zero-charac (suc n)) + ++ₘ-lCancel : {n : ℕ} (x : Fin (suc n)) → (-ₘ x) +ₘ x ≡ 0 ++ₘ-lCancel {n = n} x = +ₘ-comm (-ₘ x) x ∙ +ₘ-rCancel x + +-- TODO : Ring laws + +private + test₁ : Path (Fin 11) (5 +ₘ 10) 4 + test₁ = refl + + test₂ : Path (Fin 11) (-ₘ 7 +ₘ 5 +ₘ 10) 8 + test₂ = refl diff --git a/Cubical/Data/Fin/Base.agda b/Cubical/Data/Fin/Base.agda new file mode 100644 index 0000000000..c3ad6c533d --- /dev/null +++ b/Cubical/Data/Fin/Base.agda @@ -0,0 +1,97 @@ +{-# OPTIONS --safe #-} + +module Cubical.Data.Fin.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat using (ℕ; zero; suc) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Nat.Order.Recursive using () renaming (_≤_ to _≤′_) +open import Cubical.Data.Sigma +open import Cubical.Data.Sum using (_⊎_; _⊎?_; inl; inr) + +open import Cubical.Relation.Nullary + +-- Finite types. +-- +-- Currently it is most convenient to define these as a subtype of the +-- natural numbers, because indexed inductive definitions don't behave +-- well with cubical Agda. This definition also has some more general +-- attractive properties, of course, such as easy conversion back to +-- ℕ. +Fin : ℕ → Type₀ +Fin n = Σ[ k ∈ ℕ ] k < n + +private + variable + ℓ : Level + k : ℕ + +fzero : Fin (suc k) +fzero = (0 , suc-≤-suc zero-≤) + +-- It is easy, using this representation, to take the successor of a +-- number as a number in the next largest finite type. +fsuc : Fin k → Fin (suc k) +fsuc (k , l) = (suc k , suc-≤-suc l) + +-- Conversion back to ℕ is trivial... +toℕ : Fin k → ℕ +toℕ = fst + +-- ... and injective. +toℕ-injective : ∀{fj fk : Fin k} → toℕ fj ≡ toℕ fk → fj ≡ fk +toℕ-injective {fj = fj} {fk} = Σ≡Prop (λ _ → m≤n-isProp) + +-- Conversion from ℕ with a recursive definition of ≤ + +fromℕ≤ : (m n : ℕ) → m ≤′ n → Fin (suc n) +fromℕ≤ zero _ _ = fzero +fromℕ≤ (suc m) (suc n) m≤n = fsuc (fromℕ≤ m n m≤n) + +-- A case analysis helper for induction. +fsplit + : ∀(fj : Fin (suc k)) + → (fzero ≡ fj) ⊎ (Σ[ fk ∈ Fin k ] fsuc fk ≡ fj) +fsplit (0 , km = Empty.rec (Fin-inj′ n>m p) + +≤-·sk-cancel : ∀ {m} {k} {n} → m · suc k ≤ n · suc k → m ≤ n +≤-·sk-cancel {m} {k} {n} (d , p) = o , inj-·sm {m = k} goal where + r = d % suc k + o = d / suc k + resn·k : Residue (suc k) (n · suc k) + resn·k = ((r , n%sk_ : ∀ {x y : Node G} → Edge G x y → Edge G' (_$_ x) (_$_ y) + +open GraphHom public + +GraphGr : ∀ ℓv ℓe → Graph _ _ +Node (GraphGr ℓv ℓe) = Graph ℓv ℓe +Edge (GraphGr ℓv ℓe) G G' = GraphHom G G' + +-- Diagrams are (graph) functors with codomain Type + +Diag : ∀ ℓd (G : Graph ℓv ℓe) → Type (ℓ-suc (ℓ-max (ℓ-max ℓv ℓe) (ℓ-suc ℓd))) +Diag ℓd G = GraphHom G (TypeGr ℓd) + +record DiagMor {G : Graph ℓv ℓe} (F : Diag ℓd G) (F' : Diag ℓd' G) + : Type (ℓ-suc (ℓ-max (ℓ-max ℓv ℓe) (ℓ-suc (ℓ-max ℓd ℓd')))) where + field + nat : ∀ (x : Node G) → F $ x → F' $ x + comSq : ∀ {x y : Node G} (f : Edge G x y) → nat y ∘ F <$> f ≡ F' <$> f ∘ nat x + +open DiagMor public + +DiagGr : ∀ ℓd (G : Graph ℓv ℓe) → Graph _ _ +Node (DiagGr ℓd G) = Diag ℓd G +Edge (DiagGr ℓd G) = DiagMor diff --git a/Cubical/Data/Graph/Examples.agda b/Cubical/Data/Graph/Examples.agda new file mode 100644 index 0000000000..b94a303b9b --- /dev/null +++ b/Cubical/Data/Graph/Examples.agda @@ -0,0 +1,161 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Graph.Examples where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Empty +open import Cubical.Data.Unit renaming (Unit to ⊤) +open import Cubical.Data.Nat +open import Cubical.Data.SumFin +open import Cubical.Relation.Nullary + +open import Cubical.Data.Sum +open import Cubical.Data.Sigma + +open import Cubical.Data.Graph.Base + + +-- Some small graphs of common shape + +⇒⇐ : Graph ℓ-zero ℓ-zero +Node ⇒⇐ = Fin 3 +Edge ⇒⇐ fzero (fsuc fzero) = ⊤ +Edge ⇒⇐ (fsuc (fsuc fzero)) (fsuc fzero) = ⊤ +Edge ⇒⇐ _ _ = ⊥ + +⇐⇒ : Graph ℓ-zero ℓ-zero +Node ⇐⇒ = Fin 3 +Edge ⇐⇒ (fsuc fzero) fzero = ⊤ +Edge ⇐⇒ (fsuc fzero) (fsuc (fsuc fzero)) = ⊤ +Edge ⇐⇒ _ _ = ⊥ + +-- paralell pair graph +⇉ : Graph ℓ-zero ℓ-zero +Node ⇉ = Fin 2 +Edge ⇉ fzero (fsuc fzero) = Fin 2 +Edge ⇉ _ _ = ⊥ + + +-- The graph ω = 0 → 1 → 2 → ··· + +data Adj : ℕ → ℕ → Type₀ where + adj : ∀ n → Adj n (suc n) + +areAdj : ∀ m n → Dec (Adj m n) +areAdj zero zero = no λ () +areAdj zero (suc zero) = yes (adj zero) +areAdj zero (suc (suc n)) = no λ () +areAdj (suc m) zero = no λ () +areAdj (suc m) (suc n) = mapDec (λ { (adj .m) → adj (suc m) }) + (λ { ¬a (adj .(suc m)) → ¬a (adj m) }) + (areAdj m n) + +ωGr : Graph ℓ-zero ℓ-zero +Node ωGr = ℕ +Edge ωGr m n with areAdj m n +... | yes _ = ⊤ -- if n ≡ (suc m) +... | no _ = ⊥ -- otherwise + +record ωDiag ℓ : Type (ℓ-suc ℓ) where + field + ωNode : ℕ → Type ℓ + ωEdge : ∀ n → ωNode n → ωNode (suc n) + + asDiag : Diag ℓ ωGr + asDiag $ n = ωNode n + _<$>_ asDiag {m} {n} f with areAdj m n + asDiag <$> tt | yes (adj m) = ωEdge m + + +-- The finite connected subgraphs of ω: 𝟘,𝟙,𝟚,𝟛,... + +data AdjFin : ∀ {k} → Fin k → Fin k → Type₀ where + adj : ∀ {k} (n : Fin k) → AdjFin (finj n) (fsuc n) + +adj-fsuc : ∀ {k} {m n : Fin k} → AdjFin (fsuc m) (fsuc n) → AdjFin m n +adj-fsuc {suc k} {.(finj n)} {fsuc n} (adj .(fsuc n)) = adj n + +areAdjFin : ∀ {k} (m n : Fin k) → Dec (AdjFin m n) +areAdjFin {suc k} fzero fzero = no λ () +areAdjFin {suc (suc k)} fzero (fsuc fzero) = yes (adj fzero) +areAdjFin {suc (suc k)} fzero (fsuc (fsuc n)) = no λ () +areAdjFin {suc k} (fsuc m) fzero = no λ () +areAdjFin {suc k} (fsuc m) (fsuc n) = mapDec (λ { (adj m) → adj (fsuc m) }) + (λ { ¬a a → ¬a (adj-fsuc a) }) + (areAdjFin {k} m n) + +[_]Gr : ℕ → Graph ℓ-zero ℓ-zero +Node [ k ]Gr = Fin k +Edge [ k ]Gr m n with areAdjFin m n +... | yes _ = ⊤ -- if n ≡ (suc m) +... | no _ = ⊥ -- otherwise + +𝟘Gr 𝟙Gr 𝟚Gr 𝟛Gr : Graph ℓ-zero ℓ-zero +𝟘Gr = [ 0 ]Gr; 𝟙Gr = [ 1 ]Gr; 𝟚Gr = [ 2 ]Gr; 𝟛Gr = [ 3 ]Gr + +record [_]Diag ℓ (k : ℕ) : Type (ℓ-suc ℓ) where + field + []Node : Fin (suc k) → Type ℓ + []Edge : ∀ (n : Fin k) → []Node (finj n) → []Node (fsuc n) + + asDiag : Diag ℓ [ suc k ]Gr + asDiag $ n = []Node n + _<$>_ asDiag {m} {n} f with areAdjFin m n + _<$>_ asDiag {.(finj n)} {fsuc n} f | yes (adj .n) = []Edge n + + +-- Disjoint union of graphs + +module _ {ℓv ℓe ℓv' ℓe'} where + + _⊎Gr_ : ∀ (G : Graph ℓv ℓe) (G' : Graph ℓv' ℓe') → Graph (ℓ-max ℓv ℓv') (ℓ-max ℓe ℓe') + Node (G ⊎Gr G') = Node G ⊎ Node G' + Edge (G ⊎Gr G') (inl x) (inl y) = Lift {j = ℓe'} (Edge G x y) + Edge (G ⊎Gr G') (inr x) (inr y) = Lift {j = ℓe } (Edge G' x y) + Edge (G ⊎Gr G') _ _ = Lift ⊥ + + record ⊎Diag ℓ (G : Graph ℓv ℓe) (G' : Graph ℓv' ℓe') + : Type (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-max ℓv ℓv') (ℓ-max ℓe ℓe'))) where + field + ⊎Node : Node G ⊎ Node G' → Type ℓ + ⊎Edgel : ∀ {x y} → Edge G x y → ⊎Node (inl x) → ⊎Node (inl y) + ⊎Edger : ∀ {x y} → Edge G' x y → ⊎Node (inr x) → ⊎Node (inr y) + + asDiag : Diag ℓ (G ⊎Gr G') + asDiag $ x = ⊎Node x + _<$>_ asDiag {inl x} {inl y} f = ⊎Edgel (lower f) + _<$>_ asDiag {inr x} {inr y} f = ⊎Edger (lower f) + + +-- Cartesian product of graphs + +module _ {ℓv ℓe ℓv' ℓe'} where + + -- We need decidable equality in order to define the cartesian product + DecGraph : ∀ ℓv ℓe → Type (ℓ-suc (ℓ-max ℓv ℓe)) + DecGraph ℓv ℓe = Σ[ G ∈ Graph ℓv ℓe ] Discrete (Node G) + + _×Gr_ : (G : DecGraph ℓv ℓe) (G' : DecGraph ℓv' ℓe') → Graph (ℓ-max ℓv ℓv') (ℓ-max ℓe ℓe') + Node (G ×Gr G') = Node (fst G) × Node (fst G') + Edge (G ×Gr G') (x , x') (y , y') with snd G x y | snd G' x' y' + ... | yes _ | yes _ = Edge (fst G) x y ⊎ Edge (fst G') x' y' + ... | yes _ | no _ = Lift {j = ℓe } (Edge (fst G') x' y') + ... | no _ | yes _ = Lift {j = ℓe'} (Edge (fst G) x y) + ... | no _ | no _ = Lift ⊥ + + record ×Diag ℓ (G : DecGraph ℓv ℓe) (G' : DecGraph ℓv' ℓe') + : Type (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-max ℓv ℓv') (ℓ-max ℓe ℓe'))) where + field + ×Node : Node (fst G) × Node (fst G') → Type ℓ + ×Edge₁ : ∀ {x y} (f : Edge (fst G) x y) (x' : Node (fst G')) → ×Node (x , x') → ×Node (y , x') + ×Edge₂ : ∀ (x : Node (fst G)) {x' y'} (f : Edge (fst G') x' y') → ×Node (x , x') → ×Node (x , y') + + asDiag : Diag ℓ (G ×Gr G') + asDiag $ x = ×Node x + _<$>_ asDiag {x , x'} {y , y'} f with snd G x y | snd G' x' y' + _<$>_ asDiag {x , x'} {y , y'} (inl f) | yes _ | yes p' = subst _ p' (×Edge₁ f x') + _<$>_ asDiag {x , x'} {y , y'} (inr f) | yes p | yes _ = subst _ p (×Edge₂ x f ) + _<$>_ asDiag {x , x'} {y , y'} f | yes p | no _ = subst _ p (×Edge₂ x (lower f) ) + _<$>_ asDiag {x , x'} {y , y'} f | no _ | yes p' = subst _ p' (×Edge₁ (lower f) x') diff --git a/Cubical/Data/HomotopyGroup.agda b/Cubical/Data/HomotopyGroup.agda new file mode 100644 index 0000000000..04ab4f13b2 --- /dev/null +++ b/Cubical/Data/HomotopyGroup.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.HomotopyGroup where + +open import Cubical.Data.HomotopyGroup.Base public diff --git a/Cubical/Data/HomotopyGroup/Base.agda b/Cubical/Data/HomotopyGroup/Base.agda new file mode 100644 index 0000000000..997b31dfe4 --- /dev/null +++ b/Cubical/Data/HomotopyGroup/Base.agda @@ -0,0 +1,52 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.HomotopyGroup.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +import Cubical.Foundations.GroupoidLaws as GL +open import Cubical.Foundations.Pointed + +open import Cubical.Data.Nat +open import Cubical.Algebra.Group + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.SetTruncation as SetTrunc + +π^_ : ∀ {ℓ} → ℕ → Pointed ℓ → Group ℓ +π^_ {ℓ} n p = makeGroup e _⨀_ _⁻¹ isSetSetTrunc assoc rUnit lUnit rCancel lCancel + where + n' : ℕ + n' = suc n + + A : Type ℓ + A = typ ((Ω^ n') p) + + e : ∥ A ∥₂ + e = ∣ pt ((Ω^ n') p) ∣₂ + + _⁻¹ : ∥ A ∥₂ → ∥ A ∥₂ + _⁻¹ = SetTrunc.elim {B = λ _ → ∥ A ∥₂} (λ x → squash₂) λ a → ∣ sym a ∣₂ + + _⨀_ : ∥ A ∥₂ → ∥ A ∥₂ → ∥ A ∥₂ + _⨀_ = SetTrunc.elim2 (λ _ _ → squash₂) λ a₀ a₁ → ∣ a₀ ∙ a₁ ∣₂ + + lUnit : (a : ∥ A ∥₂) → (e ⨀ a) ≡ a + lUnit = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + (λ a → cong ∣_∣₂ (sym (GL.lUnit a) )) + + rUnit : (a : ∥ A ∥₂) → a ⨀ e ≡ a + rUnit = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + (λ a → cong ∣_∣₂ (sym (GL.rUnit a) )) + + assoc : (a b c : ∥ A ∥₂) → a ⨀ (b ⨀ c) ≡ (a ⨀ b) ⨀ c + assoc = SetTrunc.elim3 (λ _ _ _ → isProp→isSet (squash₂ _ _)) + (λ a b c → cong ∣_∣₂ (GL.assoc _ _ _)) + + lCancel : (a : ∥ A ∥₂) → ((a ⁻¹) ⨀ a) ≡ e + lCancel = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + λ a → cong ∣_∣₂ (GL.lCancel _) + + rCancel : (a : ∥ A ∥₂) → (a ⨀ (a ⁻¹)) ≡ e + rCancel = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + λ a → cong ∣_∣₂ (GL.rCancel _) diff --git a/Cubical/Data/InfNat.agda b/Cubical/Data/InfNat.agda new file mode 100644 index 0000000000..26cf4093d9 --- /dev/null +++ b/Cubical/Data/InfNat.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Data.InfNat where + +open import Cubical.Data.InfNat.Base public +open import Cubical.Data.InfNat.Properties public diff --git a/Cubical/Data/InfNat/Base.agda b/Cubical/Data/InfNat/Base.agda new file mode 100644 index 0000000000..d368954645 --- /dev/null +++ b/Cubical/Data/InfNat/Base.agda @@ -0,0 +1,36 @@ +{-# OPTIONS --no-exact-split --safe #-} + +module Cubical.Data.InfNat.Base where + +open import Cubical.Data.Nat as ℕ using (ℕ) +open import Cubical.Core.Primitives + +data ℕ+∞ : Type₀ where + ∞ : ℕ+∞ + fin : ℕ → ℕ+∞ + +suc : ℕ+∞ → ℕ+∞ +suc ∞ = ∞ +suc (fin n) = fin (ℕ.suc n) + +zero : ℕ+∞ +zero = fin ℕ.zero + +caseInfNat : ∀ {ℓ} → {A : Type ℓ} → (aF aI : A) → ℕ+∞ → A +caseInfNat aF aI (fin n) = aF +caseInfNat aF aI ∞ = aI + +infixl 6 _+_ +_+_ : ℕ+∞ → ℕ+∞ → ℕ+∞ +∞ + m = ∞ +fin n + ∞ = ∞ +fin n + fin m = fin (n ℕ.+ m) + +infixl 7 _·_ +_·_ : ℕ+∞ → ℕ+∞ → ℕ+∞ +fin m · fin n = fin (m ℕ.· n) +∞ · fin ℕ.zero = zero +fin ℕ.zero · ∞ = zero +∞ · ∞ = ∞ +∞ · fin (ℕ.suc _) = ∞ +fin (ℕ.suc _) · ∞ = ∞ diff --git a/Cubical/Data/InfNat/Properties.agda b/Cubical/Data/InfNat/Properties.agda new file mode 100644 index 0000000000..653de73ffb --- /dev/null +++ b/Cubical/Data/InfNat/Properties.agda @@ -0,0 +1,26 @@ +{-# OPTIONS --no-exact-split --safe #-} + +module Cubical.Data.InfNat.Properties where + +open import Cubical.Data.Nat as ℕ using (ℕ) +open import Cubical.Data.InfNat.Base +open import Cubical.Core.Primitives +open import Cubical.Foundations.Prelude +open import Cubical.Relation.Nullary +open import Cubical.Data.Unit +open import Cubical.Data.Empty + +fromInf-def : ℕ → ℕ+∞ → ℕ +fromInf-def n ∞ = n +fromInf-def _ (fin n) = n + +fin-inj : (n m : ℕ) → fin n ≡ fin m → n ≡ m +fin-inj x _ eq = cong (fromInf-def x) eq + +discreteInfNat : Discrete ℕ+∞ +discreteInfNat ∞ ∞ = yes (λ i → ∞) +discreteInfNat ∞ (fin _) = no λ p → subst (caseInfNat ⊥ Unit) p tt +discreteInfNat (fin _) ∞ = no λ p → subst (caseInfNat Unit ⊥) p tt +discreteInfNat (fin n) (fin m) with ℕ.discreteℕ n m +discreteInfNat (fin n) (fin m) | yes p = yes (cong fin p) +discreteInfNat (fin n) (fin m) | no ¬p = no (λ p → ¬p (fin-inj n m p)) diff --git a/Cubical/Data/Int.agda b/Cubical/Data/Int.agda new file mode 100644 index 0000000000..57147d5796 --- /dev/null +++ b/Cubical/Data/Int.agda @@ -0,0 +1,7 @@ +-- This is the preferred version of the integers in the library. Other +-- versions can be found in the MoreInts directory. +{-# OPTIONS --safe #-} +module Cubical.Data.Int where + +open import Cubical.Data.Int.Base public +open import Cubical.Data.Int.Properties public diff --git a/Cubical/Data/Int/Base.agda b/Cubical/Data/Int/Base.agda new file mode 100644 index 0000000000..b7d9c48b52 --- /dev/null +++ b/Cubical/Data/Int/Base.agda @@ -0,0 +1,85 @@ +-- This is the preferred version of the integers in the library. Other +-- versions can be found in the MoreInts directory. +{-# OPTIONS --safe #-} +module Cubical.Data.Int.Base where + +open import Cubical.Core.Everything +open import Cubical.Data.Bool +open import Cubical.Data.Nat hiding (_+_ ; _·_) renaming (isEven to isEvenℕ ; isOdd to isOddℕ) + +infix 8 -_ +infixl 7 _·_ +infixl 6 _+_ _-_ + +data ℤ : Type₀ where + pos : (n : ℕ) → ℤ + negsuc : (n : ℕ) → ℤ + +neg : (n : ℕ) → ℤ +neg zero = pos zero +neg (suc n) = negsuc n + +sucℤ : ℤ → ℤ +sucℤ (pos n) = pos (suc n) +sucℤ (negsuc zero) = pos zero +sucℤ (negsuc (suc n)) = negsuc n + +predℤ : ℤ → ℤ +predℤ (pos zero) = negsuc zero +predℤ (pos (suc n)) = pos n +predℤ (negsuc n) = negsuc (suc n) + +isEven : ℤ → Bool +isEven (pos n) = isEvenℕ n +isEven (negsuc n) = isOddℕ n + +isOdd : ℤ → Bool +isOdd (pos n) = isOddℕ n +isOdd (negsuc n) = isEvenℕ n + +abs : ℤ → ℕ +abs (pos n) = n +abs (negsuc n) = suc n + +_ℕ-_ : ℕ → ℕ → ℤ +a ℕ- 0 = pos a +0 ℕ- suc b = negsuc b +suc a ℕ- suc b = a ℕ- b + +_+pos_ : ℤ → ℕ → ℤ +z +pos 0 = z +z +pos (suc n) = sucℤ (z +pos n) + +_+negsuc_ : ℤ → ℕ → ℤ +z +negsuc 0 = predℤ z +z +negsuc (suc n) = predℤ (z +negsuc n) + +_+_ : ℤ → ℤ → ℤ +m + pos n = m +pos n +m + negsuc n = m +negsuc n + +-_ : ℤ → ℤ +- pos zero = pos zero +- pos (suc n) = negsuc n +- negsuc n = pos (suc n) + +_-_ : ℤ → ℤ → ℤ +m - n = m + (- n) + +_·_ : ℤ → ℤ → ℤ +pos zero · m = pos zero +pos (suc n) · m = m + pos n · m +negsuc zero · m = - m +negsuc (suc n) · m = - m + negsuc n · m + +-- Natural number and negative integer literals for ℤ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℤ : HasFromNat ℤ + fromNatℤ = record { Constraint = λ _ → Unit ; fromNat = λ n → pos n } + +instance + fromNegℤ : HasFromNeg ℤ + fromNegℤ = record { Constraint = λ _ → Unit ; fromNeg = λ n → neg n } diff --git a/Cubical/Data/Int/MoreInts/BiInvInt.agda b/Cubical/Data/Int/MoreInts/BiInvInt.agda new file mode 100644 index 0000000000..60adfd54dc --- /dev/null +++ b/Cubical/Data/Int/MoreInts/BiInvInt.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.BiInvInt where + +open import Cubical.Data.Int.MoreInts.BiInvInt.Base public + +open import Cubical.Data.Int.MoreInts.BiInvInt.Properties public diff --git a/Cubical/Data/Int/MoreInts/BiInvInt/Base.agda b/Cubical/Data/Int/MoreInts/BiInvInt/Base.agda new file mode 100644 index 0000000000..83af6718b4 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/BiInvInt/Base.agda @@ -0,0 +1,310 @@ +{- + +Definition of the integers as a HIT ported from the redtt library: + + https://github.com/RedPRL/redtt/blob/master/library/cool/biinv-int.red + +For the naive, but incorrect, way to define the integers as a HIT, see HITs.IsoInt + +This file contains: +- definition of BiInvInt +- proof that Int ≡ BiInvInt +- [discreteBiInvInt] and [isSetBiInvInt] +- versions of the point constructors of BiInvInt which satisfy the path constructors judgmentally + +-} +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.BiInvInt.Base where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Data.Nat +open import Cubical.Data.Int + +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Equiv.BiInvertible +open import Cubical.Foundations.Equiv.HalfAdjoint + +open import Cubical.Relation.Nullary + + +data BiInvℤ : Type₀ where + zero : BiInvℤ + suc : BiInvℤ -> BiInvℤ + + -- suc is a bi-invertible equivalence: + predr : BiInvℤ -> BiInvℤ + suc-predr : ∀ z -> suc (predr z) ≡ z + predl : BiInvℤ -> BiInvℤ + predl-suc : ∀ z -> predl (suc z) ≡ z + + +suc-biinvequiv : BiInvEquiv BiInvℤ BiInvℤ +suc-biinvequiv = record { fun = suc ; invr = predr ; invr-rightInv = suc-predr + ; invl = predl ; invl-leftInv = predl-suc } + +predr-suc : ∀ z -> predr (suc z) ≡ z +predr-suc = BiInvEquiv.invr-leftInv suc-biinvequiv + +-- since we want to use suc-adj and pred-adj (defined below) later on, we will need the +-- definition of suc-predl taken from HAEquiv instead of from BiInvEquiv + +suc-haequiv : HAEquiv BiInvℤ BiInvℤ +suc-haequiv = biInvEquiv→HAEquiv suc-biinvequiv + +suc-predl : ∀ z -> suc (predl z) ≡ z +suc-predl = isHAEquiv.rinv (snd suc-haequiv) + +-- predr and predl (as well as suc-predr and suc-predl - up to a path) are indistinguishable, +-- so we can safely define 'pred' to just be predl + +predl≡predr : ∀ z -> predl z ≡ predr z +predl≡predr z i = cong fst (isContr→isProp (isContr-hasSection (biInvEquiv→Equiv-left suc-biinvequiv)) + (predl , suc-predl) (predr , suc-predr)) i z + +suc-predl≡predr : ∀ z -> PathP (λ j → suc (predl≡predr z j) ≡ z) (suc-predl z) (suc-predr z) +suc-predl≡predr z i = cong snd (isContr→isProp (isContr-hasSection (biInvEquiv→Equiv-left suc-biinvequiv)) + (predl , suc-predl) (predr , suc-predr)) i z + +pred : BiInvℤ -> BiInvℤ +pred = predl + +suc-pred = suc-predl +pred-suc = predl-suc + +-- these paths from HAEquiv will be useful later +suc-adj : ∀ z → (λ i → suc (pred-suc z i)) ≡ suc-pred (suc z) +pred-adj : ∀ z → (λ i → pred (suc-pred z i)) ≡ pred-suc (pred z) +suc-adj z = isHAEquiv.com (snd suc-haequiv) z +pred-adj z = isHAEquiv.com-op (snd suc-haequiv) z + + + +-- ℤ ≡ BiInvℤ + +fwd : ℤ -> BiInvℤ +fwd (pos zero) = zero +fwd (pos (suc n)) = suc (fwd (pos n)) +fwd (negsuc zero) = pred zero +fwd (negsuc (suc n)) = pred (fwd (negsuc n)) + +bwd : BiInvℤ -> ℤ +bwd zero = pos zero +bwd (suc z) = sucℤ (bwd z) +bwd (predr z) = predℤ (bwd z) +bwd (suc-predr z i) = sucPred (bwd z) i +bwd (predl z) = predℤ (bwd z) +bwd (predl-suc z i) = predSuc (bwd z) i + +bwd-fwd : ∀ (x : ℤ) -> bwd (fwd x) ≡ x +bwd-fwd (pos zero) = refl +bwd-fwd (pos (suc n)) = cong sucℤ (bwd-fwd (pos n)) +bwd-fwd (negsuc zero) = refl +bwd-fwd (negsuc (suc n)) = cong predℤ (bwd-fwd (negsuc n)) + + +fwd-sucℤ : ∀ (x : ℤ) → fwd (sucℤ x) ≡ suc (fwd x) +fwd-sucℤ (pos n) = refl +fwd-sucℤ (negsuc zero) = sym (suc-pred (fwd (pos zero))) +fwd-sucℤ (negsuc (suc n)) = sym (suc-pred (fwd (negsuc n))) + +fwd-predℤ : ∀ (x : ℤ) → fwd (predℤ x) ≡ pred (fwd x) +fwd-predℤ (pos zero) = refl +fwd-predℤ (pos (suc n)) = sym (predl-suc (fwd (pos n))) +fwd-predℤ (negsuc n) = refl + +private + sym-filler : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) + → Square {- (i = i0) -} (sym p) + {- (i = i1) -} refl + {- (j = i0) -} refl + {- (j = i1) -} p + sym-filler {y = y} p i j = hcomp (λ k → λ { (j = i0) → y + ; (i = i0) → p ((~ j) ∨ (~ k)) + ; (i = i1) → y + ; (j = i1) → p (i ∨ (~ k)) }) y + +fwd-sucPred : ∀ (x : ℤ) + → Square {- (i = i0) -} (fwd-sucℤ (predℤ x) ∙ (λ i → suc (fwd-predℤ x i))) + {- (i = i1) -} (λ _ → fwd x) + {- (j = i0) -} (λ i → fwd (sucPred x i)) + {- (j = i1) -} (suc-pred (fwd x)) + + +fwd-sucPred (pos zero) i j + = hcomp (λ k → λ { (j = i0) → fwd (pos zero) + ; (i = i0) → rUnit (sym (suc-pred (fwd (pos zero)))) k j + -- because fwd-sucℤ (predℤ (pos zero)) ≡ sym (suc-pred (fwd (pos zero))) + ; (i = i1) → fwd (pos zero) + ; (j = i1) → suc-pred (fwd (pos zero)) i + }) + (sym-filler (suc-pred (fwd (pos zero))) i j) + +fwd-sucPred (pos (suc n)) i j + = hcomp (λ k → λ { (j = i0) → suc (fwd (pos n)) + ; (i = i0) → lUnit (λ i → suc (sym (predl-suc (fwd (pos n))) i)) k j + -- because fwd-predℤ (pos (suc n)) ≡ sym (predl-suc (fwd (pos n))) + ; (i = i1) → suc (fwd (pos n)) + ; (j = i1) → suc-adj (fwd (pos n)) k i + }) + (suc (sym-filler (pred-suc (fwd (pos n))) i j)) + +fwd-sucPred (negsuc n) i j + = hcomp (λ k → λ { (j = i0) → fwd (negsuc n) + ; (i = i0) → rUnit (sym (suc-pred (fwd (negsuc n)))) k j + -- because fwd-sucℤ (predℤ (negsuc n)) ≡ sym (suc-pred (fwd (negsuc n))) + ; (i = i1) → fwd (negsuc n) + ; (j = i1) → suc-pred (fwd (negsuc n)) i + }) + (sym-filler (suc-pred (fwd (negsuc n))) i j) + + +fwd-predSuc : ∀ (x : ℤ) + → Square {- (i = i0) -} (fwd-predℤ (sucℤ x) ∙ (λ i → pred (fwd-sucℤ x i))) + {- (i = i1) -} (λ _ → fwd x) + {- (j = i0) -} (λ i → fwd (predSuc x i)) + {- (j = i1) -} (pred-suc (fwd x)) + +fwd-predSuc (pos n) i j + = hcomp (λ k → λ { (j = i0) → fwd (pos n) + ; (i = i0) → rUnit (sym (pred-suc (fwd (pos n)))) k j + -- because fwd-predℤ (sucℤ (pos n)) ≡ sym (pred-suc (fwd (pos n))) + ; (i = i1) → fwd (pos n) + ; (j = i1) → pred-suc (fwd (pos n)) i + }) + (sym-filler (pred-suc (fwd (pos n))) i j) + +fwd-predSuc (negsuc zero) i j + = hcomp (λ k → λ { (j = i0) → fwd (negsuc zero) + ; (i = i0) → lUnit (λ i → pred (sym (suc-pred (fwd (pos zero))) i)) k j + -- because fwd-sucℤ (negsuc zero) ≡ sym (suc-pred (fwd (pos zero))) + ; (i = i1) → fwd (negsuc zero) + ; (j = i1) → pred-adj (fwd (pos zero)) k i + }) + (pred (sym-filler (suc-pred (fwd (pos zero))) i j)) + +fwd-predSuc (negsuc (suc n)) i j + = hcomp (λ k → λ { (j = i0) → fwd (negsuc (suc n)) + ; (i = i0) → lUnit (λ i → pred (sym (suc-pred (fwd (negsuc n))) i)) k j + -- because fwd-sucℤ (negsuc (suc n)) ≡ sym (suc-pred (fwd (negsuc n))) + ; (i = i1) → fwd (negsuc (suc n)) + ; (j = i1) → pred-adj (fwd (negsuc n)) k i + }) + (pred (sym-filler (suc-pred (fwd (negsuc n))) i j)) + + +fwd-bwd : ∀ (z : BiInvℤ) -> fwd (bwd z) ≡ z +fwd-bwd zero = refl +fwd-bwd (suc z) = fwd-sucℤ (bwd z) ∙ (λ i → suc (fwd-bwd z i)) +fwd-bwd (predr z) = fwd-predℤ (bwd z) ∙ (λ i → predl≡predr (fwd-bwd z i) i) +fwd-bwd (predl z) = fwd-predℤ (bwd z) ∙ (λ i → pred (fwd-bwd z i)) +fwd-bwd (suc-predr z i) j + = hcomp (λ k → λ { (j = i0) → fwd (sucPred (bwd z) i) + ; (i = i0) → (fwd-sucℤ (predℤ (bwd z)) + ∙ (λ i → suc (compPath-filler (fwd-predℤ (bwd z)) + (λ i' → predl≡predr (fwd-bwd z i') i') + k i))) j + ; (i = i1) → fwd-bwd z (j ∧ k) + ; (j = i1) → suc-predl≡predr (fwd-bwd z k) k i }) + (fwd-sucPred (bwd z) i j) +fwd-bwd (predl-suc z i) j + = hcomp (λ k → λ { (j = i0) → fwd (predSuc (bwd z) i) + ; (i = i0) → (fwd-predℤ (sucℤ (bwd z)) + ∙ (λ i → pred (compPath-filler (fwd-sucℤ (bwd z)) + (λ i' → suc (fwd-bwd z i')) + k i))) j + ; (i = i1) → fwd-bwd z (j ∧ k) + ; (j = i1) → pred-suc (fwd-bwd z k) i }) + (fwd-predSuc (bwd z) i j) + + +ℤ≡BiInvℤ : ℤ ≡ BiInvℤ +ℤ≡BiInvℤ = isoToPath (iso fwd bwd fwd-bwd bwd-fwd) + +discreteBiInvℤ : Discrete BiInvℤ +discreteBiInvℤ = subst Discrete ℤ≡BiInvℤ discreteℤ + +isSetBiInvℤ : isSet BiInvℤ +isSetBiInvℤ = subst isSet ℤ≡BiInvℤ isSetℤ + + + +-- suc' is equal to suc (suc≡suc') but cancels pred *exactly* (see suc'-pred) + +suc'ᵖ : (z : BiInvℤ) -> Σ BiInvℤ (suc z ≡_) +suc' = fst ∘ suc'ᵖ +suc≡suc' = snd ∘ suc'ᵖ + +suc'ᵖ zero = suc zero , refl +suc'ᵖ (suc z) = suc (suc z) , refl +suc'ᵖ (predr z) = z , suc-predr z +suc'ᵖ (suc-predr z i) = let filler : I → I → BiInvℤ + filler j i = hfill (λ j → λ { (i = i0) → suc (suc (predr z)) + ; (i = i1) → suc≡suc' z j }) + (inS (suc (suc-predr z i))) j + in filler i1 i , λ j → filler j i +suc'ᵖ (predl z) = z , suc-predl z +suc'ᵖ (predl-suc z i) = let filler : I → I → BiInvℤ + filler j i = hfill (λ j → λ { (i = i0) → suc-predl (suc z) j + ; (i = i1) → suc≡suc' z j }) + (inS (suc (predl-suc z i))) j + in filler i1 i , λ j → filler j i + +private + suc'-pred : ∀ z → suc' (pred z) ≡ z + suc'-pred z = refl + + +-- pred' is equal to pred (pred≡pred') but cancels suc *exactly* (see pred'-suc) + +predr'ᵖ : (z : BiInvℤ) -> Σ BiInvℤ (predr z ≡_) +predr' = fst ∘ predr'ᵖ +predr≡predr' = snd ∘ predr'ᵖ + +predr'ᵖ zero = predr zero , refl +predr'ᵖ (suc z) = z , predr-suc z +predr'ᵖ (predr z) = predr (predr z) , refl +predr'ᵖ (suc-predr z i) = let filler : I → I → BiInvℤ + filler j i = hfill (λ j → λ { (i = i0) → predr-suc (predr z) j + ; (i = i1) → predr≡predr' z j }) + (inS (predr (suc-predr z i))) j + in filler i1 i , λ j → filler j i +predr'ᵖ (predl z) = predr (predl z) , refl +predr'ᵖ (predl-suc z i) = let filler : I → I → BiInvℤ + filler j i = hfill (λ j → λ { (i = i0) → predr (predl (suc z)) + ; (i = i1) → predr≡predr' z j }) + (inS (predr (predl-suc z i))) j + in filler i1 i , λ j → filler j i + +predl'ᵖ : (z : BiInvℤ) -> Σ BiInvℤ (predl z ≡_) +predl' = fst ∘ predl'ᵖ +predl≡predl' = snd ∘ predl'ᵖ + +predl'ᵖ zero = predl zero , refl +predl'ᵖ (suc z) = z , predl-suc z +predl'ᵖ (predr z) = predl (predr z) , refl +predl'ᵖ (suc-predr z i) = let filler : I → I → BiInvℤ + filler j i = hfill (λ j → λ { (i = i0) → predl-suc (predr z) j + ; (i = i1) → predl≡predl' z j }) + (inS (predl (suc-predr z i))) j + in filler i1 i , λ j → filler j i +predl'ᵖ (predl z) = predl (predl z) , refl +predl'ᵖ (predl-suc z i) = let filler : I → I → BiInvℤ + filler j i = hfill (λ j → λ { (i = i0) → predl (predl (suc z)) + ; (i = i1) → predl≡predl' z j }) + (inS (predl (predl-suc z i))) j + in filler i1 i , λ j → filler j i + +private + predr'-suc : ∀ z → predr' (suc z) ≡ z + predr'-suc z = refl + + predl'-suc : ∀ z → predl' (suc z) ≡ z + predl'-suc z = refl + diff --git a/Cubical/Data/Int/MoreInts/BiInvInt/Properties.agda b/Cubical/Data/Int/MoreInts/BiInvInt/Properties.agda new file mode 100644 index 0000000000..c108cc8abc --- /dev/null +++ b/Cubical/Data/Int/MoreInts/BiInvInt/Properties.agda @@ -0,0 +1,243 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.BiInvInt.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Nat using (ℕ) +import Cubical.Data.Int as ℤ +open import Cubical.Data.Bool + +open import Cubical.Data.Int.MoreInts.BiInvInt.Base + +infixl 6 _+_ _-_ +infixl 7 _·_ + +-- To prove a property P about BiInvℤ, we need to show: +-- * P zero +-- * If P n, then P (suc n) +-- * If P n, then P (pred n) +BiInvℤ-ind-prop : + ∀ {ℓ} {P : BiInvℤ → Type ℓ} → (∀ n → isProp (P n)) → + P zero → (∀ n → P n → P (suc n)) → (∀ n → P n → P (pred n)) → + (n : BiInvℤ) → P n +BiInvℤ-ind-prop {P = P} P-isProp P-zero P-suc P-pred = φ + where + P-predr : ∀ n → P n → P (predr n) + P-predr n x = subst P (predl≡predr n) (P-pred n x) + + P-predl : ∀ n → P n → P (predl n) + P-predl = P-pred + + P-isProp' : {a b : BiInvℤ} (p : a ≡ b) (x : P a) (y : P b) → PathP (λ i → P (p i)) x y + P-isProp' _ _ _ = toPathP (P-isProp _ _ _) + + φ : (n : BiInvℤ) → P n + φ zero = P-zero + φ (suc n) = P-suc n (φ n) + φ (predr n) = P-predr n (φ n) + φ (suc-predr n i) = P-isProp' (suc-predr n) (P-suc (predr n) (P-predr n (φ n))) (φ n) i + φ (predl n) = P-predl n (φ n) + φ (predl-suc n i) = P-isProp' (predl-suc n) (P-predl (suc n) (P-suc n (φ n))) (φ n) i + +-- To define a function BiInvℤ → A, we need: +-- · a point z : A for zero +-- · an equivalence s : A ≃ A for suc/pred +BiInvℤ-rec : ∀ {ℓ} {A : Type ℓ} → A → A ≃ A → BiInvℤ → A +BiInvℤ-rec {A = A} z e = φ + where + e-Iso : Iso A A + e-Iso = equivToIso e + + φ : BiInvℤ → A + φ zero = z + φ (suc n) = Iso.fun e-Iso (φ n) + φ (predr n) = Iso.inv e-Iso (φ n) + φ (suc-predr n i) = Iso.rightInv e-Iso (φ n) i + φ (predl n) = Iso.inv e-Iso (φ n) + φ (predl-suc n i) = Iso.leftInv e-Iso (φ n) i + +sucIso : Iso BiInvℤ BiInvℤ +Iso.fun sucIso = suc +Iso.inv sucIso = pred +Iso.rightInv sucIso = suc-predl +Iso.leftInv sucIso = predl-suc + +sucEquiv : BiInvℤ ≃ BiInvℤ +sucEquiv = isoToEquiv sucIso + +-- addition + +-- the following equalities hold definitionally: +-- zero + n ≡ n +-- suc m + n ≡ suc (m + n) +-- pred m + n ≡ pred (m + n) +_+_ : BiInvℤ → BiInvℤ → BiInvℤ +_+_ = BiInvℤ-rec (idfun BiInvℤ) (postCompEquiv sucEquiv) + +-- properties of addition + ++-zero : ∀ n → n + zero ≡ n ++-zero = BiInvℤ-ind-prop (λ _ → isSetBiInvℤ _ _) refl (λ n p → cong suc p) (λ n p → cong pred p) + ++-suc : ∀ m n → m + suc n ≡ suc (m + n) ++-suc = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ m → refl) + (λ m p n → cong suc (p n)) + (λ m p n → cong pred (p n) ∙ predl-suc (m + n) ∙ sym (suc-predl (m + n))) + ++-pred : ∀ m n → m + pred n ≡ pred (m + n) ++-pred = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ m → refl) + (λ m p n → cong suc (p n) ∙ suc-predl (m + n) ∙ sym (predl-suc (m + n))) + (λ m p n → cong pred (p n)) + ++-comm : ∀ m n → m + n ≡ n + m ++-comm m n = +-comm' n m + where + +-comm' : ∀ n m → m + n ≡ n + m + +-comm' = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + +-zero + (λ n p m → +-suc m n ∙ cong suc (p m)) + (λ n p m → +-pred m n ∙ cong pred (p m)) + ++-assoc : ∀ m n o → m + (n + o) ≡ (m + n) + o ++-assoc = BiInvℤ-ind-prop (λ _ → isPropΠ2 λ _ _ → isSetBiInvℤ _ _) + (λ n o → refl) + (λ m p n o → cong suc (p n o)) + (λ m p n o → cong pred (p n o)) + +-- negation / subtraction + +-- the following equalities hold definitionally: +-- - zero ≡ zero +-- - (suc m) ≡ pred m +-- - (pred m) ≡ suc m +-_ : BiInvℤ → BiInvℤ +-_ = BiInvℤ-rec zero (invEquiv sucEquiv) + +_-_ : BiInvℤ → BiInvℤ → BiInvℤ +m - n = m + (- n) + ++-invˡ : ∀ n → (- n) + n ≡ zero ++-invˡ = BiInvℤ-ind-prop (λ _ → isSetBiInvℤ _ _) + refl + (λ n p → cong pred (+-suc (- n) n) ∙ predl-suc _ ∙ p) + (λ n p → cong suc (+-pred (- n) n) ∙ suc-predl _ ∙ p) + ++-invʳ : ∀ n → n + (- n) ≡ zero ++-invʳ = BiInvℤ-ind-prop (λ _ → isSetBiInvℤ _ _) + refl + (λ n p → cong suc (+-pred n (- n)) ∙ suc-predl _ ∙ p) + (λ n p → cong pred (+-suc n (- n)) ∙ predl-suc _ ∙ p) + +inv-hom : ∀ m n → - (m + n) ≡ (- m) + (- n) +inv-hom = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → refl) + (λ m p n → cong pred (p n)) + (λ m p n → cong suc (p n)) + +-- natural injections from ℕ + +pos : ℕ → BiInvℤ +pos ℕ.zero = zero +pos (ℕ.suc n) = suc (pos n) + +neg : ℕ → BiInvℤ +neg ℕ.zero = zero +neg (ℕ.suc n) = pred (neg n) + +-- absolute value +-- (Note that there doesn't appear to be any way around using +-- bwd here! Any direct proof ends up doing the same work...) + +abs : BiInvℤ → ℕ +abs n = ℤ.abs (bwd n) + +Iso-n+ : (n : BiInvℤ) → Iso BiInvℤ BiInvℤ +Iso.fun (Iso-n+ n) = n +_ +Iso.inv (Iso-n+ n) = - n +_ +Iso.rightInv (Iso-n+ n) m = +-assoc n (- n) m ∙ cong (_+ m) (+-invʳ n) +Iso.leftInv (Iso-n+ n) m = +-assoc (- n) n m ∙ cong (_+ m) (+-invˡ n) + +isEquiv-n+ : ∀ n → isEquiv (n +_) +isEquiv-n+ n = isoToIsEquiv (Iso-n+ n) + +-- multiplication + +-- the following equalities hold definitionally: +-- zero · n ≡ zero +-- suc m · n ≡ n + m · n +-- pred m · n ≡ (- n) + m · n +_·_ : BiInvℤ → BiInvℤ → BiInvℤ +m · n = BiInvℤ-rec zero (n +_ , isEquiv-n+ n) m + +-- properties of multiplication + +·-zero : ∀ n → n · zero ≡ zero +·-zero = BiInvℤ-ind-prop (λ _ → isSetBiInvℤ _ _) refl (λ n p → p) (λ n p → p) + +·-suc : ∀ m n → m · suc n ≡ m + m · n +·-suc = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → refl) + (λ m p n → cong suc + (cong (n +_) (p n) ∙ +-assoc n m (m · n) ∙ + cong (_+ m · n) (+-comm n m) ∙ sym (+-assoc m n (m · n)))) + (λ m p n → cong pred + (cong (- n +_) (p n) ∙ +-assoc (- n) m (m · n) ∙ + cong (_+ m · n) (+-comm (- n) m) ∙ sym (+-assoc m (- n) (m · n)))) + +·-pred : ∀ m n → m · pred n ≡ (- m) + m · n +·-pred = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → refl) + (λ m p n → cong pred + (cong (n +_) (p n) ∙ +-assoc n (- m) (m · n) ∙ + cong (_+ m · n) (+-comm n (- m)) ∙ sym (+-assoc (- m) n (m · n)))) + (λ m p n → cong suc + (cong (- n +_) (p n) ∙ +-assoc (- n) (- m) (m · n) ∙ + cong (_+ m · n) (+-comm (- n) (- m)) ∙ sym (+-assoc (- m) (- n) (m · n)))) + +·-comm : ∀ m n → m · n ≡ n · m +·-comm = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → sym (·-zero n)) + (λ m p n → cong (n +_) (p n) ∙ sym (·-suc n m)) + (λ m p n → cong (- n +_) (p n) ∙ sym (·-pred n m)) + +·-identityˡ : ∀ m → suc zero · m ≡ m +·-identityˡ = +-zero + +·-identityʳ : ∀ m → m · suc zero ≡ m +·-identityʳ m = ·-comm m (suc zero) ∙ ·-identityˡ m + +·-distribʳ : ∀ m n o → (m · o) + (n · o) ≡ (m + n) · o +·-distribʳ = BiInvℤ-ind-prop (λ _ → isPropΠ2 λ _ _ → isSetBiInvℤ _ _) + (λ n o → refl) + (λ m p n o → sym (+-assoc o (m · o) (n · o)) ∙ cong (o +_) (p n o)) + (λ m p n o → sym (+-assoc (- o) (m · o) (n · o)) ∙ cong (- o +_) (p n o)) + +·-distribˡ : ∀ o m n → (o · m) + (o · n) ≡ o · (m + n) +·-distribˡ o m n = + cong (_+ o · n) (·-comm o m) ∙ cong (m · o +_) (·-comm o n) ∙ + ·-distribʳ m n o ∙ ·-comm (m + n) o + +·-inv : ∀ m n → m · (- n) ≡ - (m · n) +·-inv = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → refl) + (λ m p n → cong (- n +_) (p n) ∙ sym (inv-hom n (m · n))) + (λ m p n → cong (- (- n) +_) (p n) ∙ sym (inv-hom (- n) (m · n))) + +inv-· : ∀ m n → (- m) · n ≡ - (m · n) +inv-· m n = ·-comm (- m) n ∙ ·-inv n m ∙ cong (-_) (·-comm n m) + +·-assoc : ∀ m n o → m · (n · o) ≡ (m · n) · o +·-assoc = BiInvℤ-ind-prop (λ _ → isPropΠ2 λ _ _ → isSetBiInvℤ _ _) + (λ n o → refl) + (λ m p n o → + cong (n · o +_) (p n o) ∙ ·-distribʳ n (m · n) o) + (λ m p n o → + cong (- (n · o) +_) (p n o) ∙ cong (_+ m · n · o) (sym (inv-· n o)) ∙ + ·-distribʳ (- n) (m · n) o) diff --git a/Cubical/Data/Int/MoreInts/DeltaInt.agda b/Cubical/Data/Int/MoreInts/DeltaInt.agda new file mode 100644 index 0000000000..9c7ad09226 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/DeltaInt.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.DeltaInt where + +open import Cubical.Data.Int.MoreInts.DeltaInt.Base public + +open import Cubical.Data.Int.MoreInts.DeltaInt.Properties public diff --git a/Cubical/Data/Int/MoreInts/DeltaInt/Base.agda b/Cubical/Data/Int/MoreInts/DeltaInt/Base.agda new file mode 100644 index 0000000000..96b80d35a9 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/DeltaInt/Base.agda @@ -0,0 +1,61 @@ +{-# OPTIONS --safe #-} + +{- + +This file defines integers as equivalence classes of pairs of natural numbers +using a direct & untruncated HIT definition (cf. Data.Int.MoreInts.DiffInt) + +and some basic operations, and the zero value: + +succ : DeltaInt → DeltaInt +pred : DeltaInt → DeltaInt +zero : {a : ℕ} → DeltaInt + +and conversion function for ℕ and Int: + +fromℕ : ℕ → DeltaInt +fromℤ : Int → DeltaInt +toℤ : DeltaInt → Int + +and a generalized version of cancel: + +cancelN : ∀ a b n → a ⊖ b ≡ (n + a) ⊖ n + b + +-} + +module Cubical.Data.Int.MoreInts.DeltaInt.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Nat hiding (zero) +open import Cubical.Data.Int hiding (abs; _+_) renaming (ℤ to Int) + +infixl 5 _⊖_ +data DeltaInt : Type₀ where + _⊖_ : ℕ → ℕ → DeltaInt + cancel : ∀ a b → a ⊖ b ≡ suc a ⊖ suc b + +succ : DeltaInt → DeltaInt +succ (x ⊖ y) = suc x ⊖ y +succ (cancel a b i) = cancel (suc a) b i + +pred : DeltaInt → DeltaInt +pred (x ⊖ y) = x ⊖ suc y +pred (cancel a b i) = cancel a (suc b) i + +zero : {a : ℕ} → DeltaInt +zero {a} = a ⊖ a + +fromℕ : ℕ → DeltaInt +fromℕ n = n ⊖ 0 + +fromℤ : Int → DeltaInt +fromℤ (pos n) = fromℕ n +fromℤ (negsuc n) = 0 ⊖ suc n + +toℤ : DeltaInt → Int +toℤ (x ⊖ y) = x ℕ- y +toℤ (cancel a b i) = a ℕ- b + +cancelN : ∀ a b n → a ⊖ b ≡ (n + a) ⊖ n + b +cancelN a b 0 = refl +cancelN a b (suc n) = cancelN a b n ∙ cancel (n + a) (n + b) diff --git a/Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda b/Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda new file mode 100644 index 0000000000..f3f0119f56 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda @@ -0,0 +1,86 @@ +{-# OPTIONS --safe #-} + +{- + +This file defines + +deltaIntSec : ∀ b → toInt (fromInt b) ≡ b +succPred : ∀ n → succ (pred n) ≡ n +predSucc : ∀ n → pred (succ n) ≡ n + +and proved DeltaInt ≡ DeltaInt by the above functions + +succPredEq : DeltaInt ≡ DeltaInt + +along with some interval-relevant lemmas + +cancelDiamond : ∀ a b i → cancel a b i ≡ cancel (suc a) (suc b) i +cancelTriangle : ∀ a b i → a ⊖ b ≡ cancel a b i + +we also have DeltaInt ≡ Int proof + +DeltaInt≡Int : DeltaInt ≡ Int + +and we transported some proofs from Int to DeltaInt + +discreteDeltaInt : Discrete DeltaInt +isSetDeltaInt : isSet DeltaInt + +-} + +module Cubical.Data.Int.MoreInts.DeltaInt.Properties where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Nat hiding (zero) +open import Cubical.Data.Int hiding (abs; _+_) +open import Cubical.Data.Int.MoreInts.DeltaInt.Base +open import Cubical.Relation.Nullary using (Discrete) + +deltaIntSec : ∀ b → toℤ (fromℤ b) ≡ b +deltaIntSec (pos n) = refl +deltaIntSec (negsuc n) = refl + +cancelDiamond : ∀ a b i → cancel a b i ≡ cancel (suc a) (suc b) i +cancelDiamond a b i j = hcomp (λ k → λ + { (i = i0) → cancel a b j + ; (i = i1) → cancel (suc a) (suc b) (j ∧ k) + ; (j = i0) → cancel a b i + ; (j = i1) → cancel (suc a) (suc b) (i ∧ k) + }) (cancel a b (i ∨ j)) + +cancelTriangle : ∀ a b i → a ⊖ b ≡ cancel a b i +cancelTriangle a b i j = hcomp (λ k → λ + { (i = i0) → a ⊖ b + ; (j = i0) → a ⊖ b + ; (j = i1) → cancel a b (i ∧ k) + }) (a ⊖ b) + +deltaIntRet : ∀ a → fromℤ (toℤ a) ≡ a +deltaIntRet (x ⊖ 0) = refl +deltaIntRet (0 ⊖ suc y) = refl +deltaIntRet (suc x ⊖ suc y) = deltaIntRet (x ⊖ y) ∙ cancel x y +deltaIntRet (cancel a 0 i) = cancelTriangle a 0 i +deltaIntRet (cancel 0 (suc b) i) = cancelTriangle 0 (suc b) i +deltaIntRet (cancel (suc a) (suc b) i) = deltaIntRet (cancel a b i) ∙ cancelDiamond a b i + +DeltaInt≡ℤ : DeltaInt ≡ ℤ +DeltaInt≡ℤ = isoToPath (iso toℤ fromℤ deltaIntSec deltaIntRet) + +discreteDeltaInt : Discrete DeltaInt +discreteDeltaInt = subst Discrete (sym DeltaInt≡ℤ) discreteℤ + +isSetDeltaInt : isSet DeltaInt +isSetDeltaInt = subst isSet (sym DeltaInt≡ℤ) isSetℤ + +succPred : ∀ n → succ (pred n) ≡ n +succPred (x ⊖ y) i = cancel x y (~ i) +-- cancel (suc a) (suc b) i ≡ cancel a b i +succPred (cancel a b i) = sym (cancelDiamond a b i) + +predSucc : ∀ n → pred (succ n) ≡ n +predSucc (x ⊖ y) = succPred (x ⊖ y) +predSucc (cancel a b i) = succPred (cancel a b i) + +succPredEq : DeltaInt ≡ DeltaInt +succPredEq = isoToPath (iso succ pred succPred predSucc) + diff --git a/Cubical/Data/Int/MoreInts/DiffInt.agda b/Cubical/Data/Int/MoreInts/DiffInt.agda new file mode 100644 index 0000000000..66354a320c --- /dev/null +++ b/Cubical/Data/Int/MoreInts/DiffInt.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.DiffInt where + +open import Cubical.Data.Int.MoreInts.DiffInt.Base public +open import Cubical.Data.Int.MoreInts.DiffInt.Properties public diff --git a/Cubical/Data/Int/MoreInts/DiffInt/Base.agda b/Cubical/Data/Int/MoreInts/DiffInt/Base.agda new file mode 100644 index 0000000000..28aed1e2fb --- /dev/null +++ b/Cubical/Data/Int/MoreInts/DiffInt/Base.agda @@ -0,0 +1,84 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.DiffInt.Base where + +open import Cubical.Foundations.Prelude + +open import Cubical.HITs.SetQuotients +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat hiding (+-comm ; +-assoc) renaming (_+_ to _+ℕ_) +open import Cubical.Data.Int renaming (ℤ to Int) + +rel : (ℕ × ℕ) → (ℕ × ℕ) → Type₀ +rel (a₀ , b₀) (a₁ , b₁) = x ≡ y + where + x = a₀ +ℕ b₁ + y = a₁ +ℕ b₀ + +ℤ = (ℕ × ℕ) / rel + + +-- Proof of equivalence between Int and DiffInt + +private + -- Prove all the identities for ℕ×ℕ first and use that to build to ℤ + + Int→ℕ×ℕ : Int → (ℕ × ℕ) + Int→ℕ×ℕ (pos n) = ( n , zero ) + Int→ℕ×ℕ (negsuc n) = ( zero , suc n ) + + ℕ×ℕ→Int : (ℕ × ℕ) → Int + ℕ×ℕ→Int(n , m) = n ℕ- m + + relInt→ℕ×ℕ : ∀ m n → rel (Int→ℕ×ℕ (m ℕ- n)) (m , n) + relInt→ℕ×ℕ zero zero = refl + relInt→ℕ×ℕ zero (suc n) = refl + relInt→ℕ×ℕ (suc m) zero = refl + relInt→ℕ×ℕ (suc m) (suc n) = + fst (Int→ℕ×ℕ (m ℕ- n)) +ℕ suc n ≡⟨ +-suc (fst (Int→ℕ×ℕ (m ℕ- n))) n ⟩ + suc (fst (Int→ℕ×ℕ (m ℕ- n)) +ℕ n) ≡⟨ cong suc (relInt→ℕ×ℕ m n) ⟩ + suc (m +ℕ snd (Int→ℕ×ℕ (m ℕ- n))) ∎ + + Int→ℕ×ℕ→Int : ∀ n → ℕ×ℕ→Int (Int→ℕ×ℕ n) ≡ n + Int→ℕ×ℕ→Int (pos n) = refl + Int→ℕ×ℕ→Int (negsuc n) = refl + + ℕ×ℕ→Int→ℕ×ℕ : ∀ p → rel (Int→ℕ×ℕ (ℕ×ℕ→Int p)) p + ℕ×ℕ→Int→ℕ×ℕ (p₀ , p₁) = relInt→ℕ×ℕ p₀ p₁ + +-- Now build to ℤ using the above + +Int→ℤ : Int → ℤ +Int→ℤ n = [ Int→ℕ×ℕ n ] + +ℤ→Int : ℤ → Int +ℤ→Int [ z ] = ℕ×ℕ→Int z +ℤ→Int(eq/ a b r i) = lemℤeq a b r i + where lemℤeq : (a b : (ℕ × ℕ)) → rel a b → ℕ×ℕ→Int(a) ≡ ℕ×ℕ→Int(b) + lemℤeq (a₀ , a₁) (b₀ , b₁) r = a₀ ℕ- a₁ ≡⟨ pos- a₀ a₁ ⟩ + pos a₀ - pos a₁ ≡[ i ]⟨ ((pos a₀ - pos a₁) + -Cancel (pos b₁) (~ i)) ⟩ + (pos a₀ - pos a₁) + (pos b₁ - pos b₁) ≡⟨ +Assoc (pos a₀ + (- pos a₁)) (pos b₁) (- pos b₁) ⟩ + ((pos a₀ - pos a₁) + pos b₁) - pos b₁ ≡[ i ]⟨ +Assoc (pos a₀) (- pos a₁) (pos b₁) (~ i) + (- pos b₁) ⟩ + (pos a₀ + ((- pos a₁) + pos b₁)) - pos b₁ ≡[ i ]⟨ (pos a₀ + +Comm (- pos a₁) (pos b₁) i) - pos b₁ ⟩ + (pos a₀ + (pos b₁ - pos a₁)) - pos b₁ ≡[ i ]⟨ +Assoc (pos a₀) (pos b₁) (- pos a₁) i + (- pos b₁) ⟩ + ((pos a₀ + pos b₁) - pos a₁) - pos b₁ ≡[ i ]⟨ (pos+ a₀ b₁ (~ i) - pos a₁) - pos b₁ ⟩ + (pos (a₀ +ℕ b₁) - pos a₁) - pos b₁ ≡[ i ]⟨ (pos (r i) - pos a₁) - pos b₁ ⟩ + (pos (b₀ +ℕ a₁) - pos a₁) - pos b₁ ≡[ i ]⟨ (pos+ b₀ a₁ i - pos a₁) - pos b₁ ⟩ + ((pos b₀ + pos a₁) - pos a₁) - pos b₁ ≡[ i ]⟨ +Assoc (pos b₀) (pos a₁) (- pos a₁) (~ i) + (- pos b₁) ⟩ + (pos b₀ + (pos a₁ - pos a₁)) - pos b₁ ≡[ i ]⟨ (pos b₀ + (-Cancel (pos a₁) i)) - pos b₁ ⟩ + pos b₀ - pos b₁ ≡[ i ]⟨ pos- b₀ b₁ (~ i) ⟩ + b₀ ℕ- b₁ ∎ +ℤ→Int(squash/ x x₀ p q i j) = isSetℤ (ℤ→Int x) (ℤ→Int x₀) (cong ℤ→Int p) (cong ℤ→Int q) i j + +Int→ℤ→Int : ∀ z → ℤ→Int (Int→ℤ z) ≡ z +Int→ℤ→Int (pos n) = refl +Int→ℤ→Int (negsuc n) = refl + +ℤ→Int→ℤ : ∀ z → Int→ℤ (ℤ→Int z) ≡ z +ℤ→Int→ℤ = elimProp (λ z → squash/ (Int→ℤ (ℤ→Int z)) z) ℕ×ℕprf + where ℕ×ℕprf : (a : ℕ × ℕ) → Int→ℤ (ℤ→Int [ a ]) ≡ [ a ] + ℕ×ℕprf (a , b) = eq/ (Int→ℕ×ℕ (ℕ×ℕ→Int (a , b))) (a , b) (ℕ×ℕ→Int→ℕ×ℕ (a , b)) + +Int≡DiffInt : Int ≡ ℤ +Int≡DiffInt = isoToPath (iso Int→ℤ ℤ→Int ℤ→Int→ℤ Int→ℤ→Int) diff --git a/Cubical/Data/Int/MoreInts/DiffInt/Properties.agda b/Cubical/Data/Int/MoreInts/DiffInt/Properties.agda new file mode 100644 index 0000000000..a5f4f67810 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/DiffInt/Properties.agda @@ -0,0 +1,227 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.DiffInt.Properties where + +open import Cubical.Foundations.HLevels + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Int.MoreInts.DiffInt.Base +open import Cubical.Data.Nat as ℕ renaming (_+_ to _+ℕ_ ; _·_ to _*ℕ_) +open import Cubical.Data.Sigma +open import Cubical.Data.Bool + +open import Cubical.Relation.Binary.Base +open import Cubical.Relation.Nullary + +open import Cubical.HITs.SetQuotients as SetQuotients +open import Cubical.HITs.PropositionalTruncation as PropositionalTruncation +open import Cubical.Foundations.Isomorphism + +open BinaryRelation + +relIsEquiv : isEquivRel rel +relIsEquiv = equivRel {A = ℕ × ℕ} relIsRefl relIsSym relIsTrans + where + relIsRefl : isRefl rel + relIsRefl (a0 , a1) = refl + + relIsSym : isSym rel + relIsSym (a0 , a1) (b0 , b1) p = sym p + + relIsTrans : isTrans rel + relIsTrans (a0 , a1) (b0 , b1) (c0 , c1) p0 p1 = + inj-m+ {m = (b0 +ℕ b1)} ((b0 +ℕ b1) +ℕ (a0 +ℕ c1) ≡⟨ ℕ.+-assoc (b0 +ℕ b1) a0 c1 ⟩ + ((b0 +ℕ b1) +ℕ a0) +ℕ c1 ≡[ i ]⟨ ℕ.+-comm b0 b1 i +ℕ a0 +ℕ c1 ⟩ + ((b1 +ℕ b0) +ℕ a0) +ℕ c1 ≡[ i ]⟨ ℕ.+-comm (b1 +ℕ b0) a0 i +ℕ c1 ⟩ + (a0 +ℕ (b1 +ℕ b0)) +ℕ c1 ≡[ i ]⟨ ℕ.+-assoc a0 b1 b0 i +ℕ c1 ⟩ + (a0 +ℕ b1) +ℕ b0 +ℕ c1 ≡⟨ sym (ℕ.+-assoc (a0 +ℕ b1) b0 c1) ⟩ + (a0 +ℕ b1) +ℕ (b0 +ℕ c1) ≡⟨ cong (λ p → p . fst +ℕ p .snd) (transport ΣPath≡PathΣ (p0 , p1))⟩ + (b0 +ℕ a1) +ℕ (c0 +ℕ b1) ≡⟨ sym (ℕ.+-assoc b0 a1 (c0 +ℕ b1))⟩ + b0 +ℕ (a1 +ℕ (c0 +ℕ b1)) ≡[ i ]⟨ b0 +ℕ (a1 +ℕ ℕ.+-comm c0 b1 i) ⟩ + b0 +ℕ (a1 +ℕ (b1 +ℕ c0)) ≡[ i ]⟨ b0 +ℕ ℕ.+-comm a1 (b1 +ℕ c0) i ⟩ + b0 +ℕ ((b1 +ℕ c0) +ℕ a1) ≡[ i ]⟨ b0 +ℕ ℕ.+-assoc b1 c0 a1 (~ i) ⟩ + b0 +ℕ (b1 +ℕ (c0 +ℕ a1)) ≡⟨ ℕ.+-assoc b0 b1 (c0 +ℕ a1)⟩ + (b0 +ℕ b1) +ℕ (c0 +ℕ a1) ∎ ) + +relIsProp : BinaryRelation.isPropValued rel +relIsProp a b x y = isSetℕ _ _ _ _ + +discreteℤ : Discrete ℤ +discreteℤ = discreteSetQuotients (discreteΣ discreteℕ λ _ → discreteℕ) relIsProp relIsEquiv (λ _ _ → discreteℕ _ _) + +private + _ : Dec→Bool (discreteℤ [ (3 , 5) ] [ (4 , 6) ]) ≡ true + _ = refl + + _ : Dec→Bool (discreteℤ [ (3 , 5) ] [ (4 , 7) ]) ≡ false + _ = refl + + _+ℕ'_ : (ℕ × ℕ) → (ℕ × ℕ) → (ℕ × ℕ) + (n₁ , n₂) +ℕ' (m₁ , m₂) = (n₁ +ℕ m₁ , n₂ +ℕ m₂) + + [_-ℕ'_] : ℕ → ℕ → ℤ + [ x -ℕ' y ] = [ x , y ] + + ℤ-cancelˡ : ∀ {a b} (c : ℕ) → [ (c ℕ.+ a) -ℕ' (c +ℕ b) ] ≡ [ a -ℕ' b ] + ℤ-cancelˡ {a} {b} c = eq/ _ _ (tmp a b c) + where tmp : ∀ a b c → rel (c +ℕ a , c +ℕ b) (a , b) + tmp a b c = + c +ℕ a +ℕ b + ≡⟨ cong (λ x → x +ℕ b) (ℕ.+-comm c a) ⟩ + a +ℕ c +ℕ b + ≡⟨ sym(ℕ.+-assoc a c b) ⟩ + a +ℕ (c +ℕ b) ∎ + + ℤ-cancelʳ : ∀ {a b} (c : ℕ) → [ a +ℕ c -ℕ' b +ℕ c ] ≡ [ a -ℕ' b ] + ℤ-cancelʳ {a} {b} c = eq/ _ _ (tmp a b c) + where tmp : ∀ a b c → rel (a +ℕ c , b +ℕ c) (a , b) + tmp a b c = + a +ℕ c +ℕ b + ≡⟨ sym(ℕ.+-assoc a c b) ⟩ + a +ℕ (c +ℕ b) + ≡⟨ cong (λ x → a +ℕ x) (ℕ.+-comm c b) ⟩ + a +ℕ (b +ℕ c) ∎ + + -- right congruence + +-right-congruence : ∀ x y z → rel y z → rel (x +ℕ' y) (x +ℕ' z) + +-right-congruence x y z p = + (fst x +ℕ fst y) +ℕ (snd x +ℕ snd z) + ≡⟨ cong (λ x' → (fst x +ℕ fst y) +ℕ x') (ℕ.+-comm (snd x) (snd z)) ⟩ + (fst x +ℕ fst y) +ℕ (snd z +ℕ snd x) + ≡⟨ ℕ.+-assoc (fst x +ℕ fst y) (snd z) (snd x) ⟩ + fst x +ℕ fst y +ℕ snd z +ℕ snd x + ≡⟨ cong (λ x' → x' +ℕ snd x) (sym (ℕ.+-assoc (fst x) (fst y) (snd z))) ⟩ + fst x +ℕ (fst y +ℕ snd z) +ℕ snd x + ≡⟨ cong (λ x' → fst x +ℕ x' +ℕ snd x) p ⟩ + fst x +ℕ (fst z +ℕ snd y) +ℕ snd x + ≡⟨ cong (λ x' → x' +ℕ snd x) (ℕ.+-assoc (fst x) (fst z) (snd y)) ⟩ + fst x +ℕ fst z +ℕ snd y +ℕ snd x + ≡⟨ sym (ℕ.+-assoc (fst x +ℕ fst z) (snd y) (snd x)) ⟩ + fst x +ℕ fst z +ℕ (snd y +ℕ snd x) + ≡⟨ cong (λ x' → fst x +ℕ fst z +ℕ x') (ℕ.+-comm (snd y) (snd x)) ⟩ + fst x +ℕ fst z +ℕ (snd x +ℕ snd y) ∎ + + +-left-congruence : ∀ x y z → rel x y → rel (x +ℕ' z) (y +ℕ' z) + +-left-congruence x y z p = + (fst x +ℕ fst z) +ℕ (snd y +ℕ snd z) + ≡⟨ cong (λ x' → x' +ℕ (snd y +ℕ snd z)) (ℕ.+-comm (fst x) (fst z)) ⟩ + (fst z +ℕ fst x) +ℕ (snd y +ℕ snd z) + ≡⟨ ℕ.+-assoc (fst z +ℕ fst x) (snd y) (snd z) ⟩ + fst z +ℕ fst x +ℕ snd y +ℕ snd z + ≡⟨ cong (λ x' → x' +ℕ snd z) (sym (ℕ.+-assoc (fst z) (fst x) (snd y))) ⟩ + fst z +ℕ (fst x +ℕ snd y) +ℕ snd z + ≡⟨ cong (λ x' → fst z +ℕ x' +ℕ snd z) p ⟩ + fst z +ℕ (fst y +ℕ snd x) +ℕ snd z + ≡⟨ cong (λ x' → x' +ℕ snd z) (ℕ.+-assoc (fst z) (fst y) (snd x)) ⟩ + fst z +ℕ fst y +ℕ snd x +ℕ snd z + ≡⟨ sym (ℕ.+-assoc (fst z +ℕ fst y) (snd x) (snd z)) ⟩ + fst z +ℕ fst y +ℕ (snd x +ℕ snd z) + ≡⟨ cong (λ x' → x' +ℕ (snd x +ℕ snd z)) (ℕ.+-comm (fst z) (fst y)) ⟩ + fst y +ℕ fst z +ℕ (snd x +ℕ snd z) ∎ + +ℤ-isSet : isSet ℤ +ℤ-isSet m n p q = squash/ m n p q + +_+_ : ℤ → ℤ → ℤ +_+_ = SetQuotients.rec2 ℤ-isSet (λ x y → [ x +ℕ' y ]) feql feqr + where + feql : (a b c : ℕ × ℕ) (r : rel a b) → [ a +ℕ' c ] ≡ [ b +ℕ' c ] + feql a b c r = eq/ (a +ℕ' c) (b +ℕ' c) (+-left-congruence a b c r) + feqr : (a b c : ℕ × ℕ) (r : rel b c) → [ a +ℕ' b ] ≡ [ a +ℕ' c ] + feqr a b c r = eq/ (a +ℕ' b) (a +ℕ' c) (+-right-congruence a b c r) + +private + zero-identityˡ-lem : (a : ℕ)(b : ℤ) → (Σ (ℕ × ℕ) ( λ (c , d) → [ (c , d) ] ≡ b )) → [ a , a ] + b ≡ b + zero-identityˡ-lem a b ( (c , d) , p ) = + [ a , a ] + b ≡⟨ cong (λ x → [ a , a ] + x) (sym p) ⟩ + [ a , a ] + [ c , d ] ≡⟨ ℤ-cancelˡ a ⟩ + [ c , d ] ≡⟨ p ⟩ + b ∎ + + zero-identityʳ-lem : (a : ℕ)(b : ℤ) → (Σ (ℕ × ℕ) ( λ (c , d) → [ (c , d) ] ≡ b )) → b + [ a , a ] ≡ b + zero-identityʳ-lem a b ( (c , d) , p ) = + b + [ a , a ] ≡⟨ cong (λ x → x + [ a , a ]) (sym(p)) ⟩ + [ c , d ] + [ a , a ] ≡⟨ ℤ-cancelʳ a ⟩ + [ c , d ] ≡⟨ p ⟩ + b ∎ + +zero-identityˡ : (a : ℕ) (b : ℤ) → [ a , a ] + b ≡ b +zero-identityˡ a b = PropositionalTruncation.rec (lem2 a b) (zero-identityˡ-lem a b) ([]surjective b) + where + lem2 : (a : ℕ) (b : ℤ) → isProp ([ a , a ] + b ≡ b) + lem2 a b = ℤ-isSet ([ a , a ] + b) b + +zero-identityʳ : (a : ℕ)(b : ℤ) → b + [ a , a ] ≡ b +zero-identityʳ a b = PropositionalTruncation.rec (lem2 a b) (zero-identityʳ-lem a b) ([]surjective b) + where + lem2 : (a : ℕ) (b : ℤ) → isProp (b + [ a , a ] ≡ b) + lem2 a b = ℤ-isSet (b + [ a , a ]) b + +-ℤ_ : ℤ → ℤ +-ℤ [ a ] = [ snd a , fst a ] +-ℤ eq/ a b r i = eq/ (snd a , fst a) (snd b , fst b) (tmp a b r) i + where + tmp : ∀ a b → fst a +ℕ snd b ≡ fst b +ℕ snd a → snd a +ℕ fst b ≡ snd b +ℕ fst a + tmp a b r = + snd a +ℕ fst b ≡⟨ ℕ.+-comm (snd a) (fst b) ⟩ + fst b +ℕ snd a ≡⟨ sym r ⟩ + fst a +ℕ snd b ≡⟨ ℕ.+-comm (fst a) (snd b) ⟩ + snd b +ℕ fst a ∎ +-ℤ squash/ a a₁ p q i i₁ = squash/ (-ℤ a) (-ℤ a₁) (cong (λ x → -ℤ x) p) (cong (λ x → -ℤ x) q) i i₁ + +private + -ℤ-invʳ-lem : (b : ℤ) → (Σ (ℕ × ℕ) ( λ (c , d) → [ (c , d) ] ≡ b )) → b + (-ℤ b) ≡ [ 0 , 0 ] + -ℤ-invʳ-lem b ( (c , d) , p ) = + b + (-ℤ b) + ≡⟨ cong (λ x → x + (-ℤ b)) (sym p) ⟩ + [ c , d ] + (-ℤ b) + ≡⟨ cong (λ x → [ c , d ] + (-ℤ x)) (sym p) ⟩ + [ c +ℕ d , d +ℕ c ] + ≡⟨ cong (λ x → [ c +ℕ d , x ]) (ℕ.+-comm d c) ⟩ + [ c +ℕ d , c +ℕ d ] + ≡⟨ ℤ-cancelˡ c ⟩ + [ d -ℕ' d ] + ≡⟨ sym(ℤ-cancelʳ 0) ⟩ + [ d +ℕ 0 -ℕ' d +ℕ 0 ] + ≡⟨ ℤ-cancelˡ d ⟩ + [ 0 , 0 ] ∎ + + -ℤ-invˡ-lem : (b : ℤ) → (Σ (ℕ × ℕ) ( λ (c , d) → [ (c , d) ] ≡ b )) → (-ℤ b) + b ≡ [ 0 , 0 ] + -ℤ-invˡ-lem b ( (c , d) , p ) = + (-ℤ b) + b + ≡⟨ cong (λ x → (-ℤ b) + x) (sym p) ⟩ + (-ℤ b) + [ c , d ] + ≡⟨ cong (λ x → (-ℤ x) + [ c , d ]) (sym p) ⟩ + [ d +ℕ c , c +ℕ d ] + ≡⟨ cong (λ x → [ x , c +ℕ d ]) (ℕ.+-comm d c) ⟩ + [ c +ℕ d , c +ℕ d ] + ≡⟨ ℤ-cancelˡ c ⟩ + [ d -ℕ' d ] + ≡⟨ sym(ℤ-cancelʳ 0) ⟩ + [ d +ℕ 0 -ℕ' d +ℕ 0 ] + ≡⟨ ℤ-cancelˡ d ⟩ + [ 0 , 0 ] ∎ + +_-_ : ℤ → ℤ → ℤ +a - b = a + (-ℤ b) + +-ℤ-invʳ : (b : ℤ) → b + (-ℤ b) ≡ [ 0 , 0 ] +-ℤ-invʳ b = PropositionalTruncation.rec (lem2 b) (-ℤ-invʳ-lem b) ([]surjective b) + where + lem2 : (b : ℤ) → isProp (b + (-ℤ b) ≡ [ 0 , 0 ]) + lem2 b = ℤ-isSet (b + (-ℤ b)) [ 0 , 0 ] + +-ℤ-invˡ : (b : ℤ) → (-ℤ b) + b ≡ [ 0 , 0 ] +-ℤ-invˡ b = PropositionalTruncation.rec (lem2 b) (-ℤ-invˡ-lem b) ([]surjective b) + where + lem2 : (b : ℤ) → isProp ((-ℤ b) + b ≡ [ 0 , 0 ]) + lem2 b = ℤ-isSet ((-ℤ b) + b) [ 0 , 0 ] + ++ℤ-assoc : ∀ x y z → x + (y + z) ≡ (x + y) + z ++ℤ-assoc = elimProp3 (λ _ _ _ → ℤ-isSet _ _) + (λ { (a , b) (c , d) (e , f) i → [ ℕ.+-assoc a c e i -ℕ' ℕ.+-assoc b d f i ] }) + ++ℤ-comm : ∀ x y → x + y ≡ y + x ++ℤ-comm = elimProp2 (λ _ _ → ℤ-isSet _ _) + λ ( a , b ) ( c , d ) i → [ ℕ.+-comm a c i -ℕ' ℕ.+-comm b d i ] diff --git a/Cubical/Data/Int/MoreInts/QuoInt.agda b/Cubical/Data/Int/MoreInts/QuoInt.agda new file mode 100644 index 0000000000..a404482428 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/QuoInt.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.QuoInt where + +open import Cubical.Data.Int.MoreInts.QuoInt.Base public + +open import Cubical.Data.Int.MoreInts.QuoInt.Properties public diff --git a/Cubical/Data/Int/MoreInts/QuoInt/Base.agda b/Cubical/Data/Int/MoreInts/QuoInt/Base.agda new file mode 100644 index 0000000000..7de09c85f0 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/QuoInt/Base.agda @@ -0,0 +1,215 @@ +-- Define the integers as a HIT by identifying +0 and -0 +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.QuoInt.Base where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Relation.Nullary + +open import Cubical.Data.Int using () renaming (ℤ to Int ; discreteℤ to discreteInt ; isSetℤ to isSetInt) +open import Cubical.Data.Nat as ℕ using (ℕ; zero; suc) +open import Cubical.Data.Bool as Bool using (Bool; not; notnot) + +variable + l : Level + + +Sign : Type₀ +Sign = Bool + +pattern spos = Bool.false +pattern sneg = Bool.true + +_·S_ : Sign → Sign → Sign +_·S_ = Bool._⊕_ + + +data ℤ : Type₀ where + signed : (s : Sign) (n : ℕ) → ℤ + posneg : signed spos 0 ≡ signed sneg 0 + +pattern pos n = signed spos n +pattern neg n = signed sneg n + + +sign : ℤ → Sign +sign (signed _ zero) = spos +sign (signed s (suc _)) = s +sign (posneg i) = spos + +sign-pos : ∀ n → sign (pos n) ≡ spos +sign-pos zero = refl +sign-pos (suc n) = refl + +abs : ℤ → ℕ +abs (signed _ n) = n +abs (posneg i) = zero + +signed-inv : ∀ n → signed (sign n) (abs n) ≡ n +signed-inv (pos zero) = refl +signed-inv (neg zero) = posneg +signed-inv (signed s (suc n)) = refl +signed-inv (posneg i) j = posneg (i ∧ j) + +signed-zero : ∀ s₁ s₂ → signed s₁ zero ≡ signed s₂ zero +signed-zero spos spos = refl +signed-zero sneg sneg = refl +signed-zero spos sneg = posneg +signed-zero sneg spos = sym posneg + + +rec : ∀ {A : Type l} → (pos' neg' : ℕ → A) → pos' 0 ≡ neg' 0 → ℤ → A +rec pos' neg' eq (pos m) = pos' m +rec pos' neg' eq (neg m) = neg' m +rec pos' neg' eq (posneg i) = eq i + +elim : ∀ (P : ℤ → Type l) + → (pos' : ∀ n → P (pos n)) + → (neg' : ∀ n → P (neg n)) + → (λ i → P (posneg i)) [ pos' 0 ≡ neg' 0 ] + → ∀ z → P z +elim P pos' neg' eq (pos n) = pos' n +elim P pos' neg' eq (neg n) = neg' n +elim P pos' neg' eq (posneg i) = eq i + + +Int→ℤ : Int → ℤ +Int→ℤ (Int.pos n) = pos n +Int→ℤ (Int.negsuc n) = neg (suc n) + +ℤ→Int : ℤ → Int +ℤ→Int (pos n) = Int.pos n +ℤ→Int (neg zero) = Int.pos 0 +ℤ→Int (neg (suc n)) = Int.negsuc n +ℤ→Int (posneg _) = Int.pos 0 + +ℤ→Int→ℤ : ∀ (n : ℤ) → Int→ℤ (ℤ→Int n) ≡ n +ℤ→Int→ℤ (pos n) _ = pos n +ℤ→Int→ℤ (neg zero) i = posneg i +ℤ→Int→ℤ (neg (suc n)) _ = neg (suc n) +ℤ→Int→ℤ (posneg j) i = posneg (j ∧ i) + +Int→ℤ→Int : ∀ (n : Int) → ℤ→Int (Int→ℤ n) ≡ n +Int→ℤ→Int (Int.pos n) _ = Int.pos n +Int→ℤ→Int (Int.negsuc n) _ = Int.negsuc n + +Int≡ℤ : Int ≡ ℤ +Int≡ℤ = isoToPath (iso Int→ℤ ℤ→Int ℤ→Int→ℤ Int→ℤ→Int) + +discreteℤ : Discrete ℤ +discreteℤ = subst Discrete Int≡ℤ discreteInt + +isSetℤ : isSet ℤ +isSetℤ = subst isSet Int≡ℤ isSetInt + + +-_ : ℤ → ℤ +- signed s n = signed (not s) n +- posneg i = posneg (~ i) + +negate-invol : ∀ n → - - n ≡ n +negate-invol (signed s n) i = signed (notnot s i) n +negate-invol (posneg i) _ = posneg i + +negateEquiv : ℤ ≃ ℤ +negateEquiv = isoToEquiv (iso -_ -_ negate-invol negate-invol) + +negateEq : ℤ ≡ ℤ +negateEq = ua negateEquiv + + +infixl 6 _+_ +infixl 7 _·_ + +sucℤ : ℤ → ℤ +sucℤ (pos n) = pos (suc n) +sucℤ (neg zero) = pos 1 +sucℤ (neg (suc n)) = neg n +sucℤ (posneg _) = pos 1 + +predℤ : ℤ → ℤ +predℤ = subst (λ Z → (Z → Z)) negateEq sucℤ + -- definitionally equal to λ n → - (sucℤ (- n)) + -- strictly more useful than the direct pattern matching version, + -- see negateSuc and negatePred + +sucPredℤ : ∀ n → sucℤ (predℤ n) ≡ n +sucPredℤ (pos zero) = sym posneg +sucPredℤ (pos (suc _)) = refl +sucPredℤ (neg _) = refl +sucPredℤ (posneg i) j = posneg (i ∨ ~ j) + +predSucℤ : ∀ n → predℤ (sucℤ n) ≡ n +predSucℤ (pos _) = refl +predSucℤ (neg zero) = posneg +predSucℤ (neg (suc _)) = refl +predSucℤ (posneg i) j = posneg (i ∧ j) + +_+_ : ℤ → ℤ → ℤ +(signed _ zero) + n = n +(posneg _) + n = n +(pos (suc m)) + n = sucℤ (pos m + n) +(neg (suc m)) + n = predℤ (neg m + n) + + +sucPathℤ : ℤ ≡ ℤ +sucPathℤ = isoToPath (iso sucℤ predℤ sucPredℤ predSucℤ) + +-- We do the same trick as in Cubical.Data.Int to prove that addition +-- is an equivalence +addEqℤ : ℕ → ℤ ≡ ℤ +addEqℤ zero = refl +addEqℤ (suc n) = addEqℤ n ∙ sucPathℤ + +predPathℤ : ℤ ≡ ℤ +predPathℤ = isoToPath (iso predℤ sucℤ predSucℤ sucPredℤ) + +subEqℤ : ℕ → ℤ ≡ ℤ +subEqℤ zero = refl +subEqℤ (suc n) = subEqℤ n ∙ predPathℤ + +addℤ : ℤ → ℤ → ℤ +addℤ (pos m) n = transport (addEqℤ m) n +addℤ (neg m) n = transport (subEqℤ m) n +addℤ (posneg _) n = n + +isEquivAddℤ : (m : ℤ) → isEquiv (addℤ m) +isEquivAddℤ (pos n) = isEquivTransport (addEqℤ n) +isEquivAddℤ (neg n) = isEquivTransport (subEqℤ n) +isEquivAddℤ (posneg _) = isEquivTransport refl + +addℤ≡+ℤ : addℤ ≡ _+_ +addℤ≡+ℤ i (pos (suc m)) n = sucℤ (addℤ≡+ℤ i (pos m) n) +addℤ≡+ℤ i (neg (suc m)) n = predℤ (addℤ≡+ℤ i (neg m) n) +addℤ≡+ℤ i (pos zero) n = n +addℤ≡+ℤ i (neg zero) n = n +addℤ≡+ℤ _ (posneg _) n = n + +isEquiv+ℤ : (m : ℤ) → isEquiv (m +_) +isEquiv+ℤ = subst (λ _+_ → (m : ℤ) → isEquiv (m +_)) addℤ≡+ℤ isEquivAddℤ + + +_·_ : ℤ → ℤ → ℤ +m · n = signed (sign m ·S sign n) (abs m ℕ.· abs n) + +private + ·-abs : ∀ m n → abs (m · n) ≡ abs m ℕ.· abs n + ·-abs m n = refl + + +-- Natural number and negative integer literals for ℤ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℤ : HasFromNat ℤ + fromNatℤ = record { Constraint = λ _ → Unit ; fromNat = λ n → pos n } + +instance + fromNegℤ : HasFromNeg ℤ + fromNegℤ = record { Constraint = λ _ → Unit ; fromNeg = λ n → neg n } diff --git a/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda b/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda new file mode 100644 index 0000000000..ac61e796b4 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda @@ -0,0 +1,235 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.QuoInt.Properties where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Relation.Nullary + +open import Cubical.Data.Nat as ℕ using (ℕ; zero; suc) +open import Cubical.Data.Bool as Bool using (Bool; not; notnot) +open import Cubical.Data.Empty +open import Cubical.Data.Unit renaming (Unit to ⊤) + +open import Cubical.Data.Int.MoreInts.QuoInt.Base + +·S-comm : ∀ x y → x ·S y ≡ y ·S x +·S-comm = Bool.⊕-comm + +·S-assoc : ∀ x y z → x ·S (y ·S z) ≡ (x ·S y) ·S z +·S-assoc = Bool.⊕-assoc + +not-·Sˡ : ∀ x y → not (x ·S y) ≡ not x ·S y +not-·Sˡ = Bool.not-⊕ˡ + + +snotz : ∀ s n s' → ¬ (signed s (suc n) ≡ signed s' zero) +snotz s n s' eq = subst (λ { (signed s (suc n)) → ⊤ + ; _ → ⊥ }) eq tt + + ++-zeroʳ : ∀ s m → m + signed s zero ≡ m ++-zeroʳ s (signed s' zero) = signed-zero s s' ++-zeroʳ s (pos (suc n)) = cong sucℤ (+-zeroʳ s (pos n)) ++-zeroʳ s (neg (suc n)) = cong predℤ (+-zeroʳ s (neg n)) ++-zeroʳ spos (posneg i) j = posneg (i ∧ j) ++-zeroʳ sneg (posneg i) j = posneg (i ∨ ~ j) + ++-identityʳ : ∀ m → m + pos zero ≡ m ++-identityʳ = +-zeroʳ spos + +sucℤ-+ʳ : ∀ m n → sucℤ (m + n) ≡ m + sucℤ n +sucℤ-+ʳ (signed _ zero) _ = refl +sucℤ-+ʳ (posneg _) _ = refl +sucℤ-+ʳ (pos (suc m)) n = cong sucℤ (sucℤ-+ʳ (pos m) n) +sucℤ-+ʳ (neg (suc m)) n = + sucPredℤ (neg m + n) ∙∙ + sym (predSucℤ (neg m + n)) ∙∙ + cong predℤ (sucℤ-+ʳ (neg m) n) + +-- I wonder if we could somehow use negEq to get this for free from the above... +predℤ-+ʳ : ∀ m n → predℤ (m + n) ≡ m + predℤ n +predℤ-+ʳ (signed _ zero) _ = refl +predℤ-+ʳ (posneg _) _ = refl +predℤ-+ʳ (neg (suc m)) n = cong predℤ (predℤ-+ʳ (neg m) n) +predℤ-+ʳ (pos (suc m)) n = + predSucℤ (pos m + n) ∙∙ + sym (sucPredℤ (pos m + n)) ∙∙ + cong sucℤ (predℤ-+ʳ (pos m) n) + ++-comm : ∀ m n → m + n ≡ n + m ++-comm (signed s zero) n = sym (+-zeroʳ s n) ++-comm (pos (suc m)) n = cong sucℤ (+-comm (pos m) n) ∙ sucℤ-+ʳ n (pos m) ++-comm (neg (suc m)) n = cong predℤ (+-comm (neg m) n) ∙ predℤ-+ʳ n (neg m) ++-comm (posneg i) n = lemma n i + where + -- get some help from: + -- https://www.codewars.com/kata/reviews/5c878e3ef1abb10001e32eb1/groups/5cca3f9e840f4b0001d8ac98 + lemma : ∀ n → (λ i → (posneg i + n) ≡ (n + posneg i)) + [ sym (+-zeroʳ spos n) ≡ sym (+-zeroʳ sneg n) ] + lemma (pos zero) i j = posneg (i ∧ j) + lemma (pos (suc n)) i = cong sucℤ (lemma (pos n) i) + lemma (neg zero) i j = posneg (i ∨ ~ j) + lemma (neg (suc n)) i = cong predℤ (lemma (neg n) i) + lemma (posneg i) j k = posneg ((i ∧ ~ k) ∨ (i ∧ j) ∨ (j ∧ k)) + +sucℤ-+ˡ : ∀ m n → sucℤ (m + n) ≡ sucℤ m + n +sucℤ-+ˡ (pos _) n = refl +sucℤ-+ˡ (neg zero) n = refl +sucℤ-+ˡ (neg (suc m)) n = sucPredℤ _ +sucℤ-+ˡ (posneg _) n = refl + +-- I wonder if we could somehow use negEq to get this for free from the above... +predℤ-+ˡ : ∀ m n → predℤ (m + n) ≡ predℤ m + n +predℤ-+ˡ (pos zero) n = refl +predℤ-+ˡ (pos (suc m)) n = predSucℤ _ +predℤ-+ˡ (neg _) n = refl +predℤ-+ˡ (posneg _) n = refl + ++-assoc : ∀ m n o → m + (n + o) ≡ m + n + o ++-assoc (signed s zero) n o = refl ++-assoc (posneg i) n o = refl ++-assoc (pos (suc m)) n o = cong sucℤ (+-assoc (pos m) n o) ∙ sucℤ-+ˡ (pos m + n) o ++-assoc (neg (suc m)) n o = cong predℤ (+-assoc (neg m) n o) ∙ predℤ-+ˡ (neg m + n) o + + +sucℤ-inj : ∀ m n → sucℤ m ≡ sucℤ n → m ≡ n +sucℤ-inj m n p = sym (predSucℤ m) ∙ cong predℤ p ∙ predSucℤ n + +predℤ-inj : ∀ m n → predℤ m ≡ predℤ n → m ≡ n +predℤ-inj m n p = sym (sucPredℤ m) ∙ cong sucℤ p ∙ sucPredℤ n + ++-injˡ : ∀ o m n → o + m ≡ o + n → m ≡ n ++-injˡ (signed _ zero) _ _ p = p ++-injˡ (posneg _) _ _ p = p ++-injˡ (pos (suc o)) m n p = +-injˡ (pos o) m n (sucℤ-inj _ _ p) ++-injˡ (neg (suc o)) m n p = +-injˡ (neg o) m n (predℤ-inj _ _ p) + ++-injʳ : ∀ m n o → m + o ≡ n + o → m ≡ n ++-injʳ m n o p = +-injˡ o m n (+-comm o m ∙ p ∙ +-comm n o) + + +·-comm : ∀ m n → m · n ≡ n · m +·-comm m n i = signed (·S-comm (sign m) (sign n) i) (ℕ.·-comm (abs m) (abs n) i) + +·-identityˡ : ∀ n → pos 1 · n ≡ n +·-identityˡ n = cong (signed (sign n)) (ℕ.+-zero (abs n)) ∙ signed-inv n + +·-identityʳ : ∀ n → n · pos 1 ≡ n +·-identityʳ n = ·-comm n (pos 1) ∙ ·-identityˡ n + +·-zeroˡ : ∀ {s} n → signed s zero · n ≡ signed s zero +·-zeroˡ _ = signed-zero _ _ + +·-zeroʳ : ∀ {s} n → n · signed s zero ≡ signed s zero +·-zeroʳ n = cong (signed _) (sym (ℕ.0≡m·0 (abs n))) ∙ signed-zero _ _ + +·-signed-pos : ∀ {s} m n → signed s m · pos n ≡ signed s (m ℕ.· n) +·-signed-pos {s} zero n = ·-zeroˡ {s} (pos n) +·-signed-pos {s} (suc m) n i = signed (·S-comm s (sign-pos n i) i) (suc m ℕ.· n) + + +-- this proof is why we defined ℤ using `signed` instead of `pos` and `neg` +-- based on that in: https://github.com/danr/Agda-Numerics +·-assoc : ∀ m n o → m · (n · o) ≡ m · n · o + +·-assoc (signed s zero) n o = + ·-zeroˡ (n · o) +·-assoc m@(signed _ (suc _)) (signed s zero) o = + ·-zeroʳ {sign o} m ∙ signed-zero _ _ ∙ cong (_· o) (sym (·-zeroʳ {s} m)) +·-assoc m@(signed _ (suc _)) n@(signed _ (suc _)) (signed s zero) = + cong (m ·_) (·-zeroʳ {s} n) ∙ ·-zeroʳ {s} m ∙ sym (·-zeroʳ {s} (m · n)) + +·-assoc (signed sm (suc m)) (signed sn (suc n)) (signed so (suc o)) i = + signed (·S-assoc sm sn so i) (ℕ.·-assoc (suc m) (suc n) (suc o) i) + +·-assoc (posneg i) n o j = + isSet→isSet' isSetℤ (·-assoc (pos zero) n o) (·-assoc (neg zero) n o) + (λ i → posneg i · (n · o)) (λ i → posneg i · n · o) i j +·-assoc m@(signed _ (suc _)) (posneg i) o j = + isSet→isSet' isSetℤ (·-assoc m (pos zero) o) (·-assoc m (neg zero) o) + (λ i → m · (posneg i · o)) (λ i → m · posneg i · o) i j +·-assoc m@(signed _ (suc _)) n@(signed _ (suc _)) (posneg i) j = + isSet→isSet' isSetℤ (·-assoc m n (pos zero)) (·-assoc m n (neg zero)) + (λ i → m · (n · posneg i)) (λ i → m · n · posneg i) i j + + +negateSuc : ∀ n → - sucℤ n ≡ predℤ (- n) +negateSuc n i = - sucℤ (negate-invol n (~ i)) + +negatePred : ∀ n → - predℤ n ≡ sucℤ (- n) +negatePred n i = negate-invol (sucℤ (- n)) i + +negate-+ : ∀ m n → - (m + n) ≡ (- m) + (- n) +negate-+ (signed _ zero) n = refl +negate-+ (posneg _) n = refl +negate-+ (pos (suc m)) n = negateSuc (pos m + n) ∙ cong predℤ (negate-+ (pos m) n) +negate-+ (neg (suc m)) n = negatePred (neg m + n) ∙ cong sucℤ (negate-+ (neg m) n) + + +negate-·ˡ : ∀ m n → - (m · n) ≡ (- m) · n +negate-·ˡ (signed _ zero) n = signed-zero (not (sign n)) (sign n) +negate-·ˡ (signed ss (suc m)) n i = signed (not-·Sˡ ss (sign n) i) (suc m ℕ.· abs n) +negate-·ˡ (posneg i) n j = + isSet→isSet' isSetℤ (signed-zero (not (sign n)) _) (signed-zero _ _) + refl (λ i → posneg (~ i) · n) i j + + +signed-distrib : ∀ s m n → signed s (m ℕ.+ n) ≡ signed s m + signed s n +signed-distrib s zero n = refl +signed-distrib spos (suc m) n = cong sucℤ (signed-distrib spos m n) +signed-distrib sneg (suc m) n = cong predℤ (signed-distrib sneg m n) + +·-pos-suc : ∀ m n → pos (suc m) · n ≡ n + pos m · n +·-pos-suc m n = signed-distrib (sign n) (abs n) (m ℕ.· abs n) + ∙ (λ i → signed-inv n i + signed (sign-pos m (~ i) ·S sign n) (m ℕ.· abs n)) + + +-- the below is based on that in: https://github.com/danr/Agda-Numerics + +·-distribˡ-pos : ∀ o m n → (pos o · m) + (pos o · n) ≡ pos o · (m + n) +·-distribˡ-pos zero m n = signed-zero (sign n) (sign (m + n)) +·-distribˡ-pos (suc o) m n = + pos (suc o) · m + pos (suc o) · n ≡[ i ]⟨ ·-pos-suc o m i + ·-pos-suc o n i ⟩ + m + pos o · m + (n + pos o · n) ≡⟨ +-assoc (m + pos o · m) n (pos o · n) ⟩ + m + pos o · m + n + pos o · n ≡[ i ]⟨ +-assoc m (pos o · m) n (~ i) + pos o · n ⟩ + m + (pos o · m + n) + pos o · n ≡[ i ]⟨ m + +-comm (pos o · m) n i + pos o · n ⟩ + m + (n + pos o · m) + pos o · n ≡[ i ]⟨ +-assoc m n (pos o · m) i + pos o · n ⟩ + m + n + pos o · m + pos o · n ≡⟨ sym (+-assoc (m + n) (pos o · m) (pos o · n)) ⟩ + m + n + (pos o · m + pos o · n) ≡⟨ cong ((m + n) +_) (·-distribˡ-pos o m n) ⟩ + m + n + pos o · (m + n) ≡⟨ sym (·-pos-suc o (m + n)) ⟩ + pos (suc o) · (m + n) ∎ + +·-distribˡ-neg : ∀ o m n → (neg o · m) + (neg o · n) ≡ neg o · (m + n) +·-distribˡ-neg o m n = + neg o · m + neg o · n ≡[ i ]⟨ negate-·ˡ (pos o) m (~ i) + negate-·ˡ (pos o) n (~ i) ⟩ + - (pos o · m) + - (pos o · n) ≡⟨ sym (negate-+ (pos o · m) (pos o · n)) ⟩ + - (pos o · m + pos o · n) ≡⟨ cong -_ (·-distribˡ-pos o m n) ⟩ + - (pos o · (m + n)) ≡⟨ negate-·ˡ (pos o) (m + n) ⟩ + neg o · (m + n) ∎ + +·-distribˡ : ∀ o m n → (o · m) + (o · n) ≡ o · (m + n) +·-distribˡ (pos o) m n = ·-distribˡ-pos o m n +·-distribˡ (neg o) m n = ·-distribˡ-neg o m n +·-distribˡ (posneg i) m n j = + isSet→isSet' isSetℤ (·-distribˡ-pos zero m n) (·-distribˡ-neg zero m n) + (λ i → posneg i · n) (λ i → posneg i · (m + n)) i j + +·-distribʳ : ∀ m n o → (m · o) + (n · o) ≡ (m + n) · o +·-distribʳ m n o = (λ i → ·-comm m o i + ·-comm n o i) ∙ ·-distribˡ o m n ∙ ·-comm o (m + n) + + +sign-pos-suc-· : ∀ m n → sign (pos (suc m) · n) ≡ sign n +sign-pos-suc-· m (signed s zero) = sign-pos (suc m ℕ.· zero) +sign-pos-suc-· m (posneg i) = sign-pos (suc m ℕ.· zero) +sign-pos-suc-· m (signed s (suc n)) = refl + +·-injˡ : ∀ o m n → pos (suc o) · m ≡ pos (suc o) · n → m ≡ n +·-injˡ o m n p = sym (signed-inv m) ∙ (λ i → signed (sign-eq i) (abs-eq i)) ∙ signed-inv n + where sign-eq = sym (sign-pos-suc-· o m) ∙ cong sign p ∙ sign-pos-suc-· o n + abs-eq = ℕ.inj-sm· {o} (cong abs p) + +·-injʳ : ∀ m n o → m · pos (suc o) ≡ n · pos (suc o) → m ≡ n +·-injʳ m n o p = ·-injˡ o m n (·-comm (pos (suc o)) m ∙ p ∙ ·-comm n (pos (suc o))) diff --git a/Cubical/Data/Int/Properties.agda b/Cubical/Data/Int/Properties.agda new file mode 100644 index 0000000000..88a5a16d31 --- /dev/null +++ b/Cubical/Data/Int/Properties.agda @@ -0,0 +1,487 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Empty +open import Cubical.Data.Nat + hiding (+-assoc ; +-comm ; ·-comm) + renaming (_·_ to _·ℕ_; _+_ to _+ℕ_ ; ·-assoc to ·ℕ-assoc) +open import Cubical.Data.Bool +open import Cubical.Data.Sum +open import Cubical.Data.Int.Base + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidableEq + +sucPred : ∀ i → sucℤ (predℤ i) ≡ i +sucPred (pos zero) = refl +sucPred (pos (suc n)) = refl +sucPred (negsuc n) = refl + +predSuc : ∀ i → predℤ (sucℤ i) ≡ i +predSuc (pos n) = refl +predSuc (negsuc zero) = refl +predSuc (negsuc (suc n)) = refl + +injPos : ∀ {a b : ℕ} → pos a ≡ pos b → a ≡ b +injPos {a} h = subst T h refl + where + T : ℤ → Type₀ + T (pos x) = a ≡ x + T (negsuc _) = ⊥ + +injNegsuc : ∀ {a b : ℕ} → negsuc a ≡ negsuc b → a ≡ b +injNegsuc {a} h = subst T h refl + where + T : ℤ → Type₀ + T (pos _) = ⊥ + T (negsuc x) = a ≡ x + +posNotnegsuc : ∀ (a b : ℕ) → ¬ (pos a ≡ negsuc b) +posNotnegsuc a b h = subst T h 0 + where + T : ℤ → Type₀ + T (pos _) = ℕ + T (negsuc _) = ⊥ + +negsucNotpos : ∀ (a b : ℕ) → ¬ (negsuc a ≡ pos b) +negsucNotpos a b h = subst T h 0 + where + T : ℤ → Type₀ + T (pos _) = ⊥ + T (negsuc _) = ℕ + +discreteℤ : Discrete ℤ +discreteℤ (pos n) (pos m) with discreteℕ n m +... | yes p = yes (cong pos p) +... | no p = no (λ x → p (injPos x)) +discreteℤ (pos n) (negsuc m) = no (posNotnegsuc n m) +discreteℤ (negsuc n) (pos m) = no (negsucNotpos n m) +discreteℤ (negsuc n) (negsuc m) with discreteℕ n m +... | yes p = yes (cong negsuc p) +... | no p = no (λ x → p (injNegsuc x)) + +isSetℤ : isSet ℤ +isSetℤ = Discrete→isSet discreteℤ + +-pos : ∀ n → - (pos n) ≡ neg n +-pos zero = refl +-pos (suc n) = refl + +-neg : ∀ n → - (neg n) ≡ pos n +-neg zero = refl +-neg (suc n) = refl + +-Involutive : ∀ z → - (- z) ≡ z +-Involutive (pos n) = cong -_ (-pos n) ∙ -neg n +-Involutive (negsuc n) = refl + +isEquiv- : isEquiv (-_) +isEquiv- = isoToIsEquiv (iso -_ -_ -Involutive -Involutive) + +sucℤ+pos : ∀ n m → sucℤ (m +pos n) ≡ (sucℤ m) +pos n +sucℤ+pos zero m = refl +sucℤ+pos (suc n) m = cong sucℤ (sucℤ+pos n m) + +predℤ+negsuc : ∀ n m → predℤ (m +negsuc n) ≡ (predℤ m) +negsuc n +predℤ+negsuc zero m = refl +predℤ+negsuc (suc n) m = cong predℤ (predℤ+negsuc n m) + +sucℤ+negsuc : ∀ n m → sucℤ (m +negsuc n) ≡ (sucℤ m) +negsuc n +sucℤ+negsuc zero m = (sucPred _) ∙ (sym (predSuc _)) +sucℤ+negsuc (suc n) m = _ ≡⟨ sucPred _ ⟩ + m +negsuc n ≡[ i ]⟨ predSuc m (~ i) +negsuc n ⟩ + (predℤ (sucℤ m)) +negsuc n ≡⟨ sym (predℤ+negsuc n (sucℤ m)) ⟩ + predℤ (sucℤ m +negsuc n) ∎ + +predℤ+pos : ∀ n m → predℤ (m +pos n) ≡ (predℤ m) +pos n +predℤ+pos zero m = refl +predℤ+pos (suc n) m = _ ≡⟨ predSuc _ ⟩ + m +pos n ≡[ i ]⟨ sucPred m (~ i) + pos n ⟩ + (sucℤ (predℤ m)) +pos n ≡⟨ sym (sucℤ+pos n (predℤ m))⟩ + (predℤ m) +pos (suc n) ∎ + +predℤ-pos : ∀ n → predℤ(- (pos n)) ≡ negsuc n +predℤ-pos zero = refl +predℤ-pos (suc n) = refl + +predℤ+ : ∀ m n → predℤ (m + n) ≡ (predℤ m) + n +predℤ+ m (pos n) = predℤ+pos n m +predℤ+ m (negsuc n) = predℤ+negsuc n m + ++predℤ : ∀ m n → predℤ (m + n) ≡ m + (predℤ n) ++predℤ m (pos zero) = refl ++predℤ m (pos (suc n)) = (predSuc (m + pos n)) ∙ (cong (_+_ m) (sym (predSuc (pos n)))) ++predℤ m (negsuc n) = refl + +sucℤ+ : ∀ m n → sucℤ (m + n) ≡ (sucℤ m) + n +sucℤ+ m (pos n) = sucℤ+pos n m +sucℤ+ m (negsuc n) = sucℤ+negsuc n m + ++sucℤ : ∀ m n → sucℤ (m + n) ≡ m + (sucℤ n) ++sucℤ m (pos n) = refl ++sucℤ m (negsuc zero) = sucPred _ ++sucℤ m (negsuc (suc n)) = (sucPred (m +negsuc n)) ∙ (cong (_+_ m) (sym (sucPred (negsuc n)))) + +pos0+ : ∀ z → z ≡ pos 0 + z +pos0+ (pos zero) = refl +pos0+ (pos (suc n)) = cong sucℤ (pos0+ (pos n)) +pos0+ (negsuc zero) = refl +pos0+ (negsuc (suc n)) = cong predℤ (pos0+ (negsuc n)) + +negsuc0+ : ∀ z → predℤ z ≡ negsuc 0 + z +negsuc0+ (pos zero) = refl +negsuc0+ (pos (suc n)) = (sym (sucPred (pos n))) ∙ (cong sucℤ (negsuc0+ _)) +negsuc0+ (negsuc zero) = refl +negsuc0+ (negsuc (suc n)) = cong predℤ (negsuc0+ (negsuc n)) + +ind-comm : {A : Type₀} (_∙_ : A → A → A) (f : ℕ → A) (g : A → A) + (p : ∀ {n} → f (suc n) ≡ g (f n)) + (g∙ : ∀ a b → g (a ∙ b) ≡ g a ∙ b) + (∙g : ∀ a b → g (a ∙ b) ≡ a ∙ g b) + (base : ∀ z → z ∙ f 0 ≡ f 0 ∙ z) + → ∀ z n → z ∙ f n ≡ f n ∙ z +ind-comm _∙_ f g p g∙ ∙g base z 0 = base z +ind-comm _∙_ f g p g∙ ∙g base z (suc n) = + z ∙ f (suc n) ≡[ i ]⟨ z ∙ p {n} i ⟩ + z ∙ g (f n) ≡⟨ sym ( ∙g z (f n)) ⟩ + g (z ∙ f n) ≡⟨ cong g IH ⟩ + g (f n ∙ z) ≡⟨ g∙ (f n) z ⟩ + g (f n) ∙ z ≡[ i ]⟨ p {n} (~ i) ∙ z ⟩ + f (suc n) ∙ z ∎ + where + IH = ind-comm _∙_ f g p g∙ ∙g base z n + +ind-assoc : {A : Type₀} (_·_ : A → A → A) (f : ℕ → A) + (g : A → A) (p : ∀ a b → g (a · b) ≡ a · (g b)) + (q : ∀ {c} → f (suc c) ≡ g (f c)) + (base : ∀ m n → (m · n) · (f 0) ≡ m · (n · (f 0))) + (m n : A) (o : ℕ) + → m · (n · (f o)) ≡ (m · n) · (f o) +ind-assoc _·_ f g p q base m n 0 = sym (base m n) +ind-assoc _·_ f g p q base m n (suc o) = + m · (n · (f (suc o))) ≡[ i ]⟨ m · (n · q {o} i) ⟩ + m · (n · (g (f o))) ≡[ i ]⟨ m · (p n (f o) (~ i)) ⟩ + m · (g (n · (f o))) ≡⟨ sym (p m (n · (f o)))⟩ + g (m · (n · (f o))) ≡⟨ cong g IH ⟩ + g ((m · n) · (f o)) ≡⟨ p (m · n) (f o) ⟩ + (m · n) · (g (f o)) ≡[ i ]⟨ (m · n) · q {o} (~ i) ⟩ + (m · n) · (f (suc o)) ∎ + where + IH = ind-assoc _·_ f g p q base m n o + ++Comm : ∀ m n → m + n ≡ n + m ++Comm m (pos n) = ind-comm _+_ pos sucℤ refl sucℤ+ +sucℤ pos0+ m n ++Comm m (negsuc n) = ind-comm _+_ negsuc predℤ refl predℤ+ +predℤ negsuc0+ m n + ++Assoc : ∀ m n o → m + (n + o) ≡ (m + n) + o ++Assoc m n (pos o) = ind-assoc _+_ pos sucℤ +sucℤ refl (λ _ _ → refl) m n o ++Assoc m n (negsuc o) = ind-assoc _+_ negsuc predℤ +predℤ refl +predℤ m n o + +-- Compose sucPathℤ with itself n times. Transporting along this +-- will be addition, transporting with it backwards will be subtraction. +-- Use this to define _+'_ for which is easier to prove +-- isEquiv (λ n → n +' m) since _+'_ is defined by transport + +sucPathℤ : ℤ ≡ ℤ +sucPathℤ = ua (sucℤ , isoToIsEquiv (iso sucℤ predℤ sucPred predSuc)) + +addEq : ℕ → ℤ ≡ ℤ +addEq zero = refl +addEq (suc n) = (addEq n) ∙ sucPathℤ + +predPathℤ : ℤ ≡ ℤ +predPathℤ = ua (predℤ , isoToIsEquiv (iso predℤ sucℤ predSuc sucPred)) + +subEq : ℕ → ℤ ≡ ℤ +subEq zero = refl +subEq (suc n) = (subEq n) ∙ predPathℤ + +_+'_ : ℤ → ℤ → ℤ +m +' pos n = transport (addEq n) m +m +' negsuc n = transport (subEq (suc n)) m + ++'≡+ : _+'_ ≡ _+_ ++'≡+ i m (pos zero) = m ++'≡+ i m (pos (suc n)) = sucℤ (+'≡+ i m (pos n)) ++'≡+ i m (negsuc zero) = predℤ m ++'≡+ i m (negsuc (suc n)) = predℤ (+'≡+ i m (negsuc n)) -- + -- compPath (λ i → (+'≡+ i (predℤ m) (negsuc n))) (sym (predℤ+negsuc n m)) i + +isEquivAddℤ' : (m : ℤ) → isEquiv (λ n → n +' m) +isEquivAddℤ' (pos n) = isEquivTransport (addEq n) +isEquivAddℤ' (negsuc n) = isEquivTransport (subEq (suc n)) + +isEquivAddℤ : (m : ℤ) → isEquiv (λ n → n + m) +isEquivAddℤ = subst (λ add → (m : ℤ) → isEquiv (λ n → add n m)) +'≡+ isEquivAddℤ' + +-- below is an alternate proof of isEquivAddℤ for comparison +-- We also have two useful lemma here. + +minusPlus : ∀ m n → (n - m) + m ≡ n +minusPlus (pos zero) n = refl +minusPlus (pos 1) = sucPred +minusPlus (pos (suc (suc m))) n = + sucℤ ((n +negsuc (suc m)) +pos (suc m)) ≡⟨ sucℤ+pos (suc m) _ ⟩ + sucℤ (n +negsuc (suc m)) +pos (suc m) ≡[ i ]⟨ sucPred (n +negsuc m) i +pos (suc m) ⟩ + (n - pos (suc m)) +pos (suc m) ≡⟨ minusPlus (pos (suc m)) n ⟩ + n ∎ +minusPlus (negsuc zero) = predSuc +minusPlus (negsuc (suc m)) n = + predℤ (sucℤ (sucℤ (n +pos m)) +negsuc m) ≡⟨ predℤ+negsuc m _ ⟩ + predℤ (sucℤ (sucℤ (n +pos m))) +negsuc m ≡[ i ]⟨ predSuc (sucℤ (n +pos m)) i +negsuc m ⟩ + sucℤ (n +pos m) +negsuc m ≡⟨ minusPlus (negsuc m) n ⟩ + n ∎ + +plusMinus : ∀ m n → (n + m) - m ≡ n +plusMinus (pos zero) n = refl +plusMinus (pos (suc m)) = minusPlus (negsuc m) +plusMinus (negsuc m) = minusPlus (pos (suc m)) + +private + alternateProof : (m : ℤ) → isEquiv (λ n → n + m) + alternateProof m = isoToIsEquiv (iso (λ n → n + m) + (λ n → n - m) + (minusPlus m) + (plusMinus m)) + +-Cancel : ∀ z → z - z ≡ 0 +-Cancel z = cong (_- z) (pos0+ z) ∙ plusMinus z (pos zero) + +pos+ : ∀ m n → pos (m +ℕ n) ≡ pos m + pos n +pos+ zero zero = refl +pos+ zero (suc n) = + pos (zero +ℕ suc n) ≡⟨ +Comm (pos (suc n)) (pos zero) ⟩ + pos zero + pos (suc n) ∎ +pos+ (suc m) zero = + pos (suc (m +ℕ zero)) ≡⟨ cong pos (cong suc (+-zero m)) ⟩ + pos (suc m) + pos zero ∎ +pos+ (suc m) (suc n) = + pos (suc m +ℕ suc n) ≡⟨ cong pos (cong suc (+-suc m n)) ⟩ + sucℤ (pos (suc (m +ℕ n))) ≡⟨ cong sucℤ (cong sucℤ (pos+ m n)) ⟩ + sucℤ (sucℤ (pos m + pos n)) ≡⟨ sucℤ+ (pos m) (sucℤ (pos n)) ⟩ + pos (suc m) + pos (suc n) ∎ + +negsuc+ : ∀ m n → negsuc (m +ℕ n) ≡ negsuc m - pos n +negsuc+ zero zero = refl +negsuc+ zero (suc n) = + negsuc (zero +ℕ suc n) ≡⟨ negsuc0+ (negsuc n) ⟩ + negsuc zero + negsuc n ≡⟨ cong (negsuc zero +_) (-pos (suc n)) ⟩ + negsuc zero - pos (suc n) ∎ +negsuc+ (suc m) zero = + negsuc (suc m +ℕ zero) ≡⟨ cong negsuc (cong suc (+-zero m)) ⟩ + negsuc (suc m) - pos zero ∎ +negsuc+ (suc m) (suc n) = + negsuc (suc m +ℕ suc n) ≡⟨ cong negsuc (sym (+-suc m (suc n))) ⟩ + negsuc (m +ℕ suc (suc n)) ≡⟨ negsuc+ m (suc (suc n)) ⟩ + negsuc m - pos (suc (suc n)) ≡⟨ sym (+predℤ (negsuc m) (negsuc n)) ⟩ + predℤ (negsuc m + negsuc n ) ≡⟨ predℤ+ (negsuc m) (negsuc n) ⟩ + negsuc (suc m) - pos (suc n) ∎ + +neg+ : ∀ m n → neg (m +ℕ n) ≡ neg m + neg n +neg+ zero zero = refl +neg+ zero (suc n) = neg (zero +ℕ suc n) ≡⟨ +Comm (neg (suc n)) (pos zero) ⟩ + neg zero + neg (suc n) ∎ +neg+ (suc m) zero = neg (suc (m +ℕ zero)) ≡⟨ cong neg (cong suc (+-zero m)) ⟩ + neg (suc m) + neg zero ∎ +neg+ (suc m) (suc n) = neg (suc m +ℕ suc n) ≡⟨ negsuc+ m (suc n) ⟩ + neg (suc m) + neg (suc n) ∎ + +ℕ-AntiComm : ∀ m n → m ℕ- n ≡ - (n ℕ- m) +ℕ-AntiComm zero zero = refl +ℕ-AntiComm zero (suc n) = refl +ℕ-AntiComm (suc m) zero = refl +ℕ-AntiComm (suc m) (suc n) = suc m ℕ- suc n ≡⟨ ℕ-AntiComm m n ⟩ + - (suc n ℕ- suc m) ∎ + +pos- : ∀ m n → m ℕ- n ≡ pos m - pos n +pos- zero zero = refl +pos- zero (suc n) = zero ℕ- suc n ≡⟨ +Comm (negsuc n) (pos zero) ⟩ + pos zero - pos (suc n) ∎ +pos- (suc m) zero = refl +pos- (suc m) (suc n) = + suc m ℕ- suc n ≡⟨ pos- m n ⟩ + pos m - pos n ≡⟨ sym (sucPred (pos m - pos n)) ⟩ + sucℤ (predℤ (pos m - pos n)) ≡⟨ cong sucℤ (+predℤ (pos m) (- pos n)) ⟩ + sucℤ (pos m + predℤ (- (pos n))) ≡⟨ cong sucℤ (cong (pos m +_) (predℤ-pos n)) ⟩ + sucℤ (pos m + negsuc n) ≡⟨ sucℤ+negsuc n (pos m) ⟩ + pos (suc m) - pos (suc n) ∎ + +-AntiComm : ∀ m n → m - n ≡ - (n - m) +-AntiComm (pos n) (pos m) = pos n - pos m ≡⟨ sym (pos- n m) ⟩ + n ℕ- m ≡⟨ ℕ-AntiComm n m ⟩ + - (m ℕ- n) ≡⟨ cong -_ (pos- m n) ⟩ + - (pos m - pos n) ∎ +-AntiComm (pos n) (negsuc m) = + pos n - negsuc m ≡⟨ +Comm (pos n) (pos (suc m)) ⟩ + pos (suc m) + pos n ≡⟨ sym (pos+ (suc m) n) ⟩ + pos (suc m +ℕ n) ≡⟨ sym (-neg (suc m +ℕ n)) ⟩ + - neg (suc m +ℕ n) ≡⟨ cong -_ (neg+ (suc m) n) ⟩ + - (neg (suc m) + neg n) ≡⟨ cong -_ (cong (negsuc m +_) (sym (-pos n))) ⟩ + - (negsuc m - pos n) ∎ +-AntiComm (negsuc n) (pos m) = + negsuc n - pos m ≡⟨ sym (negsuc+ n m) ⟩ + negsuc (n +ℕ m) ≡⟨ cong -_ (pos+ (suc n) m) ⟩ + - (pos (suc n) + pos m) ≡⟨ cong -_ (+Comm (pos (suc n)) (pos m)) ⟩ + - (pos m - negsuc n) ∎ +-AntiComm (negsuc n) (negsuc m) = + negsuc n - negsuc m ≡⟨ +Comm (negsuc n) (pos (suc m)) ⟩ + pos (suc m) + negsuc n ≡⟨ sym (pos- (suc m) (suc n)) ⟩ + suc m ℕ- suc n ≡⟨ ℕ-AntiComm (suc m) (suc n) ⟩ + - (suc n ℕ- suc m) ≡⟨ cong -_ (pos- (suc n) (suc m)) ⟩ + - (pos (suc n) - pos (suc m)) ≡⟨ cong -_ (+Comm (pos (suc n)) (negsuc m)) ⟩ + - (negsuc m - negsuc n) ∎ + +-Dist+ : ∀ m n → - (m + n) ≡ (- m) + (- n) +-Dist+ (pos n) (pos m) = + - (pos n + pos m) ≡⟨ cong -_ (sym (pos+ n m)) ⟩ + - pos (n +ℕ m) ≡⟨ -pos (n +ℕ m) ⟩ + neg (n +ℕ m) ≡⟨ neg+ n m ⟩ + neg n + neg m ≡⟨ cong (neg n +_) (sym (-pos m)) ⟩ + neg n + (- pos m) ≡⟨ cong (_+ (- pos m)) (sym (-pos n)) ⟩ + (- pos n) + (- pos m) ∎ +-Dist+ (pos n) (negsuc m) = + - (pos n + negsuc m) ≡⟨ sym (-AntiComm (pos (suc m)) (pos n)) ⟩ + pos (suc m) - pos n ≡⟨ +Comm (pos (suc m)) (- pos n) ⟩ + (- pos n) + (- negsuc m) ∎ +-Dist+ (negsuc n) (pos m) = + - (negsuc n + pos m) ≡⟨ cong -_ (+Comm (negsuc n) (pos m)) ⟩ + - (pos m + negsuc n) ≡⟨ sym (-AntiComm (- negsuc n) (pos m)) ⟩ + (- negsuc n) + (- pos m) ∎ +-Dist+ (negsuc n) (negsuc m) = + - (negsuc n + negsuc m) ≡⟨ cong -_ (sym (neg+ (suc n) (suc m))) ⟩ + - neg (suc n +ℕ suc m) ≡⟨ pos+ (suc n) (suc m) ⟩ + (- negsuc n) + (- negsuc m) ∎ + + +-- multiplication + +pos·negsuc : (n m : ℕ) → pos n · negsuc m ≡ - (pos n · pos (suc m)) +pos·negsuc zero m = refl +pos·negsuc (suc n) m = + (λ i → negsuc m + (pos·negsuc n m i)) + ∙ sym (-Dist+ (pos (suc m)) (pos n · pos (suc m))) + +negsuc·pos : (n m : ℕ) → negsuc n · pos m ≡ - (pos (suc n) · pos m) +negsuc·pos zero m = refl +negsuc·pos (suc n) m = + cong ((- pos m) +_) (negsuc·pos n m) + ∙ sym (-Dist+ (pos m) (pos m + (pos n · pos m))) + +negsuc·negsuc : (n m : ℕ) → negsuc n · negsuc m ≡ pos (suc n) · pos (suc m) +negsuc·negsuc zero m = refl +negsuc·negsuc (suc n) m = cong (pos (suc m) +_) (negsuc·negsuc n m) + +·Comm : (x y : ℤ) → x · y ≡ y · x +·Comm (pos n) (pos m) = lem n m + where + lem : (n m : ℕ) → (pos n · pos m) ≡ (pos m · pos n) + lem zero zero = refl + lem zero (suc m) i = +Comm (lem zero m i) (pos zero) i + lem (suc n) zero i = +Comm (pos zero) (lem n zero i) i + lem (suc n) (suc m) = + (λ i → pos (suc m) + (lem n (suc m) i)) + ∙∙ +Assoc (pos (suc m)) (pos n) (pos m · pos n) + ∙∙ (λ i → sucℤ+ (pos m) (pos n) (~ i) + (pos m · pos n)) + ∙∙ (λ i → +Comm (pos m) (pos (suc n)) i + (pos m · pos n)) + ∙∙ sym (+Assoc (pos (suc n)) (pos m) (pos m · pos n)) + ∙∙ (λ i → pos (suc n) + (pos m + (lem n m (~ i)))) + ∙∙ λ i → pos (suc n) + (lem (suc n) m i) +·Comm (pos n) (negsuc m) = + pos·negsuc n m + ∙∙ cong -_ (·Comm (pos n) (pos (suc m))) + ∙∙ sym (negsuc·pos m n) +·Comm (negsuc n) (pos m) = + sym (pos·negsuc m n + ∙∙ cong -_ (·Comm (pos m) (pos (suc n))) + ∙∙ sym (negsuc·pos n m)) +·Comm (negsuc n) (negsuc m) = + negsuc·negsuc n m ∙∙ ·Comm (pos (suc n)) (pos (suc m)) ∙∙ sym (negsuc·negsuc m n) + +·Rid : (x : ℤ) → x · 1 ≡ x +·Rid x = ·Comm x 1 + +private + distrHelper : (x y z w : ℤ) → (x + y) + (z + w) ≡ ((x + z) + (y + w)) + distrHelper x y z w = + +Assoc (x + y) z w + ∙∙ cong (_+ w) (sym (+Assoc x y z) ∙∙ cong (x +_) (+Comm y z) ∙∙ +Assoc x z y) + ∙∙ sym (+Assoc (x + z) y w) + +·DistR+ : (x y z : ℤ) → x · (y + z) ≡ x · y + x · z +·DistR+ (pos zero) y z = refl +·DistR+ (pos (suc n)) y z = + cong ((y + z) +_) (·DistR+ (pos n) y z) + ∙ distrHelper y z (pos n · y) (pos n · z) +·DistR+ (negsuc zero) y z = -Dist+ y z +·DistR+ (negsuc (suc n)) y z = + cong₂ _+_ (-Dist+ y z) (·DistR+ (negsuc n) y z) + ∙ distrHelper (- y) (- z) (negsuc n · y) (negsuc n · z) + +·DistL+ : (x y z : ℤ) → (x + y) · z ≡ x · z + y · z +·DistL+ x y z = ·Comm (x + y) z ∙∙ ·DistR+ z x y ∙∙ cong₂ _+_ (·Comm z x) (·Comm z y) + +-DistL· : (b c : ℤ) → - (b · c) ≡ - b · c +-DistL· (pos zero) c = refl +-DistL· (pos (suc n)) (pos m) = sym (negsuc·pos n m) +-DistL· (pos (suc zero)) (negsuc m) = + -Dist+ (negsuc m) (pos zero · negsuc m) + ∙ cong (pos (suc m) +_) (-DistL· (pos zero) (negsuc m)) +-DistL· (pos (suc (suc n))) (negsuc m) = + -Dist+ (negsuc m) (pos (suc n) · negsuc m) + ∙ cong (pos (suc m) +_) (-DistL· (pos (suc n)) (negsuc m)) +-DistL· (negsuc zero) c = -Involutive c +-DistL· (negsuc (suc n)) c = + -Dist+ (- c) (negsuc n · c) + ∙∙ cong (_+ (- (negsuc n · c))) (-Involutive c) + ∙∙ cong (c +_) (-DistL· (negsuc n) c) + +-DistR· : (b c : ℤ) → - (b · c) ≡ b · - c +-DistR· b c = cong (-_) (·Comm b c) ∙∙ -DistL· c b ∙∙ ·Comm (- c) b + +ℤ·negsuc : (n : ℤ) (m : ℕ) → n · negsuc m ≡ - (n · pos (suc m)) +ℤ·negsuc (pos n) m = pos·negsuc n m +ℤ·negsuc (negsuc n) m = negsuc·negsuc n m ∙ sym (-DistL· (negsuc n) (pos (suc m))) + +·Assoc : (a b c : ℤ) → (a · (b · c)) ≡ ((a · b) · c) +·Assoc (pos zero) b c = refl +·Assoc (pos (suc n)) b c = + cong ((b · c) +_) (·Assoc (pos n) b c) + ∙∙ cong₂ _+_ (·Comm b c) (·Comm (pos n · b) c) + ∙∙ sym (·DistR+ c b (pos n · b)) + ∙ sym (·Comm _ c) +·Assoc (negsuc zero) = -DistL· +·Assoc (negsuc (suc n)) b c = + cong ((- (b · c)) +_) (·Assoc (negsuc n) b c) + ∙∙ cong (_+ ((negsuc n · b) · c)) (-DistL· b c) + ∙∙ sym (·DistL+ (- b) (negsuc n · b) c) + +minus≡0- : (x : ℤ) → - x ≡ (0 - x) +minus≡0- x = +Comm (- x) 0 + +-- Absolute values +abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) +abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) + where + help : (x : ℤ) → (x ≡ pos (abs x)) ⊎ (x ≡ - pos (abs x)) + help (pos n) = inl refl + help (negsuc n) = inr refl + +⊎→abs : (x : ℤ) (n : ℕ) → (x ≡ pos n) ⊎ (x ≡ - pos n) → abs x ≡ n +⊎→abs (pos n₁) n (inl x₁) = cong abs x₁ +⊎→abs (negsuc n₁) n (inl x₁) = cong abs x₁ +⊎→abs x zero (inr x₁) = cong abs x₁ +⊎→abs x (suc n) (inr x₁) = cong abs x₁ + +abs- : (x : ℤ) → abs (- x) ≡ abs x +abs- (pos zero) = refl +abs- (pos (suc n)) = refl +abs- (negsuc n) = refl diff --git a/Cubical/Data/List.agda b/Cubical/Data/List.agda new file mode 100644 index 0000000000..ec82b6a1b7 --- /dev/null +++ b/Cubical/Data/List.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.List where + +open import Cubical.Data.List.Base public +open import Cubical.Data.List.Properties public diff --git a/Cubical/Data/List/Base.agda b/Cubical/Data/List/Base.agda new file mode 100644 index 0000000000..0222265954 --- /dev/null +++ b/Cubical/Data/List/Base.agda @@ -0,0 +1,52 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.List.Base where + +open import Agda.Builtin.List public +open import Cubical.Core.Everything +open import Cubical.Data.Maybe.Base as Maybe +open import Cubical.Data.Nat.Base + +module _ {ℓ} {A : Type ℓ} where + + infixr 5 _++_ + infixl 5 _∷ʳ_ + + [_] : A → List A + [ a ] = a ∷ [] + + _++_ : List A → List A → List A + [] ++ ys = ys + (x ∷ xs) ++ ys = x ∷ xs ++ ys + + rev : List A → List A + rev [] = [] + rev (x ∷ xs) = rev xs ++ [ x ] + + _∷ʳ_ : List A → A → List A + xs ∷ʳ x = xs ++ x ∷ [] + + length : List A → ℕ + length [] = 0 + length (x ∷ l) = 1 + length l + + map : ∀ {ℓ'} {B : Type ℓ'} → (A → B) → List A → List B + map f [] = [] + map f (x ∷ xs) = f x ∷ map f xs + + map2 : ∀ {ℓ' ℓ''} {B : Type ℓ'} {C : Type ℓ''} + → (A → B → C) → List A → List B → List C + map2 f [] _ = [] + map2 f _ [] = [] + map2 f (x ∷ xs) (y ∷ ys) = f x y ∷ map2 f xs ys + + filterMap : ∀ {ℓ'} {B : Type ℓ'} → (A → Maybe B) → List A → List B + filterMap f [] = [] + filterMap f (x ∷ xs) = Maybe.rec (filterMap f xs) (_∷ filterMap f xs) (f x) + + foldr : ∀ {ℓ'} {B : Type ℓ'} → (A → B → B) → B → List A → B + foldr f b [] = b + foldr f b (x ∷ xs) = f x (foldr f b xs) + + foldl : ∀ {ℓ'} {B : Type ℓ'} → (B → A → B) → B → List A → B + foldl f b [] = b + foldl f b (x ∷ xs) = foldl f (f b x) xs diff --git a/Cubical/Data/List/Properties.agda b/Cubical/Data/List/Properties.agda new file mode 100644 index 0000000000..b7278a38ce --- /dev/null +++ b/Cubical/Data/List/Properties.agda @@ -0,0 +1,176 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.List.Properties where + +open import Agda.Builtin.List +open import Cubical.Core.Everything +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Relation.Nullary + +open import Cubical.Data.List.Base + +module _ {ℓ} {A : Type ℓ} where + + ++-unit-r : (xs : List A) → xs ++ [] ≡ xs + ++-unit-r [] = refl + ++-unit-r (x ∷ xs) = cong (_∷_ x) (++-unit-r xs) + + ++-assoc : (xs ys zs : List A) → (xs ++ ys) ++ zs ≡ xs ++ ys ++ zs + ++-assoc [] ys zs = refl + ++-assoc (x ∷ xs) ys zs = cong (_∷_ x) (++-assoc xs ys zs) + + rev-snoc : (xs : List A) (y : A) → rev (xs ++ [ y ]) ≡ y ∷ rev xs + rev-snoc [] y = refl + rev-snoc (x ∷ xs) y = cong (_++ [ x ]) (rev-snoc xs y) + + rev-++ : (xs ys : List A) → rev (xs ++ ys) ≡ rev ys ++ rev xs + rev-++ [] ys = sym (++-unit-r (rev ys)) + rev-++ (x ∷ xs) ys = + cong (λ zs → zs ++ [ x ]) (rev-++ xs ys) + ∙ ++-assoc (rev ys) (rev xs) [ x ] + + rev-rev : (xs : List A) → rev (rev xs) ≡ xs + rev-rev [] = refl + rev-rev (x ∷ xs) = rev-snoc (rev xs) x ∙ cong (_∷_ x) (rev-rev xs) + + rev-rev-snoc : (xs : List A) (y : A) → + Square (rev-rev (xs ++ [ y ])) (cong (_++ [ y ]) (rev-rev xs)) (cong rev (rev-snoc xs y)) refl + rev-rev-snoc [] y = sym (lUnit refl) + rev-rev-snoc (x ∷ xs) y i j = + hcomp + (λ k → λ + { (i = i1) → compPath-filler (rev-snoc (rev xs) x) (cong (x ∷_) (rev-rev xs)) k j ++ [ y ] + ; (j = i0) → rev (rev-snoc xs y i ++ [ x ]) + ; (j = i1) → x ∷ rev-rev-snoc xs y i k + }) + (rev-snoc (rev-snoc xs y i) x j) + + data SnocView : List A → Type ℓ where + nil : SnocView [] + snoc : (x : A) → (xs : List A) → (sx : SnocView xs) → SnocView (xs ∷ʳ x) + + snocView : (xs : List A) → SnocView xs + snocView xs = helper nil xs + where + helper : {l : List A} -> SnocView l -> (r : List A) -> SnocView (l ++ r) + helper {l} sl [] = subst SnocView (sym (++-unit-r l)) sl + helper {l} sl (x ∷ r) = subst SnocView (++-assoc l (x ∷ []) r) (helper (snoc x l sl) r) + +-- Path space of list type +module ListPath {ℓ} {A : Type ℓ} where + + Cover : List A → List A → Type ℓ + Cover [] [] = Lift Unit + Cover [] (_ ∷ _) = Lift ⊥ + Cover (_ ∷ _) [] = Lift ⊥ + Cover (x ∷ xs) (y ∷ ys) = (x ≡ y) × Cover xs ys + + reflCode : ∀ xs → Cover xs xs + reflCode [] = lift tt + reflCode (_ ∷ xs) = refl , reflCode xs + + encode : ∀ xs ys → (p : xs ≡ ys) → Cover xs ys + encode xs _ = J (λ ys _ → Cover xs ys) (reflCode xs) + + encodeRefl : ∀ xs → encode xs xs refl ≡ reflCode xs + encodeRefl xs = JRefl (λ ys _ → Cover xs ys) (reflCode xs) + + decode : ∀ xs ys → Cover xs ys → xs ≡ ys + decode [] [] _ = refl + decode [] (_ ∷ _) (lift ()) + decode (x ∷ xs) [] (lift ()) + decode (x ∷ xs) (y ∷ ys) (p , c) = cong₂ _∷_ p (decode xs ys c) + + decodeRefl : ∀ xs → decode xs xs (reflCode xs) ≡ refl + decodeRefl [] = refl + decodeRefl (x ∷ xs) = cong (cong₂ _∷_ refl) (decodeRefl xs) + + decodeEncode : ∀ xs ys → (p : xs ≡ ys) → decode xs ys (encode xs ys p) ≡ p + decodeEncode xs _ = + J (λ ys p → decode xs ys (encode xs ys p) ≡ p) + (cong (decode xs xs) (encodeRefl xs) ∙ decodeRefl xs) + + isOfHLevelCover : (n : HLevel) (p : isOfHLevel (suc (suc n)) A) + (xs ys : List A) → isOfHLevel (suc n) (Cover xs ys) + isOfHLevelCover n p [] [] = + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isPropUnit) + isOfHLevelCover n p [] (y ∷ ys) = + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p (x ∷ xs) [] = + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p (x ∷ xs) (y ∷ ys) = + isOfHLevelΣ (suc n) (p x y) (\ _ → isOfHLevelCover n p xs ys) + +isOfHLevelList : ∀ {ℓ} (n : HLevel) {A : Type ℓ} + → isOfHLevel (suc (suc n)) A → isOfHLevel (suc (suc n)) (List A) +isOfHLevelList n ofLevel xs ys = + isOfHLevelRetract (suc n) + (ListPath.encode xs ys) + (ListPath.decode xs ys) + (ListPath.decodeEncode xs ys) + (ListPath.isOfHLevelCover n ofLevel xs ys) + +private + variable + ℓ : Level + A : Type ℓ + + caseList : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (n c : B) → List A → B + caseList n _ [] = n + caseList _ c (_ ∷ _) = c + + safe-head : A → List A → A + safe-head x [] = x + safe-head _ (x ∷ _) = x + + safe-tail : List A → List A + safe-tail [] = [] + safe-tail (_ ∷ xs) = xs + +cons-inj₁ : ∀ {x y : A} {xs ys} → x ∷ xs ≡ y ∷ ys → x ≡ y +cons-inj₁ {x = x} p = cong (safe-head x) p + +cons-inj₂ : ∀ {x y : A} {xs ys} → x ∷ xs ≡ y ∷ ys → xs ≡ ys +cons-inj₂ = cong safe-tail + +¬cons≡nil : ∀ {x : A} {xs} → ¬ (x ∷ xs ≡ []) +¬cons≡nil {A = A} p = lower (subst (caseList (Lift ⊥) (List A)) p []) + +¬nil≡cons : ∀ {x : A} {xs} → ¬ ([] ≡ x ∷ xs) +¬nil≡cons {A = A} p = lower (subst (caseList (List A) (Lift ⊥)) p []) + +¬snoc≡nil : ∀ {x : A} {xs} → ¬ (xs ∷ʳ x ≡ []) +¬snoc≡nil {xs = []} contra = ¬cons≡nil contra +¬snoc≡nil {xs = x ∷ xs} contra = ¬cons≡nil contra + +¬nil≡snoc : ∀ {x : A} {xs} → ¬ ([] ≡ xs ∷ʳ x) +¬nil≡snoc contra = ¬snoc≡nil (sym contra) + +cons≡rev-snoc : (x : A) → (xs : List A) → x ∷ rev xs ≡ rev (xs ∷ʳ x) +cons≡rev-snoc _ [] = refl +cons≡rev-snoc x (y ∷ ys) = λ i → cons≡rev-snoc x ys i ++ y ∷ [] + +isContr[]≡[] : isContr (Path (List A) [] []) +isContr[]≡[] = refl , ListPath.decodeEncode [] [] + +isPropXs≡[] : {xs : List A} → isProp (xs ≡ []) +isPropXs≡[] {xs = []} = isOfHLevelSuc 0 isContr[]≡[] +isPropXs≡[] {xs = x ∷ xs} = λ p _ → ⊥.rec (¬cons≡nil p) + +discreteList : Discrete A → Discrete (List A) +discreteList eqA [] [] = yes refl +discreteList eqA [] (y ∷ ys) = no ¬nil≡cons +discreteList eqA (x ∷ xs) [] = no ¬cons≡nil +discreteList eqA (x ∷ xs) (y ∷ ys) with eqA x y | discreteList eqA xs ys +... | yes p | yes q = yes (λ i → p i ∷ q i) +... | yes _ | no ¬q = no (λ p → ¬q (cons-inj₂ p)) +... | no ¬p | _ = no (λ q → ¬p (cons-inj₁ q)) + +foldrCons : (xs : List A) → foldr _∷_ [] xs ≡ xs +foldrCons [] = refl +foldrCons (x ∷ xs) = cong (x ∷_) (foldrCons xs) diff --git a/Cubical/Data/Maybe.agda b/Cubical/Data/Maybe.agda new file mode 100644 index 0000000000..bd40e1d2cc --- /dev/null +++ b/Cubical/Data/Maybe.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Maybe where + +open import Cubical.Data.Maybe.Base public +open import Cubical.Data.Maybe.Properties public diff --git a/Cubical/Data/Maybe/Base.agda b/Cubical/Data/Maybe/Base.agda new file mode 100644 index 0000000000..0f230b6233 --- /dev/null +++ b/Cubical/Data/Maybe/Base.agda @@ -0,0 +1,25 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Maybe.Base where + +open import Cubical.Core.Everything + +private + variable + ℓ : Level + A B : Type ℓ + +data Maybe (A : Type ℓ) : Type ℓ where + nothing : Maybe A + just : A → Maybe A + +caseMaybe : (n j : B) → Maybe A → B +caseMaybe n _ nothing = n +caseMaybe _ j (just _) = j + +map-Maybe : (A → B) → Maybe A → Maybe B +map-Maybe _ nothing = nothing +map-Maybe f (just x) = just (f x) + +rec : B → (A → B) → Maybe A → B +rec n j nothing = n +rec n j (just a) = j a diff --git a/Cubical/Data/Maybe/Properties.agda b/Cubical/Data/Maybe/Properties.agda new file mode 100644 index 0000000000..23da4995c2 --- /dev/null +++ b/Cubical/Data/Maybe/Properties.agda @@ -0,0 +1,158 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Maybe.Properties where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Functions.Embedding +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Relation.Nullary +open import Cubical.Data.Sum + +open import Cubical.Data.Maybe.Base + +map-Maybe-id : ∀ {ℓ} {A : Type ℓ} → ∀ m → map-Maybe (idfun A) m ≡ m +map-Maybe-id nothing = refl +map-Maybe-id (just _) = refl + +-- Path space of Maybe type +module MaybePath {ℓ} {A : Type ℓ} where + Cover : Maybe A → Maybe A → Type ℓ + Cover nothing nothing = Lift Unit + Cover nothing (just _) = Lift ⊥ + Cover (just _) nothing = Lift ⊥ + Cover (just a) (just a') = a ≡ a' + + reflCode : (c : Maybe A) → Cover c c + reflCode nothing = lift tt + reflCode (just b) = refl + + encode : ∀ c c' → c ≡ c' → Cover c c' + encode c _ = J (λ c' _ → Cover c c') (reflCode c) + + encodeRefl : ∀ c → encode c c refl ≡ reflCode c + encodeRefl c = JRefl (λ c' _ → Cover c c') (reflCode c) + + decode : ∀ c c' → Cover c c' → c ≡ c' + decode nothing nothing _ = refl + decode (just _) (just _) p = cong just p + + decodeRefl : ∀ c → decode c c (reflCode c) ≡ refl + decodeRefl nothing = refl + decodeRefl (just _) = refl + + decodeEncode : ∀ c c' → (p : c ≡ c') → decode c c' (encode c c' p) ≡ p + decodeEncode c _ = + J (λ c' p → decode c c' (encode c c' p) ≡ p) + (cong (decode c c) (encodeRefl c) ∙ decodeRefl c) + + encodeDecode : ∀ c c' → (d : Cover c c') → encode c c' (decode c c' d) ≡ d + encodeDecode nothing nothing _ = refl + encodeDecode (just a) (just a') = + J (λ a' p → encode (just a) (just a') (cong just p) ≡ p) (encodeRefl (just a)) + + Cover≃Path : ∀ c c' → Cover c c' ≃ (c ≡ c') + Cover≃Path c c' = isoToEquiv + (iso (decode c c') (encode c c') (decodeEncode c c') (encodeDecode c c')) + + Cover≡Path : ∀ c c' → Cover c c' ≡ (c ≡ c') + Cover≡Path c c' = isoToPath + (iso (decode c c') (encode c c') (decodeEncode c c') (encodeDecode c c')) + + isOfHLevelCover : (n : HLevel) + → isOfHLevel (suc (suc n)) A + → ∀ c c' → isOfHLevel (suc n) (Cover c c') + isOfHLevelCover n p nothing nothing = isOfHLevelLift (suc n) (isOfHLevelUnit (suc n)) + isOfHLevelCover n p nothing (just a') = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p (just a) nothing = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p (just a) (just a') = p a a' + +isOfHLevelMaybe : ∀ {ℓ} (n : HLevel) {A : Type ℓ} + → isOfHLevel (suc (suc n)) A + → isOfHLevel (suc (suc n)) (Maybe A) +isOfHLevelMaybe n lA c c' = + isOfHLevelRetract (suc n) + (MaybePath.encode c c') + (MaybePath.decode c c') + (MaybePath.decodeEncode c c') + (MaybePath.isOfHLevelCover n lA c c') + +private + variable + ℓ : Level + A : Type ℓ + +fromJust-def : A → Maybe A → A +fromJust-def a nothing = a +fromJust-def _ (just a) = a + +just-inj : (x y : A) → just x ≡ just y → x ≡ y +just-inj x _ eq = cong (fromJust-def x) eq + +isEmbedding-just : isEmbedding (just {A = A}) +isEmbedding-just w z = MaybePath.Cover≃Path (just w) (just z) .snd + +¬nothing≡just : ∀ {x : A} → ¬ (nothing ≡ just x) +¬nothing≡just {A = A} {x = x} p = lower (subst (caseMaybe (Maybe A) (Lift ⊥)) p (just x)) + +¬just≡nothing : ∀ {x : A} → ¬ (just x ≡ nothing) +¬just≡nothing {A = A} {x = x} p = lower (subst (caseMaybe (Lift ⊥) (Maybe A)) p (just x)) + +isProp-x≡nothing : (x : Maybe A) → isProp (x ≡ nothing) +isProp-x≡nothing nothing x w = + subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w +isProp-x≡nothing (just _) p _ = ⊥.rec (¬just≡nothing p) + +isProp-nothing≡x : (x : Maybe A) → isProp (nothing ≡ x) +isProp-nothing≡x nothing x w = + subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w +isProp-nothing≡x (just _) p _ = ⊥.rec (¬nothing≡just p) + +isContr-nothing≡nothing : isContr (nothing {A = A} ≡ nothing) +isContr-nothing≡nothing = inhProp→isContr refl (isProp-x≡nothing _) + +discreteMaybe : Discrete A → Discrete (Maybe A) +discreteMaybe eqA nothing nothing = yes refl +discreteMaybe eqA nothing (just a') = no ¬nothing≡just +discreteMaybe eqA (just a) nothing = no ¬just≡nothing +discreteMaybe eqA (just a) (just a') with eqA a a' +... | yes p = yes (cong just p) +... | no ¬p = no (λ p → ¬p (just-inj _ _ p)) + +module SumUnit where + Maybe→SumUnit : Maybe A → Unit ⊎ A + Maybe→SumUnit nothing = inl tt + Maybe→SumUnit (just a) = inr a + + SumUnit→Maybe : Unit ⊎ A → Maybe A + SumUnit→Maybe (inl _) = nothing + SumUnit→Maybe (inr a) = just a + + Maybe→SumUnit→Maybe : (x : Maybe A) → SumUnit→Maybe (Maybe→SumUnit x) ≡ x + Maybe→SumUnit→Maybe nothing = refl + Maybe→SumUnit→Maybe (just _) = refl + + SumUnit→Maybe→SumUnit : (x : Unit ⊎ A) → Maybe→SumUnit (SumUnit→Maybe x) ≡ x + SumUnit→Maybe→SumUnit (inl _) = refl + SumUnit→Maybe→SumUnit (inr _) = refl + +Maybe≡SumUnit : Maybe A ≡ Unit ⊎ A +Maybe≡SumUnit = isoToPath (iso SumUnit.Maybe→SumUnit SumUnit.SumUnit→Maybe SumUnit.SumUnit→Maybe→SumUnit SumUnit.Maybe→SumUnit→Maybe) + +congMaybeEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → A ≃ B → Maybe A ≃ Maybe B +congMaybeEquiv e = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = map-Maybe (equivFun e) + isom .inv = map-Maybe (invEq e) + isom .rightInv nothing = refl + isom .rightInv (just b) = cong just (secEq e b) + isom .leftInv nothing = refl + isom .leftInv (just a) = cong just (retEq e a) diff --git a/Cubical/Data/Nat.agda b/Cubical/Data/Nat.agda new file mode 100644 index 0000000000..b00c8ec4fa --- /dev/null +++ b/Cubical/Data/Nat.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Nat where + +open import Cubical.Data.Nat.Base public + +open import Cubical.Data.Nat.Properties public diff --git a/Cubical/Data/Nat/Algebra.agda b/Cubical/Data/Nat/Algebra.agda new file mode 100644 index 0000000000..0ee04806b4 --- /dev/null +++ b/Cubical/Data/Nat/Algebra.agda @@ -0,0 +1,299 @@ +{-# OPTIONS --no-exact-split --safe #-} + +{- + +This file shows that the property of the natural numbers being a homotopy-initial algebra of +the functor (1 + _) is equivalent to fulfilling a closely related inductive elimination principle. + +Proofing the latter is trivial, since the typechecker does the work for us. + +For details see the paper [Homotopy-initial algebras in type theory](https://arxiv.org/abs/1504.05531) +by Steve Awodey, Nicola Gambino and Kristina Sojakova. + +-} + + +module Cubical.Data.Nat.Algebra where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism + hiding (section) +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Data.Nat.Base + +private + variable + ℓ ℓ' : Level + +record NatAlgebra ℓ : Type (ℓ-suc ℓ) where + field + Carrier : Type ℓ + alg-zero : Carrier + alg-suc : Carrier → Carrier + +record NatMorphism (A : NatAlgebra ℓ) (B : NatAlgebra ℓ') : Type (ℓ-max ℓ ℓ') where + open NatAlgebra + field + morph : A .Carrier → B .Carrier + comm-zero : morph (A .alg-zero) ≡ B .alg-zero + comm-suc : morph ∘ A .alg-suc ≡ B .alg-suc ∘ morph + +record NatFiber (N : NatAlgebra ℓ') ℓ : Type (ℓ-max ℓ' (ℓ-suc ℓ)) where + open NatAlgebra N + field + Fiber : Carrier → Type ℓ + fib-zero : Fiber alg-zero + fib-suc : ∀ {n} → Fiber n → Fiber (alg-suc n) + +record NatSection {N : NatAlgebra ℓ'} (F : NatFiber N ℓ) : Type (ℓ-max ℓ' ℓ) where + open NatAlgebra N + open NatFiber F + field + section : ∀ n → Fiber n + sec-comm-zero : section alg-zero ≡ fib-zero + sec-comm-suc : ∀ n → section (alg-suc n) ≡ fib-suc (section n) + +isNatHInitial : NatAlgebra ℓ' → (ℓ : Level) → Type (ℓ-max ℓ' (ℓ-suc ℓ)) +isNatHInitial N ℓ = (M : NatAlgebra ℓ) → isContr (NatMorphism N M) + +isNatInductive : NatAlgebra ℓ' → (ℓ : Level) → Type (ℓ-max ℓ' (ℓ-suc ℓ)) +isNatInductive N ℓ = (S : NatFiber N ℓ) → NatSection S + +module AlgebraPropositionality {N : NatAlgebra ℓ'} where + open NatAlgebra N + isPropIsNatHInitial : isProp (isNatHInitial N ℓ) + isPropIsNatHInitial = isPropΠ (λ _ → isPropIsContr) + + -- under the assumption that some shape is nat-inductive, the type of sections over any fiber + -- is propositional + module SectionProp (ind : isNatInductive N ℓ) {F : NatFiber N ℓ} (S T : NatSection F) where + open NatFiber + open NatSection + + ConnectingFiber : NatFiber N ℓ + Fiber ConnectingFiber n = S .section n ≡ T .section n + fib-zero ConnectingFiber = S .sec-comm-zero ∙∙ refl ∙∙ sym (T .sec-comm-zero) + fib-suc ConnectingFiber {n} sntn = S .sec-comm-suc n ∙∙ (λ i → F .fib-suc (sntn i)) ∙∙ sym (T .sec-comm-suc n) + + open NatSection (ind ConnectingFiber) + renaming (section to α ; sec-comm-zero to ζ ; sec-comm-suc to σ) + + squeezeSquare : ∀{a}{A : Type a}{w x y z : A} (p : w ≡ x) {q : x ≡ y} (r : z ≡ y) + → (P : w ≡ z) → (sq : P ≡ p ∙∙ q ∙∙ sym r) → I → I → A + squeezeSquare p {q} r P sq i j = transport (sym (PathP≡doubleCompPathʳ p P q r)) sq i j + + S≡T : S ≡ T + section (S≡T i) n = α n i + sec-comm-zero (S≡T i) j = squeezeSquare (S .sec-comm-zero) (T .sec-comm-zero) (α alg-zero) ζ j i + sec-comm-suc (S≡T i) n j = squeezeSquare (S .sec-comm-suc n) (T .sec-comm-suc n) (α (alg-suc n)) (σ n) j i + + isPropIsNatInductive : isProp (isNatInductive N ℓ) + isPropIsNatInductive a b i F = SectionProp.S≡T a (a F) (b F) i + +module AlgebraHInd→HInit {N : NatAlgebra ℓ'} (ind : isNatInductive N ℓ) (M : NatAlgebra ℓ) where + open NatAlgebra + open NatFiber + + ConstFiberM : NatFiber N ℓ + Fiber ConstFiberM _ = M .Carrier + fib-zero ConstFiberM = M .alg-zero + fib-suc ConstFiberM = M .alg-suc + + morph→section : NatMorphism N M → NatSection ConstFiberM + morph→section x = record { section = morph ; sec-comm-zero = comm-zero ; sec-comm-suc = λ i n → comm-suc n i } + where open NatMorphism x + section→morph : NatSection ConstFiberM → NatMorphism N M + section→morph x = record { morph = section ; comm-zero = sec-comm-zero ; comm-suc = λ n i → sec-comm-suc i n } + where open NatSection x + + Morph≡Section : NatSection ConstFiberM ≡ NatMorphism N M + Morph≡Section = ua e + where + unquoteDecl e = declStrictEquiv e section→morph morph→section + + isContrMorph : isContr (NatMorphism N M) + isContrMorph = subst isContr Morph≡Section (inhProp→isContr (ind ConstFiberM) (AlgebraPropositionality.SectionProp.S≡T ind)) + +open NatAlgebra +open NatFiber +open NatSection +open NatMorphism + +module AlgebraHInit→Ind (N : NatAlgebra ℓ') ℓ (hinit : isNatHInitial N (ℓ-max ℓ' ℓ)) (F : NatFiber N (ℓ-max ℓ' ℓ)) where + + ΣAlgebra : NatAlgebra (ℓ-max ℓ' ℓ) + Carrier ΣAlgebra = Σ (N .Carrier) (F .Fiber) + alg-zero ΣAlgebra = N .alg-zero , F .fib-zero + alg-suc ΣAlgebra (n , fn) = N .alg-suc n , F .fib-suc fn + + -- the fact that we have to lift the Carrier obstructs readability a bit + -- this is the same algebra as N, but lifted into the correct universe + LiftN : NatAlgebra (ℓ-max ℓ' ℓ) + Carrier LiftN = Lift {_} {ℓ} (N .Carrier) + alg-zero LiftN = lift (N .alg-zero) + alg-suc LiftN = lift ∘ N .alg-suc ∘ lower + + _!_ : ∀ {x y} → x ≡ y → F .Fiber x → F .Fiber y + _!_ = subst (F .Fiber) + + -- from homotopy initiality of N we get + -- 1) an algebra morphism μ from N → Σ N F together with proofs of commutativity with the algebras + -- 2) projecting out the first component after μ, called α, will turn out to be the identity function + -- 3) witnesses that μ respects the definitions given in ΣAlgebra + -- a) at zero the witnesses are ζ and ζ-h + -- b) at suc the witnesses are σ and σ-h + open NatMorphism (hinit ΣAlgebra .fst) renaming (morph to μ ; comm-zero to μ-zc ; comm-suc to μ-sc) + module _ n where open Σ (μ n) public renaming (fst to α ; snd to α-h) + -- module _ i where open Σ (μ-zc i) public renaming (fst to ζ ; snd to ζ-h) + ζ : α (N .alg-zero) ≡ N .alg-zero + ζ i = μ-zc i .fst + ζ-h : PathP (λ i → F .Fiber (ζ i)) (α-h (N .alg-zero)) (F .fib-zero) + ζ-h i = μ-zc i .snd + -- module _ n i where open Σ (μ-sc i n) public renaming (fst to σ ; snd to σ-h) + σ : ∀ n → α (N .alg-suc n) ≡ N .alg-suc (α n) + σ n i = μ-sc i n .fst + σ-h : ∀ n → PathP (λ i → F .Fiber (σ n i)) (α-h (N .alg-suc n)) (F .fib-suc (α-h n)) + σ-h n i = μ-sc i n .snd + + -- liftMorph would be the identity morphism if it weren't for size issues + liftMorph : NatMorphism N LiftN + liftMorph = record { morph = lift ; comm-zero = refl ; comm-suc = refl } + -- instead of abstractly defining morphism composition and a projection algebra morphism + -- from Σ N F → N, define the composite directly. comm-zero and comm-suc thus are + -- defined without path composition + fst∘μ : NatMorphism N LiftN + morph fst∘μ = lift ∘ α + comm-zero fst∘μ i = lift (ζ i) + comm-suc fst∘μ i n = lift (σ n i) + + fst∘μ≡id : fst∘μ ≡ liftMorph + fst∘μ≡id = isContr→isProp (hinit LiftN) _ _ + + -- we get a proof that the index is preserved uniformly + P : ∀ n → α n ≡ n + P n i = lower (fst∘μ≡id i .morph n) + + -- we also have proofs that α cancels after the algebra of N + Q-zero : α (N .alg-zero) ≡ N .alg-zero + Q-zero = ζ + Q-suc : ∀ n → α (N .alg-suc n) ≡ N .alg-suc n + Q-suc n = σ n ∙ cong (N .alg-suc) (P n) + + -- but P and Q are the same up to homotopy + P-zero : P (N .alg-zero) ≡ Q-zero + P-zero i j = hcomp (λ k → λ where + (i = i0) → lower (fst∘μ≡id j .comm-zero (~ k)) + (i = i1) → ζ (j ∨ ~ k) + (j = i0) → ζ (~ k) + (j = i1) → N .alg-zero + ) (N .alg-zero) + P-suc : ∀ n → P (N .alg-suc n) ≡ Q-suc n + P-suc n i j = hcomp (λ k → λ where + (i = i0) → lower (fst∘μ≡id j .comm-suc (~ k) n) + (i = i1) → compPath-filler' (σ n) (cong (N .alg-suc) (P n)) k j + (j = i0) → σ n (~ k) + (j = i1) → N .alg-suc n + ) (N .alg-suc (P n j)) + + Fsection : NatSection F + section Fsection n = P n ! α-h n + sec-comm-zero Fsection = + P (N .alg-zero) ! α-h (N .alg-zero) + ≡[ i ]⟨ P-zero i ! α-h _ ⟩ + Q-zero ! α-h (N .alg-zero) + ≡⟨ fromPathP ζ-h ⟩ + F .fib-zero + ∎ + sec-comm-suc Fsection n = + P (N .alg-suc n) ! α-h (N .alg-suc n) + ≡[ i ]⟨ P-suc n i ! α-h _ ⟩ + Q-suc n ! α-h (N .alg-suc n) + ≡⟨ substComposite (F .Fiber) (σ n) (cong (N .alg-suc) (P n)) _ ⟩ + cong (N .alg-suc) (P n) ! (σ n ! α-h (N .alg-suc n)) + ≡[ i ]⟨ cong (N .alg-suc) (P n) ! fromPathP (σ-h n) i ⟩ + cong (N .alg-suc) (P n) ! (F .fib-suc (α-h n)) + ≡⟨ substCommSlice (F .Fiber) (F .Fiber ∘ N .alg-suc) (λ _ → F .fib-suc) (P n) (α-h n) ⟩ + F .fib-suc (P n ! α-h n) + ∎ + +isNatInductive≡isNatHInitial : {N : NatAlgebra ℓ'} (ℓ : Level) + → isNatInductive N (ℓ-max ℓ' ℓ) ≡ isNatHInitial N (ℓ-max ℓ' ℓ) +isNatInductive≡isNatHInitial {_} {N} ℓ = hPropExt isPropIsNatInductive isPropIsNatHInitial ind→init init→ind where + open AlgebraPropositionality + open AlgebraHInit→Ind N ℓ renaming (Fsection to init→ind) + open AlgebraHInd→HInit renaming (isContrMorph to ind→init) + +-- given two homotopy initial algebras there is a path between the algebras +-- this implies moreover that the carrier types are isomorphic +-- according to 5.16 in the paper this could be strengthened to isContr (N ≡ M) +isNatHInitial→algebraPath : {N M : NatAlgebra ℓ} + → (hinitN : isNatHInitial N ℓ) (hinitM : isNatHInitial M ℓ) + → N ≡ M +isNatHInitial→algebraPath {N = N} {M} hinitN hinitM = N≡M where + open Σ (hinitN M) renaming (fst to N→M) + open Σ (hinitM N) renaming (fst to M→N) + idN : NatMorphism N N + idN = record { morph = λ x → x ; comm-zero = refl ; comm-suc = refl } + idM : NatMorphism M M + idM = record { morph = λ x → x ; comm-zero = refl ; comm-suc = refl } + N→M→N : NatMorphism N N + morph N→M→N = morph M→N ∘ morph N→M + comm-zero N→M→N = (λ i → morph M→N (comm-zero N→M i)) ∙ comm-zero M→N + comm-suc N→M→N = (λ i → morph M→N ∘ comm-suc N→M i) ∙ (λ i → comm-suc M→N i ∘ morph N→M) + nmn≡idn : N→M→N ≡ idN + nmn≡idn = isContr→isProp (hinitN N) _ _ + + M→N→M : NatMorphism M M + morph M→N→M = morph N→M ∘ morph M→N + comm-zero M→N→M = (λ i → morph N→M (comm-zero M→N i)) ∙ comm-zero N→M + comm-suc M→N→M = (λ i → morph N→M ∘ comm-suc M→N i) ∙ (λ i → comm-suc N→M i ∘ morph M→N) + mnm≡idm : M→N→M ≡ idM + mnm≡idm = isContr→isProp (hinitM M) _ _ + + carrier≡ : N .Carrier ≡ M .Carrier + carrier≡ = isoToPath (iso (N→M .morph) (M→N .morph) (λ x i → mnm≡idm i .morph x) (λ x i → nmn≡idn i .morph x)) + zero≡ : PathP (λ i → carrier≡ i) (N .alg-zero) (M .alg-zero) + zero≡ = toPathP λ i → transportRefl (N→M .comm-zero i) i + suc≡ : PathP (λ i → carrier≡ i → carrier≡ i) (N .alg-suc) (M .alg-suc) + suc≡ = toPathP ( + transport refl ∘ N→M .morph ∘ N .alg-suc ∘ M→N .morph ∘ transport refl + ≡[ i ]⟨ transportReflF i ∘ N→M .morph ∘ N .alg-suc ∘ M→N .morph ∘ transportReflF i ⟩ + N→M .morph ∘ N .alg-suc ∘ M→N .morph + ≡[ i ]⟨ N→M .comm-suc i ∘ M→N .morph ⟩ + M .alg-suc ∘ N→M .morph ∘ M→N .morph + ≡[ i ]⟨ M .alg-suc ∘ mnm≡idm i .morph ⟩ + M .alg-suc + ∎) where + transportReflF : transport refl ≡ (λ x → x) + transportReflF = funExt transportRefl + + N≡M : N ≡ M + Carrier (N≡M i) = carrier≡ i + alg-zero (N≡M i) = zero≡ i + alg-suc (N≡M i) = suc≡ i + +-- the payoff, it is straight forward to define the algebra and show inductiveness of ℕ +NatAlgebraℕ : NatAlgebra ℓ-zero +Carrier NatAlgebraℕ = ℕ +alg-zero NatAlgebraℕ = zero +alg-suc NatAlgebraℕ = suc + +isNatInductiveℕ : isNatInductive NatAlgebraℕ ℓ +section (isNatInductiveℕ F) = nat-sec where + nat-sec : ∀ n → F .Fiber n + nat-sec zero = F .fib-zero + nat-sec (suc n) = F .fib-suc (nat-sec n) +sec-comm-zero (isNatInductiveℕ F) = refl +sec-comm-suc (isNatInductiveℕ F) n = refl + +isNatHInitialℕ : isNatHInitial NatAlgebraℕ ℓ +isNatHInitialℕ = transport (isNatInductive≡isNatHInitial _) isNatInductiveℕ diff --git a/Cubical/Data/Nat/Base.agda b/Cubical/Data/Nat/Base.agda new file mode 100644 index 0000000000..798728362c --- /dev/null +++ b/Cubical/Data/Nat/Base.agda @@ -0,0 +1,61 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Data.Nat.Base where + +open import Cubical.Core.Primitives + +open import Agda.Builtin.Nat public + using (zero; suc; _+_) + renaming (Nat to ℕ; _-_ to _∸_; _*_ to _·_) + +open import Cubical.Data.Nat.Literals public +open import Cubical.Data.Bool.Base +open import Cubical.Data.Sum.Base hiding (elim) +open import Cubical.Data.Empty.Base hiding (elim) +open import Cubical.Data.Unit.Base + +predℕ : ℕ → ℕ +predℕ zero = zero +predℕ (suc n) = n + +caseNat : ∀ {ℓ} → {A : Type ℓ} → (a0 aS : A) → ℕ → A +caseNat a0 aS zero = a0 +caseNat a0 aS (suc n) = aS + +doubleℕ : ℕ → ℕ +doubleℕ zero = zero +doubleℕ (suc x) = suc (suc (doubleℕ x)) + +-- doublesℕ n m = 2^n · m +doublesℕ : ℕ → ℕ → ℕ +doublesℕ zero m = m +doublesℕ (suc n) m = doublesℕ n (doubleℕ m) + +-- iterate +iter : ∀ {ℓ} {A : Type ℓ} → ℕ → (A → A) → A → A +iter zero f z = z +iter (suc n) f z = f (iter n f z) + +elim : ∀ {ℓ} {A : ℕ → Type ℓ} + → A zero + → ((n : ℕ) → A n → A (suc n)) + → (n : ℕ) → A n +elim a₀ _ zero = a₀ +elim a₀ f (suc n) = f n (elim a₀ f n) + +isEven isOdd : ℕ → Bool +isEven zero = true +isEven (suc n) = isOdd n +isOdd zero = false +isOdd (suc n) = isEven n + +--Typed version +private + toType : Bool → Type + toType false = ⊥ + toType true = Unit + +isEvenT : ℕ → Type +isEvenT n = toType (isEven n) + +isOddT : ℕ → Type +isOddT n = isEvenT (suc n) diff --git a/Cubical/Data/Nat/Coprime.agda b/Cubical/Data/Nat/Coprime.agda new file mode 100644 index 0000000000..d551ae28c3 --- /dev/null +++ b/Cubical/Data/Nat/Coprime.agda @@ -0,0 +1,95 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Nat.Coprime where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Data.Sigma +open import Cubical.Data.NatPlusOne + +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +open import Cubical.Data.Nat.Base +open import Cubical.Data.Nat.Properties +open import Cubical.Data.Nat.Divisibility +open import Cubical.Data.Nat.GCD + +areCoprime : ℕ × ℕ → Type₀ +areCoprime (m , n) = isGCD m n 1 + +-- Any pair (m , n) can be converted to a coprime pair (m' , n') s.t. +-- m' ∣ m, n' ∣ n if and only if one of m or n is nonzero + +module ToCoprime ((m , n) : ℕ × ℕ₊₁) where + d = gcd m (ℕ₊₁→ℕ n) + d∣m = gcdIsGCD m (ℕ₊₁→ℕ n) .fst .fst + d∣n = gcdIsGCD m (ℕ₊₁→ℕ n) .fst .snd + gr = gcdIsGCD m (ℕ₊₁→ℕ n) .snd + + c₁ : ℕ + p₁ : c₁ · d ≡ m + c₁ = ∣-untrunc d∣m .fst; p₁ = ∣-untrunc d∣m .snd + + c₂ : ℕ₊₁ + p₂ : (ℕ₊₁→ℕ c₂) · d ≡ (ℕ₊₁→ℕ n) + c₂ = 1+ (∣s-untrunc d∣n .fst); p₂ = ∣s-untrunc d∣n .snd + + toCoprime : ℕ × ℕ₊₁ + toCoprime = (c₁ , c₂) + + private + lem : ∀ a {b c d e} → a · b ≡ c → c · d ≡ e → a · (b · d) ≡ e + lem a p q = ·-assoc a _ _ ∙ cong (_· _) p ∙ q + + gr' : ∀ d' → prediv d' c₁ → prediv d' (ℕ₊₁→ℕ c₂) → (d' · d) ∣ d + gr' d' (b₁ , q₁) (b₂ , q₂) = gr (d' · d) ((∣ b₁ , lem b₁ q₁ p₁ ∣) , + (∣ b₂ , lem b₂ q₂ p₂ ∣)) + + d-1 = m∣sn→z 0 (<=> m > 0 or n > 0) + d-cancelʳ : ∀ d' → (d' · d) ∣ d → d' ∣ 1 + d-cancelʳ d' div = ∣-cancelʳ d-1 (∣-trans (subst (λ x → (d' · x) ∣ x) q div) + (∣-refl (sym (·-identityˡ _)))) + + toCoprimeAreCoprime : areCoprime (c₁ , ℕ₊₁→ℕ c₂) + fst toCoprimeAreCoprime = ∣-oneˡ c₁ , ∣-oneˡ (ℕ₊₁→ℕ c₂) + snd toCoprimeAreCoprime d' (d'∣c₁ , d'∣c₂) = PropTrunc.rec isProp∣ (λ a → + PropTrunc.rec isProp∣ (λ b → + d-cancelʳ d' (gr' d' a b)) d'∣c₂) d'∣c₁ + + toCoprime∣ : (c₁ ∣ m) × (ℕ₊₁→ℕ c₂ ∣ ℕ₊₁→ℕ n) + toCoprime∣ = ∣ d , ·-comm d c₁ ∙ p₁ ∣ , ∣ d , ·-comm d (ℕ₊₁→ℕ c₂) ∙ p₂ ∣ + + toCoprime-idem : areCoprime (m , ℕ₊₁→ℕ n) → (c₁ , c₂) ≡ (m , n) + toCoprime-idem cp i = q₁ i , ℕ₊₁→ℕ-inj q₂ i + where q₁ = sym (·-identityʳ c₁) ∙ cong (c₁ ·_) (sym (isGCD→gcd≡ cp)) ∙ p₁ + q₂ = sym (·-identityʳ (ℕ₊₁→ℕ c₂)) ∙ cong (ℕ₊₁→ℕ c₂ ·_) (sym (isGCD→gcd≡ cp)) ∙ p₂ + +open ToCoprime using (toCoprime; toCoprimeAreCoprime; toCoprime∣; toCoprime-idem) public + + +toCoprime-cancelʳ : ∀ ((m , n) : ℕ × ℕ₊₁) k + → toCoprime (m · ℕ₊₁→ℕ k , n ·₊₁ k) ≡ toCoprime (m , n) +toCoprime-cancelʳ (m , n) (1+ k) i = + inj-·sm {c₁'} {d-1} {c₁} r₁ i , ℕ₊₁→ℕ-inj (inj-·sm {ℕ₊₁→ℕ c₂'} {d-1} {ℕ₊₁→ℕ c₂} r₂) i + where open ToCoprime (m , n) + open ToCoprime (m · suc k , n ·₊₁ (1+ k)) using () + renaming (c₁ to c₁'; p₁ to p₁'; c₂ to c₂'; p₂ to p₂') + + q₁ : c₁' · d · suc k ≡ m · suc k + q₁ = sym (·-assoc c₁' (ToCoprime.d (m , n)) (suc k)) + ∙ cong (c₁' ·_) (sym (gcd-factorʳ m (ℕ₊₁→ℕ n) (suc k))) + ∙ p₁' + q₂ : ℕ₊₁→ℕ c₂' · (ToCoprime.d (m , n)) · suc k ≡ ℕ₊₁→ℕ n · suc k + q₂ = sym (·-assoc (ℕ₊₁→ℕ c₂') (ToCoprime.d (m , n)) (suc k)) + ∙ cong (ℕ₊₁→ℕ c₂' ·_) (sym (gcd-factorʳ m (ℕ₊₁→ℕ n) (suc k))) + ∙ p₂' + + r₁ : c₁' · suc d-1 ≡ c₁ · suc d-1 + r₁ = subst (λ z → c₁' · z ≡ c₁ · z) q (inj-·sm q₁ ∙ sym p₁) + r₂ : ℕ₊₁→ℕ c₂' · suc d-1 ≡ ℕ₊₁→ℕ c₂ · suc d-1 + r₂ = subst (λ z → ℕ₊₁→ℕ c₂' · z ≡ ℕ₊₁→ℕ c₂ · z) q (inj-·sm q₂ ∙ sym p₂) diff --git a/Cubical/Data/Nat/Divisibility.agda b/Cubical/Data/Nat/Divisibility.agda new file mode 100644 index 0000000000..27354cf552 --- /dev/null +++ b/Cubical/Data/Nat/Divisibility.agda @@ -0,0 +1,104 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Nat.Divisibility where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ + +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +open import Cubical.Data.Nat.Base +open import Cubical.Data.Nat.Properties +open import Cubical.Data.Nat.Order + +private + variable + l m n : ℕ + +_∣_ : ℕ → ℕ → Type₀ +m ∣ n = ∃[ c ∈ ℕ ] c · m ≡ n + +isProp∣ : isProp (m ∣ n) +isProp∣ = squash + +prediv : ℕ → ℕ → Type₀ +prediv m n = Σ[ c ∈ ℕ ] c · m ≡ n + +-- an alternate definition of m ∣ n that doesn't use truncation + +_∣'_ : ℕ → ℕ → Type₀ +zero ∣' n = zero ≡ n +suc m ∣' n = Σ[ c ∈ ℕ ] c · suc m ≡ n + +isProp∣' : isProp (m ∣' n) +isProp∣' {zero} {n} = isSetℕ _ _ +isProp∣' {suc m} {n} (c₁ , p₁) (c₂ , p₂) = + Σ≡Prop (λ _ → isSetℕ _ _) (inj-·sm {c₁} {m} {c₂} (p₁ ∙ sym p₂)) + +∣≃∣' : (m ∣ n) ≃ (m ∣' n) +∣≃∣' {zero} = propBiimpl→Equiv isProp∣ isProp∣' + (PropTrunc.rec (isSetℕ _ _) λ { (c , p) → 0≡m·0 c ∙ p }) + (λ p → ∣ zero , p ∣) +∣≃∣' {suc m} = propTruncIdempotent≃ isProp∣' + +∣-untrunc : m ∣ n → Σ[ c ∈ ℕ ] c · m ≡ n +∣-untrunc {zero} p = zero , equivFun ∣≃∣' p +∣-untrunc {suc m} = equivFun ∣≃∣' + + +-- basic properties of ∣ + +∣-refl : m ≡ n → m ∣ n +∣-refl p = ∣ 1 , +-zero _ ∙ p ∣ + +∣-trans : l ∣ m → m ∣ n → l ∣ n +∣-trans = PropTrunc.map2 λ { + (c₁ , p) (c₂ , q) → (c₂ · c₁ , sym (·-assoc c₂ c₁ _) ∙ cong (c₂ ·_) p ∙ q) } + +∣-left : ∀ k → m ∣ (m · k) +∣-left k = ∣ k , ·-comm k _ ∣ + +∣-right : ∀ k → m ∣ (k · m) +∣-right k = ∣ k , refl ∣ + +∣-cancelʳ : ∀ k → (m · suc k) ∣ (n · suc k) → m ∣ n +∣-cancelʳ k = PropTrunc.map λ { + (c , p) → (c , inj-·sm (sym (·-assoc c _ (suc k)) ∙ p)) } + +∣-multʳ : ∀ k → m ∣ n → (m · k) ∣ (n · k) +∣-multʳ k = PropTrunc.map λ { + (c , p) → (c , ·-assoc c _ k ∙ cong (_· k) p) } + +∣-zeroˡ : zero ∣ m → zero ≡ m +∣-zeroˡ = equivFun ∣≃∣' + +∣-zeroʳ : ∀ m → m ∣ zero +∣-zeroʳ m = ∣ zero , refl ∣ + +∣-oneˡ : ∀ m → 1 ∣ m +∣-oneˡ m = ∣ m , ·-identityʳ m ∣ + +-- if n > 0, then the constant c (s.t. c · m ≡ n) is also > 0 +∣s-untrunc : m ∣ suc n → Σ[ c ∈ ℕ ] (suc c) · m ≡ suc n +∣s-untrunc ∣p∣ with ∣-untrunc ∣p∣ +... | (zero , p) = ⊥.rec (znots p) +... | (suc c , p) = (c , p) + +m∣sn→m≤sn : m ∣ suc n → m ≤ suc n +m∣sn→m≤sn {m} {n} = f ∘ ∣s-untrunc + where f : Σ[ c ∈ ℕ ] (suc c) · m ≡ suc n → Σ[ c ∈ ℕ ] c + m ≡ suc n + f (c , p) = (c · m) , (+-comm (c · m) m ∙ p) + +m∣sn→z trySigma <|> tryParam <|> tryPi <|> tryUnit <|> useGeneric + where + tryUniv : R.TC Unit + tryUniv = R.unify (R.con (quote UARelDesc.univ) [ varg R.unknown ]) hole + + tryBinary : R.Name → R.TC Unit + tryBinary name = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ hole₂ → + R.unify (R.con name (hole₁ v∷ hole₂ v∷ [])) hole >> + autoUARelDesc n hole₁ >> + autoDUARelDesc n hole₂ + + tryParam : R.TC Unit + tryParam = + newMeta R.unknown >>= λ paramTy → + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con (quote UARelDesc.param) (paramTy v∷ hole₁ v∷ [])) hole >> + autoUARelDesc n hole₁ + + trySigma = tryBinary (quote UARelDesc.sigma) + tryPi = tryBinary (quote UARelDesc.pi) + + tryUnit : R.TC Unit + tryUnit = R.unify (R.con (quote UARelDesc.unit) []) hole + + useGeneric : R.TC Unit + useGeneric = R.unify (R.con (quote UARelDesc.generic) []) hole + + autoUARelReindex : ℕ → R.Term → R.TC Unit + autoUARelReindex zero hole = R.typeError [ R.strErr "Out of fuel" ] + autoUARelReindex (suc n) hole = + tryId <|> tryFst <|> trySnd <|> tryApp + where + tryId : R.TC Unit + tryId = R.unify (R.con (quote UARelReindex.id) []) hole + + tryUnary : R.Name → R.TC Unit + tryUnary name = + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con name [ varg hole₁ ]) hole >> + autoUARelReindex n hole₁ + + tryFst = tryUnary (quote UARelReindex.∘fst) + trySnd = tryUnary (quote UARelReindex.∘snd) + + tryApp : R.TC Unit + tryApp = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ param → + R.unify (R.con (quote UARelReindex.∘app) (hole₁ v∷ param v∷ [])) hole >> + autoUARelReindex n hole₁ + + autoSubstRelDesc : ℕ → R.Term → R.TC Unit + autoSubstRelDesc zero hole = R.typeError [ R.strErr "Out of fuel" ] + autoSubstRelDesc (suc n) hole = + tryConstant <|> tryEl <|> tryEl <|> trySigma <|> tryPi <|> useGeneric + where + tryConstant : R.TC Unit + tryConstant = + R.unify (R.con (quote SubstRelDesc.constant) []) hole + + tryEl : R.TC Unit + tryEl = + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con (quote SubstRelDesc.el) [ varg hole₁ ]) hole >> + autoUARelReindex n hole₁ + + tryBinary : R.Name → R.TC Unit + tryBinary name = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ hole₂ → + R.unify (R.con name (hole₁ v∷ hole₂ v∷ [])) hole >> + autoSubstRelDesc n hole₁ >> + autoSubstRelDesc n hole₂ + + trySigma = tryBinary (quote SubstRelDesc.sigma) + tryPi = tryBinary (quote SubstRelDesc.pi) + + useGeneric : R.TC Unit + useGeneric = R.unify (R.con (quote SubstRelDesc.generic) []) hole + + autoDUARelDesc : ℕ → R.Term → R.TC Unit + autoDUARelDesc zero hole = R.typeError [ R.strErr "Out of fuel" ] + autoDUARelDesc (suc n) hole = + tryConstant <|> tryEl <|> trySigma <|> tryPiˢ <|> tryPi <|> useGeneric + where + tryConstant : R.TC Unit + tryConstant = + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con (quote DUARelDesc.constant) [ varg hole₁ ]) hole >> + autoUARelDesc n hole₁ + + tryEl : R.TC Unit + tryEl = + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con (quote DUARelDesc.el) [ varg hole₁ ]) hole >> + autoUARelReindex n hole₁ + + tryBinary : R.Name → R.TC Unit + tryBinary name = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ hole₂ → + R.unify (R.con name (hole₁ v∷ hole₂ v∷ [])) hole >> + autoDUARelDesc n hole₁ >> + autoDUARelDesc n hole₂ + + trySigma = tryBinary (quote DUARelDesc.sigma) + tryPi = tryBinary (quote DUARelDesc.pi) + + tryPiˢ : R.TC Unit + tryPiˢ = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ hole₂ → + R.unify (R.con (quote DUARelDesc.piˢ) (hole₁ v∷ hole₂ v∷ [])) hole >> + autoSubstRelDesc n hole₁ >> + autoDUARelDesc n hole₂ + + useGeneric : R.TC Unit + useGeneric = R.unify (R.con (quote DUARelDesc.generic) []) hole + +module DisplayedAutoMacro where + autoUARel : ∀ {ℓA} (A : Type ℓA) → ℕ → R.Term → R.TC Unit + autoUARel A n hole = + R.quoteTC A >>= λ `A` → + newMeta R.unknown >>= λ desc → + makeAuxiliaryDef "autoUA" + (R.def (quote UARel) (`A` v∷ R.unknown v∷ [])) + (R.def (quote getUARel) [ varg desc ]) + >>= λ uaTerm → + R.unify hole uaTerm >> + autoUARelDesc n desc + + autoDUARel : ∀ {ℓA ℓ≅A ℓB} {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) + → ℕ → R.Term → R.TC Unit + autoDUARel 𝒮-A B n hole = + R.quoteTC 𝒮-A >>= λ `𝒮-A` → + R.quoteTC B >>= λ `B` → + newMeta R.unknown >>= λ desc → + makeAuxiliaryDef "autoDUA" + (R.def (quote DUARel) (`𝒮-A` v∷ `B` v∷ R.unknown v∷ [])) + (R.def (quote getDUARel) [ varg desc ]) + >>= λ duaTerm → + R.unify hole duaTerm >> + autoDUARelDesc n desc + +macro + autoUARel : ∀ {ℓA} (A : Type ℓA) → R.Term → R.TC Unit + autoUARel A = DisplayedAutoMacro.autoUARel A FUEL + + autoDUARel : ∀ {ℓA ℓ≅A ℓB} {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) + → R.Term → R.TC Unit + autoDUARel 𝒮-A B = DisplayedAutoMacro.autoDUARel 𝒮-A B FUEL + +private + module Example (A : Type) (a₀ : A) where + + example0 : DUARel (autoUARel Type) (λ X → X → A × X) ℓ-zero + example0 = autoDUARel _ _ + + example0' : {X Y : Type} (e : X ≃ Y) + (f : X → A × X) (g : Y → A × Y) + → (∀ x → (f x .fst ≡ g (e .fst x) .fst) × (e .fst (f x .snd) ≡ g (e .fst x) .snd)) + → PathP (λ i → ua e i → A × ua e i) f g + example0' e f g = example0 .DUARel.uaᴰ f e g .fst + + -- An example where a DUARel is parameterized over a pair of types + + example1 : DUARel (autoUARel (Type × Type)) (λ (X , Z) → X → Z) ℓ-zero + example1 = autoDUARel _ _ + + example1' : {X Y : Type} (e : X ≃ Y) {Z W : Type} (h : Z ≃ W) + (f : X → Z) (g : Y → W) + → (∀ x → h .fst (f x) ≡ g (e .fst x)) + → PathP (λ i → ua e i → ua h i) f g + example1' e h f g = example1 .DUARel.uaᴰ f (e , h) g .fst + + -- An example where a DUARel is parameterized over a family of types + + example2 : DUARel (autoUARel (A → Type)) (λ B → B a₀) ℓ-zero + example2 = autoDUARel _ _ + + example2' : {B C : A → Type} (e : (a : A) → B a ≃ C a) + (b : B a₀) (c : C a₀) + → e a₀ .fst b ≡ c + → PathP (λ i → ua (e a₀) i) b c + example2' e b c = example2 .DUARel.uaᴰ b e c .fst diff --git a/Cubical/Displayed/Base.agda b/Cubical/Displayed/Base.agda new file mode 100644 index 0000000000..e91565a559 --- /dev/null +++ b/Cubical/Displayed/Base.agda @@ -0,0 +1,87 @@ +{- + + Definition of univalent and displayed univalent relations + +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport + +open import Cubical.Data.Sigma + +open import Cubical.Relation.Binary + +private + variable + ℓ ℓA ℓA' ℓ≅A ℓB ℓB' ℓ≅B ℓC ℓ≅C : Level + +record UARel (A : Type ℓA) (ℓ≅A : Level) : Type (ℓ-max ℓA (ℓ-suc ℓ≅A)) where + no-eta-equality + constructor uarel + field + _≅_ : A → A → Type ℓ≅A + ua : (a a' : A) → (a ≅ a') ≃ (a ≡ a') + + uaIso : (a a' : A) → Iso (a ≅ a') (a ≡ a') + uaIso a a' = equivToIso (ua a a') + + ≅→≡ : {a a' : A} (p : a ≅ a') → a ≡ a' + ≅→≡ {a} {a'} = Iso.fun (uaIso a a') + ≡→≅ : {a a' : A} (p : a ≡ a') → a ≅ a' + ≡→≅ {a} {a'} = Iso.inv (uaIso a a') + + ρ : (a : A) → a ≅ a + ρ a = ≡→≅ refl + +open BinaryRelation + +-- another constructor for UARel using contractibility of relational singletons +make-𝒮 : {A : Type ℓA} {_≅_ : A → A → Type ℓ≅A} + (ρ : isRefl _≅_) (contrTotal : contrRelSingl _≅_) → UARel A ℓ≅A +UARel._≅_ (make-𝒮 {_≅_ = _≅_} _ _) = _≅_ +UARel.ua (make-𝒮 {_≅_ = _≅_} ρ c) = contrRelSingl→isUnivalent _≅_ ρ c + +record DUARel {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) + (B : A → Type ℓB) (ℓ≅B : Level) : Type (ℓ-max (ℓ-max ℓA ℓB) (ℓ-max ℓ≅A (ℓ-suc ℓ≅B))) where + no-eta-equality + constructor duarel + open UARel 𝒮-A + + field + _≅ᴰ⟨_⟩_ : {a a' : A} → B a → a ≅ a' → B a' → Type ℓ≅B + uaᴰ : {a a' : A} (b : B a) (p : a ≅ a') (b' : B a') → (b ≅ᴰ⟨ p ⟩ b') ≃ PathP (λ i → B (≅→≡ p i)) b b' + + fiberRel : (a : A) → Rel (B a) (B a) ℓ≅B + fiberRel a = _≅ᴰ⟨ ρ a ⟩_ + + uaᴰρ : {a : A} (b b' : B a) → b ≅ᴰ⟨ ρ a ⟩ b' ≃ (b ≡ b') + uaᴰρ {a} b b' = + compEquiv + (uaᴰ b (ρ _) b') + (substEquiv (λ q → PathP (λ i → B (q i)) b b') (secEq (ua a a) refl)) + + ρᴰ : {a : A} → (b : B a) → b ≅ᴰ⟨ ρ a ⟩ b + ρᴰ {a} b = invEq (uaᴰρ b b) refl + + +-- total UARel induced by a DUARel + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} {ℓ≅B : Level} + (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + where + + open UARel 𝒮-A + open DUARel 𝒮ᴰ-B + + ∫ : UARel (Σ A B) (ℓ-max ℓ≅A ℓ≅B) + UARel._≅_ ∫ (a , b) (a' , b') = Σ[ p ∈ a ≅ a' ] (b ≅ᴰ⟨ p ⟩ b') + UARel.ua ∫ (a , b) (a' , b') = + compEquiv + (Σ-cong-equiv (ua a a') (λ p → uaᴰ b p b')) + ΣPath≃PathΣ + diff --git a/Cubical/Displayed/Constant.agda b/Cubical/Displayed/Constant.agda new file mode 100644 index 0000000000..0689104c28 --- /dev/null +++ b/Cubical/Displayed/Constant.agda @@ -0,0 +1,39 @@ +{- + + Functions building DUARels on constant families + +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Constant where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst + +private + variable + ℓ ℓA ℓA' ℓP ℓ≅A ℓ≅A' ℓB ℓB' ℓ≅B ℓ≅B' ℓC ℓ≅C : Level + +-- constant DUARel + +module _ {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) + {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) where + + open UARel 𝒮-B + open DUARel + + 𝒮ᴰ-const : DUARel 𝒮-A (λ _ → B) ℓ≅B + 𝒮ᴰ-const ._≅ᴰ⟨_⟩_ b _ b' = b ≅ b' + 𝒮ᴰ-const .uaᴰ b p b' = ua b b' + +-- SubstRel for an arbitrary constant family + +module _ {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : Type ℓB) where + + open SubstRel + + 𝒮ˢ-const : SubstRel 𝒮-A (λ _ → B) + 𝒮ˢ-const .SubstRel.act _ = idEquiv B + 𝒮ˢ-const .SubstRel.uaˢ p b = transportRefl b diff --git a/Cubical/Displayed/Function.agda b/Cubical/Displayed/Function.agda new file mode 100644 index 0000000000..4c977fc792 --- /dev/null +++ b/Cubical/Displayed/Function.agda @@ -0,0 +1,135 @@ +{- + + Functions building UARels and DUARels on function types + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Function where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path + +open import Cubical.Functions.FunExtEquiv +open import Cubical.Functions.Implicit + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Constant +open import Cubical.Displayed.Morphism +open import Cubical.Displayed.Subst +open import Cubical.Displayed.Sigma + +private + variable + ℓA ℓ≅A ℓB ℓ≅B ℓC ℓ≅C : Level + +-- UARel on dependent function type +-- from UARel on domain and DUARel on codomain + +module _ {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) where + + open UARel 𝒮-A + open DUARel 𝒮ᴰ-B + + 𝒮-Π : UARel ((a : A) → B a) (ℓ-max ℓA ℓ≅B) + UARel._≅_ 𝒮-Π f f' = ∀ a → f a ≅ᴰ⟨ ρ a ⟩ f' a + UARel.ua 𝒮-Π f f' = + compEquiv + (equivΠCod λ a → uaᴰρ (f a) (f' a)) + funExtEquiv + +-- Parameterize UARel by type + +_→𝒮_ : (A : Type ℓA) {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) → UARel (A → B) (ℓ-max ℓA ℓ≅B) +(A →𝒮 𝒮-B) .UARel._≅_ f f' = ∀ a → 𝒮-B .UARel._≅_ (f a) (f' a) +(A →𝒮 𝒮-B) .UARel.ua f f' = + compEquiv + (equivΠCod λ a → 𝒮-B .UARel.ua (f a) (f' a)) + funExtEquiv + +𝒮-app : {A : Type ℓA} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} + → A → UARelHom (A →𝒮 𝒮-B) 𝒮-B +𝒮-app a .UARelHom.fun f = f a +𝒮-app a .UARelHom.rel h = h a +𝒮-app a .UARelHom.ua h = refl + +-- DUARel on dependent function type +-- from DUARels on domain and codomain + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : (a : A) → B a → Type ℓC} (𝒮ᴰ-C : DUARel (∫ 𝒮ᴰ-B) (uncurry C) ℓ≅C) + where + + open UARel 𝒮-A + private + module B = DUARel 𝒮ᴰ-B + module C = DUARel 𝒮ᴰ-C + + 𝒮ᴰ-Π : DUARel 𝒮-A (λ a → (b : B a) → C a b) (ℓ-max (ℓ-max ℓB ℓ≅B) ℓ≅C) + DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-Π f p f' = + ∀ {b b'} (q : b B.≅ᴰ⟨ p ⟩ b') → f b C.≅ᴰ⟨ p , q ⟩ f' b' + DUARel.uaᴰ 𝒮ᴰ-Π f p f' = + compEquiv + (equivImplicitΠCod λ {b} → + (equivImplicitΠCod λ {b'} → + (equivΠ (B.uaᴰ b p b') (λ q → C.uaᴰ (f b) (p , q) (f' b'))))) + funExtDepEquiv + +_→𝒮ᴰ_ : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : A → Type ℓC} (𝒮ᴰ-C : DUARel 𝒮-A C ℓ≅C) + → DUARel 𝒮-A (λ a → B a → C a) (ℓ-max (ℓ-max ℓB ℓ≅B) ℓ≅C) +𝒮ᴰ-B →𝒮ᴰ 𝒮ᴰ-C = + 𝒮ᴰ-Π 𝒮ᴰ-B (𝒮ᴰ-Lift _ 𝒮ᴰ-C 𝒮ᴰ-B) + +-- DUARel on dependent function type +-- from a SubstRel on the domain and DUARel on the codomain + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ˢ-B : SubstRel 𝒮-A B) + {C : (a : A) → B a → Type ℓC} (𝒮ᴰ-C : DUARel (∫ˢ 𝒮ˢ-B) (uncurry C) ℓ≅C) + where + + open UARel 𝒮-A + open SubstRel 𝒮ˢ-B + open DUARel 𝒮ᴰ-C + + 𝒮ᴰ-Πˢ : DUARel 𝒮-A (λ a → (b : B a) → C a b) (ℓ-max ℓB ℓ≅C) + DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-Πˢ f p f' = + (b : B _) → f b ≅ᴰ⟨ p , refl ⟩ f' (act p .fst b) + DUARel.uaᴰ 𝒮ᴰ-Πˢ f p f' = + compEquiv + (compEquiv + (equivΠCod λ b → Jequiv (λ b' q → f b ≅ᴰ⟨ p , q ⟩ f' b')) + (invEquiv implicit≃Explicit)) + (DUARel.uaᴰ (𝒮ᴰ-Π (Subst→DUA 𝒮ˢ-B) 𝒮ᴰ-C) f p f') + +-- SubstRel on a dependent function type +-- from a SubstRel on the domain and SubstRel on the codomain + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ˢ-B : SubstRel 𝒮-A B) + {C : Σ A B → Type ℓC} (𝒮ˢ-C : SubstRel (∫ˢ 𝒮ˢ-B) C) + where + + open UARel 𝒮-A + open SubstRel + private + module B = SubstRel 𝒮ˢ-B + module C = SubstRel 𝒮ˢ-C + + 𝒮ˢ-Π : SubstRel 𝒮-A (λ a → (b : B a) → C (a , b)) + 𝒮ˢ-Π .act p = equivΠ' (B.act p) (λ q → C.act (p , q)) + 𝒮ˢ-Π .uaˢ p f = + fromPathP + (DUARel.uaᴰ (𝒮ᴰ-Π (Subst→DUA 𝒮ˢ-B) (Subst→DUA 𝒮ˢ-C)) f p (equivFun (𝒮ˢ-Π .act p) f) .fst + (λ {b} → + J (λ b' q → + equivFun (C.act (p , q)) (f b) + ≡ equivFun (equivΠ' (𝒮ˢ-B .act p) (λ q → C.act (p , q))) f b') + (λ i → + C.act (p , λ j → commSqIsEq (𝒮ˢ-B .act p .snd) b (~ i) j) .fst + (f (retEq (𝒮ˢ-B .act p) b (~ i)))))) diff --git a/Cubical/Displayed/Generic.agda b/Cubical/Displayed/Generic.agda new file mode 100644 index 0000000000..d28d705bcf --- /dev/null +++ b/Cubical/Displayed/Generic.agda @@ -0,0 +1,36 @@ +{- + + Generic UARel, DUARel, and SubstRel: the path relation is always univalent + +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Generic where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst + +private + variable + ℓ ℓA ℓA' ℓP ℓ≅A ℓ≅A' ℓB ℓB' ℓ≅B ℓ≅B' ℓC ℓ≅C : Level + +-- UARel for an arbitrary type + +𝒮-generic : (A : Type ℓA) → UARel A ℓA +UARel._≅_ (𝒮-generic A) = _≡_ +UARel.ua (𝒮-generic A) a a' = idEquiv (a ≡ a') + +-- DUARel for an arbitrary family + +𝒮ᴰ-generic : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) → DUARel 𝒮-A B ℓB +𝒮ᴰ-generic 𝒮-A B .DUARel._≅ᴰ⟨_⟩_ b p b' = PathP (λ i → B (UARel.≅→≡ 𝒮-A p i)) b b' +𝒮ᴰ-generic 𝒮-A B .DUARel.uaᴰ b p b' = idEquiv _ + +-- SubstRel for an arbitrary family + +𝒮ˢ-generic : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) → SubstRel 𝒮-A B +𝒮ˢ-generic 𝒮-A B .SubstRel.act p = substEquiv B (UARel.≅→≡ 𝒮-A p) +𝒮ˢ-generic 𝒮-A B .SubstRel.uaˢ p b = refl diff --git a/Cubical/Displayed/Morphism.agda b/Cubical/Displayed/Morphism.agda new file mode 100644 index 0000000000..94b463a36a --- /dev/null +++ b/Cubical/Displayed/Morphism.agda @@ -0,0 +1,67 @@ +{- + A morphism of UARels is a function between the structures with an action on the relations that + commutes with the equivalence to PathP. + + We can reindex a DUARel or SubstRel along one of these. +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Morphism where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst + +private + variable + ℓ ℓA ℓA' ℓ≅A ℓB ℓB' ℓ≅B ℓC ℓ≅C : Level + +record UARelHom {A : Type ℓA} {B : Type ℓB} (𝒮-A : UARel A ℓ≅A) (𝒮-B : UARel B ℓ≅B) + : Type (ℓ-max (ℓ-max ℓA ℓ≅A) (ℓ-max ℓB ℓ≅B)) where + no-eta-equality + constructor uarelhom + field + fun : A → B + rel : ∀ {a a'} → UARel._≅_ 𝒮-A a a' → UARel._≅_ 𝒮-B (fun a) (fun a') + ua : ∀ {a a'} (p : UARel._≅_ 𝒮-A a a') + → cong fun (UARel.≅→≡ 𝒮-A p) ≡ UARel.≅→≡ 𝒮-B (rel p) + +open UARelHom + +𝒮-id : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) → UARelHom 𝒮-A 𝒮-A +𝒮-id 𝒮-A .fun = idfun _ +𝒮-id 𝒮-A .rel = idfun _ +𝒮-id 𝒮-A .ua _ = refl + +𝒮-∘ : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} + {C : Type ℓC} {𝒮-C : UARel C ℓ≅C} + → UARelHom 𝒮-B 𝒮-C + → UARelHom 𝒮-A 𝒮-B + → UARelHom 𝒮-A 𝒮-C +𝒮-∘ g f .fun = g .fun ∘ f .fun +𝒮-∘ g f .rel = g .rel ∘ f .rel +𝒮-∘ {𝒮-A = 𝒮-A} {𝒮-B = 𝒮-B} {𝒮-C = 𝒮-C} g f .ua p = + cong (cong (g .fun)) (f .ua p) ∙ g .ua (f .rel p) + +𝒮ᴰ-reindex : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} {C : B → Type ℓC} + (f : UARelHom 𝒮-A 𝒮-B) + → DUARel 𝒮-B C ℓ≅C + → DUARel 𝒮-A (C ∘ fun f) ℓ≅C +𝒮ᴰ-reindex f 𝒮ᴰ-C .DUARel._≅ᴰ⟨_⟩_ c p c' = 𝒮ᴰ-C .DUARel._≅ᴰ⟨_⟩_ c (f .rel p) c' +𝒮ᴰ-reindex {C = C} f 𝒮ᴰ-C .DUARel.uaᴰ c p c' = + compEquiv + (𝒮ᴰ-C .DUARel.uaᴰ c (f .rel p) c') + (substEquiv (λ q → PathP (λ i → C (q i)) c c') (sym (f .ua p))) + +𝒮ˢ-reindex : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} {C : B → Type ℓC} + (f : UARelHom 𝒮-A 𝒮-B) + → SubstRel 𝒮-B C + → SubstRel 𝒮-A (C ∘ fun f) +𝒮ˢ-reindex f 𝒮ˢ-C .SubstRel.act p = 𝒮ˢ-C .SubstRel.act (f .rel p) +𝒮ˢ-reindex {C = C} f 𝒮ˢ-C .SubstRel.uaˢ p c = + cong (λ q → subst C q c) (f .ua p) + ∙ 𝒮ˢ-C .SubstRel.uaˢ (f .rel p) c diff --git a/Cubical/Displayed/Prop.agda b/Cubical/Displayed/Prop.agda new file mode 100644 index 0000000000..d57a0d5e2a --- /dev/null +++ b/Cubical/Displayed/Prop.agda @@ -0,0 +1,52 @@ +{- + + Functions building UARels and DUARels on propositions / propositional families + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Prop where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.Displayed.Base + +private + variable + ℓA ℓ≅A ℓB ℓ≅B ℓP : Level + +𝒮-prop : (P : hProp ℓP) → UARel (P .fst) ℓ-zero +𝒮-prop P .UARel._≅_ _ _ = Unit +𝒮-prop P .UARel.ua _ _ = + invEquiv (isContr→≃Unit (isOfHLevelPath' 0 (P .snd) _ _)) + +𝒮ᴰ-prop : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (P : A → hProp ℓP) + → DUARel 𝒮-A (λ a → P a .fst) ℓ-zero +𝒮ᴰ-prop 𝒮-A P .DUARel._≅ᴰ⟨_⟩_ _ _ _ = Unit +𝒮ᴰ-prop 𝒮-A P .DUARel.uaᴰ _ _ _ = + invEquiv (isContr→≃Unit (isOfHLevelPathP' 0 (P _ .snd) _ _)) + +𝒮-subtype : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {P : A → Type ℓP} + → (∀ a → isProp (P a)) + → UARel (Σ A P) ℓ≅A +𝒮-subtype 𝒮-A propP .UARel._≅_ (a , _) (a' , _) = 𝒮-A .UARel._≅_ a a' +𝒮-subtype 𝒮-A propP .UARel.ua (a , _) (a' , _) = + compEquiv (𝒮-A .UARel.ua a a') (Σ≡PropEquiv propP) + +𝒮ᴰ-subtype : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {P : (a : A) → B a → Type ℓP} + → (∀ a b → isProp (P a b)) + → DUARel 𝒮-A (λ a → Σ[ b ∈ B a ] (P a b)) ℓ≅B +𝒮ᴰ-subtype 𝒮ᴰ-B propP .DUARel._≅ᴰ⟨_⟩_ (b , _) p (b' , _) = 𝒮ᴰ-B .DUARel._≅ᴰ⟨_⟩_ b p b' +𝒮ᴰ-subtype 𝒮ᴰ-B propP .DUARel.uaᴰ (b , _) p (b' , _) = + compEquiv + (𝒮ᴰ-B .DUARel.uaᴰ b p b') + (compEquiv + (invEquiv (Σ-contractSnd λ _ → isOfHLevelPathP' 0 (propP _ b') _ _)) + ΣPath≃PathΣ) diff --git a/Cubical/Displayed/Properties.agda b/Cubical/Displayed/Properties.agda new file mode 100644 index 0000000000..8d59bf8dd1 --- /dev/null +++ b/Cubical/Displayed/Properties.agda @@ -0,0 +1,141 @@ +{-# OPTIONS --safe #-} +module Cubical.Displayed.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence using (pathToEquiv; univalence; ua-ungluePath-Equiv) + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.Relation.Binary +open BinaryRelation + +open import Cubical.Displayed.Base + +private + variable + ℓ ℓA ℓA' ℓP ℓ≅A ℓ≅A' ℓB ℓB' ℓ≅B ℓ≅B' ℓC ℓ≅C : Level + +-- induction principles + +module _ {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) where + open UARel 𝒮-A + 𝒮-J : {a : A} + (P : (a' : A) → (p : a ≡ a') → Type ℓ) + (d : P a refl) + {a' : A} + (p : a ≅ a') + → P a' (≅→≡ p) + 𝒮-J {a} P d {a'} p + = J (λ y q → P y q) + d + (≅→≡ p) + + 𝒮-J-2 : {a : A} + (P : (a' : A) → (p : a ≅ a') → Type ℓ) + (d : P a (ρ a)) + {a' : A} + (p : a ≅ a') + → P a' p + 𝒮-J-2 {a = a} P d {a'} p + = subst (λ r → P a' r) (Iso.leftInv (uaIso a a') p) g + where + g : P a' (≡→≅ (≅→≡ p)) + g = 𝒮-J (λ y q → P y (≡→≅ q)) d p + + +-- constructors + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} + (_≅ᴰ⟨_⟩_ : {a a' : A} → B a → UARel._≅_ 𝒮-A a a' → B a' → Type ℓ≅B) + where + + open UARel 𝒮-A + + -- constructor that reduces ua to the case where p = ρ a by induction on p + private + 𝒮ᴰ-make-aux : (uni : {a : A} (b b' : B a) → b ≅ᴰ⟨ ρ a ⟩ b' ≃ (b ≡ b')) + → ({a a' : A} (b : B a) (p : a ≅ a') (b' : B a') → (b ≅ᴰ⟨ p ⟩ b') ≃ PathP (λ i → B (≅→≡ p i)) b b') + 𝒮ᴰ-make-aux uni {a} {a'} b p + = 𝒮-J-2 𝒮-A + (λ y q → (b' : B y) → (b ≅ᴰ⟨ q ⟩ b') ≃ PathP (λ i → B (≅→≡ q i)) b b') + (λ b' → uni' b') + p + where + g : (b' : B a) → (b ≡ b') ≡ PathP (λ i → B (≅→≡ (ρ a) i)) b b' + g b' = subst (λ r → (b ≡ b') ≡ PathP (λ i → B (r i)) b b') + (sym (Iso.rightInv (uaIso a a) refl)) + refl + uni' : (b' : B a) → b ≅ᴰ⟨ ρ a ⟩ b' ≃ PathP (λ i → B (≅→≡ (ρ a) i)) b b' + uni' b' = compEquiv (uni b b') (pathToEquiv (g b')) + + 𝒮ᴰ-make-1 : (uni : {a : A} (b b' : B a) → b ≅ᴰ⟨ ρ a ⟩ b' ≃ (b ≡ b')) + → DUARel 𝒮-A B ℓ≅B + DUARel._≅ᴰ⟨_⟩_ (𝒮ᴰ-make-1 uni) = _≅ᴰ⟨_⟩_ + DUARel.uaᴰ (𝒮ᴰ-make-1 uni) = 𝒮ᴰ-make-aux uni + + -- constructor that reduces univalence further to contractibility of relational singletons + + 𝒮ᴰ-make-2 : (ρᴰ : {a : A} → isRefl _≅ᴰ⟨ ρ a ⟩_) + (contrTotal : (a : A) → contrRelSingl _≅ᴰ⟨ ρ a ⟩_) + → DUARel 𝒮-A B ℓ≅B + 𝒮ᴰ-make-2 ρᴰ contrTotal = 𝒮ᴰ-make-1 (contrRelSingl→isUnivalent _ ρᴰ (contrTotal _)) + +-- relational isomorphisms + +𝒮-iso→iso : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) + {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) + (F : RelIso (UARel._≅_ 𝒮-A) (UARel._≅_ 𝒮-B)) + → Iso A B +𝒮-iso→iso 𝒮-A 𝒮-B F + = RelIso→Iso (UARel._≅_ 𝒮-A) + (UARel._≅_ 𝒮-B) + (UARel.≅→≡ 𝒮-A) + (UARel.≅→≡ 𝒮-B) + F + +-- fiberwise relational isomorphisms + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {A' : Type ℓA'} {𝒮-A' : UARel A' ℓ≅A'} + (F : Iso A A') + {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {B' : A' → Type ℓB'} (𝒮ᴰ-B' : DUARel 𝒮-A' B' ℓ≅B') where + + open UARel 𝒮-A + open DUARel 𝒮ᴰ-B renaming (_≅ᴰ⟨_⟩_ to _≅B⟨_⟩_ + ; uaᴰ to uaB + ; fiberRel to fiberRelB + ; uaᴰρ to uaᴰρB) + open DUARel 𝒮ᴰ-B' renaming (_≅ᴰ⟨_⟩_ to _≅B'⟨_⟩_ + ; uaᴰ to uaB' + ; fiberRel to fiberRelB' + ; uaᴰρ to uaᴰρB') + + private + f = Iso.fun F + + -- the following can of course be done slightly more generally + -- for fiberwise binary relations + + F*fiberRelB' : (a : A) → Rel (B' (f a)) (B' (f a)) ℓ≅B' + F*fiberRelB' a = fiberRelB' (f a) + + module _ (G : (a : A) → RelIso (fiberRelB a) (F*fiberRelB' a)) where + private + fiberIsoOver : (a : A) → Iso (B a) (B' (f a)) + fiberIsoOver a + = RelIso→Iso (fiberRelB a) + (F*fiberRelB' a) + (equivFun (uaᴰρB _ _)) + (equivFun (uaᴰρB' _ _)) + (G a) + + -- DUARelFiberIsoOver→TotalIso produces an isomorphism of total spaces + -- from a relational isomorphism between B a and (F * B) a + 𝒮ᴰ-fiberIsoOver→totalIso : Iso (Σ A B) (Σ A' B') + 𝒮ᴰ-fiberIsoOver→totalIso = Σ-cong-iso F fiberIsoOver diff --git a/Cubical/Displayed/Record.agda b/Cubical/Displayed/Record.agda new file mode 100644 index 0000000000..bfec660d99 --- /dev/null +++ b/Cubical/Displayed/Record.agda @@ -0,0 +1,231 @@ +{- + +Generate univalent reflexive graph characterizations for record types from +characterizations of the field types using reflection. + +See end of file for an example. + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Record where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Data.Sigma +open import Cubical.Data.List as List +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Data.Maybe as Maybe + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Properties +open import Cubical.Displayed.Prop +open import Cubical.Displayed.Sigma +open import Cubical.Displayed.Unit +open import Cubical.Displayed.Universe +open import Cubical.Displayed.Auto + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base +import Cubical.Reflection.RecordEquiv as RE + +{- + `DUAFields` + A collection of DURG characterizations for fields of a record is described by an element of this inductive + family. If you just want to see how to use it, have a look at the end of the file first. + + An element of `DUAFields 𝒮-A R _≅R⟨_⟩_ πS 𝒮ᴰ-S πS≅` describes a mapping + - from a structure `R : A → Type _` and notion of structured equivalence `_≅R⟨_⟩_`, + which are meant to be defined as parameterized record types, + - to a DURG `𝒮ᴰ-S`, + the underlying structure of which will be an iterated Σ-type, + - via projections `πS` and `πS≅`. + + `𝒮-A`, `R`, and `_≅R⟨_⟩_` are parameters, while `πS`, `𝒮ᴰ-S`, and `πS≅` are indices; + the user builds up the Σ-type representation of the record using the DUAFields constructors. + + A DUAFields representation is _total_ when the projections `πS` and `πS≅` are equivalences, in which case we + obtain a DURG on `R` with `_≅R⟨_⟩_` as the notion of structured equivalence---see `𝒮ᴰ-Fields` below. + + When `R`, and `_≅R⟨_⟩_` are defined by record types, we can use reflection to automatically generate proofs + `πS` and `πS≅` are equivalences---see `𝒮ᴰ-Record` below. + +-} +data DUAFields {ℓA ℓ≅A ℓR ℓ≅R} {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) + (R : A → Type ℓR) (_≅R⟨_⟩_ : {a a' : A} → R a → UARel._≅_ 𝒮-A a a' → R a' → Type ℓ≅R) + : ∀ {ℓS ℓ≅S} {S : A → Type ℓS} + (πS : ∀ {a} → R a → S a) (𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S) + (πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → r ≅R⟨ e ⟩ r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')) + → Typeω + where + + -- `fields:` + -- Base case, no fields yet recorded in `𝒮ᴰ-S`. + fields: : DUAFields 𝒮-A R _≅R⟨_⟩_ (λ _ → tt) (𝒮ᴰ-Unit 𝒮-A) (λ _ → tt) + + -- `… data[ πF ∣ 𝒮ᴰ-F ∣ πF≅ ]` + -- Add a new field with a DURG. `πF` should be the name of the field in the structure record `R` and `πF≅` + -- the name of the corresponding field in the equivalence record `_≅R⟨_⟩_`, while `𝒮ᴰ-F` is a DURG for the + -- field's type over `𝒮-A`. Data fields that depend on previous fields of the record are not currently + -- supported. + _data[_∣_∣_] : ∀ {ℓS ℓ≅S} {S : A → Type ℓS} + {πS : ∀ {a} → R a → S a} {𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S} + {πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → r ≅R⟨ e ⟩ r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')} + → DUAFields 𝒮-A R _≅R⟨_⟩_ πS 𝒮ᴰ-S πS≅ + → ∀ {ℓF ℓ≅F} {F : A → Type ℓF} + (πF : ∀ {a} → (r : R a) → F a) + (𝒮ᴰ-F : DUARel 𝒮-A F ℓ≅F) + (πF≅ : ∀ {a} {r : R a} {e} {r' : R a} (p : r ≅R⟨ e ⟩ r') → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-F (πF r) e (πF r')) + → DUAFields 𝒮-A R _≅R⟨_⟩_ (λ r → πS r , πF r) (𝒮ᴰ-S ×𝒮ᴰ 𝒮ᴰ-F) (λ p → πS≅ p , πF≅ p) + + -- `… prop[ πF ∣ propF ]` + -- Add a new propositional field. `πF` should be the name of the field in the structure record `R`, while + -- propF is a proof that this field is a proposition. + _prop[_∣_] : ∀ {ℓS ℓ≅S} {S : A → Type ℓS} + {πS : ∀ {a} → R a → S a} {𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S} + {πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → r ≅R⟨ e ⟩ r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')} + → DUAFields 𝒮-A R _≅R⟨_⟩_ πS 𝒮ᴰ-S πS≅ + → ∀ {ℓF} {F : (a : A) → S a → Type ℓF} + (πF : ∀ {a} → (r : R a) → F a (πS r)) + (propF : ∀ a s → isProp (F a s)) + → DUAFields 𝒮-A R _≅R⟨_⟩_ (λ r → πS r , πF r) (𝒮ᴰ-subtype 𝒮ᴰ-S propF) (λ p → πS≅ p) + +module _ {ℓA ℓ≅A} {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {ℓR ℓ≅R} {R : A → Type ℓR} (_≅R⟨_⟩_ : {a a' : A} → R a → UARel._≅_ 𝒮-A a a' → R a' → Type ℓ≅R) + {ℓS ℓ≅S} {S : A → Type ℓS} + {πS : ∀ {a} → R a → S a} {𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S} + {πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → r ≅R⟨ e ⟩ r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')} + (fs : DUAFields 𝒮-A R _≅R⟨_⟩_ πS 𝒮ᴰ-S πS≅) + where + + open UARel 𝒮-A + open DUARel 𝒮ᴰ-S + + 𝒮ᴰ-Fields : + (e : ∀ a → Iso (R a) (S a)) + (e≅ : ∀ a a' (r : R a) p (r' : R a') → Iso (r ≅R⟨ p ⟩ r') ((e a .Iso.fun r ≅ᴰ⟨ p ⟩ e a' .Iso.fun r'))) + → DUARel 𝒮-A R ℓ≅R + DUARel._≅ᴰ⟨_⟩_ (𝒮ᴰ-Fields e e≅) r p r' = r ≅R⟨ p ⟩ r' + DUARel.uaᴰ (𝒮ᴰ-Fields e e≅) r p r' = + isoToEquiv + (compIso + (e≅ _ _ r p r') + (compIso + (equivToIso (uaᴰ (e _ .Iso.fun r) p (e _ .Iso.fun r'))) + (invIso (congPathIso λ i → isoToEquiv (e _))))) + +module DisplayedRecordMacro where + + -- Extract a name from a term + findName : R.Term → R.TC R.Name + findName t = + Maybe.rec + (R.typeError (R.strErr "Not a name: " ∷ R.termErr t ∷ [])) + (λ s → s) + (go t) + where + go : R.Term → Maybe (R.TC R.Name) + go (R.meta x _) = just (R.blockOnMeta x) + go (R.def name _) = just (R.returnTC name) + go (R.lam _ (R.abs _ t)) = go t + go t = nothing + + -- ℓA ℓ≅A ℓR ℓ≅R A 𝒮-A R _≅R⟨_⟩_ + pattern family∷ hole = _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ hole + + -- ℓS ℓ≅S S πS 𝒮ᴰ-S πS≅ + pattern indices∷ hole = _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ hole + + {- + Takes a reflected DUAFields term as input and collects lists of structure field names and equivalence + field names. (These are returned in reverse order. + -} + parseFields : R.Term → R.TC (List R.Name × List R.Name) + parseFields (R.con (quote fields:) _) = R.returnTC ([] , []) + parseFields (R.con (quote _data[_∣_∣_]) (family∷ (indices∷ (fs v∷ ℓF h∷ ℓ≅F h∷ F h∷ πF v∷ 𝒮ᴰ-F v∷ πF≅ v∷ _)))) = + parseFields fs >>= λ (fs , f≅s) → + findName πF >>= λ f → + findName πF≅ >>= λ f≅ → + R.returnTC (f ∷ fs , f≅ ∷ f≅s) + parseFields (R.con (quote _prop[_∣_]) (family∷ (indices∷ (fs v∷ ℓF h∷ F h∷ πF v∷ _)))) = + parseFields fs >>= λ (fs , f≅s) → + findName πF >>= λ f → + R.returnTC (f ∷ fs , f≅s) + parseFields (R.meta x _) = R.blockOnMeta x + parseFields t = R.typeError (R.strErr "Malformed specification: " ∷ R.termErr t ∷ []) + + {- + Given a list of record field names (in reverse order), generates a ΣFormat (in the sense of + Cubical.Reflection.RecordEquiv) associating the record fields with the fields of a left-associated + iterated Σ-type + -} + List→LeftAssoc : List R.Name → RE.ΣFormat + List→LeftAssoc [] = RE.unit + List→LeftAssoc (x ∷ xs) = List→LeftAssoc xs RE., RE.leaf x + + module _ {ℓA ℓ≅A} {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) + {ℓR ℓ≅R} {R : A → Type ℓR} (≅R : {a a' : A} → R a → UARel._≅_ 𝒮-A a a' → R a' → Type ℓ≅R) + {ℓS ℓ≅S} {S : A → Type ℓS} + {πS : ∀ {a} → R a → S a} {𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S} + {πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → ≅R r e r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')} + where + + {- + "𝒮ᴰ-Record ... : DUARel 𝒮-A R ℓ≅R" + Requires that `R` and `_≅R⟨_⟩_` are defined by records and `πS` and `πS≅` are equivalences. + The proofs of equivalence are generated using Cubical.Reflection.RecordEquiv and then + `𝒮ᴰ-Fields` is applied. + -} + 𝒮ᴰ-Record : DUAFields 𝒮-A R ≅R πS 𝒮ᴰ-S πS≅ → R.Term → R.TC Unit + 𝒮ᴰ-Record fs hole = + R.quoteTC (DUARel 𝒮-A R ℓ≅R) >>= R.checkType hole >>= λ hole → + R.quoteωTC fs >>= λ `fs` → + parseFields `fs` >>= λ (fields , ≅fields) → + R.freshName "fieldsIso" >>= λ fieldsIso → + R.freshName "≅fieldsIso" >>= λ ≅fieldsIso → + R.quoteTC R >>= R.normalise >>= λ `R` → + R.quoteTC {A = {a a' : A} → R a → UARel._≅_ 𝒮-A a a' → R a' → Type ℓ≅R} ≅R >>= R.normalise >>= λ `≅R` → + findName `R` >>= RE.declareRecordIsoΣ' fieldsIso (List→LeftAssoc fields) >> + findName `≅R` >>= RE.declareRecordIsoΣ' ≅fieldsIso (List→LeftAssoc ≅fields) >> + R.unify hole + (R.def (quote 𝒮ᴰ-Fields) + (`≅R` v∷ `fs` v∷ + vlam "_" (R.def fieldsIso []) v∷ + vlam "a" (vlam "a'" (vlam "r" (vlam "p" (vlam "r'" (R.def ≅fieldsIso []))))) v∷ + [])) + +macro + 𝒮ᴰ-Record = DisplayedRecordMacro.𝒮ᴰ-Record + +-- Example + +private + module Example where + + record Example (A : Type) : Type where + no-eta-equality -- works with or without eta equality + field + dog : A → A → A + cat : A → A → A + mouse : Unit + + open Example + + record ExampleEquiv {A B : Type} (x : Example A) (e : A ≃ B) (y : Example B) : Type where + no-eta-equality -- works with or without eta equality + field + dogEq : ∀ a a' → e .fst (x .dog a a') ≡ y .dog (e .fst a) (e .fst a') + catEq : ∀ a a' → e .fst (x .cat a a') ≡ y .cat (e .fst a) (e .fst a') + + open ExampleEquiv + + example : DUARel (𝒮-Univ ℓ-zero) Example ℓ-zero + example = + 𝒮ᴰ-Record (𝒮-Univ ℓ-zero) ExampleEquiv + (fields: + data[ dog ∣ autoDUARel _ _ ∣ dogEq ] + data[ cat ∣ autoDUARel _ _ ∣ catEq ] + prop[ mouse ∣ (λ _ _ → isPropUnit) ]) diff --git a/Cubical/Displayed/Sigma.agda b/Cubical/Displayed/Sigma.agda new file mode 100644 index 0000000000..8eb37612b4 --- /dev/null +++ b/Cubical/Displayed/Sigma.agda @@ -0,0 +1,101 @@ +{- + + Functions building UARels and DUARels on Σ-types + +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Sigma where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst +open import Cubical.Displayed.Morphism +open import Cubical.Displayed.Constant + +private + variable + ℓ ℓA ℓA' ℓP ℓ≅A ℓ≅A' ℓB ℓB' ℓ≅B ℓ≅B' ℓC ℓ≅C : Level + +-- UARel on a Σ-type + +∫ˢ : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : A → Type ℓB} (𝒮ˢ-B : SubstRel 𝒮-A B) + → UARel (Σ A B) (ℓ-max ℓ≅A ℓB) +∫ˢ 𝒮ˢ-B = ∫ (Subst→DUA 𝒮ˢ-B) + +_×𝒮_ : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) + → UARel (A × B) (ℓ-max ℓ≅A ℓ≅B) +𝒮-A ×𝒮 𝒮-B = ∫ (𝒮ᴰ-const 𝒮-A 𝒮-B) + +-- Projection UARel morphisms + +𝒮-fst : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : A → Type ℓB} {𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B} + → UARelHom (∫ 𝒮ᴰ-B) 𝒮-A +𝒮-fst .UARelHom.fun = fst +𝒮-fst .UARelHom.rel = fst +𝒮-fst .UARelHom.ua p = refl + +𝒮-snd : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} + → UARelHom (𝒮-A ×𝒮 𝒮-B) 𝒮-B +𝒮-snd .UARelHom.fun = snd +𝒮-snd .UARelHom.rel = snd +𝒮-snd .UARelHom.ua p = refl + +-- Lift a DUARel to live over a Σ-type + +𝒮ᴰ-Lift : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : A → Type ℓC} (𝒮ᴰ-C : DUARel 𝒮-A C ℓ≅C) + → DUARel (∫ 𝒮ᴰ-C) (λ (a , _) → B a) ℓ≅B +𝒮ᴰ-Lift _ 𝒮ᴰ-B _ = 𝒮ᴰ-reindex 𝒮-fst 𝒮ᴰ-B + +-- DUARel on a Σ-type + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} {ℓ≅B : Level} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : Σ A B → Type ℓC} {ℓ≅C : Level} (𝒮ᴰ-C : DUARel (∫ 𝒮ᴰ-B) C ℓ≅C) + where + + open UARel 𝒮-A + private + module B = DUARel 𝒮ᴰ-B + module C = DUARel 𝒮ᴰ-C + + 𝒮ᴰ-Σ : DUARel 𝒮-A (λ a → Σ[ b ∈ B a ] C (a , b)) (ℓ-max ℓ≅B ℓ≅C) + DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-Σ (b , c) p (b' , c') = + Σ[ q ∈ b B.≅ᴰ⟨ p ⟩ b' ] (c C.≅ᴰ⟨ p , q ⟩ c') + DUARel.uaᴰ 𝒮ᴰ-Σ (b , c) p (b' , c') = + compEquiv + (Σ-cong-equiv (B.uaᴰ b p b') (λ q → C.uaᴰ c (p , q) c')) + ΣPath≃PathΣ + +-- DUARel on a non-dependent Σ-type + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} {ℓ≅B : Level} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : A → Type ℓC} {ℓ≅C : Level} (𝒮ᴰ-C : DUARel 𝒮-A C ℓ≅C) + where + + _×𝒮ᴰ_ : DUARel 𝒮-A (λ a → B a × C a) (ℓ-max ℓ≅B ℓ≅C) + _×𝒮ᴰ_ = 𝒮ᴰ-Σ 𝒮ᴰ-B (𝒮ᴰ-Lift 𝒮-A 𝒮ᴰ-C 𝒮ᴰ-B) + +-- SubstRel on a Σ-type + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ˢ-B : SubstRel 𝒮-A B) + {C : Σ A B → Type ℓC} (𝒮ˢ-C : SubstRel (∫ˢ 𝒮ˢ-B) C) + where + + open UARel 𝒮-A + open SubstRel + private + module B = SubstRel 𝒮ˢ-B + module C = SubstRel 𝒮ˢ-C + + 𝒮ˢ-Σ : SubstRel 𝒮-A (λ a → Σ[ b ∈ B a ] C (a , b)) + 𝒮ˢ-Σ .act p = Σ-cong-equiv (B.act p) (λ b → C.act (p , refl)) + 𝒮ˢ-Σ .uaˢ p _ = + fromPathP + (DUARel.uaᴰ (𝒮ᴰ-Σ (Subst→DUA 𝒮ˢ-B) (Subst→DUA 𝒮ˢ-C)) _ p _ .fst (refl , refl)) diff --git a/Cubical/Displayed/Subst.agda b/Cubical/Displayed/Subst.agda new file mode 100644 index 0000000000..a05712db58 --- /dev/null +++ b/Cubical/Displayed/Subst.agda @@ -0,0 +1,60 @@ +{- + Given a type A with a UARel and a family B over A, + a SubstRel on B is a family of functions a ≅ a' → B a ≃ B a' path-equal to transport in that family. + + Any SubstRel gives rise to a DUARel in which b and b' are related over p when the transport of b along p is + equial to b'. +-} + +{-# OPTIONS --safe #-} +module Cubical.Displayed.Subst where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence + +open import Cubical.Displayed.Base + +private + variable + ℓA ℓ≅A ℓB : Level + +record SubstRel {A : Type ℓA} {ℓ≅A : Level} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) + : Type (ℓ-max (ℓ-max ℓA ℓB) ℓ≅A) + where + + no-eta-equality + constructor substrel + open UARel 𝒮-A + + field + act : {a a' : A} → a ≅ a' → B a ≃ B a' + uaˢ : {a a' : A} (p : a ≅ a') (b : B a) → subst B (≅→≡ p) b ≡ equivFun (act p) b + + uaˢ⁻ : {a a' : A} (p : a ≅ a') (b : B a') → subst B (sym (≅→≡ p)) b ≡ invEq (act p) b + uaˢ⁻ p b = + subst B (sym (≅→≡ p)) b + ≡⟨ cong (subst B (sym (≅→≡ p))) (sym (secEq (act p) b)) ⟩ + subst B (sym (≅→≡ p)) (equivFun (act p) (invEq (act p) b)) + ≡⟨ cong (subst B (sym (≅→≡ p))) (sym (uaˢ p (invEq (act p) b))) ⟩ + subst B (sym (≅→≡ p)) (subst B (≅→≡ p) (invEq (act p) b)) + ≡⟨ pathToIso (cong B (≅→≡ p)) .Iso.leftInv (invEq (act p) b) ⟩ + invEq (act p) b + ∎ + +Subst→DUA : {A : Type ℓA} {ℓ≅A : Level} {𝒮-A : UARel A ℓ≅A} {B : A → Type ℓB} + → SubstRel 𝒮-A B → DUARel 𝒮-A B ℓB +DUARel._≅ᴰ⟨_⟩_ (Subst→DUA 𝒮ˢ-B) b p b' = + equivFun (SubstRel.act 𝒮ˢ-B p) b ≡ b' +DUARel.uaᴰ (Subst→DUA {𝒮-A = 𝒮-A} {B = B} 𝒮ˢ-B) b p b' = + equivFun (SubstRel.act 𝒮ˢ-B p) b ≡ b' + ≃⟨ invEquiv (compPathlEquiv (sym (SubstRel.uaˢ 𝒮ˢ-B p b))) ⟩ + subst B (≅→≡ p) b ≡ b' + ≃⟨ invEquiv (PathP≃Path (λ i → B (≅→≡ p i)) b b') ⟩ + PathP (λ i → B (≅→≡ p i)) b b' + ■ + where + open UARel 𝒮-A diff --git a/Cubical/Displayed/Unit.agda b/Cubical/Displayed/Unit.agda new file mode 100644 index 0000000000..6d24f8539e --- /dev/null +++ b/Cubical/Displayed/Unit.agda @@ -0,0 +1,27 @@ +{- + + DUARel for the constant unit family + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Unit where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Unit + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Constant + +private + variable + ℓA ℓ≅A : Level + +𝒮-Unit : UARel Unit ℓ-zero +𝒮-Unit .UARel._≅_ _ _ = Unit +𝒮-Unit .UARel.ua _ _ = invEquiv (isContr→≃Unit (isProp→isContrPath isPropUnit _ _)) + +𝒮ᴰ-Unit : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) → DUARel 𝒮-A (λ _ → Unit) ℓ-zero +𝒮ᴰ-Unit 𝒮-A = 𝒮ᴰ-const 𝒮-A 𝒮-Unit diff --git a/Cubical/Displayed/Universe.agda b/Cubical/Displayed/Universe.agda new file mode 100644 index 0000000000..5a2da6f0ee --- /dev/null +++ b/Cubical/Displayed/Universe.agda @@ -0,0 +1,32 @@ +{- + + - UARel given by a universe and equivalences + - SubstRel and DUARel for the element family over the universe + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Universe where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst + +private + variable + ℓA ℓ≅A ℓB ℓ≅B ℓP : Level + +𝒮-Univ : ∀ ℓ → UARel (Type ℓ) ℓ +𝒮-Univ ℓ .UARel._≅_ = _≃_ +𝒮-Univ ℓ .UARel.ua _ _ = isoToEquiv (invIso univalenceIso) + +𝒮ˢ-El : ∀ ℓ → SubstRel (𝒮-Univ ℓ) (λ X → X) +𝒮ˢ-El ℓ .SubstRel.act e = e +𝒮ˢ-El ℓ .SubstRel.uaˢ e a = uaβ e a + +𝒮ᴰ-El : ∀ ℓ → DUARel (𝒮-Univ ℓ) (λ X → X) ℓ +𝒮ᴰ-El ℓ .DUARel._≅ᴰ⟨_⟩_ a e a' = e .fst a ≡ a' +𝒮ᴰ-El ℓ .DUARel.uaᴰ a e a' = invEquiv (ua-ungluePath-Equiv e) diff --git a/Cubical/Experiments/Brunerie.agda b/Cubical/Experiments/Brunerie.agda new file mode 100644 index 0000000000..c7f1303506 --- /dev/null +++ b/Cubical/Experiments/Brunerie.agda @@ -0,0 +1,298 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.Brunerie where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Bool +open import Cubical.Data.Nat +open import Cubical.Data.Int +open import Cubical.Data.HomotopyGroup +open import Cubical.HITs.S1 +open import Cubical.HITs.S2 +open import Cubical.HITs.S3 +open import Cubical.HITs.Join +open import Cubical.HITs.SetTruncation as SetTrunc +open import Cubical.HITs.GroupoidTruncation as GroupoidTrunc +open import Cubical.HITs.2GroupoidTruncation as 2GroupoidTrunc +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Hopf +open S¹Hopf + +-- This code is adapted from examples/brunerie3.ctt on the pi4s3_nobug branch of cubicaltt + +Bool∙ S¹∙ S³∙ : Pointed₀ +Bool∙ = (Bool , true) +S¹∙ = (S¹ , base) +S³∙ = (S³ , base) + +∥_∥₃∙ ∥_∥₄∙ : Pointed₀ → Pointed₀ +∥ A , a ∥₃∙ = ∥ A ∥₃ , ∣ a ∣₃ +∥ A , a ∥₄∙ = ∥ A ∥₄ , ∣ a ∣₄ + +join∙ : Pointed₀ → Type₀ → Pointed₀ +join∙ (A , a) B = join A B , inl a + +Ω² Ω³ : Pointed₀ → Pointed₀ +Ω² = Ω^ 2 +Ω³ = Ω^ 3 + +mapΩrefl : {A : Pointed₀} {B : Type₀} (f : A .fst → B) → Ω A .fst → Ω (B , f (pt A)) .fst +mapΩrefl f p i = f (p i) + +mapΩ²refl : {A : Pointed₀} {B : Type₀} (f : A .fst → B) → Ω² A .fst → Ω² (B , f (pt A)) .fst +mapΩ²refl f p i j = f (p i j) + +mapΩ³refl : {A : Pointed₀} {B : Type₀} (f : A .fst → B) → Ω³ A .fst → Ω³ (B , f (pt A)) .fst +mapΩ³refl f p i j k = f (p i j k) + +meridS² : S¹ → Path S² base base +meridS² base _ = base +meridS² (loop i) j = surf i j + +alpha : join S¹ S¹ → S² +alpha (inl x) = base +alpha (inr y) = base +alpha (push x y i) = (meridS² y ∙ meridS² x) i + +connectionBoth : {A : Type₀} {a : A} (p : Path A a a) → PathP (λ i → Path A (p i) (p i)) p p +connectionBoth {a = a} p i j = + hcomp + (λ k → λ + { (i = i0) → p (j ∨ ~ k) + ; (i = i1) → p (j ∧ k) + ; (j = i0) → p (i ∨ ~ k) + ; (j = i1) → p (i ∧ k) + }) + a + +data PostTotalHopf : Type₀ where + base : S¹ → PostTotalHopf + loop : (x : S¹) → PathP (λ i → Path PostTotalHopf (base x) (base (rotLoop x (~ i)))) refl refl + +tee12 : (x : S²) → HopfS² x → PostTotalHopf +tee12 base y = base y +tee12 (surf i j) y = + hcomp + (λ k → λ + { (i = i0) → base y + ; (i = i1) → base y + ; (j = i0) → base y + ; (j = i1) → base (rotLoopInv y (~ i) k) + }) + (loop (unglue (i ∨ ~ i ∨ j ∨ ~ j) y) i j) + +tee34 : PostTotalHopf → join S¹ S¹ +tee34 (base x) = inl x +tee34 (loop x i j) = + hcomp + (λ k → λ + { (i = i0) → push x x (j ∧ ~ k) + ; (i = i1) → push x x (j ∧ ~ k) + ; (j = i0) → inl x + ; (j = i1) → push (rotLoop x (~ i)) x (~ k) + }) + (push x x j) + +tee : (x : S²) → HopfS² x → join S¹ S¹ +tee x y = tee34 (tee12 x y) + +fibΩ : {B : Pointed₀} (P : B .fst → Type₀) → P (pt B) → Ω B .fst → Type₀ +fibΩ P f p = PathP (λ i → P (p i)) f f + +fibΩ² : {B : Pointed₀} (P : B .fst → Type₀) → P (pt B) → Ω² B .fst → Type₀ +fibΩ² P f = fibΩ (fibΩ P f) refl + +fibΩ³ : {B : Pointed₀} (P : B .fst → Type₀) → P (pt B) → Ω³ B .fst → Type₀ +fibΩ³ P f = fibΩ² (fibΩ P f) refl + +Ω³Hopf : Ω³ S²∙ .fst → Type₀ +Ω³Hopf = fibΩ³ HopfS² base + +fibContrΩ³Hopf : ∀ p → Ω³Hopf p +fibContrΩ³Hopf p i j k = + hcomp + (λ m → λ + { (i = i0) → base + ; (i = i1) → base + ; (j = i0) → base + ; (j = i1) → base + ; (k = i0) → base + ; (k = i1) → + isSetΩS¹ refl refl + (λ i j → transp (λ n → HopfS² (p i j n)) (i ∨ ~ i ∨ j ∨ ~ j) base) + (λ _ _ → base) + m i j + }) + (transp (λ n → HopfS² (p i j (k ∧ n))) (i ∨ ~ i ∨ j ∨ ~ j ∨ ~ k) base) + +h : Ω³ S²∙ .fst → Ω³ (join∙ S¹∙ S¹) .fst +h p i j k = tee (p i j k) (fibContrΩ³Hopf p i j k) + +multTwoAux : (x : S²) → Path (Path ∥ S² ∥₄ ∣ x ∣₄ ∣ x ∣₄) refl refl +multTwoAux base i j = ∣ surf i j ∣₄ +multTwoAux (surf k l) i j = + hcomp + (λ m → λ + { (i = i0) → ∣ surf k l ∣₄ + ; (i = i1) → ∣ surf k l ∣₄ + ; (j = i0) → ∣ surf k l ∣₄ + ; (j = i1) → ∣ surf k l ∣₄ + ; (k = i0) → ∣ surf i j ∣₄ + ; (k = i1) → ∣ surf i j ∣₄ + ; (l = i0) → ∣ surf i j ∣₄ + ; (l = i1) → squash₄ _ _ _ _ _ _ (λ k i j → step₁ k i j) refl m k i j + }) + (step₁ k i j) + + where + step₁ : I → I → I → ∥ S² ∥₄ + step₁ k i j = + hcomp {A = ∥ S² ∥₄} + (λ m → λ + { (i = i0) → ∣ surf k (l ∧ m) ∣₄ + ; (i = i1) → ∣ surf k (l ∧ m) ∣₄ + ; (j = i0) → ∣ surf k (l ∧ m) ∣₄ + ; (j = i1) → ∣ surf k (l ∧ m) ∣₄ + ; (k = i0) → ∣ surf i j ∣₄ + ; (k = i1) → ∣ surf i j ∣₄ + ; (l = i0) → ∣ surf i j ∣₄ + }) + ∣ surf i j ∣₄ + +multTwoTildeAux : (t : ∥ S² ∥₄) → Path (Path ∥ S² ∥₄ t t) refl refl +multTwoTildeAux ∣ x ∣₄ = multTwoAux x +multTwoTildeAux (squash₄ _ _ _ _ _ _ t u k l m n) i j = + squash₄ _ _ _ _ _ _ + (λ k l m → multTwoTildeAux (t k l m) i j) + (λ k l m → multTwoTildeAux (u k l m) i j) + k l m n + +multTwoEquivAux : Path (Path (∥ S² ∥₄ ≃ ∥ S² ∥₄) (idEquiv _) (idEquiv _)) refl refl +multTwoEquivAux i j = + ( f i j + , hcomp + (λ l → λ + { (i = i0) → isPropIsEquiv _ (idIsEquiv _) (idIsEquiv _) l + ; (i = i1) → isPropIsEquiv _ (idIsEquiv _) (idIsEquiv _) l + ; (j = i0) → isPropIsEquiv _ (idIsEquiv _) (idIsEquiv _) l + ; (j = i1) → + isPropIsEquiv _ + (transp (λ k → isEquiv (f i k)) (i ∨ ~ i) (idIsEquiv _)) + (idIsEquiv _) + l + }) + (transp (λ k → isEquiv (f i (j ∧ k))) (i ∨ ~ i ∨ ~ j) (idIsEquiv _)) + ) + where + f : I → I → ∥ S² ∥₄ → ∥ S² ∥₄ + f i j t = multTwoTildeAux t i j + +tHopf³ : S³ → Type₀ +tHopf³ base = ∥ S² ∥₄ +tHopf³ (surf i j k) = + Glue ∥ S² ∥₄ + (λ { (i = i0) → (∥ S² ∥₄ , idEquiv _) + ; (i = i1) → (∥ S² ∥₄ , idEquiv _) + ; (j = i0) → (∥ S² ∥₄ , idEquiv _) + ; (j = i1) → (∥ S² ∥₄ , idEquiv _) + ; (k = i0) → (∥ S² ∥₄ , multTwoEquivAux i j) + ; (k = i1) → (∥ S² ∥₄ , idEquiv _) + }) + +π₃S³ : Ω³ S³∙ .fst → Ω² ∥ S²∙ ∥₄∙ .fst +π₃S³ p i j = transp (λ k → tHopf³ (p j k i)) i0 ∣ base ∣₄ + +codeS² : S² → hGroupoid _ +codeS² s = ∥ HopfS² s ∥₃ , squash₃ + +codeTruncS² : ∥ S² ∥₄ → hGroupoid _ +codeTruncS² = 2GroupoidTrunc.rec (isOfHLevelTypeOfHLevel 3) codeS² + +encodeTruncS² : Ω ∥ S²∙ ∥₄∙ .fst → ∥ S¹ ∥₃ +encodeTruncS² p = transp (λ i → codeTruncS² (p i) .fst) i0 ∣ base ∣₃ + +codeS¹ : S¹ → hSet _ +codeS¹ s = ∥ helix s ∥₂ , squash₂ + +codeTruncS¹ : ∥ S¹ ∥₃ → hSet _ +codeTruncS¹ = GroupoidTrunc.rec (isOfHLevelTypeOfHLevel 2) codeS¹ + +encodeTruncS¹ : Ω ∥ S¹∙ ∥₃∙ .fst → ∥ ℤ ∥₂ +encodeTruncS¹ p = transp (λ i → codeTruncS¹ (p i) .fst) i0 ∣ pos zero ∣₂ + + +-- THE BIG GAME + +f3 : Ω³ S³∙ .fst → Ω³ (join∙ S¹∙ S¹) .fst +f3 = mapΩ³refl S³→joinS¹S¹ + +f4 : Ω³ (join∙ S¹∙ S¹) .fst → Ω³ S²∙ .fst +f4 = mapΩ³refl alpha + +f5 : Ω³ S²∙ .fst → Ω³ (join∙ S¹∙ S¹) .fst +f5 = h + +f6 : Ω³ (join∙ S¹∙ S¹) .fst → Ω³ S³∙ .fst +f6 = mapΩ³refl joinS¹S¹→S³ + +f7 : Ω³ S³∙ .fst → Ω² ∥ S²∙ ∥₄∙ .fst +f7 = π₃S³ + +g8 : Ω² ∥ S²∙ ∥₄∙ .fst → Ω ∥ S¹∙ ∥₃∙ .fst +g8 = mapΩrefl encodeTruncS² + +g9 : Ω ∥ S¹∙ ∥₃∙ .fst → ∥ ℤ ∥₂ +g9 = encodeTruncS¹ + +g10 : ∥ ℤ ∥₂ → ℤ +g10 = SetTrunc.rec isSetℤ (idfun ℤ) + +-- don't run me +brunerie : ℤ +brunerie = g10 (g9 (g8 (f7 (f6 (f5 (f4 (f3 (λ i j k → surf i j k)))))))) + +-- simpler tests + +test63 : ℕ → ℤ +test63 n = g10 (g9 (g8 (f7 (63n n)))) + where + 63n : ℕ → Ω³ S³∙ .fst + 63n zero i j k = surf i j k + 63n (suc n) = f6 (f3 (63n n)) + +foo : Ω³ S²∙ .fst +foo i j k = + hcomp + (λ l → λ + { (i = i0) → surf l l + ; (i = i1) → surf l l + ; (j = i0) → surf l l + ; (j = i1) → surf l l + ; (k = i0) → surf l l + ; (k = i1) → surf l l + }) + base + +sorghum : Ω³ S²∙ .fst +sorghum i j k = + hcomp + (λ l → λ + { (i = i0) → surf j l + ; (i = i1) → surf k (~ l) + ; (j = i0) → surf k (i ∧ ~ l) + ; (j = i1) → surf k (i ∧ ~ l) + ; (k = i0) → surf j (i ∨ l) + ; (k = i1) → surf j (i ∨ l) + }) + (hcomp + (λ l → λ + { (i = i0) → base + ; (i = i1) → surf j l + ; (j = i0) → surf k i + ; (j = i1) → surf k i + ; (k = i0) → surf j (i ∧ l) + ; (k = i1) → surf j (i ∧ l) + }) + (surf k i)) + +goo : Ω³ S²∙ .fst → ℤ +goo x = g10 (g9 (g8 (f7 (f6 (f5 x))))) diff --git a/Cubical/Experiments/CohomologyGroups.agda b/Cubical/Experiments/CohomologyGroups.agda new file mode 100644 index 0000000000..1f2403aac3 --- /dev/null +++ b/Cubical/Experiments/CohomologyGroups.agda @@ -0,0 +1,140 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.CohomologyGroups where + +open import Cubical.Experiments.ZCohomologyOld.Base +open import Cubical.Experiments.ZCohomologyOld.Properties +open import Cubical.Experiments.ZCohomologyOld.MayerVietorisUnreduced +open import Cubical.Experiments.ZCohomologyOld.Groups.Unit +open import Cubical.Experiments.ZCohomologyOld.KcompPrelims +open import Cubical.Experiments.ZCohomologyOld.Groups.Sn + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Isomorphism + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) + hiding (map) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) + hiding (map) + +open import Cubical.Data.Bool +open import Cubical.Data.Sigma +open import Cubical.Data.Int + +open import Cubical.Algebra.Group + +open GroupIso +open GroupHom +open BijectionIso + +-- --------------------------H¹(S¹) ----------------------------------- +{- +In order to apply Mayer-Vietoris, we need the following lemma. +Given the following diagram + a ↦ (a , 0) ψ ϕ + A --> A × A -------> B ---> C +If ψ is an isomorphism and ϕ is surjective with ker ϕ ≡ {ψ (a , a) ∣ a ∈ A}, then C ≅ B +-} + + +diagonalIso : ∀ {ℓ ℓ' ℓ''} {A : Group {ℓ}} (B : Group {ℓ'}) {C : Group {ℓ''}} + (ψ : GroupIso (dirProd A A) B) (ϕ : GroupHom B C) + → isSurjective _ _ ϕ + → ((x : ⟨ B ⟩) → isInKer B C ϕ x + → ∃[ y ∈ ⟨ A ⟩ ] x ≡ (fun (map ψ)) (y , y)) + → ((x : ⟨ B ⟩) → (∃[ y ∈ ⟨ A ⟩ ] x ≡ (fun (map ψ)) (y , y)) + → isInKer B C ϕ x) + → GroupIso A C +diagonalIso {A = A} B {C = C} ψ ϕ issurj ker→diag diag→ker = BijectionIsoToGroupIso bijIso + where + open GroupStr + module A = GroupStr (snd A) + module B = GroupStr (snd B) + module C = GroupStr (snd C) + module A×A = GroupStr (snd (dirProd A A)) + module ψ = GroupIso ψ + module ϕ = GroupHom ϕ + ψ⁻ = inv ψ + + fstProj : GroupHom A (dirProd A A) + fun fstProj a = a , GroupStr.0g (snd A) + isHom fstProj g0 g1 i = (g0 A.+ g1) , GroupStr.lid (snd A) (GroupStr.0g (snd A)) (~ i) + + bijIso : BijectionIso A C + map' bijIso = compGroupHom fstProj (compGroupHom (map ψ) ϕ) + inj bijIso a inker = pRec (isSetCarrier A _ _) + (λ {(a' , id) → (cong fst (sym (leftInv ψ (a , GroupStr.0g (snd A))) ∙∙ cong ψ⁻ id ∙∙ leftInv ψ (a' , a'))) + ∙ cong snd (sym (leftInv ψ (a' , a')) ∙∙ cong ψ⁻ (sym id) ∙∙ leftInv ψ (a , GroupStr.0g (snd A)))}) + (ker→diag _ inker) + surj bijIso c = + pRec isPropPropTrunc + (λ { (b , id) → ∣ (fst (ψ⁻ b) A.+ (A.- snd (ψ⁻ b))) + , ((sym (GroupStr.rid (snd C) _) + ∙∙ cong ((fun ϕ) ((fun (map ψ)) (fst (ψ⁻ b) A.+ (A.- snd (ψ⁻ b)) , GroupStr.0g (snd A))) C.+_) + (sym (diag→ker (fun (map ψ) ((snd (ψ⁻ b)) , (snd (ψ⁻ b)))) + ∣ (snd (ψ⁻ b)) , refl ∣₁)) + ∙∙ sym ((isHom ϕ) _ _)) + ∙∙ cong (fun ϕ) (sym ((isHom (map ψ)) _ _) + ∙∙ cong (fun (map ψ)) (ΣPathP (sym (GroupStr.assoc (snd A) _ _ _) + ∙∙ cong (fst (ψ⁻ b) A.+_) (GroupStr.invl (snd A) _) + ∙∙ GroupStr.rid (snd A) _ + , (GroupStr.lid (snd A) _))) + ∙∙ rightInv ψ b) + ∙∙ id) ∣₁ }) + (issurj c) + +H¹-S¹≅ℤ : GroupIso intGroup (coHomGr 1 (S₊ 1)) +H¹-S¹≅ℤ = + diagonalIso (coHomGr 0 (S₊ 0)) + (invGroupIso H⁰-S⁰≅ℤ×ℤ) + (K.d 0) + (λ x → K.Ker-i⊂Im-d 0 x + (ΣPathP (isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _ + , isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _))) + ((sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + (λ x inker + → pRec isPropPropTrunc + (λ {((f , g) , id') → helper x f g id' inker}) + ((K.Ker-d⊂Im-Δ 0 ∣ x ∣₂ inker))))) + ((sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ F surj + → pRec (isSetSetTrunc _ _) + (λ { (x , id) → K.Im-Δ⊂Ker-d 0 ∣ F ∣₂ + ∣ (∣ (λ _ → x) ∣₂ , ∣ (λ _ → 0) ∣₂) , + (cong ∣_∣₂ (funExt (surjHelper x))) ∙ sym id ∣₁ }) + surj) ) + □ invGroupIso (coHomPushout≅coHomSn 0 1) + where + module K = MV Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) + + surjHelper : (x : Int) (x₁ : S₊ 0) → x -[ 0 ]ₖ 0 ≡ S0→Int (x , x) x₁ + surjHelper x true = Iso.leftInv (Iso-Kn-ΩKn+1 0) x + surjHelper x false = Iso.leftInv (Iso-Kn-ΩKn+1 0) x + + helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₂) + (id : GroupHom.fun (K.Δ 0) (f , g) ≡ ∣ F ∣₂) + → isInKer (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) + (K.d 0) + ∣ F ∣₂ + → ∃[ x ∈ Int ] ∣ F ∣₂ ≡ inv H⁰-S⁰≅ℤ×ℤ (x , x) + helper F = + sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + λ f g id inker + → pRec isPropPropTrunc + (λ ((a , b) , id2) + → sElim2 {C = λ f g → GroupHom.fun (K.Δ 0) (f , g) ≡ ∣ F ∣₂ → _ } + (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₁) + a b id2) + (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₂ inker) + where + helper2 : (f g : Unit → Int) + → Σ[ x ∈ Int ] (inv H⁰-S⁰≅ℤ×ℤ (x , x)) + ≡ GroupHom.fun (K.Δ 0) (∣ f ∣₂ , ∣ g ∣₂) + helper2 f g = (f _ -[ 0 ]ₖ g _) , cong ∣_∣₂ (funExt λ {true → refl ; false → refl}) diff --git a/Cubical/Experiments/Combinatorics.agda b/Cubical/Experiments/Combinatorics.agda new file mode 100644 index 0000000000..2c45bda19e --- /dev/null +++ b/Cubical/Experiments/Combinatorics.agda @@ -0,0 +1,59 @@ +{- + +Computable stuff constructed from the Combinatorics of Finite Sets + +-} +{-# OPTIONS --safe #-} + +module Cubical.Experiments.Combinatorics where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Nat +open import Cubical.Data.Sum +open import Cubical.Data.Sigma + +open import Cubical.Data.Fin +open import Cubical.Data.FinSet +open import Cubical.Data.FinSet.Constructors + +open import Cubical.Functions.Embedding +open import Cubical.Functions.Surjection + +-- some computable functions + +-- this should be addtion +card+ : ℕ → ℕ → ℕ +card+ m n = card (Fin m ⊎ Fin n , isFinSet⊎ (Fin m , isFinSetFin) (Fin n , isFinSetFin)) + +-- this should be multiplication +card× : ℕ → ℕ → ℕ +card× m n = card (Fin m × Fin n , isFinSet× (Fin m , isFinSetFin) (Fin n , isFinSetFin)) + +-- this should be exponential +card→ : ℕ → ℕ → ℕ +card→ m n = card ((Fin m → Fin n) , isFinSet→ (Fin m , isFinSetFin) (Fin n , isFinSetFin)) + +-- this should be factorial or zero +card≃ : ℕ → ℕ → ℕ +card≃ m n = card ((Fin m ≃ Fin n) , isFinSet≃ (Fin m , isFinSetFin) (Fin n , isFinSetFin)) + +-- this should be binomial coeffient +card↪ : ℕ → ℕ → ℕ +card↪ m n = card ((Fin m ↪ Fin n) , isFinSet↪ (Fin m , isFinSetFin) (Fin n , isFinSetFin)) + +-- this should be something that I don't know the name +card↠ : ℕ → ℕ → ℕ +card↠ m n = card ((Fin m ↠ Fin n) , isFinSet↠ (Fin m , isFinSetFin) (Fin n , isFinSetFin)) + +-- explicit numbers + +s2 : card≃ 2 2 ≡ 2 +s2 = refl + +s3 : card≃ 3 3 ≡ 6 +s3 = refl + +c3,2 : card↪ 2 3 ≡ 6 +c3,2 = refl diff --git a/Cubical/Experiments/EscardoSIP.agda b/Cubical/Experiments/EscardoSIP.agda new file mode 100644 index 0000000000..56f3acfa6d --- /dev/null +++ b/Cubical/Experiments/EscardoSIP.agda @@ -0,0 +1,218 @@ +{- +This is a rather literal translation of Martin Hötzel-Escardó's structure identity principle into cubical Agda. See +https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#sns + +All the needed preliminary results from the lecture notes are stated and proven in this file. +It would be interesting to compare the proves with the one in Cubical.Foundations.SIP +-} +{-# OPTIONS --safe #-} +module Cubical.Experiments.EscardoSIP where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Data.Sigma.Properties + +private + variable + ℓ ℓ' ℓ'' : Level + S : Type ℓ → Type ℓ' + +-- We prove several useful equalities and equivalences between Σ-types all the proofs are taken from +-- Martin Hötzel-Escardó's lecture notes. + +-- The next result is just a reformulation of pathSigma≡sigmaPath from Sigma.Properties. + +Σ-≡-≃ : {X : Type ℓ} {A : X → Type ℓ'} + → (σ τ : Σ X A) → ((σ ≡ τ) ≃ (Σ[ p ∈ (σ .fst) ≡ (τ .fst) ] (subst A p (σ .snd) ≡ (τ .snd)))) +Σ-≡-≃ {A = A} σ τ = invEquiv (ΣPathTransport≃PathΣ σ τ) + + + +-- This cubical proof is much shorter than in HoTT but requires that A, B live in the same universe. +Σ-cong : {X : Type ℓ} {A B : X → Type ℓ'} → + ((x : X) → (A x ≡ B x)) → (Σ X A ≡ Σ X B) +Σ-cong {X = X} p i = Σ[ x ∈ X ] (p x i) + +-- Two lemmas for the more general formulation using equivalences +NatΣ : {X : Type ℓ} {A : X → Type ℓ'} {B : X → Type ℓ''} + → ((x : X) → (A x) → (B x)) → (Σ X A) → (Σ X B) +NatΣ τ (x , a) = (x , τ x a) + +Σ-to-PathP : {X : Type ℓ} {A : X → Type ℓ'} {x : X} {a b : A x} + → (a ≡ b) → PathP (λ i → Σ X A) (x , a) (x , b) +Σ-to-PathP {x = x} p i = (x , p i) + + +Σ-cong-≃ : {X : Type ℓ} {A : X → Type ℓ'} {B : X → Type ℓ''} → + ((x : X) → (A x ≃ B x)) → (Σ X A ≃ Σ X B) +Σ-cong-≃ {X = X} {A = A} {B = B} φ = isoToEquiv (iso (NatΣ f) (NatΣ g) NatΣ-ε NatΣ-η) + where + f : (x : X) → (A x) → (B x) + f x = equivFun (φ x) + + g : (x : X) → (B x) → (A x) + g x = equivFun (invEquiv (φ x)) + + η : (x : X) → (a : A x) → (g x) ((f x) a) ≡ a + η x = secEq (invEquiv (φ x)) + + ε : (x : X) → (b : B x) → f x (g x b) ≡ b + ε x = retEq (invEquiv (φ x)) + + NatΣ-η : (w : Σ X A) → NatΣ g (NatΣ f w) ≡ w + NatΣ-η (x , a) = (x , g x (f x a)) ≡⟨ Σ-to-PathP (η x a) ⟩ + (x , a) ∎ + + NatΣ-ε : (u : Σ X B) → NatΣ f (NatΣ g u) ≡ u + NatΣ-ε (x , b) = (x , f x (g x b)) ≡⟨ Σ-to-PathP (ε x b) ⟩ + (x , b) ∎ + + + +-- The next result is stated a bit awkwardly but is rather straightforward to prove. +Σ-change-of-variable-Iso : {X : Type ℓ} {Y : Type ℓ'} {A : Y → Type ℓ''} (f : X → Y) + → (isHAEquiv f) → (Iso (Σ X (A ∘ f)) (Σ Y A)) +Σ-change-of-variable-Iso {ℓ = ℓ} {ℓ' = ℓ'} {X = X} {Y = Y} {A = A} f isHAEquivf = iso φ ψ φψ ψφ + where + g : Y → X + g = isHAEquiv.g isHAEquivf + ε : (x : X) → (g (f x)) ≡ x + ε = isHAEquiv.linv isHAEquivf + η : (y : Y) → f (g y) ≡ y + η = isHAEquiv.rinv isHAEquivf + τ : (x : X) → cong f (ε x) ≡ η (f x) + τ = isHAEquiv.com isHAEquivf + + φ : (Σ X (A ∘ f)) → (Σ Y A) + φ (x , a) = (f x , a) + + ψ : (Σ Y A) → (Σ X (A ∘ f)) + ψ (y , a) = (g y , subst A (sym (η y)) a) + + φψ : (z : (Σ Y A)) → φ (ψ z) ≡ z + φψ (y , a) = + ΣPathTransport→PathΣ _ _ (η y , transportTransport⁻ (λ i → A (η y i)) a) + -- last term proves transp (λ i → A (η y i)) i0 (transp (λ i → A (η y (~ i))) i0 a) ≡ a + + ψφ : (z : (Σ X (A ∘ f))) → ψ (φ z) ≡ z + ψφ (x , a) = ΣPathTransport→PathΣ _ _ (ε x , q) + where + b : A (f (g (f x))) + b = (transp (λ i → A (η (f x) (~ i))) i0 a) + + q : transp (λ i → A (f (ε x i))) i0 (transp (λ i → A (η (f x) (~ i))) i0 a) ≡ a + q = transp (λ i → A (f (ε x i))) i0 b ≡⟨ i ⟩ + transp (λ i → A (η (f x) i)) i0 b ≡⟨ transportTransport⁻ (λ i → A (η (f x) i)) a ⟩ + a ∎ + where + i : (transp (λ i → A (f (ε x i))) i0 b) ≡ (transp (λ i → A (η (f x) i)) i0 b) + i = subst (λ p → (transp (λ i → A (f (ε x i))) i0 b) ≡ (transp (λ i → A (p i)) i0 b)) + (τ x) refl + + +-- Using the result above we can prove the following quite useful result. +Σ-change-of-variable-≃ : {X : Type ℓ} {Y : Type ℓ'} {A : Y → Type ℓ''} (f : X → Y) + → (isEquiv f) → ((Σ X (A ∘ f)) ≃ (Σ Y A)) +Σ-change-of-variable-≃ f isEquivf = + isoToEquiv (Σ-change-of-variable-Iso f (equiv→HAEquiv (f , isEquivf) .snd)) + + +-- A structure is a type-family S : Type ℓ → Type ℓ', i.e. for X : Type ℓ and s : S X, the pair (X , s) +-- means that X is equipped with a S-structure, which is witnessed by s. +-- An S-structure should have a notion of S-homomorphism, or rather S-isomorphism. +-- This will be implemented by a function ι +-- that gives us for any two types with S-structure (X , s) and (Y , t) a family: +-- ι (X , s) (Y , t) : (X ≃ Y) → Type ℓ'' +-- Note that for any equivalence (f , e) : X ≃ Y the type ι (X , s) (Y , t) (f , e) need not to be +-- a proposition. Indeed this type should correspond to the ways s and t can be identified +-- as S-structures. This we call a standard notion of structure. + + +SNS : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') → Type (ℓ-max (ℓ-max (ℓ-suc ℓ) ℓ') ℓ'') +SNS {ℓ = ℓ} S ι = ∀ {X : (Type ℓ)} (s t : S X) → ((s ≡ t) ≃ ι (X , s) (X , t) (idEquiv X)) + + +-- Escardo's ρ can actually be defined from this: +ρ : {ι : StrEquiv S ℓ''} (θ : SNS S ι) (A : TypeWithStr ℓ S) → (ι A A (idEquiv (typ A))) +ρ θ A = equivFun (θ (str A) (str A)) refl + +-- We introduce the notation a bit differently: +_≃[_]_ : (A : TypeWithStr ℓ S) (ι : StrEquiv S ℓ'') (B : TypeWithStr ℓ S) → (Type (ℓ-max ℓ ℓ'')) +A ≃[ ι ] B = Σ[ f ∈ ((typ A) ≃ (typ B)) ] (ι A B f) + + + +Id→homEq : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') + → (ρ : (A : TypeWithStr ℓ S) → ι A A (idEquiv (typ A))) + → (A B : TypeWithStr ℓ S) → A ≡ B → (A ≃[ ι ] B) +Id→homEq S ι ρ A B p = J (λ y x → A ≃[ ι ] y) (idEquiv (typ A) , ρ A) p + + +-- Use a PathP version of Escardó's homomorphism-lemma +hom-lemma-dep : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') (θ : SNS S ι) + → (A B : TypeWithStr ℓ S) + → (p : (typ A) ≡ (typ B)) + → (PathP (λ i → S (p i)) (str A) (str B)) ≃ (ι A B (pathToEquiv p)) +hom-lemma-dep S ι θ A B p = (J P (λ s → γ s) p) (str B) + where + P = (λ y x → (s : S y) → PathP (λ i → S (x i)) (str A) s ≃ ι A (y , s) (pathToEquiv x)) + + γ : (s : S (typ A)) → ((str A) ≡ s) ≃ ι A ((typ A) , s) (pathToEquiv refl) + γ s = subst (λ f → ((str A) ≡ s) ≃ ι A ((typ A) , s) f) (sym pathToEquivRefl) (θ (str A) s) + + +-- Define the inverse of Id→homEq directly. +ua-lemma : (A B : Type ℓ) (e : A ≃ B) → (pathToEquiv (ua e)) ≡ e +ua-lemma A B e = EquivJ (λ A f → (pathToEquiv (ua f)) ≡ f) + (subst (λ r → pathToEquiv r ≡ idEquiv B) (sym uaIdEquiv) pathToEquivRefl) + e + + +homEq→Id : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') (θ : SNS S ι) + → (A B : TypeWithStr ℓ S) → (A ≃[ ι ] B) → A ≡ B +homEq→Id S ι θ A B (f , φ) = ΣPathP (p , q) + where + p = ua f + + ψ : ι A B (pathToEquiv p) + ψ = subst (λ g → ι A B g) (sym (ua-lemma (typ A) (typ B) f)) φ + + q : PathP (λ i → S (p i)) (str A) (str B) + q = equivFun (invEquiv (hom-lemma-dep S ι θ A B p)) ψ + + +-- Proof of the SIP: +SIP : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') (θ : SNS S ι) + → (A B : TypeWithStr ℓ S) → ((A ≡ B) ≃ (A ≃[ ι ] B)) +SIP S ι θ A B = + (A ≡ B) ≃⟨ i ⟩ + (Σ[ p ∈ (typ A) ≡ (typ B) ] PathP (λ i → S (p i)) (str A) (str B)) ≃⟨ ii ⟩ + (Σ[ p ∈ (typ A) ≡ (typ B) ] (ι A B (pathToEquiv p))) ≃⟨ iii ⟩ + (A ≃[ ι ] B) ■ + where + i = invEquiv ΣPath≃PathΣ + ii = Σ-cong-≃ (hom-lemma-dep S ι θ A B) + iii = Σ-change-of-variable-≃ pathToEquiv (equivIsEquiv univalence) + + + + + + +-- A simple example: pointed types +pointed-structure : Type ℓ → Type ℓ +pointed-structure X = X + +Pointed-Type : Type (ℓ-suc ℓ) +Pointed-Type {ℓ = ℓ} = Σ (Type ℓ) pointed-structure + +pointed-ι : (A B : Pointed-Type) → (A .fst) ≃ (B. fst) → Type ℓ +pointed-ι (X , x) (Y , y) f = (equivFun f) x ≡ y + +pointed-is-sns : SNS {ℓ = ℓ} pointed-structure pointed-ι +pointed-is-sns s t = idEquiv (s ≡ t) + +pointed-type-sip : (X Y : Type ℓ) (x : X) (y : Y) + → (Σ[ f ∈ X ≃ Y ] (f .fst) x ≡ y) ≃ ((X , x) ≡ (Y , y)) +pointed-type-sip X Y x y = invEquiv (SIP pointed-structure pointed-ι pointed-is-sns (X , x) (Y , y)) diff --git a/Cubical/Experiments/Everything.agda b/Cubical/Experiments/Everything.agda new file mode 100644 index 0000000000..2caa258d79 --- /dev/null +++ b/Cubical/Experiments/Everything.agda @@ -0,0 +1,17 @@ +-- Export only the experiments that are expected to compile (without +-- any holes) +module Cubical.Experiments.Everything where + +open import Cubical.Experiments.Brunerie public +open import Cubical.Experiments.EscardoSIP public +open import Cubical.Experiments.Generic public +open import Cubical.Experiments.NatMinusTwo +open import Cubical.Experiments.IsoInt +open import Cubical.Experiments.HAEquivInt +open import Cubical.Experiments.Poset +open import Cubical.Experiments.Problem +open import Cubical.Experiments.FunExtFromUA public +open import Cubical.Experiments.HoTT-UF +open import Cubical.Experiments.ZCohomologyOld.Base +open import Cubical.Experiments.ZCohomologyOld.KcompPrelims +open import Cubical.Experiments.ZCohomologyOld.Properties diff --git a/Cubical/Experiments/FunExtFromUA.agda b/Cubical/Experiments/FunExtFromUA.agda new file mode 100644 index 0000000000..2cf2ae5cf8 --- /dev/null +++ b/Cubical/Experiments/FunExtFromUA.agda @@ -0,0 +1,116 @@ +{- Voevodsky's proof that univalence implies funext -} + +{-# OPTIONS --safe #-} + +module Cubical.Experiments.FunExtFromUA where + +open import Cubical.Foundations.Everything + +variable + ℓ ℓ' : Level +_∼_ : {X : Type ℓ} {A : X → Type ℓ'} → (f g : (x : X) → A x) → Type (ℓ-max ℓ ℓ') +f ∼ g = ∀ x → f x ≡ g x + +funext : ∀ ℓ ℓ' → Type (ℓ-suc(ℓ-max ℓ ℓ')) +funext ℓ ℓ' = {X : Type ℓ} {Y : Type ℓ'} {f g : X → Y} → f ∼ g → f ≡ g + + +pre-comp-is-equiv : (X Y : Type ℓ) (f : X → Y) (Z : Type ℓ) + → isEquiv f + → isEquiv (λ (g : Y → Z) → g ∘ f) +pre-comp-is-equiv {ℓ} X Y f Z e = elimEquivFun P r (f , e) + where + P : (X : Type ℓ) → (X → Y) → Type ℓ + P X f = isEquiv (λ (g : Y → Z) → g ∘ f) + r : P Y (λ x → x) + r = idIsEquiv (Y → Z) + +leftCancellable : {X : Type ℓ} {Y : Type ℓ'} → (X → Y) → Type (ℓ-max ℓ ℓ') +leftCancellable f = ∀ {x x'} → f x ≡ f x' → x ≡ x' + +equivLC : {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) → isEquiv f → leftCancellable f +equivLC f e {x} {x'} p i = hcomp (λ j → \ {(i = i0) → retEq (f , e) x j ; + (i = i1) → retEq (f , e) x' j}) + (invEq (f , e) (p i)) + +univalence-gives-funext : funext ℓ' ℓ +univalence-gives-funext {ℓ'} {ℓ} {X} {Y} {f₀} {f₁} = γ + where + Δ = Σ[ y₀ ∈ Y ] Σ[ y₁ ∈ Y ] y₀ ≡ y₁ + + δ : Y → Δ + δ y = (y , y , refl) + + π₀ π₁ : Δ → Y + π₀ (y₀ , y₁ , p) = y₀ + π₁ (y₀ , y₁ , p) = y₁ + + δ-is-equiv : isEquiv δ + δ-is-equiv = isoToIsEquiv (iso δ π₀ ε η) + where + η : (y : Y) → π₀ (δ y) ≡ y + η y = refl + ε : (d : Δ) → δ (π₀ d) ≡ d + ε (y₀ , y₁ , p) i = y₀ , p i , λ j → p (i ∧ j) + + φ : (Δ → Y) → (Y → Y) + φ π = π ∘ δ + + e : isEquiv φ + e = pre-comp-is-equiv Y Δ δ Y δ-is-equiv + + p : φ π₀ ≡ φ π₁ + p = refl + + q : π₀ ≡ π₁ + q = equivLC φ e p + + g : (h : f₀ ∼ f₁) (π : Δ → Y) (x : X) → Y + g = λ h π x → π (f₀ x , f₁ x , h x) + + γ : f₀ ∼ f₁ → f₀ ≡ f₁ + γ h = cong (g h) q + + γ' : f₀ ∼ f₁ → f₀ ≡ f₁ + γ' h = f₀ ≡⟨ refl ⟩ + (λ x → f₀ x) ≡⟨ refl ⟩ + (λ x → π₀ (f₀ x , f₁ x , h x)) ≡⟨ cong (g h) q ⟩ + (λ x → π₁ (f₀ x , f₁ x , h x)) ≡⟨ refl ⟩ + (λ x → f₁ x) ≡⟨ refl ⟩ + f₁ ∎ + +{- Experiment testing univalence via funext -} + +private + + data ℕ : Type₀ where + zero : ℕ + succ : ℕ → ℕ + + f g : ℕ → ℕ + + f n = n + + g zero = zero + g (succ n) = succ (g n) + + h : (n : ℕ) → f n ≡ g n + h zero = refl + h (succ n) = cong succ (h n) + + p : f ≡ g + p = univalence-gives-funext h + + five : ℕ + five = succ (succ (succ (succ (succ zero)))) + + a : Σ ℕ (λ n → f n ≡ five) + a = five , refl + + b : Σ ℕ (λ n → g n ≡ five) + b = subst (λ - → Σ ℕ (λ n → - n ≡ five)) p a + + c : fst b ≡ five + c = refl + +{- It works, fast. -} diff --git a/Cubical/Experiments/Generic.agda b/Cubical/Experiments/Generic.agda new file mode 100644 index 0000000000..97a247578c --- /dev/null +++ b/Cubical/Experiments/Generic.agda @@ -0,0 +1,120 @@ +-- Two fun examples of generic programming using univalence + +{-# OPTIONS --safe #-} +module Cubical.Experiments.Generic where + +open import Agda.Builtin.String +open import Agda.Builtin.List + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Nat +open import Cubical.Data.Int +open import Cubical.Data.Prod + +---------------------------------------------------------------------------- +-- +-- Dan Licata's example from slide 47 of: +-- http://dlicata.web.wesleyan.edu/pubs/l13jobtalk/l13jobtalk.pdf + +There : Type₀ → Type₀ +There X = List (ℕ × String × (X × ℕ)) + +Database : Type₀ +Database = There (ℕ × ℕ) + +us : Database +us = (4 , "John" , (30 , 5) , 1956) + ∷ (8 , "Hugo" , (29 , 12) , 1978) + ∷ (15 , "James" , (1 , 7) , 1968) + ∷ (16 , "Sayid" , (2 , 10) , 1967) + ∷ (23 , "Jack" , (3 , 12) , 1969) + ∷ (42 , "Sun" , (20 , 3) , 1980) + ∷ [] + +convert : Database → Database +convert d = transport (λ i → There (swapEq ℕ ℕ i)) d + +-- Swap the dates of the American database to get the European format +eu : Database +eu = convert us + +-- A sanity check +_ : us ≡ convert eu +_ = refl + + +---------------------------------------------------------------------------- +-- +-- Example inspired by: +-- +-- Scrap Your Boilerplate: A Practical Design Pattern for Generic Programming +-- Ralf Lämmel & Simon Peyton Jones, TLDI'03 + +Address : Type₀ +Address = String + +Name : Type₀ +Name = String + +data Person : Type₀ where + P : Name → Address → Person + +data Salary (A : Type₀) : Type₀ where + S : A → Salary A + +data Employee (A : Type₀) : Type₀ where + E : Person → Salary A → Employee A + +Manager : Type₀ → Type₀ +Manager A = Employee A + +-- First test of "mutual" +mutual + data Dept (A : Type₀) : Type₀ where + D : Name → Manager A → List (SubUnit A) → Dept A + + data SubUnit (A : Type₀) : Type₀ where + PU : Employee A → SubUnit A + DU : Dept A → SubUnit A + +data Company (A : Type₀) : Type₀ where + C : List (Dept A) → Company A + + +-- A small example + +anders : Employee ℤ +anders = E (P "Anders" "Pittsburgh") (S (pos 2500)) + +andrea : Employee ℤ +andrea = E (P "Andrea" "Copenhagen") (S (pos 2000)) + +andreas : Employee ℤ +andreas = E (P "Andreas" "Gothenburg") (S (pos 3000)) + +-- For now we have a small company +genCom : Company ℤ +genCom = + C ( D "Research" andreas (PU anders ∷ PU andrea ∷ []) + ∷ []) + +-- Increase the salary for everyone by 1 +incSalary : Company ℤ → Company ℤ +incSalary c = transport (λ i → Company (sucPathℤ i)) c + +genCom1 : Company ℤ +genCom1 = incSalary genCom + + +-- Increase the salary more +incSalaryℕ : ℕ → Company ℤ → Company ℤ +incSalaryℕ n c = transport (λ i → Company (addEq n i)) c + +genCom2 : Company ℤ +genCom2 = incSalaryℕ 2 genCom + +genCom10 : Company ℤ +genCom10 = incSalaryℕ 10 genCom diff --git a/Cubical/Experiments/HAEquivInt.agda b/Cubical/Experiments/HAEquivInt.agda new file mode 100644 index 0000000000..819a03aac3 --- /dev/null +++ b/Cubical/Experiments/HAEquivInt.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.HAEquivInt where + +open import Cubical.Experiments.HAEquivInt.Base public + diff --git a/Cubical/Experiments/HAEquivInt/Base.agda b/Cubical/Experiments/HAEquivInt/Base.agda new file mode 100644 index 0000000000..7c7c96db37 --- /dev/null +++ b/Cubical/Experiments/HAEquivInt/Base.agda @@ -0,0 +1,23 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.HAEquivInt.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv.HalfAdjoint + + +data HAEquivInt : Type₀ where + zero : HAEquivInt + suc : HAEquivInt -> HAEquivInt + + -- suc is a HAEquiv: + pred : HAEquivInt -> HAEquivInt + suc-pred : ∀ z -> suc (pred z) ≡ z + pred-suc : ∀ z -> pred (suc z) ≡ z + coh : ∀ z → (λ i → suc (pred-suc z i)) ≡ suc-pred (suc z) + + +suc-haequiv : HAEquiv HAEquivInt HAEquivInt +suc-haequiv = suc , record { g = pred ; linv = pred-suc ; rinv = suc-pred ; com = coh } + + +-- OPEN: prove HAEquivInt ≃ Int! See Experiments/HInt.agda diff --git a/Cubical/Experiments/HInt.agda b/Cubical/Experiments/HInt.agda new file mode 100644 index 0000000000..7532ab79f8 --- /dev/null +++ b/Cubical/Experiments/HInt.agda @@ -0,0 +1,98 @@ +{- + +Definition of the integers as a HIT inspired by slide 10 of (original +idea due to Paolo Capriotti): + +http://www.cs.nott.ac.uk/~psztxa/talks/bonn18.pdf + +Disclaimer: this definition is very hard to work with and I have not +been able to prove that it is equivalent to nat + nat or that it is a +set. + +For a variation that relies on a different notion of equivalence +(without any 2-cell) and which seems easier to work with see: + +https://github.com/RedPRL/redtt/blob/master/library/cool/biinv-int.red + +It might be interesting to port that example one day. + +-} +module Cubical.Experiments.HInt where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.Int +open import Cubical.Data.Nat + +data ℤ : Type₀ where + zero : ℤ + suc : ℤ → ℤ + pred : ℤ → ℤ + sucpred : (n : ℤ) → suc (pred n) ≡ n + predsuc : (n : ℤ) → pred (suc n) ≡ n + -- Coherence: could also be written "sucpred (suc n) ≡ cong (suc (predsuc n))" + coh : (n : ℤ) → Path (suc (pred (suc n)) ≡ suc n) + (sucpred (suc n)) + (λ i → suc (predsuc n i)) + +-- This is equivalent to Int: +coherence : (n : Int) → Path (Path Int (sucInt (predInt (sucInt n))) (sucInt n)) + (sucPred (sucInt n)) + (cong sucInt (predSuc n)) +coherence (pos zero) = refl +coherence (pos (suc n)) = refl +coherence (negsuc zero) = refl +coherence (negsuc (suc zero)) = refl +coherence (negsuc (suc (suc n))) = refl + +ℤ→Int : ℤ → Int +ℤ→Int zero = pos 0 +ℤ→Int (suc n) = sucInt (ℤ→Int n) +ℤ→Int (pred n) = predInt (ℤ→Int n) +ℤ→Int (sucpred n i) = sucPred (ℤ→Int n) i +ℤ→Int (predsuc n i) = predSuc (ℤ→Int n) i +ℤ→Int (coh n i j) = coherence (ℤ→Int n) i j + +ℕ→ℤ : ℕ → ℤ +ℕ→ℤ zero = zero +ℕ→ℤ (suc n) = suc (ℕ→ℤ n) + +negsucℕ→ℤ : ℕ → ℤ +negsucℕ→ℤ zero = pred zero +negsucℕ→ℤ (suc n) = pred (negsucℕ→ℤ n) + +Int→ℤ : Int → ℤ +Int→ℤ (pos n) = ℕ→ℤ n +Int→ℤ (negsuc n) = negsucℕ→ℤ n + +lem1 : ∀ n → Int→ℤ (sucInt n) ≡ suc (Int→ℤ n) +lem1 (pos n) = refl +lem1 (negsuc zero) = sym (sucpred zero) +lem1 (negsuc (suc n)) = sym (sucpred (negsucℕ→ℤ n)) + +lem2 : ∀ n → Int→ℤ (predInt n) ≡ pred (Int→ℤ n) +lem2 (pos zero) = refl +lem2 (pos (suc n)) = sym (predsuc (ℕ→ℤ n)) +lem2 (negsuc n) = refl + +-- I don't see how to fill these holes, especially the last one +ℤ→Int→ℤ : ∀ (n : ℤ) → Int→ℤ (ℤ→Int n) ≡ n +ℤ→Int→ℤ zero = refl +ℤ→Int→ℤ (suc n) = (lem1 (ℤ→Int n)) ∙ (cong suc (ℤ→Int→ℤ n)) +ℤ→Int→ℤ (pred n) = (lem2 (ℤ→Int n)) ∙ (cong pred (ℤ→Int→ℤ n)) +ℤ→Int→ℤ (sucpred n i) = {!!} +ℤ→Int→ℤ (predsuc n i) = {!!} +ℤ→Int→ℤ (coh n i j) = {!!} + +Int→ℤ→Int : ∀ n → ℤ→Int (Int→ℤ n) ≡ n +Int→ℤ→Int (pos zero) = refl +Int→ℤ→Int (pos (suc n)) = cong sucInt (Int→ℤ→Int (pos n)) +Int→ℤ→Int (negsuc zero) = refl +Int→ℤ→Int (negsuc (suc n)) = cong predInt (Int→ℤ→Int (negsuc n)) + +Int≡ℤ : Int ≡ ℤ +Int≡ℤ = isoToPath (iso Int→ℤ ℤ→Int ℤ→Int→ℤ Int→ℤ→Int) + +isSetℤ : isSet ℤ +isSetℤ = subst isSet Int≡ℤ isSetInt diff --git a/Cubical/Experiments/HoTT-UF.agda b/Cubical/Experiments/HoTT-UF.agda new file mode 100644 index 0000000000..d5c7d169a8 --- /dev/null +++ b/Cubical/Experiments/HoTT-UF.agda @@ -0,0 +1,104 @@ +{- + +This is a HoTT-UF core library based on cubical type theory, where the +cubical machinery is hidden, using the HoTT Book terminology and +notation. + +The point is that function extensionality, propositional truncation +and univalence compute (an example is given below). + +For the moment, this requires the development version of Agda. + +-} + +{-# OPTIONS --exact-split --safe #-} + +module Cubical.Experiments.HoTT-UF where + +open import Cubical.Core.Primitives hiding ( _≡_ ) +open import Cubical.Core.Id public + +open import Cubical.Foundations.Id public + using ( _≡_ -- The identity type. + ; refl -- Unfortunately, pattern matching on refl is not available. + ; J -- Until it is, you have to use the induction principle J. + + ; transport -- As in the HoTT Book. + ; ap + ; _∙_ + ; _⁻¹ + + ; _≡⟨_⟩_ -- Standard equational reasoning. + ; _∎ + + ; funExt -- Function extensionality + -- (can also be derived from univalence). + + ; Σ -- Sum type. Needed to define contractible types, equivalences + ; _,_ -- and univalence. + ; pr₁ -- The eta rule is available. + ; pr₂ + + ; isProp -- The usual notions of proposition, contractible type, set. + ; isContr + ; isSet + + ; isEquiv -- A map with contractible fibers + -- (Voevodsky's version of the notion). + ; _≃_ -- The type of equivalences between two given types. + ; EquivContr -- A formulation of univalence. + + ; ∥_∥ -- Propositional truncation. + ; ∣_∣ -- Map into the propositional truncation. + ; ∥∥-isProp -- A truncated type is a proposition. + ; ∥∥-recursion -- Non-dependent elimination. + ; ∥∥-induction -- Dependent elimination. + ) + +{- + +Here is an illustration of how function extensionality computes. + +-} + +private + + data ℕ : Type₀ where + zero : ℕ + succ : ℕ → ℕ + + f g : ℕ → ℕ + + f n = n + + g zero = zero + g (succ n) = succ (g n) + + h : (n : ℕ) → f n ≡ g n + h zero = refl + h (succ n) = ap succ (h n) + + p : f ≡ g + p = funExt h + + five : ℕ + five = succ (succ (succ (succ (succ zero)))) + + a : Σ ℕ (λ n → f n ≡ five) + a = five , refl + + b : Σ ℕ (λ n → g n ≡ five) + b = transport (λ - → Σ ℕ (λ n → - n ≡ five)) p a + + c : pr₁ b ≡ five + c = refl + +{- + +If we had funExt as a postulate, then the definition of c would not +type check. Moreover, the term pr₁ b would not evaluate to five, as it +does with the cubical type theory implementation of funext. + +TODO. A similar computational example with univalence. + +-} diff --git a/Cubical/Experiments/IsoInt.agda b/Cubical/Experiments/IsoInt.agda new file mode 100644 index 0000000000..b3dbb78724 --- /dev/null +++ b/Cubical/Experiments/IsoInt.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.IsoInt where + +open import Cubical.Experiments.IsoInt.Base public + diff --git a/Cubical/Experiments/IsoInt/Base.agda b/Cubical/Experiments/IsoInt/Base.agda new file mode 100644 index 0000000000..487f3d61f4 --- /dev/null +++ b/Cubical/Experiments/IsoInt/Base.agda @@ -0,0 +1,112 @@ +{- + +The naive, but incorrect, way to define the integers as a HIT. + +This file mainly contains a proof that IsoInt ≢ Int, and ends with a + demonstration of how the same proof strategy fails for BiInvℤ. + +-} +{-# OPTIONS --safe #-} +module Cubical.Experiments.IsoInt.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Int +open import Cubical.Data.Nat +open import Cubical.Data.Empty + +open import Cubical.Relation.Nullary + + +data IsoInt : Type₀ where + zero : IsoInt + suc : IsoInt -> IsoInt + + -- suc is an isomorphism: + pred : IsoInt -> IsoInt + suc-pred : ∀ z -> suc (pred z) ≡ z + pred-suc : ∀ z -> pred (suc z) ≡ z + + +suc-iso : Iso IsoInt IsoInt +suc-iso = record { fun = suc ; inv = pred ; rightInv = suc-pred ; leftInv = pred-suc } + + +-- this submodule is adapted from Section 5 of +-- http://www.cs.ru.nl/~herman/PUBS/HIT-programming.pdf + +module NonTrivial where + + -- these two paths are distinct! + p₁ p₂ : Path IsoInt (suc (pred (suc zero))) (suc zero) + p₁ i = suc-pred (suc zero) i + p₂ i = suc (pred-suc zero i) + + -- to prove this we map into S¹, sending p₁ to refl and p₂ to loop + + open import Cubical.HITs.S1 + + toS¹ : IsoInt → S¹ + toS¹ zero = base + toS¹ (suc x) = toS¹ x + toS¹ (pred x) = toS¹ x + toS¹ (suc-pred x i) = refl {x = toS¹ x} i + toS¹ (pred-suc x i) = rotLoop (toS¹ x) i + + p₁≡refl : cong toS¹ p₁ ≡ refl + p₁≡refl = refl + + p₂≡loop : cong toS¹ p₂ ≡ loop + p₂≡loop = refl + + -- this is enough to show that p₁ and p₂ cannot be equal + p₁≢p₂ : ¬ (p₁ ≡ p₂) + p₁≢p₂ eq = znots 0≡1 + where -- using winding numbers, p₁ ≡ p₂ implies 0 ≡ 1 + 0≡1 : 0 ≡ 1 + 0≡1 = injPos (cong (winding ∘ cong toS¹) eq) + + +¬isSet-IsoInt : ¬ (isSet IsoInt) +¬isSet-IsoInt pf = NonTrivial.p₁≢p₂ (pf _ _ NonTrivial.p₁ NonTrivial.p₂) + +¬Int≡IsoInt : ¬ (ℤ ≡ IsoInt) +¬Int≡IsoInt p = ¬isSet-IsoInt (subst isSet p isSetℤ) + + + +private + + -- Note: this same proof strategy fails for BiInvℤ! + + open import Cubical.Data.Int.MoreInts.BiInvInt hiding (zero; suc; pred; suc-pred; pred-suc) + import Cubical.Data.Int.MoreInts.BiInvInt as BiI + + p₁ p₂ : Path BiInvℤ (BiI.suc (BiI.pred (BiI.suc BiI.zero))) (BiI.suc BiI.zero) + p₁ i = BiI.suc-pred (BiI.suc BiI.zero) i + p₂ i = BiI.suc (BiI.pred-suc BiI.zero i) + + open import Cubical.HITs.S1 + + toS¹ : BiInvℤ → S¹ + toS¹ BiI.zero = base + toS¹ (BiI.suc x) = toS¹ x + toS¹ (BiI.predr x) = toS¹ x + toS¹ (BiI.predl x) = toS¹ x + toS¹ (BiI.suc-predr x i) = refl {x = toS¹ x} i + toS¹ (BiI.predl-suc x i) = rotLoop (toS¹ x) i + + -- still p₂ maps to loop... + p₂≡loop : cong toS¹ p₂ ≡ loop + p₂≡loop = refl + + open import Cubical.Foundations.GroupoidLaws + + -- ...but now so does p₁! + p₁≡loop : cong toS¹ p₁ ≡ loop + p₁≡loop = sym (decodeEncode base (cong toS¹ p₁)) ∙ sym (lUnit loop) + + -- if we use BiI.predr instead of BiI.pred (≡ BiI.predl) in p₁ and p₂, + -- both paths in S¹ are refl diff --git a/Cubical/Experiments/List.agda b/Cubical/Experiments/List.agda new file mode 100644 index 0000000000..9421961e2b --- /dev/null +++ b/Cubical/Experiments/List.agda @@ -0,0 +1,285 @@ +{- + +An experiment of transporting rev-++-distr from lists to lists where +the arguments to cons have been flipped inspired by section 2 of +https://arxiv.org/abs/2010.00774 + +Note that Agda doesn't care about the order of constructors so we +can't do exactly the same example. + +-} +{-# OPTIONS --safe #-} +module Cubical.Experiments.List where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma + +infixr 5 _∷_ +infixl 5 _∷'_ +infixr 5 _++_ + +-- Normal lists +data List (A : Type) : Type where + [] : List A + _∷_ : (x : A) (xs : List A) → List A + +-- Lists where the arguments to cons have been flipped +data List' (A : Type) : Type where + [] : List' A + _∷'_ : (xs : List' A) (x : A) → List' A + +variable + A : Type + + +-- Some operations and properties for List + +_++_ : List A → List A → List A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ xs ++ ys + +rev : List A → List A +rev [] = [] +rev (x ∷ xs) = rev xs ++ (x ∷ []) + +++-unit-r : (xs : List A) → xs ++ [] ≡ xs +++-unit-r [] = refl +++-unit-r (x ∷ xs) = cong (_∷_ x) (++-unit-r xs) + +++-assoc : (xs ys zs : List A) → (xs ++ ys) ++ zs ≡ xs ++ ys ++ zs +++-assoc [] ys zs = refl +++-assoc (x ∷ xs) ys zs = cong (_∷_ x) (++-assoc xs ys zs) + +rev-++-distr : (xs ys : List A) → rev (xs ++ ys) ≡ rev ys ++ rev xs +rev-++-distr [] ys = sym (++-unit-r (rev ys)) +rev-++-distr (x ∷ xs) ys = cong (_++ _) (rev-++-distr xs ys) ∙ ++-assoc (rev ys) (rev xs) (x ∷ []) + + + +-- We now want to transport this to List'. For this we first establish +-- an isomorphism of the types. + +toList' : List A → List' A +toList' [] = [] +toList' (x ∷ xs) = toList' xs ∷' x + +fromList' : List' A → List A +fromList' [] = [] +fromList' (xs ∷' x) = x ∷ fromList' xs + +toFrom : (xs : List' A) → toList' (fromList' xs) ≡ xs +toFrom [] = refl +toFrom (xs ∷' x) i = toFrom xs i ∷' x + +fromTo : (xs : List A) → fromList' (toList' xs) ≡ xs +fromTo [] = refl +fromTo (x ∷ xs) i = x ∷ fromTo xs i + +ListIso : Iso (List A) (List' A) +ListIso = iso toList' fromList' toFrom fromTo + +ListEquiv : List A ≃ List' A +ListEquiv = isoToEquiv ListIso + +-- We then use univalence to turn this into a path +ListPath : (A : Type) → List A ≡ List' A +ListPath A = isoToPath (ListIso {A = A}) + + +-- We can now use this path to transport the operations and properties +-- from List to List' +module transport where + + -- First make a suitable Σ-type packaging what we need for the + -- transport (note that _++_ and rev here are part of the Σ-type). + -- It should be possible to automatically generate this given a module/file. + T : Type → Type + T X = Σ[ _++_ ∈ (X → X → X) ] Σ[ rev ∈ (X → X) ] ((xs ys : X) → rev (xs ++ ys) ≡ rev ys ++ rev xs) + + -- We can now transport the instance of T for List to List' + T-List' : T (List' A) + T-List' {A = A} = transport (λ i → T (ListPath A i)) (_++_ , rev , rev-++-distr) + + -- Getting the operations and property for List' is then just a matter of projecting them out + _++'_ : List' A → List' A → List' A + _++'_ = T-List' .fst + + rev' : List' A → List' A + rev' = T-List' .snd .fst + + rev-++-distr' : (xs ys : List' A) → rev' (xs ++' ys) ≡ rev' ys ++' rev' xs + rev-++-distr' = T-List' .snd .snd + + +-- To connect this with the Cubical Agda paper consider the following +-- (painfully) manual transport. This is really what the above code +-- unfolds to. +module manualtransport where + + _++'_ : List' A → List' A → List' A + _++'_ {A = A} = transport (λ i → ListPath A i → ListPath A i → ListPath A i) _++_ + + rev' : List' A → List' A + rev' {A = A} = transport (λ i → ListPath A i → ListPath A i) rev + + rev-++-distr' : (xs ys : List' A) → rev' (xs ++' ys) ≡ rev' ys ++' rev' xs + rev-++-distr' {A = A} = transport (λ i → (xs ys : ListPath A i) + → revP i (appP i xs ys) ≡ appP i (revP i ys) (revP i xs)) + rev-++-distr + where + appP : PathP (λ i → ListPath A i → ListPath A i → ListPath A i) _++_ _++'_ + appP i = transp (λ j → ListPath A (i ∧ j) → ListPath A (i ∧ j) → ListPath A (i ∧ j)) (~ i) _++_ + + revP : PathP (λ i → ListPath A i → ListPath A i) rev rev' + revP i = transp (λ j → ListPath A (i ∧ j) → ListPath A (i ∧ j)) (~ i) rev + + + +-- The above operations for List' are derived by going back and +-- forth. With the SIP we can do better and transport properties for +-- user defined operations (assuming that the operations are +-- well-defined wrt to the forward direction of the equivalence). +open import Cubical.Foundations.SIP +open import Cubical.Structures.Axioms +open import Cubical.Structures.Product +open import Cubical.Structures.Pointed +open import Cubical.Structures.Function + + +-- For illustrative purposes we first apply the SIP manually. This +-- requires quite a bit of boilerplate code which is automated in the +-- next module. +module manualSIP (A : Type) where + + -- First define the raw structure without axioms. This is just the + -- signature of _++_ and rev. + RawStruct : Type → Type + RawStruct X = (X → X → X) × (X → X) + + -- Some boilerplate code which can be automated + e1 : StrEquiv (λ x → x → x → x) ℓ-zero + e1 = FunctionEquivStr+ pointedEquivAction + (FunctionEquivStr+ pointedEquivAction PointedEquivStr) + + e2 : StrEquiv (λ x → x → x) ℓ-zero + e2 = FunctionEquivStr+ pointedEquivAction PointedEquivStr + + RawEquivStr : StrEquiv RawStruct _ + RawEquivStr = ProductEquivStr e1 e2 + + rawUnivalentStr : UnivalentStr _ RawEquivStr + rawUnivalentStr = productUnivalentStr e1 he1 e2 he2 + where + he2 : UnivalentStr (λ z → z → z) e2 + he2 = functionUnivalentStr+ pointedEquivAction pointedTransportStr + PointedEquivStr pointedUnivalentStr + + he1 : UnivalentStr (λ z → z → z → z) e1 + he1 = functionUnivalentStr+ pointedEquivAction pointedTransportStr e2 he2 + + -- Now the property that we want to transport + P : (X : Type) → RawStruct X → Type + P X (_++_ , rev) = ((xs ys : X) → rev (xs ++ ys) ≡ rev ys ++ rev xs) + + -- Package things up for List + List-Struct : Σ[ X ∈ Type ] (Σ[ s ∈ RawStruct X ] (P X s)) + List-Struct = List A , (_++_ , rev) , rev-++-distr + + + -- We now give direct definitions of ++' and rev' for List' + _++'_ : List' A → List' A → List' A + [] ++' ys = ys + (xs ∷' x) ++' ys = (xs ++' ys) ∷' x + + rev' : List' A → List' A + rev' [] = [] + rev' (xs ∷' x) = rev' xs ++' ([] ∷' x) + + -- We then package this up into a raw structure on List' + List'-RawStruct : Σ[ X ∈ Type ] (RawStruct X) + List'-RawStruct = List' A , (_++'_ , rev') + + -- Finally we prove that toList' commutes with _++_ and rev. Note + -- that this could be a lot more complex, see for example the Matrix + -- example (Cubical.Algebra.Matrix). + toList'-++ : (xs ys : List A) → toList' (xs ++ ys) ≡ toList' xs ++' toList' ys + toList'-++ [] ys = refl + toList'-++ (x ∷ xs) ys i = toList'-++ xs ys i ∷' x + + toList'-rev : (xs : List A) → toList' (rev xs) ≡ rev' (toList' xs) + toList'-rev [] = refl + toList'-rev (x ∷ xs) = toList'-++ (rev xs) (x ∷ []) ∙ cong (_++' ([] ∷' x)) (toList'-rev xs) + + -- We can now get the property for ++' and rev' via the SIP + rev-++-distr' : (xs ys : List' A) → rev' (xs ++' ys) ≡ rev' ys ++' rev' xs + rev-++-distr' = transferAxioms rawUnivalentStr List-Struct List'-RawStruct + (ListEquiv , toList'-++ , toList'-rev) + + -- Note that rev-++-distr' is really talking about the direct + -- definitions of ++' and rev', not the transported operations as in + -- the previous attempt. + + + +-- We now automate parts of the above construction +open import Cubical.Structures.Auto + +module SIP-auto (A : Type) where + + -- First define the raw structure without axioms. This is just the + -- signature of _++_ and rev. + RawStruct : Type → Type + RawStruct X = (X → X → X) × (X → X) + + -- Some automated SIP magic + RawEquivStr : _ + RawEquivStr = AutoEquivStr RawStruct + + rawUnivalentStr : UnivalentStr _ RawEquivStr + rawUnivalentStr = autoUnivalentStr RawStruct + + -- Now the property that we want to transport + P : (X : Type) → RawStruct X → Type + P X (_++_ , rev) = ((xs ys : X) → rev (xs ++ ys) ≡ rev ys ++ rev xs) + + -- Package things up for List + List-Struct : Σ[ X ∈ Type ] (Σ[ s ∈ RawStruct X ] (P X s)) + List-Struct = List A , (_++_ , rev) , rev-++-distr + + + -- We now give direct definitions of ++' and rev' for List' + _++'_ : List' A → List' A → List' A + [] ++' ys = ys + (xs ∷' x) ++' ys = (xs ++' ys) ∷' x + + rev' : List' A → List' A + rev' [] = [] + rev' (xs ∷' x) = rev' xs ++' ([] ∷' x) + + -- We then package this up into a raw structure on List' + List'-RawStruct : Σ[ X ∈ Type ] (RawStruct X) + List'-RawStruct = List' A , (_++'_ , rev') + + -- Finally we prove that toList' commutes with _++_ and rev. Note + -- that this could be a lot more complex, see for example the Matrix + -- example (Cubical.Algebra.Matrix). + toList'-++ : (xs ys : List A) → toList' (xs ++ ys) ≡ toList' xs ++' toList' ys + toList'-++ [] ys = refl + toList'-++ (x ∷ xs) ys i = toList'-++ xs ys i ∷' x + + toList'-rev : (xs : List A) → toList' (rev xs) ≡ rev' (toList' xs) + toList'-rev [] = refl + toList'-rev (x ∷ xs) = toList'-++ (rev xs) (x ∷ []) ∙ cong (_++' ([] ∷' x)) (toList'-rev xs) + + -- We can now get the property for ++' and rev' via the SIP + rev-++-distr' : (xs ys : List' A) → rev' (xs ++' ys) ≡ rev' ys ++' rev' xs + rev-++-distr' = transferAxioms rawUnivalentStr List-Struct List'-RawStruct + (ListEquiv , toList'-++ , toList'-rev) + + -- Note that rev-++-distr' is really talking about the direct + -- definitions of ++' and rev', not the transported operations as in + -- the previous attempt. diff --git a/Cubical/Experiments/LocalisationDefs.agda b/Cubical/Experiments/LocalisationDefs.agda new file mode 100644 index 0000000000..f0e2da1c09 --- /dev/null +++ b/Cubical/Experiments/LocalisationDefs.agda @@ -0,0 +1,146 @@ +-- This file contains several ways to define localisation +-- and proves them all to be equivalent + +{-# OPTIONS --safe #-} +module Cubical.Experiments.LocalisationDefs where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Transport +open import Cubical.Functions.FunExtEquiv + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Vec +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Algebra.Group +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open Iso + +private + variable + ℓ ℓ' : Level + A : Type ℓ + +record isMultClosedSubset (R' : CommRing {ℓ}) (S' : ℙ (R' .fst)) : Type ℓ where + constructor + multclosedsubset + field + containsOne : (R' .snd .CommRingStr.1r) ∈ S' + multClosed : ∀ {s t} → s ∈ S' → t ∈ S' → (R' .snd .CommRingStr._·_ s t) ∈ S' + +module _ (R' : CommRing {ℓ}) (S' : ℙ (R' .fst)) (SMultClosedSubset : isMultClosedSubset R' S') where + open isMultClosedSubset + private R = R' .fst + open CommRingStr (R' .snd) + open Theory (CommRing→Ring R') + + S = Σ[ s ∈ R ] (s ∈ S') + + -- HIT definition + data S⁻¹R : Type ℓ where + _/ₗ_ : R → S → S⁻¹R + zd : (r₁ r₂ : R) (s s₁ s₂ : S) + → fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁ + → r₁ /ₗ s₁ ≡ r₂ /ₗ s₂ + trunc : isSet S⁻¹R + + infixr 5 _/ₗ_ + + + module Elim {ℓ'} {B : S⁻¹R → Type ℓ'} + (_/*_ : (r : R) (s : S) → B (r /ₗ s)) + (zd* : (r₁ r₂ : R) (s s₁ s₂ : S) + → (p : fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + → PathP (λ i → B (zd r₁ r₂ s s₁ s₂ p i)) (r₁ /* s₁) (r₂ /* s₂)) + (trunc* : (q : S⁻¹R) → isSet (B q)) where + + + f : (q : S⁻¹R) → B q + f (r /ₗ s) = r /* s + f (zd r₁ r₂ s s₁ s₂ p i) = zd* r₁ r₂ s s₁ s₂ p i + f (trunc q₁ q₂ x y i j) = isOfHLevel→isOfHLevelDep 2 trunc* (f q₁) (f q₂) (cong f x) (cong f y) + (trunc q₁ q₂ x y) i j + + + module ElimProp {ℓ'} {B : S⁻¹R → Type ℓ'} (Bprop : {q : S⁻¹R} → isProp (B q)) + (_/*_ : (r : R) → (s : S) → B (r /ₗ s)) where + + + f : (q : S⁻¹R) → B q + f = Elim.f _/*_ (λ r₁ r₂ s s₁ s₂ p + → toPathP (Bprop (transp (λ i → B (zd r₁ r₂ s s₁ s₂ p i)) i0 (r₁ /* s₁)) + (r₂ /* s₂))) + λ q → isProp→isSet Bprop + + + module Rec {ℓ'} {B : Type ℓ'} (BType : isSet B) + (_/*_ : R → S → B) + (zd* : (r₁ r₂ : R) (s s₁ s₂ : S) + → (p : fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + → r₁ /* s₁ ≡ r₂ /* s₂) + where + + f : S⁻¹R → B + f = Elim.f _/*_ zd* (λ _ → BType) + + + -- approach using set quotients + _≈_ : R × S → R × S → Type ℓ + (r₁ , s₁) ≈ (r₂ , s₂) = ∃[ s ∈ S ] (fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + + S⁻¹R/ = (R × S) / _≈_ + + -- proving equivalence of the two types + φ : S⁻¹R/ → S⁻¹R + φ = SQ.rec trunc (λ (r , s) → r /ₗ s) β + where + α : ((r₁ , s₁) (r₂ , s₂) : R × S) → Σ[ s ∈ S ] (fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + → r₁ /ₗ s₁ ≡ r₂ /ₗ s₂ + α _ _ (s , p) = zd _ _ s _ _ p + + β : ((r₁ , s₁) (r₂ , s₂) : R × S) → ∃[ s ∈ S ] (fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + → r₁ /ₗ s₁ ≡ r₂ /ₗ s₂ + β _ _ = PT.rec (trunc _ _) (α _ _) + + ψ : S⁻¹R → S⁻¹R/ + ψ (r /ₗ s) = [ r , s ] + ψ (zd r₁ r₂ s s₁ s₂ p i) = eq/ (r₁ , s₁) (r₂ , s₂) ∣ s , p ∣ i + ψ (trunc x y p q i j) = squash/ (ψ x) (ψ y) (cong ψ p) (cong ψ q) i j + + η : section φ ψ + η = ElimProp.f (trunc _ _) λ _ _ → refl + + ε : retract φ ψ + ε = elimProp (λ _ → squash/ _ _) λ _ → refl + + S⁻¹R/≃S⁻¹R : S⁻¹R/ ≃ S⁻¹R + S⁻¹R/≃S⁻¹R = isoToEquiv (iso φ ψ η ε) + + + -- Set quotients but with Σ, this is the type used in Algebra.Localisation.Base + -- as this is easiest to use + _≈'_ : R × S → R × S → Type ℓ + (r₁ , s₁) ≈' (r₂ , s₂) = Σ[ s ∈ S ] (fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + + S⁻¹R/' = (R × S) / _≈'_ + + S⁻¹R/'≃S⁻¹R/ : S⁻¹R/' ≃ S⁻¹R/ + S⁻¹R/'≃S⁻¹R/ = SQ.truncRelEquiv diff --git a/Cubical/Experiments/NatMinusTwo.agda b/Cubical/Experiments/NatMinusTwo.agda new file mode 100644 index 0000000000..e0a676e45d --- /dev/null +++ b/Cubical/Experiments/NatMinusTwo.agda @@ -0,0 +1,21 @@ +{- + This type ℕ₋₂ was originally used as the index to n-truncation in order to + be consistent with the notation in the HoTT book. However, ℕ was already + being used as an analogous index in Foundations.HLevels, and it became + clear that having two different indexing schemes for truncation levels was + very inconvenient. In the end, having slightly nicer notation was not worth + the hassle of having to use this type everywhere where truncation levels + were needed. So for this library, use the type `HLevel = ℕ` instead. + + See the discussions below for more context: + - https://github.com/agda/cubical/issues/266 + - https://github.com/agda/cubical/pull/238 +-} +{-# OPTIONS --safe #-} +module Cubical.Experiments.NatMinusTwo where + +open import Cubical.Experiments.NatMinusTwo.Base public + +open import Cubical.Experiments.NatMinusTwo.Properties public + +open import Cubical.Experiments.NatMinusTwo.ToNatMinusOne using (1+_; ℕ₋₁→ℕ₋₂; -1+Path) public diff --git a/Cubical/Experiments/NatMinusTwo/Base.agda b/Cubical/Experiments/NatMinusTwo/Base.agda new file mode 100644 index 0000000000..e450c7adbf --- /dev/null +++ b/Cubical/Experiments/NatMinusTwo/Base.agda @@ -0,0 +1,41 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Experiments.NatMinusTwo.Base where + +open import Cubical.Core.Primitives +open import Cubical.Data.Nat +open import Cubical.Data.Empty + +record ℕ₋₂ : Type₀ where + constructor -2+_ + field + n : ℕ + +pattern neg2 = -2+ zero +pattern neg1 = -2+ (suc zero) +pattern ℕ→ℕ₋₂ n = -2+ (suc (suc n)) +pattern -1+_ n = -2+ (suc n) + +2+_ : ℕ₋₂ → ℕ +2+ (-2+ n) = n + +pred₋₂ : ℕ₋₂ → ℕ₋₂ +pred₋₂ neg2 = neg2 +pred₋₂ neg1 = neg2 +pred₋₂ (ℕ→ℕ₋₂ zero) = neg1 +pred₋₂ (ℕ→ℕ₋₂ (suc n)) = (ℕ→ℕ₋₂ n) + +suc₋₂ : ℕ₋₂ → ℕ₋₂ +suc₋₂ (-2+ n) = -2+ (suc n) + +-- Natural number and negative integer literals for ℕ₋₂ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℕ₋₂ : HasFromNat ℕ₋₂ + fromNatℕ₋₂ = record { Constraint = λ _ → Unit ; fromNat = ℕ→ℕ₋₂ } + +instance + fromNegℕ₋₂ : HasFromNeg ℕ₋₂ + fromNegℕ₋₂ = record { Constraint = λ { (suc (suc (suc _))) → ⊥ ; _ → Unit } + ; fromNeg = λ { zero → 0 ; (suc zero) → neg1 ; (suc (suc zero)) → neg2 } } diff --git a/Cubical/Experiments/NatMinusTwo/Properties.agda b/Cubical/Experiments/NatMinusTwo/Properties.agda new file mode 100644 index 0000000000..d722348437 --- /dev/null +++ b/Cubical/Experiments/NatMinusTwo/Properties.agda @@ -0,0 +1,11 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Experiments.NatMinusTwo.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.Nat +open import Cubical.Data.NatMinusOne using (ℕ₋₁) +open import Cubical.Experiments.NatMinusTwo.Base + +-2+Path : ℕ ≡ ℕ₋₂ +-2+Path = isoToPath (iso -2+_ 2+_ (λ _ → refl) (λ _ → refl)) diff --git a/Cubical/Experiments/NatMinusTwo/ToNatMinusOne.agda b/Cubical/Experiments/NatMinusTwo/ToNatMinusOne.agda new file mode 100644 index 0000000000..49392bb13c --- /dev/null +++ b/Cubical/Experiments/NatMinusTwo/ToNatMinusOne.agda @@ -0,0 +1,24 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Experiments.NatMinusTwo.ToNatMinusOne where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.NatMinusOne as ℕ₋₁ using (ℕ₋₁) +open import Cubical.Experiments.NatMinusTwo.Base as ℕ₋₂ using (ℕ₋₂; -2+_) + +-- Conversions to/from ℕ₋₁ + +-1+_ : ℕ₋₁ → ℕ₋₂ +-1+ (ℕ₋₁.-1+ n) = -2+ n + +1+_ : ℕ₋₂ → ℕ₋₁ +1+ (-2+ n) = ℕ₋₁.-1+ n + +ℕ₋₁→ℕ₋₂ : ℕ₋₁ → ℕ₋₂ +ℕ₋₁→ℕ₋₂ (ℕ₋₁.-1+ n) = ℕ₋₂.-1+ n + +-- Properties + +-1+Path : ℕ₋₁ ≡ ℕ₋₂ +-1+Path = isoToPath (iso -1+_ 1+_ (λ _ → refl) (λ _ → refl)) diff --git a/Cubical/Experiments/Poset.agda b/Cubical/Experiments/Poset.agda new file mode 100644 index 0000000000..b6b1f4e7a0 --- /dev/null +++ b/Cubical/Experiments/Poset.agda @@ -0,0 +1,326 @@ +{- + This an earlier formalization of posets using the old SIP instead of DURGs +-} + +{-# OPTIONS --safe #-} +module Cubical.Experiments.Poset where + +open import Cubical.Foundations.Prelude +open import Cubical.Functions.Logic +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv renaming (_■ to _QED) +open import Cubical.Foundations.SIP +open import Cubical.Functions.FunExtEquiv +open import Cubical.Foundations.Function +open import Cubical.Core.Primitives +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.Sigma.Properties + +open import Cubical.Structures.Axioms + +-- We will adopt the convention of denoting the level of the carrier +-- set by ℓ₀ and the level of the relation result by ℓ₁. +private + variable + ℓ ℓ₀ ℓ₁ ℓ₀′ ℓ₁′ ℓ₀′′ ℓ₁′′ : Level + +-- Max: Should just use PropRel from Base +Order : (ℓ₁ : Level) → Type ℓ₀ → Type (ℓ-max ℓ₀ (ℓ-suc ℓ₁)) +Order ℓ₁ A = A → A → hProp ℓ₁ + +isSetOrder : (ℓ₁ : Level) (A : Type ℓ₀) → isSet (Order ℓ₁ A) +isSetOrder ℓ₁ A = isSetΠ2 λ _ _ → isSetHProp + +-- We first start by defining what it means a for a function to be +-- order-preserving. The name "monotonic" is reserved for partial orders. + +isOrderPreserving : (M : TypeWithStr ℓ₀ (Order ℓ₁)) (N : TypeWithStr ℓ₀′ (Order ℓ₁′)) + → (fst M → fst N) → Type _ +isOrderPreserving (A , _⊑₀_) (B , _⊑₁_) f = + (x y : A) → ⟨ x ⊑₀ y ⟩ → ⟨ f x ⊑₁ f y ⟩ + +isPropIsOrderPreserving : (M : TypeWithStr ℓ₀ (Order ℓ₁)) + (N : TypeWithStr ℓ₀′ (Order ℓ₁′)) + → (f : fst M → fst N) + → isProp (isOrderPreserving M N f) +isPropIsOrderPreserving M (_ , _⊑₁_) f = isPropΠ3 λ x y p → snd (f x ⊑₁ f y) + +-- We then define what it means for an equivalence to order-preserving which is +-- nothing but the property that both directions of the equivalence are +-- order-preserving. + +isAnOrderPreservingEqv : (M : TypeWithStr ℓ₀ (Order ℓ₁)) + (N : TypeWithStr ℓ₀′ (Order ℓ₁′)) + → fst M ≃ fst N → Type _ +isAnOrderPreservingEqv M N e@(f , _) = + isOrderPreserving M N f × isOrderPreserving N M g + where + g = equivFun (invEquiv e) + +orderUnivalentStr : SNS {ℓ} (Order ℓ₁) isAnOrderPreservingEqv +orderUnivalentStr {ℓ = ℓ} {ℓ₁ = ℓ₁} {X = X} _⊑₀_ _⊑₁_ = + f , record { equiv-proof = f-equiv } + where + f : isAnOrderPreservingEqv (X , _⊑₀_) (X , _⊑₁_) (idEquiv X) → _⊑₀_ ≡ _⊑₁_ + f e@(φ , ψ) = funExt₂ λ x y → ⇔toPath (φ x y) (ψ x y) + + g : _⊑₀_ ≡ _⊑₁_ → isAnOrderPreservingEqv (X , _⊑₀_) (X , _⊑₁_) (idEquiv X) + g p = + subst + (λ _⊑_ → isAnOrderPreservingEqv (X , _⊑₀_) (X , _⊑_) (idEquiv X)) + p + ((λ _ _ x⊑₁y → x⊑₁y) , λ _ _ x⊑₁y → x⊑₁y) + + ret-f-g : retract f g + ret-f-g (φ , ψ) = + isPropΣ + (isPropIsOrderPreserving (X , _⊑₀_) (X , _⊑₁_) (idfun X)) + (λ _ → isPropIsOrderPreserving (X , _⊑₁_) (X , _⊑₀_) (idfun X)) + (g (f (φ , ψ))) (φ , ψ) + + f-equiv : (p : _⊑₀_ ≡ _⊑₁_) → isContr (fiber f p) + f-equiv p = ((to , from) , eq) , NTS + where + to : isOrderPreserving (X , _⊑₀_) (X , _⊑₁_) (idfun _) + to x y = subst (λ _⊑_ → ⟨ x ⊑₀ y ⟩ → ⟨ x ⊑ y ⟩) p (idfun _) + + from : isOrderPreserving (X , _⊑₁_) (X , _⊑₀_) (idfun _) + from x y = subst (λ _⊑_ → ⟨ x ⊑ y ⟩ → ⟨ x ⊑₀ y ⟩) p (idfun _) + + eq : f (to , from) ≡ p + eq = isSetOrder ℓ₁ X _⊑₀_ _⊑₁_ (f (to , from)) p + + NTS : (fib : fiber f p) → ((to , from) , eq) ≡ fib + NTS ((φ , ψ) , eq) = + Σ≡Prop + (λ i′ → isOfHLevelSuc 2 (isSetOrder ℓ₁ X) _⊑₀_ _⊑₁_ (f i′) p) + (Σ≡Prop + (λ _ → isPropIsOrderPreserving (X , _⊑₁_) (X , _⊑₀_) (idfun _)) + (isPropIsOrderPreserving (X , _⊑₀_) (X , _⊑₁_) (idfun _) to φ)) + +-- We now write down the axioms for a partial order and define posets on top of +-- raw ordered structures. + +isReflexive : {A : Type ℓ₀} → Order ℓ₁ A → hProp (ℓ-max ℓ₀ ℓ₁) +isReflexive {A = X} _⊑_ = ((x : X) → ⟨ x ⊑ x ⟩) , isPropΠ λ x → snd (x ⊑ x) + +isTransitive : {A : Type ℓ₀} → Order ℓ₁ A → hProp (ℓ-max ℓ₀ ℓ₁) +isTransitive {ℓ₀ = ℓ₀} {ℓ₁ = ℓ₁} {A = X} _⊑_ = φ , φ-prop + where + φ : Type (ℓ-max ℓ₀ ℓ₁) + φ = ((x y z : X) → ⟨ x ⊑ y ⇒ y ⊑ z ⇒ x ⊑ z ⟩) + φ-prop : isProp φ + φ-prop = isPropΠ3 λ x y z → snd (x ⊑ y ⇒ y ⊑ z ⇒ x ⊑ z) + +isAntisym : {A : Type ℓ₀} → isSet A → Order ℓ₁ A → hProp (ℓ-max ℓ₀ ℓ₁) +isAntisym {ℓ₀ = ℓ₀} {ℓ₁ = ℓ₁} {A = X} A-set _⊑_ = φ , φ-prop + where + φ : Type (ℓ-max ℓ₀ ℓ₁) + φ = ((x y : X) → ⟨ x ⊑ y ⟩ → ⟨ y ⊑ x ⟩ → x ≡ y) + φ-prop : isProp φ + φ-prop = isPropΠ3 λ x y z → isPropΠ λ _ → A-set x y + +-- The predicate expressing that a given order satisfies the partial order +-- axioms. +satPosetAx : (ℓ₁ : Level) (A : Type ℓ₀) → Order ℓ₁ A → hProp (ℓ-max ℓ₀ ℓ₁) +satPosetAx {ℓ₀ = ℓ₀} ℓ₁ A _⊑_ = φ , φ-prop + where + isPartial : isSet A → hProp (ℓ-max ℓ₀ ℓ₁) + isPartial A-set = isReflexive _⊑_ ⊓ isTransitive _⊑_ ⊓ isAntisym A-set _⊑_ + + φ = Σ[ A-set ∈ isSet A ] ⟨ isPartial A-set ⟩ + φ-prop = isOfHLevelΣ 1 isPropIsSet (λ x → snd (isPartial x)) + +-- The poset structure. +PosetStructure : (ℓ₁ : Level) → Type ℓ₀ → Type (ℓ-max ℓ₀ (ℓ-suc ℓ₁)) +PosetStructure ℓ₁ = AxiomsStructure (Order ℓ₁) λ A _⊑_ → ⟨ satPosetAx ℓ₁ A _⊑_ ⟩ + +isSetPosetStructure : (ℓ₁ : Level) (A : Type ℓ₀) → isSet (PosetStructure ℓ₁ A) +isSetPosetStructure ℓ₁ A = + isSetΣ + (isSetΠ2 λ _ _ → isSetHProp) λ _⊑_ → + isProp→isSet (snd (satPosetAx ℓ₁ A _⊑_)) + +Poset : (ℓ₀ ℓ₁ : Level) → Type (ℓ-max (ℓ-suc ℓ₀) (ℓ-suc ℓ₁)) +Poset ℓ₀ ℓ₁ = TypeWithStr ℓ₀ (PosetStructure ℓ₁) + +-- Some projections for syntactic convenience. + +-- Carrier set of a poset. +∣_∣ₚ : Poset ℓ₀ ℓ₁ → Type ℓ₀ +∣ X , _ ∣ₚ = X + +strₚ : (P : Poset ℓ₀ ℓ₁) → PosetStructure ℓ₁ ∣ P ∣ₚ +strₚ (_ , s) = s + +rel : (P : Poset ℓ₀ ℓ₁) → ∣ P ∣ₚ → ∣ P ∣ₚ → hProp ℓ₁ +rel (_ , _⊑_ , _) = _⊑_ + +syntax rel P x y = x ⊑[ P ] y + +⊑[_]-refl : (P : Poset ℓ₀ ℓ₁) → (x : ∣ P ∣ₚ) → ⟨ x ⊑[ P ] x ⟩ +⊑[_]-refl (_ , _ , _ , ⊑-refl , _) = ⊑-refl + +⊑[_]-trans : (P : Poset ℓ₀ ℓ₁) (x y z : ∣ P ∣ₚ) + → ⟨ x ⊑[ P ] y ⟩ → ⟨ y ⊑[ P ] z ⟩ → ⟨ x ⊑[ P ] z ⟩ +⊑[_]-trans (_ , _ , _ , _ , ⊑-trans , _) = ⊑-trans + +⊑[_]-antisym : (P : Poset ℓ₀ ℓ₁) (x y : ∣ P ∣ₚ) + → ⟨ x ⊑[ P ] y ⟩ → ⟨ y ⊑[ P ] x ⟩ → x ≡ y +⊑[_]-antisym (_ , _ , _ , _ , _ , ⊑-antisym) = ⊑-antisym + +carrier-is-set : (P : Poset ℓ₀ ℓ₁) → isSet ∣ P ∣ₚ +carrier-is-set (_ , _ , is-set , _) = is-set + +-- Definition of a monotonic map amounts to forgetting the partial order axioms. +isMonotonic : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀′ ℓ₁′) → (∣ P ∣ₚ → ∣ Q ∣ₚ) → Type _ +isMonotonic (A , (_⊑₀_ , _)) (B , (_⊑₁_ , _)) = + isOrderPreserving (A , _⊑₀_) (B , _⊑₁_) + +isPropIsMonotonic : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀′ ℓ₁′) + → (f : ∣ P ∣ₚ → ∣ Q ∣ₚ) + → isProp (isMonotonic P Q f) +isPropIsMonotonic (A , (_⊑₀_ , _)) (B , (_⊑₁_ , _)) f = + isPropIsOrderPreserving (A , _⊑₀_) (B , _⊑₁_) f + +-- We collect the type of monotonic maps between two posets in the following +-- type. + +_─m→_ : Poset ℓ₀ ℓ₁ → Poset ℓ₀′ ℓ₁′ → Type _ +_─m→_ P Q = Σ[ f ∈ (∣ P ∣ₚ → ∣ Q ∣ₚ) ] (isMonotonic P Q f) + +-- The identity monotonic map and composition of monotonic maps. + +𝟏m : (P : Poset ℓ₀ ℓ₁) → P ─m→ P +𝟏m P = idfun ∣ P ∣ₚ , (λ x y x⊑y → x⊑y) + +_∘m_ : {P : Poset ℓ₀ ℓ₁} {Q : Poset ℓ₀′ ℓ₁′} {R : Poset ℓ₀′′ ℓ₁′′} + → (Q ─m→ R) → (P ─m→ Q) → (P ─m→ R) +(g , pg) ∘m (f , pf) = g ∘ f , λ x y p → pg (f x) (f y) (pf x y p) + +forget-mono : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀′ ℓ₁′) + ((f , f-mono) (g , g-mono) : P ─m→ Q) + → f ≡ g + → (f , f-mono) ≡ (g , g-mono) +forget-mono P Q (f , f-mono) (g , g-mono) = + Σ≡Prop (λ f → isPropΠ3 λ x y x⊑y → snd (f x ⊑[ Q ] f y)) + +module PosetReasoning (P : Poset ℓ₀ ℓ₁) where + + _⊑⟨_⟩_ : (x : ∣ P ∣ₚ) {y z : ∣ P ∣ₚ} + → ⟨ x ⊑[ P ] y ⟩ → ⟨ y ⊑[ P ] z ⟩ → ⟨ x ⊑[ P ] z ⟩ + _ ⊑⟨ p ⟩ q = ⊑[ P ]-trans _ _ _ p q + + _■ : (x : ∣ P ∣ₚ) → ⟨ x ⊑[ P ] x ⟩ + _■ = ⊑[ P ]-refl + + infixr 0 _⊑⟨_⟩_ + infix 1 _■ + +-- Univalence for posets. + +isAMonotonicEqv : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀′ ℓ₁′) + → ∣ P ∣ₚ ≃ ∣ Q ∣ₚ → Type _ +isAMonotonicEqv (A , (_⊑₀_ , _)) (B , (_⊑₁_ , _)) = + isAnOrderPreservingEqv (A , _⊑₀_) (B , _⊑₁_) + +isPropIsAMonotonicEqv : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀ ℓ₁′) + → (eqv : ∣ P ∣ₚ ≃ ∣ Q ∣ₚ) + → isProp (isAMonotonicEqv P Q eqv) +isPropIsAMonotonicEqv P Q e@(f , _) = + isPropΣ (isPropIsMonotonic P Q f) λ _ → isPropIsMonotonic Q P g + where + g = equivFun (invEquiv e) + +-- We denote by `_≃ₚ_` the type of monotonic poset equivalences. + +_≃ₚ_ : Poset ℓ₀ ℓ₁ → Poset ℓ₀ ℓ₁ → Type _ +_≃ₚ_ P Q = Σ[ i ∈ ∣ P ∣ₚ ≃ ∣ Q ∣ₚ ] isAMonotonicEqv P Q i + +-- From this, we can already establish that posets form an SNS and prove that +-- the category of posets is univalent. + +posetUnivalentStr : SNS {ℓ} (PosetStructure ℓ₁) isAMonotonicEqv +posetUnivalentStr {ℓ₁ = ℓ₁} = + UnivalentStr→SNS + (PosetStructure ℓ₁) + isAMonotonicEqv + (axiomsUnivalentStr _ NTS (SNS→UnivalentStr isAnOrderPreservingEqv orderUnivalentStr)) + where + NTS : (A : Type ℓ) (_⊑_ : Order ℓ₁ A) → isProp ⟨ satPosetAx ℓ₁ A _⊑_ ⟩ + NTS A _⊑_ = snd (satPosetAx ℓ₁ A _⊑_) + +poset-univ₀ : (P Q : Poset ℓ₀ ℓ₁) → (P ≃ₚ Q) ≃ (P ≡ Q) +poset-univ₀ = SIP (SNS→UnivalentStr isAMonotonicEqv posetUnivalentStr) + +-- This result is almost what we want but it is better talk directly about poset +-- _isomorphisms_ rather than equivalences. In the case when types `A` and `B` +-- are sets, the type of isomorphisms between `A` and `B` is equivalent to the +-- type of equivalences betwee them. + +-- Let us start by writing down what a poset isomorphisms is. + +isPosetIso : (P Q : Poset ℓ₀ ℓ₁) → (P ─m→ Q) → Type _ +isPosetIso P Q (f , _) = Σ[ (g , _) ∈ (Q ─m→ P) ] section f g × retract f g + +isPosetIso-prop : (P Q : Poset ℓ₀ ℓ₁) (f : P ─m→ Q) + → isProp (isPosetIso P Q f) +isPosetIso-prop P Q (f , f-mono) (g₀ , sec₀ , ret₀) (g₁ , sec₁ , ret₁) = + Σ≡Prop NTS g₀=g₁ + where + NTS : ((g , _) : Q ─m→ P) → isProp (section f g × retract f g) + NTS (g , g-mono) = isPropΣ + (isPropΠ λ x → carrier-is-set Q (f (g x)) x) λ _ → + isPropΠ λ x → carrier-is-set P (g (f x)) x + + g₀=g₁ : g₀ ≡ g₁ + g₀=g₁ = + forget-mono Q P g₀ g₁ (funExt λ x → + fst g₀ x ≡⟨ sym (cong (λ - → fst g₀ -) (sec₁ x)) ⟩ + fst g₀ (f (fst g₁ x)) ≡⟨ ret₀ (fst g₁ x) ⟩ + fst g₁ x ∎) + +-- We will denote by `P ≅ₚ Q` the type of isomorphisms between posets `P` and +-- `Q`. + +_≅ₚ_ : Poset ℓ₀ ℓ₁ → Poset ℓ₀ ℓ₁ → Type _ +P ≅ₚ Q = Σ[ f ∈ P ─m→ Q ] isPosetIso P Q f + +-- ≅ₚ is equivalent to ≃ₚ. + +≃ₚ≃≅ₚ : (P Q : Poset ℓ₀ ℓ₁) → (P ≅ₚ Q) ≃ (P ≃ₚ Q) +≃ₚ≃≅ₚ P Q = isoToEquiv (iso from to ret sec) + where + to : P ≃ₚ Q → P ≅ₚ Q + to (e@(f , _) , (f-mono , g-mono)) = + (f , f-mono) , (g , g-mono) , sec-f-g , ret-f-g + where + is = equivToIso e + g = equivFun (invEquiv e) + + sec-f-g : section f g + sec-f-g = Iso.rightInv (equivToIso e) + + ret-f-g : retract f g + ret-f-g = Iso.leftInv (equivToIso e) + + from : P ≅ₚ Q → P ≃ₚ Q + from ((f , f-mono) , ((g , g-mono) , sec , ret)) = + isoToEquiv is , f-mono , g-mono + where + is : Iso ∣ P ∣ₚ ∣ Q ∣ₚ + is = iso f g sec ret + + sec : section to from + sec (f , _) = Σ≡Prop (isPosetIso-prop P Q) refl + + ret : retract to from + ret (e , _) = Σ≡Prop (isPropIsAMonotonicEqv P Q) (Σ≡Prop isPropIsEquiv refl) + +-- Once we have this equivalence, the main result is then: the type of poset +-- isomorphisms between `P` and `Q` is equivalent to the type of identity proofs +-- between `P` and `Q` + +poset-univ : (P Q : Poset ℓ₀ ℓ₁) → (P ≅ₚ Q) ≃ (P ≡ Q) +poset-univ P Q = P ≅ₚ Q ≃⟨ ≃ₚ≃≅ₚ P Q ⟩ P ≃ₚ Q ≃⟨ poset-univ₀ P Q ⟩ P ≡ Q QED diff --git a/Cubical/Experiments/Problem.agda b/Cubical/Experiments/Problem.agda new file mode 100644 index 0000000000..76a8876e12 --- /dev/null +++ b/Cubical/Experiments/Problem.agda @@ -0,0 +1,150 @@ +-- An example of something where normalization is surprisingly slow +{-# OPTIONS --safe #-} +module Cubical.Experiments.Problem where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Int + +open import Cubical.HITs.S1 +open import Cubical.HITs.S2 +open import Cubical.HITs.S3 +open import Cubical.HITs.Join +open import Cubical.Homotopy.Hopf + +open S¹Hopf + +ptType : Type _ +ptType = Σ Type₀ \ A → A + +pt : (A : ptType) → A .fst +pt A = A .snd + +S¹pt : ptType +S¹pt = (S¹ , base) +S²pt : ptType +S²pt = (S² , base) +S³pt : ptType +S³pt = (S³ , base) +joinpt : ptType +joinpt = (join S¹ S¹ , inl base) + +Ω : (A : ptType) → ptType +Ω A = Path _ (pt A) (pt A) , refl +Ω² : (A : ptType) → ptType +Ω² A = Ω (Ω A) +Ω³ : (A : ptType) → ptType +Ω³ A = Ω² (Ω A) + + +α : join S¹ S¹ → S² +α (inl _) = base +α (inr _) = base +α (push x y i) = (merid y ∙ merid x) i + where + merid : S¹ → Path S² base base + merid base = refl + merid (loop i) = λ j → surf i j + +-- The tests + +test0To2 : Ω³ S³pt .fst +test0To2 i j k = surf i j k + +f3 : Ω³ S³pt .fst → Ω³ joinpt .fst +f3 p i j k = S³→joinS¹S¹ (p i j k) + +test0To3 : Ω³ joinpt .fst +test0To3 = f3 test0To2 + +f4 : Ω³ joinpt .fst → Ω³ S²pt .fst +f4 p i j k = α (p i j k) + +test0To4 : Ω³ S²pt .fst +test0To4 = f4 test0To3 + +innerpath : ∀ i j → HopfS² (test0To4 i j i1) +innerpath i j = transp (λ k → HopfS² (test0To4 i j k)) i0 base + +-- C-c C-n problem uses a lot of memory +problem : pos 0 ≡ pos 0 +problem i = transp (λ j → helix (innerpath i j)) i0 (pos 0) + + +-- Lots of tests: (thanks Evan!) + +winding2 : Path (Path S² base base) refl refl → ℤ +winding2 p = winding (λ j → transp (λ i → HopfS² (p i j)) i0 base) + +test0 : ℤ +test0 = winding2 (λ i j → surf i j) + +test1 : ℤ +test1 = winding2 (λ i j → surf j i) + +test2 : ℤ +test2 = winding2 (λ i j → hcomp (λ _ → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) (surf i j)) + +test3 : ℤ +test3 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) base) + +test4 : ℤ +test4 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) base) + +test5 : ℤ +test5 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → surf k i ; (j = i1) → base}) base) + +test6 : ℤ +test6 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → surf k i}) base) + +test7 : ℤ +test7 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → surf j k ; (j = i0) → base ; (j = i1) → base}) (surf i j)) + +test8 : ℤ +test8 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → surf k i ; (j = i1) → base}) (surf i j)) + +test9 : ℤ +test9 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → surf k i}) (surf i j)) + +test10 : ℤ +test10 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) (surf i j)) + + + +-- Tests using HopfS²' + +winding2' : Path (Path S² base base) refl refl → ℤ +winding2' p = winding (λ j → transp (λ i → HopfS²' (p i j)) i0 base) + +test0' : ℤ +test0' = winding2' (λ i j → surf i j) + +test1' : ℤ +test1' = winding2' (λ i j → surf j i) + +test2' : ℤ +test2' = winding2' (λ i j → hcomp (λ _ → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) (surf i j)) + +test3' : ℤ +test3' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) base) + +test4' : ℤ +test4' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) base) + +test5' : ℤ +test5' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → surf k i ; (j = i1) → base}) base) + +test6' : ℤ +test6' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → surf k i}) base) + +test7' : ℤ +test7' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → surf j k ; (j = i0) → base ; (j = i1) → base}) (surf i j)) + +test8' : ℤ +test8' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → surf k i ; (j = i1) → base}) (surf i j)) + +test9' : ℤ +test9' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → surf k i}) (surf i j)) + +test10' : ℤ +test10' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) (surf i j)) diff --git a/Cubical/Experiments/ZCohomology/Benchmarks.agda b/Cubical/Experiments/ZCohomology/Benchmarks.agda new file mode 100644 index 0000000000..aeb08a96af --- /dev/null +++ b/Cubical/Experiments/ZCohomology/Benchmarks.agda @@ -0,0 +1,388 @@ +{- + +Please do not move this file. Changes should only be made if +necessary. + +This file contains benchmarks for the paper: + +Synthetic Integral Cohomology in Cubical Agda +Guillaume Brunerie, Axel Ljungström, Anders Mörtberg +Computer Science Logic (CSL) 2022 + +Command to run the benchmarks and get timings: + +agda -v profile.definitions:10 Benchmarks.agda + +This assumes that there is no Benchmarks.agdai file. If there is one, +then it should be removed before the above command is run. + +-} + +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomology.Benchmarks where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Int +open import Cubical.HITs.Sn +open import Cubical.Algebra.Group hiding (ℤ ; Bool) +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure hiding (_+ₕ_) renaming (_+'ₕ_ to _+ₕ_) +{- _+'ₕ_ is just (λ x y → (x +ₕ 0ₕ) +ₕ (y +ₕ 0ₕ)) + For technical reason, this gives nicer reductions and computes better in + higher dimensions. -} +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Torus +open import Cubical.ZCohomology.Groups.KleinBottle +open import Cubical.ZCohomology.Groups.WedgeOfSpheres +open import Cubical.ZCohomology.Groups.RP2 +open import Cubical.ZCohomology.Groups.CP2 +open import Cubical.Data.Sigma + +open import Cubical.HITs.KleinBottle +open import Cubical.HITs.RPn.Base +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Pushout +open import Cubical.Homotopy.Hopf +open S¹Hopf +open import Cubical.HITs.Truncation +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 + + +open IsGroupHom +open Iso + +-- S¹ (everything fast) +module S1-tests where + + ϕ : coHom 1 (S₊ 1) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 0)) + + ϕ⁻¹ : ℤ → coHom 1 (S₊ 1) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 0)) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- <10ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- <10ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- <10ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 12ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 13ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ -3 +ₕ ϕ⁻¹ 4) ≡ 1 -- 37ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ -5 +ₕ ϕ⁻¹ -2) ≡ -7 -- 38ms + test₇ = refl + +-- S² +module S2-tests where + + ϕ : coHom 2 (S₊ 2) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 1)) + + ϕ⁻¹ : ℤ → coHom 2 (S₊ 2) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 1)) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 13ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- 17ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 1,152ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 1,235ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 1,208ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 1,153ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- 1,365ms + test₇ = refl + +-- S³ +module S3-tests where + + ϕ : coHom 3 (S₊ 3) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 2)) + + ϕ⁻¹ : ℤ → coHom 3 (S₊ 3) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 2)) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 228ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- 231ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 325ms + test₃ = refl + +{- + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- nope + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- nope + test₅ = refl +-} + +-- S⁴ +module S4-tests where + + ϕ : coHom 4 (S₊ 4) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 3)) + + ϕ⁻¹ : ℤ → coHom 4 (S₊ 4) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 3)) + +{- _+ₕ_ Fails already here... + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- nope + test₁ = refl +-} + +-- ϕ can handle 0ₕ fast + test₂ : ϕ (0ₕ _) ≡ 0 -- < 10ms + test₂ = refl + +{- It fails to map the generator to 1, however. + test₂ : ϕ (∣ ∣_∣ ∣₂) ≡ 1 -- nope + test₂ = refl +-} + +module S1∨S1∨S2-tests₁ where -- everything fast + + ϕ : coHom 1 S²⋁S¹⋁S¹ → ℤ × ℤ + ϕ = fun (fst H¹-S²⋁S¹⋁S¹) + + ϕ⁻¹ : ℤ × ℤ → coHom 1 S²⋁S¹⋁S¹ + ϕ⁻¹ = inv (fst H¹-S²⋁S¹⋁S¹) + + test₁ : ϕ (ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 11ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ (3 , 1)) ≡ (3 , 1) -- 23ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ (0 , 0) +ₕ ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 19ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ (0 , 1) +ₕ ϕ⁻¹ (1 , 0)) ≡ (1 , 1) -- 26ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ (3 , 2) +ₕ ϕ⁻¹ (-1 , 5)) ≡ (2 , 7) -- 62ms + test₅ = refl + + +module S1∨S1∨S2-tests₂ where + + ϕ : coHom 2 S²⋁S¹⋁S¹ → ℤ + ϕ = fun (fst H²-S²⋁S¹⋁S¹) + + ϕ⁻¹ : ℤ → coHom 2 S²⋁S¹⋁S¹ + ϕ⁻¹ = inv (fst H²-S²⋁S¹⋁S¹) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 106ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 125ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 9,689ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 9,235ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 9,748ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 9,136ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- 9,557ms + test₇ = refl + + +module Torus-test₁ where -- fast + + ϕ : coHom 1 (S₊ 1 × S₊ 1) → ℤ × ℤ + ϕ = fun (fst H¹-T²≅ℤ×ℤ) + + ϕ⁻¹ : ℤ × ℤ → coHom 1 (S₊ 1 × S₊ 1) + ϕ⁻¹ = inv (fst H¹-T²≅ℤ×ℤ) + + test₁ : ϕ (ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 11ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ (3 , 1)) ≡ (3 , 1) -- 17ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ (0 , 0) +ₕ ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 19ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ (0 , 1) +ₕ ϕ⁻¹ (1 , 0)) ≡ (1 , 1) -- 26ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ (-3 , 2) +ₕ ϕ⁻¹ (-1 , 5)) ≡ (-4 , 7) -- 61ms + test₅ = refl + + +module Torus-test₂ where + + ϕ : coHom 2 (S₊ 1 × S₊ 1) → ℤ + ϕ = fun (fst H²-T²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 2 (S₊ 1 × S₊ 1) + ϕ⁻¹ = inv (fst H²-T²≅ℤ) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 136sm + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 154ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 12,790ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 12,366ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 12,257ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 13,092ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- 12,528ms + test₇ = refl + + +module Klein-test₁ where -- fast + + ϕ : coHom 1 KleinBottle → ℤ + ϕ = fun (fst H¹-𝕂²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 1 KleinBottle + ϕ⁻¹ = inv (fst H¹-𝕂²≅ℤ) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- <10ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 13ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 10ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 14ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 14ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ -3 +ₕ ϕ⁻¹ 4) ≡ 1 -- 38ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ -5 +ₕ ϕ⁻¹ -2) ≡ -7 -- 38ms + test₇ = refl + + -- The example in the paper: + test : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 2) ≡ 3 -- 22ms + test = refl + + +module Klein-test₂ where + ϕ : coHom 2 KleinBottle → Bool + ϕ = fun (fst H²-𝕂²≅Bool) + + ϕ⁻¹ : Bool → coHom 2 KleinBottle + ϕ⁻¹ = inv (fst H²-𝕂²≅Bool) + +{- + test₀ : ϕ (0ₕ _) ≡ true -- fails already here... + test₀ = refl +-} + + +module RP2-test₂ where + ϕ : coHom 2 RP² → Bool + ϕ = fun (fst H²-RP²≅Bool) + + ϕ⁻¹ : Bool → coHom 2 RP² + ϕ⁻¹ = inv (fst H²-RP²≅Bool) + + test₀ : ϕ (0ₕ _) ≡ true -- 1,328ms (unlike for Klein, this works) + test₀ = refl + +{- + test₁ : ϕ (ϕ⁻¹ true) ≡ true -- nope + test₁ = refl +-} + + +module CP2-test₂ where + ϕ : coHom 2 CP² → ℤ + ϕ = fun (fst H²CP²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 2 CP² + ϕ⁻¹ = inv (fst H²CP²≅ℤ) + + -- For explicitly constructed elements g : H²CP², ϕ works well + test₀ : ϕ (0ₕ _) ≡ 0 -- <10ms + test₀ = refl + + generator : coHom 2 CP² + generator = ∣ (λ { (inl x) → ∣ x ∣ ; (inr x) → 0ₖ _ ; (push a i) → p a i}) ∣₂ + where + ind : (B : TotalHopf → Type) → ((x : _) → isOfHLevel 3 (B x)) → B (north , base) → (x : _) → B x + ind = + transport (λ i → (B : isoToPath IsoS³TotalHopf i → Type) + → ((x : _) → isOfHLevel 3 (B x)) + → B (transp (λ j → isoToPath IsoS³TotalHopf (i ∨ ~ j)) i (north , base)) → (x : _) → B x) + λ B hLev ind → sphereElim _ (λ _ → hLev _) ind + + p : (a : TotalHopf) → ∣ fst a ∣ ≡ 0ₖ 2 + p = ind _ (λ _ → isOfHLevelTrunc 4 _ _) refl + + test₁ : ϕ generator ≡ 1 -- 24ms + test₁ = refl + + + -- For _+ₕ_ too + test₂ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 1,343ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 1,302ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 2) ≡ 4 -- 1,410ms + test₄ = refl + + +module CP2-test₄ where + ϕ : coHom 4 CP² → ℤ + ϕ = fun (fst H⁴CP²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 4 CP² + ϕ⁻¹ = inv (fst H⁴CP²≅ℤ) + +{- + test₀ : ϕ (0ₕ _) ≡ 0 -- fails already here... + test₀ = refl +-} diff --git a/Cubical/Experiments/ZCohomologyExperiments.agda b/Cubical/Experiments/ZCohomologyExperiments.agda new file mode 100644 index 0000000000..62ab820d19 --- /dev/null +++ b/Cubical/Experiments/ZCohomologyExperiments.agda @@ -0,0 +1,79 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyExperiments where +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.WedgeOfSpheres +open import Cubical.Foundations.Prelude + +open import Cubical.HITs.Truncation +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.Data.Int +open import Cubical.Foundations.Equiv +open import Cubical.Data.Sigma +open import Cubical.Foundations.Isomorphism +open import Cubical.Algebra.Group hiding (Int) + +private + + ⋁-to : coHom 2 S²⋁S¹⋁S¹ → Int + ⋁-to = Iso.fun (GroupIso.isom H²-S²⋁S¹⋁S¹) + ⋁-from : Int → coHom 2 S²⋁S¹⋁S¹ + ⋁-from = Iso.inv (GroupIso.isom H²-S²⋁S¹⋁S¹) + + g : S²⋁S¹⋁S¹ → ∥ S₊ 2 ∥ 4 + g (inl x) = ∣ x ∣ + g (inr x) = 0ₖ 2 + g (push a i) = 0ₖ 2 + + G = ∣ g ∣₂ + +{- +This computes: +⋁-to (⋁-from 1 +ₕ ⋁-from 1) ≡ 2 +⋁-to = refl + +We have ⋁-from 1 ≡ G and G is much simpler + +But this does not compute: +test₀ : ⋁-to (G +ₕ G) ≡ 2 +test₀ = refl + +With the strange addition it does: +test₁ : ⋁-to (G +'ₕ G) ≡ 2 +test₁ = refl +-} + +-- Similar stuff happens with Sⁿ. +private + S²-to : coHom 2 (S₊ 2) → Int + S²-to = Iso.fun (GroupIso.isom (Hⁿ-Sⁿ≅ℤ 1)) + + S²-from : Int → coHom 2 (S₊ 2) + S²-from = Iso.inv (GroupIso.isom (Hⁿ-Sⁿ≅ℤ 1)) + + one : coHom 2 (S₊ 2) + one = ∣ ∣_∣ ∣₂ + +{- +Does not compute +test₂ : S²-to (S²-from 1 +ₕ S²-from 1) ≡ 2 +test₂ = refl + +But this does +test₂ : S²-to (S²-from 1 +'ₕ S²-from 1) ≡ 2 +test₂ = refl + +This doesn't +test₃ : S²-to (one +ₕ one) ≡ 2 +test₃ = refl + +But this does +test₃ : S²-to (one +'ₕ one) ≡ 2 +test₃ = refl +-} diff --git a/Cubical/Experiments/ZCohomologyOld/Base.agda b/Cubical/Experiments/ZCohomologyOld/Base.agda new file mode 100644 index 0000000000..83854ddbfc --- /dev/null +++ b/Cubical/Experiments/ZCohomologyOld/Base.agda @@ -0,0 +1,46 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyOld.Base where + +open import Cubical.Data.Int.Base hiding (_+_) +open import Cubical.Data.Nat.Base +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Pointed.Base + +open import Cubical.HITs.Nullification.Base +open import Cubical.HITs.SetTruncation.Base +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.S1.Base +open import Cubical.HITs.Susp.Base +open import Cubical.HITs.Truncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +--- Cohomology --- + +{- EM-spaces Kₙ from Brunerie 2016 -} +coHomK : (n : ℕ) → Type₀ +coHomK zero = ℤ +coHomK (suc n) = ∥ S₊ (suc n) ∥ (2 + suc n) + +{- Cohomology -} +coHom : (n : ℕ) → Type ℓ → Type ℓ +coHom n A = ∥ (A → coHomK n) ∥₂ + +--- Reduced cohomology --- + +coHom-pt : (n : ℕ) → coHomK n +coHom-pt 0 = 0 +coHom-pt 1 = ∣ base ∣ +coHom-pt (suc (suc n)) = ∣ north ∣ + +{- Pointed version of Kₙ -} +coHomK-ptd : (n : ℕ) → Pointed (ℓ-zero) +coHomK-ptd n = coHomK n , coHom-pt n + +{- Reduced cohomology -} +coHomRed : (n : ℕ) → (A : Pointed ℓ) → Type ℓ +coHomRed n A = ∥ A →∙ coHomK-ptd n ∥₂ diff --git a/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda b/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda new file mode 100644 index 0000000000..9ff19c27ac --- /dev/null +++ b/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda @@ -0,0 +1,213 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyOld.KcompPrelims where + +open import Cubical.ZCohomology.Base +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Hopf +open S¹Hopf +-- open import Cubical.Homotopy.Freudenthal hiding (encode) +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; map to trMap) + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv.HalfAdjoint + +open import Cubical.Data.Int renaming (_+_ to +Int) hiding (_·_) +open import Cubical.Data.Nat hiding (_·_) +open import Cubical.Data.Unit + +open import Cubical.HITs.Susp +open import Cubical.HITs.Nullification +open import Cubical.Data.Prod.Base +open import Cubical.Homotopy.Loopspace +open import Cubical.Data.Bool +open import Cubical.Data.Sum.Base +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.Foundations.Function +open import Cubical.Foundations.Pointed +open import Cubical.HITs.S3 + +private + variable + ℓ : Level + A : Type ℓ + +{- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} + +-- Proof of Kₙ ≃ ∥ ΩSⁿ⁺¹ ∥ₙ for $n ≥ 2$ +-- Entirely based on Cavallos proof of Freudenthal in Cubical.Homotopy.Freudenthal +module miniFreudenthal (n : HLevel) where + σ : S₊ (2 + n) → typ (Ω (S₊∙ (3 + n))) + σ a = merid a ∙ merid north ⁻¹ + + S2+n = S₊ (2 + n) + 4n+2 = (2 + n) + (2 + n) + + module WC-S (p : north ≡ north) where + P : (a b : S₊ (2 + n)) → Type₀ + P a b = σ b ≡ p → hLevelTrunc 4n+2 (fiber (λ x → merid x ∙ merid a ⁻¹) p) + + hLevelP : (a b : S₊ (2 + n)) → isOfHLevel 4n+2 (P a b) + hLevelP _ _ = isOfHLevelΠ 4n+2 λ _ → isOfHLevelTrunc 4n+2 + + leftFun : (a : S₊ (2 + n)) → P a north + leftFun a r = ∣ a , (rCancel' (merid a) ∙ rCancel' (merid north) ⁻¹) ∙ r ∣ + + rightFun : (b : S₊ (2 + n)) → P north b + rightFun b r = ∣ b , r ∣ + + funsAgree : leftFun north ≡ rightFun north + funsAgree = + funExt λ r → (λ i → ∣ north , rCancel' (rCancel' (merid north)) i ∙ r ∣) + ∙ λ i → ∣ north , lUnit r (~ i) ∣ + + totalFun : (a b : S2+n) → P a b + totalFun = wedgeconFun (suc n) (suc n) hLevelP rightFun leftFun funsAgree + + leftId : (λ x → totalFun x north) ≡ leftFun + leftId x i = wedgeconRight (suc n) (suc n) hLevelP rightFun leftFun funsAgree i x + + fwd : (p : north ≡ north) (a : S2+n) + → hLevelTrunc 4n+2 (fiber σ p) + → hLevelTrunc 4n+2 (fiber (λ x → merid x ∙ merid a ⁻¹) p) + fwd p a = trRec (isOfHLevelTrunc 4n+2) (uncurry (WC-S.totalFun p a)) + + fwdnorth : (p : north ≡ north) → fwd p north ≡ idfun _ + fwdnorth p = funExt (trElim (λ _ → isOfHLevelPath 4n+2 (isOfHLevelTrunc 4n+2) _ _) + λ p → refl) + + isEquivFwd : (p : north ≡ north) (a : S2+n) → isEquiv (fwd p a) + isEquivFwd p = + suspToPropElim (ptSn (suc n)) + (λ _ → isPropIsEquiv _) + helper + where + helper : isEquiv (fwd p north) + helper = subst isEquiv (sym (fwdnorth p)) (idIsEquiv _) + + interpolate : (a : S2+n) + → PathP (λ i → S2+n → north ≡ merid a i) (λ x → merid x ∙ merid a ⁻¹) merid + interpolate a i x j = compPath-filler (merid x) (merid a ⁻¹) (~ i) j + + Code : (y : Susp S2+n) → north ≡ y → Type₀ + Code north p = hLevelTrunc 4n+2 (fiber σ p) + Code south q = hLevelTrunc 4n+2 (fiber merid q) + Code (merid a i) p = + Glue + (hLevelTrunc 4n+2 (fiber (interpolate a i) p)) + (λ + { (i = i0) → _ , (fwd p a , isEquivFwd p a) + ; (i = i1) → _ , idEquiv _ + }) + + encodeS : (y : S₊ (3 + n)) (p : north ≡ y) → Code y p + encodeS y = J Code ∣ north , rCancel' (merid north) ∣ + + encodeMerid : (a : S2+n) → encodeS south (merid a) ≡ ∣ a , refl ∣ + encodeMerid a = + cong (transport (λ i → gluePath i)) + (funExt⁻ (funExt⁻ (WC-S.leftId refl) a) _ ∙ λ i → ∣ a , lem (rCancel' (merid a)) (rCancel' (merid north)) i ∣) + ∙ transport (PathP≡Path gluePath _ _) + (λ i → ∣ a , (λ j k → rCancel-filler' (merid a) i j k) ∣) + where + gluePath : I → Type _ + gluePath i = hLevelTrunc 4n+2 (fiber (interpolate a i) (λ j → merid a (i ∧ j))) + + lem : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : z ≡ y) → (p ∙ q ⁻¹) ∙ q ≡ p + lem p q = assoc p (q ⁻¹) q ⁻¹ ∙ cong (p ∙_) (lCancel q) ∙ rUnit p ⁻¹ + + contractCodeSouth : (p : north ≡ south) (c : Code south p) → encodeS south p ≡ c + contractCodeSouth p = + trElim + (λ _ → isOfHLevelPath 4n+2 (isOfHLevelTrunc 4n+2) _ _) + (uncurry λ a → + J (λ p r → encodeS south p ≡ ∣ a , r ∣) (encodeMerid a)) + + isConnectedMerid : isConnectedFun 4n+2 (merid {A = S2+n}) + isConnectedMerid p = encodeS south p , contractCodeSouth p + + isConnectedσ : isConnectedFun 4n+2 σ + isConnectedσ = + transport (λ i → isConnectedFun 4n+2 (interpolate north (~ i))) isConnectedMerid + +isConnectedσ-Sn : (n : ℕ) → isConnectedFun (4 + n) (miniFreudenthal.σ n) +isConnectedσ-Sn n = + isConnectedFunSubtr _ n _ + (subst (λ x → isConnectedFun x (miniFreudenthal.σ n)) + helper + (miniFreudenthal.isConnectedσ n)) + where + helper : 2 + (n + (2 + n)) ≡ n + (4 + n) + helper = cong suc (sym (+-suc n _)) ∙ sym (+-suc n _) + +stabSpheres-n≥2 : (n : ℕ) → Iso (hLevelTrunc (4 + n) (S₊ (2 + n))) + (hLevelTrunc (4 + n) (typ (Ω (S₊∙ (3 + n))))) +stabSpheres-n≥2 n = connectedTruncIso (4 + n) (miniFreudenthal.σ n) (isConnectedσ-Sn n) + +-- + +ϕ : (pt a : A) → typ (Ω (Susp A , north)) +ϕ pt a = (merid a) ∙ sym (merid pt) + +private + Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) + Kn→ΩKn+1 zero x i = ∣ intLoop x i ∣ + Kn→ΩKn+1 (suc zero) = trRec (isOfHLevelTrunc 4 ∣ north ∣ ∣ north ∣) + λ a i → ∣ ϕ base a i ∣ + Kn→ΩKn+1 (suc (suc n)) = trRec (isOfHLevelTrunc (2 + (3 + n)) ∣ north ∣ ∣ north ∣) + λ a i → ∣ ϕ north a i ∣ + + d-map : typ (Ω ((Susp S¹) , north)) → S¹ + d-map p = subst HopfSuspS¹ p base + + d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r + d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ + rotLemma r + where + rotLemma : (r : S¹) → r · base ≡ r + rotLemma base = refl + rotLemma (loop i) = refl + +sphereConnectedSpecCase : isConnected 4 (Susp (Susp S¹)) +sphereConnectedSpecCase = sphereConnected 3 + +d-mapComp : Iso (fiber d-map base) (Path (S₊ 3) north north) +d-mapComp = compIso (IsoΣPathTransportPathΣ {B = HopfSuspS¹} _ _) + (congIso (invIso IsoS³TotalHopf)) + +is1Connected-dmap : isConnectedFun 3 d-map +is1Connected-dmap = toPropElim (λ _ → isPropIsOfHLevel 0) + (isConnectedRetractFromIso 3 d-mapComp + (isOfHLevelRetractFromIso 0 (invIso (PathIdTruncIso 3)) + contrHelper)) + where + contrHelper : isContr (Path (∥ Susp (Susp S¹) ∥ 4) ∣ north ∣ ∣ north ∣) + fst contrHelper = refl + snd contrHelper = isOfHLevelPlus {n = 0} 2 (sphereConnected 3) ∣ north ∣ ∣ north ∣ refl + +d-Iso : Iso (∥ Path (S₊ 2) north north ∥ 3) (coHomK 1) +d-Iso = connectedTruncIso _ d-map is1Connected-dmap + +d-mapId2 : Iso.fun d-Iso ∘ trMap (ϕ base) ≡ idfun (coHomK 1) +d-mapId2 = funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a i → ∣ d-mapId a i ∣) + +Iso∥ϕ₁∥ : Iso (coHomK 1) (∥ Path (S₊ 2) north north ∥ 3) +Iso∥ϕ₁∥ = composesToId→Iso d-Iso (trMap (ϕ base)) d-mapId2 + +Iso-Kn-ΩKn+1 : (n : HLevel) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +Iso-Kn-ΩKn+1 zero = invIso (compIso (congIso (truncIdempotentIso _ isGroupoidS¹)) ΩS¹Isoℤ) +Iso-Kn-ΩKn+1 (suc zero) = compIso Iso∥ϕ₁∥ (invIso (PathIdTruncIso 3)) +Iso-Kn-ΩKn+1 (suc (suc n)) = compIso (stabSpheres-n≥2 n) + (invIso (PathIdTruncIso (4 + n))) + where + helper : n + (4 + n) ≡ 2 + (n + (2 + n)) + helper = +-suc n (3 + n) ∙ (λ i → suc (+-suc n (2 + n) i)) diff --git a/Cubical/Experiments/ZCohomologyOld/MayerVietorisUnreduced.agda b/Cubical/Experiments/ZCohomologyOld/MayerVietorisUnreduced.agda new file mode 100644 index 0000000000..8ad1449b2e --- /dev/null +++ b/Cubical/Experiments/ZCohomologyOld/MayerVietorisUnreduced.agda @@ -0,0 +1,384 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyOld.MayerVietorisUnreduced where + +open import Cubical.ZCohomology.Base +open import Cubical.Experiments.ZCohomologyOld.Properties +open import Cubical.Experiments.ZCohomologyOld.KcompPrelims hiding (ϕ) + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Sigma +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) +open import Cubical.Data.Nat +open import Cubical.Data.Prod hiding (_×_) +open import Cubical.Algebra.Group +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) + +open GroupHom + +-- "Very" short exact sequences +-- i.e. an exact sequence A → B → C → D where A and D are trivial +record vSES {ℓ ℓ' ℓ'' ℓ'''} (A : Group {ℓ}) (B : Group {ℓ'}) (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) + : Type (ℓ-suc (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')))) where + constructor vses + + field + isTrivialLeft : isProp ⟨ leftGr ⟩ + isTrivialRight : isProp ⟨ rightGr ⟩ + + left : GroupHom leftGr A + right : GroupHom B rightGr + ϕ : GroupHom A B + + Ker-ϕ⊂Im-left : (x : ⟨ A ⟩) + → isInKer ϕ x + → isInIm left x + Ker-right⊂Im-ϕ : (x : ⟨ B ⟩) + → isInKer right x + → isInIm ϕ x + +open BijectionIso +open vSES + +vSES→GroupIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) + → vSES A B leftGr rightGr + → GroupIso A B +vSES→GroupIso {A = A} lGr rGr isvses = BijectionIsoToGroupIso theIso + where + theIso : BijectionIso _ _ + fun theIso = vSES.ϕ isvses + inj theIso a inker = pRec (isSetCarrier A _ _) + (λ (a , p) → sym p + ∙∙ cong (fun (left isvses)) (isTrivialLeft isvses a _) + ∙∙ morph1g→1g lGr A (left isvses)) + (Ker-ϕ⊂Im-left isvses a inker) + surj theIso a = Ker-right⊂Im-ϕ isvses a (isTrivialRight isvses _ _) + +vSES→GroupEquiv : {ℓ ℓ₁ ℓ₂ ℓ₃ : Level} {A : Group {ℓ}} {B : Group {ℓ₁}} (leftGr : Group {ℓ₂}) (rightGr : Group {ℓ₃}) + → vSES A B leftGr rightGr + → GroupEquiv A B +vSES→GroupEquiv lGr rGr isvses = GrIsoToGrEquiv (vSES→GroupIso lGr rGr isvses) + +module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where + -- Proof from Brunerie 2016. + -- We first define the three morphisms involved: i, Δ and d. + + private + i* : (n : ℕ) → coHom n (Pushout f g) → coHom n A × coHom n B + i* _ = sRec (isSet× isSetSetTrunc isSetSetTrunc) λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ + + iIsHom : (n : ℕ) → isGroupHom (coHomGr n (Pushout f g)) (×coHomGr n A B) (i* n) + iIsHom _ = sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ → refl + + i : (n : ℕ) → GroupHom (coHomGr n (Pushout f g)) (×coHomGr n A B) + GroupHom.fun (i n) = i* n + GroupHom.isHom (i n) = iIsHom n + + + private + distrLem : (n : ℕ) (x y z w : coHomK n) → (x +[ n ]ₖ y) -[ n ]ₖ (z +[ n ]ₖ w) ≡ (x -[ n ]ₖ z) +[ n ]ₖ (y -[ n ]ₖ w) + distrLem n x y z w = + cong (ΩKn+1→Kn n) (cong₂ (λ q p → q ∙ sym p) (+ₖ→∙ n x y) (+ₖ→∙ n z w) + ∙∙ cong ((Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) ∙_) (symDistr (Kn→ΩKn+1 n z) (Kn→ΩKn+1 n w)) + ∙∙ ((sym (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) _)) + ∙∙ cong (Kn→ΩKn+1 n x ∙_) (assoc (Kn→ΩKn+1 n y) (sym (Kn→ΩKn+1 n w)) (sym (Kn→ΩKn+1 n z))) + ∙∙ (cong (Kn→ΩKn+1 n x ∙_) (isCommΩK (suc n) _ _) + ∙∙ assoc _ _ _ + ∙∙ cong₂ _∙_ (sym (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ sym (Kn→ΩKn+1 n z)))) + (sym (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n y ∙ sym (Kn→ΩKn+1 n w))))))) + + + Δ' : (n : ℕ) → ⟨ ×coHomGr n A B ⟩ → ⟨ coHomGr n C ⟩ + Δ' n (α , β) = coHomFun n f α -[ n ]ₕ coHomFun n g β + + Δ'-isMorph : (n : ℕ) → isGroupHom (×coHomGr n A B) (coHomGr n C) (Δ' n) + Δ'-isMorph n = + prodElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _ ) + λ f' x1 g' x2 i → ∣ (λ x → distrLem n (f' (f x)) (g' (f x)) (x1 (g x)) (x2 (g x)) i) ∣₂ + + Δ : (n : ℕ) → GroupHom (×coHomGr n A B) (coHomGr n C) + GroupHom.fun (Δ n) = Δ' n + GroupHom.isHom (Δ n) = Δ'-isMorph n + + d-pre : (n : ℕ) → (C → coHomK n) → Pushout f g → coHomK (suc n) + d-pre n γ (inl x) = 0ₖ (suc n) + d-pre n γ (inr x) = 0ₖ (suc n) + d-pre zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i + d-pre (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i + + dHomHelperPath : (n : ℕ) (h l : C → coHomK n) (a : C) → I → I → coHomK (suc n) + dHomHelperPath zero h l a i j = + hcomp (λ k → λ { (i = i0) → lUnitₖ 1 (0ₖ 1) (~ j) + ; (i = i1) → lUnitₖ 1 (0ₖ 1) (~ j) + ; (j = i0) → +ₖ→∙ 0 (h a) (l a) (~ k) i + ; (j = i1) → cong₂Funct (λ x y → x +[ 1 ]ₖ y) + (Kn→ΩKn+1 0 (h a)) (Kn→ΩKn+1 0 (l a)) (~ k) i}) + (bottom i j) + where + bottom : I → I → coHomK 1 + bottom i j = hcomp (λ k → λ { (i = i0) → lUnitₖ 1 (0ₖ 1) (~ j) + ; (i = i1) → lUnitₖ 1 (Kn→ΩKn+1 0 (l a) k) (~ j) }) + (anotherbottom i j) + + where + anotherbottom : I → I → coHomK 1 + anotherbottom i j = hcomp (λ k → λ { (i = i0) → rUnitlUnit0 1 k (~ j) + ; (i = i1) → rUnitlUnit0 1 k (~ j) + ; (j = i0) → Kn→ΩKn+1 0 (h a) i + ; (j = i1) → Kn→ΩKn+1 0 (h a) i +[ 1 ]ₖ 0ₖ 1 }) + (rUnitₖ 1 (Kn→ΩKn+1 0 (h a) i) (~ j)) + dHomHelperPath (suc n) h l a i j = + hcomp (λ k → λ { (i = i0) → lUnitₖ (2 + n) (0ₖ (2 + n)) (~ j) + ; (i = i1) → lUnitₖ (2 + n) (0ₖ (2 + n)) (~ j) + ; (j = i0) → +ₖ→∙ (suc n) (h a) (l a) (~ k) i + ; (j = i1) → cong₂Funct (λ x y → x +[ 2 + n ]ₖ y) + (Kn→ΩKn+1 (suc n) (h a)) (Kn→ΩKn+1 (suc n) (l a)) (~ k) i}) + (bottom i j) + where + bottom : I → I → coHomK (2 + n) + bottom i j = hcomp (λ k → λ { (i = i0) → lUnitₖ (2 + n) (0ₖ (2 + n)) (~ j) + ; (i = i1) → lUnitₖ (2 + n) (Kn→ΩKn+1 (suc n) (l a) k) (~ j) }) + (anotherbottom i j) + + where + anotherbottom : I → I → coHomK (2 + n) + anotherbottom i j = hcomp (λ k → λ { (i = i0) → rUnitlUnit0 (2 + n) k (~ j) + ; (i = i1) → rUnitlUnit0 (2 + n) k (~ j) + ; (j = i0) → Kn→ΩKn+1 (suc n) (h a) i + ; (j = i1) → Kn→ΩKn+1 (suc n) (h a) i +[ 2 + n ]ₖ (0ₖ (2 + n)) }) + (rUnitₖ (2 + n) (Kn→ΩKn+1 (suc n) (h a) i) (~ j)) + + dHomHelper : (n : ℕ) (h l : C → coHomK n) (x : Pushout f g) + → d-pre n (λ x → h x +[ n ]ₖ l x) x ≡ d-pre n h x +[ suc n ]ₖ d-pre n l x + dHomHelper n h l (inl x) = sym (lUnitₖ (suc n) (0ₖ (suc n))) + dHomHelper n h l (inr x) = sym (lUnitₖ (suc n) (0ₖ (suc n))) + dHomHelper zero h l (push a i) j = dHomHelperPath zero h l a i j + dHomHelper (suc n) h l (push a i) j = dHomHelperPath (suc n) h l a i j + + dIsHom : (n : ℕ) → isGroupHom (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (sRec isSetSetTrunc λ a → ∣ d-pre n a ∣₂) + dIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g i → ∣ funExt (λ x → dHomHelper zero f g x) i ∣₂ + dIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g i → ∣ funExt (λ x → dHomHelper (suc n) f g x) i ∣₂ + + d : (n : ℕ) → GroupHom (coHomGr n C) (coHomGr (suc n) (Pushout f g)) + GroupHom.fun (d n) = sRec isSetSetTrunc λ a → ∣ d-pre n a ∣₂ + GroupHom.isHom (d n) = dIsHom n + + -- The long exact sequence + Im-d⊂Ker-i : (n : ℕ) (x : ⟨ (coHomGr (suc n) (Pushout f g)) ⟩) + → isInIm (d n) x + → isInKer (i (suc n)) x + Im-d⊂Ker-i n = sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ a → pRec (isOfHLevelPath' 1 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ δ b i → sRec (isSet× isSetSetTrunc isSetSetTrunc) + (λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ ) (b (~ i))) + + + Ker-i⊂Im-d : (n : ℕ) (x : ⟨ coHomGr (suc n) (Pushout f g) ⟩) + → isInKer (i (suc n)) x + → isInIm (d n) x + Ker-i⊂Im-d zero = + sElim (λ _ → isSetΠ λ _ → isProp→isSet isPropPropTrunc) + λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ 1} + (isProp→ isPropPropTrunc) + (λ p1 → pRec isPropPropTrunc λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn 0 (sym (cong (λ F → F (f c)) p1) + ∙∙ cong a (push c) + ∙∙ cong (λ F → F (g c)) p2)) ∣₂ + , cong ∣_∣₂ (funExt (λ δ → helper a p1 p2 δ)) ∣₁) + (Iso.fun PathIdTrunc₀Iso (cong fst p)) + (Iso.fun PathIdTrunc₀Iso (cong snd p)) + + where + helper : (F : (Pushout f g) → hLevelTrunc 3 (S₊ 1)) + (p1 : Path (_ → hLevelTrunc 3 (S₊ 1)) (λ a₁ → F (inl a₁)) (λ _ → ∣ base ∣)) + (p2 : Path (_ → hLevelTrunc 3 (S₊ 1)) (λ a₁ → F (inr a₁)) (λ _ → ∣ base ∣)) + → (δ : Pushout f g) + → d-pre 0 (λ c → ΩKn+1→Kn 0 ((λ i₁ → p1 (~ i₁) (f c)) + ∙∙ cong F (push c) + ∙∙ cong (λ F → F (g c)) p2)) δ + ≡ F δ + helper F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + helper F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso-Kn-ΩKn+1 0) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (doubleCompPath-filler (sym (cong (λ F → F (f a)) p1)) (cong F (push a)) (cong (λ F → F (g a)) p2) (~ j) i) + Ker-i⊂Im-d (suc n) = + sElim (λ _ → isSetΠ λ _ → isProp→isSet isPropPropTrunc) + λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ (2 + n)} (isProp→ isPropPropTrunc) + (λ p1 → pRec isPropPropTrunc λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn (suc n) (sym (cong (λ F → F (f c)) p1) + ∙∙ cong a (push c) + ∙∙ cong (λ F → F (g c)) p2)) ∣₂ + , cong ∣_∣₂ (funExt (λ δ → helper a p1 p2 δ)) ∣₁) + (Iso.fun PathIdTrunc₀Iso (cong fst p)) + (Iso.fun PathIdTrunc₀Iso (cong snd p)) + + where + helper : (F : (Pushout f g) → hLevelTrunc (4 + n) (S₊ (2 + n))) + (p1 : Path (_ → hLevelTrunc (4 + n) (S₊ (2 + n))) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) + (p2 : Path (_ → hLevelTrunc (4 + n) (S₊ (2 + n))) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) + → (δ : (Pushout f g)) + → d-pre (suc n) (λ c → ΩKn+1→Kn (suc n) ((λ i₁ → p1 (~ i₁) (f c)) + ∙∙ cong F (push c) + ∙∙ cong (λ F → F (g c)) p2)) δ + ≡ F δ + helper F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + helper F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (doubleCompPath-filler (sym (cong (λ F → F (f a)) p1)) (cong F (push a)) (cong (λ F → F (g a)) p2) (~ j) i) + + open GroupHom + + Im-i⊂Ker-Δ : (n : ℕ) (x : ⟨ ×coHomGr n A B ⟩) + → isInIm (i n) x + → isInKer (Δ n) x + Im-i⊂Ker-Δ n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInIm (i n) (Fa , Fb) + → isInKer (Δ n) (Fa , Fb)} + (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ Fa → sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fb → pRec (isSetSetTrunc _ _) + (sigmaElim (λ x → isProp→isSet (isSetSetTrunc _ _)) + λ Fd p → helper n Fa Fb Fd p)) + Fa + Fb + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : (Pushout f g) → coHomK n) + → (fun (i n) ∣ Fd ∣₂ ≡ (∣ Fa ∣₂ , ∣ Fb ∣₂)) + → (fun (Δ n)) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ n + helper zero Fa Fb Fd p = cong (fun (Δ zero)) (sym p) + ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ -[ 0 ]ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂ ) + ∙∙ cancelₕ 0 ∣ (λ x → Fd (inl (f x))) ∣₂ + helper (suc n) Fa Fb Fd p = cong (fun (Δ (suc n))) (sym p) + ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ -[ (suc n) ]ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂) + ∙∙ cancelₕ (suc n) ∣ (λ x → Fd (inl (f x))) ∣₂ + + Ker-Δ⊂Im-i : (n : ℕ) (a : ⟨ ×coHomGr n A B ⟩) + → isInKer (Δ n) a + → isInIm (i n) a + Ker-Δ⊂Im-i n = prodElim (λ _ → isSetΠ (λ _ → isProp→isSet isPropPropTrunc)) + (λ Fa Fb p → pRec isPropPropTrunc + (λ q → ∣ ∣ helpFun Fa Fb q ∣₂ , refl ∣₁) + (helper Fa Fb p)) + where + helper : (Fa : A → coHomK n) (Fb : B → coHomK n) + → fun (Δ n) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ n + → ∥ Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c)) ∥₁ + helper Fa Fb p = Iso.fun PathIdTrunc₀Iso + (sym (-+cancelₕ n ∣ (λ c → Fa (f c)) ∣₂ ∣ (λ c → Fb (g c)) ∣₂) + ∙∙ cong (λ x → x +[ n ]ₕ ∣ (λ c → Fb (g c)) ∣₂) p + ∙∙ lUnitₕ n _) + + helpFun : (Fa : A → coHomK n) (Fb : B → coHomK n) + → ((λ c → Fa (f c)) ≡ (λ c → Fb (g c))) + → Pushout f g → coHomK n + helpFun Fa Fb p (inl x) = Fa x + helpFun Fa Fb p (inr x) = Fb x + helpFun Fa Fb p (push a i) = p i a + + + private + distrHelper : (n : ℕ) (p q : _) + → ΩKn+1→Kn n p +[ n ]ₖ (-[ n ]ₖ ΩKn+1→Kn n q) ≡ ΩKn+1→Kn n (p ∙ sym q) + distrHelper n p q i = + ΩKn+1→Kn n (Iso.rightInv (Iso-Kn-ΩKn+1 n) p i + ∙ Iso.rightInv (Iso-Kn-ΩKn+1 n) (sym (Iso.rightInv (Iso-Kn-ΩKn+1 n) q i)) i) + + Ker-d⊂Im-Δ : (n : ℕ) (a : coHom n C) + → isInKer (d n) a + → isInIm (Δ n) a + Ker-d⊂Im-Δ zero = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + λ Fc p → pRec isPropPropTrunc (λ p → ∣ (∣ (λ a → ΩKn+1→Kn 0 (cong (λ f → f (inl a)) p)) ∣₂ , + ∣ (λ b → ΩKn+1→Kn 0 (cong (λ f → f (inr b)) p)) ∣₂) , + Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 Fc p c) ∣₁ ∣₁) + (Iso.fun (PathIdTrunc₀Iso) p) + + where + + helper2 : (Fc : C → coHomK 0) + (p : d-pre 0 Fc ≡ (λ _ → ∣ base ∣)) (c : C) + → ΩKn+1→Kn 0 (λ i₁ → p i₁ (inl (f c))) -[ 0 ]ₖ (ΩKn+1→Kn 0 (λ i₁ → p i₁ (inr (g c)))) ≡ Fc c + helper2 Fc p c = cong₂ (λ x y → ΩKn+1→Kn 0 (x ∙ sym y)) (Iso.rightInv (Iso-Kn-ΩKn+1 0) (λ i₁ → p i₁ (inl (f c)))) + (Iso.rightInv (Iso-Kn-ΩKn+1 0) (λ i₁ → p i₁ (inr (g c)))) + ∙∙ cong (ΩKn+1→Kn 0) (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) + ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) + ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 zero) (Fc c) + Ker-d⊂Im-Δ (suc n) = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + λ Fc p → pRec isPropPropTrunc (λ p → ∣ (∣ (λ a → ΩKn+1→Kn (suc n) (cong (λ f → f (inl a)) p)) ∣₂ , + ∣ (λ b → ΩKn+1→Kn (suc n) (cong (λ f → f (inr b)) p)) ∣₂) , + Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 Fc p c) ∣₁ ∣₁) + (Iso.fun (PathIdTrunc₀Iso) p) + + where + + helper2 : (Fc : C → coHomK (suc n)) + (p : d-pre (suc n) Fc ≡ (λ _ → ∣ north ∣)) (c : C) + → ΩKn+1→Kn (suc n) (λ i₁ → p i₁ (inl (f c))) -[ (suc n) ]ₖ (ΩKn+1→Kn (suc n) (λ i₁ → p i₁ (inr (g c)))) ≡ Fc c + helper2 Fc p c = cong₂ (λ x y → ΩKn+1→Kn (suc n) (x ∙ sym y)) (Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) (λ i₁ → p i₁ (inl (f c)))) + (Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) (λ i₁ → p i₁ (inr (g c)))) + ∙∙ cong (ΩKn+1→Kn (suc n)) (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) + ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) + ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) (Fc c) + + Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) + → isInIm (Δ n) a + → isInKer (d n) a + Im-Δ⊂Ker-d n = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fc → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + (sigmaProdElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fa Fb p → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + (λ q → ((λ i → fun (d n) ∣ (q (~ i)) ∣₂) ∙ dΔ-Id n Fa Fb)) + (Iso.fun (PathIdTrunc₀Iso) p)) + + where + d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : (Pushout f g)) + → d-pre n (Fa ∘ f) d ≡ 0ₖ (suc n) + d-preLeftId zero Fa (inl x) = Kn→ΩKn+1 0 (Fa x) + d-preLeftId (suc n) Fa (inl x) = Kn→ΩKn+1 (suc n) (Fa x) + d-preLeftId zero Fa (inr x) = refl + d-preLeftId (suc n) Fa (inr x) = refl + d-preLeftId zero Fa (push a i) j = Kn→ΩKn+1 zero (Fa (f a)) (j ∨ i) + d-preLeftId (suc n) Fa (push a i) j = Kn→ΩKn+1 (suc n) (Fa (f a)) (j ∨ i) + + d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : (Pushout f g)) + → d-pre n (Fb ∘ g) d ≡ 0ₖ (suc n) + d-preRightId n Fb (inl x) = refl + d-preRightId zero Fb (inr x) = sym (Kn→ΩKn+1 0 (Fb x)) + d-preRightId (suc n) Fb (inr x) = sym (Kn→ΩKn+1 (suc n) (Fb x)) + d-preRightId zero Fb (push a i) j = Kn→ΩKn+1 zero (Fb (g a)) (~ j ∧ i) + d-preRightId (suc n) Fb (push a i) j = Kn→ΩKn+1 (suc n) (Fb (g a)) (~ j ∧ i) + + dΔ-Id : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → fun (d n) (fun (Δ n) (∣ Fa ∣₂ , ∣ Fb ∣₂)) ≡ 0ₕ (suc n) + dΔ-Id zero Fa Fb = -distrLemma 0 1 (d zero) ∣ Fa ∘ f ∣₂ ∣ Fb ∘ g ∣₂ + ∙∙ (λ i → ∣ (λ x → d-preLeftId zero Fa x i) ∣₂ -[ 1 ]ₕ ∣ (λ x → d-preRightId zero Fb x i) ∣₂) + ∙∙ cancelₕ 1 (0ₕ 1) + dΔ-Id (suc n) Fa Fb = -distrLemma (suc n) (2 + n) (d (suc n)) ∣ Fa ∘ f ∣₂ ∣ Fb ∘ g ∣₂ + ∙∙ (λ i → ∣ (λ x → d-preLeftId (suc n) Fa x i) ∣₂ -[ (2 + n) ]ₕ ∣ (λ x → d-preRightId (suc n) Fb x i) ∣₂) + ∙∙ cancelₕ (2 + n) (0ₕ (2 + n)) diff --git a/Cubical/Experiments/ZCohomologyOld/Properties.agda b/Cubical/Experiments/ZCohomologyOld/Properties.agda new file mode 100644 index 0000000000..6f92b8e4a1 --- /dev/null +++ b/Cubical/Experiments/ZCohomologyOld/Properties.agda @@ -0,0 +1,634 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyOld.Properties where + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed hiding (id) +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; isSetSetTrunc to §) +open import Cubical.Data.Int renaming (_+_ to _ℤ+_) hiding (-_) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) hiding (map2) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.Algebra.Group renaming (Unit to trivialGroup ; ℤ to ℤGroup) +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Data.NatMinusOne + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup + +open import Cubical.Experiments.ZCohomologyOld.KcompPrelims + +open Iso renaming (inv to inv') + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + A' : Pointed ℓ + +infixr 34 _+ₖ_ +infixr 34 _+ₕ_ + +is2ConnectedKn : (n : ℕ) → isConnected 2 (coHomK (suc n)) +is2ConnectedKn zero = ∣ ∣ base ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) + (toPropElim (λ _ → isOfHLevelTrunc 2 _ _) refl)) +is2ConnectedKn (suc n) = ∣ ∣ north ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isOfHLevelTrunc 2 _ _)) + (suspToPropElim (ptSn (suc n)) (λ _ → isOfHLevelTrunc 2 _ _) refl)) + +isConnectedKn : (n : ℕ) → isConnected (2 + n) (coHomK (suc n)) +isConnectedKn n = isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (2 + n) 1)) (sphereConnected (suc n)) + +-- Induction principles for cohomology groups +-- If we want to show a proposition about some x : Hⁿ(A), it suffices to show it under the +-- assumption that x = ∣f∣₂ and that f is pointed + +coHomPointedElim : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (x : coHom (suc n) A) → B x +coHomPointedElim {ℓ' = ℓ'} {A = A} n a isprop indp = + sElim (λ _ → isOfHLevelSuc 1 (isprop _)) + λ f → helper n isprop indp f (f a) refl + where + helper : (n : ℕ) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (f : A → coHomK (suc n)) + → (x : coHomK (suc n)) + → f a ≡ x → B ∣ f ∣₂ + -- pattern matching a bit extra to avoid isOfHLevelPlus' + helper zero isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ λ _ → isprop _)) + (toPropElim (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc zero) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ λ _ → isprop _)) + (suspToPropElim base (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc zero)) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc (suc n))) isprop ind f = + trElim (λ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + +coHomPointedElim2 : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (x y : coHom (suc n) A) → B x y +coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = sElim2 (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) + λ f g → helper n a isprop indp f g (f a) (g a) refl refl + where + helper : (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (f g : A → coHomK (suc n)) + → (x y : coHomK (suc n)) + → f a ≡ x → g a ≡ y + → B ∣ f ∣₂ ∣ g ∣₂ + helper zero a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ2 λ _ _ → isprop _ _)) + (toPropElim2 (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc zero) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 base (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc zero)) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc (suc n))) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + + +{- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} +coHomRed+1Equiv : (n : ℕ) → + (A : Type ℓ) → + (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHomRed+1Equiv zero A i = ∥ helpLemma {C = (_ , pos 0)} i ∥₂ + module coHomRed+1 where + helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →∙ C))) + helpLemma {C = C} = isoToPath (iso map1 + map2 + (λ b → linvPf b) + (λ _ → refl)) + where + map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) + map1 f = map1' , refl + module helpmap where + map1' : A ⊎ Unit → fst C + map1' (inl x) = f x + map1' (inr x) = pt C + + map2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) + map2 (g , pf) x = g (inl x) + + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) + where + helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x + helper (inl x) = refl + helper (inr tt) = sym snd +coHomRed+1Equiv (suc zero) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK 1 , ∣ base ∣)} i ∥₂ +coHomRed+1Equiv (suc (suc n)) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK (2 + n) , ∣ north ∣)} i ∥₂ + +----------- + +Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +Kn→ΩKn+1 n = Iso.fun (Iso-Kn-ΩKn+1 n) + +ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n +ΩKn+1→Kn n = Iso.inv (Iso-Kn-ΩKn+1 n) + +Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) + +---------- Algebra/Group stuff -------- + +0ₖ : (n : ℕ) → coHomK n +0ₖ = coHom-pt + +_+ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n +_+ₖ_ {n = n} x y = ΩKn+1→Kn n (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) + +-ₖ_ : {n : ℕ} → coHomK n → coHomK n +-ₖ_ {n = n} x = ΩKn+1→Kn n (sym (Kn→ΩKn+1 n x)) + +-- subtraction as a binary operator +_-ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n +_-ₖ_ {n = n} x y = ΩKn+1→Kn n (Kn→ΩKn+1 n x ∙ sym (Kn→ΩKn+1 n y)) + ++ₖ-syntax : (n : ℕ) → coHomK n → coHomK n → coHomK n ++ₖ-syntax n = _+ₖ_ {n = n} + +-ₖ-syntax : (n : ℕ) → coHomK n → coHomK n +-ₖ-syntax n = -ₖ_ {n = n} + +-'ₖ-syntax : (n : ℕ) → coHomK n → coHomK n → coHomK n +-'ₖ-syntax n = _-ₖ_ {n = n} + +syntax +ₖ-syntax n x y = x +[ n ]ₖ y +syntax -ₖ-syntax n x = -[ n ]ₖ x +syntax -'ₖ-syntax n x y = x -[ n ]ₖ y + +Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n (0ₖ n) ≡ refl +Kn→ΩKn+10ₖ zero = sym (rUnit refl) +Kn→ΩKn+10ₖ (suc zero) i j = ∣ (rCancel (merid base) i j) ∣ +Kn→ΩKn+10ₖ (suc (suc n)) i j = ∣ (rCancel (merid north) i j) ∣ + +ΩKn+1→Kn-refl : (n : ℕ) → ΩKn+1→Kn n refl ≡ 0ₖ n +ΩKn+1→Kn-refl zero = refl +ΩKn+1→Kn-refl (suc zero) = refl +ΩKn+1→Kn-refl (suc (suc zero)) = refl +ΩKn+1→Kn-refl (suc (suc (suc zero))) = refl +ΩKn+1→Kn-refl (suc (suc (suc (suc zero)))) = refl +ΩKn+1→Kn-refl (suc (suc (suc (suc (suc n))))) = refl + +-0ₖ : {n : ℕ} → -[ n ]ₖ (0ₖ n) ≡ (0ₖ n) +-0ₖ {n = n} = (λ i → ΩKn+1→Kn n (sym (Kn→ΩKn+10ₖ n i))) + ∙∙ (λ i → ΩKn+1→Kn n (Kn→ΩKn+10ₖ n (~ i))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) (0ₖ n) + ++ₖ→∙ : (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (a +[ n ]ₖ b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b ++ₖ→∙ n a b = Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b) + +lUnitₖ : (n : ℕ) (x : coHomK n) → (0ₖ n) +[ n ]ₖ x ≡ x +lUnitₖ 0 x = Iso.leftInv (Iso-Kn-ΩKn+1 zero) x +lUnitₖ (suc zero) = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ x → Iso.leftInv (Iso-Kn-ΩKn+1 1) ∣ x ∣ +lUnitₖ (suc (suc n)) x = + (λ i → ΩKn+1→Kn (2 + n) (Kn→ΩKn+10ₖ (2 + n) i ∙ Kn→ΩKn+1 (2 + n) x)) ∙∙ + (cong (ΩKn+1→Kn (2 + n)) (sym (lUnit (Kn→ΩKn+1 (2 + n) x)))) ∙∙ + Iso.leftInv (Iso-Kn-ΩKn+1 (2 + n)) x +rUnitₖ : (n : ℕ) (x : coHomK n) → x +[ n ]ₖ (0ₖ n) ≡ x +rUnitₖ 0 x = Iso.leftInv (Iso-Kn-ΩKn+1 zero) x +rUnitₖ (suc zero) = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ x → Iso.leftInv (Iso-Kn-ΩKn+1 1) ∣ x ∣ +rUnitₖ (suc (suc n)) x = + (λ i → ΩKn+1→Kn (2 + n) (Kn→ΩKn+1 (2 + n) x ∙ Kn→ΩKn+10ₖ (2 + n) i)) + ∙∙ (cong (ΩKn+1→Kn (2 + n)) (sym (rUnit (Kn→ΩKn+1 (2 + n) x)))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 (2 + n)) x + +rCancelₖ : (n : ℕ) (x : coHomK n) → x +[ n ]ₖ (-[ n ]ₖ x) ≡ (0ₖ n) +rCancelₖ zero x = (λ i → ΩKn+1→Kn 0 (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ + cong (ΩKn+1→Kn 0) (rCancel (Kn→ΩKn+1 zero x)) +rCancelₖ (suc n) x = (λ i → ΩKn+1→Kn (suc n) (Kn→ΩKn+1 (1 + n) x ∙ Iso.rightInv (Iso-Kn-ΩKn+1 (1 + n)) (sym (Kn→ΩKn+1 (1 + n) x)) i)) ∙ + cong (ΩKn+1→Kn (suc n)) (rCancel (Kn→ΩKn+1 (1 + n) x)) ∙ + (λ i → ΩKn+1→Kn (suc n) (Kn→ΩKn+10ₖ (suc n) (~ i))) ∙ + Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) (0ₖ (suc n)) + +lCancelₖ : (n : ℕ) (x : coHomK n) → (-[ n ]ₖ x) +[ n ]ₖ x ≡ (0ₖ n) +lCancelₖ 0 x = (λ i → ΩKn+1→Kn 0 (Iso.rightInv (Iso-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ + cong (ΩKn+1→Kn 0) (lCancel (Kn→ΩKn+1 zero x)) +lCancelₖ (suc n) x = (λ i → ΩKn+1→Kn (suc n) (Iso.rightInv (Iso-Kn-ΩKn+1 (1 + n)) (sym (Kn→ΩKn+1 (1 + n) x)) i ∙ Kn→ΩKn+1 (1 + n) x)) ∙ + cong (ΩKn+1→Kn (suc n)) (lCancel (Kn→ΩKn+1 (1 + n) x)) ∙ + (λ i → (ΩKn+1→Kn (suc n)) (Kn→ΩKn+10ₖ (suc n) (~ i))) ∙ + Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) (0ₖ (suc n)) + +assocₖ : (n : ℕ) (x y z : coHomK n) → ((x +[ n ]ₖ y) +[ n ]ₖ z) ≡ (x +[ n ]ₖ (y +[ n ]ₖ z)) +assocₖ n x y z = ((λ i → ΩKn+1→Kn n (Kn→ΩKn+1 n (ΩKn+1→Kn n (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y)) ∙ Kn→ΩKn+1 n z)) ∙∙ + (λ i → ΩKn+1→Kn n (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙∙ + (λ i → ΩKn+1→Kn n (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) (Kn→ΩKn+1 n z) (~ i)))) ∙ + (λ i → ΩKn+1→Kn n ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) + +cancelₖ : (n : ℕ) (x : coHomK n) → x -[ n ]ₖ x ≡ (0ₖ n) +cancelₖ zero x = cong (ΩKn+1→Kn 0) (rCancel (Kn→ΩKn+1 zero x)) +cancelₖ (suc zero) x = cong (ΩKn+1→Kn 1) (rCancel (Kn→ΩKn+1 1 x)) +cancelₖ (suc (suc zero)) x = cong (ΩKn+1→Kn 2) (rCancel (Kn→ΩKn+1 2 x)) +cancelₖ (suc (suc (suc zero))) x = cong (ΩKn+1→Kn 3) (rCancel (Kn→ΩKn+1 3 x)) +cancelₖ (suc (suc (suc (suc zero)))) x = cong (ΩKn+1→Kn 4) (rCancel (Kn→ΩKn+1 4 x)) +cancelₖ (suc (suc (suc (suc (suc n))))) x = cong (ΩKn+1→Kn (5 + n)) (rCancel (Kn→ΩKn+1 (5 + n) x)) + +-rUnitₖ : (n : ℕ) (x : coHomK n) → x -[ n ]ₖ 0ₖ n ≡ x +-rUnitₖ zero x = rUnitₖ zero x +-rUnitₖ (suc n) x = cong (λ y → ΩKn+1→Kn (suc n) (Kn→ΩKn+1 (suc n) x ∙ sym y)) (Kn→ΩKn+10ₖ (suc n)) + ∙∙ cong (ΩKn+1→Kn (suc n)) (sym (rUnit (Kn→ΩKn+1 (suc n) x))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) x + +abstract + isCommΩK1 : (n : ℕ) → isComm∙ ((Ω^ n) (coHomK-ptd 1)) + isCommΩK1 zero = isCommA→isCommTrunc 2 comm-ΩS¹ isGroupoidS¹ + isCommΩK1 (suc n) = EH n + + open Iso renaming (inv to inv') + isCommΩK : (n : ℕ) → isComm∙ (coHomK-ptd n) + isCommΩK zero p q = isSetℤ _ _ (p ∙ q) (q ∙ p) + isCommΩK (suc zero) = isCommA→isCommTrunc 2 comm-ΩS¹ isGroupoidS¹ + isCommΩK (suc (suc n)) = subst isComm∙ (λ i → coHomK (2 + n) , ΩKn+1→Kn-refl (2 + n) i) (ptdIso→comm {A = (_ , _)} (invIso (Iso-Kn-ΩKn+1 (2 + n))) (EH 0)) + +commₖ : (n : ℕ) (x y : coHomK n) → (x +[ n ]ₖ y) ≡ (y +[ n ]ₖ x) +commₖ 0 x y i = ΩKn+1→Kn 0 (isCommΩK1 0 (Kn→ΩKn+1 0 x) (Kn→ΩKn+1 0 y) i) +commₖ 1 x y i = ΩKn+1→Kn 1 (ptdIso→comm {A = ((∣ north ∣ ≡ ∣ north ∣) , snd ((Ω^ 1) (coHomK 3 , ∣ north ∣)))} + {B = coHomK 2} + (invIso (Iso-Kn-ΩKn+1 2)) (EH 0) (Kn→ΩKn+1 1 x) (Kn→ΩKn+1 1 y) i) +commₖ 2 x y i = ΩKn+1→Kn 2 (ptdIso→comm {A = (∣ north ∣ ≡ ∣ north ∣) , snd ((Ω^ 1) (coHomK 4 , ∣ north ∣))} + {B = coHomK 3} + (invIso (Iso-Kn-ΩKn+1 3)) (EH 0) (Kn→ΩKn+1 2 x) (Kn→ΩKn+1 2 y) i) +commₖ 3 x y i = ΩKn+1→Kn 3 (ptdIso→comm {A = (∣ north ∣ ≡ ∣ north ∣) , snd ((Ω^ 1) (coHomK 5 , ∣ north ∣))} + {B = coHomK 4} + (invIso (Iso-Kn-ΩKn+1 4)) (EH 0) (Kn→ΩKn+1 3 x) (Kn→ΩKn+1 3 y) i) +commₖ (suc (suc (suc (suc n)))) x y i = + ΩKn+1→Kn (4 + n) (ptdIso→comm {A = (∣ north ∣ ≡ ∣ north ∣) , snd ((Ω^ 1) (coHomK (6 + n) , ∣ north ∣))} + {B = coHomK (5 + n)} + (invIso (Iso-Kn-ΩKn+1 (5 + n))) (EH 0) (Kn→ΩKn+1 (4 + n) x) (Kn→ΩKn+1 (4 + n) y) i) + + +rUnitₖ' : (n : ℕ) (x : coHomK n) → x +[ n ]ₖ (0ₖ n) ≡ x +rUnitₖ' n x = commₖ n x (0ₖ n) ∙ lUnitₖ n x + +-distrₖ : (n : ℕ) (x y : coHomK n) → -[ n ]ₖ (x +[ n ]ₖ y) ≡ (-[ n ]ₖ x) +[ n ]ₖ (-[ n ]ₖ y) +-distrₖ n x y = ((λ i → ΩKn+1→Kn n (sym (Kn→ΩKn+1 n (ΩKn+1→Kn n (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y))))) ∙∙ + (λ i → ΩKn+1→Kn n (sym (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i))) ∙∙ + (λ i → ΩKn+1→Kn n (symDistr (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i))) ∙∙ + (λ i → ΩKn+1→Kn n (Iso.rightInv (Iso-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n y)) (~ i) ∙ (Iso.rightInv (Iso-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x)) (~ i)))) ∙∙ + commₖ n (-[ n ]ₖ y) (-[ n ]ₖ x) + +private + rCancelLem : (n : ℕ) (x : coHomK n) → ΩKn+1→Kn n ((Kn→ΩKn+1 n x) ∙ refl) ≡ ΩKn+1→Kn n (Kn→ΩKn+1 n x) + rCancelLem zero x = refl + rCancelLem (suc n) x = cong (ΩKn+1→Kn (suc n)) (sym (rUnit (Kn→ΩKn+1 (suc n) x))) + + lCancelLem : (n : ℕ) (x : coHomK n) → ΩKn+1→Kn n (refl ∙ (Kn→ΩKn+1 n x)) ≡ ΩKn+1→Kn n (Kn→ΩKn+1 n x) + lCancelLem zero x = refl + lCancelLem (suc n) x = cong (ΩKn+1→Kn (suc n)) (sym (lUnit (Kn→ΩKn+1 (suc n) x))) + + +-cancelRₖ : (n : ℕ) (x y : coHomK n) → (y +[ n ]ₖ x) -[ n ]ₖ x ≡ y +-cancelRₖ n x y = (cong (ΩKn+1→Kn n) ((cong (_∙ sym (Kn→ΩKn+1 n x)) (+ₖ→∙ n y x)) + ∙∙ sym (assoc _ _ _) + ∙∙ cong (Kn→ΩKn+1 n y ∙_) (rCancel _))) + ∙∙ rCancelLem n y + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) y + +-cancelLₖ : (n : ℕ) (x y : coHomK n) → (x +[ n ]ₖ y) -[ n ]ₖ x ≡ y +-cancelLₖ n x y = cong (λ z → z -[ n ]ₖ x) (commₖ n x y) ∙ -cancelRₖ n x y + +-+cancelₖ : (n : ℕ) (x y : coHomK n) → (x -[ n ]ₖ y) +[ n ]ₖ y ≡ x +-+cancelₖ n x y = (cong (ΩKn+1→Kn n) ((cong (_∙ (Kn→ΩKn+1 n y)) (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ sym (Kn→ΩKn+1 n y)))) + ∙∙ sym (assoc _ _ _) + ∙∙ cong (Kn→ΩKn+1 n x ∙_) (lCancel _))) + ∙∙ rCancelLem n x + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) x + +---- Group structure of cohomology groups --- + +_+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_+ₕ_ {n = n} = sRec2 § λ a b → ∣ (λ x → a x +[ n ]ₖ b x) ∣₂ + +-ₕ_ : {n : ℕ} → coHom n A → coHom n A +-ₕ_ {n = n} = sRec § λ a → ∣ (λ x → -[ n ]ₖ a x) ∣₂ + +_-ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_-ₕ_ {n = n} = sRec2 § λ a b → ∣ (λ x → a x -[ n ]ₖ b x) ∣₂ + ++ₕ-syntax : (n : ℕ) → coHom n A → coHom n A → coHom n A ++ₕ-syntax n = _+ₕ_ {n = n} + +-ₕ-syntax : (n : ℕ) → coHom n A → coHom n A +-ₕ-syntax n = -ₕ_ {n = n} + +-ₕ'-syntax : (n : ℕ) → coHom n A → coHom n A → coHom n A +-ₕ'-syntax n = _-ₕ_ {n = n} + +syntax +ₕ-syntax n x y = x +[ n ]ₕ y +syntax -ₕ-syntax n x = -[ n ]ₕ x +syntax -ₕ'-syntax n x y = x -[ n ]ₕ y + +0ₕ : (n : ℕ) → coHom n A +0ₕ n = ∣ (λ _ → (0ₖ n)) ∣₂ + +rUnitₕ : (n : ℕ) (x : coHom n A) → x +[ n ]ₕ (0ₕ n) ≡ x +rUnitₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitₖ n (a x)) i ∣₂ + +lUnitₕ : (n : ℕ) (x : coHom n A) → (0ₕ n) +[ n ]ₕ x ≡ x +lUnitₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitₖ n (a x)) i ∣₂ + +rCancelₕ : (n : ℕ) (x : coHom n A) → x +[ n ]ₕ (-[ n ]ₕ x) ≡ 0ₕ n +rCancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelₖ n (a x)) i ∣₂ + +lCancelₕ : (n : ℕ) (x : coHom n A) → (-[ n ]ₕ x) +[ n ]ₕ x ≡ 0ₕ n +lCancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelₖ n (a x)) i ∣₂ + +assocₕ : (n : ℕ) (x y z : coHom n A) → ((x +[ n ]ₕ y) +[ n ]ₕ z) ≡ (x +[ n ]ₕ (y +[ n ]ₕ z)) +assocₕ n = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocₖ n (a x) (b x) (c x)) i ∣₂ + +commₕ : (n : ℕ) (x y : coHom n A) → (x +[ n ]ₕ y) ≡ (y +[ n ]ₕ x) +commₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commₖ n (a x) (b x)) i ∣₂ + +cancelₕ : (n : ℕ) (x : coHom n A) → x -[ n ]ₕ x ≡ 0ₕ n +cancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → cancelₖ n (a x)) i ∣₂ + +-ₖ-ₖ : (n : ℕ) (x : coHomK n) → (-[ n ]ₖ (-[ n ]ₖ x)) ≡ x +-ₖ-ₖ n x = cong ((ΩKn+1→Kn n) ∘ sym) (Iso.rightInv (Iso-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x))) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) x + +-- Proof that rUnitₖ and lUnitₖ agree on 0ₖ. Needed for Mayer-Vietoris. +private + rUnitlUnitGen : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {b : B} (e : Iso A (b ≡ b)) + (0A : A) + (0fun : fun e 0A ≡ refl) + → Path (inv' e (fun e 0A ∙ fun e 0A) ≡ 0A) + (cong (inv' e) (cong (_∙ fun e 0A) 0fun) ∙∙ cong (inv' e) (sym (lUnit (fun e 0A))) ∙∙ Iso.leftInv e 0A) + (cong (inv' e) (cong (fun e 0A ∙_) 0fun) ∙∙ cong (inv' e) (sym (rUnit (fun e 0A))) ∙∙ Iso.leftInv e 0A) + rUnitlUnitGen e 0A 0fun = + (λ i → cong (inv' e) (cong (_∙ fun e 0A) 0fun) ∙∙ rUnit (cong (inv' e) (sym (lUnit (fun e 0A)))) i ∙∙ Iso.leftInv e 0A) + ∙ ((λ i → (λ j → inv' e (0fun (~ i ∧ j) ∙ 0fun (j ∧ i))) + ∙∙ ((λ j → inv' e (0fun (~ i ∨ j) ∙ 0fun i)) + ∙∙ cong (inv' e) (sym (lUnit (0fun i))) + ∙∙ λ j → inv' e (0fun (i ∧ (~ j)))) + ∙∙ Iso.leftInv e 0A) + ∙∙ (λ i → (λ j → inv' e (fun e 0A ∙ 0fun j)) + ∙∙ (λ j → inv' e (0fun (j ∧ ~ i) ∙ refl)) + ∙∙ cong (inv' e) (sym (rUnit (0fun (~ i)))) + ∙∙ (λ j → inv' e (0fun (~ i ∧ ~ j))) + ∙∙ Iso.leftInv e 0A) + ∙∙ λ i → cong (inv' e) (cong (fun e 0A ∙_) 0fun) + ∙∙ rUnit (cong (inv' e) (sym (rUnit (fun e 0A)))) (~ i) + ∙∙ Iso.leftInv e 0A) + +rUnitlUnit0 : (n : ℕ) → rUnitₖ n (0ₖ n) ≡ lUnitₖ n (0ₖ n) +rUnitlUnit0 0 = refl +rUnitlUnit0 (suc zero) = refl +rUnitlUnit0 (suc (suc n)) = + sym (rUnitlUnitGen (Iso-Kn-ΩKn+1 (2 + n)) + (0ₖ (2 + n)) + (Kn→ΩKn+10ₖ (2 + n))) + +-cancelLₕ : (n : ℕ) (x y : coHom n A) → (x +[ n ]ₕ y) -[ n ]ₕ x ≡ y +-cancelLₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelLₖ n (a x) (b x) i) ∣₂ + +-cancelRₕ : (n : ℕ) (x y : coHom n A) → (y +[ n ]ₕ x) -[ n ]ₕ x ≡ y +-cancelRₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelRₖ n (a x) (b x) i) ∣₂ + +-+cancelₕ : (n : ℕ) (x y : coHom n A) → (x -[ n ]ₕ y) +[ n ]ₕ y ≡ x +-+cancelₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -+cancelₖ n (a x) (b x) i) ∣₂ + + +-- Group structure of reduced cohomology groups (in progress - might need K to compute properly first) --- + ++ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A ++ₕ∙ zero = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ zero ]ₖ b x) , (λ i → (pa i +[ zero ]ₖ pb i)) ∣₂ } ++ₕ∙ (suc zero) = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ 1 ]ₖ b x) , (λ i → pa i +[ 1 ]ₖ pb i) ∙ lUnitₖ 1 (0ₖ 1) ∣₂ } ++ₕ∙ (suc (suc n)) = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ (2 + n) ]ₖ b x) , (λ i → pa i +[ (2 + n) ]ₖ pb i) ∙ lUnitₖ (2 + n) (0ₖ (2 + n)) ∣₂ } + +open IsSemigroup +open IsMonoid +open GroupStr +open IsGroupHom + +coHomGr : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Group ℓ +coHomGr n A = coHom n A , coHomGrnA + where + coHomGrnA : GroupStr (coHom n A) + 1g coHomGrnA = 0ₕ n + GroupStr._·_ coHomGrnA = λ x y → x +[ n ]ₕ y + inv coHomGrnA = λ x → -[ n ]ₕ x + isGroup coHomGrnA = helper + where + abstract + helper : IsGroup {G = coHom n A} (0ₕ n) (λ x y → x +[ n ]ₕ y) (λ x → -[ n ]ₕ x) + helper = makeIsGroup § (λ x y z → sym (assocₕ n x y z)) (rUnitₕ n) (lUnitₕ n) (rCancelₕ n) (lCancelₕ n) + +×coHomGr : (n : ℕ) (A : Type ℓ) (B : Type ℓ') → Group _ +×coHomGr n A B = DirProd (coHomGr n A) (coHomGr n B) + +coHomFun : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) (f : A → B) → coHom n B → coHom n A +coHomFun n f = sRec § λ β → ∣ β ∘ f ∣₂ + +-distrLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : GroupHom (coHomGr n A) (coHomGr m B)) + (x y : coHom n A) + → fst f (x -[ n ]ₕ y) ≡ fst f x -[ m ]ₕ fst f y +-distrLemma n m f' x y = sym (-cancelRₕ m (f y) (f (x -[ n ]ₕ y))) + ∙∙ cong (λ x → x -[ m ]ₕ f y) (sym (f' .snd .pres· (x -[ n ]ₕ y) y)) + ∙∙ cong (λ x → x -[ m ]ₕ f y) ( cong f (-+cancelₕ n _ _)) + where + f = fst f' + +--- the loopspace of Kₙ is commutative regardless of base + +addIso : (n : ℕ) (x : coHomK n) → Iso (coHomK n) (coHomK n) +fun (addIso n x) y = y +[ n ]ₖ x +inv' (addIso n x) y = y -[ n ]ₖ x +rightInv (addIso n x) y = -+cancelₖ n y x +leftInv (addIso n x) y = -cancelRₖ n x y + +isCommΩK-based : (n : ℕ) (x : coHomK n) → isComm∙ (coHomK n , x) +isCommΩK-based zero x p q = isSetℤ _ _ (p ∙ q) (q ∙ p) +isCommΩK-based (suc zero) x = + subst isComm∙ (λ i → coHomK 1 , lUnitₖ 1 x i) + (ptdIso→comm {A = (_ , 0ₖ 1)} (addIso 1 x) + (isCommΩK 1)) +isCommΩK-based (suc (suc n)) x = + subst isComm∙ (λ i → coHomK (suc (suc n)) , lUnitₖ (suc (suc n)) x i) + (ptdIso→comm {A = (_ , 0ₖ (suc (suc n)))} (addIso (suc (suc n)) x) + (isCommΩK (suc (suc n)))) + +addLemma : (a b : ℤ) → a +[ 0 ]ₖ b ≡ (a ℤ+ b) +addLemma a b = (cong (ΩKn+1→Kn 0) (sym (congFunct ∣_∣ (intLoop a) (intLoop b)))) + ∙∙ (λ i → ΩKn+1→Kn 0 (cong ∣_∣ (intLoop-hom a b i))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 0) (a ℤ+ b) + +--- +-- hidden versions of cohom stuff using the "lock" hack. The locked versions can be used when proving things. +-- Swapping "key" for "tt*" will then give computing functions. + +Unit' : Type₀ +Unit' = lockUnit {ℓ-zero} + +lock : ∀ {ℓ} {A : Type ℓ} → Unit' → A → A +lock unlock = λ x → x + +module lockedCohom (key : Unit') where + +K : (n : ℕ) → coHomK n → coHomK n → coHomK n + +K n = lock key (_+ₖ_ {n = n}) + + -K : (n : ℕ) → coHomK n → coHomK n + -K n = lock key (-ₖ_ {n = n}) + + -Kbin : (n : ℕ) → coHomK n → coHomK n → coHomK n + -Kbin n = lock key (_-ₖ_ {n = n}) + + rUnitK : (n : ℕ) (x : coHomK n) → +K n x (0ₖ n) ≡ x + rUnitK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (0ₖ n) ≡ x + pm unlock = rUnitₖ n x + + lUnitK : (n : ℕ) (x : coHomK n) → +K n (0ₖ n) x ≡ x + lUnitK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (0ₖ n) x ≡ x + pm unlock = lUnitₖ n x + + rCancelK : (n : ℕ) (x : coHomK n) → +K n x (-K n x) ≡ 0ₖ n + rCancelK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (lock t (-ₖ_ {n = n}) x) ≡ 0ₖ n + pm unlock = rCancelₖ n x + + lCancelK : (n : ℕ) (x : coHomK n) → +K n (-K n x) x ≡ 0ₖ n + lCancelK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (-ₖ_ {n = n}) x) x ≡ 0ₖ n + pm unlock = lCancelₖ n x + + -cancelRK : (n : ℕ) (x y : coHomK n) → -Kbin n (+K n y x) x ≡ y + -cancelRK n x y = pm key + where + pm : (t : Unit') → lock t (_-ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) y x) x ≡ y + pm unlock = -cancelRₖ n x y + + -cancelLK : (n : ℕ) (x y : coHomK n) → -Kbin n (+K n x y) x ≡ y + -cancelLK n x y = pm key + where + pm : (t : Unit') → lock t (_-ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x y) x ≡ y + pm unlock = -cancelLₖ n x y + + -+cancelK : (n : ℕ) (x y : coHomK n) → +K n (-Kbin n x y) y ≡ x + -+cancelK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_-ₖ_ {n = n}) x y) y ≡ x + pm unlock = -+cancelₖ n x y + + cancelK : (n : ℕ) (x : coHomK n) → -Kbin n x x ≡ 0ₖ n + cancelK n x = pm key + where + pm : (t : Unit') → (lock t (_-ₖ_ {n = n}) x x) ≡ 0ₖ n + pm unlock = cancelₖ n x + + assocK : (n : ℕ) (x y z : coHomK n) → +K n (+K n x y) z ≡ +K n x (+K n y z) + assocK n x y z = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x y) z + ≡ lock t (_+ₖ_ {n = n}) x (lock t (_+ₖ_ {n = n}) y z) + pm unlock = assocₖ n x y z + + commK : (n : ℕ) (x y : coHomK n) → +K n x y ≡ +K n y x + commK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x y ≡ lock t (_+ₖ_ {n = n}) y x + pm unlock = commₖ n x y + + -- cohom + + +H : (n : ℕ) (x y : coHom n A) → coHom n A + +H n = sRec2 § λ a b → ∣ (λ x → +K n (a x) (b x)) ∣₂ + + -H : (n : ℕ) (x : coHom n A) → coHom n A + -H n = sRec § λ a → ∣ (λ x → -K n (a x)) ∣₂ + + -Hbin : (n : ℕ) → coHom n A → coHom n A → coHom n A + -Hbin n = sRec2 § λ a b → ∣ (λ x → -Kbin n (a x) (b x)) ∣₂ + + rUnitH : (n : ℕ) (x : coHom n A) → +H n x (0ₕ n) ≡ x + rUnitH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitK n (a x)) i ∣₂ + + lUnitH : (n : ℕ) (x : coHom n A) → +H n (0ₕ n) x ≡ x + lUnitH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitK n (a x)) i ∣₂ + + rCancelH : (n : ℕ) (x : coHom n A) → +H n x (-H n x) ≡ 0ₕ n + rCancelH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelK n (a x)) i ∣₂ + + lCancelH : (n : ℕ) (x : coHom n A) → +H n (-H n x) x ≡ 0ₕ n + lCancelH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelK n (a x)) i ∣₂ + + assocH : (n : ℕ) (x y z : coHom n A) → (+H n (+H n x y) z) ≡ (+H n x (+H n y z)) + assocH n = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocK n (a x) (b x) (c x)) i ∣₂ + + commH : (n : ℕ) (x y : coHom n A) → (+H n x y) ≡ (+H n y x) + commH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commK n (a x) (b x)) i ∣₂ + + -cancelRH : (n : ℕ) (x y : coHom n A) → -Hbin n (+H n y x) x ≡ y + -cancelRH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelRK n (a x) (b x) i) ∣₂ + + -cancelLH : (n : ℕ) (x y : coHom n A) → -Hbin n (+H n x y) x ≡ y + -cancelLH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelLK n (a x) (b x) i) ∣₂ + + -+cancelH : (n : ℕ) (x y : coHom n A) → +H n (-Hbin n x y) y ≡ x + -+cancelH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -+cancelK n (a x) (b x) i) ∣₂ + + ++K→∙ : (key : Unit') (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (lockedCohom.+K key n a b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b ++K→∙ unlock = +ₖ→∙ + ++H≡+ₕ : (key : Unit') (n : ℕ) → lockedCohom.+H key {A = A} n ≡ _+ₕ_ {n = n} ++H≡+ₕ unlock _ = refl + +rUnitlUnit0K : (key : Unit') (n : ℕ) → lockedCohom.rUnitK key n (0ₖ n) ≡ lockedCohom.lUnitK key n (0ₖ n) +rUnitlUnit0K unlock = rUnitlUnit0 diff --git a/Cubical/Experiments/ZariskiLatticeBasicOpens.agda b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda new file mode 100644 index 0000000000..50c15cb110 --- /dev/null +++ b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda @@ -0,0 +1,249 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Experiments.ZariskiLatticeBasicOpens where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Structure +open import Cubical.Functions.FunExtEquiv + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary +open import Cubical.Relation.Binary.Poset + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Localisation.Base +open import Cubical.Algebra.CommRing.Localisation.UniversalProperty +open import Cubical.Algebra.CommRing.Localisation.InvertingElements +open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Properties +open import Cubical.Algebra.CommAlgebra.Localisation +open import Cubical.Algebra.RingSolver.Reflection +open import Cubical.Algebra.Semilattice + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open Iso +open BinaryRelation +open isEquivRel + +private + variable + ℓ ℓ' : Level + + +module Presheaf (A' : CommRing ℓ) where + open CommRingStr (snd A') renaming (_·_ to _·r_ ; ·Comm to ·r-comm ; ·Assoc to ·rAssoc + ; ·Lid to ·rLid ; ·Rid to ·rRid) + open Exponentiation A' + open CommRingTheory A' + open InvertingElementsBase A' + open isMultClosedSubset + open CommAlgebraStr ⦃...⦄ + private + A = fst A' + + A[1/_] : A → CommAlgebra A' ℓ + A[1/ x ] = AlgLoc.S⁻¹RAsCommAlg A' [ x ⁿ|n≥0] (powersFormMultClosedSubset _) + + A[1/_]ˣ : (x : A) → ℙ (fst A[1/ x ]) + A[1/ x ]ˣ = (CommAlgebra→CommRing A[1/ x ]) ˣ + + + _≼_ : A → A → Type ℓ + x ≼ y = ∃[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r y -- rad(x) ⊆ rad(y) + +-- ≼ is a pre-order: + Refl≼ : isRefl _≼_ + Refl≼ x = PT.∣ 1 , 1r , ·r-comm _ _ ∣ + + Trans≼ : isTrans _≼_ + Trans≼ x y z = map2 Trans≼Σ + where + Trans≼Σ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y + → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] y ^ n ≡ a ·r z + → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r z + Trans≼Σ (n , a , p) (m , b , q) = n ·ℕ m , (a ^ m ·r b) , path + where + path : x ^ (n ·ℕ m) ≡ a ^ m ·r b ·r z + path = x ^ (n ·ℕ m) ≡⟨ ^-rdist-·ℕ x n m ⟩ + (x ^ n) ^ m ≡⟨ cong (_^ m) p ⟩ + (a ·r y) ^ m ≡⟨ ^-ldist-· a y m ⟩ + a ^ m ·r y ^ m ≡⟨ cong (a ^ m ·r_) q ⟩ + a ^ m ·r (b ·r z) ≡⟨ ·rAssoc _ _ _ ⟩ + a ^ m ·r b ·r z ∎ + + + R : A → A → Type ℓ + R x y = x ≼ y × y ≼ x -- rad(x) ≡ rad(y) + + RequivRel : isEquivRel R + RequivRel .reflexive x = Refl≼ x , Refl≼ x + RequivRel .symmetric _ _ Rxy = (Rxy .snd) , (Rxy .fst) + RequivRel .transitive _ _ _ Rxy Ryz = Trans≼ _ _ _ (Rxy .fst) (Ryz .fst) + , Trans≼ _ _ _ (Ryz .snd) (Rxy .snd) + + RpropValued : isPropValued R + RpropValued x y = isProp× isPropPropTrunc isPropPropTrunc + + powerIs≽ : (x a : A) → x ∈ [ a ⁿ|n≥0] → a ≼ x + powerIs≽ x a = map powerIs≽Σ + where + powerIs≽Σ : Σ[ n ∈ ℕ ] (x ≡ a ^ n) → Σ[ n ∈ ℕ ] Σ[ z ∈ A ] (a ^ n ≡ z ·r x) + powerIs≽Σ (n , p) = n , 1r , sym p ∙ sym (·rLid _) + + module ≼ToLoc (x y : A) where + private + instance + _ = snd A[1/ x ] + + lemma : x ≼ y → y ⋆ 1a ∈ A[1/ x ]ˣ -- y/1 ∈ A[1/x]ˣ + lemma = PT.rec (A[1/ x ]ˣ (y ⋆ 1a) .snd) lemmaΣ + where + path1 : (y z : A) → 1r ·r (y ·r 1r ·r z) ·r 1r ≡ z ·r y + path1 = solve A' + path2 : (xn : A) → xn ≡ 1r ·r 1r ·r (1r ·r 1r ·r xn) + path2 = solve A' + + lemmaΣ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y → y ⋆ 1a ∈ A[1/ x ]ˣ + lemmaΣ (n , z , p) = [ z , (x ^ n) , PT.∣ n , refl ∣ ] -- xⁿ≡zy → y⁻¹ ≡ z/xⁿ + , eq/ _ _ ((1r , powersFormMultClosedSubset _ .containsOne) + , (path1 _ _ ∙∙ sym p ∙∙ path2 _)) + + module ≼PowerToLoc (x y : A) (x≼y : x ≼ y) where + private + [yⁿ|n≥0] = [ y ⁿ|n≥0] + instance + _ = snd A[1/ x ] + lemma : ∀ (s : A) → s ∈ [yⁿ|n≥0] → s ⋆ 1a ∈ A[1/ x ]ˣ + lemma _ s∈[yⁿ|n≥0] = ≼ToLoc.lemma _ _ (Trans≼ _ y _ x≼y (powerIs≽ _ _ s∈[yⁿ|n≥0])) + + + + 𝓞ᴰ : A / R → CommAlgebra A' ℓ + 𝓞ᴰ = rec→Gpd.fun isGroupoidCommAlgebra (λ a → A[1/ a ]) RCoh LocPathProp + where + RCoh : ∀ a b → R a b → A[1/ a ] ≡ A[1/ b ] + RCoh a b (a≼b , b≼a) = fst (isContrS₁⁻¹R≡S₂⁻¹R (≼PowerToLoc.lemma _ _ b≼a) + (≼PowerToLoc.lemma _ _ a≼b)) + where + open AlgLocTwoSubsets A' [ a ⁿ|n≥0] (powersFormMultClosedSubset _) + [ b ⁿ|n≥0] (powersFormMultClosedSubset _) + + LocPathProp : ∀ a b → isProp (A[1/ a ] ≡ A[1/ b ]) + LocPathProp a b = isPropS₁⁻¹R≡S₂⁻¹R + where + open AlgLocTwoSubsets A' [ a ⁿ|n≥0] (powersFormMultClosedSubset _) + [ b ⁿ|n≥0] (powersFormMultClosedSubset _) + + + -- The quotient A/R corresponds to the basic opens of the Zariski topology. + -- Multiplication lifts to the quotient and corresponds to intersection + -- of basic opens, i.e. we get a meet-semilattice with: + _∧/_ : A / R → A / R → A / R + _∧/_ = setQuotSymmBinOp (RequivRel .reflexive) (RequivRel .transitive) _·r_ + (λ a b → subst (λ x → R (a ·r b) x) (·r-comm a b) (RequivRel .reflexive (a ·r b))) ·r-lcoh + where + ·r-lcoh-≼ : (x y z : A) → x ≼ y → (x ·r z) ≼ (y ·r z) + ·r-lcoh-≼ x y z = map ·r-lcoh-≼Σ + where + path : (x z a y zn : A) → x ·r z ·r (a ·r y ·r zn) ≡ x ·r zn ·r a ·r (y ·r z) + path = solve A' + + ·r-lcoh-≼Σ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y + → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] (x ·r z) ^ n ≡ a ·r (y ·r z) + ·r-lcoh-≼Σ (n , a , p) = suc n , (x ·r z ^ n ·r a) , (cong (x ·r z ·r_) (^-ldist-· _ _ _) + ∙∙ cong (λ v → x ·r z ·r (v ·r z ^ n)) p + ∙∙ path _ _ _ _ _) + + ·r-lcoh : (x y z : A) → R x y → R (x ·r z) (y ·r z) + ·r-lcoh x y z Rxy = ·r-lcoh-≼ x y z (Rxy .fst) , ·r-lcoh-≼ y x z (Rxy .snd) + + BasicOpens : Semilattice ℓ + BasicOpens = makeSemilattice [ 1r ] _∧/_ squash/ + (elimProp3 (λ _ _ _ → squash/ _ _) λ _ _ _ → cong [_] (·rAssoc _ _ _)) + (elimProp (λ _ → squash/ _ _) λ _ → cong [_] (·rRid _)) + (elimProp (λ _ → squash/ _ _) λ _ → cong [_] (·rLid _)) + (elimProp2 (λ _ _ → squash/ _ _) λ _ _ → cong [_] (·r-comm _ _)) + (elimProp (λ _ → squash/ _ _) λ a → eq/ _ _ -- R a a² + (∣ 1 , a , ·rRid _ ∣ , ∣ 2 , 1r , cong (a ·r_) (·rRid a) ∙ sym (·rLid _) ∣)) + + -- The induced partial order + open MeetSemilattice BasicOpens renaming (_≤_ to _≼/_ ; IndPoset to BasicOpensAsPoset) + + -- coincides with our ≼ + ≼/CoincidesWith≼ : ∀ (x y : A) → ([ x ] ≼/ [ y ]) ≡ (x ≼ y) + ≼/CoincidesWith≼ x y = [ x ] ≼/ [ y ] -- ≡⟨ refl ⟩ [ x ·r y ] ≡ [ x ] + ≡⟨ isoToPath (isEquivRel→effectiveIso RpropValued RequivRel _ _) ⟩ + R (x ·r y) x + ≡⟨ isoToPath Σ-swap-Iso ⟩ + R x (x ·r y) + ≡⟨ hPropExt (RpropValued _ _) isPropPropTrunc ·To≼ ≼To· ⟩ + x ≼ y ∎ + where + x≼xy→x≼yΣ : Σ[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r (x ·r y) + → Σ[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r y + x≼xy→x≼yΣ (n , z , p) = n , (z ·r x) , p ∙ ·rAssoc _ _ _ + + ·To≼ : R x (x ·r y) → x ≼ y + ·To≼ (x≼xy , _) = PT.map x≼xy→x≼yΣ x≼xy + + x≼y→x≼xyΣ : Σ[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r y + → Σ[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r (x ·r y) + x≼y→x≼xyΣ (n , z , p) = suc n , z , cong (x ·r_) p ∙ ·CommAssocl _ _ _ + + ≼To· : x ≼ y → R x ( x ·r y) + ≼To· x≼y = PT.map x≼y→x≼xyΣ x≼y , PT.∣ 1 , y , ·rRid _ ∙ ·r-comm _ _ ∣ + + open IsPoset + open PosetStr + Refl≼/ : isRefl _≼/_ + Refl≼/ = BasicOpensAsPoset .snd .isPoset .is-refl + + Trans≼/ : isTrans _≼/_ + Trans≼/ = BasicOpensAsPoset .snd .isPoset .is-trans + + -- The restrictions: + ρᴰᴬ : (a b : A) → a ≼ b → isContr (CommAlgebraHom A[1/ b ] A[1/ a ]) + ρᴰᴬ _ b a≼b = A[1/b]HasUniversalProp _ (≼PowerToLoc.lemma _ _ a≼b) + where + open AlgLoc A' [ b ⁿ|n≥0] (powersFormMultClosedSubset _) + renaming (S⁻¹RHasAlgUniversalProp to A[1/b]HasUniversalProp) + + ρᴰᴬId : ∀ (a : A) (r : a ≼ a) → ρᴰᴬ a a r .fst ≡ idAlgHom + ρᴰᴬId a r = ρᴰᴬ a a r .snd _ + + ρᴰᴬComp : ∀ (a b c : A) (l : a ≼ b) (m : b ≼ c) + → ρᴰᴬ a c (Trans≼ _ _ _ l m) .fst ≡ ρᴰᴬ a b l .fst ∘a ρᴰᴬ b c m .fst + ρᴰᴬComp a _ c l m = ρᴰᴬ a c (Trans≼ _ _ _ l m) .snd _ + + + ρᴰ : (x y : A / R) → x ≼/ y → CommAlgebraHom (𝓞ᴰ y) (𝓞ᴰ x) + ρᴰ = elimContr2 λ _ _ → isContrΠ + λ [a]≼/[b] → ρᴰᴬ _ _ (transport (≼/CoincidesWith≼ _ _) [a]≼/[b]) + + ρᴰId : ∀ (x : A / R) (r : x ≼/ x) → ρᴰ x x r ≡ idAlgHom + ρᴰId = SQ.elimProp (λ _ → isPropΠ (λ _ → isSetAlgebraHom _ _ _ _)) + λ a r → ρᴰᴬId a (transport (≼/CoincidesWith≼ _ _) r) + + ρᴰComp : ∀ (x y z : A / R) (l : x ≼/ y) (m : y ≼/ z) + → ρᴰ x z (Trans≼/ _ _ _ l m) ≡ ρᴰ x y l ∘a ρᴰ y z m + ρᴰComp = SQ.elimProp3 (λ _ _ _ → isPropΠ2 (λ _ _ → isSetAlgebraHom _ _ _ _)) + λ a b c _ _ → sym (ρᴰᴬ a c _ .snd _) ∙ ρᴰᴬComp a b c _ _ diff --git a/Cubical/Foundations/CartesianKanOps.agda b/Cubical/Foundations/CartesianKanOps.agda new file mode 100644 index 0000000000..6aa33b8701 --- /dev/null +++ b/Cubical/Foundations/CartesianKanOps.agda @@ -0,0 +1,149 @@ +-- This file derives some of the Cartesian Kan operations using transp +{-# OPTIONS --safe #-} +module Cubical.Foundations.CartesianKanOps where + +open import Cubical.Foundations.Prelude + +coe0→1 : ∀ {ℓ} (A : I → Type ℓ) → A i0 → A i1 +coe0→1 A a = transp (\ i → A i) i0 a + +-- "coe filler" +coe0→i : ∀ {ℓ} (A : I → Type ℓ) (i : I) → A i0 → A i +coe0→i A i a = transp (λ j → A (i ∧ j)) (~ i) a + +-- Check the equations for the coe filler +coe0→i1 : ∀ {ℓ} (A : I → Type ℓ) (a : A i0) → coe0→i A i1 a ≡ coe0→1 A a +coe0→i1 A a = refl + +coe0→i0 : ∀ {ℓ} (A : I → Type ℓ) (a : A i0) → coe0→i A i0 a ≡ a +coe0→i0 A a = refl + +-- coe backwards +coe1→0 : ∀ {ℓ} (A : I → Type ℓ) → A i1 → A i0 +coe1→0 A a = transp (λ i → A (~ i)) i0 a + +-- coe backwards filler +coe1→i : ∀ {ℓ} (A : I → Type ℓ) (i : I) → A i1 → A i +coe1→i A i a = transp (λ j → A (i ∨ ~ j)) i a + +-- Check the equations for the coe backwards filler +coe1→i0 : ∀ {ℓ} (A : I → Type ℓ) (a : A i1) → coe1→i A i0 a ≡ coe1→0 A a +coe1→i0 A a = refl + +coe1→i1 : ∀ {ℓ} (A : I → Type ℓ) (a : A i1) → coe1→i A i1 a ≡ a +coe1→i1 A a = refl + +-- "squeezeNeg" +coei→0 : ∀ {ℓ} (A : I → Type ℓ) (i : I) → A i → A i0 +coei→0 A i a = transp (λ j → A (i ∧ ~ j)) (~ i) a + +coei0→0 : ∀ {ℓ} (A : I → Type ℓ) (a : A i0) → coei→0 A i0 a ≡ a +coei0→0 A a = refl + +coei1→0 : ∀ {ℓ} (A : I → Type ℓ) (a : A i1) → coei→0 A i1 a ≡ coe1→0 A a +coei1→0 A a = refl + +-- "master coe" +-- defined as the filler of coei→0, coe0→i, and coe1→i +-- unlike in cartesian cubes, we don't get coei→i = id definitionally +coei→j : ∀ {ℓ} (A : I → Type ℓ) (i j : I) → A i → A j +coei→j A i j a = + fill (\ i → A i) + (λ j → λ { (i = i0) → coe0→i A j a + ; (i = i1) → coe1→i A j a + }) + (inS (coei→0 A i a)) + j + +-- "squeeze" +-- this is just defined as the composite face of the master coe +coei→1 : ∀ {ℓ} (A : I → Type ℓ) (i : I) → A i → A i1 +coei→1 A i a = coei→j A i i1 a + +coei0→1 : ∀ {ℓ} (A : I → Type ℓ) (a : A i0) → coei→1 A i0 a ≡ coe0→1 A a +coei0→1 A a = refl + +coei1→1 : ∀ {ℓ} (A : I → Type ℓ) (a : A i1) → coei→1 A i1 a ≡ a +coei1→1 A a = refl + +-- equations for "master coe" +coei→i0 : ∀ {ℓ} (A : I → Type ℓ) (i : I) (a : A i) → coei→j A i i0 a ≡ coei→0 A i a +coei→i0 A i a = refl + +coei0→i : ∀ {ℓ} (A : I → Type ℓ) (i : I) (a : A i0) → coei→j A i0 i a ≡ coe0→i A i a +coei0→i A i a = refl + +coei→i1 : ∀ {ℓ} (A : I → Type ℓ) (i : I) (a : A i) → coei→j A i i1 a ≡ coei→1 A i a +coei→i1 A i a = refl + +coei1→i : ∀ {ℓ} (A : I → Type ℓ) (i : I) (a : A i1) → coei→j A i1 i a ≡ coe1→i A i a +coei1→i A i a = refl + +-- only non-definitional equation +coei→i : ∀ {ℓ} (A : I → Type ℓ) (i : I) (a : A i) → coei→j A i i a ≡ a +coei→i A i = coe0→i (λ i → (a : A i) → coei→j A i i a ≡ a) i (λ _ → refl) + +-- do the same for fill + +fill1→i : ∀ {ℓ} (A : ∀ i → Type ℓ) + {φ : I} + (u : ∀ i → Partial φ (A i)) + (u1 : A i1 [ φ ↦ u i1 ]) + --------------------------- + (i : I) → A i +fill1→i A {φ = φ} u u1 i = + comp (λ j → A (i ∨ ~ j)) + (λ j → λ { (φ = i1) → u (i ∨ ~ j) 1=1 + ; (i = i1) → outS u1 }) + (outS u1) + +filli→0 : ∀ {ℓ} (A : ∀ i → Type ℓ) + {φ : I} + (u : ∀ i → Partial φ (A i)) + (i : I) + (ui : A i [ φ ↦ u i ]) + --------------------------- + → A i0 +filli→0 A {φ = φ} u i ui = + comp (λ j → A (i ∧ ~ j)) + (λ j → λ { (φ = i1) → u (i ∧ ~ j) 1=1 + ; (i = i0) → outS ui }) + (outS ui) + +filli→j : ∀ {ℓ} (A : ∀ i → Type ℓ) + {φ : I} + (u : ∀ i → Partial φ (A i)) + (i : I) + (ui : A i [ φ ↦ u i ]) + --------------------------- + (j : I) → A j +filli→j A {φ = φ} u i ui j = + fill (\ i → A i) + (λ j → λ { (φ = i1) → u j 1=1 + ; (i = i0) → fill (\ i → A i) (\ i → u i) ui j + ; (i = i1) → fill1→i A u ui j + }) + (inS (filli→0 A u i ui)) + j + +-- We can reconstruct fill from hfill, coei→j, and the path coei→i ≡ id. +-- The definition does not rely on the computational content of the coei→i path. +fill' : ∀ {ℓ} (A : I → Type ℓ) + {φ : I} + (u : ∀ i → Partial φ (A i)) + (u0 : A i0 [ φ ↦ u i0 ]) + --------------------------- + (i : I) → A i [ φ ↦ u i ] +fill' A {φ = φ} u u0 i = + inS (hcomp (λ j → λ {(φ = i1) → coei→i A i (u i 1=1) j; (i = i0) → coei→i A i (outS u0) j}) t) + where + t : A i + t = hfill {φ = φ} (λ j v → coei→j A j i (u j v)) (inS (coe0→i A i (outS u0))) i + +fill'-cap : ∀ {ℓ} (A : I → Type ℓ) + {φ : I} + (u : ∀ i → Partial φ (A i)) + (u0 : A i0 [ φ ↦ u i0 ]) + --------------------------- + → outS (fill' A u u0 i0) ≡ outS (u0) +fill'-cap A u u0 = refl diff --git a/Cubical/Foundations/Equiv.agda b/Cubical/Foundations/Equiv.agda new file mode 100644 index 0000000000..8b35a27899 --- /dev/null +++ b/Cubical/Foundations/Equiv.agda @@ -0,0 +1,307 @@ +{- + +Theory about equivalences + +Definitions are in Core/Glue.agda but re-exported by this module + +- isEquiv is a proposition ([isPropIsEquiv]) +- Any isomorphism is an equivalence ([isoToEquiv]) + +There are more statements about equivalences in Equiv/Properties.agda: + +- if f is an equivalence then (cong f) is an equivalence +- if f is an equivalence then precomposition with f is an equivalence +- if f is an equivalence then postcomposition with f is an equivalence + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv where + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Foundations.Equiv.Base public +open import Cubical.Data.Sigma.Base + +private + variable + ℓ ℓ' ℓ'' : Level + A B C D : Type ℓ + +equivIsEquiv : (e : A ≃ B) → isEquiv (equivFun e) +equivIsEquiv e = snd e + +equivCtr : (e : A ≃ B) (y : B) → fiber (equivFun e) y +equivCtr e y = e .snd .equiv-proof y .fst + +equivCtrPath : (e : A ≃ B) (y : B) → + (v : fiber (equivFun e) y) → Path _ (equivCtr e y) v +equivCtrPath e y = e .snd .equiv-proof y .snd + + +-- Proof using isPropIsContr. This is slow and the direct proof below is better + +isPropIsEquiv' : (f : A → B) → isProp (isEquiv f) +equiv-proof (isPropIsEquiv' f u0 u1 i) y = + isPropIsContr (u0 .equiv-proof y) (u1 .equiv-proof y) i + +-- Direct proof that computes quite ok (can be optimized further if +-- necessary, see: +-- https://github.com/mortberg/cubicaltt/blob/pi4s3_dimclosures/examples/brunerie2.ctt#L562 + +isPropIsEquiv : (f : A → B) → isProp (isEquiv f) +equiv-proof (isPropIsEquiv f p q i) y = + let p2 = p .equiv-proof y .snd + q2 = q .equiv-proof y .snd + in p2 (q .equiv-proof y .fst) i + , λ w j → hcomp (λ k → λ { (i = i0) → p2 w j + ; (i = i1) → q2 w (j ∨ ~ k) + ; (j = i0) → p2 (q2 w (~ k)) i + ; (j = i1) → w }) + (p2 w (i ∨ j)) + +equivEq : {e f : A ≃ B} → (h : e .fst ≡ f .fst) → e ≡ f +equivEq {e = e} {f = f} h = λ i → (h i) , isProp→PathP (λ i → isPropIsEquiv (h i)) (e .snd) (f .snd) i + +module _ {f : A → B} (equivF : isEquiv f) where + funIsEq : A → B + funIsEq = f + + invIsEq : B → A + invIsEq y = equivF .equiv-proof y .fst .fst + + secIsEq : section f invIsEq + secIsEq y = equivF .equiv-proof y .fst .snd + + retIsEq : retract f invIsEq + retIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .fst + + commSqIsEq : ∀ a → Square (secIsEq (f a)) refl (cong f (retIsEq a)) refl + commSqIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .snd + + commPathIsEq : ∀ a → secIsEq (f a) ≡ cong f (retIsEq a) + commPathIsEq a i j = + hcomp + (λ k → λ + { (i = i0) → secIsEq (f a) j + ; (i = i1) → f (retIsEq a (j ∨ ~ k)) + ; (j = i0) → f (retIsEq a (i ∧ ~ k)) + ; (j = i1) → f a + }) + (commSqIsEq a i j) + +module _ (w : A ≃ B) where + invEq : B → A + invEq = invIsEq (snd w) + + retEq : retract (w .fst) invEq + retEq = retIsEq (snd w) + + secEq : section (w .fst) invEq + secEq = secIsEq (snd w) + +open Iso + +equivToIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ≃ B → Iso A B +fun (equivToIso e) = e .fst +inv (equivToIso e) = invEq e +rightInv (equivToIso e) = secEq e +leftInv (equivToIso e) = retEq e + +-- TODO: there should be a direct proof of this that doesn't use equivToIso +invEquiv : A ≃ B → B ≃ A +invEquiv e = isoToEquiv (invIso (equivToIso e)) + +invEquivIdEquiv : (A : Type ℓ) → invEquiv (idEquiv A) ≡ idEquiv A +invEquivIdEquiv _ = equivEq refl + +compEquiv : A ≃ B → B ≃ C → A ≃ C +compEquiv f g .fst = g .fst ∘ f .fst +compEquiv {A = A} {C = C} f g .snd .equiv-proof c = contr + where + contractG = g .snd .equiv-proof c .snd + + secFiller : (a : A) (p : g .fst (f .fst a) ≡ c) → _ {- square in A -} + secFiller a p = compPath-filler (cong (invEq f ∘ fst) (contractG (_ , p))) (retEq f a) + + contr : isContr (fiber (g .fst ∘ f .fst) c) + contr .fst .fst = invEq f (invEq g c) + contr .fst .snd = cong (g .fst) (secEq f (invEq g c)) ∙ secEq g c + contr .snd (a , p) i .fst = secFiller a p i1 i + contr .snd (a , p) i .snd j = + hcomp + (λ k → λ + { (i = i1) → fSquare k + ; (j = i0) → g .fst (f .fst (secFiller a p k i)) + ; (j = i1) → contractG (_ , p) i .snd k + }) + (g .fst (secEq f (contractG (_ , p) i .fst) j)) + where + fSquare : I → C + fSquare k = + hcomp + (λ l → λ + { (j = i0) → g .fst (f .fst (retEq f a k)) + ; (j = i1) → p (k ∧ l) + ; (k = i0) → g .fst (secEq f (f .fst a) j) + ; (k = i1) → p (j ∧ l) + }) + (g .fst (f .snd .equiv-proof (f .fst a) .snd (a , refl) k .snd j)) + +compEquivIdEquiv : (e : A ≃ B) → compEquiv (idEquiv A) e ≡ e +compEquivIdEquiv e = equivEq refl + +compEquivEquivId : (e : A ≃ B) → compEquiv e (idEquiv B) ≡ e +compEquivEquivId e = equivEq refl + +invEquiv-is-rinv : (e : A ≃ B) → compEquiv e (invEquiv e) ≡ idEquiv A +invEquiv-is-rinv e = equivEq (funExt (retEq e)) + +invEquiv-is-linv : (e : A ≃ B) → compEquiv (invEquiv e) e ≡ idEquiv B +invEquiv-is-linv e = equivEq (funExt (secEq e)) + +compEquiv-assoc : (f : A ≃ B) (g : B ≃ C) (h : C ≃ D) + → compEquiv f (compEquiv g h) ≡ compEquiv (compEquiv f g) h +compEquiv-assoc f g h = equivEq refl + +LiftEquiv : A ≃ Lift {i = ℓ} {j = ℓ'} A +LiftEquiv .fst a .lower = a +LiftEquiv .snd .equiv-proof = strictContrFibers lower + +Lift≃Lift : (e : A ≃ B) → Lift {j = ℓ'} A ≃ Lift {j = ℓ''} B +Lift≃Lift e .fst a .lower = e .fst (a .lower) +Lift≃Lift e .snd .equiv-proof b .fst .fst .lower = invEq e (b .lower) +Lift≃Lift e .snd .equiv-proof b .fst .snd i .lower = + e .snd .equiv-proof (b .lower) .fst .snd i +Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .fst .lower = + e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .fst +Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .snd j .lower = + e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .snd j + +isContr→Equiv : isContr A → isContr B → A ≃ B +isContr→Equiv Actr Bctr = isoToEquiv (isContr→Iso Actr Bctr) + +propBiimpl→Equiv : (Aprop : isProp A) (Bprop : isProp B) (f : A → B) (g : B → A) → A ≃ B +propBiimpl→Equiv Aprop Bprop f g = f , hf + where + hf : isEquiv f + hf .equiv-proof y .fst = (g y , Bprop (f (g y)) y) + hf .equiv-proof y .snd h i .fst = Aprop (g y) (h .fst) i + hf .equiv-proof y .snd h i .snd = isProp→isSet' Bprop (Bprop (f (g y)) y) (h .snd) + (cong f (Aprop (g y) (h .fst))) refl i + +isEquivPropBiimpl→Equiv : isProp A → isProp B + → ((A → B) × (B → A)) ≃ (A ≃ B) +isEquivPropBiimpl→Equiv {A = A} {B = B} Aprop Bprop = isoToEquiv isom where + isom : Iso (Σ (A → B) (λ _ → B → A)) (A ≃ B) + isom .fun (f , g) = propBiimpl→Equiv Aprop Bprop f g + isom .inv e = equivFun e , invEq e + isom .rightInv e = equivEq refl + isom .leftInv _ = refl + +equivΠCod : ∀ {F : A → Type ℓ} {G : A → Type ℓ'} + → ((x : A) → F x ≃ G x) → ((x : A) → F x) ≃ ((x : A) → G x) +equivΠCod k .fst f x = k x .fst (f x) +equivΠCod k .snd .equiv-proof f .fst .fst x = equivCtr (k x) (f x) .fst +equivΠCod k .snd .equiv-proof f .fst .snd i x = equivCtr (k x) (f x) .snd i +equivΠCod k .snd .equiv-proof f .snd (g , p) i .fst x = + equivCtrPath (k x) (f x) (g x , λ j → p j x) i .fst +equivΠCod k .snd .equiv-proof f .snd (g , p) i .snd j x = + equivCtrPath (k x) (f x) (g x , λ k → p k x) i .snd j + +equivImplicitΠCod : ∀ {F : A → Type ℓ} {G : A → Type ℓ'} + → ({x : A} → F x ≃ G x) → ({x : A} → F x) ≃ ({x : A} → G x) +equivImplicitΠCod k .fst f {x} = k {x} .fst (f {x}) +equivImplicitΠCod k .snd .equiv-proof f .fst .fst {x} = equivCtr (k {x}) (f {x}) .fst +equivImplicitΠCod k .snd .equiv-proof f .fst .snd i {x} = equivCtr (k {x}) (f {x}) .snd i +equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .fst {x} = + equivCtrPath (k {x}) (f {x}) (g {x} , λ j → p j {x}) i .fst +equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .snd j {x} = + equivCtrPath (k {x}) (f {x}) (g {x} , λ k → p k {x}) i .snd j + +equiv→Iso : (A ≃ B) → (C ≃ D) → Iso (A → C) (B → D) +equiv→Iso h k .Iso.fun f b = equivFun k (f (invEq h b)) +equiv→Iso h k .Iso.inv g a = invEq k (g (equivFun h a)) +equiv→Iso h k .Iso.rightInv g = funExt λ b → secEq k _ ∙ cong g (secEq h b) +equiv→Iso h k .Iso.leftInv f = funExt λ a → retEq k _ ∙ cong f (retEq h a) + +equiv→ : (A ≃ B) → (C ≃ D) → (A → C) ≃ (B → D) +equiv→ h k = isoToEquiv (equiv→Iso h k) + + +equivΠ' : ∀ {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + {B : A → Type ℓB} {B' : A' → Type ℓB'} + (eA : A ≃ A') + (eB : {a : A} {a' : A'} → eA .fst a ≡ a' → B a ≃ B' a') + → ((a : A) → B a) ≃ ((a' : A') → B' a') +equivΠ' {B' = B'} eA eB = isoToEquiv isom + where + open Iso + + isom : Iso _ _ + isom .fun f a' = + eB (secEq eA a') .fst (f (invEq eA a')) + isom .inv f' a = + invEq (eB refl) (f' (eA .fst a)) + isom .rightInv f' = + funExt λ a' → + J (λ a'' p → eB p .fst (invEq (eB refl) (f' (p i0))) ≡ f' a'') + (secEq (eB refl) (f' (eA .fst (invEq eA a')))) + (secEq eA a') + isom .leftInv f = + funExt λ a → + subst + (λ p → invEq (eB refl) (eB p .fst (f (invEq eA (eA .fst a)))) ≡ f a) + (sym (commPathIsEq (eA .snd) a)) + (J (λ a'' p → invEq (eB refl) (eB (cong (eA .fst) p) .fst (f (invEq eA (eA .fst a)))) ≡ f a'') + (retEq (eB refl) (f (invEq eA (eA .fst a)))) + (retEq eA a)) + +equivΠ : ∀ {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + {B : A → Type ℓB} {B' : A' → Type ℓB'} + (eA : A ≃ A') + (eB : (a : A) → B a ≃ B' (eA .fst a)) + → ((a : A) → B a) ≃ ((a' : A') → B' a') +equivΠ {B = B} {B' = B'} eA eB = equivΠ' eA (λ {a = a} p → J (λ a' p → B a ≃ B' a') (eB a) p) + + +equivCompIso : (A ≃ B) → (C ≃ D) → Iso (A ≃ C) (B ≃ D) +equivCompIso h k .Iso.fun f = compEquiv (compEquiv (invEquiv h) f) k +equivCompIso h k .Iso.inv g = compEquiv (compEquiv h g) (invEquiv k) +equivCompIso h k .Iso.rightInv g = equivEq (equiv→Iso h k .Iso.rightInv (equivFun g)) +equivCompIso h k .Iso.leftInv f = equivEq (equiv→Iso h k .Iso.leftInv (equivFun f)) + +equivComp : (A ≃ B) → (C ≃ D) → (A ≃ C) ≃ (B ≃ D) +equivComp h k = isoToEquiv (equivCompIso h k) + +-- Some helpful notation: +_≃⟨_⟩_ : (X : Type ℓ) → (X ≃ B) → (B ≃ C) → (X ≃ C) +_ ≃⟨ f ⟩ g = compEquiv f g + +_■ : (X : Type ℓ) → (X ≃ X) +_■ = idEquiv + +infixr 0 _≃⟨_⟩_ +infix 1 _■ + +composesToId→Equiv : (f : A → B) (g : B → A) → f ∘ g ≡ idfun B → isEquiv f → isEquiv g +composesToId→Equiv f g id iseqf = + isoToIsEquiv + (iso g f + (λ b → (λ i → equiv-proof iseqf (f b) .snd (g (f b) , cong (λ h → h (f b)) id) (~ i) .fst) + ∙∙ cong (λ x → equiv-proof iseqf (f b) .fst .fst) id + ∙∙ λ i → equiv-proof iseqf (f b) .snd (b , refl) i .fst) + λ a i → id i a) + +-- equivalence between isEquiv and isEquiv' + +isEquiv-isEquiv'-Iso : (f : A → B) → Iso (isEquiv f) (isEquiv' f) +isEquiv-isEquiv'-Iso f .fun p = p .equiv-proof +isEquiv-isEquiv'-Iso f .inv q .equiv-proof = q +isEquiv-isEquiv'-Iso f .rightInv q = refl +isEquiv-isEquiv'-Iso f .leftInv p i .equiv-proof = p .equiv-proof + +isEquiv≃isEquiv' : (f : A → B) → isEquiv f ≃ isEquiv' f +isEquiv≃isEquiv' f = isoToEquiv (isEquiv-isEquiv'-Iso f) diff --git a/Cubical/Foundations/Equiv/Base.agda b/Cubical/Foundations/Equiv/Base.agda new file mode 100644 index 0000000000..351f2b6b1e --- /dev/null +++ b/Cubical/Foundations/Equiv/Base.agda @@ -0,0 +1,31 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.Base where + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude + +open import Cubical.Core.Glue public + using ( isEquiv ; equiv-proof ; _≃_ ; equivFun ; equivProof ) + +fiber : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (y : B) → Type (ℓ-max ℓ ℓ') +fiber {A = A} f y = Σ[ x ∈ A ] f x ≡ y + +-- Helper function for constructing equivalences from pairs (f,g) that cancel each other up to definitional +-- equality. For such (f,g), the result type simplifies to isContr (fiber f b). +strictContrFibers : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {f : A → B} (g : B → A) (b : B) + → Σ[ t ∈ fiber f (f (g b)) ] + ((t' : fiber f b) → Path (fiber f (f (g b))) t (g (f (t' .fst)) , cong (f ∘ g) (t' .snd))) +strictContrFibers {f = f} g b .fst = (g b , refl) +strictContrFibers {f = f} g b .snd (a , p) i = (g (p (~ i)) , λ j → f (g (p (~ i ∨ j)))) + +-- The identity equivalence +idIsEquiv : ∀ {ℓ} (A : Type ℓ) → isEquiv (idfun A) +idIsEquiv A .equiv-proof = strictContrFibers (idfun A) + +idEquiv : ∀ {ℓ} (A : Type ℓ) → A ≃ A +idEquiv A .fst = idfun A +idEquiv A .snd = idIsEquiv A + +-- the definition using Π-type +isEquiv' : ∀ {ℓ}{ℓ'}{A : Type ℓ}{B : Type ℓ'} → (A → B) → Type (ℓ-max ℓ ℓ') +isEquiv' {B = B} f = (y : B) → isContr (fiber f y) diff --git a/Cubical/Foundations/Equiv/BiInvertible.agda b/Cubical/Foundations/Equiv/BiInvertible.agda new file mode 100644 index 0000000000..383852814e --- /dev/null +++ b/Cubical/Foundations/Equiv/BiInvertible.agda @@ -0,0 +1,82 @@ +{- + +Some theory about Bi-Invertible Equivalences + +- BiInvEquiv to Iso +- BiInvEquiv to Equiv +- BiInvEquiv to HAEquiv +- Iso to BiInvEquiv + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.BiInvertible where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint + + +record BiInvEquiv {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') where + constructor biInvEquiv + field + fun : A → B + invr : B → A + invr-rightInv : section fun invr + invl : B → A + invl-leftInv : retract fun invl + + invr≡invl : ∀ b → invr b ≡ invl b + invr≡invl b = invr b ≡⟨ sym (invl-leftInv (invr b)) ⟩ + invl (fun (invr b)) ≡⟨ cong invl (invr-rightInv b) ⟩ + invl b ∎ + + invr-leftInv : retract fun invr + invr-leftInv a = invr≡invl (fun a) ∙ (invl-leftInv a) + + invr≡invl-leftInv : ∀ a → PathP (λ j → invr≡invl (fun a) j ≡ a) (invr-leftInv a) (invl-leftInv a) + invr≡invl-leftInv a j i = compPath-filler' (invr≡invl (fun a)) (invl-leftInv a) (~ j) i + + invl-rightInv : section fun invl + invl-rightInv a = sym (cong fun (invr≡invl a)) ∙ (invr-rightInv a) + + invr≡invl-rightInv : ∀ a → PathP (λ j → fun (invr≡invl a j) ≡ a) (invr-rightInv a) (invl-rightInv a) + invr≡invl-rightInv a j i = compPath-filler' (sym (cong fun (invr≡invl a))) (invr-rightInv a) j i + + +module _ {ℓ} {A B : Type ℓ} (e : BiInvEquiv A B) where + open BiInvEquiv e + + biInvEquiv→Iso-right : Iso A B + Iso.fun biInvEquiv→Iso-right = fun + Iso.inv biInvEquiv→Iso-right = invr + Iso.rightInv biInvEquiv→Iso-right = invr-rightInv + Iso.leftInv biInvEquiv→Iso-right = invr-leftInv + + biInvEquiv→Iso-left : Iso A B + Iso.fun biInvEquiv→Iso-left = fun + Iso.inv biInvEquiv→Iso-left = invl + Iso.rightInv biInvEquiv→Iso-left = invl-rightInv + Iso.leftInv biInvEquiv→Iso-left = invl-leftInv + + biInvEquiv→Equiv-right biInvEquiv→Equiv-left : A ≃ B + biInvEquiv→Equiv-right = fun , isoToIsEquiv biInvEquiv→Iso-right + biInvEquiv→Equiv-left = fun , isoToIsEquiv biInvEquiv→Iso-left + + -- since Iso.rightInv ends up getting modified during iso→HAEquiv, in some sense biInvEquiv→Iso-left + -- is the most natural choice for forming a HAEquiv from a BiInvEquiv + biInvEquiv→HAEquiv : HAEquiv A B + biInvEquiv→HAEquiv = iso→HAEquiv biInvEquiv→Iso-left + + +module _ {ℓ} {A B : Type ℓ} (i : Iso A B) where + open Iso i + + iso→BiInvEquiv : BiInvEquiv A B + BiInvEquiv.fun iso→BiInvEquiv = fun + BiInvEquiv.invr iso→BiInvEquiv = inv + BiInvEquiv.invr-rightInv iso→BiInvEquiv = rightInv + BiInvEquiv.invl iso→BiInvEquiv = inv + BiInvEquiv.invl-leftInv iso→BiInvEquiv = leftInv + diff --git a/Cubical/Foundations/Equiv/Fiberwise.agda b/Cubical/Foundations/Equiv/Fiberwise.agda new file mode 100644 index 0000000000..928853c6f2 --- /dev/null +++ b/Cubical/Foundations/Equiv/Fiberwise.agda @@ -0,0 +1,85 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.Fiberwise where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ' ℓ'' : Level + +module _ {A : Type ℓ} (P : A → Type ℓ') (Q : A → Type ℓ'') + (f : ∀ x → P x → Q x) + where + private + total : (Σ A P) → (Σ A Q) + total = (\ p → p .fst , f (p .fst) (p .snd)) + + -- Thm 4.7.6 + fibers-total : ∀ {xv} → Iso (fiber total (xv)) (fiber (f (xv .fst)) (xv .snd)) + fibers-total {xv} = iso h g h-g g-h + where + h : ∀ {xv} → fiber total xv → fiber (f (xv .fst)) (xv .snd) + h {xv} (p , eq) = J (\ xv eq → fiber (f (xv .fst)) (xv .snd)) ((p .snd) , refl) eq + g : ∀ {xv} → fiber (f (xv .fst)) (xv .snd) → fiber total xv + g {xv} (p , eq) = (xv .fst , p) , (\ i → _ , eq i) + h-g : ∀ {xv} y → h {xv} (g {xv} y) ≡ y + h-g {x , v} (p , eq) = J (λ _ eq₁ → h (g (p , eq₁)) ≡ (p , eq₁)) (JRefl (λ xv₁ eq₁ → fiber (f (xv₁ .fst)) (xv₁ .snd)) ((p , refl))) (eq) + g-h : ∀ {xv} y → g {xv} (h {xv} y) ≡ y + g-h {xv} ((a , p) , eq) = J (λ _ eq₁ → g (h ((a , p) , eq₁)) ≡ ((a , p) , eq₁)) + (cong g (JRefl (λ xv₁ eq₁ → fiber (f (xv₁ .fst)) (xv₁ .snd)) (p , refl))) + eq + -- Thm 4.7.7 (fiberwise equivalences) + fiberEquiv : ([tf] : isEquiv total) + → ∀ x → isEquiv (f x) + fiberEquiv [tf] x .equiv-proof y = isContrRetract (fibers-total .Iso.inv) (fibers-total .Iso.fun) (fibers-total .Iso.rightInv) + ([tf] .equiv-proof (x , y)) + + totalEquiv : (fx-equiv : ∀ x → isEquiv (f x)) + → isEquiv total + totalEquiv fx-equiv .equiv-proof (x , v) = isContrRetract (fibers-total .Iso.fun) (fibers-total .Iso.inv) (fibers-total .Iso.leftInv) + (fx-equiv x .equiv-proof v) + + +module _ {U : Type ℓ} (_~_ : U → U → Type ℓ') + (idTo~ : ∀ {A B} → A ≡ B → A ~ B) + (c : ∀ A → ∃![ X ∈ U ] (A ~ X)) + where + + isContrToUniv : ∀ {A B} → isEquiv (idTo~ {A} {B}) + isContrToUniv {A} {B} + = fiberEquiv (λ z → A ≡ z) (λ z → A ~ z) (\ B → idTo~ {A} {B}) + (λ { .equiv-proof y + → isContrΣ (isContrSingl _) + \ a → isContr→isContrPath (c A) _ _ + }) + B + + +{- + The following is called fundamental theorem of identity types in Egbert Rijke's + introduction to homotopy type theory. +-} +fundamentalTheoremOfId : {A : Type ℓ} (Eq : A → A → Type ℓ') + → ((x : A) → Eq x x) + → ((x : A) → isContr (Σ[ y ∈ A ] Eq x y)) + → (x y : A) → (x ≡ y) ≃ (Eq x y) +fundamentalTheoremOfId {A = A} Eq eqRefl eqContr x y = (fiberMap x y) , (isEquivFiberMap x y) + where + fiberMap : (x y : A) → x ≡ y → Eq x y + fiberMap x y = J (λ y p → Eq x y) (eqRefl x) + + mapOnSigma : (x : A) → Σ[ y ∈ A ] x ≡ y → Σ[ y ∈ A ] Eq x y + mapOnSigma x pair = fst pair , fiberMap x (fst pair) (snd pair) + + equivOnSigma : (x : A) → isEquiv (mapOnSigma x) + equivOnSigma x = isEquivFromIsContr (mapOnSigma x) (isContrSingl x) (eqContr x) + + isEquivFiberMap : (x y : A) → isEquiv (fiberMap x y) + isEquivFiberMap x = fiberEquiv (λ y → x ≡ y) (Eq x) (fiberMap x) (equivOnSigma x) diff --git a/Cubical/Foundations/Equiv/HalfAdjoint.agda b/Cubical/Foundations/Equiv/HalfAdjoint.agda new file mode 100644 index 0000000000..e589d5d068 --- /dev/null +++ b/Cubical/Foundations/Equiv/HalfAdjoint.agda @@ -0,0 +1,166 @@ +{- + +Half adjoint equivalences ([HAEquiv]) + +- Iso to HAEquiv ([iso→HAEquiv]) +- Equiv to HAEquiv ([equiv→HAEquiv]) +- Cong is an equivalence ([congEquiv]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.HalfAdjoint where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + +record isHAEquiv {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) : Type (ℓ-max ℓ ℓ') where + field + g : B → A + linv : ∀ a → g (f a) ≡ a + rinv : ∀ b → f (g b) ≡ b + com : ∀ a → cong f (linv a) ≡ rinv (f a) + + -- from redtt's ha-equiv/symm + com-op : ∀ b → cong g (rinv b) ≡ linv (g b) + com-op b j i = hcomp (λ k → λ { (i = i0) → linv (g b) (j ∧ (~ k)) + ; (j = i0) → g (rinv b i) + ; (j = i1) → linv (g b) (i ∨ (~ k)) + ; (i = i1) → g b }) + (cap1 j i) + + where cap0 : Square {- (j = i0) -} (λ i → f (g (rinv b i))) + {- (j = i1) -} (λ i → rinv b i) + {- (i = i0) -} (λ j → f (linv (g b) j)) + {- (i = i1) -} (λ j → rinv b j) + + cap0 j i = hcomp (λ k → λ { (i = i0) → com (g b) (~ k) j + ; (j = i0) → f (g (rinv b i)) + ; (j = i1) → rinv b i + ; (i = i1) → rinv b j }) + (rinv (rinv b i) j) + + filler : I → I → A + filler j i = hfill (λ k → λ { (i = i0) → g (rinv b k) + ; (i = i1) → g b }) + (inS (linv (g b) i)) j + + cap1 : Square {- (j = i0) -} (λ i → g (rinv b i)) + {- (j = i1) -} (λ i → g b) + {- (i = i0) -} (λ j → linv (g b) j) + {- (i = i1) -} (λ j → g b) + + cap1 j i = hcomp (λ k → λ { (i = i0) → linv (linv (g b) j) k + ; (j = i0) → linv (g (rinv b i)) k + ; (j = i1) → filler i k + ; (i = i1) → filler j k }) + (g (cap0 j i)) + + isHAEquiv→Iso : Iso A B + Iso.fun isHAEquiv→Iso = f + Iso.inv isHAEquiv→Iso = g + Iso.rightInv isHAEquiv→Iso = rinv + Iso.leftInv isHAEquiv→Iso = linv + + isHAEquiv→isEquiv : isEquiv f + isHAEquiv→isEquiv .equiv-proof y = (g y , rinv y) , isCenter where + isCenter : ∀ xp → (g y , rinv y) ≡ xp + isCenter (x , p) i = gy≡x i , ry≡p i where + gy≡x : g y ≡ x + gy≡x = sym (cong g p) ∙∙ refl ∙∙ linv x + + lem0 : Square (cong f (linv x)) p (cong f (linv x)) p + lem0 i j = invSides-filler p (sym (cong f (linv x))) (~ i) j + + ry≡p : Square (rinv y) p (cong f gy≡x) refl + ry≡p i j = hcomp (λ k → λ { (i = i0) → cong rinv p k j + ; (i = i1) → lem0 k j + ; (j = i0) → f (doubleCompPath-filler (sym (cong g p)) refl (linv x) k i) + ; (j = i1) → p k }) + (com x (~ i) j) + +open isHAEquiv using (isHAEquiv→Iso; isHAEquiv→isEquiv) public + +HAEquiv : (A : Type ℓ) (B : Type ℓ') → Type (ℓ-max ℓ ℓ') +HAEquiv A B = Σ[ f ∈ (A → B) ] isHAEquiv f + +-- vogt's lemma (https://ncatlab.org/nlab/show/homotopy+equivalence#vogts_lemma) +iso→HAEquiv : Iso A B → HAEquiv A B +iso→HAEquiv e = f , isHAEquivf + where + f = Iso.fun e + g = Iso.inv e + ε = Iso.rightInv e + η = Iso.leftInv e + + Hfa≡fHa : (f : A → A) → (H : ∀ a → f a ≡ a) → ∀ a → H (f a) ≡ cong f (H a) + Hfa≡fHa f H = J (λ f p → ∀ a → funExt⁻ (sym p) (f a) ≡ cong f (funExt⁻ (sym p) a)) + (λ a → refl) + (sym (funExt H)) + + isHAEquivf : isHAEquiv f + isHAEquiv.g isHAEquivf = g + isHAEquiv.linv isHAEquivf = η + isHAEquiv.rinv isHAEquivf b i = + hcomp (λ j → λ { (i = i0) → ε (f (g b)) j + ; (i = i1) → ε b j }) + (f (η (g b) i)) + isHAEquiv.com isHAEquivf a i j = + hcomp (λ k → λ { (i = i0) → ε (f (η a j)) k + ; (j = i0) → ε (f (g (f a))) k + ; (j = i1) → ε (f a) k}) + (f (Hfa≡fHa (λ x → g (f x)) η a (~ i) j)) + +equiv→HAEquiv : A ≃ B → HAEquiv A B +equiv→HAEquiv e = e .fst , λ where + .isHAEquiv.g → invIsEq (snd e) + .isHAEquiv.linv → retIsEq (snd e) + .isHAEquiv.rinv → secIsEq (snd e) + .isHAEquiv.com a → sym (commPathIsEq (snd e) a) + +congIso : {x y : A} (e : Iso A B) → Iso (x ≡ y) (Iso.fun e x ≡ Iso.fun e y) +congIso {x = x} {y} e = goal + where + open isHAEquiv (iso→HAEquiv e .snd) + open Iso + + goal : Iso (x ≡ y) (Iso.fun e x ≡ Iso.fun e y) + fun goal = cong (iso→HAEquiv e .fst) + inv goal p = sym (linv x) ∙∙ cong g p ∙∙ linv y + rightInv goal p i j = + hcomp (λ k → λ { (i = i0) → iso→HAEquiv e .fst + (doubleCompPath-filler (sym (linv x)) (cong g p) (linv y) k j) + ; (i = i1) → rinv (p j) k + ; (j = i0) → com x i k + ; (j = i1) → com y i k }) + (iso→HAEquiv e .fst (g (p j))) + leftInv goal p i j = + hcomp (λ k → λ { (i = i1) → p j + ; (j = i0) → Iso.leftInv e x (i ∨ k) + ; (j = i1) → Iso.leftInv e y (i ∨ k) }) + (Iso.leftInv e (p j) i) + +invCongFunct : {x : A} (e : Iso A B) (p : Iso.fun e x ≡ Iso.fun e x) (q : Iso.fun e x ≡ Iso.fun e x) + → Iso.inv (congIso e) (p ∙ q) ≡ Iso.inv (congIso e) p ∙ Iso.inv (congIso e) q +invCongFunct {x = x} e p q = helper (Iso.inv e) _ _ _ + where + helper : {x : A} {y : B} (f : A → B) (r : f x ≡ y) (p q : x ≡ x) + → (sym r ∙∙ cong f (p ∙ q) ∙∙ r) ≡ (sym r ∙∙ cong f p ∙∙ r) ∙ (sym r ∙∙ cong f q ∙∙ r) + helper {x = x} f = + J (λ y r → (p q : x ≡ x) + → (sym r ∙∙ cong f (p ∙ q) ∙∙ r) ≡ (sym r ∙∙ cong f p ∙∙ r) ∙ (sym r ∙∙ cong f q ∙∙ r)) + λ p q → (λ i → rUnit (congFunct f p q i) (~ i)) + ∙ λ i → rUnit (cong f p) i ∙ rUnit (cong f q) i + +invCongRefl : {x : A} (e : Iso A B) → Iso.inv (congIso {x = x} {y = x} e) refl ≡ refl +invCongRefl {x = x} e = (λ i → (λ j → Iso.leftInv e x (i ∨ ~ j)) ∙∙ refl ∙∙ (λ j → Iso.leftInv e x (i ∨ j))) ∙ sym (rUnit refl) diff --git a/Cubical/Foundations/Equiv/PathSplit.agda b/Cubical/Foundations/Equiv/PathSplit.agda new file mode 100644 index 0000000000..8e1f4ca60a --- /dev/null +++ b/Cubical/Foundations/Equiv/PathSplit.agda @@ -0,0 +1,140 @@ +{- + +Theory about path split equivalences. +They are convenient to construct localization HITs as in +(the "modalities paper") +https://arxiv.org/abs/1706.07526 + +- there are construction from and to equivalences ([pathSplitToEquiv] , [equivToPathSplit]) +- the structure of a path split equivalence is actually a proposition ([isPropIsPathSplitEquiv]) + +The module starts with a couple of general facts about equivalences: + +- if f is an equivalence then (cong f) is an equivalence ([equivCong]) +- if f is an equivalence then pre- and postcomposition with f are equivalences ([preCompEquiv], [postCompEquiv]) + +(those are not in 'Equiv.agda' because they need Univalence.agda (which imports Equiv.agda)) +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.PathSplit where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism + +open import Cubical.Foundations.Equiv.Properties + +open import Cubical.Data.Sigma + +record isPathSplitEquiv {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) : Type (ℓ-max ℓ ℓ') where + field + sec : hasSection f + secCong : (x y : A) → hasSection (λ (p : x ≡ y) → cong f p) + +PathSplitEquiv : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Type (ℓ-max ℓ ℓ') +PathSplitEquiv A B = Σ[ f ∈ (A → B) ] isPathSplitEquiv f + +open isPathSplitEquiv + +idIsPathSplitEquiv : ∀ {ℓ} {A : Type ℓ} → isPathSplitEquiv (λ (x : A) → x) +sec idIsPathSplitEquiv = (λ x → x) , (λ _ → refl) +secCong idIsPathSplitEquiv = λ _ _ → (λ p → p) , λ p _ → p + +module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where + + open Iso + + toIso : (f : A → B) → isPathSplitEquiv f → Iso A B + fun (toIso f _) = f + inv (toIso _ p) = p .sec .fst + rightInv (toIso _ p) = p .sec .snd + leftInv (toIso f p) x = p .secCong (p .sec .fst (f x)) x .fst (p .sec .snd (f x)) + + toIsEquiv : (f : A → B) → isPathSplitEquiv f → isEquiv f + toIsEquiv f p = isoToIsEquiv (toIso f p) + + sectionOfEquiv' : (f : A → B) → isEquiv f → B → A + sectionOfEquiv' f isEqv x = isEqv .equiv-proof x .fst .fst + + isSec : (f : A → B) → (pf : isEquiv f) → section f (sectionOfEquiv' f pf) + isSec f isEqv x = isEqv .equiv-proof x .fst .snd + + sectionOfEquiv : (f : A → B) → isEquiv f → hasSection f + sectionOfEquiv f e = sectionOfEquiv' f e , isSec f e + +module _ {ℓ} {A B : Type ℓ} where + + fromIsEquiv : (f : A → B) → isEquiv f → isPathSplitEquiv f + sec (fromIsEquiv f pf) = sectionOfEquiv' f pf , isSec f pf + secCong (fromIsEquiv f pf) x y = sectionOfEquiv (cong f) (isEquivCong (f , pf)) + + pathSplitToEquiv : PathSplitEquiv A B → A ≃ B + fst (pathSplitToEquiv (f , _)) = f + snd (pathSplitToEquiv (_ , e)) = toIsEquiv _ e + + equivToPathSplit : A ≃ B → PathSplitEquiv A B + fst (equivToPathSplit (f , _)) = f + snd (equivToPathSplit (_ , e)) = fromIsEquiv _ e + + equivHasUniqueSection : (f : A → B) → isEquiv f → ∃![ g ∈ (B → A) ] section f g + equivHasUniqueSection f eq = helper' + where + helper : isContr (fiber (λ (φ : B → A) → f ∘ φ) (idfun B)) + helper = (equiv-proof (snd (postCompEquiv (f , eq)))) (idfun B) + + helper' : ∃![ φ ∈ (B → A) ] ((x : B) → f (φ x) ≡ x) + fst helper' = (helper .fst .fst , λ x i → helper .fst .snd i x) + snd helper' y i = (fst (η i) , λ b j → snd (η i) j b) + where η = helper .snd (fst y , λ i b → snd y b i) + +{- + PathSplitEquiv is a proposition and the type + of path split equivs is equivalent to the type of equivalences +-} +isPropIsPathSplitEquiv : ∀ {ℓ} {A B : Type ℓ} (f : A → B) → isProp (isPathSplitEquiv f) +isPropIsPathSplitEquiv {_} {A} {B} f + record { sec = sec-φ ; secCong = secCong-φ } + record { sec = sec-ψ ; secCong = secCong-ψ } i + = + record { + sec = sectionsAreEqual i ; + secCong = λ x y → congSectionsAreEqual x y (secCong-φ x y) (secCong-ψ x y) i + } + where + φ' = record { sec = sec-φ ; secCong = secCong-φ } + ψ' = record { sec = sec-ψ ; secCong = secCong-ψ } + sectionsAreEqual : sec-φ ≡ sec-ψ + sectionsAreEqual = (sym (contraction sec-φ)) ∙ (contraction sec-ψ) + where contraction = snd (equivHasUniqueSection f (toIsEquiv f φ')) + congSectionsAreEqual : (x y : A) (l u : hasSection (λ (p : x ≡ y) → cong f p)) → l ≡ u + congSectionsAreEqual x y l u = (sym (contraction l)) ∙ (contraction u) + where contraction = snd (equivHasUniqueSection + (λ (p : x ≡ y) → cong f p) + (isEquivCong (pathSplitToEquiv (f , φ')))) + +module _ {ℓ} {A B : Type ℓ} where + isEquivIsPathSplitToIsEquiv : (f : A → B) → isEquiv (fromIsEquiv f) + isEquivIsPathSplitToIsEquiv f = + isoToIsEquiv + (iso (fromIsEquiv f) (toIsEquiv f) (λ b → isPropIsPathSplitEquiv f _ _) (λ a → isPropIsEquiv f _ _ )) + + isEquivPathSplitToEquiv : isEquiv (pathSplitToEquiv {A = A} {B = B}) + isEquivPathSplitToEquiv = + isoToIsEquiv + (iso pathSplitToEquiv equivToPathSplit + (λ {(f , e) i → (f , isPropIsEquiv f (toIsEquiv f (fromIsEquiv f e)) e i)}) + (λ {(f , e) i → (f , isPropIsPathSplitEquiv f (fromIsEquiv f (toIsEquiv f e)) e i)})) + + equivPathSplitToEquiv : (PathSplitEquiv A B) ≃ (A ≃ B) + equivPathSplitToEquiv = (pathSplitToEquiv , isEquivPathSplitToEquiv) + + +secCongDep : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} + → (f : ∀ a → B a → C a) {a a' : A} (q : a ≡ a') + → (∀ a (x y : B a) → hasSection (λ (p : x ≡ y) → cong (f a) p)) + → (∀ (x : B a) (y : B a') → hasSection (λ (p : PathP (λ i → B (q i)) x y) → cong₂ f q p)) +secCongDep {B = B} f {a} p secCong + = J (λ a' q → (x : B a) (y : B a') → hasSection (λ (p : PathP (λ i → B (q i)) x y) → cong₂ f q p)) + (secCong a) p diff --git a/Cubical/Foundations/Equiv/Properties.agda b/Cubical/Foundations/Equiv/Properties.agda new file mode 100644 index 0000000000..29c09ffb6a --- /dev/null +++ b/Cubical/Foundations/Equiv/Properties.agda @@ -0,0 +1,200 @@ +{- + +A couple of general facts about equivalences: + +- if f is an equivalence then (cong f) is an equivalence ([equivCong]) +- if f is an equivalence then pre- and postcomposition with f are equivalences ([preCompEquiv], [postCompEquiv]) +- if f is an equivalence then (Σ[ g ] section f g) and (Σ[ g ] retract f g) are contractible ([isContr-section], [isContr-retract]) + +- isHAEquiv is a proposition [isPropIsHAEquiv] +(these are not in 'Equiv.agda' because they need Univalence.agda (which imports Equiv.agda)) +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels + +open import Cubical.Functions.FunExtEquiv + +private + variable + ℓ ℓ′ : Level + A B C : Type ℓ + +isEquivInvEquiv : isEquiv (λ (e : A ≃ B) → invEquiv e) +isEquivInvEquiv = isoToIsEquiv goal where + open Iso + goal : Iso (A ≃ B) (B ≃ A) + goal .fun = invEquiv + goal .inv = invEquiv + goal .rightInv g = equivEq refl + goal .leftInv f = equivEq refl + +invEquivEquiv : (A ≃ B) ≃ (B ≃ A) +invEquivEquiv = _ , isEquivInvEquiv + +isEquivCong : {x y : A} (e : A ≃ B) → isEquiv (λ (p : x ≡ y) → cong (equivFun e) p) +isEquivCong e = isoToIsEquiv (congIso (equivToIso e)) + +congEquiv : {x y : A} (e : A ≃ B) → (x ≡ y) ≃ (equivFun e x ≡ equivFun e y) +congEquiv e = isoToEquiv (congIso (equivToIso e)) + +equivAdjointEquiv : (e : A ≃ B) → ∀ {a b} → (a ≡ invEq e b) ≃ (equivFun e a ≡ b) +equivAdjointEquiv e = compEquiv (congEquiv e) (compPathrEquiv (secEq e _)) + +invEq≡→equivFun≡ : (e : A ≃ B) → ∀ {a b} → invEq e b ≡ a → equivFun e a ≡ b +invEq≡→equivFun≡ e = equivFun (equivAdjointEquiv e) ∘ sym + +isEquivPreComp : (e : A ≃ B) → isEquiv (λ (φ : B → C) → φ ∘ equivFun e) +isEquivPreComp e = snd (equiv→ (invEquiv e) (idEquiv _)) + +preCompEquiv : (e : A ≃ B) → (B → C) ≃ (A → C) +preCompEquiv e = (λ φ → φ ∘ fst e) , isEquivPreComp e + +isEquivPostComp : (e : A ≃ B) → isEquiv (λ (φ : C → A) → e .fst ∘ φ) +isEquivPostComp e = snd (equivΠCod (λ _ → e)) + +postCompEquiv : (e : A ≃ B) → (C → A) ≃ (C → B) +postCompEquiv e = _ , isEquivPostComp e + +-- see also: equivΠCod for a dependent version of postCompEquiv + +hasSection : (A → B) → Type _ +hasSection {A = A} {B = B} f = Σ[ g ∈ (B → A) ] section f g + +hasRetract : (A → B) → Type _ +hasRetract {A = A} {B = B} f = Σ[ g ∈ (B → A) ] retract f g + +isEquiv→isContrHasSection : {f : A → B} → isEquiv f → isContr (hasSection f) +fst (isEquiv→isContrHasSection isEq) = invIsEq isEq , secIsEq isEq +snd (isEquiv→isContrHasSection isEq) (f , ε) i = (λ b → fst (p b i)) , (λ b → snd (p b i)) + where p : ∀ b → (invIsEq isEq b , secIsEq isEq b) ≡ (f b , ε b) + p b = isEq .equiv-proof b .snd (f b , ε b) + +isEquiv→hasSection : {f : A → B} → isEquiv f → hasSection f +isEquiv→hasSection = fst ∘ isEquiv→isContrHasSection + +isContr-hasSection : (e : A ≃ B) → isContr (hasSection (fst e)) +isContr-hasSection e = isEquiv→isContrHasSection (snd e) + +isEquiv→isContrHasRetract : {f : A → B} → isEquiv f → isContr (hasRetract f) +fst (isEquiv→isContrHasRetract isEq) = invIsEq isEq , retIsEq isEq +snd (isEquiv→isContrHasRetract {f = f} isEq) (g , η) = + λ i → (λ b → p b i) , (λ a → q a i) + where p : ∀ b → invIsEq isEq b ≡ g b + p b = sym (η (invIsEq isEq b)) ∙' cong g (secIsEq isEq b) + -- one square from the definition of invIsEq + ieSq : ∀ a → Square (cong g (secIsEq isEq (f a))) + refl + (cong (g ∘ f) (retIsEq isEq a)) + refl + ieSq a k j = g (commSqIsEq isEq a k j) + -- one square from η + ηSq : ∀ a → Square (η (invIsEq isEq (f a))) + (η a) + (cong (g ∘ f) (retIsEq isEq a)) + (retIsEq isEq a) + ηSq a i j = η (retIsEq isEq a i) j + -- and one last square from the definition of p + pSq : ∀ b → Square (η (invIsEq isEq b)) + refl + (cong g (secIsEq isEq b)) + (p b) + pSq b i j = compPath'-filler (sym (η (invIsEq isEq b))) (cong g (secIsEq isEq b)) j i + q : ∀ a → Square (retIsEq isEq a) (η a) (p (f a)) refl + q a i j = hcomp (λ k → λ { (i = i0) → ηSq a j k + ; (i = i1) → η a (j ∧ k) + ; (j = i0) → pSq (f a) i k + ; (j = i1) → η a k + }) + (ieSq a j i) + +isEquiv→hasRetract : {f : A → B} → isEquiv f → hasRetract f +isEquiv→hasRetract = fst ∘ isEquiv→isContrHasRetract + +isContr-hasRetract : (e : A ≃ B) → isContr (hasRetract (fst e)) +isContr-hasRetract e = isEquiv→isContrHasRetract (snd e) + +cong≃ : (F : Type ℓ → Type ℓ′) → (A ≃ B) → F A ≃ F B +cong≃ F e = pathToEquiv (cong F (ua e)) + +cong≃-char : (F : Type ℓ → Type ℓ′) {A B : Type ℓ} (e : A ≃ B) → ua (cong≃ F e) ≡ cong F (ua e) +cong≃-char F e = ua-pathToEquiv (cong F (ua e)) + +cong≃-idEquiv : (F : Type ℓ → Type ℓ′) (A : Type ℓ) → cong≃ F (idEquiv A) ≡ idEquiv (F A) +cong≃-idEquiv F A = cong≃ F (idEquiv A) ≡⟨ cong (λ p → pathToEquiv (cong F p)) uaIdEquiv ⟩ + pathToEquiv refl ≡⟨ pathToEquivRefl ⟩ + idEquiv (F A) ∎ + +isPropIsHAEquiv : {f : A → B} → isProp (isHAEquiv f) +isPropIsHAEquiv {f = f} ishaef = goal ishaef where + equivF : isEquiv f + equivF = isHAEquiv→isEquiv ishaef + + rCoh1 : (sec : hasSection f) → Type _ + rCoh1 (g , ε) = Σ[ η ∈ retract f g ] ∀ x → cong f (η x) ≡ ε (f x) + + rCoh2 : (sec : hasSection f) → Type _ + rCoh2 (g , ε) = Σ[ η ∈ retract f g ] ∀ x → Square (ε (f x)) refl (cong f (η x)) refl + + rCoh3 : (sec : hasSection f) → Type _ + rCoh3 (g , ε) = ∀ x → Σ[ ηx ∈ g (f x) ≡ x ] Square (ε (f x)) refl (cong f ηx) refl + + rCoh4 : (sec : hasSection f) → Type _ + rCoh4 (g , ε) = ∀ x → Path (fiber f (f x)) (g (f x) , ε (f x)) (x , refl) + + characterization : isHAEquiv f ≃ Σ _ rCoh4 + characterization = + isHAEquiv f + -- first convert between Σ and record + ≃⟨ isoToEquiv (iso (λ e → (e .g , e .rinv) , (e .linv , e .com)) + (λ e → record { g = e .fst .fst ; rinv = e .fst .snd + ; linv = e .snd .fst ; com = e .snd .snd }) + (λ _ → refl) λ _ → refl) ⟩ + Σ _ rCoh1 + -- secondly, convert the path into a dependent path for later convenience + ≃⟨ Σ-cong-equiv-snd (λ s → Σ-cong-equiv-snd + λ η → equivΠCod + λ x → compEquiv (flipSquareEquiv {a₀₀ = f x}) (invEquiv slideSquareEquiv)) ⟩ + Σ _ rCoh2 + ≃⟨ Σ-cong-equiv-snd (λ s → invEquiv Σ-Π-≃) ⟩ + Σ _ rCoh3 + ≃⟨ Σ-cong-equiv-snd (λ s → equivΠCod λ x → ΣPath≃PathΣ) ⟩ + Σ _ rCoh4 + ■ + where open isHAEquiv + + goal : isProp (isHAEquiv f) + goal = subst isProp (sym (ua characterization)) + (isPropΣ (isContr→isProp (isEquiv→isContrHasSection equivF)) + λ s → isPropΠ λ x → isProp→isSet (isContr→isProp (equivF .equiv-proof (f x))) _ _) + +-- loop spaces connected by a path are equivalent +conjugatePathEquiv : {A : Type ℓ} {a b : A} (p : a ≡ b) → (a ≡ a) ≃ (b ≡ b) +conjugatePathEquiv p = compEquiv (compPathrEquiv p) (compPathlEquiv (sym p)) + +-- composition on the right induces an equivalence of path types +compr≡Equiv : {A : Type ℓ} {a b c : A} (p q : a ≡ b) (r : b ≡ c) → (p ≡ q) ≃ (p ∙ r ≡ q ∙ r) +compr≡Equiv p q r = congEquiv ((λ s → s ∙ r) , compPathr-isEquiv r) + +-- composition on the left induces an equivalence of path types +compl≡Equiv : {A : Type ℓ} {a b c : A} (p : a ≡ b) (q r : b ≡ c) → (q ≡ r) ≃ (p ∙ q ≡ p ∙ r) +compl≡Equiv p q r = congEquiv ((λ s → p ∙ s) , (compPathl-isEquiv p)) + +isEquivFromIsContr : {A : Type ℓ} {B : Type ℓ′} + → (f : A → B) → isContr A → isContr B + → isEquiv f +isEquivFromIsContr f isContrA isContrB = + subst isEquiv (λ i x → isContr→isProp isContrB (fst B≃A x) (f x) i) (snd B≃A) + where B≃A = isContr→Equiv isContrA isContrB diff --git a/Cubical/Foundations/Everything.agda b/Cubical/Foundations/Everything.agda new file mode 100644 index 0000000000..8810a1338e --- /dev/null +++ b/Cubical/Foundations/Everything.agda @@ -0,0 +1,51 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Everything where + +-- Basic cubical prelude +open import Cubical.Foundations.Prelude public + +-- Definition of Identity types and definitions of J, funExt, +-- univalence and propositional truncation using Id instead of Path +open import Cubical.Foundations.Id + hiding ( _≡_ ; _≡⟨_⟩_ ; _∎ ) + renaming ( _≃_ to EquivId + ; EquivContr to EquivContrId + ; J to JId + ; ap to apId + ; equivFun to equivFunId + ; equivCtr to equivCtrId + ; fiber to fiberId + ; funExt to funExtId + ; isContr to isContrId + ; isProp to isPropId + ; isSet to isSetId + ; isEquiv to isEquivId + ; equivIsEquiv to equivIsEquivId + ; refl to reflId + ; ∥_∥ to propTruncId + ; ∣_∣ to incId + ; isPropIsContr to isPropIsContrId + ; isPropIsEquiv to isPropIsEquivId + ) + +open import Cubical.Foundations.GroupoidLaws public +open import Cubical.Foundations.Function public +open import Cubical.Foundations.Equiv public +open import Cubical.Foundations.Equiv.Properties public +open import Cubical.Foundations.Equiv.Fiberwise +open import Cubical.Foundations.Equiv.PathSplit public +open import Cubical.Foundations.Equiv.BiInvertible public +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels public +open import Cubical.Foundations.Path public +open import Cubical.Foundations.Pointed public +open import Cubical.Foundations.RelationalStructure public +open import Cubical.Foundations.Structure public +open import Cubical.Foundations.Transport public +open import Cubical.Foundations.Univalence public +open import Cubical.Foundations.Univalence.Universe +open import Cubical.Foundations.GroupoidLaws public +open import Cubical.Foundations.Isomorphism public +open import Cubical.Foundations.CartesianKanOps +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.SIP diff --git a/Cubical/Foundations/Function.agda b/Cubical/Foundations/Function.agda new file mode 100644 index 0000000000..60c460c91c --- /dev/null +++ b/Cubical/Foundations/Function.agda @@ -0,0 +1,154 @@ +{- + Definitions for functions +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Function where + +open import Cubical.Foundations.Prelude + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + B : A → Type ℓ + C : (a : A) → B a → Type ℓ + D : (a : A) (b : B a) → C a b → Type ℓ + E : (x : A) → (y : B x) → (z : C x y) → (w : D x y z) → Type ℓ + F : (x : A) → (y : B x) → (z : C x y) → (w : D x y z) → (u : E x y z w) → Type ℓ + +-- The identity function +idfun : (A : Type ℓ) → A → A +idfun _ x = x + +infixr 9 _∘_ + +_∘_ : (g : {a : A} → (b : B a) → C a b) → (f : (a : A) → B a) → (a : A) → C a (f a) +g ∘ f = λ x → g (f x) +{-# INLINE _∘_ #-} + +∘-assoc : (h : {a : A} {b : B a} → (c : C a b) → D a b c) + (g : {a : A} → (b : B a) → C a b) + (f : (a : A) → B a) + → (h ∘ g) ∘ f ≡ h ∘ (g ∘ f) +∘-assoc h g f i x = h (g (f x)) + +∘-idˡ : (f : (a : A) → B a) → f ∘ idfun A ≡ f +∘-idˡ f i x = f x + +∘-idʳ : (f : (a : A) → B a) → (λ {a} → idfun (B a)) ∘ f ≡ f +∘-idʳ f i x = f x + +flip : {B : Type ℓ'} {C : A → B → Type ℓ''} → + ((x : A) (y : B) → C x y) → ((y : B) (x : A) → C x y) +flip f y x = f x y +{-# INLINE flip #-} + +const : {B : Type ℓ} → A → B → A +const x = λ _ → x +{-# INLINE const #-} + +case_of_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} → (x : A) → (∀ x → B x) → B x +case x of f = f x +{-# INLINE case_of_ #-} + +case_return_of_ : ∀ {ℓ ℓ'} {A : Type ℓ} (x : A) (B : A → Type ℓ') → (∀ x → B x) → B x +case x return P of f = f x +{-# INLINE case_return_of_ #-} + +uncurry : ((x : A) → (y : B x) → C x y) → (p : Σ A B) → C (fst p) (snd p) +uncurry f (x , y) = f x y + +uncurry2 : ((x : A) → (y : B x) → (z : C x y) → D x y z) + → (p : Σ A (λ x → Σ (B x) (C x))) → D (p .fst) (p .snd .fst) (p .snd .snd) +uncurry2 f (x , y , z) = f x y z + +uncurry3 : ((x : A) → (y : B x) → (z : C x y) → (w : D x y z) → E x y z w) + → (p : Σ A (λ x → Σ (B x) (λ y → Σ (C x y) (D x y)))) + → E (p .fst) (p .snd .fst) (p .snd .snd .fst) (p .snd .snd .snd) +uncurry3 f (x , y , z , w) = f x y z w + +uncurry4 : ((x : A) → (y : B x) → (z : C x y) → (w : D x y z) → (u : E x y z w) → F x y z w u) + → (p : Σ A (λ x → Σ (B x) (λ y → Σ (C x y) (λ z → Σ (D x y z) (E x y z))))) + → F (p .fst) (p .snd .fst) (p .snd .snd .fst) (p .snd .snd .snd .fst) (p .snd .snd .snd .snd) +uncurry4 f (x , y , z , w , u) = f x y z w u + + +curry : ((p : Σ A B) → C (fst p) (snd p)) → (x : A) → (y : B x) → C x y +curry f x y = f (x , y) + +module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where + -- Notions of 'coherently constant' functions for low dimensions. + -- These are the properties of functions necessary to e.g. eliminate + -- from the propositional truncation. + + -- 2-Constant functions are coherently constant if B is a set. + 2-Constant : (A → B) → Type _ + 2-Constant f = ∀ x y → f x ≡ f y + + 2-Constant-isProp : isSet B → (f : A → B) → isProp (2-Constant f) + 2-Constant-isProp Bset f link1 link2 i x y j + = Bset (f x) (f y) (link1 x y) (link2 x y) i j + + -- 3-Constant functions are coherently constant if B is a groupoid. + record 3-Constant (f : A → B) : Type (ℓ-max ℓ ℓ') where + field + link : 2-Constant f + coh₁ : ∀ x y z → Square (link x y) (link x z) refl (link y z) + + coh₂ : ∀ x y z → Square (link x z) (link y z) (link x y) refl + coh₂ x y z i j + = hcomp (λ k → λ + { (j = i0) → link x y i + ; (i = i0) → link x z (j ∧ k) + ; (j = i1) → link x z (i ∨ k) + ; (i = i1) → link y z j + }) + (coh₁ x y z j i) + + link≡refl : ∀ x → refl ≡ link x x + link≡refl x i j + = hcomp (λ k → λ + { (i = i0) → link x x (j ∧ ~ k) + ; (i = i1) → link x x j + ; (j = i0) → f x + ; (j = i1) → link x x (~ i ∧ ~ k) + }) + (coh₁ x x x (~ i) j) + + downleft : ∀ x y → Square (link x y) refl refl (link y x) + downleft x y i j + = hcomp (λ k → λ + { (i = i0) → link x y j + ; (i = i1) → link≡refl x (~ k) j + ; (j = i0) → f x + ; (j = i1) → link y x i + }) + (coh₁ x y x i j) + + link≡symlink : ∀ x y → link x y ≡ sym (link y x) + link≡symlink x y i j + = hcomp (λ k → λ + { (i = i0) → link x y j + ; (i = i1) → link y x (~ j ∨ ~ k) + ; (j = i0) → f x + ; (j = i1) → link y x (i ∧ ~ k) + }) + (downleft x y i j) + +homotopyNatural : {B : Type ℓ'} {f g : A → B} (H : ∀ a → f a ≡ g a) {x y : A} (p : x ≡ y) → + H x ∙ cong g p ≡ cong f p ∙ H y +homotopyNatural {f = f} {g = g} H {x} {y} p i j = + hcomp (λ k → λ { (i = i0) → compPath-filler (H x) (cong g p) k j + ; (i = i1) → compPath-filler' (cong f p) (H y) k j + ; (j = i0) → cong f p (i ∧ ~ k) + ; (j = i1) → cong g p (i ∨ k) }) + (H (p i) j) + +homotopySymInv : {f : A → A} (H : ∀ a → f a ≡ a) (a : A) + → Path (f a ≡ f a) (λ i → H (H a (~ i)) i) refl +homotopySymInv {f = f} H a j i = + hcomp (λ k → λ { (i = i0) → f a + ; (i = i1) → H a (j ∧ ~ k) + ; (j = i0) → H (H a (~ i)) i + ; (j = i1) → H a (i ∧ ~ k)}) + (H (H a (~ i ∨ j)) i) diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda new file mode 100644 index 0000000000..3dafd58f2f --- /dev/null +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -0,0 +1,497 @@ +{- + +This file proves the higher groupoid structure of types +for homogeneous and heterogeneous paths + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.GroupoidLaws where + +open import Cubical.Foundations.Prelude + +private + variable + ℓ : Level + A : Type ℓ + x y z w v : A + +_⁻¹ : (x ≡ y) → (y ≡ x) +x≡y ⁻¹ = sym x≡y + +infix 40 _⁻¹ + +-- homogeneous groupoid laws + +symInvo : (p : x ≡ y) → p ≡ p ⁻¹ ⁻¹ +symInvo p = refl + +rUnit : (p : x ≡ y) → p ≡ p ∙ refl +rUnit p j i = compPath-filler p refl j i + +-- The filler of left unit: lUnit-filler p = +-- PathP (λ i → PathP (λ j → PathP (λ k → A) x (p (~ j ∨ i))) +-- (refl i) (λ j → compPath-filler refl p i j)) (λ k i → (p (~ k ∧ i ))) (lUnit p) + +lUnit-filler : {x y : A} (p : x ≡ y) → I → I → I → A +lUnit-filler {x = x} p j k i = + hfill (λ j → λ { (i = i0) → x + ; (i = i1) → p (~ k ∨ j ) + ; (k = i0) → p i + -- ; (k = i1) → compPath-filler refl p j i + }) (inS (p (~ k ∧ i ))) j + +lUnit : (p : x ≡ y) → p ≡ refl ∙ p +lUnit p j i = lUnit-filler p i1 j i + +symRefl : refl {x = x} ≡ refl ⁻¹ +symRefl i = refl + +compPathRefl : refl {x = x} ≡ refl ∙ refl +compPathRefl = rUnit refl + +-- The filler of right cancellation: rCancel-filler p = +-- PathP (λ i → PathP (λ j → PathP (λ k → A) x (p (~ j ∧ ~ i))) +-- (λ j → compPath-filler p (p ⁻¹) i j) (refl i)) (λ j i → (p (i ∧ ~ j))) (rCancel p) + +rCancel-filler : ∀ {x y : A} (p : x ≡ y) → (k j i : I) → A +rCancel-filler {x = x} p k j i = + hfill (λ k → λ { (i = i0) → x + ; (i = i1) → p (~ k ∧ ~ j) + -- ; (j = i0) → compPath-filler p (p ⁻¹) k i + ; (j = i1) → x + }) (inS (p (i ∧ ~ j))) k + +rCancel : (p : x ≡ y) → p ∙ p ⁻¹ ≡ refl +rCancel {x = x} p j i = rCancel-filler p i1 j i + +rCancel-filler' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → (i j k : I) → A +rCancel-filler' {x = x} {y} p i j k = + hfill + (λ i → λ + { (j = i1) → p (~ i ∧ k) + ; (k = i0) → x + ; (k = i1) → p (~ i) + }) + (inS (p k)) + (~ i) + +rCancel' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → p ∙ p ⁻¹ ≡ refl +rCancel' p j k = rCancel-filler' p i0 j k + +lCancel : (p : x ≡ y) → p ⁻¹ ∙ p ≡ refl +lCancel p = rCancel (p ⁻¹) + +assoc : (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → + p ∙ q ∙ r ≡ (p ∙ q) ∙ r +assoc p q r k = (compPath-filler p q k) ∙ compPath-filler' q r (~ k) + + +-- heterogeneous groupoid laws + +symInvoP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → (p : PathP A x y) → + PathP (λ j → PathP (λ i → symInvo (λ i → A i) j i) x y) p (symP (symP p)) +symInvoP p = refl + +rUnitP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → (p : PathP A x y) → + PathP (λ j → PathP (λ i → rUnit (λ i → A i) j i) x y) p (compPathP p refl) +rUnitP p j i = compPathP-filler p refl j i + +lUnitP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → (p : PathP A x y) → + PathP (λ j → PathP (λ i → lUnit (λ i → A i) j i) x y) p (compPathP refl p) +lUnitP {A = A} {x = x} p k i = + comp (λ j → lUnit-filler (λ i → A i) j k i) + (λ j → λ { (i = i0) → x + ; (i = i1) → p (~ k ∨ j ) + ; (k = i0) → p i + }) (p (~ k ∧ i )) + +rCancelP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → (p : PathP A x y) → + PathP (λ j → PathP (λ i → rCancel (λ i → A i) j i) x x) (compPathP p (symP p)) refl +rCancelP {A = A} {x = x} p j i = + comp (λ k → rCancel-filler (λ i → A i) k j i) + (λ k → λ { (i = i0) → x + ; (i = i1) → p (~ k ∧ ~ j) + ; (j = i1) → x + }) (p (i ∧ ~ j)) + +lCancelP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → (p : PathP A x y) → + PathP (λ j → PathP (λ i → lCancel (λ i → A i) j i) y y) (compPathP (symP p) p) refl +lCancelP p = rCancelP (symP p) + + + +assocP : {A : I → Type ℓ} {x : A i0} {y : A i1} {B_i1 : Type ℓ} {B : (A i1) ≡ B_i1} {z : B i1} + {C_i1 : Type ℓ} {C : (B i1) ≡ C_i1} {w : C i1} (p : PathP A x y) (q : PathP (λ i → B i) y z) (r : PathP (λ i → C i) z w) → + PathP (λ j → PathP (λ i → assoc (λ i → A i) B C j i) x w) (compPathP p (compPathP q r)) (compPathP (compPathP p q) r) +assocP {A = A} {B = B} {C = C} p q r k i = + comp (\ j' → hfill (λ j → λ { + (i = i0) → A i0 + ; (i = i1) → compPath-filler' (λ i₁ → B i₁) (λ i₁ → C i₁) (~ k) j }) + (inS (compPath-filler (λ i₁ → A i₁) (λ i₁ → B i₁) k i)) j') + (λ j → λ + { (i = i0) → p i0 + ; (i = i1) → + comp (\ j' → hfill ((λ l → λ + { (j = i0) → B k + ; (j = i1) → C l + ; (k = i1) → C (j ∧ l) + })) (inS (B ( j ∨ k)) ) j') + (λ l → λ + { (j = i0) → q k + ; (j = i1) → r l + ; (k = i1) → r (j ∧ l) + }) + (q (j ∨ k)) + }) + (compPathP-filler p q k i) + + + +-- Loic's code below + +-- some exchange law for doubleCompPath and refl + +invSides-filler : {x y z : A} (p : x ≡ y) (q : x ≡ z) → Square p (sym q) q (sym p) +invSides-filler {x = x} p q i j = + hcomp (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (~ j ∧ k) + ; (j = i0) → q (i ∧ k) + ; (j = i1) → p (~ i ∧ k)}) + x + +leftright : {ℓ : Level} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → + (refl ∙∙ p ∙∙ q) ≡ (p ∙∙ q ∙∙ refl) +leftright p q i j = + hcomp (λ t → λ { (j = i0) → p (i ∧ (~ t)) + ; (j = i1) → q (t ∨ i) }) + (invSides-filler q (sym p) (~ i) j) + +-- equating doubleCompPath and a succession of two compPath + +split-leftright : {ℓ : Level} {A : Type ℓ} {w x y z : A} (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → + (p ∙∙ q ∙∙ r) ≡ (refl ∙∙ (p ∙∙ q ∙∙ refl) ∙∙ r) +split-leftright p q r j i = + hcomp (λ t → λ { (i = i0) → p (~ j ∧ ~ t) + ; (i = i1) → r t }) + (doubleCompPath-filler p q refl j i) + +split-leftright' : {ℓ : Level} {A : Type ℓ} {w x y z : A} (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → + (p ∙∙ q ∙∙ r) ≡ (p ∙∙ (refl ∙∙ q ∙∙ r) ∙∙ refl) +split-leftright' p q r j i = + hcomp (λ t → λ { (i = i0) → p (~ t) + ; (i = i1) → r (j ∨ t) }) + (doubleCompPath-filler refl q r j i) + +doubleCompPath-elim : {ℓ : Level} {A : Type ℓ} {w x y z : A} (p : w ≡ x) (q : x ≡ y) + (r : y ≡ z) → (p ∙∙ q ∙∙ r) ≡ (p ∙ q) ∙ r +doubleCompPath-elim p q r = (split-leftright p q r) ∙ (λ i → (leftright p q (~ i)) ∙ r) + +doubleCompPath-elim' : {ℓ : Level} {A : Type ℓ} {w x y z : A} (p : w ≡ x) (q : x ≡ y) + (r : y ≡ z) → (p ∙∙ q ∙∙ r) ≡ p ∙ (q ∙ r) +doubleCompPath-elim' p q r = (split-leftright' p q r) ∙ (sym (leftright p (q ∙ r))) + +cong-∙∙-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y z w : A} + (f : A → B) (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) + → I → I → I → B +cong-∙∙-filler {A = A} f p q r k j i = + hfill ((λ k → λ { (j = i1) → doubleCompPath-filler (cong f p) (cong f q) (cong f r) k i + ; (j = i0) → f (doubleCompPath-filler p q r k i) + ; (i = i0) → f (p (~ k)) + ; (i = i1) → f (r k) })) + (inS (f (q i))) + k + +cong-∙∙ : ∀ {B : Type ℓ} (f : A → B) (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) + → cong f (p ∙∙ q ∙∙ r) ≡ (cong f p) ∙∙ (cong f q) ∙∙ (cong f r) +cong-∙∙ f p q r j i = cong-∙∙-filler f p q r i1 j i + +cong-∙ : ∀ {B : Type ℓ} (f : A → B) (p : x ≡ y) (q : y ≡ z) + → cong f (p ∙ q) ≡ (cong f p) ∙ (cong f q) +cong-∙ f p q = cong-∙∙ f refl p q + +hcomp-unique : ∀ {ℓ} {A : Type ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → + (h2 : ∀ i → A [ (φ ∨ ~ i) ↦ (\ { (φ = i1) → u i 1=1; (i = i0) → outS u0}) ]) + → (hcomp u (outS u0) ≡ outS (h2 i1)) [ φ ↦ (\ { (φ = i1) → (\ i → u i1 1=1)}) ] +hcomp-unique {φ = φ} u u0 h2 = inS (\ i → hcomp (\ k → \ { (φ = i1) → u k 1=1 + ; (i = i1) → outS (h2 k) }) + (outS u0)) + + +lid-unique : ∀ {ℓ} {A : Type ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → + (h1 h2 : ∀ i → A [ (φ ∨ ~ i) ↦ (\ { (φ = i1) → u i 1=1; (i = i0) → outS u0}) ]) + → (outS (h1 i1) ≡ outS (h2 i1)) [ φ ↦ (\ { (φ = i1) → (\ i → u i1 1=1)}) ] +lid-unique {φ = φ} u u0 h1 h2 = inS (\ i → hcomp (\ k → \ { (φ = i1) → u k 1=1 + ; (i = i0) → outS (h1 k) + ; (i = i1) → outS (h2 k) }) + (outS u0)) + + +transp-hcomp : ∀ {ℓ} (φ : I) {A' : Type ℓ} + (A : (i : I) → Type ℓ [ φ ↦ (λ _ → A') ]) (let B = \ (i : I) → outS (A i)) + → ∀ {ψ} (u : I → Partial ψ (B i0)) → (u0 : B i0 [ ψ ↦ u i0 ]) → + (transp (\ i → B i) φ (hcomp u (outS u0)) ≡ hcomp (\ i o → transp (\ i → B i) φ (u i o)) (transp (\ i → B i) φ (outS u0))) + [ ψ ↦ (\ { (ψ = i1) → (\ i → transp (\ i → B i) φ (u i1 1=1))}) ] +transp-hcomp φ A u u0 = inS (sym (outS (hcomp-unique + ((\ i o → transp (\ i → B i) φ (u i o))) (inS (transp (\ i → B i) φ (outS u0))) + \ i → inS (transp (\ i → B i) φ (hfill u u0 i))))) + where + B = \ (i : I) → outS (A i) + + +hcomp-cong : ∀ {ℓ} {A : Type ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → + (u' : I → Partial φ A) → (u0' : A [ φ ↦ u' i0 ]) → + (ueq : ∀ i → PartialP φ (\ o → u i o ≡ u' i o)) → (outS u0 ≡ outS u0') [ φ ↦ (\ { (φ = i1) → ueq i0 1=1}) ] + → (hcomp u (outS u0) ≡ hcomp u' (outS u0')) [ φ ↦ (\ { (φ = i1) → ueq i1 1=1 }) ] +hcomp-cong u u0 u' u0' ueq 0eq = inS (\ j → hcomp (\ i o → ueq i o j) (outS 0eq j)) + + +congFunct-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y z : A} (f : A → B) (p : x ≡ y) (q : y ≡ z) + → I → I → I → B +congFunct-filler {x = x} f p q i j z = + hfill (λ k → λ { (i = i0) → f x + ; (i = i1) → f (q k) + ; (j = i0) → f (compPath-filler p q k i)}) + (inS (f (p i))) + z + +congFunct : ∀ {ℓ} {B : Type ℓ} (f : A → B) (p : x ≡ y) (q : y ≡ z) → cong f (p ∙ q) ≡ cong f p ∙ cong f q +congFunct f p q j i = congFunct-filler f p q i j i1 + + +-- congFunct for dependent types +congFunct-dep : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} (f : (a : A) → B a) (p : x ≡ y) (q : y ≡ z) + → PathP (λ i → PathP (λ j → B (compPath-filler p q i j)) (f x) (f (q i))) (cong f p) (cong f (p ∙ q)) +congFunct-dep {B = B} {x = x} f p q i j = f (compPath-filler p q i j) + +cong₂Funct : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} {B : Type ℓ'} (f : A → A → B) → + (p : x ≡ y) → + {u v : A} (q : u ≡ v) → + cong₂ f p q ≡ cong (λ x → f x u) p ∙ cong (f y) q +cong₂Funct {x = x} {y = y} f p {u = u} {v = v} q j i = + hcomp (λ k → λ { (i = i0) → f x u + ; (i = i1) → f y (q k) + ; (j = i0) → f (p i) (q (i ∧ k))}) + (f (p i) u) + +symDistr-filler : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → I → I → I → A +symDistr-filler {A = A} {z = z} p q i j k = + hfill (λ k → λ { (i = i0) → q (k ∨ j) + ; (i = i1) → p (~ k ∧ j) }) + (inS (invSides-filler q (sym p) i j)) + k + +symDistr : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → sym (p ∙ q) ≡ sym q ∙ sym p +symDistr p q i j = symDistr-filler p q j i i1 + +-- we can not write hcomp-isEquiv : {ϕ : I} → (p : I → Partial ϕ A) → isEquiv (λ (a : A [ ϕ ↦ p i0 ]) → hcomp p a) +-- due to size issues. But what we can write (compare to hfill) is: +hcomp-equivFillerSub : {ϕ : I} → (p : I → Partial ϕ A) → (a : A [ ϕ ↦ p i0 ]) + → (i : I) + → A [ ϕ ∨ i ∨ ~ i ↦ (λ { (i = i0) → outS a + ; (i = i1) → hcomp (λ i → p (~ i)) (hcomp p (outS a)) + ; (ϕ = i1) → p i0 1=1 }) ] +hcomp-equivFillerSub {ϕ = ϕ} p a i = + inS (hcomp (λ k → λ { (i = i1) → hfill (λ j → p (~ j)) (inS (hcomp p (outS a))) k + ; (i = i0) → outS a + ; (ϕ = i1) → p (~ k ∧ i) 1=1 }) + (hfill p a i)) + +hcomp-equivFiller : {ϕ : I} → (p : I → Partial ϕ A) → (a : A [ ϕ ↦ p i0 ]) + → (i : I) → A +hcomp-equivFiller p a i = outS (hcomp-equivFillerSub p a i) + + +pentagonIdentity : (p : x ≡ y) → (q : y ≡ z) → (r : z ≡ w) → (s : w ≡ v) + → + (assoc p q (r ∙ s) ∙ assoc (p ∙ q) r s) + ≡ + cong (p ∙_) (assoc q r s) ∙∙ assoc p (q ∙ r) s ∙∙ cong (_∙ s) (assoc p q r) + +pentagonIdentity {x = x} {y} p q r s = + (λ i → + (λ j → cong (p ∙_) (assoc q r s) (i ∧ j)) + ∙∙ (λ j → lemma₀₀ i j ∙ lemma₀₁ i j) + ∙∙ (λ j → lemma₁₀ i j ∙ lemma₁₁ i j) + ) + where + + + lemma₀₀ : ( i j : I) → _ ≡ _ + lemma₀₀ i j i₁ = + hcomp + (λ k → λ { (j = i0) → p i₁ + ; (i₁ = i0) → x + ; (i₁ = i1) → hcomp + (λ k₁ → λ { (i = i0) → (q (j ∧ k)) + ; (k = i0) → y + ; (j = i0) → y + ; (j = i1)(k = i1) → r (k₁ ∧ i)}) + (q (j ∧ k)) + }) (p i₁) + + lemma₀₁ : ( i j : I) → hcomp + (λ k → λ {(i = i0) → q j + ; (j = i0) → y + ; (j = i1) → r (k ∧ i) + }) + (q j) ≡ _ + lemma₀₁ i j i₁ = (hcomp + (λ k → λ { (j = i1) → hcomp + (λ k₁ → λ { (i₁ = i0) → r i + ; (k = i0) → r i + ; (i = i1) → s (k₁ ∧ k ∧ i₁) + ; (i₁ = i1)(k = i1) → s k₁ }) + (r ((i₁ ∧ k) ∨ i)) + ; (i₁ = i0) → compPath-filler q r i j + ; (i₁ = i1) → hcomp + (λ k₁ → λ { (k = i0) → r i + ; (k = i1) → s k₁ + ; (i = i1) → s (k ∧ k₁)}) + (r (i ∨ k))}) + (hfill + (λ k → λ { (j = i1) → r k + ; (i₁ = i1) → r k + ; (i₁ = i0)(j = i0) → y }) + (inS (q (i₁ ∨ j))) i)) + + lemma₁₁ : ( i j : I) → (r (i ∨ j)) ≡ _ + lemma₁₁ i j i₁ = + hcomp + (λ k → λ { (i = i1) → s (i₁ ∧ k) + ; (j = i1) → s (i₁ ∧ k) + ; (i₁ = i0) → r (i ∨ j) + ; (i₁ = i1) → s k + }) (r (i ∨ j ∨ i₁)) + + + lemma₁₀-back : I → I → I → _ + lemma₁₀-back i j i₁ = + hcomp + (λ k → λ { + (i₁ = i0) → x + ; (i₁ = i1) → hcomp + (λ k₁ → λ { (k = i0) → q (j ∨ ~ i) + ; (k = i1) → r (k₁ ∧ j) + ; (j = i0) → q (k ∨ ~ i) + ; (j = i1) → r (k₁ ∧ k) + ; (i = i0) → r (k ∧ j ∧ k₁) + }) + (q (k ∨ j ∨ ~ i)) + ; (i = i0)(j = i0) → (p ∙ q) i₁ + }) + (hcomp + (λ k → λ { (i₁ = i0) → x + ; (i₁ = i1) → q ((j ∨ ~ i ) ∧ k) + ; (j = i0)(i = i1) → p i₁ + }) + (p i₁)) + + + lemma₁₀-front : I → I → I → _ + lemma₁₀-front i j i₁ = + (((λ _ → x) ∙∙ compPath-filler p q j ∙∙ + (λ i₁ → + hcomp + (λ k → λ { (i₁ = i0) → q j + ; (i₁ = i1) → r (k ∧ (j ∨ i)) + ; (j = i0)(i = i0) → q i₁ + ; (j = i1) → r (i₁ ∧ k) + }) + (q (j ∨ i₁)) + )) i₁) + + compPath-filler-in-filler : + (p : _ ≡ y) → (q : _ ≡ _ ) + → _≡_ {A = Square (p ∙ q) (p ∙ q) (λ _ → x) (λ _ → z)} + (λ i j → hcomp + (λ i₂ → + λ { (j = i0) → x + ; (j = i1) → q (i₂ ∨ ~ i) + ; (i = i0) → (p ∙ q) j + }) + (compPath-filler p q (~ i) j)) + (λ _ → p ∙ q) + compPath-filler-in-filler p q z i j = + hcomp + (λ k → λ { + (j = i0) → p i0 + ; (j = i1) → q (k ∨ ~ i ∧ ~ z) + ; (i = i0) → hcomp + (λ i₂ → λ { + (j = i0) → p i0 + ;(j = i1) → q ((k ∨ ~ z) ∧ i₂) + ;(z = i1) (k = i0) → p j + }) + (p j) + ; (i = i1) → compPath-filler p (λ i₁ → q (k ∧ i₁)) k j + ; (z = i0) → hfill + ((λ i₂ → λ { (j = i0) → p i0 + ; (j = i1) → q (i₂ ∨ ~ i) + ; (i = i0) → (p ∙ q) j + })) + (inS ((compPath-filler p q (~ i) j))) k + ; (z = i1) → compPath-filler p q k j + }) + (compPath-filler p q (~ i ∧ ~ z) j) + + + cube-comp₋₀₋ : + (c : I → I → I → A) + → {a' : Square _ _ _ _} + → (λ i i₁ → c i i0 i₁) ≡ a' + → (I → I → I → A) + cube-comp₋₀₋ c p i j k = + hcomp + (λ l → λ { + (i = i0) → c i0 j k + ;(i = i1) → c i1 j k + ;(j = i0) → p l i k + ;(j = i1) → c i i1 k + ;(k = i0) → c i j i0 + ;(k = i1) → c i j i1 + }) + (c i j k) + + cube-comp₀₋₋ : + (c : I → I → I → A) + → {a' : Square _ _ _ _} + → (λ i i₁ → c i0 i i₁) ≡ a' + → (I → I → I → A) + cube-comp₀₋₋ c p i j k = + hcomp + (λ l → λ { + (i = i0) → p l j k + ;(i = i1) → c i1 j k + ;(j = i0) → c i i0 k + ;(j = i1) → c i i1 k + ;(k = i0) → c i j i0 + ;(k = i1) → c i j i1 + }) + (c i j k) + + + + lemma₁₀-back' : _ + lemma₁₀-back' k j i₁ = + (cube-comp₋₀₋ (lemma₁₀-back) + (compPath-filler-in-filler p q)) k j i₁ + + + lemma₁₀ : ( i j : I) → _ ≡ _ + lemma₁₀ i j i₁ = + (cube-comp₀₋₋ lemma₁₀-front (sym lemma₁₀-back')) i j i₁ + +-- misc. +∙∙lCancel-fill : ∀ {ℓ} {A : Type ℓ} {x y : A} + → (p : x ≡ y) + → I → I → I → A +∙∙lCancel-fill p i j k = + hfill (λ k → λ { (i = i1) → p k + ; (j = i0) → p k + ; (j = i1) → p k}) + (inS (p i0)) k + +∙∙lCancel : ∀ {ℓ} {A : Type ℓ} {x y : A} + → (p : x ≡ y) + → sym p ∙∙ refl ∙∙ p ≡ refl +∙∙lCancel p i j = ∙∙lCancel-fill p i j i1 diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda new file mode 100644 index 0000000000..9b10de92cd --- /dev/null +++ b/Cubical/Foundations/HLevels.agda @@ -0,0 +1,658 @@ +{- + +Basic theory about h-levels/n-types: + +- Basic properties of isContr, isProp and isSet (definitions are in Prelude) + +- Hedberg's theorem can be found in Cubical/Relation/Nullary/DecidableEq + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.HLevels where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Structure +open import Cubical.Functions.FunExtEquiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence using (ua ; univalenceIso) + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ; zero; suc; _+_; +-zero; +-comm) + +HLevel : Type₀ +HLevel = ℕ + +private + variable + ℓ ℓ' ℓ'' ℓ''' ℓ'''' ℓ''''' : Level + A : Type ℓ + B : A → Type ℓ + C : (x : A) → B x → Type ℓ + D : (x : A) (y : B x) → C x y → Type ℓ + E : (x : A) (y : B x) → (z : C x y) → D x y z → Type ℓ + F : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) → Type ℓ + w x y z : A + n : HLevel + +isOfHLevel : HLevel → Type ℓ → Type ℓ +isOfHLevel 0 A = isContr A +isOfHLevel 1 A = isProp A +isOfHLevel (suc (suc n)) A = (x y : A) → isOfHLevel (suc n) (x ≡ y) + +isOfHLevelFun : (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') +isOfHLevelFun n f = ∀ b → isOfHLevel n (fiber f b) + +isOfHLevelΩ→isOfHLevel : + ∀ {ℓ} {A : Type ℓ} (n : ℕ) + → ((x : A) → isOfHLevel (suc n) (x ≡ x)) → isOfHLevel (2 + n) A +isOfHLevelΩ→isOfHLevel zero hΩ x y = + J (λ y p → (q : x ≡ y) → p ≡ q) (hΩ x refl) +isOfHLevelΩ→isOfHLevel (suc n) hΩ x y = + J (λ y p → (q : x ≡ y) → isOfHLevel (suc n) (p ≡ q)) (hΩ x refl) + +TypeOfHLevel : ∀ ℓ → HLevel → Type (ℓ-suc ℓ) +TypeOfHLevel ℓ n = TypeWithStr ℓ (isOfHLevel n) + +hProp hSet hGroupoid h2Groupoid : ∀ ℓ → Type (ℓ-suc ℓ) +hProp ℓ = TypeOfHLevel ℓ 1 +hSet ℓ = TypeOfHLevel ℓ 2 +hGroupoid ℓ = TypeOfHLevel ℓ 3 +h2Groupoid ℓ = TypeOfHLevel ℓ 4 + +-- lower h-levels imply higher h-levels + +isOfHLevelSuc : (n : HLevel) → isOfHLevel n A → isOfHLevel (suc n) A +isOfHLevelSuc 0 = isContr→isProp +isOfHLevelSuc 1 = isProp→isSet +isOfHLevelSuc (suc (suc n)) h a b = isOfHLevelSuc (suc n) (h a b) + +isSet→isGroupoid : isSet A → isGroupoid A +isSet→isGroupoid = isOfHLevelSuc 2 + +isGroupoid→is2Groupoid : isGroupoid A → is2Groupoid A +isGroupoid→is2Groupoid = isOfHLevelSuc 3 + +isOfHLevelPlus : (m : HLevel) → isOfHLevel n A → isOfHLevel (m + n) A +isOfHLevelPlus zero hA = hA +isOfHLevelPlus (suc m) hA = isOfHLevelSuc _ (isOfHLevelPlus m hA) + +isContr→isOfHLevel : (n : HLevel) → isContr A → isOfHLevel n A +isContr→isOfHLevel zero cA = cA +isContr→isOfHLevel (suc n) cA = isOfHLevelSuc _ (isContr→isOfHLevel n cA) + +isProp→isOfHLevelSuc : (n : HLevel) → isProp A → isOfHLevel (suc n) A +isProp→isOfHLevelSuc zero pA = pA +isProp→isOfHLevelSuc (suc n) pA = isOfHLevelSuc _ (isProp→isOfHLevelSuc n pA) + +isOfHLevelPlus' : (m : HLevel) → isOfHLevel m A → isOfHLevel (m + n) A +isOfHLevelPlus' {n = n} 0 = isContr→isOfHLevel n +isOfHLevelPlus' {n = n} 1 = isProp→isOfHLevelSuc n +isOfHLevelPlus' {n = n} (suc (suc m)) hA a₀ a₁ = isOfHLevelPlus' (suc m) (hA a₀ a₁) + +-- hlevel of path types + +isProp→isContrPath : isProp A → (x y : A) → isContr (x ≡ y) +isProp→isContrPath h x y = h x y , isProp→isSet h x y _ + +isContr→isContrPath : isContr A → (x y : A) → isContr (x ≡ y) +isContr→isContrPath cA = isProp→isContrPath (isContr→isProp cA) + +isOfHLevelPath' : (n : HLevel) → isOfHLevel (suc n) A → (x y : A) → isOfHLevel n (x ≡ y) +isOfHLevelPath' 0 = isProp→isContrPath +isOfHLevelPath' (suc n) h x y = h x y + +isOfHLevelPath'⁻ : (n : HLevel) → ((x y : A) → isOfHLevel n (x ≡ y)) → isOfHLevel (suc n) A +isOfHLevelPath'⁻ zero h x y = h x y .fst +isOfHLevelPath'⁻ (suc n) h = h + +isOfHLevelPath : (n : HLevel) → isOfHLevel n A → (x y : A) → isOfHLevel n (x ≡ y) +isOfHLevelPath 0 h x y = isContr→isContrPath h x y +isOfHLevelPath (suc n) h x y = isOfHLevelSuc n (isOfHLevelPath' n h x y) + +-- h-level of isOfHLevel + +isPropIsOfHLevel : (n : HLevel) → isProp (isOfHLevel n A) +isPropIsOfHLevel 0 = isPropIsContr +isPropIsOfHLevel 1 = isPropIsProp +isPropIsOfHLevel (suc (suc n)) f g i a b = + isPropIsOfHLevel (suc n) (f a b) (g a b) i + +isPropIsSet : isProp (isSet A) +isPropIsSet = isPropIsOfHLevel 2 + +isPropIsGroupoid : isProp (isGroupoid A) +isPropIsGroupoid = isPropIsOfHLevel 3 + +isPropIs2Groupoid : isProp (is2Groupoid A) +isPropIs2Groupoid = isPropIsOfHLevel 4 + +TypeOfHLevel≡ : (n : HLevel) {X Y : TypeOfHLevel ℓ n} → ⟨ X ⟩ ≡ ⟨ Y ⟩ → X ≡ Y +TypeOfHLevel≡ n = Σ≡Prop (λ _ → isPropIsOfHLevel n) + +-- hlevels are preserved by retracts (and consequently equivalences) + +isContrRetract + : ∀ {B : Type ℓ} + → (f : A → B) (g : B → A) + → (h : retract f g) + → (v : isContr B) → isContr A +fst (isContrRetract f g h (b , p)) = g b +snd (isContrRetract f g h (b , p)) x = (cong g (p (f x))) ∙ (h x) + +isPropRetract + : {B : Type ℓ} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → isProp B → isProp A +isPropRetract f g h p x y i = + hcomp + (λ j → λ + { (i = i0) → h x j + ; (i = i1) → h y j}) + (g (p (f x) (f y) i)) + +isSetRetract + : {B : Type ℓ} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → isSet B → isSet A +isSetRetract f g h set x y p q i j = + hcomp (λ k → λ { (i = i0) → h (p j) k + ; (i = i1) → h (q j) k + ; (j = i0) → h x k + ; (j = i1) → h y k}) + (g (set (f x) (f y) + (cong f p) (cong f q) i j)) + +isGroupoidRetract + : {B : Type ℓ} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → isGroupoid B → isGroupoid A +isGroupoidRetract f g h grp x y p q P Q i j k = + hcomp ((λ l → λ { (i = i0) → h (P j k) l + ; (i = i1) → h (Q j k) l + ; (j = i0) → h (p k) l + ; (j = i1) → h (q k) l + ; (k = i0) → h x l + ; (k = i1) → h y l})) + (g (grp (f x) (f y) (cong f p) (cong f q) + (cong (cong f) P) (cong (cong f) Q) i j k)) + +is2GroupoidRetract + : {B : Type ℓ} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → is2Groupoid B → is2Groupoid A +is2GroupoidRetract f g h grp x y p q P Q R S i j k l = + hcomp (λ r → λ { (i = i0) → h (R j k l) r + ; (i = i1) → h (S j k l) r + ; (j = i0) → h (P k l) r + ; (j = i1) → h (Q k l) r + ; (k = i0) → h (p l) r + ; (k = i1) → h (q l) r + ; (l = i0) → h x r + ; (l = i1) → h y r}) + (g (grp (f x) (f y) (cong f p) (cong f q) + (cong (cong f) P) (cong (cong f) Q) + (cong (cong (cong f)) R) (cong (cong (cong f)) S) i j k l)) + +isOfHLevelRetract + : (n : HLevel) {B : Type ℓ} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → isOfHLevel n B → isOfHLevel n A +isOfHLevelRetract 0 = isContrRetract +isOfHLevelRetract 1 = isPropRetract +isOfHLevelRetract 2 = isSetRetract +isOfHLevelRetract 3 = isGroupoidRetract +isOfHLevelRetract 4 = is2GroupoidRetract +isOfHLevelRetract (suc (suc (suc (suc (suc n))))) f g h ofLevel x y p q P Q R S = + isOfHLevelRetract (suc n) (cong (cong (cong (cong f)))) + (λ s i j k l → + hcomp (λ r → λ { (i = i0) → h (R j k l) r + ; (i = i1) → h (S j k l) r + ; (j = i0) → h (P k l) r + ; (j = i1) → h (Q k l) r + ; (k = i0) → h (p l) r + ; (k = i1) → h (q l) r + ; (l = i0) → h x r + ; (l = i1) → h y r}) + (g (s i j k l))) + (λ s i j k l m → + hcomp (λ n → λ { (i = i1) → s j k l m + ; (j = i0) → h (R k l m) (i ∨ n) + ; (j = i1) → h (S k l m) (i ∨ n) + ; (k = i0) → h (P l m) (i ∨ n) + ; (k = i1) → h (Q l m) (i ∨ n) + ; (l = i0) → h (p m) (i ∨ n) + ; (l = i1) → h (q m) (i ∨ n) + ; (m = i0) → h x (i ∨ n) + ; (m = i1) → h y (i ∨ n) }) + (h (s j k l m) i)) + (ofLevel (f x) (f y) + (cong f p) (cong f q) + (cong (cong f) P) (cong (cong f) Q) + (cong (cong (cong f)) R) (cong (cong (cong f)) S)) + +isOfHLevelRetractFromIso : {A : Type ℓ} {B : Type ℓ'} (n : HLevel) → Iso A B → isOfHLevel n B → isOfHLevel n A +isOfHLevelRetractFromIso n e hlev = isOfHLevelRetract n (Iso.fun e) (Iso.inv e) (Iso.leftInv e) hlev + +isOfHLevelRespectEquiv : {A : Type ℓ} {B : Type ℓ'} → (n : HLevel) → A ≃ B → isOfHLevel n A → isOfHLevel n B +isOfHLevelRespectEquiv n eq = isOfHLevelRetract n (invEq eq) (eq .fst) (secEq eq) + +isContrRetractOfConstFun : {A : Type ℓ} {B : Type ℓ'} (b₀ : B) + → Σ[ f ∈ (B → A) ] ((x : A) → (f ∘ (λ _ → b₀)) x ≡ x) + → isContr A +fst (isContrRetractOfConstFun b₀ ret) = ret .fst b₀ +snd (isContrRetractOfConstFun b₀ ret) y = ret .snd y + +-- h-level of dependent path types + +isOfHLevelPathP' : {A : I → Type ℓ} (n : HLevel) + → isOfHLevel (suc n) (A i1) + → (x : A i0) (y : A i1) → isOfHLevel n (PathP A x y) +isOfHLevelPathP' {A = A} n h x y = + isOfHLevelRetractFromIso n (PathPIsoPath _ x y) (isOfHLevelPath' n h _ _) + +isOfHLevelPathP : {A : I → Type ℓ} (n : HLevel) + → isOfHLevel n (A i1) + → (x : A i0) (y : A i1) → isOfHLevel n (PathP A x y) +isOfHLevelPathP {A = A} n h x y = + isOfHLevelRetractFromIso n (PathPIsoPath _ x y) (isOfHLevelPath n h _ _) + +-- Fillers for cubes from h-level + +isSet→isSet' : isSet A → isSet' A +isSet→isSet' Aset _ _ _ _ = toPathP (Aset _ _ _ _) + +isSet'→isSet : isSet' A → isSet A +isSet'→isSet {A = A} Aset' x y p q = Aset' p q refl refl + +isSet→SquareP : + {A : I → I → Type ℓ} + (isSet : (i j : I) → isSet (A i j)) + {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} (a₀₋ : PathP (λ j → A i0 j) a₀₀ a₀₁) + {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} (a₁₋ : PathP (λ j → A i1 j) a₁₀ a₁₁) + (a₋₀ : PathP (λ i → A i i0) a₀₀ a₁₀) (a₋₁ : PathP (λ i → A i i1) a₀₁ a₁₁) + → SquareP A a₀₋ a₁₋ a₋₀ a₋₁ +isSet→SquareP isset a₀₋ a₁₋ a₋₀ a₋₁ = + PathPIsoPath _ _ _ .Iso.inv (isOfHLevelPathP' 1 (isset _ _) _ _ _ _ ) + +isGroupoid→isGroupoid' : isGroupoid A → isGroupoid' A +isGroupoid→isGroupoid' {A = A} Agpd a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = + PathPIsoPath (λ i → Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ .Iso.inv + (isGroupoid→isPropSquare _ _ _ _ _ _) + where + isGroupoid→isPropSquare : + {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ ≡ a₀₁) + {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ ≡ a₁₁) + (a₋₀ : a₀₀ ≡ a₁₀) (a₋₁ : a₀₁ ≡ a₁₁) + → isProp (Square a₀₋ a₁₋ a₋₀ a₋₁) + isGroupoid→isPropSquare a₀₋ a₁₋ a₋₀ a₋₁ = + isOfHLevelRetractFromIso 1 (PathPIsoPath (λ i → a₋₀ i ≡ a₋₁ i) a₀₋ a₁₋) (Agpd _ _ _ _) + +isGroupoid'→isGroupoid : isGroupoid' A → isGroupoid A +isGroupoid'→isGroupoid Agpd' x y p q r s = Agpd' r s refl refl refl refl +-- h-level of Σ-types + +isProp∃! : isProp (∃! A B) +isProp∃! = isPropIsContr + +isContrΣ : isContr A → ((x : A) → isContr (B x)) → isContr (Σ A B) +isContrΣ {A = A} {B = B} (a , p) q = + let h : (x : A) (y : B x) → (q x) .fst ≡ y + h x y = (q x) .snd y + in (( a , q a .fst) + , ( λ x i → p (x .fst) i + , h (p (x .fst) i) (transp (λ j → B (p (x .fst) (i ∨ ~ j))) i (x .snd)) i)) + +isContrΣ′ : (ca : isContr A) → isContr (B (fst ca)) → isContr (Σ A B) +isContrΣ′ ca cb = isContrΣ ca (λ x → subst _ (snd ca x) cb) + +section-Σ≡Prop + : (pB : (x : A) → isProp (B x)) {u v : Σ A B} + → section (Σ≡Prop pB {u} {v}) (cong fst) +section-Σ≡Prop {A = A} pB {u} {v} p j i = + (p i .fst) , isProp→PathP (λ i → isOfHLevelPath 1 (pB (fst (p i))) + (Σ≡Prop pB {u} {v} (cong fst p) i .snd) + (p i .snd) ) + refl refl i j + +isEquiv-Σ≡Prop + : (pB : (x : A) → isProp (B x)) {u v : Σ A B} + → isEquiv (Σ≡Prop pB {u} {v}) +isEquiv-Σ≡Prop {A = A} pB {u} {v} = isoToIsEquiv (iso (Σ≡Prop pB) (cong fst) (section-Σ≡Prop pB) (λ _ → refl)) + +isPropΣ : isProp A → ((x : A) → isProp (B x)) → isProp (Σ A B) +isPropΣ pA pB t u = Σ≡Prop pB (pA (t .fst) (u .fst)) + +isOfHLevelΣ : ∀ n → isOfHLevel n A → ((x : A) → isOfHLevel n (B x)) + → isOfHLevel n (Σ A B) +isOfHLevelΣ 0 = isContrΣ +isOfHLevelΣ 1 = isPropΣ +isOfHLevelΣ {B = B} (suc (suc n)) h1 h2 x y = + isOfHLevelRetractFromIso (suc n) + (invIso (IsoΣPathTransportPathΣ _ _)) + (isOfHLevelΣ (suc n) (h1 (fst x) (fst y)) λ x → h2 _ _ _) + +isSetΣ : isSet A → ((x : A) → isSet (B x)) → isSet (Σ A B) +isSetΣ = isOfHLevelΣ 2 + +-- Useful consequence +isSetΣSndProp : isSet A → ((x : A) → isProp (B x)) → isSet (Σ A B) +isSetΣSndProp h p = isSetΣ h (λ x → isProp→isSet (p x)) + +isGroupoidΣ : isGroupoid A → ((x : A) → isGroupoid (B x)) → isGroupoid (Σ A B) +isGroupoidΣ = isOfHLevelΣ 3 + +is2GroupoidΣ : is2Groupoid A → ((x : A) → is2Groupoid (B x)) → is2Groupoid (Σ A B) +is2GroupoidΣ = isOfHLevelΣ 4 + +-- h-level of × + +isProp× : {A : Type ℓ} {B : Type ℓ'} → isProp A → isProp B → isProp (A × B) +isProp× pA pB = isPropΣ pA (λ _ → pB) + +isProp×2 : {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → isProp A → isProp B → isProp C → isProp (A × B × C) +isProp×2 pA pB pC = isProp× pA (isProp× pB pC) + +isProp×3 : {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → isProp A → isProp B → isProp C → isProp D → isProp (A × B × C × D) +isProp×3 pA pB pC pD = isProp×2 pA pB (isProp× pC pD) + +isProp×4 : {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} {E : Type ℓ''''} + → isProp A → isProp B → isProp C → isProp D → isProp E → isProp (A × B × C × D × E) +isProp×4 pA pB pC pD pE = isProp×3 pA pB pC (isProp× pD pE) + +isProp×5 : {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} {E : Type ℓ''''} {F : Type ℓ'''''} + → isProp A → isProp B → isProp C → isProp D → isProp E → isProp F + → isProp (A × B × C × D × E × F) +isProp×5 pA pB pC pD pE pF = isProp×4 pA pB pC pD (isProp× pE pF) + + +isOfHLevel× : ∀ {A : Type ℓ} {B : Type ℓ'} n → isOfHLevel n A → isOfHLevel n B + → isOfHLevel n (A × B) +isOfHLevel× n hA hB = isOfHLevelΣ n hA (λ _ → hB) + +isSet× : ∀ {A : Type ℓ} {B : Type ℓ'} → isSet A → isSet B → isSet (A × B) +isSet× = isOfHLevel× 2 + +isGroupoid× : ∀ {A : Type ℓ} {B : Type ℓ'} → isGroupoid A → isGroupoid B + → isGroupoid (A × B) +isGroupoid× = isOfHLevel× 3 + +is2Groupoid× : ∀ {A : Type ℓ} {B : Type ℓ'} → is2Groupoid A → is2Groupoid B + → is2Groupoid (A × B) +is2Groupoid× = isOfHLevel× 4 + +-- h-level of Π-types + +isOfHLevelΠ : ∀ n → ((x : A) → isOfHLevel n (B x)) + → isOfHLevel n ((x : A) → B x) +isOfHLevelΠ 0 h = (λ x → fst (h x)) , λ f i y → snd (h y) (f y) i +isOfHLevelΠ 1 h f g i x = (h x) (f x) (g x) i +isOfHLevelΠ 2 h f g F G i j z = h z (f z) (g z) (funExt⁻ F z) (funExt⁻ G z) i j +isOfHLevelΠ 3 h f g p q P Q i j k z = + h z (f z) (g z) + (funExt⁻ p z) (funExt⁻ q z) + (cong (λ f → funExt⁻ f z) P) (cong (λ f → funExt⁻ f z) Q) i j k +isOfHLevelΠ 4 h f g p q P Q R S i j k l z = + h z (f z) (g z) + (funExt⁻ p z) (funExt⁻ q z) + (cong (λ f → funExt⁻ f z) P) (cong (λ f → funExt⁻ f z) Q) + (cong (cong (λ f → funExt⁻ f z)) R) (cong (cong (λ f → funExt⁻ f z)) S) i j k l +isOfHLevelΠ (suc (suc (suc (suc (suc n))))) h f g p q P Q R S = + isOfHLevelRetract (suc n) + (cong (cong (cong funExt⁻))) (cong (cong (cong funExt))) (λ _ → refl) + (isOfHLevelΠ (suc (suc (suc (suc n)))) (λ x → h x (f x) (g x)) + (funExt⁻ p) (funExt⁻ q) + (cong funExt⁻ P) (cong funExt⁻ Q) + (cong (cong funExt⁻) R) (cong (cong funExt⁻) S)) + +isOfHLevelΠ2 : (n : HLevel) → ((x : A) → (y : B x) → isOfHLevel n (C x y)) + → isOfHLevel n ((x : A) → (y : B x) → C x y) +isOfHLevelΠ2 n f = isOfHLevelΠ n (λ x → isOfHLevelΠ n (f x)) + +isContrΠ : (h : (x : A) → isContr (B x)) → isContr ((x : A) → B x) +isContrΠ = isOfHLevelΠ 0 + +isPropΠ : (h : (x : A) → isProp (B x)) → isProp ((x : A) → B x) +isPropΠ = isOfHLevelΠ 1 + +isPropΠ2 : (h : (x : A) (y : B x) → isProp (C x y)) + → isProp ((x : A) (y : B x) → C x y) +isPropΠ2 h = isPropΠ λ x → isPropΠ λ y → h x y + +isPropΠ3 : (h : (x : A) (y : B x) (z : C x y) → isProp (D x y z)) + → isProp ((x : A) (y : B x) (z : C x y) → D x y z) +isPropΠ3 h = isPropΠ λ x → isPropΠ λ y → isPropΠ λ z → h x y z + +isPropΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) → isProp (E x y z w)) + → isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) → E x y z w) +isPropΠ4 h = isPropΠ λ _ → isPropΠ3 (h _) + +isPropΠ5 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) → isProp (F x y z w v)) + → isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) → F x y z w v) +isPropΠ5 h = isPropΠ λ _ → isPropΠ4 (h _) + +isPropImplicitΠ : (h : (x : A) → isProp (B x)) → isProp ({x : A} → B x) +isPropImplicitΠ h f g i {x} = h x (f {x}) (g {x}) i + +isPropImplicitΠ2 : (h : (x : A) (y : B x) → isProp (C x y)) → isProp ({x : A} {y : B x} → C x y) +isPropImplicitΠ2 h = isPropImplicitΠ (λ x → isPropImplicitΠ (λ y → h x y)) + +isProp→ : {A : Type ℓ} {B : Type ℓ'} → isProp B → isProp (A → B) +isProp→ pB = isPropΠ λ _ → pB + +isSetΠ : ((x : A) → isSet (B x)) → isSet ((x : A) → B x) +isSetΠ = isOfHLevelΠ 2 + +isSetΠ2 : (h : (x : A) (y : B x) → isSet (C x y)) + → isSet ((x : A) (y : B x) → C x y) +isSetΠ2 h = isSetΠ λ x → isSetΠ λ y → h x y + +isSetΠ3 : (h : (x : A) (y : B x) (z : C x y) → isSet (D x y z)) + → isSet ((x : A) (y : B x) (z : C x y) → D x y z) +isSetΠ3 h = isSetΠ λ x → isSetΠ λ y → isSetΠ λ z → h x y z + +isGroupoidΠ : ((x : A) → isGroupoid (B x)) → isGroupoid ((x : A) → B x) +isGroupoidΠ = isOfHLevelΠ 3 + +isGroupoidΠ2 : (h : (x : A) (y : B x) → isGroupoid (C x y)) → isGroupoid ((x : A) (y : B x) → C x y) +isGroupoidΠ2 h = isGroupoidΠ λ _ → isGroupoidΠ λ _ → h _ _ + +isGroupoidΠ3 : (h : (x : A) (y : B x) (z : C x y) → isGroupoid (D x y z)) + → isGroupoid ((x : A) (y : B x) (z : C x y) → D x y z) +isGroupoidΠ3 h = isGroupoidΠ λ _ → isGroupoidΠ2 λ _ → h _ _ + +isGroupoidΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) → isGroupoid (E x y z w)) + → isGroupoid ((x : A) (y : B x) (z : C x y) (w : D x y z) → E x y z w) +isGroupoidΠ4 h = isGroupoidΠ λ _ → isGroupoidΠ3 λ _ → h _ _ + +is2GroupoidΠ : ((x : A) → is2Groupoid (B x)) → is2Groupoid ((x : A) → B x) +is2GroupoidΠ = isOfHLevelΠ 4 + +isOfHLevelΠ⁻ : ∀ {A : Type ℓ} {B : Type ℓ'} n + → isOfHLevel n (A → B) → (A → isOfHLevel n B) +isOfHLevelΠ⁻ 0 h x = fst h x , λ y → funExt⁻ (snd h (const y)) x +isOfHLevelΠ⁻ 1 h x y z = funExt⁻ (h (const y) (const z)) x +isOfHLevelΠ⁻ (suc (suc n)) h x y z = + isOfHLevelΠ⁻ (suc n) (isOfHLevelRetractFromIso (suc n) funExtIso (h _ _)) x + +-- h-level of A ≃ B and A ≡ B + +isOfHLevel≃ + : ∀ n {A : Type ℓ} {B : Type ℓ'} + → (hA : isOfHLevel n A) (hB : isOfHLevel n B) → isOfHLevel n (A ≃ B) +isOfHLevel≃ zero {A = A} {B = B} hA hB = isContr→Equiv hA hB , contr + where + contr : (y : A ≃ B) → isContr→Equiv hA hB ≡ y + contr y = Σ≡Prop isPropIsEquiv (funExt (λ a → snd hB (fst y a))) + +isOfHLevel≃ (suc n) {A = A} {B = B} hA hB = + isOfHLevelΣ (suc n) (isOfHLevelΠ _ λ _ → hB) + (λ f → isProp→isOfHLevelSuc n (isPropIsEquiv f)) + +isOfHLevel≡ : ∀ n → {A B : Type ℓ} (hA : isOfHLevel n A) (hB : isOfHLevel n B) → + isOfHLevel n (A ≡ B) +isOfHLevel≡ n hA hB = isOfHLevelRetractFromIso n univalenceIso (isOfHLevel≃ n hA hB) + +isOfHLevel⁺≃ₗ + : ∀ n {A : Type ℓ} {B : Type ℓ'} + → isOfHLevel (suc n) A → isOfHLevel (suc n) (A ≃ B) +isOfHLevel⁺≃ₗ zero pA e = isOfHLevel≃ 1 pA (isOfHLevelRespectEquiv 1 e pA) e +isOfHLevel⁺≃ₗ (suc n) hA e = isOfHLevel≃ m hA (isOfHLevelRespectEquiv m e hA) e + where + m = suc (suc n) + +isOfHLevel⁺≃ᵣ + : ∀ n {A : Type ℓ} {B : Type ℓ'} + → isOfHLevel (suc n) B → isOfHLevel (suc n) (A ≃ B) +isOfHLevel⁺≃ᵣ zero pB e + = isOfHLevel≃ 1 (isPropRetract (e .fst) (invEq e) (retEq e) pB) pB e +isOfHLevel⁺≃ᵣ (suc n) hB e + = isOfHLevel≃ m (isOfHLevelRetract m (e .fst) (invEq e) (retEq e) hB) hB e + where + m = suc (suc n) + +isOfHLevel⁺≡ₗ + : ∀ n → {A B : Type ℓ} + → isOfHLevel (suc n) A → isOfHLevel (suc n) (A ≡ B) +isOfHLevel⁺≡ₗ zero pA P = isOfHLevel≡ 1 pA (subst isProp P pA) P +isOfHLevel⁺≡ₗ (suc n) hA P + = isOfHLevel≡ m hA (subst (isOfHLevel m) P hA) P + where + m = suc (suc n) + +isOfHLevel⁺≡ᵣ + : ∀ n → {A B : Type ℓ} + → isOfHLevel (suc n) B → isOfHLevel (suc n) (A ≡ B) +isOfHLevel⁺≡ᵣ zero pB P = isOfHLevel≡ 1 (subst⁻ isProp P pB) pB P +isOfHLevel⁺≡ᵣ (suc n) hB P + = isOfHLevel≡ m (subst⁻ (isOfHLevel m) P hB) hB P + where + m = suc (suc n) + +-- h-level of TypeOfHLevel + +isPropHContr : isProp (TypeOfHLevel ℓ 0) +isPropHContr x y = Σ≡Prop (λ _ → isPropIsContr) (isOfHLevel≡ 0 (x .snd) (y .snd) .fst) + +isOfHLevelTypeOfHLevel : ∀ n → isOfHLevel (suc n) (TypeOfHLevel ℓ n) +isOfHLevelTypeOfHLevel zero = isPropHContr +isOfHLevelTypeOfHLevel (suc n) (X , a) (Y , b) = + isOfHLevelRetract (suc n) (cong fst) (Σ≡Prop λ _ → isPropIsOfHLevel (suc n)) + (section-Σ≡Prop λ _ → isPropIsOfHLevel (suc n)) + (isOfHLevel≡ (suc n) a b) + +isSetHProp : isSet (hProp ℓ) +isSetHProp = isOfHLevelTypeOfHLevel 1 + +-- h-level of lifted type + +isOfHLevelLift : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} → isOfHLevel n A → isOfHLevel n (Lift {j = ℓ'} A) +isOfHLevelLift n = isOfHLevelRetract n lower lift λ _ → refl + +---------------------------- + +-- More consequences of isProp and isContr + +inhProp→isContr : A → isProp A → isContr A +inhProp→isContr x h = x , h x + +extend : isContr A → (∀ φ → (u : Partial φ A) → Sub A φ u) +extend (x , p) φ u = inS (hcomp (λ { j (φ = i1) → p (u 1=1) j }) x) + +isContrPartial→isContr : ∀ {ℓ} {A : Type ℓ} + → (extend : ∀ φ → Partial φ A → A) + → (∀ u → u ≡ (extend i1 λ { _ → u})) + → isContr A +isContrPartial→isContr {A = A} extend law + = ex , λ y → law ex ∙ (λ i → Aux.v y i) ∙ sym (law y) + where ex = extend i0 empty + module Aux (y : A) (i : I) where + φ = ~ i ∨ i + u : Partial φ A + u = λ { (i = i0) → ex ; (i = i1) → y } + v = extend φ u + +-- Dependent h-level over a type + +isOfHLevelDep : HLevel → {A : Type ℓ} (B : A → Type ℓ') → Type (ℓ-max ℓ ℓ') +isOfHLevelDep 0 {A = A} B = {a : A} → Σ[ b ∈ B a ] ({a' : A} (b' : B a') (p : a ≡ a') → PathP (λ i → B (p i)) b b') +isOfHLevelDep 1 {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) (p : a0 ≡ a1) → PathP (λ i → B (p i)) b0 b1 +isOfHLevelDep (suc (suc n)) {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) → isOfHLevelDep (suc n) {A = a0 ≡ a1} (λ p → PathP (λ i → B (p i)) b0 b1) + +isOfHLevel→isOfHLevelDep : (n : HLevel) + → {A : Type ℓ} {B : A → Type ℓ'} (h : (a : A) → isOfHLevel n (B a)) → isOfHLevelDep n {A = A} B +isOfHLevel→isOfHLevelDep 0 h {a} = + (h a .fst , λ b' p → isProp→PathP (λ i → isContr→isProp (h (p i))) (h a .fst) b') +isOfHLevel→isOfHLevelDep 1 h = λ b0 b1 p → isProp→PathP (λ i → h (p i)) b0 b1 +isOfHLevel→isOfHLevelDep (suc (suc n)) {A = A} {B} h {a0} {a1} b0 b1 = + isOfHLevel→isOfHLevelDep (suc n) (λ p → helper p) + where + helper : (p : a0 ≡ a1) → + isOfHLevel (suc n) (PathP (λ i → B (p i)) b0 b1) + helper p = J (λ a1 p → ∀ b1 → isOfHLevel (suc n) (PathP (λ i → B (p i)) b0 b1)) + (λ _ → h _ _ _) p b1 + +isContrDep→isPropDep : isOfHLevelDep 0 B → isOfHLevelDep 1 B +isContrDep→isPropDep {B = B} Bctr {a0 = a0} b0 b1 p i + = comp (λ k → B (p (i ∧ k))) (λ k → λ where + (i = i0) → Bctr .snd b0 refl k + (i = i1) → Bctr .snd b1 p k) + (c0 .fst) + where + c0 = Bctr {a0} + +isPropDep→isSetDep : isOfHLevelDep 1 B → isOfHLevelDep 2 B +isPropDep→isSetDep {B = B} Bprp b0 b1 b2 b3 p i j + = comp (λ k → B (p (i ∧ k) (j ∧ k))) (λ k → λ where + (j = i0) → Bprp b0 b0 refl k + (i = i0) → Bprp b0 (b2 j) (λ k → p i0 (j ∧ k)) k + (i = i1) → Bprp b0 (b3 j) (λ k → p k (j ∧ k)) k + (j = i1) → Bprp b0 b1 (λ k → p (i ∧ k) (j ∧ k)) k) + b0 + +isOfHLevelDepSuc : (n : HLevel) → isOfHLevelDep n B → isOfHLevelDep (suc n) B +isOfHLevelDepSuc 0 = isContrDep→isPropDep +isOfHLevelDepSuc 1 = isPropDep→isSetDep +isOfHLevelDepSuc (suc (suc n)) Blvl b0 b1 = isOfHLevelDepSuc (suc n) (Blvl b0 b1) + +isPropDep→isSetDep' + : isOfHLevelDep 1 B + → {p : w ≡ x} {q : y ≡ z} {r : w ≡ y} {s : x ≡ z} + → {tw : B w} {tx : B x} {ty : B y} {tz : B z} + → (sq : Square p q r s) + → (tp : PathP (λ i → B (p i)) tw tx) + → (tq : PathP (λ i → B (q i)) ty tz) + → (tr : PathP (λ i → B (r i)) tw ty) + → (ts : PathP (λ i → B (s i)) tx tz) + → SquareP (λ i j → B (sq i j)) tp tq tr ts +isPropDep→isSetDep' {B = B} Bprp {p} {q} {r} {s} {tw} sq tp tq tr ts i j + = comp (λ k → B (sq (i ∧ k) (j ∧ k))) (λ k → λ where + (i = i0) → Bprp tw (tp j) (λ k → p (k ∧ j)) k + (i = i1) → Bprp tw (tq j) (λ k → sq (i ∧ k) (j ∧ k)) k + (j = i0) → Bprp tw (tr i) (λ k → r (k ∧ i)) k + (j = i1) → Bprp tw (ts i) (λ k → sq (k ∧ i) (j ∧ k)) k) + tw + +isOfHLevelΣ' : ∀ n → isOfHLevel n A → isOfHLevelDep n B → isOfHLevel n (Σ A B) +isOfHLevelΣ' 0 Actr Bctr .fst = (Actr .fst , Bctr .fst) +isOfHLevelΣ' 0 Actr Bctr .snd (x , y) i + = Actr .snd x i , Bctr .snd y (Actr .snd x) i +isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .fst = Alvl w x i +isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .snd = Blvl y z (Alvl w x) i +isOfHLevelΣ' {A = A} {B = B} (suc (suc n)) Alvl Blvl (w , y) (x , z) + = isOfHLevelRetract (suc n) + (λ p → (λ i → p i .fst) , λ i → p i .snd) + ΣPathP + (λ x → refl) + (isOfHLevelΣ' (suc n) (Alvl w x) (Blvl y z)) diff --git a/Cubical/Foundations/Id.agda b/Cubical/Foundations/Id.agda new file mode 100644 index 0000000000..505b0f0c81 --- /dev/null +++ b/Cubical/Foundations/Id.agda @@ -0,0 +1,313 @@ +{- + +This file contains: + +- Id, refl and J (with definitional computation rule) + +- Basic theory about Id, proved using J + +- Lemmas for going back and forth between Path and Id + +- Function extensionality for Id + +- fiber, isContr, equiv all defined using Id + +- The univalence axiom expressed using only Id ([EquivContr]) + +- Propositional truncation and its elimination principle + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Id where + +open import Cubical.Foundations.Prelude public + hiding ( _≡_ ; _≡⟨_⟩_ ; _∎ ; isPropIsContr) + renaming ( refl to reflPath + ; transport to transportPath + ; J to JPath + ; JRefl to JPathRefl + ; sym to symPath + ; _∙_ to compPath + ; cong to congPath + ; funExt to funExtPath + ; isContr to isContrPath + ; isProp to isPropPath + ; isSet to isSetPath + ; fst to pr₁ -- as in the HoTT book + ; snd to pr₂ ) + +open import Cubical.Foundations.Equiv + renaming ( fiber to fiberPath + ; isEquiv to isEquivPath + ; _≃_ to EquivPath + ; equivFun to equivFunPath + ; isPropIsEquiv to isPropIsEquivPath ) + hiding ( equivCtr + ; equivIsEquiv ) + +open import Cubical.Foundations.Univalence + renaming ( EquivContr to EquivContrPath ) +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.PropositionalTruncation public + renaming ( squash to squashPath + ; rec to recPropTruncPath + ; elim to elimPropTruncPath ) +open import Cubical.Core.Id public + using (Id; ⟨_,_⟩; faceId; pathId; elimId; _≡_) + +private + variable + ℓ ℓ' : Level + A : Type ℓ + +-- Version of the constructor for Id where the y is also +-- explicit. This is sometimes useful when it is needed for +-- typechecking (see JId below). +conId : ∀ {x : A} φ (y : A [ φ ↦ (λ _ → x) ]) + (w : (Path _ x (outS y)) [ φ ↦ (λ { (φ = i1) → λ _ → x}) ]) → + x ≡ outS y +conId φ _ w = ⟨ φ , outS w ⟩ + +-- Reflexivity +refl : ∀ {x : A} → x ≡ x +refl {x = x} = ⟨ i1 , (λ _ → x) ⟩ + + +-- Definition of J for Id +module _ {x : A} (P : ∀ (y : A) → Id x y → Type ℓ') (d : P x refl) where + J : ∀ {y : A} (w : x ≡ y) → P y w + J {y = y} = elimId P (λ φ y w → comp (λ i → P _ (conId (φ ∨ ~ i) (inS (outS w i)) + (inS (λ j → outS w (i ∧ j))))) + (λ i → λ { (φ = i1) → d}) d) {y = y} + + -- Check that J of refl is the identity function + Jdefeq : Path _ (J refl) d + Jdefeq _ = d + + +-- Basic theory about Id, proved using J +transport : ∀ (B : A → Type ℓ') {x y : A} + → x ≡ y → B x → B y +transport B {x} p b = J (λ y p → B y) b p + +_⁻¹ : {x y : A} → x ≡ y → y ≡ x +_⁻¹ {x = x} p = J (λ z _ → z ≡ x) refl p + +ap : ∀ {B : Type ℓ'} (f : A → B) → ∀ {x y : A} → x ≡ y → f x ≡ f y +ap f {x} = J (λ z _ → f x ≡ f z) refl + +_∙_ : ∀ {x y z : A} → x ≡ y → y ≡ z → x ≡ z +_∙_ {x = x} p = J (λ y _ → x ≡ y) p + +infix 4 _∙_ +infix 3 _∎ +infixr 2 _≡⟨_⟩_ + +_≡⟨_⟩_ : (x : A) {y z : A} → x ≡ y → y ≡ z → x ≡ z +_ ≡⟨ p ⟩ q = p ∙ q + +_∎ : (x : A) → x ≡ x +_ ∎ = refl + +-- Convert between Path and Id +pathToId : ∀ {x y : A} → Path _ x y → Id x y +pathToId {x = x} = JPath (λ y _ → Id x y) refl + +pathToIdRefl : ∀ {x : A} → Path _ (pathToId (λ _ → x)) refl +pathToIdRefl {x = x} = JPathRefl (λ y _ → Id x y) refl + +idToPath : ∀ {x y : A} → Id x y → Path _ x y +idToPath {x = x} = J (λ y _ → Path _ x y) (λ _ → x) + +idToPathRefl : ∀ {x : A} → Path _ (idToPath {x = x} refl) reflPath +idToPathRefl {x = x} _ _ = x + +pathToIdToPath : ∀ {x y : A} → (p : Path _ x y) → Path _ (idToPath (pathToId p)) p +pathToIdToPath {x = x} = JPath (λ y p → Path _ (idToPath (pathToId p)) p) + (λ i → idToPath (pathToIdRefl i)) + +idToPathToId : ∀ {x y : A} → (p : Id x y) → Path _ (pathToId (idToPath p)) p +idToPathToId {x = x} = J (λ b p → Path _ (pathToId (idToPath p)) p) pathToIdRefl + + +-- We get function extensionality by going back and forth between Path and Id +funExt : ∀ {B : A → Type ℓ'} {f g : (x : A) → B x} → + ((x : A) → f x ≡ g x) → f ≡ g +funExt p = pathToId (λ i x → idToPath (p x) i) + + +-- Equivalences expressed using Id + +fiber : ∀ {A : Type ℓ} {B : Type ℓ'} (f : A → B) (y : B) → Type (ℓ-max ℓ ℓ') +fiber {A = A} f y = Σ[ x ∈ A ] f x ≡ y + +isContr : Type ℓ → Type ℓ +isContr A = Σ[ x ∈ A ] (∀ y → x ≡ y) + +isProp : Type ℓ → Type ℓ +isProp A = (x y : A) → x ≡ y + +isSet : Type ℓ → Type ℓ +isSet A = (x y : A) → isProp (x ≡ y) + +record isEquiv {A : Type ℓ} {B : Type ℓ'} (f : A → B) : Type (ℓ-max ℓ ℓ') where + field + equiv-proof : (y : B) → isContr (fiber f y) + +open isEquiv public + +infix 4 _≃_ + +_≃_ : ∀ (A : Type ℓ) (B : Type ℓ') → Type (ℓ-max ℓ ℓ') +A ≃ B = Σ[ f ∈ (A → B) ] (isEquiv f) + +equivFun : ∀ {B : Type ℓ'} → A ≃ B → A → B +equivFun e = pr₁ e + +equivIsEquiv : ∀ {B : Type ℓ'} (e : A ≃ B) → isEquiv (equivFun e) +equivIsEquiv e = pr₂ e + +equivCtr : ∀ {B : Type ℓ'} (e : A ≃ B) (y : B) → fiber (equivFun e) y +equivCtr e y = e .pr₂ .equiv-proof y .pr₁ + + +-- Functions for going between the various definitions. This could +-- also be achieved by making lines in the universe and transporting +-- back and forth along them. + +fiberPathToFiber : ∀ {B : Type ℓ'} {f : A → B} {y : B} → + fiberPath f y → fiber f y +fiberPathToFiber (x , p) = (x , pathToId p) + +fiberToFiberPath : ∀ {B : Type ℓ'} {f : A → B} {y : B} → + fiber f y → fiberPath f y +fiberToFiberPath (x , p) = (x , idToPath p) + +fiberToFiber : ∀ {B : Type ℓ'} {f : A → B} {y : B} + (p : fiber f y) → Path _ (fiberPathToFiber (fiberToFiberPath p)) p +fiberToFiber (x , p) = λ i → x , idToPathToId p i + +fiberPathToFiberPath : ∀ {B : Type ℓ'} {f : A → B} {y : B} + (p : fiberPath f y) → Path _ (fiberToFiberPath (fiberPathToFiber p)) p +fiberPathToFiberPath (x , p) = λ i → x , pathToIdToPath p i + +isContrPathToIsContr : isContrPath A → isContr A +isContrPathToIsContr (ctr , p) = (ctr , λ y → pathToId (p y)) + +isContrToIsContrPath : isContr A → isContrPath A +isContrToIsContrPath (ctr , p) = (ctr , λ y → idToPath (p y)) + +isPropPathToIsProp : isPropPath A → isProp A +isPropPathToIsProp H x y = pathToId (H x y) + +isPropToIsPropPath : isProp A → isPropPath A +isPropToIsPropPath H x y i = idToPath (H x y) i + +-- Specialized helper lemmas for going back and forth between +-- isContrPath and isContr: + +helper1 : ∀ {A B : Type ℓ} (f : A → B) (g : B → A) + (h : (y : B) → Path B (f (g y)) y) → isContrPath A → isContr B +helper1 f g h (x , p) = + (f x , λ y → pathToId (λ i → hcomp (λ j → λ { (i = i0) → f x + ; (i = i1) → h y j }) + (f (p (g y) i)))) + +helper2 : ∀ {A B : Type ℓ} (f : A → B) (g : B → A) + (h : (y : A) → Path A (g (f y)) y) → isContr B → isContrPath A +helper2 {A = A} f g h (x , p) = (g x , λ y → idToPath (rem y)) + where + rem : ∀ (y : A) → g x ≡ y + rem y = + g x ≡⟨ ap g (p (f y)) ⟩ + g (f y) ≡⟨ pathToId (h y) ⟩ + y ∎ + +-- This proof is essentially the one for proving that isContr with +-- Path is a proposition, but as we are working with Id we have to +-- insert a lof of conversion functions. It is still nice that is +-- works like this though. +isPropIsContr : ∀ (p1 p2 : isContr A) → Path (isContr A) p1 p2 +isPropIsContr (a0 , p0) (a1 , p1) j = + ( idToPath (p0 a1) j , + hcomp (λ i → λ { (j = i0) → λ x → idToPathToId (p0 x) i + ; (j = i1) → λ x → idToPathToId (p1 x) i }) + (λ x → pathToId (λ i → hcomp (λ k → λ { (i = i0) → idToPath (p0 a1) j + ; (i = i1) → idToPath (p0 x) (j ∨ k) + ; (j = i0) → idToPath (p0 x) (i ∧ k) + ; (j = i1) → idToPath (p1 x) i }) + (idToPath (p0 (idToPath (p1 x) i)) j)))) + +-- We now prove that isEquiv is a proposition +isPropIsEquiv : ∀ {A : Type ℓ} {B : Type ℓ} → {f : A → B} → (h1 h2 : isEquiv f) → Path _ h1 h2 +equiv-proof (isPropIsEquiv {f = f} h1 h2 i) y = + isPropIsContr {A = fiber f y} (h1 .equiv-proof y) (h2 .equiv-proof y) i + +-- Go from a Path equivalence to an Id equivalence +equivPathToEquiv : ∀ {A : Type ℓ} {B : Type ℓ'} → EquivPath A B → A ≃ B +equivPathToEquiv (f , p) = + (f , λ { .equiv-proof y → helper1 fiberPathToFiber fiberToFiberPath fiberToFiber (p .equiv-proof y) }) + +-- Go from an Id equivalence to a Path equivalence +equivToEquivPath : ∀ {A : Type ℓ} {B : Type ℓ'} → A ≃ B → EquivPath A B +equivToEquivPath (f , p) = + (f , λ { .equiv-proof y → helper2 fiberPathToFiber fiberToFiberPath fiberPathToFiberPath (p .equiv-proof y) }) + +equivToEquiv : ∀ {A : Type ℓ} {B : Type ℓ} → (p : A ≃ B) → Path _ (equivPathToEquiv (equivToEquivPath p)) p +equivToEquiv (f , p) i = + (f , isPropIsEquiv (λ { .equiv-proof y → helper1 fiberPathToFiber fiberToFiberPath fiberToFiber + (helper2 fiberPathToFiber fiberToFiberPath fiberPathToFiberPath (p .equiv-proof y)) }) p i) + + +-- We can finally prove univalence with Id everywhere from the one for Path +EquivContr : ∀ (A : Type ℓ) → isContr (Σ[ T ∈ Type ℓ ] (T ≃ A)) +EquivContr {ℓ = ℓ} A = helper1 f1 f2 f12 (EquivContrPath A) + where + f1 : {A : Type ℓ} → Σ[ T ∈ Type ℓ ] (EquivPath T A) → Σ[ T ∈ Type ℓ ] (T ≃ A) + f1 (x , p) = x , equivPathToEquiv p + + f2 : {A : Type ℓ} → Σ[ T ∈ Type ℓ ] (T ≃ A) → Σ[ T ∈ Type ℓ ] (EquivPath T A) + f2 (x , p) = x , equivToEquivPath p + + f12 : (y : Σ[ T ∈ Type ℓ ] (T ≃ A)) → Path (Σ[ T ∈ Type ℓ ] (T ≃ A)) (f1 (f2 y)) y + f12 (x , p) i = x , equivToEquiv {A = x} {B = A} p i + +-- Propositional truncation + +∥∥-isProp : ∀ (x y : ∥ A ∥) → x ≡ y +∥∥-isProp x y = pathToId (squashPath x y) + +∥∥-recursion : ∀ {A : Type ℓ} {P : Type ℓ} → isProp P → (A → P) → ∥ A ∥ → P +∥∥-recursion Pprop f x = recPropTruncPath (isPropToIsPropPath Pprop) f x + +∥∥-induction : ∀ {A : Type ℓ} {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → + ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a +∥∥-induction Pprop f x = elimPropTruncPath (λ a → isPropToIsPropPath (Pprop a)) f x + + +-- Univalence + +path≡Id : ∀ {ℓ} {A B : Type ℓ} → Path _ (Path _ A B) (Id A B) +path≡Id = isoToPath (iso pathToId idToPath idToPathToId pathToIdToPath ) + +equivPathToEquivPath : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} → (p : EquivPath A B) → + Path _ (equivToEquivPath (equivPathToEquiv p)) p +equivPathToEquivPath (f , p) i = + ( f , isPropIsEquivPath f (equivToEquivPath (equivPathToEquiv (f , p)) .pr₂) p i ) + +equivPath≡Equiv : ∀ {ℓ} {A B : Type ℓ} → Path _ (EquivPath A B) (A ≃ B) +equivPath≡Equiv {ℓ} = isoToPath (iso (equivPathToEquiv {ℓ}) equivToEquivPath equivToEquiv equivPathToEquivPath) + +univalenceId : ∀ {ℓ} {A B : Type ℓ} → (A ≡ B) ≃ (A ≃ B) +univalenceId {ℓ} {A = A} {B = B} = equivPathToEquiv rem + where + rem0 : Path _ (Lift (EquivPath A B)) (Lift (A ≃ B)) + rem0 = congPath Lift equivPath≡Equiv + + rem1 : Path _ (Id A B) (Lift (A ≃ B)) + rem1 i = hcomp (λ j → λ { (i = i0) → path≡Id {A = A} {B = B} j + ; (i = i1) → rem0 j }) + (univalencePath {A = A} {B = B} i) + + rem : EquivPath (Id A B) (A ≃ B) + rem = compEquiv (eqweqmap rem1) (invEquiv LiftEquiv) diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda new file mode 100644 index 0000000000..a6a9d874f0 --- /dev/null +++ b/Cubical/Foundations/Isomorphism.agda @@ -0,0 +1,185 @@ +{- + +Theory about isomorphisms + +- Definitions of [section] and [retract] +- Definition of isomorphisms ([Iso]) +- Any isomorphism is an equivalence ([isoToEquiv]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Isomorphism where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.Base + +open import Cubical.Foundations.Function + +private + variable + ℓ ℓ' : Level + A B C : Type ℓ + +-- Section and retract +module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where + section : (f : A → B) → (g : B → A) → Type ℓ' + section f g = ∀ b → f (g b) ≡ b + + -- NB: `g` is the retraction! + retract : (f : A → B) → (g : B → A) → Type ℓ + retract f g = ∀ a → g (f a) ≡ a + +record Iso {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality + constructor iso + field + fun : A → B + inv : B → A + rightInv : section fun inv + leftInv : retract fun inv + +isIso : (A → B) → Type _ +isIso {A = A} {B = B} f = Σ[ g ∈ (B → A) ] Σ[ _ ∈ section f g ] retract f g + +isoFunInjective : (f : Iso A B) → (x y : A) → Iso.fun f x ≡ Iso.fun f y → x ≡ y +isoFunInjective f x y h = sym (Iso.leftInv f x) ∙∙ cong (Iso.inv f) h ∙∙ Iso.leftInv f y + +isoInvInjective : (f : Iso A B) → (x y : B) → Iso.inv f x ≡ Iso.inv f y → x ≡ y +isoInvInjective f x y h = sym (Iso.rightInv f x) ∙∙ cong (Iso.fun f) h ∙∙ Iso.rightInv f y + +-- Any iso is an equivalence +module _ (i : Iso A B) where + open Iso i renaming ( fun to f + ; inv to g + ; rightInv to s + ; leftInv to t) + + private + module _ (y : B) (x0 x1 : A) (p0 : f x0 ≡ y) (p1 : f x1 ≡ y) where + fill0 : I → I → A + fill0 i = hfill (λ k → λ { (i = i1) → t x0 k + ; (i = i0) → g y }) + (inS (g (p0 (~ i)))) + + fill1 : I → I → A + fill1 i = hfill (λ k → λ { (i = i1) → t x1 k + ; (i = i0) → g y }) + (inS (g (p1 (~ i)))) + + fill2 : I → I → A + fill2 i = hfill (λ k → λ { (i = i1) → fill1 k i1 + ; (i = i0) → fill0 k i1 }) + (inS (g y)) + + p : x0 ≡ x1 + p i = fill2 i i1 + + sq : I → I → A + sq i j = hcomp (λ k → λ { (i = i1) → fill1 j (~ k) + ; (i = i0) → fill0 j (~ k) + ; (j = i1) → t (fill2 i i1) (~ k) + ; (j = i0) → g y }) + (fill2 i j) + + sq1 : I → I → B + sq1 i j = hcomp (λ k → λ { (i = i1) → s (p1 (~ j)) k + ; (i = i0) → s (p0 (~ j)) k + ; (j = i1) → s (f (p i)) k + ; (j = i0) → s y k }) + (f (sq i j)) + + lemIso : (x0 , p0) ≡ (x1 , p1) + lemIso i .fst = p i + lemIso i .snd = λ j → sq1 i (~ j) + + isoToIsEquiv : isEquiv f + isoToIsEquiv .equiv-proof y .fst .fst = g y + isoToIsEquiv .equiv-proof y .fst .snd = s y + isoToIsEquiv .equiv-proof y .snd z = lemIso y (g y) (fst z) (s y) (snd z) + + +isoToEquiv : Iso A B → A ≃ B +isoToEquiv i .fst = _ +isoToEquiv i .snd = isoToIsEquiv i + +isoToPath : Iso A B → A ≡ B +isoToPath {A = A} {B = B} f i = + Glue B (λ { (i = i0) → (A , isoToEquiv f) + ; (i = i1) → (B , idEquiv B) }) + +open Iso + +invIso : Iso A B → Iso B A +fun (invIso f) = inv f +inv (invIso f) = fun f +rightInv (invIso f) = leftInv f +leftInv (invIso f) = rightInv f + +compIso : Iso A B → Iso B C → Iso A C +fun (compIso i j) = fun j ∘ fun i +inv (compIso i j) = inv i ∘ inv j +rightInv (compIso i j) b = cong (fun j) (rightInv i (inv j b)) ∙ rightInv j b +leftInv (compIso i j) a = cong (inv i) (leftInv j (fun i a)) ∙ leftInv i a + +composesToId→Iso : (G : Iso A B) (g : B → A) → G .fun ∘ g ≡ idfun B → Iso B A +fun (composesToId→Iso _ g _) = g +inv (composesToId→Iso j _ _) = fun j +rightInv (composesToId→Iso i g path) b = + sym (leftInv i (g (fun i b))) ∙∙ cong (λ g → inv i (g (fun i b))) path ∙∙ leftInv i b +leftInv (composesToId→Iso _ _ path) b i = path i b + +idIso : Iso A A +fun idIso = idfun _ +inv idIso = idfun _ +rightInv idIso _ = refl +leftInv idIso _ = refl + +LiftIso : Iso A (Lift {i = ℓ} {j = ℓ'} A) +fun LiftIso = lift +inv LiftIso = lower +rightInv LiftIso _ = refl +leftInv LiftIso _ = refl + +isContr→Iso : isContr A → isContr B → Iso A B +fun (isContr→Iso _ Bctr) _ = Bctr .fst +inv (isContr→Iso Actr _) _ = Actr .fst +rightInv (isContr→Iso _ Bctr) = Bctr .snd +leftInv (isContr→Iso Actr _) = Actr .snd + +isProp→Iso : (Aprop : isProp A) (Bprop : isProp B) (f : A → B) (g : B → A) → Iso A B +fun (isProp→Iso _ _ f _) = f +inv (isProp→Iso _ _ _ g) = g +rightInv (isProp→Iso _ Bprop f g) b = Bprop (f (g b)) b +leftInv (isProp→Iso Aprop _ f g) a = Aprop (g (f a)) a + +domIso : ∀ {ℓ} {C : Type ℓ} → Iso A B → Iso (A → C) (B → C) +fun (domIso e) f b = f (inv e b) +inv (domIso e) f a = f (fun e a) +rightInv (domIso e) f i x = f (rightInv e x i) +leftInv (domIso e) f i x = f (leftInv e x i) + +-- Helpful notation +_Iso⟨_⟩_ : ∀ {ℓ ℓ' ℓ''} {B : Type ℓ'} {C : Type ℓ''} (X : Type ℓ) → Iso X B → Iso B C → Iso X C +_ Iso⟨ f ⟩ g = compIso f g + +_∎Iso : ∀ {ℓ} (A : Type ℓ) → Iso A A +A ∎Iso = idIso {A = A} + +infixr 0 _Iso⟨_⟩_ +infix 1 _∎Iso + +codomainIsoDep : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} + → ((a : A) → Iso (B a) (C a)) + → Iso ((a : A) → B a) ((a : A) → C a) +Iso.fun (codomainIsoDep is) f a = Iso.fun (is a) (f a) +Iso.inv (codomainIsoDep is) f a = Iso.inv (is a) (f a) +Iso.rightInv (codomainIsoDep is) f = funExt λ a → Iso.rightInv (is a) (f a) +Iso.leftInv (codomainIsoDep is) f = funExt λ a → Iso.leftInv (is a) (f a) + +codomainIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → Iso B C + → Iso (A → B) (A → C) +codomainIso z = codomainIsoDep λ _ → z diff --git a/Cubical/Foundations/Path.agda b/Cubical/Foundations/Path.agda new file mode 100644 index 0000000000..399b223705 --- /dev/null +++ b/Cubical/Foundations/Path.agda @@ -0,0 +1,362 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Path where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence + +open import Cubical.Reflection.StrictEquiv + +private + variable + ℓ ℓ' : Level + A : Type ℓ + +-- Less polymorphic version of `cong`, to avoid some unresolved metas +cong′ : ∀ {B : Type ℓ'} (f : A → B) {x y : A} (p : x ≡ y) + → Path B (f x) (f y) +cong′ f = cong f +{-# INLINE cong′ #-} + +PathP≡Path : ∀ (P : I → Type ℓ) (p : P i0) (q : P i1) → + PathP P p q ≡ Path (P i1) (transport (λ i → P i) p) q +PathP≡Path P p q i = PathP (λ j → P (i ∨ j)) (transport-filler (λ j → P j) p i) q + +PathP≡Path⁻ : ∀ (P : I → Type ℓ) (p : P i0) (q : P i1) → + PathP P p q ≡ Path (P i0) p (transport⁻ (λ i → P i) q) +PathP≡Path⁻ P p q i = PathP (λ j → P (~ i ∧ j)) p (transport⁻-filler (λ j → P j) q i) + +PathPIsoPath : ∀ (A : I → Type ℓ) (x : A i0) (y : A i1) → Iso (PathP A x y) (transport (λ i → A i) x ≡ y) +PathPIsoPath A x y .Iso.fun = fromPathP +PathPIsoPath A x y .Iso.inv = toPathP +PathPIsoPath A x y .Iso.rightInv q k i = + hcomp + (λ j → λ + { (i = i0) → slide (j ∨ ~ k) + ; (i = i1) → q j + ; (k = i0) → transp (λ l → A (i ∨ l)) i (fromPathPFiller j) + ; (k = i1) → ∧∨Square i j + }) + (transp (λ l → A (i ∨ ~ k ∨ l)) (i ∨ ~ k) + (transp (λ l → (A (i ∨ (~ k ∧ l)))) (k ∨ i) + (transp (λ l → A (i ∧ l)) (~ i) + x))) + where + fromPathPFiller : _ + fromPathPFiller = + hfill + (λ j → λ + { (i = i0) → x + ; (i = i1) → q j }) + (inS (transp (λ j → A (i ∧ j)) (~ i) x)) + + slide : I → _ + slide i = transp (λ l → A (i ∨ l)) i (transp (λ l → A (i ∧ l)) (~ i) x) + + ∧∨Square : I → I → _ + ∧∨Square i j = + hcomp + (λ l → λ + { (i = i0) → slide j + ; (i = i1) → q (j ∧ l) + ; (j = i0) → slide i + ; (j = i1) → q (i ∧ l) + }) + (slide (i ∨ j)) +PathPIsoPath A x y .Iso.leftInv q k i = + outS + (hcomp-unique + (λ j → λ + { (i = i0) → x + ; (i = i1) → transp (λ l → A (j ∨ l)) j (q j) + }) + (inS (transp (λ l → A (i ∧ l)) (~ i) x)) + (λ j → inS (transp (λ l → A (i ∧ (j ∨ l))) (~ i ∨ j) (q (i ∧ j))))) + k + +PathP≃Path : (A : I → Type ℓ) (x : A i0) (y : A i1) → + PathP A x y ≃ (transport (λ i → A i) x ≡ y) +PathP≃Path A x y = isoToEquiv (PathPIsoPath A x y) + +PathP≡compPath : ∀ {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) (r : x ≡ z) + → (PathP (λ i → x ≡ q i) p r) ≡ (p ∙ q ≡ r) +PathP≡compPath p q r k = PathP (λ i → p i0 ≡ q (i ∨ k)) (λ j → compPath-filler p q k j) r + +-- a quick corollary for 3-constant functions +3-ConstantCompChar : {A : Type ℓ} {B : Type ℓ'} (f : A → B) (link : 2-Constant f) + → (∀ x y z → link x y ∙ link y z ≡ link x z) + → 3-Constant f +3-Constant.link (3-ConstantCompChar f link coh₂) = link +3-Constant.coh₁ (3-ConstantCompChar f link coh₂) _ _ _ = + transport⁻ (PathP≡compPath _ _ _) (coh₂ _ _ _) + +PathP≡doubleCompPathˡ : ∀ {A : Type ℓ} {w x y z : A} (p : w ≡ y) (q : w ≡ x) (r : y ≡ z) (s : x ≡ z) + → (PathP (λ i → p i ≡ s i) q r) ≡ (p ⁻¹ ∙∙ q ∙∙ s ≡ r) +PathP≡doubleCompPathˡ p q r s k = PathP (λ i → p (i ∨ k) ≡ s (i ∨ k)) + (λ j → doubleCompPath-filler (p ⁻¹) q s k j) r + +PathP≡doubleCompPathʳ : ∀ {A : Type ℓ} {w x y z : A} (p : w ≡ y) (q : w ≡ x) (r : y ≡ z) (s : x ≡ z) + → (PathP (λ i → p i ≡ s i) q r) ≡ (q ≡ p ∙∙ r ∙∙ s ⁻¹) +PathP≡doubleCompPathʳ p q r s k = PathP (λ i → p (i ∧ (~ k)) ≡ s (i ∧ (~ k))) + q (λ j → doubleCompPath-filler p r (s ⁻¹) k j) + +compPathl-cancel : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : x ≡ z) → p ∙ (sym p ∙ q) ≡ q +compPathl-cancel p q = p ∙ (sym p ∙ q) ≡⟨ assoc p (sym p) q ⟩ + (p ∙ sym p) ∙ q ≡⟨ cong (_∙ q) (rCancel p) ⟩ + refl ∙ q ≡⟨ sym (lUnit q) ⟩ + q ∎ + +compPathr-cancel : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : z ≡ y) (q : x ≡ y) → (q ∙ sym p) ∙ p ≡ q +compPathr-cancel {x = x} p q i j = + hcomp-equivFiller (doubleComp-faces (λ _ → x) (sym p) j) (inS (q j)) (~ i) + +compPathl-isEquiv : {x y z : A} (p : x ≡ y) → isEquiv (λ (q : y ≡ z) → p ∙ q) +compPathl-isEquiv p = isoToIsEquiv (iso (p ∙_) (sym p ∙_) (compPathl-cancel p) (compPathl-cancel (sym p))) + +compPathlEquiv : {x y z : A} (p : x ≡ y) → (y ≡ z) ≃ (x ≡ z) +compPathlEquiv p = (p ∙_) , compPathl-isEquiv p + +compPathr-isEquiv : {x y z : A} (p : y ≡ z) → isEquiv (λ (q : x ≡ y) → q ∙ p) +compPathr-isEquiv p = isoToIsEquiv (iso (_∙ p) (_∙ sym p) (compPathr-cancel p) (compPathr-cancel (sym p))) + +compPathrEquiv : {x y z : A} (p : y ≡ z) → (x ≡ y) ≃ (x ≡ z) +compPathrEquiv p = (_∙ p) , compPathr-isEquiv p + +-- Variations of isProp→isSet for PathP +isProp→SquareP : ∀ {B : I → I → Type ℓ} → ((i j : I) → isProp (B i j)) + → {a : B i0 i0} {b : B i0 i1} {c : B i1 i0} {d : B i1 i1} + → (r : PathP (λ j → B j i0) a c) (s : PathP (λ j → B j i1) b d) + → (t : PathP (λ j → B i0 j) a b) (u : PathP (λ j → B i1 j) c d) + → SquareP B t u r s +isProp→SquareP {B = B} isPropB {a = a} r s t u i j = + hcomp (λ { k (i = i0) → isPropB i0 j (base i0 j) (t j) k + ; k (i = i1) → isPropB i1 j (base i1 j) (u j) k + ; k (j = i0) → isPropB i i0 (base i i0) (r i) k + ; k (j = i1) → isPropB i i1 (base i i1) (s i) k + }) (base i j) where + base : (i j : I) → B i j + base i j = transport (λ k → B (i ∧ k) (j ∧ k)) a + +isProp→isPropPathP : ∀ {ℓ} {B : I → Type ℓ} → ((i : I) → isProp (B i)) + → (b0 : B i0) (b1 : B i1) + → isProp (PathP (λ i → B i) b0 b1) +isProp→isPropPathP {B = B} hB b0 b1 = isProp→SquareP (λ _ → hB) refl refl + +isProp→isContrPathP : {A : I → Type ℓ} → (∀ i → isProp (A i)) + → (x : A i0) (y : A i1) + → isContr (PathP A x y) +isProp→isContrPathP h x y = isProp→PathP h x y , isProp→isPropPathP h x y _ + +-- Flipping a square along its diagonal + +flipSquare : {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ ≡ a₀₁} + {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ ≡ a₁₁} + {a₋₀ : a₀₀ ≡ a₁₀} {a₋₁ : a₀₁ ≡ a₁₁} + → Square a₀₋ a₁₋ a₋₀ a₋₁ → Square a₋₀ a₋₁ a₀₋ a₁₋ +flipSquare sq i j = sq j i + +module _ {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ ≡ a₀₁} {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ ≡ a₁₁} + {a₋₀ : a₀₀ ≡ a₁₀} {a₋₁ : a₀₁ ≡ a₁₁} + where + + flipSquareEquiv : Square a₀₋ a₁₋ a₋₀ a₋₁ ≃ Square a₋₀ a₋₁ a₀₋ a₁₋ + unquoteDef flipSquareEquiv = defStrictEquiv flipSquareEquiv flipSquare flipSquare + + flipSquarePath : Square a₀₋ a₁₋ a₋₀ a₋₁ ≡ Square a₋₀ a₋₁ a₀₋ a₁₋ + flipSquarePath = ua flipSquareEquiv + +module _ {a₀₀ a₁₁ : A} {a₋ : a₀₀ ≡ a₁₁} + {a₁₀ : A} {a₁₋ : a₁₀ ≡ a₁₁} {a₋₀ : a₀₀ ≡ a₁₀} where + + slideSquareFaces : (i j k : I) → Partial (i ∨ ~ i ∨ j ∨ ~ j) A + slideSquareFaces i j k (i = i0) = a₋ (j ∧ ~ k) + slideSquareFaces i j k (i = i1) = a₁₋ j + slideSquareFaces i j k (j = i0) = a₋₀ i + slideSquareFaces i j k (j = i1) = a₋ (i ∨ ~ k) + + slideSquare : Square a₋ a₁₋ a₋₀ refl → Square refl a₁₋ a₋₀ a₋ + slideSquare sq i j = hcomp (slideSquareFaces i j) (sq i j) + + slideSquareEquiv : (Square a₋ a₁₋ a₋₀ refl) ≃ (Square refl a₁₋ a₋₀ a₋) + slideSquareEquiv = isoToEquiv (iso slideSquare slideSquareInv fillerTo fillerFrom) where + slideSquareInv : Square refl a₁₋ a₋₀ a₋ → Square a₋ a₁₋ a₋₀ refl + slideSquareInv sq i j = hcomp (λ k → slideSquareFaces i j (~ k)) (sq i j) + fillerTo : ∀ p → slideSquare (slideSquareInv p) ≡ p + fillerTo p k i j = hcomp-equivFiller (λ k → slideSquareFaces i j (~ k)) (inS (p i j)) (~ k) + fillerFrom : ∀ p → slideSquareInv (slideSquare p) ≡ p + fillerFrom p k i j = hcomp-equivFiller (slideSquareFaces i j) (inS (p i j)) (~ k) + +-- The type of fillers of a square is equivalent to the double composition identites +Square≃doubleComp : {a₀₀ a₀₁ a₁₀ a₁₁ : A} + (a₀₋ : a₀₀ ≡ a₀₁) (a₁₋ : a₁₀ ≡ a₁₁) + (a₋₀ : a₀₀ ≡ a₁₀) (a₋₁ : a₀₁ ≡ a₁₁) + → Square a₀₋ a₁₋ a₋₀ a₋₁ ≃ (a₋₀ ⁻¹ ∙∙ a₀₋ ∙∙ a₋₁ ≡ a₁₋) +Square≃doubleComp a₀₋ a₁₋ a₋₀ a₋₁ = transportEquiv (PathP≡doubleCompPathˡ a₋₀ a₀₋ a₁₋ a₋₁) + +-- Flipping a square in Ω²A is the same as inverting it +sym≡flipSquare : {x : A} (P : Square (refl {x = x}) refl refl refl) + → sym P ≡ flipSquare P +sym≡flipSquare {x = x} P = sym (main refl P) + where + B : (q : x ≡ x) → I → Type _ + B q i = PathP (λ j → x ≡ q (i ∨ j)) (λ k → q (i ∧ k)) refl + + main : (q : x ≡ x) (p : refl ≡ q) → PathP (λ i → B q i) (λ i j → p j i) (sym p) + main q = J (λ q p → PathP (λ i → B q i) (λ i j → p j i) (sym p)) refl + +-- Inverting both interval arguments of a square in Ω²A is the same as doing nothing +sym-cong-sym≡id : {x : A} (P : Square (refl {x = x}) refl refl refl) + → P ≡ λ i j → P (~ i) (~ j) +sym-cong-sym≡id {x = x} P = sym (main refl P) + where + B : (q : x ≡ x) → I → Type _ + B q i = Path (x ≡ q i) (λ j → q (i ∨ ~ j)) λ j → q (i ∧ j) + + main : (q : x ≡ x) (p : refl ≡ q) → PathP (λ i → B q i) (λ i j → p (~ i) (~ j)) p + main q = J (λ q p → PathP (λ i → B q i) (λ i j → p (~ i) (~ j)) p) refl + +-- Applying cong sym is the same as flipping a square in Ω²A +flipSquare≡cong-sym : ∀ {ℓ} {A : Type ℓ} {x : A} (P : Square (refl {x = x}) refl refl refl) + → flipSquare P ≡ λ i j → P i (~ j) +flipSquare≡cong-sym P = sym (sym≡flipSquare P) ∙ sym (sym-cong-sym≡id (cong sym P)) + +-- Applying cong sym is the same as inverting a square in Ω²A +sym≡cong-sym : ∀ {ℓ} {A : Type ℓ} {x : A} (P : Square (refl {x = x}) refl refl refl) + → sym P ≡ cong sym P +sym≡cong-sym P = sym-cong-sym≡id (sym P) + +-- sym induces an equivalence on identity types of paths +symIso : {a b : A} (p q : a ≡ b) → Iso (p ≡ q) (q ≡ p) +symIso p q = iso sym sym (λ _ → refl) λ _ → refl + +-- J is an equivalence +Jequiv : {x : A} (P : ∀ y → x ≡ y → Type ℓ') → P x refl ≃ (∀ {y} (p : x ≡ y) → P y p) +Jequiv P = isoToEquiv isom + where + isom : Iso _ _ + Iso.fun isom = J P + Iso.inv isom f = f refl + Iso.rightInv isom f = + implicitFunExt λ {_} → + funExt λ t → + J (λ _ t → J P (f refl) t ≡ f t) (JRefl P (f refl)) t + Iso.leftInv isom = JRefl P + +-- Action of PathP on equivalences (without relying on univalence) + +congPathIso : ∀ {ℓ ℓ'} {A : I → Type ℓ} {B : I → Type ℓ'} + (e : ∀ i → A i ≃ B i) {a₀ : A i0} {a₁ : A i1} + → Iso (PathP A a₀ a₁) (PathP B (e i0 .fst a₀) (e i1 .fst a₁)) +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.fun p i = e i .fst (p i) +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.inv q i = + hcomp + (λ j → λ + { (i = i0) → retEq (e i0) a₀ j + ; (i = i1) → retEq (e i1) a₁ j + }) + (invEq (e i) (q i)) +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.rightInv q k i = + hcomp + (λ j → λ + { (i = i0) → commSqIsEq (e i0 .snd) a₀ j k + ; (i = i1) → commSqIsEq (e i1 .snd) a₁ j k + ; (k = i0) → + e i .fst + (hfill + (λ j → λ + { (i = i0) → retEq (e i0) a₀ j + ; (i = i1) → retEq (e i1) a₁ j + }) + (inS (invEq (e i) (q i))) + j) + ; (k = i1) → q i + }) + (secEq (e i) (q i) k) + where b = commSqIsEq +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.leftInv p k i = + hcomp + (λ j → λ + { (i = i0) → retEq (e i0) a₀ (j ∨ k) + ; (i = i1) → retEq (e i1) a₁ (j ∨ k) + ; (k = i1) → p i + }) + (retEq (e i) (p i) k) + +congPathEquiv : ∀ {ℓ ℓ'} {A : I → Type ℓ} {B : I → Type ℓ'} + (e : ∀ i → A i ≃ B i) {a₀ : A i0} {a₁ : A i1} + → PathP A a₀ a₁ ≃ PathP B (e i0 .fst a₀) (e i1 .fst a₁) +congPathEquiv e = isoToEquiv (congPathIso e) + +-- Characterizations of dependent paths in path types + +doubleCompPath-filler∙ : {a b c d : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) + → PathP (λ i → p i ≡ r (~ i)) (p ∙ q ∙ r) q +doubleCompPath-filler∙ {A = A} {b = b} p q r j i = + hcomp (λ k → λ { (i = i0) → p j + ; (i = i1) → side j k + ; (j = i1) → q (i ∧ k)}) + (p (j ∨ i)) + where + side : I → I → A + side i j = + hcomp (λ k → λ { (i = i1) → q j + ; (j = i0) → b + ; (j = i1) → r (~ i ∧ k)}) + (q j) + +PathP→compPathL : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → PathP (λ i → p i ≡ q i) r s + → sym p ∙ r ∙ q ≡ s +PathP→compPathL {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (j ∨ k) + ; (i = i1) → q (j ∨ k) + ; (j = i0) → doubleCompPath-filler∙ (sym p) r q (~ k) i + ; (j = i1) → s i }) + (P j i) + +PathP→compPathR : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → PathP (λ i → p i ≡ q i) r s + → r ≡ p ∙ s ∙ sym q +PathP→compPathR {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (j ∧ (~ k)) + ; (i = i1) → q (j ∧ (~ k)) + ; (j = i0) → r i + ; (j = i1) → doubleCompPath-filler∙ p s (sym q) (~ k) i}) + (P j i) + + +-- Other direction + +compPathL→PathP : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → sym p ∙ r ∙ q ≡ s + → PathP (λ i → p i ≡ q i) r s +compPathL→PathP {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (~ k ∨ j) + ; (i = i1) → q (~ k ∨ j) + ; (j = i0) → doubleCompPath-filler∙ (sym p) r q k i + ; (j = i1) → s i}) + (P j i) + +compPathR→PathP : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → r ≡ p ∙ s ∙ sym q + → PathP (λ i → p i ≡ q i) r s +compPathR→PathP {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (k ∧ j) + ; (j = i0) → r i + ; (j = i1) → doubleCompPath-filler∙ p s (sym q) k i}) + (P j i) + +compPathR→PathP∙∙ : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → r ≡ p ∙∙ s ∙∙ sym q + → PathP (λ i → p i ≡ q i) r s +compPathR→PathP∙∙ {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (k ∧ j) + ; (j = i0) → r i + ; (j = i1) → doubleCompPath-filler p s (sym q) (~ k) i}) + (P j i) diff --git a/Cubical/Foundations/Pointed.agda b/Cubical/Foundations/Pointed.agda new file mode 100644 index 0000000000..f81c75846c --- /dev/null +++ b/Cubical/Foundations/Pointed.agda @@ -0,0 +1,9 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed where + +open import Cubical.Foundations.Pointed.Base public +open import Cubical.Foundations.Pointed.Properties public +open import Cubical.Foundations.Pointed.FunExt public +open import Cubical.Foundations.Pointed.Homotopy public + +open import Cubical.Foundations.Pointed.Homogeneous diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda new file mode 100644 index 0000000000..1cc03d57e3 --- /dev/null +++ b/Cubical/Foundations/Pointed/Base.agda @@ -0,0 +1,128 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Structure using (typ) public +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function + +Pointed : (ℓ : Level) → Type (ℓ-suc ℓ) +Pointed ℓ = TypeWithStr ℓ (λ x → x) + +pt : ∀ {ℓ} (A∙ : Pointed ℓ) → typ A∙ +pt = str + +Pointed₀ = Pointed ℓ-zero + +{- Pointed functions -} +_→∙_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') +(A , a) →∙ (B , b) = Σ[ f ∈ (A → B) ] f a ≡ b + +_→∙_∙ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +(A →∙ B ∙) .fst = A →∙ B +(A →∙ B ∙) .snd .fst x = pt B +(A →∙ B ∙) .snd .snd = refl + +idfun∙ : ∀ {ℓ} (A : Pointed ℓ) → A →∙ A +idfun∙ A .fst x = x +idfun∙ A .snd = refl + +{-Pointed equivalences -} +_≃∙_ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') +A ≃∙ B = Σ[ e ∈ fst A ≃ fst B ] fst e (pt A) ≡ pt B + +{- Underlying pointed map of an equivalence -} +≃∙map : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ≃∙ B → A →∙ B +fst (≃∙map e) = fst (fst e) +snd (≃∙map e) = snd e + +invEquiv∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ≃∙ B → B ≃∙ A +fst (invEquiv∙ x) = invEquiv (fst x) +snd (invEquiv∙ {A = A} x) = + sym (cong (fst (invEquiv (fst x))) (snd x)) ∙ retEq (fst x) (pt A) + +compEquiv∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → A ≃∙ B → B ≃∙ C → A ≃∙ C +fst (compEquiv∙ e1 e2) = compEquiv (fst e1) (fst e2) +snd (compEquiv∙ e1 e2) = cong (fst (fst e2)) (snd e1) ∙ snd e2 + +Equiv∙J : ∀ {ℓ ℓ'} {B : Pointed ℓ} (C : (A : Pointed ℓ) → A ≃∙ B → Type ℓ') + → C B (idEquiv (fst B) , refl) + → {A : _} (e : A ≃∙ B) → C A e +Equiv∙J {ℓ} {ℓ'} {B = B} C ind {A = A} = + uncurry λ e p → help e (pt A) (pt B) p C ind + where + help : ∀ {A : Type ℓ} (e : A ≃ typ B) (a : A) (b : typ B) + → (p : fst e a ≡ b) + → (C : (A : Pointed ℓ) → A ≃∙ (fst B , b) → Type ℓ') + → C (fst B , b) (idEquiv (fst B) , refl) + → C (A , a) (e , p) + help = EquivJ (λ A e → (a : A) (b : typ B) + → (p : fst e a ≡ b) + → (C : (A : Pointed ℓ) → A ≃∙ (fst B , b) → Type ℓ') + → C (fst B , b) (idEquiv (fst B) , refl) + → C (A , a) (e , p)) + λ a b → J (λ b p + → (C : (A : Pointed ℓ) → A ≃∙ (fst B , b) → Type ℓ') + → C (fst B , b) + (idEquiv (fst B) , refl) → + C (typ B , a) (idEquiv (typ B) , p)) + λ _ p → p + +ua∙ : ∀ {ℓ} {A B : Pointed ℓ} (e : fst A ≃ fst B) + → fst e (snd A) ≡ snd B → A ≡ B +fst (ua∙ e p i) = ua e i +snd (ua∙ {A = A} e p i) = ua-gluePath e p i + +{- J for pointed function types -} +→∙J : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Type ℓ'} + → (P : (b₀ : B) (f : A →∙ (B , b₀)) → Type ℓ'') + → ((f : fst A → B) → P (f (pt A)) (f , refl)) + → {b₀ : B} (f : A →∙ (B , b₀)) → P b₀ f +→∙J {A = A} P ind = + uncurry λ f → J (λ b₀ y → P b₀ (f , y)) (ind f) + +{- HIT allowing for pattern matching on pointed types -} +data Pointer {ℓ} (A : Pointed ℓ) : Type ℓ where + pt₀ : Pointer A + ⌊_⌋ : typ A → Pointer A + id : ⌊ pt A ⌋ ≡ pt₀ + +IsoPointedPointer : ∀ {ℓ} {A : Pointed ℓ} → Iso (typ A) (Pointer A) +Iso.fun IsoPointedPointer = ⌊_⌋ +Iso.inv (IsoPointedPointer {A = A}) pt₀ = pt A +Iso.inv IsoPointedPointer ⌊ x ⌋ = x +Iso.inv (IsoPointedPointer {A = A}) (id i) = pt A +Iso.rightInv IsoPointedPointer pt₀ = id +Iso.rightInv IsoPointedPointer ⌊ x ⌋ = refl +Iso.rightInv IsoPointedPointer (id i) j = id (i ∧ j) +Iso.leftInv IsoPointedPointer x = refl + +Pointed≡Pointer : ∀ {ℓ} {A : Pointed ℓ} → typ A ≡ Pointer A +Pointed≡Pointer = isoToPath IsoPointedPointer + +Pointer∙ : ∀ {ℓ} (A : Pointed ℓ) → Pointed ℓ +Pointer∙ A .fst = Pointer A +Pointer∙ A .snd = pt₀ + +Pointed≡∙Pointer : ∀ {ℓ} {A : Pointed ℓ} → A ≡ (Pointer A , pt₀) +Pointed≡∙Pointer {A = A} i = (Pointed≡Pointer {A = A} i) , helper i + where + helper : PathP (λ i → Pointed≡Pointer {A = A} i) (pt A) pt₀ + helper = ua-gluePath (isoToEquiv (IsoPointedPointer {A = A})) id + +pointerFun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Pointer A → Pointer B +pointerFun f pt₀ = pt₀ +pointerFun f ⌊ x ⌋ = ⌊ fst f x ⌋ +pointerFun f (id i) = (cong ⌊_⌋ (snd f) ∙ id) i + +pointerFun∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Pointer∙ A →∙ Pointer∙ B +pointerFun∙ f .fst = pointerFun f +pointerFun∙ f .snd = refl diff --git a/Cubical/Foundations/Pointed/FunExt.agda b/Cubical/Foundations/Pointed/FunExt.agda new file mode 100644 index 0000000000..69530642c3 --- /dev/null +++ b/Cubical/Foundations/Pointed/FunExt.agda @@ -0,0 +1,48 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.FunExt where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Pointed.Properties +open import Cubical.Foundations.Pointed.Homotopy + +private + variable + ℓ ℓ' : Level + +module _ {A : Pointed ℓ} {B : typ A → Type ℓ'} {ptB : B (pt A)} where + + -- pointed function extensionality + funExt∙P : {f g : Π∙ A B ptB} → f ∙∼P g → f ≡ g + funExt∙P (h , h∙) i .fst x = h x i + funExt∙P (h , h∙) i .snd = h∙ i + + -- inverse of pointed function extensionality + funExt∙P⁻ : {f g : Π∙ A B ptB} → f ≡ g → f ∙∼P g + funExt∙P⁻ p .fst a i = p i .fst a + funExt∙P⁻ p .snd i = p i .snd + + -- function extensionality is an isomorphism, PathP version + funExt∙PIso : (f g : Π∙ A B ptB) → Iso (f ∙∼P g) (f ≡ g) + Iso.fun (funExt∙PIso f g) = funExt∙P {f = f} {g = g} + Iso.inv (funExt∙PIso f g) = funExt∙P⁻ {f = f} {g = g} + Iso.rightInv (funExt∙PIso f g) p i j = p j + Iso.leftInv (funExt∙PIso f g) h _ = h + + -- transformed to equivalence + funExt∙P≃ : (f g : Π∙ A B ptB) → (f ∙∼P g) ≃ (f ≡ g) + funExt∙P≃ f g = isoToEquiv (funExt∙PIso f g) + + -- funExt∙≃ using the other kind of pointed homotopy + funExt∙≃ : (f g : Π∙ A B ptB) → (f ∙∼ g) ≃ (f ≡ g) + funExt∙≃ f g = compEquiv (∙∼≃∙∼P f g) (funExt∙P≃ f g) + + -- standard pointed function extensionality and its inverse + funExt∙ : {f g : Π∙ A B ptB} → f ∙∼ g → f ≡ g + funExt∙ {f = f} {g = g} = equivFun (funExt∙≃ f g) + + funExt∙⁻ : {f g : Π∙ A B ptB} → f ≡ g → f ∙∼ g + funExt∙⁻ {f = f} {g = g} = equivFun (invEquiv (funExt∙≃ f g)) diff --git a/Cubical/Foundations/Pointed/Homogeneous.agda b/Cubical/Foundations/Pointed/Homogeneous.agda new file mode 100644 index 0000000000..4c00385cad --- /dev/null +++ b/Cubical/Foundations/Pointed/Homogeneous.agda @@ -0,0 +1,141 @@ +{- + +Definition of a homogeneous pointed type, and proofs that pi, product, path, and discrete types are homogeneous + +Portions of this file adapted from Nicolai Kraus' code here: + https://bitbucket.org/nicolaikraus/agda/src/e30d70c72c6af8e62b72eefabcc57623dd921f04/trunc-inverse.lagda + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.Homogeneous where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Relation.Nullary + +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Pointed.Properties +open import Cubical.Structures.Pointed + +isHomogeneous : ∀ {ℓ} → Pointed ℓ → Type (ℓ-suc ℓ) +isHomogeneous {ℓ} (A , x) = ∀ y → Path (Pointed ℓ) (A , x) (A , y) + +-- Pointed functions into a homogeneous type are equal as soon as they are equal +-- as unpointed functions +→∙Homogeneous≡ : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} + (h : isHomogeneous B∙) → f∙ .fst ≡ g∙ .fst → f∙ ≡ g∙ +→∙Homogeneous≡ {A∙ = A∙@(_ , a₀)} {B∙@(B , _)} {f∙@(_ , f₀)} {g∙@(_ , g₀)} h p = + subst (λ Q∙ → PathP (λ i → A∙ →∙ Q∙ i) f∙ g∙) (sym (flipSquare fix)) badPath + where + badPath : PathP (λ i → A∙ →∙ (B , (sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) f∙ g∙ + badPath i .fst = p i + badPath i .snd j = doubleCompPath-filler (sym f₀) (funExt⁻ p a₀) g₀ j i + + fix : PathP (λ i → B∙ ≡ (B , (sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) refl refl + fix i = + hcomp + (λ j → λ + { (i = i0) → lCancel (h (pt B∙)) j + ; (i = i1) → lCancel (h (pt B∙)) j + }) + (sym (h (pt B∙)) ∙ h ((sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) + +isHomogeneousPi : ∀ {ℓ ℓ'} {A : Type ℓ} {B∙ : A → Pointed ℓ'} + → (∀ a → isHomogeneous (B∙ a)) → isHomogeneous (Πᵘ∙ A B∙) +isHomogeneousPi h f i .fst = ∀ a → typ (h a (f a) i) +isHomogeneousPi h f i .snd a = pt (h a (f a) i) + +isHomogeneousΠ∙ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : typ A → Type ℓ') + → (b₀ : B (pt A)) + → ((a : typ A) (x : B a) → isHomogeneous (B a , x)) + → (f : Π∙ A B b₀) + → isHomogeneous (Π∙ A B b₀ , f) +fst (isHomogeneousΠ∙ A B b₀ h f g i) = + Σ[ r ∈ ((a : typ A) → fst ((h a (fst f a) (fst g a)) i)) ] + r (pt A) ≡ hcomp (λ k → λ {(i = i0) → snd f k + ; (i = i1) → snd g k}) + (snd (h (pt A) (fst f (pt A)) (fst g (pt A)) i)) +snd (isHomogeneousΠ∙ A B b₀ h f g i) = + (λ a → snd (h a (fst f a) (fst g a) i)) + , λ j → hcomp (λ k → λ { (i = i0) → snd f (k ∧ j) + ; (i = i1) → snd g (k ∧ j) + ; (j = i0) → snd (h (pt A) (fst f (pt A)) + (fst g (pt A)) i)}) + (snd (h (pt A) (fst f (pt A)) (fst g (pt A)) i)) + +isHomogeneous→∙ : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} + → isHomogeneous B∙ → isHomogeneous (A∙ →∙ B∙ ∙) +isHomogeneous→∙ {A∙ = A∙} {B∙} h f∙ = + ΣPathP + ( (λ i → Π∙ A∙ (λ a → T a i) (t₀ i)) + , PathPIsoPath _ _ _ .Iso.inv + (→∙Homogeneous≡ h + (PathPIsoPath (λ i → (a : typ A∙) → T a i) (λ _ → pt B∙) _ .Iso.fun + (λ i a → pt (h (f∙ .fst a) i)))) + ) + where + T : ∀ a → typ B∙ ≡ typ B∙ + T a i = typ (h (f∙ .fst a) i) + + t₀ : PathP (λ i → T (pt A∙) i) (pt B∙) (pt B∙) + t₀ = cong pt (h (f∙ .fst (pt A∙))) ▷ f∙ .snd + +isHomogeneousProd : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} + → isHomogeneous A∙ → isHomogeneous B∙ → isHomogeneous (A∙ ×∙ B∙) +isHomogeneousProd hA hB (a , b) i .fst = typ (hA a i) × typ (hB b i) +isHomogeneousProd hA hB (a , b) i .snd .fst = pt (hA a i) +isHomogeneousProd hA hB (a , b) i .snd .snd = pt (hB b i) + +isHomogeneousPath : ∀ {ℓ} (A : Type ℓ) {x y : A} (p : x ≡ y) → isHomogeneous ((x ≡ y) , p) +isHomogeneousPath A {x} {y} p q + = pointed-sip ((x ≡ y) , p) ((x ≡ y) , q) (eqv , compPathr-cancel p q) + where eqv : (x ≡ y) ≃ (x ≡ y) + eqv = compPathlEquiv (q ∙ sym p) + +module HomogeneousDiscrete {ℓ} {A∙ : Pointed ℓ} (dA : Discrete (typ A∙)) (y : typ A∙) where + + -- switches pt A∙ with y + switch : typ A∙ → typ A∙ + switch x with dA x (pt A∙) + ... | yes _ = y + ... | no _ with dA x y + ... | yes _ = pt A∙ + ... | no _ = x + + switch-ptA∙ : switch (pt A∙) ≡ y + switch-ptA∙ with dA (pt A∙) (pt A∙) + ... | yes _ = refl + ... | no ¬p = ⊥.rec (¬p refl) + + switch-idp : ∀ x → switch (switch x) ≡ x + switch-idp x with dA x (pt A∙) + switch-idp x | yes p with dA y (pt A∙) + switch-idp x | yes p | yes q = q ∙ sym p + switch-idp x | yes p | no _ with dA y y + switch-idp x | yes p | no _ | yes _ = sym p + switch-idp x | yes p | no _ | no ¬p = ⊥.rec (¬p refl) + switch-idp x | no ¬p with dA x y + switch-idp x | no ¬p | yes p with dA y (pt A∙) + switch-idp x | no ¬p | yes p | yes q = ⊥.rec (¬p (p ∙ q)) + switch-idp x | no ¬p | yes p | no _ with dA (pt A∙) (pt A∙) + switch-idp x | no ¬p | yes p | no _ | yes _ = sym p + switch-idp x | no ¬p | yes p | no _ | no ¬q = ⊥.rec (¬q refl) + switch-idp x | no ¬p | no ¬q with dA x (pt A∙) + switch-idp x | no ¬p | no ¬q | yes p = ⊥.rec (¬p p) + switch-idp x | no ¬p | no ¬q | no _ with dA x y + switch-idp x | no ¬p | no ¬q | no _ | yes q = ⊥.rec (¬q q) + switch-idp x | no ¬p | no ¬q | no _ | no _ = refl + + switch-eqv : typ A∙ ≃ typ A∙ + switch-eqv = isoToEquiv (iso switch switch switch-idp switch-idp) + +isHomogeneousDiscrete : ∀ {ℓ} {A∙ : Pointed ℓ} (dA : Discrete (typ A∙)) → isHomogeneous A∙ +isHomogeneousDiscrete {ℓ} {A∙} dA y + = pointed-sip (typ A∙ , pt A∙) (typ A∙ , y) (switch-eqv , switch-ptA∙) + where open HomogeneousDiscrete {ℓ} {A∙} dA y diff --git a/Cubical/Foundations/Pointed/Homotopy.agda b/Cubical/Foundations/Pointed/Homotopy.agda new file mode 100644 index 0000000000..06612f6b2d --- /dev/null +++ b/Cubical/Foundations/Pointed/Homotopy.agda @@ -0,0 +1,119 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.Homotopy where + +{- + This module defines two kinds of pointed homotopies, + ∙∼ and ∙∼P, and proves their equivalence +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Fiberwise +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Pointed.Properties +open import Cubical.Homotopy.Base +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ' : Level + +module _ {A : Pointed ℓ} {B : typ A → Type ℓ'} {ptB : B (pt A)} where + + ⋆ = pt A + + -- pointed homotopy as pointed Π. This is just a Σ-type, see ∙∼Σ + _∙∼_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ') + (f₁ , f₂) ∙∼ (g₁ , g₂) = Π∙ A (λ x → f₁ x ≡ g₁ x) (f₂ ∙ g₂ ⁻¹) + + -- pointed homotopy with PathP. Also a Σ-type, see ∙∼PΣ + _∙∼P_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ') + (f₁ , f₂) ∙∼P (g₁ , g₂) = Σ[ h ∈ f₁ ∼ g₁ ] PathP (λ i → h ⋆ i ≡ ptB) f₂ g₂ + + -- Proof that f ∙∼ g ≃ f ∙∼P g + -- using equivalence of the total map of φ + private + module _ (f g : Π∙ A B ptB) (H : f .fst ∼ g .fst) where + -- convenient notation + f₁ = fst f + f₂ = snd f + g₁ = fst g + g₂ = snd g + + -- P is the predicate on a homotopy H to be pointed of the ∙∼ kind + P : Type ℓ' + P = H ⋆ ≡ f₂ ∙ g₂ ⁻¹ + + -- Q is the predicate on a homotopy H to be pointed of the ∙∼P kind + Q : Type ℓ' + Q = PathP (λ i → H ⋆ i ≡ ptB) f₂ g₂ + + -- simplify the notation even more to see that P≡Q + -- is just a jingle of paths + p = H ⋆ + r = f₂ + s = g₂ + P≡Q : P ≡ Q + P≡Q = p ≡ r ∙ s ⁻¹ + ≡⟨ isoToPath (symIso p (r ∙ s ⁻¹)) ⟩ + r ∙ s ⁻¹ ≡ p + ≡⟨ cong (r ∙ s ⁻¹ ≡_) (rUnit p ∙∙ cong (p ∙_) (sym (rCancel s)) ∙∙ assoc p s (s ⁻¹)) ⟩ + r ∙ s ⁻¹ ≡ (p ∙ s) ∙ s ⁻¹ + ≡⟨ sym (ua (compr≡Equiv r (p ∙ s) (s ⁻¹))) ⟩ + r ≡ p ∙ s + ≡⟨ ua (compl≡Equiv (p ⁻¹) r (p ∙ s)) ⟩ + p ⁻¹ ∙ r ≡ p ⁻¹ ∙ (p ∙ s) + ≡⟨ cong (p ⁻¹ ∙ r ≡_ ) (assoc (p ⁻¹) p s ∙∙ (cong (_∙ s) (lCancel p)) ∙∙ sym (lUnit s)) ⟩ + p ⁻¹ ∙ r ≡ s + ≡⟨ cong (λ z → p ⁻¹ ∙ z ≡ s) (rUnit r) ⟩ + p ⁻¹ ∙ (r ∙ refl) ≡ s + ≡⟨ cong (_≡ s) (sym (doubleCompPath-elim' (p ⁻¹) r refl)) ⟩ + p ⁻¹ ∙∙ r ∙∙ refl ≡ s + ≡⟨ sym (ua (Square≃doubleComp r s p refl)) ⟩ + PathP (λ i → p i ≡ ptB) r s ∎ + + -- φ is a fiberwise transformation (H : f ∼ g) → P H → Q H + -- φ is even a fiberwise equivalence by P≡Q + φ : P → Q + φ = transport P≡Q + + -- The total map corresponding to φ + totφ : (f g : Π∙ A B ptB) → f ∙∼ g → f ∙∼P g + totφ f g p .fst = p .fst + totφ f g p .snd = φ f g (p .fst) (p .snd) + + -- transformation of the homotopies using totφ + ∙∼→∙∼P : (f g : Π∙ A B ptB) → (f ∙∼ g) → (f ∙∼P g) + ∙∼→∙∼P f g = totφ f g + + -- Proof that ∙∼ and ∙∼P are equivalent using the fiberwise equivalence φ + ∙∼≃∙∼P : (f g : Π∙ A B ptB) → (f ∙∼ g) ≃ (f ∙∼P g) + ∙∼≃∙∼P f g = Σ-cong-equiv-snd (λ H → transportEquiv (P≡Q f g H)) + + -- inverse of ∙∼→∙∼P extracted from the equivalence + ∙∼P→∙∼ : {f g : Π∙ A B ptB} → f ∙∼P g → f ∙∼ g + ∙∼P→∙∼ {f = f} {g = g} = invEq (∙∼≃∙∼P f g) + + -- ∙∼≃∙∼P transformed to a path + ∙∼≡∙∼P : (f g : Π∙ A B ptB) → (f ∙∼ g) ≡ (f ∙∼P g) + ∙∼≡∙∼P f g = ua (∙∼≃∙∼P f g) + + -- Verifies that the pointed homotopies actually correspond + -- to their Σ-type versions + _∙∼Σ_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ') + f ∙∼Σ g = Σ[ H ∈ f .fst ∼ g .fst ] (P f g H) + + _∙∼PΣ_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ') + f ∙∼PΣ g = Σ[ H ∈ f .fst ∼ g .fst ] (Q f g H) + + ∙∼≡∙∼Σ : (f g : Π∙ A B ptB) → f ∙∼ g ≡ f ∙∼Σ g + ∙∼≡∙∼Σ f g = refl + + ∙∼P≡∙∼PΣ : (f g : Π∙ A B ptB) → f ∙∼P g ≡ f ∙∼PΣ g + ∙∼P≡∙∼PΣ f g = refl diff --git a/Cubical/Foundations/Pointed/Properties.agda b/Cubical/Foundations/Pointed/Properties.agda new file mode 100644 index 0000000000..a96ca1be0a --- /dev/null +++ b/Cubical/Foundations/Pointed/Properties.agda @@ -0,0 +1,99 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ' ℓA ℓB ℓC ℓD : Level + +-- the default pointed Π-type: A is pointed, and B has a base point in the chosen fiber +Π∙ : (A : Pointed ℓ) (B : typ A → Type ℓ') (ptB : B (pt A)) → Type (ℓ-max ℓ ℓ') +Π∙ A B ptB = Σ[ f ∈ ((a : typ A) → B a) ] f (pt A) ≡ ptB + +-- the unpointed Π-type becomes a pointed type if the fibers are all pointed +Πᵘ∙ : (A : Type ℓ) (B : A → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +Πᵘ∙ A B .fst = ∀ a → typ (B a) +Πᵘ∙ A B .snd a = pt (B a) + +-- if the base and all fibers are pointed, we have the pointed pointed Π-type +Πᵖ∙ : (A : Pointed ℓ) (B : typ A → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +Πᵖ∙ A B .fst = Π∙ A (typ ∘ B) (pt (B (pt A))) +Πᵖ∙ A B .snd .fst a = pt (B a) +Πᵖ∙ A B .snd .snd = refl + +-- the default pointed Σ-type is just the Σ-type, but as a pointed type +Σ∙ : (A : Pointed ℓ) (B : typ A → Type ℓ') (ptB : B (pt A)) → Pointed (ℓ-max ℓ ℓ') +Σ∙ A B ptB .fst = Σ[ a ∈ typ A ] B a +Σ∙ A B ptB .snd .fst = pt A +Σ∙ A B ptB .snd .snd = ptB + +-- version if B is a family of pointed types +Σᵖ∙ : (A : Pointed ℓ) (B : typ A → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +Σᵖ∙ A B = Σ∙ A (typ ∘ B) (pt (B (pt A))) + +_×∙_ : (A∙ : Pointed ℓ) (B∙ : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +(A∙ ×∙ B∙) .fst = (typ A∙) × (typ B∙) +(A∙ ×∙ B∙) .snd .fst = pt A∙ +(A∙ ×∙ B∙) .snd .snd = pt B∙ + +-- composition of pointed maps +_∘∙_ : {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} + (g : B →∙ C) (f : A →∙ B) → (A →∙ C) +((g , g∙) ∘∙ (f , f∙)) .fst x = g (f x) +((g , g∙) ∘∙ (f , f∙)) .snd = (cong g f∙) ∙ g∙ + +-- post composition +post∘∙ : ∀ {ℓX ℓ ℓ'} (X : Pointed ℓX) {A : Pointed ℓ} {B : Pointed ℓ'} + → (A →∙ B) → ((X →∙ A ∙) →∙ (X →∙ B ∙)) +post∘∙ X f .fst g = f ∘∙ g +post∘∙ X f .snd = + ΣPathP + ( (funExt λ _ → f .snd) + , (sym (lUnit (f .snd)) ◁ λ i j → f .snd (i ∨ j))) + +-- pointed identity +id∙ : (A : Pointed ℓA) → (A →∙ A) +id∙ A .fst x = x +id∙ A .snd = refl + +-- constant pointed map +const∙ : (A : Pointed ℓA) (B : Pointed ℓB) → (A →∙ B) +const∙ _ B .fst _ = B .snd +const∙ _ B .snd = refl + +-- left identity law for pointed maps +∘∙-idˡ : {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) → f ∘∙ id∙ A ≡ f +∘∙-idˡ f = ΣPathP ( refl , (lUnit (f .snd)) ⁻¹ ) + +-- right identity law for pointed maps +∘∙-idʳ : {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) → id∙ B ∘∙ f ≡ f +∘∙-idʳ f = ΣPathP ( refl , (rUnit (f .snd)) ⁻¹ ) + +-- associativity for composition of pointed maps +∘∙-assoc : {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} {D : Pointed ℓD} + (h : C →∙ D) (g : B →∙ C) (f : A →∙ B) + → (h ∘∙ g) ∘∙ f ≡ h ∘∙ (g ∘∙ f) +∘∙-assoc (h , h∙) (g , g∙) (f , f∙) = ΣPathP (refl , q) + where + q : (cong (h ∘ g) f∙) ∙ (cong h g∙ ∙ h∙) ≡ cong h (cong g f∙ ∙ g∙) ∙ h∙ + q = ( (cong (h ∘ g) f∙) ∙ (cong h g∙ ∙ h∙) + ≡⟨ refl ⟩ + (cong h (cong g f∙)) ∙ (cong h g∙ ∙ h∙) + ≡⟨ assoc (cong h (cong g f∙)) (cong h g∙) h∙ ⟩ + (cong h (cong g f∙) ∙ cong h g∙) ∙ h∙ + ≡⟨ cong (λ p → p ∙ h∙) ((cong-∙ h (cong g f∙) g∙) ⁻¹) ⟩ + (cong h (cong g f∙ ∙ g∙) ∙ h∙) ∎ ) + +module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where + isInIm∙ : (x : typ B) → Type (ℓ-max ℓ ℓ') + isInIm∙ x = Σ[ z ∈ typ A ] fst f z ≡ x + + isInKer∙ : (x : fst A) → Type ℓ' + isInKer∙ x = fst f x ≡ snd B diff --git a/Cubical/Foundations/Powerset.agda b/Cubical/Foundations/Powerset.agda new file mode 100644 index 0000000000..043a0dfed8 --- /dev/null +++ b/Cubical/Foundations/Powerset.agda @@ -0,0 +1,65 @@ +{- + +This file introduces the "powerset" of a type in the style of +Escardó's lecture notes: + +https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#propositionalextensionality + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Powerset where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence using (hPropExt) + +open import Cubical.Data.Sigma + +private + variable + ℓ : Level + X : Type ℓ + +ℙ : Type ℓ → Type (ℓ-suc ℓ) +ℙ X = X → hProp _ + +isSetℙ : isSet (ℙ X) +isSetℙ = isSetΠ λ x → isSetHProp + +infix 5 _∈_ + +_∈_ : {X : Type ℓ} → X → ℙ X → Type ℓ +x ∈ A = ⟨ A x ⟩ + +_⊆_ : {X : Type ℓ} → ℙ X → ℙ X → Type ℓ +A ⊆ B = ∀ x → x ∈ A → x ∈ B + +∈-isProp : (A : ℙ X) (x : X) → isProp (x ∈ A) +∈-isProp A = snd ∘ A + +⊆-isProp : (A B : ℙ X) → isProp (A ⊆ B) +⊆-isProp A B = isPropΠ2 (λ x _ → ∈-isProp B x) + +⊆-refl : (A : ℙ X) → A ⊆ A +⊆-refl A x = idfun (x ∈ A) + +subst-∈ : (A : ℙ X) {x y : X} → x ≡ y → x ∈ A → y ∈ A +subst-∈ A = subst (_∈ A) + +⊆-refl-consequence : (A B : ℙ X) → A ≡ B → (A ⊆ B) × (B ⊆ A) +⊆-refl-consequence A B p = subst (A ⊆_) p (⊆-refl A) + , subst (B ⊆_) (sym p) (⊆-refl B) + +⊆-extensionality : (A B : ℙ X) → (A ⊆ B) × (B ⊆ A) → A ≡ B +⊆-extensionality A B (φ , ψ) = + funExt (λ x → TypeOfHLevel≡ 1 (hPropExt (A x .snd) (B x .snd) (φ x) (ψ x))) + +⊆-extensionalityEquiv : (A B : ℙ X) → (A ⊆ B) × (B ⊆ A) ≃ (A ≡ B) +⊆-extensionalityEquiv A B = isoToEquiv (iso (⊆-extensionality A B) + (⊆-refl-consequence A B) + (λ _ → isSetℙ A B _ _) + (λ _ → isPropΣ (⊆-isProp A B) (λ _ → ⊆-isProp B A) _ _)) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda new file mode 100644 index 0000000000..0f15d63467 --- /dev/null +++ b/Cubical/Foundations/Prelude.agda @@ -0,0 +1,510 @@ +{- + +This file proves a variety of basic results about paths: + +- refl, sym, cong and composition of paths. This is used to set up + equational reasoning. + +- Transport, subst and functional extensionality + +- J and its computation rule (up to a path) + +- Σ-types and contractibility of singletons + +- Converting PathP to and from a homogeneous path with transp + +- Direct definitions of lower h-levels + +- Export natural numbers + +- Export universe lifting + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Prelude where + +open import Cubical.Core.Primitives public + +infixr 30 _∙_ +infixr 30 _∙₂_ +infix 3 _∎ +infixr 2 _≡⟨_⟩_ +infixr 2.5 _≡⟨_⟩≡⟨_⟩_ + +-- Basic theory about paths. These proofs should typically be +-- inlined. This module also makes equational reasoning work with +-- (non-dependent) paths. + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + B : A → Type ℓ + x y z w : A + +refl : x ≡ x +refl {x = x} _ = x +{-# INLINE refl #-} + +sym : x ≡ y → y ≡ x +sym p i = p (~ i) +{-# INLINE sym #-} + +symP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → + (p : PathP A x y) → PathP (λ i → A (~ i)) y x +symP p j = p (~ j) + +cong : (f : (a : A) → B a) (p : x ≡ y) → + PathP (λ i → B (p i)) (f x) (f y) +cong f p i = f (p i) +{-# INLINE cong #-} + +congP : {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ'} + (f : (i : I) → (a : A i) → B i a) {x : A i0} {y : A i1} + (p : PathP A x y) → PathP (λ i → B i (p i)) (f i0 x) (f i1 y) +congP f p i = f i (p i) +{-# INLINE congP #-} + +cong₂ : {C : (a : A) → (b : B a) → Type ℓ} → + (f : (a : A) → (b : B a) → C a b) → + (p : x ≡ y) → + {u : B x} {v : B y} (q : PathP (λ i → B (p i)) u v) → + PathP (λ i → C (p i) (q i)) (f x u) (f y v) +cong₂ f p q i = f (p i) (q i) +{-# INLINE cong₂ #-} + +congP₂ : {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ'} + {C : (i : I) (a : A i) → B i a → Type ℓ''} + (f : (i : I) → (a : A i) → (b : B i a) → C i a b) + {x : A i0} {y : A i1} {u : B i0 x} {v : B i1 y} + (p : PathP A x y) (q : PathP (λ i → B i (p i)) u v) + → PathP (λ i → C i (p i) (q i)) (f i0 x u) (f i1 y v) +congP₂ f p q i = f i (p i) (q i) +{-# INLINE congP₂ #-} + +{- The most natural notion of homogenous path composition + in a cubical setting is double composition: + + x ∙ ∙ ∙ > w + ^ ^ + p⁻¹ | | r ^ + | | j | + y — — — > z ∙ — > + q i + + `p ∙∙ q ∙∙ r` gives the line at the top, + `doubleCompPath-filler p q r` gives the whole square +-} + +doubleComp-faces : {x y z w : A } (p : x ≡ y) (r : z ≡ w) + → (i : I) (j : I) → Partial (i ∨ ~ i) A +doubleComp-faces p r i j (i = i0) = p (~ j) +doubleComp-faces p r i j (i = i1) = r j + +_∙∙_∙∙_ : w ≡ x → x ≡ y → y ≡ z → w ≡ z +(p ∙∙ q ∙∙ r) i = + hcomp (doubleComp-faces p r i) (q i) + +doubleCompPath-filler : (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → PathP (λ j → p (~ j) ≡ r j) q (p ∙∙ q ∙∙ r) +doubleCompPath-filler p q r j i = + hfill (doubleComp-faces p r i) (inS (q i)) j + +-- any two definitions of double composition are equal +compPath-unique : (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → (α β : Σ[ s ∈ x ≡ w ] PathP (λ j → p (~ j) ≡ r j) q s) + → α ≡ β +compPath-unique p q r (α , α-filler) (β , β-filler) t + = (λ i → cb i1 i) , (λ j i → cb j i) + where cb : I → I → _ + cb j i = hfill (λ j → λ { (t = i0) → α-filler j i + ; (t = i1) → β-filler j i + ; (i = i0) → p (~ j) + ; (i = i1) → r j }) + (inS (q i)) j + +{- For single homogenous path composition, we take `p = refl`: + + x ∙ ∙ ∙ > z + ‖ ^ + ‖ | r ^ + ‖ | j | + x — — — > y ∙ — > + q i + + `q ∙ r` gives the line at the top, + `compPath-filler q r` gives the whole square +-} + +_∙_ : x ≡ y → y ≡ z → x ≡ z +p ∙ q = refl ∙∙ p ∙∙ q + +compPath-filler : (p : x ≡ y) (q : y ≡ z) → PathP (λ j → x ≡ q j) p (p ∙ q) +compPath-filler p q = doubleCompPath-filler refl p q + +-- We could have also defined single composition by taking `r = refl`: + +_∙'_ : x ≡ y → y ≡ z → x ≡ z +p ∙' q = p ∙∙ q ∙∙ refl + +compPath'-filler : (p : x ≡ y) (q : y ≡ z) → PathP (λ j → p (~ j) ≡ z) q (p ∙' q) +compPath'-filler p q = doubleCompPath-filler p q refl + +-- It's easy to show that `p ∙ q` also has such a filler: +compPath-filler' : (p : x ≡ y) (q : y ≡ z) → PathP (λ j → p (~ j) ≡ z) q (p ∙ q) +compPath-filler' {z = z} p q j i = + hcomp (λ k → λ { (i = i0) → p (~ j) + ; (i = i1) → q k + ; (j = i0) → q (i ∧ k) }) + (p (i ∨ ~ j)) +-- Note: We can omit a (j = i1) case here since when (j = i1), the whole expression is +-- definitionally equal to `p ∙ q`. (Notice that `p ∙ q` is also an hcomp.) Nevertheless, +-- we could have given `compPath-filler p q k i` as the (j = i1) case. + +-- From this, we can show that these two notions of composition are the same +compPath≡compPath' : (p : x ≡ y) (q : y ≡ z) → p ∙ q ≡ p ∙' q +compPath≡compPath' p q j = + compPath-unique p q refl (p ∙ q , compPath-filler' p q) + (p ∙' q , compPath'-filler p q) j .fst + +-- Double composition agrees with iterated single composition +doubleCompPath≡compPath : {x y z w : A} + (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r +doubleCompPath≡compPath p q r i j = + hcomp (λ k → λ { (i = i1) → compPath-filler' p (q ∙ r) k j + ; (j = i0) → p (~ k) + ; (j = i1) → r (i ∨ k)}) + (compPath-filler q r i j) + +-- Heterogeneous path composition and its filler: + +-- Composition in a family indexed by the interval +compPathP : {A : I → Type ℓ} {x : A i0} {y : A i1} {B_i1 : Type ℓ} {B : (A i1) ≡ B_i1} {z : B i1} + → PathP A x y → PathP (λ i → B i) y z → PathP (λ j → ((λ i → A i) ∙ B) j) x z +compPathP {A = A} {x = x} {B = B} p q i = + comp (λ j → compPath-filler (λ i → A i) B j i) + (λ j → λ { (i = i0) → x ; + (i = i1) → q j }) + (p i) + +-- Composition in a family indexed by a type +compPathP' : {B : A → Type ℓ'} {x' : B x} {y' : B y} {z' : B z} {p : x ≡ y} {q : y ≡ z} + (P : PathP (λ i → B (p i)) x' y') (Q : PathP (λ i → B (q i)) y' z') + → PathP (λ i → B ((p ∙ q) i)) x' z' +compPathP' {B = B} {x' = x'} {p = p} {q = q} P Q i = + comp (λ j → B (compPath-filler p q j i)) + (λ j → λ { (i = i0) → x' ; + (i = i1) → Q j }) + (P i) + +compPathP-filler : {A : I → Type ℓ} {x : A i0} {y : A i1} {B_i1 : Type ℓ} {B : A i1 ≡ B_i1} {z : B i1} + (p : PathP A x y) (q : PathP (λ i → B i) y z) + → PathP (λ j → PathP (λ k → (compPath-filler (λ i → A i) B j k)) x (q j)) p (compPathP p q) +compPathP-filler {A = A} {x = x} {B = B} p q j i = + fill (λ j → compPath-filler (λ i → A i) B j i) + (λ j → λ { (i = i0) → x ; + (i = i1) → q j }) + (inS (p i)) j + +compPathP'-filler : {B : A → Type ℓ'} {x' : B x} {y' : B y} {z' : B z} {p : x ≡ y} {q : y ≡ z} + (P : PathP (λ i → B (p i)) x' y') (Q : PathP (λ i → B (q i)) y' z') + → PathP (λ j → PathP (λ i → B (compPath-filler p q j i)) x' (Q j)) P (compPathP' {B = B} P Q) +compPathP'-filler {B = B} {x' = x'} {p = p} {q = q} P Q j i = + fill (λ j → B (compPath-filler p q j i)) + (λ j → λ { (i = i0) → x' ; + (i = i1) → Q j }) + (inS (P i)) + j + +-- Syntax for chains of equational reasoning + +_≡⟨_⟩_ : (x : A) → x ≡ y → y ≡ z → x ≡ z +_ ≡⟨ x≡y ⟩ y≡z = x≡y ∙ y≡z + +≡⟨⟩-syntax : (x : A) → x ≡ y → y ≡ z → x ≡ z +≡⟨⟩-syntax = _≡⟨_⟩_ +infixr 2 ≡⟨⟩-syntax +syntax ≡⟨⟩-syntax x (λ i → B) y = x ≡[ i ]⟨ B ⟩ y + +≡⟨⟩⟨⟩-syntax : (x y : A) → x ≡ y → y ≡ z → z ≡ w → x ≡ w +≡⟨⟩⟨⟩-syntax x y p q r = p ∙∙ q ∙∙ r +infixr 3 ≡⟨⟩⟨⟩-syntax +syntax ≡⟨⟩⟨⟩-syntax x y B C = x ≡⟨ B ⟩≡ y ≡⟨ C ⟩≡ + +_≡⟨_⟩≡⟨_⟩_ : (x : A) → x ≡ y → y ≡ z → z ≡ w → x ≡ w +_ ≡⟨ x≡y ⟩≡⟨ y≡z ⟩ z≡w = x≡y ∙∙ y≡z ∙∙ z≡w + +_∎ : (x : A) → x ≡ x +_ ∎ = refl + +-- Transport and subst + +-- transport is a special case of transp +transport : {A B : Type ℓ} → A ≡ B → A → B +transport p a = transp (λ i → p i) i0 a + +-- Transporting in a constant family is the identity function (up to a +-- path). If we would have regularity this would be definitional. +transportRefl : (x : A) → transport refl x ≡ x +transportRefl {A = A} x i = transp (λ _ → A) i x + +transport-filler : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) (x : A) + → PathP (λ i → p i) x (transport p x) +transport-filler p x i = transp (λ j → p (i ∧ j)) (~ i) x + +-- We want B to be explicit in subst +subst : (B : A → Type ℓ') (p : x ≡ y) → B x → B y +subst B p pa = transport (λ i → B (p i)) pa + +subst2 : ∀ {ℓ' ℓ''} {B : Type ℓ'} {z w : B} (C : A → B → Type ℓ'') + (p : x ≡ y) (q : z ≡ w) → C x z → C y w +subst2 B p q b = transport (λ i → B (p i) (q i)) b + +substRefl : (px : B x) → subst B refl px ≡ px +substRefl px = transportRefl px + +subst-filler : (B : A → Type ℓ') (p : x ≡ y) (b : B x) + → PathP (λ i → B (p i)) b (subst B p b) +subst-filler B p = transport-filler (cong B p) + +-- Function extensionality + +funExt : {B : A → I → Type ℓ'} + {f : (x : A) → B x i0} {g : (x : A) → B x i1} + → ((x : A) → PathP (B x) (f x) (g x)) + → PathP (λ i → (x : A) → B x i) f g +funExt p i x = p x i + +implicitFunExt : {B : A → I → Type ℓ'} + {f : {x : A} → B x i0} {g : {x : A} → B x i1} + → ({x : A} → PathP (B x) (f {x}) (g {x})) + → PathP (λ i → {x : A} → B x i) f g +implicitFunExt p i {x} = p {x} i + +-- the inverse to funExt (see Functions.FunExtEquiv), converting paths +-- between functions to homotopies; `funExt⁻` is called `happly` and +-- defined by path induction in the HoTT book (see function 2.9.2 in +-- section 2.9) +funExt⁻ : {B : A → I → Type ℓ'} + {f : (x : A) → B x i0} {g : (x : A) → B x i1} + → PathP (λ i → (x : A) → B x i) f g + → ((x : A) → PathP (B x) (f x) (g x)) +funExt⁻ eq x i = eq i x + +-- J for paths and its computation rule + +module _ (P : ∀ y → x ≡ y → Type ℓ') (d : P x refl) where + J : (p : x ≡ y) → P y p + J p = transport (λ i → P (p i) (λ j → p (i ∧ j))) d + + JRefl : J refl ≡ d + JRefl = transportRefl d + + J-∙ : (p : x ≡ y) (q : y ≡ z) + → J (p ∙ q) ≡ transport (λ i → P (q i) (λ j → compPath-filler p q i j)) (J p) + J-∙ p q k = + transp + (λ i → P (q (i ∨ ~ k)) + (λ j → compPath-filler p q (i ∨ ~ k) j)) (~ k) + (J (λ j → compPath-filler p q (~ k) j)) + +-- Converting to and from a PathP + +module _ {A : I → Type ℓ} {x : A i0} {y : A i1} where + toPathP : transport (λ i → A i) x ≡ y → PathP A x y + toPathP p i = hcomp (λ j → λ { (i = i0) → x + ; (i = i1) → p j }) + (transp (λ j → A (i ∧ j)) (~ i) x) + + fromPathP : PathP A x y → transport (λ i → A i) x ≡ y + fromPathP p i = transp (λ j → A (i ∨ j)) i (p i) + +-- Whiskering a dependent path by a path +-- Double whiskering +_◁_▷_ : ∀ {ℓ} {A : I → Type ℓ} {a₀ a₀' : A i0} {a₁ a₁' : A i1} + → a₀ ≡ a₀' → PathP A a₀' a₁ → a₁ ≡ a₁' + → PathP A a₀ a₁' +(p ◁ P ▷ q) i = + hcomp (λ j → λ {(i = i0) → p (~ j) ; (i = i1) → q j}) (P i) + +doubleWhiskFiller : + ∀ {ℓ} {A : I → Type ℓ} {a₀ a₀' : A i0} {a₁ a₁' : A i1} + → (p : a₀ ≡ a₀') (pq : PathP A a₀' a₁) (q : a₁ ≡ a₁') + → PathP (λ i → PathP A (p (~ i)) (q i)) + pq + (p ◁ pq ▷ q) +doubleWhiskFiller p pq q k i = + hfill (λ j → λ {(i = i0) → p (~ j) ; (i = i1) → q j}) + (inS (pq i)) + k + +_◁_ : ∀ {ℓ} {A : I → Type ℓ} {a₀ a₀' : A i0} {a₁ : A i1} + → a₀ ≡ a₀' → PathP A a₀' a₁ → PathP A a₀ a₁ +(p ◁ q) = p ◁ q ▷ refl + +_▷_ : ∀ {ℓ} {A : I → Type ℓ} {a₀ : A i0} {a₁ a₁' : A i1} + → PathP A a₀ a₁ → a₁ ≡ a₁' → PathP A a₀ a₁' +p ▷ q = refl ◁ p ▷ q + +-- Direct definitions of lower h-levels + +isContr : Type ℓ → Type ℓ +isContr A = Σ[ x ∈ A ] (∀ y → x ≡ y) + +isProp : Type ℓ → Type ℓ +isProp A = (x y : A) → x ≡ y + +isSet : Type ℓ → Type ℓ +isSet A = (x y : A) → isProp (x ≡ y) + +isGroupoid : Type ℓ → Type ℓ +isGroupoid A = ∀ a b → isSet (Path A a b) + +is2Groupoid : Type ℓ → Type ℓ +is2Groupoid A = ∀ a b → isGroupoid (Path A a b) + +-- Contractibility of singletons + +singlP : (A : I → Type ℓ) (a : A i0) → Type _ +singlP A a = Σ[ x ∈ A i1 ] PathP A a x + +singl : (a : A) → Type _ +singl {A = A} a = singlP (λ _ → A) a + +isContrSingl : (a : A) → isContr (singl a) +isContrSingl a .fst = (a , refl) +isContrSingl a .snd p i .fst = p .snd i +isContrSingl a .snd p i .snd j = p .snd (i ∧ j) + +isContrSinglP : (A : I → Type ℓ) (a : A i0) → isContr (singlP A a) +isContrSinglP A a .fst = _ , transport-filler (λ i → A i) a +isContrSinglP A a .snd (x , p) i = + _ , λ j → fill A (λ j → λ {(i = i0) → transport-filler (λ i → A i) a j; (i = i1) → p j}) (inS a) j + +-- Higher cube types + +SquareP : + (A : I → I → Type ℓ) + {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} (a₀₋ : PathP (λ j → A i0 j) a₀₀ a₀₁) + {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} (a₁₋ : PathP (λ j → A i1 j) a₁₀ a₁₁) + (a₋₀ : PathP (λ i → A i i0) a₀₀ a₁₀) (a₋₁ : PathP (λ i → A i i1) a₀₁ a₁₁) + → Type ℓ +SquareP A a₀₋ a₁₋ a₋₀ a₋₁ = PathP (λ i → PathP (λ j → A i j) (a₋₀ i) (a₋₁ i)) a₀₋ a₁₋ + +Square : + {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ ≡ a₀₁) + {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ ≡ a₁₁) + (a₋₀ : a₀₀ ≡ a₁₀) (a₋₁ : a₀₁ ≡ a₁₁) + → Type _ +Square a₀₋ a₁₋ a₋₀ a₋₁ = PathP (λ i → a₋₀ i ≡ a₋₁ i) a₀₋ a₁₋ + +Cube : + {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ ≡ a₀₀₁} + {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ ≡ a₀₁₁} + {a₀₋₀ : a₀₀₀ ≡ a₀₁₀} {a₀₋₁ : a₀₀₁ ≡ a₀₁₁} + (a₀₋₋ : Square a₀₀₋ a₀₁₋ a₀₋₀ a₀₋₁) + {a₁₀₀ a₁₀₁ : A} {a₁₀₋ : a₁₀₀ ≡ a₁₀₁} + {a₁₁₀ a₁₁₁ : A} {a₁₁₋ : a₁₁₀ ≡ a₁₁₁} + {a₁₋₀ : a₁₀₀ ≡ a₁₁₀} {a₁₋₁ : a₁₀₁ ≡ a₁₁₁} + (a₁₋₋ : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁) + {a₋₀₀ : a₀₀₀ ≡ a₁₀₀} {a₋₀₁ : a₀₀₁ ≡ a₁₀₁} + (a₋₀₋ : Square a₀₀₋ a₁₀₋ a₋₀₀ a₋₀₁) + {a₋₁₀ : a₀₁₀ ≡ a₁₁₀} {a₋₁₁ : a₀₁₁ ≡ a₁₁₁} + (a₋₁₋ : Square a₀₁₋ a₁₁₋ a₋₁₀ a₋₁₁) + (a₋₋₀ : Square a₀₋₀ a₁₋₀ a₋₀₀ a₋₁₀) + (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) + → Type _ +Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = + PathP (λ i → Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ + +-- Vertical composition of squares + +_∙₂_ : + {a₀₀ a₀₁ a₀₂ : A} {a₀₋ : a₀₀ ≡ a₀₁} {b₀₋ : a₀₁ ≡ a₀₂} + {a₁₀ a₁₁ a₁₂ : A} {a₁₋ : a₁₀ ≡ a₁₁} {b₁₋ : a₁₁ ≡ a₁₂} + {a₋₀ : a₀₀ ≡ a₁₀} {a₋₁ : a₀₁ ≡ a₁₁} {a₋₂ : a₀₂ ≡ a₁₂} + (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square b₀₋ b₁₋ a₋₁ a₋₂) + → Square (a₀₋ ∙ b₀₋) (a₁₋ ∙ b₁₋) a₋₀ a₋₂ +_∙₂_ = congP₂ (λ _ → _∙_) + +-- Alternative (equivalent) definitions of hlevel n that give fillers for n-cubes instead of n-globes + +isSet' : Type ℓ → Type ℓ +isSet' A = + {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ ≡ a₀₁) + {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ ≡ a₁₁) + (a₋₀ : a₀₀ ≡ a₁₀) (a₋₁ : a₀₁ ≡ a₁₁) + → Square a₀₋ a₁₋ a₋₀ a₋₁ + +isGroupoid' : Type ℓ → Type ℓ +isGroupoid' A = + {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ ≡ a₀₀₁} + {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ ≡ a₀₁₁} + {a₀₋₀ : a₀₀₀ ≡ a₀₁₀} {a₀₋₁ : a₀₀₁ ≡ a₀₁₁} + (a₀₋₋ : Square a₀₀₋ a₀₁₋ a₀₋₀ a₀₋₁) + {a₁₀₀ a₁₀₁ : A} {a₁₀₋ : a₁₀₀ ≡ a₁₀₁} + {a₁₁₀ a₁₁₁ : A} {a₁₁₋ : a₁₁₀ ≡ a₁₁₁} + {a₁₋₀ : a₁₀₀ ≡ a₁₁₀} {a₁₋₁ : a₁₀₁ ≡ a₁₁₁} + (a₁₋₋ : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁) + {a₋₀₀ : a₀₀₀ ≡ a₁₀₀} {a₋₀₁ : a₀₀₁ ≡ a₁₀₁} + (a₋₀₋ : Square a₀₀₋ a₁₀₋ a₋₀₀ a₋₀₁) + {a₋₁₀ : a₀₁₀ ≡ a₁₁₀} {a₋₁₁ : a₀₁₁ ≡ a₁₁₁} + (a₋₁₋ : Square a₀₁₋ a₁₁₋ a₋₁₀ a₋₁₁) + (a₋₋₀ : Square a₀₋₀ a₁₋₀ a₋₀₀ a₋₁₀) + (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) + → Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ + +-- Essential properties of isProp and isContr + +isProp→PathP : ∀ {B : I → Type ℓ} → ((i : I) → isProp (B i)) + → (b0 : B i0) (b1 : B i1) + → PathP B b0 b1 +isProp→PathP hB b0 b1 = toPathP (hB _ _ _) + +isPropIsContr : isProp (isContr A) +isPropIsContr (c0 , h0) (c1 , h1) j .fst = h0 c1 j +isPropIsContr (c0 , h0) (c1 , h1) j .snd y i = + hcomp (λ k → λ { (i = i0) → h0 (h0 c1 j) k; + (i = i1) → h0 y k; + (j = i0) → h0 (h0 y i) k; + (j = i1) → h0 (h1 y i) k}) + c0 + +isContr→isProp : isContr A → isProp A +isContr→isProp (x , p) a b = sym (p a) ∙ p b + +isProp→isSet : isProp A → isSet A +isProp→isSet h a b p q j i = + hcomp (λ k → λ { (i = i0) → h a a k + ; (i = i1) → h a b k + ; (j = i0) → h a (p i) k + ; (j = i1) → h a (q i) k }) a + +isProp→isSet' : isProp A → isSet' A +isProp→isSet' h {a} p q r s i j = + hcomp (λ k → λ { (i = i0) → h a (p j) k + ; (i = i1) → h a (q j) k + ; (j = i0) → h a (r i) k + ; (j = i1) → h a (s i) k}) a + +isPropIsProp : isProp (isProp A) +isPropIsProp f g i a b = isProp→isSet f a b (f a b) (g a b) i + +isPropSingl : {a : A} → isProp (singl a) +isPropSingl = isContr→isProp (isContrSingl _) + +isPropSinglP : {A : I → Type ℓ} {a : A i0} → isProp (singlP A a) +isPropSinglP = isContr→isProp (isContrSinglP _ _) + +-- Universe lifting + +record Lift {i j} (A : Type i) : Type (ℓ-max i j) where + constructor lift + field + lower : A + +open Lift public + +liftExt : ∀ {A : Type ℓ} {a b : Lift {ℓ} {ℓ'} A} → (lower a ≡ lower b) → a ≡ b +liftExt x i = lift (x i) diff --git a/Cubical/Foundations/RelationalStructure.agda b/Cubical/Foundations/RelationalStructure.agda new file mode 100644 index 0000000000..2b8320fce5 --- /dev/null +++ b/Cubical/Foundations/RelationalStructure.agda @@ -0,0 +1,268 @@ +{- + +Definition of what it means to be a notion of relational structure + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.RelationalStructure where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Univalence +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.HITs.SetQuotients +open import Cubical.Relation.Binary.Base +open import Cubical.Relation.ZigZag.Base + +open BinaryRelation +open isEquivRel +open isQuasiEquivRel + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + +-- A notion of structured relation for a structure S assigns a relation on S X and S Y to every relation on X +-- and Y. We require the output to be proposition-valued when the input is proposition-valued. +StrRel : (S : Type ℓ → Type ℓ') (ℓ'' : Level) → Type (ℓ-max (ℓ-suc (ℓ-max ℓ ℓ'')) ℓ') +StrRel {ℓ = ℓ} S ℓ'' = ∀ {A B} (R : Rel A B ℓ) → Rel (S A) (S B) ℓ'' + +-- Given a type A and relation R, a quotient structure is a structure on the set quotient A/R such that +-- the graph of [_] : A → A/R is a structured relation +InducedQuotientStr : (S : Type ℓ → Type ℓ') (ρ : StrRel S ℓ'') + (A : TypeWithStr ℓ S) (R : Rel (typ A) (typ A) ℓ) + → Type (ℓ-max ℓ' ℓ'') +InducedQuotientStr S ρ A R = + Σ[ s ∈ S (typ A / R) ] ρ (graphRel [_]) (A .snd) s + +-- A structured equivalence relation R on a structured type A should induce a structure on A/R +InducesQuotientStr : (S : Type ℓ → Type ℓ') (ρ : StrRel S ℓ'') → Type _ +InducesQuotientStr {ℓ = ℓ} S ρ = + (A : TypeWithStr ℓ S) (R : EquivPropRel (typ A) ℓ) + → ρ (R .fst .fst) (A .snd) (A .snd) + → ∃![ s ∈ S (typ A / R .fst .fst) ] ρ (graphRel [_]) (A .snd) s + +-- The identity should be a structured relation +isReflexiveStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +isReflexiveStrRel {ℓ = ℓ} {S = S} ρ = + {X : Type ℓ} → (s : S X) → ρ (idPropRel X .fst) s s + +-- The inverse of a structured relation should be structured +isSymmetricStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +isSymmetricStrRel {ℓ = ℓ} {S = S} ρ = + {X Y : Type ℓ} (R : PropRel X Y ℓ) + {sx : S X} {sy : S Y} + → ρ (R .fst) sx sy + → ρ (invPropRel R .fst) sy sx + +-- The composite of structured relations should be structured +isTransitiveStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +isTransitiveStrRel {ℓ = ℓ} {S = S} ρ = + {X Y Z : Type ℓ} + (R : PropRel X Y ℓ) (R' : PropRel Y Z ℓ) + {sx : S X} {sy : S Y} {sz : S Z} + → ρ (R .fst) sx sy + → ρ (R' .fst) sy sz + → ρ (compPropRel R R' .fst) sx sz + +-- The type of structures on a set should be a set +preservesSetsStr : (S : Type ℓ → Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +preservesSetsStr S = ∀ {X} → isSet X → isSet (S X) + +-- The type of structures on a prop-valued relation should be a prop +preservesPropsStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +preservesPropsStrRel {ℓ = ℓ} {S = S} ρ = + {X Y : Type ℓ} {R : Rel X Y ℓ} + → (∀ x y → isProp (R x y)) + → (sx : S X) (sy : S Y) + → isProp (ρ R sx sy) + +record SuitableStrRel (S : Type ℓ → Type ℓ') (ρ : StrRel S ℓ'') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) ℓ') ℓ'') + where + field + quo : InducesQuotientStr S ρ + symmetric : isSymmetricStrRel ρ + transitive : isTransitiveStrRel ρ + set : preservesSetsStr S + prop : preservesPropsStrRel ρ + +open SuitableStrRel public + +quotientPropRel : ∀ {ℓ} {A : Type ℓ} (R : Rel A A ℓ) → PropRel A (A / R) ℓ +quotientPropRel R .fst a t = [ a ] ≡ t +quotientPropRel R .snd _ _ = squash/ _ _ + +-- We can also ask for a notion of structured relations to agree with some notion of structured equivalences. + +StrRelMatchesEquiv : {S : Type ℓ → Type ℓ'} + → StrRel S ℓ'' → StrEquiv S ℓ''' → Type _ +StrRelMatchesEquiv {S = S} ρ ι = + (A B : TypeWithStr _ S) (e : typ A ≃ typ B) → + ρ (graphRel (e .fst)) (A .snd) (B .snd) ≃ ι A B e + +-- Additional conditions for a "positive" notion of structured relation + +isDetransitiveStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +isDetransitiveStrRel {ℓ = ℓ} {S = S} ρ = + {X Y Z : Type ℓ} + (R : PropRel X Y ℓ) (R' : PropRel Y Z ℓ) + {sx : S X} {sz : S Z} + → ρ (compPropRel R R' .fst) sx sz + → ∃[ sy ∈ S Y ] ρ (R .fst) sx sy × ρ (R' .fst) sy sz + +record StrRelAction {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') + : Type (ℓ-max (ℓ-suc ℓ) (ℓ-max ℓ' ℓ'')) + where + field + actStr : ∀ {X Y} → (X → Y) → S X → S Y + actStrId : ∀ {X} (s : S X) → actStr (idfun X) s ≡ s + actRel : ∀ {X₀ Y₀ X₁ Y₁} {f : X₀ → X₁} {g : Y₀ → Y₁} + {R₀ : X₀ → Y₀ → Type ℓ} {R₁ : X₁ → Y₁ → Type ℓ} + → (∀ x y → R₀ x y → R₁ (f x) (g y)) + → ∀ sx sy → ρ R₀ sx sy → ρ R₁ (actStr f sx) (actStr g sy) + +open StrRelAction public + +strRelQuotientComparison : {S : Type ℓ → Type ℓ'} {ρ : StrRel S ℓ''} + (θ : SuitableStrRel S ρ) (α : StrRelAction ρ) + {X : Type ℓ} (R : EquivPropRel X ℓ) + → (S X / ρ (R .fst .fst)) → S (X / R .fst .fst) +strRelQuotientComparison θ α R [ s ] = α .actStr [_] s +strRelQuotientComparison {ρ = ρ} θ α R (eq/ s t r i) = + (sym leftEq ∙ rightEq) i + where + ρs : ρ (R .fst .fst) s s + ρs = + subst (λ R' → ρ R' s s) + (funExt₂ λ x x' → + hPropExt squash (R .fst .snd x x') + (Trunc.rec (R .fst .snd x x') + (λ {(_ , r , r') → R .snd .transitive _ _ _ r (R .snd .symmetric _ _ r')})) + (λ r → ∣ x' , r , R .snd .reflexive x' ∣)) + (θ .transitive (R .fst) (invPropRel (R .fst)) r (θ .symmetric (R .fst) r)) + + leftEq : θ .quo (_ , s) R ρs .fst .fst ≡ α .actStr [_] s + leftEq = + cong fst + (θ .quo (_ , s) R ρs .snd + ( α .actStr [_] s + , subst + (λ s' → ρ (graphRel [_]) s' (α .actStr [_] s)) + (α .actStrId s) + (α .actRel eq/ s s ρs) + )) + + rightEq : θ .quo (_ , s) R ρs .fst .fst ≡ α .actStr [_] t + rightEq = + cong fst + (θ .quo (_ , s) R ρs .snd + ( α .actStr [_] t + , subst + (λ s' → ρ (graphRel [_]) s' (α .actStr [_] t)) + (α .actStrId s) + (α .actRel eq/ s t r) + )) +strRelQuotientComparison θ α R (squash/ _ _ p q i j) = + θ .set squash/ _ _ + (cong (strRelQuotientComparison θ α R) p) + (cong (strRelQuotientComparison θ α R) q) + i j + +record PositiveStrRel {S : Type ℓ → Type ℓ'} {ρ : StrRel S ℓ''} (θ : SuitableStrRel S ρ) + : Type (ℓ-max (ℓ-suc ℓ) (ℓ-max ℓ' ℓ'')) + where + field + act : StrRelAction ρ + reflexive : isReflexiveStrRel ρ + detransitive : isDetransitiveStrRel ρ + quo : {X : Type ℓ} (R : EquivPropRel X ℓ) → isEquiv (strRelQuotientComparison θ act R) + +open PositiveStrRel public + +posRelReflexive : {S : Type ℓ → Type ℓ'} {ρ : StrRel S ℓ''} {θ : SuitableStrRel S ρ} + → PositiveStrRel θ + → {X : Type ℓ} (R : EquivPropRel X ℓ) + → (s : S X) → ρ (R .fst .fst) s s +posRelReflexive {ρ = ρ} σ R s = + subst + (uncurry (ρ (R .fst .fst))) + (ΣPathP (σ .act .actStrId s , σ .act .actStrId s)) + (σ .act .actRel + (λ x y → + Trunc.rec (R .fst .snd _ _) + (λ p → subst (R .fst .fst x) p (R .snd .reflexive x))) + s s + (σ .reflexive s)) + +-- Given a suitable notion of structured relation, if we have a structured quasi equivalence relation R +-- between structured types A and B, we get induced structures on the quotients A/(R ∙ R⁻¹) and B/(R⁻¹ ∙ R), +-- and the induced equivalence e : A/(R ∙ R⁻¹) ≃ B/(R⁻¹ ∙ R) is structured with respect to those quotient +-- structures. + +record QERDescends (S : Type ℓ → Type ℓ') (ρ : StrRel S ℓ'') + (A B : TypeWithStr ℓ S) (R : QuasiEquivRel (typ A) (typ B) ℓ) : Type (ℓ-max ℓ' ℓ'') + where + private + module E = QER→Equiv R + + field + quoᴸ : InducedQuotientStr S ρ A E.Rᴸ + quoᴿ : InducedQuotientStr S ρ B E.Rᴿ + rel : ρ (graphRel (E.Thm .fst)) (quoᴸ .fst) (quoᴿ .fst) + +open QERDescends + +structuredQER→structuredEquiv : {S : Type ℓ → Type ℓ'} {ρ : StrRel S ℓ''} + (θ : SuitableStrRel S ρ) + (A B : TypeWithStr ℓ S) (R : QuasiEquivRel (typ A) (typ B) ℓ) + → ρ (R .fst .fst) (A .snd) (B .snd) + → QERDescends S ρ A B R +structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .quoᴸ = + θ .quo (X , s) (QER→EquivRel R) + (θ .transitive (R .fst) (invPropRel (R .fst)) r (θ .symmetric (R .fst) r)) + .fst + +structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .quoᴿ = + θ .quo (Y , t) (QER→EquivRel (invQER R)) + (θ .transitive (invPropRel (R .fst)) (R .fst) (θ .symmetric (R .fst) r) r) + .fst + +structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .rel = + subst (λ R' → ρ R' (quol .fst) (quor .fst)) correction + (θ .transitive (compPropRel (invPropRel (quotientPropRel E.Rᴸ)) (R .fst)) (quotientPropRel E.Rᴿ) + (θ .transitive (invPropRel (quotientPropRel E.Rᴸ)) (R .fst) + (θ .symmetric (quotientPropRel E.Rᴸ) (quol .snd)) + r) + (quor .snd)) + where + module E = QER→Equiv R + quol = structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .quoᴸ + quor = structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .quoᴿ + [R] = compPropRel (compPropRel (invPropRel (quotientPropRel E.Rᴸ)) (R .fst)) (quotientPropRel E.Rᴿ) + + correction : [R] .fst ≡ graphRel (E.Thm .fst) + correction = + funExt₂ λ qx qy → + (hPropExt squash (squash/ _ _) + (Trunc.rec (squash/ _ _) + (λ {(y , qr , py) → + Trunc.rec + (squash/ _ _) + (λ {(x , px , r) → + (cong (E.Thm .fst) (sym px) + ∙ E.relToFwd≡ r) + ∙ py}) + qr})) + (elimProp {P = λ qx → E.Thm .fst qx ≡ qy → [R] .fst qx qy} + (λ _ → isPropΠ λ _ → squash) + (λ x → + elimProp {P = λ qy → E.Thm .fst [ x ] ≡ qy → [R] .fst [ x ] qy} + (λ _ → isPropΠ λ _ → squash) + (λ y p → ∣ _ , ∣ _ , refl , E.fwd≡ToRel p ∣ , refl ∣) + qy) + qx)) diff --git a/Cubical/Foundations/SIP.agda b/Cubical/Foundations/SIP.agda new file mode 100644 index 0000000000..b68515407f --- /dev/null +++ b/Cubical/Foundations/SIP.agda @@ -0,0 +1,124 @@ +{- + +In this file we apply the cubical machinery to Martin Hötzel-Escardó's +structure identity principle: + +https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#sns + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.SIP where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence renaming (ua-pathToEquiv to ua-pathToEquiv') +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Structure public + +private + variable + ℓ ℓ₁ ℓ₂ ℓ₃ ℓ₄ ℓ₅ : Level + S : Type ℓ₁ → Type ℓ₂ + +-- Note that for any equivalence (f , e) : X ≃ Y the type ι (X , s) (Y , t) (f , e) need not to be +-- a proposition. Indeed this type should correspond to the ways s and t can be identified +-- as S-structures. This we call a standard notion of structure or SNS. +-- We will use a different definition, but the two definitions are interchangeable. +SNS : (S : Type ℓ₁ → Type ℓ₂) (ι : StrEquiv S ℓ₃) → Type (ℓ-max (ℓ-max (ℓ-suc ℓ₁) ℓ₂) ℓ₃) +SNS {ℓ₁} S ι = ∀ {X : Type ℓ₁} (s t : S X) → ι (X , s) (X , t) (idEquiv X) ≃ (s ≡ t) + +-- We introduce the notation for structure preserving equivalences a +-- bit differently, but this definition doesn't actually change from +-- Escardó's notes. +_≃[_]_ : (A : TypeWithStr ℓ₁ S) (ι : StrEquiv S ℓ₂) (B : TypeWithStr ℓ₁ S) → Type (ℓ-max ℓ₁ ℓ₂) +A ≃[ ι ] B = Σ[ e ∈ typ A ≃ typ B ] (ι A B e) + + + +-- The following PathP version of SNS is a bit easier to work with +-- for the proof of the SIP +UnivalentStr : (S : Type ℓ₁ → Type ℓ₂) (ι : StrEquiv S ℓ₃) → Type (ℓ-max (ℓ-max (ℓ-suc ℓ₁) ℓ₂) ℓ₃) +UnivalentStr {ℓ₁} S ι = + {A B : TypeWithStr ℓ₁ S} (e : typ A ≃ typ B) + → ι A B e ≃ PathP (λ i → S (ua e i)) (str A) (str B) + +-- A quick sanity-check that our definition is interchangeable with +-- Escardó's. The direction SNS→UnivalentStr corresponds more or less +-- to a dependent EquivJ formulation of Escardó's homomorphism-lemma. +UnivalentStr→SNS : (S : Type ℓ₁ → Type ℓ₂) (ι : StrEquiv S ℓ₃) + → UnivalentStr S ι → SNS S ι +UnivalentStr→SNS S ι θ {X = X} s t = + ι (X , s) (X , t) (idEquiv X) + ≃⟨ θ (idEquiv X) ⟩ + PathP (λ i → S (ua (idEquiv X) i)) s t + ≃⟨ transportEquiv (λ j → PathP (λ i → S (uaIdEquiv {A = X} j i)) s t) ⟩ + s ≡ t + ■ + + +SNS→UnivalentStr : (ι : StrEquiv S ℓ₃) → SNS S ι → UnivalentStr S ι +SNS→UnivalentStr {S = S} ι θ {A = A} {B = B} e = EquivJ P C e (str A) (str B) + where + Y = typ B + + P : (X : Type _) → X ≃ Y → Type _ + P X e' = (s : S X) (t : S Y) → ι (X , s) (Y , t) e' ≃ PathP (λ i → S (ua e' i)) s t + + C : (s t : S Y) → ι (Y , s) (Y , t) (idEquiv Y) ≃ PathP (λ i → S (ua (idEquiv Y) i)) s t + C s t = + ι (Y , s) (Y , t) (idEquiv Y) + ≃⟨ θ s t ⟩ + s ≡ t + ≃⟨ transportEquiv (λ j → PathP (λ i → S (uaIdEquiv {A = Y} (~ j) i)) s t) ⟩ + PathP (λ i → S (ua (idEquiv Y) i)) s t + ■ + +TransportStr : {S : Type ℓ → Type ℓ₁} (α : EquivAction S) → Type (ℓ-max (ℓ-suc ℓ) ℓ₁) +TransportStr {ℓ} {S = S} α = + {X Y : Type ℓ} (e : X ≃ Y) (s : S X) → equivFun (α e) s ≡ subst S (ua e) s + +TransportStr→UnivalentStr : {S : Type ℓ → Type ℓ₁} (α : EquivAction S) + → TransportStr α → UnivalentStr S (EquivAction→StrEquiv α) +TransportStr→UnivalentStr {S = S} α τ {X , s} {Y , t} e = + equivFun (α e) s ≡ t + ≃⟨ pathToEquiv (cong (_≡ t) (τ e s)) ⟩ + subst S (ua e) s ≡ t + ≃⟨ invEquiv (PathP≃Path _ _ _) ⟩ + PathP (λ i → S (ua e i)) s t + ■ + +UnivalentStr→TransportStr : {S : Type ℓ → Type ℓ₁} (α : EquivAction S) + → UnivalentStr S (EquivAction→StrEquiv α) → TransportStr α +UnivalentStr→TransportStr {S = S} α θ e s = + invEq (θ e) (transport-filler (cong S (ua e)) s) + +invTransportStr : {S : Type ℓ → Type ℓ₂} (α : EquivAction S) (τ : TransportStr α) + {X Y : Type ℓ} (e : X ≃ Y) (t : S Y) → invEq (α e) t ≡ subst⁻ S (ua e) t +invTransportStr {S = S} α τ e t = + sym (transport⁻Transport (cong S (ua e)) (invEq (α e) t)) + ∙∙ sym (cong (subst⁻ S (ua e)) (τ e (invEq (α e) t))) + ∙∙ cong (subst⁻ S (ua e)) (secEq (α e) t) + + +--- We can now define an invertible function +--- +--- sip : A ≃[ ι ] B → A ≡ B + +module _ {S : Type ℓ₁ → Type ℓ₂} {ι : StrEquiv S ℓ₃} + (θ : UnivalentStr S ι) (A B : TypeWithStr ℓ₁ S) + where + + sip : A ≃[ ι ] B → A ≡ B + sip (e , p) i = ua e i , θ e .fst p i + + SIP : A ≃[ ι ] B ≃ (A ≡ B) + SIP = + sip , isoToIsEquiv (compIso (Σ-cong-iso (invIso univalenceIso) (equivToIso ∘ θ)) ΣPathIsoPathΣ) + + sip⁻ : A ≡ B → A ≃[ ι ] B + sip⁻ = invEq SIP diff --git a/Cubical/Foundations/Structure.agda b/Cubical/Foundations/Structure.agda new file mode 100644 index 0000000000..6de3bf9e1b --- /dev/null +++ b/Cubical/Foundations/Structure.agda @@ -0,0 +1,49 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Structure where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Syntax.⟨⟩ + +private + variable + ℓ ℓ' ℓ'' : Level + S : Type ℓ → Type ℓ' + +-- A structure is a type-family S : Type ℓ → Type ℓ', i.e. for X : Type ℓ and s : S X, +-- the pair (X , s) : TypeWithStr ℓ S means that X is equipped with a S-structure, witnessed by s. + +TypeWithStr : (ℓ : Level) (S : Type ℓ → Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +TypeWithStr ℓ S = Σ[ X ∈ Type ℓ ] S X + +typ : TypeWithStr ℓ S → Type ℓ +typ = fst + +str : (A : TypeWithStr ℓ S) → S (typ A) +str = snd + +instance + TypeWithStr-has-⟨⟩ : ∀ {ℓ S} → has-⟨⟩ (TypeWithStr {ℓ' = ℓ'} ℓ S) + ⟨_⟩ ⦃ TypeWithStr-has-⟨⟩ ⦄ = typ + +-- Allow users to avoid importing the syntax module directly for +-- backwards compatibility. +open import Cubical.Syntax.⟨⟩ using (⟨_⟩) public + +-- An S-structure should have a notion of S-homomorphism, or rather S-isomorphism. +-- This will be implemented by a function ι : StrEquiv S ℓ' +-- that gives us for any two types with S-structure (X , s) and (Y , t) a family: +-- ι (X , s) (Y , t) : (X ≃ Y) → Type ℓ'' +StrEquiv : (S : Type ℓ → Type ℓ'') (ℓ' : Level) → Type (ℓ-max (ℓ-suc (ℓ-max ℓ ℓ')) ℓ'') +StrEquiv {ℓ} S ℓ' = (A B : TypeWithStr ℓ S) → typ A ≃ typ B → Type ℓ' + +-- An S-structure may instead be equipped with an action on equivalences, which will +-- induce a notion of S-homomorphism + +EquivAction : (S : Type ℓ → Type ℓ'') → Type (ℓ-max (ℓ-suc ℓ) ℓ'') +EquivAction {ℓ} S = {X Y : Type ℓ} → X ≃ Y → S X ≃ S Y + +EquivAction→StrEquiv : {S : Type ℓ → Type ℓ''} + → EquivAction S → StrEquiv S ℓ'' +EquivAction→StrEquiv α (X , s) (Y , t) e = equivFun (α e) s ≡ t + diff --git a/Cubical/Foundations/Transport.agda b/Cubical/Foundations/Transport.agda new file mode 100644 index 0000000000..502e8960bc --- /dev/null +++ b/Cubical/Foundations/Transport.agda @@ -0,0 +1,154 @@ +{- Basic theory about transport: + +- transport is invertible +- transport is an equivalence ([transportEquiv]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Transport where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function using (_∘_) + +-- Direct definition of transport filler, note that we have to +-- explicitly tell Agda that the type is constant (like in CHM) +transpFill : ∀ {ℓ} {A : Type ℓ} + (φ : I) + (A : (i : I) → Type ℓ [ φ ↦ (λ _ → A) ]) + (u0 : outS (A i0)) + → -------------------------------------- + PathP (λ i → outS (A i)) u0 (transp (λ i → outS (A i)) φ u0) +transpFill φ A u0 i = transp (λ j → outS (A (i ∧ j))) (~ i ∨ φ) u0 + +transport⁻ : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → B → A +transport⁻ p = transport (λ i → p (~ i)) + +subst⁻ : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} (B : A → Type ℓ') (p : x ≡ y) → B y → B x +subst⁻ B p pa = transport⁻ (λ i → B (p i)) pa + +transport-fillerExt : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) + → PathP (λ i → A → p i) (λ x → x) (transport p) +transport-fillerExt p i x = transport-filler p x i + +transport⁻-fillerExt : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) + → PathP (λ i → p i → A) (λ x → x) (transport⁻ p) +transport⁻-fillerExt p i x = transp (λ j → p (i ∧ ~ j)) (~ i) x + +transport-fillerExt⁻ : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) + → PathP (λ i → p i → B) (transport p) (λ x → x) +transport-fillerExt⁻ p = symP (transport⁻-fillerExt (sym p)) + +transport⁻-fillerExt⁻ : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) + → PathP (λ i → B → p i) (transport⁻ p) (λ x → x) +transport⁻-fillerExt⁻ p = symP (transport-fillerExt (sym p)) + + +transport⁻-filler : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) (x : B) + → PathP (λ i → p (~ i)) x (transport⁻ p x) +transport⁻-filler p x = transport-filler (λ i → p (~ i)) x + +transport⁻Transport : ∀ {ℓ} {A B : Type ℓ} → (p : A ≡ B) → (a : A) → + transport⁻ p (transport p a) ≡ a +transport⁻Transport p a j = transport⁻-fillerExt p (~ j) (transport-fillerExt p (~ j) a) + +transportTransport⁻ : ∀ {ℓ} {A B : Type ℓ} → (p : A ≡ B) → (b : B) → + transport p (transport⁻ p b) ≡ b +transportTransport⁻ p b j = transport-fillerExt⁻ p j (transport⁻-fillerExt⁻ p j b) + +-- Transport is an equivalence +isEquivTransport : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) → isEquiv (transport p) +isEquivTransport {A = A} {B = B} p = + transport (λ i → isEquiv (transport-fillerExt p i)) (idIsEquiv A) + +transportEquiv : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → A ≃ B +transportEquiv p = (transport p , isEquivTransport p) + +substEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {a a' : A} (P : A → Type ℓ') (p : a ≡ a') → P a ≃ P a' +substEquiv P p = (subst P p , isEquivTransport (λ i → P (p i))) + +liftEquiv : ∀ {ℓ ℓ'} {A B : Type ℓ} (P : Type ℓ → Type ℓ') (e : A ≃ B) → P A ≃ P B +liftEquiv P e = substEquiv P (ua e) + +transpEquiv : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) → ∀ i → p i ≃ B +transpEquiv P i .fst = transp (λ j → P (i ∨ j)) i +transpEquiv P i .snd + = transp (λ k → isEquiv (transp (λ j → P (i ∨ (j ∧ k))) (i ∨ ~ k))) + i (idIsEquiv (P i)) + +uaTransportη : ∀ {ℓ} {A B : Type ℓ} (P : A ≡ B) → ua (transportEquiv P) ≡ P +uaTransportη P i j + = Glue (P i1) λ where + (j = i0) → P i0 , transportEquiv P + (i = i1) → P j , transpEquiv P j + (j = i1) → P i1 , idEquiv (P i1) + +pathToIso : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → Iso A B +Iso.fun (pathToIso x) = transport x +Iso.inv (pathToIso x) = transport⁻ x +Iso.rightInv (pathToIso x) = transportTransport⁻ x +Iso.leftInv (pathToIso x) = transport⁻Transport x + +isInjectiveTransport : ∀ {ℓ : Level} {A B : Type ℓ} {p q : A ≡ B} + → transport p ≡ transport q → p ≡ q +isInjectiveTransport {p = p} {q} α i = + hcomp + (λ j → λ + { (i = i0) → retEq univalence p j + ; (i = i1) → retEq univalence q j + }) + (invEq univalence ((λ a → α i a) , t i)) + where + t : PathP (λ i → isEquiv (λ a → α i a)) (pathToEquiv p .snd) (pathToEquiv q .snd) + t = isProp→PathP (λ i → isPropIsEquiv (λ a → α i a)) _ _ + +transportUaInv : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → transport (ua (invEquiv e)) ≡ transport (sym (ua e)) +transportUaInv e = cong transport (uaInvEquiv e) +-- notice that transport (ua e) would reduce, thus an alternative definition using EquivJ can give +-- refl for the case of idEquiv: +-- transportUaInv e = EquivJ (λ _ e → transport (ua (invEquiv e)) ≡ transport (sym (ua e))) refl e + +isSet-subst : ∀ {ℓ ℓ′} {A : Type ℓ} {B : A → Type ℓ′} + → (isSet-A : isSet A) + → ∀ {a : A} + → (p : a ≡ a) → (x : B a) → subst B p x ≡ x +isSet-subst {B = B} isSet-A p x = subst (λ p′ → subst B p′ x ≡ x) (isSet-A _ _ refl p) (substRefl {B = B} x) + +-- substituting along a composite path is equivalent to substituting twice +substComposite : ∀ {ℓ ℓ′} {A : Type ℓ} → (B : A → Type ℓ′) + → {x y z : A} (p : x ≡ y) (q : y ≡ z) (u : B x) + → subst B (p ∙ q) u ≡ subst B q (subst B p u) +substComposite B p q Bx i = + transport (cong B (compPath-filler' p q (~ i))) (transport-fillerExt (cong B p) i Bx) + +-- transporting along a composite path is equivalent to transporting twice +transportComposite : ∀ {ℓ} {A B C : Type ℓ} (p : A ≡ B) (q : B ≡ C) (x : A) + → transport (p ∙ q) x ≡ transport q (transport p x) +transportComposite = substComposite (λ D → D) + +-- substitution commutes with morphisms in slices +substCommSlice : ∀ {ℓ ℓ′} {A : Type ℓ} + → (B C : A → Type ℓ′) + → (F : ∀ i → B i → C i) + → {x y : A} (p : x ≡ y) (u : B x) + → subst C p (F x u) ≡ F y (subst B p u) +substCommSlice B C F p Bx i = + transport-fillerExt⁻ (cong C p) i (F _ (transport-fillerExt (cong B p) i Bx)) + +-- transporting over (λ i → B (p i) → C (p i)) divides the transport into +-- transports over (λ i → C (p i)) and (λ i → B (p (~ i))) +funTypeTransp : ∀ {ℓ ℓ'} {A : Type ℓ} (B C : A → Type ℓ') {x y : A} (p : x ≡ y) (f : B x → C x) + → PathP (λ i → B (p i) → C (p i)) f (subst C p ∘ f ∘ subst B (sym p)) +funTypeTransp B C {x = x} p f i b = + transp (λ j → C (p (j ∧ i))) (~ i) (f (transp (λ j → B (p (i ∧ ~ j))) (~ i) b)) + +-- transports between loop spaces preserve path composition +overPathFunct : ∀ {ℓ} {A : Type ℓ} {x y : A} (p q : x ≡ x) (P : x ≡ y) + → transport (λ i → P i ≡ P i) (p ∙ q) + ≡ transport (λ i → P i ≡ P i) p ∙ transport (λ i → P i ≡ P i) q +overPathFunct p q = + J (λ y P → transport (λ i → P i ≡ P i) (p ∙ q) ≡ transport (λ i → P i ≡ P i) p ∙ transport (λ i → P i ≡ P i) q) + (transportRefl (p ∙ q) ∙ cong₂ _∙_ (sym (transportRefl p)) (sym (transportRefl q))) diff --git a/Cubical/Foundations/Univalence.agda b/Cubical/Foundations/Univalence.agda new file mode 100644 index 0000000000..b113d83bd6 --- /dev/null +++ b/Cubical/Foundations/Univalence.agda @@ -0,0 +1,309 @@ +{- + +Proof of the standard formulation of the univalence theorem and +various consequences of univalence + +- Re-exports Glue types from Cubical.Core.Glue +- The ua constant and its computation rule (up to a path) +- Proof of univalence using that unglue is an equivalence ([EquivContr]) +- Equivalence induction ([EquivJ], [elimEquiv]) +- Univalence theorem ([univalence]) +- The computation rule for ua ([uaβ]) +- Isomorphism induction ([elimIso]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Univalence where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Sigma.Base + +open import Cubical.Core.Glue public + using ( Glue ; glue ; unglue ; lineToEquiv ) + +open import Cubical.Reflection.StrictEquiv + +private + variable + ℓ ℓ' : Level + +-- The ua constant +ua : ∀ {A B : Type ℓ} → A ≃ B → A ≡ B +ua {A = A} {B = B} e i = Glue B (λ { (i = i0) → (A , e) + ; (i = i1) → (B , idEquiv B) }) + +uaIdEquiv : {A : Type ℓ} → ua (idEquiv A) ≡ refl +uaIdEquiv {A = A} i j = Glue A {φ = i ∨ ~ j ∨ j} (λ _ → A , idEquiv A) + +-- Propositional extensionality +hPropExt : {A B : Type ℓ} → isProp A → isProp B → (A → B) → (B → A) → A ≡ B +hPropExt Aprop Bprop f g = ua (propBiimpl→Equiv Aprop Bprop f g) + +-- the unglue and glue primitives specialized to the case of ua + +ua-unglue : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : ua e i) + → B {- [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → x }) ] -} +ua-unglue e i x = unglue (i ∨ ~ i) x + +ua-glue : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : Partial (~ i) A) + (y : B [ _ ↦ (λ { (i = i0) → e .fst (x 1=1) }) ]) + → ua e i {- [ _ ↦ (λ { (i = i0) → x 1=1 ; (i = i1) → outS y }) ] -} +ua-glue e i x y = glue {φ = i ∨ ~ i} (λ { (i = i0) → x 1=1 ; (i = i1) → outS y }) (outS y) + +module _ {A B : Type ℓ} (e : A ≃ B) {x : A} {y : B} where + -- sometimes more useful are versions of these functions with the (i : I) factored in + + ua-ungluePath : PathP (λ i → ua e i) x y → e .fst x ≡ y + ua-ungluePath p i = ua-unglue e i (p i) + + ua-gluePath : e .fst x ≡ y → PathP (λ i → ua e i) x y + ua-gluePath p i = ua-glue e i (λ { (i = i0) → x }) (inS (p i)) + + -- ua-ungluePath and ua-gluePath are definitional inverses + ua-ungluePath-Equiv : (PathP (λ i → ua e i) x y) ≃ (e .fst x ≡ y) + unquoteDef ua-ungluePath-Equiv = + defStrictEquiv ua-ungluePath-Equiv ua-ungluePath ua-gluePath + +-- ua-unglue and ua-glue are also definitional inverses, in a way +-- strengthening the types of ua-unglue and ua-glue gives a nicer formulation of this, see below + +ua-unglue-glue : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : Partial (~ i) A) (y : B [ _ ↦ _ ]) + → ua-unglue e i (ua-glue e i x y) ≡ outS y +ua-unglue-glue _ _ _ _ = refl + +ua-glue-unglue : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : ua e i) + → ua-glue e i (λ { (i = i0) → x }) (inS (ua-unglue e i x)) ≡ x +ua-glue-unglue _ _ _ = refl + +-- mainly for documentation purposes, ua-unglue and ua-glue wrapped in cubical subtypes + +ua-unglueS : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) (y : B) + → ua e i [ _ ↦ (λ { (i = i0) → x ; (i = i1) → y }) ] + → B [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → y }) ] +ua-unglueS e i x y s = inS (ua-unglue e i (outS s)) + +ua-glueS : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) (y : B) + → B [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → y }) ] + → ua e i [ _ ↦ (λ { (i = i0) → x ; (i = i1) → y }) ] +ua-glueS e i x y s = inS (ua-glue e i (λ { (i = i0) → x }) (inS (outS s))) + +ua-unglueS-glueS : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) (y : B) + (s : B [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → y }) ]) + → outS (ua-unglueS e i x y (ua-glueS e i x y s)) ≡ outS s +ua-unglueS-glueS _ _ _ _ _ = refl + +ua-glueS-unglueS : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) (y : B) + (s : ua e i [ _ ↦ (λ { (i = i0) → x ; (i = i1) → y }) ]) + → outS (ua-glueS e i x y (ua-unglueS e i x y s)) ≡ outS s +ua-glueS-unglueS _ _ _ _ _ = refl + + +-- a version of ua-glue with a single endpoint, identical to `ua-gluePath e {x} refl i` +ua-gluePt : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) + → ua e i {- [ _ ↦ (λ { (i = i0) → x ; (i = i1) → e .fst x }) ] -} +ua-gluePt e i x = ua-glue e i (λ { (i = i0) → x }) (inS (e .fst x)) + + +-- Proof of univalence using that unglue is an equivalence: + +-- unglue is an equivalence +unglueIsEquiv : ∀ (A : Type ℓ) (φ : I) + (f : PartialP φ (λ o → Σ[ T ∈ Type ℓ ] T ≃ A)) → + isEquiv {A = Glue A f} (unglue φ) +equiv-proof (unglueIsEquiv A φ f) = λ (b : A) → + let u : I → Partial φ A + u i = λ{ (φ = i1) → equivCtr (f 1=1 .snd) b .snd (~ i) } + ctr : fiber (unglue φ) b + ctr = ( glue (λ { (φ = i1) → equivCtr (f 1=1 .snd) b .fst }) (hcomp u b) + , λ j → hfill u (inS b) (~ j)) + in ( ctr + , λ (v : fiber (unglue φ) b) i → + let u' : I → Partial (φ ∨ ~ i ∨ i) A + u' j = λ { (φ = i1) → equivCtrPath (f 1=1 .snd) b v i .snd (~ j) + ; (i = i0) → hfill u (inS b) j + ; (i = i1) → v .snd (~ j) } + in ( glue (λ { (φ = i1) → equivCtrPath (f 1=1 .snd) b v i .fst }) (hcomp u' b) + , λ j → hfill u' (inS b) (~ j))) + +-- Any partial family of equivalences can be extended to a total one +-- from Glue [ φ ↦ (T,f) ] A to A +unglueEquiv : ∀ (A : Type ℓ) (φ : I) + (f : PartialP φ (λ o → Σ[ T ∈ Type ℓ ] T ≃ A)) → + (Glue A f) ≃ A +unglueEquiv A φ f = ( unglue φ , unglueIsEquiv A φ f ) + + +-- The following is a formulation of univalence proposed by Martín Escardó: +-- https://groups.google.com/forum/#!msg/homotopytypetheory/HfCB_b-PNEU/Ibb48LvUMeUJ +-- See also Theorem 5.8.4 of the HoTT Book. +-- +-- The reason we have this formulation in the core library and not the +-- standard one is that this one is more direct to prove using that +-- unglue is an equivalence. The standard formulation can be found in +-- Cubical/Basics/Univalence. +-- +EquivContr : ∀ (A : Type ℓ) → ∃![ T ∈ Type ℓ ] (T ≃ A) +EquivContr {ℓ = ℓ} A = + ( (A , idEquiv A) + , idEquiv≡ ) + where + idEquiv≡ : (y : Σ (Type ℓ) (λ T → T ≃ A)) → (A , idEquiv A) ≡ y + idEquiv≡ w = \ { i .fst → Glue A (f i) + ; i .snd .fst → unglueEquiv _ _ (f i) .fst + ; i .snd .snd .equiv-proof → unglueEquiv _ _ (f i) .snd .equiv-proof + } + where + f : ∀ i → PartialP (~ i ∨ i) (λ x → Σ[ T ∈ Type ℓ ] T ≃ A) + f i = λ { (i = i0) → A , idEquiv A ; (i = i1) → w } + +contrSinglEquiv : {A B : Type ℓ} (e : A ≃ B) → (B , idEquiv B) ≡ (A , e) +contrSinglEquiv {A = A} {B = B} e = + isContr→isProp (EquivContr B) (B , idEquiv B) (A , e) + +-- Equivalence induction +EquivJ : {A B : Type ℓ} (P : (A : Type ℓ) → (e : A ≃ B) → Type ℓ') + → (r : P B (idEquiv B)) → (e : A ≃ B) → P A e +EquivJ P r e = subst (λ x → P (x .fst) (x .snd)) (contrSinglEquiv e) r + +-- Assuming that we have an inverse to ua we can easily prove univalence +module Univalence (au : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → A ≃ B) + (aurefl : ∀ {ℓ} {A B : Type ℓ} → au refl ≡ idEquiv A) where + + ua-au : {A B : Type ℓ} (p : A ≡ B) → ua (au p) ≡ p + ua-au {B = B} = J (λ _ p → ua (au p) ≡ p) + (cong ua (aurefl {B = B}) ∙ uaIdEquiv) + + au-ua : {A B : Type ℓ} (e : A ≃ B) → au (ua e) ≡ e + au-ua {B = B} = EquivJ (λ _ f → au (ua f) ≡ f) + (subst (λ r → au r ≡ idEquiv _) (sym uaIdEquiv) (aurefl {B = B})) + + isoThm : ∀ {ℓ} {A B : Type ℓ} → Iso (A ≡ B) (A ≃ B) + isoThm .Iso.fun = au + isoThm .Iso.inv = ua + isoThm .Iso.rightInv = au-ua + isoThm .Iso.leftInv = ua-au + + thm : ∀ {ℓ} {A B : Type ℓ} → isEquiv au + thm {A = A} {B = B} = isoToIsEquiv {B = A ≃ B} isoThm + +pathToEquiv : {A B : Type ℓ} → A ≡ B → A ≃ B +pathToEquiv p = lineToEquiv (λ i → p i) + +pathToEquivRefl : {A : Type ℓ} → pathToEquiv refl ≡ idEquiv A +pathToEquivRefl {A = A} = equivEq (λ i x → transp (λ _ → A) i x) + +pathToEquiv-ua : {A B : Type ℓ} (e : A ≃ B) → pathToEquiv (ua e) ≡ e +pathToEquiv-ua = Univalence.au-ua pathToEquiv pathToEquivRefl + +ua-pathToEquiv : {A B : Type ℓ} (p : A ≡ B) → ua (pathToEquiv p) ≡ p +ua-pathToEquiv = Univalence.ua-au pathToEquiv pathToEquivRefl + +-- Univalence +univalenceIso : {A B : Type ℓ} → Iso (A ≡ B) (A ≃ B) +univalenceIso = Univalence.isoThm pathToEquiv pathToEquivRefl + +univalence : {A B : Type ℓ} → (A ≡ B) ≃ (A ≃ B) +univalence = ( pathToEquiv , Univalence.thm pathToEquiv pathToEquivRefl ) + +-- The original map from UniMath/Foundations +eqweqmap : {A B : Type ℓ} → A ≡ B → A ≃ B +eqweqmap {A = A} e = J (λ X _ → A ≃ X) (idEquiv A) e + +eqweqmapid : {A : Type ℓ} → eqweqmap refl ≡ idEquiv A +eqweqmapid {A = A} = JRefl (λ X _ → A ≃ X) (idEquiv A) + +univalenceStatement : {A B : Type ℓ} → isEquiv (eqweqmap {ℓ} {A} {B}) +univalenceStatement = Univalence.thm eqweqmap eqweqmapid + +univalenceUAH : {A B : Type ℓ} → (A ≡ B) ≃ (A ≃ B) +univalenceUAH = ( _ , univalenceStatement ) + +univalencePath : {A B : Type ℓ} → (A ≡ B) ≡ Lift (A ≃ B) +univalencePath = ua (compEquiv univalence LiftEquiv) + +-- The computation rule for ua. Because of "ghcomp" it is now very +-- simple compared to cubicaltt: +-- https://github.com/mortberg/cubicaltt/blob/master/examples/univalence.ctt#L202 +uaβ : {A B : Type ℓ} (e : A ≃ B) (x : A) → transport (ua e) x ≡ equivFun e x +uaβ e x = transportRefl (equivFun e x) + +uaη : ∀ {A B : Type ℓ} → (P : A ≡ B) → ua (pathToEquiv P) ≡ P +uaη = J (λ _ q → ua (pathToEquiv q) ≡ q) (cong ua pathToEquivRefl ∙ uaIdEquiv) + +-- Lemmas for constructing and destructing dependent paths in a function type where the domain is ua. +ua→ : ∀ {ℓ ℓ'} {A₀ A₁ : Type ℓ} {e : A₀ ≃ A₁} {B : (i : I) → Type ℓ'} + {f₀ : A₀ → B i0} {f₁ : A₁ → B i1} + → ((a : A₀) → PathP B (f₀ a) (f₁ (e .fst a))) + → PathP (λ i → ua e i → B i) f₀ f₁ +ua→ {e = e} {f₀ = f₀} {f₁} h i a = + hcomp + (λ j → λ + { (i = i0) → f₀ a + ; (i = i1) → f₁ (lem a j) + }) + (h (transp (λ j → ua e (~ j ∧ i)) (~ i) a) i) + where + lem : ∀ a₁ → e .fst (transport (sym (ua e)) a₁) ≡ a₁ + lem a₁ = secEq e _ ∙ transportRefl _ + +ua→⁻ : ∀ {ℓ ℓ'} {A₀ A₁ : Type ℓ} {e : A₀ ≃ A₁} {B : (i : I) → Type ℓ'} + {f₀ : A₀ → B i0} {f₁ : A₁ → B i1} + → PathP (λ i → ua e i → B i) f₀ f₁ + → ((a : A₀) → PathP B (f₀ a) (f₁ (e .fst a))) +ua→⁻ {e = e} {f₀ = f₀} {f₁} p a i = + hcomp + (λ k → λ + { (i = i0) → f₀ a + ; (i = i1) → f₁ (uaβ e a k) + }) + (p i (transp (λ j → ua e (j ∧ i)) (~ i) a)) + +-- Useful lemma for unfolding a transported function over ua +-- If we would have regularity this would be refl +transportUAop₁ : ∀ {A B : Type ℓ} → (e : A ≃ B) (f : A → A) (x : B) + → transport (λ i → ua e i → ua e i) f x ≡ equivFun e (f (invEq e x)) +transportUAop₁ e f x i = transportRefl (equivFun e (f (invEq e (transportRefl x i)))) i + +-- Binary version +transportUAop₂ : ∀ {ℓ} {A B : Type ℓ} → (e : A ≃ B) (f : A → A → A) (x y : B) + → transport (λ i → ua e i → ua e i → ua e i) f x y ≡ + equivFun e (f (invEq e x) (invEq e y)) +transportUAop₂ e f x y i = + transportRefl (equivFun e (f (invEq e (transportRefl x i)) + (invEq e (transportRefl y i)))) i + +-- Alternative version of EquivJ that only requires a predicate on functions +elimEquivFun : {A B : Type ℓ} (P : (A : Type ℓ) → (A → B) → Type ℓ') + → (r : P B (idfun B)) → (e : A ≃ B) → P A (e .fst) +elimEquivFun P r e = subst (λ x → P (x .fst) (x .snd .fst)) (contrSinglEquiv e) r + +-- Isomorphism induction +elimIso : {B : Type ℓ} → (Q : {A : Type ℓ} → (A → B) → (B → A) → Type ℓ') → + (h : Q (idfun B) (idfun B)) → {A : Type ℓ} → + (f : A → B) → (g : B → A) → section f g → retract f g → Q f g +elimIso {ℓ} {ℓ'} {B} Q h {A} f g sfg rfg = rem1 f g sfg rfg + where + P : (A : Type ℓ) → (f : A → B) → Type (ℓ-max ℓ' ℓ) + P A f = (g : B → A) → section f g → retract f g → Q f g + + rem : P B (idfun B) + rem g sfg rfg = subst (Q (idfun B)) (λ i b → (sfg b) (~ i)) h + + rem1 : {A : Type ℓ} → (f : A → B) → P A f + rem1 f g sfg rfg = elimEquivFun P rem (f , isoToIsEquiv (iso f g sfg rfg)) g sfg rfg + + +uaInvEquiv : ∀ {A B : Type ℓ} → (e : A ≃ B) → ua (invEquiv e) ≡ sym (ua e) +uaInvEquiv {B = B} = EquivJ (λ _ e → ua (invEquiv e) ≡ sym (ua e)) + (cong ua (invEquivIdEquiv B)) + +uaCompEquiv : ∀ {A B C : Type ℓ} → (e : A ≃ B) (f : B ≃ C) → ua (compEquiv e f) ≡ ua e ∙ ua f +uaCompEquiv {B = B} {C} = EquivJ (λ _ e → (f : B ≃ C) → ua (compEquiv e f) ≡ ua e ∙ ua f) + (λ f → cong ua (compEquivIdEquiv f) + ∙ sym (cong (λ x → x ∙ ua f) uaIdEquiv + ∙ sym (lUnit (ua f)))) diff --git a/Cubical/Foundations/Univalence/Universe.agda b/Cubical/Foundations/Univalence/Universe.agda new file mode 100644 index 0000000000..e41e2db5ac --- /dev/null +++ b/Cubical/Foundations/Univalence/Universe.agda @@ -0,0 +1,108 @@ +{-# OPTIONS --safe --postfix-projections #-} + +open import Cubical.Core.Everything + +open import Cubical.Functions.Embedding +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport + +-- A helper module for deriving univalence for a higher inductive-recursive +-- universe. +-- +-- U is the type of codes +-- El is the decoding +-- un is a higher constructor that requires paths between codes to exist +-- for equivalences of decodings +-- comp is intended to be the computational behavior of El on un, although +-- it seems that being a path is sufficient. +-- +-- Given a universe defined as above, it's possible to show that the path +-- space of the code type is equivalent to the path space of the actual +-- decodings, which are themselves determined by equivalences. +-- +-- The levels are left independent, but of course it will generally be +-- impossible to define this sort of universe unless ℓ' < ℓ, because El will +-- be too big to go in a constructor of U. The exception would be if U could +-- be defined independently of El, though it might be tricky to get the right +-- higher structure in such a case. +module Cubical.Foundations.Univalence.Universe {ℓ ℓ'} + (U : Type ℓ) + (El : U → Type ℓ') + (un : ∀ s t → El s ≃ El t → s ≡ t) + (comp : ∀{s t} (e : El s ≃ El t) → cong El (un s t e) ≡ ua e) + where +private + variable + A : Type ℓ' + +module UU-Lemmas where + reg : transport (λ _ → A) ≡ idfun A + reg {A} i z = transp (λ _ → A) i z + + nu : ∀ x y → x ≡ y → El x ≃ El y + nu x y p = transportEquiv (cong El p) + + cong-un-te + : ∀ x y (p : El x ≡ El y) + → cong El (un x y (transportEquiv p)) ≡ p + cong-un-te x y p + = comp (transportEquiv p) ∙ uaTransportη p + + nu-un : ∀ x y (e : El x ≃ El y) → nu x y (un x y e) ≡ e + nu-un x y e + = equivEq {e = nu x y (un x y e)} {f = e} λ i z + → (cong (λ p → transport p z) (comp e) ∙ uaβ e z) i + + El-un-equiv : ∀ x i → El (un x x (idEquiv _) i) ≃ El x + El-un-equiv x i = λ where + .fst → transp (λ j → p j) (i ∨ ~ i) + .snd → transp (λ j → isEquiv (transp (λ k → p (j ∧ k)) (~ j ∨ i ∨ ~ i))) + (i ∨ ~ i) (idIsEquiv T) + where + T = El (un x x (idEquiv _) i) + p : T ≡ El x + p j = (comp (idEquiv _) ∙ uaIdEquiv {A = El x}) j i + + un-refl : ∀ x → un x x (idEquiv (El x)) ≡ refl + un-refl x i j + = hcomp (λ k → λ where + (i = i0) → un x x (idEquiv (El x)) j + (i = i1) → un x x (idEquiv (El x)) (j ∨ k) + (j = i0) → un x x (idEquiv (El x)) (~ i ∨ k) + (j = i1) → x) + (un (un x x (idEquiv (El x)) (~ i)) x (El-un-equiv x (~ i)) j) + + nu-refl : ∀ x → nu x x refl ≡ idEquiv (El x) + nu-refl x = equivEq {e = nu x x refl} {f = idEquiv (El x)} reg + + un-nu : ∀ x y (p : x ≡ y) → un x y (nu x y p) ≡ p + un-nu x y p + = J (λ z q → un x z (nu x z q) ≡ q) (cong (un x x) (nu-refl x) ∙ un-refl x) p + +open UU-Lemmas +open Iso + +equivIso : ∀ s t → Iso (s ≡ t) (El s ≃ El t) +equivIso s t .fun = nu s t +equivIso s t .inv = un s t +equivIso s t .rightInv = nu-un s t +equivIso s t .leftInv = un-nu s t + +pathIso : ∀ s t → Iso (s ≡ t) (El s ≡ El t) +pathIso s t .fun = cong El +pathIso s t .inv = un s t ∘ transportEquiv +pathIso s t .rightInv = cong-un-te s t +pathIso s t .leftInv = un-nu s t + +minivalence : ∀{s t} → (s ≡ t) ≃ (El s ≃ El t) +minivalence {s} {t} = isoToEquiv (equivIso s t) + +path-reflection : ∀{s t} → (s ≡ t) ≃ (El s ≡ El t) +path-reflection {s} {t} = isoToEquiv (pathIso s t) + +isEmbeddingEl : isEmbedding El +isEmbeddingEl s t = snd path-reflection diff --git a/Cubical/Functions/Bundle.agda b/Cubical/Functions/Bundle.agda new file mode 100644 index 0000000000..98638451f0 --- /dev/null +++ b/Cubical/Functions/Bundle.agda @@ -0,0 +1,50 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Bundle where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Fibration + +open import Cubical.Foundations.Structure +open import Cubical.Structures.TypeEqvTo + +open import Cubical.Data.Sigma.Properties +open import Cubical.HITs.PropositionalTruncation + +module _ {ℓb ℓf} {B : Type ℓb} {F : Type ℓf} {ℓ} where + +{- + A fiber bundle with base space B and fibers F is a map `p⁻¹ : B → TypeEqvTo F` + taking points in the base space to their respective fibers. + + e.g. a double cover is a map `B → TypeEqvTo Bool` +-} + + Total : (p⁻¹ : B → TypeEqvTo ℓ F) → Type (ℓ-max ℓb ℓ) + Total p⁻¹ = Σ[ x ∈ B ] p⁻¹ x .fst + + pr : (p⁻¹ : B → TypeEqvTo ℓ F) → Total p⁻¹ → B + pr p⁻¹ = fst + + inc : (p⁻¹ : B → TypeEqvTo ℓ F) (x : B) → p⁻¹ x .fst → Total p⁻¹ + inc p⁻¹ x = (x ,_) + + fibPrEquiv : (p⁻¹ : B → TypeEqvTo ℓ F) (x : B) → fiber (pr p⁻¹) x ≃ p⁻¹ x .fst + fibPrEquiv p⁻¹ x = fiberEquiv (λ x → typ (p⁻¹ x)) x + +module _ {ℓb ℓf} (B : Type ℓb) (ℓ : Level) (F : Type ℓf) where + private + ℓ' = ℓ-max ℓb ℓ + +{- + Equivalently, a fiber bundle with base space B and fibers F is a type E and + a map `p : E → B` with fibers merely equivalent to F. +-} + + bundleEquiv : (B → TypeEqvTo ℓ' F) ≃ (Σ[ E ∈ Type ℓ' ] Σ[ p ∈ (E → B) ] ∀ x → ∥ fiber p x ≃ F ∥) + bundleEquiv = compEquiv (compEquiv Σ-Π-≃ (pathToEquiv p)) Σ-assoc-≃ + where e = fibrationEquiv B ℓ + p : (Σ[ p⁻¹ ∈ (B → Type ℓ') ] ∀ x → ∥ p⁻¹ x ≃ F ∥) + ≡ (Σ[ p ∈ (Σ[ E ∈ Type ℓ' ] (E → B)) ] ∀ x → ∥ fiber (snd p) x ≃ F ∥ ) + p i = Σ[ q ∈ ua e (~ i) ] ∀ x → ∥ ua-unglue e (~ i) q x ≃ F ∥ diff --git a/Cubical/Functions/Embedding.agda b/Cubical/Functions/Embedding.agda new file mode 100644 index 0000000000..3f83f2e55d --- /dev/null +++ b/Cubical/Functions/Embedding.agda @@ -0,0 +1,427 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Embedding where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence using (ua; univalence) +open import Cubical.Functions.Fibration + +open import Cubical.Data.Sigma +open import Cubical.Functions.Fibration +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Nullary using (Discrete; yes; no) +open import Cubical.Structures.Axioms + +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Data.Nat using (ℕ; zero; suc) +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ₁ ℓ₂ : Level + A B : Type ℓ + f h : A → B + w x : A + y z : B + +-- Embeddings are generalizations of injections. The usual +-- definition of injection as: +-- +-- f x ≡ f y → x ≡ y +-- +-- is not well-behaved with higher h-levels, while embeddings +-- are. +isEmbedding : (A → B) → Type _ +isEmbedding f = ∀ w x → isEquiv {A = w ≡ x} (cong f) + +isPropIsEmbedding : isProp (isEmbedding f) +isPropIsEmbedding {f = f} = isPropΠ2 λ _ _ → isPropIsEquiv (cong f) + +-- If A and B are h-sets, then injective functions between +-- them are embeddings. +-- +-- Note: It doesn't appear to be possible to omit either of +-- the `isSet` hypotheses. +injEmbedding + : {f : A → B} + → isSet A → isSet B + → (∀{w x} → f w ≡ f x → w ≡ x) + → isEmbedding f +injEmbedding {f = f} iSA iSB inj w x + = isoToIsEquiv (iso (cong f) inj sect retr) + where + sect : section (cong f) inj + sect p = iSB (f w) (f x) _ p + + retr : retract (cong f) inj + retr p = iSA w x _ p + +-- If `f` is an embedding, we'd expect the fibers of `f` to be +-- propositions, like an injective function. +hasPropFibers : (A → B) → Type _ +hasPropFibers f = ∀ y → isProp (fiber f y) + +-- This can be relaxed to having all prop fibers over the image, see [hasPropFibersOfImage→isEmbedding] +hasPropFibersOfImage : (A → B) → Type _ +hasPropFibersOfImage f = ∀ x → isProp (fiber f (f x)) + +-- some notation +_↪_ : Type ℓ₁ → Type ℓ₂ → Type (ℓ-max ℓ₁ ℓ₂) +A ↪ B = Σ[ f ∈ (A → B) ] isEmbedding f + +hasPropFibersIsProp : isProp (hasPropFibers f) +hasPropFibersIsProp = isPropΠ (λ _ → isPropIsProp) + +private + lemma₀ : (p : y ≡ z) → fiber f y ≡ fiber f z + lemma₀ {f = f} p = λ i → fiber f (p i) + + lemma₁ : isEmbedding f → ∀ x → isContr (fiber f (f x)) + lemma₁ {f = f} iE x = value , path + where + value : fiber f (f x) + value = (x , refl) + + path : ∀(fi : fiber f (f x)) → value ≡ fi + path (w , p) i + = case equiv-proof (iE w x) p of λ + { ((q , sq) , _) + → hfill (λ j → λ { (i = i0) → (x , refl) + ; (i = i1) → (w , sq j) + }) + (inS (q (~ i) , λ j → f (q (~ i ∨ j)))) + i1 + } + +isEmbedding→hasPropFibers : isEmbedding f → hasPropFibers f +isEmbedding→hasPropFibers iE y (x , p) + = subst (λ f → isProp f) (lemma₀ p) (isContr→isProp (lemma₁ iE x)) (x , p) + +private + fibCong→PathP + : {f : A → B} + → (p : f w ≡ f x) + → (fi : fiber (cong f) p) + → PathP (λ i → fiber f (p i)) (w , refl) (x , refl) + fibCong→PathP p (q , r) i = q i , λ j → r j i + + PathP→fibCong + : {f : A → B} + → (p : f w ≡ f x) + → (pp : PathP (λ i → fiber f (p i)) (w , refl) (x , refl)) + → fiber (cong f) p + PathP→fibCong p pp = (λ i → fst (pp i)) , (λ j i → snd (pp i) j) + +PathP≡fibCong + : {f : A → B} + → (p : f w ≡ f x) + → PathP (λ i → fiber f (p i)) (w , refl) (x , refl) ≡ fiber (cong f) p +PathP≡fibCong p + = isoToPath (iso (PathP→fibCong p) (fibCong→PathP p) (λ _ → refl) (λ _ → refl)) + +hasPropFibers→isEmbedding : hasPropFibers f → isEmbedding f +hasPropFibers→isEmbedding {f = f} iP w x .equiv-proof p + = subst isContr (PathP≡fibCong p) (isProp→isContrPathP (λ i → iP (p i)) fw fx) + where + fw : fiber f (f w) + fw = (w , refl) + + fx : fiber f (f x) + fx = (x , refl) + +hasPropFibersOfImage→hasPropFibers : hasPropFibersOfImage f → hasPropFibers f +hasPropFibersOfImage→hasPropFibers {f = f} fibImg y a b = + subst (λ y → isProp (fiber f y)) (snd a) (fibImg (fst a)) a b + +hasPropFibersOfImage→isEmbedding : hasPropFibersOfImage f → isEmbedding f +hasPropFibersOfImage→isEmbedding = hasPropFibers→isEmbedding ∘ hasPropFibersOfImage→hasPropFibers + +isEmbedding≡hasPropFibers : isEmbedding f ≡ hasPropFibers f +isEmbedding≡hasPropFibers + = isoToPath + (iso isEmbedding→hasPropFibers + hasPropFibers→isEmbedding + (λ _ → hasPropFibersIsProp _ _) + (λ _ → isPropIsEmbedding _ _)) + +isEquiv→hasPropFibers : isEquiv f → hasPropFibers f +isEquiv→hasPropFibers e b = isContr→isProp (equiv-proof e b) + +isEquiv→isEmbedding : isEquiv f → isEmbedding f +isEquiv→isEmbedding e = λ _ _ → congEquiv (_ , e) .snd + +Equiv→Embedding : A ≃ B → A ↪ B +Equiv→Embedding (f , isEquivF) = (f , isEquiv→isEmbedding isEquivF) + +iso→isEmbedding : ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + ------------------------------- + → isEmbedding (Iso.fun isom) +iso→isEmbedding {A = A} {B} isom = (isEquiv→isEmbedding (equivIsEquiv (isoToEquiv isom))) + +isEmbedding→Injection : + ∀ {ℓ} {A B C : Type ℓ} + → (a : A → B) + → (e : isEmbedding a) + ---------------------- + → ∀ {f g : C → A} → + ∀ x → (a (f x) ≡ a (g x)) ≡ (f x ≡ g x) +isEmbedding→Injection a e {f = f} {g} x = sym (ua (cong a , e (f x) (g x))) + +-- if `f` has a retract, then `cong f` has, as well. If `B` is a set, then `cong f` +-- further has a section, making `f` an embedding. +module _ {f : A → B} (retf : hasRetract f) where + open Σ retf renaming (fst to g ; snd to ϕ) + + congRetract : f w ≡ f x → w ≡ x + congRetract {w = w} {x = x} p = sym (ϕ w) ∙∙ cong g p ∙∙ ϕ x + + isRetractCongRetract : retract (cong {x = w} {y = x} f) congRetract + isRetractCongRetract p = transport (PathP≡doubleCompPathˡ _ _ _ _) (λ i j → ϕ (p j) i) + + hasRetract→hasRetractCong : hasRetract (cong {x = w} {y = x} f) + hasRetract→hasRetractCong = congRetract , isRetractCongRetract + + retractableIntoSet→isEmbedding : isSet B → isEmbedding f + retractableIntoSet→isEmbedding setB w x = + isoToIsEquiv (iso (cong f) congRetract (λ _ → setB _ _ _ _) (hasRetract→hasRetractCong .snd)) + +Embedding-into-Discrete→Discrete : A ↪ B → Discrete B → Discrete A +Embedding-into-Discrete→Discrete (f , isEmbeddingF) _≟_ x y with f x ≟ f y +... | yes p = yes (invIsEq (isEmbeddingF x y) p) +... | no ¬p = no (¬p ∘ cong f) + +Embedding-into-isProp→isProp : A ↪ B → isProp B → isProp A +Embedding-into-isProp→isProp (f , isEmbeddingF) isProp-B x y + = invIsEq (isEmbeddingF x y) (isProp-B (f x) (f y)) + +Embedding-into-isSet→isSet : A ↪ B → isSet B → isSet A +Embedding-into-isSet→isSet (f , isEmbeddingF) isSet-B x y p q = + p ≡⟨ sym (retIsEq isEquiv-cong-f p) ⟩ + cong-f⁻¹ (cong f p) ≡⟨ cong cong-f⁻¹ cong-f-p≡cong-f-q ⟩ + cong-f⁻¹ (cong f q) ≡⟨ retIsEq isEquiv-cong-f q ⟩ + q ∎ + where + cong-f-p≡cong-f-q = isSet-B (f x) (f y) (cong f p) (cong f q) + isEquiv-cong-f = isEmbeddingF x y + cong-f⁻¹ = invIsEq isEquiv-cong-f + +Embedding-into-hLevel→hLevel + : ∀ n → A ↪ B → isOfHLevel (suc n) B → isOfHLevel (suc n) A +Embedding-into-hLevel→hLevel zero = Embedding-into-isProp→isProp +Embedding-into-hLevel→hLevel (suc n) (f , isEmbeddingF) Blvl x y + = isOfHLevelRespectEquiv (suc n) (invEquiv equiv) subLvl + where + equiv : (x ≡ y) ≃ (f x ≡ f y) + equiv .fst = cong f + equiv .snd = isEmbeddingF x y + subLvl : isOfHLevel (suc n) (f x ≡ f y) + subLvl = Blvl (f x) (f y) + +-- We now show that the powerset is the subtype classifier +-- i.e. ℙ X ≃ Σ[A ∈ Type ℓ] (A ↪ X) +Embedding→Subset : {X : Type ℓ} → Σ[ A ∈ Type ℓ ] (A ↪ X) → ℙ X +Embedding→Subset (_ , f , isEmbeddingF) x = fiber f x , isEmbedding→hasPropFibers isEmbeddingF x + +Subset→Embedding : {X : Type ℓ} → ℙ X → Σ[ A ∈ Type ℓ ] (A ↪ X) +Subset→Embedding {X = X} A = D , fst , Ψ + where + D = Σ[ x ∈ X ] x ∈ A + + Ψ : isEmbedding fst + Ψ w x = isEmbeddingFstΣProp (∈-isProp A) + +Subset→Embedding→Subset : {X : Type ℓ} → section (Embedding→Subset {ℓ} {X}) (Subset→Embedding {ℓ} {X}) +Subset→Embedding→Subset _ = funExt λ x → Σ≡Prop (λ _ → isPropIsProp) (ua (FiberIso.fiberEquiv _ x)) + +Embedding→Subset→Embedding : {X : Type ℓ} → retract (Embedding→Subset {ℓ} {X}) (Subset→Embedding {ℓ} {X}) +Embedding→Subset→Embedding {ℓ = ℓ} {X = X} (A , f , ψ) = + cong (equivFun Σ-assoc-≃) (Σ≡Prop (λ _ → isPropIsEmbedding) (retEq (fibrationEquiv X ℓ) (A , f))) + +Subset≃Embedding : {X : Type ℓ} → ℙ X ≃ (Σ[ A ∈ Type ℓ ] (A ↪ X)) +Subset≃Embedding = isoToEquiv (iso Subset→Embedding Embedding→Subset + Embedding→Subset→Embedding Subset→Embedding→Subset) + +Subset≡Embedding : {X : Type ℓ} → ℙ X ≡ (Σ[ A ∈ Type ℓ ] (A ↪ X)) +Subset≡Embedding = ua Subset≃Embedding + +isEmbedding-∘ : isEmbedding f → isEmbedding h → isEmbedding (f ∘ h) +isEmbedding-∘ {f = f} {h = h} Embf Embh w x + = compEquiv (cong h , Embh w x) (cong f , Embf (h w) (h x)) .snd + +isEmbedding→embedsFibersIntoSingl + : isEmbedding f + → ∀ z → fiber f z ↪ singl z +isEmbedding→embedsFibersIntoSingl {f = f} isE z = e , isEmbE where + e : fiber f z → singl z + e x = f (fst x) , sym (snd x) + + isEmbE : isEmbedding e + isEmbE u v = goal where + -- "adjust" ΣeqCf by trivial equivalences that hold judgementally, which should save compositions + Dom′ : ∀ u v → Type _ + Dom′ u v = Σ[ p ∈ fst u ≡ fst v ] PathP (λ i → f (p i) ≡ z) (snd u) (snd v) + Cod′ : ∀ u v → Type _ + Cod′ u v = Σ[ p ∈ f (fst u) ≡ f (fst v) ] PathP (λ i → p i ≡ z) (snd u) (snd v) + ΣeqCf : Dom′ u v ≃ Cod′ u v + ΣeqCf = Σ-cong-equiv-fst (_ , isE _ _) + + dom→ : u ≡ v → Dom′ u v + dom→ p = cong fst p , cong snd p + dom← : Dom′ u v → u ≡ v + dom← p i = p .fst i , p .snd i + + cod→ : e u ≡ e v → Cod′ u v + cod→ p = cong fst p , cong (sym ∘ snd) p + cod← : Cod′ u v → e u ≡ e v + cod← p i = p .fst i , sym (p .snd i) + + goal : isEquiv (cong e) + goal .equiv-proof x .fst .fst = + dom← (equivCtr ΣeqCf (cod→ x) .fst) + goal .equiv-proof x .fst .snd j = + cod← (equivCtr ΣeqCf (cod→ x) .snd j) + goal .equiv-proof x .snd (g , p) i .fst = + dom← (equivCtrPath ΣeqCf (cod→ x) (dom→ g , cong cod→ p) i .fst) + goal .equiv-proof x .snd (g , p) i .snd j = + cod← (equivCtrPath ΣeqCf (cod→ x) (dom→ g , cong cod→ p) i .snd j) + +isEmbedding→hasPropFibers′ : isEmbedding f → hasPropFibers f +isEmbedding→hasPropFibers′ {f = f} iE z = + Embedding-into-isProp→isProp (isEmbedding→embedsFibersIntoSingl iE z) isPropSingl + +universeEmbedding : + ∀ {ℓ ℓ₁ : Level} + → (F : Type ℓ → Type ℓ₁) + → (∀ X → F X ≃ X) + → isEmbedding F +universeEmbedding F liftingEquiv = hasPropFibersOfImage→isEmbedding propFibersF where + lemma : ∀ A B → (F A ≡ F B) ≃ (B ≡ A) + lemma A B = (F A ≡ F B) ≃⟨ univalence ⟩ + (F A ≃ F B) ≃⟨ equivComp (liftingEquiv A) (liftingEquiv B) ⟩ + (A ≃ B) ≃⟨ invEquivEquiv ⟩ + (B ≃ A) ≃⟨ invEquiv univalence ⟩ + (B ≡ A) ■ + fiberSingl : ∀ X → fiber F (F X) ≃ singl X + fiberSingl X = Σ-cong-equiv-snd (λ _ → lemma _ _) + propFibersF : hasPropFibersOfImage F + propFibersF X = Embedding-into-isProp→isProp (Equiv→Embedding (fiberSingl X)) isPropSingl + +liftEmbedding : (ℓ ℓ₁ : Level) + → isEmbedding (Lift {i = ℓ} {j = ℓ₁}) +liftEmbedding ℓ ℓ₁ = universeEmbedding (Lift {j = ℓ₁}) (λ _ → invEquiv LiftEquiv) + +module FibrationIdentityPrinciple {B : Type ℓ} {ℓ₁} where + -- note that fibrationEquiv (for good reason) uses ℓ₁ = ℓ-max ℓ ℓ₁, so we have to work + -- some universe magic to achieve good universe polymorphism + + -- First, prove it for the case that's dealt with in fibrationEquiv + Fibration′ = Fibration B (ℓ-max ℓ ℓ₁) + + module Lifted (f g : Fibration′) where + f≃g′ : Type (ℓ-max ℓ ℓ₁) + f≃g′ = ∀ b → fiber (f .snd) b ≃ fiber (g .snd) b + + Fibration′IP : f≃g′ ≃ (f ≡ g) + Fibration′IP = + f≃g′ + ≃⟨ equivΠCod (λ _ → invEquiv univalence) ⟩ + (∀ b → fiber (f .snd) b ≡ fiber (g .snd) b) + ≃⟨ funExtEquiv ⟩ + fiber (f .snd) ≡ fiber (g .snd) + ≃⟨ invEquiv (congEquiv (fibrationEquiv B ℓ₁)) ⟩ + f ≡ g + ■ + + -- Then embed into the above case by lifting the type + L : Type _ → Type _ -- local synonym fixing the levels of Lift + L = Lift {i = ℓ₁} {j = ℓ} + + liftFibration : Fibration B ℓ₁ → Fibration′ + liftFibration (A , f) = L A , f ∘ lower + + hasPropFibersLiftFibration : hasPropFibers liftFibration + hasPropFibersLiftFibration (A , f) = + Embedding-into-isProp→isProp (Equiv→Embedding fiberChar) + (isPropΣ (isEmbedding→hasPropFibers (liftEmbedding _ _) A) + λ _ → isEquiv→hasPropFibers (snd (invEquiv (preCompEquiv LiftEquiv))) _) + where + fiberChar : fiber liftFibration (A , f) + ≃ (Σ[ (E , eq) ∈ fiber L A ] fiber (_∘ lower) (transport⁻ (λ i → eq i → B) f)) + fiberChar = + fiber liftFibration (A , f) + ≃⟨ Σ-cong-equiv-snd (λ _ → invEquiv ΣPath≃PathΣ) ⟩ + (Σ[ (E , g) ∈ Fibration B ℓ₁ ] Σ[ eq ∈ (L E ≡ A) ] PathP (λ i → eq i → B) (g ∘ lower) f) + ≃⟨ boringSwap ⟩ + (Σ[ (E , eq) ∈ fiber L A ] Σ[ g ∈ (E → B) ] PathP (λ i → eq i → B) (g ∘ lower) f) + ≃⟨ Σ-cong-equiv-snd (λ _ → Σ-cong-equiv-snd λ _ → transportEquiv (PathP≡Path⁻ _ _ _)) ⟩ + (Σ[ (E , eq) ∈ fiber L A ] fiber (_∘ lower) (transport⁻ (λ i → eq i → B) f)) + ■ where + unquoteDecl boringSwap = + declStrictEquiv boringSwap + (λ ((E , g) , (eq , p)) → ((E , eq) , (g , p))) + (λ ((E , g) , (eq , p)) → ((E , eq) , (g , p))) + + isEmbeddingLiftFibration : isEmbedding liftFibration + isEmbeddingLiftFibration = hasPropFibers→isEmbedding hasPropFibersLiftFibration + + -- and finish off + module _ (f g : Fibration B ℓ₁) where + open Lifted (liftFibration f) (liftFibration g) + f≃g : Type (ℓ-max ℓ ℓ₁) + f≃g = ∀ b → fiber (f .snd) b ≃ fiber (g .snd) b + + FibrationIP : f≃g ≃ (f ≡ g) + FibrationIP = + f≃g ≃⟨ equivΠCod (λ b → equivComp (Σ-cong-equiv-fst LiftEquiv) + (Σ-cong-equiv-fst LiftEquiv)) ⟩ + f≃g′ ≃⟨ Fibration′IP ⟩ + (liftFibration f ≡ liftFibration g) ≃⟨ invEquiv (_ , isEmbeddingLiftFibration _ _) ⟩ + (f ≡ g) ■ +open FibrationIdentityPrinciple renaming (f≃g to _≃Fib_) using (FibrationIP) public + +Embedding : (B : Type ℓ₁) → (ℓ : Level) → Type (ℓ-max ℓ₁ (ℓ-suc ℓ)) +Embedding B ℓ = Σ[ A ∈ Type ℓ ] A ↪ B + +module EmbeddingIdentityPrinciple {B : Type ℓ} {ℓ₁} (f g : Embedding B ℓ₁) where + module _ where + open Σ f renaming (fst to F) public + open Σ g renaming (fst to G) public + open Σ (f .snd) renaming (fst to ffun; snd to isEmbF) public + open Σ (g .snd) renaming (fst to gfun; snd to isEmbG) public + f≃g : Type _ + f≃g = (∀ b → fiber ffun b → fiber gfun b) × + (∀ b → fiber gfun b → fiber ffun b) + toFibr : Embedding B ℓ₁ → Fibration B ℓ₁ + toFibr (A , (f , _)) = (A , f) + + isEmbeddingToFibr : isEmbedding toFibr + isEmbeddingToFibr w x = fullEquiv .snd where + -- carefully managed such that (cong toFibr) is the equivalence + fullEquiv : (w ≡ x) ≃ (toFibr w ≡ toFibr x) + fullEquiv = compEquiv (congEquiv (invEquiv Σ-assoc-≃)) (invEquiv (Σ≡PropEquiv (λ _ → isPropIsEmbedding))) + + EmbeddingIP : f≃g ≃ (f ≡ g) + EmbeddingIP = + f≃g + ≃⟨ strictIsoToEquiv (invIso toProdIso) ⟩ + (∀ b → (fiber ffun b → fiber gfun b) × (fiber gfun b → fiber ffun b)) + ≃⟨ equivΠCod (λ _ → isEquivPropBiimpl→Equiv (isEmbedding→hasPropFibers isEmbF _) + (isEmbedding→hasPropFibers isEmbG _)) ⟩ + (∀ b → (fiber (f .snd .fst) b) ≃ (fiber (g .snd .fst) b)) + ≃⟨ FibrationIP (toFibr f) (toFibr g) ⟩ + (toFibr f ≡ toFibr g) + ≃⟨ invEquiv (_ , isEmbeddingToFibr _ _) ⟩ + f ≡ g + ■ +open EmbeddingIdentityPrinciple renaming (f≃g to _≃Emb_) using (EmbeddingIP) public diff --git a/Cubical/Functions/Fibration.agda b/Cubical/Functions/Fibration.agda new file mode 100644 index 0000000000..602be2d34e --- /dev/null +++ b/Cubical/Functions/Fibration.agda @@ -0,0 +1,111 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Fibration where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels using (isOfHLevel→isOfHLevelDep) +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Data.Sigma + +private + variable + ℓ ℓb : Level + B : Type ℓb + +module FiberIso {ℓ} (p⁻¹ : B → Type ℓ) (x : B) where + + p : Σ B p⁻¹ → B + p = fst + + fwd : fiber p x → p⁻¹ x + fwd ((x' , y) , q) = subst (λ z → p⁻¹ z) q y + + bwd : p⁻¹ x → fiber p x + bwd y = (x , y) , refl + + fwd-bwd : ∀ x → fwd (bwd x) ≡ x + fwd-bwd y = transportRefl y + + bwd-fwd : ∀ x → bwd (fwd x) ≡ x + bwd-fwd ((x' , y) , q) i = h (r i) + where h : Σ[ s ∈ singl x ] p⁻¹ (s .fst) → fiber p x + h ((x , p) , y) = (x , y) , sym p + r : Path (Σ[ s ∈ singl x ] p⁻¹ (s .fst)) + ((x , refl ) , subst p⁻¹ q y) + ((x' , sym q) , y ) + r = ΣPathP (isContrSingl x .snd (x' , sym q) + , toPathP (transport⁻Transport (λ i → p⁻¹ (q i)) y)) + + -- HoTT Lemma 4.8.1 + fiberEquiv : fiber p x ≃ p⁻¹ x + fiberEquiv = isoToEquiv (iso fwd bwd fwd-bwd bwd-fwd) + +open FiberIso using (fiberEquiv) public + +module _ {ℓ} {E : Type ℓ} (p : E → B) where + + -- HoTT Lemma 4.8.2 + totalEquiv : E ≃ Σ B (fiber p) + totalEquiv = isoToEquiv isom + where isom : Iso E (Σ B (fiber p)) + Iso.fun isom x = p x , x , refl + Iso.inv isom (b , x , q) = x + Iso.leftInv isom x i = x + Iso.rightInv isom (b , x , q) i = q i , x , λ j → q (i ∧ j) + +module _ (B : Type ℓb) (ℓ : Level) where + private + ℓ' = ℓ-max ℓb ℓ + + -- HoTT Theorem 4.8.3 + fibrationEquiv : (Σ[ E ∈ Type ℓ' ] (E → B)) ≃ (B → Type ℓ') + fibrationEquiv = isoToEquiv isom + where isom : Iso (Σ[ E ∈ Type ℓ' ] (E → B)) (B → Type ℓ') + Iso.fun isom (E , p) = fiber p + Iso.inv isom p⁻¹ = Σ B p⁻¹ , fst + Iso.rightInv isom p⁻¹ i x = ua (fiberEquiv p⁻¹ x) i + Iso.leftInv isom (E , p) i = ua e (~ i) , fst ∘ ua-unglue e (~ i) + where e = totalEquiv p + + +module ForSets {E : Type ℓ} {isSetB : isSet B} (f : E → B) where + module _ {x x'} {px : x ≡ x'} {a' : fiber f x} {b' : fiber f x'} where + -- fibers are equal when their representatives are equal + fibersEqIfRepsEq : fst a' ≡ fst b' + → PathP (λ i → fiber f (px i)) a' b' + fibersEqIfRepsEq p = ΣPathP (p , + (isOfHLevel→isOfHLevelDep 1 + (λ (v , w) → isSetB (f v) w) + (snd a') (snd b') + (λ i → (p i , px i)))) +-- The path type in a fiber of f is equivalent to a fiber of (cong f) +open import Cubical.Foundations.Function + +fiberPath : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {f : A → B} {b : B} (h h' : fiber f b) → + (Σ[ p ∈ (fst h ≡ fst h') ] (PathP (λ i → f (p i) ≡ b) (snd h) (snd h'))) + ≡ fiber (cong f) (h .snd ∙∙ refl ∙∙ sym (h' .snd)) +fiberPath h h' = cong (Σ (h .fst ≡ h' .fst)) (funExt λ p → flipSquarePath ∙ PathP≡doubleCompPathʳ _ _ _ _) + +fiber≡ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {f : A → B} {b : B} (h h' : fiber f b) + → (h ≡ h') ≡ fiber (cong f) (h .snd ∙∙ refl ∙∙ sym (h' .snd)) +fiber≡ {f = f} {b = b} h h' = + ΣPath≡PathΣ ⁻¹ ∙ + fiberPath h h' + +fiberCong : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) {a₀ a₁ : A} (q : f a₀ ≡ f a₁) + → fiber (cong f) q ≡ Path (fiber f (f a₁)) (a₀ , q) (a₁ , refl) +fiberCong f q = + cong (fiber (cong f)) (cong sym (lUnit (sym q))) + ∙ sym (fiber≡ (_ , q) (_ , refl)) + +FibrationStr : (B : Type ℓb) → Type ℓ → Type (ℓ-max ℓ ℓb) +FibrationStr B A = A → B + +Fibration : (B : Type ℓb) → (ℓ : Level) → Type (ℓ-max ℓb (ℓ-suc ℓ)) +Fibration {ℓb = ℓb} B ℓ = Σ[ A ∈ Type ℓ ] FibrationStr B A diff --git a/Cubical/Functions/Fixpoint.agda b/Cubical/Functions/Fixpoint.agda new file mode 100644 index 0000000000..24eb44d951 --- /dev/null +++ b/Cubical/Functions/Fixpoint.agda @@ -0,0 +1,43 @@ +{- + Definition of function fixpoint and Kraus' lemma +-} +{-# OPTIONS --safe #-} +module Cubical.Functions.Fixpoint where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws + +private + variable + ℓ : Level + A : Type ℓ + +Fixpoint : (A → A) → Type _ +Fixpoint {A = A} f = Σ A (λ x → f x ≡ x) + +fixpoint : {f : A → A} → Fixpoint f → A +fixpoint = fst + +fixpointPath : {f : A → A} → (p : Fixpoint f) → f (fixpoint p) ≡ fixpoint p +fixpointPath = snd + +-- Kraus' lemma +-- a version not using cubical features can be found at +-- https://www.cs.bham.ac.uk/~mhe/GeneralizedHedberg/html/GeneralizedHedberg.html#21576 +2-Constant→isPropFixpoint : (f : A → A) → 2-Constant f → isProp (Fixpoint f) +2-Constant→isPropFixpoint f fconst (x , p) (y , q) i = s i , t i where + noose : ∀ x y → f x ≡ f y + noose x y = sym (fconst x x) ∙ fconst x y + -- the main idea is that for any path p, cong f p does not depend on p + -- but only on its endpoints and the structure of 2-Constant f + KrausInsight : ∀ {x y} → (p : x ≡ y) → noose x y ≡ cong f p + KrausInsight {x} = J (λ y p → noose x y ≡ cong f p) (lCancel (fconst x x)) + -- Need to solve for a path s : x ≡ y, such that: + -- transport (λ i → cong f s i ≡ s i) p ≡ q + s : x ≡ y + s = sym p ∙∙ noose x y ∙∙ q + t' : PathP (λ i → noose x y i ≡ s i) p q + t' i j = doubleCompPath-filler (sym p) (noose x y) q j i + t : PathP (λ i → cong f s i ≡ s i) p q + t = subst (λ kraus → PathP (λ i → kraus i ≡ s i) p q) (KrausInsight s) t' diff --git a/Cubical/Functions/FunExtEquiv.agda b/Cubical/Functions/FunExtEquiv.agda new file mode 100644 index 0000000000..21a5efdab7 --- /dev/null +++ b/Cubical/Functions/FunExtEquiv.agda @@ -0,0 +1,193 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.FunExtEquiv where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.CartesianKanOps +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Vec.Base +open import Cubical.Data.Vec.NAry +open import Cubical.Data.Nat + +open import Cubical.Reflection.StrictEquiv + +private + variable + ℓ ℓ₁ ℓ₂ ℓ₃ : Level + +-- Function extensionality is an equivalence +module _ {A : Type ℓ} {B : A → I → Type ℓ₁} + {f : (x : A) → B x i0} {g : (x : A) → B x i1} where + + funExtEquiv : (∀ x → PathP (B x) (f x) (g x)) ≃ PathP (λ i → ∀ x → B x i) f g + unquoteDef funExtEquiv = defStrictEquiv funExtEquiv funExt funExt⁻ + + funExtPath : (∀ x → PathP (B x) (f x) (g x)) ≡ PathP (λ i → ∀ x → B x i) f g + funExtPath = ua funExtEquiv + + funExtIso : Iso (∀ x → PathP (B x) (f x) (g x)) (PathP (λ i → ∀ x → B x i) f g) + funExtIso = iso funExt funExt⁻ (λ x → refl {x = x}) (λ x → refl {x = x}) + +-- Function extensionality for binary functions +funExt₂ : {A : Type ℓ} {B : A → Type ℓ₁} {C : (x : A) → B x → I → Type ℓ₂} + {f : (x : A) → (y : B x) → C x y i0} + {g : (x : A) → (y : B x) → C x y i1} + → ((x : A) (y : B x) → PathP (C x y) (f x y) (g x y)) + → PathP (λ i → ∀ x y → C x y i) f g +funExt₂ p i x y = p x y i + +-- Function extensionality for binary functions is an equivalence +module _ {A : Type ℓ} {B : A → Type ℓ₁} {C : (x : A) → B x → I → Type ℓ₂} + {f : (x : A) → (y : B x) → C x y i0} + {g : (x : A) → (y : B x) → C x y i1} where + private + appl₂ : PathP (λ i → ∀ x y → C x y i) f g → ∀ x y → PathP (C x y) (f x y) (g x y) + appl₂ eq x y i = eq i x y + + funExt₂Equiv : (∀ x y → PathP (C x y) (f x y) (g x y)) ≃ (PathP (λ i → ∀ x y → C x y i) f g) + unquoteDef funExt₂Equiv = defStrictEquiv funExt₂Equiv funExt₂ appl₂ + + funExt₂Path : (∀ x y → PathP (C x y) (f x y) (g x y)) ≡ (PathP (λ i → ∀ x y → C x y i) f g) + funExt₂Path = ua funExt₂Equiv + + +-- Function extensionality for ternary functions +funExt₃ : {A : Type ℓ} {B : A → Type ℓ₁} {C : (x : A) → B x → Type ℓ₂} + {D : (x : A) → (y : B x) → C x y → I → Type ℓ₃} + {f : (x : A) → (y : B x) → (z : C x y) → D x y z i0} + {g : (x : A) → (y : B x) → (z : C x y) → D x y z i1} + → ((x : A) (y : B x) (z : C x y) → PathP (D x y z) (f x y z) (g x y z)) + → PathP (λ i → ∀ x y z → D x y z i) f g +funExt₃ p i x y z = p x y z i + +-- Function extensionality for ternary functions is an equivalence +module _ {A : Type ℓ} {B : A → Type ℓ₁} {C : (x : A) → B x → Type ℓ₂} + {D : (x : A) → (y : B x) → C x y → I → Type ℓ₃} + {f : (x : A) → (y : B x) → (z : C x y) → D x y z i0} + {g : (x : A) → (y : B x) → (z : C x y) → D x y z i1} where + private + appl₃ : PathP (λ i → ∀ x y z → D x y z i) f g → ∀ x y z → PathP (D x y z) (f x y z) (g x y z) + appl₃ eq x y z i = eq i x y z + + funExt₃Equiv : (∀ x y z → PathP (D x y z) (f x y z) (g x y z)) ≃ (PathP (λ i → ∀ x y z → D x y z i) f g) + unquoteDef funExt₃Equiv = defStrictEquiv funExt₃Equiv funExt₃ appl₃ + + funExt₃Path : (∀ x y z → PathP (D x y z) (f x y z) (g x y z)) ≡ (PathP (λ i → ∀ x y z → D x y z i) f g) + funExt₃Path = ua funExt₃Equiv + + +-- n-ary non-dependent funext +nAryFunExt : {X : Type ℓ} {Y : I → Type ℓ₁} (n : ℕ) (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + → ((xs : Vec X n) → PathP Y (fX $ⁿ xs) (fY $ⁿ map (λ x → x) xs)) + → PathP (λ i → nAryOp n X (Y i)) fX fY +nAryFunExt zero fX fY p = p [] +nAryFunExt (suc n) fX fY p i x = nAryFunExt n (fX x) (fY x) (λ xs → p (x ∷ xs)) i + +-- n-ary funext⁻ +nAryFunExt⁻ : (n : ℕ) {X : Type ℓ} {Y : I → Type ℓ₁} (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + → PathP (λ i → nAryOp n X (Y i)) fX fY + → ((xs : Vec X n) → PathP Y (fX $ⁿ xs) (fY $ⁿ map (λ x → x) xs)) +nAryFunExt⁻ zero fX fY p [] = p +nAryFunExt⁻ (suc n) fX fY p (x ∷ xs) = nAryFunExt⁻ n (fX x) (fY x) (λ i → p i x) xs + +nAryFunExtEquiv : (n : ℕ) {X : Type ℓ} {Y : I → Type ℓ₁} (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + → ((xs : Vec X n) → PathP Y (fX $ⁿ xs) (fY $ⁿ map (λ x → x) xs)) ≃ PathP (λ i → nAryOp n X (Y i)) fX fY +nAryFunExtEquiv n {X} {Y} fX fY = isoToEquiv (iso (nAryFunExt n fX fY) (nAryFunExt⁻ n fX fY) + (linv n fX fY) (rinv n fX fY)) + where + linv : (n : ℕ) (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + (p : PathP (λ i → nAryOp n X (Y i)) fX fY) + → nAryFunExt n fX fY (nAryFunExt⁻ n fX fY p) ≡ p + linv zero fX fY p = refl + linv (suc n) fX fY p i j x = linv n (fX x) (fY x) (λ k → p k x) i j + + rinv : (n : ℕ) (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + (p : (xs : Vec X n) → PathP Y (fX $ⁿ xs) (fY $ⁿ map (λ x → x) xs)) + → nAryFunExt⁻ n fX fY (nAryFunExt n fX fY p) ≡ p + rinv zero fX fY p i [] = p [] + rinv (suc n) fX fY p i (x ∷ xs) = rinv n (fX x) (fY x) (λ ys i → p (x ∷ ys) i) i xs + +-- Funext when the domain also depends on the interval + +funExtDep : {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ₁} + {f : (x : A i0) → B i0 x} {g : (x : A i1) → B i1 x} + → ({x₀ : A i0} {x₁ : A i1} (p : PathP A x₀ x₁) → PathP (λ i → B i (p i)) (f x₀) (g x₁)) + → PathP (λ i → (x : A i) → B i x) f g +funExtDep {A = A} {B} {f} {g} h i x = + comp + (λ k → B i (coei→i A i x k)) + (λ k → λ + { (i = i0) → f (coei→i A i0 x k) + ; (i = i1) → g (coei→i A i1 x k) + }) + (h (λ j → coei→j A i j x) i) + +funExtDep⁻ : {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ₁} + {f : (x : A i0) → B i0 x} {g : (x : A i1) → B i1 x} + → PathP (λ i → (x : A i) → B i x) f g + → ({x₀ : A i0} {x₁ : A i1} (p : PathP A x₀ x₁) → PathP (λ i → B i (p i)) (f x₀) (g x₁)) +funExtDep⁻ q p i = q i (p i) + +funExtDepEquiv : {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ₁} + {f : (x : A i0) → B i0 x} {g : (x : A i1) → B i1 x} + → ({x₀ : A i0} {x₁ : A i1} (p : PathP A x₀ x₁) → PathP (λ i → B i (p i)) (f x₀) (g x₁)) + ≃ PathP (λ i → (x : A i) → B i x) f g +funExtDepEquiv {A = A} {B} {f} {g} = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = funExtDep + isom .inv = funExtDep⁻ + isom .rightInv q m i x = + comp + (λ k → B i (coei→i A i x (k ∨ m))) + (λ k → λ + { (i = i0) → f (coei→i A i0 x (k ∨ m)) + ; (i = i1) → g (coei→i A i1 x (k ∨ m)) + ; (m = i1) → q i x + }) + (q i (coei→i A i x m)) + isom .leftInv h m p i = + comp + (λ k → B i (lemi→i m k)) + (λ k → λ + { (i = i0) → f (lemi→i m k) + ; (i = i1) → g (lemi→i m k) + ; (m = i1) → h p i + }) + (h (λ j → lemi→j j m) i) + where + lemi→j : ∀ j → coei→j A i j (p i) ≡ p j + lemi→j j = + coei→j (λ k → coei→j A i k (p i) ≡ p k) i j (coei→i A i (p i)) + + lemi→i : PathP (λ m → lemi→j i m ≡ p i) (coei→i A i (p i)) refl + lemi→i = + sym (coei→i (λ k → coei→j A i k (p i) ≡ p k) i (coei→i A i (p i))) + ◁ λ m k → lemi→j i (m ∨ k) + +heteroHomotopy≃Homotopy : {A : I → Type ℓ} {B : (i : I) → Type ℓ₁} + {f : A i0 → B i0} {g : A i1 → B i1} + → ({x₀ : A i0} {x₁ : A i1} → PathP A x₀ x₁ → PathP B (f x₀) (g x₁)) + ≃ ((x₀ : A i0) → PathP B (f x₀) (g (transport (λ i → A i) x₀))) +heteroHomotopy≃Homotopy {A = A} {B} {f} {g} = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun h x₀ = h (isContrSinglP A x₀ .fst .snd) + isom .inv k {x₀} {x₁} p = + subst (λ fib → PathP B (f x₀) (g (fib .fst))) (isContrSinglP A x₀ .snd (x₁ , p)) (k x₀) + isom .rightInv k = funExt λ x₀ → + cong (λ α → subst (λ fib → PathP B (f x₀) (g (fib .fst))) α (k x₀)) + (isProp→isSet isPropSinglP (isContrSinglP A x₀ .fst) _ + (isContrSinglP A x₀ .snd (isContrSinglP A x₀ .fst)) + refl) + ∙ transportRefl (k x₀) + isom .leftInv h j {x₀} {x₁} p = + transp + (λ i → PathP B (f x₀) (g (isContrSinglP A x₀ .snd (x₁ , p) (i ∨ j) .fst))) + j + (h (isContrSinglP A x₀ .snd (x₁ , p) j .snd)) diff --git a/Cubical/Functions/Implicit.agda b/Cubical/Functions/Implicit.agda new file mode 100644 index 0000000000..a98a2d3a65 --- /dev/null +++ b/Cubical/Functions/Implicit.agda @@ -0,0 +1,17 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Implicit where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +implicit≃Explicit : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} + → ({a : A} → B a) ≃ ((a : A) → B a) +implicit≃Explicit = isoToEquiv isom + where + isom : Iso _ _ + Iso.fun isom f a = f + Iso.inv isom f = f _ + Iso.rightInv isom f = funExt λ _ → refl + Iso.leftInv isom f = implicitFunExt refl + diff --git a/Cubical/Functions/Involution.agda b/Cubical/Functions/Involution.agda new file mode 100644 index 0000000000..9f9e3bfa9a --- /dev/null +++ b/Cubical/Functions/Involution.agda @@ -0,0 +1,42 @@ +{-# OPTIONS --safe #-} + +module Cubical.Functions.Involution where + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence + +isInvolution : ∀{ℓ} {A : Type ℓ} → (A → A) → Type _ +isInvolution f = ∀ x → f (f x) ≡ x + +module _ {ℓ} {A : Type ℓ} {f : A → A} (invol : isInvolution f) where + open Iso + + involIso : Iso A A + involIso .fun = f + involIso .inv = f + involIso .rightInv = invol + involIso .leftInv = invol + + involIsEquiv : isEquiv f + involIsEquiv = isoToIsEquiv involIso + + involEquiv : A ≃ A + involEquiv = f , involIsEquiv + + involPath : A ≡ A + involPath = ua involEquiv + + involEquivComp : compEquiv involEquiv involEquiv ≡ idEquiv A + involEquivComp + = equivEq (λ i x → invol x i) + + involPathComp : involPath ∙ involPath ≡ refl + involPathComp + = sym (uaCompEquiv involEquiv involEquiv) ∙∙ cong ua involEquivComp ∙∙ uaIdEquiv + + involPath² : Square involPath refl refl involPath + involPath² + = subst (λ s → Square involPath s refl involPath) + involPathComp (compPath-filler involPath involPath) diff --git a/Cubical/Functions/Logic.agda b/Cubical/Functions/Logic.agda new file mode 100644 index 0000000000..3e5439c208 --- /dev/null +++ b/Cubical/Functions/Logic.agda @@ -0,0 +1,290 @@ +-- Various functions for manipulating hProps. +-- +-- This file used to be part of Foundations, but it turned out to be +-- not very useful so moved here. Feel free to upstream content. +-- +-- Note that it is often a bad idea to use hProp instead of having the +-- isProp proof separate. The reason is that Agda can rarely infer +-- isProp proofs making it easier to just give them explicitly instead +-- of having them bundled up with the type. +-- +{-# OPTIONS --safe #-} +module Cubical.Functions.Logic where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence using (hPropExt) + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum as ⊎ using (_⊎_) +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +open import Cubical.Relation.Nullary hiding (¬_) + + +-------------------------------------------------------------------------------- +-- The type hProp of mere propositions +-- the definition hProp is given in Foundations.HLevels +-- hProp ℓ = Σ (Type ℓ) isProp + +private + variable + ℓ ℓ' ℓ'' : Level + P Q R : hProp ℓ + A B C : Type ℓ + +infix 10 ¬_ +infixr 8 _⊔_ +infixr 8 _⊔′_ +infixr 8 _⊓_ +infixr 8 _⊓′_ +infixr 6 _⇒_ +infixr 4 _⇔_ +infix 30 _≡ₚ_ +infix 30 _≢ₚ_ + +infix 2 ∃[]-syntax +infix 2 ∃[∶]-syntax + +infix 2 ∀[∶]-syntax +infix 2 ∀[]-syntax + +infix 2 ⇒∶_⇐∶_ +infix 2 ⇐∶_⇒∶_ + +∥_∥ₚ : Type ℓ → hProp ℓ +∥ A ∥ₚ = ∥ A ∥ , isPropPropTrunc + +_≡ₚ_ : (x y : A) → hProp _ +x ≡ₚ y = ∥ x ≡ y ∥ₚ + +hProp≡ : ⟨ P ⟩ ≡ ⟨ Q ⟩ → P ≡ Q +hProp≡ = TypeOfHLevel≡ 1 + +isProp⟨⟩ : (A : hProp ℓ) → isProp ⟨ A ⟩ +isProp⟨⟩ = snd + +-------------------------------------------------------------------------------- +-- Logical implication of mere propositions + +_⇒_ : (A : hProp ℓ) → (B : hProp ℓ') → hProp _ +A ⇒ B = (⟨ A ⟩ → ⟨ B ⟩) , isPropΠ λ _ → isProp⟨⟩ B + +⇔toPath : ⟨ P ⇒ Q ⟩ → ⟨ Q ⇒ P ⟩ → P ≡ Q +⇔toPath {P = P} {Q = Q} P⇒Q Q⇒P = hProp≡ (hPropExt (isProp⟨⟩ P) (isProp⟨⟩ Q) P⇒Q Q⇒P) + +pathTo⇒ : P ≡ Q → ⟨ P ⇒ Q ⟩ +pathTo⇒ p x = subst fst p x + +pathTo⇐ : P ≡ Q → ⟨ Q ⇒ P ⟩ +pathTo⇐ p x = subst fst (sym p) x + +substₚ : {x y : A} (B : A → hProp ℓ) → ⟨ x ≡ₚ y ⇒ B x ⇒ B y ⟩ +substₚ {x = x} {y = y} B = PropTrunc.elim (λ _ → isPropΠ λ _ → isProp⟨⟩ (B y)) (subst (fst ∘ B)) + +-------------------------------------------------------------------------------- +-- Mixfix notations for ⇔-toPath +-- see ⊔-identityˡ and ⊔-identityʳ for the difference + +⇒∶_⇐∶_ : ⟨ P ⇒ Q ⟩ → ⟨ Q ⇒ P ⟩ → P ≡ Q +⇒∶_⇐∶_ = ⇔toPath + +⇐∶_⇒∶_ : ⟨ Q ⇒ P ⟩ → ⟨ P ⇒ Q ⟩ → P ≡ Q +⇐∶ g ⇒∶ f = ⇔toPath f g +-------------------------------------------------------------------------------- +-- False and True + +⊥ : hProp _ +⊥ = ⊥.⊥ , λ () + +⊤ : ∀ {ℓ} → hProp ℓ +⊤ = Unit* , (λ _ _ _ → tt*) + +-------------------------------------------------------------------------------- +-- Pseudo-complement of mere propositions + +¬_ : hProp ℓ → hProp _ +¬ A = (⟨ A ⟩ → ⊥.⊥) , isPropΠ λ _ → ⊥.isProp⊥ + +_≢ₚ_ : (x y : A) → hProp _ +x ≢ₚ y = ¬ x ≡ₚ y + +-------------------------------------------------------------------------------- +-- Disjunction of mere propositions + +_⊔′_ : Type ℓ → Type ℓ' → Type _ +A ⊔′ B = ∥ A ⊎ B ∥ + +_⊔_ : hProp ℓ → hProp ℓ' → hProp _ +P ⊔ Q = ∥ ⟨ P ⟩ ⊎ ⟨ Q ⟩ ∥ₚ + +inl : A → A ⊔′ B +inl x = ∣ ⊎.inl x ∣ + +inr : B → A ⊔′ B +inr x = ∣ ⊎.inr x ∣ + +⊔-elim : (P : hProp ℓ) (Q : hProp ℓ') (R : ⟨ P ⊔ Q ⟩ → hProp ℓ'') + → (∀ x → ⟨ R (inl x) ⟩) → (∀ y → ⟨ R (inr y) ⟩) → (∀ z → ⟨ R z ⟩) +⊔-elim _ _ R P⇒R Q⇒R = PropTrunc.elim (snd ∘ R) (⊎.elim P⇒R Q⇒R) + +-------------------------------------------------------------------------------- +-- Conjunction of mere propositions +_⊓′_ : Type ℓ → Type ℓ' → Type _ +A ⊓′ B = A × B + +_⊓_ : hProp ℓ → hProp ℓ' → hProp _ +A ⊓ B = ⟨ A ⟩ ⊓′ ⟨ B ⟩ , isOfHLevelΣ 1 (isProp⟨⟩ A) (\ _ → isProp⟨⟩ B) + +⊓-intro : (P : hProp ℓ) (Q : ⟨ P ⟩ → hProp ℓ') (R : ⟨ P ⟩ → hProp ℓ'') + → (∀ a → ⟨ Q a ⟩) → (∀ a → ⟨ R a ⟩) → (∀ (a : ⟨ P ⟩) → ⟨ Q a ⊓ R a ⟩ ) +⊓-intro _ _ _ = \ f g a → f a , g a + +-------------------------------------------------------------------------------- +-- Logical bi-implication of mere propositions + +_⇔_ : hProp ℓ → hProp ℓ' → hProp _ +A ⇔ B = (A ⇒ B) ⊓ (B ⇒ A) + +⇔-id : (P : hProp ℓ) → ⟨ P ⇔ P ⟩ +⇔-id P = (idfun ⟨ P ⟩) , (idfun ⟨ P ⟩) + +-------------------------------------------------------------------------------- +-- Universal Quantifier + + +∀[∶]-syntax : (A → hProp ℓ) → hProp _ +∀[∶]-syntax {A = A} P = (∀ x → ⟨ P x ⟩) , isPropΠ (isProp⟨⟩ ∘ P) + +∀[]-syntax : (A → hProp ℓ) → hProp _ +∀[]-syntax {A = A} P = (∀ x → ⟨ P x ⟩) , isPropΠ (isProp⟨⟩ ∘ P) + +syntax ∀[∶]-syntax {A = A} (λ a → P) = ∀[ a ∶ A ] P +syntax ∀[]-syntax (λ a → P) = ∀[ a ] P + +-------------------------------------------------------------------------------- +-- Existential Quantifier + +∃[]-syntax : (A → hProp ℓ) → hProp _ +∃[]-syntax {A = A} P = ∥ Σ A (⟨_⟩ ∘ P) ∥ₚ + +∃[∶]-syntax : (A → hProp ℓ) → hProp _ +∃[∶]-syntax {A = A} P = ∥ Σ A (⟨_⟩ ∘ P) ∥ₚ + +syntax ∃[∶]-syntax {A = A} (λ x → P) = ∃[ x ∶ A ] P +syntax ∃[]-syntax (λ x → P) = ∃[ x ] P + +-------------------------------------------------------------------------------- +-- Decidable mere proposition + +Decₚ : (P : hProp ℓ) → hProp ℓ +Decₚ P = Dec ⟨ P ⟩ , isPropDec (isProp⟨⟩ P) + +-------------------------------------------------------------------------------- +-- Negation commutes with truncation + +∥¬A∥≡¬∥A∥ : (A : Type ℓ) → ∥ (A → ⊥.⊥) ∥ₚ ≡ (¬ ∥ A ∥ₚ) +∥¬A∥≡¬∥A∥ _ = + ⇒∶ (λ ¬A A → PropTrunc.elim (λ _ → ⊥.isProp⊥) + (PropTrunc.elim (λ _ → isPropΠ λ _ → ⊥.isProp⊥) (λ ¬p p → ¬p p) ¬A) A) + ⇐∶ λ ¬p → ∣ (λ p → ¬p ∣ p ∣) ∣ + +-------------------------------------------------------------------------------- +-- (hProp, ⊔, ⊥) is a bounded ⊔-semilattice + +⊔-assoc : (P : hProp ℓ) (Q : hProp ℓ') (R : hProp ℓ'') + → P ⊔ (Q ⊔ R) ≡ (P ⊔ Q) ⊔ R +⊔-assoc P Q R = + ⇒∶ ⊔-elim P (Q ⊔ R) (λ _ → (P ⊔ Q) ⊔ R) + (inl ∘ inl) + (⊔-elim Q R (λ _ → (P ⊔ Q) ⊔ R) (inl ∘ inr) inr) + + ⇐∶ assoc2 + where + assoc2 : (A ⊔′ B) ⊔′ C → A ⊔′ (B ⊔′ C) + assoc2 ∣ ⊎.inr a ∣ = ∣ ⊎.inr ∣ ⊎.inr a ∣ ∣ + assoc2 ∣ ⊎.inl ∣ ⊎.inr b ∣ ∣ = ∣ ⊎.inr ∣ ⊎.inl b ∣ ∣ + assoc2 ∣ ⊎.inl ∣ ⊎.inl c ∣ ∣ = ∣ ⊎.inl c ∣ + assoc2 ∣ ⊎.inl (squash x y i) ∣ = isPropPropTrunc (assoc2 ∣ ⊎.inl x ∣) (assoc2 ∣ ⊎.inl y ∣) i + assoc2 (squash x y i) = isPropPropTrunc (assoc2 x) (assoc2 y) i + +⊔-idem : (P : hProp ℓ) → P ⊔ P ≡ P +⊔-idem P = + ⇒∶ (⊔-elim P P (\ _ → P) (\ x → x) (\ x → x)) + ⇐∶ inl + +⊔-comm : (P : hProp ℓ) (Q : hProp ℓ') → P ⊔ Q ≡ Q ⊔ P +⊔-comm P Q = + ⇒∶ (⊔-elim P Q (\ _ → (Q ⊔ P)) inr inl) + ⇐∶ (⊔-elim Q P (\ _ → (P ⊔ Q)) inr inl) + +⊔-identityˡ : (P : hProp ℓ) → ⊥ ⊔ P ≡ P +⊔-identityˡ P = + ⇒∶ (⊔-elim ⊥ P (λ _ → P) (λ ()) (λ x → x)) + ⇐∶ inr + +⊔-identityʳ : (P : hProp ℓ) → P ⊔ ⊥ ≡ P +⊔-identityʳ P = ⇔toPath (⊔-elim P ⊥ (λ _ → P) (λ x → x) λ ()) inl + +-------------------------------------------------------------------------------- +-- (hProp, ⊓, ⊤) is a bounded ⊓-lattice + +⊓-assoc : (P : hProp ℓ) (Q : hProp ℓ') (R : hProp ℓ'') + → P ⊓ Q ⊓ R ≡ (P ⊓ Q) ⊓ R +⊓-assoc _ _ _ = + ⇒∶ (λ {(x , (y , z)) → (x , y) , z}) + ⇐∶ (λ {((x , y) , z) → x , (y , z) }) + +⊓-comm : (P : hProp ℓ) (Q : hProp ℓ') → P ⊓ Q ≡ Q ⊓ P +⊓-comm _ _ = ⇔toPath (\ p → p .snd , p .fst) (\ p → p .snd , p .fst) + +⊓-idem : (P : hProp ℓ) → P ⊓ P ≡ P +⊓-idem _ = ⇔toPath fst (λ x → x , x) + +⊓-identityˡ : (P : hProp ℓ) → ⊤ {ℓ} ⊓ P ≡ P +⊓-identityˡ _ = ⇔toPath snd λ x → tt* , x + +⊓-identityʳ : (P : hProp ℓ) → P ⊓ ⊤ {ℓ} ≡ P +⊓-identityʳ _ = ⇔toPath fst λ x → x , tt* + +-------------------------------------------------------------------------------- +-- Distributive laws + +⇒-⊓-distrib : (P : hProp ℓ) (Q : hProp ℓ')(R : hProp ℓ'') + → P ⇒ (Q ⊓ R) ≡ (P ⇒ Q) ⊓ (P ⇒ R) +⇒-⊓-distrib _ _ _ = + ⇒∶ (λ f → (fst ∘ f) , (snd ∘ f)) + ⇐∶ (λ { (f , g) x → f x , g x}) + +⊓-⊔-distribˡ : (P : hProp ℓ) (Q : hProp ℓ')(R : hProp ℓ'') + → P ⊓ (Q ⊔ R) ≡ (P ⊓ Q) ⊔ (P ⊓ R) +⊓-⊔-distribˡ P Q R = + ⇒∶ (λ { (x , a) → ⊔-elim Q R (λ _ → (P ⊓ Q) ⊔ (P ⊓ R)) + (λ y → ∣ ⊎.inl (x , y) ∣ ) + (λ z → ∣ ⊎.inr (x , z) ∣ ) a }) + + ⇐∶ ⊔-elim (P ⊓ Q) (P ⊓ R) (λ _ → P ⊓ Q ⊔ R) + (λ y → fst y , inl (snd y)) + (λ z → fst z , inr (snd z)) + +⊔-⊓-distribˡ : (P : hProp ℓ) (Q : hProp ℓ')(R : hProp ℓ'') + → P ⊔ (Q ⊓ R) ≡ (P ⊔ Q) ⊓ (P ⊔ R) +⊔-⊓-distribˡ P Q R = + ⇒∶ ⊔-elim P (Q ⊓ R) (λ _ → (P ⊔ Q) ⊓ (P ⊔ R) ) + (\ x → inl x , inl x) (\ p → inr (p .fst) , inr (p .snd)) + + ⇐∶ (λ { (x , y) → ⊔-elim P R (λ _ → P ⊔ Q ⊓ R) inl + (λ z → ⊔-elim P Q (λ _ → P ⊔ Q ⊓ R) inl (λ y → inr (y , z)) x) y }) + +⊓-∀-distrib : (P : A → hProp ℓ) (Q : A → hProp ℓ') + → (∀[ a ∶ A ] P a) ⊓ (∀[ a ∶ A ] Q a) ≡ (∀[ a ∶ A ] (P a ⊓ Q a)) +⊓-∀-distrib P Q = + ⇒∶ (λ {(p , q) a → p a , q a}) + ⇐∶ λ pq → (fst ∘ pq ) , (snd ∘ pq) diff --git a/Cubical/Functions/Morphism.agda b/Cubical/Functions/Morphism.agda new file mode 100644 index 0000000000..c4886cfdfb --- /dev/null +++ b/Cubical/Functions/Morphism.agda @@ -0,0 +1,84 @@ +{- + General lemmas about morphisms (defined as loosely as possible) +-} +{-# OPTIONS --safe #-} +module Cubical.Functions.Morphism where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open Iso +module ax {ℓ : Level} (A : Type ℓ) (_+A_ : A → A → A) (a₀ : A) where + rUnit = (a : A) → a +A a₀ ≡ a + lUnit = (a : A) → a₀ +A a ≡ a + + rCancel : (-A_ : A → A) → Type ℓ + rCancel -A_ = (a : A) → a +A (-A a) ≡ a₀ + + lCancel : (-A_ : A → A) → Type ℓ + lCancel -A_ = (a : A) → (-A a) +A a ≡ a₀ + + assoc = (x y z : A) → x +A (y +A z) ≡ ((x +A y) +A z) + + comm = (x y : A) → x +A y ≡ y +A x + +module morphLemmas {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} + (_+A_ : A → A → A) (_+B_ : B → B → B) + (f : A → B) (f-hom : (x y : A) → f (x +A y) ≡ f x +B f y) + where + + 0↦0 : (a₀ : A) (b₀ : B) (-A_ : A → A) (-B_ : B → B) + → ax.rUnit A _+A_ a₀ + → ax.rUnit B _+B_ b₀ + → ax.rCancel B _+B_ b₀ -B_ + → ax.assoc B _+B_ b₀ + → f a₀ ≡ b₀ + 0↦0 a₀ b₀ -A_ -B_ rUnitA rUnitB rCancelB assocB = + sym (rUnitB (f a₀)) + ∙∙ cong (f a₀ +B_) (sym (rCancelB (f a₀))) + ∙∙ assocB (f a₀) (f a₀) (-B f a₀) + ∙∙ sym (cong (_+B (-B f a₀)) (cong f (sym (rUnitA a₀)) ∙ f-hom a₀ a₀)) + ∙∙ rCancelB (f a₀) + + distrMinus : (a₀ : A) (b₀ : B) (-A_ : A → A) (-B_ : B → B) + → ax.lUnit B _+B_ b₀ + → ax.rUnit B _+B_ b₀ + → ax.lCancel A _+A_ a₀ -A_ + → ax.rCancel B _+B_ b₀ -B_ + → ax.assoc B _+B_ b₀ + → (0↦0 : f a₀ ≡ b₀) + → (x : A) → f (-A x) ≡ -B (f x) + distrMinus a₀ b₀ -A_ -B_ lUnitB rUnitB lCancelA rCancelB assocB 0↦0 x = + sym (rUnitB _) + ∙∙ cong (f (-A x) +B_) (sym (rCancelB (f x))) + ∙∙ assocB _ _ _ + ∙∙ cong (_+B (-B (f x))) (sym (f-hom (-A x) x) ∙∙ cong f (lCancelA x) ∙∙ 0↦0) + ∙∙ lUnitB _ + + distrMinus' : (a₀ : A) (b₀ : B) (-A_ : A → A) (-B_ : B → B) + → ax.lUnit B _+B_ b₀ + → ax.rUnit B _+B_ b₀ + → ax.rUnit A _+A_ a₀ + → ax.lCancel A _+A_ a₀ -A_ + → ax.rCancel B _+B_ b₀ -B_ + → ax.assoc A _+A_ a₀ + → ax.assoc B _+B_ b₀ + → f a₀ ≡ b₀ -- not really needed, but it can be useful to specify the proof yourself + → (x y : A) + → f (x +A (-A y)) ≡ (f x +B (-B (f y))) + distrMinus' a₀ b₀ -A_ -B_ lUnitB rUnitB rUnitA lCancelA rCancelB assocA assocB 0↦0 x y = + sym (rUnitB _) + ∙∙ cong (f (x +A (-A y)) +B_) (sym (rCancelB (f y))) ∙ assocB _ _ _ + ∙∙ cong (_+B (-B f y)) (sym (f-hom (x +A (-A y)) y) + ∙ cong f (sym (assocA x (-A y) y) + ∙∙ cong (x +A_) (lCancelA y) + ∙∙ rUnitA x)) + + isMorphInv : (g : B → A) → section f g → retract f g + → (x y : B) + → (g (x +B y) ≡ (g x +A g y)) + isMorphInv g sect retr x y = + cong g (cong₂ _+B_ (sym (sect x)) (sym (sect y))) + ∙∙ cong g (sym (f-hom (g x) (g y))) + ∙∙ retr (g x +A g y) diff --git a/Cubical/Functions/Surjection.agda b/Cubical/Functions/Surjection.agda new file mode 100644 index 0000000000..1c833049ad --- /dev/null +++ b/Cubical/Functions/Surjection.agda @@ -0,0 +1,80 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Surjection where + +open import Cubical.Core.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Embedding +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + f : A → B + +isSurjection : (A → B) → Type _ +isSurjection f = ∀ b → ∥ fiber f b ∥ + +_↠_ : Type ℓ → Type ℓ' → Type (ℓ-max ℓ ℓ') +A ↠ B = Σ[ f ∈ (A → B) ] isSurjection f + +section→isSurjection : {g : B → A} → section f g → isSurjection f +section→isSurjection {g = g} s b = ∣ g b , s b ∣ + +isPropIsSurjection : isProp (isSurjection f) +isPropIsSurjection = isPropΠ λ _ → squash + +isEquiv→isSurjection : isEquiv f → isSurjection f +isEquiv→isSurjection e b = ∣ fst (equiv-proof e b) ∣ + +isEquiv→isEmbedding×isSurjection : isEquiv f → isEmbedding f × isSurjection f +isEquiv→isEmbedding×isSurjection e = isEquiv→isEmbedding e , isEquiv→isSurjection e + +isEmbedding×isSurjection→isEquiv : isEmbedding f × isSurjection f → isEquiv f +equiv-proof (isEmbedding×isSurjection→isEquiv {f = f} (emb , sur)) b = + inhProp→isContr (PropTrunc.rec fib' (λ x → x) fib) fib' + where + hpf : hasPropFibers f + hpf = isEmbedding→hasPropFibers emb + + fib : ∥ fiber f b ∥ + fib = sur b + + fib' : isProp (fiber f b) + fib' = hpf b + +isEquiv≃isEmbedding×isSurjection : isEquiv f ≃ isEmbedding f × isSurjection f +isEquiv≃isEmbedding×isSurjection = isoToEquiv (iso + isEquiv→isEmbedding×isSurjection + isEmbedding×isSurjection→isEquiv + (λ _ → isOfHLevelΣ 1 isPropIsEmbedding (\ _ → isPropIsSurjection) _ _) + (λ _ → isPropIsEquiv _ _ _)) + +-- obs: for epi⇒surjective to go through we require a stronger +-- hypothesis that one would expect: +-- f must cancel functions from a higher universe. +rightCancellable : (f : A → B) → Type _ +rightCancellable {ℓ} {A} {ℓ'} {B} f = ∀ {C : Type (ℓ-suc (ℓ-max ℓ ℓ'))} + → ∀ (g g' : B → C) → (∀ x → g (f x) ≡ g' (f x)) → ∀ y → g y ≡ g' y + +-- This statement is in Mac Lane & Moerdijk (page 143, corollary 5). +epi⇒surjective : (f : A → B) → rightCancellable f → isSurjection f +epi⇒surjective f rc y = transport (fact₂ y) tt* + where hasPreimage : (A → B) → B → _ + hasPreimage f y = ∥ fiber f y ∥ + + fact₁ : ∀ x → Unit* ≡ hasPreimage f (f x) + fact₁ x = hPropExt isPropUnit* + isPropPropTrunc + (λ _ → ∣ (x , refl) ∣) + (λ _ → tt*) + + fact₂ : ∀ y → Unit* ≡ hasPreimage f y + fact₂ = rc _ _ fact₁ diff --git a/Cubical/HITs/2GroupoidTruncation.agda b/Cubical/HITs/2GroupoidTruncation.agda new file mode 100644 index 0000000000..8178890b8d --- /dev/null +++ b/Cubical/HITs/2GroupoidTruncation.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.2GroupoidTruncation where + +open import Cubical.HITs.2GroupoidTruncation.Base public +open import Cubical.HITs.2GroupoidTruncation.Properties public diff --git a/Cubical/HITs/2GroupoidTruncation/Base.agda b/Cubical/HITs/2GroupoidTruncation/Base.agda new file mode 100644 index 0000000000..0a83ce4ca6 --- /dev/null +++ b/Cubical/HITs/2GroupoidTruncation/Base.agda @@ -0,0 +1,17 @@ +{- + +This file contains: + +- Definition of 2-groupoid truncations + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.2GroupoidTruncation.Base where + +open import Cubical.Foundations.Prelude + +-- 2-groupoid truncation as a higher inductive type: + +data ∥_∥₄ {ℓ} (A : Type ℓ) : Type ℓ where + ∣_∣₄ : A → ∥ A ∥₄ + squash₄ : ∀ (x y : ∥ A ∥₄) (p q : x ≡ y) (r s : p ≡ q) (t u : r ≡ s) → t ≡ u diff --git a/Cubical/HITs/2GroupoidTruncation/Properties.agda b/Cubical/HITs/2GroupoidTruncation/Properties.agda new file mode 100644 index 0000000000..30dd85a57d --- /dev/null +++ b/Cubical/HITs/2GroupoidTruncation/Properties.agda @@ -0,0 +1,68 @@ +{- + +This file contains: + +- Properties of 2-groupoid truncations + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.2GroupoidTruncation.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.HITs.2GroupoidTruncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +rec : ∀ {B : Type ℓ} → is2Groupoid B → (A → B) → ∥ A ∥₄ → B +rec gB f ∣ x ∣₄ = f x +rec gB f (squash₄ _ _ _ _ _ _ t u i j k l) = + gB _ _ _ _ _ _ (λ m n o → rec gB f (t m n o)) (λ m n o → rec gB f (u m n o)) + i j k l + +elim : {B : ∥ A ∥₄ → Type ℓ} + (bG : (x : ∥ A ∥₄) → is2Groupoid (B x)) + (f : (x : A) → B ∣ x ∣₄) (x : ∥ A ∥₄) → B x +elim bG f ∣ x ∣₄ = f x +elim bG f (squash₄ x y p q r s u v i j k l) = + isOfHLevel→isOfHLevelDep 4 bG _ _ _ _ _ _ + (λ j k l → elim bG f (u j k l)) (λ j k l → elim bG f (v j k l)) + (squash₄ x y p q r s u v) + i j k l + +elim2 : {B : ∥ A ∥₄ → ∥ A ∥₄ → Type ℓ} + (gB : ((x y : ∥ A ∥₄) → is2Groupoid (B x y))) + (g : (a b : A) → B ∣ a ∣₄ ∣ b ∣₄) + (x y : ∥ A ∥₄) → B x y +elim2 gB g = elim (λ _ → is2GroupoidΠ (λ _ → gB _ _)) + (λ a → elim (λ _ → gB _ _) (g a)) + +elim3 : {B : (x y z : ∥ A ∥₄) → Type ℓ} + (gB : ((x y z : ∥ A ∥₄) → is2Groupoid (B x y z))) + (g : (a b c : A) → B ∣ a ∣₄ ∣ b ∣₄ ∣ c ∣₄) + (x y z : ∥ A ∥₄) → B x y z +elim3 gB g = elim2 (λ _ _ → is2GroupoidΠ (λ _ → gB _ _ _)) + (λ a b → elim (λ _ → gB _ _ _) (g a b)) + +is2Groupoid2GroupoidTrunc : is2Groupoid ∥ A ∥₄ +is2Groupoid2GroupoidTrunc a b p q r s = squash₄ a b p q r s + +2GroupoidTruncIdempotent≃ : is2Groupoid A → ∥ A ∥₄ ≃ A +2GroupoidTruncIdempotent≃ {A = A} hA = isoToEquiv f + where + f : Iso ∥ A ∥₄ A + Iso.fun f = rec hA (idfun A) + Iso.inv f x = ∣ x ∣₄ + Iso.rightInv f _ = refl + Iso.leftInv f = elim (λ _ → isOfHLevelSuc 4 is2Groupoid2GroupoidTrunc _ _) (λ _ → refl) + +2GroupoidTruncIdempotent : is2Groupoid A → ∥ A ∥₄ ≡ A +2GroupoidTruncIdempotent hA = ua (2GroupoidTruncIdempotent≃ hA) diff --git a/Cubical/HITs/AssocList.agda b/Cubical/HITs/AssocList.agda new file mode 100644 index 0000000000..746a04c1fa --- /dev/null +++ b/Cubical/HITs/AssocList.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.AssocList where + + +open import Cubical.HITs.AssocList.Base public +open import Cubical.HITs.AssocList.Properties public diff --git a/Cubical/HITs/AssocList/Base.agda b/Cubical/HITs/AssocList/Base.agda new file mode 100644 index 0000000000..26f82b0406 --- /dev/null +++ b/Cubical/HITs/AssocList/Base.agda @@ -0,0 +1,69 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.AssocList.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Nat using (ℕ; _+_) + +private + variable + ℓ : Level + A : Type ℓ + +infixr 5 ⟨_,_⟩∷_ + +data AssocList (A : Type ℓ) : Type ℓ where + ⟨⟩ : AssocList A + ⟨_,_⟩∷_ : (a : A) (n : ℕ) (xs : AssocList A) → AssocList A + per : ∀ a b xs → ⟨ a , 1 ⟩∷ ⟨ b , 1 ⟩∷ xs + ≡ ⟨ b , 1 ⟩∷ ⟨ a , 1 ⟩∷ xs + agg : ∀ a m n xs → ⟨ a , m ⟩∷ ⟨ a , n ⟩∷ xs + ≡ ⟨ a , m + n ⟩∷ xs + del : ∀ a xs → ⟨ a , 0 ⟩∷ xs ≡ xs + trunc : isSet (AssocList A) + +pattern ⟨_⟩ a = ⟨ a , 1 ⟩∷ ⟨⟩ + + +-- Elimination and recursion principle for association lists +module Elim {ℓ'} {B : AssocList A → Type ℓ'} + (⟨⟩* : B ⟨⟩) (⟨_,_⟩∷*_ : (x : A) (n : ℕ) {xs : AssocList A} → B xs → B (⟨ x , n ⟩∷ xs)) + (per* : (x y : A) {xs : AssocList A} (b : B xs) + → PathP (λ i → B (per x y xs i)) (⟨ x , 1 ⟩∷* ⟨ y , 1 ⟩∷* b) (⟨ y , 1 ⟩∷* ⟨ x , 1 ⟩∷* b)) + (agg* : (x : A) (m n : ℕ) {xs : AssocList A} (b : B xs) + → PathP (λ i → B (agg x m n xs i)) (⟨ x , m ⟩∷* ⟨ x , n ⟩∷* b) (⟨ x , m + n ⟩∷* b)) + (del* : (x : A) {xs : AssocList A} (b : B xs) + → PathP (λ i → B (del x xs i)) (⟨ x , 0 ⟩∷* b) b) + (trunc* : (xs : AssocList A) → isSet (B xs)) where + + f : (xs : AssocList A) → B xs + f ⟨⟩ = ⟨⟩* + f (⟨ a , n ⟩∷ xs) = ⟨ a , n ⟩∷* f xs + f (per a b xs i) = per* a b (f xs) i + f (agg a m n xs i) = agg* a m n (f xs) i + f (del a xs i) = del* a (f xs) i + f (trunc xs ys p q i j) = isOfHLevel→isOfHLevelDep 2 trunc* (f xs) (f ys) (cong f p) (cong f q) (trunc xs ys p q) i j + + + +module ElimProp {ℓ'} {B : AssocList A → Type ℓ'} (BProp : {xs : AssocList A} → isProp (B xs)) + (⟨⟩* : B ⟨⟩) (⟨_,_⟩∷*_ : (x : A) (n : ℕ) {xs : AssocList A} → B xs → B (⟨ x , n ⟩∷ xs)) where + + f : (xs : AssocList A) → B xs + f = Elim.f ⟨⟩* ⟨_,_⟩∷*_ + (λ x y {xs} b → toPathP (BProp (transp (λ i → B (per x y xs i)) i0 (⟨ x , 1 ⟩∷* ⟨ y , 1 ⟩∷* b)) (⟨ y , 1 ⟩∷* ⟨ x , 1 ⟩∷* b))) + (λ x m n {xs} b → toPathP (BProp (transp (λ i → B (agg x m n xs i)) i0 (⟨ x , m ⟩∷* ⟨ x , n ⟩∷* b)) (⟨ x , m + n ⟩∷* b))) + (λ x {xs} b → toPathP (BProp (transp (λ i → B (del x xs i)) i0 (⟨ x , 0 ⟩∷* b)) b)) + (λ xs → isProp→isSet BProp) + + + +module Rec {ℓ'} {B : Type ℓ'} (BType : isSet B) + (⟨⟩* : B) (⟨_,_⟩∷*_ : (x : A) (n : ℕ) → B → B) + (per* : (x y : A) (b : B) → (⟨ x , 1 ⟩∷* ⟨ y , 1 ⟩∷* b) ≡ (⟨ y , 1 ⟩∷* ⟨ x , 1 ⟩∷* b)) + (agg* : (x : A) (m n : ℕ) (b : B) → (⟨ x , m ⟩∷* ⟨ x , n ⟩∷* b) ≡ (⟨ x , m + n ⟩∷* b)) + (del* : (x : A) (b : B) → (⟨ x , 0 ⟩∷* b) ≡ b) where + + f : AssocList A → B + f = Elim.f ⟨⟩* (λ x n b → ⟨ x , n ⟩∷* b) (λ x y b → per* x y b) (λ x m n b → agg* x m n b) (λ x b → del* x b) (λ _ → BType) diff --git a/Cubical/HITs/AssocList/Properties.agda b/Cubical/HITs/AssocList/Properties.agda new file mode 100644 index 0000000000..2133961615 --- /dev/null +++ b/Cubical/HITs/AssocList/Properties.agda @@ -0,0 +1,162 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.AssocList.Properties where + +open import Cubical.HITs.AssocList.Base as AL +open import Cubical.Foundations.Everything +open import Cubical.Foundations.SIP +open import Cubical.HITs.FiniteMultiset as FMS +open import Cubical.Data.Nat using (ℕ; zero; suc; _+_; +-assoc; isSetℕ) +open import Cubical.Structures.MultiSet +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidableEq + +private + variable + ℓ : Level + A : Type ℓ + + + +multiPer : (a b : A) (m n : ℕ) (xs : AssocList A) + → ⟨ a , m ⟩∷ ⟨ b , n ⟩∷ xs ≡ ⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs +multiPer a b zero n xs = del a (⟨ b , n ⟩∷ xs) ∙ cong (λ ys → ⟨ b , n ⟩∷ ys) (sym (del a xs)) +multiPer a b (suc m) zero xs = cong (λ ys → ⟨ a , suc m ⟩∷ ys) (del b xs) ∙ sym (del b (⟨ a , suc m ⟩∷ xs)) +multiPer a b (suc m) (suc n) xs = + + ⟨ a , suc m ⟩∷ ⟨ b , suc n ⟩∷ xs ≡⟨ sym (agg a 1 m (⟨ b , suc n ⟩∷ xs)) ⟩ + ⟨ a , 1 ⟩∷ ⟨ a , m ⟩∷ ⟨ b , suc n ⟩∷ xs ≡⟨ cong (λ ys → ⟨ a , 1 ⟩∷ ys) (multiPer a b m (suc n) xs) ⟩ + ⟨ a , 1 ⟩∷ ⟨ b , suc n ⟩∷ ⟨ a , m ⟩∷ xs ≡⟨ cong (λ ys → ⟨ a , 1 ⟩∷ ys) (sym (agg b 1 n (⟨ a , m ⟩∷ xs))) ⟩ + ⟨ a , 1 ⟩∷ ⟨ b , 1 ⟩∷ ⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs ≡⟨ per a b (⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs) ⟩ + ⟨ b , 1 ⟩∷ ⟨ a , 1 ⟩∷ ⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs ≡⟨ cong (λ ys → ⟨ b , 1 ⟩∷ ⟨ a , 1 ⟩∷ ys) (multiPer b a n m xs) ⟩ + ⟨ b , 1 ⟩∷ ⟨ a , 1 ⟩∷ ⟨ a , m ⟩∷ ⟨ b , n ⟩∷ xs ≡⟨ cong (λ ys → ⟨ b , 1 ⟩∷ ys) (agg a 1 m (⟨ b , n ⟩∷ xs)) ⟩ + ⟨ b , 1 ⟩∷ ⟨ a , suc m ⟩∷ ⟨ b , n ⟩∷ xs ≡⟨ cong (λ ys → ⟨ b , 1 ⟩∷ ys) (multiPer a b (suc m) n xs) ⟩ + ⟨ b , 1 ⟩∷ ⟨ b , n ⟩∷ ⟨ a , suc m ⟩∷ xs ≡⟨ agg b 1 n (⟨ a , suc m ⟩∷ xs) ⟩ + ⟨ b , suc n ⟩∷ ⟨ a , suc m ⟩∷ xs ∎ + + + + + +-- Show that association lists and finite multisets are equivalent +multi-∷ : A → ℕ → FMSet A → FMSet A +multi-∷ x zero xs = xs +multi-∷ x (suc n) xs = x ∷ multi-∷ x n xs + +multi-∷-agg : (x : A) (m n : ℕ) (b : FMSet A) → multi-∷ x m (multi-∷ x n b) ≡ multi-∷ x (m + n) b +multi-∷-agg x zero n b = refl +multi-∷-agg x (suc m) n b i = x ∷ (multi-∷-agg x m n b i) + +AL→FMS : AssocList A → FMSet A +AL→FMS = AL.Rec.f FMS.trunc [] multi-∷ comm multi-∷-agg λ _ _ → refl + + +FMS→AL : FMSet A → AssocList A +FMS→AL = FMS.Rec.f AL.trunc ⟨⟩ (λ x xs → ⟨ x , 1 ⟩∷ xs) per + + + +AL→FMS∘FMS→AL≡id : section {A = AssocList A} AL→FMS FMS→AL +AL→FMS∘FMS→AL≡id = FMS.ElimProp.f (FMS.trunc _ _) refl (λ x p → cong (λ ys → x ∷ ys) p) + + +-- need a little lemma for other direction +multi-∷-id : (x : A) (n : ℕ) (u : FMSet A) + → FMS→AL (multi-∷ x n u) ≡ ⟨ x , n ⟩∷ FMS→AL u +multi-∷-id x zero u = sym (del x (FMS→AL u)) +multi-∷-id x (suc n) u = FMS→AL (multi-∷ x (suc n) u) ≡⟨ cong (λ ys → ⟨ x , 1 ⟩∷ ys) (multi-∷-id x n u) ⟩ + ⟨ x , 1 ⟩∷ ⟨ x , n ⟩∷ (FMS→AL u) ≡⟨ agg x 1 n (FMS→AL u) ⟩ + ⟨ x , (suc n) ⟩∷ (FMS→AL u) ∎ + +FMS→AL∘AL→FMS≡id : retract {A = AssocList A} AL→FMS FMS→AL +FMS→AL∘AL→FMS≡id = AL.ElimProp.f (AL.trunc _ _) refl (λ x n {xs} p → (multi-∷-id x n (AL→FMS xs)) ∙ cong (λ ys → ⟨ x , n ⟩∷ ys) p) + + + + +AssocList≃FMSet : AssocList A ≃ FMSet A +AssocList≃FMSet = isoToEquiv (iso AL→FMS FMS→AL AL→FMS∘FMS→AL≡id FMS→AL∘AL→FMS≡id) + +FMSet≃AssocList : FMSet A ≃ AssocList A +FMSet≃AssocList = isoToEquiv (iso FMS→AL AL→FMS FMS→AL∘AL→FMS≡id AL→FMS∘FMS→AL≡id) + + +AssocList≡FMSet : AssocList A ≡ FMSet A +AssocList≡FMSet = ua AssocList≃FMSet + + + + +-- We want to define a multiset structure on AssocList A, we use the recursor to define the count-function + + +module _(discA : Discrete A) where + setA = Discrete→isSet discA + + + + ALcount-⟨,⟩∷* : A → A → ℕ → ℕ → ℕ + ALcount-⟨,⟩∷* a x n xs with discA a x + ... | yes a≡x = n + xs + ... | no a≢x = xs + + + ALcount-per* : (a x y : A) (xs : ℕ) + → ALcount-⟨,⟩∷* a x 1 (ALcount-⟨,⟩∷* a y 1 xs) + ≡ ALcount-⟨,⟩∷* a y 1 (ALcount-⟨,⟩∷* a x 1 xs) + ALcount-per* a x y xs with discA a x | discA a y + ALcount-per* a x y xs | yes a≡x | yes a≡y = refl + ALcount-per* a x y xs | yes a≡x | no a≢y = refl + ALcount-per* a x y xs | no a≢x | yes a≡y = refl + ALcount-per* a x y xs | no a≢x | no a≢y = refl + + + ALcount-agg* : (a x : A) (m n xs : ℕ) + → ALcount-⟨,⟩∷* a x m (ALcount-⟨,⟩∷* a x n xs) + ≡ ALcount-⟨,⟩∷* a x (m + n) xs + ALcount-agg* a x m n xs with discA a x + ... | yes _ = +-assoc m n xs + ... | no _ = refl + + ALcount-del* : (a x : A) (xs : ℕ) → ALcount-⟨,⟩∷* a x 0 xs ≡ xs + ALcount-del* a x xs with discA a x + ... | yes _ = refl + ... | no _ = refl + + + ALcount : A → AssocList A → ℕ + ALcount a = AL.Rec.f isSetℕ 0 (ALcount-⟨,⟩∷* a) (ALcount-per* a) (ALcount-agg* a) (ALcount-del* a) + + + + AL-with-str : MultiSet A setA + AL-with-str = (AssocList A , ⟨⟩ , ⟨_, 1 ⟩∷_ , ALcount) + + +-- We want to show that Al-with-str ≅ FMS-with-str as multiset-structures + FMS→AL-EquivStr : MultiSetEquivStr A setA (FMS-with-str discA) (AL-with-str) FMSet≃AssocList + FMS→AL-EquivStr = refl , (λ a xs → refl) , φ + where + φ : ∀ a xs → FMScount discA a xs ≡ ALcount a (FMS→AL xs) + φ a = FMS.ElimProp.f (isSetℕ _ _) refl ψ + where + ψ : (x : A) {xs : FMSet A} + → FMScount discA a xs ≡ ALcount a (FMS→AL xs) + → FMScount discA a (x ∷ xs) ≡ ALcount a (FMS→AL (x ∷ xs)) + ψ x {xs} p = subst B α θ + where + B = λ ys → FMScount discA a (x ∷ xs) ≡ ALcount a ys + + α : ⟨ x , 1 ⟩∷ FMS→AL xs ≡ FMS→AL (x ∷ xs) + α = sym (multi-∷-id x 1 xs) + + θ : FMScount discA a (x ∷ xs) ≡ ALcount a (⟨ x , 1 ⟩∷ (FMS→AL xs)) + θ with discA a x + ... | yes _ = cong suc p + ... | no ¬p = p + + + + FMS-with-str≡AL-with-str : FMS-with-str discA ≡ AL-with-str + FMS-with-str≡AL-with-str = sip (multiSetUnivalentStr A setA) + (FMS-with-str discA) AL-with-str + (FMSet≃AssocList , FMS→AL-EquivStr) diff --git a/Cubical/HITs/Colimit.agda b/Cubical/HITs/Colimit.agda new file mode 100644 index 0000000000..18c956c576 --- /dev/null +++ b/Cubical/HITs/Colimit.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Colimit where + +open import Cubical.HITs.Colimit.Base public +open import Cubical.HITs.Colimit.Examples public diff --git a/Cubical/HITs/Colimit/Base.agda b/Cubical/HITs/Colimit/Base.agda new file mode 100644 index 0000000000..004335fef0 --- /dev/null +++ b/Cubical/HITs/Colimit/Base.agda @@ -0,0 +1,132 @@ +{- + + Homotopy colimits of graphs + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.Colimit.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Graph + +-- Cones under a diagram + +record Cocone ℓ {ℓd ℓv ℓe} {I : Graph ℓv ℓe} (F : Diag ℓd I) (X : Type ℓ) + : Type (ℓ-suc (ℓ-max ℓ (ℓ-max (ℓ-max ℓv ℓe) (ℓ-suc ℓd)))) where + field + leg : ∀ (j : Node I) → F $ j → X + com : ∀ {j k} (f : Edge I j k) → leg k ∘ F <$> f ≡ leg j + + postcomp : ∀ {ℓ'} {Y : Type ℓ'} → (X → Y) → Cocone ℓ' F Y + leg (postcomp h) j = h ∘ leg j + com (postcomp h) f = cong (h ∘_) (com f) + +open Cocone public + + +-- Σ (Type ℓ) (Cocone ℓ F) forms a category: + +module _ {ℓd ℓv ℓe} {I : Graph ℓv ℓe} {F : Diag ℓd I} where + + private + -- the "lower star" functor + _* : ∀ {ℓ ℓ'} {X : Type ℓ} {Y : Type ℓ'} → (X → Y) → Cocone _ F X → Cocone _ F Y + (h *) C = postcomp C h + + CoconeMor : ∀ {ℓ ℓ'} → Σ (Type ℓ) (Cocone ℓ F) → Σ (Type ℓ') (Cocone ℓ' F) → Type _ + CoconeMor (X , C) (Y , D) = Σ[ h ∈ (X → Y) ] (h *) C ≡ D + + idCoconeMor : ∀ {ℓ} (Cp : Σ (Type ℓ) (Cocone ℓ F)) → CoconeMor Cp Cp + idCoconeMor Cp = (λ x → x) , refl + + compCoconeMor : ∀ {ℓ ℓ' ℓ''} {C : Σ (Type ℓ) (Cocone ℓ F)} {D : Σ (Type ℓ') (Cocone ℓ' F)} + {E : Σ (Type ℓ'') (Cocone ℓ'' F)} + → CoconeMor D E → CoconeMor C D → CoconeMor C E + compCoconeMor (g , q) (f , p) = g ∘ f , (cong (g *) p) ∙ q + + +-- Universal cocones are initial objects in the category Σ (Type ℓ) (Cocone ℓ F) + +module _ {ℓ ℓd ℓv ℓe} {I : Graph ℓv ℓe} {F : Diag ℓd I} {X : Type ℓ} where + + isUniversalAt : ∀ ℓq → Cocone ℓ F X → Type (ℓ-max ℓ (ℓ-suc (ℓ-max ℓq (ℓ-max (ℓ-max ℓv ℓe) (ℓ-suc ℓd))))) + isUniversalAt ℓq C = ∀ (Y : Type ℓq) → isEquiv {A = (X → Y)} {B = Cocone ℓq F Y} (postcomp C) + -- (unfolding isEquiv, this ^ is equivalent to what one might expect:) + -- ∀ (Y : Type ℓ) (D : Cocone ℓ F Y) → ∃![ h ∈ (X → Y) ] (h *) C ≡ D + -- (≡ isContr (CoconeMor (X , C) (Y , D))) + + isPropIsUniversalAt : ∀ ℓq (C : Cocone ℓ F X) → isProp (isUniversalAt ℓq C) + isPropIsUniversalAt ℓq C = isPropΠ (λ Y → isPropIsEquiv (postcomp C)) + + isUniversal : Cocone ℓ F X → Typeω + isUniversal C = ∀ ℓq → isUniversalAt ℓq C + + +-- Colimits are universal cocones + +record isColimit {ℓ ℓd ℓv ℓe} {I : Graph ℓv ℓe} (F : Diag ℓd I) (X : Type ℓ) : Typeω where + field + cone : Cocone ℓ F X + univ : isUniversal cone +open isColimit public + +module _ {ℓ ℓ' ℓd ℓv ℓe} {I : Graph ℓv ℓe} {F : Diag ℓd I} {X : Type ℓ} {Y : Type ℓ'} where + + postcomp⁻¹ : isColimit F X → Cocone ℓ' F Y → (X → Y) + postcomp⁻¹ cl = invEq (_ , univ cl _ Y) + + postcomp⁻¹-inv : (cl : isColimit F X) (D : Cocone ℓ' F Y) → (postcomp (cone cl) (postcomp⁻¹ cl D)) ≡ D + postcomp⁻¹-inv cl D = secEq (_ , univ cl _ Y) D + + postcomp⁻¹-mor : (cl : isColimit F X) (D : Cocone ℓ' F Y) → CoconeMor (X , cone cl) (Y , D) + postcomp⁻¹-mor cl D = (postcomp⁻¹ cl D) , (postcomp⁻¹-inv cl D) + +-- Colimits are unique + +module _ {ℓ ℓ' ℓd ℓv ℓe} {I : Graph ℓv ℓe} {F : Diag ℓd I} {X : Type ℓ} {Y : Type ℓ'} where + + uniqColimit : isColimit F X → isColimit F Y → X ≃ Y + uniqColimit cl cl' + = isoToEquiv (iso (fst fwd) (fst bwd) + (λ x i → fst (isContr→isProp (equiv-proof (univ cl' ℓ' Y) (cone cl')) + (compCoconeMor fwd bwd) + (idCoconeMor (Y , cone cl')) i) x) + (λ x i → fst (isContr→isProp (equiv-proof (univ cl ℓ X) (cone cl)) + (compCoconeMor bwd fwd) + (idCoconeMor (X , cone cl)) i) x)) + where fwd : CoconeMor (X , cone cl ) (Y , cone cl') + bwd : CoconeMor (Y , cone cl') (X , cone cl ) + fwd = postcomp⁻¹-mor cl (cone cl') + bwd = postcomp⁻¹-mor cl' (cone cl) + +-- Colimits always exist + +data colim {ℓd ℓe ℓv} {I : Graph ℓv ℓe} (F : Diag ℓd I) : Type (ℓ-suc (ℓ-max (ℓ-max ℓv ℓe) (ℓ-suc ℓd))) where + colim-leg : ∀ (j : Node I) → F $ j → colim F + colim-com : ∀ {j k} (f : Edge I j k) → colim-leg k ∘ F <$> f ≡ colim-leg j + +module _ {ℓd ℓv ℓe} {I : Graph ℓv ℓe} {F : Diag ℓd I} where + + colimCone : Cocone _ F (colim F) + leg colimCone = colim-leg + com colimCone = colim-com + + rec : ∀ {ℓ} {X : Type ℓ} → Cocone ℓ F X → (colim F → X) + rec C (colim-leg j A) = leg C j A + rec C (colim-com f i A) = com C f i A + + colimIsColimit : isColimit F (colim F) + cone colimIsColimit = colimCone + univ colimIsColimit ℓq Y + = isoToIsEquiv record { fun = postcomp colimCone + ; inv = rec + ; rightInv = λ C → refl + ; leftInv = λ h → funExt (eq h) } + where eq : ∀ h (x : colim _) → rec (postcomp colimCone h) x ≡ h x + eq h (colim-leg j A) = refl + eq h (colim-com f i A) = refl diff --git a/Cubical/HITs/Colimit/Examples.agda b/Cubical/HITs/Colimit/Examples.agda new file mode 100644 index 0000000000..e2e926d2f0 --- /dev/null +++ b/Cubical/HITs/Colimit/Examples.agda @@ -0,0 +1,76 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Colimit.Examples where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.SumFin + +open import Cubical.Data.Graph +open import Cubical.HITs.Colimit.Base + +open import Cubical.HITs.Pushout + + +-- Pushouts are colimits over the graph ⇐⇒ + +module _ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} where + + PushoutDiag : (A → B) → (A → C) → Diag (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) ⇐⇒ + (PushoutDiag f g) $ fzero = Lift {j = ℓ-max ℓ ℓ'' } B + (PushoutDiag f g) $ fsuc fzero = Lift {j = ℓ-max ℓ' ℓ'' } A + (PushoutDiag f g) $ fsuc (fsuc fzero) = Lift {j = ℓ-max ℓ ℓ' } C + _<$>_ (PushoutDiag f g) {fsuc fzero} {fzero} tt (lift a) = lift (f a) + _<$>_ (PushoutDiag f g) {fsuc fzero} {fsuc (fsuc fzero)} tt (lift a) = lift (g a) + +module _ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {f : A → B} {g : A → C} where + + PushoutCocone : Cocone _ (PushoutDiag f g) (Pushout f g) + leg PushoutCocone fzero (lift b) = inl b + leg PushoutCocone (fsuc fzero) (lift a) = inr (g a) + leg PushoutCocone (fsuc (fsuc fzero)) (lift c) = inr c + com PushoutCocone {fsuc fzero} {fzero} tt i (lift a) = push a i + com PushoutCocone {fsuc fzero} {fsuc (fsuc fzero)} tt i (lift a) = inr (g a) + + private + module _ ℓq (Y : Type ℓq) where + + fwd : (Pushout f g → Y) → Cocone ℓq (PushoutDiag f g) Y + fwd = postcomp PushoutCocone + + module _ (C : Cocone ℓq (PushoutDiag f g) Y) where + coml : ∀ a → leg C fzero (lift (f a)) ≡ leg C (fsuc fzero) (lift a) + comr : ∀ a → leg C (fsuc (fsuc fzero)) (lift (g a)) ≡ leg C (fsuc fzero) (lift a) + coml a i = com C {j = fsuc fzero} {k = fzero} tt i (lift a) + comr a i = com C {j = fsuc fzero} {k = fsuc (fsuc fzero)} tt i (lift a) + + bwd : Cocone ℓq (PushoutDiag f g) Y → (Pushout f g → Y) + bwd C (inl b) = leg C fzero (lift b) + bwd C (inr c) = leg C (fsuc (fsuc fzero)) (lift c) + bwd C (push a i) = (coml C a ∙ sym (comr C a)) i + + bwd-fwd : ∀ F → bwd (fwd F) ≡ F + bwd-fwd F i (inl b) = F (inl b) + bwd-fwd F i (inr c) = F (inr c) + bwd-fwd F i (push a j) = compPath-filler (coml (fwd F) a) (sym (comr (fwd F) a)) (~ i) j + + fwd-bwd : ∀ C → fwd (bwd C) ≡ C + leg (fwd-bwd C i) fzero (lift b) = leg C fzero (lift b) + leg (fwd-bwd C i) (fsuc fzero) (lift a) = comr C a i + leg (fwd-bwd C i) (fsuc (fsuc fzero)) (lift c) = leg C (fsuc (fsuc fzero)) (lift c) + com (fwd-bwd C i) {fsuc fzero} {fzero} tt j (lift a) -- coml (fwd-bwd C i) = ... + = compPath-filler (coml C a) (sym (comr C a)) (~ i) j + com (fwd-bwd C i) {fsuc fzero} {fsuc (fsuc fzero)} tt j (lift a) -- comr (fwd-bwd C i) = ... + = comr C a (i ∧ j) + + eqv : isEquiv {A = (Pushout f g → Y)} {B = Cocone ℓq (PushoutDiag f g) Y} (postcomp PushoutCocone) + eqv = isoToIsEquiv (iso fwd bwd fwd-bwd bwd-fwd) + + isColimPushout : isColimit (PushoutDiag f g) (Pushout f g) + cone isColimPushout = PushoutCocone + univ isColimPushout = eqv + + colim≃Pushout : colim (PushoutDiag f g) ≃ Pushout f g + colim≃Pushout = uniqColimit colimIsColimit isColimPushout + diff --git a/Cubical/HITs/Cost.agda b/Cubical/HITs/Cost.agda new file mode 100644 index 0000000000..6e9e6dc965 --- /dev/null +++ b/Cubical/HITs/Cost.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Cost where + +open import Cubical.HITs.Cost.Base diff --git a/Cubical/HITs/Cost/Base.agda b/Cubical/HITs/Cost/Base.agda new file mode 100644 index 0000000000..483d1bc7f7 --- /dev/null +++ b/Cubical/HITs/Cost/Base.agda @@ -0,0 +1,124 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Cost.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.HITs.PropositionalTruncation +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +private + variable + ℓ : Level + A B C : Type ℓ + +Cost : (A : Type ℓ) → Type ℓ +Cost A = A × ∥ ℕ ∥ + +-- To compare two elements of Cost A we only need to look at the first parts +Cost≡ : (x y : Cost A) → x .fst ≡ y .fst → x ≡ y +Cost≡ _ _ = Σ≡Prop λ _ → squash + +-- To make it easier to program with Cost A we prove that it forms a +-- monad which counts the number of calls to >>=. We could also turn +-- it into a proper state monad which tracks the cost, but for the +-- programs we write later this simple version suffices +_>>=_ : Cost A → (A → Cost B) → Cost B +(x , m) >>= g with g x +... | (y , n) = (y , map suc (map2 _+_ m n)) + +return : A → Cost A +return x = (x , ∣ 0 ∣) + +-- The monad laws all hold by Cost≡ + +>>=-return : (f : Cost A) → f >>= return ≡ f +>>=-return f = Cost≡ _ _ refl + +return->>= : (a : A) (f : A → Cost B) → return a >>= f ≡ f a +return->>= a f = Cost≡ _ _ refl + +>>=-assoc : (f : Cost A) (g : A → Cost B) (h : B → Cost C) + → (f >>= g) >>= h ≡ f >>= (λ x → g x >>= h) +>>=-assoc f g h = Cost≡ _ _ refl + + +-- An inefficient version of addition which recurses on both arguments +addBad : ℕ → ℕ → Cost ℕ +addBad 0 0 = return 0 +addBad 0 (suc y) = do + x ← addBad 0 y + return (suc x) +addBad (suc x) y = do + z ← addBad x y + return (suc z) + +-- More efficient addition which only recurse on the first argument +add : ℕ → ℕ → Cost ℕ +add 0 y = return y +add (suc x) y = do + z ← add x y + return (suc z) + +private + -- addBad x y will do x + y calls + _ : addBad 3 5 ≡ (8 , ∣ 8 ∣) + _ = refl + + -- add x y will only do x recursive calls + _ : add 3 5 ≡ (8 , ∣ 3 ∣) + _ = refl + +-- Despite addBad and add having different cost we can still prove +-- that they are equal functions +add≡addBad : add ≡ addBad +add≡addBad i x y = Cost≡ (add x y) (addBad x y) (rem x y) i + where + rem : (x y : ℕ) → add x y .fst ≡ addBad x y .fst + rem 0 0 = refl + rem 0 (suc y) = cong suc (rem 0 y) + rem (suc x) y = cong suc (rem x y) + + +-- Another example: computing Fibonacci numbers + +fib : ℕ → Cost ℕ +fib 0 = return 0 +fib 1 = return 1 +fib (suc (suc x)) = do + y ← fib (suc x) + z ← fib x + return (y + z) + +fibTail : ℕ → Cost ℕ +fibTail 0 = return 0 +fibTail (suc x) = fibAux x 1 0 + where + fibAux : ℕ → ℕ → ℕ → Cost ℕ + fibAux 0 res _ = return res + fibAux (suc x) res prev = do + r ← fibAux x (res + prev) res + return r + + +private + _ : fib 10 ≡ (55 , ∣ 176 ∣) + _ = refl + + _ : fibTail 10 ≡ (55 , ∣ 9 ∣) + _ = refl + + _ : fib 20 ≡ (6765 , ∣ 21890 ∣) + _ = refl + + _ : fibTail 20 ≡ (6765 , ∣ 19 ∣) + _ = refl + +-- Exercise left to the reader: prove that fib and fibTail are equal functions +-- fib≡fibTail : fib ≡ fibTail +-- fib≡fibTail i x = Cost≡ (fib x) (fibTail x) (rem x) i +-- where +-- rem : (x : ℕ) → fib x .fst ≡ fibTail x .fst +-- rem zero = refl +-- rem (suc zero) = refl +-- rem (suc (suc x)) = {!!} diff --git a/Cubical/HITs/Cylinder.agda b/Cubical/HITs/Cylinder.agda new file mode 100644 index 0000000000..36c27f0103 --- /dev/null +++ b/Cubical/HITs/Cylinder.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Cylinder where + +open import Cubical.HITs.Cylinder.Base public + +-- open import Cubical.HITs.Cylinder.Properties public diff --git a/Cubical/HITs/Cylinder/Base.agda b/Cubical/HITs/Cylinder/Base.agda new file mode 100644 index 0000000000..6331d0384b --- /dev/null +++ b/Cubical/HITs/Cylinder/Base.agda @@ -0,0 +1,265 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.Cylinder.Base where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Everything + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Data.Sum using (_⊎_; inl; inr) + +open import Cubical.HITs.PropositionalTruncation using (∥_∥; ∣_∣; squash) +open import Cubical.HITs.Interval using (Interval; zero; one; seg) + +-- Cylinder A is a cylinder object in the category of cubical types. +-- +-- https://ncatlab.org/nlab/show/cylinder+object +data Cylinder {ℓ} (A : Type ℓ) : Type ℓ where + inl : A → Cylinder A + inr : A → Cylinder A + cross : ∀ x → inl x ≡ inr x + +-- Dual to this is the cocylinder or path space object. +-- +-- https://ncatlab.org/nlab/show/path+space+object +Cocylinder : ∀ {ℓ} → Type ℓ → Type ℓ +Cocylinder A = Interval → A + +module _ {ℓ} {A : Type ℓ} where + -- The cylinder is part of a factorization of the obvious mapping + -- of type A ⊎ A → A into a pair of mappings: + -- + -- A ⊎ A → Cylinder A ≃ A + -- + -- include is the first part of the factorization. + include : A ⊎ A → Cylinder A + include (inl x) = inl x + include (inr x) = inr x + + -- The above inclusion is surjective + includeSurjective : ∀ c → ∃[ s ∈ A ⊎ A ] include s ≡ c + includeSurjective (inl x) = ∣ inl x , refl ∣ + includeSurjective (inr x) = ∣ inr x , refl ∣ + includeSurjective (cross x i) = + squash + ∣ inl x , (λ j → cross x (i ∧ j)) ∣ + ∣ inr x , (λ j → cross x (i ∨ ~ j)) ∣ + i + + elim + : ∀{ℓ'} {B : Cylinder A → Type ℓ'} + → (f : (x : A) → B (inl x)) + → (g : (x : A) → B (inr x)) + → (p : ∀ x → PathP (λ i → B (cross x i)) (f x) (g x)) + → (c : Cylinder A) → B c + elim f _ _ (inl x) = f x + elim _ g _ (inr x) = g x + elim _ _ p (cross x i) = p x i + + private + out : Cylinder A → A + out (inl x) = x + out (inr x) = x + out (cross x i) = x + + inl-out : (c : Cylinder A) → inl (out c) ≡ c + inl-out (inl x) = refl + inl-out (inr x) = cross x + inl-out (cross x i) = λ j → cross x (i ∧ j) + + out-inl : ∀(x : A) → out (inl x) ≡ x + out-inl x = refl + + -- The second part of the factorization above. + CylinderA≃A : Cylinder A ≃ A + CylinderA≃A = isoToEquiv (iso out inl out-inl inl-out) + + -- The cocylinder has a similar equivalence that is part + -- of factorizing the diagonal mapping. + private + inco : A → Cocylinder A + inco x _ = x + + outco : Cocylinder A → A + outco f = f zero + + A→CocylinderA→A : (x : A) → outco (inco x) ≡ x + A→CocylinderA→A x = refl + + CocylinderA→A→CocylinderA : (c : Cocylinder A) → inco (outco c) ≡ c + CocylinderA→A→CocylinderA c j zero = c zero + CocylinderA→A→CocylinderA c j one = c (seg j) + CocylinderA→A→CocylinderA c j (seg i) = c (seg (j ∧ i)) + + A≃CocylinderA : A ≃ Cocylinder A + A≃CocylinderA = + isoToEquiv (iso inco outco CocylinderA→A→CocylinderA A→CocylinderA→A) + + project : Cocylinder A → A × A + project c = c zero , c one + + +-- Since we can construct cylinders for every type, Cylinder actually +-- constitutes a cylinder functor: +-- +-- https://ncatlab.org/nlab/show/cylinder+functor +-- +-- e₀ = inl +-- e₁ = inr +-- σ = out +module Functorial where + private + variable + ℓa ℓb ℓc : Level + A : Type ℓa + B : Type ℓb + C : Type ℓc + + map : (A → B) → Cylinder A → Cylinder B + map f (inl x) = inl (f x) + map f (inr x) = inr (f x) + map f (cross x i) = cross (f x) i + + mapId : map (λ(x : A) → x) ≡ (λ x → x) + mapId i (inl x) = inl x + mapId i (inr x) = inr x + mapId i (cross x j) = cross x j + + map∘ + : (f : A → B) → (g : B → C) + → map (λ x → g (f x)) ≡ (λ x → map g (map f x)) + map∘ f g i (inl x) = inl (g (f x)) + map∘ f g i (inr x) = inr (g (f x)) + map∘ f g i (cross x j) = cross (g (f x)) j + + -- There is an adjunction between the cylinder and coyclinder + -- functors. + -- + -- Cylinder ⊣ Cocylinder + adj₁ : (Cylinder A → B) → A → Cocylinder B + adj₁ f x zero = f (inl x) + adj₁ f x one = f (inr x) + adj₁ f x (seg i) = f (cross x i) + + adj₂ : (A → Cocylinder B) → Cylinder A → B + adj₂ g (inl x) = g x zero + adj₂ g (inr x) = g x one + adj₂ g (cross x i) = g x (seg i) + + adj₁₂ : (g : A → Cocylinder B) → adj₁ (adj₂ g) ≡ g + adj₁₂ g _ x zero = g x zero + adj₁₂ g _ x one = g x one + adj₁₂ g _ x (seg i) = g x (seg i) + + adj₂₁ : (f : Cylinder A → B) → adj₂ (adj₁ f) ≡ f + adj₂₁ f j (inl x) = f (inl x) + adj₂₁ f j (inr x) = f (inr x) + adj₂₁ f j (cross x i) = f (cross x i) + +module IntervalEquiv where + -- There is an equivalence between the interval and the + -- cylinder over the unit type. + Interval→CylinderUnit : Interval → Cylinder Unit + Interval→CylinderUnit zero = inl _ + Interval→CylinderUnit one = inr _ + Interval→CylinderUnit (seg i) = cross _ i + + CylinderUnit→Interval : Cylinder Unit → Interval + CylinderUnit→Interval (inl _) = zero + CylinderUnit→Interval (inr _) = one + CylinderUnit→Interval (cross _ i) = seg i + + Interval→CylinderUnit→Interval + : ∀ i → CylinderUnit→Interval (Interval→CylinderUnit i) ≡ i + Interval→CylinderUnit→Interval zero = refl + Interval→CylinderUnit→Interval one = refl + Interval→CylinderUnit→Interval (seg i) = refl + + CylinderUnit→Interval→CylinderUnit + : ∀ c → Interval→CylinderUnit (CylinderUnit→Interval c) ≡ c + CylinderUnit→Interval→CylinderUnit (inl _) = refl + CylinderUnit→Interval→CylinderUnit (inr _) = refl + CylinderUnit→Interval→CylinderUnit (cross _ i) = refl + + CylinderUnit≃Interval : Cylinder Unit ≃ Interval + CylinderUnit≃Interval = + isoToEquiv (iso CylinderUnit→Interval Interval→CylinderUnit Interval→CylinderUnit→Interval CylinderUnit→Interval→CylinderUnit) + + + -- More generally, there is an equivalence between the cylinder + -- over any type A and the product of A and the interval. + module _ {ℓ} {A : Type ℓ} where + private + Cyl : Type ℓ + Cyl = A × Interval + + CylinderA→A×Interval : Cylinder A → Cyl + CylinderA→A×Interval (inl x) = x , zero + CylinderA→A×Interval (inr x) = x , one + CylinderA→A×Interval (cross x i) = x , seg i + + A×Interval→CylinderA : Cyl → Cylinder A + A×Interval→CylinderA (x , zero) = inl x + A×Interval→CylinderA (x , one) = inr x + A×Interval→CylinderA (x , seg i) = cross x i + + A×Interval→CylinderA→A×Interval + : ∀ c → CylinderA→A×Interval (A×Interval→CylinderA c) ≡ c + A×Interval→CylinderA→A×Interval (x , zero) = refl + A×Interval→CylinderA→A×Interval (x , one) = refl + A×Interval→CylinderA→A×Interval (x , seg i) = refl + + CylinderA→A×Interval→CylinderA + : ∀ c → A×Interval→CylinderA (CylinderA→A×Interval c) ≡ c + CylinderA→A×Interval→CylinderA (inl x) = refl + CylinderA→A×Interval→CylinderA (inr x) = refl + CylinderA→A×Interval→CylinderA (cross x i) = refl + + CylinderA≃A×Interval : Cylinder A ≃ Cyl + CylinderA≃A×Interval = + isoToEquiv + (iso CylinderA→A×Interval + A×Interval→CylinderA + A×Interval→CylinderA→A×Interval + CylinderA→A×Interval→CylinderA) + +-- The cylinder is also the pushout of the identity on A with itself. +module Push {ℓ} {A : Type ℓ} where + open import Cubical.HITs.Pushout + + private + Push : Type ℓ + Push = Pushout (λ(x : A) → x) (λ x → x) + + Cyl : Type ℓ + Cyl = Cylinder A + + Cylinder→Pushout : Cyl → Push + Cylinder→Pushout (inl x) = inl x + Cylinder→Pushout (inr x) = inr x + Cylinder→Pushout (cross x i) = push x i + + Pushout→Cylinder : Push → Cyl + Pushout→Cylinder (inl x) = inl x + Pushout→Cylinder (inr x) = inr x + Pushout→Cylinder (push x i) = cross x i + + Pushout→Cylinder→Pushout : ∀ p → Cylinder→Pushout (Pushout→Cylinder p) ≡ p + Pushout→Cylinder→Pushout (inl x) = refl + Pushout→Cylinder→Pushout (inr x) = refl + Pushout→Cylinder→Pushout (push x i) = refl + + Cylinder→Pushout→Cylinder : ∀ c → Pushout→Cylinder (Cylinder→Pushout c) ≡ c + Cylinder→Pushout→Cylinder (inl x) = refl + Cylinder→Pushout→Cylinder (inr x) = refl + Cylinder→Pushout→Cylinder (cross x i) = refl + + Pushout≃Cylinder : Push ≃ Cyl + Pushout≃Cylinder = + isoToEquiv + (iso Pushout→Cylinder + Cylinder→Pushout + Cylinder→Pushout→Cylinder + Pushout→Cylinder→Pushout) diff --git a/Cubical/HITs/Delooping/Two/Base.agda b/Cubical/HITs/Delooping/Two/Base.agda new file mode 100644 index 0000000000..655cadf9c2 --- /dev/null +++ b/Cubical/HITs/Delooping/Two/Base.agda @@ -0,0 +1,55 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.Delooping.Two.Base where + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude + +-- Explicit construction of the delooping of the two-element group +data Bℤ₂ : Type₀ where + base : Bℤ₂ + loop : base ≡ base + loop² : Square loop refl refl loop + trunc : isGroupoid Bℤ₂ + +private + variable + ℓ : Level + A : Type ℓ + +module Elim where + rec : (x : A) + → (p : x ≡ x) + → (sq : Square p refl refl p) + → isGroupoid A + → Bℤ₂ → A + rec x p sq Agpd = go + where + go : _ → _ + go base = x + go (loop i) = p i + go (loop² i j) = sq i j + go (trunc x y p q r s i j k) + = Agpd + (go x) (go y) + (cong go p) (cong go q) + (cong (cong go) r) (cong (cong go) s) + i j k + + elim : (P : Bℤ₂ → Type ℓ) + → (x : P base) + → (p : PathP (λ i → P (loop i)) x x) + → (sq : SquareP (λ i j → P (loop² i j)) p (λ i → x) (λ i → x) p) + → isOfHLevelDep 3 P + → (z : Bℤ₂) → P z + elim P x p sq Pgpd = go + where + go : (z : Bℤ₂) → P z + go base = x + go (loop i) = p i + go (loop² i j) = sq i j + go (trunc x y p q r s i j k) + = Pgpd (go x) (go y) + (cong go p) (cong go q) + (cong (cong go) r) (cong (cong go) s) + (trunc x y p q r s) i j k diff --git a/Cubical/HITs/Delooping/Two/Properties.agda b/Cubical/HITs/Delooping/Two/Properties.agda new file mode 100644 index 0000000000..9a5b7bb96f --- /dev/null +++ b/Cubical/HITs/Delooping/Two/Properties.agda @@ -0,0 +1,198 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.Delooping.Two.Properties where + +open import Cubical.Functions.Involution + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Bool +open import Cubical.Data.Empty +open import Cubical.Data.Unit + +open import Cubical.HITs.Delooping.Two.Base +open import Cubical.HITs.PropositionalTruncation + +private + variable + ℓ : Level + +module Embed where + isPropDepIsSet : isOfHLevelDep {ℓ' = ℓ} 1 isSet + isPropDepIsSet = isOfHLevel→isOfHLevelDep 1 (λ _ → isPropIsSet) + + notSet : PathP (λ i → isSet (notEq i)) isSetBool isSetBool + notSet = isPropDepIsSet isSetBool isSetBool notEq + + notNot² : Square notEq refl refl notEq + notNot² = involPath² notnot + + notNotSet + : SquareP (λ i j → isSet (notNot² i j)) notSet refl refl notSet + notNotSet = isPropDep→isSetDep' + isPropDepIsSet + (involPath² notnot) + notSet refl refl notSet + + Code : Bℤ₂ → hSet ℓ-zero + Code = Elim.rec + (Bool , isSetBool) + (λ i → notEq i , notSet i) + (λ i j → λ where + .fst → notNot² i j + .snd → notNotSet i j) + (isOfHLevelTypeOfHLevel 2) + + El : Bℤ₂ → Type₀ + El b = Code b .fst + +module BINARY where + open import Cubical.Data.FinSet.Binary.Large + + sem : Bℤ₂ → Binary _ + sem = Elim.rec Base Loop Loop² isGroupoidBinary + + loop? : Bool → base ≡ base + loop? false = refl + loop? true = loop + + Loop²-coh : (a b c : Bool) → Type₀ + Loop²-coh false false false = Unit + Loop²-coh false true true = Unit + Loop²-coh true false true = Unit + Loop²-coh true true false = Unit + Loop²-coh _ _ _ = ⊥ + + rf : Bool ≡ Bool → Bool + rf P = transport P false + + Loop²-coh-lemma₀ + : ∀(p q r : Bool) + → r ⊕ p ≡ q + → Loop²-coh p q r + Loop²-coh-lemma₀ false false false sq = _ + Loop²-coh-lemma₀ false true true sq = _ + Loop²-coh-lemma₀ true false true sq = _ + Loop²-coh-lemma₀ true true false sq = _ + Loop²-coh-lemma₀ false true false = false≢true + Loop²-coh-lemma₀ false false true = true≢false + Loop²-coh-lemma₀ true false false = true≢false + Loop²-coh-lemma₀ true true true = false≢true + + Loop²-coh-lemma + : ∀(P Q R : Bool ≡ Bool) + → Square P Q refl R + → Loop²-coh (rf P) (rf Q) (rf R) + Loop²-coh-lemma P Q R sq = Loop²-coh-lemma₀ p q r eqn + where + p = rf P + q = rf Q + r = rf R + + open BoolReflection + + cmp : P ∙ R ≡ Q + cmp i j + = hcomp (λ k → λ where + (i = i0) → compPath-filler P R k j + (i = i1) → Q j + (j = i0) → Bool + (j = i1) → R (i ∨ k)) + (sq i j) + + rcmp : ⊕-Path (r ⊕ p) ≡ ⊕-Path q + rcmp = ⊕-Path (r ⊕ p) + ≡[ i ]⟨ ⊕-comp p r (~ i) ⟩ + ⊕-Path p ∙ ⊕-Path r + ≡[ i ]⟨ ⊕-complete P (~ i) ∙ ⊕-complete R (~ i) ⟩ + P ∙ R + ≡⟨ cmp ⟩ + Q + ≡⟨ ⊕-complete Q ⟩ + ⊕-Path q ∎ + open Iso + eqn : r ⊕ p ≡ q + eqn = transport (λ i → + reflectIso .leftInv (r ⊕ p) i ≡ reflectIso .leftInv q i) + (cong (reflectIso .inv) rcmp) + + loop²? + : ∀ p q r → Loop²-coh p q r + → Square (loop? p) (loop? q) refl (loop? r) + loop²? false false false _ = refl + loop²? false true true _ = λ i j → loop (i ∧ j) + loop²? true false true _ = loop² + loop²? true true false _ = refl + + module _ (B : Type₀) where + based : (P : Bool ≃ B) → Bℤ₂ + based _ = base + + pull₀ : (P Q : Bool ≃ B) → Bool ≡ Bool + pull₀ P Q i + = hcomp (λ k → λ where + (i = i0) → ua P (~ k) + (i = i1) → ua Q (~ k)) + B + + pull₁ : (P Q : Bool ≃ B) → Square (ua P) (ua Q) (pull₀ P Q) refl + pull₁ P Q i j + = hcomp (λ k → λ where + (i = i0) → ua P (~ k ∨ j) + (i = i1) → ua Q (~ k ∨ j) + (j = i1) → B) + B + + pull₂ + : (P Q R : Bool ≃ B) + → Square (pull₀ P Q) (pull₀ P R) refl (pull₀ Q R) + pull₂ P Q R i j + = hcomp (λ k → λ where + (j = i0) → ua P (~ k) + (i = i0) (j = i1) → ua Q (~ k) + (i = i1) (j = i1) → ua R (~ k)) + B + + pull₃ + : (P Q R : Bool ≃ B) + → Cube (pull₁ P Q) (pull₁ P R) + (λ _ → ua P) (pull₁ Q R) + (pull₂ P Q R) (λ _ _ → B) + pull₃ P Q R i j k + = hcomp (λ τ → λ where + (j = i0) → ua P (~ τ ∨ k) + (i = i0) (j = i1) → ua Q (~ τ ∨ k) + (i = i1) (j = i1) → ua R (~ τ ∨ k) + (k = i1) → B) + B + + looped : (P Q : Bool ≃ B) → based P ≡ based Q + looped P Q = loop? b + where + b : Bool + b = rf (pull₀ P Q) + + looped² + : (P Q R : Bool ≃ B) + → Square (looped P Q) (looped P R) refl (looped Q R) + looped² P Q R = loop²? pq pr qr pqr + where + pq = rf (pull₀ P Q) + pr = rf (pull₀ P R) + qr = rf (pull₀ Q R) + + pqr : Loop²-coh pq pr qr + pqr = Loop²-coh-lemma (pull₀ P Q) (pull₀ P R) (pull₀ Q R) (pull₂ P Q R) + + syn : Binary _ → Bℤ₂ + syn (B , tP) = rec→Gpd trunc (based B) 3k tP + where + open 3-Constant + 3k : 3-Constant (based B) + 3k .link = looped B + 3k .coh₁ = looped² B diff --git a/Cubical/HITs/DunceCap.agda b/Cubical/HITs/DunceCap.agda new file mode 100644 index 0000000000..42c01490a7 --- /dev/null +++ b/Cubical/HITs/DunceCap.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.DunceCap where + +open import Cubical.HITs.DunceCap.Base public +open import Cubical.HITs.DunceCap.Properties public diff --git a/Cubical/HITs/DunceCap/Base.agda b/Cubical/HITs/DunceCap/Base.agda new file mode 100644 index 0000000000..57f87839b9 --- /dev/null +++ b/Cubical/HITs/DunceCap/Base.agda @@ -0,0 +1,27 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.DunceCap.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.HITs.S1 using (S¹; base) +import Cubical.HITs.S1 as S¹ + +open import Cubical.HITs.MappingCones + +-- definition of the dunce cap as a HIT + +data Dunce : Type₀ where + base : Dunce + loop : base ≡ base + surf : PathP (λ i → loop i ≡ loop i) loop refl + +-- definition of the dunce cap as a mapping cone + +dunceMap : S¹ → S¹ +dunceMap base = base +dunceMap (S¹.loop i) = (S¹.loop ⁻¹ ∙∙ S¹.loop ∙∙ S¹.loop) i + +DunceCone : Type₀ +DunceCone = Cone dunceMap diff --git a/Cubical/HITs/DunceCap/Properties.agda b/Cubical/HITs/DunceCap/Properties.agda new file mode 100644 index 0000000000..a655e4d410 --- /dev/null +++ b/Cubical/HITs/DunceCap/Properties.agda @@ -0,0 +1,64 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.DunceCap.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.HITs.S1 using (S¹; base; decodeEncode) +import Cubical.HITs.S1 as S¹ + +open import Cubical.HITs.MappingCones +open import Cubical.HITs.DunceCap.Base + +-- DunceCone is contractible (easy) + +Disk : Type₀ +Disk = Cone (idfun S¹) + +isContr-Disk : isContr Disk +fst isContr-Disk = hub +snd isContr-Disk (inj x) i = spoke x i +snd isContr-Disk hub i = hub +snd isContr-Disk (spoke x j) i = spoke x (i ∧ j) + +dunceMap≡id : dunceMap ≡ idfun S¹ +dunceMap≡id i base = base +dunceMap≡id i (S¹.loop j) = p i j + where p : S¹.loop ⁻¹ ∙∙ S¹.loop ∙∙ S¹.loop ≡ S¹.loop + p = sym (decodeEncode base (S¹.loop ⁻¹ ∙∙ S¹.loop ∙∙ S¹.loop)) ∙ sym (lUnit S¹.loop) + +isContr-DunceCone : isContr DunceCone +isContr-DunceCone = subst isContr (cong Cone (sym dunceMap≡id)) isContr-Disk + +-- Dunce is contractible (harder) + +contrDunce : (d : Dunce) → d ≡ base +contrDunce base k = loop k +contrDunce (loop i) k = surf k i +contrDunce (surf i j) k + = hcomp (λ l → λ { (i = i0) → surf k j + ; (i = i1) → cube l + ; (j = i0) → cube l + ; (j = i1) → cube l + ; (k = i0) → surf i j + ; (k = i1) → base }) + (surf (k ∨ i) j) + where cube : I → Dunce + cube l = hfill (λ i → λ { (k = i0) → loop i + ; (k = i1) → base + ; (l = i0) → loop (k ∨ i) + ; (l = i1) → surf k i }) + (inS (loop k)) i + +isContr-Dunce : isContr Dunce +fst isContr-Dunce = base +snd isContr-Dunce = sym ∘ contrDunce + + +Dunce≡DunceCone : Dunce ≡ DunceCone +Dunce≡DunceCone = ua (isContr→Equiv isContr-Dunce isContr-DunceCone) diff --git a/Cubical/HITs/EilenbergMacLane1.agda b/Cubical/HITs/EilenbergMacLane1.agda new file mode 100644 index 0000000000..28ea97a274 --- /dev/null +++ b/Cubical/HITs/EilenbergMacLane1.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.EilenbergMacLane1 where + +open import Cubical.HITs.EilenbergMacLane1.Base public +open import Cubical.HITs.EilenbergMacLane1.Properties public diff --git a/Cubical/HITs/EilenbergMacLane1/Base.agda b/Cubical/HITs/EilenbergMacLane1/Base.agda new file mode 100644 index 0000000000..b28a48d271 --- /dev/null +++ b/Cubical/HITs/EilenbergMacLane1/Base.agda @@ -0,0 +1,43 @@ +{- + +This file contains: + +- The first Eilenberg–Mac Lane type as a HIT + +-} +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.EilenbergMacLane1.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Algebra.Group.Base + +private + variable ℓ : Level + +module _ (Group@(G , str) : Group ℓ) where + open GroupStr str + + data EM₁ : Type ℓ where + embase : EM₁ + emloop : G → embase ≡ embase + emcomp : (g h : G) + → PathP (λ j → embase ≡ emloop h j) (emloop g) (emloop (g · h)) + emsquash : ∀ (x y : EM₁) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s + + {- The emcomp constructor fills the square: +^ + emloop (g · h) + [a]— — — >[c] + ‖ ^ + ‖ | emloop h ^ + ‖ | j | + [a]— — — >[b] ∙ — > + emloop g i + + We use this to give another constructor-like construction: + -} + + emloop-comp : (g h : G) → emloop (g · h) ≡ emloop g ∙ emloop h + emloop-comp g h i = compPath-unique refl (emloop g) (emloop h) + (emloop (g · h) , emcomp g h) + (emloop g ∙ emloop h , compPath-filler (emloop g) (emloop h)) i .fst diff --git a/Cubical/HITs/EilenbergMacLane1/Properties.agda b/Cubical/HITs/EilenbergMacLane1/Properties.agda new file mode 100644 index 0000000000..7af601e5d4 --- /dev/null +++ b/Cubical/HITs/EilenbergMacLane1/Properties.agda @@ -0,0 +1,100 @@ +{- + +Eilenberg–Mac Lane type K(G, 1) + +-} + +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.EilenbergMacLane1.Properties where + +open import Cubical.HITs.EilenbergMacLane1.Base + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Group.Base + +open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥; ∣_∣; squash) +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) + +private + variable + ℓG ℓ : Level + +module _ ((G , str) : Group ℓG) where + + open GroupStr str + + elimSet : {B : EM₁ (G , str) → Type ℓ} + → ((x : EM₁ (G , str)) → isSet (B x)) + → (b : B embase) + → ((g : G) → PathP (λ i → B (emloop g i)) b b) + → (x : EM₁ (G , str)) + → B x + elimSet Bset b bloop embase = b + elimSet Bset b bloop (emloop g i) = bloop g i + elimSet Bset b bloop (emcomp g h i j) = + isSet→SquareP + (λ i j → Bset (emcomp g h i j)) + (λ j → bloop g j) (λ j → bloop (g · h) j) + (λ i → b) (λ i → bloop h i) + i j + elimSet Bset b bloop (emsquash x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 (λ x → isSet→isGroupoid (Bset x)) + _ _ _ _ (λ j k → g (r j k)) (λ j k → g (s j k)) (emsquash x y p q r s) i j k + where + g = elimSet Bset b bloop + + elimProp : {B : EM₁ (G , str) → Type ℓ} + → ((x : EM₁ (G , str)) → isProp (B x)) + → B embase + → (x : EM₁ (G , str)) + → B x + elimProp Bprop b x = + elimSet + (λ x → isProp→isSet (Bprop x)) + b + (λ g → isProp→PathP (λ i → Bprop ((emloop g) i)) b b) + x + + elimProp2 : {C : EM₁ (G , str) → EM₁ (G , str) → Type ℓ} + → ((x y : EM₁ (G , str)) → isProp (C x y)) + → C embase embase + → (x y : EM₁ (G , str)) + → C x y + elimProp2 Cprop c = + elimProp + (λ x → isPropΠ (λ y → Cprop x y)) + (elimProp (λ y → Cprop embase y) c) + + elim : {B : EM₁ (G , str) → Type ℓ} + → ((x : EM₁ (G , str)) → isGroupoid (B x)) + → (b : B embase) + → (bloop : (g : G) → PathP (λ i → B (emloop g i)) b b) + → ((g h : G) → SquareP (λ i j → B (emcomp g h i j)) + (bloop g) (bloop (g · h)) (λ j → b) (bloop h)) + → (x : EM₁ (G , str)) + → B x + elim Bgpd b bloop bcomp embase = b + elim Bgpd b bloop bcomp (emloop g i) = bloop g i + elim Bgpd b bloop bcomp (emcomp g h i j) = bcomp g h i j + elim Bgpd b bloop bcomp (emsquash x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 Bgpd + _ _ _ _ (λ j k → g (r j k)) (λ j k → g (s j k)) (emsquash x y p q r s) i j k + where + g = elim Bgpd b bloop bcomp + + rec : {B : Type ℓ} + → isGroupoid B + → (b : B) + → (bloop : G → b ≡ b) + → ((g h : G) → Square (bloop g) (bloop (g · h)) refl (bloop h)) + → (x : EM₁ (G , str)) + → B + rec Bgpd = elim (λ _ → Bgpd) diff --git a/Cubical/HITs/FiniteMultiset.agda b/Cubical/HITs/FiniteMultiset.agda new file mode 100644 index 0000000000..cc015b38f3 --- /dev/null +++ b/Cubical/HITs/FiniteMultiset.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.FiniteMultiset where + + +open import Cubical.HITs.FiniteMultiset.Base public +open import Cubical.HITs.FiniteMultiset.Properties public diff --git a/Cubical/HITs/FiniteMultiset/Base.agda b/Cubical/HITs/FiniteMultiset/Base.agda new file mode 100644 index 0000000000..8ca45984c5 --- /dev/null +++ b/Cubical/HITs/FiniteMultiset/Base.agda @@ -0,0 +1,50 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.FiniteMultiset.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.HITs.SetTruncation +open import Cubical.Foundations.HLevels + +private + variable + ℓ : Level + A : Type ℓ + +infixr 5 _∷_ + +data FMSet (A : Type ℓ) : Type ℓ where + [] : FMSet A + _∷_ : (x : A) → (xs : FMSet A) → FMSet A + comm : ∀ x y xs → x ∷ y ∷ xs ≡ y ∷ x ∷ xs + trunc : isSet (FMSet A) + +pattern [_] x = x ∷ [] + +module Elim {ℓ'} {B : FMSet A → Type ℓ'} + ([]* : B []) (_∷*_ : (x : A) {xs : FMSet A} → B xs → B (x ∷ xs)) + (comm* : (x y : A) {xs : FMSet A} (b : B xs) + → PathP (λ i → B (comm x y xs i)) (x ∷* (y ∷* b)) (y ∷* (x ∷* b))) + (trunc* : (xs : FMSet A) → isSet (B xs)) where + + f : (xs : FMSet A) → B xs + f [] = []* + f (x ∷ xs) = x ∷* f xs + f (comm x y xs i) = comm* x y (f xs) i + f (trunc xs zs p q i j) = + isOfHLevel→isOfHLevelDep 2 trunc* (f xs) (f zs) (cong f p) (cong f q) (trunc xs zs p q) i j + +module ElimProp {ℓ'} {B : FMSet A → Type ℓ'} (BProp : {xs : FMSet A} → isProp (B xs)) + ([]* : B []) (_∷*_ : (x : A) {xs : FMSet A} → B xs → B (x ∷ xs)) where + + f : (xs : FMSet A) → B xs + f = Elim.f []* _∷*_ + (λ x y {xs} b → + toPathP (BProp (transp (λ i → B (comm x y xs i)) i0 (x ∷* (y ∷* b))) (y ∷* (x ∷* b)))) + (λ xs → isProp→isSet BProp) + +module Rec {ℓ'} {B : Type ℓ'} (BType : isSet B) + ([]* : B) (_∷*_ : A → B → B) + (comm* : (x y : A) (b : B) → x ∷* (y ∷* b) ≡ y ∷* (x ∷* b)) where + + f : FMSet A → B + f = Elim.f []* (λ x b → x ∷* b) (λ x y b → comm* x y b) (λ _ → BType) diff --git a/Cubical/HITs/FiniteMultiset/CountExtensionality.agda b/Cubical/HITs/FiniteMultiset/CountExtensionality.agda new file mode 100644 index 0000000000..79bea106e0 --- /dev/null +++ b/Cubical/HITs/FiniteMultiset/CountExtensionality.agda @@ -0,0 +1,191 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.FiniteMultiset.CountExtensionality where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum +open import Cubical.Relation.Nullary +open import Cubical.HITs.FiniteMultiset.Base +open import Cubical.HITs.FiniteMultiset.Properties as FMS +open import Cubical.Structures.MultiSet +open import Cubical.Relation.Nullary.DecidableEq + + +private + variable + ℓ : Level + + +-- We define a partial order on FMSet A and use it to proof +-- a strong induction principle for finite multi-sets. +-- Finally, we use this stronger elimination principle to show +-- that any two FMSets can be identified, if they have the same count for every a : A +module _{A : Type ℓ} (discA : Discrete A) where + _≼_ : FMSet A → FMSet A → Type ℓ + xs ≼ ys = ∀ a → FMScount discA a xs ≤ FMScount discA a ys + + ≼-refl : ∀ xs → xs ≼ xs + ≼-refl xs a = ≤-refl + + ≼-trans : ∀ xs ys zs → xs ≼ ys → ys ≼ zs → xs ≼ zs + ≼-trans xs ys zs xs≼ys ys≼zs a = ≤-trans (xs≼ys a) (ys≼zs a) + + + ≼[]→≡[] : ∀ xs → xs ≼ [] → xs ≡ [] + ≼[]→≡[] xs xs≼[] = FMScount-0-lemma discA xs λ a → ≤0→≡0 (xs≼[] a) + + + + + ≼-remove1 : ∀ a xs → remove1 discA a xs ≼ xs + ≼-remove1 a xs b with discA a b + ... | yes a≡b = subst (λ n → n ≤ FMScount discA b xs) (sym path) (≤-predℕ) + where + path : FMScount discA b (remove1 discA a xs) ≡ predℕ (FMScount discA b xs) + path = cong (λ c → FMScount discA b (remove1 discA c xs)) a≡b ∙ remove1-predℕ-lemma discA b xs + ... | no a≢b = subst (λ n → n ≤ FMScount discA b xs) + (sym (FMScount-remove1-≢-lemma discA xs λ b≡a → a≢b (sym b≡a))) ≤-refl + + + ≼-remove1-lemma : ∀ x xs ys → ys ≼ (x ∷ xs) → (remove1 discA x ys) ≼ xs + ≼-remove1-lemma x xs ys ys≼x∷xs a with discA a x + ... | yes a≡x = ≤-trans (≤-trans (0 , p₁) (predℕ-≤-predℕ (ys≼x∷xs a))) + (0 , cong predℕ (FMScount-≡-lemma discA xs a≡x)) + where + p₁ : FMScount discA a (remove1 discA x ys) ≡ predℕ (FMScount discA a ys) + p₁ = subst (λ b → FMScount discA a (remove1 discA b ys) ≡ predℕ (FMScount discA a ys)) a≡x (remove1-predℕ-lemma discA a ys) + ... | no a≢x = ≤-trans (≤-trans (0 , FMScount-remove1-≢-lemma discA ys a≢x) (ys≼x∷xs a)) + (0 , FMScount-≢-lemma discA xs a≢x) + + + ≼-Dichotomy : ∀ x xs ys → ys ≼ (x ∷ xs) → (ys ≼ xs) ⊎ (ys ≡ x ∷ (remove1 discA x ys)) + ≼-Dichotomy x xs ys ys≼x∷xs with (FMScount discA x ys) ≟ suc (FMScount discA x xs) + ... | lt suc = ⊥.rec (¬msuc (ys≼x∷xs x)) (0 , FMScount-≡-lemma-refl discA xs) + + + + -- proving a strong elimination principle for finite multisets + module ≼-ElimProp {ℓ'} {B : FMSet A → Type ℓ'} + (BisProp : ∀ {xs} → isProp (B xs)) (b₀ : B []) + (B-≼-hyp : ∀ x xs → (∀ ys → ys ≼ xs → B ys) → B (x ∷ xs)) where + + C : FMSet A → Type (ℓ-max ℓ ℓ') + C xs = ∀ ys → ys ≼ xs → B ys + + g : ∀ xs → C xs + g = ElimProp.f (isPropΠ2 (λ _ _ → BisProp)) c₀ θ + where + c₀ : C [] + c₀ ys ys≼[] = subst B (sym (≼[]→≡[] ys ys≼[])) b₀ + + θ : ∀ x {xs} → C xs → C (x ∷ xs) + θ x {xs} hyp ys ys≼x∷xs with ≼-Dichotomy x xs ys ys≼x∷xs + ... | inl ys≼xs = hyp ys ys≼xs + ... | inr ys≡x∷zs = subst B (sym ys≡x∷zs) (B-≼-hyp x zs χ) + where + zs = remove1 discA x ys + χ : ∀ vs → vs ≼ zs → B vs + χ vs vs≼zs = hyp vs (≼-trans vs zs xs vs≼zs (≼-remove1-lemma x xs ys ys≼x∷xs)) + + f : ∀ xs → B xs + f = C→B g + where + C→B : (∀ xs → C xs) → (∀ xs → B xs) + C→B C-hyp xs = C-hyp xs xs (≼-refl xs) + + + ≼-ElimPropBin : ∀ {ℓ'} {B : FMSet A → FMSet A → Type ℓ'} + → (∀ {xs} {ys} → isProp (B xs ys)) + → (B [] []) + → (∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B (x ∷ xs) ys) + → (∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B xs (x ∷ ys)) + ------------------------------------------------------------------------------- + → (∀ xs ys → B xs ys) + ≼-ElimPropBin {B = B} BisProp b₀₀ left-hyp right-hyp = ≼-ElimProp.f (isPropΠ (λ _ → BisProp)) θ χ + where + θ : ∀ ys → B [] ys + θ = ≼-ElimProp.f BisProp b₀₀ h₁ + where + h₁ : ∀ x ys → (∀ ws → ws ≼ ys → B [] ws) → B [] (x ∷ ys) + h₁ x ys mini-h = right-hyp x [] ys h₂ + where + h₂ : ∀ vs ws → vs ≼ [] → ws ≼ ys → B vs ws + h₂ vs ws vs≼[] ws≼ys = subst (λ zs → B zs ws) (sym (≼[]→≡[] vs vs≼[])) (mini-h ws ws≼ys) + + χ : ∀ x xs → (∀ zs → zs ≼ xs → (∀ ys → B zs ys)) → ∀ ys → B (x ∷ xs) ys + χ x xs h ys = left-hyp x xs ys λ vs ws vs≼xs _ → h vs vs≼xs ws + + + + ≼-ElimPropBinSym : ∀ {ℓ'} {B : FMSet A → FMSet A → Type ℓ'} + → (∀ {xs} {ys} → isProp (B xs ys)) + → (∀ {xs} {ys} → B xs ys → B ys xs) + → (B [] []) + → (∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B (x ∷ xs) ys) + ---------------------------------------------------------------------------- + → (∀ xs ys → B xs ys) + ≼-ElimPropBinSym {B = B} BisProp BisSym b₀₀ left-hyp = ≼-ElimPropBin BisProp b₀₀ left-hyp right-hyp + where + right-hyp : ∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B xs (x ∷ ys) + right-hyp x xs ys h₁ = BisSym (left-hyp x ys xs (λ vs ws vs≼ys ws≼xs → BisSym (h₁ ws vs ws≼xs vs≼ys))) + + + -- The main result + module FMScountExt where + B : FMSet A → FMSet A → Type ℓ + B xs ys = (∀ a → FMScount discA a xs ≡ FMScount discA a ys) → xs ≡ ys + + BisProp : ∀ {xs} {ys} → isProp (B xs ys) + BisProp = isPropΠ λ _ → trunc _ _ + + BisSym : ∀ {xs} {ys} → B xs ys → B ys xs + BisSym {xs} {ys} b h = sym (b (λ a → sym (h a))) + + b₀₀ : B [] [] + b₀₀ _ = refl + + left-hyp : ∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B (x ∷ xs) ys + left-hyp x xs ys hyp h₁ = (λ i → x ∷ (hyp-path i)) ∙ sym path + where + eq₁ : FMScount discA x ys ≡ suc (FMScount discA x xs) + eq₁ = sym (h₁ x) ∙ FMScount-≡-lemma-refl discA xs + + path : ys ≡ x ∷ (remove1 discA x ys) + path = remove1-suc-lemma discA x (FMScount discA x xs) ys eq₁ + + hyp-path : xs ≡ remove1 discA x ys + hyp-path = hyp xs (remove1 discA x ys) (≼-refl xs) (≼-remove1 x ys) θ + where + θ : ∀ a → FMScount discA a xs ≡ FMScount discA a (remove1 discA x ys) + θ a with discA a x + ... | yes a≡x = subst (λ b → FMScount discA b xs ≡ FMScount discA b (remove1 discA x ys)) (sym a≡x) eq₂ + where + eq₂ : FMScount discA x xs ≡ FMScount discA x (remove1 discA x ys) + eq₂ = FMScount discA x xs ≡⟨ cong predℕ (sym (FMScount-≡-lemma-refl discA xs)) ⟩ + predℕ (FMScount discA x (x ∷ xs)) ≡⟨ cong predℕ (h₁ x) ⟩ + predℕ (FMScount discA x ys) ≡⟨ sym (remove1-predℕ-lemma discA x ys) ⟩ + FMScount discA x (remove1 discA x ys) ∎ + ... | no a≢x = + FMScount discA a xs ≡⟨ sym (FMScount-≢-lemma discA xs a≢x) ⟩ + FMScount discA a (x ∷ xs) ≡⟨ h₁ a ⟩ + FMScount discA a ys ≡⟨ cong (FMScount discA a) path ⟩ + FMScount discA a (x ∷ (remove1 discA x ys)) ≡⟨ FMScount-≢-lemma discA (remove1 discA x ys) a≢x ⟩ + FMScount discA a (remove1 discA x ys) ∎ + + + Thm : ∀ xs ys → (∀ a → FMScount discA a xs ≡ FMScount discA a ys) → xs ≡ ys + Thm = ≼-ElimPropBinSym BisProp BisSym b₀₀ left-hyp diff --git a/Cubical/HITs/FiniteMultiset/Properties.agda b/Cubical/HITs/FiniteMultiset/Properties.agda new file mode 100644 index 0000000000..1a4e240285 --- /dev/null +++ b/Cubical/HITs/FiniteMultiset/Properties.agda @@ -0,0 +1,219 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.FiniteMultiset.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Nat +open import Cubical.Data.Empty as ⊥ +open import Cubical.Relation.Nullary +open import Cubical.HITs.FiniteMultiset.Base +open import Cubical.Structures.MultiSet +open import Cubical.Relation.Nullary.DecidableEq + +private + variable + ℓ : Level + A : Type ℓ + +infixr 30 _++_ + +_++_ : ∀ (xs ys : FMSet A) → FMSet A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ xs ++ ys +comm x y xs i ++ ys = comm x y (xs ++ ys) i +trunc xs zs p q i j ++ ys = + trunc (xs ++ ys) (zs ++ ys) (cong (_++ ys) p) (cong (_++ ys) q) i j + +unitl-++ : ∀ (xs : FMSet A) → [] ++ xs ≡ xs +unitl-++ xs = refl + +unitr-++ : ∀ (xs : FMSet A) → xs ++ [] ≡ xs +unitr-++ = ElimProp.f (trunc _ _) refl (λ x p → cong (_∷_ x) p) + +assoc-++ : ∀ (xs ys zs : FMSet A) → xs ++ (ys ++ zs) ≡ (xs ++ ys) ++ zs +assoc-++ = ElimProp.f (isPropΠ2 (λ _ _ → trunc _ _)) + (λ ys zs → refl) + (λ x p ys zs → cong (_∷_ x) (p ys zs)) + +cons-++ : ∀ (x : A) (xs : FMSet A) → x ∷ xs ≡ xs ++ [ x ] +cons-++ x = ElimProp.f (trunc _ _) + refl + (λ y {xs} p → comm x y xs ∙ cong (_∷_ y) p) + +comm-++ : ∀ (xs ys : FMSet A) → xs ++ ys ≡ ys ++ xs +comm-++ = ElimProp.f (isPropΠ (λ _ → trunc _ _)) + (λ ys → sym (unitr-++ ys)) + (λ x {xs} p ys → cong (x ∷_) (p ys) + ∙ cong (_++ xs) (cons-++ x ys) + ∙ sym (assoc-++ ys [ x ] xs)) + +module FMSetUniversal {ℓ} {M : Type ℓ} (MSet : isSet M) + (e : M) (_⊗_ : M → M → M) + (comm-⊗ : ∀ x y → x ⊗ y ≡ y ⊗ x) (assoc-⊗ : ∀ x y z → x ⊗ (y ⊗ z) ≡ (x ⊗ y) ⊗ z) + (unit-⊗ : ∀ x → e ⊗ x ≡ x) + (f : A → M) where + + f-extend : FMSet A → M + f-extend = Rec.f MSet e (λ x m → f x ⊗ m) + (λ x y m → comm-⊗ (f x) (f y ⊗ m) ∙ sym (assoc-⊗ (f y) m (f x)) ∙ cong (f y ⊗_) (comm-⊗ m (f x))) + + f-extend-nil : f-extend [] ≡ e + f-extend-nil = refl + + f-extend-cons : ∀ x xs → f-extend (x ∷ xs) ≡ f x ⊗ f-extend xs + f-extend-cons x xs = refl + + f-extend-sing : ∀ x → f-extend [ x ] ≡ f x + f-extend-sing x = comm-⊗ (f x) e ∙ unit-⊗ (f x) + + f-extend-++ : ∀ xs ys → f-extend (xs ++ ys) ≡ f-extend xs ⊗ f-extend ys + f-extend-++ = ElimProp.f (isPropΠ λ _ → MSet _ _) + (λ ys → sym (unit-⊗ (f-extend ys))) + (λ x {xs} p ys → cong (f x ⊗_) (p ys) ∙ assoc-⊗ (f x) (f-extend xs) (f-extend ys)) + + module _ (h : FMSet A → M) (h-nil : h [] ≡ e) (h-sing : ∀ x → h [ x ] ≡ f x) + (h-++ : ∀ xs ys → h (xs ++ ys) ≡ h xs ⊗ h ys) where + + f-extend-unique : h ≡ f-extend + f-extend-unique = funExt (ElimProp.f (MSet _ _) + h-nil + (λ x {xs} p → (h-++ [ x ] xs) ∙ cong (_⊗ h xs) (h-sing x) ∙ cong (f x ⊗_) p)) + + + +-- We want to construct a multiset-structure on FMSet A, the empty set and insertion are given by the constructors, +-- for the count part we use the recursor + + +-- If A has decidable equality we can define the count-function: +module _(discA : Discrete A) where + FMScount-∷* : A → A → ℕ → ℕ + FMScount-∷* a x n with discA a x + ... | yes a≡x = suc n + ... | no a≢x = n + + FMScount-comm* : (a x y : A) (n : ℕ) + → FMScount-∷* a x (FMScount-∷* a y n) + ≡ FMScount-∷* a y (FMScount-∷* a x n) + FMScount-comm* a x y n with discA a x | discA a y + ... | yes a≡x | yes a≡y = refl + ... | yes a≡x | no a≢y = refl + ... | no a≢x | yes a≡y = refl + ... | no a≢x | no a≢y = refl + + FMScount : A → FMSet A → ℕ + FMScount a = Rec.f isSetℕ 0 (FMScount-∷* a) (FMScount-comm* a) + + + FMS-with-str : MultiSet A (Discrete→isSet discA) + FMS-with-str = (FMSet A , [] , _∷_ , FMScount) + + + + -- We prove some useful properties of the FMScount function + FMScount-≡-lemma : ∀ {a} {x} xs → a ≡ x → FMScount a (x ∷ xs) ≡ suc (FMScount a xs) + FMScount-≡-lemma {a} {x} xs a≡x with discA a x + ... | yes _ = refl + ... | no a≢x = ⊥.rec (a≢x a≡x) + + + FMScount-≡-lemma-refl : ∀ {x} xs → FMScount x (x ∷ xs) ≡ suc (FMScount x xs) + FMScount-≡-lemma-refl {x} xs = FMScount-≡-lemma xs refl + + + FMScount-≢-lemma : ∀ {a} {x} xs → ¬ a ≡ x → FMScount a (x ∷ xs) ≡ FMScount a xs + FMScount-≢-lemma {a} {x} xs a≢x with discA a x + ... | yes a≡x = ⊥.rec (a≢x a≡x) + ... | no _ = refl + + + FMScount-0-lemma : ∀ xs → (∀ a → FMScount a xs ≡ 0) → xs ≡ [] + FMScount-0-lemma = ElimProp.f (isPropΠ λ _ → trunc _ _) (λ _ → refl) θ + where + θ : ∀ x {xs} → ((∀ a → FMScount a xs ≡ 0) → xs ≡ []) + → ((∀ a → FMScount a (x ∷ xs) ≡ 0) → (x ∷ xs) ≡ []) + θ x {xs} _ p = ⊥.rec (snotz (sym (FMScount-≡-lemma-refl xs) ∙ p x)) + + + -- we define a function that removes an element a from a finite multiset once + -- by simultaneously defining two lemmas about it + remove1 : A → FMSet A → FMSet A + remove1 a [] = [] + remove1 a (x ∷ xs) with (discA a x) + ... | yes _ = xs + ... | no _ = x ∷ remove1 a xs + remove1 a (comm x y xs i) with discA a x with discA a y + ... | yes a≡x | yes a≡y = ((sym a≡y ∙ a≡x) i) ∷ xs + ... | yes a≡x | no _ = y ∷ (eq i) + where + eq : xs ≡ remove1 a (x ∷ xs) + eq with discA a x + eq | yes _ = refl + eq | no a≢x = ⊥.rec (a≢x a≡x) + remove1 a (comm x y xs i) | no _ | yes a≡y = x ∷ (eq i) + where + eq : remove1 a (y ∷ xs) ≡ xs + eq with discA a y + eq | yes _ = refl + eq | no a≢y = ⊥.rec (a≢y a≡y) + remove1 a (comm x y xs i) | no a≢x | no a≢y = ((λ i → x ∷ (p i)) ∙∙ comm x y (remove1 a xs) ∙∙ (λ i → y ∷ (sym q i))) i + where + p : remove1 a (y ∷ xs) ≡ y ∷ (remove1 a xs) + p with discA a y + p | yes a≡y = ⊥.rec (a≢y a≡y) + p | no _ = refl + q : remove1 a (x ∷ xs) ≡ x ∷ (remove1 a xs) + q with discA a x + q | yes a≡x = ⊥.rec (a≢x a≡x) + q | no _ = refl + remove1 a (trunc xs ys p q i j) = trunc (remove1 a xs) (remove1 a ys) (cong (remove1 a) p) (cong (remove1 a) q) i j + + + remove1-≡-lemma : ∀ {a} {x} xs → a ≡ x → xs ≡ remove1 a (x ∷ xs) + remove1-≡-lemma {a} {x} xs a≡x with discA a x + ... | yes _ = refl + ... | no a≢x = ⊥.rec (a≢x a≡x) + + remove1-≢-lemma : ∀ {a} {x} xs → ¬ a ≡ x → remove1 a (x ∷ xs) ≡ x ∷ remove1 a xs + remove1-≢-lemma {a} {x} xs a≢x with discA a x + ... | yes a≡x = ⊥.rec (a≢x a≡x) + ... | no _ = refl + + + remove1-predℕ-lemma : ∀ a xs → FMScount a (remove1 a xs) ≡ predℕ (FMScount a xs) + remove1-predℕ-lemma a = ElimProp.f (isSetℕ _ _) refl θ + where + θ : ∀ x {xs} → FMScount a (remove1 a xs) ≡ predℕ (FMScount a xs) + → FMScount a (remove1 a (x ∷ xs)) ≡ predℕ (FMScount a (x ∷ xs)) + θ x {xs} p with discA a x + ... | yes _ = refl + ... | no a≢x = FMScount-≢-lemma (remove1 a xs) a≢x ∙ p + + + remove1-zero-lemma : ∀ a xs → FMScount a xs ≡ zero → xs ≡ remove1 a xs + remove1-zero-lemma a = ElimProp.f (isPropΠ λ _ → trunc _ _) (λ _ → refl) θ + where + θ : ∀ x {xs} → (FMScount a xs ≡ zero → xs ≡ remove1 a xs) + → FMScount a (x ∷ xs) ≡ zero → x ∷ xs ≡ remove1 a (x ∷ xs) + θ x {xs} hyp p with discA a x + ... | yes _ = ⊥.rec (snotz p) + ... | no _ = cong (x ∷_) (hyp p) + + + remove1-suc-lemma : ∀ a n xs → FMScount a xs ≡ suc n → xs ≡ a ∷ (remove1 a xs) + remove1-suc-lemma a n = ElimProp.f (isPropΠ λ _ → trunc _ _) (λ p → ⊥.rec (znots p)) θ + where + θ : ∀ x {xs} → (FMScount a xs ≡ suc n → xs ≡ a ∷ (remove1 a xs)) + → FMScount a (x ∷ xs) ≡ suc n → x ∷ xs ≡ a ∷ (remove1 a (x ∷ xs)) + θ x {xs} hyp p with discA a x + ... | yes a≡x = (λ i → (sym a≡x i) ∷ xs) + ... | no a≢x = cong (x ∷_) (hyp p) ∙ comm x a (remove1 a xs) + + + FMScount-remove1-≢-lemma : ∀ {a} {x} xs → ¬ a ≡ x → FMScount a (remove1 x xs) ≡ FMScount a xs + FMScount-remove1-≢-lemma {a} {x} xs a≢x with discreteℕ (FMScount x xs) zero + ... | yes p = cong (FMScount a) (sym (remove1-zero-lemma x xs p)) + ... | no ¬p = sym (FMScount-≢-lemma (remove1 x xs) a≢x) ∙ cong (FMScount a) eq₁ + where + eq₁ : (x ∷ remove1 x xs) ≡ xs + eq₁ = sym (remove1-suc-lemma x (predℕ (FMScount x xs)) xs (suc-predℕ (FMScount x xs) ¬p)) diff --git a/Cubical/HITs/GroupoidQuotients.agda b/Cubical/HITs/GroupoidQuotients.agda new file mode 100644 index 0000000000..ef2758b565 --- /dev/null +++ b/Cubical/HITs/GroupoidQuotients.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.GroupoidQuotients where + +open import Cubical.HITs.GroupoidQuotients.Base public +open import Cubical.HITs.GroupoidQuotients.Properties public diff --git a/Cubical/HITs/GroupoidQuotients/Base.agda b/Cubical/HITs/GroupoidQuotients/Base.agda new file mode 100644 index 0000000000..0e141fea7a --- /dev/null +++ b/Cubical/HITs/GroupoidQuotients/Base.agda @@ -0,0 +1,43 @@ +{- + +This file contains: + +- Definition of groupoid quotients + +-} +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.GroupoidQuotients.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Relation.Binary.Base + +-- Groupoid quotients as a higher inductive type: +-- For the definition, only transitivity is needed +data _//_ {ℓ ℓ'} (A : Type ℓ) {R : A → A → Type ℓ'} + (Rt : BinaryRelation.isTrans R) : Type (ℓ-max ℓ ℓ') where + [_] : (a : A) → A // Rt + eq// : {a b : A} → (r : R a b) → [ a ] ≡ [ b ] + comp// : {a b c : A} → (r : R a b) → (s : R b c) + → PathP (λ j → [ a ] ≡ eq// s j) (eq// r) (eq// (Rt a b c r s)) + squash// : ∀ (x y : A // Rt) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s + +{- The comp// constructor fills the square: + + eq// (Rt a b c r s) + [a]— — — >[c] + ‖ ^ + ‖ | eq// s ^ + ‖ | j | + [a]— — — >[b] ∙ — > + eq// r i + + We use this to give another constructor-like construction: +-} + +comp'// : {ℓ ℓ' : Level} (A : Type ℓ) {R : A → A → Type ℓ'} + (Rt : BinaryRelation.isTrans R) + {a b c : A} → (r : R a b) → (s : R b c) + → eq// {Rt = Rt} (Rt a b c r s) ≡ eq// r ∙ eq// s +comp'// A Rt r s i = compPath-unique refl (eq// r) (eq// s) + (eq// {Rt = Rt} (Rt _ _ _ r s) , comp// r s) + (eq// r ∙ eq// s , compPath-filler (eq// r) (eq// s)) i .fst diff --git a/Cubical/HITs/GroupoidQuotients/Properties.agda b/Cubical/HITs/GroupoidQuotients/Properties.agda new file mode 100644 index 0000000000..aff8f33131 --- /dev/null +++ b/Cubical/HITs/GroupoidQuotients/Properties.agda @@ -0,0 +1,106 @@ +{- + +Groupoid quotients: + +-} + +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.GroupoidQuotients.Properties where + +open import Cubical.HITs.GroupoidQuotients.Base + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels + +open import Cubical.Functions.Surjection + +open import Cubical.Data.Sigma + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary.Base + +open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥; ∣_∣; squash) +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) + +-- Type quotients + +private + variable + ℓA ℓR ℓ : Level + A : Type ℓA + R : A → A → Type ℓR + +elimSet : (Rt : BinaryRelation.isTrans R) + → {B : A // Rt → Type ℓ} + → ((x : A // Rt) → isSet (B x)) + → (f : (a : A) → B [ a ]) + → ({a b : A} (r : R a b) → PathP (λ i → B (eq// r i)) (f a) (f b)) + → (x : A // Rt) + → B x +elimSet Rt Bset f feq [ a ] = f a +elimSet Rt Bset f feq (eq// r i) = feq r i +elimSet Rt Bset f feq (comp// {a} {b} {c} r s i j) = + isSet→SquareP (λ i j → Bset (comp// r s i j)) + (λ j → feq r j) (λ j → feq (Rt a b c r s) j) + (λ i → f a) (λ i → feq s i) i j +elimSet Rt Bset f feq (squash// x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 (λ x → isSet→isGroupoid (Bset x)) + _ _ _ _ (λ j k → g (r j k)) (λ j k → g (s j k)) (squash// x y p q r s) i j k + where + g = elimSet Rt Bset f feq + +elimProp : (Rt : BinaryRelation.isTrans R) + → {B : A // Rt → Type ℓ} + → ((x : A // Rt) → isProp (B x)) + → ((a : A) → B [ a ]) + → (x : A // Rt) + → B x +elimProp Rt Brop f x = + elimSet Rt (λ x → isProp→isSet (Brop x)) f (λ r → isProp→PathP (λ i → Brop (eq// r i)) (f _) (f _)) x + +elimProp2 : (Rt : BinaryRelation.isTrans R) + → {C : A // Rt → A // Rt → Type ℓ} + → ((x y : A // Rt) → isProp (C x y)) + → ((a b : A) → C [ a ] [ b ]) + → (x y : A // Rt) + → C x y +elimProp2 Rt Cprop f = elimProp Rt (λ x → isPropΠ (λ y → Cprop x y)) + (λ x → elimProp Rt (λ y → Cprop [ x ] y) (f x)) + +isSurjective[] : (Rt : BinaryRelation.isTrans R) + → isSurjection (λ a → [ a ]) +isSurjective[] Rt = elimProp Rt (λ x → squash) (λ a → ∣ a , refl ∣) + +elim : (Rt : BinaryRelation.isTrans R) + → {B : A // Rt → Type ℓ} + → ((x : A // Rt) → isGroupoid (B x)) + → (f : (a : A) → B [ a ]) + → (feq : {a b : A} (r : R a b) → PathP (λ i → B (eq// r i)) (f a) (f b)) + → ({a b c : A} (r : R a b) (s : R b c) + → SquareP (λ i j → B (comp// r s i j)) + (feq r) (feq (Rt a b c r s)) (λ j → f a) (feq s)) + → (x : A // Rt) + → B x +elim Rt Bgpd f feq fcomp [ a ] = f a +elim Rt Bgpd f feq fcomp (eq// r i) = feq r i +elim Rt Bgpd f feq fcomp (comp// r s i j) = fcomp r s i j +elim Rt Bgpd f feq fcomp (squash// x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 Bgpd + _ _ _ _ (λ j k → g (r j k)) (λ j k → g (s j k)) (squash// x y p q r s) i j k + where + g = elim Rt Bgpd f feq fcomp + +rec : (Rt : BinaryRelation.isTrans R) + → {B : Type ℓ} + → isGroupoid B + → (f : A → B) + → (feq : {a b : A} (r : R a b) → f a ≡ f b) + → ({a b c : A} (r : R a b) (s : R b c) + → Square (feq r) (feq (Rt a b c r s)) refl (feq s)) + → (x : A // Rt) + → B +rec Rt Bgpd = elim Rt (λ _ → Bgpd) diff --git a/Cubical/HITs/GroupoidTruncation.agda b/Cubical/HITs/GroupoidTruncation.agda new file mode 100644 index 0000000000..62e8c4db05 --- /dev/null +++ b/Cubical/HITs/GroupoidTruncation.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.GroupoidTruncation where + +open import Cubical.HITs.GroupoidTruncation.Base public +open import Cubical.HITs.GroupoidTruncation.Properties public diff --git a/Cubical/HITs/GroupoidTruncation/Base.agda b/Cubical/HITs/GroupoidTruncation/Base.agda new file mode 100644 index 0000000000..80ec0f201c --- /dev/null +++ b/Cubical/HITs/GroupoidTruncation/Base.agda @@ -0,0 +1,17 @@ +{- + +This file contains: + +- Definition of groupoid truncations + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.GroupoidTruncation.Base where + +open import Cubical.Core.Primitives + +-- groupoid truncation as a higher inductive type: + +data ∥_∥₃ {ℓ} (A : Type ℓ) : Type ℓ where + ∣_∣₃ : A → ∥ A ∥₃ + squash₃ : ∀ (x y : ∥ A ∥₃) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s diff --git a/Cubical/HITs/GroupoidTruncation/Properties.agda b/Cubical/HITs/GroupoidTruncation/Properties.agda new file mode 100644 index 0000000000..c46f41a385 --- /dev/null +++ b/Cubical/HITs/GroupoidTruncation/Properties.agda @@ -0,0 +1,67 @@ +{- + +This file contains: + +- Properties of groupoid truncations + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.GroupoidTruncation.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.HITs.GroupoidTruncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +rec : ∀ {B : Type ℓ} → isGroupoid B → (A → B) → ∥ A ∥₃ → B +rec gB f ∣ x ∣₃ = f x +rec gB f (squash₃ x y p q r s i j k) = + gB _ _ _ _ (λ m n → rec gB f (r m n)) (λ m n → rec gB f (s m n)) i j k + +elim : {B : ∥ A ∥₃ → Type ℓ} + (bG : (x : ∥ A ∥₃) → isGroupoid (B x)) + (f : (x : A) → B ∣ x ∣₃) (x : ∥ A ∥₃) → B x +elim bG f (∣ x ∣₃) = f x +elim bG f (squash₃ x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 bG _ _ _ _ + (λ j k → elim bG f (r j k)) (λ j k → elim bG f (s j k)) + (squash₃ x y p q r s) + i j k + +elim2 : {B : ∥ A ∥₃ → ∥ A ∥₃ → Type ℓ} + (gB : ((x y : ∥ A ∥₃) → isGroupoid (B x y))) + (g : (a b : A) → B ∣ a ∣₃ ∣ b ∣₃) + (x y : ∥ A ∥₃) → B x y +elim2 gB g = elim (λ _ → isGroupoidΠ (λ _ → gB _ _)) + (λ a → elim (λ _ → gB _ _) (g a)) + +elim3 : {B : (x y z : ∥ A ∥₃) → Type ℓ} + (gB : ((x y z : ∥ A ∥₃) → isGroupoid (B x y z))) + (g : (a b c : A) → B ∣ a ∣₃ ∣ b ∣₃ ∣ c ∣₃) + (x y z : ∥ A ∥₃) → B x y z +elim3 gB g = elim2 (λ _ _ → isGroupoidΠ (λ _ → gB _ _ _)) + (λ a b → elim (λ _ → gB _ _ _) (g a b)) + +isGroupoidGroupoidTrunc : isGroupoid ∥ A ∥₃ +isGroupoidGroupoidTrunc a b p q r s = squash₃ a b p q r s + +groupoidTruncIdempotent≃ : isGroupoid A → ∥ A ∥₃ ≃ A +groupoidTruncIdempotent≃ {A = A} hA = isoToEquiv f + where + f : Iso ∥ A ∥₃ A + Iso.fun f = rec hA (idfun A) + Iso.inv f x = ∣ x ∣₃ + Iso.rightInv f _ = refl + Iso.leftInv f = elim (λ _ → isGroupoid→is2Groupoid isGroupoidGroupoidTrunc _ _) (λ _ → refl) + +groupoidTruncIdempotent : isGroupoid A → ∥ A ∥₃ ≡ A +groupoidTruncIdempotent hA = ua (groupoidTruncIdempotent≃ hA) diff --git a/Cubical/HITs/InfNat.agda b/Cubical/HITs/InfNat.agda new file mode 100644 index 0000000000..d45d38cb40 --- /dev/null +++ b/Cubical/HITs/InfNat.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --no-exact-split --safe #-} + +module Cubical.HITs.InfNat where + +open import Cubical.HITs.InfNat.Base public +open import Cubical.HITs.InfNat.Properties public diff --git a/Cubical/HITs/InfNat/Base.agda b/Cubical/HITs/InfNat/Base.agda new file mode 100644 index 0000000000..992ab56fcf --- /dev/null +++ b/Cubical/HITs/InfNat/Base.agda @@ -0,0 +1,16 @@ +{-# OPTIONS --no-exact-split --safe #-} + +module Cubical.HITs.InfNat.Base where + +open import Cubical.Core.Everything +open import Cubical.Data.Maybe +open import Cubical.Data.Nat + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism + +data ℕ+∞ : Type₀ where + zero : ℕ+∞ + suc : ℕ+∞ → ℕ+∞ + ∞ : ℕ+∞ + suc-inf : ∞ ≡ suc ∞ diff --git a/Cubical/HITs/InfNat/Properties.agda b/Cubical/HITs/InfNat/Properties.agda new file mode 100644 index 0000000000..f974ba7a9e --- /dev/null +++ b/Cubical/HITs/InfNat/Properties.agda @@ -0,0 +1,61 @@ +{-# OPTIONS --no-exact-split --safe #-} + +module Cubical.HITs.InfNat.Properties where + +open import Cubical.Core.Everything +open import Cubical.Data.Maybe +open import Cubical.Data.Nat + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism + +open import Cubical.HITs.InfNat.Base +import Cubical.Data.InfNat as Coprod + +ℕ+∞→Cℕ+∞ : ℕ+∞ → Coprod.ℕ+∞ +ℕ+∞→Cℕ+∞ zero = Coprod.zero +ℕ+∞→Cℕ+∞ (suc n) = Coprod.suc (ℕ+∞→Cℕ+∞ n) +ℕ+∞→Cℕ+∞ ∞ = Coprod.∞ +ℕ+∞→Cℕ+∞ (suc-inf i) = Coprod.∞ + +ℕ→ℕ+∞ : ℕ → ℕ+∞ +ℕ→ℕ+∞ zero = zero +ℕ→ℕ+∞ (suc n) = suc (ℕ→ℕ+∞ n) + +Cℕ+∞→ℕ+∞ : Coprod.ℕ+∞ → ℕ+∞ +Cℕ+∞→ℕ+∞ Coprod.∞ = ∞ +Cℕ+∞→ℕ+∞ (Coprod.fin n) = ℕ→ℕ+∞ n + +ℕ→ℕ+∞→Cℕ+∞ : ∀ n → ℕ+∞→Cℕ+∞ (ℕ→ℕ+∞ n) ≡ Coprod.fin n +ℕ→ℕ+∞→Cℕ+∞ zero = refl +ℕ→ℕ+∞→Cℕ+∞ (suc n) = cong Coprod.suc (ℕ→ℕ+∞→Cℕ+∞ n) + +Cℕ+∞→ℕ+∞→Cℕ+∞ : ∀ n → ℕ+∞→Cℕ+∞ (Cℕ+∞→ℕ+∞ n) ≡ n +Cℕ+∞→ℕ+∞→Cℕ+∞ Coprod.∞ = refl +Cℕ+∞→ℕ+∞→Cℕ+∞ (Coprod.fin n) = ℕ→ℕ+∞→Cℕ+∞ n + +suc-hom : ∀ n → Cℕ+∞→ℕ+∞ (Coprod.suc n) ≡ suc (Cℕ+∞→ℕ+∞ n) +suc-hom Coprod.∞ = suc-inf +suc-hom (Coprod.fin x) = refl + +ℕ+∞→Cℕ+∞→ℕ+∞ : ∀ n → Cℕ+∞→ℕ+∞ (ℕ+∞→Cℕ+∞ n) ≡ n +ℕ+∞→Cℕ+∞→ℕ+∞ zero = refl +ℕ+∞→Cℕ+∞→ℕ+∞ ∞ = refl +ℕ+∞→Cℕ+∞→ℕ+∞ (suc n) = suc-hom (ℕ+∞→Cℕ+∞ n) ∙ cong suc (ℕ+∞→Cℕ+∞→ℕ+∞ n) +ℕ+∞→Cℕ+∞→ℕ+∞ (suc-inf i) = lemma i + where + lemma : (λ i → ∞ ≡ suc-inf i) [ refl ≡ suc-inf ∙ refl ] + lemma i j = hcomp (λ k → λ + { (i = i0) → ∞ + ; (i = i1) → compPath-filler suc-inf refl k j + ; (j = i0) → ∞ + ; (j = i1) → suc-inf i + }) (suc-inf (i ∧ j)) + +open Iso + +ℕ+∞⇔Cℕ+∞ : Iso ℕ+∞ Coprod.ℕ+∞ +ℕ+∞⇔Cℕ+∞ .fun = ℕ+∞→Cℕ+∞ +ℕ+∞⇔Cℕ+∞ .inv = Cℕ+∞→ℕ+∞ +ℕ+∞⇔Cℕ+∞ .leftInv = ℕ+∞→Cℕ+∞→ℕ+∞ +ℕ+∞⇔Cℕ+∞ .rightInv = Cℕ+∞→ℕ+∞→Cℕ+∞ diff --git a/Cubical/HITs/Interval.agda b/Cubical/HITs/Interval.agda new file mode 100644 index 0000000000..ca584ce135 --- /dev/null +++ b/Cubical/HITs/Interval.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Interval where + +open import Cubical.HITs.Interval.Base public + +-- open import Cubical.HITs.Interval.Properties public diff --git a/Cubical/HITs/Interval/Base.agda b/Cubical/HITs/Interval/Base.agda new file mode 100644 index 0000000000..54865c63b3 --- /dev/null +++ b/Cubical/HITs/Interval/Base.agda @@ -0,0 +1,39 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Interval.Base where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude + +data Interval : Type₀ where + zero : Interval + one : Interval + seg : zero ≡ one + +isContrInterval : isContr Interval +isContrInterval = (zero , (λ x → rem x)) + where + rem : (x : Interval) → zero ≡ x + rem zero = refl + rem one = seg + rem (seg i) j = seg (i ∧ j) + +funExtInterval : ∀ {ℓ} (A B : Type ℓ) (f g : A → B) → ((x : A) → f x ≡ g x) → f ≡ g +funExtInterval A B f g p = λ i x → hmtpy x (seg i) + where + hmtpy : A → Interval → B + hmtpy x zero = f x + hmtpy x one = g x + hmtpy x (seg i) = p x i + +elim : (A : Interval → Type₀) (x : A zero) (y : A one) + (p : PathP (λ i → A (seg i)) x y) → (x : Interval) → A x +elim A x y p zero = x +elim A x y p one = y +elim A x y p (seg i) = p i + +-- Note that this is not definitional (it is not proved by refl) +intervalEta : ∀ {A : Type₀} (f : Interval → A) → elim _ (f zero) (f one) (λ i → f (seg i)) ≡ f +intervalEta f i zero = f zero +intervalEta f i one = f one +intervalEta f i (seg j) = f (seg j) diff --git a/Cubical/HITs/Join.agda b/Cubical/HITs/Join.agda new file mode 100644 index 0000000000..9fce340098 --- /dev/null +++ b/Cubical/HITs/Join.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Join where + +open import Cubical.HITs.Join.Base public +open import Cubical.HITs.Join.Properties public diff --git a/Cubical/HITs/Join/Base.agda b/Cubical/HITs/Join/Base.agda new file mode 100644 index 0000000000..d94f9551a6 --- /dev/null +++ b/Cubical/HITs/Join/Base.agda @@ -0,0 +1,87 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Join.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.HITs.S1 +open import Cubical.HITs.S3 + +-- redtt version : https://github.com/RedPRL/redtt/blob/master/library/cool/s3-to-join.red + +data join {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') where + inl : A → join A B + inr : B → join A B + push : ∀ a b → inl a ≡ inr b + +facek01 : I → I → I → join S¹ S¹ +facek01 i j k = hfill (λ l → λ { (j = i0) → push base base (~ l ∧ ~ k) + ; (j = i1) → push base base (~ l ∧ ~ k) + ; (k = i0) → push (loop j) base (~ l) + ; (k = i1) → inl base }) + (inS (push base base (~ k))) i + +border-contraction : I → I → I → I → join S¹ S¹ +border-contraction i j k m = + hfill (λ l → λ { (i = i0) → facek01 i1 j l + ; (i = i1) → push base (loop k) (~ l) + ; (j = i0) → push base (loop k) (i ∧ ~ l) + ; (j = i1) → push base (loop k) (i ∧ ~ l) + ; (k = i0) → facek01 (~ i) j l + ; (k = i1) → facek01 (~ i) j l }) + (inS (push (loop j) (loop k) i)) m + +S³→joinS¹S¹ : S³ → join S¹ S¹ +S³→joinS¹S¹ base = inl base +S³→joinS¹S¹ (surf j k i) = border-contraction i j k i1 + +joinS¹S¹→S³ : join S¹ S¹ → S³ +joinS¹S¹→S³ (inl x) = base +joinS¹S¹→S³ (inr x) = base +joinS¹S¹→S³ (push base b i) = base +joinS¹S¹→S³ (push (loop x) base i) = base +joinS¹S¹→S³ (push (loop i) (loop j) k) = surf i j k + +connection : I → I → I → I → S³ +connection i j k l = + hfill (λ m → λ { (k = i0) → joinS¹S¹→S³ (facek01 m i j) + ; (k = i1) → base + ; (j = i0) → base + ; (j = i1) → base + ; (i = i0) → base + ; (i = i1) → base }) + (inS base) l + +S³→joinS¹S¹→S³ : ∀ x → joinS¹S¹→S³ (S³→joinS¹S¹ x) ≡ x +S³→joinS¹S¹→S³ base l = base +S³→joinS¹S¹→S³ (surf j k i) l = + hcomp (λ m → λ { (l = i0) → joinS¹S¹→S³ (border-contraction i j k m) + ; (l = i1) → surf j k i + ; (i = i0) → connection j m l i1 + ; (i = i1) → base + ; (j = i0) → base + ; (j = i1) → base + ; (k = i0) → connection j m l (~ i) + ; (k = i1) → connection j m l (~ i) }) + (surf j k i) + +joinS¹S¹→S³→joinS¹S¹ : ∀ x → S³→joinS¹S¹ (joinS¹S¹→S³ x) ≡ x +joinS¹S¹→S³→joinS¹S¹ (inl base) l = inl base +joinS¹S¹→S³→joinS¹S¹ (inl (loop i)) l = facek01 i1 i (~ l) +joinS¹S¹→S³→joinS¹S¹ (inr base) l = push base base l +joinS¹S¹→S³→joinS¹S¹ (inr (loop i)) l = push base (loop i) l +joinS¹S¹→S³→joinS¹S¹ (push base base i) l = push base base (i ∧ l) +joinS¹S¹→S³→joinS¹S¹ (push base (loop k) i) l = push base (loop k) (i ∧ l) +joinS¹S¹→S³→joinS¹S¹ (push (loop k) base i) l = facek01 (~ i) k (~ l) +joinS¹S¹→S³→joinS¹S¹ (push (loop j) (loop k) i) l = border-contraction i j k (~ l) + +S³IsojoinS¹S¹ : Iso S³ (join S¹ S¹) +Iso.fun S³IsojoinS¹S¹ = S³→joinS¹S¹ +Iso.inv S³IsojoinS¹S¹ = joinS¹S¹→S³ +Iso.rightInv S³IsojoinS¹S¹ = joinS¹S¹→S³→joinS¹S¹ +Iso.leftInv S³IsojoinS¹S¹ = S³→joinS¹S¹→S³ + + +S³≡joinS¹S¹ : S³ ≡ join S¹ S¹ +S³≡joinS¹S¹ = isoToPath S³IsojoinS¹S¹ diff --git a/Cubical/HITs/Join/Properties.agda b/Cubical/HITs/Join/Properties.agda new file mode 100644 index 0000000000..f032bce49f --- /dev/null +++ b/Cubical/HITs/Join/Properties.agda @@ -0,0 +1,476 @@ +{- + +This file contains: + +- Equivalence with the pushout definition + Written by: Loïc Pujet, September 2019 + +- Associativity of the join + Written by: Loïc Pujet, September 2019 + +-} + +{-# OPTIONS --safe #-} + +module Cubical.HITs.Join.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function + +open import Cubical.Data.Sigma renaming (fst to proj₁; snd to proj₂) + +open import Cubical.HITs.Join.Base +open import Cubical.HITs.Pushout + +private + variable + ℓ ℓ' : Level + +open Iso + +-- Characterisation of function type join A B → C +IsoFunSpaceJoin : ∀ {ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → Iso (join A B → C) + (Σ[ f ∈ (A → C) ] Σ[ g ∈ (B → C) ] + ((a : A) (b : B) → f a ≡ g b)) +fun IsoFunSpaceJoin f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) +inv IsoFunSpaceJoin (f , g , p) (inl x) = f x +inv IsoFunSpaceJoin (f , g , p) (inr x) = g x +inv IsoFunSpaceJoin (f , g , p) (push a b i) = p a b i +rightInv IsoFunSpaceJoin (f , g , p) = refl +leftInv IsoFunSpaceJoin f = + funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} + +-- Alternative definition of the join using a pushout +joinPushout : (A : Type ℓ) → (B : Type ℓ') → Type (ℓ-max ℓ ℓ') +joinPushout A B = Pushout {A = A × B} proj₁ proj₂ + +-- Proof that it is equal +joinPushout-iso-join : (A : Type ℓ) → (B : Type ℓ') → Iso (joinPushout A B) (join A B) +joinPushout-iso-join A B = iso joinPushout→join join→joinPushout join→joinPushout→join joinPushout→join→joinPushout + where + joinPushout→join : joinPushout A B → join A B + joinPushout→join (inl x) = inl x + joinPushout→join (inr x) = inr x + joinPushout→join (push x i) = push (proj₁ x) (proj₂ x) i + + join→joinPushout : join A B → joinPushout A B + join→joinPushout (inl x) = inl x + join→joinPushout (inr x) = inr x + join→joinPushout (push a b i) = push (a , b) i + + joinPushout→join→joinPushout : ∀ x → join→joinPushout (joinPushout→join x) ≡ x + joinPushout→join→joinPushout (inl x) = refl + joinPushout→join→joinPushout (inr x) = refl + joinPushout→join→joinPushout (push (a , b) j) = refl + + join→joinPushout→join : ∀ x → joinPushout→join (join→joinPushout x) ≡ x + join→joinPushout→join (inl x) = refl + join→joinPushout→join (inr x) = refl + join→joinPushout→join (push a b j) = refl + +-- We will need both the equivalence and path version +joinPushout≃join : (A : Type ℓ) → (B : Type ℓ') → joinPushout A B ≃ join A B +joinPushout≃join A B = isoToEquiv (joinPushout-iso-join A B) + +joinPushout≡join : (A : Type ℓ) → (B : Type ℓ') → joinPushout A B ≡ join A B +joinPushout≡join A B = isoToPath (joinPushout-iso-join A B) + + +{- + Proof of associativity of the join +-} +join-assoc : (A B C : Type₀) → join (join A B) C ≡ join A (join B C) +join-assoc A B C = (joinPushout≡join (join A B) C) ⁻¹ + ∙ (spanEquivToPushoutPath sp3≃sp4) ⁻¹ + ∙ (3x3-span.3x3-lemma span) ⁻¹ + ∙ (spanEquivToPushoutPath sp1≃sp2) + ∙ (joinPushout≡join A (join B C)) + where + -- the meat of the proof is handled by the 3x3 lemma applied to this diagram + span : 3x3-span + span = record { + A00 = A; A02 = A × B; A04 = B; + A20 = A × C; A22 = A × B × C; A24 = B × C; + A40 = A × C; A42 = A × C; A44 = C; + f10 = proj₁; f12 = proj₁₂; f14 = proj₁; + f30 = λ x → x; f32 = proj₁₃; f34 = proj₂; + f01 = proj₁; f21 = proj₁₃; f41 = λ x → x; + f03 = proj₂; f23 = proj₂; f43 = proj₂; + H11 = H11; H13 = H13; H31 = H31; H33 = H33 } + where + proj₁₃ : A × B × C → A × C + proj₁₃ (a , (b , c)) = a , c + + proj₁₂ : A × B × C → A × B + proj₁₂ (a , (b , c)) = a , b + + H11 : (x : A × B × C) → proj₁ (proj₁₂ x) ≡ proj₁ (proj₁₃ x) + H11 (a , (b , c)) = refl + + H13 : (x : A × B × C) → proj₂ (proj₁₂ x) ≡ proj₁ (proj₂ x) + H13 (a , (b , c)) = refl + + H31 : (x : A × B × C) → proj₁₃ x ≡ proj₁₃ x + H31 (a , (b , c)) = refl + + H33 : (x : A × B × C) → proj₂ (proj₁₃ x) ≡ proj₂ (proj₂ x) + H33 (a , (b , c)) = refl + + -- the first pushout span appearing in the 3x3 lemma + sp1 : 3-span + sp1 = record { + A0 = 3x3-span.A□0 span; + A2 = 3x3-span.A□2 span; + A4 = 3x3-span.A□4 span; + f1 = 3x3-span.f□1 span; + f3 = 3x3-span.f□3 span } + + -- the first span we are interested in + sp2 : 3-span + sp2 = record { + A0 = A ; + A2 = A × (join B C) ; + A4 = join B C ; + f1 = proj₁ ; + f3 = proj₂ } + + -- proof that they are in fact equivalent + sp1≃sp2 : 3-span-equiv sp1 sp2 + sp1≃sp2 = record { + e0 = A□0≃A; + e2 = A□2≃A×join; + e4 = joinPushout≃join B C; + H1 = H1; + H3 = H2 } + where + A×join : Type₀ + A×join = A × (join B C) + + A□2→A×join : 3x3-span.A□2 span → A×join + A□2→A×join (inl (a , b)) = a , inl b + A□2→A×join (inr (a , c)) = a , inr c + A□2→A×join (push (a , (b , c)) i) = a , push b c i + + A×join→A□2 : A×join → 3x3-span.A□2 span + A×join→A□2 (a , inl b) = inl (a , b) + A×join→A□2 (a , inr c) = inr (a , c) + A×join→A□2 (a , push b c i) = push (a , (b , c)) i + + A×join→A□2→A×join : ∀ x → A×join→A□2 (A□2→A×join x) ≡ x + A×join→A□2→A×join (inl (a , b)) = refl + A×join→A□2→A×join (inr (a , c)) = refl + A×join→A□2→A×join (push (a , (b , c)) i) = refl + + A□2→A×join→A□2 : ∀ x → A□2→A×join (A×join→A□2 x) ≡ x + A□2→A×join→A□2 (a , inl b) = refl + A□2→A×join→A□2 (a , inr c) = refl + A□2→A×join→A□2 (a , push b c i) = refl + + A□2≃A×join : 3x3-span.A□2 span ≃ A×join + A□2≃A×join = isoToEquiv (iso A□2→A×join A×join→A□2 A□2→A×join→A□2 A×join→A□2→A×join) + + A→A□0 : A → 3x3-span.A□0 span + A→A□0 b = inl b + + A□0→A : 3x3-span.A□0 span → A + A□0→A (inl b) = b + A□0→A (inr a) = proj₁ a + A□0→A (push a i) = proj₁ a + + A→A□0→A : ∀ x → A□0→A (A→A□0 x) ≡ x + A→A□0→A x = refl + + A□0→A→A□0 : ∀ x → A→A□0 (A□0→A x) ≡ x + A□0→A→A□0 (inl b) = refl + A□0→A→A□0 (inr a) j = push a j + A□0→A→A□0 (push a i) j = push a (j ∧ i) + + A□0≃A : 3x3-span.A□0 span ≃ A + A□0≃A = isoToEquiv (iso A□0→A A→A□0 A→A□0→A A□0→A→A□0) + + H1 : (x : 3x3-span.A□2 span) → proj₁ (A□2→A×join x) ≡ A□0→A (3x3-span.f□1 span x) + H1 (inl (a , b)) = refl + H1 (inr (a , c)) = refl + H1 (push (a , (b , c)) i) j = A□0→A (doubleCompPath-filler refl (λ i → push (a , c) i) refl j i) + + H2 : (x : 3x3-span.A□2 span) → proj₂ (A□2→A×join x) ≡ fst (joinPushout≃join _ _) (3x3-span.f□3 span x) + H2 (inl (a , b)) = refl + H2 (inr (a , c)) = refl + H2 (push (a , (b , c)) i) j = fst (joinPushout≃join _ _) (doubleCompPath-filler refl (λ i → push (b , c) i) refl j i) + + -- the second span appearing in 3x3 lemma + sp3 : 3-span + sp3 = record { + A0 = 3x3-span.A0□ span; + A2 = 3x3-span.A2□ span; + A4 = 3x3-span.A4□ span; + f1 = 3x3-span.f1□ span; + f3 = 3x3-span.f3□ span } + + -- the second span we are interested in + sp4 : 3-span + sp4 = record { + A0 = join A B ; + A2 = (join A B) × C ; + A4 = C ; + f1 = proj₁ ; + f3 = proj₂ } + + -- proof that they are in fact equivalent + sp3≃sp4 : 3-span-equiv sp3 sp4 + sp3≃sp4 = record { + e0 = joinPushout≃join A B; + e2 = A2□≃join×C; + e4 = A4□≃C; + H1 = H4; + H3 = H3 } + where + join×C : Type₀ + join×C = (join A B) × C + + A2□→join×C : 3x3-span.A2□ span → join×C + A2□→join×C (inl (a , c)) = (inl a) , c + A2□→join×C (inr (b , c)) = (inr b) , c + A2□→join×C (push (a , (b , c)) i) = push a b i , c + + join×C→A2□ : join×C → 3x3-span.A2□ span + join×C→A2□ (inl a , c) = inl (a , c) + join×C→A2□ (inr b , c) = inr (b , c) + join×C→A2□ (push a b i , c) = push (a , (b , c)) i + + join×C→A2□→join×C : ∀ x → join×C→A2□ (A2□→join×C x) ≡ x + join×C→A2□→join×C (inl (a , c)) = refl + join×C→A2□→join×C (inr (b , c)) = refl + join×C→A2□→join×C (push (a , (b , c)) j) = refl + + A2□→join×C→A2□ : ∀ x → A2□→join×C (join×C→A2□ x) ≡ x + A2□→join×C→A2□ (inl a , c) = refl + A2□→join×C→A2□ (inr b , c) = refl + A2□→join×C→A2□ (push a b i , c) = refl + + A2□≃join×C : 3x3-span.A2□ span ≃ join×C + A2□≃join×C = isoToEquiv (iso A2□→join×C join×C→A2□ A2□→join×C→A2□ join×C→A2□→join×C) + + C→A4□ : C → 3x3-span.A4□ span + C→A4□ b = inr b + + A4□→C : 3x3-span.A4□ span → C + A4□→C (inl x) = proj₂ x + A4□→C (inr c) = c + A4□→C (push x i) = proj₂ x + + C→A4□→C : ∀ x → A4□→C (C→A4□ x) ≡ x + C→A4□→C x = refl + + A4□→C→A4□ : ∀ x → C→A4□ (A4□→C x) ≡ x + A4□→C→A4□ (inl x) j = push x (~ j) + A4□→C→A4□ (inr c) = refl + A4□→C→A4□ (push x i) j = push x (~ j ∨ i) + + A4□≃C : 3x3-span.A4□ span ≃ C + A4□≃C = isoToEquiv (iso A4□→C C→A4□ C→A4□→C A4□→C→A4□) + + H3 : (x : 3x3-span.A2□ span) → proj₂ (A2□→join×C x) ≡ A4□→C (3x3-span.f3□ span x) + H3 (inl (a , c)) = refl + H3 (inr (b , c)) = refl + H3 (push (a , (b , c)) i) j = A4□→C (doubleCompPath-filler refl (λ i → push (a , c) i) refl j i) + + H4 : (x : 3x3-span.A2□ span) → proj₁ (A2□→join×C x) ≡ fst (joinPushout≃join _ _) (3x3-span.f1□ span x) + H4 (inl (a , c)) = refl + H4 (inr (b , c)) = refl + H4 (push (a , (b , c)) i) j = fst (joinPushout≃join _ _) (doubleCompPath-filler refl (λ i → push (a , b) i) refl j i) + +{- + Direct proof of an associativity-related property. Combined with + commutativity, this implies that the join is associative. +-} +joinSwitch : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → join (join A B) C ≃ join (join C B) A +joinSwitch = isoToEquiv (iso switch switch invol invol) + where + switch : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → join (join A B) C → join (join C B) A + switch (inl (inl a)) = inr a + switch (inl (inr b)) = inl (inr b) + switch (inl (push a b i)) = push (inr b) a (~ i) + switch (inr c) = inl (inl c) + switch (push (inl a) c j) = push (inl c) a (~ j) + switch (push (inr b) c j) = inl (push c b (~ j)) + switch (push (push a b i) c j) = + hcomp + (λ k → λ + { (i = i0) → push (inl c) a (~ j ∨ ~ k) + ; (i = i1) → inl (push c b (~ j)) + ; (j = i0) → push (inr b) a (~ i) + ; (j = i1) → push (inl c) a (~ i ∧ ~ k) + }) + (push (push c b (~ j)) a (~ i)) + + invol : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (u : join (join A B) C) → switch (switch u) ≡ u + invol (inl (inl a)) = refl + invol (inl (inr b)) = refl + invol (inl (push a b i)) = refl + invol (inr c) = refl + invol (push (inl a) c j) = refl + invol (push (inr b) c j) = refl + invol {A = A} {B} {C} (push (push a b i) c j) l = + comp + (λ _ → join (join A B) C) + (λ k → λ + { (i = i0) → push (inl a) c (j ∧ (k ∨ l)) + ; (i = i1) → push (inr b) c j + ; (j = i0) → inl (push a b i) + ; (j = i1) → push (inl a) c (i ∨ (k ∨ l)) + ; (l = i1) → push (push a b i) c j + }) + (hcomp + (λ k → λ + { (i = i0) → push (inl a) c (j ∧ (~ k ∨ l)) + ; (i = i1) → push (inr b) c j + ; (j = i0) → inl (push a b i) + ; (j = i1) → push (inl a) c (i ∨ (~ k ∨ l)) + ; (l = i1) → push (push a b i) c j + }) + (push (push a b i) c j)) + +{- + Direct proof of associativity. +-} +joinAssocDirect : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → join (join A B) C ≃ join A (join B C) +joinAssocDirect {A = A} {B} {C} = + isoToEquiv (iso forward back forwardBack backForward) + where + forward : join (join A B) C → join A (join B C) + forward (inl (inl a)) = inl a + forward (inl (inr b)) = inr (inl b) + forward (inl (push a b i)) = push a (inl b) i + forward (inr c) = inr (inr c) + forward (push (inl a) c j) = push a (inr c) j + forward (push (inr b) c j) = inr (push b c j) + forward (push (push a b i) c j) = + hcomp + (λ k → λ + { (i = i0) → push a (inr c) (j ∧ k) + ; (i = i1) → inr (push b c j) + ; (j = i0) → push a (inl b) i + ; (j = i1) → push a (inr c) (i ∨ k) + }) + (push a (push b c j) i) + + back : join A (join B C) → join (join A B) C + back (inl a) = inl (inl a) + back (inr (inl b)) = inl (inr b) + back (inr (inr c)) = inr c + back (inr (push b c j)) = push (inr b) c j + back (push a (inl b) i) = inl (push a b i) + back (push a (inr c) i) = push (inl a) c i + back (push a (push b c j) i) = + hcomp + (λ k → λ + { (i = i0) → push (inl a) c (j ∧ ~ k) + ; (i = i1) → push (inr b) c j + ; (j = i0) → inl (push a b i) + ; (j = i1) → push (inl a) c (i ∨ ~ k) + }) + (push (push a b i) c j) + + forwardBack : ∀ u → forward (back u) ≡ u + forwardBack (inl a) = refl + forwardBack (inr (inl b)) = refl + forwardBack (inr (inr c)) = refl + forwardBack (inr (push b c j)) = refl + forwardBack (push a (inl b) i) = refl + forwardBack (push a (inr c) i) = refl + forwardBack (push a (push b c j) i) l = + comp + (λ _ → join A (join B C)) + (λ k → λ + { (i = i0) → push a (inr c) (j ∧ (~ k ∧ ~ l)) + ; (i = i1) → inr (push b c j) + ; (j = i0) → push a (inl b) i + ; (j = i1) → push a (inr c) (i ∨ (~ k ∧ ~ l)) + ; (l = i1) → push a (push b c j) i + }) + (hcomp + (λ k → λ + { (i = i0) → push a (inr c) (j ∧ (k ∧ ~ l)) + ; (i = i1) → inr (push b c j) + ; (j = i0) → push a (inl b) i + ; (j = i1) → push a (inr c) (i ∨ (k ∧ ~ l)) + ; (l = i1) → push a (push b c j) i + }) + (push a (push b c j) i)) + + backForward : ∀ u → back (forward u) ≡ u + backForward (inl (inl a)) = refl + backForward (inl (inr b)) = refl + backForward (inl (push a b i)) = refl + backForward (inr c) = refl + backForward (push (inl a) c j) = refl + backForward (push (inr b) c j) = refl + backForward (push (push a b i) c j) l = + comp + (λ _ → join (join A B) C) + (λ k → λ + { (i = i0) → push (inl a) c (j ∧ (k ∨ l)) + ; (i = i1) → push (inr b) c j + ; (j = i0) → inl (push a b i) + ; (j = i1) → push (inl a) c (i ∨ (k ∨ l)) + ; (l = i1) → push (push a b i) c j + }) + (hcomp + (λ k → λ + { (i = i0) → push (inl a) c (j ∧ (~ k ∨ l)) + ; (i = i1) → push (inr b) c j + ; (j = i0) → inl (push a b i) + ; (j = i1) → push (inl a) c (i ∨ (~ k ∨ l)) + ; (l = i1) → push (push a b i) c j + }) + (push (push a b i) c j)) + +-- commutativity +join-commFun : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} → join A B → join B A +join-commFun (inl x) = inr x +join-commFun (inr x) = inl x +join-commFun (push a b i) = push b a (~ i) + +join-commFun² : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} (x : join A B) + → join-commFun (join-commFun x) ≡ x +join-commFun² (inl x) = refl +join-commFun² (inr x) = refl +join-commFun² (push a b i) = refl + +join-comm : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Iso (join A B) (join B A) +fun join-comm = join-commFun +inv join-comm = join-commFun +rightInv join-comm = join-commFun² +leftInv join-comm = join-commFun² + +join→ : ∀ {ℓ'' ℓ'''} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → (A → C) → (B → D) → join A B → join C D +join→ f g (inl x) = inl (f x) +join→ f g (inr x) = inr (g x) +join→ f g (push a b i) = push (f a) (g b) i + +-- Applying Isos to joins (more efficient than transports) +Iso→joinIso : ∀ {ℓ'' ℓ'''} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → Iso A C → Iso B D → Iso (join A B) (join C D) +fun (Iso→joinIso is1 is2) x = join→ (Iso.fun is1) (Iso.fun is2) x +inv (Iso→joinIso is1 is2) x = join→ (Iso.inv is1) (Iso.inv is2) x +rightInv (Iso→joinIso is1 is2) (inl x) i = inl (rightInv is1 x i) +rightInv (Iso→joinIso is1 is2) (inr x) i = inr (rightInv is2 x i) +rightInv (Iso→joinIso is1 is2) (push a b j) i = + push (rightInv is1 a i) (rightInv is2 b i) j +leftInv (Iso→joinIso is1 is2) (inl x) i = inl (leftInv is1 x i) +leftInv (Iso→joinIso is1 is2) (inr x) i = inr (leftInv is2 x i) +leftInv (Iso→joinIso is1 is2) (push a b i) j = + push (leftInv is1 a j) (leftInv is2 b j) i diff --git a/Cubical/HITs/KleinBottle.agda b/Cubical/HITs/KleinBottle.agda new file mode 100644 index 0000000000..bc1b0a0bd0 --- /dev/null +++ b/Cubical/HITs/KleinBottle.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.KleinBottle where + +open import Cubical.HITs.KleinBottle.Base public +open import Cubical.HITs.KleinBottle.Properties public diff --git a/Cubical/HITs/KleinBottle/Base.agda b/Cubical/HITs/KleinBottle/Base.agda new file mode 100644 index 0000000000..3e9645c0b2 --- /dev/null +++ b/Cubical/HITs/KleinBottle/Base.agda @@ -0,0 +1,15 @@ +{- + +Definition of the Klein bottle as a HIT + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.KleinBottle.Base where + +open import Cubical.Core.Everything + +data KleinBottle : Type where + point : KleinBottle + line1 : point ≡ point + line2 : point ≡ point + square : PathP (λ i → line1 (~ i) ≡ line1 i) line2 line2 diff --git a/Cubical/HITs/KleinBottle/Properties.agda b/Cubical/HITs/KleinBottle/Properties.agda new file mode 100644 index 0000000000..0bb163bfaa --- /dev/null +++ b/Cubical/HITs/KleinBottle/Properties.agda @@ -0,0 +1,159 @@ +{- + +Definition of the Klein bottle as a HIT + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.KleinBottle.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Data.Int +open import Cubical.Data.Sigma +open import Cubical.HITs.S1 +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +open import Cubical.HITs.KleinBottle.Base + +loop1 : S¹ → KleinBottle +loop1 base = point +loop1 (loop i) = line1 i + +invS¹Loop : S¹ → Type +invS¹Loop base = S¹ +invS¹Loop (loop i) = invS¹Path i + +loop1Inv : (s : S¹) → loop1 (invLooper s) ≡ loop1 s +loop1Inv base = line2 +loop1Inv (loop i) = square i + +twist : (s : S¹) → PathP (λ i → invS¹Path i) s (invLooper s) +twist s i = glue (λ {(i = i0) → s; (i = i1) → invLooper s}) (invLooper s) + +twistBaseLoop : (s : S¹) → invS¹Loop s +twistBaseLoop base = base +twistBaseLoop (loop i) = twist base i + +kleinBottle≃Σ : KleinBottle ≃ Σ S¹ invS¹Loop +kleinBottle≃Σ = isoToEquiv (iso fro to froTo toFro) + where + fro : KleinBottle → Σ S¹ invS¹Loop + fro point = (base , base) + fro (line1 i) = (base , loop i) + fro (line2 j) = (loop (~ j) , twist base (~ j)) + fro (square i j) = (loop (~ j) , twist (loop i) (~ j)) + + toLoopFiller : (j : I) → ua invS¹Equiv j → I → KleinBottle + toLoopFiller j g l = + hfill + (λ l → λ + { (j = i0) → loop1Inv g l + ; (j = i1) → loop1 g + }) + (inS (loop1 (unglue (j ∨ ~ j) g))) + l + + to : Σ S¹ invS¹Loop → KleinBottle + to (base , s) = loop1 s + to (loop j , g) = toLoopFiller j g i1 + + toFro : (a : KleinBottle) → to (fro a) ≡ a + toFro point = refl + toFro (line1 i) = refl + toFro (line2 j) k = lUnit line2 (~ k) j + toFro (square i j) k = lUnit (square i) (~ k) j + + froLoop1 : (s : S¹) → fro (loop1 s) ≡ (base , s) + froLoop1 base = refl + froLoop1 (loop i) = refl + + froLoop1Inv : + PathP (λ k → (s : S¹) → froLoop1 (invLooper s) k ≡ froLoop1 s k) + (λ s l → fro (loop1Inv s l)) + (λ s l → loop (~ l) , twist s (~ l)) + froLoop1Inv k base l = loop (~ l) , twist base (~ l) + froLoop1Inv k (loop i) l = loop (~ l) , twist (loop i) (~ l) + + froTo : (a : Σ S¹ invS¹Loop) → fro (to a) ≡ a + froTo (base , s) = froLoop1 s + froTo (loop j , g) k = + hcomp + (λ l → λ + { (j = i0) → froLoop1Inv k g l + ; (j = i1) → froLoop1 g k + ; (k = i0) → fro (toLoopFiller j g l) + ; (k = i1) → + ( loop (j ∨ ~ l) + , glue + (λ + { (j = i0) (l = i1) → g + ; (j = i1) → g + ; (l = i0) → unglue (j ∨ ~ j) g + }) + (unglue (j ∨ ~ j) g) + ) + }) + (froLoop1 (unglue (j ∨ ~ j) g) k) + +isGroupoidKleinBottle : isGroupoid KleinBottle +isGroupoidKleinBottle = + transport (λ i → isGroupoid (ua kleinBottle≃Σ (~ i))) + (isOfHLevelΣ 3 isGroupoidS¹ + (λ s → + PropTrunc.rec + (isPropIsOfHLevel 3) + (λ p → subst (λ s → isGroupoid (invS¹Loop s)) p isGroupoidS¹) + (isConnectedS¹ s))) + +-- Transport across the following is too slow :( + +ΩKlein≡ℤ² : Path KleinBottle point point ≡ ℤ × ℤ +ΩKlein≡ℤ² = + Path KleinBottle point point + ≡⟨ (λ i → basePath i ≡ basePath i) ⟩ + Path (Σ S¹ invS¹Loop) (base , base) (base , base) + ≡⟨ sym ΣPath≡PathΣ ⟩ + Σ ΩS¹ (λ p → PathP (λ j → invS¹Loop (p j)) base base) + ≡⟨ (λ i → Σ ΩS¹ (λ p → PathP (λ j → invS¹Loop (p (j ∨ i))) (twistBaseLoop (p i)) base)) ⟩ + ΩS¹ × ΩS¹ + ≡⟨ (λ i → ΩS¹≡ℤ i × ΩS¹≡ℤ i) ⟩ + ℤ × ℤ ∎ + where + basePath : PathP (λ i → ua kleinBottle≃Σ i) point (base , base) + basePath i = glue (λ {(i = i0) → point; (i = i1) → base , base}) (base , base) + +-- We can at least define the winding function directly and get results on small examples + +windingKlein : Path KleinBottle point point → ℤ × ℤ +windingKlein p = (z₀ , z₁) + where + step₀ : Path (Σ S¹ invS¹Loop) (base , base) (base , base) + step₀ = (λ i → kleinBottle≃Σ .fst (p i)) + + z₀ : ℤ + z₀ = winding (λ i → kleinBottle≃Σ .fst (p i) .fst) + + z₁ : ℤ + z₁ = winding + (transport + (λ i → PathP (λ j → invS¹Loop (step₀ (j ∨ i) .fst)) (twistBaseLoop (step₀ i .fst)) base) + (cong snd step₀)) + +_ : windingKlein line1 ≡ (pos 0 , pos 1) +_ = refl + +_ : windingKlein line2 ≡ (negsuc 0 , pos 0) +_ = refl + +_ : windingKlein (line1 ∙ line2) ≡ (negsuc 0 , negsuc 0) +_ = refl + +_ : windingKlein (line1 ∙ line2 ∙ line1) ≡ (negsuc 0 , pos 0) +_ = refl diff --git a/Cubical/HITs/ListedFiniteSet.agda b/Cubical/HITs/ListedFiniteSet.agda new file mode 100644 index 0000000000..4a1d5e8a54 --- /dev/null +++ b/Cubical/HITs/ListedFiniteSet.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.ListedFiniteSet where + +open import Cubical.HITs.ListedFiniteSet.Base public +open import Cubical.HITs.ListedFiniteSet.Properties public diff --git a/Cubical/HITs/ListedFiniteSet/Base.agda b/Cubical/HITs/ListedFiniteSet/Base.agda new file mode 100644 index 0000000000..d1f6e73816 --- /dev/null +++ b/Cubical/HITs/ListedFiniteSet/Base.agda @@ -0,0 +1,115 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.ListedFiniteSet.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum as ⊎ using (_⊎_; inl; inr) + +open import Cubical.Functions.Logic + +private + variable + ℓ : Level + A B : Type ℓ + +infixr 20 _∷_ +-- infix 30 _∈_ +infixr 5 _++_ + +data LFSet (A : Type ℓ) : Type ℓ where + [] : LFSet A + _∷_ : (x : A) → (xs : LFSet A) → LFSet A + dup : ∀ x xs → x ∷ x ∷ xs ≡ x ∷ xs + comm : ∀ x y xs → x ∷ y ∷ xs ≡ y ∷ x ∷ xs + trunc : isSet (LFSet A) + +-- Membership. +-- +-- Doing some proofs with equational reasoning adds an extra "_∙ refl" +-- at the end. +-- We might want to avoid it, or come up with a more clever equational reasoning. +_∈_ : {A : Type ℓ} → A → LFSet A → hProp ℓ +z ∈ [] = Lift ⊥.⊥ , isOfHLevelLift 1 isProp⊥ +z ∈ (y ∷ xs) = (z ≡ₚ y) ⊔ (z ∈ xs) +z ∈ dup x xs i = proof i + where + -- proof : z ∈ (x ∷ x ∷ xs) ≡ z ∈ (x ∷ xs) + proof = z ≡ₚ x ⊔ (z ≡ₚ x ⊔ z ∈ xs) ≡⟨ ⊔-assoc (z ≡ₚ x) (z ≡ₚ x) (z ∈ xs) ⟩ + (z ≡ₚ x ⊔ z ≡ₚ x) ⊔ z ∈ xs ≡⟨ cong (_⊔ (z ∈ xs)) (⊔-idem (z ≡ₚ x)) ⟩ + z ≡ₚ x ⊔ z ∈ xs ∎ +z ∈ comm x y xs i = proof i + where + -- proof : z ∈ (x ∷ y ∷ xs) ≡ z ∈ (y ∷ x ∷ xs) + proof = z ≡ₚ x ⊔ (z ≡ₚ y ⊔ z ∈ xs) ≡⟨ ⊔-assoc (z ≡ₚ x) (z ≡ₚ y) (z ∈ xs) ⟩ + (z ≡ₚ x ⊔ z ≡ₚ y) ⊔ z ∈ xs ≡⟨ cong (_⊔ (z ∈ xs)) (⊔-comm (z ≡ₚ x) (z ≡ₚ y)) ⟩ + (z ≡ₚ y ⊔ z ≡ₚ x) ⊔ z ∈ xs ≡⟨ sym (⊔-assoc (z ≡ₚ y) (z ≡ₚ x) (z ∈ xs)) ⟩ + z ≡ₚ y ⊔ (z ≡ₚ x ⊔ z ∈ xs) ∎ +x ∈ trunc xs ys p q i j = isSetHProp (x ∈ xs) (x ∈ ys) (cong (x ∈_) p) (cong (x ∈_) q) i j + +module Elim {ℓ} + {B : LFSet A → Type ℓ} + ([]* : B []) + (_∷*_ : (x : A) {xs : LFSet A} → B xs → B (x ∷ xs)) + (comm* : (x y : A) {xs : LFSet A} (b : B xs) + → PathP (λ i → B (comm x y xs i)) (x ∷* (y ∷* b)) (y ∷* (x ∷* b))) + (dup* : (x : A) {xs : LFSet A} (b : B xs) + → PathP (λ i → B (dup x xs i)) (x ∷* (x ∷* b)) (x ∷* b)) + (trunc* : (xs : LFSet A) → isSet (B xs)) where + + f : ∀ x → B x + f [] = []* + f (x ∷ xs) = x ∷* f xs + f (dup x xs i) = dup* x (f xs) i + f (comm x y xs i) = comm* x y (f xs) i + f (trunc x y p q i j) = + isOfHLevel→isOfHLevelDep 2 trunc* + (f x) (f y) + (λ i → f (p i)) (λ i → f (q i)) + (trunc x y p q) i j + +module Rec {ℓ} {B : Type ℓ} + ([]* : B) + (_∷*_ : (x : A) → B → B) + (comm* : (x y : A) (xs : B) → (x ∷* (y ∷* xs)) ≡ (y ∷* (x ∷* xs))) + (dup* : (x : A) (b : B) → (x ∷* (x ∷* b)) ≡ (x ∷* b)) + (trunc* : isSet B) where + + f : LFSet A → B + f = + Elim.f + []* (λ x xs → x ∷* xs) + (λ x y b → comm* x y b) (λ x b → dup* x b) + λ _ → trunc* + +module PropElim {ℓ} + {B : LFSet A → Type ℓ} + ([]* : B []) (_∷*_ : (x : A) {xs : LFSet A} → B xs → B (x ∷ xs)) + (trunc* : (xs : LFSet A) → isProp (B xs)) where + + f : ∀ x → B x + f = + Elim.f + []* _∷*_ + (λ _ _ _ → isOfHLevel→isOfHLevelDep 1 trunc* _ _ _) + (λ _ _ → isOfHLevel→isOfHLevelDep 1 trunc* _ _ _) + λ xs → isProp→isSet (trunc* xs) + +_++_ : ∀ (xs ys : LFSet A) → LFSet A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ (xs ++ ys) +dup x xs i ++ ys = dup x (xs ++ ys) i +comm x y xs i ++ ys = comm x y (xs ++ ys) i +trunc xs zs p q i j ++ ys = + trunc (xs ++ ys) (zs ++ ys) (cong (_++ ys) p) (cong (_++ ys) q) i j + +map : (A → B) → LFSet A → LFSet B +map f [] = [] +map f (x ∷ xs) = f x ∷ map f xs +map f (dup x xs i) = dup (f x) (map f xs) i +map f (comm x y xs i) = comm (f x) (f y) (map f xs) i +map f (trunc xs ys p q i j) = + trunc (map f xs) (map f ys) (cong (map f) p) (cong (map f) q) i j + +disj-union : LFSet A → LFSet B → LFSet (A ⊎ B) +disj-union xs ys = map ⊎.inl xs ++ map ⊎.inr ys diff --git a/Cubical/HITs/ListedFiniteSet/Properties.agda b/Cubical/HITs/ListedFiniteSet/Properties.agda new file mode 100644 index 0000000000..7bee9f313f --- /dev/null +++ b/Cubical/HITs/ListedFiniteSet/Properties.agda @@ -0,0 +1,106 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.ListedFiniteSet.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Prod using (_×_; _,_) + +open import Cubical.HITs.ListedFiniteSet.Base + +private + variable + ℓ : Level + A B : Type ℓ + +assoc-++ : ∀ (xs : LFSet A) ys zs → xs ++ (ys ++ zs) ≡ (xs ++ ys) ++ zs +assoc-++ [] ys zs = refl +assoc-++ (x ∷ xs) ys zs = cong (x ∷_) (assoc-++ xs ys zs) +assoc-++ (dup x xs i) ys zs j = dup x (assoc-++ xs ys zs j) i +assoc-++ (comm x y xs i) ys zs j = comm x y (assoc-++ xs ys zs j) i +assoc-++ (trunc xs xs' p q i k) ys zs j = + trunc + (assoc-++ xs ys zs j) (assoc-++ xs' ys zs j) + (cong (λ xs → assoc-++ xs ys zs j) p) (cong (λ xs → assoc-++ xs ys zs j) q) + i k + +comm-++-[] : ∀ (xs : LFSet A) → xs ++ [] ≡ [] ++ xs +comm-++-[] xs = + PropElim.f + refl + (λ x {xs} ind → + (x ∷ xs) ++ [] ≡⟨ refl ⟩ + x ∷ (xs ++ []) ≡⟨ cong (x ∷_) ind ⟩ + x ∷ xs ≡⟨ refl ⟩ + [] ++ (x ∷ xs) ∎ + ) + (λ _ → trunc _ _) + xs + +comm-++-∷ + : ∀ (z : A) xs ys + → xs ++ (z ∷ ys) ≡ (z ∷ xs) ++ ys +comm-++-∷ z xs ys = + PropElim.f + refl + (λ x {xs} ind → + x ∷ (xs ++ (z ∷ ys)) ≡⟨ cong (x ∷_) ind ⟩ + x ∷ z ∷ (xs ++ ys) ≡⟨ comm x z (xs ++ ys) ⟩ + z ∷ x ∷ (xs ++ ys) ∎ + ) + (λ _ → trunc _ _) + xs + +comm-++ : (xs ys : LFSet A) → xs ++ ys ≡ ys ++ xs +comm-++ xs ys = + PropElim.f + (comm-++-[] xs) + (λ y {ys} ind → + xs ++ (y ∷ ys) ≡⟨ comm-++-∷ y xs ys ⟩ + y ∷ (xs ++ ys) ≡⟨ cong (y ∷_) ind ⟩ + y ∷ (ys ++ xs) ≡⟨ refl ⟩ + (y ∷ ys) ++ xs ∎ + ) + (λ _ → trunc _ _) + ys + +idem-++ : (xs : LFSet A) → xs ++ xs ≡ xs +idem-++ = + PropElim.f + refl + (λ x {xs} ind → + (x ∷ xs) ++ (x ∷ xs) ≡⟨ refl ⟩ + x ∷ (xs ++ (x ∷ xs)) ≡⟨ refl ⟩ + x ∷ (xs ++ ((x ∷ []) ++ xs)) ≡⟨ cong (x ∷_) (assoc-++ xs (x ∷ []) xs) ⟩ + x ∷ ((xs ++ (x ∷ [])) ++ xs) + ≡⟨ cong (λ h → x ∷ (h ++ xs)) (comm-++ xs (x ∷ [])) ⟩ + x ∷ x ∷ (xs ++ xs) ≡⟨ cong (λ ys → x ∷ x ∷ ys) ind ⟩ + x ∷ x ∷ xs ≡⟨ dup x xs ⟩ + x ∷ xs ∎ + ) + (λ xs → trunc (xs ++ xs) xs) + +cart-product : LFSet A → LFSet B → LFSet (A × B) +cart-product [] ys = [] +cart-product (x ∷ xs) ys = map (x ,_) ys ++ cart-product xs ys +cart-product (dup x xs i) ys = + ( map (x ,_) ys ++ (map (x ,_) ys ++ cart-product xs ys) + ≡⟨ assoc-++ (map (x ,_) ys) (map (x ,_) ys) (cart-product xs ys) ⟩ + (map (x ,_) ys ++ map (x ,_) ys) ++ cart-product xs ys + ≡⟨ cong (_++ cart-product xs ys) (idem-++ (map (x ,_) ys)) ⟩ + map (x ,_) ys ++ cart-product xs ys + ∎ + ) i +cart-product (comm x y xs i) ys = + ( map (x ,_) ys ++ (map (y ,_) ys ++ cart-product xs ys) + ≡⟨ assoc-++ (map (x ,_) ys) (map (y ,_) ys) (cart-product xs ys) ⟩ + (map (x ,_) ys ++ map (y ,_) ys) ++ cart-product xs ys + ≡⟨ cong (_++ cart-product xs ys) (comm-++ (map (x ,_) ys) (map (y ,_) ys)) ⟩ + (map (y ,_) ys ++ map (x ,_) ys) ++ cart-product xs ys + ≡⟨ sym (assoc-++ (map (y ,_) ys) (map (x ,_) ys) (cart-product xs ys)) ⟩ + map (y ,_) ys ++ (map (x ,_) ys ++ cart-product xs ys) + ∎ + ) i +cart-product (trunc xs xs′ p q i j) ys = + trunc + (cart-product xs ys) (cart-product xs′ ys) + (λ k → cart-product (p k) ys) (λ k → cart-product (q k) ys) + i j diff --git a/Cubical/HITs/Localization.agda b/Cubical/HITs/Localization.agda new file mode 100644 index 0000000000..e8bd1f0178 --- /dev/null +++ b/Cubical/HITs/Localization.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Localization where + +open import Cubical.HITs.Localization.Base public + +open import Cubical.HITs.Localization.Properties public diff --git a/Cubical/HITs/Localization/Base.agda b/Cubical/HITs/Localization/Base.agda new file mode 100644 index 0000000000..d6c378d72e --- /dev/null +++ b/Cubical/HITs/Localization/Base.agda @@ -0,0 +1,26 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Localization.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.PathSplit +open isPathSplitEquiv + +module _ {ℓα ℓs ℓt} {A : Type ℓα} {S : A → Type ℓs} {T : A → Type ℓt} where + + isLocal : ∀ (F : ∀ α → S α → T α) {ℓ} (X : Type ℓ) → Type _ + isLocal F X = ∀ α → isPathSplitEquiv (λ (g : T α → X) → g ∘ F α) + + data Localize (F : ∀ α → S α → T α) {ℓ} (X : Type ℓ) : Type (ℓ-max ℓ (ℓ-max ℓα (ℓ-max ℓs ℓt))) where + ∣_∣ : X → Localize F X + -- (_∘ F α) : (T α → Localize F X) → (S α → Localize F X) is a path-split equivalence ∀ α + ext : ∀ α → (S α → Localize F X) → (T α → Localize F X) + isExt : ∀ α (f : S α → Localize F X) (s : S α) → ext α f (F α s) ≡ f s + ≡ext : ∀ α (g h : T α → Localize F X) → ((s : S α) → g (F α s) ≡ h (F α s)) → ((t : T α) → g t ≡ h t) + ≡isExt : ∀ α g h (p : (s : S α) → g (F α s) ≡ h (F α s)) (s : S α) → ≡ext α g h p (F α s) ≡ p s + + isLocal-Localize : ∀ (F : ∀ α → S α → T α) {ℓ} (X : Type ℓ) → isLocal F (Localize F X) + fst (sec (isLocal-Localize F X α)) f t = ext α f t + snd (sec (isLocal-Localize F X α)) f i s = isExt α f s i + fst (secCong (isLocal-Localize F X α) g h) p i t = ≡ext α g h (funExt⁻ p) t i + snd (secCong (isLocal-Localize F X α) g h) p i j t = ≡isExt α g h (funExt⁻ p) t i j diff --git a/Cubical/HITs/Localization/Properties.agda b/Cubical/HITs/Localization/Properties.agda new file mode 100644 index 0000000000..0a36c9429e --- /dev/null +++ b/Cubical/HITs/Localization/Properties.agda @@ -0,0 +1,23 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Localization.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv.PathSplit +open isPathSplitEquiv + +open import Cubical.HITs.Localization.Base + +module _ {ℓα ℓs ℓt} {A : Type ℓα} {S : A → Type ℓs} {T : A → Type ℓt} where + + rec : ∀ {F : ∀ α → S α → T α} {ℓ ℓ'} {X : Type ℓ} {Y : Type ℓ'} + → (lY : isLocal F Y) → (X → Y) → Localize F X → Y + rec lY g ∣ x ∣ = g x + rec lY g (ext α f t) = fst (sec (lY α)) (λ s → rec lY g (f s)) t + rec lY g (isExt α f s i) = snd (sec (lY α)) (λ s → rec lY g (f s)) i s + rec lY f (≡ext α g h p t i) = fst (secCong (lY α) (λ t → rec lY f (g t)) (λ t → rec lY f (h t))) + (λ i s → rec lY f (p s i)) i t + rec lY f (≡isExt α g h p t i j) = snd (secCong (lY α) (λ t → rec lY f (g t)) (λ t → rec lY f (h t))) + (λ i s → rec lY f (p s i)) i j t diff --git a/Cubical/HITs/MappingCones.agda b/Cubical/HITs/MappingCones.agda new file mode 100644 index 0000000000..29543c1de1 --- /dev/null +++ b/Cubical/HITs/MappingCones.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.MappingCones where + +open import Cubical.HITs.MappingCones.Base public +open import Cubical.HITs.MappingCones.Properties public diff --git a/Cubical/HITs/MappingCones/Base.agda b/Cubical/HITs/MappingCones/Base.agda new file mode 100644 index 0000000000..0c533ea3ef --- /dev/null +++ b/Cubical/HITs/MappingCones/Base.agda @@ -0,0 +1,26 @@ +{- + +Mapping cones or the homotopy cofiber/cokernel + +-} +{-# OPTIONS --safe #-} + +module Cubical.HITs.MappingCones.Base where + +open import Cubical.Foundations.Prelude + +private + variable + ℓ ℓ' ℓ'' : Level + +data Cone {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) : Type (ℓ-max ℓ ℓ') where + inj : Y → Cone f + hub : Cone f + spoke : (x : X) → hub ≡ inj (f x) + +-- the attachment of multiple mapping cones + +data Cones {X : Type ℓ} {Y : Type ℓ'} (A : Type ℓ'') (f : A → X → Y) : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') where + inj : Y → Cones A f + hub : A → Cones A f + spoke : (a : A) (x : X) → hub a ≡ inj (f a x) diff --git a/Cubical/HITs/MappingCones/Properties.agda b/Cubical/HITs/MappingCones/Properties.agda new file mode 100644 index 0000000000..3c344b2fd9 --- /dev/null +++ b/Cubical/HITs/MappingCones/Properties.agda @@ -0,0 +1,51 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.MappingCones.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Unit +open import Cubical.Data.Sum +open import Cubical.HITs.Pushout + +open import Cubical.HITs.MappingCones.Base + +private + variable + ℓ ℓ' ℓ'' : Level + +PushoutUnit-iso-Cone : ∀ {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) → Iso (Pushout (const tt) f) (Cone f) +Iso.fun (PushoutUnit-iso-Cone f) (inl tt) = hub +Iso.fun (PushoutUnit-iso-Cone f) (inr x) = inj x +Iso.fun (PushoutUnit-iso-Cone f) (push x i) = spoke x i +Iso.inv (PushoutUnit-iso-Cone f) (inj x) = inr x +Iso.inv (PushoutUnit-iso-Cone f) hub = inl tt +Iso.inv (PushoutUnit-iso-Cone f) (spoke x i) = push x i +Iso.rightInv (PushoutUnit-iso-Cone f) (inj x) = refl +Iso.rightInv (PushoutUnit-iso-Cone f) hub = refl +Iso.rightInv (PushoutUnit-iso-Cone f) (spoke x i) = refl +Iso.leftInv (PushoutUnit-iso-Cone f) (inl tt) = refl +Iso.leftInv (PushoutUnit-iso-Cone f) (inr x) = refl +Iso.leftInv (PushoutUnit-iso-Cone f) (push x i) = refl + +PushoutUnit≡Cone : ∀ {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) → Pushout (const tt) f ≡ Cone f +PushoutUnit≡Cone f = isoToPath (PushoutUnit-iso-Cone f) + +ConesUnit-iso-Cone : ∀ {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) → Iso (Cones Unit (λ { tt → f })) (Cone f) +Iso.fun (ConesUnit-iso-Cone f) (inj x) = inj x +Iso.fun (ConesUnit-iso-Cone f) (hub tt) = hub +Iso.fun (ConesUnit-iso-Cone f) (spoke tt x i) = spoke x i +Iso.inv (ConesUnit-iso-Cone f) (inj x) = inj x +Iso.inv (ConesUnit-iso-Cone f) hub = hub tt +Iso.inv (ConesUnit-iso-Cone f) (spoke x i) = spoke tt x i +Iso.rightInv (ConesUnit-iso-Cone f) (inj x) = refl +Iso.rightInv (ConesUnit-iso-Cone f) hub = refl +Iso.rightInv (ConesUnit-iso-Cone f) (spoke x i) = refl +Iso.leftInv (ConesUnit-iso-Cone f) (inj x) = refl +Iso.leftInv (ConesUnit-iso-Cone f) (hub x) = refl +Iso.leftInv (ConesUnit-iso-Cone f) (spoke a x i) = refl + +ConesUnit≡Cone : ∀ {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) → (Cones Unit (λ { tt → f })) ≡ (Cone f) +ConesUnit≡Cone f = isoToPath (ConesUnit-iso-Cone f) diff --git a/Cubical/HITs/Modulo.agda b/Cubical/HITs/Modulo.agda new file mode 100644 index 0000000000..50dfa2e9b1 --- /dev/null +++ b/Cubical/HITs/Modulo.agda @@ -0,0 +1,8 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.Modulo where + +open import Cubical.HITs.Modulo.Base public +open import Cubical.HITs.Modulo.Properties public + +open import Cubical.HITs.Modulo.FinEquiv public diff --git a/Cubical/HITs/Modulo/Base.agda b/Cubical/HITs/Modulo/Base.agda new file mode 100644 index 0000000000..ef61722803 --- /dev/null +++ b/Cubical/HITs/Modulo/Base.agda @@ -0,0 +1,60 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.Modulo.Base where + +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Empty using (⊥) +open import Cubical.Data.Nat using (ℕ; zero; suc; _+_) +open import Cubical.Data.Unit renaming (Unit to ⊤) + +open import Cubical.Relation.Nullary + +NonZero : ℕ → Type₀ +NonZero 0 = ⊥ +NonZero _ = ⊤ + +private + variable + ℓ : Level + k : ℕ + +-- The Modulo type is similar to the Fin type, but instead of being +-- inhabited by canonical values, the inhabitants are all naturals, +-- and paths are added between numbers that have the same residue. +-- +-- This representation makes it easier to do certain arithmetic +-- without changing the modulus. For instance, we can just add any +-- natural to a Modulo k to get another, whereas with Fin k, we must +-- calculate the canonical representative. +-- +-- The reason the path constructor is guarded is to avoid adding +-- non-trivial path structure to the k=0 case. If it were not guarded, +-- each `Modulo 0` would become like the circle, and guarding the +-- constructor is somewhat easier to work with than truncation. +-- +-- Note also that unlike `Fin 0`, `Modulo 0` is equivalent to the naturals. +data Modulo (k : ℕ) : Type₀ where + embed : (n : ℕ) → Modulo k + pre-step : NonZero k → (n : ℕ) → embed n ≡ embed (k + n) + +-- When we are working with k = suc k₀, the `step` alias is much +-- we can use this alias. +pattern step n i = pre-step _ n i + +-- Helper to avoid having to case on `k` in certain places. +ztep : ∀{k} n → Path (Modulo k) (embed n) (embed (k + n)) +ztep {0} n = refl +ztep {suc k} n = step n + +-- The standard eliminator for `Modulo`. +elim + : (P : ∀ k → Modulo k → Type ℓ) + → (e : ∀ k n → P k (embed n)) + → (st : ∀ k n → PathP (λ i → P (suc k) (step n i)) (e (suc k) n) (e (suc k) (suc k + n))) + → (m : Modulo k) → P k m +elim P e st (embed n) = e _ n +elim {k = suc k} P e st (step n i) = st k n i diff --git a/Cubical/HITs/Modulo/FinEquiv.agda b/Cubical/HITs/Modulo/FinEquiv.agda new file mode 100644 index 0000000000..e7ee0bb7d8 --- /dev/null +++ b/Cubical/HITs/Modulo/FinEquiv.agda @@ -0,0 +1,77 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.Modulo.FinEquiv where + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Fin +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order + +open import Cubical.HITs.Modulo.Base + +-- For positive `k`, every `Modulo k` can be reduced to its +-- residue modulo `k`, given by a value of type `Fin k`. This +-- forms an equivalence between `Fin k` and `Modulo k`. +module Reduction {k₀ : ℕ} where + private + k = suc k₀ + + fembed : Fin k → Modulo k + fembed = embed ∘ toℕ + + ResiduePath : ℕ → Type₀ + ResiduePath n = Σ[ f ∈ Fin k ] fembed f ≡ embed n + + rbase : ∀ n (n ‖ℕ‖ is nearly Invertible + https://homotopytypetheory.org/2013/10/28/the-truncation-map-_-ℕ-‖ℕ‖-is-nearly-invertible/ + +Defines [recover], which definitionally satisfies `recover ∣ x ∣ ≡ x` ([recover∣∣]) for homogeneous types + +Also see the follow-up post by Jason Gross: + Composition is not what you think it is! Why “nearly invertible” isn’t. + https://homotopytypetheory.org/2014/02/24/composition-is-not-what-you-think-it-is-why-nearly-invertible-isnt/ + +-} +{-# OPTIONS --safe #-} + +module Cubical.HITs.PropositionalTruncation.MagicTrick where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.HITs.PropositionalTruncation.Base +open import Cubical.HITs.PropositionalTruncation.Properties + +module Recover {ℓ} (A∙ : Pointed ℓ) (h : isHomogeneous A∙) where + private + A = typ A∙ + a = pt A∙ + + toEquivPtd : ∥ A ∥ → Σ[ B∙ ∈ Pointed ℓ ] (A , a) ≡ B∙ + toEquivPtd = rec isPropSingl (λ x → (A , x) , h x) + private + B∙ : ∥ A ∥ → Pointed ℓ + B∙ tx = fst (toEquivPtd tx) + + -- the key observation is that B∙ ∣ x ∣ is definitionally equal to (A , x) + private + obvs : ∀ x → B∙ ∣ x ∣ ≡ (A , x) + obvs x = refl -- try it: `C-c C-n B∙ ∣ x ∣` gives `(A , x)` + + -- thus any truncated element (of a homogeneous type) can be recovered by agda's normalizer! + + recover : ∀ (tx : ∥ A ∥) → typ (B∙ tx) + recover tx = pt (B∙ tx) + + recover∣∣ : ∀ (x : A) → recover ∣ x ∣ ≡ x + recover∣∣ x = refl -- try it: `C-c C-n recover ∣ x ∣` gives `x` + + private + -- notice that the following typechecks because typ (B∙ ∣ x ∣) is definitionally equal to to A, but + -- `recover : ∥ A ∥ → A` does not because typ (B∙ tx) is not definitionally equal to A (though it is + -- judegmentally equal to A by cong typ (snd (toEquivPtd tx)) : A ≡ typ (B∙ tx)) + obvs2 : A → A + obvs2 = recover ∘ ∣_∣ + + -- one might wonder if (cong recover (squash ∣ x ∣ ∣ y ∣)) therefore has type x ≡ y, but thankfully + -- typ (B∙ (squash ∣ x ∣ ∣ y ∣ i)) is *not* A (it's a messy hcomp involving h x and h y) + recover-squash : ∀ x y → -- x ≡ y -- this raises an error + PathP (λ i → typ (B∙ (squash ∣ x ∣ ∣ y ∣ i))) x y + recover-squash x y = cong recover (squash ∣ x ∣ ∣ y ∣) + + +-- Demo, adapted from: +-- https://bitbucket.org/nicolaikraus/agda/src/e30d70c72c6af8e62b72eefabcc57623dd921f04/trunc-inverse.lagda + +private + open import Cubical.Data.Nat + open Recover (ℕ , zero) (isHomogeneousDiscrete discreteℕ) + + -- only `∣hidden∣` is exported, `hidden` is no longer in scope + module _ where + private + hidden : ℕ + hidden = 17 + + ∣hidden∣ : ∥ ℕ ∥ + ∣hidden∣ = ∣ hidden ∣ + + -- we can still recover the value, even though agda can no longer see `hidden`! + test : recover ∣hidden∣ ≡ 17 + test = refl -- try it: `C-c C-n recover ∣hidden∣` gives `17` + -- `C-c C-n hidden` gives an error + + -- Finally, note that the definition of recover is independent of the proof that A is homogeneous. Thus we + -- still can definitionally recover information hidden by ∣_∣ as long as we permit holes. Try replacing + -- `isHomogeneousDiscrete discreteℕ` above with a hole (`?`) and notice that everything still works diff --git a/Cubical/HITs/PropositionalTruncation/Monad.agda b/Cubical/HITs/PropositionalTruncation/Monad.agda new file mode 100644 index 0000000000..7ecb829a5a --- /dev/null +++ b/Cubical/HITs/PropositionalTruncation/Monad.agda @@ -0,0 +1,33 @@ +{-# OPTIONS --safe #-} +{- +Implements the monadic interface of propositional truncation, for reasoning in do-syntax. +-} + +module Cubical.HITs.PropositionalTruncation.Monad where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Functions.Logic +open import Cubical.HITs.PropositionalTruncation + +private + variable + ℓ : Level + P Q : Type ℓ + +infix 1 proof_by_ +proof_by_ : (P : hProp ℓ) → ∥ ⟨ P ⟩ ∥ → ⟨ P ⟩ +proof P by p = rec (isProp⟨⟩ P) (λ p → p) p + +return : P → ∥ P ∥ +return p = ∣ p ∣ + +exact_ : ∥ P ∥ → ∥ P ∥ +exact p = p + +_>>=_ : ∥ P ∥ → (P → ∥ Q ∥) → ∥ Q ∥ +p >>= f = rec isPropPropTrunc f p + +_>>_ : ∥ P ∥ → ∥ Q ∥ → ∥ Q ∥ +_ >> q = q diff --git a/Cubical/HITs/PropositionalTruncation/Properties.agda b/Cubical/HITs/PropositionalTruncation/Properties.agda new file mode 100644 index 0000000000..0594340184 --- /dev/null +++ b/Cubical/HITs/PropositionalTruncation/Properties.agda @@ -0,0 +1,499 @@ +{- + +This file contains: + +- Eliminator for propositional truncation + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.PropositionalTruncation.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma +open import Cubical.Data.Sum hiding (rec ; elim ; map) + +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓ ℓ′ : Level + A B C : Type ℓ + A′ : Type ℓ′ + +∥∥-isPropDep : (P : A → Type ℓ) → isOfHLevelDep 1 (λ x → ∥ P x ∥) +∥∥-isPropDep P = isOfHLevel→isOfHLevelDep 1 (λ _ → squash) + +rec : {P : Type ℓ} → isProp P → (A → P) → ∥ A ∥ → P +rec Pprop f ∣ x ∣ = f x +rec Pprop f (squash x y i) = Pprop (rec Pprop f x) (rec Pprop f y) i + +rec2 : {P : Type ℓ} → isProp P → (A → B → P) → ∥ A ∥ → ∥ B ∥ → P +rec2 Pprop f ∣ x ∣ ∣ y ∣ = f x y +rec2 Pprop f ∣ x ∣ (squash y z i) = Pprop (rec2 Pprop f ∣ x ∣ y) (rec2 Pprop f ∣ x ∣ z) i +rec2 Pprop f (squash x y i) z = Pprop (rec2 Pprop f x z) (rec2 Pprop f y z) i + +-- Old version +-- rec2 : ∀ {P : Type ℓ} → isProp P → (A → A → P) → ∥ A ∥ → ∥ A ∥ → P +-- rec2 Pprop f = rec (isProp→ Pprop) (λ a → rec Pprop (f a)) + +elim : {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) + → ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a +elim Pprop f ∣ x ∣ = f x +elim Pprop f (squash x y i) = + isOfHLevel→isOfHLevelDep 1 Pprop + (elim Pprop f x) (elim Pprop f y) (squash x y) i + +elim2 : {P : ∥ A ∥ → ∥ B ∥ → Type ℓ} + (Pprop : (x : ∥ A ∥) (y : ∥ B ∥) → isProp (P x y)) + (f : (a : A) (b : B) → P ∣ a ∣ ∣ b ∣) + (x : ∥ A ∥) (y : ∥ B ∥) → P x y +elim2 Pprop f = + elim (λ _ → isPropΠ (λ _ → Pprop _ _)) + (λ a → elim (λ _ → Pprop _ _) (f a)) + +elim3 : {P : ∥ A ∥ → ∥ B ∥ → ∥ C ∥ → Type ℓ} + (Pprop : ((x : ∥ A ∥) (y : ∥ B ∥) (z : ∥ C ∥) → isProp (P x y z))) + (g : (a : A) (b : B) (c : C) → P (∣ a ∣) ∣ b ∣ ∣ c ∣) + (x : ∥ A ∥) (y : ∥ B ∥) (z : ∥ C ∥) → P x y z +elim3 Pprop g = elim2 (λ _ _ → isPropΠ (λ _ → Pprop _ _ _)) + (λ a b → elim (λ _ → Pprop _ _ _) (g a b)) + +isPropPropTrunc : isProp ∥ A ∥ +isPropPropTrunc x y = squash x y + +propTrunc≃ : A ≃ B → ∥ A ∥ ≃ ∥ B ∥ +propTrunc≃ e = + propBiimpl→Equiv + isPropPropTrunc isPropPropTrunc + (rec isPropPropTrunc (λ a → ∣ e .fst a ∣)) + (rec isPropPropTrunc (λ b → ∣ invEq e b ∣)) + +propTruncIdempotent≃ : isProp A → ∥ A ∥ ≃ A +propTruncIdempotent≃ {A = A} hA = isoToEquiv f + where + f : Iso ∥ A ∥ A + Iso.fun f = rec hA (idfun A) + Iso.inv f x = ∣ x ∣ + Iso.rightInv f _ = refl + Iso.leftInv f = elim (λ _ → isProp→isSet isPropPropTrunc _ _) (λ _ → refl) + +propTruncIdempotent : isProp A → ∥ A ∥ ≡ A +propTruncIdempotent hA = ua (propTruncIdempotent≃ hA) + +-- We could also define the eliminator using the recursor +elim' : {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → + ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a +elim' {P = P} Pprop f a = + rec (Pprop a) (λ x → transp (λ i → P (squash ∣ x ∣ a i)) i0 (f x)) a + +map : (A → B) → (∥ A ∥ → ∥ B ∥) +map f = rec squash (∣_∣ ∘ f) + +map2 : (A → B → C) → (∥ A ∥ → ∥ B ∥ → ∥ C ∥) +map2 f = rec (isPropΠ λ _ → squash) (map ∘ f) + +-- The propositional truncation can be eliminated into non-propositional +-- types as long as the function used in the eliminator is 'coherently +-- constant.' The details of this can be found in the following paper: +-- +-- https://arxiv.org/pdf/1411.2682.pdf +module SetElim (Bset : isSet B) where + Bset' : isSet' B + Bset' = isSet→isSet' Bset + + rec→Set : (f : A → B) (kf : 2-Constant f) → ∥ A ∥ → B + helper : (f : A → B) (kf : 2-Constant f) → (t u : ∥ A ∥) + → rec→Set f kf t ≡ rec→Set f kf u + + rec→Set f kf ∣ x ∣ = f x + rec→Set f kf (squash t u i) = helper f kf t u i + + helper f kf ∣ x ∣ ∣ y ∣ = kf x y + helper f kf (squash t u i) v + = Bset' (helper f kf t v) (helper f kf u v) (helper f kf t u) refl i + helper f kf t (squash u v i) + = Bset' (helper f kf t u) (helper f kf t v) refl (helper f kf u v) i + + kcomp : (f : ∥ A ∥ → B) → 2-Constant (f ∘ ∣_∣) + kcomp f x y = cong f (squash ∣ x ∣ ∣ y ∣) + + Fset : isSet (A → B) + Fset = isSetΠ (const Bset) + + Kset : (f : A → B) → isSet (2-Constant f) + Kset f = isSetΠ (λ _ → isSetΠ (λ _ → isProp→isSet (Bset _ _))) + + setRecLemma + : (f : ∥ A ∥ → B) + → rec→Set (f ∘ ∣_∣) (kcomp f) ≡ f + setRecLemma f i t + = elim {P = λ t → rec→Set (f ∘ ∣_∣) (kcomp f) t ≡ f t} + (λ t → Bset _ _) (λ x → refl) t i + + mkKmap : (∥ A ∥ → B) → Σ (A → B) 2-Constant + mkKmap f = f ∘ ∣_∣ , kcomp f + + fib : (g : Σ (A → B) 2-Constant) → fiber mkKmap g + fib (g , kg) = rec→Set g kg , refl + + eqv : (g : Σ (A → B) 2-Constant) → ∀ fi → fib g ≡ fi + eqv g (f , p) = + Σ≡Prop (λ f → isOfHLevelΣ 2 Fset Kset _ _) + (cong (uncurry rec→Set) (sym p) ∙ setRecLemma f) + + trunc→Set≃ : (∥ A ∥ → B) ≃ (Σ (A → B) 2-Constant) + trunc→Set≃ .fst = mkKmap + trunc→Set≃ .snd .equiv-proof g = fib g , eqv g + + -- The strategy of this equivalence proof follows the paper more closely. + -- It is used further down for the groupoid version, because the above + -- strategy does not generalize so easily. + e : B → Σ (A → B) 2-Constant + e b = const b , λ _ _ → refl + + eval : A → (γ : Σ (A → B) 2-Constant) → B + eval a₀ (g , _) = g a₀ + + e-eval : ∀ (a₀ : A) γ → e (eval a₀ γ) ≡ γ + e-eval a₀ (g , kg) i .fst a₁ = kg a₀ a₁ i + e-eval a₀ (g , kg) i .snd a₁ a₂ = Bset' refl (kg a₁ a₂) (kg a₀ a₁) (kg a₀ a₂) i + + e-isEquiv : A → isEquiv (e {A = A}) + e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ → refl) + + preEquiv₁ : ∥ A ∥ → B ≃ Σ (A → B) 2-Constant + preEquiv₁ t = e , rec (isPropIsEquiv e) e-isEquiv t + + preEquiv₂ : (∥ A ∥ → Σ (A → B) 2-Constant) ≃ Σ (A → B) 2-Constant + preEquiv₂ = isoToEquiv (iso to const (λ _ → refl) retr) + where + to : (∥ A ∥ → Σ (A → B) 2-Constant) → Σ (A → B) 2-Constant + to f .fst x = f ∣ x ∣ .fst x + to f .snd x y i = f (squash ∣ x ∣ ∣ y ∣ i) .snd x y i + + retr : retract to const + retr f i t .fst x = f (squash ∣ x ∣ t i) .fst x + retr f i t .snd x y + = Bset' + (λ j → f (squash ∣ x ∣ ∣ y ∣ j) .snd x y j) + (f t .snd x y) + (λ j → f (squash ∣ x ∣ t j) .fst x) + (λ j → f (squash ∣ y ∣ t j) .fst y) + i + + trunc→Set≃₂ : (∥ A ∥ → B) ≃ Σ (A → B) 2-Constant + trunc→Set≃₂ = compEquiv (equivΠCod preEquiv₁) preEquiv₂ + +open SetElim public using (rec→Set; trunc→Set≃) + +elim→Set + : {P : ∥ A ∥ → Type ℓ} + → (∀ t → isSet (P t)) + → (f : (x : A) → P ∣ x ∣) + → (kf : ∀ x y → PathP (λ i → P (squash ∣ x ∣ ∣ y ∣ i)) (f x) (f y)) + → (t : ∥ A ∥) → P t +elim→Set {A = A} {P = P} Pset f kf t + = rec→Set (Pset t) g gk t + where + g : A → P t + g x = transp (λ i → P (squash ∣ x ∣ t i)) i0 (f x) + + gk : 2-Constant g + gk x y i = transp (λ j → P (squash (squash ∣ x ∣ ∣ y ∣ i) t j)) i0 (kf x y i) + +elim2→Set : + {P : ∥ A ∥ → ∥ B ∥ → Type ℓ} + → (∀ t u → isSet (P t u)) + → (f : (x : A) (y : B) → P ∣ x ∣ ∣ y ∣) + → (kf₁ : ∀ x y v → PathP (λ i → P (squash ∣ x ∣ ∣ y ∣ i) ∣ v ∣) (f x v) (f y v)) + → (kf₂ : ∀ x v w → PathP (λ i → P ∣ x ∣ (squash ∣ v ∣ ∣ w ∣ i)) (f x v) (f x w)) + → (sf : ∀ x y v w → SquareP (λ i j → P (squash ∣ x ∣ ∣ y ∣ i) (squash ∣ v ∣ ∣ w ∣ j)) + (kf₂ x v w) (kf₂ y v w) (kf₁ x y v) (kf₁ x y w)) + → (t : ∥ A ∥) → (u : ∥ B ∥) → P t u +elim2→Set {A = A} {B = B} {P = P} Pset f kf₁ kf₂ sf = + elim→Set (λ _ → isSetΠ (λ _ → Pset _ _)) mapHelper squareHelper + where + mapHelper : (x : A) (u : ∥ B ∥) → P ∣ x ∣ u + mapHelper x = elim→Set (λ _ → Pset _ _) (f x) (kf₂ x) + + squareHelper : (x y : A) + → PathP (λ i → (u : ∥ B ∥) → P (squash ∣ x ∣ ∣ y ∣ i) u) (mapHelper x) (mapHelper y) + squareHelper x y i = elim→Set (λ _ → Pset _ _) (λ v → kf₁ x y v i) λ v w → sf x y v w i + +RecHProp : (P : A → hProp ℓ) (kP : ∀ x y → P x ≡ P y) → ∥ A ∥ → hProp ℓ +RecHProp P kP = rec→Set isSetHProp P kP + +module GpdElim (Bgpd : isGroupoid B) where + Bgpd' : isGroupoid' B + Bgpd' = isGroupoid→isGroupoid' Bgpd + + module _ (f : A → B) (3kf : 3-Constant f) where + open 3-Constant 3kf + + rec→Gpd : ∥ A ∥ → B + pathHelper : (t u : ∥ A ∥) → rec→Gpd t ≡ rec→Gpd u + triHelper₁ + : (t u v : ∥ A ∥) + → Square (pathHelper t u) (pathHelper t v) refl (pathHelper u v) + triHelper₂ + : (t u v : ∥ A ∥) + → Square (pathHelper t v) (pathHelper u v) (pathHelper t u) refl + + rec→Gpd ∣ x ∣ = f x + rec→Gpd (squash t u i) = pathHelper t u i + + pathHelper ∣ x ∣ ∣ y ∣ = link x y + pathHelper (squash t u j) v = triHelper₂ t u v j + pathHelper ∣ x ∣ (squash u v j) = triHelper₁ ∣ x ∣ u v j + + triHelper₁ ∣ x ∣ ∣ y ∣ ∣ z ∣ = coh₁ x y z + triHelper₁ (squash s t i) u v + = Bgpd' + (triHelper₁ s u v) + (triHelper₁ t u v) + (triHelper₂ s t u) + (triHelper₂ s t v) + (λ i → refl) + (λ i → pathHelper u v) + i + triHelper₁ ∣ x ∣ (squash t u i) v + = Bgpd' + (triHelper₁ ∣ x ∣ t v) + (triHelper₁ ∣ x ∣ u v) + (triHelper₁ ∣ x ∣ t u) + (λ i → pathHelper ∣ x ∣ v) + (λ i → refl) + (triHelper₂ t u v) + i + triHelper₁ ∣ x ∣ ∣ y ∣ (squash u v i) + = Bgpd' + (triHelper₁ ∣ x ∣ ∣ y ∣ u) + (triHelper₁ ∣ x ∣ ∣ y ∣ v) + (λ i → link x y) + (triHelper₁ ∣ x ∣ u v) + (λ i → refl) + (triHelper₁ ∣ y ∣ u v) + i + + triHelper₂ ∣ x ∣ ∣ y ∣ ∣ z ∣ = coh₂ x y z + triHelper₂ (squash s t i) u v + = Bgpd' + (triHelper₂ s u v) + (triHelper₂ t u v) + (triHelper₂ s t v) + (λ i → pathHelper u v) + (triHelper₂ s t u) + (λ i → refl) + i + triHelper₂ ∣ x ∣ (squash t u i) v + = Bgpd' + (triHelper₂ ∣ x ∣ t v) + (triHelper₂ ∣ x ∣ u v) + (λ i → pathHelper ∣ x ∣ v) + (triHelper₂ t u v) + (triHelper₁ ∣ x ∣ t u) + (λ i → refl) + i + triHelper₂ ∣ x ∣ ∣ y ∣ (squash u v i) + = Bgpd' + (triHelper₂ ∣ x ∣ ∣ y ∣ u) + (triHelper₂ ∣ x ∣ ∣ y ∣ v) + (triHelper₁ ∣ x ∣ u v) + (triHelper₁ ∣ y ∣ u v) + (λ i → link x y) + (λ i → refl) + i + + preEquiv₁ : (∥ A ∥ → Σ (A → B) 3-Constant) ≃ Σ (A → B) 3-Constant + preEquiv₁ = isoToEquiv (iso fn const (λ _ → refl) retr) + where + open 3-Constant + + fn : (∥ A ∥ → Σ (A → B) 3-Constant) → Σ (A → B) 3-Constant + fn f .fst x = f ∣ x ∣ .fst x + fn f .snd .link x y i = f (squash ∣ x ∣ ∣ y ∣ i) .snd .link x y i + fn f .snd .coh₁ x y z i j + = f (squash ∣ x ∣ (squash ∣ y ∣ ∣ z ∣ i) j) .snd .coh₁ x y z i j + + retr : retract fn const + retr f i t .fst x = f (squash ∣ x ∣ t i) .fst x + retr f i t .snd .link x y j + = f (squash (squash ∣ x ∣ ∣ y ∣ j) t i) .snd .link x y j + retr f i t .snd .coh₁ x y z + = Bgpd' + (λ k j → f (cb k j i0) .snd .coh₁ x y z k j ) + (λ k j → f (cb k j i1) .snd .coh₁ x y z k j) + (λ k j → f (cb i0 j k) .snd .link x y j) + (λ k j → f (cb i1 j k) .snd .link x z j) + (λ _ → refl) + (λ k j → f (cb j i1 k) .snd .link y z j) + i + where + cb : I → I → I → ∥ _ ∥ + cb i j k = squash (squash ∣ x ∣ (squash ∣ y ∣ ∣ z ∣ i) j) t k + + e : B → Σ (A → B) 3-Constant + e b .fst _ = b + e b .snd = record + { link = λ _ _ _ → b + ; coh₁ = λ _ _ _ _ _ → b + } + + eval : A → Σ (A → B) 3-Constant → B + eval a₀ (g , _) = g a₀ + + module _ where + open 3-Constant + e-eval : ∀(a₀ : A) γ → e (eval a₀ γ) ≡ γ + e-eval a₀ (g , 3kg) i .fst x = 3kg .link a₀ x i + e-eval a₀ (g , 3kg) i .snd .link x y = λ j → 3kg .coh₁ a₀ x y j i + e-eval a₀ (g , 3kg) i .snd .coh₁ x y z + = Bgpd' + (λ _ _ → g a₀) + (3kg .coh₁ x y z) + (λ k j → 3kg .coh₁ a₀ x y j k) + (λ k j → 3kg .coh₁ a₀ x z j k) + (λ _ → refl) + (λ k j → 3kg .coh₁ a₀ y z j k) + i + + e-isEquiv : A → isEquiv (e {A = A}) + e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ → refl) + + preEquiv₂ : ∥ A ∥ → B ≃ Σ (A → B) 3-Constant + preEquiv₂ t = e , rec (isPropIsEquiv e) e-isEquiv t + + trunc→Gpd≃ : (∥ A ∥ → B) ≃ Σ (A → B) 3-Constant + trunc→Gpd≃ = compEquiv (equivΠCod preEquiv₂) preEquiv₁ + +open GpdElim using (rec→Gpd; trunc→Gpd≃) public + +squashᵗ + : ∀(x y z : A) + → Square (squash ∣ x ∣ ∣ y ∣) (squash ∣ x ∣ ∣ z ∣) refl (squash ∣ y ∣ ∣ z ∣) +squashᵗ x y z i = squash ∣ x ∣ (squash ∣ y ∣ ∣ z ∣ i) + +elim→Gpd + : (P : ∥ A ∥ → Type ℓ) + → (∀ t → isGroupoid (P t)) + → (f : (x : A) → P ∣ x ∣) + → (kf : ∀ x y → PathP (λ i → P (squash ∣ x ∣ ∣ y ∣ i)) (f x) (f y)) + → (3kf : ∀ x y z + → SquareP (λ i j → P (squashᵗ x y z i j)) (kf x y) (kf x z) refl (kf y z)) + → (t : ∥ A ∥) → P t +elim→Gpd {A = A} P Pgpd f kf 3kf t = rec→Gpd (Pgpd t) g 3kg t + where + g : A → P t + g x = transp (λ i → P (squash ∣ x ∣ t i)) i0 (f x) + + open 3-Constant + + 3kg : 3-Constant g + 3kg .link x y i + = transp (λ j → P (squash (squash ∣ x ∣ ∣ y ∣ i) t j)) i0 (kf x y i) + 3kg .coh₁ x y z i j + = transp (λ k → P (squash (squashᵗ x y z i j) t k)) i0 (3kf x y z i j) + +RecHSet : (P : A → TypeOfHLevel ℓ 2) → 3-Constant P → ∥ A ∥ → TypeOfHLevel ℓ 2 +RecHSet P 3kP = rec→Gpd (isOfHLevelTypeOfHLevel 2) P 3kP + +∥∥-IdempotentL-⊎-≃ : ∥ ∥ A ∥ ⊎ A′ ∥ ≃ ∥ A ⊎ A′ ∥ +∥∥-IdempotentL-⊎-≃ = isoToEquiv ∥∥-IdempotentL-⊎-Iso + where ∥∥-IdempotentL-⊎-Iso : Iso (∥ ∥ A ∥ ⊎ A′ ∥) (∥ A ⊎ A′ ∥) + Iso.fun ∥∥-IdempotentL-⊎-Iso x = rec squash lem x + where lem : ∥ A ∥ ⊎ A′ → ∥ A ⊎ A′ ∥ + lem (inl x) = map (λ a → inl a) x + lem (inr x) = ∣ inr x ∣ + Iso.inv ∥∥-IdempotentL-⊎-Iso x = map lem x + where lem : A ⊎ A′ → ∥ A ∥ ⊎ A′ + lem (inl x) = inl ∣ x ∣ + lem (inr x) = inr x + Iso.rightInv ∥∥-IdempotentL-⊎-Iso x = squash (Iso.fun ∥∥-IdempotentL-⊎-Iso (Iso.inv ∥∥-IdempotentL-⊎-Iso x)) x + Iso.leftInv ∥∥-IdempotentL-⊎-Iso x = squash (Iso.inv ∥∥-IdempotentL-⊎-Iso (Iso.fun ∥∥-IdempotentL-⊎-Iso x)) x + +∥∥-IdempotentL-⊎ : ∥ ∥ A ∥ ⊎ A′ ∥ ≡ ∥ A ⊎ A′ ∥ +∥∥-IdempotentL-⊎ = ua ∥∥-IdempotentL-⊎-≃ + +∥∥-IdempotentR-⊎-≃ : ∥ A ⊎ ∥ A′ ∥ ∥ ≃ ∥ A ⊎ A′ ∥ +∥∥-IdempotentR-⊎-≃ = isoToEquiv ∥∥-IdempotentR-⊎-Iso + where ∥∥-IdempotentR-⊎-Iso : Iso (∥ A ⊎ ∥ A′ ∥ ∥) (∥ A ⊎ A′ ∥) + Iso.fun ∥∥-IdempotentR-⊎-Iso x = rec squash lem x + where lem : A ⊎ ∥ A′ ∥ → ∥ A ⊎ A′ ∥ + lem (inl x) = ∣ inl x ∣ + lem (inr x) = map (λ a → inr a) x + Iso.inv ∥∥-IdempotentR-⊎-Iso x = map lem x + where lem : A ⊎ A′ → A ⊎ ∥ A′ ∥ + lem (inl x) = inl x + lem (inr x) = inr ∣ x ∣ + Iso.rightInv ∥∥-IdempotentR-⊎-Iso x = squash (Iso.fun ∥∥-IdempotentR-⊎-Iso (Iso.inv ∥∥-IdempotentR-⊎-Iso x)) x + Iso.leftInv ∥∥-IdempotentR-⊎-Iso x = squash (Iso.inv ∥∥-IdempotentR-⊎-Iso (Iso.fun ∥∥-IdempotentR-⊎-Iso x)) x + +∥∥-IdempotentR-⊎ : ∥ A ⊎ ∥ A′ ∥ ∥ ≡ ∥ A ⊎ A′ ∥ +∥∥-IdempotentR-⊎ = ua ∥∥-IdempotentR-⊎-≃ + +∥∥-Idempotent-⊎ : {A : Type ℓ} {A′ : Type ℓ′} → ∥ ∥ A ∥ ⊎ ∥ A′ ∥ ∥ ≡ ∥ A ⊎ A′ ∥ +∥∥-Idempotent-⊎ {A = A} {A′} = ∥ ∥ A ∥ ⊎ ∥ A′ ∥ ∥ ≡⟨ ∥∥-IdempotentR-⊎ ⟩ + ∥ ∥ A ∥ ⊎ A′ ∥ ≡⟨ ∥∥-IdempotentL-⊎ ⟩ + ∥ A ⊎ A′ ∥ ∎ + +∥∥-IdempotentL-×-≃ : ∥ ∥ A ∥ × A′ ∥ ≃ ∥ A × A′ ∥ +∥∥-IdempotentL-×-≃ = isoToEquiv ∥∥-IdempotentL-×-Iso + where ∥∥-IdempotentL-×-Iso : Iso (∥ ∥ A ∥ × A′ ∥) (∥ A × A′ ∥) + Iso.fun ∥∥-IdempotentL-×-Iso x = rec squash lem x + where lem : ∥ A ∥ × A′ → ∥ A × A′ ∥ + lem (a , a′) = map2 (λ a a′ → a , a′) a ∣ a′ ∣ + Iso.inv ∥∥-IdempotentL-×-Iso x = map lem x + where lem : A × A′ → ∥ A ∥ × A′ + lem (a , a′) = ∣ a ∣ , a′ + Iso.rightInv ∥∥-IdempotentL-×-Iso x = squash (Iso.fun ∥∥-IdempotentL-×-Iso (Iso.inv ∥∥-IdempotentL-×-Iso x)) x + Iso.leftInv ∥∥-IdempotentL-×-Iso x = squash (Iso.inv ∥∥-IdempotentL-×-Iso (Iso.fun ∥∥-IdempotentL-×-Iso x)) x + +∥∥-IdempotentL-× : ∥ ∥ A ∥ × A′ ∥ ≡ ∥ A × A′ ∥ +∥∥-IdempotentL-× = ua ∥∥-IdempotentL-×-≃ + +∥∥-IdempotentR-×-≃ : ∥ A × ∥ A′ ∥ ∥ ≃ ∥ A × A′ ∥ +∥∥-IdempotentR-×-≃ = isoToEquiv ∥∥-IdempotentR-×-Iso + where ∥∥-IdempotentR-×-Iso : Iso (∥ A × ∥ A′ ∥ ∥) (∥ A × A′ ∥) + Iso.fun ∥∥-IdempotentR-×-Iso x = rec squash lem x + where lem : A × ∥ A′ ∥ → ∥ A × A′ ∥ + lem (a , a′) = map2 (λ a a′ → a , a′) ∣ a ∣ a′ + Iso.inv ∥∥-IdempotentR-×-Iso x = map lem x + where lem : A × A′ → A × ∥ A′ ∥ + lem (a , a′) = a , ∣ a′ ∣ + Iso.rightInv ∥∥-IdempotentR-×-Iso x = squash (Iso.fun ∥∥-IdempotentR-×-Iso (Iso.inv ∥∥-IdempotentR-×-Iso x)) x + Iso.leftInv ∥∥-IdempotentR-×-Iso x = squash (Iso.inv ∥∥-IdempotentR-×-Iso (Iso.fun ∥∥-IdempotentR-×-Iso x)) x + +∥∥-IdempotentR-× : ∥ A × ∥ A′ ∥ ∥ ≡ ∥ A × A′ ∥ +∥∥-IdempotentR-× = ua ∥∥-IdempotentR-×-≃ + +∥∥-Idempotent-× : {A : Type ℓ} {A′ : Type ℓ′} → ∥ ∥ A ∥ × ∥ A′ ∥ ∥ ≡ ∥ A × A′ ∥ +∥∥-Idempotent-× {A = A} {A′} = ∥ ∥ A ∥ × ∥ A′ ∥ ∥ ≡⟨ ∥∥-IdempotentR-× ⟩ + ∥ ∥ A ∥ × A′ ∥ ≡⟨ ∥∥-IdempotentL-× ⟩ + ∥ A × A′ ∥ ∎ + +∥∥-Idempotent-×-≃ : {A : Type ℓ} {A′ : Type ℓ′} → ∥ ∥ A ∥ × ∥ A′ ∥ ∥ ≃ ∥ A × A′ ∥ +∥∥-Idempotent-×-≃ {A = A} {A′} = compEquiv ∥∥-IdempotentR-×-≃ ∥∥-IdempotentL-×-≃ + +∥∥-×-≃ : {A : Type ℓ} {A′ : Type ℓ′} → ∥ A ∥ × ∥ A′ ∥ ≃ ∥ A × A′ ∥ +∥∥-×-≃ {A = A} {A′} = compEquiv (invEquiv (propTruncIdempotent≃ (isProp× isPropPropTrunc isPropPropTrunc))) ∥∥-Idempotent-×-≃ + +∥∥-× : {A : Type ℓ} {A′ : Type ℓ′} → ∥ A ∥ × ∥ A′ ∥ ≡ ∥ A × A′ ∥ +∥∥-× = ua ∥∥-×-≃ + +-- using this we get a convenient recursor/eliminator for binary functions into sets +rec2→Set : {A B C : Type ℓ} (Cset : isSet C) + → (f : A → B → C) + → (∀ (a a' : A) (b b' : B) → f a b ≡ f a' b') + → ∥ A ∥ → ∥ B ∥ → C +rec2→Set {A = A} {B = B} {C = C} Cset f fconst = curry (g ∘ ∥∥-×-≃ .fst) + where + g : ∥ A × B ∥ → C + g = rec→Set Cset (uncurry f) λ x y → fconst (fst x) (fst y) (snd x) (snd y) diff --git a/Cubical/HITs/Pushout.agda b/Cubical/HITs/Pushout.agda new file mode 100644 index 0000000000..e1b04e6505 --- /dev/null +++ b/Cubical/HITs/Pushout.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Pushout where + +open import Cubical.HITs.Pushout.Base public +open import Cubical.HITs.Pushout.Properties public +open import Cubical.HITs.Pushout.KrausVonRaumer public diff --git a/Cubical/HITs/Pushout/Base.agda b/Cubical/HITs/Pushout/Base.agda new file mode 100644 index 0000000000..5a26c78e7f --- /dev/null +++ b/Cubical/HITs/Pushout/Base.agda @@ -0,0 +1,117 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Pushout.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.HITs.Susp.Base + +data Pushout {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : A → B) (g : A → C) : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) where + inl : B → Pushout f g + inr : C → Pushout f g + push : (a : A) → inl (f a) ≡ inr (g a) + +-- cofiber (equivalent to Cone in Cubical.HITs.MappingCones.Base) +cofib : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type _ +cofib f = Pushout (λ _ → tt) f + +cfcod : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → B → cofib f +cfcod f = inr + +-- Suspension defined as a pushout + +PushoutSusp : ∀ {ℓ} (A : Type ℓ) → Type ℓ +PushoutSusp A = Pushout {A = A} {B = Unit} {C = Unit} (λ _ → tt) (λ _ → tt) + +PushoutSusp→Susp : ∀ {ℓ} {A : Type ℓ} → PushoutSusp A → Susp A +PushoutSusp→Susp (inl _) = north +PushoutSusp→Susp (inr _) = south +PushoutSusp→Susp (push a i) = merid a i + +Susp→PushoutSusp : ∀ {ℓ} {A : Type ℓ} → Susp A → PushoutSusp A +Susp→PushoutSusp north = inl tt +Susp→PushoutSusp south = inr tt +Susp→PushoutSusp (merid a i) = push a i + +Susp→PushoutSusp→Susp : ∀ {ℓ} {A : Type ℓ} (x : Susp A) → + PushoutSusp→Susp (Susp→PushoutSusp x) ≡ x +Susp→PushoutSusp→Susp north = refl +Susp→PushoutSusp→Susp south = refl +Susp→PushoutSusp→Susp (merid _ _) = refl + +PushoutSusp→Susp→PushoutSusp : ∀ {ℓ} {A : Type ℓ} (x : PushoutSusp A) → + Susp→PushoutSusp (PushoutSusp→Susp x) ≡ x +PushoutSusp→Susp→PushoutSusp (inl _) = refl +PushoutSusp→Susp→PushoutSusp (inr _) = refl +PushoutSusp→Susp→PushoutSusp (push _ _) = refl + +PushoutSuspIsoSusp : ∀ {ℓ} {A : Type ℓ} → Iso (PushoutSusp A) (Susp A) +Iso.fun PushoutSuspIsoSusp = PushoutSusp→Susp +Iso.inv PushoutSuspIsoSusp = Susp→PushoutSusp +Iso.rightInv PushoutSuspIsoSusp = Susp→PushoutSusp→Susp +Iso.leftInv PushoutSuspIsoSusp = PushoutSusp→Susp→PushoutSusp + + +PushoutSusp≃Susp : ∀ {ℓ} {A : Type ℓ} → PushoutSusp A ≃ Susp A +PushoutSusp≃Susp = isoToEquiv PushoutSuspIsoSusp + +PushoutSusp≡Susp : ∀ {ℓ} {A : Type ℓ} → PushoutSusp A ≡ Susp A +PushoutSusp≡Susp = isoToPath PushoutSuspIsoSusp + +-- Generalised pushout, used in e.g. BlakersMassey +data PushoutGen {ℓ₁ ℓ₂ ℓ₃ : Level} {X : Type ℓ₁} {Y : Type ℓ₂} + (Q : X → Y → Type ℓ₃) : Type (ℓ-max (ℓ-max ℓ₁ ℓ₂) ℓ₃) + where + inl : X → PushoutGen Q + inr : Y → PushoutGen Q + push : {x : X} {y : Y} → Q x y → inl x ≡ inr y + +-- The usual pushout is a special case of the above +module _ {ℓ₁ ℓ₂ ℓ₃ : Level} {A : Type ℓ₁} {B : Type ℓ₂} {C : Type ℓ₃} + (f : A → B) (g : A → C) where + open Iso + + doubleFib : B → C → Type _ + doubleFib b c = Σ[ a ∈ A ] (f a ≡ b) × (g a ≡ c) + + PushoutGenFib = PushoutGen doubleFib + + Pushout→PushoutGen : Pushout f g → PushoutGenFib + Pushout→PushoutGen (inl x) = inl x + Pushout→PushoutGen (inr x) = inr x + Pushout→PushoutGen (push a i) = push (a , refl , refl) i + + PushoutGen→Pushout : PushoutGenFib → Pushout f g + PushoutGen→Pushout (inl x) = inl x + PushoutGen→Pushout (inr x) = inr x + PushoutGen→Pushout (push (x , p , q) i) = + ((λ i → inl (p (~ i))) ∙∙ push x ∙∙ (λ i → inr (q i))) i + + IsoPushoutPushoutGen : Iso (Pushout f g) (PushoutGenFib) + fun IsoPushoutPushoutGen = Pushout→PushoutGen + inv IsoPushoutPushoutGen = PushoutGen→Pushout + rightInv IsoPushoutPushoutGen (inl x) = refl + rightInv IsoPushoutPushoutGen (inr x) = refl + rightInv IsoPushoutPushoutGen (push (x , p , q) i) j = lem x p q j i + where + lem : {b : B} {c : C} (x : A) (p : f x ≡ b) (q : g x ≡ c) + → cong Pushout→PushoutGen (cong PushoutGen→Pushout (push (x , p , q))) + ≡ push (x , p , q) + lem {c = c} x = + J (λ b p → (q : g x ≡ c) + → cong Pushout→PushoutGen + (cong PushoutGen→Pushout (push (x , p , q))) + ≡ push (x , p , q)) + (J (λ c q → cong Pushout→PushoutGen + (cong PushoutGen→Pushout (push (x , refl , q))) + ≡ push (x , refl , q)) + (cong (cong Pushout→PushoutGen) (sym (rUnit (push x))))) + leftInv IsoPushoutPushoutGen (inl x) = refl + leftInv IsoPushoutPushoutGen (inr x) = refl + leftInv IsoPushoutPushoutGen (push a i) j = rUnit (push a) (~ j) i diff --git a/Cubical/HITs/Pushout/Flattening.agda b/Cubical/HITs/Pushout/Flattening.agda new file mode 100644 index 0000000000..f9528296fd --- /dev/null +++ b/Cubical/HITs/Pushout/Flattening.agda @@ -0,0 +1,91 @@ +{- + + The flattening lemma for pushouts (Lemma 8.5.3 in the HoTT book) proved in a cubical style. + + The proof in the HoTT book (the core lying in Lemma 6.12.2, the flattening lemma for coequalizers) + consists mostly of long strings of equalities about transport. This proof follows almost + entirely from definitional equalities involving glue/unglue. + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.Pushout.Flattening where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma + +open import Cubical.HITs.Pushout.Base + +module FlatteningLemma {ℓa ℓb ℓc} {A : Type ℓa} {B : Type ℓb} {C : Type ℓc} (f : A → B) (g : A → C) + {ℓ} (F : B → Type ℓ) (G : C → Type ℓ) (e : ∀ a → F (f a) ≃ G (g a)) where + + E : Pushout f g → Type ℓ + E (inl x) = F x + E (inr x) = G x + E (push a i) = ua (e a) i + + Σf : Σ[ a ∈ A ] F (f a) → Σ[ b ∈ B ] F b + Σf (a , x) = (f a , x) + + Σg : Σ[ a ∈ A ] F (f a) → Σ[ c ∈ C ] G c + Σg (a , x) = (g a , (e a) .fst x) + + module FlattenIso where + + fwd : Pushout Σf Σg → Σ (Pushout f g) E + fwd (inl (b , x)) = (inl b , x) + fwd (inr (c , x)) = (inr c , x) + fwd (push (a , x) i) = (push a i , ua-gluePt (e a) i x) + + bwd : Σ (Pushout f g) E → Pushout Σf Σg + bwd (inl b , x) = inl (b , x) + bwd (inr c , x) = inr (c , x) + bwd (push a i , x) = hcomp (λ j → λ { (i = i0) → push (a , x) (~ j) + ; (i = i1) → inr (g a , x) }) + (inr (g a , ua-unglue (e a) i x)) + + bwd-fwd : ∀ x → bwd (fwd x) ≡ x + bwd-fwd (inl (b , x)) = refl + bwd-fwd (inr (c , x)) = refl + bwd-fwd (push (a , x) i) j = + hcomp (λ k → λ { (i = i0) → push (a , ua-gluePt (e a) i0 x) (~ k) + ; (i = i1) → inr (g a , ua-gluePt (e a) i1 x) + ; (j = i1) → push (a , x) (i ∨ ~ k) }) + (inr (g a , ua-unglue (e a) i (ua-gluePt (e a) i x))) + -- Note: the (j = i1) case typechecks because of the definitional equalities: + -- ua-gluePt e i0 x ≡ x , ua-gluePt e i1 x ≡ e .fst x, + -- ua-unglue-glue : ua-unglue e i (ua-gluePt e i x) ≡ e .fst x + + -- essentially: ua-glue e (i ∨ ~ k) ∘ ua-unglue e i + sq : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) + → SquareP (λ i k → ua e i → ua e (i ∨ ~ k)) + {- i = i0 -} (λ k x → ua-gluePt e (~ k) x) + {- i = i1 -} (λ k x → x) + {- k = i0 -} (λ i x → ua-unglue e i x) + {- k = i1 -} (λ i x → x) + sq e i k x = ua-glue e (i ∨ ~ k) (λ { ((i ∨ ~ k) = i0) → x }) + (inS (ua-unglue e i x)) + -- Note: this typechecks because of the definitional equalities: + -- ua-unglue e i0 x ≡ e .fst x, ua-glue e i1 _ (inS y) ≡ y, ua-unglue e i1 x ≡ x, + -- ua-glue-unglue : ua-glue e i (λ { (i = i0) → x }) (inS (ua-unglue e i x)) ≡ x + + fwd-bwd : ∀ x → fwd (bwd x) ≡ x + fwd-bwd (inl b , x) = refl + fwd-bwd (inr c , x) = refl + fwd-bwd (push a i , x) j = + -- `fwd` (or any function) takes hcomps to comps on a constant family, so we must use a comp here + comp (λ _ → Σ (Pushout f g) E) + (λ k → λ { (i = i0) → push a (~ k) , ua-gluePt (e a) (~ k) x + ; (i = i1) → inr (g a) , x + ; (j = i1) → push a (i ∨ ~ k) , sq (e a) i k x }) + (inr (g a) , ua-unglue (e a) i x) + + isom : Iso (Σ (Pushout f g) E) (Pushout Σf Σg) + isom = iso bwd fwd bwd-fwd fwd-bwd + + flatten : Σ (Pushout f g) E ≃ Pushout Σf Σg + flatten = isoToEquiv FlattenIso.isom diff --git a/Cubical/HITs/Pushout/KrausVonRaumer.agda b/Cubical/HITs/Pushout/KrausVonRaumer.agda new file mode 100644 index 0000000000..35f48f4557 --- /dev/null +++ b/Cubical/HITs/Pushout/KrausVonRaumer.agda @@ -0,0 +1,198 @@ +{- + +An induction principle for paths in a pushout, described in +Kraus and von Raumer, "Path Spaces of Higher Inductive Types in Homotopy Type Theory" +https://arxiv.org/abs/1901.06022 + +-} + +{-# OPTIONS --safe #-} +module Cubical.HITs.Pushout.KrausVonRaumer where + +open import Cubical.Foundations.Everything +open import Cubical.Functions.Embedding +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Sum as ⊎ +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.HITs.Pushout.Base as ⊔ +open import Cubical.HITs.Pushout.Properties + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + A : Type ℓ + B : Type ℓ' + C : Type ℓ'' + +private + interpolate : {x y z : A} (q : y ≡ z) + → PathP (λ i → x ≡ q i → x ≡ z) (_∙ q) (idfun _) + interpolate q i p j = + hcomp + (λ k → λ + { (j = i0) → p i0 + ; (j = i1) → q (i ∨ k) + ; (i = i1) → p j + }) + (p j) + + interpolateCompPath : {x y : A} (p : x ≡ y) {z : A} (q : y ≡ z) + → (λ i → interpolate q i (λ j → compPath-filler p q i j)) ≡ refl + interpolateCompPath p = + J (λ z q → (λ i → interpolate q i (λ j → compPath-filler p q i j)) ≡ refl) + (homotopySymInv (λ p i j → compPath-filler p refl (~ i) j) p) + +module ElimL + {f : A → B} {g : A → C} {b₀ : B} + (P : ∀ b → Path (Pushout f g) (inl b₀) (inl b) → Type ℓ''') + (Q : ∀ c → Path (Pushout f g) (inl b₀) (inr c) → Type ℓ''') + (r : P b₀ refl) + (e : (a : A) (q : inl b₀ ≡ inl (f a)) → P (f a) q ≃ Q (g a) (q ∙ push a)) + where + + Codes : (d : Pushout f g) (q : inl b₀ ≡ d) → Type ℓ''' + Codes (inl b) q = P b q + Codes (inr c) q = Q c q + Codes (push a i) q = + Glue + (Q (g a) (interpolate (push a) i q)) + (λ + { (i = i0) → _ , e a q + ; (i = i1) → _ , idEquiv (Q (g a) q) + }) + + elimL : ∀ b q → P b q + elimL _ = J Codes r + + elimR : ∀ c q → Q c q + elimR _ = J Codes r + + refl-β : elimL b₀ refl ≡ r + refl-β = transportRefl _ + + push-β : (a : A) (q : inl b₀ ≡ inl (f a)) + → elimR (g a) (q ∙ push a) ≡ e a q .fst (elimL (f a) q) + push-β a q = + J-∙ Codes r q (push a) + ∙ fromPathP + (subst + (λ α → PathP (λ i → Q (g a) (α i)) (e a q .fst (elimL (f a) q)) (e a q .fst (elimL (f a) q))) + (interpolateCompPath q (push a) ⁻¹) + refl) + +module ElimR + {f : A → B} {g : A → C} {c₀ : C} + (P : ∀ b → Path (Pushout f g) (inr c₀) (inl b) → Type ℓ''') + (Q : ∀ c → Path (Pushout f g) (inr c₀) (inr c) → Type ℓ''') + (r : Q c₀ refl) + (e : (a : A) (q : inr c₀ ≡ inl (f a)) → P (f a) q ≃ Q (g a) (q ∙ push a)) + where + + Codes : (d : Pushout f g) (q : inr c₀ ≡ d) → Type ℓ''' + Codes (inl b) q = P b q + Codes (inr c) q = Q c q + Codes (push a i) q = + Glue + (Q (g a) (interpolate (push a) i q)) + (λ + { (i = i0) → _ , e a q + ; (i = i1) → _ , idEquiv (Q (g a) q) + }) + + elimL : ∀ b q → P b q + elimL _ = J Codes r + + elimR : ∀ c q → Q c q + elimR _ = J Codes r + + refl-β : elimR c₀ refl ≡ r + refl-β = transportRefl _ + + push-β : (a : A) (q : inr c₀ ≡ inl (f a)) + → elimR (g a) (q ∙ push a) ≡ e a q .fst (elimL (f a) q) + push-β a q = + J-∙ Codes r q (push a) + ∙ fromPathP + (subst + (λ α → PathP (λ i → Q (g a) (α i)) (e a q .fst (elimL (f a) q)) (e a q .fst (elimL (f a) q))) + (interpolateCompPath q (push a) ⁻¹) + refl) + +-- Example application: pushouts preserve embeddings + +isEmbeddingInr : + (f : A → B) (g : A → C) + → isEmbedding f → isEmbedding (⊔.inr {f = f} {g = g}) +isEmbeddingInr f g fEmb c₀ c₁ = + isoToIsEquiv (iso _ (fst ∘ bwd c₁) (snd ∘ bwd c₁) bwdCong) + where + Q : ∀ c → ⊔.inr c₀ ≡ ⊔.inr c → Type _ + Q _ q = fiber (cong ⊔.inr) q + + P : ∀ b → ⊔.inr c₀ ≡ ⊔.inl b → Type _ + P b p = Σ[ u ∈ fiber f b ] Q _ (p ∙ cong ⊔.inl (u .snd ⁻¹) ∙ push (u .fst)) + + module Bwd = ElimR P Q + (refl , refl) + (λ a p → + subst + (P (f a) p ≃_) + (cong (λ w → fiber (cong ⊔.inr) (p ∙ w)) (lUnit (push a) ⁻¹)) + (Σ-contractFst (inhProp→isContr (a , refl) (isEmbedding→hasPropFibers fEmb (f a))))) + + bwd : ∀ c → (t : ⊔.inr c₀ ≡ ⊔.inr c) → fiber (cong ⊔.inr) t + bwd = Bwd.elimR + + bwdCong : ∀ {c} → (r : c₀ ≡ c) → bwd c (cong ⊔.inr r) .fst ≡ r + bwdCong = J (λ c r → bwd c (cong ⊔.inr r) .fst ≡ r) (cong fst Bwd.refl-β) + + +-- Further Application: Pushouts of emedding-spans of n-Types are n-Types, for n≥0 +module _ (f : A → B) (g : A → C) where + inlrJointlySurjective : + (z : Pushout f g) → ∥ Σ[ x ∈ (B ⊎ C) ] (⊎.rec inl inr x) ≡ z ∥ + inlrJointlySurjective = + elimProp _ + (λ _ → isPropPropTrunc) + (λ b → ∣ ⊎.inl b , refl ∣) + (λ c → ∣ ⊎.inr c , refl ∣) + + preserveHLevelEmbedding : + {n : HLevel} + → isEmbedding f + → isEmbedding g + → isOfHLevel (2 + n) B + → isOfHLevel (2 + n) C + → isOfHLevel (2 + n) (Pushout f g) + preserveHLevelEmbedding {n = n} fEmb gEmb isOfHLB isOfHLC = + isOfHLevelΩ→isOfHLevel n ΩHLevelPushout + where isEmbInr = isEmbeddingInr f g fEmb + isEmbInrSwitched = isEmbeddingInr g f gEmb + + equivΩC : {x : Pushout f g} (c : C) (p : inr c ≡ x) + → (c ≡ c) ≃ (x ≡ x) + equivΩC c p = compEquiv (_ , isEmbInr c c) (conjugatePathEquiv p) + + equivΩB : {x : Pushout f g} (b : B) (p : inl b ≡ x) + → (b ≡ b) ≃ (x ≡ x) + equivΩB b p = compEquiv + (compEquiv (_ , isEmbInrSwitched b b) + (congEquiv pushoutSwitchEquiv)) + (conjugatePathEquiv p) + + ΩHLevelPushout : (x : Pushout f g) → isOfHLevel (suc n) (x ≡ x) + ΩHLevelPushout x = + Trunc.elim + (λ _ → isPropIsOfHLevel {A = (x ≡ x)} (suc n)) + (λ {(⊎.inl b , p) → + isOfHLevelRespectEquiv + (suc n) + (equivΩB b p) + (isOfHLB b b); + (⊎.inr c , p) → + isOfHLevelRespectEquiv + (suc n) + (equivΩC c p) + (isOfHLC c c)}) + (inlrJointlySurjective x) diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda new file mode 100644 index 0000000000..61fccbabc9 --- /dev/null +++ b/Cubical/HITs/Pushout/Properties.agda @@ -0,0 +1,543 @@ +{- + +This file contains: +- Elimination principle for pushouts + +- Homotopy natural equivalences of pushout spans + Written by: Loïc Pujet, September 2019 + +- 3×3 lemma for pushouts + Written by: Loïc Pujet, April 2019 + +- Homotopy natural equivalences of pushout spans + (unpacked and avoiding transports) + +-} + +{-# OPTIONS --safe #-} + +module Cubical.HITs.Pushout.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.HITs.Pushout.Base + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + A : Type ℓ + B : Type ℓ' + C : Type ℓ'' + +{- + Elimination for propositions +-} +elimProp : {f : A → B} {g : A → C} + → (P : Pushout f g → Type ℓ''') + → ((x : Pushout f g) → isProp (P x)) + → ((b : B) → P (inl b)) + → ((c : C) → P (inr c)) + → (x : Pushout f g) → P x +elimProp P isPropP PB PC (inl x) = PB x +elimProp P isPropP PB PC (inr x) = PC x +elimProp {f = f} {g = g} P isPropP PB PC (push a i) = + isOfHLevel→isOfHLevelDep 1 isPropP (PB (f a)) (PC (g a)) (push a) i + + +{- + Switching the span does not change the pushout +-} +pushoutSwitchEquiv : {f : A → B} {g : A → C} + → Pushout f g ≃ Pushout g f +pushoutSwitchEquiv = isoToEquiv (iso f inv leftInv rightInv) + where f = λ {(inr x) → inl x; (inl x) → inr x; (push a i) → push a (~ i)} + inv = λ {(inr x) → inl x; (inl x) → inr x; (push a i) → push a (~ i)} + leftInv = λ {(inl x) → refl; (inr x) → refl; (push a i) → refl} + rightInv = λ {(inl x) → refl; (inr x) → refl; (push a i) → refl} + +{- + Definition of pushout diagrams +-} + +record 3-span : Type₁ where + field + A0 A2 A4 : Type₀ + f1 : A2 → A0 + f3 : A2 → A4 + +3span : {A0 A2 A4 : Type₀} → (A2 → A0) → (A2 → A4) → 3-span +3span f1 f3 = record { f1 = f1 ; f3 = f3 } + +spanPushout : (s : 3-span) → Type₀ +spanPushout s = Pushout (3-span.f1 s) (3-span.f3 s) + +{- + Definition of a homotopy natural diagram equivalence +-} + +record 3-span-equiv (s1 : 3-span) (s2 : 3-span) : Type₀ where + field + e0 : 3-span.A0 s1 ≃ 3-span.A0 s2 + e2 : 3-span.A2 s1 ≃ 3-span.A2 s2 + e4 : 3-span.A4 s1 ≃ 3-span.A4 s2 + H1 : ∀ x → 3-span.f1 s2 (e2 .fst x) ≡ e0 .fst (3-span.f1 s1 x) + H3 : ∀ x → 3-span.f3 s2 (e2 .fst x) ≡ e4 .fst (3-span.f3 s1 x) + +{- + Proof that homotopy equivalent spans are in fact equal +-} +spanEquivToPath : {s1 : 3-span} → {s2 : 3-span} → (e : 3-span-equiv s1 s2) → s1 ≡ s2 +spanEquivToPath {s1} {s2} e = spanPath + where + open 3-span-equiv e + open 3-span + + path0 : A0 s1 ≡ A0 s2 + path0 = ua e0 + + path2 : A2 s1 ≡ A2 s2 + path2 = ua e2 + + path4 : A4 s1 ≡ A4 s2 + path4 = ua e4 + + spanPath1 : I → 3-span + spanPath1 i = record { A0 = path0 i ; A2 = path2 i ; A4 = path4 i ; + f1 = λ x → (transp (λ j → path0 (i ∧ j)) (~ i) (f1 s1 (transp (λ j → path2 (~ j ∧ i)) (~ i) x))) ; + f3 = λ x → (transp (λ j → path4 (i ∧ j)) (~ i) (f3 s1 (transp (λ j → path2 (~ j ∧ i)) (~ i) x))) } + + spanPath2 : I → 3-span + spanPath2 i = record { A0 = A0 s2 ; A2 = A2 s2 ; A4 = A4 s2 ; f1 = f1Path i ; f3 = f3Path i } + where + f1Path : I → A2 s2 → A0 s2 + f1Path i x = ((uaβ e0 (f1 s1 (transport (λ j → path2 (~ j)) x))) + ∙ (H1 (transport (λ j → path2 (~ j)) x)) ⁻¹ + ∙ (λ j → f1 s2 (uaβ e2 (transport (λ j → path2 (~ j)) x) (~ j))) + ∙ (λ j → f1 s2 (transportTransport⁻ path2 x j))) i + + f3Path : I → A2 s2 → A4 s2 + f3Path i x = ((uaβ e4 (f3 s1 (transport (λ j → path2 (~ j)) x))) + ∙ (H3 (transport (λ j → path2 (~ j)) x)) ⁻¹ + ∙ (λ j → f3 s2 (uaβ e2 (transport (λ j → path2 (~ j)) x) (~ j))) + ∙ (λ j → f3 s2 (transportTransport⁻ path2 x j))) i + + spanPath : s1 ≡ s2 + spanPath = (λ i → spanPath1 i) ∙ (λ i → spanPath2 i) + +-- as a corollary, they have the same pushout +spanEquivToPushoutPath : {s1 : 3-span} → {s2 : 3-span} → (e : 3-span-equiv s1 s2) + → spanPushout s1 ≡ spanPushout s2 +spanEquivToPushoutPath {s1} {s2} e = cong spanPushout (spanEquivToPath e) + +{- + + 3×3 lemma for pushouts + +Let Aᵢⱼ denote a type, fᵢⱼ a map and Hᵢⱼ a homotopy. Given a diagram of the following shape: + +A00 ←—f01—— A02 ——f03—→ A04 + ↑ ↑ ↑ +f10 H11 f12 H13 f14 + | | | +A20 ←—f21—— A22 ——f23—→ A24 + | | | +f30 H31 f32 H33 f34 + ↓ ↓ ↓ +A40 ←—f41—— A42 ——f43—→ A44 + +one can build its colimit from pushouts in two ways: +- either build pushouts A□0, A□2, A□4 of the lines and then build the pushout A□○ of the resulting + diagram A□0 ←—f□1—— A□2 ——f□3—→ A□4 ; +- or build pushouts of the columns and build the pushout A○□ of the resulting diagram A0□ ← A2□ → A4□. + +Then the lemma states there is an equivalence A□○ ≃ A○□. + +-} + +record 3x3-span : Type₁ where + field + A00 A02 A04 A20 A22 A24 A40 A42 A44 : Type₀ + + f10 : A20 → A00 + f12 : A22 → A02 + f14 : A24 → A04 + + f30 : A20 → A40 + f32 : A22 → A42 + f34 : A24 → A44 + + f01 : A02 → A00 + f21 : A22 → A20 + f41 : A42 → A40 + + f03 : A02 → A04 + f23 : A22 → A24 + f43 : A42 → A44 + + H11 : ∀ x → f01 (f12 x) ≡ f10 (f21 x) + H13 : ∀ x → f03 (f12 x) ≡ f14 (f23 x) + H31 : ∀ x → f41 (f32 x) ≡ f30 (f21 x) + H33 : ∀ x → f43 (f32 x) ≡ f34 (f23 x) + + -- pushouts of the lines + A□0 : Type₀ + A□0 = Pushout f10 f30 + + A□2 : Type₀ + A□2 = Pushout f12 f32 + + A□4 : Type₀ + A□4 = Pushout f14 f34 + + -- maps between pushouts + f□1 : A□2 → A□0 + f□1 (inl x) = inl (f01 x) + f□1 (inr x) = inr (f41 x) + f□1 (push a j) = ((λ i → inl (H11 a i)) + ∙∙ push (f21 a) + ∙∙ (λ i → inr (H31 a (~ i)))) j + + f□3 : A□2 → A□4 + f□3 (inl x) = inl (f03 x) + f□3 (inr x) = inr (f43 x) + f□3 (push a j) = ((λ i → inl (H13 a i)) + ∙∙ push (f23 a) + ∙∙ (λ i → inr (H33 a (~ i)))) j + + -- total pushout + A□○ : Type₀ + A□○ = Pushout f□1 f□3 + + -- pushouts of the columns + A0□ : Type₀ + A0□ = Pushout f01 f03 + + A2□ : Type₀ + A2□ = Pushout f21 f23 + + A4□ : Type₀ + A4□ = Pushout f41 f43 + + -- maps between pushouts + f1□ : A2□ → A0□ + f1□ (inl x) = inl (f10 x) + f1□ (inr x) = inr (f14 x) + f1□ (push a j) = ((λ i → inl (H11 a (~ i))) + ∙∙ push (f12 a) + ∙∙ (λ i → inr (H13 a i))) j + + f3□ : A2□ → A4□ + f3□ (inl x) = inl (f30 x) + f3□ (inr x) = inr (f34 x) + f3□ (push a j) = ((λ i → inl (H31 a (~ i))) + ∙∙ push (f32 a) + ∙∙ (λ i → inr (H33 a i))) j + + -- total pushout + A○□ : Type₀ + A○□ = Pushout f1□ f3□ + + -- forward map of the equivalence A□○ ≃ A○□ + forward-l : A□0 → A○□ + forward-l (inl x) = inl (inl x) + forward-l (inr x) = inr (inl x) + forward-l (push a i) = push (inl a) i + + forward-r : A□4 → A○□ + forward-r (inl x) = inl (inr x) + forward-r (inr x) = inr (inr x) + forward-r (push a i) = push (inr a) i + + forward-filler : A22 → I → I → I → A○□ + forward-filler a i j = hfill (λ t → λ { (i = i0) → inl (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) (~ t) j) + ; (i = i1) → inr (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) (~ t) j) + ; (j = i0) → forward-l (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) t i) + ; (j = i1) → forward-r (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) t i) }) + (inS (push (push a j) i)) + + A□○→A○□ : A□○ → A○□ + A□○→A○□ (inl x) = forward-l x + A□○→A○□ (inr x) = forward-r x + A□○→A○□ (push (inl x) i) = inl (push x i) + A□○→A○□ (push (inr x) i) = inr (push x i) + A□○→A○□ (push (push a i) j) = forward-filler a i j i1 + + -- backward map + backward-l : A0□ → A□○ + backward-l (inl x) = inl (inl x) + backward-l (inr x) = inr (inl x) + backward-l (push a i) = push (inl a) i + + backward-r : A4□ → A□○ + backward-r (inl x) = inl (inr x) + backward-r (inr x) = inr (inr x) + backward-r (push a i) = push (inr a) i + + backward-filler : A22 → I → I → I → A□○ + backward-filler a i j = hfill (λ t → λ { (i = i0) → inl (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) (~ t) j) + ; (i = i1) → inr (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) (~ t) j) + ; (j = i0) → backward-l (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) t i) + ; (j = i1) → backward-r (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) t i) }) + (inS (push (push a j) i)) + + A○□→A□○ : A○□ → A□○ + A○□→A□○ (inl x) = backward-l x + A○□→A□○ (inr x) = backward-r x + A○□→A□○ (push (inl x) i) = inl (push x i) + A○□→A□○ (push (inr x) i) = inr (push x i) + A○□→A□○ (push (push a i) j) = backward-filler a i j i1 + + -- first homotopy + homotopy1-l : ∀ x → A□○→A○□ (backward-l x) ≡ inl x + homotopy1-l (inl x) = refl + homotopy1-l (inr x) = refl + homotopy1-l (push a i) = refl + + homotopy1-r : ∀ x → A□○→A○□ (backward-r x) ≡ inr x + homotopy1-r (inl x) = refl + homotopy1-r (inr x) = refl + homotopy1-r (push a i) = refl + + A○□→A□○→A○□ : ∀ x → A□○→A○□ (A○□→A□○ x) ≡ x + A○□→A□○→A○□ (inl x) = homotopy1-l x + A○□→A□○→A○□ (inr x) = homotopy1-r x + A○□→A□○→A○□ (push (inl x) i) k = push (inl x) i + A○□→A□○→A○□ (push (inr x) i) k = push (inr x) i + A○□→A□○→A○□ (push (push a i) j) k = + hcomp (λ t → λ { (i = i0) → forward-l (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) (~ t) j) + ; (i = i1) → forward-r (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) (~ t) j) + ; (j = i0) → homotopy1-l (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) t i) k + ; (j = i1) → homotopy1-r (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) t i) k + ; (k = i0) → A□○→A○□ (backward-filler a i j t) + ; (k = i1) → forward-filler a j i (~ t) }) + (forward-filler a j i i1) + + -- second homotopy + homotopy2-l : ∀ x → A○□→A□○ (forward-l x) ≡ inl x + homotopy2-l (inl x) = refl + homotopy2-l (inr x) = refl + homotopy2-l (push a i) = refl + + homotopy2-r : ∀ x → A○□→A□○ (forward-r x) ≡ inr x + homotopy2-r (inl x) = refl + homotopy2-r (inr x) = refl + homotopy2-r (push a i) = refl + + A□○→A○□→A□○ : ∀ x → A○□→A□○ (A□○→A○□ x) ≡ x + A□○→A○□→A□○ (inl x) = homotopy2-l x + A□○→A○□→A□○ (inr x) = homotopy2-r x + A□○→A○□→A□○ (push (inl x) i) k = push (inl x) i + A□○→A○□→A□○ (push (inr x) i) k = push (inr x) i + A□○→A○□→A□○ (push (push a i) j) k = + hcomp (λ t → λ { (i = i0) → backward-l (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) (~ t) j) + ; (i = i1) → backward-r (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) (~ t) j) + ; (j = i0) → homotopy2-l (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) t i) k + ; (j = i1) → homotopy2-r (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) t i) k + ; (k = i0) → A○□→A□○ (forward-filler a i j t) + ; (k = i1) → backward-filler a j i (~ t) }) + (backward-filler a j i i1) + + 3x3-Iso : Iso A□○ A○□ + Iso.fun 3x3-Iso = A□○→A○□ + Iso.inv 3x3-Iso = A○□→A□○ + Iso.rightInv 3x3-Iso = A○□→A□○→A○□ + Iso.leftInv 3x3-Iso = A□○→A○□→A□○ + + -- the claimed result + 3x3-lemma : A□○ ≡ A○□ + 3x3-lemma = isoToPath 3x3-Iso + +PushoutToProp : {f : A → B} {g : A → C} + {D : Pushout f g → Type ℓ'''} + → ((x : Pushout f g) → isProp (D x)) + → ((a : B) → D (inl a)) + → ((c : C) → D (inr c)) + → (x : Pushout f g) → D x +PushoutToProp isset baseB baseC (inl x) = baseB x +PushoutToProp isset baseB baseC (inr x) = baseC x +PushoutToProp {f = f} {g = g} isset baseB baseC (push a i) = + isOfHLevel→isOfHLevelDep 1 isset (baseB (f a)) (baseC (g a)) (push a) i + +-- explicit characterisation of equivalences +-- Pushout f₁ g₁ ≃ Pushout f₂ g₂ avoiding +-- transports + +open Iso +private + module PushoutIso {ℓ : Level} {A₁ B₁ C₁ A₂ B₂ C₂ : Type ℓ} + (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) + (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) + (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) + (id1 : (fst B≃) ∘ f₁ ≡ f₂ ∘ (fst A≃)) + (id2 : (fst C≃) ∘ g₁ ≡ g₂ ∘ (fst A≃)) + where + F : Pushout f₁ g₁ → Pushout f₂ g₂ + F (inl x) = inl (fst B≃ x) + F (inr x) = inr (fst C≃ x) + F (push a i) = + ((λ i → inl (funExt⁻ id1 a i)) + ∙∙ push (fst A≃ a) + ∙∙ λ i → inr (sym (funExt⁻ id2 a) i)) i + + G : Pushout f₂ g₂ → Pushout f₁ g₁ + G (inl x) = inl (invEq B≃ x) + G (inr x) = inr (invEq C≃ x) + G (push a i) = + ((λ i → inl ((sym (cong (invEq B≃) (funExt⁻ id1 (invEq A≃ a) + ∙ cong f₂ (secEq A≃ a))) + ∙ retEq B≃ (f₁ (invEq A≃ a))) i)) + ∙∙ push (invEq A≃ a) + ∙∙ λ i → inr (((sym (retEq C≃ (g₁ (invEq A≃ a))) + ∙ (cong (invEq C≃) ((funExt⁻ id2 (invEq A≃ a))))) + ∙ cong (invEq C≃) (cong g₂ (secEq A≃ a))) i)) i + + + abbrType₁ : {ℓ : Level} {A₁ B₁ C₁ A₂ B₂ C₂ : Type ℓ} + (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) + → (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) + (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) + → (id1 : (fst B≃) ∘ f₁ ≡ f₂ ∘ (fst A≃)) + (id2 : (fst C≃) ∘ g₁ ≡ g₂ ∘ (fst A≃)) + → Type _ + abbrType₁ A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 = + ((x : _) → PushoutIso.F A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 + (PushoutIso.G A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 x) ≡ x) + × ((x : _) → PushoutIso.G A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 + (PushoutIso.F A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 x) ≡ x) + + abbrType : {ℓ : Level} {A₁ B₁ C₁ A₂ B₂ C₂ : Type ℓ} + (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) + → Type _ + abbrType {A₁ = A₁} {B₁ = B₁} {C₁ = C₁} {A₂ = A₂} {B₂ = B₂} {C₂ = C₂} + A≃ B≃ C≃ = + (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) + (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) + (id1 : (fst B≃) ∘ f₁ ≡ f₂ ∘ (fst A≃)) + (id2 : (fst C≃) ∘ g₁ ≡ g₂ ∘ (fst A≃)) + → abbrType₁ A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 + + F-G-cancel : {ℓ : Level} {A₁ B₁ C₁ A₂ B₂ C₂ : Type ℓ} + (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) + → abbrType A≃ B≃ C≃ + F-G-cancel {A₁ = A₁} {B₁ = B₁} {C₁ = C₁} {A₂ = A₂} {B₂ = B₂} {C₂ = C₂} = + EquivJ (λ A₁ A≃ → (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) → + abbrType A≃ B≃ C≃) + (EquivJ (λ B₁ B≃ → (C≃ : C₁ ≃ C₂) → + abbrType (idEquiv A₂) B≃ C≃) + (EquivJ (λ C₁ C≃ → abbrType (idEquiv A₂) (idEquiv B₂) C≃) + λ f₁ g₁ f₂ g₂ + → J (λ f₂ id1 → (id2 : g₁ ≡ g₂) + → abbrType₁ (idEquiv A₂) (idEquiv B₂) (idEquiv C₂) + f₁ g₁ f₂ g₂ id1 id2) + (J (λ g₂ id2 → abbrType₁ (idEquiv A₂) (idEquiv B₂) (idEquiv C₂) + f₁ g₁ f₁ g₂ refl id2) + (postJ f₁ g₁)))) + + where + postJ : (f₁ : A₂ → B₂) (g₁ : A₂ → C₂) + → abbrType₁ (idEquiv A₂) (idEquiv B₂) (idEquiv C₂) + f₁ g₁ f₁ g₁ refl refl + postJ f₁ g₁ = FF-GG , GG-FF + where + refl-lem : ∀ {ℓ} {A : Type ℓ} (x : A) + → (refl {x = x} ∙ refl) ∙ refl ≡ refl + refl-lem x = sym (rUnit _) ∙ sym (rUnit _) + + FF = PushoutIso.F (idEquiv A₂) (idEquiv B₂) (idEquiv C₂) + f₁ g₁ f₁ g₁ refl refl + GG = PushoutIso.G (idEquiv A₂) (idEquiv B₂) (idEquiv C₂) + f₁ g₁ f₁ g₁ refl refl + + FF-GG : (x : Pushout f₁ g₁) → FF (GG x) ≡ x + FF-GG (inl x) = refl + FF-GG (inr x) = refl + FF-GG (push a i) j = lem j i + where + lem : Path (Path (Pushout f₁ g₁) (inl (f₁ a)) (inr (g₁ a))) + (cong FF ((λ i → inl (((refl ∙ refl) ∙ (refl {x = f₁ a})) i )) + ∙∙ push {f = f₁} {g = g₁} a + ∙∙ λ i → inr (((refl ∙ refl) ∙ (refl {x = g₁ a})) i))) + (push a) + lem = (λ i → cong FF ((λ j → inl (refl-lem (f₁ a) i j)) + ∙∙ push a + ∙∙ λ j → inr (refl-lem (g₁ a) i j))) + ∙∙ cong (cong FF) (sym (rUnit (push a))) + ∙∙ sym (rUnit (push a)) + + GG-FF : (x : _) → GG (FF x) ≡ x + GG-FF (inl x) = refl + GG-FF (inr x) = refl + GG-FF (push a i) j = lem j i + where + lem : cong GG (refl ∙∙ push a ∙∙ refl) ≡ push a + lem = cong (cong GG) (sym (rUnit (push a))) + ∙∙ (λ i → ((λ j → inl (refl-lem (f₁ a) i j)) + ∙∙ push a + ∙∙ λ j → inr (refl-lem (g₁ a) i j))) + ∙∙ sym (rUnit (push a)) + + +module _ {ℓ : Level} {A₁ B₁ C₁ A₂ B₂ C₂ : Type ℓ} + (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) + (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) + (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) + (id1 : fst B≃ ∘ f₁ ≡ f₂ ∘ fst A≃) + (id2 : fst C≃ ∘ g₁ ≡ g₂ ∘ fst A≃) + where + private + module P = PushoutIso A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 + l-r-cancel = F-G-cancel A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 + + pushoutIso : Iso (Pushout f₁ g₁) (Pushout f₂ g₂) + fun pushoutIso = P.F + inv pushoutIso = P.G + rightInv pushoutIso = fst l-r-cancel + leftInv pushoutIso = snd l-r-cancel + + pushoutEquiv : Pushout f₁ g₁ ≃ Pushout f₂ g₂ + pushoutEquiv = isoToEquiv pushoutIso + + +open import Cubical.Data.Unit +open import Cubical.Foundations.Pointed + +module PushoutDistr {ℓ ℓ' ℓ'' ℓ''' : Level} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + (f : B → A) (g : C → B) (h : C → D) where + PushoutDistrFun : Pushout {C = Pushout g h} f inl → Pushout (f ∘ g) h + PushoutDistrFun (inl x) = inl x + PushoutDistrFun (inr (inl x)) = inl (f x) + PushoutDistrFun (inr (inr x)) = inr x + PushoutDistrFun (inr (push a i)) = push a i + PushoutDistrFun (push a i) = inl (f a) + + PushoutDistrInv : Pushout (f ∘ g) h → Pushout {C = Pushout g h} f inl + PushoutDistrInv (inl x) = inl x + PushoutDistrInv (inr x) = inr (inr x) + PushoutDistrInv (push c i) = (push (g c) ∙ λ j → inr (push c j)) i + + PushoutDistrIso : Iso (Pushout {C = Pushout g h} f inl) (Pushout (f ∘ g) h) + fun PushoutDistrIso = PushoutDistrFun + inv PushoutDistrIso = PushoutDistrInv + rightInv PushoutDistrIso (inl x) = refl + rightInv PushoutDistrIso (inr x) = refl + rightInv PushoutDistrIso (push a i) j = + (cong-∙ (fun PushoutDistrIso) (push (g a)) (λ j → inr (push a j)) + ∙ sym (lUnit _)) j i + leftInv PushoutDistrIso (inl x) = refl + leftInv PushoutDistrIso (inr (inl x)) = push x + leftInv PushoutDistrIso (inr (inr x)) = refl + leftInv PushoutDistrIso (inr (push a i)) j = + compPath-filler' (push (g a)) (λ j → inr (push a j)) (~ j) i + leftInv PushoutDistrIso (push a i) j = push a (i ∧ j) diff --git a/Cubical/HITs/RPn.agda b/Cubical/HITs/RPn.agda new file mode 100644 index 0000000000..019b3ee2a8 --- /dev/null +++ b/Cubical/HITs/RPn.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.RPn where + +open import Cubical.HITs.RPn.Base public diff --git a/Cubical/HITs/RPn/Base.agda b/Cubical/HITs/RPn/Base.agda new file mode 100644 index 0000000000..f3df232e61 --- /dev/null +++ b/Cubical/HITs/RPn/Base.agda @@ -0,0 +1,320 @@ +{- + + A defintion of the real projective spaces following: + + [BR17] U. Buchholtz, E. Rijke, The real projective spaces in homotopy type theory. + (2017) https://arxiv.org/abs/1704.05770 + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.RPn.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Functions.Bundle + +open import Cubical.Foundations.SIP +open import Cubical.Structures.Pointed +open import Cubical.Structures.TypeEqvTo + +open import Cubical.Data.Bool +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.NatMinusOne +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ hiding (elim) +open import Cubical.Data.Sum as ⊎ hiding (elim) + +open import Cubical.HITs.PropositionalTruncation as PropTrunc hiding (elim) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Join +open import Cubical.HITs.Pushout +open import Cubical.HITs.Pushout.Flattening + +private + variable + ℓ ℓ' ℓ'' : Level + +-- PR² as a HIT +data RP² : Type₀ where + point : RP² + line : point ≡ point + square : line ≡ sym line + +-- Definition II.1 in [BR17], see also Cubical.Functions.Bundle + +2-EltType₀ = TypeEqvTo ℓ-zero Bool -- Σ[ X ∈ Type₀ ] ∥ X ≃ Bool ∥ +2-EltPointed₀ = PointedEqvTo ℓ-zero Bool -- Σ[ X ∈ Type₀ ] X × ∥ X ≃ Bool ∥ + +Bool* : 2-EltType₀ +Bool* = Bool , ∣ idEquiv _ ∣ + + +-- Our first goal is to 'lift' `_⊕_ : Bool → Bool ≃ Bool` to a function `_⊕_ : A → A ≃ Bool` +-- for any 2-element type (A, ∣e∣). + +-- `isContrBoolPointedEquiv` and `isContr-2-EltPointedEquiv` are contained in the proof +-- of Lemma II.2 in [BR17], though we prove `isContr-BoolPointedEquiv` more directly +-- with ⊕ -- [BR17] proves it for just the x = false case and uses notEquiv to get +-- the x = true case. + +-- (λ y → x ⊕ y) is the unqiue pointed isomorphism (Bool , false) ≃ (Bool , x) +isContrBoolPointedEquiv : ∀ x → isContr ((Bool , false) ≃[ PointedEquivStr ] (Bool , x)) +fst (isContrBoolPointedEquiv x) = ((λ y → x ⊕ y) , isEquiv-⊕ x) , ⊕-comm x false +snd (isContrBoolPointedEquiv x) (e , p) + = Σ≡Prop (λ e → isSetBool (equivFun e false) x) + (Σ≡Prop isPropIsEquiv (funExt λ { false → ⊕-comm x false ∙ sym p + ; true → ⊕-comm x true ∙ sym q })) + where q : e .fst true ≡ not x + q with dichotomyBool (invEq e (not x)) + ... | inl q = invEq≡→equivFun≡ e q + ... | inr q = ⊥.rec (not≢const x (sym (invEq≡→equivFun≡ e q) ∙ p)) + +-- Since isContr is a mere proposition, we can eliminate a witness ∣e∣ : ∥ X ≃ Bool ∥ to get +-- that there is therefore a unique pointed isomorphism (Bool , false) ≃ (X , x) for any +-- 2-element pointed type (X , x, ∣e∣). +isContr-2-EltPointedEquiv : (X∙ : 2-EltPointed₀) + → isContr ((Bool , false , ∣ idEquiv Bool ∣) ≃[ PointedEqvToEquivStr Bool ] X∙) +isContr-2-EltPointedEquiv (X , x , ∣e∣) + = PropTrunc.rec isPropIsContr + (λ e → J (λ X∙ _ → isContr ((Bool , false) ≃[ PointedEquivStr ] X∙)) + (isContrBoolPointedEquiv (e .fst x)) + (sym (pointed-sip _ _ (e , refl)))) + ∣e∣ + +-- This unique isomorphism must be _⊕_ 'lifted' to X. This idea is alluded to at the end of the +-- proof of Theorem III.4 in [BR17], where the authors reference needing ⊕-comm. +module ⊕* (X : 2-EltType₀) where + + _⊕*_ : typ X → typ X → Bool + y ⊕* z = invEquiv (fst (fst (isContr-2-EltPointedEquiv (fst X , y , snd X)))) .fst z + + -- we've already shown that this map is an equivalence on the right + + isEquivʳ : (y : typ X) → isEquiv (y ⊕*_) + isEquivʳ y = invEquiv (fst (fst (isContr-2-EltPointedEquiv (fst X , y , snd X)))) .snd + + Equivʳ : typ X → typ X ≃ Bool + Equivʳ y = (y ⊕*_) , isEquivʳ y + + -- any mere proposition that holds for (Bool, _⊕_) holds for (typ X, _⊕*_) + -- this amounts to just carefully unfolding the PropTrunc.elim and J in isContr-2-EltPointedEquiv + elim : ∀ {ℓ'} (P : (A : Type₀) (_⊕'_ : A → A → Bool) → Type ℓ') (propP : ∀ A _⊕'_ → isProp (P A _⊕'_)) + → P Bool _⊕_ → P (typ X) _⊕*_ + elim {ℓ'} P propP r = PropTrunc.elim {P = λ ∣e∣ → P (typ X) (R₁ ∣e∣)} (λ _ → propP _ _) + (λ e → EquivJ (λ A e → P A (R₂ A e)) r e) + (snd X) + where R₁ : ∥ fst X ≃ Bool ∥ → typ X → typ X → Bool + R₁ ∣e∣ y = invEq (fst (fst (isContr-2-EltPointedEquiv (fst X , y , ∣e∣)))) + R₂ : (B : Type₀) → B ≃ Bool → B → B → Bool + R₂ A e y = invEq (fst (fst (J (λ A∙ _ → isContr ((Bool , false) ≃[ PointedEquivStr ] A∙)) + (isContrBoolPointedEquiv (e .fst y)) + (sym (pointed-sip (A , y) (Bool , e .fst y) (e , refl)))))) + + -- as a consequence, we get that ⊕* is commutative, and is therefore also an equivalence on the left + + comm : (y z : typ X) → y ⊕* z ≡ z ⊕* y + comm = elim (λ A _⊕'_ → (x y : A) → x ⊕' y ≡ y ⊕' x) + (λ _ _ → isPropΠ2 (λ _ _ → isSetBool _ _)) + ⊕-comm + + isEquivˡ : (y : typ X) → isEquiv (_⊕* y) + isEquivˡ y = subst isEquiv (funExt (λ z → comm y z)) (isEquivʳ y) + + Equivˡ : typ X → typ X ≃ Bool + Equivˡ y = (_⊕* y) , isEquivˡ y + +-- Lemma II.2 in [BR17], though we do not use it here +-- Note: Lemma II.3 is `pointed-sip`, used in `PointedEqvTo-sip` +isContr-2-EltPointed : isContr (2-EltPointed₀) +fst isContr-2-EltPointed = (Bool , false , ∣ idEquiv Bool ∣) +snd isContr-2-EltPointed A∙ = PointedEqvTo-sip Bool _ A∙ (fst (isContr-2-EltPointedEquiv A∙)) + + +-------------------------------------------------------------------------------- + +-- Now we mutually define RP n and its double cover (Definition III.1 in [BR17]), +-- and show that the total space of this double cover is S n (Theorem III.4). + +RP : ℕ₋₁ → Type₀ +cov⁻¹ : (n : ℕ₋₁) → RP n → 2-EltType₀ -- (see Cubical.Functions.Bundle) + +RP neg1 = ⊥ +RP (ℕ→ℕ₋₁ n) = Pushout (pr (cov⁻¹ (-1+ n))) (λ _ → tt) +{- + tt + Total (cov⁻¹ (n-1)) — — — > Unit + | ∙ + pr | ∙ inr + | ∙ + V V + RP (n-1) ∙ ∙ ∙ ∙ ∙ ∙ > RP n := Pushout pr (const tt) + inl +-} + +cov⁻¹ neg1 x = Bool* +cov⁻¹ (ℕ→ℕ₋₁ n) (inl x) = cov⁻¹ (-1+ n) x +cov⁻¹ (ℕ→ℕ₋₁ n) (inr _) = Bool* +cov⁻¹ (ℕ→ℕ₋₁ n) (push (x , y) i) = ua ((λ z → y ⊕* z) , ⊕*.isEquivʳ (cov⁻¹ (-1+ n) x) y) i , ∣p∣ i + where open ⊕* (cov⁻¹ (-1+ n) x) + ∣p∣ = isProp→PathP (λ i → squash {A = ua (⊕*.Equivʳ (cov⁻¹ (-1+ n) x) y) i ≃ Bool}) + (str (cov⁻¹ (-1+ n) x)) (∣ idEquiv _ ∣) +{- + tt + Total (cov⁻¹ (n-1)) — — — > Unit + | | + pr | // ua α // | Bool + | | + V V + RP (n-1) - - - - - - > Type + cov⁻¹ (n-1) + + where α : ∀ (x : Total (cov⁻¹ (n-1))) → cov⁻¹ (n-1) (pr x) ≃ Bool + α (x , y) = (λ z → y ⊕* z) , ⊕*.isEquivʳ y +-} + +TotalCov≃Sn : ∀ n → Total (cov⁻¹ n) ≃ S n +TotalCov≃Sn neg1 = isoToEquiv (iso (λ { () }) (λ { () }) (λ { () }) (λ { () })) +TotalCov≃Sn (ℕ→ℕ₋₁ n) = + Total (cov⁻¹ (ℕ→ℕ₋₁ n)) ≃⟨ i ⟩ + Pushout Σf Σg ≃⟨ ii ⟩ + join (Total (cov⁻¹ (-1+ n))) Bool ≃⟨ iii ⟩ + S (ℕ→ℕ₋₁ n) ■ + where +{- + (i) First we want to show that `Total (cov⁻¹ (ℕ→ℕ₋₁ n))` is equivalent to a pushout. + We do this using the flattening lemma, which states: + + Given f,g,F,G,e such that the following square commutes: + + g + A — — — — > C Define: E : Pushout f g → Type + | | E (inl b) = F b + f | ua e | G E (inr c) = G c + V V E (push a i) = ua (e a) i + B — — — — > Type + F + + Then, the total space `Σ (Pushout f g) E` is the following pushout: + + Σg := (g , e a) + Σ[ a ∈ A ] F (f a) — — — — — — — — > Σ[ c ∈ C ] G c + | ∙ + Σf := (f , id) | ∙ + V V + Σ[ b ∈ B ] F b ∙ ∙ ∙ ∙ ∙ ∙ ∙ ∙ ∙ ∙ > Σ (Pushout f g) E + + In our case, setting `f = pr (cov⁻¹ (n-1))`, `g = λ _ → tt`, `F = cov⁻¹ (n-1)`, `G = λ _ → Bool`, + and `e = λ (x , y) → ⊕*.Equivʳ (cov⁻¹ (n-1) x) y` makes E equal (up to funExt) to `cov⁻¹ n`. + + Thus the flattening lemma gives us that `Total (cov⁻¹ n) ≃ Pushout Σf Σg`. +-} + open FlatteningLemma {- f = -} (λ x → pr (cov⁻¹ (-1+ n)) x) {- g = -} (λ _ → tt) + {- F = -} (λ x → typ (cov⁻¹ (-1+ n) x)) {- G = -} (λ _ → Bool) + {- e = -} (λ { (x , y) → ⊕*.Equivʳ (cov⁻¹ (-1+ n) x) y }) + hiding (Σf ; Σg) + + cov⁻¹≃E : ∀ x → typ (cov⁻¹ (ℕ→ℕ₋₁ n) x) ≃ E x + cov⁻¹≃E (inl x) = idEquiv _ + cov⁻¹≃E (inr x) = idEquiv _ + cov⁻¹≃E (push a i) = idEquiv _ + + -- for easier reference, we copy these definitons here + Σf : Σ[ x ∈ Total (cov⁻¹ (-1+ n)) ] typ (cov⁻¹ (-1+ n) (fst x)) → Total (cov⁻¹ (-1+ n)) + Σg : Σ[ x ∈ Total (cov⁻¹ (-1+ n)) ] typ (cov⁻¹ (-1+ n) (fst x)) → Unit × Bool + Σf ((x , y) , z) = (x , z) -- ≡ (f a , x) + Σg ((x , y) , z) = (tt , y ⊕* z) -- ≡ (g a , (e a) .fst x) + where open ⊕* (cov⁻¹ (-1+ n) x) + + i : Total (cov⁻¹ (ℕ→ℕ₋₁ n)) ≃ Pushout Σf Σg + i = (Σ[ x ∈ RP (ℕ→ℕ₋₁ n) ] typ (cov⁻¹ (ℕ→ℕ₋₁ n) x)) ≃⟨ Σ-cong-equiv-snd cov⁻¹≃E ⟩ + (Σ[ x ∈ RP (ℕ→ℕ₋₁ n) ] E x) ≃⟨ flatten ⟩ + Pushout Σf Σg ■ +{- + (ii) Next we want to show that `Pushout Σf Σg` is equivalent to `join (Total (cov⁻¹ (n-1))) Bool`. + Since both are pushouts, this can be done by defining a diagram equivalence: + + Σf Σg + Total (cov⁻¹ (n-1)) < — — Σ[ x ∈ Total (cov⁻¹ (n-1)) ] cov⁻¹ (n-1) (pr x) — — > Unit × Bool + | ∙ | + id |≃ u ∙≃ snd |≃ + V V V + Total (cov⁻¹ (n-1)) < — — — — — — — Total (cov⁻¹ (n-1)) × Bool — — — — — — — — — > Bool + proj₁ proj₂ + + where the equivalence u above must therefore satisfy: `u .fst x ≡ (Σf x , snd (Σg x))` + Unfolding this, we get: `u .fst ((x , y) , z) ≡ ((x , z) , (y ⊕* z))` + + It suffices to show that the map y ↦ y ⊕* z is an equivalence, since we can then express u as + the following composition of equivalences: + ((x , y) , z) ↦ (x , (y , z)) ↦ (x , (z , y)) ↦ (x , (z , y ⊕* z)) ↦ ((x , z) , y ⊕* z) + + This was proved above by ⊕*.isEquivˡ. +-} + u : ∀ {n} → (Σ[ x ∈ Total (cov⁻¹ n) ] typ (cov⁻¹ n (fst x))) ≃ (Total (cov⁻¹ n) × Bool) + u {n} = Σ[ x ∈ Total (cov⁻¹ n) ] typ (cov⁻¹ n (fst x)) ≃⟨ Σ-assoc-≃ ⟩ + Σ[ x ∈ RP n ] (typ (cov⁻¹ n x)) × (typ (cov⁻¹ n x)) ≃⟨ Σ-cong-equiv-snd (λ x → Σ-swap-≃) ⟩ + Σ[ x ∈ RP n ] (typ (cov⁻¹ n x)) × (typ (cov⁻¹ n x)) ≃⟨ Σ-cong-equiv-snd + (λ x → Σ-cong-equiv-snd + (λ y → ⊕*.Equivˡ (cov⁻¹ n x) y)) ⟩ + Σ[ x ∈ RP n ] (typ (cov⁻¹ n x)) × Bool ≃⟨ invEquiv Σ-assoc-≃ ⟩ + Total (cov⁻¹ n) × Bool ■ + + H : ∀ x → u .fst x ≡ (Σf x , snd (Σg x)) + H x = refl + + nat : 3-span-equiv (3span Σf Σg) (3span {A2 = Total (cov⁻¹ (-1+ n)) × Bool} fst snd) + nat = record { e0 = idEquiv _ ; e2 = u ; e4 = ΣUnit _ + ; H1 = λ x → cong fst (H x) + ; H3 = λ x → cong snd (H x) } + + ii : Pushout Σf Σg ≃ join (Total (cov⁻¹ (-1+ n))) Bool + ii = compEquiv (pathToEquiv (spanEquivToPushoutPath nat)) (joinPushout≃join _ _) + +{- + (iii) Finally, it's trivial to show that `join (Total (cov⁻¹ (n-1))) Bool` is equivalent to + `Susp (Total (cov⁻¹ (n-1)))`. Induction then gives us that `Susp (Total (cov⁻¹ (n-1)))` + is equivalent to `S n`, which completes the proof. +-} + + iii : join (Total (cov⁻¹ (-1+ n))) Bool ≃ S (ℕ→ℕ₋₁ n) + iii = join (Total (cov⁻¹ (-1+ n))) Bool ≃⟨ invEquiv Susp≃joinBool ⟩ + Susp (Total (cov⁻¹ (-1+ n))) ≃⟨ congSuspEquiv (TotalCov≃Sn (-1+ n)) ⟩ + S (ℕ→ℕ₋₁ n) ■ + +-- the usual covering map S n → RP n, with fibers exactly cov⁻¹ + +cov : (n : ℕ₋₁) → S n → RP n +cov n = pr (cov⁻¹ n) ∘ invEq (TotalCov≃Sn n) + +fibcov≡cov⁻¹ : ∀ n (x : RP n) → fiber (cov n) x ≡ cov⁻¹ n x .fst +fibcov≡cov⁻¹ n x = + fiber (cov n) x ≡[ i ]⟨ fiber {A = ua e i} (pr (cov⁻¹ n) ∘ ua-unglue e i) x ⟩ + fiber (pr (cov⁻¹ n)) x ≡⟨ ua (fibPrEquiv (cov⁻¹ n) x) ⟩ + cov⁻¹ n x .fst ∎ + where e = invEquiv (TotalCov≃Sn n) + + +-------------------------------------------------------------------------------- + +-- Finally, we state the trivial equivalences for RP 0 and RP 1 (Example III.3 in [BR17]) + +RP0≃Unit : RP 0 ≃ Unit +RP0≃Unit = isoToEquiv (iso (λ _ → tt) (λ _ → inr tt) (λ _ → refl) (λ { (inr tt) → refl })) + +RP1≡S1 : RP 1 ≡ S 1 +RP1≡S1 = Pushout {A = Total (cov⁻¹ 0)} {B = RP 0} (pr (cov⁻¹ 0)) (λ _ → tt) ≡⟨ i ⟩ + Pushout {A = Total (cov⁻¹ 0)} {B = Unit} (λ _ → tt) (λ _ → tt) ≡⟨ ii ⟩ + Pushout {A = S 0} {B = Unit} (λ _ → tt) (λ _ → tt) ≡⟨ PushoutSusp≡Susp ⟩ + S 1 ∎ + where i = λ i → Pushout {A = Total (cov⁻¹ 0)} + {B = ua RP0≃Unit i} + (λ x → ua-gluePt RP0≃Unit i (pr (cov⁻¹ 0) x)) + (λ _ → tt) + ii = λ j → Pushout {A = ua (TotalCov≃Sn 0) j} (λ _ → tt) (λ _ → tt) diff --git a/Cubical/HITs/Rationals/HITQ.agda b/Cubical/HITs/Rationals/HITQ.agda new file mode 100644 index 0000000000..ce7734b93f --- /dev/null +++ b/Cubical/HITs/Rationals/HITQ.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.HITQ where + +open import Cubical.HITs.Rationals.HITQ.Base public + +-- open import Cubical.HITs.Rationals.HITQ.Properties public diff --git a/Cubical/HITs/Rationals/HITQ/Base.agda b/Cubical/HITs/Rationals/HITQ/Base.agda new file mode 100644 index 0000000000..f93a022ae4 --- /dev/null +++ b/Cubical/HITs/Rationals/HITQ/Base.agda @@ -0,0 +1,36 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.HITQ.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Relation.Nullary + +open import Cubical.Data.Int + +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma + + +-- ℚ as a higher inductive type + +data ℚ : Type₀ where + con : (u : ℤ) (a : ℤ) → ¬ (a ≡ pos 0) → ℚ + path : ∀ u a v b {p q} → (u · b) ≡ (v · a) → con u a p ≡ con v b q + trunc : isSet ℚ + +[_] : ℤ × ℕ₊₁ → ℚ +[ a , 1+ b ] = con a (pos (suc b)) (ℕ.snotz ∘ cong abs) + + +-- Natural number and negative integer literals for ℚ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℚ : HasFromNat ℚ + fromNatℚ = record { Constraint = λ _ → Unit ; fromNat = λ n → [ pos n , 1 ] } + +instance + fromNegℚ : HasFromNeg ℚ + fromNegℚ = record { Constraint = λ _ → Unit ; fromNeg = λ n → [ neg n , 1 ] } diff --git a/Cubical/HITs/Rationals/QuoQ.agda b/Cubical/HITs/Rationals/QuoQ.agda new file mode 100644 index 0000000000..3cd9432665 --- /dev/null +++ b/Cubical/HITs/Rationals/QuoQ.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.QuoQ where + +open import Cubical.HITs.Rationals.QuoQ.Base public + +open import Cubical.HITs.Rationals.QuoQ.Properties public diff --git a/Cubical/HITs/Rationals/QuoQ/Base.agda b/Cubical/HITs/Rationals/QuoQ/Base.agda new file mode 100644 index 0000000000..5af5d6d571 --- /dev/null +++ b/Cubical/HITs/Rationals/QuoQ/Base.agda @@ -0,0 +1,76 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.QuoQ.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Nat as ℕ using (discreteℕ) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma +open import Cubical.Data.Int.MoreInts.QuoInt + +open import Cubical.HITs.SetQuotients as SetQuotient + using ([_]; eq/; discreteSetQuotients) renaming (_/_ to _//_) public + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary.Base +open BinaryRelation + +ℕ₊₁→ℤ : ℕ₊₁ → ℤ +ℕ₊₁→ℤ n = pos (ℕ₊₁→ℕ n) + +private + ℕ₊₁→ℤ-hom : ∀ m n → ℕ₊₁→ℤ (m ·₊₁ n) ≡ ℕ₊₁→ℤ m · ℕ₊₁→ℤ n + ℕ₊₁→ℤ-hom _ _ = refl + + +-- ℚ as a set quotient of ℤ × ℕ₊₁ (as in the HoTT book) + +_∼_ : ℤ × ℕ₊₁ → ℤ × ℕ₊₁ → Type₀ +(a , b) ∼ (c , d) = a · ℕ₊₁→ℤ d ≡ c · ℕ₊₁→ℤ b + +ℚ : Type₀ +ℚ = (ℤ × ℕ₊₁) // _∼_ + + +isSetℚ : isSet ℚ +isSetℚ = SetQuotient.squash/ + +[_/_] : ℤ → ℕ₊₁ → ℚ +[ a / b ] = [ a , b ] + + +isEquivRel∼ : isEquivRel _∼_ +isEquivRel.reflexive isEquivRel∼ (a , b) = refl +isEquivRel.symmetric isEquivRel∼ (a , b) (c , d) = sym +isEquivRel.transitive isEquivRel∼ (a , b) (c , d) (e , f) p q = ·-injʳ _ _ _ r + where r = (a · ℕ₊₁→ℤ f) · ℕ₊₁→ℤ d ≡[ i ]⟨ ·-comm a (ℕ₊₁→ℤ f) i · ℕ₊₁→ℤ d ⟩ + (ℕ₊₁→ℤ f · a) · ℕ₊₁→ℤ d ≡⟨ sym (·-assoc (ℕ₊₁→ℤ f) a (ℕ₊₁→ℤ d)) ⟩ + ℕ₊₁→ℤ f · (a · ℕ₊₁→ℤ d) ≡[ i ]⟨ ℕ₊₁→ℤ f · p i ⟩ + ℕ₊₁→ℤ f · (c · ℕ₊₁→ℤ b) ≡⟨ ·-assoc (ℕ₊₁→ℤ f) c (ℕ₊₁→ℤ b) ⟩ + (ℕ₊₁→ℤ f · c) · ℕ₊₁→ℤ b ≡[ i ]⟨ ·-comm (ℕ₊₁→ℤ f) c i · ℕ₊₁→ℤ b ⟩ + (c · ℕ₊₁→ℤ f) · ℕ₊₁→ℤ b ≡[ i ]⟨ q i · ℕ₊₁→ℤ b ⟩ + (e · ℕ₊₁→ℤ d) · ℕ₊₁→ℤ b ≡⟨ sym (·-assoc e (ℕ₊₁→ℤ d) (ℕ₊₁→ℤ b)) ⟩ + e · (ℕ₊₁→ℤ d · ℕ₊₁→ℤ b) ≡[ i ]⟨ e · ·-comm (ℕ₊₁→ℤ d) (ℕ₊₁→ℤ b) i ⟩ + e · (ℕ₊₁→ℤ b · ℕ₊₁→ℤ d) ≡⟨ ·-assoc e (ℕ₊₁→ℤ b) (ℕ₊₁→ℤ d) ⟩ + (e · ℕ₊₁→ℤ b) · ℕ₊₁→ℤ d ∎ + +eq/⁻¹ : ∀ x y → Path ℚ [ x ] [ y ] → x ∼ y +eq/⁻¹ = SetQuotient.effective (λ _ _ → isSetℤ _ _) isEquivRel∼ + +discreteℚ : Discrete ℚ +discreteℚ = discreteSetQuotients (discreteΣ discreteℤ (λ _ → subst Discrete 1+Path discreteℕ)) + (λ _ _ → isSetℤ _ _) isEquivRel∼ (λ _ _ → discreteℤ _ _) + + +-- Natural number and negative integer literals for ℚ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℚ : HasFromNat ℚ + fromNatℚ = record { Constraint = λ _ → Unit ; fromNat = λ n → [ pos n / 1 ] } + +instance + fromNegℚ : HasFromNeg ℚ + fromNegℚ = record { Constraint = λ _ → Unit ; fromNeg = λ n → [ neg n / 1 ] } diff --git a/Cubical/HITs/Rationals/QuoQ/Properties.agda b/Cubical/HITs/Rationals/QuoQ/Properties.agda new file mode 100644 index 0000000000..cc8932b221 --- /dev/null +++ b/Cubical/HITs/Rationals/QuoQ/Properties.agda @@ -0,0 +1,245 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.QuoQ.Properties where + +open import Cubical.Foundations.Everything hiding (_⁻¹) + +open import Cubical.Data.Int.MoreInts.QuoInt as ℤ using (ℤ; Sign; signed; pos; neg; posneg; sign) +open import Cubical.HITs.SetQuotients as SetQuotient using () renaming (_/_ to _//_) + +open import Cubical.Data.Nat as ℕ using (ℕ; zero; suc) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma + +open import Cubical.Data.Sum +open import Cubical.Relation.Nullary + +open import Cubical.HITs.Rationals.QuoQ.Base + +ℚ-cancelˡ : ∀ {a b} (c : ℕ₊₁) → [ ℕ₊₁→ℤ c ℤ.· a / c ·₊₁ b ] ≡ [ a / b ] +ℚ-cancelˡ {a} {b} c = eq/ _ _ + (cong (ℤ._· ℕ₊₁→ℤ b) (ℤ.·-comm (ℕ₊₁→ℤ c) a) ∙ sym (ℤ.·-assoc a (ℕ₊₁→ℤ c) (ℕ₊₁→ℤ b))) + +ℚ-cancelʳ : ∀ {a b} (c : ℕ₊₁) → [ a ℤ.· ℕ₊₁→ℤ c / b ·₊₁ c ] ≡ [ a / b ] +ℚ-cancelʳ {a} {b} c = eq/ _ _ + (sym (ℤ.·-assoc a (ℕ₊₁→ℤ c) (ℕ₊₁→ℤ b)) ∙ cong (a ℤ.·_) (ℤ.·-comm (ℕ₊₁→ℤ c) (ℕ₊₁→ℤ b))) + +-- useful functions for defining operations on ℚ + +onCommonDenom : + (g : ℤ × ℕ₊₁ → ℤ × ℕ₊₁ → ℤ) + (g-eql : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : a ℤ.· ℕ₊₁→ℤ d ≡ c ℤ.· ℕ₊₁→ℤ b) + → ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) ≡ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f))) + (g-eqr : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : c ℤ.· ℕ₊₁→ℤ f ≡ e ℤ.· ℕ₊₁→ℤ d) + → (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f ≡ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d) + → ℚ → ℚ → ℚ +onCommonDenom g g-eql g-eqr = SetQuotient.rec2 isSetℚ + (λ { (a , b) (c , d) → [ g (a , b) (c , d) / b ·₊₁ d ] }) + (λ { (a , b) (c , d) (e , f) p → eql (a , b) (c , d) (e , f) p }) + (λ { (a , b) (c , d) (e , f) p → eqr (a , b) (c , d) (e , f) p }) + where eql : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : a ℤ.· ℕ₊₁→ℤ d ≡ c ℤ.· ℕ₊₁→ℤ b) + → [ g (a , b) (e , f) / b ·₊₁ f ] ≡ [ g (c , d) (e , f) / d ·₊₁ f ] + eql (a , b) (c , d) (e , f) p = + [ g (a , b) (e , f) / b ·₊₁ f ] + ≡⟨ sym (ℚ-cancelˡ d) ⟩ + [ ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) / d ·₊₁ (b ·₊₁ f) ] + ≡[ i ]⟨ [ ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) / ·₊₁-assoc d b f i ] ⟩ + [ ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) / (d ·₊₁ b) ·₊₁ f ] + ≡[ i ]⟨ [ g-eql (a , b) (c , d) (e , f) p i / ·₊₁-comm d b i ·₊₁ f ] ⟩ + [ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f)) / (b ·₊₁ d) ·₊₁ f ] + ≡[ i ]⟨ [ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f)) / ·₊₁-assoc b d f (~ i) ] ⟩ + [ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f)) / b ·₊₁ (d ·₊₁ f) ] + ≡⟨ ℚ-cancelˡ b ⟩ + [ g (c , d) (e , f) / d ·₊₁ f ] ∎ + eqr : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : c ℤ.· ℕ₊₁→ℤ f ≡ e ℤ.· ℕ₊₁→ℤ d) + → [ g (a , b) (c , d) / b ·₊₁ d ] ≡ [ g (a , b) (e , f) / b ·₊₁ f ] + eqr (a , b) (c , d) (e , f) p = + [ g (a , b) (c , d) / b ·₊₁ d ] + ≡⟨ sym (ℚ-cancelʳ f) ⟩ + [ (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f / (b ·₊₁ d) ·₊₁ f ] + ≡[ i ]⟨ [ (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f / ·₊₁-assoc b d f (~ i) ] ⟩ + [ (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f / b ·₊₁ (d ·₊₁ f) ] + ≡[ i ]⟨ [ g-eqr (a , b) (c , d) (e , f) p i / b ·₊₁ ·₊₁-comm d f i ] ⟩ + [ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d / b ·₊₁ (f ·₊₁ d) ] + ≡[ i ]⟨ [ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d / ·₊₁-assoc b f d i ] ⟩ + [ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d / (b ·₊₁ f) ·₊₁ d ] + ≡⟨ ℚ-cancelʳ d ⟩ + [ g (a , b) (e , f) / b ·₊₁ f ] ∎ + +onCommonDenomSym : + (g : ℤ × ℕ₊₁ → ℤ × ℕ₊₁ → ℤ) + (g-sym : ∀ x y → g x y ≡ g y x) + (g-eql : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : a ℤ.· ℕ₊₁→ℤ d ≡ c ℤ.· ℕ₊₁→ℤ b) + → ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) ≡ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f))) + → ℚ → ℚ → ℚ +onCommonDenomSym g g-sym g-eql = onCommonDenom g g-eql q-eqr + where q-eqr : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : c ℤ.· ℕ₊₁→ℤ f ≡ e ℤ.· ℕ₊₁→ℤ d) + → (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f ≡ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d + q-eqr (a , b) (c , d) (e , f) p = + (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f ≡[ i ]⟨ ℤ.·-comm (g-sym (a , b) (c , d) i) (ℕ₊₁→ℤ f) i ⟩ + ℕ₊₁→ℤ f ℤ.· (g (c , d) (a , b)) ≡⟨ g-eql (c , d) (e , f) (a , b) p ⟩ + ℕ₊₁→ℤ d ℤ.· (g (e , f) (a , b)) ≡[ i ]⟨ ℤ.·-comm (ℕ₊₁→ℤ d) (g-sym (e , f) (a , b) i) i ⟩ + (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d ∎ + +onCommonDenomSym-comm : ∀ {g} g-sym {g-eql} (x y : ℚ) + → onCommonDenomSym g g-sym g-eql x y ≡ + onCommonDenomSym g g-sym g-eql y x +onCommonDenomSym-comm g-sym = SetQuotient.elimProp2 (λ _ _ → isSetℚ _ _) + (λ { (a , b) (c , d) i → [ g-sym (a , b) (c , d) i / ·₊₁-comm b d i ] }) + + +-- basic arithmetic operations on ℚ + +infixl 6 _+_ +infixl 7 _·_ + +private + lem₁ : ∀ a b c d e (p : a ℤ.· b ≡ c ℤ.· d) → b ℤ.· (a ℤ.· e) ≡ d ℤ.· (c ℤ.· e) + lem₁ a b c d e p = ℤ.·-assoc b a e + ∙ cong (ℤ._· e) (ℤ.·-comm b a ∙ p ∙ ℤ.·-comm c d) + ∙ sym (ℤ.·-assoc d c e) + + lem₂ : ∀ a b c → a ℤ.· (b ℤ.· c) ≡ c ℤ.· (b ℤ.· a) + lem₂ a b c = cong (a ℤ.·_) (ℤ.·-comm b c) ∙ ℤ.·-assoc a c b + ∙ cong (ℤ._· b) (ℤ.·-comm a c) ∙ sym (ℤ.·-assoc c a b) + ∙ cong (c ℤ.·_) (ℤ.·-comm a b) + +_+_ : ℚ → ℚ → ℚ +_+_ = onCommonDenomSym + (λ { (a , b) (c , d) → a ℤ.· (ℕ₊₁→ℤ d) ℤ.+ c ℤ.· (ℕ₊₁→ℤ b) }) + (λ { (a , b) (c , d) → ℤ.+-comm (a ℤ.· (ℕ₊₁→ℤ d)) (c ℤ.· (ℕ₊₁→ℤ b)) }) + (λ { (a , b) (c , d) (e , f) p → + ℕ₊₁→ℤ d ℤ.· (a ℤ.· ℕ₊₁→ℤ f ℤ.+ e ℤ.· ℕ₊₁→ℤ b) + ≡⟨ sym (ℤ.·-distribˡ (ℕ₊₁→ℤ d) (a ℤ.· ℕ₊₁→ℤ f) (e ℤ.· ℕ₊₁→ℤ b)) ⟩ + ℕ₊₁→ℤ d ℤ.· (a ℤ.· ℕ₊₁→ℤ f) ℤ.+ ℕ₊₁→ℤ d ℤ.· (e ℤ.· ℕ₊₁→ℤ b) + ≡[ i ]⟨ lem₁ a (ℕ₊₁→ℤ d) c (ℕ₊₁→ℤ b) (ℕ₊₁→ℤ f) p i ℤ.+ lem₂ (ℕ₊₁→ℤ d) e (ℕ₊₁→ℤ b) i ⟩ + ℕ₊₁→ℤ b ℤ.· (c ℤ.· ℕ₊₁→ℤ f) ℤ.+ ℕ₊₁→ℤ b ℤ.· (e ℤ.· ℕ₊₁→ℤ d) + ≡⟨ ℤ.·-distribˡ (ℕ₊₁→ℤ b) (c ℤ.· ℕ₊₁→ℤ f) (e ℤ.· ℕ₊₁→ℤ d) ⟩ + ℕ₊₁→ℤ b ℤ.· (c ℤ.· ℕ₊₁→ℤ f ℤ.+ e ℤ.· ℕ₊₁→ℤ d) ∎ }) + ++-comm : ∀ x y → x + y ≡ y + x ++-comm = onCommonDenomSym-comm + (λ { (a , b) (c , d) → ℤ.+-comm (a ℤ.· (ℕ₊₁→ℤ d)) (c ℤ.· (ℕ₊₁→ℤ b)) }) + ++-identityˡ : ∀ x → 0 + x ≡ x ++-identityˡ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) i → [ ℤ.·-identityʳ a i / ·₊₁-identityˡ b i ] }) + ++-identityʳ : ∀ x → x + 0 ≡ x ++-identityʳ x = +-comm x _ ∙ +-identityˡ x + ++-assoc : ∀ x y z → x + (y + z) ≡ (x + y) + z ++-assoc = SetQuotient.elimProp3 (λ _ _ _ → isSetℚ _ _) + (λ { (a , b) (c , d) (e , f) i → [ eq a (ℕ₊₁→ℤ b) c (ℕ₊₁→ℤ d) e (ℕ₊₁→ℤ f) i / ·₊₁-assoc b d f i ] }) + where eq₁ : ∀ a b c → (a ℤ.· b) ℤ.· c ≡ a ℤ.· (c ℤ.· b) + eq₁ a b c = sym (ℤ.·-assoc a b c) ∙ cong (a ℤ.·_) (ℤ.·-comm b c) + eq₂ : ∀ a b c → (a ℤ.· b) ℤ.· c ≡ (a ℤ.· c) ℤ.· b + eq₂ a b c = eq₁ a b c ∙ ℤ.·-assoc a c b + + eq : ∀ a b c d e f → Path ℤ _ _ + eq a b c d e f = + a ℤ.· (d ℤ.· f) ℤ.+ (c ℤ.· f ℤ.+ e ℤ.· d) ℤ.· b + ≡[ i ]⟨ a ℤ.· (d ℤ.· f) ℤ.+ ℤ.·-distribʳ (c ℤ.· f) (e ℤ.· d) b (~ i) ⟩ + a ℤ.· (d ℤ.· f) ℤ.+ ((c ℤ.· f) ℤ.· b ℤ.+ (e ℤ.· d) ℤ.· b) + ≡[ i ]⟨ ℤ.+-assoc (ℤ.·-assoc a d f i) (eq₂ c f b i) (eq₁ e d b i) i ⟩ + ((a ℤ.· d) ℤ.· f ℤ.+ (c ℤ.· b) ℤ.· f) ℤ.+ e ℤ.· (b ℤ.· d) + ≡[ i ]⟨ ℤ.·-distribʳ (a ℤ.· d) (c ℤ.· b) f i ℤ.+ e ℤ.· (b ℤ.· d) ⟩ + (a ℤ.· d ℤ.+ c ℤ.· b) ℤ.· f ℤ.+ e ℤ.· (b ℤ.· d) ∎ + + +_·_ : ℚ → ℚ → ℚ +_·_ = onCommonDenomSym + (λ { (a , _) (c , _) → a ℤ.· c }) + (λ { (a , _) (c , _) → ℤ.·-comm a c }) + (λ { (a , b) (c , d) (e , _) p → lem₁ a (ℕ₊₁→ℤ d) c (ℕ₊₁→ℤ b) e p }) + +·-comm : ∀ x y → x · y ≡ y · x +·-comm = onCommonDenomSym-comm (λ { (a , _) (c , _) → ℤ.·-comm a c }) + +·-identityˡ : ∀ x → 1 · x ≡ x +·-identityˡ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) i → [ ℤ.·-identityˡ a i / ·₊₁-identityˡ b i ] }) + +·-identityʳ : ∀ x → x · 1 ≡ x +·-identityʳ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) i → [ ℤ.·-identityʳ a i / ·₊₁-identityʳ b i ] }) + +·-zeroˡ : ∀ x → 0 · x ≡ 0 +·-zeroˡ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) → (λ i → [ p a b i / 1 ·₊₁ b ]) ∙ ℚ-cancelʳ b }) + where p : ∀ a b → 0 ℤ.· a ≡ 0 ℤ.· ℕ₊₁→ℤ b + p a b = ℤ.·-zeroˡ {ℤ.spos} a ∙ sym (ℤ.·-zeroˡ {ℤ.spos} (ℕ₊₁→ℤ b)) + +·-zeroʳ : ∀ x → x · 0 ≡ 0 +·-zeroʳ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) → (λ i → [ p a b i / b ·₊₁ 1 ]) ∙ ℚ-cancelˡ b }) + where p : ∀ a b → a ℤ.· 0 ≡ ℕ₊₁→ℤ b ℤ.· 0 + p a b = ℤ.·-zeroʳ {ℤ.spos} a ∙ sym (ℤ.·-zeroʳ {ℤ.spos} (ℕ₊₁→ℤ b)) + +·-assoc : ∀ x y z → x · (y · z) ≡ (x · y) · z +·-assoc = SetQuotient.elimProp3 (λ _ _ _ → isSetℚ _ _) + (λ { (a , b) (c , d) (e , f) i → [ ℤ.·-assoc a c e i / ·₊₁-assoc b d f i ] }) + +·-distribˡ : ∀ x y z → (x · y) + (x · z) ≡ x · (y + z) +·-distribˡ = SetQuotient.elimProp3 (λ _ _ _ → isSetℚ _ _) + (λ { (a , b) (c , d) (e , f) → eq a b c d e f }) + where lem : ∀ {ℓ} {A : Type ℓ} (_·_ : A → A → A) + (·-comm : ∀ x y → x · y ≡ y · x) + (·-assoc : ∀ x y z → x · (y · z) ≡ (x · y) · z) + a c b d + → (a · c) · (b · d) ≡ (a · (c · d)) · b + lem _·_ ·-comm ·-assoc a c b d = + (a · c) · (b · d) ≡[ i ]⟨ (a · c) · ·-comm b d i ⟩ + (a · c) · (d · b) ≡⟨ ·-assoc (a · c) d b ⟩ + ((a · c) · d) · b ≡[ i ]⟨ ·-assoc a c d (~ i) · b ⟩ + (a · (c · d)) · b ∎ + + lemℤ = lem ℤ._·_ ℤ.·-comm ℤ.·-assoc + lemℕ₊₁ = lem _·₊₁_ ·₊₁-comm ·₊₁-assoc + + eq : ∀ a b c d e f → + [ (a ℤ.· c) ℤ.· ℕ₊₁→ℤ (b ·₊₁ f) ℤ.+ (a ℤ.· e) ℤ.· ℕ₊₁→ℤ (b ·₊₁ d) + / (b ·₊₁ d) ·₊₁ (b ·₊₁ f) ] + ≡ [ a ℤ.· (c ℤ.· ℕ₊₁→ℤ f ℤ.+ e ℤ.· ℕ₊₁→ℤ d) + / b ·₊₁ (d ·₊₁ f) ] + eq a b c d e f = + (λ i → [ lemℤ a c (ℕ₊₁→ℤ b) (ℕ₊₁→ℤ f) i ℤ.+ lemℤ a e (ℕ₊₁→ℤ b) (ℕ₊₁→ℤ d) i + / lemℕ₊₁ b d b f i ]) ∙ + (λ i → [ ℤ.·-distribʳ (a ℤ.· (c ℤ.· ℕ₊₁→ℤ f)) (a ℤ.· (e ℤ.· ℕ₊₁→ℤ d)) (ℕ₊₁→ℤ b) i + / (b ·₊₁ (d ·₊₁ f)) ·₊₁ b ]) ∙ + ℚ-cancelʳ {a ℤ.· (c ℤ.· ℕ₊₁→ℤ f) ℤ.+ a ℤ.· (e ℤ.· ℕ₊₁→ℤ d)} {b ·₊₁ (d ·₊₁ f)} b ∙ + (λ i → [ ℤ.·-distribˡ a (c ℤ.· ℕ₊₁→ℤ f) (e ℤ.· ℕ₊₁→ℤ d) i + / b ·₊₁ (d ·₊₁ f) ]) + +·-distribʳ : ∀ x y z → (x · z) + (y · z) ≡ (x + y) · z +·-distribʳ x y z = (λ i → ·-comm x z i + ·-comm y z i) ∙ ·-distribˡ z x y ∙ ·-comm z (x + y) + + +-_ : ℚ → ℚ +- x = -1 · x + +negate-invol : ∀ x → - - x ≡ x +negate-invol x = ·-assoc -1 -1 x ∙ ·-identityˡ x + +negateEquiv : ℚ ≃ ℚ +negateEquiv = isoToEquiv (iso -_ -_ negate-invol negate-invol) + +negateEq : ℚ ≡ ℚ +negateEq = ua negateEquiv + ++-inverseˡ : ∀ x → (- x) + x ≡ 0 ++-inverseˡ x = (λ i → (-1 · x) + ·-identityˡ x (~ i)) ∙ ·-distribʳ -1 1 x ∙ ·-zeroˡ x + +_-_ : ℚ → ℚ → ℚ +x - y = x + (- y) + ++-inverseʳ : ∀ x → x - x ≡ 0 ++-inverseʳ x = +-comm x (- x) ∙ +-inverseˡ x + ++-injˡ : ∀ x y z → x + y ≡ x + z → y ≡ z ++-injˡ x y z p = sym (q y) ∙ cong ((- x) +_) p ∙ q z + where q : ∀ y → (- x) + (x + y) ≡ y + q y = +-assoc (- x) x y ∙ cong (_+ y) (+-inverseˡ x) ∙ +-identityˡ y + ++-injʳ : ∀ x y z → x + y ≡ z + y → x ≡ z ++-injʳ x y z p = +-injˡ y x z (+-comm y x ∙ p ∙ +-comm z y) + diff --git a/Cubical/HITs/Rationals/SigmaQ.agda b/Cubical/HITs/Rationals/SigmaQ.agda new file mode 100644 index 0000000000..46a9efbd49 --- /dev/null +++ b/Cubical/HITs/Rationals/SigmaQ.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.SigmaQ where + +open import Cubical.HITs.Rationals.SigmaQ.Base public + +open import Cubical.HITs.Rationals.SigmaQ.Properties public diff --git a/Cubical/HITs/Rationals/SigmaQ/Base.agda b/Cubical/HITs/Rationals/SigmaQ/Base.agda new file mode 100644 index 0000000000..8dda85f8bf --- /dev/null +++ b/Cubical/HITs/Rationals/SigmaQ/Base.agda @@ -0,0 +1,54 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.SigmaQ.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Int.MoreInts.QuoInt + +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma + +open import Cubical.Data.Nat.GCD +open import Cubical.Data.Nat.Coprime + + +-- ℚ as the set of coprime pairs in ℤ × ℕ₊₁ + +ℚ : Type₀ +ℚ = Σ[ (a , b) ∈ ℤ × ℕ₊₁ ] areCoprime (abs a , ℕ₊₁→ℕ b) + +isSetℚ : isSet ℚ +isSetℚ = isSetΣ (isSet× isSetℤ (subst isSet 1+Path isSetℕ)) (λ _ → isProp→isSet isPropIsGCD) + + +signedPair : Sign → ℕ × ℕ₊₁ → ℤ × ℕ₊₁ +signedPair s (a , b) = (signed s a , b) + +[_] : ℤ × ℕ₊₁ → ℚ +[ signed s a , b ] = signedPair s (toCoprime (a , b)) , toCoprimeAreCoprime (a , b) +[ posneg i , b ] = (posneg i , 1) , toCoprimeAreCoprime (0 , b) + +[]-cancelʳ : ∀ ((a , b) : ℤ × ℕ₊₁) k → [ a · pos (ℕ₊₁→ℕ k) , b ·₊₁ k ] ≡ [ a , b ] +[]-cancelʳ (signed s zero , b) k = + Σ≡Prop (λ _ → isPropIsGCD) (λ i → signed-zero spos s i , 1) +[]-cancelʳ (signed s (suc a) , b) k = + Σ≡Prop (λ _ → isPropIsGCD) (λ i → signedPair (·S-comm s spos i) + (toCoprime-cancelʳ (suc a , b) k i)) +[]-cancelʳ (posneg i , b) k j = + isSet→isSet' isSetℚ ([]-cancelʳ (pos zero , b) k) ([]-cancelʳ (neg zero , b) k) + (λ i → [ posneg i · pos (ℕ₊₁→ℕ k) , b ·₊₁ k ]) (λ i → [ posneg i , b ]) i j + + +-- Natural number and negative integer literals for ℚ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℚ : HasFromNat ℚ + fromNatℚ = record { Constraint = λ _ → Unit ; fromNat = λ n → (pos n , 1) , oneGCD n } + +instance + fromNegℚ : HasFromNeg ℚ + fromNegℚ = record { Constraint = λ _ → Unit ; fromNeg = λ n → (neg n , 1) , oneGCD n } diff --git a/Cubical/HITs/Rationals/SigmaQ/Properties.agda b/Cubical/HITs/Rationals/SigmaQ/Properties.agda new file mode 100644 index 0000000000..8dcedfe4cf --- /dev/null +++ b/Cubical/HITs/Rationals/SigmaQ/Properties.agda @@ -0,0 +1,78 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.SigmaQ.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Int.MoreInts.QuoInt +import Cubical.HITs.SetQuotients as SetQuotient + +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma + +open import Cubical.Data.Nat.GCD +open import Cubical.Data.Nat.Coprime + +open import Cubical.HITs.Rationals.QuoQ as Quo using (ℕ₊₁→ℤ) +import Cubical.HITs.Rationals.SigmaQ.Base as Sigma + + +reduce : Quo.ℚ → Sigma.ℚ +reduce = SetQuotient.rec Sigma.isSetℚ Sigma.[_] + (λ { (a , b) (c , d) p → sym (Sigma.[]-cancelʳ (a , b) d) + ∙ (λ i → Sigma.[ p i , ·₊₁-comm b d i ]) + ∙ Sigma.[]-cancelʳ (c , d) b }) + +private + toCoprime-eq₁ : ∀ s ((a , b) : ℕ × ℕ₊₁) + → signed s (toCoprime (a , b) .fst) · ℕ₊₁→ℤ b + ≡ signed s a · ℕ₊₁→ℤ (toCoprime (a , b) .snd) + toCoprime-eq₁ s (a , b) = + signed s c₁ · ℕ₊₁→ℤ b ≡⟨ ·-signed-pos c₁ (ℕ₊₁→ℕ b) ⟩ + signed s (c₁ ℕ.· ℕ₊₁→ℕ b) ≡[ i ]⟨ signed s (c₁ ℕ.· p₂ (~ i)) ⟩ + signed s (c₁ ℕ.· (ℕ₊₁→ℕ c₂ ℕ.· d)) ≡[ i ]⟨ signed s (c₁ ℕ.· ℕ.·-comm (ℕ₊₁→ℕ c₂) d i) ⟩ + signed s (c₁ ℕ.· (d ℕ.· ℕ₊₁→ℕ c₂)) ≡[ i ]⟨ signed s (ℕ.·-assoc c₁ d (ℕ₊₁→ℕ c₂) i) ⟩ + signed s ((c₁ ℕ.· d) ℕ.· ℕ₊₁→ℕ c₂) ≡[ i ]⟨ signed s (p₁ i ℕ.· ℕ₊₁→ℕ c₂) ⟩ + signed s (a ℕ.· ℕ₊₁→ℕ c₂) ≡⟨ sym (·-signed-pos a (ℕ₊₁→ℕ c₂)) ⟩ + signed s a · ℕ₊₁→ℤ c₂ ∎ + where open ToCoprime (a , b) + +[]-reduce : ∀ x → Quo.[ reduce x .fst ] ≡ x +[]-reduce = SetQuotient.elimProp (λ _ → Quo.isSetℚ _ _) + (λ { (signed s a , b) → Quo.eq/ _ _ (toCoprime-eq₁ s (a , b)) + ; (posneg i , b) j → isSet→isSet' Quo.isSetℚ + (Quo.eq/ (fst (reduce Quo.[ pos 0 , b ])) (pos 0 , b) (toCoprime-eq₁ spos (0 , b))) + (Quo.eq/ (fst (reduce Quo.[ neg 0 , b ])) (neg 0 , b) (toCoprime-eq₁ sneg (0 , b))) + (λ i → Quo.[ reduce (Quo.[ posneg i , b ]) .fst ]) (λ i → Quo.[ posneg i , b ]) i j }) + +private + toCoprime-eq₂ : ∀ s ((a , b) : ℕ × ℕ₊₁) (cp : areCoprime (a , ℕ₊₁→ℕ b)) + → Sigma.signedPair s (toCoprime (a , b)) ≡ Sigma.signedPair s (a , b) + toCoprime-eq₂ s (a , b) cp i = Sigma.signedPair s (toCoprime-idem (a , b) cp i) + +reduce-[] : ∀ x → reduce Quo.[ x .fst ] ≡ x + -- equivalent to: Sigma.[ s .fst ] ≡ x +reduce-[] ((signed s a , b) , cp) = + Σ≡Prop (λ _ → isPropIsGCD) (toCoprime-eq₂ s (a , b) cp) +reduce-[] ((posneg i , b) , cp) j = + isSet→isSet' Sigma.isSetℚ + (Σ≡Prop (λ _ → isPropIsGCD) (toCoprime-eq₂ spos (0 , b) cp)) + (Σ≡Prop (λ _ → isPropIsGCD) (toCoprime-eq₂ sneg (0 , b) cp)) + (λ i → Sigma.[ posneg i , b ]) (λ i → (posneg i , b) , cp) i j + + +Quoℚ-iso-Sigmaℚ : Iso Quo.ℚ Sigma.ℚ +Iso.fun Quoℚ-iso-Sigmaℚ = reduce +Iso.inv Quoℚ-iso-Sigmaℚ = Quo.[_] ∘ fst +Iso.rightInv Quoℚ-iso-Sigmaℚ = reduce-[] +Iso.leftInv Quoℚ-iso-Sigmaℚ = []-reduce + +Quoℚ≃Sigmaℚ : Quo.ℚ ≃ Sigma.ℚ +Quoℚ≃Sigmaℚ = isoToEquiv Quoℚ-iso-Sigmaℚ + +Quoℚ≡Sigmaℚ : Quo.ℚ ≡ Sigma.ℚ +Quoℚ≡Sigmaℚ = isoToPath Quoℚ-iso-Sigmaℚ diff --git a/Cubical/HITs/S1.agda b/Cubical/HITs/S1.agda new file mode 100644 index 0000000000..9490001b24 --- /dev/null +++ b/Cubical/HITs/S1.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.S1 where + +open import Cubical.HITs.S1.Base public +open import Cubical.HITs.S1.Properties public diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda new file mode 100644 index 0000000000..cfea54dc7b --- /dev/null +++ b/Cubical/HITs/S1/Base.agda @@ -0,0 +1,493 @@ +{- + +Definition of the circle as a HIT with a proof that Ω(S¹) ≡ ℤ + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.S1.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Nat + hiding (_+_ ; _·_ ; +-assoc ; +-comm) +open import Cubical.Data.Int hiding (_·_) + +data S¹ : Type₀ where + base : S¹ + loop : base ≡ base + +-- Check that transp is the identity function for S¹ +module _ where + transpS¹ : ∀ (φ : I) (u0 : S¹) → transp (λ _ → S¹) φ u0 ≡ u0 + transpS¹ φ u0 = refl + + compS1 : ∀ (φ : I) (u : ∀ i → Partial φ S¹) (u0 : S¹ [ φ ↦ u i0 ]) → + comp (λ _ → S¹) (\ i → u i) (outS u0) ≡ hcomp u (outS u0) + compS1 φ u u0 = refl + +-- ΩS¹ ≡ ℤ +helix : S¹ → Type₀ +helix base = ℤ +helix (loop i) = sucPathℤ i + +ΩS¹ : Type₀ +ΩS¹ = base ≡ base + +encode : ∀ x → base ≡ x → helix x +encode x p = subst helix p (pos zero) + +winding : ΩS¹ → ℤ +winding = encode base + +intLoop : ℤ → ΩS¹ +intLoop (pos zero) = refl +intLoop (pos (suc n)) = intLoop (pos n) ∙ loop +intLoop (negsuc zero) = sym loop +intLoop (negsuc (suc n)) = intLoop (negsuc n) ∙ sym loop + +decodeSquare : (n : ℤ) → PathP (λ i → base ≡ loop i) (intLoop (predℤ n)) (intLoop n) +decodeSquare (pos zero) i j = loop (i ∨ ~ j) +decodeSquare (pos (suc n)) i j = hfill (λ k → λ { (j = i0) → base + ; (j = i1) → loop k } ) + (inS (intLoop (pos n) j)) i +decodeSquare (negsuc n) i j = hfill (λ k → λ { (j = i0) → base + ; (j = i1) → loop (~ k) }) + (inS (intLoop (negsuc n) j)) (~ i) + +decode : (x : S¹) → helix x → base ≡ x +decode base = intLoop +decode (loop i) y j = + let n : ℤ + n = unglue (i ∨ ~ i) y + in hcomp (λ k → λ { (i = i0) → intLoop (predSuc y k) j + ; (i = i1) → intLoop y j + ; (j = i0) → base + ; (j = i1) → loop i }) + (decodeSquare n i j) + +decodeEncode : (x : S¹) (p : base ≡ x) → decode x (encode x p) ≡ p +decodeEncode x p = J (λ y q → decode y (encode y q) ≡ q) (λ _ → refl) p + +isSetΩS¹ : isSet ΩS¹ +isSetΩS¹ p q r s j i = + hcomp (λ k → λ { (i = i0) → decodeEncode base p k + ; (i = i1) → decodeEncode base q k + ; (j = i0) → decodeEncode base (r i) k + ; (j = i1) → decodeEncode base (s i) k }) + (decode base (isSetℤ (winding p) (winding q) (cong winding r) (cong winding s) j i)) + +-- This proof does not rely on rewriting hcomp with empty systems in +-- ℤ as ghcomp has been implemented! +windingℤLoop : (n : ℤ) → winding (intLoop n) ≡ n +windingℤLoop (pos zero) = refl +windingℤLoop (pos (suc n)) = cong sucℤ (windingℤLoop (pos n)) +windingℤLoop (negsuc zero) = refl +windingℤLoop (negsuc (suc n)) = cong predℤ (windingℤLoop (negsuc n)) + +ΩS¹Isoℤ : Iso ΩS¹ ℤ +Iso.fun ΩS¹Isoℤ = winding +Iso.inv ΩS¹Isoℤ = intLoop +Iso.rightInv ΩS¹Isoℤ = windingℤLoop +Iso.leftInv ΩS¹Isoℤ = decodeEncode base + +ΩS¹≡ℤ : ΩS¹ ≡ ℤ +ΩS¹≡ℤ = isoToPath ΩS¹Isoℤ + +-- intLoop and winding are group homomorphisms +private + intLoop-sucℤ : (z : ℤ) → intLoop (sucℤ z) ≡ intLoop z ∙ loop + intLoop-sucℤ (pos n) = refl + intLoop-sucℤ (negsuc zero) = sym (lCancel loop) + intLoop-sucℤ (negsuc (suc n)) = + rUnit (intLoop (negsuc n)) + ∙ (λ i → intLoop (negsuc n) ∙ lCancel loop (~ i)) + ∙ assoc (intLoop (negsuc n)) (sym loop) loop + + intLoop-predℤ : (z : ℤ) → intLoop (predℤ z) ≡ intLoop z ∙ sym loop + intLoop-predℤ (pos zero) = lUnit (sym loop) + intLoop-predℤ (pos (suc n)) = + rUnit (intLoop (pos n)) + ∙ (λ i → intLoop (pos n) ∙ (rCancel loop (~ i))) + ∙ assoc (intLoop (pos n)) loop (sym loop) + intLoop-predℤ (negsuc n) = refl + +intLoop-hom : (a b : ℤ) → (intLoop a) ∙ (intLoop b) ≡ intLoop (a + b) +intLoop-hom a (pos zero) = sym (rUnit (intLoop a)) +intLoop-hom a (pos (suc n)) = + assoc (intLoop a) (intLoop (pos n)) loop + ∙ (λ i → (intLoop-hom a (pos n) i) ∙ loop) + ∙ sym (intLoop-sucℤ (a + pos n)) +intLoop-hom a (negsuc zero) = sym (intLoop-predℤ a) +intLoop-hom a (negsuc (suc n)) = + assoc (intLoop a) (intLoop (negsuc n)) (sym loop) + ∙ (λ i → (intLoop-hom a (negsuc n) i) ∙ (sym loop)) + ∙ sym (intLoop-predℤ (a + negsuc n)) + +winding-hom : (a b : ΩS¹) → winding (a ∙ b) ≡ (winding a) + (winding b) +winding-hom a b i = + hcomp (λ t → λ { (i = i0) → winding (decodeEncode base a t ∙ decodeEncode base b t) + ; (i = i1) → windingℤLoop (winding a + winding b) t }) + (winding (intLoop-hom (winding a) (winding b) i)) + +-- Commutativity +comm-ΩS¹ : (p q : ΩS¹) → p ∙ q ≡ q ∙ p +comm-ΩS¹ p q = sym (cong₂ (_∙_) (decodeEncode base p) (decodeEncode base q)) + ∙ intLoop-hom (winding p) (winding q) + ∙ cong intLoop (+Comm (winding p) (winding q)) + ∙ sym (intLoop-hom (winding q) (winding p)) + ∙ (cong₂ (_∙_) (decodeEncode base q) (decodeEncode base p)) + +-- Based homotopy group + +basedΩS¹ : (x : S¹) → Type₀ +basedΩS¹ x = x ≡ x + +-- Proof that the homotopy group is actually independent on the basepoint +-- first, give a quasi-inverse to the basechange basedΩS¹→ΩS¹ for any loop i +-- (which does *not* match at endpoints) +private + ΩS¹→basedΩS¹-filler : I → I → ΩS¹ → I → S¹ + ΩS¹→basedΩS¹-filler l i x j = + hfill (λ t → λ { (j = i0) → loop (i ∧ t) + ; (j = i1) → loop (i ∧ t) }) + (inS (x j)) l + + basedΩS¹→ΩS¹-filler : (_ i : I) → basedΩS¹ (loop i) → I → S¹ + basedΩS¹→ΩS¹-filler l i x j = + hfill (λ t → λ { (j = i0) → loop (i ∧ (~ t)) + ; (j = i1) → loop (i ∧ (~ t)) }) + (inS (x j)) l + + ΩS¹→basedΩS¹ : (i : I) → ΩS¹ → basedΩS¹ (loop i) + ΩS¹→basedΩS¹ i x j = ΩS¹→basedΩS¹-filler i1 i x j + + basedΩS¹→ΩS¹ : (i : I) → basedΩS¹ (loop i) → ΩS¹ + basedΩS¹→ΩS¹ i x j = basedΩS¹→ΩS¹-filler i1 i x j + + + + basedΩS¹→ΩS¹→basedΩS¹ : (i : I) → (x : basedΩS¹ (loop i)) + → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) ≡ x + basedΩS¹→ΩS¹→basedΩS¹ i x j k = + hcomp (λ t → λ { (j = i1) → basedΩS¹→ΩS¹-filler (~ t) i x k + ; (j = i0) → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) k + ; (k = i0) → loop (i ∧ (t ∨ (~ j))) + ; (k = i1) → loop (i ∧ (t ∨ (~ j))) }) + (ΩS¹→basedΩS¹-filler (~ j) i (basedΩS¹→ΩS¹ i x) k) + + ΩS¹→basedΩS¹→ΩS¹ : (i : I) → (x : ΩS¹) + → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) ≡ x + ΩS¹→basedΩS¹→ΩS¹ i x j k = + hcomp (λ t → λ { (j = i1) → ΩS¹→basedΩS¹-filler (~ t) i x k + ; (j = i0) → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) k + ; (k = i0) → loop (i ∧ ((~ t) ∧ j)) + ; (k = i1) → loop (i ∧ ((~ t) ∧ j)) }) + (basedΩS¹→ΩS¹-filler (~ j) i (ΩS¹→basedΩS¹ i x) k) + + -- from the existence of our quasi-inverse, we deduce that the basechange is an equivalence + -- for all loop i + + basedΩS¹→ΩS¹-isequiv : (i : I) → isEquiv (basedΩS¹→ΩS¹ i) + basedΩS¹→ΩS¹-isequiv i = isoToIsEquiv (iso (basedΩS¹→ΩS¹ i) (ΩS¹→basedΩS¹ i) + (ΩS¹→basedΩS¹→ΩS¹ i) (basedΩS¹→ΩS¹→basedΩS¹ i)) + + + -- now extend the basechange so that both ends match + -- (and therefore we get a basechange for any x : S¹) + + loop-conjugation : basedΩS¹→ΩS¹ i1 ≡ λ x → x + loop-conjugation i x = + let p = (doubleCompPath-elim loop x (sym loop)) + ∙ (λ i → (lUnit loop i ∙ x) ∙ sym loop) + in + ((sym (decodeEncode base (basedΩS¹→ΩS¹ i1 x))) + ∙ (λ t → intLoop (winding (p t))) + ∙ (λ t → intLoop (winding-hom (intLoop (pos (suc zero)) ∙ x) + (intLoop (negsuc zero)) t)) + ∙ (λ t → intLoop ((winding-hom (intLoop (pos (suc zero))) x t) + + (windingℤLoop (negsuc zero) t))) + ∙ (λ t → intLoop (((windingℤLoop (pos (suc zero)) t) + (winding x)) + (negsuc zero))) + ∙ (λ t → intLoop ((+Comm (pos (suc zero)) (winding x) t) + (negsuc zero))) + ∙ (λ t → intLoop (+Assoc (winding x) (pos (suc zero)) (negsuc zero) (~ t))) + ∙ (decodeEncode base x)) i + + refl-conjugation : basedΩS¹→ΩS¹ i0 ≡ λ x → x + refl-conjugation i x j = + hfill (λ t → λ { (j = i0) → base + ; (j = i1) → base }) + (inS (x j)) (~ i) + + basechange : (x : S¹) → basedΩS¹ x → ΩS¹ + basechange base y = y + basechange (loop i) y = + hcomp (λ t → λ { (i = i0) → refl-conjugation t y + ; (i = i1) → loop-conjugation t y }) + (basedΩS¹→ΩS¹ i y) + + -- for any loop i, the old basechange is equal to the new one + basedΩS¹→ΩS¹≡basechange : (i : I) → basedΩS¹→ΩS¹ i ≡ basechange (loop i) + basedΩS¹→ΩS¹≡basechange i j y = + hfill (λ t → λ { (i = i0) → refl-conjugation t y + ; (i = i1) → loop-conjugation t y }) + (inS (basedΩS¹→ΩS¹ i y)) j + + -- so for any loop i, the extended basechange is an equivalence + basechange-isequiv-aux : (i : I) → isEquiv (basechange (loop i)) + basechange-isequiv-aux i = + transport (λ j → isEquiv (basedΩS¹→ΩS¹≡basechange i j)) (basedΩS¹→ΩS¹-isequiv i) + + + -- as being an equivalence is contractible, basechange is an equivalence for all x : S¹ + basechange-isequiv : (x : S¹) → isEquiv (basechange x) + basechange-isequiv base = basechange-isequiv-aux i0 + basechange-isequiv (loop i) = + hcomp (λ t → λ { (i = i0) → basechange-isequiv-aux i0 + ; (i = i1) → isPropIsEquiv (basechange base) (basechange-isequiv-aux i1) + (basechange-isequiv-aux i0) t }) + (basechange-isequiv-aux i) + + basedΩS¹≡ΩS¹' : (x : S¹) → basedΩS¹ x ≡ ΩS¹ + basedΩS¹≡ΩS¹' x = ua (basechange x , basechange-isequiv x) + + basedΩS¹≡ℤ' : (x : S¹) → basedΩS¹ x ≡ ℤ + basedΩS¹≡ℤ' x = (basedΩS¹≡ΩS¹' x) ∙ ΩS¹≡ℤ + + +---- Alternative proof of the same thing ----- + +toPropElim : ∀ {ℓ} {B : S¹ → Type ℓ} + → ((s : S¹) → isProp (B s)) + → B base + → (s : S¹) → B s +toPropElim isprop b base = b +toPropElim isprop b (loop i) = + isOfHLevel→isOfHLevelDep 1 isprop b b loop i + +toPropElim2 : ∀ {ℓ} {B : S¹ → S¹ → Type ℓ} + → ((s t : S¹) → isProp (B s t)) + → B base base + → (s t : S¹) → B s t +toPropElim2 isprop b base base = b +toPropElim2 isprop b base (loop i) = + isProp→PathP (λ i → isprop base (loop i)) b b i +toPropElim2 isprop b (loop i) base = + isProp→PathP (λ i → isprop (loop i) base) b b i +toPropElim2 {B = B} isprop b (loop i) (loop j) = + isSet→SquareP (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) + (isProp→PathP (λ i₁ → isprop base (loop i₁)) b b) + (isProp→PathP (λ i₁ → isprop base (loop i₁)) b b) + (isProp→PathP (λ i₁ → isprop (loop i₁) base) b b) + (isProp→PathP (λ i₁ → isprop (loop i₁) base) b b) i j + +toSetElim2 : ∀ {ℓ} {B : S¹ → S¹ → Type ℓ} + → ((s t : S¹) → isSet (B s t)) + → (x : B base base) + → PathP (λ i → B base (loop i)) x x + → PathP (λ i → B (loop i) base) x x + → (s t : S¹) → B s t +toSetElim2 isset b right left base base = b +toSetElim2 isset b right left base (loop i) = right i +toSetElim2 isset b right left (loop i) base = left i +toSetElim2 isset b right left (loop i) (loop j) = + isSet→SquareP (λ _ _ → isset _ _) right right left left i j + +isSetΩx : (x : S¹) → isSet (x ≡ x) +isSetΩx = toPropElim (λ _ → isPropIsSet) isSetΩS¹ + +basechange2 : (x : S¹) → ΩS¹ → (x ≡ x) +basechange2 base p = p +basechange2 (loop i) p = + hcomp (λ k → λ { (i = i0) → lUnit (rUnit p (~ k)) (~ k) + ; (i = i1) → (cong ((sym loop) ∙_) (comm-ΩS¹ p loop) + ∙ assoc (sym loop) loop p + ∙ cong (_∙ p) (lCancel loop) + ∙ sym (lUnit _)) k }) + ((λ j → loop (~ j ∧ i)) ∙ p ∙ λ j → loop (j ∧ i)) +basechange2⁻ : (x : S¹) → (x ≡ x) → ΩS¹ +basechange2⁻ base p = p +basechange2⁻ (loop i) p = + hcomp (λ k → λ { (i = i0) → lUnit (rUnit p (~ k)) (~ k) + ; (i = i1) → (cong (loop ∙_) (comm-ΩS¹ p (sym loop)) + ∙ assoc loop (sym loop) p + ∙ cong (_∙ p) (rCancel loop) + ∙ sym (lUnit _)) k }) + ((λ j → loop (i ∧ j)) ∙ p ∙ λ j → loop (i ∧ (~ j))) +basechange2-sect : (x : S¹) → section (basechange2 x) (basechange2⁻ x) +basechange2-sect = + toPropElim (λ _ → isOfHLevelΠ 1 λ _ → isSetΩx _ _ _ ) + λ _ → refl + +basechange2-retr : (x : S¹) → retract (basechange2 x) (basechange2⁻ x) +basechange2-retr = + toPropElim (λ s → isOfHLevelΠ 1 λ x → isSetΩx _ _ _) + λ _ → refl + +Iso-basedΩS¹-ΩS¹ : (x : S¹) → Iso (basedΩS¹ x) ΩS¹ +Iso.fun (Iso-basedΩS¹-ΩS¹ x) = basechange2⁻ x +Iso.inv (Iso-basedΩS¹-ΩS¹ x) = basechange2 x +Iso.rightInv (Iso-basedΩS¹-ΩS¹ x) = basechange2-retr x +Iso.leftInv (Iso-basedΩS¹-ΩS¹ x) = basechange2-sect x + +Iso-basedΩS¹-ℤ : (x : S¹) → Iso (basedΩS¹ x) ℤ +Iso-basedΩS¹-ℤ x = compIso (Iso-basedΩS¹-ΩS¹ x) ΩS¹Isoℤ + +basedΩS¹≡ℤ : (x : S¹) → basedΩS¹ x ≡ ℤ +basedΩS¹≡ℤ x = isoToPath (Iso-basedΩS¹-ℤ x) + +-- baschange2⁻ is a morphism + +basechange2⁻-morph : (x : S¹) (p q : x ≡ x) → + basechange2⁻ x (p ∙ q) ≡ basechange2⁻ x p ∙ basechange2⁻ x q +basechange2⁻-morph = + toPropElim {B = λ x → (p q : x ≡ x) → basechange2⁻ x (p ∙ q) ≡ basechange2⁻ x p ∙ basechange2⁻ x q} + (λ _ → isPropΠ2 λ _ _ → isSetΩS¹ _ _) + λ _ _ → refl + +-- Some tests +module _ where + private + test-winding-pos : winding (intLoop (pos 5)) ≡ pos 5 + test-winding-pos = refl + + test-winding-neg : winding (intLoop (negsuc 5)) ≡ negsuc 5 + test-winding-neg = refl + +-- the inverse when S¹ is seen as a group + +invLooper : S¹ → S¹ +invLooper base = base +invLooper (loop i) = loop (~ i) + +invInvolutive : section invLooper invLooper +invInvolutive base = refl +invInvolutive (loop i) = refl + +invS¹Equiv : S¹ ≃ S¹ +invS¹Equiv = isoToEquiv theIso + where + theIso : Iso S¹ S¹ + Iso.fun theIso = invLooper + Iso.inv theIso = invLooper + Iso.rightInv theIso = invInvolutive + Iso.leftInv theIso = invInvolutive + +invS¹Path : S¹ ≡ S¹ +invS¹Path = ua invS¹Equiv + +-- rot, used in the Hopf fibration +-- we will denote rotation by _·_ +-- it is called μ in the HoTT-book in "8.5.2 The Hopf construction" + +rotLoop : (a : S¹) → a ≡ a +rotLoop base = loop +rotLoop (loop i) j = + hcomp (λ k → λ { (i = i0) → loop (j ∨ ~ k) + ; (i = i1) → loop (j ∧ k) + ; (j = i0) → loop (i ∨ ~ k) + ; (j = i1) → loop (i ∧ k)}) base + +_·_ : S¹ → S¹ → S¹ +base · x = x +(loop i) · x = rotLoop x i + +infixl 30 _·_ + + +-- rot i j = filler-rot i j i1 +filler-rot : I → I → I → S¹ +filler-rot i j = hfill (λ k → λ { (i = i0) → loop (j ∨ ~ k) + ; (i = i1) → loop (j ∧ k) + ; (j = i0) → loop (i ∨ ~ k) + ; (j = i1) → loop (i ∧ k) }) (inS base) + +isPropFamS¹ : ∀ {ℓ} (P : S¹ → Type ℓ) (pP : (x : S¹) → isProp (P x)) (b0 : P base) → + PathP (λ i → P (loop i)) b0 b0 +isPropFamS¹ P pP b0 i = pP (loop i) (transp (λ j → P (loop (i ∧ j))) (~ i) b0) + (transp (λ j → P (loop (i ∨ ~ j))) i b0) i + +rotIsEquiv : (a : S¹) → isEquiv (a ·_) +rotIsEquiv base = idIsEquiv S¹ +rotIsEquiv (loop i) = isPropFamS¹ (λ x → isEquiv (x ·_)) + (λ x → isPropIsEquiv (x ·_)) + (idIsEquiv _) i + +-- more direct definition of the rot (loop i) equivalence + +rotLoopInv : (a : S¹) → PathP (λ i → rotLoop (rotLoop a (~ i)) i ≡ a) refl refl +rotLoopInv a i j = homotopySymInv rotLoop a j i + +rotLoopEquiv : (i : I) → S¹ ≃ S¹ +rotLoopEquiv i = + isoToEquiv + (iso (λ a → rotLoop a i) + (λ a → rotLoop a (~ i)) + (λ a → rotLoopInv a i) + (λ a → rotLoopInv a (~ i))) + +-- some cancellation laws, used in the Hopf fibration +private + rotInv-aux-1 : I → I → I → I → S¹ + rotInv-aux-1 j k i = + hfill (λ l → λ { (k = i0) → (loop (i ∧ ~ l)) · loop j + ; (k = i1) → loop j + ; (i = i0) → (loop k · loop j) · loop (~ k) + ; (i = i1) → loop (~ k ∧ ~ l) · loop j }) + (inS ((loop (k ∨ i) · loop j) · loop (~ k))) + + rotInv-aux-2 : I → I → I → S¹ + rotInv-aux-2 i j k = + hcomp (λ l → λ { (k = i0) → invLooper (filler-rot (~ i) (~ j) l) + ; (k = i1) → loop (j ∧ l) + ; (i = i0) → filler-rot k j l + ; (i = i1) → loop (j ∧ l) + ; (j = i0) → loop (i ∨ k ∨ (~ l)) + ; (j = i1) → loop ((i ∨ k) ∧ l) }) + (base) + + rotInv-aux-3 : I → I → I → I → S¹ + rotInv-aux-3 j k i = + hfill (λ l → λ { (k = i0) → rotInv-aux-2 i j l + ; (k = i1) → loop j + ; (i = i0) → loop (k ∨ l) · loop j + ; (i = i1) → loop k · (invLooper (loop (~ j) · loop k)) }) + (inS (loop k · (invLooper (loop (~ j) · loop (k ∨ ~ i))))) + + rotInv-aux-4 : I → I → I → I → S¹ + rotInv-aux-4 j k i = + hfill (λ l → λ { (k = i0) → rotInv-aux-2 i j l + ; (k = i1) → loop j + ; (i = i0) → loop j · loop (k ∨ l) + ; (i = i1) → (invLooper (loop (~ j) · loop k)) · loop k }) + (inS ((invLooper (loop (~ j) · loop (k ∨ ~ i))) · loop k)) + +rotInv-1 : (a b : S¹) → b · a · invLooper b ≡ a +rotInv-1 base base i = base +rotInv-1 base (loop k) i = rotInv-aux-1 i0 k i i1 +rotInv-1 (loop j) base i = loop j +rotInv-1 (loop j) (loop k) i = rotInv-aux-1 j k i i1 + +rotInv-2 : (a b : S¹) → invLooper b · a · b ≡ a +rotInv-2 base base i = base +rotInv-2 base (loop k) i = rotInv-aux-1 i0 (~ k) i i1 +rotInv-2 (loop j) base i = loop j +rotInv-2 (loop j) (loop k) i = rotInv-aux-1 j (~ k) i i1 + +rotInv-3 : (a b : S¹) → b · (invLooper (invLooper a · b)) ≡ a +rotInv-3 base base i = base +rotInv-3 base (loop k) i = rotInv-aux-3 i0 k (~ i) i1 +rotInv-3 (loop j) base i = loop j +rotInv-3 (loop j) (loop k) i = rotInv-aux-3 j k (~ i) i1 + +rotInv-4 : (a b : S¹) → invLooper (b · invLooper a) · b ≡ a +rotInv-4 base base i = base +rotInv-4 base (loop k) i = rotInv-aux-4 i0 k (~ i) i1 +rotInv-4 (loop j) base i = loop j +rotInv-4 (loop j) (loop k) i = rotInv-aux-4 j k (~ i) i1 diff --git a/Cubical/HITs/S1/Properties.agda b/Cubical/HITs/S1/Properties.agda new file mode 100644 index 0000000000..7ba091704a --- /dev/null +++ b/Cubical/HITs/S1/Properties.agda @@ -0,0 +1,34 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.S1.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.HITs.S1.Base +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +isConnectedS¹ : (s : S¹) → ∥ base ≡ s ∥ +isConnectedS¹ base = ∣ refl ∣ +isConnectedS¹ (loop i) = + squash ∣ (λ j → loop (i ∧ j)) ∣ ∣ (λ j → loop (i ∨ ~ j)) ∣ i + +isGroupoidS¹ : isGroupoid S¹ +isGroupoidS¹ s t = + PropTrunc.rec isPropIsSet + (λ p → + subst (λ s → isSet (s ≡ t)) p + (PropTrunc.rec isPropIsSet + (λ q → subst (λ t → isSet (base ≡ t)) q isSetΩS¹) + (isConnectedS¹ t))) + (isConnectedS¹ s) + +IsoFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → Iso (S¹ → A) (Σ[ x ∈ A ] x ≡ x) +Iso.fun IsoFunSpaceS¹ f = (f base) , (cong f loop) +Iso.inv IsoFunSpaceS¹ (x , p) base = x +Iso.inv IsoFunSpaceS¹ (x , p) (loop i) = p i +Iso.rightInv IsoFunSpaceS¹ (x , p) = refl +Iso.leftInv IsoFunSpaceS¹ f = funExt λ {base → refl ; (loop i) → refl} diff --git a/Cubical/HITs/S2.agda b/Cubical/HITs/S2.agda new file mode 100644 index 0000000000..1c783bef7c --- /dev/null +++ b/Cubical/HITs/S2.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.S2 where + +open import Cubical.HITs.S2.Base public + +-- open import Cubical.HITs.S2.Properties public diff --git a/Cubical/HITs/S2/Base.agda b/Cubical/HITs/S2/Base.agda new file mode 100644 index 0000000000..6b156b5542 --- /dev/null +++ b/Cubical/HITs/S2/Base.agda @@ -0,0 +1,57 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.S2.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Pointed + +data S² : Type₀ where + base : S² + surf : PathP (λ i → base ≡ base) refl refl + +S²∙ : Pointed ℓ-zero +S²∙ = S² , base + +S²ToSetElim : ∀ {ℓ} {A : S² → Type ℓ} + → ((x : S²) → isSet (A x)) + → A base + → (x : S²) → A x +S²ToSetElim set b base = b +S²ToSetElim set b (surf i j) = + isOfHLevel→isOfHLevelDep 2 set b b {a0 = refl} {a1 = refl} refl refl surf i j + +-- Wedge connectivity lemmas for S² (binary maps 2-groupoids) +wedgeconFunS² : ∀ {ℓ} {P : S² → S² → Type ℓ} + → ((x y : _) → isOfHLevel 4 (P x y)) + → (l : ((x : S²) → P x base)) + → (r : (x : S²) → P base x) + → l base ≡ r base + → (x y : _) → P x y +wedgeconFunS² {P = P} hlev l r p base y = r y +wedgeconFunS² {P = P} hlev l r p (surf i i₁) y = help y i i₁ + where + help : (y : S²) → SquareP (λ i j → P (surf i j) y) + (λ _ → r y) (λ _ → r y) + (λ _ → r y) λ _ → r y + help = + S²ToSetElim (λ _ → isOfHLevelPathP' 2 (isOfHLevelPathP' 3 (hlev _ _) _ _) _ _) + λ w j → hcomp (λ k → λ { (j = i0) → p k + ; (j = i1) → p k + ; (w = i0) → p k + ; (w = i1) → p k}) + (l (surf w j)) + +wedgeconFunS²Id : ∀ {ℓ} {P : S² → S² → Type ℓ} + → (h : ((x y : _) → isOfHLevel 4 (P x y))) + → (l : ((x : S²) → P x base)) + → (r : (x : S²) → P base x) + → (p : l base ≡ r base) + → (x : S²) → wedgeconFunS² h l r p x base ≡ l x +wedgeconFunS²Id h l r p base = sym p +wedgeconFunS²Id h l r p (surf i j) k = + hcomp (λ w → λ {(i = i0) → p (~ k ∧ w) + ; (i = i1) → p (~ k ∧ w) + ; (j = i0) → p (~ k ∧ w) + ; (j = i1) → p (~ k ∧ w) + ; (k = i1) → l (surf i j)}) + (l (surf i j)) diff --git a/Cubical/HITs/S3.agda b/Cubical/HITs/S3.agda new file mode 100644 index 0000000000..e5c16e1cc8 --- /dev/null +++ b/Cubical/HITs/S3.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.S3 where + +open import Cubical.HITs.S3.Base public + +-- open import Cubical.HITs.S3.Properties public diff --git a/Cubical/HITs/S3/Base.agda b/Cubical/HITs/S3/Base.agda new file mode 100644 index 0000000000..dd765533ba --- /dev/null +++ b/Cubical/HITs/S3/Base.agda @@ -0,0 +1,8 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.S3.Base where + +open import Cubical.Foundations.Prelude + +data S³ : Type₀ where + base : S³ + surf : PathP (λ j → PathP (λ i → base ≡ base) refl refl) refl refl diff --git a/Cubical/HITs/SetQuotients.agda b/Cubical/HITs/SetQuotients.agda new file mode 100644 index 0000000000..7339c4344a --- /dev/null +++ b/Cubical/HITs/SetQuotients.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.SetQuotients where + +open import Cubical.HITs.SetQuotients.Base public +open import Cubical.HITs.SetQuotients.Properties public diff --git a/Cubical/HITs/SetQuotients/Base.agda b/Cubical/HITs/SetQuotients/Base.agda new file mode 100644 index 0000000000..2c47ec7d26 --- /dev/null +++ b/Cubical/HITs/SetQuotients/Base.agda @@ -0,0 +1,17 @@ +{- + +This file contains: + +- Definition of set quotients + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.SetQuotients.Base where + +open import Cubical.Core.Primitives + +-- Set quotients as a higher inductive type: +data _/_ {ℓ ℓ'} (A : Type ℓ) (R : A → A → Type ℓ') : Type (ℓ-max ℓ ℓ') where + [_] : (a : A) → A / R + eq/ : (a b : A) → (r : R a b) → [ a ] ≡ [ b ] + squash/ : (x y : A / R) → (p q : x ≡ y) → p ≡ q diff --git a/Cubical/HITs/SetQuotients/Properties.agda b/Cubical/HITs/SetQuotients/Properties.agda new file mode 100644 index 0000000000..1ee6891688 --- /dev/null +++ b/Cubical/HITs/SetQuotients/Properties.agda @@ -0,0 +1,334 @@ +{- + +Set quotients: + +-} + +{-# OPTIONS --safe #-} +module Cubical.HITs.SetQuotients.Properties where + +open import Cubical.HITs.SetQuotients.Base + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Univalence + +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Data.Sigma + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary.Base + +open import Cubical.HITs.TypeQuotients as TypeQuot using (_/ₜ_ ; [_] ; eq/) +open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥ ; ∣_∣ ; squash) +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂ ; ∣_∣₂ ; squash₂ + ; isSetSetTrunc) + + +private + variable + ℓ ℓ' ℓ'' : Level + A B C : Type ℓ + R S T : A → A → Type ℓ + +elimProp : {P : A / R → Type ℓ} + → (∀ x → isProp (P x)) + → (∀ a → P [ a ]) + → ∀ x → P x +elimProp prop f [ x ] = f x +elimProp prop f (squash/ x y p q i j) = + isOfHLevel→isOfHLevelDep 2 (λ x → isProp→isSet (prop x)) + (g x) (g y) (cong g p) (cong g q) (squash/ x y p q) i j + where + g = elimProp prop f +elimProp prop f (eq/ a b r i) = + isProp→PathP (λ i → prop (eq/ a b r i)) (f a) (f b) i + +elimProp2 : {P : A / R → B / S → Type ℓ} + → (∀ x y → isProp (P x y)) + → (∀ a b → P [ a ] [ b ]) + → ∀ x y → P x y +elimProp2 prop f = + elimProp (λ x → isPropΠ (prop x)) λ a → + elimProp (prop [ a ]) (f a) + +elimProp3 : {P : A / R → B / S → C / T → Type ℓ} + → (∀ x y z → isProp (P x y z)) + → (∀ a b c → P [ a ] [ b ] [ c ]) + → ∀ x y z → P x y z +elimProp3 prop f = + elimProp (λ x → isPropΠ2 (prop x)) λ a → + elimProp2 (prop [ a ]) (f a) + +-- sometimes more convenient: +elimContr : {P : A / R → Type ℓ} + → (∀ a → isContr (P [ a ])) + → ∀ x → P x +elimContr contr = + elimProp (elimProp (λ _ → isPropIsProp) λ _ → isContr→isProp (contr _)) λ _ → + contr _ .fst + +elimContr2 : {P : A / R → B / S → Type ℓ} + → (∀ a b → isContr (P [ a ] [ b ])) + → ∀ x y → P x y +elimContr2 contr = + elimContr λ _ → + isOfHLevelΠ 0 (elimContr λ _ → inhProp→isContr (contr _ _) isPropIsContr) + +-- lemma 6.10.2 in hott book +[]surjective : (x : A / R) → ∃[ a ∈ A ] [ a ] ≡ x +[]surjective = elimProp (λ x → squash) (λ a → ∣ a , refl ∣) + +elim : {P : A / R → Type ℓ} + → (∀ x → isSet (P x)) + → (f : (a : A) → (P [ a ])) + → ((a b : A) (r : R a b) → PathP (λ i → P (eq/ a b r i)) (f a) (f b)) + → ∀ x → P x +elim set f feq [ a ] = f a +elim set f feq (eq/ a b r i) = feq a b r i +elim set f feq (squash/ x y p q i j) = + isOfHLevel→isOfHLevelDep 2 set + (g x) (g y) (cong g p) (cong g q) (squash/ x y p q) i j + where + g = elim set f feq + +rec : isSet B + → (f : A → B) + → ((a b : A) (r : R a b) → f a ≡ f b) + → A / R → B +rec set f feq [ a ] = f a +rec set f feq (eq/ a b r i) = feq a b r i +rec set f feq (squash/ x y p q i j) = set (g x) (g y) (cong g p) (cong g q) i j + where + g = rec set f feq + +rec2 : isSet C + → (f : A → B → C) + → (∀ a b c → R a b → f a c ≡ f b c) + → (∀ a b c → S b c → f a b ≡ f a c) + → A / R → B / S → C +rec2 set f feql feqr = + rec (isSetΠ (λ _ → set)) + (λ a → rec set (f a) (feqr a)) + (λ a b r → funExt (elimProp (λ _ → set _ _) (λ c → feql a b c r))) + +-- the recursor for maps into groupoids: +-- i.e. for any type A with a binary relation R and groupoid B, +-- we can construct a map A / R → B from a map A → B satisfying the conditions +-- (i) ∀ (a b : A) → R a b → f a ≡ f b +-- (ii) ∀ (a b : A) → isProp (f a ≡ f b) + +-- We start by proving that we can recover the set-quotient +-- by set-truncating the (non-truncated type quotient) +typeQuotSetTruncIso : Iso (A / R) ∥ A /ₜ R ∥₂ +Iso.fun typeQuotSetTruncIso = rec isSetSetTrunc (λ a → ∣ [ a ] ∣₂) + λ a b r → cong ∣_∣₂ (eq/ a b r) +Iso.inv typeQuotSetTruncIso = SetTrunc.rec squash/ (TypeQuot.rec [_] eq/) +Iso.rightInv typeQuotSetTruncIso = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + (TypeQuot.elimProp (λ _ → squash₂ _ _) λ _ → refl) +Iso.leftInv typeQuotSetTruncIso = elimProp (λ _ → squash/ _ _) λ _ → refl + +module rec→Gpd {B : Type ℓ''} (Bgpd : isGroupoid B) + (f : A → B) + (feq : ∀ (a b : A) → R a b → f a ≡ f b) + (fprop : ∀ (a b : A) → isProp (f a ≡ f b)) + where + + fun : A / R → B + fun = f₁ ∘ f₂ + where + f₁ : ∥ A /ₜ R ∥₂ → B + f₁ = SetTrunc.rec→Gpd.fun Bgpd f/ congF/Const + where + f/ : A /ₜ R → B + f/ = TypeQuot.rec f feq + + congF/Const : (a b : A /ₜ R) (p q : a ≡ b) → cong f/ p ≡ cong f/ q + congF/Const = + TypeQuot.elimProp2 + (λ _ _ → isPropΠ2 λ _ _ → Bgpd _ _ _ _) + (λ a b p q → fprop a b (cong f/ p) (cong f/ q)) + + f₂ : A / R → ∥ A /ₜ R ∥₂ + f₂ = Iso.fun typeQuotSetTruncIso + + +setQuotUniversalIso : isSet B + → Iso (A / R → B) (Σ[ f ∈ (A → B) ] ((a b : A) → R a b → f a ≡ f b)) +Iso.fun (setQuotUniversalIso Bset) g = (λ a → g [ a ]) , λ a b r i → g (eq/ a b r i) +Iso.inv (setQuotUniversalIso Bset) h = rec Bset (fst h) (snd h) +Iso.rightInv (setQuotUniversalIso Bset) h = refl +Iso.leftInv (setQuotUniversalIso Bset) g = + funExt λ x → + PropTrunc.rec + (Bset (out (intro g) x) (g x)) + (λ sur → cong (out (intro g)) (sym (snd sur)) ∙ (cong g (snd sur))) + ([]surjective x) + where + intro = Iso.fun (setQuotUniversalIso Bset) + out = Iso.inv (setQuotUniversalIso Bset) + +setQuotUniversal : isSet B + → (A / R → B) ≃ (Σ[ f ∈ (A → B) ] ((a b : A) → R a b → f a ≡ f b)) +setQuotUniversal Bset = isoToEquiv (setQuotUniversalIso Bset) + +open BinaryRelation + +setQuotUnaryOp : (-_ : A → A) + → (∀ a a' → R a a' → R (- a) (- a')) + → (A / R → A / R) +setQuotUnaryOp -_ h = rec squash/ (λ a → [ - a ]) (λ a b x → eq/ _ _ (h _ _ x)) + +-- characterisation of binary functions/operations on set-quotients +setQuotUniversal2Iso : isSet C → isRefl R → isRefl S + → Iso (A / R → B / S → C) + (Σ[ _∗_ ∈ (A → B → C) ] (∀ a a' b b' → R a a' → S b b' → a ∗ b ≡ a' ∗ b')) +Iso.fun (setQuotUniversal2Iso {R = R} {S = S} Bset isReflR isReflS) _∗/_ = _∗_ , h + where + _∗_ = λ a b → [ a ] ∗/ [ b ] + + h : ∀ a a' b b' → R a a' → S b b' → a ∗ b ≡ a' ∗ b' + h a a' b b' r s = cong (_∗/ [ b ]) (eq/ _ _ r) ∙ cong ([ a' ] ∗/_) (eq/ _ _ s) +Iso.inv (setQuotUniversal2Iso {R = R} {S = S} Bset isReflR isReflS) (_∗_ , h) = + rec2 Bset _∗_ hleft hright + where + hleft : ∀ a a' b → R a a' → (a ∗ b) ≡ (a' ∗ b) + hleft _ _ b r = h _ _ _ _ r (isReflS b) + + hright : ∀ a b b' → S b b' → (a ∗ b) ≡ (a ∗ b') + hright a _ _ r = h _ _ _ _ (isReflR a) r +Iso.rightInv (setQuotUniversal2Iso Bset isReflR isReflS) (_∗_ , h) = + Σ≡Prop (λ _ → isPropΠ4 λ _ _ _ _ → isPropΠ2 λ _ _ → Bset _ _) refl +Iso.leftInv (setQuotUniversal2Iso Bset isReflR isReflS) _∗/_ = + funExt₂ (elimProp2 (λ _ _ → Bset _ _) λ _ _ → refl) + +setQuotUniversal2 : isSet C → isRefl R → isRefl S + → (A / R → B / S → C) + ≃ (Σ[ _∗_ ∈ (A → B → C) ] (∀ a a' b b' → R a a' → S b b' → a ∗ b ≡ a' ∗ b')) +setQuotUniversal2 Bset isReflR isReflS = + isoToEquiv (setQuotUniversal2Iso Bset isReflR isReflS) + +-- corollary for binary operations +-- TODO: prove truncated inverse for effective relations +setQuotBinOp : isRefl R → isRefl S + → (_∗_ : A → B → C) + → (∀ a a' b b' → R a a' → S b b' → T (a ∗ b) (a' ∗ b')) + → (A / R → B / S → C / T) +setQuotBinOp isReflR isReflS _∗_ h = + rec2 squash/ (λ a b → [ a ∗ b ]) + (λ _ _ _ r → eq/ _ _ (h _ _ _ _ r (isReflS _))) + (λ _ _ _ s → eq/ _ _ (h _ _ _ _ (isReflR _) s)) + +setQuotSymmBinOp : isRefl R → isTrans R + → (_∗_ : A → A → A) + → (∀ a b → R (a ∗ b) (b ∗ a)) + → (∀ a a' b → R a a' → R (a ∗ b) (a' ∗ b)) + → (A / R → A / R → A / R) +setQuotSymmBinOp {A = A} {R = R} isReflR isTransR _∗_ ∗Rsymm h = + setQuotBinOp isReflR isReflR _∗_ h' + where + h' : ∀ a a' b b' → R a a' → R b b' → R (a ∗ b) (a' ∗ b') + h' a a' b b' ra rb = + isTransR _ _ _ (h a a' b ra) + (isTransR _ _ _ (∗Rsymm a' b) + (isTransR _ _ _ (h b b' a' rb) (∗Rsymm b' a'))) + +effective : (Rprop : isPropValued R) (Requiv : isEquivRel R) + → (a b : A) → [ a ] ≡ [ b ] → R a b +effective {A = A} {R = R} Rprop (equivRel R/refl R/sym R/trans) a b p = + transport aa≡ab (R/refl _) + where + helper : A / R → hProp _ + helper = + rec isSetHProp + (λ c → (R a c , Rprop a c)) + (λ c d cd → + Σ≡Prop (λ _ → isPropIsProp) + (hPropExt (Rprop a c) (Rprop a d) + (λ ac → R/trans _ _ _ ac cd) + (λ ad → R/trans _ _ _ ad (R/sym _ _ cd)))) + + aa≡ab : R a a ≡ R a b + aa≡ab i = helper (p i) .fst + +isEquivRel→effectiveIso : isPropValued R → isEquivRel R + → (a b : A) → Iso ([ a ] ≡ [ b ]) (R a b) +Iso.fun (isEquivRel→effectiveIso {R = R} Rprop Req a b) = effective Rprop Req a b +Iso.inv (isEquivRel→effectiveIso {R = R} Rprop Req a b) = eq/ a b +Iso.rightInv (isEquivRel→effectiveIso {R = R} Rprop Req a b) _ = Rprop a b _ _ +Iso.leftInv (isEquivRel→effectiveIso {R = R} Rprop Req a b) _ = squash/ _ _ _ _ + +isEquivRel→isEffective : isPropValued R → isEquivRel R → isEffective R +isEquivRel→isEffective Rprop Req a b = + isoToIsEquiv (invIso (isEquivRel→effectiveIso Rprop Req a b)) + +discreteSetQuotients : Discrete A → isPropValued R → isEquivRel R + → (∀ a₀ a₁ → Dec (R a₀ a₁)) + → Discrete (A / R) +discreteSetQuotients {A = A} {R = R} Adis Rprop Req Rdec = + elim (λ a₀ → isSetΠ (λ a₁ → isProp→isSet (isPropDec (squash/ a₀ a₁)))) + discreteSetQuotients' discreteSetQuotients'-eq + where + discreteSetQuotients' : (a : A) (y : A / R) → Dec ([ a ] ≡ y) + discreteSetQuotients' a₀ = + elim (λ a₁ → isProp→isSet (isPropDec (squash/ [ a₀ ] a₁))) dis dis-eq + where + dis : (a₁ : A) → Dec ([ a₀ ] ≡ [ a₁ ]) + dis a₁ with Rdec a₀ a₁ + ... | (yes p) = yes (eq/ a₀ a₁ p) + ... | (no ¬p) = no λ eq → ¬p (effective Rprop Req a₀ a₁ eq ) + + dis-eq : (a b : A) (r : R a b) → + PathP (λ i → Dec ([ a₀ ] ≡ eq/ a b r i)) (dis a) (dis b) + dis-eq a b ab = J (λ b ab → ∀ k → PathP (λ i → Dec ([ a₀ ] ≡ ab i)) (dis a) k) + (λ k → isPropDec (squash/ _ _) _ _) (eq/ a b ab) (dis b) + + discreteSetQuotients'-eq : (a b : A) (r : R a b) → + PathP (λ i → (y : A / R) → Dec (eq/ a b r i ≡ y)) + (discreteSetQuotients' a) (discreteSetQuotients' b) + discreteSetQuotients'-eq a b ab = + J (λ b ab → ∀ k → PathP (λ i → (y : A / R) → Dec (ab i ≡ y)) + (discreteSetQuotients' a) k) + (λ k → funExt (λ x → isPropDec (squash/ _ _) _ _)) (eq/ a b ab) (discreteSetQuotients' b) + + +-- Quotienting by the truncated relation is equivalent to quotienting by untruncated relation +truncRelIso : Iso (A / R) (A / (λ a b → ∥ R a b ∥)) +Iso.fun truncRelIso = rec squash/ [_] λ _ _ r → eq/ _ _ ∣ r ∣ +Iso.inv truncRelIso = rec squash/ [_] λ _ _ → PropTrunc.rec (squash/ _ _) λ r → eq/ _ _ r +Iso.rightInv truncRelIso = elimProp (λ _ → squash/ _ _) λ _ → refl +Iso.leftInv truncRelIso = elimProp (λ _ → squash/ _ _) λ _ → refl + +truncRelEquiv : A / R ≃ A / (λ a b → ∥ R a b ∥) +truncRelEquiv = isoToEquiv truncRelIso + +-- Using this we can obtain a useful characterization of +-- path-types for equivalence relations (not prop-valued) +-- and their quotients + +isEquivRel→TruncIso : isEquivRel R → (a b : A) → Iso ([ a ] ≡ [ b ]) ∥ R a b ∥ +isEquivRel→TruncIso {A = A} {R = R} Req a b = + compIso + (isProp→Iso (squash/ _ _) (squash/ _ _) + (cong (Iso.fun truncRelIso)) (cong (Iso.inv truncRelIso))) + (isEquivRel→effectiveIso (λ _ _ → PropTrunc.isPropPropTrunc) ∥R∥eq a b) + where + open isEquivRel + ∥R∥eq : isEquivRel λ a b → ∥ R a b ∥ + reflexive ∥R∥eq a = ∣ reflexive Req a ∣ + symmetric ∥R∥eq a b = PropTrunc.map (symmetric Req a b) + transitive ∥R∥eq a b c = PropTrunc.map2 (transitive Req a b c) + +-- quotienting by 'logically equivalent' relations gives the same quotient +relBiimpl→TruncIso : ({a b : A} → R a b → S a b) → ({a b : A} → S a b → R a b) → Iso (A / R) (A / S) +Iso.fun (relBiimpl→TruncIso R→S S→R) = rec squash/ [_] λ _ _ Rab → eq/ _ _ (R→S Rab) +Iso.inv (relBiimpl→TruncIso R→S S→R) = rec squash/ [_] λ _ _ Sab → eq/ _ _ (S→R Sab) +Iso.rightInv (relBiimpl→TruncIso R→S S→R) = elimProp (λ _ → squash/ _ _) λ _ → refl +Iso.leftInv (relBiimpl→TruncIso R→S S→R) = elimProp (λ _ → squash/ _ _) λ _ → refl diff --git a/Cubical/HITs/SetTruncation.agda b/Cubical/HITs/SetTruncation.agda new file mode 100644 index 0000000000..402435485d --- /dev/null +++ b/Cubical/HITs/SetTruncation.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.SetTruncation where + +open import Cubical.HITs.SetTruncation.Base public +open import Cubical.HITs.SetTruncation.Properties public diff --git a/Cubical/HITs/SetTruncation/Base.agda b/Cubical/HITs/SetTruncation/Base.agda new file mode 100644 index 0000000000..033d865063 --- /dev/null +++ b/Cubical/HITs/SetTruncation/Base.agda @@ -0,0 +1,23 @@ +{- + +This file contains: + +- Definition of set truncations + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.SetTruncation.Base where + +open import Cubical.Core.Primitives +open import Cubical.Foundations.Pointed + +-- set truncation as a higher inductive type: + +data ∥_∥₂ {ℓ} (A : Type ℓ) : Type ℓ where + ∣_∣₂ : A → ∥ A ∥₂ + squash₂ : ∀ (x y : ∥ A ∥₂) (p q : x ≡ y) → p ≡ q + +-- Pointed version +∥_∥₂∙ : ∀ {ℓ} (A : Pointed ℓ) → Pointed ℓ +fst ∥ A ∥₂∙ = ∥ fst A ∥₂ +snd ∥ A ∥₂∙ = ∣ pt A ∣₂ diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda new file mode 100644 index 0000000000..b6ed287133 --- /dev/null +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -0,0 +1,328 @@ +{- + +This file contains: + +- Properties of set truncations + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.SetTruncation.Properties where + +open import Cubical.HITs.SetTruncation.Base + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed.Base +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) hiding (elim2 ; elim3 ; rec2 ; map) + +private + variable + ℓ ℓ' ℓ'' : Level + A B C D : Type ℓ + +isSetPathImplicit : {x y : ∥ A ∥₂} → isSet (x ≡ y) +isSetPathImplicit = isOfHLevelPath 2 squash₂ _ _ + +rec : isSet B → (A → B) → ∥ A ∥₂ → B +rec Bset f ∣ x ∣₂ = f x +rec Bset f (squash₂ x y p q i j) = + Bset _ _ (cong (rec Bset f) p) (cong (rec Bset f) q) i j + +rec2 : isSet C → (A → B → C) → ∥ A ∥₂ → ∥ B ∥₂ → C +rec2 Cset f ∣ x ∣₂ ∣ y ∣₂ = f x y +rec2 Cset f ∣ x ∣₂ (squash₂ y z p q i j) = + Cset _ _ (cong (rec2 Cset f ∣ x ∣₂) p) (cong (rec2 Cset f ∣ x ∣₂) q) i j +rec2 Cset f (squash₂ x y p q i j) z = + Cset _ _ (cong (λ a → rec2 Cset f a z) p) (cong (λ a → rec2 Cset f a z) q) i j + +-- Old version: +-- rec2 Cset f = rec (isSetΠ λ _ → Cset) λ x → rec Cset (f x) + +-- lemma 6.9.1 in HoTT book +elim : {B : ∥ A ∥₂ → Type ℓ} + (Bset : (x : ∥ A ∥₂) → isSet (B x)) + (f : (a : A) → B (∣ a ∣₂)) + (x : ∥ A ∥₂) → B x +elim Bset f ∣ a ∣₂ = f a +elim Bset f (squash₂ x y p q i j) = + isOfHLevel→isOfHLevelDep 2 Bset _ _ + (cong (elim Bset f) p) (cong (elim Bset f) q) (squash₂ x y p q) i j + +elim2 : {C : ∥ A ∥₂ → ∥ B ∥₂ → Type ℓ} + (Cset : ((x : ∥ A ∥₂) (y : ∥ B ∥₂) → isSet (C x y))) + (f : (a : A) (b : B) → C ∣ a ∣₂ ∣ b ∣₂) + (x : ∥ A ∥₂) (y : ∥ B ∥₂) → C x y +elim2 Cset f ∣ x ∣₂ ∣ y ∣₂ = f x y +elim2 Cset f ∣ x ∣₂ (squash₂ y z p q i j) = + isOfHLevel→isOfHLevelDep 2 (λ a → Cset ∣ x ∣₂ a) _ _ + (cong (elim2 Cset f ∣ x ∣₂) p) (cong (elim2 Cset f ∣ x ∣₂) q) (squash₂ y z p q) i j +elim2 Cset f (squash₂ x y p q i j) z = + isOfHLevel→isOfHLevelDep 2 (λ a → Cset a z) _ _ + (cong (λ a → elim2 Cset f a z) p) (cong (λ a → elim2 Cset f a z) q) (squash₂ x y p q) i j + +-- Old version: +-- elim2 Cset f = elim (λ _ → isSetΠ (λ _ → Cset _ _)) +-- (λ a → elim (λ _ → Cset _ _) (f a)) + +-- TODO: generalize +elim3 : {B : (x y z : ∥ A ∥₂) → Type ℓ} + (Bset : ((x y z : ∥ A ∥₂) → isSet (B x y z))) + (g : (a b c : A) → B ∣ a ∣₂ ∣ b ∣₂ ∣ c ∣₂) + (x y z : ∥ A ∥₂) → B x y z +elim3 Bset g = elim2 (λ _ _ → isSetΠ (λ _ → Bset _ _ _)) + (λ a b → elim (λ _ → Bset _ _ _) (g a b)) + + +-- the recursor for maps into groupoids following the "HIT proof" in: +-- https://arxiv.org/abs/1507.01150 +-- i.e. for any type A and groupoid B we can construct a map ∥ A ∥₂ → B +-- from a map A → B satisfying the condition +-- ∀ (a b : A) (p q : a ≡ b) → cong f p ≡ cong f q +-- TODO: prove that this is an equivalence +module rec→Gpd {A : Type ℓ} {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A → B) + (congFConst : ∀ (a b : A) (p q : a ≡ b) → cong f p ≡ cong f q) where + + data H : Type ℓ where + η : A → H + ε : ∀ (a b : A) → ∥ a ≡ b ∥ → η a ≡ η b -- prop. trunc. of a≡b + δ : ∀ (a b : A) (p : a ≡ b) → ε a b ∣ p ∣ ≡ cong η p + gtrunc : isGroupoid H + + -- write elimination principle for H + module Helim {P : H → Type ℓ''} (Pgpd : ∀ h → isGroupoid (P h)) + (η* : (a : A) → P (η a)) + (ε* : ∀ (a b : A) (∣p∣ : ∥ a ≡ b ∥) + → PathP (λ i → P (ε a b ∣p∣ i)) (η* a) (η* b)) + (δ* : ∀ (a b : A) (p : a ≡ b) + → PathP (λ i → PathP (λ j → P (δ a b p i j)) (η* a) (η* b)) + (ε* a b ∣ p ∣) (cong η* p)) where + + fun : (h : H) → P h + fun (η a) = η* a + fun (ε a b ∣p∣ i) = ε* a b ∣p∣ i + fun (δ a b p i j) = δ* a b p i j + fun (gtrunc x y p q α β i j k) = isOfHLevel→isOfHLevelDep 3 Pgpd + (fun x) (fun y) + (cong fun p) (cong fun q) + (cong (cong fun) α) (cong (cong fun) β) + (gtrunc x y p q α β) i j k + + module Hrec {C : Type ℓ''} (Cgpd : isGroupoid C) + (η* : A → C) + (ε* : ∀ (a b : A) → ∥ a ≡ b ∥ → η* a ≡ η* b) + (δ* : ∀ (a b : A) (p : a ≡ b) → ε* a b ∣ p ∣ ≡ cong η* p) where + + fun : H → C + fun (η a) = η* a + fun (ε a b ∣p∣ i) = ε* a b ∣p∣ i + fun (δ a b p i j) = δ* a b p i j + fun (gtrunc x y p q α β i j k) = Cgpd (fun x) (fun y) (cong fun p) (cong fun q) + (cong (cong fun) α) (cong (cong fun) β) i j k + + module HelimProp {P : H → Type ℓ''} (Pprop : ∀ h → isProp (P h)) + (η* : (a : A) → P (η a)) where + + fun : ∀ h → P h + fun = Helim.fun (λ _ → isSet→isGroupoid (isProp→isSet (Pprop _))) η* + (λ a b ∣p∣ → isOfHLevel→isOfHLevelDep 1 Pprop _ _ (ε a b ∣p∣)) + λ a b p → isOfHLevel→isOfHLevelDep 1 + {B = λ p → PathP (λ i → P (p i)) (η* a) (η* b)} + (λ _ → isOfHLevelPathP 1 (Pprop _) _ _) _ _ (δ a b p) + + -- The main trick: eliminating into hsets is easy + -- i.e. H has the universal property of set truncation... + module HelimSet {P : H → Type ℓ''} (Pset : ∀ h → isSet (P h)) + (η* : ∀ a → P (η a)) where + + fun : (h : H) → P h + fun = Helim.fun (λ _ → isSet→isGroupoid (Pset _)) η* ε* + λ a b p → isOfHLevel→isOfHLevelDep 1 + {B = λ p → PathP (λ i → P (p i)) (η* a) (η* b)} + (λ _ → isOfHLevelPathP' 1 (Pset _) _ _) _ _ (δ a b p) + where + ε* : (a b : A) (∣p∣ : ∥ a ≡ b ∥) → PathP (λ i → P (ε a b ∣p∣ i)) (η* a) (η* b) + ε* a b = pElim (λ _ → isOfHLevelPathP' 1 (Pset _) (η* a) (η* b)) + λ p → subst (λ x → PathP (λ i → P (x i)) (η* a) (η* b)) + (sym (δ a b p)) (cong η* p) + + + -- Now we need to prove that H is a set. + -- We start with a little lemma: + localHedbergLemma : {X : Type ℓ''} (P : X → Type ℓ'') + → (∀ x → isProp (P x)) + → ((x : X) → P x → (y : X) → P y → x ≡ y) + -------------------------------------------------- + → (x : X) → P x → (y : X) → isProp (x ≡ y) + localHedbergLemma {X = X} P Pprop P→≡ x px y = isPropRetract + (λ p → subst P p px) (λ py → sym (P→≡ x px x px) ∙ P→≡ x px y py) + isRetract (Pprop y) + where + isRetract : (p : x ≡ y) → (sym (P→≡ x px x px)) ∙ P→≡ x px y (subst P p px) ≡ p + isRetract = J (λ y' p' → (sym (P→≡ x px x px)) ∙ P→≡ x px y' (subst P p' px) ≡ p') + (subst (λ px' → sym (P→≡ x px x px) ∙ P→≡ x px x px' ≡ refl) + (sym (substRefl {B = P} px)) (lCancel (P→≡ x px x px))) + + Hset : isSet H + Hset = HelimProp.fun (λ _ → isPropΠ λ _ → isPropIsProp) baseCaseLeft + where + baseCaseLeft : (a₀ : A) (y : H) → isProp (η a₀ ≡ y) + baseCaseLeft a₀ = localHedbergLemma (λ x → Q x .fst) (λ x → Q x .snd) Q→≡ _ ∣ refl ∣ + where + Q : H → hProp ℓ + Q = HelimSet.fun (λ _ → isSetHProp) λ b → ∥ a₀ ≡ b ∥ , isPropPropTrunc + -- Q (η b) = ∥ a ≡ b ∥ + + Q→≡ : (x : H) → Q x .fst → (y : H) → Q y .fst → x ≡ y + Q→≡ = HelimSet.fun (λ _ → isSetΠ3 λ _ _ _ → gtrunc _ _) + λ a p → HelimSet.fun (λ _ → isSetΠ λ _ → gtrunc _ _) + λ b q → sym (ε a₀ a p) ∙ ε a₀ b q + + -- our desired function will split through H, + -- i.e. we get a function ∥ A ∥₂ → H → B + fun : ∥ A ∥₂ → B + fun = f₁ ∘ f₂ + where + f₁ : H → B + f₁ = Hrec.fun Bgpd f εᶠ λ _ _ _ → refl + where + εᶠ : (a b : A) → ∥ a ≡ b ∥ → f a ≡ f b + εᶠ a b = rec→Set (Bgpd _ _) (cong f) λ p q → congFConst a b p q + -- this is the inductive step, + -- we use that maps ∥ A ∥ → B for an hset B + -- correspond to 2-Constant maps A → B (which cong f is by assumption) + f₂ : ∥ A ∥₂ → H + f₂ = rec Hset η + + +map : (A → B) → ∥ A ∥₂ → ∥ B ∥₂ +map f = rec squash₂ λ x → ∣ f x ∣₂ + +map∙ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (f : A →∙ B) → ∥ A ∥₂∙ →∙ ∥ B ∥₂∙ +fst (map∙ f) = map (fst f) +snd (map∙ f) = cong ∣_∣₂ (snd f) + +setTruncUniversal : isSet B → (∥ A ∥₂ → B) ≃ (A → B) +setTruncUniversal {B = B} Bset = + isoToEquiv (iso (λ h x → h ∣ x ∣₂) (rec Bset) (λ _ → refl) rinv) + where + rinv : (f : ∥ A ∥₂ → B) → rec Bset (λ x → f ∣ x ∣₂) ≡ f + rinv f i x = + elim (λ x → isProp→isSet (Bset (rec Bset (λ x → f ∣ x ∣₂) x) (f x))) + (λ _ → refl) x i + +isSetSetTrunc : isSet ∥ A ∥₂ +isSetSetTrunc a b p q = squash₂ a b p q + +setTruncIdempotentIso : isSet A → Iso ∥ A ∥₂ A +Iso.fun (setTruncIdempotentIso hA) = rec hA (idfun _) +Iso.inv (setTruncIdempotentIso hA) x = ∣ x ∣₂ +Iso.rightInv (setTruncIdempotentIso hA) _ = refl +Iso.leftInv (setTruncIdempotentIso hA) = elim (λ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ _ → refl) + +setTruncIdempotent≃ : isSet A → ∥ A ∥₂ ≃ A +setTruncIdempotent≃ {A = A} hA = isoToEquiv (setTruncIdempotentIso hA) + +setTruncIdempotent : isSet A → ∥ A ∥₂ ≡ A +setTruncIdempotent hA = ua (setTruncIdempotent≃ hA) + +isContr→isContrSetTrunc : isContr A → isContr (∥ A ∥₂) +isContr→isContrSetTrunc contr = ∣ fst contr ∣₂ + , elim (λ _ → isOfHLevelPath 2 (isSetSetTrunc) _ _) + λ a → cong ∣_∣₂ (snd contr a) + + +setTruncIso : Iso A B → Iso ∥ A ∥₂ ∥ B ∥₂ +Iso.fun (setTruncIso is) = rec isSetSetTrunc (λ x → ∣ Iso.fun is x ∣₂) +Iso.inv (setTruncIso is) = rec isSetSetTrunc (λ x → ∣ Iso.inv is x ∣₂) +Iso.rightInv (setTruncIso is) = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ a → cong ∣_∣₂ (Iso.rightInv is a) +Iso.leftInv (setTruncIso is) = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ a → cong ∣_∣₂ (Iso.leftInv is a) + +setSigmaIso : {B : A → Type ℓ} → Iso ∥ Σ A B ∥₂ ∥ Σ A (λ x → ∥ B x ∥₂) ∥₂ +setSigmaIso {A = A} {B = B} = iso fun funinv sect retr + where + {- writing it out explicitly to avoid yellow highlighting -} + fun : ∥ Σ A B ∥₂ → ∥ Σ A (λ x → ∥ B x ∥₂) ∥₂ + fun = rec isSetSetTrunc λ {(a , p) → ∣ a , ∣ p ∣₂ ∣₂} + funinv : ∥ Σ A (λ x → ∥ B x ∥₂) ∥₂ → ∥ Σ A B ∥₂ + funinv = rec isSetSetTrunc (λ {(a , p) → rec isSetSetTrunc (λ p → ∣ a , p ∣₂) p}) + sect : section fun funinv + sect = elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ { (a , p) → elim {B = λ p → fun (funinv ∣ a , p ∣₂) ≡ ∣ a , p ∣₂} + (λ p → isOfHLevelPath 2 isSetSetTrunc _ _) (λ _ → refl) p } + retr : retract fun funinv + retr = elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ { _ → refl } + +sigmaElim : {B : ∥ A ∥₂ → Type ℓ} {C : Σ ∥ A ∥₂ B → Type ℓ'} + (Bset : (x : Σ ∥ A ∥₂ B) → isSet (C x)) + (g : (a : A) (b : B ∣ a ∣₂) → C (∣ a ∣₂ , b)) + (x : Σ ∥ A ∥₂ B) → C x +sigmaElim {B = B} {C = C} set g (x , y) = + elim {B = λ x → (y : B x) → C (x , y)} (λ _ → isSetΠ λ _ → set _) g x y + +sigmaProdElim : {C : ∥ A ∥₂ × ∥ B ∥₂ → Type ℓ} {D : Σ (∥ A ∥₂ × ∥ B ∥₂) C → Type ℓ'} + (Bset : (x : Σ (∥ A ∥₂ × ∥ B ∥₂) C) → isSet (D x)) + (g : (a : A) (b : B) (c : C (∣ a ∣₂ , ∣ b ∣₂)) → D ((∣ a ∣₂ , ∣ b ∣₂) , c)) + (x : Σ (∥ A ∥₂ × ∥ B ∥₂) C) → D x +sigmaProdElim {B = B} {C = C} {D = D} set g ((x , y) , c) = + elim {B = λ x → (y : ∥ B ∥₂) (c : C (x , y)) → D ((x , y) , c)} + (λ _ → isSetΠ λ _ → isSetΠ λ _ → set _) + (λ x → elim (λ _ → isSetΠ λ _ → set _) (g x)) + x y c + +prodElim : {C : ∥ A ∥₂ × ∥ B ∥₂ → Type ℓ} + → ((x : ∥ A ∥₂ × ∥ B ∥₂) → isSet (C x)) + → ((a : A) (b : B) → C (∣ a ∣₂ , ∣ b ∣₂)) + → (x : ∥ A ∥₂ × ∥ B ∥₂) → C x +prodElim setC f (a , b) = elim2 (λ x y → setC (x , y)) f a b + +prodRec : {C : Type ℓ} → isSet C → (A → B → C) → ∥ A ∥₂ × ∥ B ∥₂ → C +prodRec setC f (a , b) = rec2 setC f a b + +prodElim2 : {E : (∥ A ∥₂ × ∥ B ∥₂) → (∥ C ∥₂ × ∥ D ∥₂) → Type ℓ} + → ((x : ∥ A ∥₂ × ∥ B ∥₂) (y : ∥ C ∥₂ × ∥ D ∥₂) → isSet (E x y)) + → ((a : A) (b : B) (c : C) (d : D) → E (∣ a ∣₂ , ∣ b ∣₂) (∣ c ∣₂ , ∣ d ∣₂)) + → ((x : ∥ A ∥₂ × ∥ B ∥₂) (y : ∥ C ∥₂ × ∥ D ∥₂) → (E x y)) +prodElim2 isset f = prodElim (λ _ → isSetΠ λ _ → isset _ _) + λ a b → prodElim (λ _ → isset _ _) + λ c d → f a b c d + +setTruncOfProdIso : Iso ∥ A × B ∥₂ (∥ A ∥₂ × ∥ B ∥₂) +Iso.fun setTruncOfProdIso = rec (isSet× isSetSetTrunc isSetSetTrunc) λ { (a , b) → ∣ a ∣₂ , ∣ b ∣₂ } +Iso.inv setTruncOfProdIso = prodRec isSetSetTrunc λ a b → ∣ a , b ∣₂ +Iso.rightInv setTruncOfProdIso = + prodElim (λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ → refl +Iso.leftInv setTruncOfProdIso = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ {(a , b) → refl} + +IsoSetTruncateSndΣ : {A : Type ℓ} {B : A → Type ℓ'} → Iso ∥ Σ A B ∥₂ ∥ Σ A (λ x → ∥ B x ∥₂) ∥₂ +Iso.fun IsoSetTruncateSndΣ = map λ a → (fst a) , ∣ snd a ∣₂ +Iso.inv IsoSetTruncateSndΣ = rec isSetSetTrunc (uncurry λ x → map λ b → x , b) +Iso.rightInv IsoSetTruncateSndΣ = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry λ a → elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ _ → refl) +Iso.leftInv IsoSetTruncateSndΣ = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ _ → refl + +PathIdTrunc₀Iso : {a b : A} → Iso (∣ a ∣₂ ≡ ∣ b ∣₂) ∥ a ≡ b ∥ +Iso.fun (PathIdTrunc₀Iso {b = b}) p = + transport (λ i → rec {B = TypeOfHLevel _ 1} (isOfHLevelTypeOfHLevel 1) + (λ a → ∥ a ≡ b ∥ , squash) (p (~ i)) .fst) + ∣ refl ∣ +Iso.inv PathIdTrunc₀Iso = pRec (squash₂ _ _) (cong ∣_∣₂) +Iso.rightInv PathIdTrunc₀Iso _ = squash _ _ +Iso.leftInv PathIdTrunc₀Iso _ = squash₂ _ _ _ _ diff --git a/Cubical/HITs/SmashProduct.agda b/Cubical/HITs/SmashProduct.agda new file mode 100644 index 0000000000..00c7b1ceb2 --- /dev/null +++ b/Cubical/HITs/SmashProduct.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.SmashProduct where + +open import Cubical.HITs.SmashProduct.Base public + +-- open import Cubical.HITs.SmashProduct.Properties public diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda new file mode 100644 index 0000000000..59cd416982 --- /dev/null +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -0,0 +1,213 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.SmashProduct.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.Pushout.Base +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.HITs.Wedge +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws + +data Smash {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') : Type (ℓ-max ℓ ℓ') where + basel : Smash A B + baser : Smash A B + proj : (x : typ A) → (y : typ B) → Smash A B + gluel : (a : typ A) → proj a (pt B) ≡ basel + gluer : (b : typ B) → proj (pt A) b ≡ baser + +private + variable + ℓ ℓ' : Level + A B C D : Pointed ℓ + +Smash-map : (f : A →∙ C) (g : B →∙ D) → Smash A B → Smash C D +Smash-map f g basel = basel +Smash-map f g baser = baser +Smash-map (f , fpt) (g , gpt) (proj x y) = proj (f x) (g y) +Smash-map (f , fpt) (g , gpt) (gluel a i) = ((λ j → proj (f a) (gpt j)) ∙ gluel (f a)) i +Smash-map (f , fpt) (g , gpt) (gluer b i) = ((λ j → proj (fpt j) (g b)) ∙ gluer (g b)) i + +-- Commutativity +comm : Smash A B → Smash B A +comm basel = baser +comm baser = basel +comm (proj x y) = proj y x +comm (gluel a i) = gluer a i +comm (gluer b i) = gluel b i + +commK : (x : Smash A B) → comm (comm x) ≡ x +commK basel = refl +commK baser = refl +commK (proj x y) = refl +commK (gluel a x) = refl +commK (gluer b x) = refl + +-- WIP below + +SmashPt : (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +SmashPt A B = (Smash A B , basel) + +SmashPtProj : (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +SmashPtProj A B = Smash A B , (proj (snd A) (snd B)) + +--- Alternative definition + +i∧ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ⋁ B → (typ A) × (typ B) +i∧ {A = A , ptA} {B = B , ptB} (inl x) = x , ptB +i∧ {A = A , ptA} {B = B , ptB} (inr x) = ptA , x +i∧ {A = A , ptA} {B = B , ptB} (push tt i) = ptA , ptB + +_⋀_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +A ⋀ B = Pushout {A = (A ⋁ B)} (λ _ → tt) i∧ + +_⋀∙_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') +A ⋀∙ B = Pushout {A = (A ⋁ B)} (λ _ → tt) i∧ , (inl tt) + + +_⋀→_ : (f : A →∙ C) (g : B →∙ D) → A ⋀ B → C ⋀ D +(f ⋀→ g) (inl tt) = inl tt +((f , fpt) ⋀→ (g , gpt)) (inr (x , x₁)) = inr (f x , g x₁) +_⋀→_ {B = B} {D = D} (f , fpt) (b , gpt) (push (inl x) i) = (push (inl (f x)) ∙ (λ i → inr (f x , gpt (~ i)))) i +_⋀→_ (f , fpt) (g , gpt) (push (inr x) i) = (push (inr (g x)) ∙ (λ i → inr (fpt (~ i) , g x))) i +_⋀→_ {A = A} {C = C} {B = B} {D = D} (f , fpt) (g , gpt) (push (push tt j) i) = + hcomp (λ k → λ { (i = i0) → inl tt + ; (i = i1) → inr (fpt (~ k) , gpt (~ k)) + ; (j = i0) → compPath-filler (push (inl (fpt (~ k)))) + ((λ i → inr (fpt (~ k) , gpt (~ i)))) k i + ; (j = i1) → compPath-filler (push (inr (gpt (~ k)))) + ((λ i → inr (fpt (~ i) , gpt (~ k)))) k i}) + (push (push tt j) i) + +⋀→Smash : A ⋀ B → Smash A B +⋀→Smash (inl x) = basel +⋀→Smash (inr (x , x₁)) = proj x x₁ +⋀→Smash (push (inl x) i) = gluel x (~ i) +⋀→Smash {A = A} {B = B} (push (inr x) i) = (sym (gluel (snd A)) ∙∙ gluer (snd B) ∙∙ sym (gluer x)) i +⋀→Smash {A = A} {B = B} (push (push a j) i) = + hcomp (λ k → λ { (i = i0) → gluel (snd A) (k ∨ ~ j) + ; (i = i1) → gluer (snd B) (~ k ∧ j) + ; (j = i0) → gluel (snd A) (~ i)}) + (invSides-filler (gluel (snd A)) (gluer (snd B)) j (~ i)) + +Smash→⋀ : Smash A B → A ⋀ B +Smash→⋀ basel = inl tt +Smash→⋀ baser = inl tt +Smash→⋀ (proj x y) = inr (x , y) +Smash→⋀ (gluel a i) = push (inl a) (~ i) +Smash→⋀ (gluer b i) = push (inr b) (~ i) + +{- associativity maps for smash produts. Proof pretty much direcly translated from https://github.com/ecavallo/redtt/blob/master/library/pointed/smash.red -} +private + pivotl : (b b' : typ B) + → Path (Smash A B) (proj (snd A) b) (proj (snd A) b') + pivotl b b' i = (gluer b ∙ sym (gluer b')) i + + pivotr : (a a' : typ A) + → Path (Smash A B) (proj a (snd B)) (proj a' (snd B)) + pivotr a a' i = (gluel a ∙ sym (gluel a')) i + + pivotlrId : {A : Pointed ℓ} {B : Pointed ℓ'} → _ + pivotlrId {A = A} {B = B} = rCancel (gluer (snd B)) ∙ sym (rCancel (gluel (snd A))) + + rearrange-proj : (c : fst C) + → (Smash A B) → Smash (SmashPtProj C B) A + rearrange-proj c basel = baser + rearrange-proj c baser = basel + rearrange-proj c (proj x y) = proj (proj c y) x + rearrange-proj {C = C} c (gluel a i) = + hcomp (λ k → λ { (i = i0) → proj (pivotr (snd C) c k) a + ; (i = i1) → baser}) + (gluer a i) + rearrange-proj c (gluer b i) = gluel (proj c b) i + + rearrange-gluel : (s : Smash A B) + → Path (Smash (SmashPtProj C B) A) basel (rearrange-proj (snd C) s) + rearrange-gluel {A = A} {B = B} {C = C} basel = sym (gluel (proj (snd C) (snd B))) ∙ + gluer (snd A) + rearrange-gluel baser = refl + rearrange-gluel {A = A} {B = B} {C = C} (proj a b) i = + hcomp (λ k → λ { (i = i0) → (sym (gluel (proj (snd C) (snd B))) ∙ + gluer (snd A)) (~ k) + ; (i = i1) → proj (pivotl (snd B) b k) a}) + (gluer a (~ i)) + rearrange-gluel {A = A} {B = B} {C = C} (gluel a i) j = + hcomp (λ k → λ { (i = i1) → ((λ i₁ → gluel (proj (snd C) (snd B)) (~ i₁)) ∙ + gluer (snd A)) (~ k ∨ j) + ; (j = i0) → ((λ i₁ → gluel (proj (snd C) (snd B)) (~ i₁)) ∙ + gluer (snd A)) (~ k) + ; (j = i1) → top-filler i k}) + (gluer a (i ∨ ~ j)) + where + top-filler : I → I → Smash (SmashPtProj C B) A + top-filler i j = + hcomp (λ k → λ { (i = i0) → side-filler k j + ; (i = i1) → gluer a (j ∨ k) + ; (j = i0) → gluer a (i ∧ k)}) + (gluer a (i ∧ j)) + where + side-filler : I → I → Smash (SmashPtProj C B) A + side-filler i j = + hcomp (λ k → λ { (i = i0) → proj (proj (snd C) (snd B)) a + ; (i = i1) → proj ((rCancel (gluel (snd C)) ∙ sym (rCancel (gluer (snd B)))) k j) a + ; (j = i0) → proj (proj (snd C) (snd B)) a + ; (j = i1) → (proj ((gluel (snd C) ∙ sym (gluel (snd C))) i) a)}) + (proj ((gluel (snd C) ∙ sym (gluel (snd C))) (j ∧ i)) a) + rearrange-gluel {A = A} {B = B} {C = C} (gluer b i) j = + hcomp (λ k → λ {(i = i1) → ((sym (gluel (proj (snd C) (snd B)))) ∙ gluer (snd A)) (~ k) + ; (j = i0) → ((sym (gluel (proj (snd C) (snd B)))) ∙ gluer (snd A)) (~ k) + ; (j = i1) → top-filler1 i k}) + (gluer (snd A) (i ∨ (~ j))) + where + top-filler1 : I → I → Smash (SmashPtProj C B) A + top-filler1 i j = + hcomp (λ k → λ { (i = i0) → congFunct (λ x → proj x (snd A)) (gluer (snd B)) (sym (gluer b)) (~ k) j + ; (i = i1) → (sym (gluel (proj (snd C) (snd B))) ∙ gluer (snd A)) (~ j) + ; (j = i0) → gluer (snd A) i + ; (j = i1) → gluel (proj (snd C) b) i}) + (top-filler2 i j) + where + top-filler2 : I → I → Smash (SmashPtProj C B) A + top-filler2 i j = + hcomp (λ k → λ { (j = i0) → gluer (snd A) (i ∧ k) + ; (j = i1) → gluel (gluer b (~ k)) i}) + (hcomp (λ k → λ { (j = i0) → gluel (gluer (snd B) i0) (~ k ∧ (~ i)) + ; (j = i1) → gluel (baser) (~ k ∨ i) + ; (i = i0) → gluel (gluer (snd B) j) (~ k) + ; (i = i1) → gluel (proj (snd C) (snd B)) j }) + (gluel (proj (snd C) (snd B)) (j ∨ (~ i)))) + + rearrange : Smash (SmashPtProj A B) C → Smash (SmashPtProj C B) A + rearrange basel = basel + rearrange baser = baser + rearrange (proj x y) = rearrange-proj y x + rearrange (gluel a i) = rearrange-gluel a (~ i) + rearrange {A = A} {B = B} {C = C} (gluer b i) = ((λ j → proj (pivotr b (snd C) j) (snd A)) ∙ + gluer (snd A)) i + + ⋀∙→SmashPtProj : (A ⋀∙ B) →∙ SmashPtProj A B + ⋀∙→SmashPtProj {A = A} {B = B} = fun , refl + where + fun : (A ⋀ B) → Smash A B + fun (inl x) = proj (snd A) (snd B) + fun (inr (x , x₁)) = proj x x₁ + fun (push (inl x) i) = pivotr (snd A) x i + fun (push (inr x) i) = pivotl (snd B) x i + fun (push (push a j) i) = pivotlrId (~ j) i + + SmashPtProj→⋀∙ : (SmashPtProj A B) →∙ (A ⋀∙ B) + SmashPtProj→⋀∙ {A = A} {B = B} = Smash→⋀ , sym (push (inr (snd B))) + +SmashAssociate : Smash (SmashPtProj A B) C → Smash A (SmashPtProj B C) +SmashAssociate = comm ∘ Smash-map (comm , refl) (idfun∙ _) ∘ rearrange + +SmashAssociate⁻ : Smash A (SmashPtProj B C) → Smash (SmashPtProj A B) C +SmashAssociate⁻ = rearrange ∘ comm ∘ Smash-map (idfun∙ _) (comm , refl) + +⋀-associate : (A ⋀∙ B) ⋀ C → A ⋀ (B ⋀∙ C) +⋀-associate = (idfun∙ _ ⋀→ SmashPtProj→⋀∙) ∘ Smash→⋀ ∘ SmashAssociate ∘ ⋀→Smash ∘ (⋀∙→SmashPtProj ⋀→ idfun∙ _) + +⋀-associate⁻ : A ⋀ (B ⋀∙ C) → (A ⋀∙ B) ⋀ C +⋀-associate⁻ = (SmashPtProj→⋀∙ ⋀→ idfun∙ _) ∘ Smash→⋀ ∘ SmashAssociate⁻ ∘ ⋀→Smash ∘ (idfun∙ _ ⋀→ ⋀∙→SmashPtProj) diff --git a/Cubical/HITs/Sn.agda b/Cubical/HITs/Sn.agda new file mode 100644 index 0000000000..069cfd5e0c --- /dev/null +++ b/Cubical/HITs/Sn.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Sn where + +open import Cubical.HITs.Sn.Base public +open import Cubical.HITs.Sn.Properties public diff --git a/Cubical/HITs/Sn/Base.agda b/Cubical/HITs/Sn/Base.agda new file mode 100644 index 0000000000..9234d84ae1 --- /dev/null +++ b/Cubical/HITs/Sn/Base.agda @@ -0,0 +1,29 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Sn.Base where + +open import Cubical.HITs.Susp.Base +open import Cubical.Foundations.Pointed +open import Cubical.Data.Nat +open import Cubical.Data.NatMinusOne +open import Cubical.Data.Empty +open import Cubical.Foundations.Prelude +open import Cubical.Data.Bool +open import Cubical.HITs.S1 + +S : ℕ₋₁ → Type₀ +S neg1 = ⊥ +S (ℕ→ℕ₋₁ n) = Susp (S (-1+ n)) + +S₊ : ℕ → Type₀ +S₊ 0 = Bool +S₊ 1 = S¹ +S₊ (suc (suc n)) = Susp (S₊ (suc n)) + +ptSn : (n : ℕ) → S₊ n +ptSn zero = true +ptSn (suc zero) = base +ptSn (suc (suc n)) = north + +-- Pointed version +S₊∙ : (n : ℕ) → Pointed₀ +S₊∙ n = (S₊ n) , ptSn n diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda new file mode 100644 index 0000000000..43add34f73 --- /dev/null +++ b/Cubical/HITs/Sn/Properties.agda @@ -0,0 +1,476 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Sn.Properties where + +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Path +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.Sigma +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation +-- open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Join +open import Cubical.Data.Bool + +private + variable + ℓ : Level + +IsoSucSphereSusp : (n : ℕ) → Iso (S₊ (suc n)) (Susp (S₊ n)) +IsoSucSphereSusp zero = S¹IsoSuspBool +IsoSucSphereSusp (suc n) = idIso + +IsoSucSphereSusp∙ : (n : ℕ) + → Iso.inv (IsoSucSphereSusp n) north ≡ ptSn (suc n) +IsoSucSphereSusp∙ zero = refl +IsoSucSphereSusp∙ (suc n) = refl + +-- Elimination principles for spheres +sphereElim : (n : ℕ) {A : (S₊ (suc n)) → Type ℓ} → ((x : S₊ (suc n)) → isOfHLevel (suc n) (A x)) + → A (ptSn (suc n)) + → (x : S₊ (suc n)) → A x +sphereElim zero hlev pt = toPropElim hlev pt +sphereElim (suc n) hlev pt north = pt +sphereElim (suc n) {A = A} hlev pt south = subst A (merid (ptSn (suc n))) pt +sphereElim (suc n) {A = A} hlev pt (merid a i) = + sphereElim n {A = λ a → PathP (λ i → A (merid a i)) pt (subst A (merid (ptSn (suc n))) pt)} + (λ a → isOfHLevelPathP' (suc n) (hlev south) _ _) + (λ i → transp (λ j → A (merid (ptSn (suc n)) (i ∧ j))) (~ i) pt) + a i + +sphereElim2 : ∀ {ℓ} (n : ℕ) {A : (S₊ (suc n)) → (S₊ (suc n)) → Type ℓ} + → ((x y : S₊ (suc n)) → isOfHLevel (suc n) (A x y)) + → A (ptSn (suc n)) (ptSn (suc n)) + → (x y : S₊ (suc n)) → A x y +sphereElim2 n hlev pt = sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → hlev _ _) + (sphereElim n (hlev _ ) pt) + +private + compPath-lem : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : z ≡ y) + → PathP (λ i → (p ∙ sym q) i ≡ y) p q + compPath-lem {y = y} p q i j = + hcomp (λ k → λ { (i = i0) → p j + ; (i = i1) → q (~ k ∨ j) + ; (j = i1) → y }) + (p (j ∨ i)) + +sphereToPropElim : (n : ℕ) {A : (S₊ (suc n)) → Type ℓ} → ((x : S₊ (suc n)) → isProp (A x)) + → A (ptSn (suc n)) + → (x : S₊ (suc n)) → A x +sphereToPropElim zero = toPropElim +sphereToPropElim (suc n) hlev pt north = pt +sphereToPropElim (suc n) {A = A} hlev pt south = subst A (merid (ptSn (suc n))) pt +sphereToPropElim (suc n) {A = A} hlev pt (merid a i) = + isProp→PathP {B = λ i → A (merid a i)} (λ _ → hlev _) pt (subst A (merid (ptSn (suc n))) pt) i + +-- Elimination rule for fibrations (x : Sⁿ) → (y : Sᵐ) → A x y of h-Level (n + m). +-- The following principle is just the special case of the "Wedge Connectivity Lemma" +-- for spheres (See Cubical.Homotopy.WedgeConnectivity or chapter 8.6 in the HoTT book). +-- We prove it directly here for three reasons: +-- (i) it should perform better +-- (ii) we get a slightly stronger statement for spheres: one of the homotopies will, by design, be refl +-- (iii) the fact that the two homotopies only differ by (composition with) the homotopy leftFunction(base) ≡ rightFunction(base) +-- is close to trivial + +wedgeconFun : (n m : ℕ) {A : (S₊ (suc n)) → (S₊ (suc m)) → Type ℓ} + → ((x : S₊ (suc n)) (y : S₊ (suc m)) → isOfHLevel ((suc n) + (suc m)) (A x y)) + → (f : (x : _) → A (ptSn (suc n)) x) + → (g : (x : _) → A x (ptSn (suc m))) + → (g (ptSn (suc n)) ≡ f (ptSn (suc m))) + → (x : S₊ (suc n)) (y : S₊ (suc m)) → A x y +wedgeconLeft : (n m : ℕ) {A : (S₊ (suc n)) → (S₊ (suc m)) → Type ℓ} + → (hLev : ((x : S₊ (suc n)) (y : S₊ (suc m)) → isOfHLevel ((suc n) + (suc m)) (A x y))) + → (f : (x : _) → A (ptSn (suc n)) x) + → (g : (x : _) → A x (ptSn (suc m))) + → (hom : g (ptSn (suc n)) ≡ f (ptSn (suc m))) + → (x : _) → wedgeconFun n m hLev f g hom (ptSn (suc n)) x ≡ f x +wedgeconRight : (n m : ℕ) {A : (S₊ (suc n)) → (S₊ (suc m)) → Type ℓ} + → (hLev : ((x : S₊ (suc n)) (y : S₊ (suc m)) → isOfHLevel ((suc n) + (suc m)) (A x y))) + → (f : (x : _) → A (ptSn (suc n)) x) + → (g : (x : _) → A x (ptSn (suc m))) + → (hom : g (ptSn (suc n)) ≡ f (ptSn (suc m))) + → (x : _) → wedgeconFun n m hLev f g hom x (ptSn (suc m)) ≡ g x +wedgeconFun zero zero {A = A} hlev f g hom = F + where + helper : SquareP (λ i j → A (loop i) (loop j)) (cong f loop) (cong f loop) + (λ i → hcomp (λ k → λ { (i = i0) → hom k + ; (i = i1) → hom k }) + (g (loop i))) + λ i → hcomp (λ k → λ { (i = i0) → hom k + ; (i = i1) → hom k }) + (g (loop i)) + helper = toPathP (isOfHLevelPathP' 1 (hlev _ _) _ _ _ _) + + F : (x y : S¹) → A x y + F base y = f y + F (loop i) base = hcomp (λ k → λ { (i = i0) → hom k + ; (i = i1) → hom k }) + (g (loop i)) + F (loop i) (loop j) = helper i j + +wedgeconFun zero (suc m) {A = A} hlev f g hom = F₀ + module _ where + transpLemma₀ : (x : S₊ (suc m)) → transport (λ i₁ → A base (merid x i₁)) (g base) ≡ f south + transpLemma₀ x = cong (transport (λ i₁ → A base (merid x i₁))) + hom + ∙ (λ i → transp (λ j → A base (merid x (i ∨ j))) i + (f (merid x i))) + + pathOverMerid₀ : (x : S₊ (suc m)) → PathP (λ i₁ → A base (merid x i₁)) + (g base) + (transport (λ i₁ → A base (merid (ptSn (suc m)) i₁)) + (g base)) + pathOverMerid₀ x i = hcomp (λ k → λ { (i = i0) → g base + ; (i = i1) → (transpLemma₀ x ∙ sym (transpLemma₀ (ptSn (suc m)))) k}) + (transp (λ i₁ → A base (merid x (i₁ ∧ i))) (~ i) + (g base)) + + pathOverMeridId₀ : pathOverMerid₀ (ptSn (suc m)) ≡ λ i → transp (λ i₁ → A base (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g base) + pathOverMeridId₀ = + (λ j i → hcomp (λ k → λ {(i = i0) → g base + ; (i = i1) → rCancel (transpLemma₀ (ptSn (suc m))) j k}) + (transp (λ i₁ → A base (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g base))) + ∙ λ j i → hfill (λ k → λ { (i = i0) → g base + ; (i = i1) → transport (λ i₁ → A base (merid (ptSn (suc m)) i₁)) + (g base)}) + (inS (transp (λ i₁ → A base (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g base))) (~ j) + + indStep₀ : (x : _) (a : _) → PathP (λ i → A x (merid a i)) + (g x) + (subst (λ y → A x y) (merid (ptSn (suc m))) + (g x)) + indStep₀ = wedgeconFun zero m (λ _ _ → isOfHLevelPathP' (2 + m) (hlev _ _) _ _) + pathOverMerid₀ + (λ a i → transp (λ i₁ → A a (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g a)) + (sym pathOverMeridId₀) + + F₀ : (x : S¹) (y : Susp (S₊ (suc m))) → A x y + F₀ x north = g x + F₀ x south = subst (λ y → A x y) (merid (ptSn (suc m))) (g x) + F₀ x (merid a i) = indStep₀ x a i +wedgeconFun (suc n) m {A = A} hlev f g hom = F₁ + module _ where + transpLemma₁ : (x : S₊ (suc n)) → transport (λ i₁ → A (merid x i₁) (ptSn (suc m))) (f (ptSn (suc m))) ≡ g south + transpLemma₁ x = cong (transport (λ i₁ → A (merid x i₁) (ptSn (suc m)))) + (sym hom) + ∙ (λ i → transp (λ j → A (merid x (i ∨ j)) (ptSn (suc m))) i + (g (merid x i))) + + pathOverMerid₁ : (x : S₊ (suc n)) → PathP (λ i₁ → A (merid x i₁) (ptSn (suc m))) + (f (ptSn (suc m))) + (transport (λ i₁ → A (merid (ptSn (suc n)) i₁) (ptSn (suc m))) + (f (ptSn (suc m)))) + pathOverMerid₁ x i = hcomp (λ k → λ { (i = i0) → f (ptSn (suc m)) + ; (i = i1) → (transpLemma₁ x ∙ sym (transpLemma₁ (ptSn (suc n)))) k }) + (transp (λ i₁ → A (merid x (i₁ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m)))) + + pathOverMeridId₁ : pathOverMerid₁ (ptSn (suc n)) ≡ λ i → transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m))) + pathOverMeridId₁ = + (λ j i → hcomp (λ k → λ { (i = i0) → f (ptSn (suc m)) + ; (i = i1) → rCancel (transpLemma₁ (ptSn (suc n))) j k }) + (transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m))))) + ∙ λ j i → hfill (λ k → λ { (i = i0) → f (ptSn (suc m)) + ; (i = i1) → transport (λ i₁ → A (merid (ptSn (suc n)) i₁) (ptSn (suc m))) + (f (ptSn (suc m))) }) + (inS (transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m))))) (~ j) + + indStep₁ : (a : _) (y : _) → PathP (λ i → A (merid a i) y) + (f y) + (subst (λ x → A x y) (merid (ptSn (suc n))) + (f y)) + indStep₁ = wedgeconFun n m (λ _ _ → isOfHLevelPathP' (suc (n + suc m)) (hlev _ _) _ _) + (λ a i → transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) a) (~ i) + (f a)) + pathOverMerid₁ + pathOverMeridId₁ + + F₁ : (x : Susp (S₊ (suc n))) (y : S₊ (suc m)) → A x y + F₁ north y = f y + F₁ south y = subst (λ x → A x y) (merid (ptSn (suc n))) (f y) + F₁ (merid a i) y = indStep₁ a y i +wedgeconRight zero zero {A = A} hlev f g hom = right + where + right : (x : S¹) → _ + right base = sym hom + right (loop i) j = hcomp (λ k → λ { (i = i0) → hom (~ j ∧ k) + ; (i = i1) → hom (~ j ∧ k) + ; (j = i1) → g (loop i) }) + (g (loop i)) +wedgeconRight zero (suc m) {A = A} hlev f g hom x = refl +wedgeconRight (suc n) m {A = A} hlev f g hom = right + where + lem : (x : _) → indStep₁ n m hlev f g hom x (ptSn (suc m)) ≡ _ + lem = wedgeconRight n m (λ _ _ → isOfHLevelPathP' (suc (n + suc m)) (hlev _ _) _ _) + (λ a i → transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) a) (~ i) + (f a)) + (pathOverMerid₁ n m hlev f g hom) + (pathOverMeridId₁ n m hlev f g hom) + + right : (x : Susp (S₊ (suc n))) → _ ≡ g x + right north = sym hom + right south = cong (subst (λ x → A x (ptSn (suc m))) + (merid (ptSn (suc n)))) + (sym hom) + ∙ λ i → transp (λ j → A (merid (ptSn (suc n)) (i ∨ j)) (ptSn (suc m))) i + (g (merid (ptSn (suc n)) i)) + right (merid a i) j = + hcomp (λ k → λ { (i = i0) → hom (~ j) + ; (i = i1) → transpLemma₁ n m hlev f g hom (ptSn (suc n)) j + ; (j = i0) → lem a (~ k) i + ; (j = i1) → g (merid a i)}) + (hcomp (λ k → λ { (i = i0) → hom (~ j) + ; (i = i1) → compPath-lem (transpLemma₁ n m hlev f g hom a) (transpLemma₁ n m hlev f g hom (ptSn (suc n))) k j + ; (j = i1) → g (merid a i)}) + (hcomp (λ k → λ { (i = i0) → hom (~ j) + ; (j = i0) → transp (λ i₂ → A (merid a (i₂ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m))) + ; (j = i1) → transp (λ j → A (merid a (i ∧ (j ∨ k))) (ptSn (suc m))) (k ∨ ~ i) + (g (merid a (i ∧ k))) }) + (transp (λ i₂ → A (merid a (i₂ ∧ i)) (ptSn (suc m))) (~ i) + (hom (~ j))))) +wedgeconLeft zero zero {A = A} hlev f g hom x = refl +wedgeconLeft zero (suc m) {A = A} hlev f g hom = help + where + left₁ : (x : _) → indStep₀ m hlev f g hom base x ≡ _ + left₁ = wedgeconLeft zero m (λ _ _ → isOfHLevelPathP' (2 + m) (hlev _ _) _ _) + (pathOverMerid₀ m hlev f g hom) + (λ a i → transp (λ i₁ → A a (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g a)) + (sym (pathOverMeridId₀ m hlev f g hom)) + + help : (x : S₊ (suc (suc m))) → _ + help north = hom + help south = cong (subst (A base) (merid (ptSn (suc m)))) hom + ∙ λ i → transp (λ j → A base (merid (ptSn (suc m)) (i ∨ j))) i + (f (merid (ptSn (suc m)) i)) + help (merid a i) j = + hcomp (λ k → λ { (i = i0) → hom j + ; (i = i1) → transpLemma₀ m hlev f g hom (ptSn (suc m)) j + ; (j = i0) → left₁ a (~ k) i + ; (j = i1) → f (merid a i)}) + (hcomp (λ k → λ { (i = i0) → hom j + ; (i = i1) → compPath-lem (transpLemma₀ m hlev f g hom a) + (transpLemma₀ m hlev f g hom (ptSn (suc m))) k j + ; (j = i1) → f (merid a i)}) + (hcomp (λ k → λ { (i = i0) → hom j + ; (j = i0) → transp (λ i₂ → A base (merid a (i₂ ∧ i))) (~ i) + (g base) + ; (j = i1) → transp (λ j → A base (merid a (i ∧ (j ∨ k)))) (k ∨ ~ i) + (f (merid a (i ∧ k)))}) + (transp (λ i₂ → A base (merid a (i₂ ∧ i))) (~ i) + (hom j)))) +wedgeconLeft (suc n) m {A = A} hlev f g hom _ = refl + +---------- Connectedness ----------- + +sphereConnected : (n : HLevel) → isConnected (suc n) (S₊ n) +sphereConnected n = ∣ ptSn n ∣ , elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (λ a → sym (spoke ∣_∣ (ptSn n)) ∙ spoke ∣_∣ a) + +-- The fact that path spaces of Sn are connected can be proved directly for Sⁿ. +-- (Unfortunately, this does not work for higher paths) +pathIdTruncSⁿ : (n : ℕ) (x y : S₊ (suc n)) + → Path (hLevelTrunc (2 + n) (S₊ (suc n))) ∣ x ∣ ∣ y ∣ + → hLevelTrunc (suc n) (x ≡ y) +pathIdTruncSⁿ n = sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) + (sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) + λ _ → ∣ refl ∣) + +pathIdTruncSⁿ⁻ : (n : ℕ) (x y : S₊ (suc n)) + → hLevelTrunc (suc n) (x ≡ y) + → Path (hLevelTrunc (2 + n) (S₊ (suc n))) ∣ x ∣ ∣ y ∣ +pathIdTruncSⁿ⁻ n x y = rec (isOfHLevelTrunc (2 + n) _ _) + (J (λ y _ → Path (hLevelTrunc (2 + n) (S₊ (suc n))) ∣ x ∣ ∣ y ∣) refl) + +pathIdTruncSⁿretract : (n : ℕ) (x y : S₊ (suc n)) → (p : hLevelTrunc (suc n) (x ≡ y)) → pathIdTruncSⁿ n x y (pathIdTruncSⁿ⁻ n x y p) ≡ p +pathIdTruncSⁿretract n = + sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ y → elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (J (λ y p → pathIdTruncSⁿ n (ptSn (suc n)) y (pathIdTruncSⁿ⁻ n (ptSn (suc n)) y ∣ p ∣) ≡ ∣ p ∣) + (cong (pathIdTruncSⁿ n (ptSn (suc n)) (ptSn (suc n))) (transportRefl refl) ∙ pm-help n)) + where + pm-help : (n : ℕ) → pathIdTruncSⁿ n (ptSn (suc n)) (ptSn (suc n)) refl ≡ ∣ refl ∣ + pm-help zero = refl + pm-help (suc n) = refl + +isConnectedPathSⁿ : (n : ℕ) (x y : S₊ (suc n)) → isConnected (suc n) (x ≡ y) +isConnectedPathSⁿ n x y = + isContrRetract + (pathIdTruncSⁿ⁻ n x y) + (pathIdTruncSⁿ n x y) + (pathIdTruncSⁿretract n x y) + ((isContr→isProp (sphereConnected (suc n)) ∣ x ∣ ∣ y ∣) + , isProp→isSet (isContr→isProp (sphereConnected (suc n))) _ _ _) + + +-- Equivalence Sⁿ*Sᵐ≃Sⁿ⁺ᵐ⁺¹ +IsoSphereJoin : (n m : ℕ) + → Iso (join (S₊ n) (S₊ m)) (S₊ (suc (n + m))) +IsoSphereJoin zero m = + compIso join-comm + (compIso (invIso Susp-iso-joinBool) + (invIso (IsoSucSphereSusp m))) +IsoSphereJoin (suc n) m = + compIso (Iso→joinIso + (compIso (pathToIso (cong S₊ (cong suc (+-comm zero n)))) + (invIso (IsoSphereJoin n 0))) + idIso) + (compIso (equivToIso joinAssocDirect) + (compIso (Iso→joinIso idIso + (compIso join-comm + (compIso (invIso Susp-iso-joinBool) + (invIso (IsoSucSphereSusp m))))) + (compIso + (IsoSphereJoin n (suc m)) + (pathToIso λ i → S₊ (suc (+-suc n m i)))))) + +IsoSphereJoinPres∙' : (n m : ℕ) + → Iso.fun (IsoSphereJoin n m) (inl (ptSn n)) ≡ ptSn (suc (n + m)) +IsoSphereJoinPres∙' zero zero = refl +IsoSphereJoinPres∙' zero (suc m) = refl +IsoSphereJoinPres∙' (suc n) m = + cong (transport (λ i → S₊ (suc (+-suc n m i)))) + (cong (Iso.fun (IsoSphereJoin n (suc m))) + (cong (join→ (idfun (S₊ n)) + (λ x → + Iso.inv (IsoSucSphereSusp m) + (Iso.inv Susp-iso-joinBool (join-commFun x)))) + (cong (joinAssocDirect {C = S₊ m} .fst) + (cong (inl ∘ (Iso.inv (IsoSphereJoin n 0))) + (transportS∙ (suc n) _ (cong suc (+-comm 0 n))) + ∙ cong inl (sym (cong (Iso.inv (IsoSphereJoin n 0)) + (IsoSphereJoinPres∙' n 0)) + ∙ Iso.leftInv (IsoSphereJoin n 0) (inl (ptSn _)))))) + ∙ IsoSphereJoinPres∙' n (suc m)) + ∙ transportS∙ _ _ (cong suc (+-suc n m)) + where + transportS∙ : (n m : ℕ) (p : n ≡ m) → transport (λ i → S₊ (p i)) (ptSn n) + ≡ ptSn _ + transportS∙ zero m = + J (λ m p → transport (λ i → S₊ (p i)) true ≡ ptSn m) refl + transportS∙ (suc zero) m = + J (λ m p → transport (λ i → S₊ (p i)) base ≡ ptSn m) refl + transportS∙ (suc (suc n)) m = + J (λ m p → transport (λ i → S₊ (p i)) north ≡ ptSn m) refl + +IsoSphereJoinPres∙ : (n m : ℕ) + → Iso.fun (IsoSphereJoin n m) (inl (ptSn n)) ≡ ptSn (suc (n + m)) +IsoSphereJoinPres∙ zero zero = refl +IsoSphereJoinPres∙ zero (suc m) = refl +IsoSphereJoinPres∙ (suc n) zero = IsoSphereJoinPres∙' (suc n) zero +IsoSphereJoinPres∙ (suc zero) (suc zero) = refl +IsoSphereJoinPres∙ (suc zero) (suc (suc m)) = IsoSphereJoinPres∙' (suc zero) (suc (suc m)) +IsoSphereJoinPres∙ (suc (suc n)) (suc m) = IsoSphereJoinPres∙' (suc (suc n)) (suc m) + +{- +IsoSphereJoinPres∙ : (n m : ℕ) + → Iso.fun (IsoSphereJoin n m) (inl (ptSn n)) ≡ ptSn (suc (n + m)) +IsoSphereJoinPres∙ zero zero = refl +IsoSphereJoinPres∙ zero (suc m) = refl +IsoSphereJoinPres∙ (suc n) m = + cong (transport (λ i → S₊ (suc (+-suc n m i)))) + (cong (Iso.fun (IsoSphereJoin n (suc m))) + (cong (join→ (idfun (S₊ n)) + (λ x → + Iso.inv (IsoSucSphereSusp m) + (Iso.inv Susp-iso-joinBool (join-commFun x)))) + (cong (joinAssocDirect {C = S₊ m} .fst) + (cong (inl ∘ (Iso.inv (IsoSphereJoin n 0))) + (transportS∙ (suc n) _ (cong suc (+-comm 0 n))) + ∙ cong inl (sym (cong (Iso.inv (IsoSphereJoin n 0)) + (IsoSphereJoinPres∙ n 0)) + ∙ Iso.leftInv (IsoSphereJoin n 0) (inl (ptSn _)))))) + ∙ IsoSphereJoinPres∙ n (suc m)) + ∙ transportS∙ _ _ (cong suc (+-suc n m)) + where + transportS∙ : (n m : ℕ) (p : n ≡ m) → transport (λ i → S₊ (p i)) (ptSn n) + ≡ ptSn _ + transportS∙ zero m = + J (λ m p → transport (λ i → S₊ (p i)) true ≡ ptSn m) refl + transportS∙ (suc zero) m = + J (λ m p → transport (λ i → S₊ (p i)) base ≡ ptSn m) refl + transportS∙ (suc (suc n)) m = + J (λ m p → transport (λ i → S₊ (p i)) north ≡ ptSn m) refl +-} + +IsoSphereJoin⁻Pres∙ : (n m : ℕ) + → Iso.inv (IsoSphereJoin n m) (ptSn (suc (n + m))) ≡ inl (ptSn n) +IsoSphereJoin⁻Pres∙ n m = + cong (Iso.inv (IsoSphereJoin n m)) (sym (IsoSphereJoinPres∙ n m)) + ∙ Iso.leftInv (IsoSphereJoin n m) (inl (ptSn n)) + +-- Some lemmas on the H +rUnitS¹ : (x : S¹) → x * base ≡ x +rUnitS¹ base = refl +rUnitS¹ (loop i₁) = refl + +commS¹ : (a x : S¹) → a * x ≡ x * a +commS¹ = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) + (sym ∘ rUnitS¹) + rUnitS¹ + refl + +SuspS¹-hom : (a x : S¹) + → Path (Path (hLevelTrunc 4 (S₊ 2)) _ _) + (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) + (cong ∣_∣ₕ (merid a ∙ sym (merid base)) + ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) +SuspS¹-hom = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) + (λ x → lUnit _ + ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) + (λ x → (λ i → cong ∣_∣ₕ (merid (rUnitS¹ x i) ∙ sym (merid base))) + ∙∙ rUnit _ + ∙∙ cong (cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙_) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) + (sym (l (cong ∣_∣ₕ (merid base ∙ sym (merid base))) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))))) + where + l : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (P : refl ≡ p) + → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P + l p = J (λ p P → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P) refl + +rCancelS¹ : (x : S¹) → ptSn 1 ≡ x * (invLooper x) +rCancelS¹ base = refl +rCancelS¹ (loop i) j = + hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) + base + +SuspS¹-inv : (x : S¹) → Path (Path (hLevelTrunc 4 (S₊ 2)) _ _) + (cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) + (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) +SuspS¹-inv x = (lUnit _ + ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) + (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) + ∙∙ sym (assoc _ _ _)) + ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) lem + ∙∙ (assoc _ _ _ + ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) + (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) + ∙∙ sym (lUnit _)) + where + lem : cong ∣_∣ₕ (merid x ∙ sym (merid base)) + ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) + ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) + ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) + lem = sym (SuspS¹-hom x (invLooper x)) + ∙ ((λ i → cong ∣_∣ₕ (merid (rCancelS¹ x (~ i)) ∙ sym (merid base))) + ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) diff --git a/Cubical/HITs/Susp.agda b/Cubical/HITs/Susp.agda new file mode 100644 index 0000000000..fe041e6f5c --- /dev/null +++ b/Cubical/HITs/Susp.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Susp where + +open import Cubical.HITs.Susp.Base public + +open import Cubical.HITs.Susp.Properties public diff --git a/Cubical/HITs/Susp/Base.agda b/Cubical/HITs/Susp/Base.agda new file mode 100644 index 0000000000..c774f268d6 --- /dev/null +++ b/Cubical/HITs/Susp/Base.agda @@ -0,0 +1,195 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Susp.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed + +open import Cubical.Data.Bool +open import Cubical.Data.Empty + +open import Cubical.HITs.S1 +open import Cubical.HITs.S2 +open import Cubical.HITs.S3 + +open Iso + +private + variable + ℓ ℓ′ : Level + +data Susp (A : Type ℓ) : Type ℓ where + north : Susp A + south : Susp A + merid : (a : A) → north ≡ south + +Susp∙ : (A : Type ℓ) → Pointed ℓ +Susp∙ A = Susp A , north + +-- induced function +suspFun : {A : Type ℓ} {B : Type ℓ′} (f : A → B) + → Susp A → Susp B +suspFun f north = north +suspFun f south = south +suspFun f (merid a i) = merid (f a) i + +BoolIsoSusp⊥ : Iso Bool (Susp ⊥) +fun BoolIsoSusp⊥ = λ {true → north; false → south} +inv BoolIsoSusp⊥ = λ {north → true; south → false} +rightInv BoolIsoSusp⊥ = λ {north → refl; south → refl} +leftInv BoolIsoSusp⊥ = λ {true → refl; false → refl} + +Bool≃Susp⊥ : Bool ≃ Susp ⊥ +Bool≃Susp⊥ = isoToEquiv BoolIsoSusp⊥ + +SuspBool : Type₀ +SuspBool = Susp Bool + +SuspBool→S¹ : SuspBool → S¹ +SuspBool→S¹ north = base +SuspBool→S¹ south = base +SuspBool→S¹ (merid false i) = loop i +SuspBool→S¹ (merid true i) = base + +S¹→SuspBool : S¹ → SuspBool +S¹→SuspBool base = north +S¹→SuspBool (loop i) = (merid false ∙ sym (merid true)) i + +SuspBool→S¹→SuspBool : (x : SuspBool) → Path _ (S¹→SuspBool (SuspBool→S¹ x)) x +SuspBool→S¹→SuspBool north = refl +SuspBool→S¹→SuspBool south = merid true +SuspBool→S¹→SuspBool (merid false i) j = hcomp (λ k → (λ { (j = i1) → merid false i + ; (i = i0) → north + ; (i = i1) → merid true (j ∨ ~ k)})) + (merid false i) +SuspBool→S¹→SuspBool (merid true i) j = merid true (i ∧ j) + +S¹→SuspBool→S¹ : (x : S¹) → SuspBool→S¹ (S¹→SuspBool x) ≡ x +S¹→SuspBool→S¹ base = refl +S¹→SuspBool→S¹ (loop i) j = hfill (λ k → λ { (i = i0) → base + ; (i = i1) → base }) + (inS (loop i)) (~ j) + +S¹IsoSuspBool : Iso S¹ SuspBool +fun S¹IsoSuspBool = S¹→SuspBool +inv S¹IsoSuspBool = SuspBool→S¹ +rightInv S¹IsoSuspBool = SuspBool→S¹→SuspBool +leftInv S¹IsoSuspBool = S¹→SuspBool→S¹ + +S¹≃SuspBool : S¹ ≃ SuspBool +S¹≃SuspBool = isoToEquiv S¹IsoSuspBool + +S¹≡SuspBool : S¹ ≡ SuspBool +S¹≡SuspBool = ua S¹≃SuspBool + +-- Now the sphere + +SuspS¹ : Type₀ +SuspS¹ = Susp S¹ + +SuspS¹→S² : SuspS¹ → S² +SuspS¹→S² north = base +SuspS¹→S² south = base +SuspS¹→S² (merid base i) = base +SuspS¹→S² (merid (loop j) i) = surf i j + +meridian-contraction : I → I → I → SuspS¹ +meridian-contraction i j l = hfill (λ k → λ { (i = i0) → north + ; (i = i1) → merid base (~ k) + ; (j = i0) → merid base (~ k ∧ i) + ; (j = i1) → merid base (~ k ∧ i) }) + (inS (merid (loop j) i)) l + +S²→SuspS¹ : S² → SuspS¹ +S²→SuspS¹ base = north +S²→SuspS¹ (surf i j) = meridian-contraction i j i1 + +S²→SuspS¹→S² : ∀ x → SuspS¹→S² (S²→SuspS¹ x) ≡ x +S²→SuspS¹→S² base k = base +S²→SuspS¹→S² (surf i j) k = SuspS¹→S² (meridian-contraction i j (~ k)) + +SuspS¹→S²→SuspS¹ : ∀ x → S²→SuspS¹ (SuspS¹→S² x) ≡ x +SuspS¹→S²→SuspS¹ north k = north +SuspS¹→S²→SuspS¹ south k = merid base k +SuspS¹→S²→SuspS¹ (merid base j) k = merid base (k ∧ j) +SuspS¹→S²→SuspS¹ (merid (loop j) i) k = meridian-contraction i j (~ k) + +S²IsoSuspS¹ : Iso S² SuspS¹ +fun S²IsoSuspS¹ = S²→SuspS¹ +inv S²IsoSuspS¹ = SuspS¹→S² +rightInv S²IsoSuspS¹ = SuspS¹→S²→SuspS¹ +leftInv S²IsoSuspS¹ = S²→SuspS¹→S² + +S²≃SuspS¹ : S² ≃ SuspS¹ +S²≃SuspS¹ = isoToEquiv S²IsoSuspS¹ + +S²≡SuspS¹ : S² ≡ SuspS¹ +S²≡SuspS¹ = ua S²≃SuspS¹ + +-- And the 3-sphere + +SuspS² : Type₀ +SuspS² = Susp S² + +SuspS²→S³ : SuspS² → S³ +SuspS²→S³ north = base +SuspS²→S³ south = base +SuspS²→S³ (merid base i) = base +SuspS²→S³ (merid (surf j k) i) = surf i j k + +meridian-contraction-2 : I → I → I → I → SuspS² +meridian-contraction-2 i j k l = hfill (λ m → λ { (i = i0) → north + ; (i = i1) → merid base (~ m) + ; (j = i0) → merid base (~ m ∧ i) + ; (j = i1) → merid base (~ m ∧ i) + ; (k = i0) → merid base (~ m ∧ i) + ; (k = i1) → merid base (~ m ∧ i) }) + (inS (merid (surf j k) i)) l + +S³→SuspS² : S³ → SuspS² +S³→SuspS² base = north +S³→SuspS² (surf i j k) = meridian-contraction-2 i j k i1 + +S³→SuspS²→S³ : ∀ x → SuspS²→S³ (S³→SuspS² x) ≡ x +S³→SuspS²→S³ base l = base +S³→SuspS²→S³ (surf i j k) l = SuspS²→S³ (meridian-contraction-2 i j k (~ l)) + +SuspS²→S³→SuspS² : ∀ x → S³→SuspS² (SuspS²→S³ x) ≡ x +SuspS²→S³→SuspS² north l = north +SuspS²→S³→SuspS² south l = merid base l +SuspS²→S³→SuspS² (merid base j) l = merid base (l ∧ j) +SuspS²→S³→SuspS² (merid (surf j k) i) l = meridian-contraction-2 i j k (~ l) + +S³IsoSuspS² : Iso S³ SuspS² +fun S³IsoSuspS² = S³→SuspS² +inv S³IsoSuspS² = SuspS²→S³ +rightInv S³IsoSuspS² = SuspS²→S³→SuspS² +leftInv S³IsoSuspS² = S³→SuspS²→S³ + +S³≃SuspS² : S³ ≃ SuspS² +S³≃SuspS² = isoToEquiv S³IsoSuspS² + +S³≡SuspS² : S³ ≡ SuspS² +S³≡SuspS² = ua S³≃SuspS² + +IsoType→IsoSusp : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso (Susp A) (Susp B) +fun (IsoType→IsoSusp is) north = north +fun (IsoType→IsoSusp is) south = south +fun (IsoType→IsoSusp is) (merid a i) = merid (fun is a) i +inv (IsoType→IsoSusp is) north = north +inv (IsoType→IsoSusp is) south = south +inv (IsoType→IsoSusp is) (merid a i) = merid (inv is a) i +rightInv (IsoType→IsoSusp is) north = refl +rightInv (IsoType→IsoSusp is) south = refl +rightInv (IsoType→IsoSusp is) (merid a i) j = merid (rightInv is a j) i +leftInv (IsoType→IsoSusp is) north = refl +leftInv (IsoType→IsoSusp is) south = refl +leftInv (IsoType→IsoSusp is) (merid a i) j = merid (leftInv is a j) i + +IsoSuspS²SuspSuspS¹ : Iso (Susp S²) (Susp (Susp S¹)) +IsoSuspS²SuspSuspS¹ = IsoType→IsoSusp S²IsoSuspS¹ + +IsoS³S3 : Iso S³ (Susp (Susp S¹)) +IsoS³S3 = compIso S³IsoSuspS² IsoSuspS²SuspSuspS¹ diff --git a/Cubical/HITs/Susp/LoopAdjunction.agda b/Cubical/HITs/Susp/LoopAdjunction.agda new file mode 100644 index 0000000000..91be404b7b --- /dev/null +++ b/Cubical/HITs/Susp/LoopAdjunction.agda @@ -0,0 +1,113 @@ +{- Adjunction between suspension and loop space (Rongji Kang, Oct. 2021) + +Main results: + +- Ω∙ : ℕ → Pointed ℓ → Pointed ℓ +- ΣΩAdjunction : ((X , x₀) : Pointed ℓ) (Y : Pointed ℓ') → (∙Susp X →∙ Y) ≃ ((X , x₀) →∙ Ω∙ 1 Y) +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.Susp.LoopAdjunction where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Foundations.Pointed.Base +open import Cubical.Data.Nat.Base +open import Cubical.HITs.Susp.Base + +private + variable + ℓ ℓ' : Level + +Ω∙ : ℕ → Pointed ℓ → Pointed ℓ +Ω∙ 0 X = X +Ω∙ 1 (X , x) = (x ≡ x) , refl +Ω∙ (suc (suc n)) X = Ω∙ 1 (Ω∙ (suc n) X) + +ΣX→∙YEquiv : ((X , x₀) : Pointed ℓ) ((Y , y₀) : Pointed ℓ') + → (Susp∙ X →∙ (Y , y₀)) ≃ (Σ[ y ∈ Y ] (X → (y₀ ≡ y))) +ΣX→∙YEquiv (X , x₀) (Y , y₀) = + isoToEquiv (iso left→right right→left right→left→right left→right→left) + where + left→right : (Susp∙ X →∙ (Y , y₀)) → Σ[ y ∈ Y ] (X → (y₀ ≡ y)) + left→right (f , b) .fst = f south + left→right (f , b) .snd x = sym b ∙ cong f (merid x) + + right→left : (Σ[ y ∈ Y ] (X → (y₀ ≡ y))) → (Susp∙ X →∙ (Y , y₀)) + right→left (y , g) .fst north = y₀ + right→left (y , g) .fst south = y + right→left (y , g) .fst (merid x i) = g x i + right→left (y , g) .snd = refl + + right→left→right : (g : Σ[ y ∈ Y ] (X → (y₀ ≡ y))) → left→right (right→left g) ≡ g + right→left→right (y , g) i .fst = y + right→left→right (y , g) i .snd x = lUnit (g x) (~ i) + + left→right→left : (f : Susp∙ X →∙ (Y , y₀)) → right→left (left→right f) ≡ f + left→right→left (f , b) i .fst north = b (~ i) + left→right→left (f , b) i .fst south = f south + left→right→left (f , b) i .fst (merid x j) = + hcomp (λ k → λ { (i = i0) → compPath-filler (sym b) (cong f (merid x)) k j + ; (i = i1) → f (merid x (j ∧ k)) + ; (j = i0) → b (~ i) + ; (j = i1) → f (merid x k) }) + (b (~ (i ∨ j))) + left→right→left (f , b) i .snd j = b (~ i ∨ j) + +X→∙ΩYEquiv : ((X , x₀) : Pointed ℓ)((Y , y₀) : Pointed ℓ') + → ((X , x₀) →∙ Ω∙ 1 (Y , y₀)) ≃ (Σ[ y ∈ Y ] (X → (y₀ ≡ y))) +X→∙ΩYEquiv (X , x₀) (Y , y₀) = + isoToEquiv (iso left→right right→left right→left→right left→right→left) + where + left→right : ((X , x₀) →∙ Ω∙ 1 (Y , y₀)) → Σ[ y ∈ Y ] (X → (y₀ ≡ y)) + left→right _ .fst = y₀ + left→right (f , b) .snd = f + + right→left : (Σ[ y ∈ Y ] (X → (y₀ ≡ y))) → ((X , x₀) →∙ Ω∙ 1 (Y , y₀)) + right→left (y , g) .fst x = g x ∙ sym (g x₀) + right→left (y , g) .snd = rCancel (g x₀) + + right→left→right : (g : Σ[ y ∈ Y ] (X → (y₀ ≡ y))) → left→right (right→left g) ≡ g + right→left→right (y , g) i .fst = g x₀ i + right→left→right (y , g) i .snd x j = + hcomp (λ k → λ { (i = i0) → compPath-filler (g x) (sym (g x₀)) k j + ; (i = i1) → g x j + ; (j = i0) → y₀ + ; (j = i1) → g x₀ (i ∨ ~ k) }) + (g x j) + + f-filler : ((X , x₀) →∙ Ω∙ 1 (Y , y₀)) → X → (i j k : I) → Y + f-filler (f , b) x i j k = + hfill (λ k' → λ { (i = i0) → compPath-filler (f x) (sym (f x₀)) k' j + ; (i = i1) → f x j + ; (j = i0) → y₀ + ; (j = i1) → b i (~ k') }) + (inS (f x j)) k + + bottom : ((X , x₀) →∙ Ω∙ 1 (Y , y₀)) → (i j k : I) → Y + bottom (f , b) i j k = + hcomp (λ l → λ { (i = i0) → b (~ l) (k ∧ ~ j) + ; (i = i1) → b (j ∨ ~ l) k + ; (j = i0) → b (~ l) k + ; (j = i1) → y₀ + ; (k = i0) → y₀ + ; (k = i1) → b (~ l ∨ i) (~ j) }) + y₀ + + left→right→left : (f : (X , x₀) →∙ Ω∙ 1 (Y , y₀)) → right→left (left→right f) ≡ f + left→right→left (f , b) i .fst x j = + f-filler (f , b) x i j i1 + left→right→left (f , b) i .snd j k = + hcomp (λ l → λ { (i = i0) → rCancel-filler (f x₀) l j k + ; (i = i1) → b j k + ; (j = i0) → f-filler (f , b) x₀ i k l + ; (j = i1) → y₀ + ; (k = i0) → y₀ + ; (k = i1) → b i (~ l ∧ ~ j) }) + (bottom (f , b) i j k) + +{- The Main Theorem -} +ΣΩAdjunction : ((X , x₀) : Pointed ℓ) (Y : Pointed ℓ') → (Susp∙ X →∙ Y) ≃ ((X , x₀) →∙ Ω∙ 1 Y) +ΣΩAdjunction X Y = compEquiv (ΣX→∙YEquiv X Y) (invEquiv (X→∙ΩYEquiv X Y)) diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda new file mode 100644 index 0000000000..e751ac8904 --- /dev/null +++ b/Cubical/HITs/Susp/Properties.agda @@ -0,0 +1,193 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Susp.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Path +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Bool +open import Cubical.Data.Sigma + +open import Cubical.HITs.Join +open import Cubical.HITs.Susp.Base +open import Cubical.Homotopy.Loopspace + +private + variable + ℓ : Level + +open Iso + +Susp-iso-joinBool : ∀ {ℓ} {A : Type ℓ} → Iso (Susp A) (join A Bool) +fun Susp-iso-joinBool north = inr true +fun Susp-iso-joinBool south = inr false +fun Susp-iso-joinBool (merid a i) = (sym (push a true) ∙ push a false) i +inv Susp-iso-joinBool (inr true ) = north +inv Susp-iso-joinBool (inr false) = south +inv Susp-iso-joinBool (inl _) = north +inv Susp-iso-joinBool (push a true i) = north +inv Susp-iso-joinBool (push a false i) = merid a i +rightInv Susp-iso-joinBool (inr true ) = refl +rightInv Susp-iso-joinBool (inr false) = refl +rightInv Susp-iso-joinBool (inl a) = sym (push a true) +rightInv Susp-iso-joinBool (push a true i) j = push a true (i ∨ ~ j) +rightInv Susp-iso-joinBool (push a false i) j + = hcomp (λ k → λ { (i = i0) → push a true (~ j) + ; (i = i1) → push a false k + ; (j = i1) → push a false (i ∧ k) }) + (push a true (~ i ∧ ~ j)) +leftInv Susp-iso-joinBool north = refl +leftInv Susp-iso-joinBool south = refl +leftInv (Susp-iso-joinBool {A = A}) (merid a i) j + = hcomp (λ k → λ { (i = i0) → transp (λ _ → Susp A) (k ∨ j) north + ; (i = i1) → transp (λ _ → Susp A) (k ∨ j) (merid a k) + ; (j = i1) → merid a (i ∧ k) }) + (transp (λ _ → Susp A) j north) + +Susp≃joinBool : ∀ {ℓ} {A : Type ℓ} → Susp A ≃ join A Bool +Susp≃joinBool = isoToEquiv Susp-iso-joinBool + +Susp≡joinBool : ∀ {ℓ} {A : Type ℓ} → Susp A ≡ join A Bool +Susp≡joinBool = isoToPath Susp-iso-joinBool + +congSuspIso : ∀ {ℓ} {A B : Type ℓ} → Iso A B → Iso (Susp A) (Susp B) +fun (congSuspIso is) = suspFun (fun is) +inv (congSuspIso is) = suspFun (inv is) +rightInv (congSuspIso is) north = refl +rightInv (congSuspIso is) south = refl +rightInv (congSuspIso is) (merid a i) j = merid (rightInv is a j) i +leftInv (congSuspIso is) north = refl +leftInv (congSuspIso is) south = refl +leftInv (congSuspIso is) (merid a i) j = merid (leftInv is a j) i + +congSuspEquiv : ∀ {ℓ} {A B : Type ℓ} → A ≃ B → Susp A ≃ Susp B +congSuspEquiv {ℓ} {A} {B} h = isoToEquiv (congSuspIso (equivToIso h)) + +suspToPropElim : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Type ℓ'} (a : A) + → ((x : Susp A) → isProp (B x)) + → B north + → (x : Susp A) → B x +suspToPropElim a isProp Bnorth north = Bnorth +suspToPropElim {B = B} a isProp Bnorth south = subst B (merid a) Bnorth +suspToPropElim {B = B} a isProp Bnorth (merid a₁ i) = + isOfHLevel→isOfHLevelDep 1 isProp Bnorth (subst B (merid a) Bnorth) (merid a₁) i + +suspToPropElim2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Susp A → Type ℓ'} (a : A) + → ((x y : Susp A) → isProp (B x y)) + → B north north + → (x y : Susp A) → B x y +suspToPropElim2 _ _ Bnorth north north = Bnorth +suspToPropElim2 {B = B} a _ Bnorth north south = subst (B north) (merid a) Bnorth +suspToPropElim2 {B = B} a isprop Bnorth north (merid x i) = + isProp→PathP (λ i → isprop north (merid x i)) + Bnorth (subst (B north) (merid a) Bnorth) i +suspToPropElim2 {B = B} a _ Bnorth south north = subst (λ x → B x north) (merid a) Bnorth +suspToPropElim2 {B = B} a _ Bnorth south south = subst (λ x → B x x) (merid a) Bnorth +suspToPropElim2 {B = B} a isprop Bnorth south (merid x i) = + isProp→PathP (λ i → isprop south (merid x i)) + (subst (λ x → B x north) (merid a) Bnorth) + (subst (λ x → B x x) (merid a) Bnorth) i +suspToPropElim2 {B = B} a isprop Bnorth (merid x i) north = + isProp→PathP (λ i → isprop (merid x i) north) + Bnorth (subst (λ x → B x north) (merid a) Bnorth) i +suspToPropElim2 {B = B} a isprop Bnorth (merid x i) south = + isProp→PathP (λ i → isprop (merid x i) south) + (subst (B north) (merid a) Bnorth) + (subst (λ x → B x x) (merid a) Bnorth) i +suspToPropElim2 {B = B} a isprop Bnorth (merid x i) (merid y j) = + isSet→SquareP (λ i j → isOfHLevelSuc 1 (isprop _ _)) + (isProp→PathP (λ i₁ → isprop north (merid y i₁)) Bnorth + (subst (B north) (merid a) Bnorth)) + (isProp→PathP (λ i₁ → isprop south (merid y i₁)) + (subst (λ x₁ → B x₁ north) (merid a) Bnorth) + (subst (λ x₁ → B x₁ x₁) (merid a) Bnorth)) + (isProp→PathP (λ i₁ → isprop (merid x i₁) north) Bnorth + (subst (λ x₁ → B x₁ north) (merid a) Bnorth)) + (isProp→PathP (λ i₁ → isprop (merid x i₁) south) + (subst (B north) (merid a) Bnorth) + (subst (λ x₁ → B x₁ x₁) (merid a) Bnorth)) i j +{- Clever proof: +suspToPropElim2 a isProp Bnorth = + suspToPropElim a (λ x → isOfHLevelΠ 1 λ y → isProp x y) + (suspToPropElim a (λ x → isProp north x) Bnorth) +-} + +funSpaceSuspIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Iso (Σ[ x ∈ B ] Σ[ y ∈ B ] (A → x ≡ y)) (Susp A → B) +Iso.fun funSpaceSuspIso (x , y , f) north = x +Iso.fun funSpaceSuspIso (x , y , f) south = y +Iso.fun funSpaceSuspIso (x , y , f) (merid a i) = f a i +Iso.inv funSpaceSuspIso f = (f north) , (f south , (λ x → cong f (merid x))) +Iso.rightInv funSpaceSuspIso f = funExt λ {north → refl + ; south → refl + ; (merid a i) → refl} +Iso.leftInv funSpaceSuspIso _ = refl + +toSusp : (A : Pointed ℓ) → typ A → typ (Ω (Susp∙ (typ A))) +toSusp A x = merid x ∙ merid (pt A) ⁻¹ + +toSuspPointed : (A : Pointed ℓ) → A →∙ Ω (Susp∙ (typ A)) +fst (toSuspPointed A) = toSusp A +snd (toSuspPointed A) = rCancel (merid (pt A)) + +module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} where + fromSusp→toΩ : Susp∙ (typ A) →∙ B → (A →∙ Ω B) + fst (fromSusp→toΩ f) x = sym (snd f) ∙∙ cong (fst f) (toSusp A x) ∙∙ snd f + snd (fromSusp→toΩ f) = + cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong (cong (fst f)) + (rCancel (merid (pt A)))) + ∙ ∙∙lCancel (snd f) + + toΩ→fromSusp : A →∙ Ω B → Susp∙ (typ A) →∙ B + fst (toΩ→fromSusp f) north = pt B + fst (toΩ→fromSusp f) south = pt B + fst (toΩ→fromSusp f) (merid a i) = fst f a i + snd (toΩ→fromSusp f) = refl + + ΩSuspAdjointIso : Iso (A →∙ Ω B) (Susp∙ (typ A) →∙ B) + fun ΩSuspAdjointIso = toΩ→fromSusp + inv ΩSuspAdjointIso = fromSusp→toΩ + rightInv ΩSuspAdjointIso f = + ΣPathP (funExt + (λ { north → sym (snd f) + ; south → sym (snd f) ∙ cong (fst f) (merid (pt A)) + ; (merid a i) j → + hcomp (λ k → λ { (i = i0) → snd f (~ j ∧ k) + ; (i = i1) → compPath-filler' + (sym (snd f)) + (cong (fst f) (merid (pt A))) k j + ; (j = i1) → fst f (merid a i)}) + (fst f (compPath-filler (merid a) (sym (merid (pt A))) (~ j) i))}) + , λ i j → snd f (~ i ∨ j)) + leftInv ΩSuspAdjointIso f = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ x → sym (rUnit _) + ∙ cong-∙ (fst (toΩ→fromSusp f)) (merid x) (sym (merid (pt A))) + ∙ cong (fst f x ∙_) (cong sym (snd f)) + ∙ sym (rUnit _)) + + IsoΩFunSuspFun : Iso (typ (Ω (A →∙ B ∙))) (Susp∙ (typ A) →∙ B) + IsoΩFunSuspFun = compIso (ΩfunExtIso A B) ΩSuspAdjointIso + +-- inversion +invSusp : ∀ {ℓ} {A : Type ℓ} → Susp A → Susp A +invSusp north = south +invSusp south = north +invSusp (merid a i) = merid a (~ i) + +invSusp² : ∀ {ℓ} {A : Type ℓ} (x : Susp A) → invSusp (invSusp x) ≡ x +invSusp² north = refl +invSusp² south = refl +invSusp² (merid a i) = refl + +invSuspIso : ∀ {ℓ} {A : Type ℓ} → Iso (Susp A) (Susp A) +fun invSuspIso = invSusp +inv invSuspIso = invSusp +rightInv invSuspIso = invSusp² +leftInv invSuspIso = invSusp² diff --git a/Cubical/HITs/Torus.agda b/Cubical/HITs/Torus.agda new file mode 100644 index 0000000000..18acd0243d --- /dev/null +++ b/Cubical/HITs/Torus.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Torus where + +open import Cubical.HITs.Torus.Base public + +-- open import Cubical.HITs.Torus.Properties public diff --git a/Cubical/HITs/Torus/Base.agda b/Cubical/HITs/Torus/Base.agda new file mode 100644 index 0000000000..662cda6836 --- /dev/null +++ b/Cubical/HITs/Torus/Base.agda @@ -0,0 +1,95 @@ +{- + +Definition of the torus as a HIT together with a proof that it is +equivalent to two circles + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.Torus.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Nat +open import Cubical.Data.Int +open import Cubical.Data.Sigma + +open import Cubical.HITs.S1 + +data Torus : Type₀ where + point : Torus + line1 : point ≡ point + line2 : point ≡ point + square : PathP (λ i → line1 i ≡ line1 i) line2 line2 + +t2c : Torus → S¹ × S¹ +t2c point = ( base , base ) +t2c (line1 i) = ( loop i , base ) +t2c (line2 j) = ( base , loop j ) +t2c (square i j) = ( loop i , loop j ) + +c2t : S¹ × S¹ → Torus +c2t (base , base) = point +c2t (loop i , base) = line1 i +c2t (base , loop j) = line2 j +c2t (loop i , loop j) = square i j + +c2t-t2c : ∀ (t : Torus) → c2t (t2c t) ≡ t +c2t-t2c point = refl +c2t-t2c (line1 _) = refl +c2t-t2c (line2 _) = refl +c2t-t2c (square _ _) = refl + +t2c-c2t : ∀ (p : S¹ × S¹) → t2c (c2t p) ≡ p +t2c-c2t (base , base) = refl +t2c-c2t (base , loop _) = refl +t2c-c2t (loop _ , base) = refl +t2c-c2t (loop _ , loop _) = refl + +Torus≡S¹×S¹ : Torus ≡ S¹ × S¹ +Torus≡S¹×S¹ = isoToPath (iso t2c c2t t2c-c2t c2t-t2c) + +point-path : PathP (λ i → Torus≡S¹×S¹ i) point (base , base) +point-path i = + (glue (λ { + (i = i0) → point; + (i = i1) → (base , base) }) (base , base)) + +Loop : {A : Type₀} (p : A) → Type₀ +Loop p = p ≡ p + +ΩTorus : Type₀ +ΩTorus = Loop point + +-- TODO: upstream +lemPathAnd : ∀ {ℓ} {A B : Type ℓ} (t u : A × B) → + Path _ (t ≡ u) ((t .fst ≡ u .fst) × ((t .snd) ≡ (u .snd))) +lemPathAnd t u = isoToPath (iso (λ tu → (λ i → tu i .fst) , λ i → tu i .snd) + (λ tu i → tu .fst i , tu .snd i) + (λ y → refl) + (λ x → refl)) + +ΩTorus≡ℤ×ℤ : ΩTorus ≡ ℤ × ℤ +ΩTorus≡ℤ×ℤ = + ΩTorus + ≡⟨ (λ i → Loop (point-path i)) ⟩ + Loop (base , base) + ≡⟨ lemPathAnd (base , base) (base , base) ⟩ + ΩS¹ × ΩS¹ + ≡⟨ (λ i → ΩS¹≡ℤ i × ΩS¹≡ℤ i) ⟩ + ℤ × ℤ ∎ + +-- Computing the winding numbers on the torus +windingTorus : ΩTorus → ℤ × ℤ +windingTorus l = ( winding (λ i → t2c (l i) .fst) + , winding (λ i → t2c (l i) .snd)) + +module _ where + private + test1 : windingTorus (line1 ∙ line2) ≡ (pos 1 , pos 1) + test1 = refl + + test2 : windingTorus (line1 ∙ line2 ∙ sym line1 ∙ sym line1) ≡ (negsuc 0 , pos 1) + test2 = refl diff --git a/Cubical/HITs/Truncation.agda b/Cubical/HITs/Truncation.agda new file mode 100644 index 0000000000..3b716c88a9 --- /dev/null +++ b/Cubical/HITs/Truncation.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Truncation where + +open import Cubical.HITs.Truncation.Base public +open import Cubical.HITs.Truncation.Properties public + +import Cubical.HITs.Truncation.FromNegTwo diff --git a/Cubical/HITs/Truncation/Base.agda b/Cubical/HITs/Truncation/Base.agda new file mode 100644 index 0000000000..fdd0be75c0 --- /dev/null +++ b/Cubical/HITs/Truncation/Base.agda @@ -0,0 +1,43 @@ +{- + +A simpler definition of truncation ∥ A ∥ n from n ≥ -1 + +Note that this uses the HoTT book's indexing, so it will be off + from `∥_∥_` in HITs.Truncation.Base by -2 + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.Truncation.Base where + +open import Cubical.Data.NatMinusOne +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Pointed +open import Cubical.HITs.Sn.Base +open import Cubical.Data.Nat.Base +open import Cubical.Data.Unit.Base +open import Cubical.Data.Empty + +-- this definition is off by one. Use hLevelTrunc or ∥_∥ for truncations +-- (off by 2 w.r.t. the HoTT-book) +data HubAndSpoke {ℓ} (A : Type ℓ) (n : ℕ) : Type ℓ where + ∣_∣ : A → HubAndSpoke A n + hub : (f : S₊ n → HubAndSpoke A n) → HubAndSpoke A n + spoke : (f : S₊ n → HubAndSpoke A n) (x : S₊ n) → hub f ≡ f x + +hLevelTrunc : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Type ℓ +hLevelTrunc zero A = Unit* +hLevelTrunc (suc n) A = HubAndSpoke A n + +∥_∥_ : ∀ {ℓ} (A : Type ℓ) (n : ℕ) → Type ℓ +∥ A ∥ n = hLevelTrunc n A + +∣_∣ₕ : ∀ {ℓ} {A : Type ℓ} {n : ℕ} → A → ∥ A ∥ n +∣_∣ₕ {n = zero} a = tt* +∣_∣ₕ {n = suc n} a = ∣ a ∣ + +-- Pointed version +hLevelTrunc∙ : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Pointed ℓ +fst (hLevelTrunc∙ n A) = hLevelTrunc n (typ A) +snd (hLevelTrunc∙ zero A) = tt* +snd (hLevelTrunc∙ (suc n) A) = ∣ pt A ∣ₕ diff --git a/Cubical/HITs/Truncation/FromNegTwo.agda b/Cubical/HITs/Truncation/FromNegTwo.agda new file mode 100644 index 0000000000..257017b24b --- /dev/null +++ b/Cubical/HITs/Truncation/FromNegTwo.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Truncation.FromNegTwo where + +open import Cubical.HITs.Truncation.FromNegTwo.Base public +open import Cubical.HITs.Truncation.FromNegTwo.Properties public diff --git a/Cubical/HITs/Truncation/FromNegTwo/Base.agda b/Cubical/HITs/Truncation/FromNegTwo/Base.agda new file mode 100644 index 0000000000..605b23d88f --- /dev/null +++ b/Cubical/HITs/Truncation/FromNegTwo/Base.agda @@ -0,0 +1,29 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Truncation.FromNegTwo.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Nat +open import Cubical.Data.NatMinusOne +open import Cubical.HITs.Nullification +open import Cubical.HITs.Sn.Base + +-- For the hub-and-spoke construction discussed in the HoTT book, which doesn't work in the base case +-- of contractibility, see `HITs.Truncation.Base`. The definition of truncation here contains +-- two more constructors which are redundant when n ≥ 1 but give contractibility when n = 0. + +-- data hLevelTrunc {ℓ} (n : HLevel) (A : Type ℓ) : Type (ℓ-max ℓ ℓ') where +-- -- the hub-and-spoke definition in `Truncation.Base` +-- ∣_∣ : A → hLevelTrunc n A +-- hub : (f : S (-1+ n) → hLevelTrunc n A) → hLevelTrunc n A +-- spoke : (f : S (-1+ n) → hLevelTrunc n A) (s : S) → hub f ≡ f s +-- -- two additional constructors needed to ensure that hLevelTrunc 0 A is contractible +-- ≡hub : ∀ {x y} (p : S (-1+ n) → x ≡ y) → x ≡ y +-- ≡spoke : ∀ {x y} (p : S (-1+ n) → x ≡ y) (s : S (-1+ n)) → ≡hub p ≡ p s + +hLevelTrunc : ∀ {ℓ} → HLevel → Type ℓ → Type ℓ +hLevelTrunc n A = Null (S (-1+ n)) A + +-- Note that relative to the HoTT book, this notation is off by +2 +∥_∥_ : ∀ {ℓ} → Type ℓ → HLevel → Type ℓ +∥ A ∥ n = hLevelTrunc n A diff --git a/Cubical/HITs/Truncation/FromNegTwo/Properties.agda b/Cubical/HITs/Truncation/FromNegTwo/Properties.agda new file mode 100644 index 0000000000..cb97ab7df7 --- /dev/null +++ b/Cubical/HITs/Truncation/FromNegTwo/Properties.agda @@ -0,0 +1,396 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Truncation.FromNegTwo.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Equiv.PathSplit +open isPathSplitEquiv +open import Cubical.Modalities.Modality +open Modality + +open import Cubical.Data.Empty as ⊥ using (⊥) +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.NatMinusOne as ℕ₋₁ +open import Cubical.Data.Sigma +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.Susp.Base +open import Cubical.HITs.Nullification as Null hiding (rec; elim) + +open import Cubical.HITs.Truncation.FromNegTwo.Base + +open import Cubical.HITs.PropositionalTruncation as PropTrunc + renaming (∥_∥ to ∥_∥₁; ∣_∣ to ∣_∣₁; squash to squash₁) using () +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) +open import Cubical.HITs.GroupoidTruncation as GpdTrunc using (∥_∥₃; ∣_∣₃; squash₃) +open import Cubical.HITs.2GroupoidTruncation as 2GpdTrunc using (∥_∥₄; ∣_∣₄; squash₄) + +private + variable + ℓ ℓ' : Level + A B : Type ℓ + +sphereFill : (n : ℕ₋₁) (f : S n → A) → Type _ +sphereFill {A = A} n f = Σ[ top ∈ A ] ((x : S n) → top ≡ f x) + +isSphereFilled : ℕ₋₁ → Type ℓ → Type ℓ +isSphereFilled n A = (f : S n → A) → sphereFill n f + +isSphereFilledTrunc : {n : HLevel} → isSphereFilled (-1+ n) (hLevelTrunc n A) +isSphereFilledTrunc {n = zero} f = hub f , ⊥.elim +isSphereFilledTrunc {n = suc n} f = hub f , spoke f + +isSphereFilled→isOfHLevelSuc : {n : HLevel} → isSphereFilled (ℕ→ℕ₋₁ n) A → isOfHLevel (suc n) A +isSphereFilled→isOfHLevelSuc {A = A} {zero} h x y = sym (snd (h f) north) ∙ snd (h f) south + where + f : Susp ⊥ → A + f north = x + f south = y +isSphereFilled→isOfHLevelSuc {A = A} {suc n} h x y = isSphereFilled→isOfHLevelSuc (helper h x y) + where + helper : isSphereFilled (ℕ→ℕ₋₁ (suc n)) A → (x y : A) → isSphereFilled (ℕ→ℕ₋₁ n) (x ≡ y) + helper h x y f = sym p ∙ q , r + where + f' : Susp (S (ℕ→ℕ₋₁ n)) → A + f' north = x + f' south = y + f' (merid u i) = f u i + + p = snd (h f') north + q = snd (h f') south + + r : (s : S (ℕ→ℕ₋₁ n)) → sym p ∙ q ≡ f s + r s i j = hcomp (λ k → λ { (i = i0) → compPath-filler (sym p) q k j + ; (i = i1) → snd (h f') (merid s j) k + ; (j = i0) → p (k ∨ ~ i) + ; (j = i1) → q k }) + (p (~ i ∧ ~ j)) + +isOfHLevel→isSphereFilled : {n : HLevel} → isOfHLevel n A → isSphereFilled (-1+ n) A +isOfHLevel→isSphereFilled {n = zero} h f = fst h , λ _ → snd h _ +isOfHLevel→isSphereFilled {n = suc zero} h f = f north , λ _ → h _ _ +isOfHLevel→isSphereFilled {A = A} {suc (suc n)} h = helper λ x y → isOfHLevel→isSphereFilled (h x y) + where + helper : {n : HLevel} → ((x y : A) → isSphereFilled (-1+ n) (x ≡ y)) → isSphereFilled (suc₋₁ (-1+ n)) A + helper {n} h f = l , r + where + l : A + l = f north + + f' : S (-1+ n) → f north ≡ f south + f' x i = f (merid x i) + + h' : sphereFill (-1+ n) f' + h' = h (f north) (f south) f' + + r : (x : S (suc₋₁ (-1+ n))) → l ≡ f x + r north = refl + r south = h' .fst + r (merid x i) j = hcomp (λ k → λ { (i = i0) → f north + ; (i = i1) → h' .snd x (~ k) j + ; (j = i0) → f north + ; (j = i1) → f (merid x i) }) + (f (merid x (i ∧ j))) + +-- isNull (S n) A ≃ (isSphereFilled n A) × (∀ (x y : A) → isSphereFilled n (x ≡ y)) + +isOfHLevel→isSnNull : {n : HLevel} → isOfHLevel n A → isNull (S (-1+ n)) A +fst (sec (isOfHLevel→isSnNull h)) f = fst (isOfHLevel→isSphereFilled h f) +snd (sec (isOfHLevel→isSnNull h)) f i s = snd (isOfHLevel→isSphereFilled h f) s i +fst (secCong (isOfHLevel→isSnNull h) x y) p = fst (isOfHLevel→isSphereFilled (isOfHLevelPath _ h x y) (funExt⁻ p)) +snd (secCong (isOfHLevel→isSnNull h) x y) p i j s = snd (isOfHLevel→isSphereFilled (isOfHLevelPath _ h x y) (funExt⁻ p)) s i j + +isSnNull→isOfHLevel : {n : HLevel} → isNull (S (-1+ n)) A → isOfHLevel n A +isSnNull→isOfHLevel {n = zero} nA = fst (sec nA) ⊥.rec , λ y → fst (secCong nA _ y) (funExt ⊥.elim) +isSnNull→isOfHLevel {n = suc n} nA = isSphereFilled→isOfHLevelSuc (λ f → fst (sec nA) f , λ s i → snd (sec nA) f i s) + +isOfHLevelTrunc : (n : HLevel) → isOfHLevel n (hLevelTrunc n A) +isOfHLevelTrunc zero = hub ⊥.rec , λ _ → ≡hub ⊥.rec +isOfHLevelTrunc (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilledTrunc + +-- isOfHLevelTrunc n = isSnNull→isOfHLevel isNull-Null + +-- hLevelTrunc n is a modality + +rec : {n : HLevel} + {B : Type ℓ'} → + isOfHLevel n B → + (A → B) → + hLevelTrunc n A → + B +rec h = Null.rec (isOfHLevel→isSnNull h) + +elim : {n : HLevel} + {B : hLevelTrunc n A → Type ℓ'} + (hB : (x : hLevelTrunc n A) → isOfHLevel n (B x)) + (g : (a : A) → B (∣ a ∣)) + (x : hLevelTrunc n A) → + B x +elim hB = Null.elim (λ x → isOfHLevel→isSnNull (hB x)) + +elim2 : {n : HLevel} + {B : hLevelTrunc n A → hLevelTrunc n A → Type ℓ'} + (hB : ((x y : hLevelTrunc n A) → isOfHLevel n (B x y))) + (g : (a b : A) → B ∣ a ∣ ∣ b ∣) + (x y : hLevelTrunc n A) → + B x y +elim2 {n = n} hB g = + elim (λ _ → isOfHLevelΠ n (λ _ → hB _ _)) + (λ a → elim (λ _ → hB _ _) (λ b → g a b)) + +elim3 : {n : HLevel} + {B : (x y z : hLevelTrunc n A) → Type ℓ'} + (hB : ((x y z : hLevelTrunc n A) → isOfHLevel n (B x y z))) + (g : (a b c : A) → B (∣ a ∣) ∣ b ∣ ∣ c ∣) + (x y z : hLevelTrunc n A) → + B x y z +elim3 {n = n} hB g = + elim2 (λ _ _ → isOfHLevelΠ n (hB _ _)) + (λ a b → elim (λ _ → hB _ _ _) (λ c → g a b c)) + +HLevelTruncModality : ∀ {ℓ} (n : HLevel) → Modality ℓ +isModal (HLevelTruncModality n) = isOfHLevel n +isPropIsModal (HLevelTruncModality n) = isPropIsOfHLevel n +◯ (HLevelTruncModality n) = hLevelTrunc n +◯-isModal (HLevelTruncModality n) = isOfHLevelTrunc n +η (HLevelTruncModality n) = ∣_∣ +◯-elim (HLevelTruncModality n) = elim +◯-elim-β (HLevelTruncModality n) = λ _ _ _ → refl +◯-=-isModal (HLevelTruncModality n) = isOfHLevelPath n (isOfHLevelTrunc n) + +truncIdempotentIso : (n : HLevel) → isOfHLevel n A → Iso A (hLevelTrunc n A) +truncIdempotentIso n hA = isModalToIso (HLevelTruncModality n) hA + +truncIdempotent≃ : (n : HLevel) → isOfHLevel n A → A ≃ hLevelTrunc n A +truncIdempotent≃ n hA = ∣_∣ , isModalToIsEquiv (HLevelTruncModality n) hA + +truncIdempotent : (n : HLevel) → isOfHLevel n A → hLevelTrunc n A ≡ A +truncIdempotent n hA = ua (invEquiv (truncIdempotent≃ n hA)) + +-- universal property + +univTrunc : ∀ {ℓ} (n : HLevel) {B : TypeOfHLevel ℓ n} → Iso (hLevelTrunc n A → B .fst) (A → B .fst) +Iso.fun (univTrunc n {B , lev}) g a = g ∣ a ∣ +Iso.inv (univTrunc n {B , lev}) = elim λ _ → lev +Iso.rightInv (univTrunc n {B , lev}) b = refl +Iso.leftInv (univTrunc n {B , lev}) b = funExt (elim (λ x → isOfHLevelPath _ lev _ _) + λ a → refl) + +-- functorial action + +map : {n : HLevel} {B : Type ℓ'} (g : A → B) + → hLevelTrunc n A → hLevelTrunc n B +map g = rec (isOfHLevelTrunc _) (λ a → ∣ g a ∣) + +mapCompIso : {n : HLevel} {B : Type ℓ'} → (Iso A B) → Iso (hLevelTrunc n A) (hLevelTrunc n B) +Iso.fun (mapCompIso g) = map (Iso.fun g) +Iso.inv (mapCompIso g) = map (Iso.inv g) +Iso.rightInv (mapCompIso g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ b → cong ∣_∣ (Iso.rightInv g b) +Iso.leftInv (mapCompIso g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ a → cong ∣_∣ (Iso.leftInv g a) + +mapId : {n : HLevel} → ∀ t → map {n = n} (idfun A) t ≡ t +mapId {n = n} = + elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) (λ _ → refl) + +-- equivalences to prop/set/groupoid truncations + + +propTruncTrunc1Iso : Iso ∥ A ∥₁ (∥ A ∥ 1) +Iso.fun propTruncTrunc1Iso = PropTrunc.elim (λ _ → isOfHLevelTrunc 1) ∣_∣ +Iso.inv propTruncTrunc1Iso = elim (λ _ → squash₁) ∣_∣₁ +Iso.rightInv propTruncTrunc1Iso = elim (λ _ → isOfHLevelPath 1 (isOfHLevelTrunc 1) _ _) (λ _ → refl) +Iso.leftInv propTruncTrunc1Iso = PropTrunc.elim (λ _ → isOfHLevelPath 1 squash₁ _ _) (λ _ → refl) + +propTrunc≃Trunc1 : ∥ A ∥₁ ≃ ∥ A ∥ 1 +propTrunc≃Trunc1 = isoToEquiv propTruncTrunc1Iso + +propTrunc≡Trunc1 : ∥ A ∥₁ ≡ ∥ A ∥ 1 +propTrunc≡Trunc1 = ua propTrunc≃Trunc1 + + +setTruncTrunc2Iso : Iso ∥ A ∥₂ (∥ A ∥ 2) +Iso.fun setTruncTrunc2Iso = SetTrunc.elim (λ _ → isOfHLevelTrunc 2) ∣_∣ +Iso.inv setTruncTrunc2Iso = elim (λ _ → squash₂) ∣_∣₂ +Iso.rightInv setTruncTrunc2Iso = elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) (λ _ → refl) +Iso.leftInv setTruncTrunc2Iso = SetTrunc.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) (λ _ → refl) + +setTrunc≃Trunc2 : ∥ A ∥₂ ≃ ∥ A ∥ 2 +setTrunc≃Trunc2 = isoToEquiv setTruncTrunc2Iso + +propTrunc≡Trunc2 : ∥ A ∥₂ ≡ ∥ A ∥ 2 +propTrunc≡Trunc2 = ua setTrunc≃Trunc2 + +groupoidTrunc≃Trunc3Iso : Iso ∥ A ∥₃ (∥ A ∥ 3) +Iso.fun groupoidTrunc≃Trunc3Iso = GpdTrunc.elim (λ _ → isOfHLevelTrunc 3) ∣_∣ +Iso.inv groupoidTrunc≃Trunc3Iso = elim (λ _ → squash₃) ∣_∣₃ +Iso.rightInv groupoidTrunc≃Trunc3Iso = elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) (λ _ → refl) +Iso.leftInv groupoidTrunc≃Trunc3Iso = GpdTrunc.elim (λ _ → isOfHLevelPath 3 squash₃ _ _) (λ _ → refl) + +groupoidTrunc≃Trunc3 : ∥ A ∥₃ ≃ ∥ A ∥ 3 +groupoidTrunc≃Trunc3 = isoToEquiv groupoidTrunc≃Trunc3Iso + +groupoidTrunc≡Trunc3 : ∥ A ∥₃ ≡ ∥ A ∥ 3 +groupoidTrunc≡Trunc3 = ua groupoidTrunc≃Trunc3 + + +2GroupoidTrunc≃Trunc4Iso : Iso ∥ A ∥₄ (∥ A ∥ 4) +Iso.fun 2GroupoidTrunc≃Trunc4Iso = 2GpdTrunc.elim (λ _ → isOfHLevelTrunc 4) ∣_∣ +Iso.inv 2GroupoidTrunc≃Trunc4Iso = elim (λ _ → squash₄) ∣_∣₄ +Iso.rightInv 2GroupoidTrunc≃Trunc4Iso = elim (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (λ _ → refl) +Iso.leftInv 2GroupoidTrunc≃Trunc4Iso = 2GpdTrunc.elim (λ _ → isOfHLevelPath 4 squash₄ _ _) (λ _ → refl) + +2GroupoidTrunc≃Trunc4 : ∥ A ∥₄ ≃ ∥ A ∥ 4 +2GroupoidTrunc≃Trunc4 = + isoToEquiv + (iso + (2GpdTrunc.elim (λ _ → isOfHLevelTrunc 4) ∣_∣) + (elim (λ _ → squash₄) ∣_∣₄) + (elim (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (λ _ → refl)) + (2GpdTrunc.elim (λ _ → isOfHLevelPath 4 squash₄ _ _) (λ _ → refl))) + +2GroupoidTrunc≡Trunc4 : ∥ A ∥₄ ≡ ∥ A ∥ 4 +2GroupoidTrunc≡Trunc4 = ua 2GroupoidTrunc≃Trunc4 + + +isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) +isContr→isContrTrunc n contr = ∣ fst contr ∣ , (elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) + + +truncOfProdIso : (n : ℕ) → Iso (hLevelTrunc n (A × B)) (hLevelTrunc n A × hLevelTrunc n B) +Iso.fun (truncOfProdIso n) = rec (isOfHLevelΣ n (isOfHLevelTrunc n) (λ _ → isOfHLevelTrunc n)) λ {(a , b) → ∣ a ∣ , ∣ b ∣} +Iso.inv (truncOfProdIso n) (a , b) = rec (isOfHLevelTrunc n) + (λ a → rec (isOfHLevelTrunc n) + (λ b → ∣ a , b ∣) + b) + a +Iso.rightInv (truncOfProdIso n) (a , b) = + elim {B = λ a → Iso.fun (truncOfProdIso n) (Iso.inv (truncOfProdIso n) (a , b)) ≡ (a , b)} + (λ _ → isOfHLevelPath n (isOfHLevelΣ n (isOfHLevelTrunc n) (λ _ → isOfHLevelTrunc n)) _ _) + (λ a → elim {B = λ b → Iso.fun (truncOfProdIso n) (Iso.inv (truncOfProdIso n) (∣ a ∣ , b)) ≡ (∣ a ∣ , b)} + (λ _ → isOfHLevelPath n (isOfHLevelΣ n (isOfHLevelTrunc n) (λ _ → isOfHLevelTrunc n)) _ _) + (λ b → refl) b) a +Iso.leftInv (truncOfProdIso n) = elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +---- ∥ Ω A ∥ ₙ ≡ Ω ∥ A ∥ₙ₊₁ ---- + + +isOfHLevelTypeOfHLevel2 : ∀ n → isOfHLevel (suc n) (TypeOfHLevel ℓ n) +isOfHLevelTypeOfHLevel2 n = isOfHLevelTypeOfHLevel n + + {- Proofs of Theorem 7.3.12. and Corollary 7.3.13. in the HoTT book -} + +module ΩTrunc where + {- We define the fibration P to show a more general result -} + P : {X : Type ℓ} {n : HLevel} → ∥ X ∥ (suc n) → ∥ X ∥ (suc n) → Type ℓ + P {n = n} x y = elim2 (λ _ _ → isOfHLevelTypeOfHLevel2 (n)) + (λ a b → ∥ a ≡ b ∥ n , isOfHLevelTrunc (n)) x y .fst + + {- We will need P to be of hLevel n + 3 -} + hLevelP : {n : HLevel} (a b : ∥ B ∥ (suc n)) → isOfHLevel ((suc n)) (P a b) + hLevelP {n = n} = + elim2 (λ x y → isProp→isOfHLevelSuc (n) (isPropIsOfHLevel (suc n))) + (λ a b → isOfHLevelSuc (n) (isOfHLevelTrunc (n))) + + {- decode function from P x y to x ≡ y -} + decode-fun : {n : HLevel} (x y : ∥ B ∥ (suc n)) → P x y → x ≡ y + decode-fun {n = n} = + elim2 (λ u v → isOfHLevelΠ (suc n) + (λ _ → isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n)) u v)) + decode* + where + decode* : ∀ {n : HLevel} (u v : B) + → P {n = n} ∣ u ∣ ∣ v ∣ → Path (∥ B ∥ (suc n)) ∣ u ∣ ∣ v ∣ + decode* {B = B} {n = zero} u v = + rec ( isOfHLevelTrunc 1 ∣ u ∣ ∣ v ∣ + , λ _ → isOfHLevelSuc 1 (isOfHLevelTrunc 1) _ _ _ _) (cong ∣_∣) + decode* {n = suc n} u v = + rec (isOfHLevelTrunc (suc (suc n)) ∣ u ∣ ∣ v ∣) (cong ∣_∣) + + {- auxiliary function r used to define encode -} + r : {m : HLevel} (u : ∥ B ∥ (suc m)) → P u u + r = elim (λ x → hLevelP x x) (λ a → ∣ refl ∣) + + {- encode function from x ≡ y to P x y -} + encode-fun : {n : HLevel} (x y : ∥ B ∥ (suc n)) → x ≡ y → P x y + encode-fun x y p = transport (λ i → P x (p i)) (r x) + + {- We need the following two lemmas on the functions behaviour for refl -} + dec-refl : {n : HLevel} (x : ∥ B ∥ (suc n)) → decode-fun x x (r x) ≡ refl + dec-refl {n = zero} = + elim (λ x → isOfHLevelSuc 1 (isOfHLevelSuc 1 (isOfHLevelTrunc 1) x x) _ _) (λ _ → refl) + dec-refl {n = suc n} = + elim (λ x → isOfHLevelSuc (suc n) + (isOfHLevelSuc (suc n) + (isOfHLevelTrunc (suc (suc n)) x x) + (decode-fun x x (r x)) refl)) + (λ _ → refl) + + enc-refl : {n : HLevel} (x : ∥ B ∥ (suc n)) → encode-fun x x refl ≡ r x + enc-refl x j = transp (λ _ → P x x) j (r x) + + {- decode-fun is a right-inverse -} + P-rinv : {n : HLevel} (u v : ∥ B ∥ (suc n)) (x : Path (∥ B ∥ (suc n)) u v) + → decode-fun u v (encode-fun u v x) ≡ x + P-rinv u v = J (λ y p → decode-fun u y (encode-fun u y p) ≡ p) + (cong (decode-fun u u) (enc-refl u) ∙ dec-refl u) + + {- decode-fun is a left-inverse -} + P-linv : {n : HLevel} (u v : ∥ B ∥ (suc n )) (x : P u v) + → encode-fun u v (decode-fun u v x) ≡ x + P-linv {n = n} = + elim2 (λ x y → isOfHLevelΠ (suc n) + (λ z → isOfHLevelSuc (suc n) (hLevelP x y) _ _)) + helper + where + helper : {n : HLevel} (a b : B) (p : P {n = n} ∣ a ∣ ∣ b ∣) + → encode-fun _ _ (decode-fun ∣ a ∣ ∣ b ∣ p) ≡ p + helper {n = zero} a b = + elim (λ x → ( sym (isOfHLevelTrunc 0 .snd _) ∙ isOfHLevelTrunc 0 .snd x + , λ y → isOfHLevelSuc 1 (isOfHLevelSuc 0 (isOfHLevelTrunc 0)) _ _ _ _)) + (J (λ y p → encode-fun ∣ a ∣ ∣ y ∣ (decode-fun _ _ ∣ p ∣) ≡ ∣ p ∣) + (enc-refl ∣ a ∣)) + helper {n = suc n} a b = + elim (λ x → hLevelP {n = suc n} ∣ a ∣ ∣ b ∣ _ _) + (J (λ y p → encode-fun {n = suc n} ∣ a ∣ ∣ y ∣ (decode-fun _ _ ∣ p ∣) ≡ ∣ p ∣) + (enc-refl ∣ a ∣)) + + {- The final Iso established -} + IsoFinal : (n : HLevel) (x y : ∥ B ∥ (suc n)) → Iso (x ≡ y) (P x y) + Iso.fun (IsoFinal _ x y) = encode-fun x y + Iso.inv (IsoFinal _ x y) = decode-fun x y + Iso.rightInv (IsoFinal _ x y) = P-linv x y + Iso.leftInv (IsoFinal _ x y) = P-rinv x y + +PathIdTrunc : {a b : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) +PathIdTrunc n = isoToPath (ΩTrunc.IsoFinal n _ _) + +PathΩ : {a : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) +PathΩ n = PathIdTrunc n + +{- Special case using direct defs of truncations -} +PathIdTrunc₀Iso : {a b : A} → Iso (∣ a ∣₂ ≡ ∣ b ∣₂) ∥ a ≡ b ∥₁ +PathIdTrunc₀Iso = compIso (congIso setTruncTrunc2Iso) + (compIso (ΩTrunc.IsoFinal _ ∣ _ ∣ ∣ _ ∣) + (invIso propTruncTrunc1Iso)) + +------------------------- + + +truncOfTruncIso : (n m : HLevel) → Iso (hLevelTrunc n A) (hLevelTrunc n (hLevelTrunc (m + n) A)) +Iso.fun (truncOfTruncIso n m) = elim (λ _ → isOfHLevelTrunc n) λ a → ∣ ∣ a ∣ ∣ +Iso.inv (truncOfTruncIso {A = A} n m) = + elim (λ _ → isOfHLevelTrunc n) + (elim (λ _ → (isOfHLevelPlus m (isOfHLevelTrunc n ))) + λ a → ∣ a ∣) +Iso.rightInv (truncOfTruncIso {A = A} n m) = + elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _ ) + (elim (λ x → isOfHLevelPath (m + n) (isOfHLevelPlus m (isOfHLevelTrunc n)) _ _ ) + λ a → refl) +Iso.leftInv (truncOfTruncIso n m) = elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +truncOfTruncEq : (n m : ℕ) → (hLevelTrunc n A) ≃ (hLevelTrunc n (hLevelTrunc (m + n) A)) +truncOfTruncEq n m = isoToEquiv (truncOfTruncIso n m) diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda new file mode 100644 index 0000000000..12c16ab75c --- /dev/null +++ b/Cubical/HITs/Truncation/Properties.agda @@ -0,0 +1,566 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.Truncation.Properties where +open import Cubical.Data.NatMinusOne +open import Cubical.HITs.Truncation.Base + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Equiv.PathSplit +open isPathSplitEquiv +open import Cubical.Modalities.Modality +open Modality + +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.Sigma +open import Cubical.Data.Bool +open import Cubical.Data.Unit +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp.Base +open import Cubical.HITs.Nullification as Null hiding (rec; elim) + +open import Cubical.HITs.PropositionalTruncation as PropTrunc + renaming (∥_∥ to ∥_∥₁; ∣_∣ to ∣_∣₁; squash to squash₁) using () +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) +open import Cubical.HITs.GroupoidTruncation as GpdTrunc using (∥_∥₃; ∣_∣₃; squash₃) +open import Cubical.HITs.2GroupoidTruncation as 2GpdTrunc using (∥_∥₄; ∣_∣₄; squash₄) + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + +sphereFill : (n : ℕ) (f : S₊ n → A) → Type _ +sphereFill {A = A} n f = Σ[ top ∈ A ] ((x : S₊ n) → top ≡ f x) + +isSphereFilled : ℕ → Type ℓ → Type ℓ +isSphereFilled n A = (f : S₊ n → A) → sphereFill n f + +isSphereFilled∥∥ : {n : ℕ} → isSphereFilled n (HubAndSpoke A n) +isSphereFilled∥∥ f = (hub f) , (spoke f) + +isSphereFilled→isOfHLevel : (n : ℕ) → isSphereFilled n A → isOfHLevel (suc n) A +isSphereFilled→isOfHLevel {A = A} zero h x y = sym (snd (h f) true) ∙ snd (h f) false + where + f : Bool → A + f true = x + f false = y +isSphereFilled→isOfHLevel {A = A} (suc zero) h x y = + J (λ y q → (p : x ≡ y) → q ≡ p) (helper x) + where + helper : (x : A) (p : x ≡ x) → refl ≡ p + helper x p i j = + hcomp (λ k → λ { (i = i0) → h S¹→A .snd base k + ; (i = i1) → p j + ; (j = i0) → h S¹→A .snd base (i ∨ k) + ; (j = i1) → h S¹→A .snd base (i ∨ k)}) + (h S¹→A .snd (loop j) i) + where + S¹→A : S¹ → A + S¹→A base = x + S¹→A (loop i) = p i +isSphereFilled→isOfHLevel {A = A} (suc (suc n)) h x y = + isSphereFilled→isOfHLevel (suc n) (helper h x y) + where + helper : {n : ℕ} → isSphereFilled (suc (suc n)) A → (x y : A) → isSphereFilled (suc n) (x ≡ y) + helper {n = n} h x y f = sym (snd (h f') north) ∙ (snd (h f') south) , r + where + f' : Susp (S₊ (suc n)) → A + f' north = x + f' south = y + f' (merid u i) = f u i + + r : (s : S₊ (suc n)) → sym (snd (h f') north) ∙ (snd (h f') south) ≡ f s + r s i j = hcomp + (λ k → + λ { (i = i1) → snd (h f') (merid s j) k + ; (j = i0) → snd (h f') north (k ∨ (~ i)) + ; (j = i1) → snd (h f') south k + }) + (snd (h f') north (~ i ∧ ~ j)) + +isOfHLevel→isSphereFilled : (n : ℕ) → isOfHLevel (suc n) A → isSphereFilled n A +isOfHLevel→isSphereFilled zero h f = (f true) , (λ _ → h _ _) +isOfHLevel→isSphereFilled {A = A} (suc zero) h f = (f base) , toPropElim (λ _ → h _ _) refl +isOfHLevel→isSphereFilled {A = A} (suc (suc n)) h = + helper λ x y → isOfHLevel→isSphereFilled (suc n) (h x y) + where + helper : {n : ℕ} → ((x y : A) → isSphereFilled (suc n) (x ≡ y)) + → isSphereFilled (suc (suc n)) A + helper {n = n} h f = f north , r + where + r : (x : S₊ (suc (suc n))) → f north ≡ f x + r north = refl + r south = h (f north) (f south) (λ x → cong f (merid x)) .fst + r (merid x i) j = hcomp (λ k → λ { (i = i0) → f north + ; (i = i1) → h (f north) (f south) (λ x → cong f (merid x)) .snd x (~ k) j + ; (j = i0) → f north + ; (j = i1) → f (merid x i) }) (f (merid x (i ∧ j))) + +isOfHLevelTrunc : (n : ℕ) → isOfHLevel n (∥ A ∥ n) +isOfHLevelTrunc zero = isOfHLevelUnit* 0 +isOfHLevelTrunc (suc n) = isSphereFilled→isOfHLevel n isSphereFilled∥∥ + +isOfHLevelTruncPath : {n : ℕ} {x y : hLevelTrunc n A} → isOfHLevel n (x ≡ y) +isOfHLevelTruncPath {n = n} = isOfHLevelPath n (isOfHLevelTrunc n) _ _ + +rec₊ : {n : HLevel} + {B : Type ℓ'} → + isOfHLevel (suc n) B → + (A → B) → + hLevelTrunc (suc n) A → + B +rec₊ h g ∣ x ∣ = g x +rec₊ {n = n} {B = B} hB g (hub f) = isOfHLevel→isSphereFilled n hB (λ x → rec₊ hB g (f x)) .fst +rec₊ {n = n} hB g (spoke f x i) = + isOfHLevel→isSphereFilled n hB (λ x → rec₊ hB g (f x)) .snd x i + +rec : {n : HLevel} + {B : Type ℓ'} → + isOfHLevel n B → + (A → B) → + hLevelTrunc n A → + B +rec {n = zero} h g t = h .fst +rec {n = suc n} = rec₊ + +rec2 : {n : HLevel} + {B : Type ℓ'} → + isOfHLevel n B → + (A → A → B) → + (t u : hLevelTrunc n A) + → B +rec2 h g = rec (isOfHLevelΠ _ (λ _ → h)) (λ a → rec h (λ b → g a b)) + +rec3 : {n : HLevel} + {B : Type ℓ'} → + isOfHLevel n B → + (A → A → A → B) → + (t u v : hLevelTrunc n A) + → B +rec3 h g = rec2 (isOfHLevelΠ _ (λ _ → h)) (λ a b → rec h (λ c → g a b c)) + +elim₊ : {n : ℕ} + {B : ∥ A ∥ (suc n) → Type ℓ'} + (hB : (x : ∥ A ∥ (suc n)) → isOfHLevel (suc n) (B x)) + (g : (a : A) → B (∣ a ∣)) + (x : ∥ A ∥ (suc n)) → + B x +elim₊ hB g (∣ a ∣ ) = g a +elim₊ {B = B} hB g (hub f) = + isOfHLevel→isSphereFilled _ (hB (hub f)) (λ x → subst B (sym (spoke f x)) (elim₊ hB g (f x)) ) .fst +elim₊ {B = B} hB g (spoke f x i) = + toPathP {A = λ i → B (spoke f x (~ i))} + (sym (isOfHLevel→isSphereFilled _ (hB (hub f)) (λ x → subst B (sym (spoke f x)) (elim₊ hB g (f x))) .snd x)) + (~ i) + +elim : {n : ℕ} + {B : ∥ A ∥ n → Type ℓ'} + (hB : (x : ∥ A ∥ n) → isOfHLevel n (B x)) + (g : (a : A) → B (∣ a ∣ₕ)) + (x : ∥ A ∥ n) → + B x +elim {n = zero} h g t = h t .fst +elim {n = suc n} = elim₊ + +elim2 : {n : ℕ} + {B : ∥ A ∥ n → ∥ A ∥ n → Type ℓ'} + (hB : ((x y : ∥ A ∥ n) → isOfHLevel n (B x y))) + (g : (a b : A) → B ∣ a ∣ₕ ∣ b ∣ₕ) + (x y : ∥ A ∥ n) → + B x y +elim2 hB g = elim (λ _ → isOfHLevelΠ _ (λ _ → hB _ _)) λ a → + elim (λ _ → hB _ _) (λ b → g a b) + +elim3 : {n : ℕ} + {B : (x y z : ∥ A ∥ n) → Type ℓ'} + (hB : ((x y z : ∥ A ∥ n) → isOfHLevel n (B x y z))) + (g : (a b c : A) → B (∣ a ∣ₕ) ∣ b ∣ₕ ∣ c ∣ₕ) + (x y z : ∥ A ∥ n) → + B x y z +elim3 hB g = elim2 (λ _ _ → isOfHLevelΠ _ (hB _ _)) λ a b → + elim (λ _ → hB _ _ _) (λ c → g a b c) + +isContr→isContr∥ : (n : ℕ) → isContr A → isContr (∥ A ∥ n) +isContr→isContr∥ zero _ = tt* , (λ _ _ → tt*) +isContr→isContr∥ (suc n) contr = ∣ fst contr ∣ , (elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (λ a i → ∣ snd contr a i ∣)) + +isOfHLevelMin→isOfHLevel : {n m : ℕ} → isOfHLevel (min n m) A → isOfHLevel n A × isOfHLevel m A +isOfHLevelMin→isOfHLevel {n = zero} {m = m} h = h , isContr→isOfHLevel m h +isOfHLevelMin→isOfHLevel {n = suc n} {m = zero} h = (isContr→isOfHLevel (suc n) h) , h +isOfHLevelMin→isOfHLevel {A = A} {n = suc n} {m = suc m} h = + subst (λ x → isOfHLevel x A) (helper n m) + (isOfHLevelPlus (suc n ∸ (suc (min n m))) h) + , subst (λ x → isOfHLevel x A) ((λ i → m ∸ (minComm n m i) + suc (minComm n m i)) ∙ helper m n) + (isOfHLevelPlus (suc m ∸ (suc (min n m))) h) + where + helper : (n m : ℕ) → n ∸ min n m + suc (min n m) ≡ suc n + helper zero zero = refl + helper zero (suc m) = refl + helper (suc n) zero = cong suc (+-comm n 1) + helper (suc n) (suc m) = +-suc _ _ ∙ cong suc (helper n m) + +ΣTruncElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {n m : ℕ} + {B : (x : ∥ A ∥ n) → Type ℓ'} + {C : (Σ[ a ∈ (∥ A ∥ n) ] (∥ B a ∥ m)) → Type ℓ''} + → ((x : (Σ[ a ∈ (∥ A ∥ n) ] (∥ B a ∥ m))) → isOfHLevel (min n m) (C x)) + → ((a : A) (b : B (∣ a ∣ₕ)) → C (∣ a ∣ₕ , ∣ b ∣ₕ)) + → (x : (Σ[ a ∈ (∥ A ∥ n) ] (∥ B a ∥ m))) → C x +ΣTruncElim hB g (a , b) = + elim (λ x → isOfHLevelΠ _ λ b → isOfHLevelMin→isOfHLevel (hB (x , b)) .fst ) + (λ a → elim (λ _ → isOfHLevelMin→isOfHLevel (hB (∣ a ∣ₕ , _)) .snd) λ b → g a b) + a b + +truncIdempotentIso : (n : ℕ) → isOfHLevel n A → Iso (∥ A ∥ n) A +truncIdempotentIso zero hA = isContr→Iso (isOfHLevelUnit* 0) hA +Iso.fun (truncIdempotentIso (suc n) hA) = rec hA λ a → a +Iso.inv (truncIdempotentIso (suc n) hA) = ∣_∣ +Iso.rightInv (truncIdempotentIso (suc n) hA) _ = refl +Iso.leftInv (truncIdempotentIso (suc n) hA) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ _ → refl + +truncIdempotent≃ : (n : ℕ) → isOfHLevel n A → ∥ A ∥ n ≃ A +truncIdempotent≃ n hA = isoToEquiv (truncIdempotentIso n hA) + +truncIdempotent : (n : ℕ) → isOfHLevel n A → ∥ A ∥ n ≡ A +truncIdempotent n hA = ua (truncIdempotent≃ n hA) + +HLevelTruncModality : ∀ {ℓ} (n : HLevel) → Modality ℓ +isModal (HLevelTruncModality n) = isOfHLevel n +isPropIsModal (HLevelTruncModality n) = isPropIsOfHLevel n +◯ (HLevelTruncModality n) = hLevelTrunc n +◯-isModal (HLevelTruncModality n) = isOfHLevelTrunc n +η (HLevelTruncModality zero) _ = tt* +η (HLevelTruncModality (suc n)) = ∣_∣ +◯-elim (HLevelTruncModality zero) cB _ tt* = cB tt* .fst +◯-elim (HLevelTruncModality (suc n)) = elim +◯-elim-β (HLevelTruncModality zero) cB f a = cB tt* .snd (f a) +◯-elim-β (HLevelTruncModality (suc n)) = λ _ _ _ → refl +◯-=-isModal (HLevelTruncModality zero) x y = (isOfHLevelUnit* 1 x y) , (isOfHLevelUnit* 2 x y _) +◯-=-isModal (HLevelTruncModality (suc n)) = isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) + +-- universal property + +univTrunc : ∀ {ℓ} (n : HLevel) {B : TypeOfHLevel ℓ n} → Iso (hLevelTrunc n A → B .fst) (A → B .fst) +univTrunc zero {B , lev} = isContr→Iso (isOfHLevelΠ 0 (λ _ → lev)) (isOfHLevelΠ 0 λ _ → lev) +Iso.fun (univTrunc (suc n) {B , lev}) g a = g ∣ a ∣ +Iso.inv (univTrunc (suc n) {B , lev}) = rec lev +Iso.rightInv (univTrunc (suc n) {B , lev}) b = refl +Iso.leftInv (univTrunc (suc n) {B , lev}) b = funExt (elim (λ x → isOfHLevelPath _ lev _ _) + λ a → refl) + +-- some useful properties of recursor + +recUniq : {n : HLevel} + → (h : isOfHLevel n B) + → (g : A → B) + → (x : A) + → rec h g ∣ x ∣ₕ ≡ g x +recUniq {n = zero} h g x = h .snd (g x) +recUniq {n = suc n} _ _ _ = refl + +∘rec : ∀{ℓ''} {n : HLevel}{C : Type ℓ''} + → (h : isOfHLevel n B) + → (h' : isOfHLevel n C) + → (g : A → B) + → (f : B → C) + → (x : hLevelTrunc n A) + → rec h' (f ∘ g) x ≡ f (rec h g x) +∘rec {n = zero} h h' g f x = h' .snd (f (rec h g x)) +∘rec {n = suc n} h h' g f = elim (λ _ → isOfHLevelPath _ h' _ _) (λ _ → refl) + +recId : {n : HLevel} + → (f : A → hLevelTrunc n A) + → ((x : A) → f x ≡ ∣ x ∣ₕ) + → rec (isOfHLevelTrunc _) f ≡ idfun _ +recId {n = n} f h i x = + elim {B = λ a → rec (isOfHLevelTrunc _) f a ≡ a} + (λ _ → isOfHLevelTruncPath) (λ a → recUniq {n = n} (isOfHLevelTrunc _) f a ∙ h a) x i + +-- functorial action + +map : {n : HLevel} {B : Type ℓ'} (g : A → B) + → hLevelTrunc n A → hLevelTrunc n B +map g = rec (isOfHLevelTrunc _) (λ a → ∣ g a ∣ₕ) + +map2 : ∀ {ℓ''} {n : HLevel} {B : Type ℓ'} {C : Type ℓ''} (g : A → B → C) + → hLevelTrunc n A → hLevelTrunc n B → hLevelTrunc n C +map2 {n = zero} g = λ _ _ → tt* +map2 {n = suc n} g = rec (isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc _) + (λ a → rec (isOfHLevelTrunc _) λ b → ∣ g a b ∣) + +mapCompIso : {n : HLevel} {B : Type ℓ'} → (Iso A B) → Iso (hLevelTrunc n A) (hLevelTrunc n B) +mapCompIso {n = zero} {B} _ = isContr→Iso (isOfHLevelUnit* 0) (isOfHLevelUnit* 0) +Iso.fun (mapCompIso {n = (suc n)} g) = map (Iso.fun g) +Iso.inv (mapCompIso {n = (suc n)} g) = map (Iso.inv g) +Iso.rightInv (mapCompIso {n = (suc n)} g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ b → cong ∣_∣ (Iso.rightInv g b) +Iso.leftInv (mapCompIso {n = (suc n)} g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ a → cong ∣_∣ (Iso.leftInv g a) + +mapId : {n : HLevel} → ∀ t → map {n = n} (idfun A) t ≡ t +mapId {n = 0} tt* = refl +mapId {n = (suc n)} = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) (λ _ → refl) + +-- equivalences to prop/set/groupoid truncations + +propTruncTrunc1Iso : Iso ∥ A ∥₁ (∥ A ∥ 1) +Iso.fun propTruncTrunc1Iso = PropTrunc.rec (isOfHLevelTrunc 1) ∣_∣ +Iso.inv propTruncTrunc1Iso = rec squash₁ ∣_∣₁ +Iso.rightInv propTruncTrunc1Iso = elim (λ _ → isOfHLevelPath 1 (isOfHLevelTrunc 1) _ _) (λ _ → refl) +Iso.leftInv propTruncTrunc1Iso = PropTrunc.elim (λ _ → isOfHLevelPath 1 squash₁ _ _) (λ _ → refl) + +propTrunc≃Trunc1 : ∥ A ∥₁ ≃ ∥ A ∥ 1 +propTrunc≃Trunc1 = isoToEquiv propTruncTrunc1Iso + +propTrunc≡Trunc1 : ∥ A ∥₁ ≡ ∥ A ∥ 1 +propTrunc≡Trunc1 = ua propTrunc≃Trunc1 + + +setTruncTrunc2Iso : Iso ∥ A ∥₂ (∥ A ∥ 2) +Iso.fun setTruncTrunc2Iso = SetTrunc.rec (isOfHLevelTrunc 2) ∣_∣ +Iso.inv setTruncTrunc2Iso = rec squash₂ ∣_∣₂ +Iso.rightInv setTruncTrunc2Iso = elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) (λ _ → refl) +Iso.leftInv setTruncTrunc2Iso = SetTrunc.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) (λ _ → refl) + +setTrunc≃Trunc2 : ∥ A ∥₂ ≃ ∥ A ∥ 2 +setTrunc≃Trunc2 = isoToEquiv setTruncTrunc2Iso + +propTrunc≡Trunc2 : ∥ A ∥₂ ≡ ∥ A ∥ 2 +propTrunc≡Trunc2 = ua setTrunc≃Trunc2 + +groupoidTruncTrunc3Iso : Iso ∥ A ∥₃ (∥ A ∥ 3) +Iso.fun groupoidTruncTrunc3Iso = GpdTrunc.rec (isOfHLevelTrunc 3) ∣_∣ +Iso.inv groupoidTruncTrunc3Iso = rec squash₃ ∣_∣₃ +Iso.rightInv groupoidTruncTrunc3Iso = elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) (λ _ → refl) +Iso.leftInv groupoidTruncTrunc3Iso = GpdTrunc.elim (λ _ → isOfHLevelPath 3 squash₃ _ _) (λ _ → refl) + +groupoidTrunc≃Trunc3 : ∥ A ∥₃ ≃ ∥ A ∥ 3 +groupoidTrunc≃Trunc3 = isoToEquiv groupoidTruncTrunc3Iso + +groupoidTrunc≡Trunc3 : ∥ A ∥₃ ≡ ∥ A ∥ 3 +groupoidTrunc≡Trunc3 = ua groupoidTrunc≃Trunc3 + +2GroupoidTruncTrunc4Iso : Iso ∥ A ∥₄ (∥ A ∥ 4) +Iso.fun 2GroupoidTruncTrunc4Iso = 2GpdTrunc.rec (isOfHLevelTrunc 4) ∣_∣ +Iso.inv 2GroupoidTruncTrunc4Iso = rec squash₄ ∣_∣₄ +Iso.rightInv 2GroupoidTruncTrunc4Iso = elim (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (λ _ → refl) +Iso.leftInv 2GroupoidTruncTrunc4Iso = 2GpdTrunc.elim (λ _ → isOfHLevelPath 4 squash₄ _ _) (λ _ → refl) + +2GroupoidTrunc≃Trunc4 : ∥ A ∥₄ ≃ ∥ A ∥ 4 +2GroupoidTrunc≃Trunc4 = isoToEquiv 2GroupoidTruncTrunc4Iso + +2GroupoidTrunc≡Trunc4 : ∥ A ∥₄ ≡ ∥ A ∥ 4 +2GroupoidTrunc≡Trunc4 = ua 2GroupoidTrunc≃Trunc4 + +isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) +isContr→isContrTrunc zero contr = isOfHLevelUnit* 0 +isContr→isContrTrunc (suc n) contr = + ∣ fst contr ∣ , (elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ a → cong ∣_∣ (snd contr a)) + +truncOfProdIso : (n : ℕ) → Iso (hLevelTrunc n (A × B)) (hLevelTrunc n A × hLevelTrunc n B) +truncOfProdIso 0 = isContr→Iso (isOfHLevelUnit* 0) (isOfHLevel× 0 (isOfHLevelUnit* 0) (isOfHLevelUnit* 0)) +Iso.fun (truncOfProdIso (suc n)) = rec (isOfHLevelΣ (suc n) (isOfHLevelTrunc (suc n)) (λ _ → isOfHLevelTrunc (suc n))) λ {(a , b) → ∣ a ∣ , ∣ b ∣} +Iso.inv (truncOfProdIso (suc n)) (a , b) = rec (isOfHLevelTrunc (suc n)) + (λ a → rec (isOfHLevelTrunc (suc n)) + (λ b → ∣ a , b ∣) + b) + a +Iso.rightInv (truncOfProdIso (suc n)) (a , b) = + elim {B = λ a → Iso.fun (truncOfProdIso (suc n)) (Iso.inv (truncOfProdIso (suc n)) (a , b)) ≡ (a , b)} + (λ _ → isOfHLevelPath (suc n) (isOfHLevelΣ (suc n) (isOfHLevelTrunc (suc n)) (λ _ → isOfHLevelTrunc (suc n))) _ _) + (λ a → elim {B = λ b → Iso.fun (truncOfProdIso (suc n)) (Iso.inv (truncOfProdIso (suc n)) (∣ a ∣ , b)) ≡ (∣ a ∣ , b)} + (λ _ → isOfHLevelPath (suc n) (isOfHLevelΣ (suc n) (isOfHLevelTrunc (suc n)) (λ _ → isOfHLevelTrunc (suc n))) _ _) + (λ b → refl) b) a +Iso.leftInv (truncOfProdIso (suc n)) = elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ a → refl + +---- ∥ Ω A ∥ ₙ ≡ Ω ∥ A ∥ₙ₊₁ ---- + + + {- Proofs of Theorem 7.3.12. and Corollary 7.3.13. in the HoTT book -} + +module ΩTrunc {X : Type ℓ} {n : HLevel} where + {- We define the fibration P to show a more general result -} + Code : ∥ X ∥ (2 + n) → ∥ X ∥ (2 + n) → TypeOfHLevel ℓ (suc n) + Code x y = + rec2 (isOfHLevelTypeOfHLevel (suc n)) + (λ a b → ∥ a ≡ b ∥ (suc n) , isOfHLevelTrunc (suc n)) x y + + P : ∥ X ∥ (2 + n) → ∥ X ∥ (2 + n) → Type ℓ + P x y = Code x y .fst + + {- We will need P to be of hLevel n + 3 -} + hLevelP : (a b : ∥ X ∥ (2 + n)) → isOfHLevel (2 + n) (P a b) + hLevelP x y = isOfHLevelSuc (suc n) (Code x y .snd) + + {- decode function from P x y to x ≡ y -} + decode-fun : (x y : ∥ X ∥ (2 + n)) → P x y → x ≡ y + decode-fun = + elim2 (λ u v → isOfHLevelΠ (2 + n)(λ _ → isOfHLevelSuc (2 + n) (isOfHLevelTrunc (2 + n)) u v)) + decode* + where + decode* : (u v : X) → P ∣ u ∣ ∣ v ∣ → Path (∥ X ∥ (2 + n)) ∣ u ∣ ∣ v ∣ + decode* u v = + rec (isOfHLevelTrunc (2 + n) ∣ u ∣ ∣ v ∣) (cong ∣_∣) + + {- auxiliary function r used to define encode -} + r : (u : ∥ X ∥ (2 + n)) → P u u + r = elim (λ x → hLevelP x x) (λ a → ∣ refl ∣) + + encode-fun : (x y : ∥ X ∥ (2 + n)) → x ≡ y → P x y + encode-fun x y p = subst (P x) p (r x) + + encode-fill : (x y : ∥ X ∥ (2 + n)) (p : x ≡ y) + → PathP (λ i → P x (p i)) (r x) (encode-fun x y p) + encode-fill x y p = subst-filler (P x) p (r x) + + {- We need the following lemma on the functions behaviour for refl -} + dec-refl : (x : ∥ X ∥ (2 + n)) → decode-fun x x (r x) ≡ refl + dec-refl = + elim (λ x → isOfHLevelSuc (1 + n) + (isOfHLevelSuc (1 + n) + (isOfHLevelTrunc (2 + n) x x) + (decode-fun x x (r x)) refl)) + (λ _ → refl) + + {- decode-fun is a right-inverse -} + P-rinv : (u v : ∥ X ∥ (2 + n)) (x : Path (∥ X ∥ (2 + n)) u v) + → decode-fun u v (encode-fun u v x) ≡ x + P-rinv u v p = + transport + (λ i → decode-fun u (p i) (encode-fill u v p i) ≡ (λ j → p (i ∧ j))) + (dec-refl u) + + {- decode-fun is a left-inverse -} + P-linv : (u v : ∥ X ∥ (2 + n)) (x : P u v) + → encode-fun u v (decode-fun u v x) ≡ x + P-linv = + elim2 (λ x y → isOfHLevelΠ (2 + n) (λ z → isOfHLevelSuc (2 + n) (hLevelP x y) _ _)) + (λ a b → elim (λ p → hLevelP ∣ a ∣ ∣ b ∣ _ _) (helper a b)) + where + helper : (a b : X) (p : a ≡ b) → encode-fun _ _ (decode-fun ∣ a ∣ ∣ b ∣ (∣ p ∣)) ≡ ∣ p ∣ + helper a b p = + transport + (λ i → subst-filler (P ∣ a ∣) (cong ∣_∣ p) ∣ refl ∣ i ≡ ∣ (λ j → p (i ∧ j)) ∣) + refl + + {- The final Iso established -} + IsoFinal : (x y : ∥ X ∥ (2 + n)) → Iso (x ≡ y) (P x y) + Iso.fun (IsoFinal x y) = encode-fun x y + Iso.inv (IsoFinal x y) = decode-fun x y + Iso.rightInv (IsoFinal x y) = P-linv x y + Iso.leftInv (IsoFinal x y) = P-rinv x y + + +P : (x y z : ∥ X ∥ (2 + n)) → (P x y) → (P y z) → P x z + +P = + elim3 (λ x _ z → isOfHLevelΠ (2 + n) λ _ → isOfHLevelΠ (2 + n) λ _ → hLevelP x z) λ a b c → + rec (isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) λ p → + rec (isOfHLevelTrunc (suc n)) λ q → + ∣ p ∙ q ∣ + + +P-funct : (x y z : ∥ X ∥ (2 + n)) (p : x ≡ y) (q : y ≡ z) + → +P x y z (encode-fun x y p) (encode-fun y z q) ≡ encode-fun x z (p ∙ q) + +P-funct x y z p q = + transport + (λ i → +P x y (q i) (encode-fun x y p) (encode-fill y z q i) ≡ encode-fun x (q i) (compPath-filler p q i)) + (transport + (λ i → +P x (p i) (p i) (encode-fill x y p i) (r (p i)) ≡ encode-fill x y p i) + (elim {B = λ x → +P x x x (r x) (r x) ≡ r x} + (λ x → isOfHLevelPath (2 + n) (hLevelP x x) _ _) + (λ a i → ∣ rUnit refl (~ i) ∣) + x)) + +PathIdTruncIso : {a b : A} (n : HLevel) → Iso (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ b ∣) (∥ a ≡ b ∥ n) +PathIdTruncIso zero = + isContr→Iso + (isOfHLevelTrunc 1 _ _ , isOfHLevelPath 1 (isOfHLevelTrunc 1) ∣ _ ∣ ∣ _ ∣ _) + (isOfHLevelUnit* 0) +PathIdTruncIso (suc n) = ΩTrunc.IsoFinal ∣ _ ∣ ∣ _ ∣ + +PathIdTrunc : {a b : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) +PathIdTrunc n = isoToPath (PathIdTruncIso n) + +PathΩ : {a : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) +PathΩ n = PathIdTrunc n + +PathIdTruncIsoFunct : ∀ {A : Type ℓ} {a : A} (n : HLevel) → (p q : (Path (∥ A ∥ (2 + n)) ∣ a ∣ ∣ a ∣)) + → Iso.fun (PathIdTruncIso (suc n)) (p ∙ q) + ≡ map2 _∙_ (Iso.fun (PathIdTruncIso (suc n)) p) (Iso.fun (PathIdTruncIso (suc n)) q) +PathIdTruncIsoFunct {a = a} n p q = sym (ΩTrunc.+P-funct (∣ a ∣) ∣ a ∣ ∣ a ∣ p q) + +------------------------- + +truncOfTruncIso : (n m : HLevel) → Iso (hLevelTrunc n A) (hLevelTrunc n (hLevelTrunc (m + n) A)) +truncOfTruncIso zero m = isContr→Iso (isOfHLevelUnit* 0) (isOfHLevelUnit* 0) +Iso.fun (truncOfTruncIso (suc n) zero) = rec (isOfHLevelTrunc (suc n)) λ a → ∣ ∣ a ∣ ∣ +Iso.fun (truncOfTruncIso (suc n) (suc m)) = rec (isOfHLevelTrunc (suc n)) λ a → ∣ ∣ a ∣ ∣ +Iso.inv (truncOfTruncIso (suc n) zero) = rec (isOfHLevelTrunc (suc n)) + (rec (isOfHLevelTrunc (suc n)) + λ a → ∣ a ∣) +Iso.inv (truncOfTruncIso (suc n) (suc m)) = rec (isOfHLevelTrunc (suc n)) + (rec (isOfHLevelPlus (suc m) (isOfHLevelTrunc (suc n))) + λ a → ∣ a ∣) +Iso.rightInv (truncOfTruncIso (suc n) zero) = + elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _ ) + (elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _ ) + λ a → refl) +Iso.rightInv (truncOfTruncIso (suc n) (suc m)) = + elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _ ) + (elim (λ x → isOfHLevelPath ((suc m) + (suc n)) (isOfHLevelPlus (suc m) (isOfHLevelTrunc (suc n))) _ _ ) + λ a → refl) +Iso.leftInv (truncOfTruncIso (suc n) zero) = + elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ a → refl +Iso.leftInv (truncOfTruncIso (suc n) (suc m)) = + elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ a → refl + +truncOfTruncIso' : (n m : HLevel) → Iso (hLevelTrunc n A) (hLevelTrunc n (hLevelTrunc (n + m) A)) +truncOfTruncIso' zero m = isContr→Iso (isOfHLevelUnit* 0) (isOfHLevelUnit* 0) +Iso.fun (truncOfTruncIso' (suc n) m) = rec (isOfHLevelTrunc (suc n)) λ a → ∣ ∣ a ∣ ∣ +Iso.inv (truncOfTruncIso' {A = A} (suc n) m) = + rec (isOfHLevelTrunc (suc n)) + (rec (isOfHLevelPlus' {n = m} (suc n) (isOfHLevelTrunc (suc n))) ∣_∣) +Iso.rightInv (truncOfTruncIso' (suc n) m) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (elim (λ _ → isOfHLevelPath (suc n + m) (isOfHLevelPlus' {n = m} (suc n) (isOfHLevelTrunc (suc n))) _ _) + λ _ → refl) +Iso.leftInv (truncOfTruncIso' (suc n) m) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ _ → refl + +truncOfTruncEq : (n m : ℕ) → (hLevelTrunc n A) ≃ (hLevelTrunc n (hLevelTrunc (m + n) A)) +truncOfTruncEq n m = isoToEquiv (truncOfTruncIso n m) + +truncOfTruncIso2 : (n m : HLevel) → Iso (hLevelTrunc (m + n) (hLevelTrunc n A)) (hLevelTrunc n A) +truncOfTruncIso2 n m = truncIdempotentIso (m + n) (isOfHLevelPlus m (isOfHLevelTrunc n)) + +truncOfΣIso : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : A → Type ℓ'} + → Iso (hLevelTrunc n (Σ A B)) (hLevelTrunc n (Σ A λ x → hLevelTrunc n (B x))) +truncOfΣIso zero = idIso +Iso.fun (truncOfΣIso (suc n)) = map λ {(a , b) → a , ∣ b ∣} +Iso.inv (truncOfΣIso (suc n)) = + rec (isOfHLevelTrunc (suc n)) + (uncurry λ a → rec (isOfHLevelTrunc (suc n)) λ b → ∣ a , b ∣) +Iso.rightInv (truncOfΣIso (suc n)) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (uncurry λ a → elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ b → refl) +Iso.leftInv (truncOfΣIso (suc n)) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ {(a , b) → refl} + +{- transport along family of truncations -} + +transportTrunc : {n : HLevel}{p : A ≡ B} + → (a : A) + → transport (λ i → hLevelTrunc n (p i)) ∣ a ∣ₕ ≡ ∣ transport (λ i → p i) a ∣ₕ +transportTrunc {n = zero} a = refl +transportTrunc {n = suc n} a = refl diff --git a/Cubical/HITs/TypeQuotients.agda b/Cubical/HITs/TypeQuotients.agda new file mode 100644 index 0000000000..0059b89aad --- /dev/null +++ b/Cubical/HITs/TypeQuotients.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.TypeQuotients where + +open import Cubical.HITs.TypeQuotients.Base public +open import Cubical.HITs.TypeQuotients.Properties public diff --git a/Cubical/HITs/TypeQuotients/Base.agda b/Cubical/HITs/TypeQuotients/Base.agda new file mode 100644 index 0000000000..7a8cfb2169 --- /dev/null +++ b/Cubical/HITs/TypeQuotients/Base.agda @@ -0,0 +1,17 @@ +{- + +This file contains: + +- Definition of type quotients (i.e. non-truncated quotients) + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.TypeQuotients.Base where + +open import Cubical.Core.Primitives + +-- Type quotients as a higher inductive type: +data _/ₜ_ {ℓ ℓ'} (A : Type ℓ) (R : A → A → Type ℓ') : Type (ℓ-max ℓ ℓ') where + [_] : (a : A) → A /ₜ R + eq/ : (a b : A) → (r : R a b) → [ a ] ≡ [ b ] + diff --git a/Cubical/HITs/TypeQuotients/Properties.agda b/Cubical/HITs/TypeQuotients/Properties.agda new file mode 100644 index 0000000000..1616700612 --- /dev/null +++ b/Cubical/HITs/TypeQuotients/Properties.agda @@ -0,0 +1,60 @@ +{- + +Type quotients: + +-} + +{-# OPTIONS --safe #-} +module Cubical.HITs.TypeQuotients.Properties where + +open import Cubical.HITs.TypeQuotients.Base + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels + +open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥ ; ∣_∣ ; squash) +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂ ; ∣_∣₂ ; squash₂ + ; isSetSetTrunc) + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + R : A → A → Type ℓ' + B : A /ₜ R → Type ℓ'' + C : A /ₜ R → A /ₜ R → Type ℓ'' + + +elim : (f : (a : A) → (B [ a ])) + → ((a b : A) (r : R a b) → PathP (λ i → B (eq/ a b r i)) (f a) (f b)) + ------------------------------------------------------------------------ + → (x : A /ₜ R) → B x +elim f feq [ a ] = f a +elim f feq (eq/ a b r i) = feq a b r i + +rec : {X : Type ℓ''} + → (f : A → X) + → (∀ (a b : A) → R a b → f a ≡ f b) + ------------------------------------- + → A /ₜ R → X +rec f feq [ a ] = f a +rec f feq (eq/ a b r i) = feq a b r i + +elimProp : ((x : A /ₜ R ) → isProp (B x)) + → ((a : A) → B ( [ a ])) + --------------------------------- + → (x : A /ₜ R) → B x +elimProp Bprop f [ a ] = f a +elimProp Bprop f (eq/ a b r i) = isOfHLevel→isOfHLevelDep 1 Bprop (f a) (f b) (eq/ a b r) i + +elimProp2 : ((x y : A /ₜ R ) → isProp (C x y)) + → ((a b : A) → C [ a ] [ b ]) + -------------------------------------- + → (x y : A /ₜ R) → C x y +elimProp2 Cprop f = elimProp (λ x → isPropΠ (λ y → Cprop x y)) + (λ x → elimProp (λ y → Cprop [ x ] y) (f x)) diff --git a/Cubical/HITs/Wedge.agda b/Cubical/HITs/Wedge.agda new file mode 100644 index 0000000000..1cf2e97475 --- /dev/null +++ b/Cubical/HITs/Wedge.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Wedge where + +open import Cubical.HITs.Wedge.Base public diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda new file mode 100644 index 0000000000..887dc002f4 --- /dev/null +++ b/Cubical/HITs/Wedge/Base.agda @@ -0,0 +1,59 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Wedge.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.HITs.Pushout.Base +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Foundations.GroupoidLaws + +_⋁_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +_⋁_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ _ → ptA) (λ _ → ptB) + + +-- Pointed versions +_⋁∙ₗ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') +A ⋁∙ₗ B = (A ⋁ B) , (inl (snd A)) + +_⋁∙ᵣ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') +A ⋁∙ᵣ B = (A ⋁ B) , (inr (snd B)) + +-- Wedge sums of functions +_∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → (f : A →∙ C) (g : B →∙ C) + → A ⋁ B → fst C +(f ∨→ g) (inl x) = fst f x +(f ∨→ g) (inr x) = fst g x +(f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ + +-- Pointed version +∨→∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → (f : A →∙ C) (g : B →∙ C) → ((A ⋁∙ₗ B) →∙ C) +fst (∨→∙ {A = A} f g) = f ∨→ g +snd (∨→∙ {A = A} f g) = snd f + +-- Wedge sum of Units is contractible +isContr-Unit⋁Unit : isContr ((Unit , tt) ⋁ (Unit , tt)) +fst isContr-Unit⋁Unit = inl tt +snd isContr-Unit⋁Unit (inl tt) = refl +snd isContr-Unit⋁Unit (inr tt) = push tt +snd isContr-Unit⋁Unit (push tt i) j = push tt (i ∧ j) + +⋁↪ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → A ⋁ B → typ A × typ B +⋁↪ {B = B} (inl x) = x , pt B +⋁↪ {A = A} (inr x) = pt A , x +⋁↪ {A = A} {B = B} (push a i) = pt A , pt B + +fold⋁ : ∀ {ℓ} {A : Pointed ℓ} → A ⋁ A → typ A +fold⋁ (inl x) = x +fold⋁ (inr x) = x +fold⋁ {A = A} (push a i) = snd A + +id∨→∙id : ∀ {ℓ} {A : Pointed ℓ} → ∨→∙ (idfun∙ A) (idfun∙ A) ≡ (fold⋁ , refl) +id∨→∙id {A = A} = + ΣPathP ((funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j → rUnit (λ _ → pt A) (~ j) i})) + , refl) diff --git a/Cubical/Homotopy/Base.agda b/Cubical/Homotopy/Base.agda new file mode 100644 index 0000000000..1faa2f9bac --- /dev/null +++ b/Cubical/Homotopy/Base.agda @@ -0,0 +1,19 @@ +{-# OPTIONS --safe #-} + +module Cubical.Homotopy.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv.Properties + +private + variable + ℓ ℓ' : Level + +_∼_ : {X : Type ℓ} {Y : X → Type ℓ'} → (f g : (x : X) → Y x) → Type (ℓ-max ℓ ℓ') +_∼_ {X = X} f g = (x : X) → f x ≡ g x + +funExt∼ : {X : Type ℓ} {Y : X → Type ℓ'} {f g : (x : X) → Y x} (H : f ∼ g) → f ≡ g +funExt∼ = funExt + +∼-refl : {X : Type ℓ} {Y : X → Type ℓ'} {f : (x : X) → Y x} → f ∼ f +∼-refl {f = f} = λ x → refl {x = f x} diff --git a/Cubical/Homotopy/BlakersMassey.agda b/Cubical/Homotopy/BlakersMassey.agda new file mode 100644 index 0000000000..102791ab74 --- /dev/null +++ b/Cubical/Homotopy/BlakersMassey.agda @@ -0,0 +1,809 @@ +{- + +A Cubical proof of Blakers-Massey Theorem (KANG Rongji, Oct. 2021) + +Based on the previous type-theoretic proof described in + Kuen-Bang Hou (Favonia), Eric Finster, Dan Licata, Peter LeFanu Lumsdaine, + "A Mechanization of the Blakers–Massey Connectivity Theorem in Homotopy Type Theory" + (https://arxiv.org/abs/1605.03227) + +Also the HoTT-Agda formalization by Favonia: + (https://github.com/HoTT/HoTT-Agda/blob/master/theorems/homotopy/BlakersMassey.agda) + +Using cubes explicitly as much as possible. + +-} +{-# OPTIONS --safe #-} +module Cubical.Homotopy.BlakersMassey where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Transport +open import Cubical.Foundations.HLevels + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Isomorphism +open Iso +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.HalfAdjoint + +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.HITs.Truncation renaming (hLevelTrunc to Trunc) +open import Cubical.HITs.Pushout hiding (PushoutGenFib) + +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.WedgeConnectivity + +module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} + (X : Type ℓ₁)(Y : Type ℓ₂)(Q : X → Y → Type ℓ₃) + {m : HLevel} (leftConn : (x : X) → isConnected (1 + m) (Σ[ y ∈ Y ] Q x y)) + {n : HLevel} (rightConn : (y : Y) → isConnected (1 + n) (Σ[ x ∈ X ] Q x y)) + where + + ℓ : Level + ℓ = ℓ-max (ℓ-max ℓ₁ ℓ₂) ℓ₃ + + leftFiber : X → Type (ℓ-max ℓ₂ ℓ₃) + leftFiber x = Σ[ y ∈ Y ] Q x y + + rightFiber : Y → Type (ℓ-max ℓ₁ ℓ₃) + rightFiber y = Σ[ x ∈ X ] Q x y + + {- We use the alternative formulation of pushout with fewer parameters -} + + PushoutQ = PushoutGen Q + + {- Some preliminary definitions for convenience -} + + fiberSquare : + {x₀ x₁ : X}{y₀ : Y}{p : PushoutQ}(q₀₀ : Q x₀ y₀)(q₁₀ : Q x₁ y₀) + → inl x₁ ≡ p → inl x₀ ≡ p → Type ℓ + fiberSquare q₀₀ q₁₀ r' r = PathP (λ i → push q₀₀ (~ i) ≡ r' i) (sym (push q₁₀)) r + + fiberSquarePush : + {x₀ x₁ : X}{y₀ y₁ : Y}(q₀₀ : Q x₀ y₀)(q₁₀ : Q x₁ y₀)(q₁₁ : Q x₁ y₁) + → inl x₀ ≡ inr y₁ → Type ℓ + fiberSquarePush q₀₀ q₁₀ q₁₁ = fiberSquare q₀₀ q₁₀ (push q₁₁) + + fiber' : {x₀ : X}{y₀ : Y}(q₀₀ : Q x₀ y₀){x₁ : X}{p : PushoutQ} → inl x₁ ≡ p → inl x₀ ≡ p → Type ℓ + fiber' {y₀ = y₀} q₀₀ {x₁ = x₁} r' r = Σ[ q₁₀ ∈ Q x₁ y₀ ] fiberSquare q₀₀ q₁₀ r' r + + fiber'Push : {x₀ x₁ : X}{y₀ y₁ : Y}(q₀₀ : Q x₀ y₀)(q₁₁ : Q x₁ y₁) → inl x₀ ≡ inr y₁ → Type ℓ + fiber'Push q₀₀ q₁₁ = fiber' q₀₀ (push q₁₁) + + leftCodeExtended : + {x₀ : X}{y₀ : Y}(q₀₀ : Q x₀ y₀) + → (x₁ : X){p : PushoutQ} → inl x₁ ≡ p → inl x₀ ≡ p → Type ℓ + leftCodeExtended {y₀ = y₀} q₀₀ x₁ r' r = Trunc (m + n) (fiber' q₀₀ r' r) + + rightCode : {x₀ : X}(y : Y) → Path PushoutQ (inl x₀) (inr y) → Type ℓ + rightCode y r = Trunc (m + n) (fiber push r) + + {- Bunch of coherence data that will be used to construct Code -} + + {- Definitions of fiber→ -} + + module _ + {x₁ : X}{y₀ : Y}(q₁₀ : Q x₁ y₀) where + + {- (x₀ , q₀₀) = (x₁ , q₁₀) -} + module _ + {y₁ : Y}(q₁₁ : Q x₁ y₁) + (r : inl x₁ ≡ inr y₁) + (p : fiberSquarePush q₁₀ q₁₀ q₁₁ r) where + + fiber→[q₀₀=q₁₀]-filler : (i j k : I) → PushoutQ + fiber→[q₀₀=q₁₀]-filler i j k' = + hfill (λ k → λ { (i = i0) → push q₁₁ (j ∧ k) + ; (i = i1) → p k j + ; (j = i0) → push q₁₀ (i ∧ ~ k) + ; (j = i1) → push q₁₁ k }) + (inS (push q₁₀ (i ∧ ~ j))) k' + + fiber→[q₀₀=q₁₀] : fiber push r + fiber→[q₀₀=q₁₀] .fst = q₁₁ + fiber→[q₀₀=q₁₀] .snd i j = fiber→[q₀₀=q₁₀]-filler i j i1 + + ∣fiber→[q₀₀=q₁₀]∣ : rightCode _ r + ∣fiber→[q₀₀=q₁₀]∣ = ∣ fiber→[q₀₀=q₁₀] ∣ₕ + + {- (y₁ , q₁₁) = (y₀ , q₁₀) -} + module _ + {x₀ : X}(q₀₀ : Q x₀ y₀) + (r : inl x₀ ≡ inr y₀) + (p : fiberSquarePush q₀₀ q₁₀ q₁₀ r) where + + fiber→[q₁₁=q₁₀]-filler : (i j k : I) → PushoutQ + fiber→[q₁₁=q₁₀]-filler i j k' = + hfill (λ k → λ { (i = i0) → push q₀₀ (j ∨ ~ k) + ; (i = i1) → p k j + ; (j = i0) → push q₀₀ (~ k) + ; (j = i1) → push q₁₀ (~ i ∨ k) }) + (inS (push q₁₀ (~ i ∨ ~ j))) k' + + fiber→[q₁₁=q₁₀] : fiber push r + fiber→[q₁₁=q₁₀] .fst = q₀₀ + fiber→[q₁₁=q₁₀] .snd i j = fiber→[q₁₁=q₁₀]-filler i j i1 + + ∣fiber→[q₁₁=q₁₀]∣ : rightCode _ r + ∣fiber→[q₁₁=q₁₀]∣ = ∣ fiber→[q₁₁=q₁₀] ∣ₕ + + {- q₀₀ = q₁₁ = q₁₀ -} + fiber→[q₀₀=q₁₀=q₁₁]-filler : + (r : inl x₁ ≡ inr y₀) + → (p : fiberSquarePush q₁₀ q₁₀ q₁₀ r) + → (i j k l : I) → PushoutQ + fiber→[q₀₀=q₁₀=q₁₁]-filler r p i j k l' = + hfill (λ l → λ { (i = i0) → fiber→[q₀₀=q₁₀]-filler q₁₀ r p j k l + ; (i = i1) → fiber→[q₁₁=q₁₀]-filler q₁₀ r p j k l + ; (j = i0) → push q₁₀ ((i ∨ (k ∧ l)) ∧ (k ∨ (i ∧ ~ l))) + ; (j = i1) → p l k + ; (k = i0) → push q₁₀ ((i ∨ j) ∧ ~ l) + ; (k = i1) → push q₁₀ ((i ∧ ~ j) ∨ l) }) + (inS (push q₁₀ ((i ∨ (~ k ∧ j)) ∧ (~ k ∨ (i ∧ ~ j))))) l' + + fiber→[q₀₀=q₁₀=q₁₁] : fiber→[q₀₀=q₁₀] q₁₀ ≡ fiber→[q₁₁=q₁₀] q₁₀ + fiber→[q₀₀=q₁₀=q₁₁] i r p .fst = q₁₀ + fiber→[q₀₀=q₁₀=q₁₁] i r p .snd j k = fiber→[q₀₀=q₁₀=q₁₁]-filler r p i j k i1 + + ∣fiber→[q₀₀=q₁₀=q₁₁]∣ : ∣fiber→[q₀₀=q₁₀]∣ q₁₀ ≡ ∣fiber→[q₁₁=q₁₀]∣ q₁₀ + ∣fiber→[q₀₀=q₁₀=q₁₁]∣ i r p = ∣ fiber→[q₀₀=q₁₀=q₁₁] i r p ∣ₕ + + {- Definitions of fiber← -} + + module _ + {x₀ : X}{y₁ : Y}(q₀₁ : Q x₀ y₁) where + + {- (x₁ , q₁₁) = (x₀ , q₀₁) -} + module _ + {y₀ : Y}(q₀₀ : Q x₀ y₀) + (r : inl x₀ ≡ inr y₁) + (p : push q₀₁ ≡ r) where + + fiber←[q₁₁=q₀₁]-filler : (i j k : I) → PushoutQ + fiber←[q₁₁=q₀₁]-filler i j k' = + hfill (λ k → λ { (i = i0) → push q₀₀ (~ j ∧ k) + ; (i = i1) → p k j + ; (j = i0) → push q₀₀ (~ i ∧ k) + ; (j = i1) → push q₀₁ i }) + (inS (push q₀₁ (i ∧ j))) k' + + fiber←[q₁₁=q₀₁] : fiber'Push q₀₀ q₀₁ r + fiber←[q₁₁=q₀₁] .fst = q₀₀ + fiber←[q₁₁=q₀₁] .snd i j = fiber←[q₁₁=q₀₁]-filler i j i1 + + ∣fiber←[q₁₁=q₀₁]∣ : leftCodeExtended q₀₀ _ (push q₀₁) r + ∣fiber←[q₁₁=q₀₁]∣ = ∣ fiber←[q₁₁=q₀₁] ∣ₕ + + {- (y₀ , q₀₀) = (y₁ , q₀₁) -} + module _ + {x₁ : X}(q₁₁ : Q x₁ y₁) + (r : inl x₀ ≡ inr y₁) + (p : push q₀₁ ≡ r) where + + fiber←[q₀₀=q₀₁]-filler : (i j k : I) → PushoutQ + fiber←[q₀₀=q₀₁]-filler i j k' = + hfill (λ k → λ { (i = i0) → push q₁₁ (~ j ∨ ~ k) + ; (i = i1) → p k j + ; (j = i0) → push q₀₁ (~ i) + ; (j = i1) → push q₁₁ (i ∨ ~ k) }) + (inS (push q₀₁ (~ i ∨ j))) k' + + fiber←[q₀₀=q₀₁] : fiber'Push q₀₁ q₁₁ r + fiber←[q₀₀=q₀₁] .fst = q₁₁ + fiber←[q₀₀=q₀₁] .snd i j = fiber←[q₀₀=q₀₁]-filler i j i1 + + ∣fiber←[q₀₀=q₀₁]∣ : leftCodeExtended q₀₁ _ (push q₁₁) r + ∣fiber←[q₀₀=q₀₁]∣ = ∣ fiber←[q₀₀=q₀₁] ∣ₕ + + {- q₀₀ = q₀₁ = q₁₁ -} + fiber←[q₀₀=q₀₁=q₁₁]-filler : + (r : inl x₀ ≡ inr y₁) + → (p : push q₀₁ ≡ r) + → (i j k l : I) → PushoutQ + fiber←[q₀₀=q₀₁=q₁₁]-filler r p i j k l' = + hfill (λ l → λ { (i = i0) → fiber←[q₁₁=q₀₁]-filler q₀₁ r p j k l + ; (i = i1) → fiber←[q₀₀=q₀₁]-filler q₀₁ r p j k l + ; (j = i0) → push q₀₁ ((i ∨ (~ k ∧ l)) ∧ (~ k ∨ (i ∧ ~ l))) + ; (j = i1) → p l k + ; (k = i0) → push q₀₁ ((i ∨ l) ∧ ~ j) + ; (k = i1) → push q₀₁ ((i ∧ ~ l) ∨ j) }) + (inS (push q₀₁ ((i ∨ (k ∧ j)) ∧ (k ∨ (i ∧ ~ j))))) l' + + fiber←[q₀₀=q₀₁=q₁₁] : fiber←[q₁₁=q₀₁] q₀₁ ≡ fiber←[q₀₀=q₀₁] q₀₁ + fiber←[q₀₀=q₀₁=q₁₁] i r p .fst = q₀₁ + fiber←[q₀₀=q₀₁=q₁₁] i r p .snd j k = fiber←[q₀₀=q₀₁=q₁₁]-filler r p i j k i1 + + ∣fiber←[q₀₀=q₀₁=q₁₁]∣ : ∣fiber←[q₁₁=q₀₁]∣ q₀₁ ≡ ∣fiber←[q₀₀=q₀₁]∣ q₀₁ + ∣fiber←[q₀₀=q₀₁=q₁₁]∣ i r p = ∣ fiber←[q₀₀=q₀₁=q₁₁] i r p ∣ₕ + + {- Definitions of fiber→← -} + + module _ + {x₁ : X}{y₀ : Y}(q₁₀ : Q x₁ y₀) where + + {- (x₀ , q₀₀) = (x₁ , q₁₀) -} + module _ + {y₁ : Y}(q₁₁ : Q x₁ y₁) + (r : inl x₁ ≡ inr y₁) + (p : fiberSquarePush q₁₀ q₁₀ q₁₁ r) where + + fiber→←[q₀₀=q₁₀]-filler : (i j k l : I) → PushoutQ + fiber→←[q₀₀=q₁₀]-filler i j k l' = + let p' = fiber→[q₀₀=q₁₀] q₁₀ q₁₁ r p .snd in + hfill (λ l → λ { (i = i0) → fiber←[q₁₁=q₀₁]-filler q₁₁ q₁₀ r p' j k l + ; (i = i1) → fiber→[q₀₀=q₁₀]-filler q₁₀ q₁₁ r p l k j + ; (j = i0) → push q₁₀ (~ k ∧ l) + ; (j = i1) → p' l k + ; (k = i0) → push q₁₀ (~ j ∧ l) + ; (k = i1) → push q₁₁ j }) + (inS (push q₁₁ (j ∧ k))) l' + + fiber→←[q₀₀=q₁₀] : fiber←[q₁₁=q₀₁] q₁₁ q₁₀ r (fiber→[q₀₀=q₁₀] q₁₀ q₁₁ r p .snd) .snd ≡ p + fiber→←[q₀₀=q₁₀] i j k = fiber→←[q₀₀=q₁₀]-filler i j k i1 + + {- (y₁ , q₁₁) = (y₀ , q₁₀) -} + module _ + {x₀ : X}(q₀₀ : Q x₀ y₀) + (r : inl x₀ ≡ inr y₀) + (p : fiberSquarePush q₀₀ q₁₀ q₁₀ r) where + + fiber→←[q₁₁=q₁₀]-filler : (i j k l : I) → PushoutQ + fiber→←[q₁₁=q₁₀]-filler i j k l' = + let p' = fiber→[q₁₁=q₁₀] q₁₀ q₀₀ r p .snd in + hfill (λ l → λ { (i = i0) → fiber←[q₀₀=q₀₁]-filler q₀₀ q₁₀ r p' j k l + ; (i = i1) → fiber→[q₁₁=q₁₀]-filler q₁₀ q₀₀ r p l k j + ; (j = i0) → push q₁₀ (~ k ∨ ~ l) + ; (j = i1) → p' l k + ; (k = i0) → push q₀₀ (~ j) + ; (k = i1) → push q₁₀ (j ∨ ~ l) }) + (inS (push q₀₀ (~ j ∨ k))) l' + + fiber→←[q₁₁=q₁₀] : fiber←[q₀₀=q₀₁] q₀₀ q₁₀ r (fiber→[q₁₁=q₁₀] q₁₀ q₀₀ r p .snd) .snd ≡ p + fiber→←[q₁₁=q₁₀] i j k = fiber→←[q₁₁=q₁₀]-filler i j k i1 + + {- q₀₀ = q₁₀ = q₁₁ -} + fiber→←hypercube : + (r : inl x₁ ≡ inr y₀) + → (p : fiberSquarePush q₁₀ q₁₀ q₁₀ r) + → PathP (λ i → fiber←[q₀₀=q₀₁=q₁₁] q₁₀ i r (fiber→[q₀₀=q₁₀=q₁₁] q₁₀ i r p .snd) .snd ≡ p) + (fiber→←[q₀₀=q₁₀] q₁₀ r p) (fiber→←[q₁₁=q₁₀] q₁₀ r p) + fiber→←hypercube r p i j u v = + hcomp (λ l → λ { (i = i0) → fiber→←[q₀₀=q₁₀]-filler q₁₀ r p j u v l + ; (i = i1) → fiber→←[q₁₁=q₁₀]-filler q₁₀ r p j u v l + ; (j = i0) → fiber←[q₀₀=q₀₁=q₁₁]-filler q₁₀ r (fiber→[q₀₀=q₁₀=q₁₁] q₁₀ i r p .snd) i u v l + ; (j = i1) → fiber→[q₀₀=q₁₀=q₁₁]-filler q₁₀ r p i l v u + ; (u = i0) → push q₁₀ ((i ∨ (~ v ∧ l)) ∧ (~ v ∨ (i ∧ ~ l))) + ; (u = i1) → fiber→[q₀₀=q₁₀=q₁₁] q₁₀ i r p .snd l v + ; (v = i0) → push q₁₀ ((i ∨ l) ∧ ~ u) + ; (v = i1) → push q₁₀ ((i ∧ ~ l) ∨ u) }) + (push q₁₀ ((i ∨ (v ∧ u)) ∧ (v ∨ (i ∧ ~ u)))) + + {- Definitions of fiber←→ -} + + module _ + {x₀ : X}{y₁ : Y}(q₀₁ : Q x₀ y₁) where + + {- (x₁ , q₁₁) = (x₀ , q₀₁) -} + module _ + {y₀ : Y}(q₀₀ : Q x₀ y₀) + (r : inl x₀ ≡ inr y₁) + (p : push q₀₁ ≡ r) where + + fiber←→[q₁₁=q₀₁]-filler : (i j k l : I) → PushoutQ + fiber←→[q₁₁=q₀₁]-filler i j k l' = + let p' = fiber←[q₁₁=q₀₁] q₀₁ q₀₀ r p .snd in + hfill (λ l → λ { (i = i0) → fiber→[q₀₀=q₁₀]-filler q₀₀ q₀₁ r p' j k l + ; (i = i1) → fiber←[q₁₁=q₀₁]-filler q₀₁ q₀₀ r p l k j + ; (j = i0) → push q₀₁ (k ∧ l) + ; (j = i1) → p' l k + ; (k = i0) → push q₀₀ (j ∧ ~ l) + ; (k = i1) → push q₀₁ l }) + (inS (push q₀₀ (j ∧ ~ k))) l' + + fiber←→[q₁₁=q₀₁] : fiber→[q₀₀=q₁₀] q₀₀ q₀₁ r (fiber←[q₁₁=q₀₁] q₀₁ q₀₀ r p .snd) .snd ≡ p + fiber←→[q₁₁=q₀₁] i j k = fiber←→[q₁₁=q₀₁]-filler i j k i1 + + {- (y₀ , q₀₀) = (y₁ , q₀₁) -} + module _ + {x₁ : X}(q₁₁ : Q x₁ y₁) + (r : inl x₀ ≡ inr y₁) + (p : push q₀₁ ≡ r) where + + fiber←→[q₀₀=q₀₁]-filler : (i j k l : I) → PushoutQ + fiber←→[q₀₀=q₀₁]-filler i j k l' = + let p' = fiber←[q₀₀=q₀₁] q₀₁ q₁₁ r p .snd in + hfill (λ l → λ { (i = i0) → fiber→[q₁₁=q₁₀]-filler q₁₁ q₀₁ r p' j k l + ; (i = i1) → fiber←[q₀₀=q₀₁]-filler q₀₁ q₁₁ r p l k j + ; (j = i0) → push q₀₁ (k ∨ ~ l) + ; (j = i1) → p' l k + ; (k = i0) → push q₀₁ (~ l) + ; (k = i1) → push q₁₁ (~ j ∨ l) }) + (inS (push q₁₁ (~ j ∨ ~ k))) l' + + fiber←→[q₀₀=q₀₁] : fiber→[q₁₁=q₁₀] q₁₁ q₀₁ r (fiber←[q₀₀=q₀₁] q₀₁ q₁₁ r p .snd) .snd ≡ p + fiber←→[q₀₀=q₀₁] i j k = fiber←→[q₀₀=q₀₁]-filler i j k i1 + + {- q₀₀ = q₀₁ = q₁₁ -} + fiber←→hypercube : + (r : inl x₀ ≡ inr y₁) + → (p : push q₀₁ ≡ r) + → PathP (λ i → fiber→[q₀₀=q₁₀=q₁₁] q₀₁ i r (fiber←[q₀₀=q₀₁=q₁₁] q₀₁ i r p .snd) .snd ≡ p) + (fiber←→[q₁₁=q₀₁] q₀₁ r p) (fiber←→[q₀₀=q₀₁] q₀₁ r p) + fiber←→hypercube r p i j u v = + hcomp (λ l → λ { (i = i0) → fiber←→[q₁₁=q₀₁]-filler q₀₁ r p j u v l + ; (i = i1) → fiber←→[q₀₀=q₀₁]-filler q₀₁ r p j u v l + ; (j = i0) → fiber→[q₀₀=q₁₀=q₁₁]-filler q₀₁ r (fiber←[q₀₀=q₀₁=q₁₁] q₀₁ i r p .snd) i u v l + ; (j = i1) → fiber←[q₀₀=q₀₁=q₁₁]-filler q₀₁ r p i l v u + ; (u = i0) → push q₀₁ ((i ∨ (v ∧ l)) ∧ (v ∨ (i ∧ ~ l))) + ; (u = i1) → fiber←[q₀₀=q₀₁=q₁₁] q₀₁ i r p .snd l v + ; (v = i0) → push q₀₁ ((i ∨ u) ∧ ~ l) + ; (v = i1) → push q₀₁ ((i ∧ ~ u) ∨ l) }) + (push q₀₁ ((i ∨ (~ v ∧ u)) ∧ (~ v ∨ (i ∧ ~ u)))) + + module Fiber→ + {x₁ : X}{y₀ : Y}(q₁₀ : Q x₁ y₀) = + WedgeConnectivity m n + (leftFiber x₁ , (y₀ , q₁₀)) (leftConn x₁) + (rightFiber y₀ , (x₁ , q₁₀)) (rightConn y₀) + (λ (y₁ , q₁₁) (x₀ , q₀₀) → + (((r : inl x₀ ≡ inr y₁) → fiberSquarePush q₀₀ q₁₀ q₁₁ r → rightCode _ r) + , isOfHLevelΠ2 _ (λ x y → isOfHLevelTrunc _))) + (λ (y₁ , q₁₁) → ∣fiber→[q₀₀=q₁₀]∣ q₁₀ q₁₁) + (λ (x₀ , q₀₀) → ∣fiber→[q₁₁=q₁₀]∣ q₁₀ q₀₀) + (∣fiber→[q₀₀=q₁₀=q₁₁]∣ q₁₀) + + fiber→ : + {x₁ : X}{y₀ : Y}(q₁₀ : Q x₁ y₀) + → {x₀ : X}(q₀₀ : Q x₀ y₀) → {y₁ : Y}(q₁₁ : Q x₁ y₁) + → (r : inl x₀ ≡ inr y₁) + → fiberSquarePush q₀₀ q₁₀ q₁₁ r → rightCode _ r + fiber→ q₁₀ q₀₀ q₁₁ = Fiber→.extension q₁₀ (_ , q₁₁) (_ , q₀₀) + + module Fiber← + {x₀ : X}{y₁ : Y}(q₀₁ : Q x₀ y₁) = + WedgeConnectivity m n + (leftFiber x₀ , (y₁ , q₀₁)) (leftConn x₀) + (rightFiber y₁ , (x₀ , q₀₁)) (rightConn y₁) + (λ (y₀ , q₀₀) (x₁ , q₁₁) → + (((r : inl x₀ ≡ inr y₁) → push q₀₁ ≡ r → leftCodeExtended q₀₀ _ (push q₁₁) r) + , isOfHLevelΠ2 _ (λ x y → isOfHLevelTrunc _))) + (λ (y₀ , q₀₀) → ∣fiber←[q₁₁=q₀₁]∣ q₀₁ q₀₀) + (λ (x₁ , q₁₁) → ∣fiber←[q₀₀=q₀₁]∣ q₀₁ q₁₁) + (∣fiber←[q₀₀=q₀₁=q₁₁]∣ q₀₁) + + fiber← : + {x₀ : X}{y₁ : Y}(q₀₁ : Q x₀ y₁) + → {y₀ : Y}(q₀₀ : Q x₀ y₀) → {x₁ : X}(q₁₁ : Q x₁ y₁) + → (r : inl x₀ ≡ inr y₁) + → push q₀₁ ≡ r → leftCodeExtended q₀₀ _ (push q₁₁) r + fiber← q₀₁ q₀₀ q₁₁ = Fiber←.extension q₀₁ (_ , q₀₀) (_ , q₁₁) + + module _ + {x₀ x₁ : X}{y₀ y₁ : Y} + (q₀₀ : Q x₀ y₀)(q₁₁ : Q x₁ y₁) + (r : inl x₀ ≡ inr y₁) where + + left→rightCodeExtended : leftCodeExtended q₀₀ _ (push q₁₁) r → rightCode _ r + left→rightCodeExtended = + rec (isOfHLevelTrunc _) (λ (q₁₀ , p) → fiber→ q₁₀ q₀₀ q₁₁ r p) + + right→leftCodeExtended : rightCode _ r → leftCodeExtended q₀₀ _ (push q₁₁) r + right→leftCodeExtended = + rec (isOfHLevelTrunc _) (λ (q₀₁ , p) → fiber← q₀₁ q₀₀ q₁₁ r p) + + {- Definition of one-side homotopy -} + + module _ + {x₁ : X}{y₀ : Y}(q₁₀ : Q x₁ y₀) where + + {- (x₀ , q₀₀) = (x₁ , q₁₀) -} + module _ + {y₁ : Y}(q₁₁ : Q x₁ y₁) + (r : inl x₁ ≡ inr y₁) + (p : fiberSquarePush q₁₀ q₁₀ q₁₁ r) where + + ∣fiber→←[q₀₀=q₁₀]∣ : right→leftCodeExtended q₁₀ q₁₁ r (fiber→ q₁₀ q₁₀ q₁₁ r p) ≡ ∣ q₁₀ , p ∣ₕ + ∣fiber→←[q₀₀=q₁₀]∣ = + (λ i → right→leftCodeExtended q₁₀ q₁₁ r (Fiber→.left q₁₀ (_ , q₁₁) i r p)) + ∙ recUniq {n = m + n} _ _ _ + ∙ (λ i → Fiber←.left q₁₁ (_ , q₁₀) i r (fiber→[q₀₀=q₁₀] q₁₀ q₁₁ r p .snd)) + ∙ (λ i → ∣ q₁₀ , fiber→←[q₀₀=q₁₀] q₁₀ q₁₁ r p i ∣ₕ) + + {- (y₁ , q₁₁) = (y₀ , q₁₀) -} + module _ + {x₀ : X}(q₀₀ : Q x₀ y₀) + (r : inl x₀ ≡ inr y₀) + (p : fiberSquarePush q₀₀ q₁₀ q₁₀ r) where + + ∣fiber→←[q₁₁=q₁₀]∣ : right→leftCodeExtended q₀₀ q₁₀ r (fiber→ q₁₀ q₀₀ q₁₀ r p) ≡ ∣ q₁₀ , p ∣ₕ + ∣fiber→←[q₁₁=q₁₀]∣ = + (λ i → right→leftCodeExtended q₀₀ q₁₀ r (Fiber→.right q₁₀ (_ , q₀₀) i r p)) + ∙ recUniq {n = m + n} _ _ _ + ∙ (λ i → Fiber←.right q₀₀ (_ , q₁₀) i r (fiber→[q₁₁=q₁₀] q₁₀ q₀₀ r p .snd)) + ∙ (λ i → ∣ q₁₀ , fiber→←[q₁₁=q₁₀] q₁₀ q₀₀ r p i ∣ₕ) + + {- q₀₀ = q₁₁ = q₁₀ -} + module _ + (r : inl x₁ ≡ inr y₀) + (p : fiberSquarePush q₁₀ q₁₀ q₁₀ r) where + + path→←Square = + (λ i j → right→leftCodeExtended q₁₀ q₁₀ r (Fiber→.homSquare q₁₀ i j r p)) + ∙₂ (λ i → recUniq {n = m + n} _ _ _) + ∙₂ (λ i j → Fiber←.homSquare q₁₀ i j r (fiber→[q₀₀=q₁₀=q₁₁] q₁₀ i r p .snd)) + ∙₂ (λ i j → ∣ (q₁₀ , fiber→←hypercube q₁₀ r p i j) ∣ₕ) + + ∣fiber→←[q₀₀=q₁₀=q₁₁]∣ : ∣fiber→←[q₀₀=q₁₀]∣ q₁₀ ≡ ∣fiber→←[q₁₁=q₁₀]∣ q₁₀ + ∣fiber→←[q₀₀=q₁₀=q₁₁]∣ i r p = path→←Square r p i + + fiber→← : + {x₁ : X}{y₀ : Y}(q₁₀ : Q x₁ y₀) + → {x₀ : X}(q₀₀ : Q x₀ y₀) → {y₁ : Y}(q₁₁ : Q x₁ y₁) + → (r : inl x₀ ≡ inr y₁) + → (p : fiberSquarePush q₀₀ q₁₀ q₁₁ r) + → right→leftCodeExtended q₀₀ q₁₁ r (fiber→ q₁₀ q₀₀ q₁₁ r p) ≡ ∣ q₁₀ , p ∣ₕ + fiber→← {x₁ = x₁} {y₀ = y₀} q₁₀ q₀₀' q₁₁' = + WedgeConnectivity.extension m n + (leftFiber x₁ , (y₀ , q₁₀)) (leftConn x₁) + (rightFiber y₀ , (x₁ , q₁₀)) (rightConn y₀) + (λ (y₁ , q₁₁) (x₀ , q₀₀) → + (( (r : inl x₀ ≡ inr y₁) → (p : fiberSquarePush q₀₀ q₁₀ q₁₁ r) + → right→leftCodeExtended q₀₀ q₁₁ r (fiber→ q₁₀ q₀₀ q₁₁ r p) ≡ ∣ q₁₀ , p ∣ₕ ) + , isOfHLevelΠ2 _ (λ x y → isOfHLevelTruncPath))) + (λ (y₁ , q₁₁) → ∣fiber→←[q₀₀=q₁₀]∣ q₁₀ q₁₁) + (λ (x₀ , q₀₀) → ∣fiber→←[q₁₁=q₁₀]∣ q₁₀ q₀₀) + (∣fiber→←[q₀₀=q₁₀=q₁₁]∣ q₁₀) + (_ , q₁₁') (_ , q₀₀') + + {- Definition of the other side homotopy -} + + module _ + {x₀ : X}{y₁ : Y}(q₀₁ : Q x₀ y₁) where + + {- (x₁ , q₁₁) = (x₀ , q₀₁) -} + module _ + {y₀ : Y}(q₀₀ : Q x₀ y₀) + (r : inl x₀ ≡ inr y₁) + (p : push q₀₁ ≡ r) where + + ∣fiber←→[q₁₁=q₀₁]∣ : left→rightCodeExtended q₀₀ q₀₁ r (fiber← q₀₁ q₀₀ q₀₁ r p) ≡ ∣ q₀₁ , p ∣ₕ + ∣fiber←→[q₁₁=q₀₁]∣ = + (λ i → left→rightCodeExtended q₀₀ q₀₁ r (Fiber←.left q₀₁ (_ , q₀₀) i r p)) + ∙ recUniq {n = m + n} _ _ _ + ∙ (λ i → Fiber→.left q₀₀ (_ , q₀₁) i r (fiber←[q₁₁=q₀₁] q₀₁ q₀₀ r p .snd)) + ∙ (λ i → ∣ q₀₁ , fiber←→[q₁₁=q₀₁] q₀₁ q₀₀ r p i ∣ₕ) + + {- (y₀ , q₀₀) = (y₁ , q₀₁) -} + module _ + {x₁ : X}(q₁₁ : Q x₁ y₁) + (r : inl x₀ ≡ inr y₁) + (p : push q₀₁ ≡ r) where + + ∣fiber←→[q₀₀=q₀₁]∣ : left→rightCodeExtended q₀₁ q₁₁ r (fiber← q₀₁ q₀₁ q₁₁ r p) ≡ ∣ q₀₁ , p ∣ₕ + ∣fiber←→[q₀₀=q₀₁]∣ = + (λ i → left→rightCodeExtended q₀₁ q₁₁ r (Fiber←.right q₀₁ (_ , q₁₁) i r p)) + ∙ recUniq {n = m + n} _ _ _ + ∙ (λ i → Fiber→.right q₁₁ (_ , q₀₁) i r (fiber←[q₀₀=q₀₁] q₀₁ q₁₁ r p .snd)) + ∙ (λ i → ∣ q₀₁ , fiber←→[q₀₀=q₀₁] q₀₁ q₁₁ r p i ∣ₕ) + + {- q₀₀ = q₀₁ = q₁₁ -} + module _ + (r : inl x₀ ≡ inr y₁) + (p : push q₀₁ ≡ r) where + + path←→Square = + (λ i j → left→rightCodeExtended q₀₁ q₀₁ r (Fiber←.homSquare q₀₁ i j r p)) + ∙₂ (λ i → recUniq {n = m + n} _ _ _) + ∙₂ (λ i j → Fiber→.homSquare q₀₁ i j r (fiber←[q₀₀=q₀₁=q₁₁] q₀₁ i r p .snd)) + ∙₂ (λ i j → ∣ q₀₁ , fiber←→hypercube q₀₁ r p i j ∣ₕ) + + ∣fiber←→[q₀₀=q₀₁=q₁₁]∣ : ∣fiber←→[q₁₁=q₀₁]∣ q₀₁ ≡ ∣fiber←→[q₀₀=q₀₁]∣ q₀₁ + ∣fiber←→[q₀₀=q₀₁=q₁₁]∣ i r p = path←→Square r p i + + fiber←→ : + {x₀ : X}{y₁ : Y}(q₀₁ : Q x₀ y₁) + → {y₀ : Y}(q₀₀ : Q x₀ y₀) → {x₁ : X}(q₁₁ : Q x₁ y₁) + → (r : inl x₀ ≡ inr y₁) + → (p : push q₀₁ ≡ r) + → left→rightCodeExtended q₀₀ q₁₁ r (fiber← q₀₁ q₀₀ q₁₁ r p) ≡ ∣ q₀₁ , p ∣ₕ + fiber←→ {x₀ = x₀} {y₁ = y₁} q₀₁ q₀₀' q₁₁' = + WedgeConnectivity.extension m n + (leftFiber x₀ , (y₁ , q₀₁)) (leftConn x₀) + (rightFiber y₁ , (x₀ , q₀₁)) (rightConn y₁) + (λ (y₀ , q₀₀) (x₁ , q₁₁) → + (( (r : inl x₀ ≡ inr y₁) → (p : push q₀₁ ≡ r) + → left→rightCodeExtended q₀₀ q₁₁ r (fiber← q₀₁ q₀₀ q₁₁ r p) ≡ ∣ q₀₁ , p ∣ₕ ) + , isOfHLevelΠ2 _ (λ x y → isOfHLevelTruncPath))) + (λ (y₀ , q₀₀) → ∣fiber←→[q₁₁=q₀₁]∣ q₀₁ q₀₀) + (λ (x₁ , q₁₁) → ∣fiber←→[q₀₀=q₀₁]∣ q₀₁ q₁₁) + (∣fiber←→[q₀₀=q₀₁=q₁₁]∣ q₀₁) + (_ , q₀₀') (_ , q₁₁') + + module _ + {x₀ x₁ : X}{y₀ y₁ : Y} + (q₀₀ : Q x₀ y₀)(q₁₁ : Q x₁ y₁) + (r : inl x₀ ≡ inr y₁) where + + left→right→leftCodeExtended : + (a : leftCodeExtended q₀₀ _ (push q₁₁) r) + → right→leftCodeExtended q₀₀ q₁₁ r (left→rightCodeExtended q₀₀ q₁₁ r a) ≡ a + left→right→leftCodeExtended a = + sym (∘rec _ _ _ (right→leftCodeExtended q₀₀ q₁₁ r) a) ∙ + (λ i → recId _ (λ (q₁₀ , p) → fiber→← q₁₀ q₀₀ q₁₁ r p) i a) + + right→left→rightCodeExtended : + (a : rightCode _ r) + → left→rightCodeExtended q₀₀ q₁₁ r (right→leftCodeExtended q₀₀ q₁₁ r a) ≡ a + right→left→rightCodeExtended a = + sym (∘rec _ _ _ (left→rightCodeExtended q₀₀ q₁₁ r) a) ∙ + (λ i → recId _ (λ (q₀₁ , p) → fiber←→ q₀₁ q₀₀ q₁₁ r p) i a) + + left≃rightCodeExtended : leftCodeExtended q₀₀ _ (push q₁₁) r ≃ rightCode y₁ r + left≃rightCodeExtended = + isoToEquiv (iso (left→rightCodeExtended _ _ _) (right→leftCodeExtended _ _ _) + right→left→rightCodeExtended left→right→leftCodeExtended) + + {- Definition and properties of Code -} + + module _ (x₀ : X)(y₀ : Y)(q₀₀ : Q x₀ y₀) where + + leftCode' : (x : X){p : PushoutQ} → inl x ≡ p → inl x₀ ≡ p → Type ℓ + leftCode' x r' = leftCodeExtended q₀₀ x r' + + leftCode : (x : X) → inl x₀ ≡ inl x → Type ℓ + leftCode x = leftCode' x refl + + fiberPath : {x : X}{y : Y} → (q : Q x y) → leftCode' x (push q) ≡ rightCode y + fiberPath q i r = ua (left≃rightCodeExtended q₀₀ q r) i + + pushCode : + {x : X}{y : Y} → (q : Q x y) + → PathP (λ i → inl x₀ ≡ push q i → Type ℓ) (leftCode x) (rightCode y) + pushCode q i = + hcomp (λ j → λ { (i = i0) → leftCode _ + ; (i = i1) → fiberPath q j }) + (leftCode' _ (λ j → push q (i ∧ j))) + + Code : (p : PushoutQ) → inl x₀ ≡ p → Type ℓ + Code (inl x) = leftCode x + Code (inr y) = rightCode y + Code (push q i) = pushCode q i + + {- Transportation rule of pushCode -} + + transpLeftCode : (y : Y) → (q : Q x₀ y) → (q' : leftCodeExtended q₀₀ _ refl refl) → leftCode' _ (push q) (push q) + transpLeftCode y q q' = + transport (λ i → leftCode' _ (λ j → push q (i ∧ j)) (λ j → push q (i ∧ j))) q' + + transpPushCodeβ' : + (y : Y) → (q : Q x₀ y) → (q' : leftCodeExtended q₀₀ _ refl refl) + → transport (λ i → pushCode q i (λ j → push q (i ∧ j))) q' ≡ left→rightCodeExtended _ _ _ (transpLeftCode y q q') + transpPushCodeβ' y q q' i = transportRefl (left→rightCodeExtended _ _ _ (transpLeftCode y q (transportRefl q' i))) i + + module _ + {p : PushoutQ}(r : inl x₀ ≡ p) where + + fiber-filler : I → Type ℓ + fiber-filler i = fiber' q₀₀ (λ j → r (i ∧ j)) (λ j → r (i ∧ j)) + + module _ + (q : fiberSquare q₀₀ q₀₀ refl refl) where + + transpLeftCode-filler : (i j k : I) → PushoutQ + transpLeftCode-filler i j k' = + hfill (λ k → λ { (i = i0) → push q₀₀ (~ j) + ; (i = i1) → r (j ∧ k) + ; (j = i0) → push q₀₀ (~ i) + ; (j = i1) → r (i ∧ k) }) + (inS (q i j)) k' + + transpLeftCodeβ' : + {p : PushoutQ} → (r : inl x₀ ≡ p) → (q : fiberSquare q₀₀ q₀₀ refl refl) + → transport (λ i → fiber-filler r i) (q₀₀ , q) ≡ (q₀₀ , λ i j → transpLeftCode-filler r q i j i1) + transpLeftCodeβ' r q = + J (λ p r → transport (λ i → fiber-filler r i) (q₀₀ , q) ≡ (q₀₀ , λ i j → transpLeftCode-filler r q i j i1)) + (transportRefl _ ∙ (λ k → (q₀₀ , λ i j → transpLeftCode-filler refl q i j k))) r + + transpLeftCodeβ : + (y : Y) → (q : Q x₀ y) → (q' : fiberSquare q₀₀ q₀₀ refl refl) + → transpLeftCode y q ∣ q₀₀ , q' ∣ₕ ≡ ∣ q₀₀ , (λ i j → transpLeftCode-filler (push q) q' i j i1) ∣ₕ + transpLeftCodeβ y q q' = transportTrunc _ ∙ (λ i → ∣ transpLeftCodeβ' _ q' i ∣ₕ) + + transpPushCodeβ : + (y : Y) → (q : Q x₀ y) → (q' : fiberSquare q₀₀ q₀₀ refl refl) + → transport (λ i → pushCode q i (λ j → push q (i ∧ j))) ∣ q₀₀ , q' ∣ₕ + ≡ ∣fiber→[q₀₀=q₁₀]∣ q₀₀ q (push q) (λ i j → transpLeftCode-filler (push q) q' i j i1) + transpPushCodeβ y q q' = + transpPushCodeβ' _ _ _ + ∙ (λ i → left→rightCodeExtended _ _ _ (transpLeftCodeβ _ _ q' i)) + ∙ recUniq {n = m + n} _ _ _ + ∙ (λ i' → Fiber→.left q₀₀ (_ , q) i' (push q) (λ i j → transpLeftCode-filler (push q) q' i j i1)) + + {- The contractibility of Code -} + + centerCode : {p : PushoutQ} → (r : inl x₀ ≡ p) → Code p r + centerCode r = + transport (λ i → Code _ (λ j → r (i ∧ j))) ∣ q₀₀ , (λ i j → push q₀₀ (~ i ∧ ~ j)) ∣ₕ + + module _ + (y : Y)(q : Q x₀ y) where + + transp-filler : (i j k : I) → PushoutQ + transp-filler = transpLeftCode-filler (push q) (λ i' j' → push q₀₀ (~ i' ∧ ~ j')) + + transp-square : fiberSquare q₀₀ q₀₀ (push q) (push q) + transp-square i j = transp-filler i j i1 + + contractionCodeRefl' : + fiber→[q₀₀=q₁₀] q₀₀ q (push q) transp-square .snd ≡ refl + contractionCodeRefl' i j k = + hcomp (λ l → λ { (i = i0) → fiber→[q₀₀=q₁₀]-filler q₀₀ q (push q) transp-square j k l + ; (i = i1) → transp-square (~ j ∨ l) k + ; (j = i0) → push q (k ∧ (i ∨ l)) + ; (j = i1) → transp-square l k + ; (k = i0) → push q₀₀ (j ∧ ~ l) + ; (k = i1) → push q ((i ∧ ~ j) ∨ l) }) + (transp-filler (~ j) k i) + + contractionCodeRefl : centerCode (push q) ≡ ∣ q , refl ∣ₕ + contractionCodeRefl = transpPushCodeβ _ _ _ ∙ (λ i → ∣ q , contractionCodeRefl' i ∣ₕ) + + module _ + (y : Y)(r : inl x₀ ≡ inr y) where + + contractionCode' : (a : fiber push r) → centerCode r ≡ ∣ a ∣ₕ + contractionCode' (q , p') = J (λ r' p → centerCode r' ≡ ∣ q , p ∣ₕ) (contractionCodeRefl _ q) p' + + contractionCode : (a : Code _ r) → centerCode r ≡ a + contractionCode = elim (λ _ → isOfHLevelTruncPath) contractionCode' + + isContrCode : isContr (Code _ r) + isContrCode = centerCode r , contractionCode + + excision-helper : + (x : X) → Trunc (1 + m) (Σ[ y₀ ∈ Y ] Q x y₀) + → (y : Y) → (r : inl x ≡ inr y) → isContr (Trunc (m + n) (fiber push r)) + excision-helper x y' y r = rec (isProp→isOfHLevelSuc m isPropIsContr) (λ (y₀ , q₀₀) → isContrCode x y₀ q₀₀ y r ) y' + + {- The Main Result : Blakers-Massey Homotopy Excision Theorem -} + Excision : (x : X)(y : Y) → isConnectedFun (m + n) (push {x = x} {y = y}) + Excision x y = excision-helper x (leftConn x .fst) y + + +{- +We also give the following version of the theorem: Given a square + + g + A --------------> B + |\ ↗ | + | \ ↗ | + | \ ↗ | +f | X | inr + | / | + | / | + | / | + v / v + B -----------> Pushout f g + inl + +where X is the pullback of inl and inr + (X := Σ[ (b , c) ∈ B × C ] (inl b ≡ inr c)). + +If f in n-connected and g in m-connected, then the diagonal map +A → X is (n+m)-connected +-} + + +private + shuffleFibIso₁ : + {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : A → B) (g : A → C) (b : B) + → Iso (Σ[ c ∈ C ] Σ[ a ∈ A ] (f a ≡ b) × (g a ≡ c)) + (Σ[ a ∈ A ] ((Σ[ c ∈ C ] (g a ≡ c)) × (f a ≡ b))) + shuffleFibIso₁ f g b = + compIso (invIso Σ-assoc-Iso) + (compIso (Σ-cong-iso-fst Σ-swap-Iso) + (compIso + (Σ-cong-iso-snd (λ y → Σ-swap-Iso)) + (compIso Σ-assoc-Iso + (Σ-cong-iso-snd λ a → invIso Σ-assoc-Iso)))) + + shuffleFibIso₂ : {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : A → B) (g : A → C) (x : _) + → Iso (Σ[ a ∈ A ] ((Σ[ c ∈ C ] (g a ≡ c)) × (f a ≡ x))) + (fiber f x) + shuffleFibIso₂ f g x = Σ-cong-iso-snd + λ a → compIso (Σ-cong-iso-fst + (isContr→Iso (isContrSingl (g a)) + isContrUnit)) + lUnit×Iso + +module BlakersMassey□ {ℓ ℓ' ℓ'' : Level} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : A → B) (g : A → C) (n m : ℕ) + (con-f : isConnectedFun (suc n) f) + (con-g : isConnectedFun (suc m) g) where + + {- Some abbreviations and connectivity -} + private + fib = doubleFib f g + + B-con : (x : B) → isConnected (suc n) (Σ[ c ∈ C ] (fib x c)) + B-con x = + isConnectedRetractFromIso (suc n) + (compIso + (shuffleFibIso₁ f g x) + (shuffleFibIso₂ f g x)) + (con-f x) + + C-con : (c : C) → isConnected (suc m) (Σ[ b ∈ B ] (fib b c)) + C-con c = + isConnectedRetractFromIso (suc m) + (compIso + (compIso (Σ-cong-iso-snd + (λ _ → Σ-cong-iso-snd λ _ → Σ-swap-Iso)) + (shuffleFibIso₁ g f c)) + (shuffleFibIso₂ g f c)) + (con-g c) + + open module BM-f-g = BlakersMassey B C fib {m = n} B-con {n = m} C-con + + fib× : (B × C) → Type _ + fib× (b , c) = fib b c + + PushoutGenPath× : B × C → Type _ + PushoutGenPath× (b , c) = Path (PushoutGen fib) (inl b) (inr c) + + PushoutPath× : B × C → Type _ + PushoutPath× (b , c) = Path (Pushout f g) (inl b) (inr c) + + {- The function in question -} + toPullback : A → Σ (B × C) PushoutPath× + toPullback a = (f a , g a) , push a + + {- We redescribe toPullback as a composition of three maps, + two of which are equivs and one of which is (n+m)-connected -} + Totalfib×→Total : Σ (B × C) fib× → Σ (B × C) PushoutGenPath× + Totalfib×→Total = + TotalFun {A = B × C} {B = fib×} {C = PushoutGenPath×} (λ a → push) + + isConnectedTotalFun : isConnectedFun (n + m) Totalfib×→Total + isConnectedTotalFun = + FunConnected→TotalFunConnected (λ _ → push) (n + m) (uncurry BM-f-g.Excision) + + TotalPathGen×Iso : Iso (Σ (B × C) PushoutGenPath×) (Σ (B × C) PushoutPath×) + TotalPathGen×Iso = + Σ-cong-iso-snd λ x + → congIso (invIso (IsoPushoutPushoutGen f g)) + + Totalfib×Iso : Iso (Σ (B × C) fib×) A + fun Totalfib×Iso ((b , c) , a , p) = a + inv Totalfib×Iso a = (f a , g a) , a , refl , refl + rightInv Totalfib×Iso _ = refl + leftInv Totalfib×Iso ((b , c) , a , (p , q)) i = + ((p i) , (q i)) , (a , ((λ j → p (i ∧ j)) , (λ j → q (i ∧ j)))) + + toPullback' : A → Σ (B × C) PushoutPath× + toPullback' = + (fun TotalPathGen×Iso ∘ Totalfib×→Total) ∘ inv Totalfib×Iso + + toPullback'≡toPullback : toPullback' ≡ toPullback + toPullback'≡toPullback = + funExt λ x → ΣPathP (refl , (sym (rUnit (push x)))) + + isConnected-toPullback : isConnectedFun (n + m) toPullback + isConnected-toPullback = + subst (isConnectedFun (n + m)) toPullback'≡toPullback + (isConnectedComp + (fun TotalPathGen×Iso ∘ Totalfib×→Total) + (inv Totalfib×Iso) (n + m) + (isConnectedComp (fun TotalPathGen×Iso) Totalfib×→Total (n + m) + (isEquiv→isConnected _ (isoToIsEquiv TotalPathGen×Iso) (n + m)) + isConnectedTotalFun) + (isEquiv→isConnected _ (isoToIsEquiv (invIso Totalfib×Iso)) (n + m))) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda new file mode 100644 index 0000000000..4d34ae86dd --- /dev/null +++ b/Cubical/Homotopy/Connected.agda @@ -0,0 +1,405 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.Connected where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Fibration +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.HITs.Nullification +open import Cubical.HITs.Susp +open import Cubical.HITs.SmashProduct +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec) +open import Cubical.Homotopy.Loopspace +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.S1 +open import Cubical.Data.Bool +open import Cubical.Data.Unit + +-- Note that relative to most sources, this notation is off by +2 +isConnected : ∀ {ℓ} (n : HLevel) (A : Type ℓ) → Type ℓ +isConnected n A = isContr (hLevelTrunc n A) + +isConnectedFun : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') +isConnectedFun n f = ∀ b → isConnected n (fiber f b) + +isTruncatedFun : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') +isTruncatedFun n f = ∀ b → isOfHLevel n (fiber f b) + +isConnectedSubtr : ∀ {ℓ} {A : Type ℓ} (n m : HLevel) + → isConnected (m + n) A + → isConnected n A +isConnectedSubtr {A = A} n m iscon = + isOfHLevelRetractFromIso 0 (truncOfTruncIso n m) (helper n iscon) + where + helper : (n : ℕ) → isConnected (m + n) A → isContr (hLevelTrunc n (hLevelTrunc (m + n) A)) + helper zero iscon = isContrUnit* + helper (suc n) iscon = ∣ iscon .fst ∣ , (Trunc.elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ a → cong ∣_∣ (iscon .snd a)) + +isConnectedFunSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : HLevel) (f : A → B) + → isConnectedFun (m + n) f + → isConnectedFun n f +isConnectedFunSubtr n m f iscon b = isConnectedSubtr n m (iscon b) + +private + typeToFiberIso : ∀ {ℓ} (A : Type ℓ) → Iso A (fiber (λ (x : A) → tt) tt) + Iso.fun (typeToFiberIso A) x = x , refl + Iso.inv (typeToFiberIso A) = fst + Iso.rightInv (typeToFiberIso A) b i = fst b , (isOfHLevelSuc 1 (isPropUnit) tt tt (snd b) refl) i + Iso.leftInv (typeToFiberIso A) a = refl + + typeToFiber : ∀ {ℓ} (A : Type ℓ) → A ≡ fiber (λ (x : A) → tt) tt + typeToFiber A = isoToPath (typeToFiberIso A) + +isOfHLevelIsConnectedStable : ∀ {ℓ} {A : Type ℓ} (n : ℕ) + → isOfHLevel n (isConnected n A) +isOfHLevelIsConnectedStable {A = A} zero = + (tt* , (λ _ → refl)) , λ _ → refl +isOfHLevelIsConnectedStable {A = A} (suc n) = + isProp→isOfHLevelSuc n isPropIsContr + +module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) where + private + inv : ∀ {ℓ'''} (n : HLevel) (P : B → TypeOfHLevel ℓ''' (suc n)) + → ((a : A) → P (f a) .fst) + → (b : B) + → hLevelTrunc (suc n) (fiber f b) → P b .fst + inv n P t b = + Trunc.rec + (P b .snd) + (λ {(a , p) → subst (fst ∘ P) p (t a)}) + + isIsoPrecompose : ∀ {ℓ'''} (n : ℕ) (P : B → TypeOfHLevel ℓ''' n) + → isConnectedFun n f + → Iso ((b : B) → P b .fst) ((a : A) → P (f a) .fst) + isIsoPrecompose zero P fConn = isContr→Iso (isOfHLevelΠ _ (λ b → P b .snd)) (isOfHLevelΠ _ λ a → P (f a) .snd) + Iso.fun (isIsoPrecompose (suc n) P fConn) = _∘ f + Iso.inv (isIsoPrecompose (suc n) P fConn) t b = inv n P t b (fConn b .fst) + Iso.rightInv (isIsoPrecompose (suc n) P fConn) t = + funExt λ a → cong (inv n P t (f a)) (fConn (f a) .snd ∣ a , refl ∣) + ∙ substRefl {B = fst ∘ P} (t a) + Iso.leftInv (isIsoPrecompose (suc n) P fConn) s = + funExt λ b → + Trunc.elim + {B = λ d → inv n P (s ∘ f) b d ≡ s b} + (λ _ → isOfHLevelPath (suc n) (P b .snd) _ _) + (λ {(a , p) i → transp (λ j → P (p (j ∨ i)) .fst) i (s (p i))}) + (fConn b .fst) + + isEquivPrecompose : ∀ {ℓ'''} (n : ℕ) (P : B → TypeOfHLevel ℓ''' n) + → isConnectedFun n f + → isEquiv (λ(s : (b : B) → P b .fst) → s ∘ f) + isEquivPrecompose zero P fConn = isoToIsEquiv theIso + where + theIso : Iso ((b : B) → P b .fst) ((a : A) → P (f a) .fst) + Iso.fun theIso = λ(s : (b : B) → P b .fst) → s ∘ f + Iso.inv theIso = λ _ b → P b .snd .fst + Iso.rightInv theIso g = funExt λ x → P (f x) .snd .snd (g x) + Iso.leftInv theIso g = funExt λ x → P x .snd .snd (g x) + isEquivPrecompose (suc n) P fConn = isoToIsEquiv (isIsoPrecompose (suc n) P fConn) + + isConnectedPrecompose : (n : ℕ) → ((P : B → TypeOfHLevel (ℓ-max ℓ ℓ') n) + → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) + → isConnectedFun n f + isConnectedPrecompose zero P→sect b = isContrUnit* + isConnectedPrecompose (suc n) P→sect b = c n P→sect b , λ y → sym (fun n P→sect b y) + where + P : (n : HLevel) → ((P : B → TypeOfHLevel ℓ (suc n)) + → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) + → B → Type _ + P n s b = hLevelTrunc (suc n) (fiber f b) + + c : (n : HLevel) → ((P : B → TypeOfHLevel (ℓ-max ℓ ℓ') (suc n)) + → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) → (b : B) + → hLevelTrunc (suc n) (fiber f b) + c n s = (s λ b → (hLevelTrunc (suc n) (fiber f b) , isOfHLevelTrunc _)) .fst + λ a → ∣ a , refl ∣ + + fun : (n : HLevel) (P→sect : ((P : B → TypeOfHLevel (ℓ-max ℓ ℓ') (suc n)) + → hasSection λ(s : (b : B) → P b .fst) → s ∘ f)) + → (b : B) (w : (hLevelTrunc (suc n) (fiber f b))) + → w ≡ c n P→sect b + fun n P→sect b = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + λ a → J (λ b p → ∣ (fst a) , p ∣ ≡ c n P→sect b) + (c* (fst a)) + (snd a) + where + c* : ((a : A) → ∣ (a , refl {x = f a}) ∣ ≡ c n P→sect (f a)) + c* a = sym (cong (λ x → x a) (P→sect (λ b → hLevelTrunc (suc n) (fiber f b) , isOfHLevelTrunc _) .snd λ a → ∣ a , refl ∣)) + +isOfHLevelPrecomposeConnected : ∀ {ℓ ℓ' ℓ''} (k : HLevel) (n : HLevel) + {A : Type ℓ} {B : Type ℓ'} (P : B → TypeOfHLevel ℓ'' (k + n)) (f : A → B) + → isConnectedFun n f + → isOfHLevelFun k (λ(s : (b : B) → P b .fst) → s ∘ f) +isOfHLevelPrecomposeConnected zero n P f fConn = + elim.isEquivPrecompose f n P fConn .equiv-proof +isOfHLevelPrecomposeConnected (suc k) n P f fConn t = + isOfHLevelPath'⁻ k + λ {(s₀ , p₀) (s₁ , p₁) → + isOfHLevelRetractFromIso k (invIso ΣPathIsoPathΣ) + (subst (isOfHLevel k) + (sym (fiberPath (s₀ , p₀) (s₁ , p₁))) + (isOfHLevelRetract k + (λ {(q , α) → (funExt⁻ q) , (cong funExt⁻ α)}) + (λ {(h , β) → (funExt h) , (cong funExt β)}) + (λ _ → refl) + (isOfHLevelPrecomposeConnected k n + (λ b → (s₀ b ≡ s₁ b) , isOfHLevelPath' (k + n) (P b .snd) _ _) + f fConn + (funExt⁻ (p₀ ∙∙ refl ∙∙ sym p₁)))))} + +indMapEquiv→conType : ∀ {ℓ} {A : Type ℓ} (n : HLevel) + → ((B : TypeOfHLevel ℓ n) + → isEquiv (λ (b : (fst B)) → λ (a : A) → b)) + → isConnected n A +indMapEquiv→conType {A = A} zero BEq = isContrUnit* +indMapEquiv→conType {A = A} (suc n) BEq = + isOfHLevelRetractFromIso 0 (mapCompIso {n = (suc n)} (typeToFiberIso A)) + (elim.isConnectedPrecompose (λ _ → tt) (suc n) + (λ P → ((λ a _ → a) ∘ invIsEq (BEq (P tt))) + , λ a → equiv-proof (BEq (P tt)) a .fst .snd) + tt) + +isConnectedComp : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : B → C) (g : A → B) (n : ℕ) + → isConnectedFun n f + → isConnectedFun n g + → isConnectedFun n (f ∘ g) +isConnectedComp {C = C} f g n con-f con-g = + elim.isConnectedPrecompose (f ∘ g) n + λ P → + isEquiv→hasSection + (compEquiv + (_ , elim.isEquivPrecompose f n P con-f) + (_ , elim.isEquivPrecompose g n (λ b → P (f b)) con-g) .snd) + +isEquiv→isConnected : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + (f : A → B) + → isEquiv f + → (n : ℕ) → isConnectedFun n f +isEquiv→isConnected f iseq n b = + isContr→isContr∥ n (equiv-proof iseq b) + + +isConnectedPath : ∀ {ℓ} (n : HLevel) {A : Type ℓ} + → isConnected (suc n) A + → (a₀ a₁ : A) → isConnected n (a₀ ≡ a₁) +isConnectedPath zero connA a₀ a₁ = isContrUnit* +isConnectedPath (suc n) {A = A} connA a₀ a₁ = + isOfHLevelRetractFromIso 0 (invIso (PathIdTruncIso (suc n))) (isContr→isContrPath connA _ _) + +isConnectedPathP : ∀ {ℓ} (n : HLevel) {A : I → Type ℓ} + → isConnected (suc n) (A i1) + → (a₀ : A i0) (a₁ : A i1) → isConnected n (PathP A a₀ a₁) +isConnectedPathP n con a₀ a₁ = + subst (isConnected n) (sym (PathP≡Path _ _ _)) + (isConnectedPath n con _ _) + +isConnectedCong : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) + → isConnectedFun (suc n) f + → ∀ {a₀ a₁} → isConnectedFun n {A = a₀ ≡ a₁} (cong f) +isConnectedCong n f cf {a₀} {a₁} q = + subst (isConnected n) + (sym (fiberCong f q)) + (isConnectedPath n (cf (f a₁)) (a₀ , q) (a₁ , refl)) + +isConnectedCong' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x : A} {y : B} + (n : ℕ) (f : A → B) + → isConnectedFun (suc n) f + → (p : f x ≡ y) + → isConnectedFun n + λ (q : x ≡ x) → sym p ∙∙ cong f q ∙∙ p +isConnectedCong' {x = x} n f conf p = + transport (λ i → (isConnectedFun n + λ (q : x ≡ x) + → doubleCompPath-filler (sym p) (cong f q) p i)) + (isConnectedCong n f conf) + +isConnectedRetract : ∀ {ℓ ℓ'} (n : HLevel) + {A : Type ℓ} {B : Type ℓ'} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → isConnected n B → isConnected n A +isConnectedRetract zero _ _ _ _ = isContrUnit* +isConnectedRetract (suc n) f g h = + isContrRetract + (Trunc.map f) + (Trunc.map g) + (Trunc.elim + (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (λ a → cong ∣_∣ (h a))) + +isConnectedRetractFromIso : ∀ {ℓ ℓ'} (n : HLevel) + {A : Type ℓ} {B : Type ℓ'} + → Iso A B + → isConnected n B → isConnected n A +isConnectedRetractFromIso n e = + isConnectedRetract n + (Iso.fun e) + (Iso.inv e) + (Iso.leftInv e) + +isConnectedPoint : ∀ {ℓ} (n : HLevel) {A : Type ℓ} + → isConnected (suc n) A + → (a : A) → isConnectedFun n (λ(_ : Unit) → a) +isConnectedPoint n connA a₀ a = + isConnectedRetract n + snd (_ ,_) (λ _ → refl) + (isConnectedPath n connA a₀ a) + +isConnectedPoint2 : ∀ {ℓ} (n : HLevel) {A : Type ℓ} (a : A) + → isConnectedFun n (λ(_ : Unit) → a) + → isConnected (suc n) A +isConnectedPoint2 n {A = A} a connMap = indMapEquiv→conType _ λ B → isoToIsEquiv (theIso B) + where + module _ {ℓ' : Level} (B : TypeOfHLevel ℓ' (suc n)) + where + helper : (f : A → fst B) → (a2 : A) → f a2 ≡ f a + helper f = equiv-proof (elim.isEquivPrecompose (λ _ → a) n (λ a2 → (f a2 ≡ f a) , + isOfHLevelPath' n (snd B) (f a2) (f a)) connMap) (λ _ → refl) .fst .fst + + theIso : Iso (fst B) (A → fst B) + Iso.fun theIso b a = b + Iso.inv theIso f = f a + Iso.rightInv theIso f = funExt λ y → sym (helper f y) + Iso.leftInv theIso _ = refl + +connectedTruncIso : ∀ {ℓ} {A B : Type ℓ} (n : HLevel) (f : A → B) + → isConnectedFun n f + → Iso (hLevelTrunc n A) (hLevelTrunc n B) +connectedTruncIso {A = A} {B = B} zero f con = isContr→Iso isContrUnit* isContrUnit* +connectedTruncIso {A = A} {B = B} (suc n) f con = g + where + back : B → hLevelTrunc (suc n) A + back y = map fst ((con y) .fst) + + backSection : (b : B) → Path (hLevelTrunc (suc n) B) + (Trunc.rec (isOfHLevelTrunc (suc n)) + (λ a → ∣ f a ∣) + (Trunc.rec (isOfHLevelTrunc (suc n)) + back ∣ b ∣)) + ∣ b ∣ + backSection b = helper (λ p → map f p ≡ ∣ b ∣) + (λ x → (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n))) + (map f (map fst x)) ∣ b ∣) + (λ a p → cong ∣_∣ p) + (fst (con b)) + where + helper : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} (P : hLevelTrunc (suc n) A → Type ℓ'') + → ((x : hLevelTrunc (suc n) (Σ A B)) → isOfHLevel (suc n) (P (map fst x))) + → ((a : A) (b : B a) → P ∣ a ∣) + → (p : hLevelTrunc (suc n) (Σ A B)) + → P (map fst p) + helper P hlev pf = Trunc.elim hlev λ pair → pf (fst pair) (snd pair) + + g : Iso (hLevelTrunc (suc n) A) (hLevelTrunc (suc n) B) + Iso.fun g = map f + Iso.inv g = Trunc.rec (isOfHLevelTrunc _) back + Iso.leftInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + λ a → cong (map fst) (con (f a) .snd ∣ a , refl ∣) + Iso.rightInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + backSection + +connectedTruncIso2 : ∀ {ℓ} {A B : Type ℓ} (n m : HLevel) (f : A → B) + → Σ[ x ∈ ℕ ] x + n ≡ m + → isConnectedFun m f + → Iso (hLevelTrunc n A) (hLevelTrunc n B) +connectedTruncIso2 {A = A} {B = B} n m f (x , pf) con = + connectedTruncIso n f (isConnectedFunSubtr n x f (transport (λ i → isConnectedFun (pf (~ i)) f) con)) + +connectedTruncEquiv : ∀ {ℓ} {A B : Type ℓ} (n : HLevel) (f : A → B) + → isConnectedFun n f + → hLevelTrunc n A ≃ hLevelTrunc n B +connectedTruncEquiv {A = A} {B = B} n f con = isoToEquiv (connectedTruncIso n f con) + + +-- TODO : Reorganise the following proofs. + +inrConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : HLevel) + → (f : C → A) (g : C → B) + → isConnectedFun n f + → isConnectedFun n {A = B} {B = Pushout f g} inr +inrConnected {A = A} {B = B} {C = C} n f g iscon = + elim.isConnectedPrecompose inr n λ P → (k P) , λ b → refl + where + module _ {ℓ : Level} (P : (Pushout f g) → TypeOfHLevel ℓ n) + (h : (b : B) → typ (P (inr b))) + where + Q : A → TypeOfHLevel _ n + Q a = (P (inl a)) + + fun : (c : C) → typ (Q (f c)) + fun c = transport (λ i → typ (P (push c (~ i)))) (h (g c)) + + k : (d : Pushout f g) → typ (P d) + k (inl x) = equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .fst x + k (inr x) = h x + k (push a i) = + hcomp (λ k → λ { (i = i0) → equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .snd i0 a + ; (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k }) + (transp (λ j → typ (P (push a (i ∧ j)))) + (~ i) + (equiv-proof (elim.isEquivPrecompose f n Q iscon) + fun .fst .snd i a)) + +{- Given two fibration B , C : A → Type and a family of maps on fibres + f : (a : A) → B a → C a, we have that f a is n-connected for all (a : A) + iff the induced map on totalspaces Σ A B → Σ A C is n-connected -} +module _ {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} + (f : ((a : A) → B a → C a)) + where + open Iso + + TotalFun : Σ A B → Σ A C + TotalFun (a , b) = a , f a b + + fibTotalFun→fibFun : (x : Σ A C) + → Σ[ y ∈ Σ A B ] TotalFun y ≡ x + → Σ[ y ∈ B (fst x) ] f (fst x) y ≡ snd x + fibTotalFun→fibFun x = + uncurry + λ r → J (λ x _ → Σ[ y ∈ B (fst x) ] f (fst x) y ≡ snd x) + ((snd r) , refl) + + fibFun→fibTotalFun : (x : Σ A C) + → Σ[ y ∈ B (fst x) ] f (fst x) y ≡ snd x + → Σ[ y ∈ Σ A B ] TotalFun y ≡ x + fibFun→fibTotalFun x (b , p) = (_ , b) , ΣPathP (refl , p) + + Iso-fibTotalFun-fibFun : (x : Σ A C) + → Iso (Σ[ y ∈ Σ A B ] TotalFun y ≡ x) + (Σ[ y ∈ B (fst x) ] f (fst x) y ≡ snd x) + fun (Iso-fibTotalFun-fibFun x) = fibTotalFun→fibFun x + inv (Iso-fibTotalFun-fibFun x) = fibFun→fibTotalFun x + rightInv (Iso-fibTotalFun-fibFun x) (r , y) j = + transp (λ i → Σ[ b ∈ B (fst x) ] (f (fst x) b ≡ y (i ∨ j))) j + (r , λ i → y (i ∧ j)) + leftInv (Iso-fibTotalFun-fibFun x) = + uncurry λ r + → J (λ x y → inv (Iso-fibTotalFun-fibFun x) + (fun (Iso-fibTotalFun-fibFun x) (r , y)) + ≡ (r , y)) + (cong (fibFun→fibTotalFun (TotalFun r)) + (transportRefl (snd r , refl))) + + TotalFunConnected→FunConnected : (n : ℕ) + → isConnectedFun n TotalFun + → ((a : A) → isConnectedFun n (f a)) + TotalFunConnected→FunConnected n con a r = + isConnectedRetractFromIso n + (invIso (Iso-fibTotalFun-fibFun (a , r))) (con (a , r)) + + FunConnected→TotalFunConnected : (n : ℕ) + → ((a : A) → isConnectedFun n (f a)) + → isConnectedFun n TotalFun + FunConnected→TotalFunConnected n con r = + isConnectedRetractFromIso n + (Iso-fibTotalFun-fibFun r) (con (fst r) (snd r)) diff --git a/Cubical/Homotopy/EilenbergSteenrod.agda b/Cubical/Homotopy/EilenbergSteenrod.agda new file mode 100644 index 0000000000..cabb12e728 --- /dev/null +++ b/Cubical/Homotopy/EilenbergSteenrod.agda @@ -0,0 +1,47 @@ +{-# OPTIONS --safe #-} + +module Cubical.Homotopy.EilenbergSteenrod where + +{- +This module contains the Eilenberg-Steenrod axioms for ordinary +reduced cohomology theories with binary additivity. +The axioms are based on the ones given in Cavallo's MSc thesis +(https://www.cs.cmu.edu/~ecavallo/works/thesis15.pdf) and +Buchholtz/Favonia (2018) +-} + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.HITs.Wedge +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp + +open import Cubical.Data.Empty +open import Cubical.Relation.Nullary + +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Sigma +open import Cubical.Data.Int + +open import Cubical.Algebra.Group hiding (ℤ ; Bool) +open import Cubical.Algebra.AbGroup + +record coHomTheory {ℓ ℓ' : Level} (H : (n : ℤ) → Pointed ℓ → AbGroup ℓ') : Type (ℓ-suc (ℓ-max ℓ ℓ')) + where + Boolℓ : Pointed ℓ + Boolℓ = Lift Bool , lift true + field + Hmap : (n : ℤ) → {A B : Pointed ℓ} (f : A →∙ B) → AbGroupHom (H n B) (H n A) + Suspension : Σ[ F ∈ ((n : ℤ) {A : Pointed ℓ} → AbGroupEquiv (H (sucℤ n) (Susp (typ A) , north)) (H n A)) ] + ({A B : Pointed ℓ} (f : A →∙ B) (n : ℤ) + → fst (Hmap (sucℤ n) (suspFun (fst f) , refl)) ∘ invEq (fst (F n {A = B})) + ≡ invEq (fst (F n {A = A})) ∘ fst (Hmap n f)) + Exactness : {A B : Pointed ℓ} (f : A →∙ B) (n : ℤ) + → Ker (Hmap n f) + ≡ Im (Hmap n {B = _ , inr (pt B)} (cfcod (fst f) , refl)) + Dimension : (n : ℤ) → ¬ n ≡ 0 → isContr (fst (H n Boolℓ)) + BinaryWedge : (n : ℤ) {A B : Pointed ℓ} → AbGroupEquiv (H n (A ⋁ B , (inl (pt A)))) (dirProdAb (H n A) (H n B)) diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda new file mode 100644 index 0000000000..ba05ca7250 --- /dev/null +++ b/Cubical/Homotopy/Freudenthal.agda @@ -0,0 +1,135 @@ +{- + +Freudenthal suspension theorem + +-} +{-# OPTIONS --safe #-} +module Cubical.Homotopy.Freudenthal where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Sigma +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.HITs.Nullification +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec ; elim to trElim) +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.WedgeConnectivity +open import Cubical.Homotopy.Loopspace +open import Cubical.HITs.SmashProduct + +open import Cubical.HITs.S1 hiding (encode) +open import Cubical.HITs.Sn +open import Cubical.HITs.S2 +open import Cubical.HITs.S3 + +module _ {ℓ} (n : HLevel) {A : Pointed ℓ} (connA : isConnected (suc (suc n)) (typ A)) where + + private + 2n+2 = suc n + suc n + + module WC (p : north ≡ north) = + WedgeConnectivity (suc n) (suc n) A connA A connA + (λ a b → + ( (σ A b ≡ p → hLevelTrunc 2n+2 (fiber (λ x → merid x ∙ merid a ⁻¹) p)) + , isOfHLevelΠ 2n+2 λ _ → isOfHLevelTrunc 2n+2 + )) + (λ a r → ∣ a , (rCancel' (merid a) ∙ rCancel' (merid (pt A)) ⁻¹) ∙ r ∣) + (λ b r → ∣ b , r ∣) + (funExt λ r → + cong′ (λ w → ∣ pt A , w ∣) + (cong (_∙ r) (rCancel' (rCancel' (merid (pt A)))) + ∙ lUnit r ⁻¹)) + + fwd : (p : north ≡ north) (a : typ A) + → hLevelTrunc 2n+2 (fiber (σ A) p) + → hLevelTrunc 2n+2 (fiber (λ x → merid x ∙ merid a ⁻¹) p) + fwd p a = Trunc.rec (isOfHLevelTrunc 2n+2) (uncurry (WC.extension p a)) + + isEquivFwd : (p : north ≡ north) (a : typ A) → isEquiv (fwd p a) + isEquivFwd p a .equiv-proof = + elim.isEquivPrecompose (λ _ → pt A) (suc n) + (λ a → + ( (∀ t → isContr (fiber (fwd p a) t)) + , isProp→isOfHLevelSuc n (isPropΠ λ _ → isPropIsContr) + )) + + (isConnectedPoint (suc n) connA (pt A)) + .equiv-proof + (λ _ → Trunc.elim + (λ _ → isProp→isOfHLevelSuc (n + suc n) isPropIsContr) + (λ fib → + subst (λ k → isContr (fiber k ∣ fib ∣)) + (cong (Trunc.rec (isOfHLevelTrunc 2n+2) ∘ uncurry) + (funExt (WC.right p) ⁻¹)) + (subst isEquiv + (funExt (Trunc.mapId) ⁻¹) + (idIsEquiv _) + .equiv-proof ∣ fib ∣) + )) + .fst .fst a + + interpolate : (a : typ A) + → PathP (λ i → typ A → north ≡ merid a i) (λ x → merid x ∙ merid a ⁻¹) merid + interpolate a i x j = compPath-filler (merid x) (merid a ⁻¹) (~ i) j + + Code : (y : Susp (typ A)) → north ≡ y → Type ℓ + Code north p = hLevelTrunc 2n+2 (fiber (σ A) p) + Code south q = hLevelTrunc 2n+2 (fiber merid q) + Code (merid a i) p = + Glue + (hLevelTrunc 2n+2 (fiber (interpolate a i) p)) + (λ + { (i = i0) → _ , (fwd p a , isEquivFwd p a) + ; (i = i1) → _ , idEquiv _ + }) + + encode : (y : Susp (typ A)) (p : north ≡ y) → Code y p + encode y = J Code ∣ pt A , rCancel' (merid (pt A)) ∣ + + encodeMerid : (a : typ A) → encode south (merid a) ≡ ∣ a , refl ∣ + encodeMerid a = + cong (transport (λ i → gluePath i)) + (funExt⁻ (WC.left refl a) _ ∙ λ i → ∣ a , lem (rCancel' (merid a)) (rCancel' (merid (pt A))) i ∣) + ∙ transport (PathP≡Path gluePath _ _) + (λ i → ∣ a , (λ j k → rCancel-filler' (merid a) i j k) ∣) + where + gluePath : I → Type _ + gluePath i = hLevelTrunc 2n+2 (fiber (interpolate a i) (λ j → merid a (i ∧ j))) + + lem : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : z ≡ y) → (p ∙ q ⁻¹) ∙ q ≡ p + lem p q = assoc p (q ⁻¹) q ⁻¹ ∙∙ cong (p ∙_) (lCancel q) ∙∙ rUnit p ⁻¹ + + contractCodeSouth : (p : north ≡ south) (c : Code south p) → encode south p ≡ c + contractCodeSouth p = + Trunc.elim + (λ _ → isOfHLevelPath 2n+2 (isOfHLevelTrunc 2n+2) _ _) + (uncurry λ a → + J (λ p r → encode south p ≡ ∣ a , r ∣) (encodeMerid a)) + + isConnectedMerid : isConnectedFun 2n+2 (merid {A = typ A}) + isConnectedMerid p = encode south p , contractCodeSouth p + + isConnectedσ : isConnectedFun 2n+2 (σ A) + isConnectedσ = + transport (λ i → isConnectedFun 2n+2 (interpolate (pt A) (~ i))) isConnectedMerid + +isConn→isConnSusp : ∀ {ℓ} {A : Pointed ℓ} → isConnected 2 (typ A) → isConnected 2 (Susp (typ A)) +isConn→isConnSusp {A = A} iscon = ∣ north ∣ + , trElim (λ _ → isOfHLevelSuc 1 (isOfHLevelTrunc 2 _ _)) + (suspToPropElim (pt A) (λ _ → isOfHLevelTrunc 2 _ _) + refl) + +FreudenthalEquiv : ∀ {ℓ} (n : HLevel) (A : Pointed ℓ) + → isConnected (2 + n) (typ A) + → hLevelTrunc ((suc n) + (suc n)) (typ A) + ≃ hLevelTrunc ((suc n) + (suc n)) (typ (Ω (Susp (typ A) , north))) +FreudenthalEquiv n A iscon = connectedTruncEquiv _ + (σ A) + (isConnectedσ _ iscon) +FreudenthalIso : ∀ {ℓ} (n : HLevel) (A : Pointed ℓ) + → isConnected (2 + n) (typ A) + → Iso (hLevelTrunc ((suc n) + (suc n)) (typ A)) + (hLevelTrunc ((suc n) + (suc n)) (typ (Ω (Susp (typ A) , north)))) +FreudenthalIso n A iscon = connectedTruncIso _ (σ A) (isConnectedσ _ iscon) diff --git a/Cubical/Homotopy/Group/Base.agda b/Cubical/Homotopy/Group/Base.agda new file mode 100644 index 0000000000..6cd9148277 --- /dev/null +++ b/Cubical/Homotopy/Group/Base.agda @@ -0,0 +1,1046 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Group.Base where + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function +open import Cubical.Foundations.Transport + +open import Cubical.Functions.Morphism + +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 + ; map to sMap) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.S1 + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Unit + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + +open Iso +open IsGroup +open IsSemigroup +open IsMonoid +open GroupStr + +{- Homotopy group -} +π : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ +π n A = ∥ typ ((Ω^ n) A) ∥₂ + +{- Alternative formulation. This will be given a group structure in + the Properties file -} +π' : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ +π' n A = ∥ S₊∙ n →∙ A ∥₂ + +{- π as a group -} +1π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π n A +1π zero {A = A} = ∣ pt A ∣₂ +1π (suc n) = ∣ refl ∣₂ + +·π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π (suc n) A → π (suc n) A → π (suc n) A +·π n = sRec2 squash₂ λ p q → ∣ p ∙ q ∣₂ + +-π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π (suc n) A → π (suc n) A +-π n = sMap sym + +π-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n x (1π (suc n))) ≡ x +π-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ rUnit p (~ i) ∣₂ + +π-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n (1π (suc n)) x) ≡ x +π-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ lUnit p (~ i) ∣₂ + +π-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n x (-π n x)) ≡ 1π (suc n) +π-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ rCancel p i ∣₂ + +π-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n (-π n x) x) ≡ 1π (suc n) +π-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ lCancel p i ∣₂ + +π-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π (suc n) A) + → ·π n x (·π n y z) ≡ ·π n (·π n x y) z +π-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙assoc p q r i ∣₂ + +π-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π (suc (suc n)) A) + → ·π (suc n) x y ≡ ·π (suc n) y x +π-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ EH n p q i ∣₂ + +-- πₙ₊₁ +πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ +fst (πGr n A) = π (suc n) A +1g (snd (πGr n A)) = 1π (suc n) +GroupStr._·_ (snd (πGr n A)) = ·π n +inv (snd (πGr n A)) = -π n +is-set (isSemigroup (isMonoid (isGroup (snd (πGr n A))))) = squash₂ +assoc (isSemigroup (isMonoid (isGroup (snd (πGr n A))))) = π-assoc n +identity (isMonoid (isGroup (snd (πGr n A)))) x = (π-rUnit n x) , (π-lUnit n x) +inverse (isGroup (snd (πGr n A))) x = (π-rCancel n x) , (π-lCancel n x) + +-- Group operations on π'. +-- We define the corresponding structure on the untruncated +-- (S₊∙ n →∙ A). + +∙Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) +∙Π {A = A} {n = zero} p q = (λ _ → pt A) , refl +fst (∙Π {A = A} {n = suc zero} (f , p) (g , q)) base = pt A +fst (∙Π {A = A} {n = suc zero} (f , p) (g , q)) (loop j) = + ((sym p ∙∙ cong f loop ∙∙ p) ∙ (sym q ∙∙ cong g loop ∙∙ q)) j +snd (∙Π {A = A} {n = suc zero} (f , p) (g , q)) = refl +fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) north = pt A +fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) south = pt A +fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) (merid a j) = + ((sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ p) + ∙ (sym q ∙∙ cong g (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ q)) j +snd (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) = refl + +-Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) +-Π {n = zero} f = f +fst (-Π {A = A} {n = suc zero} f) base = fst f base +fst (-Π {A = A} {n = suc zero} f) (loop j) = fst f (loop (~ j)) +snd (-Π {A = A} {n = suc zero} f) = snd f +fst (-Π {A = A} {n = suc (suc n)} f) north = fst f north +fst (-Π {A = A} {n = suc (suc n)} f) south = fst f north +fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = + fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) +snd (-Π {A = A} {n = suc (suc n)} f) = snd f + + +-- to prove that this gives a group structure on π', we first +-- prove that Ωⁿ A ≃ (Sⁿ →∙ A). +-- We use the following map +mutual + Ω→SphereMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → typ ((Ω^ n) A) → (S₊∙ n →∙ A) + fst (Ω→SphereMap zero a) false = a + fst (Ω→SphereMap zero {A = A} a) true = pt A + snd (Ω→SphereMap zero a) = refl + fst (Ω→SphereMap (suc zero) {A = A} p) base = pt A + fst (Ω→SphereMap (suc zero) p) (loop i) = p i + snd (Ω→SphereMap (suc zero) p) = refl + fst (Ω→SphereMap (suc (suc n)) {A = A} p) north = pt A + fst (Ω→SphereMap (suc (suc n)) {A = A} p) south = pt A + fst (Ω→SphereMap (suc (suc n)) p) (merid a i) = + (sym (Ω→SphereMapId (suc n) a) + ∙∙ (λ i → Ω→SphereMap (suc n) (p i) .fst a) + ∙∙ Ω→SphereMapId (suc n) a) i + snd (Ω→SphereMap (suc (suc n)) p) = refl + + Ω→SphereMapId : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (a : _) + → Ω→SphereMap n {A = A} (pt ((Ω^ n) A)) .fst a ≡ pt A + Ω→SphereMapId zero false = refl + Ω→SphereMapId zero true = refl + Ω→SphereMapId (suc zero) base = refl + Ω→SphereMapId (suc zero) (loop i) = refl + Ω→SphereMapId (suc (suc n)) north = refl + Ω→SphereMapId (suc (suc n)) south = refl + Ω→SphereMapId (suc (suc n)) {A = A} (merid a i) j = + ∙∙lCancel (Ω→SphereMapId (suc n) {A = A} a) j i + +Ω→SphereMapId2 : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → Ω→SphereMap n {A = A} (pt ((Ω^ n) A)) ≡ ((λ _ → pt A) , refl) +fst (Ω→SphereMapId2 n {A = A} i) a = funExt (Ω→SphereMapId n {A = A}) i a +snd (Ω→SphereMapId2 zero {A = A} i) = refl +snd (Ω→SphereMapId2 (suc zero) {A = A} i) = refl +snd (Ω→SphereMapId2 (suc (suc n)) {A = A} i) = refl + +-- Pointed version +Ω→SphereMap∙ : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → ((Ω^ n) A) →∙ (S₊∙ n →∙ A ∙) +Ω→SphereMap∙ n .fst = Ω→SphereMap n +Ω→SphereMap∙ n .snd = Ω→SphereMapId2 n + +-- We define the following maps which will be used to +-- show that Ω→SphereMap is an equivalence +Ω→SphereMapSplit₁ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → typ ((Ω^ (suc n)) A) + → typ (Ω (S₊∙ n →∙ A ∙)) +Ω→SphereMapSplit₁ n = Ω→ (Ω→SphereMap∙ n) .fst + +ΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → typ (Ω (S₊∙ n →∙ A ∙)) + → (S₊∙ (suc n) →∙ A) +fst (ΩSphereMap {A = A} zero p) base = p i0 .fst false +fst (ΩSphereMap {A = A} zero p) (loop i) = p i .fst false +snd (ΩSphereMap {A = A} zero p) = refl +ΩSphereMap {A = A} (suc n) = fun IsoΩFunSuspFun + +-- Functoriality +-- The homogeneity assumption is not necessary but simplifying +isNaturalΩSphereMap : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + (homogB : isHomogeneous B) (f : A →∙ B) (n : ℕ) + → ∀ g → f ∘∙ ΩSphereMap n g ≡ ΩSphereMap n (Ω→ (post∘∙ (S₊∙ n) f) .fst g) +isNaturalΩSphereMap A B homogB f 0 g = + →∙Homogeneous≡ homogB (funExt lem) + where + lem : ∀ x → f .fst (ΩSphereMap 0 g .fst x) + ≡ ΩSphereMap 0 (Ω→ (post∘∙ (S₊∙ 0) f) .fst g) .fst x + lem base = f .snd + lem (loop i) j = + hfill + (λ j → λ + { (i = i0) → post∘∙ _ f .snd j + ; (i = i1) → post∘∙ _ f .snd j + }) + (inS (f ∘∙ g i)) + j .fst false +isNaturalΩSphereMap A B homogB f (n@(suc _)) g = + →∙Homogeneous≡ homogB (funExt lem) + where + lem : ∀ x → f .fst (ΩSphereMap n g .fst x) + ≡ ΩSphereMap n (Ω→ (post∘∙ (S₊∙ n) f) .fst g) .fst x + lem north = f .snd + lem south = f .snd + lem (merid a i) j = + hfill + (λ j → λ + { (i = i0) → post∘∙ _ f .snd j + ; (i = i1) → post∘∙ _ f .snd j + }) + (inS (f ∘∙ g i)) + j .fst a + +SphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ (suc n) →∙ A) + → typ (Ω (S₊∙ n →∙ A ∙)) +SphereMapΩ {A = A} zero (f , p) = + ΣPathP ((funExt λ { false → sym p ∙∙ cong f loop ∙∙ p + ; true → refl}) + , refl) +SphereMapΩ {A = A} (suc n) = inv IsoΩFunSuspFun + +SphereMapΩIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ (suc n) →∙ A) + (typ (Ω (S₊∙ n →∙ A ∙))) +fun (SphereMapΩIso n) = SphereMapΩ n +inv (SphereMapΩIso n) = ΩSphereMap n +fst (rightInv (SphereMapΩIso zero) f i j) false = rUnit (λ j → fst (f j) false) (~ i) j +fst (rightInv (SphereMapΩIso {A = A} zero) f i j) true = snd (f j) (~ i) +snd (rightInv (SphereMapΩIso {A = A} zero) f i j) k = snd (f j) (~ i ∨ k) +rightInv (SphereMapΩIso (suc n)) = leftInv IsoΩFunSuspFun +leftInv (SphereMapΩIso zero) f = + ΣPathP ((funExt (λ { base → sym (snd f) + ; (loop i) j → doubleCompPath-filler + (sym (snd f)) + (cong (fst f) loop) + (snd f) (~ j) i})) + , λ i j → snd f (~ i ∨ j)) +leftInv (SphereMapΩIso (suc n)) = rightInv IsoΩFunSuspFun + +{- +In order to show that Ω→SphereMap is an equivalence, we show that it factors + + Ω→SphereMapSplit₁ ΩSphereMap +Ωⁿ⁺¹A ----------------> Ω (Sⁿ →∙ A) -----------> (Sⁿ⁺¹ →∙ A) +-} + +Ω→SphereMap-split : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : typ ((Ω^ (suc n)) A)) + → Ω→SphereMap (suc n) p ≡ ΩSphereMap n (Ω→SphereMapSplit₁ n p) +Ω→SphereMap-split {A = A} zero p = + ΣPathP ((funExt (λ { base → refl + ; (loop i) j → lem (~ j) i})) + , refl) + where + lem : funExt⁻ (cong fst (Ω→SphereMapSplit₁ zero p)) false ≡ p + lem = (λ i → funExt⁻ (cong-∙∙ fst (sym (Ω→SphereMapId2 zero)) + (cong (Ω→SphereMap zero) p) + (Ω→SphereMapId2 zero) i) false) + ∙ sym (rUnit _) +Ω→SphereMap-split {A = A} (suc n) p = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → lem₂ a j i})) + , refl) + where + lem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (a : S₊ (suc n)) + → Ω→SphereMapId (suc n) {A = A} a + ≡ (λ i → fst (Ω→SphereMapId2 (suc n) {A = A} i) a) + lem zero base = refl + lem zero (loop i) = refl + lem (suc n) north = refl + lem (suc n) south = refl + lem (suc n) (merid a i) = refl + + lem₂ : (a : S₊ (suc n)) + → ((λ i₁ → Ω→SphereMapId (suc n) {A = A} a (~ i₁)) + ∙∙ (λ i₁ → Ω→SphereMap (suc n) (p i₁) .fst a) + ∙∙ Ω→SphereMapId (suc n) a) + ≡ (λ i → Ω→SphereMapSplit₁ (suc n) p i .fst a) + lem₂ a = cong (λ x → sym x + ∙∙ funExt⁻ (cong fst (λ i → Ω→SphereMap (suc n) (p i))) a + ∙∙ x) + (lem n a) + ∙∙ sym (cong-∙∙ (λ x → x a) + (cong fst (λ i → Ω→SphereMapId2 (suc n) (~ i))) + (cong fst (λ i → Ω→SphereMap (suc n) (p i))) + (cong fst (Ω→SphereMapId2 (suc n)))) + ∙∙ (λ i → funExt⁻ (cong-∙∙ fst (sym (Ω→SphereMapId2 (suc n))) + (cong (Ω→SphereMap (suc n)) p) + (Ω→SphereMapId2 (suc n)) (~ i)) a) + +isEquiv-Ω→SphereMap₀ : ∀ {ℓ} {A : Pointed ℓ} + → isEquiv (Ω→SphereMap 0 {A = A}) +isEquiv-Ω→SphereMap₀ {A = A} = + isoToIsEquiv + (iso _ (λ f → fst f false) + (λ f → ΣPathP ((funExt (λ { false → refl ; true → sym (snd f)})) + , λ i j → snd f (~ i ∨ j))) + λ p → refl) + +isEquiv-Ω→SphereMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → isEquiv (Ω→SphereMap n {A = A}) +isEquiv-Ω→SphereMap zero {A = A} = + (isoToIsEquiv + (iso _ (λ f → fst f false) + (λ f → ΣPathP ((funExt (λ { false → refl + ; true → sym (snd f)})) + , λ i j → snd f (~ i ∨ j))) + λ _ → refl)) +isEquiv-Ω→SphereMap (suc zero) {A = A} = + isoToIsEquiv (iso _ invFun sec λ p → sym (rUnit p)) + where + invFun : S₊∙ 1 →∙ A → typ (Ω A) + invFun (f , p) = sym p ∙∙ cong f loop ∙∙ p + + sec : section (Ω→SphereMap 1) invFun + sec (f , p) = + ΣPathP ((funExt (λ { base → sym p + ; (loop i) j → doubleCompPath-filler + (sym p) (cong f loop) p (~ j) i})) + , λ i j → p (~ i ∨ j)) + +isEquiv-Ω→SphereMap (suc (suc n)) = + subst isEquiv (sym (funExt (Ω→SphereMap-split (suc n)))) + (snd (compEquiv + ((Ω→SphereMapSplit₁ (suc n)) , + (isEquivΩ→ (Ω→SphereMap (suc n) , Ω→SphereMapId2 (suc n)) + (isEquiv-Ω→SphereMap (suc n)))) + (invEquiv (isoToEquiv (SphereMapΩIso (suc n)))))) + +IsoΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (typ ((Ω^ n) A)) (S₊∙ n →∙ A) +IsoΩSphereMap n = equivToIso (_ , isEquiv-Ω→SphereMap n) + +IsoSphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) +IsoSphereMapΩ {A = A} n = + invIso (IsoΩSphereMap n) + +SphereMap→Ω : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ n →∙ A → fst ((Ω^ n) A) +SphereMap→Ω n = fun (IsoSphereMapΩ n) + +isHom-Ω→SphereMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (p q : _) + → Ω→SphereMap (suc n) {A = A} (p ∙ q) + ≡ ∙Π (Ω→SphereMap (suc n) {A = A} p) + (Ω→SphereMap (suc n) {A = A} q) +isHom-Ω→SphereMap zero {A = A} p q = + ΣPathP ((funExt (λ { base → refl + ; (loop i) j → (rUnit p j ∙ rUnit q j) i})) + , refl) +isHom-Ω→SphereMap (suc n) {A = A} p q = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → main a j i})) + , refl) + where + doubleComp-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q r : y ≡ y) + → (p ∙∙ q ∙∙ sym p) ∙ (p ∙∙ r ∙∙ sym p) + ≡ (p ∙∙ (q ∙ r) ∙∙ sym p) + doubleComp-lem p q r i j = + hcomp (λ k → λ { (i = i0) → (doubleCompPath-filler p q (sym p) k + ∙ doubleCompPath-filler p r (sym p) k) j + ; (i = i1) → doubleCompPath-filler p (q ∙ r) (sym p) k j + ; (j = i0) → p (~ k) + ; (j = i1) → p (~ k)}) + ((q ∙ r) j) + + lem : (p : typ ((Ω^ (suc (suc n))) A)) + → cong (fst (Ω→SphereMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl + lem p = + cong (sym (Ω→SphereMapId (suc n) (ptSn _)) ∙∙_∙∙ Ω→SphereMapId (suc n) (ptSn _)) + (rUnit _ ∙ (λ j → (λ i → Ω→SphereMap (suc n) {A = A} refl .snd (i ∧ j)) + ∙∙ (λ i → Ω→SphereMap (suc n) {A = A} (p i) .snd j) + ∙∙ λ i → Ω→SphereMap (suc n) {A = A} refl .snd (~ i ∧ j)) + ∙ ∙∙lCancel _) + ∙ ∙∙lCancel _ + + main : (a : S₊ (suc n)) + → sym (Ω→SphereMapId (suc n) a) + ∙∙ funExt⁻ (cong fst (cong (Ω→SphereMap (suc n)) (p ∙ q))) a + ∙∙ Ω→SphereMapId (suc n) a + ≡ cong (fst (∙Π (Ω→SphereMap (suc (suc n)) p) (Ω→SphereMap (suc (suc n)) q))) (merid a) + main a = (cong (sym (Ω→SphereMapId (suc n) a) ∙∙_∙∙ (Ω→SphereMapId (suc n) a)) + (cong-∙ (λ x → Ω→SphereMap (suc n) x .fst a) p q) + ∙ sym (doubleComp-lem (sym (Ω→SphereMapId (suc n) a)) _ _)) + ∙∙ cong₂ _∙_ (sym (cong (cong (fst (Ω→SphereMap (suc (suc n)) p)) (merid a) ∙_) + (cong sym (lem p)) ∙ sym (rUnit _))) + (sym (cong (cong (fst (Ω→SphereMap (suc (suc n)) q)) (merid a) ∙_) + (cong sym (lem q)) ∙ sym (rUnit _))) + ∙∙ λ i → (rUnit (cong-∙ (fst (Ω→SphereMap (suc (suc n)) p)) + (merid a) (sym (merid (ptSn _))) (~ i)) i) + ∙ (rUnit (cong-∙ (fst (Ω→SphereMap (suc (suc n)) q)) + (merid a) (sym (merid (ptSn _)))(~ i)) i) + + +-- The iso is structure preserving +IsoSphereMapΩ-pres∙Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → SphereMap→Ω (suc n) (∙Π f g) + ≡ SphereMap→Ω (suc n) f ∙ SphereMap→Ω (suc n) g +IsoSphereMapΩ-pres∙Π n = + morphLemmas.isMorphInv _∙_ ∙Π (Ω→SphereMap (suc n)) + (isHom-Ω→SphereMap n) + (SphereMap→Ω (suc n)) + (leftInv (IsoSphereMapΩ (suc n))) + (rightInv (IsoSphereMapΩ (suc n))) + +-- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) +-- before doing it on π'. These will be the equivalents of the +-- usual groupoid laws on Ω A. +1Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) +fst (1Π {A = A}) _ = pt A +snd (1Π {A = A}) = refl + +∙Π-rUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f 1Π ≡ f +fst (∙Π-rUnit {A = A} {n = zero} f i) base = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = zero} f i) (loop j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) loop) + help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) + (sym (rUnit refl)) ∙ sym (rUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-rUnit {A = A} {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-rUnit {A = A} {n = suc n} f i) north = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i +fst (∙Π-rUnit {A = A} {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) (merid a)) + help = (cong (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) ∙_) + (sym (rUnit refl)) + ∙ sym (rUnit _)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) + (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) + (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-rUnit {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-lUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π 1Π f ≡ f +fst (∙Π-lUnit {n = zero} f i) base = snd f (~ i) +fst (∙Π-lUnit {n = zero} f i) (loop j) = s i j + where + s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (cong (fst f) loop) + s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (sym (rUnit refl)) ∙ sym (lUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-lUnit {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-lUnit {n = suc n} f i) north = snd f (~ i) +fst (∙Π-lUnit {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i +fst (∙Π-lUnit {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + ((refl ∙ refl) ∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) + (cong (fst f) (merid a)) + help = + (cong (_∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) + (sym (rUnit refl)) + ∙ sym (lUnit _)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-lUnit {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-rCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f (-Π f) ≡ 1Π +fst (∙Π-rCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j +snd (∙Π-rCancel {A = A} {n = zero} f i) = refl +fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = lem i i₁ + where + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) + + lem : pl + ∙ ((sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f)) ≡ refl + lem = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl + (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ rCancel pl +snd (∙Π-rCancel {A = A} {n = suc n} f i) = refl + +∙Π-lCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π (-Π f) f ≡ 1Π +fst (∙Π-lCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-lCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j +fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = lem i j + where + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) + + lem : (sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) ∙ pl + ≡ refl + lem = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ lCancel pl +snd (∙Π-lCancel {A = A} {n = zero} f i) = refl +snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl + +∙Π-assoc : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g h : S₊∙ (suc n) →∙ A) + → ∙Π f (∙Π g h) ≡ ∙Π (∙Π f g) h +∙Π-assoc {n = n} f g h = + sym (leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) + ∙∙ cong (Ω→SphereMap (suc n)) (IsoSphereMapΩ-pres∙Π n f (∙Π g h) + ∙∙ cong (SphereMap→Ω (suc n) f ∙_) (IsoSphereMapΩ-pres∙Π n g h) + ∙∙ ∙assoc (SphereMap→Ω (suc n) f) (SphereMap→Ω (suc n) g) (SphereMap→Ω (suc n) h) + ∙∙ cong (_∙ SphereMap→Ω (suc n) h) (sym (IsoSphereMapΩ-pres∙Π n f g)) + ∙∙ sym (IsoSphereMapΩ-pres∙Π n (∙Π f g) h)) + ∙∙ leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) + +∙Π-comm : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g : S₊∙ (suc (suc n)) →∙ A) + → ∙Π f g ≡ ∙Π g f +∙Π-comm {A = A} {n = n} f g = + sym (leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) + ∙∙ cong (Ω→SphereMap (suc (suc n))) (IsoSphereMapΩ-pres∙Π (suc n) f g + ∙∙ EH _ _ _ + ∙∙ sym (IsoSphereMapΩ-pres∙Π (suc n) g f)) + ∙∙ leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) + +{- π'' as a group -} +1π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' n A +1π' n {A = A} = ∣ 1Π ∣₂ + +·π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A → π' (suc n) A +·π' n = sRec2 squash₂ λ p q → ∣ ∙Π p q ∣₂ + +-π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A +-π' n = sMap -Π + +π'-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (1π' (suc n))) ≡ x +π'-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rUnit p i ∣₂ + +π'-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (1π' (suc n)) x) ≡ x +π'-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lUnit p i ∣₂ + +π'-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (-π' n x)) ≡ 1π' (suc n) +π'-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rCancel p i ∣₂ + +π'-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (-π' n x) x) ≡ 1π' (suc n) +π'-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lCancel p i ∣₂ + +π'-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π' (suc n) A) + → ·π' n x (·π' n y z) ≡ ·π' n (·π' n x y) z +π'-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙Π-assoc p q r i ∣₂ + +π'-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π' (suc (suc n)) A) + → ·π' (suc n) x y ≡ ·π' (suc n) y x +π'-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ ∙Π-comm p q i ∣₂ + +-- We finally get the group definition +π'Gr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ +fst (π'Gr n A) = π' (suc n) A +1g (snd (π'Gr n A)) = 1π' (suc n) +GroupStr._·_ (snd (π'Gr n A)) = ·π' n +inv (snd (π'Gr n A)) = -π' n +is-set (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = squash₂ +assoc (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = π'-assoc n +identity (isMonoid (isGroup (snd (π'Gr n A)))) x = (π'-rUnit n x) , (π'-lUnit n x) +inverse (isGroup (snd (π'Gr n A))) x = (π'-rCancel n x) , (π'-lCancel n x) + +-- and finally, the Iso +π'Gr≅πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → GroupIso (π'Gr n A) (πGr n A) +fst (π'Gr≅πGr n A) = setTruncIso (IsoSphereMapΩ (suc n)) +snd (π'Gr≅πGr n A) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) + +{- Proof of πₙ(ΩA) = πₙ₊₁(A) -} +Iso-πΩ-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (π n (Ω A)) (π (suc n) A) +Iso-πΩ-π {A = A} n = setTruncIso (invIso (flipΩIso n)) + +GrIso-πΩ-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupIso (πGr n (Ω A)) (πGr (suc n) A) +fst (GrIso-πΩ-π n) = Iso-πΩ-π _ +snd (GrIso-πΩ-π n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (flipΩIso⁻pres· n p q)) + + +{- Proof that πₙ(A) ≅ πₙ(∥ A ∥ₙ) -} +isContrΩTrunc : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → isContr (typ ((Ω^ n) (hLevelTrunc∙ n A))) +isContrΩTrunc {A = A} zero = isContrUnit* +isContrΩTrunc {A = A} (suc n) = + subst isContr main (isContrΩTrunc {A = Ω A} n) + where + lem₁ : (n : ℕ) → fun (PathIdTruncIso n) (λ _ → ∣ pt A ∣) + ≡ snd (hLevelTrunc∙ n (Ω A)) + lem₁ zero = refl + lem₁ (suc n) = transportRefl ∣ refl ∣ + + lem₂ : hLevelTrunc∙ n (Ω A) ≡ (Ω (hLevelTrunc∙ (suc n) A)) + lem₂ = sym (ua∙ (isoToEquiv (PathIdTruncIso n)) + (lem₁ n)) + + main : (typ ((Ω^ n) (hLevelTrunc∙ n (Ω A)))) + ≡ (typ ((Ω^ suc n) (hLevelTrunc∙ (suc n) A))) + main = (λ i → typ ((Ω^ n) (lem₂ i))) + ∙ sym (isoToPath (flipΩIso n)) + + +mutual + ΩTruncSwitchFun : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) → + (hLevelTrunc∙ (suc (suc m)) ((Ω^ n) A)) + →∙ ((Ω^ n) (hLevelTrunc∙ (suc n + suc m) A)) + ΩTruncSwitchFun {A = A} n m = + ((λ x → transport + (λ i → fst ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) i) A))) + (Iso.fun (ΩTruncSwitch {A = A} n (suc (suc m))) x)) + , cong (transport + (λ i → fst ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) i) A)))) + (ΩTruncSwitch∙ n (suc (suc m))) + ∙ λ j → transp + (λ i → fst ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) (i ∨ j)) A))) + j (snd ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) j) A)))) + + ΩTruncSwitchLem : + ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → Iso + (typ (Ω (hLevelTrunc∙ (suc (suc m)) ((Ω^ n) A)))) + (typ ((Ω^ suc n) (hLevelTrunc∙ (suc n + suc m) A))) + ΩTruncSwitchLem {A = A} n m = + (equivToIso + (Ω→ (ΩTruncSwitchFun n m) .fst + , isEquivΩ→ _ (compEquiv (isoToEquiv (ΩTruncSwitch {A = A} n (suc (suc m)))) + (transportEquiv + (λ i → typ ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) i) A)))) .snd))) + + ΩTruncSwitch : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → Iso (hLevelTrunc m (fst ((Ω^ n) A))) + (typ ((Ω^ n) (hLevelTrunc∙ (n + m) A))) + ΩTruncSwitch {A = A} n zero = + equivToIso + (invEquiv + (isContr→≃Unit* + (subst isContr + (λ i → (typ ((Ω^ n) (hLevelTrunc∙ (+-comm zero n i) A)))) + (isContrΩTrunc n)))) + ΩTruncSwitch {A = A} zero (suc m) = idIso + ΩTruncSwitch {A = A} (suc n) (suc m) = + compIso (invIso (PathIdTruncIso _)) + (ΩTruncSwitchLem n m) + + ΩTruncSwitch∙ : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → Iso.fun (ΩTruncSwitch {A = A} n m) (snd (hLevelTrunc∙ m ((Ω^ n) A))) + ≡ pt ((Ω^ n) (hLevelTrunc∙ (n + m) A)) + ΩTruncSwitch∙ {A = A} n zero = + isContr→isProp + ((subst isContr + (λ i → (typ ((Ω^ n) (hLevelTrunc∙ (+-comm zero n i) A)))) + (isContrΩTrunc n))) _ _ + ΩTruncSwitch∙ {A = A} zero (suc m) = refl + ΩTruncSwitch∙ {A = A} (suc n) (suc m) = ∙∙lCancel _ + + +ΩTruncSwitch-hom : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) (p q : _) + → Iso.fun (ΩTruncSwitch {A = A} (suc n) (suc m)) ∣ p ∙ q ∣ + ≡ Iso.fun (ΩTruncSwitch {A = A} (suc n) (suc m)) ∣ p ∣ + ∙ Iso.fun (ΩTruncSwitch {A = A} (suc n) (suc m)) ∣ q ∣ +ΩTruncSwitch-hom {A = A} n m p q = + cong (Iso.fun (ΩTruncSwitchLem {A = A} n m)) + (cong-∙ ∣_∣ₕ p q) + ∙ Ω→pres∙ (ΩTruncSwitchFun n m) (cong ∣_∣ₕ p) (cong ∣_∣ₕ q) + +2TruncΩIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (hLevelTrunc 2 (fst ((Ω^ n) A))) + (typ ((Ω^ n) (hLevelTrunc∙ (2 + n) A))) +2TruncΩIso zero = idIso +2TruncΩIso {A = A} (suc n) = + compIso + (ΩTruncSwitch (suc n) 2) + (pathToIso λ i → typ ((Ω^ suc n) (hLevelTrunc∙ (+-comm (suc n) 2 i) A))) + +hLevΩ+ : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → isOfHLevel (m + n) (typ A) → isOfHLevel n (typ ((Ω^ m) A)) +hLevΩ+ n zero p = p +hLevΩ+ {A = A} zero (suc zero) p = refl , λ _ → isProp→isSet p _ _ _ _ +hLevΩ+ {A = A} zero (suc (suc zero)) p = + refl , λ y → isOfHLevelSuc 2 p _ _ _ _ refl y +hLevΩ+ {A = A} zero (suc (suc (suc m))) p = + transport + (λ i → isContr (typ (Ω (ua∙ + (isoToEquiv (flipΩIso {A = A} (suc m))) (flipΩrefl m) (~ i))))) + (hLevΩ+ {A = Ω A} zero (suc (suc m)) + (subst (λ x → isOfHLevel x (typ (Ω A))) + (+-comm zero (suc (suc m))) + (lem (pt A) (pt A)))) + where + lem : isOfHLevel (3 + m) (typ A) + lem = subst (λ x → isOfHLevel x (typ A)) + (λ i → suc (+-comm (2 + m) zero i)) p +hLevΩ+ {A = A} (suc n) (suc m) p = + subst (isOfHLevel (suc n)) + (sym (ua (isoToEquiv (flipΩIso {A = A} m)))) + (hLevΩ+ {A = Ω A} (suc n) m + (isOfHLevelPath' (m + suc n) p _ _)) + +isSetΩTrunc : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (isSet (typ (Ω ((Ω^ n) (hLevelTrunc∙ (suc (suc (suc n))) A))))) +isSetΩTrunc {A = A} zero = isOfHLevelTrunc 3 _ _ +isSetΩTrunc {A = A} (suc n) = + hLevΩ+ 2 (suc (suc n)) + (transport + (λ i → isOfHLevel (+-comm 2 (2 + n) i) (hLevelTrunc (4 + n) (typ A))) + (isOfHLevelTrunc (suc (suc (suc (suc n)))))) + +πTruncIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (π n A) (π n (hLevelTrunc∙ (2 + n) A)) +πTruncIso {A = A} zero = + compIso (invIso (setTruncIdempotentIso squash₂)) + (setTruncIso setTruncTrunc2Iso) +πTruncIso {A = A} (suc n) = + compIso setTruncTrunc2Iso + (compIso + (2TruncΩIso (suc n)) + (invIso (setTruncIdempotentIso (isSetΩTrunc n)))) + +πTruncGroupIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupIso (πGr n A) (πGr n (hLevelTrunc∙ (3 + n) A)) +fst (πTruncGroupIso n) = πTruncIso (suc n) +snd (πTruncGroupIso {A = A} n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ a b + → cong (inv (setTruncIdempotentIso (isSetΩTrunc n))) + (cong + (transport + (λ i → typ ((Ω^ suc n) (hLevelTrunc∙ (+-comm (suc n) 2 i) A)))) + (ΩTruncSwitch-hom n 1 a b) + ∙ transpΩTruncSwitch + (λ w → ((Ω^ n) (hLevelTrunc∙ w A))) (+-comm (suc n) 2) _ _)) + where + transpΩTruncSwitch : ∀ {ℓ} (A : ℕ → Pointed ℓ) {n m : ℕ} + (r : n ≡ m) (p q : typ (Ω (A n))) + → subst (λ n → typ (Ω (A n))) r (p ∙ q) + ≡ subst (λ n → typ (Ω (A n))) r p + ∙ subst (λ n → typ (Ω (A n))) r q + transpΩTruncSwitch A {n = n} = + J (λ m r → (p q : typ (Ω (A n))) + → subst (λ n → typ (Ω (A n))) r (p ∙ q) + ≡ subst (λ n → typ (Ω (A n))) r p + ∙ subst (λ n → typ (Ω (A n))) r q) + λ p q → transportRefl _ ∙ cong₂ _∙_ + (sym (transportRefl p)) (sym (transportRefl q)) + +-- Often, we prefer thinking of Ωⁿ A as (Sⁿ →∙ A). +-- The goal of the following lemmas is to show that the maps +-- Ωⁿ A → Ωⁿ B and Ωⁿ (fib f) →∙ Ωⁿ A get sent to post composition +-- under the equivalence Ωⁿ A as (Sⁿ →∙ A). This also gives a proof +-- that post composition induces a homomorphism of homotopy groups. + +-- The following lemmas is not pretty but very helpful +private + bigLemma : ∀ {ℓ ℓ'} {A₁ B₁ C₁ : Type ℓ} {A₂ B₂ C₂ : Type ℓ'} + (A₁→B₁ : A₁ ≃ B₁) (B₁→C₁ : B₁ ≃ C₁) + (A₂→B₂ : A₂ ≃ B₂) (B₂→C₂ : B₂ ≃ C₂) + (A₁→A₂ : A₁ → A₂) + (B₁→B₂ : B₁ → B₂) + (C₁→C₂ : C₁ → C₂) + → (B₁→B₂ ∘ (fst A₁→B₁)) ≡ (fst A₂→B₂ ∘ A₁→A₂) + → C₁→C₂ ∘ fst B₁→C₁ ≡ fst B₂→C₂ ∘ B₁→B₂ + → C₁→C₂ ∘ fst B₁→C₁ ∘ fst A₁→B₁ + ≡ fst B₂→C₂ ∘ fst A₂→B₂ ∘ A₁→A₂ + bigLemma {B₁ = B₁} {C₁ = C₁} {A₂ = A₂} {B₂ = B₂} {C₂ = C₂} = + EquivJ + (λ A₁ A₁→B₁ → (B₁→C₁ : B₁ ≃ C₁) (A₂→B₂ : A₂ ≃ B₂) + (B₂→C₂ : B₂ ≃ C₂) (A₁→A₂ : A₁ → A₂) (B₁→B₂ : B₁ → B₂) + (C₁→C₂ : C₁ → C₂) → + B₁→B₂ ∘ fst A₁→B₁ ≡ fst A₂→B₂ ∘ A₁→A₂ → + C₁→C₂ ∘ fst B₁→C₁ ≡ fst B₂→C₂ ∘ B₁→B₂ → + C₁→C₂ ∘ fst B₁→C₁ ∘ fst A₁→B₁ ≡ fst B₂→C₂ ∘ fst A₂→B₂ ∘ A₁→A₂) + (EquivJ (λ B₁ B₁→C₁ → (A₂→B₂ : A₂ ≃ B₂) (B₂→C₂ : B₂ ≃ C₂) + (A₁→A₂ : B₁ → A₂) (B₁→B₂ : B₁ → B₂) (C₁→C₂ : C₁ → C₂) → + (B₁→B₂) ≡ (fst A₂→B₂ ∘ A₁→A₂) → + (C₁→C₂ ∘ (fst B₁→C₁)) ≡ (fst B₂→C₂ ∘ (B₁→B₂)) → + (C₁→C₂ ∘ (fst B₁→C₁)) ≡ (fst B₂→C₂ ∘ (fst A₂→B₂ ∘ A₁→A₂))) + (EquivJ (λ A₂ A₂→B₂ → (B₂→C₂ : B₂ ≃ C₂) (A₁→A₂ : C₁ → A₂) + (B₁→B₂ : C₁ → B₂) (C₁→C₂ : C₁ → C₂) → + B₁→B₂ ≡ (fst A₂→B₂ ∘ A₁→A₂) → + (C₁→C₂) ≡ (fst B₂→C₂ ∘ B₁→B₂) → + (C₁→C₂) ≡ fst B₂→C₂ ∘ (fst A₂→B₂ ∘ A₁→A₂)) + (EquivJ (λ B₂ B₂→C₂ → (A₁→A₂ B₁→B₂ : C₁ → B₂) (C₁→C₂ : C₁ → C₂) → + B₁→B₂ ≡ A₁→A₂ → + C₁→C₂ ≡ (fst B₂→C₂ ∘ B₁→B₂) → + C₁→C₂ ≡ (fst B₂→C₂ ∘ A₁→A₂)) + λ _ _ _ p q → q ∙ p))) + +{- +We want to show that the following square +commutes. + + Ωⁿ f +Ωⁿ A ----------→ Ωⁿ B +| | +| | +v f∘_ v +(Sⁿ→∙A) ------> (Sⁿ→∙B) +-} + +Ω^→≈post∘∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (f : A →∙ B) + → Path ((Ω^ (suc n)) A →∙ (S₊∙ (suc n) →∙ B ∙)) + (post∘∙ (S₊∙ (suc n)) f ∘∙ Ω→SphereMap∙ (suc n)) + (Ω→SphereMap∙ (suc n) ∘∙ Ω^→ (suc n) f) +Ω^→≈post∘∙ {A = A} {B = B} zero f = + →∙Homogeneous≡ + (subst isHomogeneous + (ua∙ (Ω→SphereMap 1 , (isEquiv-Ω→SphereMap 1)) + (Ω→SphereMap∙ 1 {A = B} .snd)) + (isHomogeneousPath _ _)) + (funExt λ p → + ΣPathP ((funExt (λ { base → snd f + ; (loop i) j → + doubleCompPath-filler + (sym (snd f)) (cong (fst f) p) (snd f) j i})) + , (sym (lUnit (snd f)) ◁ λ i j → snd f (i ∨ j)))) +Ω^→≈post∘∙ {A = A} {B = B} (suc n) f = + →∙Homogeneous≡ + (subst isHomogeneous + (ua∙ (Ω→SphereMap (2 + n) , (isEquiv-Ω→SphereMap (2 + n))) + (Ω→SphereMap∙ (2 + n) {A = B} .snd)) + (isHomogeneousPath _ _)) + ((funExt λ p + → (λ i → post∘∙ (S₊∙ (2 + n)) f .fst (Ω→SphereMap-split (suc n) p i)) + ∙∙ funExt⁻ + (bigLemma + (Ω→SphereMapSplit₁ (suc n) + , isEquivΩ→ _ (isEquiv-Ω→SphereMap (suc n))) + (ΩSphereMap (suc n) , isoToIsEquiv (invIso (SphereMapΩIso (suc n)))) + (Ω→SphereMapSplit₁ (suc n) + , isEquivΩ→ _ (isEquiv-Ω→SphereMap (suc n))) + (ΩSphereMap (suc n) , isoToIsEquiv (invIso (SphereMapΩIso (suc n)))) + (Ω^→ (2 + n) f .fst) (Ω→ (post∘∙ (S₊∙ (suc n)) f) .fst) + (post∘∙ (S₊∙ (2 + n)) f .fst) + (funExt topSquare) + (sym (funExt (bottomSquare f)))) + p + ∙∙ sym (Ω→SphereMap-split (suc n) (Ω^→ (2 + n) f .fst p)))) + where + topSquare : (p : typ ((Ω^ (2 + n)) A)) + → Path (typ (Ω ((S₊∙ (suc n)) →∙ B ∙))) + ((Ω→ (post∘∙ (S₊∙ (suc n)) f) .fst ∘ Ω→ (Ω→SphereMap∙ (suc n)) .fst) p) + (((Ω→ (Ω→SphereMap∙ (suc n))) .fst ∘ (Ω^→ (suc (suc n)) f .fst)) p) + topSquare p = sym (Ω→∘ (post∘∙ (S₊∙ (suc n)) f) (Ω→SphereMap∙ (suc n)) p) + ∙ (λ i → Ω→ (Ω^→≈post∘∙ {A = A} {B = B} n f i) .fst p) + ∙ Ω→∘ (Ω→SphereMap∙ (suc n)) (Ω^→ (suc n) f) p + + bottomSquare : (f : A →∙ B) (g : typ (Ω (S₊∙ (suc n) →∙ A ∙))) + → Path (S₊∙ (2 + n) →∙ B) + (ΩSphereMap (suc n) (Ω→ (post∘∙ (S₊∙ (suc n)) f) .fst g)) + ((post∘∙ (S₊∙ (2 + n)) f .fst ∘ ΩSphereMap (suc n)) g) + bottomSquare = + →∙J (λ b₀ f → (g : typ (Ω (S₊∙ (suc n) →∙ A ∙))) + → Path (S₊∙ (suc (suc n)) →∙ (fst B , b₀)) + (ΩSphereMap (suc n) (Ω→ (post∘∙ (S₊∙ (suc n)) f) .fst g)) + ((post∘∙ (S₊∙ (suc (suc n))) f .fst ∘ ΩSphereMap (suc n)) g)) + λ f g → ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → lem f g a j i})) + , lUnit refl) + where + lem : (f : typ A → typ B) (g : typ (Ω (S₊∙ (suc n) →∙ A ∙))) + → (a : S₊ (suc n)) + → cong (fst (ΩSphereMap (suc n) + (Ω→ (post∘∙ (S₊∙ (suc n)) (f , refl)) .fst g))) + (merid a) + ≡ cong (fst ((f , refl) ∘∙ ΩSphereMap (suc n) g)) (merid a) + lem f g a = + (λ i → funExt⁻ + (cong-∙∙ fst (sym (snd (post∘∙ (S₊∙ (suc n)) (f , (λ _ → f (snd A)))))) + (cong (fst (post∘∙ (S₊∙ (suc n)) (f , (λ _ → f (snd A))))) g) + (snd (post∘∙ (S₊∙ (suc n)) (f , (λ _ → f (snd A))))) i) a) + ∙ sym (rUnit (λ i → f (fst (g i) a))) + +{- We can use this to define prove that post composition induces a homomorphism +πₙ A → πₙ B-} + +π'∘∙fun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → π' (suc n) A → π' (suc n) B +π'∘∙fun n f = sMap (f ∘∙_) + +GroupHomπ≅π'PathP : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') (n : ℕ) + → GroupHom (πGr n A) (πGr n B) ≡ GroupHom (π'Gr n A) (π'Gr n B) +GroupHomπ≅π'PathP A B n i = + GroupHom (fst (GroupPath _ _) (GroupIso→GroupEquiv (π'Gr≅πGr n A)) (~ i)) + (fst (GroupPath _ _) (GroupIso→GroupEquiv (π'Gr≅πGr n B)) (~ i)) + +πFun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → π (suc n) A → π (suc n) B +πFun n f = sMap (fst (Ω^→ (suc n) f)) + +πHom : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → GroupHom (πGr n A) (πGr n B) +fst (πHom n f) = πFun n f +snd (πHom n f) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (Ω^→pres∙ f n p q)) + +π'∘∙Hom' : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → GroupHom (π'Gr n A) (π'Gr n B) +π'∘∙Hom' {A = A} {B = B} n f = + transport (λ i → GroupHomπ≅π'PathP A B n i) + (πHom n f) + +π'∘∙Hom'≡π'∘∙fun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) → π'∘∙Hom' n f .fst ≡ π'∘∙fun n f +π'∘∙Hom'≡π'∘∙fun n f = + funExt (sElim (λ _ → isSetPathImplicit) + λ g → cong ∣_∣₂ + ((λ i → inv (IsoSphereMapΩ (suc n)) + (transportRefl (Ω^→ (suc n) f .fst + (transportRefl (fun (IsoSphereMapΩ (suc n)) g) i)) i)) + ∙ sym (funExt⁻ (cong fst (Ω^→≈post∘∙ n f)) + (fun (IsoSphereMapΩ (suc n)) g)) + ∙ cong (f ∘∙_) (leftInv (IsoSphereMapΩ (suc n)) g))) + +π'∘∙Hom : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → GroupHom (π'Gr n A) (π'Gr n B) +fst (π'∘∙Hom n f) = sMap (f ∘∙_) +snd (π'∘∙Hom {A = A} {B = B} n f) = isHom∘∙ + where + abstract + isHom∘∙ : IsGroupHom (π'Gr n A .snd) (fst (π'∘∙Hom n f)) (π'Gr n B .snd) + isHom∘∙ = + transport (λ i → IsGroupHom (π'Gr n A .snd) + (π'∘∙Hom'≡π'∘∙fun n f i) + (π'Gr n B .snd)) + (π'∘∙Hom' n f .snd) + +-- post composition with an equivalence induces an +-- isomorphism of homotopy groups +π'eqFun : ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) + → A ≃∙ B + → (π' (suc n) A) → π' (suc n) B +π'eqFun n p = π'∘∙fun n (≃∙map p) + +π'eqFun-idEquiv : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → π'eqFun n (idEquiv (fst A) , (λ _ → pt A)) + ≡ idfun _ +π'eqFun-idEquiv n = + funExt (sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (∘∙-idʳ f)) + +π'eqFunIsEquiv : + ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) + → (e : A ≃∙ B) + → isEquiv (π'eqFun n e) +π'eqFunIsEquiv {B = B} n = + Equiv∙J (λ A e → isEquiv (π'eqFun n e)) + (subst isEquiv (sym (π'eqFun-idEquiv n)) + (idIsEquiv (π' (suc n) B))) + +π'eqFunIsHom : ∀ {ℓ} {A B : Pointed ℓ}(n : ℕ) + → (e : A ≃∙ B) + → IsGroupHom (π'Gr n A .snd) (π'eqFun n e) + (π'Gr n B .snd) +π'eqFunIsHom {B = B} n = + Equiv∙J (λ A e → IsGroupHom (π'Gr n A .snd) (π'eqFun n e) (π'Gr n B .snd)) + (subst (λ x → IsGroupHom (π'Gr n B .snd) x (π'Gr n B .snd)) + (sym (π'eqFun-idEquiv n)) + (makeIsGroupHom λ _ _ → refl)) + +π'Iso : ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) + → A ≃∙ B + → GroupEquiv (π'Gr n A) (π'Gr n B) +fst (fst (π'Iso n e)) = π'eqFun n e +snd (fst (π'Iso n e)) = π'eqFunIsEquiv n e +snd (π'Iso n e) = π'eqFunIsHom n e diff --git a/Cubical/Homotopy/Group/LES.agda b/Cubical/Homotopy/Group/LES.agda new file mode 100644 index 0000000000..c3a2f53c52 --- /dev/null +++ b/Cubical/Homotopy/Group/LES.agda @@ -0,0 +1,665 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains: +1. The long exact sequence of loop spaces Ωⁿ (fib f) → Ωⁿ A → Ωⁿ B +2. The long exact sequence of homotopy groups πₙ(fib f) → πₙ A → πₙ B +3. Some lemmas relating the map in the sequence to maps using the + other definition of πₙ (maps from Sⁿ) +-} +module Cubical.Homotopy.Group.LES where + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Group.Base + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open Iso +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function + +open import Cubical.HITs.SetTruncation + renaming (rec to sRec + ; elim to sElim ; elim2 to sElim2 + ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec) + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group + +-- We will need an explicitly defined equivalence +-- (PathP (λ i → p i ≡ y) q q) ≃ (sym q ∙∙ p ∙∙ q ≡ refl) +-- This is given by →∙∙lCancel below +module _ {ℓ : Level} {A : Type ℓ} {x y : A} (p : x ≡ x) (q : x ≡ y) where + →∙∙lCancel-fill : PathP (λ i → p i ≡ y) q q → I → I → I → A + →∙∙lCancel-fill PP k i j = + hfill (λ k → λ {(i = i0) → doubleCompPath-filler (sym q) p q k j + ; (i = i1) → y + ; (j = i0) → q (i ∨ k) + ; (j = i1) → q (i ∨ k)}) + (inS (PP j i)) + k + + ←∙∙lCancel-fill : sym q ∙∙ p ∙∙ q ≡ refl → I → I → I → A + ←∙∙lCancel-fill PP k i j = + hfill (λ k → λ {(i = i0) → q (j ∨ ~ k) + ; (i = i1) → q (j ∨ ~ k) + ; (j = i0) → doubleCompPath-filler (sym q) p q (~ k) i + ; (j = i1) → y}) + (inS (PP j i)) + k + + →∙∙lCancel : PathP (λ i → p i ≡ y) q q → sym q ∙∙ p ∙∙ q ≡ refl + →∙∙lCancel PP i j = →∙∙lCancel-fill PP i1 i j + + ←∙∙lCancel : sym q ∙∙ p ∙∙ q ≡ refl → PathP (λ i → p i ≡ y) q q + ←∙∙lCancel PP i j = ←∙∙lCancel-fill PP i1 i j + + ←∙∙lCancel→∙∙lCancel : (PP : PathP (λ i → p i ≡ y) q q) + → ←∙∙lCancel (→∙∙lCancel PP) ≡ PP + ←∙∙lCancel→∙∙lCancel PP r i j = + hcomp (λ k → λ {(r = i0) → ←∙∙lCancel-fill (→∙∙lCancel PP) k i j + ; (r = i1) → PP i j + ; (j = i0) → doubleCompPath-filler (sym q) p q (~ k ∧ ~ r) i + ; (j = i1) → y + ; (i = i0) → q (j ∨ ~ k ∧ ~ r) + ; (i = i1) → q (j ∨ ~ k ∧ ~ r)}) + (hcomp (λ k → λ {(r = i0) → →∙∙lCancel-fill PP k j i + ; (r = i1) → PP i j + ; (j = i0) → doubleCompPath-filler (sym q) p q (k ∧ ~ r) i + ; (j = i1) → y + ; (i = i0) → q (j ∨ k ∧ ~ r) + ; (i = i1) → q (j ∨ k ∧ ~ r)}) + (PP i j)) + + →∙∙lCancel←∙∙lCancel : (PP : sym q ∙∙ p ∙∙ q ≡ refl) + → →∙∙lCancel (←∙∙lCancel PP) ≡ PP + →∙∙lCancel←∙∙lCancel PP r i j = + hcomp (λ k → λ {(r = i0) → →∙∙lCancel-fill (←∙∙lCancel PP) k i j + ; (r = i1) → PP i j + ; (j = i0) → q (i ∨ k ∨ r) + ; (j = i1) → q (i ∨ k ∨ r) + ; (i = i0) → doubleCompPath-filler (sym q) p q (r ∨ k) j + ; (i = i1) → y}) + (hcomp (λ k → λ {(r = i0) → ←∙∙lCancel-fill PP k j i + ; (r = i1) → PP i j + ; (j = i0) → q (i ∨ r ∨ ~ k) + ; (j = i1) → q (i ∨ r ∨ ~ k) + ; (i = i0) → doubleCompPath-filler (sym q) p q (r ∨ ~ k) j + ; (i = i1) → y}) + (PP i j)) + +←∙∙lCancel-refl-refl : + {ℓ : Level} {A : Type ℓ} {x : A} (p : refl {x = x} ≡ refl) + → ←∙∙lCancel {x = x} {y = x} refl refl (sym (rUnit refl) ∙ p) + ≡ flipSquare p +←∙∙lCancel-refl-refl p k i j = + hcomp (λ r → λ { (i = i0) → p i0 i0 + ; (i = i1) → p i0 i0 + ; (j = i0) → rUnit (λ _ → p i0 i0) (~ r) i + ; (j = i1) → p i0 i0 + ; (k = i0) → ←∙∙lCancel-fill refl refl (sym (rUnit refl) ∙ p) r i j + ; (k = i1) → compPath-filler' (sym (rUnit refl)) p (~ r) j i}) + ((sym (rUnit refl) ∙ p) j i) + +{- We need an iso Ω(fib f) ≅ fib(Ω f) -} +ΩFibreIso : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Iso (typ (Ω (fiber (fst f) (pt B) , (pt A) , snd f))) + (fiber (Ω→ f .fst) refl) +fun (ΩFibreIso f) p = (cong fst p) , + →∙∙lCancel (cong (fst f) (cong fst p)) (snd f) + (cong snd p) +fst (inv (ΩFibreIso f) (p , q) i) = p i +snd (inv (ΩFibreIso f) (p , q) i) = ←∙∙lCancel (cong (fst f) p) (snd f) q i +rightInv (ΩFibreIso f) (p , q) = ΣPathP (refl , →∙∙lCancel←∙∙lCancel _ _ q) +fst (leftInv (ΩFibreIso f) p i j) = fst (p j) +snd (leftInv (ΩFibreIso f) p i j) k = + ←∙∙lCancel→∙∙lCancel _ _ (cong snd p) i j k + +{- Some homomorphism properties of the above iso -} +ΩFibreIsopres∙fst : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → (p q : (typ (Ω (fiber (fst f) (pt B) , (pt A) , snd f)))) + → fst (fun (ΩFibreIso f) (p ∙ q)) + ≡ fst (fun (ΩFibreIso f) p) ∙ fst (fun (ΩFibreIso f) q) +ΩFibreIsopres∙fst f p q = cong-∙ fst p q + +ΩFibreIso⁻pres∙snd : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (f : A →∙ B) (p q : typ (Ω (Ω B))) + → inv (ΩFibreIso f) (refl , (Ω→ f .snd ∙ p ∙ q)) + ≡ inv (ΩFibreIso f) (refl , Ω→ f .snd ∙ p) + ∙ inv (ΩFibreIso f) (refl , Ω→ f .snd ∙ q) +ΩFibreIso⁻pres∙snd {A = A} {B = B}= + →∙J (λ b₀ f → (p q : typ (Ω (Ω (fst B , b₀)))) + → inv (ΩFibreIso f) (refl , (Ω→ f .snd ∙ p ∙ q)) + ≡ inv (ΩFibreIso f) (refl , Ω→ f .snd ∙ p) + ∙ inv (ΩFibreIso f) (refl , Ω→ f .snd ∙ q)) + ind + where + ind : (f : typ A → typ B) (p q : typ (Ω (Ω (fst B , f (pt A))))) + → inv (ΩFibreIso (f , refl)) (refl , (sym (rUnit refl) ∙ p ∙ q)) + ≡ inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ p) + ∙ inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ q) + fst (ind f p q i j) = + (rUnit refl + ∙ sym (cong-∙ fst + (inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ p)) + (inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ q)))) i j + snd (ind f p q i j) k = + hcomp (λ r + → λ {(i = i0) → ←∙∙lCancel-refl-refl (p ∙ q) (~ r) j k -- + ; (i = i1) → + snd (compPath-filler + (inv (ΩFibreIso (f , refl)) + (refl , sym (rUnit refl) ∙ p)) + (inv (ΩFibreIso (f , refl)) + (refl , sym (rUnit refl) ∙ q)) r j) k + ; (j = i0) → f (snd A) + ; (j = i1) → snd (inv (ΩFibreIso (f , refl)) + (refl , sym (rUnit refl) ∙ q) (r ∨ ~ i)) k + ; (k = i0) → main r i j + ; (k = i1) → f (snd A)}) + (hcomp (λ r → λ {(i = i0) → (p ∙ q) k j + ; (i = i1) → ←∙∙lCancel-refl-refl p (~ r) j k + ; (j = i0) → f (snd A) + ; (j = i1) → ←∙∙lCancel-refl-refl q (~ r) (~ i) k + ; (k = i0) → f (pt A) + ; (k = i1) → f (snd A)}) + (hcomp (λ r → λ {(i = i0) → (compPath-filler' p q r) k j + ; (i = i1) → p (k ∨ ~ r) j + ; (j = i0) → f (snd A) + ; (j = i1) → q k (~ i) + ; (k = i0) → p (k ∨ ~ r) j + ; (k = i1) → f (snd A)}) + (q k (~ i ∧ j)))) + where + P = (inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ p)) + Q = (inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ q)) + + main : I → I → I → fst B + main r i j = + hcomp (λ k → λ {(i = i0) → f (snd A) + ; (i = i1) → f (fst (compPath-filler P Q (r ∨ ~ k) j)) + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (r = i0) → f (fst (compPath-filler P Q (i ∧ ~ k) j)) + ; (r = i1) → f ((rUnit refl ∙ sym (cong-∙ fst P Q)) i j)}) + (hcomp (λ k → λ {(i = i0) → f (rUnit (λ _ → pt A) (~ k ∧ r) j) + ; (i = i1) → f (fst ((P ∙ Q) j)) + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (r = i0) → f (fst (compPath-filler P Q i j)) + ; (r = i1) → f ((compPath-filler' (rUnit refl) + (sym (cong-∙ fst P Q)) k) i j)}) + (hcomp (λ k → λ {(i = i0) → f (rUnit (λ _ → pt A) (k ∧ r) j) + ; (i = i1) → f (fst (compPath-filler P Q k j)) + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (r = i0) → f (fst (compPath-filler P Q (i ∧ k) j)) + ; (r = i1) → f ((cong-∙∙-filler fst refl P Q) k (~ i) j)}) + (f (snd A)))) + +ΩFibreIso∙ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Iso.fun (ΩFibreIso f) refl ≡ (refl , (∙∙lCancel (snd f))) +ΩFibreIso∙ {A = A} {B = B} = + →∙J (λ b f → Iso.fun (ΩFibreIso f) refl ≡ (refl , (∙∙lCancel (snd f)))) + λ f → ΣPathP (refl , help f) + where + help : (f : fst A → fst B) → + →∙∙lCancel (λ i → f (snd A)) refl (λ i → refl) + ≡ ∙∙lCancel refl + help f i j r = + hcomp (λ k → λ {(i = i0) → + →∙∙lCancel-fill (λ _ → f (snd A)) refl refl k j r + ; (i = i1) → ∙∙lCancel-fill (λ _ → f (snd A)) j r k + ; (j = i0) → rUnit (λ _ → f (snd A)) k r + ; (j = i1) → f (snd A) + ; (r = i1) → f (snd A) + ; (r = i0) → f (snd A)}) + (f (snd A)) + +ΩFibreIso⁻∙ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Iso.inv (ΩFibreIso f) (refl , (∙∙lCancel (snd f))) ≡ refl +ΩFibreIso⁻∙ f = + cong (Iso.inv (ΩFibreIso f)) (sym (ΩFibreIso∙ f)) ∙ leftInv (ΩFibreIso f) refl + +{- Ωⁿ (fib f) ≃∙ fib (Ωⁿ f) -} +Ω^Fibre≃∙ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → ((Ω^ n) (fiber (fst f) (pt B) , (pt A) , snd f)) + ≃∙ ((fiber (Ω^→ n f .fst) (snd ((Ω^ n) B))) + , (snd ((Ω^ n) A)) , (Ω^→ n f .snd)) +Ω^Fibre≃∙ zero f = (idEquiv _) , refl +Ω^Fibre≃∙ (suc n) f = + compEquiv∙ + (Ω≃∙ (Ω^Fibre≃∙ n f)) + ((isoToEquiv (ΩFibreIso (Ω^→ n f))) , ΩFibreIso∙ (Ω^→ n f)) + +{- Its inverse iso directly defined -} +Ω^Fibre≃∙⁻ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → ((fiber (Ω^→ n f .fst) (snd ((Ω^ n) B))) + , (snd ((Ω^ n) A)) , (Ω^→ n f .snd)) + ≃∙ ((Ω^ n) (fiber (fst f) (pt B) , (pt A) , snd f)) +Ω^Fibre≃∙⁻ zero f = (idEquiv _) , refl +Ω^Fibre≃∙⁻ (suc n) f = + compEquiv∙ + ((isoToEquiv (invIso (ΩFibreIso (Ω^→ n f)))) + , (ΩFibreIso⁻∙ (Ω^→ n f))) + (Ω≃∙ (Ω^Fibre≃∙⁻ n f)) + +isHomogeneousΩ^→fib : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → isHomogeneous + ((fiber (Ω^→ (suc n) f .fst) (snd ((Ω^ (suc n)) B))) + , (snd ((Ω^ (suc n)) A)) , (Ω^→ (suc n) f .snd)) +isHomogeneousΩ^→fib n f = + subst isHomogeneous (ua∙ ((fst (Ω^Fibre≃∙ (suc n) f))) + (snd (Ω^Fibre≃∙ (suc n) f))) + (isHomogeneousPath _ _) + +Ω^Fibre≃∙sect : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → (≃∙map (Ω^Fibre≃∙⁻ n f) ∘∙ ≃∙map (Ω^Fibre≃∙ n f)) + ≡ idfun∙ _ +Ω^Fibre≃∙sect zero f = ΣPathP (refl , (sym (rUnit refl))) +Ω^Fibre≃∙sect (suc n) f = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt + λ p → cong (fst (fst (Ω≃∙ (Ω^Fibre≃∙⁻ n f)))) + (leftInv (ΩFibreIso (Ω^→ n f)) + ((fst (fst (Ω≃∙ (Ω^Fibre≃∙ n f))) p))) + ∙ sym (Ω→∘ (≃∙map (Ω^Fibre≃∙⁻ n f)) + (≃∙map (Ω^Fibre≃∙ n f)) p) + ∙ (λ i → (Ω→ (Ω^Fibre≃∙sect n f i)) .fst p) + ∙ sym (rUnit p)) + +Ω^Fibre≃∙retr : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → (≃∙map (Ω^Fibre≃∙ n f) ∘∙ ≃∙map (Ω^Fibre≃∙⁻ n f)) + ≡ idfun∙ _ +Ω^Fibre≃∙retr zero f = ΣPathP (refl , (sym (rUnit refl))) +Ω^Fibre≃∙retr (suc n) f = + →∙Homogeneous≡ (isHomogeneousΩ^→fib n f) + (funExt (λ p → + cong (fun (ΩFibreIso (Ω^→ n f))) + ((sym (Ω→∘ (≃∙map (Ω^Fibre≃∙ n f)) + (≃∙map (Ω^Fibre≃∙⁻ n f)) + (inv (ΩFibreIso (Ω^→ n f)) p))) + ∙ (λ i → Ω→ (Ω^Fibre≃∙retr n f i) .fst (inv (ΩFibreIso (Ω^→ n f)) p)) + ∙ sym (rUnit (inv (ΩFibreIso (Ω^→ n f)) p))) + ∙ rightInv (ΩFibreIso (Ω^→ n f)) p)) + +Ω^Fibre≃∙' : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → ((Ω^ n) (fiber (fst f) (pt B) , (pt A) , snd f)) + ≃∙ ((fiber (Ω^→ n f .fst) (snd ((Ω^ n) B))) + , (snd ((Ω^ n) A)) , (Ω^→ n f .snd)) +Ω^Fibre≃∙' zero f = idEquiv _ , refl +Ω^Fibre≃∙' (suc zero) f = + (isoToEquiv (ΩFibreIso (Ω^→ zero f))) , ΩFibreIso∙ (Ω^→ zero f) +Ω^Fibre≃∙' (suc (suc n)) f = + compEquiv∙ + (Ω≃∙ (Ω^Fibre≃∙ (suc n) f)) + ((isoToEquiv (ΩFibreIso (Ω^→ (suc n) f))) , ΩFibreIso∙ (Ω^→ (suc n) f)) + +-- The long exact sequence of loop spaces. +module ΩLES {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where + {- Fibre of f -} + fibf : Pointed _ + fibf = fiber (fst f) (pt B) , (pt A , snd f) + + {- Fibre of Ωⁿ f -} + fibΩ^f : (n : ℕ) → Pointed _ + fst (fibΩ^f n) = fiber (Ω^→ n f .fst) (snd ((Ω^ n) B)) + snd (fibΩ^f n) = (snd ((Ω^ n) A)) , (Ω^→ n f .snd) + + Ω^fibf : (n : ℕ) → Pointed _ + Ω^fibf n = (Ω^ n) fibf + + {- Helper function fib (Ωⁿ f) → Ωⁿ A -} + fibΩ^f→A : (n : ℕ) → fibΩ^f n →∙ (Ω^ n) A + fst (fibΩ^f→A n) = fst + snd (fibΩ^f→A n) = refl + + {- The main function Ωⁿ(fib f) → Ωⁿ A, which is just the composition + Ωⁿ(fib f) ≃ fib (Ωⁿ f) → Ωⁿ A, where the last function is + fibΩ^f→A. Hence most proofs will concern fibΩ^f→A, since it is easier to + work with. -} + Ω^fibf→A : (n : ℕ) → Ω^fibf n →∙ (Ω^ n) A + Ω^fibf→A n = fibΩ^f→A n ∘∙ ≃∙map (Ω^Fibre≃∙ n f) + + {- The function preserves path composition -} + Ω^fibf→A-pres∙ : (n : ℕ) → (p q : Ω^fibf (suc n) .fst) + → Ω^fibf→A (suc n) .fst (p ∙ q) + ≡ Ω^fibf→A (suc n) .fst p + ∙ Ω^fibf→A (suc n) .fst q + Ω^fibf→A-pres∙ n p q = + cong (fst (fibΩ^f→A (suc n))) + (cong (fun (ΩFibreIso (Ω^→ n f))) + (Ω→pres∙ (≃∙map (Ω^Fibre≃∙ n f)) p q)) + ∙ ΩFibreIsopres∙fst (Ω^→ n f) + (fst (Ω→ (≃∙map (Ω^Fibre≃∙ n f))) p) + (fst (Ω→ (≃∙map (Ω^Fibre≃∙ n f))) q) + + {- The function Ωⁿ A → Ωⁿ B -} + A→B : (n : ℕ) → (Ω^ n) A →∙ (Ω^ n) B + A→B n = Ω^→ n f + + {- It preserves path composition -} + A→B-pres∙ : (n : ℕ) → (p q : typ ((Ω^ (suc n)) A)) + → fst (A→B (suc n)) (p ∙ q) + ≡ fst (A→B (suc n)) p ∙ fst (A→B (suc n)) q + A→B-pres∙ n p q = Ω^→pres∙ f n p q + + {- Helper function Ωⁿ⁺¹ B → fib (Ωⁿ f) -} + ΩB→fibΩ^f : (n : ℕ) → ((Ω^ (suc n)) B) →∙ fibΩ^f n + fst (ΩB→fibΩ^f n) x = (snd ((Ω^ n) A)) , (Ω^→ n f .snd ∙ x) + snd (ΩB→fibΩ^f n) = ΣPathP (refl , (sym (rUnit _))) + + {- The main function Ωⁿ⁺¹ B → Ωⁿ (fib f), + factoring through the above function -} + ΩB→Ω^fibf : (n : ℕ) → (Ω^ (suc n)) B →∙ Ω^fibf n + ΩB→Ω^fibf n = + (≃∙map (Ω^Fibre≃∙⁻ n f)) + ∘∙ ΩB→fibΩ^f n + + {- It preserves path composition -} + ΩB→Ω^fibf-pres∙ : (n : ℕ) → (p q : typ ((Ω^ (2 + n)) B)) + → fst (ΩB→Ω^fibf (suc n)) (p ∙ q) + ≡ fst (ΩB→Ω^fibf (suc n)) p ∙ fst (ΩB→Ω^fibf (suc n)) q + ΩB→Ω^fibf-pres∙ n p q = + cong (fst (fst (Ω^Fibre≃∙⁻ (suc n) f))) + refl + ∙ cong (fst (fst (Ω≃∙ (Ω^Fibre≃∙⁻ n f)))) + (cong (fun (invIso (ΩFibreIso (Ω^→ n f)))) + (λ _ → snd ((Ω^ suc n) A) , Ω^→ (suc n) f .snd ∙ p ∙ q)) + ∙ cong (fst (fst (Ω≃∙ (Ω^Fibre≃∙⁻ n f)))) + (ΩFibreIso⁻pres∙snd (Ω^→ n f) p q) + ∙ Ω≃∙pres∙ (Ω^Fibre≃∙⁻ n f) + (inv (ΩFibreIso (Ω^→ n f)) (refl , Ω→ (Ω^→ n f) .snd ∙ p)) + (inv (ΩFibreIso (Ω^→ n f)) (refl , Ω→ (Ω^→ n f) .snd ∙ q)) + + {- Hence we have our sequence + ... → Ωⁿ⁺¹B → Ωⁿ(fib f) → Ωⁿ A → Ωⁿ B → ... (*) + We first prove the exactness properties for the helper functions + ΩB→fibΩ^f and fibΩ^f→A, and then deduce exactness of the whole sequence + by noting that the functions in (*) are just ΩB→fibΩ^f, fibΩ^f→A but + composed with equivalences + -} + private + Im-fibΩ^f→A⊂Ker-A→B : (n : ℕ) (x : _) + → isInIm∙ (fibΩ^f→A n) x + → isInKer∙ (A→B n) x + Im-fibΩ^f→A⊂Ker-A→B n x = + uncurry λ p → J (λ x _ → isInKer∙ (A→B n) x) + (snd p) + + Ker-fibΩ^f→A⊂Im-ΩB→fibΩ^f : (n : ℕ) (x : _) + → isInKer∙ (fibΩ^f→A n) x + → isInIm∙ (ΩB→fibΩ^f n) x + Ker-fibΩ^f→A⊂Im-ΩB→fibΩ^f n x ker = + (sym (Ω^→ n f .snd) + ∙ cong (Ω^→ n f .fst) (sym ker) ∙ snd x) , ΣPathP ((sym ker) , + ((∙assoc (Ω^→ n f .snd) + (sym (Ω^→ n f .snd)) + (sym (cong (Ω^→ n f .fst) ker) ∙ snd x) + ∙∙ cong (_∙ (sym (cong (Ω^→ n f .fst) ker) ∙ snd x)) + (rCancel (Ω^→ n f .snd)) + ∙∙ sym (lUnit (sym (cong (Ω^→ n f .fst) ker) ∙ snd x))) + ◁ (λ i j → compPath-filler' + (cong (Ω^→ n f .fst) (sym ker)) (snd x) (~ i) j))) + + Im-A→B⊂Ker-ΩB→fibΩ^f : (n : ℕ) (x : fst (((Ω^ (suc n)) B))) + → isInIm∙ (A→B (suc n)) x + → isInKer∙ (ΩB→fibΩ^f n) x + Im-A→B⊂Ker-ΩB→fibΩ^f n x = + uncurry λ p + → J (λ x _ → isInKer∙ (ΩB→fibΩ^f n) x) + (ΣPathP (p , (((λ i → (λ j → Ω^→ n f .snd (j ∧ ~ i)) + ∙ ((λ j → Ω^→ n f .snd (~ j ∧ ~ i)) + ∙∙ cong (Ω^→ n f .fst) p + ∙∙ Ω^→ n f .snd)) + ∙ sym (lUnit (cong (Ω^→ n f .fst) p ∙ Ω^→ n f .snd))) + ◁ λ i j → compPath-filler' + (cong (Ω^→ n f .fst) p) (Ω^→ n f .snd) (~ i) j))) + + Ker-ΩB→fibΩ^f⊂Im-A→B : (n : ℕ) (x : fst (((Ω^ (suc n)) B))) + → isInKer∙ (ΩB→fibΩ^f n) x + → isInIm∙ (A→B (suc n)) x + fst (Ker-ΩB→fibΩ^f⊂Im-A→B n x inker) = cong fst inker + snd (Ker-ΩB→fibΩ^f⊂Im-A→B n x inker) = lem + where + lem : fst (A→B (suc n)) (λ i → fst (inker i)) ≡ x + lem i j = + hcomp (λ k → λ { (i = i0) → doubleCompPath-filler + (sym (snd (Ω^→ n f))) + ((λ i → Ω^→ n f .fst (fst (inker i)))) + (snd (Ω^→ n f)) k j + ; (i = i1) → compPath-filler' (Ω^→ n f .snd) x (~ k) j + ; (j = i0) → snd (Ω^→ n f) k + ; (j = i1) → snd (Ω^→ n f) (k ∨ i)}) + (hcomp (λ k → λ { (i = i0) → (snd (inker j)) (~ k) + ; (i = i1) → ((Ω^→ n f .snd) ∙ x) (j ∨ ~ k) + ; (j = i0) → ((Ω^→ n f .snd) ∙ x) (~ k) + ; (j = i1) → snd (Ω^→ n f) (i ∨ ~ k)}) + (snd ((Ω^ n) B))) + + {- Finally, we get exactness of the sequence + we are interested in -} + Im-Ω^fibf→A⊂Ker-A→B : (n : ℕ) (x : _) + → isInIm∙ (Ω^fibf→A n) x + → isInKer∙ (A→B n) x + Im-Ω^fibf→A⊂Ker-A→B n x x₁ = + Im-fibΩ^f→A⊂Ker-A→B n x + (((fst (fst (Ω^Fibre≃∙ n f))) (fst x₁)) + , snd x₁) + + Ker-A→B⊂Im-Ω^fibf→A : (n : ℕ) (x : _) + → isInKer∙ (A→B n) x + → isInIm∙ (Ω^fibf→A n) x + Ker-A→B⊂Im-Ω^fibf→A n x ker = + invEq (fst (Ω^Fibre≃∙ n f)) (x , ker) + , (cong fst (secEq (fst (Ω^Fibre≃∙ n f)) (x , ker))) + + Ker-Ω^fibf→A⊂Im-ΩB→Ω^fibf : (n : ℕ) (x : _) + → isInKer∙ (Ω^fibf→A n) x + → isInIm∙ (ΩB→Ω^fibf n) x + Ker-Ω^fibf→A⊂Im-ΩB→Ω^fibf n x p = + fst r + , cong (fst ((fst (Ω^Fibre≃∙⁻ n f)))) (snd r) + ∙ funExt⁻ (cong fst (Ω^Fibre≃∙sect n f)) x + where + r : isInIm∙ (ΩB→fibΩ^f n) (fst (fst (Ω^Fibre≃∙ n f)) x) + r = Ker-fibΩ^f→A⊂Im-ΩB→fibΩ^f n (fst (fst (Ω^Fibre≃∙ n f)) x) p + + Im-ΩB→Ω^fibf⊂Ker-Ω^fibf→A : (n : ℕ) (x : _) + → isInIm∙ (ΩB→Ω^fibf n) x + → isInKer∙ (Ω^fibf→A n) x + Im-ΩB→Ω^fibf⊂Ker-Ω^fibf→A n x = + uncurry λ p → + J (λ x _ → isInKer∙ (Ω^fibf→A n) x) + (cong (fst (fibΩ^f→A n)) + (funExt⁻ (cong fst (Ω^Fibre≃∙retr n f)) _)) + + Im-A→B⊂Ker-ΩB→Ω^fibf : (n : ℕ) (x : fst (((Ω^ (suc n)) B))) + → isInIm∙ (A→B (suc n)) x + → isInKer∙ (ΩB→Ω^fibf n) x + Im-A→B⊂Ker-ΩB→Ω^fibf n x p = + cong (fst ((fst (Ω^Fibre≃∙⁻ n f)))) + (Im-A→B⊂Ker-ΩB→fibΩ^f n x p) + ∙ snd (Ω^Fibre≃∙⁻ n f) + + Ker-ΩB→Ω^fibf⊂Im-A→B : (n : ℕ) (x : fst (((Ω^ (suc n)) B))) + → isInKer∙ (ΩB→Ω^fibf n) x + → isInIm∙ (A→B (suc n)) x + Ker-ΩB→Ω^fibf⊂Im-A→B n x p = + Ker-ΩB→fibΩ^f⊂Im-A→B n x + (funExt⁻ (cong fst (sym (Ω^Fibre≃∙retr n f))) (ΩB→fibΩ^f n .fst x) + ∙ cong (fst (fst (Ω^Fibre≃∙ n f))) p + ∙ snd (Ω^Fibre≃∙ n f)) + +{- Some useful lemmas for converting the above sequence a +a sequence of homotopy groups -} +module setTruncLemmas {ℓ ℓ' ℓ'' : Level} + {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (n m l : ℕ) + (f : (Ω ((Ω^ n) A)) →∙ (Ω ((Ω^ m) B))) + (g : (Ω ((Ω^ m) B)) →∙ (Ω ((Ω^ l) C))) + (e₁ : IsGroupHom (snd (πGr n A)) (sMap (fst f)) (snd (πGr m B))) + (e₂ : IsGroupHom (snd (πGr m B)) (sMap (fst g)) (snd (πGr l C))) where + + ker⊂im : ((x : typ (Ω ((Ω^ m) B))) → isInKer∙ g x → isInIm∙ f x) + → (x : π (suc m) B) → isInKer (_ , e₂) x → isInIm (_ , e₁) x + ker⊂im ind = + sElim (λ _ → isSetΠ λ _ → isProp→isSet squash) + λ p ker → + pRec squash + (λ ker∙ → ∣ ∣ ind p ker∙ .fst ∣₂ , cong ∣_∣₂ (ind p ker∙ .snd) ∣ ) + (fun PathIdTrunc₀Iso ker) + + im⊂ker : ((x : typ (Ω ((Ω^ m) B))) → isInIm∙ f x → isInKer∙ g x) + → (x : π (suc m) B) → isInIm (_ , e₁) x → isInKer (_ , e₂) x + im⊂ker ind = + sElim (λ _ → isSetΠ λ _ → isSetPathImplicit) + λ p → + pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ λ _ → isSetPathImplicit) + λ a q → pRec (squash₂ _ _) + (λ q → cong ∣_∣₂ (ind p (a , q))) + (fun PathIdTrunc₀Iso q))) + +{- The long exact sequence of homotopy groups -} +module πLES {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where + module Ωs = ΩLES f + open Ωs renaming (A→B to A→B') + + fib = fibf + + fib→A : (n : ℕ) → GroupHom (πGr n fib) (πGr n A) + fst (fib→A n) = sMap (fst (Ω^fibf→A (suc n))) + snd (fib→A n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (Ω^fibf→A-pres∙ n p q)) + + A→B : (n : ℕ) → GroupHom (πGr n A) (πGr n B) + fst (A→B n) = sMap (fst (A→B' (suc n))) + snd (A→B n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ g h → cong ∣_∣₂ (Ω^→pres∙ f n g h)) + + B→fib : (n : ℕ) → GroupHom (πGr (suc n) B) (πGr n fib) + fst (B→fib n) = sMap (fst (ΩB→Ω^fibf (suc n))) + snd (B→fib n) = + makeIsGroupHom + (sElim2 + (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (ΩB→Ω^fibf-pres∙ n p q)) + + Ker-A→B⊂Im-fib→A : (n : ℕ) (x : π (suc n) A) + → isInKer (A→B n) x + → isInIm (fib→A n) x + Ker-A→B⊂Im-fib→A n = + setTruncLemmas.ker⊂im n n n + (Ω^fibf→A (suc n)) (A→B' (suc n)) + (snd (fib→A n)) (snd (A→B n)) + (Ker-A→B⊂Im-Ω^fibf→A (suc n)) + + Im-fib→A⊂Ker-A→B : (n : ℕ) (x : π (suc n) A) + → isInIm (fib→A n) x + → isInKer (A→B n) x + Im-fib→A⊂Ker-A→B n = + setTruncLemmas.im⊂ker n n n + (Ω^fibf→A (suc n)) (A→B' (suc n)) + (snd (fib→A n)) (snd (A→B n)) + (Im-Ω^fibf→A⊂Ker-A→B (suc n)) + + Ker-fib→A⊂Im-B→fib : (n : ℕ) (x : π (suc n) fib) + → isInKer (fib→A n) x + → isInIm (B→fib n) x + Ker-fib→A⊂Im-B→fib n = + setTruncLemmas.ker⊂im (suc n) n n + (ΩB→Ω^fibf (suc n)) (Ω^fibf→A (suc n)) + (snd (B→fib n)) (snd (fib→A n)) + (Ker-Ω^fibf→A⊂Im-ΩB→Ω^fibf (suc n)) + + Im-B→fib⊂Ker-fib→A : (n : ℕ) (x : π (suc n) fib) + → isInIm (B→fib n) x + → isInKer (fib→A n) x + Im-B→fib⊂Ker-fib→A n = + setTruncLemmas.im⊂ker (suc n) n n + (ΩB→Ω^fibf (suc n)) (Ω^fibf→A (suc n)) + (snd (B→fib n)) (snd (fib→A n)) + (Im-ΩB→Ω^fibf⊂Ker-Ω^fibf→A (suc n)) + + Im-A→B⊂Ker-B→fib : (n : ℕ) (x : π (suc (suc n)) B) + → isInIm (A→B (suc n)) x + → isInKer (B→fib n) x + Im-A→B⊂Ker-B→fib n = + setTruncLemmas.im⊂ker (suc n) (suc n) n + (A→B' (suc (suc n))) (ΩB→Ω^fibf (suc n)) + (snd (A→B (suc n))) (snd (B→fib n)) + (Im-A→B⊂Ker-ΩB→Ω^fibf (suc n)) + + Ker-B→fib⊂Im-A→B : (n : ℕ) (x : π (suc (suc n)) B) + → isInKer (B→fib n) x + → isInIm (A→B (suc n)) x + Ker-B→fib⊂Im-A→B n = + setTruncLemmas.ker⊂im (suc n) (suc n) n + (A→B' (suc (suc n))) (ΩB→Ω^fibf (suc n)) + (snd (A→B (suc n))) (snd (B→fib n)) + (Ker-ΩB→Ω^fibf⊂Im-A→B (suc n)) + +{- We prove that the map Ωⁿ(fib f) → Ωⁿ A indeed is just the map +Ωⁿ fst -} +private + Ω^fibf→A-ind : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → ΩLES.Ω^fibf→A f (suc n) ≡ Ω→ (ΩLES.Ω^fibf→A f n) + Ω^fibf→A-ind {A = A} {B = B} n f = + (λ _ → πLES.Ωs.fibΩ^f→A f (suc n) + ∘∙ ≃∙map (Ω^Fibre≃∙ (suc n) f)) + ∙ →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ p → + (λ j → cong fst (Ω→ (≃∙map (Ω^Fibre≃∙ n f)) .fst p)) + ∙ rUnit ((λ i → fst + (Ω→ (≃∙map (Ω^Fibre≃∙ n f)) .fst p i))) + ∙ sym (Ω→∘ (πLES.Ωs.fibΩ^f→A f n) (≃∙map (Ω^Fibre≃∙ n f)) p)) + +Ω^fibf→A≡ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → ΩLES.Ω^fibf→A f n ≡ Ω^→ n (fst , refl) +Ω^fibf→A≡ zero f = ΣPathP (refl , (sym (lUnit refl))) +Ω^fibf→A≡ (suc n) f = Ω^fibf→A-ind n f ∙ cong Ω→ (Ω^fibf→A≡ n f) + +{- We now get a nice characterisation of the functions in the induced LES +of homotopy groups defined using (Sⁿ →∙ A) -} + +π∘∙A→B-PathP : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → PathP (λ i → GroupHomπ≅π'PathP A B n i) + (πLES.A→B f n) + (π'∘∙Hom n f) +π∘∙A→B-PathP n f = + toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) (π'∘∙Hom'≡π'∘∙fun n f)) + +π∘∙fib→A-PathP : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → PathP (λ i → GroupHomπ≅π'PathP (ΩLES.fibf f) A n i) + (πLES.fib→A f n) + (π'∘∙Hom n (fst , refl)) +π∘∙fib→A-PathP {A = A} {B = B} n f = + toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (cong (transport + (λ i → (fst (GroupPath _ _) + (GroupIso→GroupEquiv (π'Gr≅πGr n (ΩLES.fibf f))) (~ i)) .fst + → (fst (GroupPath _ _) + (GroupIso→GroupEquiv (π'Gr≅πGr n A)) (~ i)) .fst)) + lem + ∙ π'∘∙Hom'≡π'∘∙fun n (fst , refl))) + where + lem : πLES.fib→A f n .fst ≡ sMap (Ω^→ (suc n) (fst , refl) .fst) + lem = cong sMap (cong fst (Ω^fibf→A≡ (suc n) f)) diff --git a/Cubical/Homotopy/Group/Pi3S2.agda b/Cubical/Homotopy/Group/Pi3S2.agda new file mode 100644 index 0000000000..c921b24aeb --- /dev/null +++ b/Cubical/Homotopy/Group/Pi3S2.agda @@ -0,0 +1,154 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains: +1. The iso π₃S²≅ℤ +2. A proof that π₃S² is generated by the Hopf map +-} +module Cubical.Homotopy.Group.Pi3S2 where + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Homotopy.Group.LES +open import Cubical.Homotopy.Group.PinSn +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.HopfInvariant.HopfMap + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open Iso +open import Cubical.Foundations.Equiv + +open import Cubical.HITs.SetTruncation renaming (elim to sElim) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact + +TotalHopf→∙S² : (Σ (S₊ 2) S¹Hopf , north , base) →∙ S₊∙ 2 +fst TotalHopf→∙S² = fst +snd TotalHopf→∙S² = refl + +IsoTotalSpaceJoin' : Iso (Σ (S₊ 2) S¹Hopf) (S₊ 3) +IsoTotalSpaceJoin' = compIso hopfS¹.IsoTotalSpaceJoin (IsoSphereJoin 1 1) + +IsoFiberTotalHopfS¹ : Iso (fiber (fst TotalHopf→∙S²) north) S¹ +fun IsoFiberTotalHopfS¹ ((x , y) , z) = subst S¹Hopf z y +inv IsoFiberTotalHopfS¹ x = (north , x) , refl +rightInv IsoFiberTotalHopfS¹ x = refl +leftInv IsoFiberTotalHopfS¹ ((x , y) , z) = + ΣPathP + ((ΣPathP + (sym z , (λ i → transp (λ j → S¹Hopf (z (~ i ∧ j))) i y))) + , (λ j i → z (i ∨ ~ j))) + +IsoFiberTotalHopfS¹∙≡ : + (fiber (fst TotalHopf→∙S²) north , (north , base) , refl) ≡ S₊∙ 1 +IsoFiberTotalHopfS¹∙≡ = ua∙ (isoToEquiv IsoFiberTotalHopfS¹) refl + +private + transportGroupEquiv : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → isEquiv (fst (πLES.A→B f n)) + → isEquiv (fst (π'∘∙Hom n f)) + transportGroupEquiv n f iseq = + transport (λ i → isEquiv (fst (π∘∙A→B-PathP n f i))) iseq + +π₃S²≅π₃TotalHopf : GroupEquiv (πGr 2 (Σ (S₊ 2) S¹Hopf , north , base)) + (πGr 2 (S₊∙ 2)) +fst (fst π₃S²≅π₃TotalHopf) = fst (πLES.A→B TotalHopf→∙S² 2) +snd (fst π₃S²≅π₃TotalHopf) = + SES→isEquiv + (isContr→≡UnitGroup + (subst isContr (cong (π 3) (sym IsoFiberTotalHopfS¹∙≡)) + (∣ refl ∣₂ , (sElim (λ _ → isSetPathImplicit) + (λ p → cong ∣_∣₂ + (isOfHLevelSuc 3 isGroupoidS¹ _ _ _ _ _ _ refl p)))))) + (isContr→≡UnitGroup + (subst isContr (cong (π 2) (sym IsoFiberTotalHopfS¹∙≡)) + (∣ refl ∣₂ , (sElim (λ _ → isSetPathImplicit) (λ p + → cong ∣_∣₂ (isGroupoidS¹ _ _ _ _ refl p)))))) + (πLES.fib→A TotalHopf→∙S² 2) + (πLES.A→B TotalHopf→∙S² 2) + (πLES.B→fib TotalHopf→∙S² 1) + (πLES.Ker-A→B⊂Im-fib→A TotalHopf→∙S² 2) + (πLES.Ker-B→fib⊂Im-A→B TotalHopf→∙S² 1) +snd π₃S²≅π₃TotalHopf = snd (πLES.A→B TotalHopf→∙S² 2) + +π'₃S²≅π'₃TotalHopf : GroupEquiv (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base)) + (π'Gr 2 (S₊∙ 2)) +fst (fst π'₃S²≅π'₃TotalHopf) = fst (π'∘∙Hom 2 TotalHopf→∙S²) +snd (fst π'₃S²≅π'₃TotalHopf) = + transportGroupEquiv 2 TotalHopf→∙S² (π₃S²≅π₃TotalHopf .fst .snd) +snd π'₃S²≅π'₃TotalHopf = snd (π'∘∙Hom 2 TotalHopf→∙S²) + +πS³≅πTotalHopf : (n : ℕ) + → GroupEquiv (π'Gr n (S₊∙ 3)) (π'Gr n (Σ (S₊ 2) S¹Hopf , north , base)) +πS³≅πTotalHopf n = + π'Iso n + ((isoToEquiv + (invIso + (compIso + (hopfS¹.IsoTotalSpaceJoin) + (IsoSphereJoin 1 1)))) + , refl) + +π₃S²≅ℤ : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤ +π₃S²≅ℤ = + compGroupEquiv + (invGroupEquiv + (compGroupEquiv (πS³≅πTotalHopf 2) π'₃S²≅π'₃TotalHopf)) + (GroupIso→GroupEquiv (πₙ'Sⁿ≅ℤ 2)) + +-- We prove that the generator is the Hopf map +π₃TotalHopf-gen' : π' 3 (Σ (Susp S¹) S¹Hopf , north , base) +π₃TotalHopf-gen' = + ∣ inv (compIso (hopfS¹.IsoTotalSpaceJoin) (IsoSphereJoin 1 1)) , refl ∣₂ + +πS³≅πTotalHopf-gen : + fst (fst (πS³≅πTotalHopf 2)) ∣ idfun∙ _ ∣₂ ≡ π₃TotalHopf-gen' +πS³≅πTotalHopf-gen = + cong ∣_∣₂ + (∘∙-idʳ (inv (compIso (hopfS¹.IsoTotalSpaceJoin) + (IsoSphereJoin 1 1)) , refl)) + +πTotalHopf-gen : + gen₁-by (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base)) + ∣ inv (compIso (hopfS¹.IsoTotalSpaceJoin) (IsoSphereJoin 1 1)) , refl ∣₂ +πTotalHopf-gen = + subst (gen₁-by (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base))) + πS³≅πTotalHopf-gen + (Iso-pres-gen₁ (π'Gr 2 (S₊∙ 3)) + (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base)) + ∣ idfun∙ _ ∣₂ + (πₙ'Sⁿ-gen-by-idfun 2) + (GroupEquiv→GroupIso (πS³≅πTotalHopf 2))) + +πTotalHopf≅πS²-gen : + fst (fst π'₃S²≅π'₃TotalHopf) + π₃TotalHopf-gen' + ≡ ∣ HopfMap' , refl ∣₂ +πTotalHopf≅πS²-gen = + cong ∣_∣₂ (ΣPathP (refl , (sym (rUnit refl)))) + +π₂S³-gen-by-HopfMap' : gen₁-by (π'Gr 2 (S₊∙ 2)) ∣ HopfMap' , refl ∣₂ +π₂S³-gen-by-HopfMap' = + subst (gen₁-by (π'Gr 2 (S₊∙ 2))) πTotalHopf≅πS²-gen + (Iso-pres-gen₁ (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base)) (π'Gr 2 (S₊∙ 2)) + ∣ inv (compIso (hopfS¹.IsoTotalSpaceJoin) (IsoSphereJoin 1 1)) , refl ∣₂ + πTotalHopf-gen + (GroupEquiv→GroupIso π'₃S²≅π'₃TotalHopf)) + +π₂S³-gen-by-HopfMap : gen₁-by (π'Gr 2 (S₊∙ 2)) ∣ HopfMap ∣₂ +π₂S³-gen-by-HopfMap = + subst (gen₁-by (π'Gr 2 (S₊∙ 2))) + (cong ∣_∣₂ (sym hopfMap≡HopfMap')) + π₂S³-gen-by-HopfMap' diff --git a/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso.agda b/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso.agda new file mode 100644 index 0000000000..fca96aa611 --- /dev/null +++ b/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso.agda @@ -0,0 +1,682 @@ +{-# OPTIONS --safe #-} +{- +This file has been split in two due to slow type checking +combined with insufficient reductions when the +experimental-lossy-unification flag is included. +Part 2: Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 + +The goal of these two files is to show that +π₄(S³) ≅ π₃((S² × S²) ⊔ᴬ S²) where A = S² ∨ S². +This is proved in Brunerie (2016) using the James construction and +via (S² × S²) ⊔ᴬ S² ≃ J₂(S²). + +In this file, we prove it directly using the encode-decode method. For +the statement of the isomorphism, see part 2. +-} +module Cubical.Homotopy.Group.Pi4S3.S3PushoutIso where + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Homotopy.Group.Base +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open Iso +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Morphism + +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 + ; map to sMap) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.S1 hiding (decode ; encode) + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Bool + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + + +open import Cubical.HITs.Join +open import Cubical.HITs.Pushout +open import Cubical.HITs.Wedge +open import Cubical.Homotopy.Freudenthal hiding (Code ; encode) +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation renaming + (rec to trRec ; elim to trElim ; elim2 to trElim2 ; map to trMap) +open import Cubical.Foundations.Function +open import Cubical.HITs.S2 + +Pushout⋁↪fold⋁ : ∀ {ℓ} (A : Pointed ℓ) → Type ℓ +Pushout⋁↪fold⋁ A = Pushout {A = A ⋁ A} ⋁↪ fold⋁ + +Pushout⋁↪fold⋁∙ : ∀ {ℓ} (A : Pointed ℓ) → Pointed ℓ +fst (Pushout⋁↪fold⋁∙ A) = Pushout⋁↪fold⋁ A +snd (Pushout⋁↪fold⋁∙ A) = inl (pt A , pt A) + +-- The type of interest -- ''almost`` equivalent to ΩS³ +Pushout⋁↪fold⋁S₊2 = Pushout {A = S₊∙ 2 ⋁ S₊∙ 2} ⋁↪ fold⋁ +-- Same type, using base/surf definition of S² (easier to work with) +Pushout⋁↪fold⋁S² = Pushout⋁↪fold⋁ S²∙ +-- Truncated version, to be proved equivalent to ∥ ΩS³ ∥₅ +∥Pushout⋁↪fold⋁S²∥₅ = hLevelTrunc 5 Pushout⋁↪fold⋁S² + +Ω∥S³∥₆ = Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ north ∣ + +-- Elimination into sets for Pushout⋁↪fold⋁S² +module _ {ℓ : Level} {P : Pushout⋁↪fold⋁S² → Type ℓ} + (hlev : ((x : Pushout⋁↪fold⋁S²) → isSet (P x))) + (b : P (inl (base , base))) where + private + ×fun : (x y : S²) → P (inl (x , y)) + ×fun = S²ToSetElim (λ _ → isSetΠ (λ _ → hlev _)) + (S²ToSetElim (λ _ → hlev _) b) + + inrfun : (x : S²) → P (inr x) + inrfun = S²ToSetElim (λ _ → hlev _) (subst P (push (inl base)) b) + + inl-pp : (x : S²) → PathP (λ i → P (push (inl x) i)) (×fun x base) (inrfun x) + inl-pp = S²ToSetElim (λ _ → isOfHLevelPathP 2 (hlev _) _ _) + λ j → transp (λ i → P (push (inl base) (i ∧ j))) (~ j) b + + inr-pp : (x : S²) → PathP (λ i → P (push (inr x) i)) (×fun base x) (inrfun x) + inr-pp = S²ToSetElim (λ _ → isOfHLevelPathP 2 (hlev _) _ _) + (transport (λ j → PathP (λ i → P (push (push tt j) i)) (×fun base base) + (inrfun base)) + (inl-pp base)) + + Pushout⋁↪fold⋁S²→Set : (x : Pushout⋁↪fold⋁S²) → P x + Pushout⋁↪fold⋁S²→Set (inl x) = ×fun (fst x) (snd x) + Pushout⋁↪fold⋁S²→Set (inr x) = inrfun x + Pushout⋁↪fold⋁S²→Set (push (inl x) i) = inl-pp x i + Pushout⋁↪fold⋁S²→Set (push (inr x) i) = inr-pp x i + Pushout⋁↪fold⋁S²→Set (push (push a j) i) = help j i + where + help : SquareP (λ j i → P (push (push tt j) i)) + (λ i → inl-pp base i) + (λ i → inr-pp base i) + (λ _ → ×fun base base) + (λ _ → inrfun base) + help = toPathP (isProp→PathP (λ _ → isOfHLevelPathP' 1 (hlev _) _ _) _ _) + +{- A wedge connectivity lemma for Pushout⋁↪fold⋁S². It can be +stated for general dependent types, but it's easier to work with +in the special case of path types -} +module Pushout⋁↪fold⋁S²WedgeCon {ℓ : Level } {A : Type ℓ} + (f g : Pushout⋁↪fold⋁S² → A) + (h : isOfHLevel 5 A) + (lp : (x : S²) → f (inl (x , base)) ≡ g (inl (x , base))) + (rp : (x : S²) → f (inl (base , x)) ≡ g (inl (base , x))) + (r≡l : (x : S²) + → (sym (cong f (push (inl x))) ∙∙ lp x ∙∙ cong g (push (inl x))) + ≡ (sym (cong f (push (inr x))) ∙∙ rp x ∙∙ cong g (push (inr x)))) + where + btm-filler : I → I → I → A + btm-filler k i j = + hfill (λ k → λ {(i = i0) + → doubleCompPath-filler (sym (cong f (push (inl base)))) + (lp base) + (cong g (push (inl base))) (~ k) j + ; (i = i1) + → doubleCompPath-filler (sym (cong f (push (inr base)))) + (rp base) + (cong g (push (inr base))) (~ k) j + ; (j = i0) → f (push (push tt i) (~ k)) + ; (j = i1) → g (push (push tt i) (~ k))}) + (inS (r≡l base i j)) + k + + lp-base≡rp-base : lp base ≡ rp base + lp-base≡rp-base = (λ i j → btm-filler i1 i j) + + f∘inl≡g∘inl : (x y : S²) → f (inl (x , y)) ≡ g (inl (x , y)) + f∘inl≡g∘inl = wedgeconFunS² (λ _ _ → h _ _) lp rp lp-base≡rp-base + + f∘inl≡g∘inl-base : (x : S²) → f∘inl≡g∘inl x base ≡ lp x + f∘inl≡g∘inl-base = wedgeconFunS²Id (λ _ _ → h _ _) lp rp lp-base≡rp-base + + f∘inr≡g∘inr : (x : S²) → f (inr x) ≡ g (inr x) + f∘inr≡g∘inr x = cong f (sym (push (inl x))) ∙∙ lp x ∙∙ cong g (push (inl x)) + + inlfill : (x : S²) → I → I → I → A + inlfill x k i j = + hfill (λ k → λ { (i = i0) → f∘inl≡g∘inl-base x (~ k) j + ; (i = i1) → doubleCompPath-filler + (cong f (sym (push (inl x)))) + (lp x) + (cong g (push (inl x))) k j + ; (j = i0) → f (push (inl x) (i ∧ k)) + ; (j = i1) → g (push (inl x) (i ∧ k))}) + (inS (lp x j)) + k + + inrfill : (x : S²) → I → I → I → A + inrfill x k i j = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler + (cong f (sym (push (inr x)))) + (rp x) + (cong g (push (inr x))) (~ k) j + ; (i = i1) → f∘inr≡g∘inr x j + ; (j = i0) → f (push (inr x) (i ∨ ~ k)) + ; (j = i1) → g (push (inr x) (i ∨ ~ k))}) + (inS (r≡l x (~ i) j)) + k + + main : (x : Pushout⋁↪fold⋁S²) → f x ≡ g x + main (inl x) = f∘inl≡g∘inl (fst x) (snd x) + main (inr x) = f∘inr≡g∘inr x + main (push (inl x) i) j = inlfill x i1 i j + main (push (inr x) i) j = inrfill x i1 i j + main (push (push a i) j) k = + hcomp (λ r → λ {(i = i0) → inlfill base i1 j k + ; (i = i1) → inrfill-side r j k + ; (j = i0) → rp base k + ; (j = i1) → r≡l base (~ r ∧ i) k + ; (k = i0) → f (push (push a i) j) + ; (k = i1) → g (push (push a i) j)}) + (hcomp (λ r → λ {(i = i0) → inlfill base r j k + ; (i = i1) → doubleCompPath-filler + (cong f (sym (push (inr base)))) + (rp base) + (cong g (push (inr base))) (j ∧ r) k + ; (j = i0) → lp-base≡rp-base (r ∨ i) k + ; (j = i1) → inlfill-side r i k + ; (k = i0) → f (push (push a i) (j ∧ r)) + ; (k = i1) → g (push (push a i) (j ∧ r))}) + (lp-base≡rp-base i k)) + where + inlfill-side : I → I → I → A + inlfill-side r i k = + hcomp (λ j → λ { (r = i0) → btm-filler j i k + ; (r = i1) → r≡l base i k + ; (i = i0) → doubleCompPath-filler + (cong f (sym (push (inl base)))) + (lp base) + (cong g (push (inl base))) (r ∨ ~ j) k + ; (i = i1) → doubleCompPath-filler + (cong f (sym (push (inr base)))) + (rp base) + (cong g (push (inr base))) (r ∨ ~ j) k + ; (k = i0) → f (push (push a i) (r ∨ ~ j)) + ; (k = i1) → g (push (push a i) (r ∨ ~ j))}) + (r≡l base i k) + + inrfill-side : I → I → I → A + inrfill-side r j k = + hcomp (λ i → λ { (r = i0) + → (doubleCompPath-filler + (cong f (sym (push (inr base)))) + (rp base) + (cong g (push (inr base)))) (j ∨ ~ i) k + ; (r = i1) → inrfill base i j k + ; (j = i0) → inrfill base i i0 k + ; (j = i1) → r≡l base (~ r) k + ; (k = i0) → f (push (inr base) (j ∨ ~ i)) + ; (k = i1) → g (push (inr base) (j ∨ ~ i))}) + (r≡l base (~ r ∨ ~ j) k) + +{- +Goal: Prove Ω ∥ SuspS² ∥₆ ≃ ∥ Pushout⋁↪fold⋁S² ∥₅ +This will by done via the encode-decode method. For this, we nede a family +of equivalences ∥ Pushout⋁↪fold⋁S² ∥₅ ≃ ∥ Pushout⋁↪fold⋁S² ∥₅, indexed by S². +In order to define this, we need the following cubes/coherences in +∥ Pushout⋁↪fold⋁S² ∥₅: +-} + +→Ω²∥Pushout⋁↪fold⋁S²∥₅ : (x y : S²) + → Square {A = ∥Pushout⋁↪fold⋁S²∥₅} (λ _ → ∣ inl (x , y) ∣ₕ) + (λ _ → ∣ inl (x , y) ∣ₕ) + (λ _ → ∣ inl (x , y) ∣ₕ) + (λ _ → ∣ inl (x , y) ∣ₕ) +→Ω²∥Pushout⋁↪fold⋁S²∥₅ = + wedgeconFunS² + (λ _ _ → isOfHLevelPath 4 (isOfHLevelTrunc 5 _ _) _ _) + (λ x → λ i j → ∣ inl (x , surf i j) ∣ₕ) + (λ x → λ i j → ∣ inl (surf i j , x) ∣ₕ) + λ r i j → ∣ hcomp (λ k → λ { (i = i0) → push (push tt (~ r)) (~ k) + ; (i = i1) → push (push tt (~ r)) (~ k) + ; (j = i1) → push (push tt (~ r)) (~ k) + ; (j = i0) → push (push tt (~ r)) (~ k) + ; (r = i0) → push (inr (surf i j)) (~ k) + ; (r = i1) → push (inl (surf i j)) (~ k)}) + (inr (surf i j)) ∣ₕ + +→Ω²∥Pushout⋁↪fold⋁S²∥₅Id : (x : S²) + → →Ω²∥Pushout⋁↪fold⋁S²∥₅ x base ≡ λ i j → ∣ inl (x , surf i j) ∣ₕ +→Ω²∥Pushout⋁↪fold⋁S²∥₅Id = + wedgeconFunS²Id + (λ _ _ → isOfHLevelPath 4 (isOfHLevelTrunc 5 _ _) _ _) + (λ x → λ i j → ∣ inl (x , surf i j) ∣ₕ) + (λ x → λ i j → ∣ inl (surf i j , x) ∣ₕ) + λ r i j → ∣ hcomp (λ k → λ { (i = i0) → push (push tt (~ r)) (~ k) + ; (i = i1) → push (push tt (~ r)) (~ k) + ; (j = i1) → push (push tt (~ r)) (~ k) + ; (j = i0) → push (push tt (~ r)) (~ k) + ; (r = i0) → push (inr (surf i j)) (~ k) + ; (r = i1) → push (inl (surf i j)) (~ k)}) + (inr (surf i j)) ∣ₕ + +push-inl∙push-inr⁻ : (y : S²) → Path Pushout⋁↪fold⋁S² (inl (y , base)) (inl (base , y)) +push-inl∙push-inr⁻ y i = (push (inl y) ∙ sym (push (inr y))) i + +push-inl∙push-inr⁻∙ : push-inl∙push-inr⁻ base ≡ refl +push-inl∙push-inr⁻∙ = + (λ k i → (push (inl base) ∙ sym (push (push tt (~ k)))) i) + ∙ rCancel (push (inl base)) + +push-inl∙push-inr⁻-filler : I → I → I → I → Pushout⋁↪fold⋁S² +push-inl∙push-inr⁻-filler r i j k = + hfill (λ r → λ { (i = i0) → push-inl∙push-inr⁻∙ (~ r) k + ; (i = i1) → push-inl∙push-inr⁻∙ (~ r) k + ; (j = i0) → push-inl∙push-inr⁻∙ (~ r) k + ; (j = i1) → push-inl∙push-inr⁻∙ (~ r) k + ; (k = i0) → inl (surf i j , base) + ; (k = i1) → inl (surf i j , base)}) + (inS (inl (surf i j , base))) + r + +push-inl∙push-inr⁻-hLevFiller : (y : S²) + → Cube {A = ∥Pushout⋁↪fold⋁S²∥₅} + (λ j k → ∣ push-inl∙push-inr⁻ y k ∣) + (λ j k → ∣ push-inl∙push-inr⁻ y k ∣) + (λ j k → ∣ push-inl∙push-inr⁻ y k ∣) + (λ j k → ∣ push-inl∙push-inr⁻ y k ∣) + (→Ω²∥Pushout⋁↪fold⋁S²∥₅ y (pt (S² , base))) + λ i j → ∣ inl (surf i j , y) ∣ +push-inl∙push-inr⁻-hLevFiller = + S²ToSetElim (λ _ → isOfHLevelPathP' 2 + (isOfHLevelPathP' 3 + (isOfHLevelPathP' 4 (isOfHLevelTrunc 5) _ _) _ _) _ _) + λ i j k → ∣ push-inl∙push-inr⁻-filler i1 i j k ∣ₕ + + +{- +We combine the above definitions. The equivalence +∥Pushout⋁↪fold⋁S²∥₅ ≃ ∥Pushout⋁↪fold⋁S²∥₅ (w.r.t. S²) is induced +by the following function +-} +S²→Pushout⋁↪fold⋁S²↺' : (x : S²) → Pushout⋁↪fold⋁S² → ∥Pushout⋁↪fold⋁S²∥₅ +S²→Pushout⋁↪fold⋁S²↺' base (inl (x , y)) = ∣ inl (x , y) ∣ +S²→Pushout⋁↪fold⋁S²↺' (surf i j) (inl (x , y)) = →Ω²∥Pushout⋁↪fold⋁S²∥₅ x y i j +S²→Pushout⋁↪fold⋁S²↺' base (inr z) = ∣ inl (base , z) ∣ₕ +S²→Pushout⋁↪fold⋁S²↺' (surf i i₁) (inr z) = ∣ inl (surf i i₁ , z) ∣ₕ +S²→Pushout⋁↪fold⋁S²↺' base (push (inl y) k) = + ∣ (push (inl y) ∙ sym (push (inr y))) k ∣ +S²→Pushout⋁↪fold⋁S²↺' (surf i j) (push (inl y) k) = + push-inl∙push-inr⁻-hLevFiller y i j k +S²→Pushout⋁↪fold⋁S²↺' base (push (inr y) i) = ∣ inl (base , y) ∣ +S²→Pushout⋁↪fold⋁S²↺' (surf i j) (push (inr y) k) = ∣ inl (surf i j , y) ∣ +S²→Pushout⋁↪fold⋁S²↺' base (push (push a i₁) i) = ∣ push-inl∙push-inr⁻∙ i₁ i ∣ +S²→Pushout⋁↪fold⋁S²↺' (surf i j) (push (push a k) l) = + ∣ hcomp (λ r → λ { (i = i0) → push-inl∙push-inr⁻∙ (k ∨ ~ r) l + ; (i = i1) → push-inl∙push-inr⁻∙ (k ∨ ~ r) l + ; (j = i0) → push-inl∙push-inr⁻∙ (k ∨ ~ r) l + ; (j = i1) → push-inl∙push-inr⁻∙ (k ∨ ~ r) l + ; (k = i0) → push-inl∙push-inr⁻-filler r i j l + ; (k = i1) → inl (surf i j , base) + ; (l = i0) → inl (surf i j , base) + ; (l = i1) → inl (surf i j , base)}) + (inl (surf i j , base)) ∣ₕ + +{- For easier treatment later, we state its inverse explicitly -} +S²→Pushout⋁↪fold⋁S²↺'⁻ : (x : S²) → Pushout⋁↪fold⋁S² → ∥Pushout⋁↪fold⋁S²∥₅ +S²→Pushout⋁↪fold⋁S²↺'⁻ base y = S²→Pushout⋁↪fold⋁S²↺' base y +S²→Pushout⋁↪fold⋁S²↺'⁻ (surf i j) y = S²→Pushout⋁↪fold⋁S²↺' (surf (~ i) j) y + + +S²→Pushout⋁↪fold⋁S²↺'≡idfun : (x : _) → S²→Pushout⋁↪fold⋁S²↺' base x ≡ ∣ x ∣ +S²→Pushout⋁↪fold⋁S²↺'≡idfun (inl x) = refl +S²→Pushout⋁↪fold⋁S²↺'≡idfun (inr x) = cong ∣_∣ₕ (push (inr x)) +S²→Pushout⋁↪fold⋁S²↺'≡idfun (push (inl x) i) j = + ∣ compPath-filler (push (inl x)) (λ i₁ → push (inr x) (~ i₁)) (~ j) i ∣ₕ +S²→Pushout⋁↪fold⋁S²↺'≡idfun (push (inr x) i) j = ∣ push (inr x) (j ∧ i) ∣ +S²→Pushout⋁↪fold⋁S²↺'≡idfun (push (push a i) j) k = + ∣ coh-lem {A = Pushout⋁↪fold⋁S²} + (push (inl base)) (push (inr base)) (λ k → push (push tt k)) i j k ∣ₕ + where + coh-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p q : x ≡ y) (r : p ≡ q) + → Cube {A = A} + (λ j k → compPath-filler p (sym q) (~ k) j) (λ k j → q (k ∧ j)) + (λ i k → x) (λ i k → q k) + (((λ k → p ∙ sym (r (~ k))) ∙ rCancel p)) + r + coh-lem {A = A} {x = x} {y = y} = + J (λ y p → (q : x ≡ y) (r : p ≡ q) + → Cube {A = A} + (λ j k → compPath-filler p (sym q) (~ k) j) (λ k j → q (k ∧ j)) + (λ i k → x) (λ i k → q k) + (((λ k → p ∙ sym (r (~ k))) ∙ rCancel p)) + r) + λ q → J (λ q r → Cube (λ j k → compPath-filler refl (λ i → q (~ i)) (~ k) j) + (λ k j → q (k ∧ j)) (λ i k → x) (λ i k → q k) + ((λ k → refl ∙ (λ i → r (~ k) (~ i))) ∙ rCancel refl) r) + ((λ z j k → lUnit (sym (rUnit (λ _ → x))) z k j) + ◁ (λ i j k → (refl ∙ (λ i₁ → rUnit (λ _ → x) (~ i₁))) (k ∨ i) j)) + +S²→Pushout⋁↪fold⋁S²↺ : (x : S²) → ∥Pushout⋁↪fold⋁S²∥₅ → ∥Pushout⋁↪fold⋁S²∥₅ +S²→Pushout⋁↪fold⋁S²↺ x = trRec (isOfHLevelTrunc 5) (S²→Pushout⋁↪fold⋁S²↺' x) + +isEq-S²→Pushout⋁↪fold⋁S²↺' : (x : _) → isEquiv (S²→Pushout⋁↪fold⋁S²↺ x) +isEq-S²→Pushout⋁↪fold⋁S²↺' = + S²ToSetElim (λ _ → isProp→isSet (isPropIsEquiv _)) + (subst isEquiv (funExt id≡) (idIsEquiv _)) + where + id≡ : (x : _) → x ≡ S²→Pushout⋁↪fold⋁S²↺ base x + id≡ = trElim (λ _ → isOfHLevelPath 5 (isOfHLevelTrunc 5) _ _) + (sym ∘ S²→Pushout⋁↪fold⋁S²↺'≡idfun) + +S²→Pushout⋁↪fold⋁S²↺⁻¹ : (x : S²) → ∥Pushout⋁↪fold⋁S²∥₅ → ∥Pushout⋁↪fold⋁S²∥₅ +S²→Pushout⋁↪fold⋁S²↺⁻¹ x = trRec (isOfHLevelTrunc 5) (S²→Pushout⋁↪fold⋁S²↺'⁻ x) + +{- S²→Pushout⋁↪fold⋁S²↺⁻¹ is a section of S²→Pushout⋁↪fold⋁S²↺ -} +secS²→Pushout⋁↪fold⋁S²↺ : (x : S²) (y : ∥Pushout⋁↪fold⋁S²∥₅) + → S²→Pushout⋁↪fold⋁S²↺ x (S²→Pushout⋁↪fold⋁S²↺⁻¹ x y) ≡ y +secS²→Pushout⋁↪fold⋁S²↺ x = + trElim (λ _ → isOfHLevelPath 5 (isOfHLevelTrunc 5) _ _) + (h x) + where + h : (x : S²) (a : Pushout⋁↪fold⋁S²) + → S²→Pushout⋁↪fold⋁S²↺ x + (S²→Pushout⋁↪fold⋁S²↺⁻¹ x ∣ a ∣) ≡ ∣ a ∣ + h base a = cong (S²→Pushout⋁↪fold⋁S²↺ base) + (S²→Pushout⋁↪fold⋁S²↺'≡idfun a) + ∙ S²→Pushout⋁↪fold⋁S²↺'≡idfun a + h (surf i j) a k = main a k i j + where + side : (a : Pushout⋁↪fold⋁S²) → _ + side a = cong (S²→Pushout⋁↪fold⋁S²↺ base) (S²→Pushout⋁↪fold⋁S²↺'≡idfun a) + ∙ S²→Pushout⋁↪fold⋁S²↺'≡idfun a + + refl-b : Path ∥Pushout⋁↪fold⋁S²∥₅ _ _ + refl-b = (refl {x = ∣ inl (base , base) ∣ₕ}) + + main : (a : Pushout⋁↪fold⋁S²) + → Cube (λ i j → S²→Pushout⋁↪fold⋁S²↺ (surf i j) + (S²→Pushout⋁↪fold⋁S²↺⁻¹ (surf i j) ∣ a ∣)) + (λ _ _ → ∣ a ∣) + (λ i _ → side a i) (λ i _ → side a i) + (λ i _ → side a i) (λ i _ → side a i) + main = + Pushout⋁↪fold⋁S²→Set (λ _ → isOfHLevelPathP' 2 + (isOfHLevelPathP' 3 (isOfHLevelTrunc 5 _ _) _ _) _ _) + λ k i j → + hcomp (λ r → λ { (i = i0) → rUnit refl-b r k + ; (i = i1) → rUnit refl-b r k + ; (j = i0) → rUnit refl-b r k + ; (j = i1) → rUnit refl-b r k + ; (k = i0) → →Ω²∥Pushout⋁↪fold⋁S²∥₅Id + (surf (~ i) j) (~ r) i j + ; (k = i1) → μ base base}) + (μ-coh k i j) + where + μ : (x y : S²) → ∥Pushout⋁↪fold⋁S²∥₅ + μ x y = ∣ inl (x , y) ∣ₕ + + μUnit : (x : S²) → μ x base ≡ μ base x + μUnit base = refl + μUnit (surf i j) k = ∣ + hcomp (λ r → λ {(i = i0) → push (push tt k) (~ r) + ; (i = i1) → push (push tt k) (~ r) + ; (j = i0) → push (push tt k) (~ r) + ; (j = i1) → push (push tt k) (~ r) + ; (k = i0) → push (inl (surf i j)) (~ r) + ; (k = i1) → push (inr (surf i j)) (~ r)}) + (inr (surf i j)) ∣ₕ + + μ-coh : Path (Square {A = ∥Pushout⋁↪fold⋁S²∥₅} + (λ _ → ∣ inl (base , base) ∣) (λ _ → ∣ inl (base , base) ∣) + (λ _ → ∣ inl (base , base) ∣) (λ _ → ∣ inl (base , base) ∣)) + (λ i j → ∣ inl (surf (~ i) j , surf i j) ∣ₕ) + refl + μ-coh = + (cong₂Funct (λ (x y : (Path S² base base)) → cong₂ μ x y) (sym surf) surf + ∙ cong (_∙ cong (cong (μ base)) surf) (λ i → cong (cong (λ x → μUnit x i)) (sym surf)) + ∙ lCancel (cong (cong (μ base)) surf)) + +retrS²→Pushout⋁↪fold⋁S²↺ : (x : S²) (y : ∥Pushout⋁↪fold⋁S²∥₅) + → S²→Pushout⋁↪fold⋁S²↺⁻¹ x (S²→Pushout⋁↪fold⋁S²↺ x y) ≡ y +retrS²→Pushout⋁↪fold⋁S²↺ x = + trElim (λ _ → isOfHLevelPath 5 (isOfHLevelTrunc 5) _ _) + (main x) + where + main : (x : S²) (a : Pushout⋁↪fold⋁S²) + → S²→Pushout⋁↪fold⋁S²↺⁻¹ x (S²→Pushout⋁↪fold⋁S²↺ x ∣ a ∣) ≡ ∣ a ∣ + main base a = secS²→Pushout⋁↪fold⋁S²↺ base ∣ a ∣ + main (surf i j) a l = secS²→Pushout⋁↪fold⋁S²↺ (surf (~ i) j) ∣ a ∣ l + +∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ : (x : S²) + → Iso ∥Pushout⋁↪fold⋁S²∥₅ ∥Pushout⋁↪fold⋁S²∥₅ +fun (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ x) = S²→Pushout⋁↪fold⋁S²↺ x +inv (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ x) = S²→Pushout⋁↪fold⋁S²↺⁻¹ x +rightInv (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ x) = secS²→Pushout⋁↪fold⋁S²↺ x +leftInv (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ x) = retrS²→Pushout⋁↪fold⋁S²↺ x + +PreCode : (x : Susp S²) → Type +PreCode north = ∥Pushout⋁↪fold⋁S²∥₅ +PreCode south = ∥Pushout⋁↪fold⋁S²∥₅ +PreCode (merid a i) = isoToPath (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ a) i + +hLevPreCode : (x : Susp S²) → isOfHLevel 5 (PreCode x) +hLevPreCode = + suspToPropElim base (λ _ → isPropIsOfHLevel 5) (isOfHLevelTrunc 5) + +Code : (hLevelTrunc 6 (Susp S²)) → Type ℓ-zero +Code = + fst ∘ trRec {B = TypeOfHLevel ℓ-zero 5} (isOfHLevelTypeOfHLevel 5) + λ x → (PreCode x) , (hLevPreCode x) + +private + cong-coherence : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (r : refl ≡ p) + → cong (p ∙_) (sym r) ∙ sym (rUnit p) + ≡ cong (_∙ p) (sym r) ∙ sym (lUnit p) + cong-coherence p = J (λ p r → cong (p ∙_) (sym r) ∙ sym (rUnit p) + ≡ cong (_∙ p) (sym r) ∙ sym (lUnit p)) refl + +{- The function Pushout⋁↪fold⋁S² → Ω∥S³∥₆ will be the obvious one -} +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ : Pushout⋁↪fold⋁S² → Ω∥S³∥₆ +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (inl x) = cong ∣_∣ₕ (σ S²∙ (fst x) ∙ σ S²∙ (snd x)) +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (inr x) = cong ∣_∣ₕ (σ S²∙ x) +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (push (inl x) i) j = + ∣ (cong (σ S²∙ x ∙_) (rCancel (merid base)) ∙ sym (rUnit (σ S²∙ x))) i j ∣ₕ +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (push (inr x) i) j = + ∣ (cong (_∙ σ S²∙ x) (rCancel (merid base)) ∙ sym (lUnit (σ S²∙ x))) i j ∣ₕ +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (push (push a k) i) j = + ∣ cong-coherence (σ S²∙ base) (sym (rCancel (merid base))) k i j ∣ₕ + +{- Truncated version (the equivalence) -} +∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ : ∥Pushout⋁↪fold⋁S²∥₅ → Ω∥S³∥₆ +∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ = + trRec (isOfHLevelTrunc 6 _ _) Pushout⋁↪fold⋁S²→Ω∥S³∥₆ + +decode' : (x : Susp S²) → Code ∣ x ∣ + → Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ x ∣ +decode' north p = ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ p +decode' south p = ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ p ∙ cong ∣_∣ₕ (merid base) +decode' (merid a i) = main a i + where + baseId : (x : Pushout⋁↪fold⋁S²) + → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ (S²→Pushout⋁↪fold⋁S²↺' base x) + ≡ ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ ∣ x ∣ + baseId x = + cong ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ (S²→Pushout⋁↪fold⋁S²↺'≡idfun x) + + mainLemma : (a : S²) (x : Pushout⋁↪fold⋁S²) + → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ (S²→Pushout⋁↪fold⋁S²↺'⁻ a x) + ∙ (λ i → ∣ merid a i ∣) + ≡ ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ ∣ x ∣ ∙ cong ∣_∣ₕ (merid base) + mainLemma base x = cong (_∙ cong ∣_∣ₕ (merid base)) (baseId x) + mainLemma (surf i j) x k = surf-filler x k i j + where + ∙ΩbaseId : + (q : typ (Ω (Susp∙ (typ S²∙)))) → q ≡ q ∙ σ S²∙ base + ∙ΩbaseId q = rUnit q ∙ cong (q ∙_) (sym (rCancel (merid base))) + + cubeCoherence : + ∀ {ℓ} {A : Type ℓ} {x : A} {p : x ≡ x} + → (refl ≡ p) + → (q : Square {A = x ≡ x} (λ _ → p) (λ _ → p) (λ _ → p) (λ _ → p)) + → Cube {A = Path A x x} + (λ i j → q i j ∙ q (~ i) j) (λ _ _ → p ∙ p) + (λ k j → p ∙ p) (λ _ _ → p ∙ p) + (λ _ _ → p ∙ p) λ _ _ → p ∙ p + cubeCoherence {A = A} {x = x} {p = p} = + J (λ p _ → (q : Square {A = x ≡ x} + (λ _ → p) (λ _ → p) (λ _ → p) (λ _ → p)) + → Cube {A = Path A x x} + (λ i j → q i j ∙ q (~ i) j) (λ _ _ → p ∙ p) + (λ k j → p ∙ p) (λ _ _ → p ∙ p) + (λ _ _ → p ∙ p) λ _ _ → p ∙ p) + (λ q → + (cong₂Funct (λ (x y : Path (x ≡ x) refl refl) → cong₂ _∙_ x y) q (sym q) + ∙ transport + (λ i → cong (λ (p : (refl {x = x}) ≡ refl) k → rUnit (p k) i) q + ∙ cong (λ (p : (refl {x = x}) ≡ refl) k → lUnit (p k) i) (sym q) + ≡ refl {x = refl {x = lUnit (refl {x = x}) i}}) + (rCancel q))) + + surf-side : (i j r l : I) → hLevelTrunc 6 (Susp S²) + surf-side i j r l = + ((cong ∣_∣ₕ (∙ΩbaseId (σ S²∙ (surf (~ i) j)) r)) + ∙ cong ∣_∣ₕ (compPath-filler (merid (surf i j)) + (sym (merid base)) (~ r))) l + + side : (r l : I) → hLevelTrunc 6 (Susp S²) + side r l = + ((cong ∣_∣ₕ (∙ΩbaseId (σ S²∙ base) r)) + ∙ cong ∣_∣ₕ (compPath-filler (merid base) + (sym (merid base)) (~ r))) l + + surf-filler : (x : Pushout⋁↪fold⋁S²) + → Cube {A = Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ south ∣} + (λ i j → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ + (S²→Pushout⋁↪fold⋁S²↺' (surf (~ i) j) x) + ∙ (λ k → ∣ merid (surf i j) k ∣)) + (λ _ _ → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ ∣ x ∣ + ∙ cong ∣_∣ₕ (merid base)) + (λ k j → baseId x k ∙ (λ i₂ → ∣ merid base i₂ ∣)) + (λ k j → baseId x k ∙ (λ i₂ → ∣ merid base i₂ ∣)) + (λ k i → baseId x k ∙ (λ i₂ → ∣ merid base i₂ ∣)) + (λ k i → baseId x k ∙ (λ i₂ → ∣ merid base i₂ ∣)) + surf-filler = + Pushout⋁↪fold⋁S²→Set + (λ _ → isOfHLevelPathP' 2 + (isOfHLevelPathP' 3 (isOfHLevelTrunc 6 _ _ _ _) _ _) _ _) + (λ k i j l → hcomp (λ r + → λ {(i = i0) → surf-side i0 i0 r l + ; (i = i1) → surf-side i0 i0 r l + ; (j = i0) → surf-side i0 i0 r l + ; (j = i1) → surf-side i0 i0 r l + ; (k = i0) → surf-side i j r l + ; (k = i1) → surf-side i0 i0 r l + ; (l = i0) → ∣ north ∣ₕ + ; (l = i1) → ∣ merid base r ∣ₕ}) + (cubeCoherence {p = cong ∣_∣ₕ (σ (S²∙) base)} + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))) + (λ i j k → ∣ σ S²∙ (surf (~ i) j) k ∣ₕ) k i j l)) + + main : (a : S²) + → PathP (λ i → Code ∣ merid a i ∣ + → Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ merid a i ∣) + ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ + λ x → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ x + ∙ cong ∣_∣ₕ (merid base) + main a = + toPathP (funExt + (trElim (λ _ → isOfHLevelSuc 4 (isOfHLevelTrunc 6 _ _ _ _)) + (λ x + → (λ j → transp (λ i → Path + (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ merid a (i ∨ j) ∣) j + (compPath-filler + (∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ + (S²→Pushout⋁↪fold⋁S²↺'⁻ a (transportRefl x j))) + (cong ∣_∣ₕ (merid a)) j)) + ∙ mainLemma a x))) + +decode : (x : hLevelTrunc 6 (Susp S²)) + → Code x → Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ x +decode = + trElim (λ _ → isOfHLevelΠ 6 λ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) + decode' + +decode∙ : decode ∣ north ∣ ∣ inl (base , base) ∣ ≡ refl +decode∙ = + cong (cong ∣_∣ₕ) (cong (λ x → x ∙ x) (rCancel (merid base)) ∙ sym (rUnit refl)) + +encode : (x : hLevelTrunc 6 (Susp S²)) + → Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ x → Code x +encode x p = subst Code p ∣ inl (base , base) ∣ + +encode-decode : (x : hLevelTrunc 6 (Susp S²)) + → (p : ∣ north ∣ ≡ x) → decode x (encode x p) ≡ p +encode-decode x = J (λ x p → decode x (encode x p) ≡ p) + (cong (decode ∣ north ∣) (transportRefl ∣ inl (base , base) ∣ₕ) + ∙ decode∙) + +decode-encode : (p : Code ∣ north ∣) → encode ∣ north ∣ (decode ∣ north ∣ p) ≡ p +decode-encode = + trElim (λ _ → isOfHLevelPath 5 (isOfHLevelTrunc 5) _ _) + (Pushout⋁↪fold⋁S²WedgeCon.main ((λ a → encode ∣ north ∣ (decode ∣ north ∣ ∣ a ∣))) ∣_∣ₕ + (isOfHLevelTrunc 5) + (λ x → cong (encode ∣ north ∣) (cong (cong ∣_∣ₕ) (cong (σ S²∙ x ∙_) + (rCancel (merid base)) ∙ sym (rUnit (σ S²∙ x)))) + ∙ surf-filler x) + (λ x → (cong (encode ∣ north ∣) (cong (cong ∣_∣ₕ) (cong (_∙ σ S²∙ x) + (rCancel (merid base)) ∙ sym (lUnit (σ S²∙ x)))) + ∙ surf-filler x + ∙ cong ∣_∣ₕ (push (inl x)) ∙ cong ∣_∣ₕ (sym (push (inr x))))) + λ x → lem (encode ∣ north ∣) + (cong (decode ∣ north ∣) (cong ∣_∣ₕ (push (inl x)))) + ((cong (decode ∣ north ∣) (cong ∣_∣ₕ (push (inr x))))) + (surf-filler x) (cong ∣_∣ₕ (push (inl x))) (cong ∣_∣ₕ (sym (push (inr x))))) + where + subber = transport (λ j → Code ∣ merid base (~ j) ∣ₕ) + + lem : ∀ {ℓ} {A B : Type ℓ} {x x' y : A} {l w u : B} + (f : A → B) (p : x ≡ y) (p' : x' ≡ y) (q : f y ≡ l) + (r : l ≡ w) (r' : w ≡ u) + → cong f (sym p) ∙∙ cong f p ∙ q ∙∙ r + ≡ (cong f (sym p') ∙∙ (cong f p' ∙ q ∙ (r ∙ r')) ∙∙ sym r') + lem {x = x} {x' = x'} {y = y'} {l = l} {w = w} {u = u} f p p' q r r' = + (λ i → (λ j → f (p (~ j ∨ i))) ∙∙ (λ j → f (p (j ∨ i))) ∙ q ∙∙ r) + ∙∙ (λ i → (λ j → f (p' (~ j ∨ ~ i))) ∙∙ (λ j → f (p' (j ∨ ~ i))) + ∙ compPath-filler q r i ∙∙ λ j → r (i ∨ j)) + ∙∙ λ i → cong f (sym p') ∙∙ cong f p' ∙ q + ∙ compPath-filler r r' i ∙∙ λ j → r' (~ j ∧ i) + + mainId : (x : S²) + → (S²→Pushout⋁↪fold⋁S²↺' x (inl (base , base))) ≡ ∣ inl (x , base) ∣ₕ + mainId base = refl + mainId (surf i i₁) = refl + + surf-filler : (x : S²) + → encode ∣ north ∣ (λ i → ∣ σ S²∙ x i ∣) ≡ ∣ inl (x , base) ∣ + surf-filler x = + (λ i → transp (λ j → Code (∣ merid base (i ∧ ~ j) ∣ₕ)) (~ i) + (encode ∣ merid base i ∣ + λ j → ∣ compPath-filler + (merid x) (sym (merid base)) (~ i) j ∣ₕ)) + ∙ cong subber + (transportRefl (S²→Pushout⋁↪fold⋁S²↺' x (inl (base , base))) + ∙ mainId x) + +IsoΩ∥SuspS²∥₅∥Pushout⋁↪fold⋁S²∥₅ : + Iso (hLevelTrunc 5 (typ (Ω (Susp∙ S²)))) + (hLevelTrunc 5 Pushout⋁↪fold⋁S²) +IsoΩ∥SuspS²∥₅∥Pushout⋁↪fold⋁S²∥₅ = + compIso (invIso (PathIdTruncIso _)) is + where + is : Iso _ _ + fun is = encode ∣ north ∣ + inv is = decode ∣ north ∣ + rightInv is = decode-encode + leftInv is = encode-decode ∣ north ∣ diff --git a/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso2.agda b/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso2.agda new file mode 100644 index 0000000000..4329e05a97 --- /dev/null +++ b/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso2.agda @@ -0,0 +1,64 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 where + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group + +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Pushout +open import Cubical.HITs.Truncation renaming + (rec to trRec ; elim to trElim ; elim2 to trElim2 ; map to trMap) +open import Cubical.HITs.S2 + +≡→Pushout⋁↪fold⋁≡ : ∀ {ℓ} {A B : Pointed ℓ} + → (A ≡ B) → Pushout⋁↪fold⋁∙ A ≡ Pushout⋁↪fold⋁∙ B +≡→Pushout⋁↪fold⋁≡ = cong Pushout⋁↪fold⋁∙ + +private + ∙≃→π≅ : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) + → (e : typ A ≃ typ B) + → fst e (pt A) ≡ pt B + → GroupEquiv (πGr n A) (πGr n B) + ∙≃→π≅ {A = A} {B = B} n e = + EquivJ (λ A e → (a : A) → fst e a ≡ pt B + → GroupEquiv (πGr n (A , a)) (πGr n B)) + (λ b p → J (λ b p → GroupEquiv (πGr n (typ B , b)) (πGr n B)) + idGroupEquiv (sym p)) + e (pt A) + +{- π₄(S³) ≅ π₃((S² × S²) ⊔ᴬ S²) + where A = S² ∨ S² -} +π₄S³≅π₃PushS² : + GroupIso (πGr 3 (S₊∙ 3)) + (πGr 2 (Pushout⋁↪fold⋁∙ (S₊∙ 2))) +π₄S³≅π₃PushS² = + compGroupIso + (GroupEquiv→GroupIso + (∙≃→π≅ 3 (compEquiv (isoToEquiv (invIso IsoS³S3)) S³≃SuspS²) refl)) + (compGroupIso + (invGroupIso (GrIso-πΩ-π 2)) + (compGroupIso + (πTruncGroupIso 2) + (compGroupIso + (GroupEquiv→GroupIso + (∙≃→π≅ {B = _ , ∣ inl (base , base) ∣ₕ} + 2 (isoToEquiv IsoΩ∥SuspS²∥₅∥Pushout⋁↪fold⋁S²∥₅) encode∙)) + (compGroupIso + (invGroupIso (πTruncGroupIso 2)) + (GroupEquiv→GroupIso (invEq (GroupPath _ _) + (cong (πGr 2) + (cong Pushout⋁↪fold⋁∙ (ua∙ S²≃SuspS¹ refl))))))))) + where + encode∙ : encode ∣ north ∣ refl ≡ ∣ inl (base , base) ∣ + encode∙ = transportRefl _ diff --git a/Cubical/Homotopy/Group/Pi4S3/Summary.agda b/Cubical/Homotopy/Group/Pi4S3/Summary.agda new file mode 100644 index 0000000000..f64d8719b2 --- /dev/null +++ b/Cubical/Homotopy/Group/Pi4S3/Summary.agda @@ -0,0 +1,133 @@ +{- + +This file contains a summary of what remains for π₄(S³) ≡ ℤ/2ℤ to be proved. + +See the module π₄S³ at the end of this file. + +-} + +{-# OPTIONS --safe #-} +module Cubical.Homotopy.Group.Pi4S3.Summary where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed + +open import Cubical.Data.Nat.Base +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Int renaming (ℤ to Int) hiding (_+_) + +open import Cubical.HITs.Sn +open import Cubical.HITs.SetTruncation + +open import Cubical.Homotopy.Group.Base hiding (π) +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.HopfInvariant.Homomorphism +open import Cubical.Homotopy.HopfInvariant.HopfMap +open import Cubical.Homotopy.HopfInvariant.Whitehead +open import Cubical.Homotopy.Whitehead +open import Cubical.Homotopy.Group.Pi3S2 +open import Cubical.Homotopy.Group.Pi4S3.Tricky hiding (hopfInvariantEquiv) + +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Instances.Bool +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Instances.Int +open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.Group.ZAction + +private + variable + ℓ : Level + +-- TODO: ideally this would not be off by one in the definition of π'Gr +π : ℕ → Pointed ℓ → Group ℓ +π n X = π'Gr (predℕ n) X + +-- Nicer notation for the spheres (as pointed types) +𝕊² 𝕊³ : Pointed₀ +𝕊² = S₊∙ 2 +𝕊³ = S₊∙ 3 + +-- Whitehead product +[_]× : {X : Pointed ℓ} {n m : ℕ} → π' (suc n) X × π' (suc m) X → π' (suc (n + m)) X +[_]× (f , g) = [ f ∣ g ]π' + +-- Some type abbreviations (unproved results) +π₄S³≡ℤ/something : GroupEquiv (π 3 𝕊²) ℤ → Type₁ +π₄S³≡ℤ/something eq = + π 4 𝕊³ ≡ ℤ/ abs (eq .fst .fst [ ∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂ ]×) + + +-- The intended proof: +module π₄S³ + (π₃S²≃ℤ : GroupEquiv (π 3 𝕊²) ℤ) + (gen-by-HopfMap : gen₁-by (π 3 𝕊²) ∣ HopfMap ∣₂) + (π₄S³≡ℤ/whitehead : π₄S³≡ℤ/something π₃S²≃ℤ) + (hopfWhitehead : abs (HopfInvariant-π' 0 ([ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×)) ≡ 2) + where + + -- Package the Hopf invariant up into a group equivalence + hopfInvariantEquiv : GroupEquiv (π 3 𝕊²) ℤ + fst (fst hopfInvariantEquiv) = HopfInvariant-π' 0 + snd (fst hopfInvariantEquiv) = + GroupEquivℤ-isEquiv (invGroupEquiv π₃S²≃ℤ) ∣ HopfMap ∣₂ + gen-by-HopfMap (GroupHom-HopfInvariant-π' 0) + (abs→⊎ _ _ HopfInvariant-HopfMap) + snd hopfInvariantEquiv = snd (GroupHom-HopfInvariant-π' 0) + + -- The two equivalences map [ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]× to + -- the same number, modulo the sign + remAbs : abs (groupEquivFun π₃S²≃ℤ [ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×) + ≡ abs (groupEquivFun hopfInvariantEquiv [ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×) + remAbs = absGroupEquivℤGroup (invGroupEquiv π₃S²≃ℤ) (invGroupEquiv hopfInvariantEquiv) _ + + -- So the image of [ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]× under π₃S²≃ℤ + -- is 2 (modulo the sign) + remAbs₂ : abs (groupEquivFun π₃S²≃ℤ [ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×) ≡ 2 + remAbs₂ = remAbs ∙ hopfWhitehead + + -- The final result then follows + π₄S³≡ℤ : π 4 𝕊³ ≡ ℤ/ 2 + π₄S³≡ℤ = π₄S³≡ℤ/whitehead ∙ cong (ℤ/_) remAbs₂ + +-- In order to instantiate the module, we need the four following lemmas: + +{- Lemma 1 -} +Lemma₁ : GroupEquiv ℤ (π'Gr 2 (S₊∙ 2)) +Lemma₁ = invGroupEquiv π₃S²≅ℤ + +{- Lemma 2 -} +Lemma₂ : gen₁-by (π 3 𝕊²) ∣ HopfMap ∣₂ +Lemma₂ = π₂S³-gen-by-HopfMap + +{- Lemma 3 -} +{- +Lemma₃ : π₄S³≡ℤ/something π₃S²≅ℤ +Lemma₃ = {!!} + +-} + + + +{- Lemma 4 -} +Lemma₄ : abs (HopfInvariant-π' 0 ([ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×)) ≡ 2 +Lemma₄ = HopfInvariantWhitehead + +-- However, we when trying to prove it, it turned out to be easier to diverge +-- from the above a bit, since we do not have enough theory about exact sequences +-- in the library instead of proving (π₄S³≡ℤ/something π₃S²≅ℤ), we have first proved +-- abs (HopfInvariant-π' 0 ([ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×)) ≡ 2) → π₄S³≅ℤ/2 + +hopfWhitehead→π₄S³≅ℤ/2 : + abs (HopfInvariant-π' 0 ([ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×)) ≡ 2 + → GroupEquiv (π 4 𝕊³) (ℤ/ 2) +hopfWhitehead→π₄S³≅ℤ/2 p = + compGroupEquiv + (compGroupEquiv (GroupIso→GroupEquiv (π'Gr≅πGr 3 𝕊³)) + (∣HopfWhitehead∣≡2→π₄S³≅Bool p)) + (GroupIso→GroupEquiv Bool≅ℤ/2) + +-- And so we get the Iso +π₄S³≅ℤ/2 : GroupEquiv (π 4 𝕊³) (ℤ/ 2) +π₄S³≅ℤ/2 = hopfWhitehead→π₄S³≅ℤ/2 Lemma₄ diff --git a/Cubical/Homotopy/Group/Pi4S3/Tricky.agda b/Cubical/Homotopy/Group/Pi4S3/Tricky.agda new file mode 100644 index 0000000000..c9ca909f0d --- /dev/null +++ b/Cubical/Homotopy/Group/Pi4S3/Tricky.agda @@ -0,0 +1,772 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Group.Pi4S3.Tricky where + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Homotopy.Group.Base +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open Iso +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Morphism + +open import Cubical.Data.Unit +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 + ; map to sMap) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.S1 hiding (decode ; encode) + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Bool + +open import Cubical.Algebra.Group renaming (Unit to UnitGr ; Bool to BoolGr) +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Group.ZAction + +open import Cubical.HITs.Join +open import Cubical.HITs.Pushout +open import Cubical.HITs.Wedge +open import Cubical.Homotopy.Freudenthal hiding (Code ; encode) +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation renaming + (rec to trRec ; elim to trElim ; elim2 to trElim2 ; map to trMap) +open import Cubical.Foundations.Function +open import Cubical.HITs.S2 + +open import Cubical.Homotopy.BlakersMassey +open import Cubical.Homotopy.Whitehead + +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim ; map to pMap) + + +W' = joinTo⋁ {A = S₊∙ 1} {B = S₊∙ 1} + +W : S₊ 3 → (S₊∙ 2 ⋁ S₊∙ 2) +W = W' ∘ Iso.inv (IsoSphereJoin 1 1) + +fold∘W : S₊ 3 → S₊ 2 +fold∘W = fold⋁ ∘ W + +thePushout : Type +thePushout = Pushout (λ _ → tt) fold∘W + +isConnectedS3→S2 : (f : S₊ 3 → S₊ 2) → isConnectedFun 2 f +isConnectedS3→S2 f p = + trRec (isProp→isOfHLevelSuc 1 isPropIsContr) + (J (λ p _ → isConnected 2 (fiber f p)) + (∣ north , refl ∣ + , (trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (uncurry (sphereElim 2 (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) + λ p → trRec (isOfHLevelTrunc 2 _ _) (λ r → cong ∣_∣ₕ (ΣPathP (refl , r))) (fun (PathIdTruncIso 1) (isContr→isProp (isConnectedPath 2 (sphereConnected 2) (f north) (f north)) ∣ refl ∣ ∣ p ∣))))))) + (fun (PathIdTruncIso 2) + (isContr→isProp (sphereConnected 2) ∣ f north ∣ ∣ p ∣)) + +fibreUnitMap : ∀ {ℓ} {A : Type ℓ} → Iso (fiber (λ (a : A) → tt) tt) A +fun fibreUnitMap (x , p) = x +inv fibreUnitMap a = a , refl +rightInv fibreUnitMap _ = refl +leftInv fibreUnitMap _ = refl + +isConnectedUnitmap : ∀ {ℓ} {A : Type ℓ} (n : ℕ) + → isConnected n A + → isConnectedFun n (λ (a : A) → tt) +isConnectedUnitmap n con tt = + subst (isConnected n) (isoToPath (invIso fibreUnitMap)) con + +module BM-inst = + BlakersMassey□ (λ _ → tt) fold∘W 3 + 1 + (isConnectedUnitmap 4 (sphereConnected 3)) + (isConnectedS3→S2 fold∘W) + +open BM-inst + +Grr : S₊ 3 → Σ (Unit × S₊ 2) PushoutPath× +Grr = toPullback + +thePullback∙ : Pointed₀ +fst thePullback∙ = Σ (Unit × S₊ 2) PushoutPath× +snd thePullback∙ = (tt , north) , (push north) + +isConGrr : isConnectedFun 4 Grr +isConGrr = isConnected-toPullback + +open import Cubical.Homotopy.Group.LES + +thePushout∙ : Pointed₀ +thePushout∙ = thePushout , inl tt + +inr' : S₊ 2 → thePushout +inr' = inr + +fibreinr' : Iso (fiber inr' (inl tt)) (Σ (Unit × S₊ 2) PushoutPath×) +fun fibreinr' (x , p) = (tt , x) , (sym p) +inv fibreinr' ((tt , x) , p) = x , (sym p) +rightInv fibreinr' _ = refl +leftInv fibreinr' _ = refl + +TotalPushoutPath×∙ : Pointed ℓ-zero +fst TotalPushoutPath×∙ = Σ (Unit × S₊ 2) PushoutPath× +snd TotalPushoutPath×∙ = (tt , north) , push north + +P : Pointed₀ +P = (fiber inr' (inl tt) , north , (sym (push north))) + + +π'P→π'TotalPath× : (n : ℕ) → GroupEquiv (π'Gr n TotalPushoutPath×∙) (π'Gr n P) +fst (fst (π'P→π'TotalPath× n)) = + π'eqFun n ((invEquiv (isoToEquiv fibreinr')) , refl) +snd (fst (π'P→π'TotalPath× n)) = π'eqFunIsEquiv n _ +snd (π'P→π'TotalPath× n) = π'eqFunIsHom n _ + +inr∙ : S₊∙ 2 →∙ thePushout∙ +fst inr∙ = inr' +snd inr∙ = sym (push north) + +module LESinst = πLES {A = S₊∙ 2} {B = thePushout∙} inr∙ + + +≃∙→πHom : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (A ≃∙ B) + → (n : ℕ) → GroupEquiv (πGr n A) (πGr n B) +fst (fst (≃∙→πHom e n)) = fst (πHom n (≃∙map e)) +snd (fst (≃∙→πHom e n)) = + isoToIsEquiv (setTruncIso (equivToIso (_ , isEquivΩ^→ (suc n) (≃∙map e) (snd (fst e))))) +snd (≃∙→πHom e n) = snd (πHom n (≃∙map e)) + +isConnectedΩ→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (f : A →∙ B) + → isConnectedFun (suc n) (fst f) + → isConnectedFun n (fst (Ω→ f)) +isConnectedΩ→ n f g = + transport (λ i → isConnectedFun n λ b + → doubleCompPath-filler (sym (snd f)) (cong (fst f) b) (snd f) i) + (isConnectedCong n _ g) + +isConnectedΩ^→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n m : ℕ) + → (f : A →∙ B) + → isConnectedFun (suc n) (fst f) + → isConnectedFun ((suc n) ∸ m) (fst (Ω^→ m f)) +isConnectedΩ^→ n zero f conf = conf +isConnectedΩ^→ n (suc zero) f conf = isConnectedΩ→ n f conf +isConnectedΩ^→ {A = A} {B = B} n (suc (suc m)) f conf = + transport (λ i → isConnectedFun (suc n ∸ suc (suc m)) + λ q → doubleCompPath-filler + (sym (snd (Ω^→ (suc m) f))) + (cong (fst (Ω^→ (suc m) f)) q) + (snd (Ω^→ (suc m) f)) i) + (help n m (help2 n (suc m)) (isConnectedΩ^→ n (suc m) f conf)) + where + open import Cubical.Data.Sum + help2 : (n m : ℕ) → ((suc n ∸ m ≡ suc (n ∸ m))) ⊎ (suc n ∸ suc m ≡ 0) + help2 n zero = inl refl + help2 zero (suc m) = inr refl + help2 (suc n) (suc m) = help2 n m + + help : (n m : ℕ) → ((suc n ∸ (suc m) ≡ suc (n ∸ suc m))) ⊎ (suc n ∸ suc (suc m) ≡ 0) + → isConnectedFun (suc n ∸ suc m) (fst (Ω^→ (suc m) f)) + → isConnectedFun (suc n ∸ suc (suc m)) + {A = Path ((Ω^ (suc m)) (_ , pt A) .fst) + refl refl} + (cong (fst (Ω^→ (suc m) f))) + help n m (inl x) conf = + isConnectedCong (n ∸ suc m) (fst (Ω^→ (suc m) f)) + (subst (λ x → isConnectedFun x (fst (Ω^→ (suc m) f))) x conf) + help n m (inr x) conf = + subst (λ x → isConnectedFun x (cong {x = refl} {y = refl} (fst (Ω^→ (suc m) f)))) + (sym x) + λ b → tt* , λ _ → refl + +π₂P≅0 : GroupEquiv (πGr 1 P) UnitGr +π₂P≅0 = compGroupEquiv (≃∙→πHom (isoToEquiv fibreinr' , refl) 1) + (GroupIso→GroupEquiv + (contrGroupIsoUnit (isOfHLevelRetractFromIso 0 (invIso map23) isContrπ₂S³))) + where + mapp : Iso (hLevelTrunc 4 (S₊ 3)) (hLevelTrunc 4 (Σ (Unit × S₊ 2) PushoutPath×)) + mapp = connectedTruncIso 4 Grr isConnected-toPullback + + map23 : Iso (π 2 (hLevelTrunc∙ 4 (S₊∙ 3))) (π 2 thePullback∙) + map23 = + (compIso (setTruncIso + (equivToIso (_ , + (isEquivΩ^→ 2 (fun mapp , refl) (isoToIsEquiv mapp))))) + (invIso (πTruncIso 2))) + + zz : isContr (π 2 (Unit , tt)) + fst zz = ∣ refl ∣₂ + snd zz = sElim (λ _ → isSetPathImplicit) λ p → refl + + isContrπ₂S³ : isContr (π 2 (hLevelTrunc∙ 4 (S₊∙ 3))) + isContrπ₂S³ = + subst (λ x → isContr (π 2 x)) + (λ i → ((sym (isContr→≡Unit (sphereConnected 3))) i) + , transp (λ j → isContr→≡Unit + (sphereConnected 3) (~ i ∧ j)) i ∣ north ∣) + zz + +testComm : (λ x → snd (fst x)) ∘ Grr ≡ fold∘W +testComm = refl +-- nice + +-- G → H → L → R +record exact {ℓ ℓ' ℓ'' ℓ''' : Level} (G : Group ℓ) (H : Group ℓ') (L : Group ℓ'') (R : Group ℓ''') + (G→H : GroupHom G H) (H→L : GroupHom H L) (L→R : GroupHom L R) + : Type (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ'''))) where + field + ImG→H⊂KerH→L : (x : fst H) → isInIm G→H x → isInKer H→L x + KerH→L⊂ImG→H : (x : fst H) → isInKer H→L x → isInIm G→H x + + ImH→L⊂KerL→R : (x : fst L) → isInIm H→L x → isInKer L→R x + KerL→R⊂ImH→L : (x : fst L) → isInKer L→R x → isInIm H→L x + +P→S²→Pushout : + exact (πGr 2 P) (πGr 2 (S₊∙ 2)) (πGr 2 thePushout∙) (πGr 1 P) + (LESinst.fib→A 2) + (LESinst.A→B 2) + (LESinst.B→fib 1) +exact.ImG→H⊂KerH→L P→S²→Pushout = LESinst.Im-fib→A⊂Ker-A→B 2 +exact.KerH→L⊂ImG→H P→S²→Pushout = LESinst.Ker-A→B⊂Im-fib→A 2 +exact.ImH→L⊂KerL→R P→S²→Pushout = LESinst.Im-A→B⊂Ker-B→fib 1 +exact.KerL→R⊂ImH→L P→S²→Pushout = LESinst.Ker-B→fib⊂Im-A→B 1 + +transportExact : {ℓ ℓ' ℓ'' ℓ''' : Level} + (G G₂ : Group ℓ) (H H₂ : Group ℓ') (L L₂ : Group ℓ'') (R R₂ : Group ℓ''') + → (G≡G₂ : G ≡ G₂) → (H≡H₂ : H ≡ H₂) (L≡L₂ : L ≡ L₂) (R≡R₂ : R ≡ R₂) + → (G→H : GroupHom G H) (G₂→H₂ : GroupHom G₂ H₂) + → (H→L : GroupHom H L) (H₂→L₂ : GroupHom H₂ L₂) + → (L→R : GroupHom L R) + → exact G H L R G→H H→L L→R + → PathP (λ i → GroupHom (G≡G₂ i) (H≡H₂ i)) G→H G₂→H₂ + → PathP (λ i → GroupHom (H≡H₂ i) (L≡L₂ i)) H→L H₂→L₂ + → Σ[ L₂→R₂ ∈ GroupHom L₂ R₂ ] + exact G₂ H₂ L₂ R₂ G₂→H₂ H₂→L₂ L₂→R₂ +transportExact G G₂ H H₂ L L₂ R R₂ = + J4 (λ G₂ H₂ L₂ R₂ G≡G₂ H≡H₂ L≡L₂ R≡R₂ + → (G→H : GroupHom G H) (G₂→H₂ : GroupHom G₂ H₂) + → (H→L : GroupHom H L) (H₂→L₂ : GroupHom H₂ L₂) + → (L→R : GroupHom L R) + → exact G H L R G→H H→L L→R + → PathP (λ i → GroupHom (G≡G₂ i) (H≡H₂ i)) G→H G₂→H₂ + → PathP (λ i → GroupHom (H≡H₂ i) (L≡L₂ i)) H→L H₂→L₂ + → Σ[ L₂→R₂ ∈ GroupHom L₂ R₂ ] + exact G₂ H₂ L₂ R₂ G₂→H₂ H₂→L₂ L₂→R₂) + (λ G→H G₂→H₂ H→L H₂→L₂ L→R ex + → J (λ G₂→H₂ _ + → H→L ≡ H₂→L₂ + → Σ[ L→R ∈ GroupHom L R ] + exact G H L R G₂→H₂ H₂→L₂ L→R) + (J (λ H₂→L₂ _ → Σ[ L→R ∈ GroupHom L R ] + exact G H L R G→H H₂→L₂ L→R) + (L→R , ex))) + G₂ H₂ L₂ R₂ + where + abstract + J4 : ∀ {ℓ ℓ₂ ℓ₃ ℓ₄ ℓ'} {A : Type ℓ} {A₂ : Type ℓ₂} {A₃ : Type ℓ₃} {A₄ : Type ℓ₄} + {x : A} {x₂ : A₂} {x₃ : A₃} {x₄ : A₄} + (B : (y : A) (z : A₂) (w : A₃) (u : A₄) → x ≡ y → x₂ ≡ z → x₃ ≡ w → x₄ ≡ u → Type ℓ') + → B x x₂ x₃ x₄ refl refl refl refl + → (y : A) (z : A₂) (w : A₃) (u : A₄) (p : x ≡ y) (q : x₂ ≡ z) (r : x₃ ≡ w) (s : x₄ ≡ u) + → B y z w u p q r s + J4 {x = x} {x₂ = x₂} {x₃ = x₃} {x₄ = x₄} B b y z w u = + J (λ y p → (q : x₂ ≡ z) (r : x₃ ≡ w) (s : x₄ ≡ u) → + B y z w u p q r s) + (J (λ z q → (r : x₃ ≡ w) (s : x₄ ≡ u) → B x z w u refl q r s) + (J (λ w r → (s : x₄ ≡ u) → B x x₂ w u refl refl r s) + (J (λ u s → B x x₂ x₃ u refl refl refl s) b))) + +ΣP→S²→Pushout' : + Σ[ F ∈ GroupHom (π'Gr 2 thePushout∙) (π'Gr 1 P) ] + (exact (π'Gr 2 P) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) (π'Gr 1 P) + (π'∘∙Hom 2 (fst , refl)) + (π'∘∙Hom 2 inr∙) + F) +ΣP→S²→Pushout' = + transportExact _ _ _ _ _ _ _ _ + (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 2 P))))) + (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 2 (S₊∙ 2)))))) + (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 2 thePushout∙))))) + (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 1 P))))) + _ _ _ _ _ + P→S²→Pushout + (toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (fromPathP λ i → fst (π∘∙fib→A-PathP 2 inr∙ i)))) + ((toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (fromPathP λ i → fst (π∘∙A→B-PathP 2 inr∙ i))))) + +abstract + π₂thePushout→π₁P : GroupHom (π'Gr 2 thePushout∙) (π'Gr 1 P) + π₂thePushout→π₁P = fst ΣP→S²→Pushout' + + P→S²→Pushout→P' : + exact (π'Gr 2 P) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) (π'Gr 1 P) + (π'∘∙Hom 2 (fst , refl)) + (π'∘∙Hom 2 inr∙) + π₂thePushout→π₁P + P→S²→Pushout→P' = ΣP→S²→Pushout' .snd + +→UnitHom : ∀ {ℓ} (G : Group ℓ) → GroupHom G UnitGr +fst (→UnitHom G) _ = tt +snd (→UnitHom G) = makeIsGroupHom λ _ _ → refl + +compSurj : ∀ {ℓ ℓ' ℓ''} {G : Group ℓ} {H : Group ℓ'} {L : Group ℓ''} + → (G→H : GroupHom G H) (H→L : GroupHom H L) + → isSurjective G→H → isSurjective H→L + → isSurjective (compGroupHom G→H H→L) +compSurj G→H H→L surj1 surj2 l = + pRec squash + (λ {(h , p) + → pMap (λ {(g , q) → g , (cong (fst H→L) q ∙ p)}) + (surj1 h)}) + (surj2 l) + +π₃S³→π₃PΩ : GroupHom (πGr 2 (S₊∙ 3)) (πGr 2 TotalPushoutPath×∙) +π₃S³→π₃PΩ = πHom 2 (Grr , refl) + +TotalPushoutPath×∙→P : TotalPushoutPath×∙ →∙ P +fst TotalPushoutPath×∙→P x = (snd (fst x)) , (sym (snd x)) +snd TotalPushoutPath×∙→P = refl + +isSurjective-π₃S³→π₃PΩ : isSurjective π₃S³→π₃PΩ +isSurjective-π₃S³→π₃PΩ = + sElim (λ _ → isProp→isSet squash) + λ p → trRec squash + (λ s → ∣ ∣ fst s ∣₂ , (cong ∣_∣₂ (snd s)) ∣) + (((isConnectedΩ^→ 3 3 (Grr , refl) isConGrr) p) .fst) + + +π₃S³→π₃P : GroupHom (π'Gr 2 (S₊∙ 3)) (π'Gr 2 TotalPushoutPath×∙) +π₃S³→π₃P = π'∘∙Hom 2 (Grr , refl) + +transportLem : PathP (λ i → GroupHomπ≅π'PathP (S₊∙ 3) TotalPushoutPath×∙ 2 i) + π₃S³→π₃PΩ π₃S³→π₃P +transportLem = + toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (π'∘∙Hom'≡π'∘∙fun {A = S₊∙ 3} {B = TotalPushoutPath×∙} 2 (Grr , refl))) + +isSurjective-π₃S³→π₃P : isSurjective π₃S³→π₃P +isSurjective-π₃S³→π₃P = + transport (λ i → isSurjective (transportLem i)) + isSurjective-π₃S³→π₃PΩ + +π₃S³→π₃S²' : GroupHom (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) +π₃S³→π₃S²' = compGroupHom π₃S³→π₃P (π'∘∙Hom 2 ((λ x → snd (fst x)) , refl)) + +test : π₃S³→π₃S²' ≡ π'∘∙Hom 2 (fold∘W , refl) +test = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (ΣPathP (refl , + (λ j → (λ i → snd (fst ((rUnit (λ k → Grr (snd f k)) (~ j)) i))) + ∙ (λ _ → north)))))) + +amap : TotalPushoutPath×∙ →∙ P +fst amap ((tt , s), p) = s , (sym p) +snd amap = refl + +π₃S³→π₃P' : GroupHom (π'Gr 2 (S₊∙ 3)) (π'Gr 2 P) +π₃S³→π₃P' = compGroupHom π₃S³→π₃P (π'∘∙Hom 2 amap) + +isSurjective-π₃S³→π₃P' : isSurjective π₃S³→π₃P' +isSurjective-π₃S³→π₃P' = + compSurj π₃S³→π₃P (π'∘∙Hom 2 amap) + isSurjective-π₃S³→π₃P + (sElim (λ _ → isProp→isSet squash) + λ {(s , p) → ∣ ∣ (λ x → (tt , s x .fst) , sym (s x .snd)) + , ΣPathP ((ΣPathP (refl , (cong fst p))) + , (λ i j → snd (p i) (~ j))) ∣₂ + , cong ∣_∣₂ (ΣPathP (refl , sym (rUnit p))) ∣}) + +-- π₃P → π₃S² → π₃ Pushout → Unit +P→S²→Pushout→Unit : + exact (π'Gr 2 P) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr + (π'∘∙Hom 2 (fst , refl)) + (π'∘∙Hom 2 inr∙) + (→UnitHom (π'Gr 2 thePushout∙)) +P→S²→Pushout→Unit = + transport (λ i → + exact (π'Gr 2 P) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) + (GroupPath _ _ .fst + (compGroupEquiv (GroupIso→GroupEquiv (π'Gr≅πGr 1 P)) π₂P≅0) i) + (π'∘∙Hom 2 (fst , refl)) + (π'∘∙Hom 2 inr∙) + (r i)) + P→S²→Pushout→P' + where + r : PathP (λ i → GroupHom (π'Gr 2 thePushout∙) + ((GroupPath _ _ .fst + (compGroupEquiv (GroupIso→GroupEquiv (π'Gr≅πGr 1 P)) π₂P≅0) i))) + π₂thePushout→π₁P + (→UnitHom (π'Gr 2 thePushout∙)) + r = toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ _ → refl)) +open exact + +extendExactSurj : {ℓ ℓ' ℓ'' ℓ''' ℓ'''' : Level} + (G : Group ℓ) (H : Group ℓ') (L : Group ℓ'') (R : Group ℓ''') (S : Group ℓ'''') + (G→H : GroupHom G H) (H→L : GroupHom H L) (L→R : GroupHom L R) (R→S : GroupHom R S) + → isSurjective G→H + → exact H L R S H→L L→R R→S + → exact G L R S (compGroupHom G→H H→L) L→R R→S +ImG→H⊂KerH→L (extendExactSurj G H L R S G→H H→L L→R R→S surj ex) x = + pRec (GroupStr.is-set (snd R) _ _) + (uncurry λ g → J (λ x _ → isInKer L→R x) + (ImG→H⊂KerH→L ex (fst H→L (fst G→H g)) + ∣ (fst G→H g) , refl ∣)) +KerH→L⊂ImG→H (extendExactSurj G H L R S G→H H→L L→R R→S surj ex) x ker = + pRec squash + (uncurry λ y → J (λ x _ → isInIm (compGroupHom G→H H→L) x) + (pMap (uncurry + (λ y → J (λ y _ → Σ[ g ∈ fst G ] fst H→L (fst G→H g) ≡ H→L .fst y) + (y , refl))) (surj y))) + (KerH→L⊂ImG→H ex x ker) +ImH→L⊂KerL→R (extendExactSurj G H L R S G→H H→L L→R R→S surj ex) = + ImH→L⊂KerL→R ex +KerL→R⊂ImH→L (extendExactSurj G H L R S G→H H→L L→R R→S surj ex) = + KerL→R⊂ImH→L ex + +P'→S²→Pushout→Unit' : + exact (π'Gr 2 TotalPushoutPath×∙) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr + (compGroupHom (π'∘∙Hom 2 TotalPushoutPath×∙→P) (π'∘∙Hom 2 (fst , refl))) + (π'∘∙Hom 2 inr∙) + (→UnitHom (π'Gr 2 thePushout∙)) +P'→S²→Pushout→Unit' = + extendExactSurj _ _ _ _ _ _ _ _ _ + (sElim (λ _ → isProp→isSet squash) + (λ f → ∣ ∣ (λ x → (tt , fst f x .fst) , sym (fst f x .snd)) + , ΣPathP ((ΣPathP (refl , cong fst (snd f))) , λ j i → snd f j .snd (~ i)) ∣₂ + , cong ∣_∣₂ (ΣPathP (refl , sym (rUnit _))) ∣)) + P→S²→Pushout→Unit + + +S³→S²→Pushout→Unit'' : + exact (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr + (compGroupHom π₃S³→π₃P + (compGroupHom (π'∘∙Hom 2 TotalPushoutPath×∙→P) (π'∘∙Hom 2 (fst , refl)))) + (π'∘∙Hom 2 inr∙) + (→UnitHom (π'Gr 2 thePushout∙)) +S³→S²→Pushout→Unit'' = + extendExactSurj _ _ _ _ _ _ _ _ _ + isSurjective-π₃S³→π₃P P'→S²→Pushout→Unit' + +tripleComp≡ : + (compGroupHom π₃S³→π₃P + (compGroupHom + (π'∘∙Hom 2 TotalPushoutPath×∙→P) (π'∘∙Hom 2 (fst , refl)))) + ≡ π'∘∙Hom 2 (fold∘W , refl) +tripleComp≡ = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (ΣPathP (refl , (cong (_∙ refl) + (λ j → cong fst (rUnit (cong (fst TotalPushoutPath×∙→P) + (rUnit (cong Grr (snd f)) (~ j))) (~ j)))))))) + +S³→S²→Pushout→Unit : + exact (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr + (π'∘∙Hom 2 (fold∘W , refl)) + (π'∘∙Hom 2 inr∙) + (→UnitHom (π'Gr 2 thePushout∙)) +S³→S²→Pushout→Unit = + subst + (λ F → exact (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr + F (π'∘∙Hom 2 inr∙) + (→UnitHom (π'Gr 2 thePushout∙))) + tripleComp≡ + S³→S²→Pushout→Unit'' + +indΠ₃S₂ : ∀ {ℓ} {A : Pointed ℓ} → (f g : A →∙ S₊∙ 2) → fst f ≡ fst g → ∥ f ≡ g ∥ +indΠ₃S₂ {A = A} f g p = + trRec squash + (λ r → ∣ ΣPathP (p , r) ∣) + (isConnectedPathP 1 {A = (λ i → p i (snd A) ≡ north)} + (isConnectedPathSⁿ 1 (fst g (pt A)) north) (snd f) (snd g) .fst ) + +Fold∘W : fst (π'∘∙Hom 2 (fold∘W , refl)) ∣ idfun∙ (S₊∙ 3) ∣₂ + ≡ ∣ [ idfun∙ (S₊∙ 2) ∣ idfun∙ (S₊∙ 2) ]₂ ∣₂ +Fold∘W = + pRec (squash₂ _ _) + (cong ∣_∣₂) + (indΠ₃S₂ _ _ + (funExt (λ x → funExt⁻ (sym (cong fst (id∨→∙id {A = S₊∙ 2}))) (W x)))) + +-- Todo: generalise this + +-- ℤ --f--> ℤ --g--> ℤ ---> 0 + +open import Cubical.Data.Int renaming (ℤ to Z ; _·_ to _·Z_ ; _+_ to _+Z_) +module exact→BoolIso (G : Group₀) (F : GroupHom ℤ ℤ) (H : GroupHom ℤ G) (P : fst F 1 ≡ 2) + (ex : exact ℤ ℤ G UnitGr F H (→UnitHom G)) where + + strG = snd G + strH = snd H + open GroupStr + + Bool→G : Bool → fst G + Bool→G false = fst H 1 + Bool→G true = 1g strG + + isSurjectiveH : isSurjective H + isSurjectiveH b = exact.KerL→R⊂ImH→L ex b refl + + Bool→GHom : IsGroupHom (snd BoolGr) Bool→G strG + Bool→GHom = + makeIsGroupHom + λ { false false → sym (exact.ImG→H⊂KerH→L ex 2 ∣ 1 , P ∣) + ∙ IsGroupHom.pres· (snd H) 1 1 + ; false true → sym (rid (snd G) _) + ; true y → sym (lid (snd G) _)} + + BoolGHom : GroupHom BoolGr G + fst BoolGHom = Bool→G + snd BoolGHom = Bool→GHom + + genG : fst G + genG = fst H 1 + + open IsGroupHom + + _·G_ = GroupStr._·_ (snd G) + + isSurjectiveBoolGHomPos : (n : ℕ) → Σ[ b ∈ Bool ] fst H (pos n) ≡ Bool→G b + isSurjectiveBoolGHomPos zero = true , pres1 strH + isSurjectiveBoolGHomPos (suc zero) = false , refl + isSurjectiveBoolGHomPos (suc (suc n)) = + isSurjectiveBoolGHomPos n .fst + , (cong (fst H) (+Comm (pos n) 2) + ∙ (pres· (snd H) 2 (pos n))) + ∙ cong (_·G (fst H (pos n))) + (sym (sym (exact.ImG→H⊂KerH→L ex 2 ∣ 1 , P ∣))) + ∙ lid (snd G) _ + ∙ isSurjectiveBoolGHomPos n .snd + + isSurjectiveBoolGHomPre : (x : Z) → Σ[ b ∈ Bool ] fst H x ≡ Bool→G b + isSurjectiveBoolGHomPre (pos n) = isSurjectiveBoolGHomPos n + isSurjectiveBoolGHomPre (negsuc n) = + (isSurjectiveBoolGHomPos (suc n) .fst) + , cong (fst H) (+Comm (- (pos (suc n))) 0) + ∙ presinv (snd H) (pos (suc n)) + ∙ cong (inv strG) (isSurjectiveBoolGHomPos (suc n) .snd) + ∙ sym (presinv Bool→GHom (isSurjectiveBoolGHomPos (suc n) .fst)) + + isSurjectiveBoolGHom : isSurjective BoolGHom + isSurjectiveBoolGHom g = + pMap (λ {(x , p) → (isSurjectiveBoolGHomPre x .fst) + , (sym (isSurjectiveBoolGHomPre x .snd) ∙ p)}) + (isSurjectiveH g) + + isInjectiveBoolGHom : isInjective BoolGHom + isInjectiveBoolGHom false p = + ⊥-rec (pRec isProp⊥ (uncurry contr) ∃1≡2) + where + H≡0 : isInKer H 1 + H≡0 = p + + H≡ : isInIm F 1 + H≡ = exact.KerH→L⊂ImG→H ex 1 p + + F≡pos : (n : ℕ) → fst F (pos n) ≡ 2 ·Z (pos n) + F≡pos zero = pres1 (snd F) + F≡pos (suc zero) = P + F≡pos (suc (suc n)) = + cong (fst F) (+Comm (pos (suc n)) (pos 1)) + ∙∙ pres· (snd F) (pos 1) (pos (suc n)) + ∙∙ cong₂ _+Z_ P (F≡pos (suc n)) + ∙∙ +Comm 2 (2 ·Z (pos (suc n))) + ∙∙ (sym (·DistR+ 2 (pos (suc n)) 1)) + + F≡ : (x : Z) → fst F x ≡ 2 ·Z x + F≡ (pos n) = F≡pos n + F≡ (negsuc n) = + cong (fst F) (+Comm (negsuc n) 0) + ∙ presinv (snd F) (pos (suc n)) + ∙ (sym (+Comm (- (fst F (pos (suc n)))) 0) + ∙ cong -_ (F≡pos (suc n))) + ∙ sym (ℤ·negsuc 2 n) + + ∃1≡2 : ∃[ x ∈ Z ] 1 ≡ 2 ·Z x + ∃1≡2 = + pMap (λ {(x , p) → x , (sym p ∙ F≡ x)}) H≡ + + contr : (x : Z) → 1 ≡ 2 ·Z x → ⊥ + contr (pos zero) p = snotz (injPos p) + contr (pos (suc n)) p = snotz (injPos (sym (cong predℤ p ∙ predSuc _ ∙ help1 n))) + where + help1 : (n : ℕ) → pos (suc n) +pos n ≡ pos (suc (n + n)) + help1 zero = refl + help1 (suc n) = cong sucℤ (sym (sucℤ+ (pos (suc n)) (pos n))) + ∙ cong (sucℤ ∘ sucℤ) (help1 n) + ∙ cong pos (cong (2 +_) (sym (+-suc n n))) + contr (negsuc n) p = negsucNotpos (suc (n + n)) 1 (sym (p ∙ negsuclem n)) + where + negsuclem : (n : ℕ) → (negsuc n +negsuc n) ≡ negsuc (suc (n + n)) + negsuclem zero = refl + negsuclem (suc n) = + cong predℤ (+Comm (negsuc (suc n)) (negsuc n) + ∙ cong predℤ (negsuclem n)) + ∙ cong (negsuc ∘ (2 +_)) (sym (+-suc n n)) + isInjectiveBoolGHom true p = refl + + BijectionIsoBoolG : BijectionIso BoolGr G + BijectionIso.fun BijectionIsoBoolG = BoolGHom + BijectionIso.inj BijectionIsoBoolG = isInjectiveBoolGHom + BijectionIso.surj BijectionIsoBoolG = isSurjectiveBoolGHom + + Bool≅G : GroupIso BoolGr G + Bool≅G = BijectionIso→GroupIso BijectionIsoBoolG + +open import Cubical.Data.Sum +exact→Bool≅G± : (G : Group₀) (F : GroupHom ℤ ℤ) (H : GroupHom ℤ G) + (P : (fst F 1 ≡ 2) ⊎ (fst F 1 ≡ - 2)) + (ex : exact ℤ ℤ G UnitGr F H (→UnitHom G)) + → GroupIso BoolGr G +exact→Bool≅G± G F H (inl x) ex = exact→BoolIso.Bool≅G G F H x ex +exact→Bool≅G± G F H (inr x) ex = + exact→BoolIso.Bool≅G _ _ _ (cong (fst flip') x) exact' + where + flip' : GroupHom ℤ ℤ + fst flip' x = GroupStr.inv (snd ℤ) x + snd flip' = makeIsGroupHom (λ x y → + (+Comm (pos 0) (- (x +Z y)) + ∙ -Dist+ x y) + ∙ λ i → GroupStr.lid (snd ℤ) (- x) (~ i) +Z GroupStr.lid (snd ℤ) (- y) (~ i)) + + exact' : exact ℤ ℤ G UnitGr (compGroupHom F flip') H (→UnitHom G) + ImG→H⊂KerH→L exact' x inim = + (cong (fst H) (sym (GroupTheory.invInv ℤ x)) + ∙ (IsGroupHom.presinv (snd H) (fst flip' x))) + ∙∙ cong (GroupStr.inv (snd G)) (ImG→H⊂KerH→L ex _ (help inim)) + ∙∙ GroupTheory.inv1g G + where + help : isInIm (compGroupHom F flip') x → isInIm F (fst flip' x) + help = + pMap λ {(x , p) → x , + ((sym (-Involutive (fst F x)) + ∙ +Comm (- (- fst F x)) (pos 0)) + ∙ (λ i → pos 0 - (+Comm (- fst F x) (pos 0) i))) + ∙ λ i → pos 0 - p i } + KerH→L⊂ImG→H exact' x inker = + pMap (λ {(x , p) → (fst flip' x) , (GroupStr.lid (snd ℤ) (- (fst F (pos 0 - x))) + ∙ cong -_ (IsGroupHom.presinv (snd F) x + ∙ GroupStr.lid (snd ℤ) (- (fst F x))) + ∙ -Involutive (fst F x) ∙ p)}) + (KerH→L⊂ImG→H ex x inker) + ImH→L⊂KerL→R exact' = ImH→L⊂KerL→R ex + KerL→R⊂ImH→L exact' = KerL→R⊂ImH→L ex + + +exact→boolIsoGenPre : + (G H L : Group₀) (Z≅G : GroupEquiv ℤ G) (Z≅H : GroupEquiv ℤ H) + → (G→H : GroupHom G H) (H→L : GroupHom H L) + → ((invEq (fst Z≅H) (fst G→H (fst (fst Z≅G) 1)) ≡ 2) + ⊎ (invEq (fst Z≅H) (fst G→H (fst (fst Z≅G) 1)) ≡ - 2)) + → exact G H L UnitGr G→H H→L (→UnitHom L) + → GroupIso BoolGr L +exact→boolIsoGenPre G H L = + GroupEquivJ (λ G Z≅G + → (Z≅H : GroupEquiv ℤ H) + → (G→H : GroupHom G H) (H→L : GroupHom H L) + → ((invEq (fst Z≅H) (fst G→H (fst (fst Z≅G) 1)) ≡ 2) + ⊎ (invEq (fst Z≅H) (fst G→H (fst (fst Z≅G) 1)) ≡ - 2)) + → exact G H L UnitGr G→H H→L (→UnitHom L) + → GroupIso BoolGr L) + (GroupEquivJ (λ H Z≅H → + (G→H : GroupHom ℤ H) + (H→L : GroupHom H L) → + (invEq (fst Z≅H) (fst G→H (idfun (typ ℤ) (pos 1))) ≡ pos 2) ⊎ + (invEq (fst Z≅H) (fst G→H (idfun (typ ℤ) (pos 1))) ≡ negsuc 1) → + exact ℤ H L UnitGr G→H H→L (→UnitHom L) → GroupIso BoolGr L) + λ Z→Z L P ex → exact→Bool≅G± _ _ L P ex) + +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.HopfInvariant.HopfMap +open import Cubical.Homotopy.HopfInvariant.Homomorphism +open import Cubical.Homotopy.Group.Pi3S2 +open import Cubical.Homotopy.Group.PinSn + + +-- π₂S³-gen-by-HopfMap + +-- abs (HopfInvariant-π' 0 ([ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×)) ≡ 2 + +hopfInvariantEquiv : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤ +fst (fst hopfInvariantEquiv) = HopfInvariant-π' 0 +snd (fst hopfInvariantEquiv) = + GroupEquivℤ-isEquiv (invGroupEquiv π₃S²≅ℤ) ∣ HopfMap ∣₂ + π₂S³-gen-by-HopfMap (GroupHom-HopfInvariant-π' 0) + (abs→⊎ _ _ HopfInvariant-HopfMap) +snd hopfInvariantEquiv = snd (GroupHom-HopfInvariant-π' 0) + +π₃PushoutCharac : + abs (HopfInvariant-π' 0 [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π') ≡ 2 + → GroupIso BoolGr (π'Gr 2 thePushout∙) +π₃PushoutCharac p = + exact→boolIsoGenPre (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) + (GroupIso→GroupEquiv (invGroupIso (πₙ'Sⁿ≅ℤ 2))) + (invGroupEquiv hopfInvariantEquiv) + (π'∘∙Hom 2 (fold∘W , refl)) + (π'∘∙Hom 2 inr∙) + (abs→⊎ _ _ (cong abs + (cong (invEq (fst (invGroupEquiv hopfInvariantEquiv))) + (cong (fst (π'∘∙Hom 2 (fold∘W , refl))) + help) + ∙ cong (HopfInvariant-π' 0) + (Fold∘W ∙ sym (cong ∣_∣₂ ([]≡[]₂ (idfun∙ (S₊∙ 2)) (idfun∙ (S₊∙ 2)))))) ∙ p)) + S³→S²→Pushout→Unit + where + help : inv (fst (πₙ'Sⁿ≅ℤ 2)) 1 ≡ ∣ idfun∙ _ ∣₂ + help = cong (inv (fst (πₙ'Sⁿ≅ℤ 2))) (sym (πₙ'Sⁿ≅ℤ-idfun∙ 2)) + ∙ leftInv (fst (πₙ'Sⁿ≅ℤ 2)) ∣ idfun∙ _ ∣₂ + +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso + +∣HopfWhitehead∣≡2→π₄S³≅Bool : + abs (HopfInvariant-π' 0 [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π') ≡ 2 + → GroupEquiv (πGr 3 (S₊∙ 3)) BoolGr +∣HopfWhitehead∣≡2→π₄S³≅Bool p = + compGroupEquiv + (compGroupEquiv + (GroupIso→GroupEquiv + (compGroupIso π₄S³≅π₃PushS² + (invGroupIso (π'Gr≅πGr 2 (Pushout⋁↪fold⋁∙ (S₊∙ 2)))))) + (compGroupEquiv (invGroupEquiv (π'Iso 2 lem∙)) + (π'Iso 2 (lem₂ , sym (push north))))) + (invGroupEquiv (GroupIso→GroupEquiv (π₃PushoutCharac p))) + where + lem₂ : Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ ≃ fst thePushout∙ + lem₂ = compEquiv + (compEquiv pushoutSwitchEquiv (isoToEquiv (PushoutDistr.PushoutDistrIso fold⋁ W λ _ → tt))) + pushoutSwitchEquiv + + lem₁ : Pushout W (λ _ → tt) ≃ cofibW S¹ S¹ base base + lem₁ = pushoutEquiv W (λ _ → tt) joinTo⋁ (λ _ → tt) + (isoToEquiv (invIso (IsoSphereJoin 1 1))) + (idEquiv _) + (idEquiv _) + refl + refl + + lem : Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ + ≃ fst (Pushout⋁↪fold⋁∙ (S₊∙ 2)) + lem = pushoutEquiv inl _ ⋁↪ fold⋁ + (idEquiv _) + (compEquiv lem₁ (isoToEquiv (invIso (Iso-Susp×Susp-cofibJoinTo⋁ S¹ S¹ base base)))) + (idEquiv _) + (Susp×Susp→cofibW≡ S¹ S¹ base base) + refl + + lem∙ : (Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ , inr north) + ≃∙ (Pushout⋁↪fold⋁∙ (S₊∙ 2)) + fst lem∙ = lem + snd lem∙ = sym (push (inl north)) diff --git a/Cubical/Homotopy/Group/PinSn.agda b/Cubical/Homotopy/Group/PinSn.agda new file mode 100644 index 0000000000..15dbfccf8c --- /dev/null +++ b/Cubical/Homotopy/Group/PinSn.agda @@ -0,0 +1,334 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains: +1. An iso πₙ'(Sⁿ) ≅ ℤ, where πₙ'(Sⁿ) = ∥ Sⁿ →∙ Sⁿ ∥₀ +2. A proof that idfun∙ : Sⁿ →∙ Sⁿ is the generator of πₙ'(Sⁿ) +-} +module Cubical.Homotopy.Group.PinSn where + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.SuspensionMap +open import Cubical.Homotopy.Connected + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism + +open import Cubical.HITs.SetTruncation + renaming (elim to sElim ; elim2 to sElim2 + ; map to sMap) +open import Cubical.HITs.Truncation renaming + (elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Int renaming (_+_ to _+ℤ_) + +open import Cubical.ZCohomology.Properties + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.ZAction + +open Iso + +-- The goal is to prove that πₙSⁿ ≅ ℤ. This is of course trivial, given +-- that ΩⁿK(ℤ,n) ≅ ℤ is already proved. However, we would like to do +-- this for πₙSⁿ defined via (Sⁿ →∙ Sⁿ) and prove that the generator +-- of this group is idfun∙ : Sⁿ → Sⁿ. + +private + suspMapS^ : (n : ℕ) → (S₊∙ (suc n) →∙ S₊∙ (suc n)) + → S₊∙ (2 + n) →∙ S₊∙ (2 + n) + suspMapS^ n = suspMap {A = S₊∙ (suc n)} n + +is2ConnectedSuspMap : (n : ℕ) → isConnectedFun 2 (suspMapS^ (suc n)) +is2ConnectedSuspMap n = + isConnectedFunSubtr 2 n _ + (subst (λ x → isConnectedFun x (suspMapS^ (suc n))) + (subtrLem n (suc (suc n)) ∙ +-comm 2 n) + (isConnectedSuspMap (suc n) (suc n))) + where + subtrLem : (n m : ℕ) → (n + m ∸ n) ≡ m + subtrLem zero m = refl + subtrLem (suc n) m = subtrLem n m + +-- From the connectedness of the suspension map, +-- we get an iso πₙSⁿ ≅ πₙ₊₁(Sⁿ⁺¹) for n ≥ 2. +SphereSuspIso : (n : ℕ) + → Iso (π' (2 + n) (S₊∙ (2 + n))) (π' (3 + n) (S₊∙ (3 + n))) +SphereSuspIso n = + compIso setTruncTrunc2Iso + (compIso + (connectedTruncIso 2 + (suspMap {A = S₊∙ (suc (suc n))} (suc n)) (is2ConnectedSuspMap n)) + (invIso (setTruncTrunc2Iso))) + +SphereSuspGrIso : (n : ℕ) + → GroupIso (π'Gr (suc n) (S₊∙ (2 + n))) (π'Gr (2 + n) (S₊∙ (3 + n))) +fst (SphereSuspGrIso n) = SphereSuspIso n +snd (SphereSuspGrIso n) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → IsGroupHom.pres· (suspMapπ'Hom (suc n) .snd) ∣ f ∣₂ ∣ g ∣₂) + +private + stLoop₁ : π 2 (S₊∙ 2) + stLoop₁ = ∣ sym (rCancel (merid base)) + ∙∙ (λ i → merid (loop i) ∙ sym (merid base)) + ∙∙ rCancel (merid base) ∣₂ + + stLoop₁flip : π 2 (S₊∙ 2) + stLoop₁flip = sMap flipSquare stLoop₁ + +flipSquareIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupIso (πGr (suc n) A) (πGr (suc n) A) +fun (fst (flipSquareIso n)) = sMap flipSquare +inv (fst (flipSquareIso n)) = sMap flipSquare +rightInv (fst (flipSquareIso n)) = + sElim (λ _ → isSetPathImplicit) λ _ → refl +leftInv (fst (flipSquareIso n)) = + sElim (λ _ → isSetPathImplicit) λ _ → refl +snd (flipSquareIso n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ + ((sym (sym≡flipSquare (f ∙ g)) + ∙∙ symDistr f g + ∙∙ cong₂ _∙_ (sym≡flipSquare g) (sym≡flipSquare f) + ∙ EH n (flipSquare g) (flipSquare f)))) + +-- We now want to prove π₁S¹≅π₂S². This will be done by passing through +-- The loop space defintion of the groups +π'₂S²≅π₂S² : GroupIso (π'Gr 1 (S₊∙ 2)) (πGr 1 (S₊∙ 2)) +π'₂S²≅π₂S² = π'Gr≅πGr 1 (S₊∙ 2) + +{- We now define the iso π₂S² ≅ π₁S¹ -} +π₂S²≅π₁S¹ : GroupIso (πGr 1 (S₊∙ 2)) (πGr 0 (S₊∙ 1)) +fst π₂S²≅π₁S¹ = + compIso setTruncTrunc2Iso + (compIso + (compIso (invIso (PathIdTruncIso 2)) + (compIso (congIso (invIso (PathIdTruncIso 3))) + (compIso + (congIso (invIso (Iso-Kn-ΩKn+1 1))) + (PathIdTruncIso 2)))) + (invIso setTruncTrunc2Iso)) +snd π₂S²≅π₁S¹ = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → + cong (inv setTruncTrunc2Iso) + (cong (fun (PathIdTruncIso 2)) + (cong (cong (ΩKn+1→Kn 1)) + (cong (cong (inv (PathIdTruncIso 3))) + (cong (inv (PathIdTruncIso 2)) + (refl {x = ∣ f ∙ g ∣}) + ∙ cong-∙ ∣_∣ₕ f g) + ∙ cong-∙ (inv (PathIdTruncIso 3)) (cong ∣_∣ₕ f) (cong ∣_∣ₕ g)) + ∙ cong-∙ (ΩKn+1→Kn 1) + (cong (inv (PathIdTruncIso 3)) (cong ∣_∣ₕ f)) + (cong (inv (PathIdTruncIso 3)) (cong ∣_∣ₕ g))) + ∙ PathIdTruncIsoFunct 1 + (cong (ΩKn+1→Kn 1) (λ i → inv (PathIdTruncIso 3) ∣ f i ∣ₕ)) + (cong (ΩKn+1→Kn 1) (λ i → inv (PathIdTruncIso 3) ∣ g i ∣ₕ))) + ∙ setTruncTrunc2IsoFunct + ((fun (PathIdTruncIso 2)) + (cong (ΩKn+1→Kn 1) (λ i → inv (PathIdTruncIso 3) ∣ f i ∣ₕ))) + ((fun (PathIdTruncIso 2)) + (cong (ΩKn+1→Kn 1) (λ i → inv (PathIdTruncIso 3) ∣ g i ∣ₕ)))) + where + setTruncTrunc2IsoFunct : ∀ {ℓ} {A : Type ℓ} {x : A} + (p q : hLevelTrunc 2 (x ≡ x)) + → inv setTruncTrunc2Iso + (Cubical.HITs.Truncation.map2 _∙_ p q) + ≡ ·π 0 (inv setTruncTrunc2Iso p) (inv setTruncTrunc2Iso q) + setTruncTrunc2IsoFunct = + trElim2 (λ _ _ → isSetPathImplicit) λ _ _ → refl + +π₂'S²≅π₁'S¹ : GroupIso (π'Gr 1 (S₊∙ 2)) (π'Gr 0 (S₊∙ 1)) +π₂'S²≅π₁'S¹ = + compGroupIso (π'Gr≅πGr 1 (S₊∙ 2)) + (compGroupIso (flipSquareIso 0) + (compGroupIso π₂S²≅π₁S¹ + (invGroupIso (π'Gr≅πGr 0 (S₊∙ 1))))) + +-- We get our iso +πₙ'Sⁿ≅ℤ : (n : ℕ) → GroupIso (π'Gr n (S₊∙ (suc n))) ℤGroup +πₙ'Sⁿ≅ℤ zero = + compGroupIso (π'Gr≅πGr zero (S₊∙ 1)) + ((compIso (setTruncIdempotentIso (isGroupoidS¹ _ _)) ΩS¹Isoℤ) + , makeIsGroupHom (sElim2 (λ _ _ → isProp→isSet (isSetℤ _ _)) + winding-hom)) +πₙ'Sⁿ≅ℤ (suc zero) = compGroupIso π₂'S²≅π₁'S¹ (πₙ'Sⁿ≅ℤ zero) +πₙ'Sⁿ≅ℤ (suc (suc n)) = + compGroupIso (invGroupIso (SphereSuspGrIso n)) (πₙ'Sⁿ≅ℤ (suc n)) + +-- Same thing for the loop space definition +πₙSⁿ≅ℤ : (n : ℕ) → GroupIso (πGr n (S₊∙ (suc n))) ℤGroup +πₙSⁿ≅ℤ n = + compGroupIso (invGroupIso (π'Gr≅πGr n (S₊∙ (suc n)))) (πₙ'Sⁿ≅ℤ n) + +-- The goal now is to show that πₙ'Sⁿ≅ℤ takes idfun∙ : Sⁿ → Sⁿ to 1. +-- For this, we need a bunch of identities: +private + Isoπ₁S¹ℤ = (compIso (setTruncIdempotentIso (isGroupoidS¹ _ _)) ΩS¹Isoℤ) + + π'₂S²≅π₂S²⁻-stLoop' : inv (fst (π'₂S²≅π₂S²)) stLoop₁flip ≡ ∣ idfun∙ _ ∣₂ + π'₂S²≅π₂S²⁻-stLoop' = + cong ∣_∣₂ (ΣPathP ((funExt + (λ { north → refl + ; south → merid base + ; (merid base i) j → + hcomp (λ k + → λ {(i = i0) → north + ; (i = i1) → merid base (j ∧ k) + ; (j = i0) → rUnit (λ _ → north) k i + ; (j = i1) → merid base (i ∧ k)}) + north + ; (merid (loop k) i) j + → hcomp (λ r + → λ {(i = i0) → north + ; (i = i1) → merid base (j ∧ r) + ; (j = i0) → + rUnit (funExt⁻ (cong fst (cong (Ω→SphereMap 1) + (flipSquare (sym (rCancel (merid base)) + ∙∙ (λ i₁ → merid (loop i₁) + ∙ sym (merid base)) + ∙∙ rCancel (merid base))))) (loop k)) r i + ; (j = i1) → lem₂ r i k}) + (((sym (rCancel (merid base)) ∙∙ + (λ i₁ → merid (loop i₁) ∙ sym (merid base)) ∙∙ + rCancel (merid base))) k i)})) , refl)) + where + genBot+side : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) + → Cube {A = A} (λ j r → x) (λ j r → p (~ j ∨ r)) + (λ i r → p i) (λ i r → p (i ∧ r)) + (λ i j → p (i ∧ ~ j)) (λ i j → p i) + × Cube {A = A} (λ j r → p (~ j ∨ r)) (λ j r → p (r ∧ j)) + (λ k r → p (~ k)) (λ k r → p r) + (λ k j → p (~ k ∧ ~ j)) λ k j → p (j ∨ ~ k) + genBot+side {A = A} {x = x} = + J (λ y p → Cube {A = A} (λ j r → x) (λ j r → p (~ j ∨ r)) + (λ i r → p i) (λ i r → p (i ∧ r)) + (λ i j → p (i ∧ ~ j)) (λ i j → p i) + × Cube {A = A} (λ j r → p (~ j ∨ r)) (λ j r → p (r ∧ j)) + (λ k r → p (~ k)) (λ k r → p r) + (λ k j → p (~ k ∧ ~ j)) λ k j → p (j ∨ ~ k)) + (refl , refl) + + lem₁ : I → I → I → S₊ 2 + lem₁ j i r = + hcomp (λ k → λ {(i = i0) → north + ; (i = i1) → genBot+side (merid base) .snd k j r + ; (j = i0) → compPath-filler + (merid base) (sym (merid base)) k i + ; (j = i1) → merid base (i ∧ r) + ; (r = i0) → rCancel-filler (merid base) k j i + ; (r = i1) → compPath-filler + (merid base) (sym (merid base)) (~ j ∧ k) i}) + (genBot+side (merid base) .fst i j r) + + lem₂ : I → I → I → S₊ 2 + lem₂ r i k = + hcomp (λ j → λ {(i = i0) → north + ; (i = i1) → merid base (r ∧ j) + ; (r = i0) → doubleCompPath-filler + (sym (rCancel (merid base))) + (λ i₁ → merid (loop i₁) ∙ sym (merid base)) + (rCancel (merid base)) j k i + ; (r = i1) → compPath-filler + (merid (loop k)) (sym (merid base)) (~ j) i + ; (k = i0) → lem₁ j i r + ; (k = i1) → lem₁ j i r}) + ((merid (loop k) ∙ sym (merid base)) i) + + π₂S²≅π₁S¹-stLoop : fun (fst π₂S²≅π₁S¹) stLoop₁ ≡ ∣ loop ∣₂ + π₂S²≅π₁S¹-stLoop = + sym (leftInv Isoπ₁S¹ℤ + (fun (fst π₂S²≅π₁S¹) stLoop₁)) + ∙∙ cong (inv Isoπ₁S¹ℤ) compute + ∙∙ leftInv (compIso (setTruncIdempotentIso (isGroupoidS¹ _ _)) ΩS¹Isoℤ) + ∣ loop ∣₂ + where + compute : fun Isoπ₁S¹ℤ (fun (fst π₂S²≅π₁S¹) stLoop₁) + ≡ fun Isoπ₁S¹ℤ ∣ loop ∣₂ + compute = refl + + π₂'S²≅π₁'S¹≡ : fun (fst π₂'S²≅π₁'S¹) ∣ idfun∙ _ ∣₂ ≡ ∣ idfun∙ _ ∣₂ + π₂'S²≅π₁'S¹≡ = + lem₁ ∣ idfun∙ _ ∣₂ stLoop₁ + (sym π'₂S²≅π₂S²⁻-stLoop') + (cong (inv (fst (π'Gr≅πGr zero (S₊∙ 1)))) π₂S²≅π₁S¹-stLoop + ∙ lem₂) + where + lem₁ : (x : _) (y : π 2 (S₊∙ 2)) + → (x ≡ inv (fst π'₂S²≅π₂S²) (fun (fst (flipSquareIso 0)) y)) + → inv (fst (π'Gr≅πGr zero (S₊∙ 1))) (fun (fst π₂S²≅π₁S¹) y) + ≡ ∣ idfun∙ _ ∣₂ + → fun (fst π₂'S²≅π₁'S¹) x ≡ ∣ idfun∙ _ ∣₂ + lem₁ x y p q = + cong (fun (fst π₂'S²≅π₁'S¹)) p + ∙∙ (λ i → inv (fst (π'Gr≅πGr zero (S₊∙ (suc zero)))) (fun (fst π₂S²≅π₁S¹) + (fun (fst (flipSquareIso zero)) + (rightInv + (fst (π'Gr≅πGr (suc zero) (S₊∙ (suc (suc zero))))) + (inv (fst (flipSquareIso zero)) y) i) + ))) + ∙∙ cong (inv (fst (π'Gr≅πGr zero (S₊∙ (suc zero))))) + (cong (fun (fst π₂S²≅π₁S¹)) + (rightInv (fst (flipSquareIso zero)) y)) + ∙ q + + lem₂ : inv (fst (π'Gr≅πGr zero (S₊∙ 1))) ∣ loop ∣₂ ≡ ∣ idfun∙ _ ∣₂ + lem₂ = + cong ∣_∣₂ (ΣPathP (funExt (λ { base → refl ; (loop i) → refl}) , refl)) + + suspPresIdfun : (n : ℕ) → suspMap n (idfun∙ (S₊∙ (suc n))) ≡ idfun∙ _ + suspPresIdfun n = + ΣPathP ((funExt + (λ { north → refl + ; south → merid (ptSn (suc n)) + ; (merid a i) j + → compPath-filler (merid a) (sym (merid (ptSn (suc n)))) (~ j) i})) + , refl) + + suspPresIdfun2 : (n : ℕ) + → fun (fst (invGroupIso (SphereSuspGrIso n))) + ∣ idfun∙ (S₊∙ (suc (suc (suc n)))) ∣₂ + ≡ ∣ idfun∙ _ ∣₂ + suspPresIdfun2 n = + sym (cong (fun (fst (invGroupIso (SphereSuspGrIso n)))) + (cong ∣_∣₂ (suspPresIdfun (suc n)))) + ∙ leftInv (SphereSuspIso n) ∣ idfun∙ _ ∣₂ + +-- We finally have our results +πₙ'Sⁿ≅ℤ-idfun∙ : (n : ℕ) + → fun (fst (πₙ'Sⁿ≅ℤ n)) ∣ idfun∙ _ ∣₂ ≡ (pos (suc zero)) +πₙ'Sⁿ≅ℤ-idfun∙ zero = refl +πₙ'Sⁿ≅ℤ-idfun∙ (suc zero) = speedUp ∣ idfun∙ _ ∣₂ π₂'S²≅π₁'S¹≡ + where + speedUp : (x : _) -- stated like this for faster type checking + → fun (fst π₂'S²≅π₁'S¹) x ≡ ∣ idfun∙ _ ∣₂ + → fun (fst (πₙ'Sⁿ≅ℤ (suc zero))) x ≡ pos (suc zero) + speedUp x p = cong (fun (fst (πₙ'Sⁿ≅ℤ zero))) p +πₙ'Sⁿ≅ℤ-idfun∙ (suc (suc n)) = + cong (fun (fst (πₙ'Sⁿ≅ℤ (suc n)))) (suspPresIdfun2 n) + ∙ πₙ'Sⁿ≅ℤ-idfun∙ (suc n) + +πₙ'Sⁿ-gen-by-idfun : (n : ℕ) → gen₁-by (π'Gr n (S₊∙ (suc n))) ∣ idfun∙ _ ∣₂ +πₙ'Sⁿ-gen-by-idfun n = + subst (gen₁-by (π'Gr n (S₊∙ (suc n)))) + (sym (cong (inv (fst (πₙ'Sⁿ≅ℤ n))) (πₙ'Sⁿ≅ℤ-idfun∙ n)) + ∙ leftInv (fst (πₙ'Sⁿ≅ℤ n)) ∣ idfun∙ _ ∣₂) + (Iso-pres-gen₁ ℤGroup (π'Gr n (S₊∙ (suc n))) + (pos (suc zero)) + (λ h → h , (sym (·Comm h (pos 1)) ∙ ℤ·≡· h (pos 1))) + (invGroupIso (πₙ'Sⁿ≅ℤ n))) diff --git a/Cubical/Homotopy/Group/SuspensionMap.agda b/Cubical/Homotopy/Group/SuspensionMap.agda new file mode 100644 index 0000000000..e1f75a1f39 --- /dev/null +++ b/Cubical/Homotopy/Group/SuspensionMap.agda @@ -0,0 +1,608 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Group.SuspensionMap where + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Freudenthal +open import Cubical.Homotopy.Connected + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport + +open import Cubical.Functions.Morphism + +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp renaming (toSusp to σ ; toSuspPointed to σ∙) +open import Cubical.HITs.S1 + +open import Cubical.Data.Bool hiding (_≟_) +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + + +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; rec2 to pRec2 ; elim to pElim) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim + ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) +open import Cubical.HITs.Truncation + renaming (rec to trRec) + +open Iso +open IsGroup +open IsSemigroup +open IsMonoid +open GroupStr + +{- +This file concerns the suspension maps +suspMapΩ : πₙA → πₙ₊₁ΣA and +suspMap : π'ₙA → π'ₙ₊₁ΣA +For instance, we want to transport freudenthal +for suspMapΩ to suspMap by establishing a dependent +path between the two functions. + +This gives, in particular, surjectivity of +πₙ₊₁(Sⁿ) → πₙ₊₂(Sⁿ⁺¹) for n ≥ 2. +-} + +-- Definition of the suspension functions +suspMap : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) + → S₊∙ (suc n) →∙ A + → S₊∙ (suc (suc n)) →∙ Susp∙ (typ A) +fst (suspMap n f) north = north +fst (suspMap n f) south = north +fst (suspMap {A = A} n f) (merid a i) = + (merid (fst f a) ∙ sym (merid (pt A))) i +snd (suspMap n f) = refl + +suspMapΩ∙ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → ((Ω^ n) A) + →∙ ((Ω^ (suc n)) (Susp∙ (typ A))) +fst (suspMapΩ∙ {A = A} zero) a = merid a ∙ sym (merid (pt A)) +snd (suspMapΩ∙ {A = A} zero) = rCancel (merid (pt A)) +suspMapΩ∙ {A = A} (suc n) = Ω→ (suspMapΩ∙ {A = A} n) + +suspMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → typ ((Ω^ n) A) → typ ((Ω^ (suc n)) (Susp∙ (typ A))) +suspMapΩ {A = A} n = suspMapΩ∙ {A = A} n .fst + + +suspMapπ' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → π' (suc n) A + → π' (suc (suc n)) (Susp∙ (typ A)) +suspMapπ' n = sMap (suspMap n) + +suspMapπ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → π n A → π (suc n) (Susp∙ (typ A)) +suspMapπ n = sMap (suspMapΩ n) + +{- suspMapΩ + Ωⁿ A --------------------> Ω¹⁺ⁿ (Susp A) + | | + | = | ≃ flipΩ + | Ωⁿ→ σ v +Ωⁿ A -------------------> Ωⁿ (Ω (Susp A)) + | | + | | + | ≃ Ω→SphereMap | ≃ Ω→SphereMap + | | + v post∘∙ . σ v + (Sⁿ →∙ A) -------------- > (Sⁿ →∙ Ω (Susp A)) + | | + | = | ≃ botᵣ + | | + v suspMap v + (Sⁿ →∙ A) -------------- > (Sⁿ⁺¹→∙ Susp A) + -} + +botᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ n →∙ Ω (Susp∙ (typ A))) + → S₊∙ (suc n) →∙ Susp∙ (typ A) +fst (botᵣ zero (f , p)) base = north +fst (botᵣ zero (f , p)) (loop i) = f false i +snd (botᵣ zero (f , p)) = refl +fst (botᵣ (suc n) (f , p)) north = north +fst (botᵣ (suc n) (f , p)) south = north +fst (botᵣ (suc n) (f , p)) (merid a i) = f a i +snd (botᵣ (suc n) (f , p)) = refl + +----- Top square filler ----- +top□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Ω^→ (suc n) (σ∙ A) + ≡ (((Iso.fun (flipΩIso (suc n))) , flipΩrefl n) + ∘∙ suspMapΩ∙ (suc n)) +top□ {A = A} zero = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ p → sym (transportRefl _)) +top□ {A = A} (suc n) = + cong Ω→ (top□ {A = A} n) + ∙ →∙Homogeneous≡ + (isHomogeneousPath _ _) + (funExt λ x + → Ω→∘ (fun (flipΩIso (suc n)) , flipΩrefl n) (suspMapΩ∙ (suc n)) x) + +----- Middle square filler ----- +module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + (homogB : isHomogeneous B) (f : A →∙ B) where + nat = isNaturalΩSphereMap A B homogB f + mutual + isNatural-Ω→SphereMap : ∀ n p → f ∘∙ Ω→SphereMap n p + ≡ Ω→SphereMap n (Ω^→ n f .fst p) + isNatural-Ω→SphereMap 0 p = + →∙Homogeneous≡ homogB (funExt λ {true → f .snd; false → refl}) + isNatural-Ω→SphereMap (n@(suc n')) p = + cong (f ∘∙_) (Ω→SphereMap-split n' p) + ∙ nat n' (Ω→SphereMapSplit₁ n' p) + ∙ cong (ΩSphereMap n') inner + ∙ sym (Ω→SphereMap-split n' (Ω^→ n f .fst p)) + where + inner : Ω→ (post∘∙ (S₊∙ n') f) .fst (Ω→ (Ω→SphereMap∙ n') .fst p) + ≡ Ω→ (Ω→SphereMap∙ n') .fst (Ω^→ (suc n') f .fst p) + inner = + sym (Ω→∘ (post∘∙ (S₊∙ n') f) (Ω→SphereMap∙ n') p) + ∙ cong (λ g∙ → Ω→ g∙ .fst p) (isNatural-Ω→SphereMap∙ n') + ∙ Ω→∘ (Ω→SphereMap∙ n') (Ω^→ n' f) p + + isNatural-Ω→SphereMap∙ : ∀ n + → post∘∙ (S₊∙ n) f ∘∙ (Ω→SphereMap∙ n) + ≡ (Ω→SphereMap∙ n {A = B} ∘∙ Ω^→ n f) + isNatural-Ω→SphereMap∙ n = + →∙Homogeneous≡ (isHomogeneous→∙ homogB) + (funExt (isNatural-Ω→SphereMap n)) + +mid□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (p : typ ((Ω^ (suc n)) A)) + → fst (post∘∙ (S₊∙ (suc n)) (σ∙ A)) (Ω→SphereMap (suc n) p) + ≡ Ω→SphereMap (suc n) (fst (Ω^→ (suc n) (σ∙ A)) p) +mid□ {A = A} n p = + funExt⁻ + (cong fst + (isNatural-Ω→SphereMap∙ + A (Ω (Susp∙ (typ A))) + (isHomogeneousPath _ _) + (σ∙ A) (suc n))) p + +----- Bottom square filler ----- +bot□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f : (S₊∙ (suc n) →∙ A)) + → suspMap n f + ≡ botᵣ {A = A} (suc n) (post∘∙ (S₊∙ (suc n)) (σ∙ A) .fst f) +bot□ {A = A} n f = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) → refl})) + , refl) + +-- We prove that botᵣ is an equivalence +botᵣ⁻' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ (suc n) →∙ Susp∙ (typ A) → (S₊ n → typ (Ω (Susp∙ (typ A)))) +botᵣ⁻' zero f false = + sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f +botᵣ⁻' zero f true = refl +botᵣ⁻' (suc n) f x = + sym (snd f) + ∙∙ cong (fst f) (merid x ∙ sym (merid (ptSn (suc n)))) + ∙∙ snd f + +botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ (suc n) →∙ Susp∙ (typ A) + → (S₊∙ n →∙ Ω (Susp∙ (typ A))) +fst (botᵣ⁻ {A = A} n f) = botᵣ⁻' {A = A} n f +snd (botᵣ⁻ {A = A} zero f) = refl +snd (botᵣ⁻ {A = A} (suc n) f) = + cong (sym (snd f) ∙∙_∙∙ snd f) + (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙ ∙∙lCancel (snd f) + +botᵣIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ Ω (Susp∙ (typ A))) + (S₊∙ (suc n) →∙ Susp∙ (typ A)) +botᵣIso {A = A} n = (iso (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) (h n) (retr n)) + where + h : (n : ℕ) → section (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) + h zero (f , p) = + ΣPathP (funExt (λ { base → sym p + ; (loop i) j → doubleCompPath-filler + (sym p) (cong f loop) p (~ j) i}) + , λ i j → p (~ i ∨ j)) + h (suc n) (f , p) = + ΣPathP (funExt (λ { north → sym p + ; south → sym p ∙ cong f (merid (ptSn _)) + ; (merid a i) j + → hcomp (λ k + → λ { (i = i0) → p (~ j ∧ k) + ; (i = i1) → compPath-filler' + (sym p) (cong f (merid (ptSn _))) k j + ; (j = i1) → f (merid a i)}) + (f (compPath-filler + (merid a) (sym (merid (ptSn _))) (~ j) i))}) + , λ i j → p (~ i ∨ j)) + + retr : (n : ℕ) → retract (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) + retr zero (f , p) = + ΣPathP ((funExt (λ { false → sym (rUnit _) ; true → sym p})) + , λ i j → p (~ i ∨ j)) + retr (suc n) (f , p) = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ x → (λ i + → rUnit (cong-∙ (fst ((botᵣ {A = A}(suc n) (f , p)))) + (merid x) + (sym (merid (ptSn (suc n)))) i) (~ i)) + ∙∙ (λ i → f x ∙ sym (p i) ) + ∙∙ sym (rUnit (f x))) + +-- Right hand composite iso +IsoΩSphereMapᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (typ ((Ω^ (suc n)) (Susp∙ (typ A)))) ((S₊∙ (suc n) →∙ Susp∙ (typ A))) +IsoΩSphereMapᵣ {A = A} n = + compIso (flipΩIso n) + (compIso (IsoΩSphereMap n) (botᵣIso {A = A} n)) + +-- The dependent path between the two suspension functions +suspMapPathP : ∀ {ℓ} (A : Pointed ℓ) (n : ℕ) + → (typ ((Ω^ (suc n)) A) → (typ ((Ω^ (suc (suc n))) (Susp∙ (typ A))))) + ≡ ((S₊∙ (suc n) →∙ A) → S₊∙ (suc (suc n)) →∙ (Susp∙ (typ A))) +suspMapPathP A n i = + isoToPath (IsoΩSphereMap {A = A} (suc n)) i + → isoToPath (IsoΩSphereMapᵣ {A = A} (suc n)) i + +Ωσ→suspMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → PathP (λ i → suspMapPathP A n i) + (suspMapΩ (suc n)) + (suspMap n) +Ωσ→suspMap {A = A} n = + toPathP (funExt (λ p → + (λ i → transportRefl + (Iso.fun (IsoΩSphereMapᵣ {A = A} (suc n)) + (suspMapΩ {A = A} (suc n) + (Iso.inv (IsoΩSphereMap {A = A} (suc n)) + (transportRefl p i)))) i) + ∙∙ cong (botᵣ {A = A} (suc n)) + (cong (Ω→SphereMap (suc n) {A = Ω (Susp∙ (typ A)) }) + ((sym (funExt⁻ (cong fst (top□ {A = A} n)) + (invEq (Ω→SphereMap (suc n) + , isEquiv-Ω→SphereMap (suc n)) p)))) + ∙ (sym (mid□ n (invEq (Ω→SphereMap (suc n) + , isEquiv-Ω→SphereMap (suc n)) p)) + ∙ cong (σ∙ (fst A , snd A) ∘∙_) + (secEq (Ω→SphereMap (suc n) + , isEquiv-Ω→SphereMap (suc n)) p))) + ∙∙ sym (bot□ n p))) + +-- Connectedness of suspFunΩ (Freudenthal) +suspMapΩ-connected : ∀ {ℓ} (n : HLevel) (m : ℕ) {A : Pointed ℓ} + (connA : isConnected (suc (suc n)) (typ A)) + → isConnectedFun ((suc n + suc n) ∸ m) (suspMapΩ {A = A} m) +suspMapΩ-connected n zero {A = A} connA = isConnectedσ n connA +suspMapΩ-connected n (suc m) {A = A} connA with ((n + suc n) ≟ m) +... | (lt p) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) + (sym (n∸m≡0 _ m p)) + λ b → tt* , (λ {tt* → refl}) +... | (eq q) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) + (sym (n∸n≡0 m) ∙ cong (_∸ m) (sym q)) + λ b → tt* , (λ {tt* → refl}) +... | (gt p) = isConnectedCong' (n + suc n ∸ m) (suspMapΩ {A = A} m) + (subst (λ x → isConnectedFun x (suspMapΩ {A = A} m)) + (sym (suc∸-fst (n + suc n) m p)) + (suspMapΩ-connected n m connA)) + (snd (suspMapΩ∙ m)) + +-- We prove that the right iso is structure preserving +private + invComp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) + fst (invComp n f g) x = (fst f x) ∙ (fst g x) + snd (invComp n f g) = cong₂ _∙_ (snd f) (snd g) ∙ sym (rUnit refl) + + -- We prove that it agrees with ∙Π + ∙Π≡invComp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc n) →∙ Ω (Susp∙ (typ A))) + → ∙Π f g ≡ invComp {A = A} (suc n) f g + ∙Π≡invComp zero f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt (λ { base → rUnit refl + ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) + ∙ cong (refl ∙_) (snd g)) + ; (loop i) j → + hcomp (λ k → + λ { (i = i0) → (rUnit refl + ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) + ∙ cong (refl ∙_) (snd g))) j + ; (i = i1) → (rUnit refl + ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) + ∙ cong (refl ∙_) (snd g))) j + ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i + ; (j = i1) → cong₂Funct _∙_ + (cong (fst f) loop) (cong (fst g) loop) (~ k) i}) + (hcomp (λ k → + λ { (i = i0) → (rUnit refl + ∙ sym (cong (_∙ snd g (~ k)) (λ j → snd f (j ∨ ~ k)) + ∙ cong (refl ∙_) (λ j → snd g (j ∨ ~ k)))) j + ; (i = i1) → (rUnit refl + ∙ sym (cong (_∙ snd g (~ k)) (λ j → snd f (j ∨ ~ k)) + ∙ cong (refl ∙_) (λ j → snd g (j ∨ ~ k)))) j + ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i + ; (j = i1) → (cong (_∙ snd g (~ k)) + (doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) + (snd f) (~ k)) ∙ + cong ((snd f (~ k)) ∙_) + (doubleCompPath-filler (sym (snd g)) + (cong (fst g) loop) (snd g) (~ k))) i}) + (hcomp (λ k → + λ { (i = i0) → (rUnit (rUnit refl) + ∙ cong (rUnit refl ∙_) (cong sym (rUnit refl))) k j + ; (i = i1) → (rUnit (rUnit refl) + ∙ cong (rUnit refl ∙_) (cong sym (rUnit refl))) k j + ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i + ; (j = i1) → (cong (_∙ refl) + ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + ∙ cong (refl ∙_) + (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i}) + ((cong (λ x → rUnit x j) + (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ cong (λ x → lUnit x j) + (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i)))})) + ∙Π≡invComp {A = A} (suc n) f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ { north → rUnit refl + ∙ sym (cong (fst f north ∙_) (snd g) + ∙ cong (_∙ refl) (snd f)) + ; south → rUnit refl + ∙ sym (cong₂ _∙_ + (cong (fst f) (sym (merid (ptSn _))) ∙ snd f) + (cong (fst g) (sym (merid (ptSn _))) ∙ snd g)) + ; (merid a i) j → p a i j}) + where + module _ (a : S₊ (suc n)) where + f-n = fst f north + g-n = fst g north + cong-f = (cong (fst f) (merid a ∙ sym (merid (ptSn _)))) + cong-g = (cong (fst g) (merid a ∙ sym (merid (ptSn _)))) + c-f = sym (snd f) ∙∙ cong-f ∙∙ snd f + c-g = sym (snd g) ∙∙ cong-g ∙∙ snd g + + p : I → I → fst (Ω (Susp∙ (typ A))) + p i j = + hcomp (λ k → + λ { (i = i0) → (rUnit (λ _ → snd (Susp∙ (typ A))) + ∙ sym ((cong (fst f north ∙_) (snd g) + ∙ cong (_∙ refl) (snd f)))) j + ; (i = i1) → (rUnit refl + ∙ sym (cong₂ _∙_ + (compPath-filler' + (cong (fst f) (sym (merid (ptSn _)))) (snd f) k) + (compPath-filler' + (cong (fst g) (sym (merid (ptSn _)))) (snd g) k))) j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → fst f (compPath-filler + (merid a) + (sym (merid (ptSn _))) (~ k) i) + ∙ fst g (compPath-filler + (merid a) + (sym (merid (ptSn _))) (~ k) i)}) + (hcomp (λ k → + λ {(i = i0) → (rUnit (λ _ → snd (Susp∙ (typ A))) + ∙ sym ((cong (fst f north ∙_) (snd g) + ∙ cong (_∙ refl) (snd f)))) j + ; (i = i1) → (rUnit refl ∙ sym (cong₂ _∙_ (snd f) (snd g))) j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → cong₂Funct _∙_ cong-f cong-g (~ k) i}) + (hcomp (λ k → + λ {(i = i0) → (rUnit refl + ∙ sym (compPath-filler' + ((cong (fst f north ∙_) (snd g))) + (cong (_∙ refl) (snd f)) k)) j + ; (i = i1) → (rUnit refl + ∙ sym (cong₂ _∙_ (λ i → snd f (i ∨ ~ k)) (snd g))) j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → (cong (λ x → x ∙ snd g (~ k)) + (doubleCompPath-filler refl + cong-f (snd f) (~ k)) + ∙ cong ((snd f (~ k)) ∙_) + (doubleCompPath-filler + (sym (snd g)) cong-g refl (~ k))) i}) + (hcomp (λ k → + λ {(i = i0) → compPath-filler + (rUnit (λ _ → snd (Susp∙ (typ A)))) + (sym ((cong (_∙ refl) (snd f)))) k j + ; (i = i1) → compPath-filler + (rUnit refl) + (sym (cong (refl ∙_) (snd g))) k j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → (cong (_∙ refl) + ((λ j → snd f (~ j ∧ ~ k)) ∙∙ cong-f ∙∙ snd f) + ∙ cong (refl ∙_) + (sym (snd g) ∙∙ cong-g ∙∙ (λ j → snd g (j ∧ ~ k)))) i}) + (((cong (λ x → rUnit x j) c-f) ∙ (cong (λ x → lUnit x j) c-g)) i)))) + +hom-botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc n) →∙ Susp∙ (typ A)) + → botᵣ⁻ {A = A} n (∙Π f g) + ≡ invComp {A = A} n (botᵣ⁻ {A = A} n f) (botᵣ⁻ {A = A} n g) +hom-botᵣ⁻ zero f g = + ΣPathP ((funExt (λ { false → sym (rUnit _) + ; true → (rUnit _)})) + , ((λ i j → rUnit refl (i ∧ ~ j)) + ▷ lUnit (sym (rUnit refl)))) +hom-botᵣ⁻ (suc n) f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt (λ x → (sym (rUnit (cong (fst (∙Π f g)) (merid x ∙ sym (merid (ptSn _)))))) + ∙∙ cong-∙ (fst (∙Π f g)) (merid x) (sym (merid (ptSn _))) + ∙∙ cong (cong (fst (∙Π f g)) (merid x) ∙_) (cong sym lem) + ∙ sym (rUnit (cong (fst (∙Π f g)) (merid x))))) + where + lem : cong (fst (∙Π f g)) (merid (ptSn (suc n))) ≡ refl + lem = (λ i → (sym (snd f) ∙∙ cong (fst f) (rCancel (merid (ptSn _)) i) ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) (rCancel (merid (ptSn _)) i) ∙∙ snd g)) + ∙ (λ i → ∙∙lCancel (snd f) i ∙ ∙∙lCancel (snd g) i) + ∙ sym (rUnit refl) + +-- We get that botᵣ⁻ (and hence botᵣ) is homomorphism +hom-botᵣ⁻' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc (suc n)) →∙ Susp∙ (typ A)) + → botᵣ⁻ {A = A} (suc n) (∙Π f g) + ≡ ∙Π (botᵣ⁻ {A = A} (suc n) f) (botᵣ⁻ {A = A} (suc n) g) +hom-botᵣ⁻' {A = A} n f g = + hom-botᵣ⁻ {A = A} (suc n) f g + ∙ sym (∙Π≡invComp {A = A} _ (botᵣ⁻ {A = A} _ f) (botᵣ⁻ {A = A} _ g)) + +hom-botᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc n) →∙ Ω (Susp∙ (typ A))) + → botᵣ {A = A} (suc n) (∙Π f g) + ≡ ∙Π (botᵣ {A = A} (suc n) f) (botᵣ {A = A} (suc n) g) +hom-botᵣ {A = A} n f g = + morphLemmas.isMorphInv ∙Π ∙Π (botᵣ⁻ {A = A} (suc n)) + (hom-botᵣ⁻' {A = A} n) + (botᵣ {A = A} (suc n)) + (leftInv (botᵣIso {A = A} (suc n))) + (rightInv (botᵣIso {A = A} (suc n))) + f g + +isHom-IsoΩSphereMapᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + (p q : typ ((Ω^ (2 + n)) (Susp∙ (typ A)))) + → Iso.fun (IsoΩSphereMapᵣ (suc n)) (p ∙ q) + ≡ ∙Π (Iso.fun (IsoΩSphereMapᵣ (suc n)) p) + (Iso.fun (IsoΩSphereMapᵣ (suc n)) q) +isHom-IsoΩSphereMapᵣ {A = A} n p q = + cong (botᵣ {A = A} (suc n)) + (cong (Ω→SphereMap (suc n) {A = Ω (Susp∙ (typ A))}) + (flipΩIsopres· n p q) + ∙ isHom-Ω→SphereMap n (fun (flipΩIso (suc n)) p) + (fun (flipΩIso (suc n)) q)) + ∙ hom-botᵣ n _ _ + +suspMapΩ→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : typ ((Ω^ (suc n)) A)) + → suspMapΩ (suc n) (p ∙ q) + ≡ suspMapΩ (suc n) p ∙ suspMapΩ (suc n) q +suspMapΩ→hom {A = A} n p q = + cong (sym (snd (suspMapΩ∙ {A = A} n)) ∙∙_∙∙ snd (suspMapΩ∙ {A = A} n)) + (cong-∙ (fst (suspMapΩ∙ {A = A} n)) p q) + ∙ help (snd (suspMapΩ∙ {A = A} n)) _ _ + where + help : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q s : x ≡ x) + → sym p ∙∙ (q ∙ s) ∙∙ p + ≡ (sym p ∙∙ q ∙∙ p) ∙ (sym p ∙∙ s ∙∙ p) + help {x = x} = + J (λ y p → (q s : x ≡ x) + → sym p ∙∙ (q ∙ s) ∙∙ p + ≡ (sym p ∙∙ q ∙∙ p ) ∙ (sym p ∙∙ s ∙∙ p)) + λ q s → sym (rUnit (q ∙ s)) + ∙ cong₂ _∙_ (rUnit q) (rUnit s) + +private + transportLem : ∀ {ℓ} {A B : Type ℓ} + (_+A_ : A → A → A) (_+B_ : B → B → B) + → (e : Iso A B) + → ((x y : A) → fun e (x +A y) ≡ fun e x +B fun e y) + → PathP (λ i → isoToPath e i → isoToPath e i → isoToPath e i) + _+A_ _+B_ + transportLem _+A_ _+B_ e hom = + toPathP (funExt (λ p → funExt λ q → + (λ i → transportRefl + (fun e (inv e (transportRefl p i) + +A inv e (transportRefl q i))) i) + ∙∙ hom (inv e p) (inv e q) + ∙∙ cong₂ _+B_ (rightInv e p) (rightInv e q))) + + pₗ : ∀ {ℓ} (A : Pointed ℓ) (n : ℕ) + → typ (Ω ((Ω^ n) A)) ≡ (S₊∙ (suc n) →∙ A) + pₗ A n = isoToPath (IsoΩSphereMap {A = A} (suc n)) + + pᵣ : ∀ {ℓ} (A : Pointed ℓ) (n : ℕ) + → typ ((Ω^ (2 + n)) (Susp∙ (typ A))) + ≡ (S₊∙ (suc (suc n)) →∙ Susp∙ (typ A)) + pᵣ A n = isoToPath (IsoΩSphereMapᵣ {A = A} (suc n)) + +∙Π→∙ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → PathP (λ i → pₗ A n i → pₗ A n i → pₗ A n i) _∙_ ∙Π +∙Π→∙ {A = A} n = + transportLem _∙_ ∙Π (IsoΩSphereMap {A = A} (suc n)) (isHom-Ω→SphereMap n) + +∙Π→∙ᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → PathP (λ i → pᵣ A n i → pᵣ A n i → pᵣ A n i) _∙_ ∙Π +∙Π→∙ᵣ {A = A} n = + transportLem _∙_ ∙Π (IsoΩSphereMapᵣ {A = A} (suc n)) (isHom-IsoΩSphereMapᵣ n) + +isHom-suspMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → suspMap n (∙Π f g) + ≡ ∙Π (suspMap n f) (suspMap n g) +isHom-suspMap {A = A} n = + transport (λ i → (f g : isoToPath (IsoΩSphereMap {A = A} (suc n)) i) + → Ωσ→suspMap n i (∙Π→∙ n i f g) + ≡ ∙Π→∙ᵣ n i (Ωσ→suspMap n i f) (Ωσ→suspMap n i g)) + (suspMapΩ→hom n) + +suspMapπHom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupHom (πGr n A) (πGr (suc n) (Susp∙ (typ A))) +fst (suspMapπHom {A = A} n) = suspMapπ (suc n) +snd (suspMapπHom {A = A} n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (suspMapΩ→hom n p q)) + +suspMapπ'Hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupHom (π'Gr n A) (π'Gr (suc n) (Susp∙ (typ A))) +fst (suspMapπ'Hom {A = A} n) = suspMapπ' n +snd (suspMapπ'Hom {A = A} n) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ (isHom-suspMap n f g)) + +πGr≅π'Grᵣ : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) + → GroupIso (πGr (suc n) (Susp∙ (typ A))) + (π'Gr (suc n) (Susp∙ (typ A))) +fst (πGr≅π'Grᵣ n A) = setTruncIso (IsoΩSphereMapᵣ (suc n)) +snd (πGr≅π'Grᵣ n A) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ (isHom-IsoΩSphereMapᵣ n f g)) + +private + isConnectedPres : ∀ {ℓ} {A : Pointed ℓ} (con n : ℕ) + → isConnectedFun con (suspMapΩ∙ {A = A} (suc n) .fst) + → isConnectedFun con (suspMap {A = A} n) + isConnectedPres {A = A} con n hyp = + transport (λ i → isConnectedFun con (Ωσ→suspMap {A = A} n i)) hyp + +isConnectedSuspMap : (n m : ℕ) + → isConnectedFun ((m + suc m) ∸ n) (suspMap {A = S₊∙ (suc m)} n) +isConnectedSuspMap n m = + isConnectedPres _ _ (suspMapΩ-connected m (suc n) (sphereConnected (suc m))) + +isSurjectiveSuspMap : (n : ℕ) + → isSurjective (suspMapπ'Hom {A = S₊∙ (2 + n)} (2 + n)) +isSurjectiveSuspMap n = + sElim (λ _ → isProp→isSet squash) + λ f → + trRec + ((subst (λ x → isOfHLevel x (isInIm (suspMapπ'Hom (2 + n)) ∣ f ∣₂)) + (sym (snd (lem n n))) + (isProp→isOfHLevelSuc {A = isInIm (suspMapπ'Hom (2 + n)) ∣ f ∣₂} + (fst (lem n n)) squash))) + (λ p → ∣ ∣ fst p ∣₂ , (cong ∣_∣₂ (snd p)) ∣) + (fst (isConnectedSuspMap (2 + n) (suc n) f)) + where + lem : (m n : ℕ) → Σ[ x ∈ ℕ ] ((m + suc (suc n) ∸ suc n) ≡ suc x) + lem zero zero = 0 , refl + lem (suc m) zero = suc m , +-comm m 2 + lem zero (suc n) = lem zero n + lem (suc m) (suc n) = fst (lem (suc m) n) , (cong (_∸ (suc n)) (+-comm m (3 + n)) + ∙∙ cong (_∸ n) (+-comm (suc (suc n)) m) + ∙∙ snd (lem (suc m) n)) diff --git a/Cubical/Homotopy/HSpace.agda b/Cubical/Homotopy/HSpace.agda new file mode 100644 index 0000000000..055f37e2ac --- /dev/null +++ b/Cubical/Homotopy/HSpace.agda @@ -0,0 +1,51 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.HSpace where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn + +record HSpace {ℓ : Level} (A : Pointed ℓ) : Type ℓ where + constructor HSp + field + μ : typ A → typ A → typ A + μₗ : (x : typ A) → μ (pt A) x ≡ x + μᵣ : (x : typ A) → μ x (pt A) ≡ x + μₗᵣ : μₗ (pt A) ≡ μᵣ (pt A) + +record AssocHSpace {ℓ : Level} {A : Pointed ℓ} (e : HSpace A) : Type ℓ where + constructor AssocHSp + field + μ-assoc : (x y z : _) + → HSpace.μ e (HSpace.μ e x y) z ≡ HSpace.μ e x (HSpace.μ e y z) + + μ-assoc-filler : (y z : typ A) + → PathP (λ i → HSpace.μ e (HSpace.μₗ e y i) z + ≡ HSpace.μₗ e (HSpace.μ e y z) i) + (μ-assoc (pt A) y z) + refl + +-- Instances +open HSpace +open AssocHSpace + +-- S¹ +S1-HSpace : HSpace (S₊∙ 1) +μ S1-HSpace = _·_ +μₗ S1-HSpace x = refl +μᵣ S1-HSpace base = refl +μᵣ S1-HSpace (loop i) = refl +μₗᵣ S1-HSpace x = refl + +S1-AssocHSpace : AssocHSpace S1-HSpace +μ-assoc S1-AssocHSpace base _ _ = refl +μ-assoc S1-AssocHSpace (loop i) x y j = h x y j i + where + h : (x y : S₊ 1) → cong (_· y) (rotLoop x) ≡ rotLoop (x · y) + h = wedgeconFun _ _ (λ _ _ → isOfHLevelPath 2 (isGroupoidS¹ _ _) _ _) + (λ x → refl) + (λ { base → refl ; (loop i) → refl}) + refl +μ-assoc-filler S1-AssocHSpace _ _ = refl diff --git a/Cubical/Homotopy/Hopf.agda b/Cubical/Homotopy/Hopf.agda new file mode 100644 index 0000000000..4063153617 --- /dev/null +++ b/Cubical/Homotopy/Hopf.agda @@ -0,0 +1,704 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.Hopf where + +open import Cubical.Homotopy.HSpace + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Data.Int hiding (_·_) + +open import Cubical.HITs.Pushout.Flattening +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.HITs.S2 +open import Cubical.HITs.S3 +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) +open import Cubical.HITs.Join +open import Cubical.HITs.Interval + renaming ( zero to I0 ; one to I1 ) + +open Iso +open HSpace +open AssocHSpace + +private + retEq≡secEq : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) + → (x : _) → secEq e (e .fst x) ≡ cong (e .fst) (retEq e x) + retEq≡secEq {A = A} = + EquivJ (λ B e → (x : _) → secEq e (e .fst x) ≡ cong (e .fst) (retEq e x)) + λ _ → refl + +module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} + (e-ass : AssocHSpace e) (conA : ((x y : typ A) → ∥ x ≡ y ∥)) where + isEquiv-μ : (x : typ A) → isEquiv (λ z → (μ e z x)) + isEquiv-μ x = pRec (isPropIsEquiv _) + (J (λ x _ → isEquiv (λ z → μ e z x)) + (subst isEquiv (funExt (λ z → sym (μᵣ e z))) + (idIsEquiv (typ A)))) + (conA (pt A) x) + + isEquiv-μ' : (x : typ A) → isEquiv (μ e x) + isEquiv-μ' x = + pRec (isPropIsEquiv _) + (J (λ x _ → isEquiv (μ e x)) + (subst isEquiv (funExt (λ x → sym (μₗ e x))) (idIsEquiv (typ A)))) + (conA (pt A) x) + + μ-eq : (x : typ A) → typ A ≃ typ A + μ-eq x = (λ z → μ e z x) , (isEquiv-μ x) + + μ-eq' : (x : typ A) → typ A ≃ typ A + μ-eq' x = μ e x , isEquiv-μ' x + + Hopf : Susp (typ A) → Type ℓ + Hopf north = typ A + Hopf south = typ A + Hopf (merid a i₁) = ua (μ-eq a) i₁ + + TotalSpaceHopfPush : Type _ + TotalSpaceHopfPush = + Pushout {A = typ A × typ A} fst λ x → μ e (fst x) (snd x) + + TotalSpaceHopfPush→TotalSpace : + TotalSpaceHopfPush → Σ[ x ∈ Susp (typ A) ] Hopf x + TotalSpaceHopfPush→TotalSpace (inl x) = north , x + TotalSpaceHopfPush→TotalSpace (inr x) = south , x + TotalSpaceHopfPush→TotalSpace (push (x , y) i₁) = + merid y i₁ , ua-gluePt (μ-eq y) i₁ x + + joinIso₁ : Iso (TotalSpaceHopfPush) (join (typ A) (typ A)) + joinIso₁ = theIso + where + F : TotalSpaceHopfPush → join (typ A) (typ A) + F (inl x) = inl x + F (inr x) = inr x + F (push (a , x) i) = push a (μ e a x) i + + G : join (typ A) (typ A) → TotalSpaceHopfPush + G (inl x) = inl x + G (inr x) = inr x + G (push a b i) = + (push (a , invEq (μ-eq' a) b) ∙ cong inr (secEq (μ-eq' a) b)) i + + s : section F G + s (inl x) = refl + s (inr x) = refl + s (push a b i) j = + hcomp (λ k → λ { (i = i0) → inl a + ; (i = i1) → inr (secEq (μ-eq' a) b (j ∨ k)) + ; (j = i0) → F (compPath-filler + (push (a , invEq (μ-eq' a) b)) + (cong inr (secEq (μ-eq' a) b)) k i) + ; (j = i1) → push a b i}) + (hcomp (λ k → λ { (i = i0) → inl a + ; (i = i1) → inr (secEq (μ-eq' a) b (~ k ∨ j)) + ; (j = i0) → push a (secEq (μ-eq' a) b (~ k)) i + ; (j = i1) → push a b i}) + (push a b i)) + + r : retract F G + r (inl x) = refl + r (inr x) = refl + r (push (x , y) i) j = + hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (μ e x y) + ; (j = i0) → (push (x , invEq (μ-eq' x) (μ e x y)) + ∙ (λ i₁ → inr (retEq≡secEq (μ-eq' x) y (~ k) i₁))) i + ; (j = i1) → push (x , y) i}) + (hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (μ e x (retEq (μ-eq' x) y k)) + ; (j = i1) → push (x , retEq (μ-eq' x) y k) i}) + ((push (x , invEq (μ-eq' x) (μ e x y))) i)) + + theIso : Iso TotalSpaceHopfPush (join (typ A) (typ A)) + fun theIso = F + inv theIso = G + rightInv theIso = s + leftInv theIso = r + + isEquivTotalSpaceHopfPush→TotalSpace : + isEquiv TotalSpaceHopfPush→TotalSpace + isEquivTotalSpaceHopfPush→TotalSpace = + isoToIsEquiv theIso + where + inv' : _ → _ + inv' (north , y) = inl y + inv' (south , y) = inr y + inv' (merid a i , y) = + hcomp (λ k → λ { (i = i0) → push (y , a) (~ k) + ; (i = i1) → inr y}) + (inr (ua-unglue (μ-eq a) i y)) + where + + pp : PathP (λ i → ua (μ-eq a) i → TotalSpaceHopfPush) + inl inr + pp = ua→ {e = μ-eq a} {B = λ _ → TotalSpaceHopfPush} λ b → push (b , a) + + sect : (x : _) → TotalSpaceHopfPush→TotalSpace (inv' x) ≡ x + sect (north , x) = refl + sect (south , x) = refl + sect (merid a i , y) j = + hcomp (λ k → λ { (i = i0) → merid a (~ k ∧ ~ j) + , ua-gluePt (μ-eq a) (~ k ∧ ~ j) y + ; (i = i1) → south , y + ; (j = i0) → + TotalSpaceHopfPush→TotalSpace + (hfill (λ k → λ { (i = i0) → push (y , a) (~ k) + ; (i = i1) → inr y}) + (inS (inr (ua-unglue (μ-eq a) i y))) + k) + ; (j = i1) → merid a i , y}) + ((merid a (i ∨ ~ j)) , lem (μ-eq a) i j y) + where + lem : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → + PathP (λ i → PathP (λ j → (y : ua e i) → ua e (i ∨ ~ j)) + (λ y → ua-unglue e i y) + λ y → y) + (λ j y → ua-gluePt e (~ j) y) + refl + lem {A = A} {B = B} = + EquivJ (λ B e → PathP (λ i → PathP (λ j → (y : ua e i) → ua e (i ∨ ~ j)) + (λ y → ua-unglue e i y) + λ y → y) + (λ j y → ua-gluePt e (~ j) y) + refl) + λ i j a → ua-gluePt (idEquiv B) (i ∨ ~ j) (ua-unglue (idEquiv B) i a) + + retr : retract TotalSpaceHopfPush→TotalSpace inv' + retr (inl x) = refl + retr (inr x) = refl + retr (push (x , y) i) j = + hcomp (λ k → λ { (i = i0) → push (x , y) (~ k) + ; (i = i1) → inr (μ e x y) + ; (j = i1) → push (x , y) (i ∨ ~ k)}) + (inr (μ e x y)) + + theIso : Iso TotalSpaceHopfPush (Σ (Susp (typ A)) Hopf) + fun theIso = TotalSpaceHopfPush→TotalSpace + inv theIso = inv' + rightInv theIso = sect + leftInv theIso = retr + + IsoTotalSpaceJoin : Iso (Σ[ x ∈ Susp (typ A) ] Hopf x) (join (typ A) (typ A)) + IsoTotalSpaceJoin = + compIso (equivToIso (invEquiv (_ , isEquivTotalSpaceHopfPush→TotalSpace))) + joinIso₁ + + induced : TotalSpaceHopfPush → Susp (typ A) + induced = fst ∘ TotalSpaceHopfPush→TotalSpace + + ua-lem : (x y z : typ A) → (i j : I) → ua (μ-eq y) i + ua-lem x y z i j = + fill (λ k → ua (μ-eq y) i) + (λ j → λ { (i = i0) → μ e z x + ; (i = i1) → μ-assoc e-ass z x y j}) + (inS (ua-gluePt (μ-eq y) i (μ e z x))) + j + + TotalSpaceHopfPush→≃Hopf : (x : TotalSpaceHopfPush) → typ A ≃ Hopf (induced x) + TotalSpaceHopfPush→≃Hopf (inl x) = μ-eq x + TotalSpaceHopfPush→≃Hopf (inr x) = μ-eq x + TotalSpaceHopfPush→≃Hopf (push (x , y) i₁) = pp x y i₁ + where + pp : (x y : _) → PathP (λ i → typ A ≃ ua (μ-eq y) i) (μ-eq x) (μ-eq (μ e x y)) + pp x y = ΣPathP (P , help) + where + P : PathP (λ z → typ A → ua (μ-eq y) z) (fst (μ-eq x)) + (fst (μ-eq (μ e x y))) + P i z = ua-lem x y z i i1 + + abstract + help : PathP (λ i₂ → isEquiv (P i₂)) (snd (μ-eq x)) + (snd (μ-eq (μ e x y))) + help = toPathP (isPropIsEquiv _ _ _) + + Push→TotalSpaceHopf : (a : typ A) (x : TotalSpaceHopfPush) + → Σ[ x ∈ Susp (typ A) ] Hopf x + Push→TotalSpaceHopf a x = (induced x) , fst (TotalSpaceHopfPush→≃Hopf x) a + + Push→TotalSpaceHopf-equiv : (a : typ A) → isEquiv (Push→TotalSpaceHopf a) + Push→TotalSpaceHopf-equiv a = pRec (isPropIsEquiv _) + (J (λ a _ → isEquiv (Push→TotalSpaceHopf a)) + (subst isEquiv (sym main) + isEquivTotalSpaceHopfPush→TotalSpace)) + (conA (pt A) a) + where + lem₁ : (x : _) → fst ((Push→TotalSpaceHopf (pt A)) x) + ≡ fst (TotalSpaceHopfPush→TotalSpace x) + lem₁ (inl x) = refl + lem₁ (inr x) = refl + lem₁ (push a i) = refl + + lem₂ : (x : _) + → PathP (λ i → Hopf (lem₁ x i)) + (snd ((Push→TotalSpaceHopf (pt A)) x)) + (snd (TotalSpaceHopfPush→TotalSpace x)) + lem₂ (inl x) = μₗ e x + lem₂ (inr x) = μₗ e x + lem₂ (push (x , y) i) j = + hcomp (λ k → λ {(i = i0) → μₗ e x j + ; (i = i1) → μ-assoc-filler e-ass x y j k + ; (j = i0) → ua-lem x y (pt A) i k + ; (j = i1) → ua-gluePt (μ-eq y) i x}) + (ua-gluePt (μ-eq y) i (μₗ e x j)) + + main : Push→TotalSpaceHopf (pt A) ≡ TotalSpaceHopfPush→TotalSpace + main i x = (lem₁ x i) , (lem₂ x i) + + TotalSpaceHopfPush² : Type _ + TotalSpaceHopfPush² = Pushout {A = TotalSpaceHopfPush} (λ _ → tt) induced + + P : TotalSpaceHopfPush² → Type _ + P (inl x) = typ A + P (inr x) = Hopf x + P (push a i) = ua (TotalSpaceHopfPush→≃Hopf a) i + + TotalSpacePush² : Type _ + TotalSpacePush² = Σ[ x ∈ TotalSpaceHopfPush² ] P x + + TotalSpacePush²' : Type _ + TotalSpacePush²' = + Pushout {A = typ A × TotalSpaceHopfPush} + {C = Σ[ x ∈ Susp (typ A) ] Hopf x} + fst + λ x → Push→TotalSpaceHopf (fst x) (snd x) + + IsoTotalSpacePush²TotalSpacePush²' : Iso TotalSpacePush² TotalSpacePush²' + IsoTotalSpacePush²TotalSpacePush²' = + compIso iso₂ (compIso (equivToIso fl.flatten) iso₁) + where + module fl = + FlatteningLemma (λ _ → tt) induced (λ x → typ A) + Hopf TotalSpaceHopfPush→≃Hopf + + iso₁ : Iso (Pushout fl.Σf fl.Σg) TotalSpacePush²' + fun iso₁ (inl x) = inl (snd x) + fun iso₁ (inr x) = inr x + fun iso₁ (push a i) = push ((snd a) , (fst a)) i + inv iso₁ (inl x) = inl (tt , x) + inv iso₁ (inr x) = inr x + inv iso₁ (push a i) = push (snd a , fst a) i + rightInv iso₁ (inl x) = refl + rightInv iso₁ (inr x) = refl + rightInv iso₁ (push a i) = refl + leftInv iso₁ (inl x) = refl + leftInv iso₁ (inr x) = refl + leftInv iso₁ (push a i) = refl + + iso₂ : Iso TotalSpacePush² (Σ (Pushout (λ _ → tt) induced) fl.E) + fun iso₂ (inl x , y) = inl x , y + fun iso₂ (inr x , y) = inr x , y + fun iso₂ (push a i , y) = push a i , y + inv iso₂ (inl x , y) = inl x , y + inv iso₂ (inr x , y) = inr x , y + inv iso₂ (push a i , y) = push a i , y + rightInv iso₂ (inl x , snd₁) = refl + rightInv iso₂ (inr x , snd₁) = refl + rightInv iso₂ (push a i , snd₁) = refl + leftInv iso₂ (inl x , snd₁) = refl + leftInv iso₂ (inr x , snd₁) = refl + leftInv iso₂ (push a i , snd₁) = refl + + F : TotalSpacePush²' + → (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} fst snd) + F (inl x) = inl x + F (inr x) = inr x + F (push (x , y) i) = push (x , Push→TotalSpaceHopf x y) i + + G : (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} fst snd) + → TotalSpacePush²' + G (inl x) = inl x + G (inr x) = inr x + G (push (x , y) i) = + hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) + → inr (secEq (_ , Push→TotalSpaceHopf-equiv x) y k)}) + (push (x , invEq (_ , Push→TotalSpaceHopf-equiv x) y) i) + + IsoTotalSpacePush²'ΣPush : Iso TotalSpacePush²' + (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} fst snd) + fun IsoTotalSpacePush²'ΣPush = F + inv IsoTotalSpacePush²'ΣPush = G + rightInv IsoTotalSpacePush²'ΣPush (inl x) = refl + rightInv IsoTotalSpacePush²'ΣPush (inr x) = refl + rightInv IsoTotalSpacePush²'ΣPush (push (x , y) i) j = + hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) + → inr (secEq (_ , Push→TotalSpaceHopf-equiv x) y k) + ; (j = i0) → F ( + hfill (λ k → + λ { (i = i0) → inl x + ; (i = i1) + → inr (secEq (_ + , Push→TotalSpaceHopf-equiv x) y k)}) + (inS (push (x + , invEq (_ + , Push→TotalSpaceHopf-equiv x) y) i)) k) + ; (j = i1) + → push (x + , (secEq (_ + , Push→TotalSpaceHopf-equiv x) y k)) i}) + (push (x , (secEq (_ , Push→TotalSpaceHopf-equiv x) y i0)) i) + leftInv IsoTotalSpacePush²'ΣPush (inl x) = refl + leftInv IsoTotalSpacePush²'ΣPush (inr x) = refl + leftInv IsoTotalSpacePush²'ΣPush (push (x , y) i) j = + hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (secEq (Push→TotalSpaceHopf x + , Push→TotalSpaceHopf-equiv x) + (Push→TotalSpaceHopf x y) (j ∨ k)) + ; (j = i1) → push (x , y) i}) + (hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (retEq≡secEq + (Push→TotalSpaceHopf x + , Push→TotalSpaceHopf-equiv x) + y (~ k) j) + ; (j = i1) → push (x , y) i + ; (j = i0) → push (x , invEq + (Push→TotalSpaceHopf x + , Push→TotalSpaceHopf-equiv x) + (Push→TotalSpaceHopf x y)) i}) + (push (x , retEq (Push→TotalSpaceHopf x + , Push→TotalSpaceHopf-equiv x) y j) i)) + + joinIso₂ : Iso TotalSpacePush² (join (typ A) (join (typ A) (typ A))) + joinIso₂ = + compIso IsoTotalSpacePush²TotalSpacePush²' + (compIso IsoTotalSpacePush²'ΣPush + (compIso (equivToIso (joinPushout≃join _ _)) + (pathToIso (cong (join (typ A)) + (isoToPath IsoTotalSpaceJoin))))) + + +-- Direct construction of Hopf fibration for S¹ +module S¹Hopf where + Border : (x : S¹) → (j : I) → Partial (j ∨ ~ j) (Σ Type₀ (λ T → T ≃ S¹)) + Border x j (j = i0) = S¹ , (x ·_) , rotIsEquiv x + Border x j (j = i1) = S¹ , idEquiv S¹ + + -- Hopf fibration using SuspS¹ + HopfSuspS¹ : SuspS¹ → Type₀ + HopfSuspS¹ north = S¹ + HopfSuspS¹ south = S¹ + HopfSuspS¹ (merid x j) = Glue S¹ (Border x j) + + -- Hopf fibration using S² + -- TODO : prove that it is equivalent to HopfSuspS¹ + HopfS² : S² → Type₀ + HopfS² base = S¹ + HopfS² (surf i j) = Glue S¹ (λ { (i = i0) → _ , idEquiv S¹ + ; (i = i1) → _ , idEquiv S¹ + ; (j = i0) → _ , idEquiv S¹ + ; (j = i1) → _ , _ , rotIsEquiv (loop i) } ) + + -- Hopf fibration using more direct definition of the rot equivalence + -- TODO : prove that it is equivalent to HopfSuspS¹ + HopfS²' : S² → Type₀ + HopfS²' base = S¹ + HopfS²' (surf i j) = Glue S¹ (λ { (i = i0) → _ , rotLoopEquiv i0 + ; (i = i1) → _ , rotLoopEquiv i0 + ; (j = i0) → _ , rotLoopEquiv i0 + ; (j = i1) → _ , rotLoopEquiv i } ) + + -- Total space of the fibration + TotalHopf : Type₀ + TotalHopf = Σ SuspS¹ HopfSuspS¹ + + -- Forward direction + filler-1 : I → (j : I) → (y : S¹) → Glue S¹ (Border y j) → join S¹ S¹ + filler-1 i j y x = hfill (λ t → λ { (j = i0) → inl (rotInv-1 x y t) + ; (j = i1) → inr x }) + (inS (push ((unglue (j ∨ ~ j) x) · invLooper y) (unglue (j ∨ ~ j) x) j)) i + + TotalHopf→JoinS¹S¹ : TotalHopf → join S¹ S¹ + TotalHopf→JoinS¹S¹ (north , x) = inl x + TotalHopf→JoinS¹S¹ (south , x) = inr x + TotalHopf→JoinS¹S¹ (merid y j , x) = filler-1 i1 j y x + + -- Backward direction + JoinS¹S¹→TotalHopf : join S¹ S¹ → TotalHopf + JoinS¹S¹→TotalHopf (inl x) = (north , x) + JoinS¹S¹→TotalHopf (inr x) = (south , x) + JoinS¹S¹→TotalHopf (push y x j) = + (merid (invLooper y · x) j + , glue (λ { (j = i0) → y ; (j = i1) → x }) (rotInv-2 x y j)) + + -- Now for the homotopies, we will need to fill squares indexed by x y : S¹ with value in S¹ + -- Some will be extremeley tough, but happen to be easy when x = y = base + -- therefore, we fill them for x = y = base and then use the connectedness of S¹ × S¹ and + -- the discreteness of ΩS¹ to get general fillers. + + -- To proceed with that strategy, we first need a lemma : + -- the sections of the trivial fibration λ (_ : S¹) (_ : S¹) → Int are constant + + -- this should be generalized to a constant fibration over a connected space with + -- discrete fiber + fibℤ : S¹ → S¹ → Type₀ + fibℤ _ _ = ℤ + + S¹→HSet : (A : Type₀) (p : isSet A) (F : S¹ → A) (x : S¹) → F base ≡ F x + S¹→HSet A p F base = refl {x = F base} + S¹→HSet A p F (loop i) = f' i + where + f : PathP (λ i → F base ≡ F (loop i)) refl (cong F loop) + f i = λ j → F (loop (i ∧ j)) + L : cong F loop ≡ refl + L = p (F base) (F base) (f i1) refl + f' : PathP (λ i → F base ≡ F (loop i)) (refl {x = F base}) (refl {x = F base}) + f' = transport (λ i → PathP (λ j → F base ≡ F (loop j)) refl (L i)) f + + constant-loop : (F : S¹ → S¹ → ℤ) → (x y : S¹) → F base base ≡ F x y + constant-loop F x y = L0 ∙ L1 + where + p : isSet (S¹ → ℤ) + p = isSetΠ (λ _ → isSetℤ) + L : F base ≡ F x + L = S¹→HSet (S¹ → ℤ) p F x + L0 : F base base ≡ F x base + L0 i = L i base + L1 : F x base ≡ F x y + L1 = S¹→HSet ℤ isSetℤ (F x) y + + discretefib : (F : S¹ → S¹ → Type₀) → Type₀ + discretefib F = (a : (x y : S¹) → F x y) → + (b : (x y : S¹) → F x y) → + (a base base ≡ b base base) → + (x y : S¹) → a x y ≡ b x y + + discretefib-fibℤ : discretefib fibℤ + discretefib-fibℤ a b h x y i = + hcomp (λ t → λ { (i = i0) → constant-loop a x y t + ; (i = i1) → constant-loop b x y t }) + (h i) + + -- first homotopy + + assocFiller-3-aux : I → I → I → I → S¹ + assocFiller-3-aux x y j i = + hfill (λ t → λ { (i = i0) → rotInv-1 (loop y) (loop (~ y) · loop x) t + ; (i = i1) → rotInv-3 (loop y) (loop x) t + ; (x = i0) (y = i0) → base + ; (x = i0) (y = i1) → base + ; (x = i1) (y = i0) → base + ; (x = i1) (y = i1) → base }) + (inS ((rotInv-2 (loop x) (loop y) i) · (invLooper (loop (~ y) · loop x)))) j + + -- assocFiller-3-endpoint is used only in the type of the next function, to specify the + -- second endpoint. + -- However, I only need the first endpoint, but I cannot specify only one of them as is. + -- TODO : use cubical extension types when available to remove assocFiller-3-endpoint + assocFiller-3-endpoint : (x : S¹) → (y : S¹) → y ≡ y + assocFiller-3-endpoint base base i = base + assocFiller-3-endpoint (loop x) base i = assocFiller-3-aux x i0 i1 i + assocFiller-3-endpoint base (loop y) i = assocFiller-3-aux i0 y i1 i + assocFiller-3-endpoint (loop x) (loop y) i = assocFiller-3-aux x y i1 i + + assocFiller-3 : (x : S¹) → (y : S¹) → + PathP (λ j → rotInv-1 y (invLooper y · x) j ≡ rotInv-3 y x j) + (λ i → ((rotInv-2 x y i) · (invLooper (invLooper y · x)))) + (assocFiller-3-endpoint x y) + assocFiller-3 base base j i = base + assocFiller-3 (loop x) base j i = assocFiller-3-aux x i0 j i + assocFiller-3 base (loop y) j i = assocFiller-3-aux i0 y j i + assocFiller-3 (loop x) (loop y) j i = assocFiller-3-aux x y j i + + assoc-3 : (_ y : S¹) → basedΩS¹ y + assoc-3 x y i = assocFiller-3 x y i1 i + + fibℤ≡fibAssoc-3 : fibℤ ≡ (λ _ y → basedΩS¹ y) + fibℤ≡fibAssoc-3 i = λ x y → basedΩS¹≡ℤ y (~ i) + + discretefib-fibAssoc-3 : discretefib (λ _ y → basedΩS¹ y) + discretefib-fibAssoc-3 = + transp (λ i → discretefib (fibℤ≡fibAssoc-3 i)) i0 discretefib-fibℤ + + assocConst-3 : (x y : S¹) → assoc-3 x y ≡ refl + assocConst-3 x y = discretefib-fibAssoc-3 assoc-3 (λ _ _ → refl) refl x y + + assocSquare-3 : I → I → S¹ → S¹ → S¹ + assocSquare-3 i j x y = hcomp (λ t → λ { (i = i0) → assocFiller-3 x y j i0 + ; (i = i1) → assocFiller-3 x y j i1 + ; (j = i0) → assocFiller-3 x y i0 i + ; (j = i1) → assocConst-3 x y t i }) + (assocFiller-3 x y j i) + + filler-3 : I → I → S¹ → S¹ → join S¹ S¹ + filler-3 i j y x = + hcomp (λ t → λ { (i = i0) → filler-1 t j (invLooper y · x) + (glue (λ { (j = i0) → y ; (j = i1) → x }) + (rotInv-2 x y j)) + ; (i = i1) → push (rotInv-3 y x t) x j + ; (j = i0) → inl (assocSquare-3 i t x y) + ; (j = i1) → inr x }) + (push ((rotInv-2 x y (i ∨ j)) · (invLooper (invLooper y · x))) (rotInv-2 x y (i ∨ j)) j) + + JoinS¹S¹→TotalHopf→JoinS¹S¹ : ∀ x → TotalHopf→JoinS¹S¹ (JoinS¹S¹→TotalHopf x) ≡ x + JoinS¹S¹→TotalHopf→JoinS¹S¹ (inl x) i = inl x + JoinS¹S¹→TotalHopf→JoinS¹S¹ (inr x) i = inr x + JoinS¹S¹→TotalHopf→JoinS¹S¹ (push y x j) i = filler-3 i j y x + + -- Second homotopy + + -- This HIT is the total space of the Hopf fibration but the ends of SuspS¹ have not been + -- glued together yet — which makes it into a cylinder. + -- This allows to write compositions that do not properly match at the endpoints. However, + -- I suspect it is unnecessary. TODO : do without PseudoHopf + + PseudoHopf : Type₀ + PseudoHopf = (S¹ × Interval) × S¹ + + PseudoHopf-π1 : PseudoHopf → S¹ + PseudoHopf-π1 ((y , _) , _) = y + + PseudoHopf-π2 : PseudoHopf → S¹ + PseudoHopf-π2 (_ , x) = x + + assocFiller-4-aux : I → I → I → I → S¹ + assocFiller-4-aux x y j i = + hfill (λ t → λ { (i = i0) → ((invLooper (loop y · loop x · loop (~ y))) · (loop y · loop x)) + · (rotInv-1 (loop x) (loop y) t) + ; (i = i1) → (rotInv-4 (loop y) (loop y · loop x) (~ t)) · loop x + ; (x = i0) (y = i0) → base + ; (x = i0) (y = i1) → base + ; (x = i1) (y = i0) → base + ; (x = i1) (y = i1) → base }) + (inS (rotInv-2 (loop y · loop x) (loop y · loop x · loop (~ y)) i)) j + + -- See assocFiller-3-endpoint + -- TODO : use cubical extension types when available to remove assocFiller-4-endpoint + assocFiller-4-endpoint : (x y : S¹) → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x) + assocFiller-4-endpoint base base i = base + assocFiller-4-endpoint (loop x) base i = assocFiller-4-aux x i0 i1 i + assocFiller-4-endpoint base (loop y) i = assocFiller-4-aux i0 y i1 i + assocFiller-4-endpoint (loop x) (loop y) i = assocFiller-4-aux x y i1 i + + assocFiller-4 : (x y : S¹) → + PathP (λ j → ((invLooper (y · x · invLooper y)) · (y · x)) · (rotInv-1 x y j) ≡ (rotInv-4 y (y · x) (~ j)) · x) + (λ i → (rotInv-2 (y · x) (y · x · invLooper y) i)) + (assocFiller-4-endpoint x y) + assocFiller-4 base base j i = base + assocFiller-4 (loop x) base j i = assocFiller-4-aux x i0 j i + assocFiller-4 base (loop y) j i = assocFiller-4-aux i0 y j i + assocFiller-4 (loop x) (loop y) j i = assocFiller-4-aux x y j i + + assoc-4 : (x y : S¹) → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x) + assoc-4 x y i = assocFiller-4 x y i1 i + + fibℤ≡fibAssoc-4 : fibℤ ≡ (λ x y → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x)) + fibℤ≡fibAssoc-4 i = λ x y → basedΩS¹≡ℤ (((invLooper (y · x · invLooper y)) · (y · x)) · x) (~ i) + + discretefib-fibAssoc-4 : discretefib (λ x y → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x)) + discretefib-fibAssoc-4 = + transp (λ i → discretefib (fibℤ≡fibAssoc-4 i)) i0 discretefib-fibℤ + + assocConst-4 : (x y : S¹) → assoc-4 x y ≡ refl + assocConst-4 x y = discretefib-fibAssoc-4 assoc-4 (λ _ _ → refl) refl x y + + assocSquare-4 : I → I → S¹ → S¹ → S¹ + assocSquare-4 i j x y = + hcomp (λ t → λ { (i = i0) → assocFiller-4 x y j i0 + ; (i = i1) → assocFiller-4 x y j i1 + ; (j = i0) → assocFiller-4 x y i0 i + ; (j = i1) → assocConst-4 x y t i }) + (assocFiller-4 x y j i) + + filler-4-0 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf + filler-4-0 i j y x = + let x' = unglue (j ∨ ~ j) x in + hfill (λ t → λ { (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) + , invLooper (y · x · invLooper y) · (y · x) · (rotInv-1 x y t)) + ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) + (inS ((invLooper (x' · invLooper y) · x' , seg j) , rotInv-2 x' (x' · invLooper y) j)) i + + filler-4-1 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf + filler-4-1 i j y x = + let x' = unglue (j ∨ ~ j) x in + hfill (λ t → λ { (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) + , (rotInv-4 y (y · x) (~ t)) · x) + ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) + (inS ((invLooper (x' · invLooper y) · x' , seg j) , unglue (j ∨ ~ j) x)) i + + filler-4-2 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → TotalHopf + filler-4-2 i j y x = + let x' = unglue (j ∨ ~ j) x in + hcomp (λ t → λ { (i = i0) → JoinS¹S¹→TotalHopf (filler-1 t j y x) + ; (i = i1) → (merid (PseudoHopf-π1 (filler-4-0 t j y x)) j + , glue (λ { (j = i0) → rotInv-1 x y t ; (j = i1) → x }) + (PseudoHopf-π2 (filler-4-0 t j y x))) + ; (j = i0) → (north , rotInv-1 x y t) + ; (j = i1) → (south , x) }) + (merid (invLooper (x' · invLooper y) · x') j + , glue (λ { (j = i0) → y · x · invLooper y ; (j = i1) → x }) (rotInv-2 x' (x' · invLooper y) j)) + + filler-4-3 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf + filler-4-3 i j y x = + let x' = unglue (j ∨ ~ j) x in + hcomp (λ t → λ { (i = i0) → filler-4-0 t j y x + ; (i = i1) → filler-4-1 t j y x + ; (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) , assocSquare-4 i t x y) + ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) + ((invLooper (x' · invLooper y) · x' , seg j) , rotInv-2 x' (x' · invLooper y) (i ∨ j)) + + filler-4-4 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf + filler-4-4 i j y x = + let x' = unglue (j ∨ ~ j) x in + hcomp (λ t → λ { (i = i0) → filler-4-1 t j y x + ; (i = i1) → ((y , seg j) , unglue (j ∨ ~ j) x) + ; (j = i0) → ((rotInv-4 y (y · x) i , I0) + , (rotInv-4 y (y · x) (i ∨ ~ t)) · x) + ; (j = i1) → ((rotInv-4 y x i , I1) , x) }) + ((rotInv-4 y x' i , seg j) , x') + + filler-4-5 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → TotalHopf + filler-4-5 i j y x = + hcomp (λ t → λ { (i = i0) → filler-4-2 (~ t) j y x + ; (i = i1) → (merid (PseudoHopf-π1 (filler-4-4 t j y x)) j + , glue (λ { (j = i0) → x ; (j = i1) → x }) + (PseudoHopf-π2 (filler-4-4 t j y x))) + ; (j = i0) → (north , x) + ; (j = i1) → (south , x) }) + (merid (PseudoHopf-π1 (filler-4-3 i j y x)) j + , glue (λ { (j = i0) → x ; (j = i1) → x }) (PseudoHopf-π2 (filler-4-3 i j y x))) + + TotalHopf→JoinS¹S¹→TotalHopf : ∀ x → JoinS¹S¹→TotalHopf (TotalHopf→JoinS¹S¹ x) ≡ x + TotalHopf→JoinS¹S¹→TotalHopf (north , x) i = (north , x) + TotalHopf→JoinS¹S¹→TotalHopf (south , x) i = (south , x) + TotalHopf→JoinS¹S¹→TotalHopf (merid y j , x) i = filler-4-5 i j y x + + + JoinS¹S¹≡TotalHopf : join S¹ S¹ ≡ TotalHopf + JoinS¹S¹≡TotalHopf = isoToPath (iso JoinS¹S¹→TotalHopf + TotalHopf→JoinS¹S¹ + TotalHopf→JoinS¹S¹→TotalHopf + JoinS¹S¹→TotalHopf→JoinS¹S¹) + + S³≡TotalHopf : S³ ≡ TotalHopf + S³≡TotalHopf = S³≡joinS¹S¹ ∙ JoinS¹S¹≡TotalHopf + + open Iso + IsoS³TotalHopf : Iso (S₊ 3) TotalHopf + fun IsoS³TotalHopf x = JoinS¹S¹→TotalHopf (S³→joinS¹S¹ (inv IsoS³S3 x)) + inv IsoS³TotalHopf x = fun IsoS³S3 (joinS¹S¹→S³ (TotalHopf→JoinS¹S¹ x)) + rightInv IsoS³TotalHopf x = + cong (JoinS¹S¹→TotalHopf ∘ S³→joinS¹S¹) + (leftInv IsoS³S3 (joinS¹S¹→S³ (TotalHopf→JoinS¹S¹ x))) + ∙∙ cong JoinS¹S¹→TotalHopf + (joinS¹S¹→S³→joinS¹S¹ (TotalHopf→JoinS¹S¹ x)) + ∙∙ TotalHopf→JoinS¹S¹→TotalHopf x + leftInv IsoS³TotalHopf x = + cong (fun IsoS³S3 ∘ joinS¹S¹→S³) + (JoinS¹S¹→TotalHopf→JoinS¹S¹ (S³→joinS¹S¹ (inv IsoS³S3 x))) + ∙∙ cong (fun IsoS³S3) (S³→joinS¹S¹→S³ (inv IsoS³S3 x)) + ∙∙ Iso.rightInv IsoS³S3 x diff --git a/Cubical/Homotopy/HopfInvariant/Base.agda b/Cubical/Homotopy/HopfInvariant/Base.agda new file mode 100644 index 0000000000..1efa57d9a8 --- /dev/null +++ b/Cubical/Homotopy/HopfInvariant/Base.agda @@ -0,0 +1,313 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.HopfInvariant.Base where + +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.SuspensionMap + +open import Cubical.Relation.Nullary + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.GradedCommutativity +open import Cubical.ZCohomology.Gysin + +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Unit + +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.Truncation + renaming (rec to trRec) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec) + +-- The pushout of the hopf invariant +HopfInvariantPush : (n : ℕ) + → (f : S₊ (3 +ℕ n +ℕ n) → S₊ (2 +ℕ n)) + → Type _ +HopfInvariantPush n f = Pushout (λ _ → tt) f + +Hopfα : (n : ℕ) + → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) + → coHom (2 +ℕ n) (HopfInvariantPush n (fst f)) +Hopfα n f = ∣ (λ { (inl x) → 0ₖ _ + ; (inr x) → ∣ x ∣ + ; (push a i) → help a (~ i)}) ∣₂ + where + help : (a : S₊ (3 +ℕ n +ℕ n)) → ∣ fst f a ∣ ≡ 0ₖ (2 +ℕ n) + help = sphereElim _ (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) + (isOfHLevelPath' (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n)) _ _)) + (cong ∣_∣ₕ (snd f)) + +Hopfβ : (n : ℕ) + → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) + → coHom (4 +ℕ (n +ℕ n)) (HopfInvariantPush n (fst f)) +Hopfβ n f = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) ∣ ∣_∣ ∣₂ + +-- To define the Hopf invariant, we need to establish the +-- non-trivial isos of Hⁿ(HopfInvariant). +module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where + module M = MV _ _ _ (λ _ → tt) (fst f) + ¬lemHopf : (n m : ℕ) → ¬ suc (suc (m +ℕ n)) ≡ suc n + ¬lemHopf zero zero p = snotz (cong predℕ p) + ¬lemHopf (suc n) zero p = ¬lemHopf n zero (cong predℕ p) + ¬lemHopf zero (suc m) p = snotz (cong predℕ p) + ¬lemHopf (suc n) (suc m) p = + ¬lemHopf n (suc m) (cong (suc ∘ suc ) + (sym (+-suc m n)) ∙ (cong predℕ p)) + + SphereHopfCohomIso : + GroupEquiv + (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n))) + ((coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f)))) + fst (fst SphereHopfCohomIso) = + fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + snd (fst SphereHopfCohomIso) = help + where + abstract + help : isEquiv (fst (MV.d _ _ _ (λ _ → tt) (fst f) + (3 +ℕ n +ℕ n))) + help = + SES→isEquiv + (isContr→≡UnitGroup + (isContrΣ (GroupIsoUnitGroup→isContr (invGroupIso (Hⁿ-Unit≅0 _))) + λ _ → GroupIsoUnitGroup→isContr + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ (¬lemHopf n n))))) + (isContr→≡UnitGroup + (isContrΣ (GroupIsoUnitGroup→isContr + (invGroupIso (Hⁿ-Unit≅0 _))) + λ _ → GroupIsoUnitGroup→isContr + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ (¬lemHopf n (suc n)))))) + (M.Δ (3 +ℕ n +ℕ n)) + (M.d (3 +ℕ n +ℕ n)) + (M.i (4 +ℕ n +ℕ n)) + (M.Ker-d⊂Im-Δ _) + (M.Ker-i⊂Im-d _) + snd SphereHopfCohomIso = d-inv + where + -- Currently, type checking is very slow without the abstract flag. + -- TODO : Remove abstract + abstract + d-inv : IsGroupHom (snd (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n)))) + (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) + (snd (coHomGr (suc (3 +ℕ n +ℕ n)) + (Pushout (λ _ → tt) (fst f)))) + d-inv = snd (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + + Hopfβ-Iso : GroupIso (coHomGr (suc (3 +ℕ n +ℕ n)) + (HopfInvariantPush _ (fst f))) + ℤGroup + fst Hopfβ-Iso = compIso (invIso (equivToIso (fst SphereHopfCohomIso))) + (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) + snd Hopfβ-Iso = grHom + where + abstract + grHom : + IsGroupHom (coHomGr (suc (3 +ℕ n +ℕ n)) + (Pushout (λ _ → tt) (fst f)) .snd) + (λ x → Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) + (invEq (fst SphereHopfCohomIso) x)) + (ℤGroup .snd) + grHom = snd (compGroupIso + (GroupEquiv→GroupIso (invGroupEquiv (SphereHopfCohomIso))) + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) + +Hⁿ-Sⁿ≅ℤ-nice-generator : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ +Hⁿ-Sⁿ≅ℤ-nice-generator zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ +Hⁿ-Sⁿ≅ℤ-nice-generator (suc n) = + (λ i → Iso.inv (fst (suspensionAx-Sn (suc n) (suc n))) + (Hⁿ-Sⁿ≅ℤ-nice-generator n i)) + ∙ cong ∣_∣₂ + (funExt λ { north → refl + ; south → cong ∣_∣ₕ (merid north) + ; (merid a i) j → ∣ compPath-filler (merid a) + (sym (merid north)) (~ j) i ∣ₕ}) + +Hopfβ↦1 : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) + → Iso.fun (fst (Hopfβ-Iso n f)) (Hopfβ n f) ≡ 1 +Hopfβ↦1 n f = + sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) lem) + ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) + where + lem : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f + lem = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + (Hⁿ-Sⁿ≅ℤ-nice-generator _ i)) + ∙ cong ∣_∣₂ (funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl})) + +module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where + private + 2+n = 2 +ℕ n + H = HopfInvariantPush n (fst f) + + H→Sphere : coHom 2+n H → coHom 2+n (S₊ (suc (suc n))) + H→Sphere = sMap (_∘ inr) + + grHom : IsGroupHom (snd (coHomGr 2+n H)) + H→Sphere (snd (coHomGr 2+n (S₊ (suc (suc n))))) + grHom = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → refl) + + preSphere→H : (g : (S₊ (suc (suc n)) → coHomK 2+n)) + → H → coHomK (2 +ℕ n) + preSphere→H g (inl x) = 0ₖ _ + preSphere→H g (inr x) = g x -ₖ g north + preSphere→H g (push a i) = lem a i + where + lem : (a : S₊ (suc (suc (suc (n +ℕ n))))) + → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north + lem = + sphereElim _ + (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) + (isOfHLevelTrunc (4 +ℕ n) _ _)) + (sym (rCancelₖ _ (g north)) + ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) + + Sphere→H : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H + Sphere→H = sMap preSphere→H + + conCohom2+n : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ + conCohom2+n = + coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ + + HIPSphereCohomIso : + Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) + (coHom (2 +ℕ n) (S₊ (2 +ℕ n))) + Iso.fun HIPSphereCohomIso = H→Sphere + Iso.inv HIPSphereCohomIso = Sphere→H + Iso.rightInv HIPSphereCohomIso = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ g → pRec (squash₂ _ _) + (λ p → cong ∣_∣₂ (funExt λ x → cong (g x +ₖ_) (cong (-ₖ_) p) + ∙ rUnitₖ _ (g x))) + (conCohom2+n (g north)) + Iso.leftInv HIPSphereCohomIso = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ g → + pRec (squash₂ _ _) + (pRec (isPropΠ (λ _ → squash₂ _ _)) + (λ gn gtt → + trRec (isProp→isOfHLevelSuc n (squash₂ _ _)) + (λ r → cong ∣_∣₂ (funExt λ { + (inl x) → sym gtt + ; (inr x) → (λ i → g (inr x) -ₖ gn i) + ∙ rUnitₖ _ (g (inr x)) + ; (push a i) + → sphereElim _ + {A = λ a → + PathP (λ i → preSphere→H (λ x → g (inr x)) + (push a i) + ≡ g (push a i)) + (sym gtt) + ((λ i → g (inr (fst f a)) -ₖ gn i) + ∙ rUnitₖ _ (g (inr (fst f a))))} + (λ _ → isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) + (isOfHLevelPath (suc (suc (suc (suc (n +ℕ n))))) + (isOfHLevelPlus' {n = n} (suc (suc (suc (suc n)))) + (isOfHLevelTrunc (suc (suc (suc (suc n)))))) _ _) _ _) + r a i})) + (push-helper g gtt gn)) + (conCohom2+n (g (inr north)))) + (conCohom2+n (g (inl tt))) + where + push-helper : (g : HopfInvariantPush n (fst f) → coHomK (suc (suc n))) + → (gtt : g (inl tt) ≡ 0ₖ (suc (suc n))) + → (gn : g (inr north) ≡ 0ₖ (suc (suc n))) + → hLevelTrunc (suc n) + (PathP (λ i → preSphere→H (λ x → g (inr x)) (push north i) + ≡ g (push north i)) + (sym gtt) + ((λ i → g (inr (fst f north)) -ₖ gn i) + ∙ rUnitₖ _ (g (inr (fst f north))))) + push-helper g gtt gn = + isConnectedPathP _ (isConnectedPath _ (isConnectedKn _) _ _) _ _ .fst + + Hopfα-Iso : GroupIso (coHomGr (2 +ℕ n) (HopfInvariantPush n (fst f))) ℤGroup + Hopfα-Iso = + compGroupIso + (HIPSphereCohomIso , grHom) + (Hⁿ-Sⁿ≅ℤ (suc n)) + +Hopfα-Iso-α : (n : ℕ) (f : _) + → Iso.fun (fst (Hopfα-Iso n f)) (Hopfα n f) + ≡ 1 +Hopfα-Iso-α n f = + sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + where + hz : Iso.fun (HIPSphereCohomIso n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ + hz = refl + +⌣-α : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ +⌣-α n f = Hopfα n f ⌣ Hopfα n f + +HopfInvariant : (n : ℕ) + → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ +HopfInvariant n f = Iso.fun (fst (Hopfβ-Iso n f)) + (subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i → suc (suc (suc (+-suc n n i)))) + (⌣-α n f)) + +HopfInvariant-π' : (n : ℕ) → π' (3 +ℕ (n +ℕ n)) (S₊∙ (2 +ℕ n)) → ℤ +HopfInvariant-π' n = sRec isSetℤ (HopfInvariant n) + +HopfInvariant-π : (n : ℕ) → π (3 +ℕ (n +ℕ n)) (S₊∙ (2 +ℕ n)) → ℤ +HopfInvariant-π n = sRec isSetℤ λ x → HopfInvariant-π' n ∣ Ω→SphereMap _ x ∣₂ + + +-- Elimination principle for the pushout defining the Hopf Invariant +HopfInvariantPushElim : + ∀ {ℓ} n → (f : _) + → {P : HopfInvariantPush n f → Type ℓ} + → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) + → (e : P (inl tt)) + (g : (x : _) → P (inr x)) + (r : PathP (λ i → P (push north i)) e (g (f north))) + → (x : _) → P x +HopfInvariantPushElim n f {P = P} hlev e g r (inl x) = e +HopfInvariantPushElim n f {P = P} hlev e g r (inr x) = g x +HopfInvariantPushElim n f {P = P} hlev e g r (push a i₁) = help a i₁ + where + help : (a : Susp (Susp (S₊ (suc (n +ℕ n))))) + → PathP (λ i → P (push a i)) e (g (f a)) + help = + sphereElim _ + (sphereElim _ + (λ _ → isProp→isOfHLevelSuc (suc (suc (n +ℕ n))) + (isPropIsOfHLevel _)) + (isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) + (subst (isOfHLevel (suc (suc (suc (suc (n +ℕ n)))))) + (cong P (push north)) + hlev) _ _)) r diff --git a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda new file mode 100644 index 0000000000..5286267555 --- /dev/null +++ b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda @@ -0,0 +1,884 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.HopfInvariant.Homomorphism where + +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.SuspensionMap + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.GradedCommutativity +open import Cubical.ZCohomology.Gysin + +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Unit + +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.Truncation +open import Cubical.HITs.SetTruncation + renaming (elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + +-- The pushout describing the hopf invariant of the multiplication (∙Π) of +-- two maps (S³⁺²ⁿ →∙ S²⁺ⁿ) +C·Π : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +C·Π n f g = Pushout (λ _ → tt) (∙Π f g .fst) + +-- Another pushout, essentially the same, but starting with +-- S³⁺²ⁿ ∨ S³⁺²ⁿ. This will be used to break up the hopf invariant +-- of ∙Π f g. +C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +C* n f g = Pushout (λ _ → tt) (fst (∨→∙ f g)) + +-- The coequaliser of ∙Π and ∨→∙ +θ : ∀ {ℓ} {A : Pointed ℓ} → Susp (fst A) + → (Susp (fst A) , north) ⋁ (Susp (fst A) , north) +θ north = inl north +θ south = inr north +θ {A = A} (merid a i₁) = + ((λ i → inl ((merid a ∙ sym (merid (pt A))) i)) + ∙∙ push tt + ∙∙ λ i → inr ((merid a ∙ sym (merid (pt A))) i)) i₁ + +∙Π≡∨→∘θ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → (x : _) → ∙Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) +∙Π≡∨→∘θ n f g north = sym (snd f) +∙Π≡∨→∘θ n f g south = sym (snd g) +∙Π≡∨→∘θ n f g (merid a i₁) j = main j i₁ + where + help : cong (f ∨→ g) (cong (θ {A = S₊∙ _}) (merid a)) + ≡ (refl ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) + help = cong-∙∙ (f ∨→ g) ((λ i → inl ((merid a ∙ sym (merid north)) i))) + (push tt) + (λ i → inr ((merid a ∙ sym (merid north)) i)) + ∙∙ doubleCompPath≡compPath _ _ _ + ∙∙ cong (cong (f ∨→ g) + (λ i₂ → inl ((merid a ∙ (λ i₃ → merid north (~ i₃))) i₂)) ∙_) + (sym (assoc _ _ _)) + ∙∙ assoc _ _ _ + ∙∙ cong₂ _∙_ refl (compPath≡compPath' _ _) + + main : PathP (λ i → snd f (~ i) + ≡ snd g (~ i)) (λ i₁ → ∙Π f g .fst (merid a i₁)) + λ i₁ → (f ∨→ g) (θ {A = S₊∙ _} (merid a i₁)) + main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) + ∙∙ cong (fst f) (merid a ∙ sym (merid north)) + ∙∙ snd f) + ∙ (sym (snd g) + ∙∙ cong (fst g) (merid a ∙ sym (merid north)) + ∙∙ λ j → snd g (~ i ∧ j))) + ▷ sym help + +private + WedgeElim : ∀ {ℓ} (n : ℕ) + {P : S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)) → Type ℓ} + → ((x : _) → isOfHLevel (3 +ℕ n) (P x)) + → P (inl north) + → (x : _) → P x + WedgeElim n {P = P} x s (inl x₁) = + sphereElim _ {A = P ∘ inl} + (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)) s x₁ + WedgeElim n {P = P} x s (inr x₁) = + sphereElim _ {A = P ∘ inr} + (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) + (isPropIsOfHLevel (suc (suc (suc (n +ℕ n)))))) + (subst (isOfHLevel ((3 +ℕ n) +ℕ n)) (cong P (push tt)) + (isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)))) + (subst P (push tt) s) x₁ + WedgeElim n {P = P} x s (push a j) = + transp (λ i → P (push tt (i ∧ j))) (~ j) s + +H²C*≅ℤ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr (2 +ℕ n) (C* n f g)) ℤGroup +H²C*≅ℤ n f g = compGroupIso is (Hⁿ-Sⁿ≅ℤ (suc n)) + where + ∘inr : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) + ∘inr = sMap λ f → f ∘ inr + + invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) + → C* n f g → coHomK (2 +ℕ n) + invMapPrim h (inl x) = h (ptSn _) + invMapPrim h (inr x) = h x + invMapPrim h (push a i₁) = + WedgeElim n {P = λ a → h north ≡ h (fst (∨→∙ f g) a)} + (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) + (cong h (sym (snd f))) a i₁ + + invMap : coHom (2 +ℕ n) (S₊ (2 +ℕ n)) → (coHom (2 +ℕ n) (C* n f g)) + invMap = sMap invMapPrim + + ∘inrSect : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → ∘inr (invMap x) ≡ x + ∘inrSect = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ a → refl + + ∘inrRetr : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (∘inr x) ≡ x + ∘inrRetr = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) + ∙ sym (push (inl north))) + ; (inr x) → refl + ; (push a i) → lem₁ h a i}) + where + lem₁ : (h : C* n f g → coHomK (2 +ℕ n)) (a : _) + → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) + (cong h ((λ i → inr ((snd f) (~ i))) + ∙ sym (push (inl north)))) + refl + lem₁ h = + WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) + (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) + lem₂ + + where + lem₂ : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) + ≡ h (push (inl north) i)) + (cong h ((λ i → inr ((snd f) (~ i))) + ∙ sym (push (inl north)))) + refl + lem₂ i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) + ; (j = i0) → inr (snd f (~ i)) + ; (j = i1) → push (inl north) (i ∨ ~ k)}) + (inr (snd f (~ i ∧ ~ j)))) + + is : GroupIso (coHomGr (2 +ℕ n) (C* n f g)) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) + Iso.fun (fst is) = ∘inr + Iso.inv (fst is) = invMap + Iso.rightInv (fst is) = ∘inrSect + Iso.leftInv (fst is) = ∘inrRetr + snd is = makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → refl) + +H⁴-C·Π : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C·Π n f g)) + ℤGroup +H⁴-C·Π n f g = compGroupIso + (transportCohomIso (cong (3 +ℕ_) (+-suc n n))) (Hopfβ-Iso n (∙Π f g)) + +H²-C·Π : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr (2 +ℕ n) (C·Π n f g)) + ℤGroup +H²-C·Π n f g = Hopfα-Iso n (∙Π f g) + + +H⁴-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + (DirProd ℤGroup ℤGroup) +H⁴-C* n f g = + compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) + (compGroupIso (transportCohomIso (cong (2 +ℕ_) (+-suc n n))) + (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) _) + (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ _) (Hⁿ-Sⁿ≅ℤ _)))) + where + module Ms = MV _ _ _ (λ _ → tt) (fst (∨→∙ f g)) + fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) + (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) + (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + fst (fst fstIso) = Ms.d (suc (suc (n +ℕ suc n))) .fst + snd (fst fstIso) = + SES→isEquiv + (sym (fst (GroupPath _ _) + (isContr→≃Unit + (isContrΣ (subst isContr (isoToPath (invIso (fst (Hⁿ-Unit≅0 _)))) + isContrUnit) + (λ _ → subst isContr (isoToPath (invIso + (fst (Hⁿ-Sᵐ≅0 _ _ λ p → + ¬lemHopf 0 ((λ _ → north) , refl) + n n (cong suc (sym (+-suc n n)) ∙ p))))) + isContrUnit)) + , makeIsGroupHom λ _ _ → refl))) + (GroupPath _ _ .fst + (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) + , makeIsGroupHom λ _ _ → refl) + (GroupEquivDirProd + (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) + (GroupIso→GroupEquiv + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → + ¬lemHopf 0 ((λ _ → north) , refl) + n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) + ∙ p))))))) + (Ms.Δ (suc (suc (n +ℕ suc n)))) + (Ms.d (suc (suc (n +ℕ suc n)))) + (Ms.i (suc (suc (suc (n +ℕ suc n))))) + (Ms.Ker-d⊂Im-Δ _) + (Ms.Ker-i⊂Im-d _) + snd fstIso = Ms.d (suc (suc (n +ℕ suc n))) .snd + +module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) where + + C·Π' = C·Π n f g + + -- The generators of the cohomology groups of C* and C·Π + + βₗ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βₗ = Iso.inv (fst (H⁴-C* n f g)) (1 , 0) + + βᵣ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βᵣ = Iso.inv (fst (H⁴-C* n f g)) (0 , 1) + + β·Π : coHom ((2 +ℕ n) +' (2 +ℕ n)) C·Π' + β·Π = Iso.inv (fst (H⁴-C·Π n f g)) 1 + + α* : coHom (2 +ℕ n) (C* n f g) + α* = Iso.inv (fst (H²C*≅ℤ n f g)) 1 + + -- They can be difficult to work with in this form. + -- We give them simpler forms and prove that these are equivalent + βᵣ'-fun : C* n f g → coHomK ((4 +ℕ (n +ℕ n))) + βᵣ'-fun (inl x) = 0ₖ _ + βᵣ'-fun (inr x) = 0ₖ _ + βᵣ'-fun (push (inl x) i₁) = 0ₖ _ + βᵣ'-fun (push (inr x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ + βᵣ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ (~ i₂) i₁ + + βₗ'-fun : C* n f g → coHomK (4 +ℕ (n +ℕ n)) + βₗ'-fun (inl x) = 0ₖ _ + βₗ'-fun (inr x) = 0ₖ _ + βₗ'-fun (push (inl x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ + βₗ'-fun (push (inr x) i₁) = 0ₖ _ + βₗ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ i₂ i₁ + + βₗ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βₗ''-fun = subst (λ m → coHom m (C* n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∣ βₗ'-fun ∣₂ + + βᵣ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βᵣ''-fun = subst (λ m → coHom m (C* n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∣ βᵣ'-fun ∣₂ + + private + 0ₖ≡∨→ : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ + 0ₖ≡∨→ = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) + (cong ∣_∣ₕ (sym (snd f))) + + 0ₖ≡∨→-inr : 0ₖ≡∨→ (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) + 0ₖ≡∨→-inr = + (λ j → transport (λ i → 0ₖ (2 +ℕ n) + ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) + λ i → ∣ snd f (~ i ∨ j) ∣ₕ) + ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j + λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ + + α*'-fun : C* n f g → coHomK (2 +ℕ n) + α*'-fun (inl x) = 0ₖ _ + α*'-fun (inr x) = ∣ x ∣ + α*'-fun (push a i₁) = 0ₖ≡∨→ a i₁ + + α*' : coHom (2 +ℕ n) (C* n f g) + α*' = ∣ α*'-fun ∣₂ + + -- We also need the following three maps + jₗ : HopfInvariantPush n (fst f) → C* n f g + jₗ (inl x) = inl x + jₗ (inr x) = inr x + jₗ (push a i₁) = push (inl a) i₁ + + jᵣ : HopfInvariantPush n (fst g) → C* n f g + jᵣ (inl x) = inl x + jᵣ (inr x) = inr x + jᵣ (push a i₁) = push (inr a) i₁ + + q : C·Π' → C* n f g + q (inl x) = inl x + q (inr x) = inr x + q (push a i₁) = (push (θ a) ∙ λ i → inr (∙Π≡∨→∘θ n f g a (~ i))) i₁ + +α↦1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → Iso.fun (fst (H²C*≅ℤ n f g)) (α*' n f g) ≡ 1 +α↦1 n f g = + sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + +α≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → α* n f g ≡ α*' n f g +α≡ n f g = sym (Iso.leftInv (fst (H²C*≅ℤ n f g)) _) + ∙∙ cong (Iso.inv (fst (H²C*≅ℤ n f g))) lem + ∙∙ Iso.leftInv (fst (H²C*≅ℤ n f g)) _ + where + lem : Iso.fun (fst (H²C*≅ℤ n f g)) (α* n f g) + ≡ Iso.fun (fst (H²C*≅ℤ n f g)) (α*' n f g) + lem = (Iso.rightInv (fst (H²C*≅ℤ n f g)) (pos 1)) ∙ sym (α↦1 n f g) + +q-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (q n f g) (α* n f g) ≡ Hopfα n (∙Π f g) +q-α n f g = (λ i → coHomFun _ (q n f g) (α≡ n f g i)) + ∙ sym (Iso.leftInv is _) + ∙∙ cong (Iso.inv is) lem + ∙∙ Iso.leftInv is _ + where + is = fst (Hopfα-Iso n (∙Π f g)) + + lem : Iso.fun is (coHomFun _ (q n f g) (α*' n f g)) + ≡ Iso.fun is (Hopfα n (∙Π f g)) + lem = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + ∙∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + ∙∙ sym (Hopfα-Iso-α n (∙Π f g)) + +βₗ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → βₗ n f g ≡ βₗ''-fun n f g +βₗ≡ n f g = cong (∨-d ∘ subber₂) + (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (((suc (suc (n +ℕ n))))))))) help + ∙ lem) + ∙ funExt⁻ ∨-d∘subber ∣ wedgeGen ∣₂ + ∙ cong subber₃ (sym βₗ'-fun≡) + where + ∨-d = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (n +ℕ suc n))) .fst + ∨-d' = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (suc (n +ℕ n)))) .fst + + help : Iso.inv (fst (GroupIsoDirProd + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) + ≡ (∣ ∣_∣ ∣₂ , 0ₕ _) + help = ΣPathP ((Hⁿ-Sⁿ≅ℤ-nice-generator _) + , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + + subber₃ = subst (λ m → coHom m (C* n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + + subber₂ = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) + ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) + (sym (cong (2 +ℕ_) (+-suc n n)))) + + ∨-d∘subber : ∨-d ∘ subber₂ ≡ subber₃ ∘ ∨-d' + ∨-d∘subber = + funExt (λ a → + (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) + (C* n f g)) (~ j) + (MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) + (suc (suc (+-suc n n j))) .fst + (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) + (S₊∙ (suc (suc (suc (n +ℕ n)))) + ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j + a)))) + + wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) + → coHomK (suc (suc (suc (n +ℕ n))))) + wedgeGen (inl x) = ∣ x ∣ + wedgeGen (inr x) = 0ₖ _ + wedgeGen (push a i₁) = 0ₖ _ + + βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ ∨-d' ∣ wedgeGen ∣₂ + βₗ'-fun≡ = + cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push (inl x) i₁) → refl + ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ + ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) + + lem : Iso.inv (fst (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (((suc (suc (n +ℕ n))))))) + (∣ ∣_∣ ∣₂ , 0ₕ _) + ≡ ∣ wedgeGen ∣₂ + lem = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ + ; (inr x) → refl + ; (push a i₁) → refl}) + +βᵣ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → βᵣ n f g ≡ βᵣ''-fun n f g +βᵣ≡ n f g = + cong (∨-d ∘ subber₂) + (cong (Iso.inv (fst (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (suc (suc (n +ℕ n)))))) + help + ∙ lem) + ∙ funExt⁻ ∨-d∘subber ∣ wedgeGen ∣₂ + ∙ cong (subber₃) (sym βᵣ'-fun≡) + where + ∨-d = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (n +ℕ suc n))) .fst + ∨-d' = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (suc (n +ℕ n)))) .fst + + help : Iso.inv (fst (GroupIsoDirProd + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) + ≡ (0ₕ _ , ∣ ∣_∣ ∣₂) + help = + ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) + , (Hⁿ-Sⁿ≅ℤ-nice-generator _)) + + subber₃ = subst (λ m → coHom m (C* n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + + subber₂ = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) + ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) + (sym (cong (2 +ℕ_) (+-suc n n)))) + + ∨-d∘subber : ∨-d ∘ subber₂ ≡ subber₃ ∘ ∨-d' + ∨-d∘subber = + funExt (λ a → + (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) + (C* n f g)) (~ j) + (MV.d _ _ _ (λ _ → tt) + (fst (∨→∙ f g)) (suc (suc (+-suc n n j))) .fst + (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) + (S₊∙ (suc (suc (suc (n +ℕ n)))) + ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j + a)))) + + wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) + → coHomK (suc (suc (suc (n +ℕ n))))) + wedgeGen (inl x) = 0ₖ _ + wedgeGen (inr x) = ∣ x ∣ + wedgeGen (push a i₁) = 0ₖ _ + + + lem : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (suc (suc (n +ℕ n)))))) + (0ₕ _ , ∣ ∣_∣ ∣₂) + ≡ ∣ wedgeGen ∣₂ + lem = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ + ; (push a i₁) → refl}) + + βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ ∨-d' ∣ wedgeGen ∣₂ + βᵣ'-fun≡ = + cong ∣_∣₂ (funExt + λ { (inl x) → refl + ; (inr x) → refl + ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ + ; (push (inr x) i₁) → refl + ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) + +q-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (q n f g) (βₗ n f g) + ≡ β·Π n f g +q-βₗ n f g = cong (coHomFun _ (q n f g)) (βₗ≡ n f g) + ∙ transportLem + ∙ cong (subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) + (Hopfβ n (∙Π f g))) + ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) + (Hopfβ↦1 n (∙Π f g))) + where + transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) + (βₗ''-fun n f g) + ≡ subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n (∙Π f g)) + transportLem = + natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙ cong (subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → push-lem a i₁})) + where + push-lem : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → + PathP (λ i₁ → βₗ'-fun n f g ((push (θ x) + ∙ λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) i₁) + ≡ MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) + (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) + (λ _ → βₗ'-fun n f g (q n f g (inl tt))) + (λ _ → βₗ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) + push-lem x = + flipSquare (cong-∙ (βₗ'-fun n f g) + (push (θ x)) (λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) + ∙∙ sym (rUnit _) + ∙∙ lem₁ x) + + where + lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) + ≡ Kn→ΩKn+1 _ ∣ x ∣ + lem₁ north = refl + lem₁ south = + sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) + lem₁ (merid a j) k = merid-lem k j + where + p = pt (S₊∙ (suc (suc (n +ℕ n)))) + + merid-lem : + PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ + ≡ (sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north))) k) + (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + merid-lem = + (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) + (λ i₁ → inl ((merid a ∙ (sym (merid p))) i₁)) + (push tt) + ((λ i₁ → inr ((merid a ∙ (sym (merid p))) i₁))) + ∙ sym (compPath≡compPath' _ _) + ∙ cong (_∙ Kn→ΩKn+10ₖ _) + (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) + (merid a) (sym (merid north))) + ∙ sym (assoc _ _ _) + ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) + (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north)))))) + ◁ λ i j → compPath-filler + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + (sym (sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north)))) + (~ i) j + +q-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (q n f g) (βᵣ n f g) + ≡ β·Π n f g +q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) + ∙ transportLem + ∙ cong (subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) + (Hopfβ n (∙Π f g))) + ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) + (Hopfβ↦1 n (∙Π f g))) + where + transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) + (βᵣ''-fun n f g) + ≡ subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n (∙Π f g)) + transportLem = + natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙ cong (subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → push-lem a i₁})) + where + push-lem : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → + PathP + (λ i₁ → + βᵣ'-fun n f g ((push (θ x) ∙ λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) i₁) ≡ + MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) + (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) + (λ _ → βᵣ'-fun n f g (q n f g (inl tt))) + (λ _ → βᵣ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) + push-lem x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) + (λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) + ∙∙ sym (rUnit _) + ∙∙ lem₁ x) + where + lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) + ≡ Kn→ΩKn+1 _ ∣ x ∣ + lem₁ north = sym (Kn→ΩKn+10ₖ _) + lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north)) + lem₁ (merid a j) k = merid-lem k j + where + p = pt (S₊∙ (suc (suc (n +ℕ n)))) + + merid-lem : + PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) + ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north))) k) + (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + merid-lem = + (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) + (λ i₁ → inl ((merid a + ∙ (sym (merid p))) i₁)) + (push tt) + (λ i₁ → inr ((merid a + ∙ (sym (merid p))) i₁)) + ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) + (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) + (merid a) (sym (merid (ptSn _))))) + ∙∙ sym (doubleCompPath≡compPath _ _ _)) + ◁ symP (doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) + (merid a)) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) + (sym (merid north)))) + +jₗ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jₗ n f g) (α* n f g) + ≡ Hopfα n f +jₗ-α n f g = + cong (coHomFun _ (jₗ n f g)) (α≡ n f g) + ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst f) + (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) + (isOfHLevelPlus' {n = n} (4 +ℕ n) + (isOfHLevelTrunc (4 +ℕ n))) _ _) + refl + (λ _ → refl) + λ j → refl)) + +jₗ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jₗ n f g) (βₗ n f g) + ≡ subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n f) +jₗ-βₗ n f g = + cong (coHomFun _ (jₗ n f g)) (βₗ≡ n f g) + ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl})) + +jₗ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jₗ n f g) (βᵣ n f g) + ≡ 0ₕ _ +jₗ-βᵣ n f g = + cong (coHomFun _ (jₗ n f g)) (βᵣ≡ n f g) + ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + cool + where + cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jₗ n f g) + ∣ βᵣ'-fun n f g ∣₂ ≡ 0ₕ _ + cool = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl}) + +jᵣ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jᵣ n f g) (α* n f g) + ≡ Hopfα n g +jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α≡ n f g) + ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst g) + (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) + (isOfHLevelPlus' {n = n} (4 +ℕ n) + (isOfHLevelTrunc (4 +ℕ n))) _ _) + refl + (λ _ → refl) + (flipSquare (0ₖ≡∨→-inr n f g)))) + +jᵣ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jᵣ n f g) (βₗ n f g) ≡ 0ₕ _ +jᵣ-βₗ n f g = + cong (coHomFun _ (jᵣ n f g)) (βₗ≡ n f g) + ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + cool + where + cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jᵣ n f g) + ∣ βₗ'-fun n f g ∣₂ ≡ 0ₕ _ + cool = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl}) + + +jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jᵣ n f g) (βᵣ n f g) + ≡ subst (λ m → coHom m (HopfInvariantPush n (fst g))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n g) +jᵣ-βᵣ n f g = + cong (coHomFun _ (jᵣ n f g)) (βᵣ≡ n f g) + ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst g))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl})) + +genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + (βₗ n f g) + (βᵣ n f g) +genH²ⁿC* n f g = + Iso-pres-gen₂ (DirProd ℤGroup ℤGroup) + (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + (1 , 0) + (0 , 1) + gen₂ℤ×ℤ + (invGroupIso (H⁴-C* n f g)) + +private + H⁴C* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Group _ + H⁴C* n f g = coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + + X : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ + X n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .fst .fst + + Y : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ + Y n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .fst .snd + + α*≡⌣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → α* n f g ⌣ α* n f g ≡ ((X n f g) ℤ[ H⁴C* n f g ]· βₗ n f g) + +ₕ ((Y n f g) ℤ[ H⁴C* n f g ]· βᵣ n f g) + α*≡⌣ n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .snd + +coHomFun⌣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) + → (n : ℕ) (x y : coHom n B) + → coHomFun _ f (x ⌣ y) ≡ coHomFun _ f x ⌣ coHomFun _ f y +coHomFun⌣ f n = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl + +isHom-HopfInvariant : + (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) + → HopfInvariant n (∙Π f g) ≡ HopfInvariant n f + HopfInvariant n g +isHom-HopfInvariant n f g = + mainL + ∙ sym (cong₂ _+_ f-id g-id) + where + eq₁ : (Hopfα n (∙Π f g)) ⌣ (Hopfα n (∙Π f g)) + ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β·Π n f g)) + eq₁ = cong (λ x → x ⌣ x) (sym (q-α n f g) + ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α≡ n f g)) + ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α≡ n f g)) + ∙ α*≡⌣ n f g) + ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ + ∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) + ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) + (q-βₗ n f g)) + (homPresℤ· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) + ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) + (q-βᵣ n f g)) + ∙ sym (distrℤ· (coHomGr _ _) (β·Π n f g) (X n f g) (Y n f g)) + + eq₂ : Hopfα n f ⌣ Hopfα n f + ≡ (X n f g ℤ[ coHomGr _ _ ]· + subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n f)) + eq₂ = cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) + ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) + ∙∙ cong (coHomFun _ (jₗ n f g)) (α*≡⌣ n f g) + ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ + ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) + (homPresℤ· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) + ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) + (jₗ-βᵣ n f g)) + ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) + ∙∙ (rUnitₕ _ _ + ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) + + eq₃ : Hopfα n g ⌣ Hopfα n g + ≡ (Y n f g ℤ[ coHomGr _ _ ]· + subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n g)) + eq₃ = cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) + ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) + ∙∙ cong (coHomFun _ (jᵣ n f g)) (α*≡⌣ n f g) + ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ + ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) + ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) + (jᵣ-βₗ n f g)) + (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) + ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl + ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) + + transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} (n m : ℕ) (x : ℤ) (p : m ≡ n) + (g : fst (G n)) + → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) + ≡ (x ℤ[ G n ]· g) + transpLem {G = G} n m x = + J (λ n p → (g : fst (G n)) + → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) + ≡ (x ℤ[ G n ]· g)) + λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) + + mainL : HopfInvariant n (∙Π f g) ≡ X n f g + Y n f g + mainL = + cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (∙Π f g)))) + λ i₁ → suc (suc (suc (+-suc n n i₁)))) + eq₁) + ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) + (transpLem {G = λ x → coHomGr x + (HopfInvariantPush n (fst (∙Π f g)))} _ _ + (X n f g + Y n f g) + (λ i₁ → suc (suc (suc (+-suc n n i₁)))) + (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1)) + ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n (∙Π f g))) + (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1) + (X n f g + Y n f g) + ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) + (Iso.rightInv ((fst (Hopfβ-Iso n (∙Π f g)))) 1) + ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) + + f-id : HopfInvariant n f ≡ X n f g + f-id = + cong (Iso.fun (fst (Hopfβ-Iso n f))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i₁ → suc (suc (suc (+-suc n n i₁))))) eq₂) + ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) + (transpLem {G = λ x → coHomGr x + (HopfInvariantPush n (fst f))} _ _ + (X n f g) + ((cong (suc ∘ suc ∘ suc) (+-suc n n))) + (Hopfβ n f)) + ∙ homPresℤ· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) + ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n f) + ∙ rUnitℤ·ℤ (X n f g) + + g-id : HopfInvariant n g ≡ Y n f g + g-id = + cong (Iso.fun (fst (Hopfβ-Iso n g))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) + (λ i₁ → suc (suc (suc (+-suc n n i₁))))) + eq₃) + ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n g))) + (transpLem {G = λ x → coHomGr x + (HopfInvariantPush n (fst g))} _ _ + (Y n f g) + ((cong (suc ∘ suc ∘ suc) (+-suc n n))) + (Hopfβ n g)) + ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) + ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n g) + ∙∙ rUnitℤ·ℤ (Y n f g) + +GroupHom-HopfInvariant-π' : (n : ℕ) + → GroupHom (π'Gr (suc (suc (n +ℕ n))) (S₊∙ (suc (suc n)))) ℤGroup +fst (GroupHom-HopfInvariant-π' n) = HopfInvariant-π' n +snd (GroupHom-HopfInvariant-π' n) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) + (isHom-HopfInvariant n)) + +GroupHom-HopfInvariant-π : (n : ℕ) + → GroupHom (πGr (suc (suc (n +ℕ n))) (S₊∙ (suc (suc n)))) ℤGroup +fst (GroupHom-HopfInvariant-π n) = HopfInvariant-π n +snd (GroupHom-HopfInvariant-π n) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) + λ p q → cong (HopfInvariant-π' n) + (IsGroupHom.pres· + (snd (invGroupIso (π'Gr≅πGr (suc (suc (n +ℕ n))) + (S₊∙ (suc (suc n)))))) + ∣ p ∣₂ ∣ q ∣₂) + ∙ IsGroupHom.pres· (GroupHom-HopfInvariant-π' n .snd) + ∣ Ω→SphereMap _ p ∣₂ ∣ Ω→SphereMap _ q ∣₂) diff --git a/Cubical/Homotopy/HopfInvariant/HopfMap.agda b/Cubical/Homotopy/HopfInvariant/HopfMap.agda new file mode 100644 index 0000000000..587d1e5f9a --- /dev/null +++ b/Cubical/Homotopy/HopfInvariant/HopfMap.agda @@ -0,0 +1,559 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains a proof that the generator of Π₃S² has +hopf invariant ±1. +-} +module Cubical.Homotopy.HopfInvariant.HopfMap where + +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.Hopf +open S¹Hopf +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.HSpace + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.Gysin + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sum +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Unit + +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Join +open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec) + +HopfMap : S₊∙ 3 →∙ S₊∙ 2 +fst HopfMap x = JoinS¹S¹→TotalHopf (Iso.inv (IsoSphereJoin 1 1) x) .fst +snd HopfMap = refl + +-- We use the Hopf fibration in order to connect it to the Gysin Sequence +module hopfS¹ = + Hopf S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) + +S¹Hopf = hopfS¹.Hopf +E* = hopfS¹.TotalSpacePush²' +IsoE*join = hopfS¹.joinIso₂ +IsoTotalHopf' = hopfS¹.joinIso₁ +CP² = hopfS¹.TotalSpaceHopfPush² +fibr = hopfS¹.P + +TotalHopf' : Type _ +TotalHopf' = Σ (S₊ 2) S¹Hopf + +IsoJoins : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) +IsoJoins = cong (join S¹) (isoToPath (IsoSphereJoin 1 1)) + +-- CP² is 1-connected +conCP² : (x y : CP²) → ∥ x ≡ y ∥₂ +conCP² x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP²' x) (conCP²' y) + where + conCP²' : (x : CP²) → ∥ x ≡ inl tt ∥₂ + conCP²' (inl x) = ∣ refl ∣₂ + conCP²' (inr x) = sphereElim 1 {A = λ x → ∥ inr x ≡ inl tt ∥₂} + (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x + conCP²' (push a i) = main a i + where + indLem : ∀ {ℓ} {A : hopfS¹.TotalSpaceHopfPush → Type ℓ} + → ((a : _) → isProp (A a)) + → A (inl base) + → ((a : hopfS¹.TotalSpaceHopfPush) → A a) + indLem {A = A} p b = + PushoutToProp p + (sphereElim 0 (λ _ → p _) b) + (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) + + main : (a : hopfS¹.TotalSpaceHopfPush) + → PathP (λ i → ∥ Path CP² (push a i) (inl tt) ∥₂) + (conCP²' (inl tt)) (conCP²' (inr (hopfS¹.induced a))) + main = indLem (λ _ → isOfHLevelPathP' 1 squash₂ _ _) + λ j → ∣ (λ i → push (inl base) (~ i ∧ j)) ∣₂ + +module GysinS² = Gysin (CP² , inl tt) fibr conCP² 2 idIso refl + +E'S⁴Iso : Iso GysinS².E' (S₊ 5) +E'S⁴Iso = + compIso IsoE*join + (compIso (Iso→joinIso idIso (IsoSphereJoin 1 1)) + (IsoSphereJoin 1 3)) + +isContrH³E : isContr (coHom 3 (GysinS².E')) +isContrH³E = + subst isContr (sym (isoToPath + (fst (Hⁿ-Sᵐ≅0 2 4 + λ p → snotz (sym (cong (predℕ ∘ predℕ) p))))) + ∙ cong (coHom 3) (sym (isoToPath E'S⁴Iso))) + isContrUnit + +isContrH⁴E : isContr (coHom 4 (GysinS².E')) +isContrH⁴E = + subst isContr (sym (isoToPath + (fst (Hⁿ-Sᵐ≅0 3 4 + λ p → snotz (sym (cong (predℕ ∘ predℕ ∘ predℕ) p))))) + ∙ cong (coHom 4) (sym (isoToPath E'S⁴Iso))) + isContrUnit + +-- We will need a bunch of elimination principles +joinS¹S¹→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (inl base) + → (x : _) → P x +joinS¹S¹→Groupoid P grp b = + transport (λ i → (x : (isoToPath (invIso (IsoSphereJoin 1 1))) i) + → P (transp (λ j → isoToPath (invIso (IsoSphereJoin 1 1)) (i ∨ j)) i x)) + (sphereElim _ (λ _ → grp _) b) + +TotalHopf→Gpd : ∀ {ℓ} (P : hopfS¹.TotalSpaceHopfPush → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (inl base) + → (x : _) → P x +TotalHopf→Gpd P grp b = + transport (λ i → (x : sym (isoToPath IsoTotalHopf') i) + → P (transp (λ j → isoToPath IsoTotalHopf' (~ i ∧ ~ j)) i x)) + (joinS¹S¹→Groupoid _ (λ _ → grp _) b) + +TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) hopfS¹.Hopf → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (north , base) + → (x : _) → P x +TotalHopf→Gpd' P grp b = + transport (λ i → (x : ua (_ , hopfS¹.isEquivTotalSpaceHopfPush→TotalSpace) i) + → P (transp (λ j → ua (_ , hopfS¹.isEquivTotalSpaceHopfPush→TotalSpace) + (i ∨ j)) i x)) + (TotalHopf→Gpd _ (λ _ → grp _) b) + +CP²→Groupoid : ∀ {ℓ} {P : CP² → Type ℓ} + → ((x : _) → is2Groupoid (P x)) + → (b : P (inl tt)) + → (s2 : ((x : _) → P (inr x))) + → PathP (λ i → P (push (inl base) i)) b (s2 north) + → (x : _) → P x +CP²→Groupoid {P = P} grp b s2 pp (inl x) = b +CP²→Groupoid {P = P} grp b s2 pp (inr x) = s2 x +CP²→Groupoid {P = P} grp b s2 pp (push a i₁) = lem a i₁ + where + lem : (a : hopfS¹.TotalSpaceHopfPush) → PathP (λ i → P (push a i)) b (s2 _) + lem = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp + +-- The function inducing the iso H²(S²) ≅ H²(CP²) +H²S²→H²CP²-raw : (S₊ 2 → coHomK 2) → (CP² → coHomK 2) +H²S²→H²CP²-raw f (inl x) = 0ₖ _ +H²S²→H²CP²-raw f (inr x) = f x -ₖ f north +H²S²→H²CP²-raw f (push a i) = + TotalHopf→Gpd (λ x → 0ₖ 2 + ≡ f (hopfS¹.TotalSpaceHopfPush→TotalSpace x .fst) -ₖ f north) + (λ _ → isOfHLevelTrunc 4 _ _) + (sym (rCancelₖ 2 (f north))) + a i + +H²S²→H²CP² : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP² .fst +H²S²→H²CP² = sMap H²S²→H²CP²-raw + +cancel-H²S²→H²CP² : (f : CP² → coHomK 2) + → ∥ H²S²→H²CP²-raw (f ∘ inr) ≡ f ∥ +cancel-H²S²→H²CP² f = + pRec squash + (λ p → pRec squash + (λ f → ∣ funExt f ∣) + (cancelLem p)) + (connLem (f (inl tt))) + where + connLem : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ + connLem = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ + + cancelLem : (p : f (inl tt) ≡ 0ₖ 2) + → ∥ ((x : CP²) → H²S²→H²CP²-raw (f ∘ inr) x ≡ f x) ∥ + cancelLem p = trRec squash (λ pp → + ∣ CP²→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + (sym p) + (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) + ∙∙ (λ i → f (inr x) -ₖ p i) + ∙∙ rUnitₖ 2 (f (inr x))) pp ∣) + help + where + help : + hLevelTrunc 1 + (PathP (λ i₁ → H²S²→H²CP²-raw (f ∘ inr) (push (inl base) i₁) + ≡ f (push (inl base) i₁)) + (sym p) + (((λ i₁ → f (inr north) -ₖ f (push (inl base) (~ i₁))) ∙∙ + (λ i₁ → f (inr north) -ₖ p i₁) ∙∙ rUnitₖ 2 (f (inr north))))) + help = isConnectedPathP 1 (isConnectedPath 2 (isConnectedKn 1) _ _) _ _ .fst + +H²CP²≅H²S² : GroupIso (coHomGr 2 CP²) (coHomGr 2 (S₊ 2)) +Iso.fun (fst H²CP²≅H²S²) = sMap (_∘ inr) +Iso.inv (fst H²CP²≅H²S²) = H²S²→H²CP² +Iso.rightInv (fst H²CP²≅H²S²) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → trRec {B = Iso.fun (fst H²CP²≅H²S²) (Iso.inv (fst H²CP²≅H²S²) ∣ f ∣₂) + ≡ ∣ f ∣₂} + (isOfHLevelPath 2 squash₂ _ _) + (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) + (Iso.fun (PathIdTruncIso _) + (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) +Iso.leftInv (fst H²CP²≅H²S²) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → pRec (squash₂ _ _) (cong ∣_∣₂) (cancel-H²S²→H²CP² f) +snd H²CP²≅H²S² = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) + +H²CP²≅ℤ : GroupIso (coHomGr 2 CP²) ℤGroup +H²CP²≅ℤ = compGroupIso H²CP²≅H²S² (Hⁿ-Sⁿ≅ℤ 1) + +-- ⌣ gives a groupEquiv H²(CP²) ≃ H⁴(CP²) +⌣Equiv : GroupEquiv (coHomGr 2 CP²) (coHomGr 4 CP²) +fst (fst ⌣Equiv) = GysinS².⌣-hom 2 .fst +snd (fst ⌣Equiv) = isEq⌣ + where + isEq⌣ : isEquiv (GysinS².⌣-hom 2 .fst) + isEq⌣ = + SES→isEquiv + (GroupPath _ _ .fst (invGroupEquiv + (isContr→≃Unit isContrH³E + , makeIsGroupHom λ _ _ → refl))) + (GroupPath _ _ .fst (invGroupEquiv + (isContr→≃Unit isContrH⁴E + , makeIsGroupHom λ _ _ → refl))) + (GysinS².susp∘ϕ 1) + (GysinS².⌣-hom 2) + (GysinS².p-hom 4) + (GysinS².Ker-⌣e⊂Im-Susp∘ϕ _) + (GysinS².Ker-p⊂Im-⌣e _) +snd ⌣Equiv = GysinS².⌣-hom 2 .snd + +-- The generator of H²(CP²) +genCP² : coHom 2 CP² +genCP² = ∣ CP²→Groupoid (λ _ → isOfHLevelTrunc 4) + (0ₖ _) + ∣_∣ + refl ∣₂ + +inrInjective : (f g : CP² → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ + → Path (coHom 2 CP²) ∣ f ∣₂ ∣ g ∣₂ +inrInjective f g = pRec (squash₂ _ _) + (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) + (λ pp → cong ∣_∣₂ + (funExt (CP²→Groupoid + (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + id + (funExt⁻ p) + (compPathR→PathP pp)))) + (conn2 (f (inl tt)) (g (inl tt)) id + (cong f (push (inl base)) + ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) + (conn (f (inl tt)) (g (inl tt)))) + + where + conn : (x y : coHomK 2) → ∥ x ≡ y ∥ + conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash) + (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) + + conn2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) + conn2 x y p q = + Iso.fun (PathIdTruncIso _) + (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) + +-- A couple of basic lemma concerning the hSpace structure on S¹ +private + invLooperLem₁ : (a x : S¹) + → (invEq (hopfS¹.μ-eq a) x) * a ≡ (invLooper a * x) * a + invLooperLem₁ a x = + secEq (hopfS¹.μ-eq a) x + ∙∙ cong (_* x) (rCancelS¹ a) + ∙∙ AssocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x + ∙ commS¹ _ _ + + invLooperLem₂ : (a x : S¹) → invEq (hopfS¹.μ-eq a) x ≡ invLooper a * x + invLooperLem₂ a x = sym (retEq (hopfS¹.μ-eq a) (invEq (hopfS¹.μ-eq a) x)) + ∙∙ cong (invEq (hopfS¹.μ-eq a)) (invLooperLem₁ a x) + ∙∙ retEq (hopfS¹.μ-eq a) (invLooper a * x) + + rotLoop² : (a : S¹) → Path (a ≡ a) (λ i → rotLoop (rotLoop a i) (~ i)) refl + rotLoop² = + sphereElim 0 (λ _ → isGroupoidS¹ _ _ _ _) + λ i j → hcomp (λ {k → λ { (i = i1) → base + ; (j = i0) → base + ; (j = i1) → base}}) + base + + +-- We prove that the generator of CP² given by Gysin is the same one +-- as genCP², which is much easier to work with +Gysin-e≡genCP² : GysinS².e ≡ genCP² +Gysin-e≡genCP² = + inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (main x)) south) ∣ + where + mainId : (x : Σ (S₊ 2) hopfS¹.Hopf) + → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ + mainId = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) + ; south → λ y → cong ∣_∣ₕ (sym (merid y)) + ; (merid a i) → main a i} + where + main : (a : S¹) → PathP (λ i → (y : hopfS¹.Hopf (merid a i)) + → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) + (λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y))) + λ y → cong ∣_∣ₕ (sym (merid y)) + main a = + toPathP + (funExt λ x → + cong (transport (λ i₁ + → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣)) + ((λ i → (λ z → cong ∣_∣ₕ + (merid base + ∙ sym (merid (transport + (λ j → uaInvEquiv (hopfS¹.μ-eq a) (~ i) j) x))) z)) + ∙ λ i → cong ∣_∣ₕ (merid base + ∙ sym (merid (transportRefl (invEq (hopfS¹.μ-eq a) x) i)))) + ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) + ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i + (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong ∣_∣ₕ (merid base + ∙ sym (merid (invLooperLem₂ a x i)))) i)) + ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) + (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) + (sym (merid (invLooper a * x))))) + ∙ cong sym (SuspS¹-hom (invLooper a) x) + ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) + ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) + ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x + ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + (sym (λ i₁ → ∣ (merid (invLooper a) + ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + ∙ cong₂ _∙_ (cong sym (SuspS¹-inv a) + ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) + (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) + ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) + ∙∙ (λ j → (λ i₁ → ∣ merid a (~ i₁ ∨ j) ∣) ∙ ((λ i₁ → ∣ merid a (i₁ ∨ j) ∣) + ∙ (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) + ∙ (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) + ∙ (λ i₁ → ∣ merid x (~ i₁) ∣ₕ)) + ∙∙ sym (lUnit _) + ∙∙ sym (assoc _ _ _) + ∙∙ (sym (lUnit _) ∙ sym (lUnit _) ∙ sym (lUnit _))) + + gen' : (x : S₊ 2) → preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2 + fst (gen' x) north = ∣ north ∣ + fst (gen' x) south = ∣ x ∣ + fst (gen' x) (merid a i₁) = mainId (x , a) (~ i₁) + snd (gen' x) = refl + + gen'Id : GysinS².c (inr north) .fst ≡ gen' north .fst + gen'Id = cong fst (GysinS².cEq (inr north) ∣ push (inl base) ∣₂) + ∙ (funExt lem) + where + lem : (qb : _) → + ∣ (subst (fst ∘ preThom.Q (CP² , inl tt) fibr) + (sym (push (inl base))) qb) ∣ + ≡ gen' north .fst qb + lem north = refl + lem south = cong ∣_∣ₕ (sym (merid base)) + lem (merid a i) j = + hcomp (λ k → + λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ + ; (i = i1) → ∣ merid base (~ j) ∣ + ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ + ; (j = i1) → ∣ compPath-filler + (merid base) (sym (merid a)) k (~ i) ∣ₕ}) + (hcomp (λ k → + λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ + ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ + ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ + ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) + ∣ south ∣) + + setHelp : (x : S₊ 2) + → isSet (preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2) + setHelp = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) + (isOfHLevel↑∙' 0 1) + + main : (x : S₊ 2) → (GysinS².c (inr x) ≡ gen' x) + main = sphereElim _ (λ x → isOfHLevelPath 2 (setHelp x) _ _) + (→∙Homogeneous≡ (isHomogeneousKn _) gen'Id) + +isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) → Type ℓ +isGenerator≃ℤ G g = + Σ[ e ∈ GroupIso G ℤGroup ] abs (Iso.fun (fst e) g) ≡ 1 + +isGenerator≃ℤ-e : isGenerator≃ℤ (coHomGr 2 CP²) GysinS².e +isGenerator≃ℤ-e = + subst (isGenerator≃ℤ (coHomGr 2 CP²)) (sym Gysin-e≡genCP²) + (H²CP²≅ℤ , refl) + +-- Alternative definition of the hopfMap +HopfMap' : S₊ 3 → S₊ 2 +HopfMap' x = + hopfS¹.TotalSpaceHopfPush→TotalSpace + (Iso.inv IsoTotalHopf' + (Iso.inv (IsoSphereJoin 1 1) x)) .fst + +hopfMap≡HopfMap' : HopfMap ≡ (HopfMap' , refl) +hopfMap≡HopfMap' = + ΣPathP ((funExt (λ x → + cong (λ x → JoinS¹S¹→TotalHopf x .fst) + (sym (Iso.rightInv IsoTotalHopf' + (Iso.inv (IsoSphereJoin 1 1) x))) + ∙ sym (lem (Iso.inv IsoTotalHopf' + (Iso.inv (IsoSphereJoin 1 1) x))))) + , flipSquare (sym (rUnit refl) ◁ λ _ _ → north)) + where + lem : (x : _) → hopfS¹.TotalSpaceHopfPush→TotalSpace x .fst + ≡ JoinS¹S¹→TotalHopf (Iso.fun IsoTotalHopf' x) .fst + lem (inl x) = refl + lem (inr x) = refl + lem (push (base , snd₁) i) = refl + lem (push (loop i₁ , a) i) k = merid (rotLoop² a (~ k) i₁) i + +CP²' : Type _ +CP²' = Pushout {A = S₊ 3} (λ _ → tt) HopfMap' + +PushoutReplaceBase : ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} + {f : A → C} {g : A → D} (e : B ≃ A) + → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g +PushoutReplaceBase {f = f} {g = g} = + EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) refl + +CP2≡CP²' : CP²' ≡ CP² +CP2≡CP²' = + PushoutReplaceBase + (isoToEquiv (compIso (invIso (IsoSphereJoin 1 1)) (invIso IsoTotalHopf'))) + +-- packaging everything up: +⌣equiv→pres1 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) + → (g₁ : coHom 2 G) (h₁ : coHom 2 H) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] + abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] + abs (fst (fst ϕ) h₁) ≡ 1))) + → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) + → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) + → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 +⌣equiv→pres1 {G = G} = J (λ H _ → (g₁ : coHom 2 G) (h₁ : coHom 2 H) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] + abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] + abs (fst (fst ϕ) h₁) ≡ 1))) + → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) + → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) + → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1) + help + where + help : (g₁ h₁ : coHom 2 G) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] + abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] + abs (fst (fst ϕ) h₁) ≡ 1))) + → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) + → (3rdEq : GroupEquiv (coHomGr 4 G) ℤGroup) + → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 + help g h (ϕ , idg) (ψ , idh) ⌣eq ξ = + ⊎→abs _ _ + (groupEquivPresGen _ + (compGroupEquiv main (compGroupEquiv (invGroupEquiv main) ψ)) + h (abs→⊎ _ _ (cong abs (cong (fst (fst ψ)) (retEq (fst main) h)) + ∙ idh)) (compGroupEquiv main ξ)) + where + lem₁ : ((fst (fst ψ) h) ≡ 1) ⊎ (fst (fst ψ) h ≡ -1) + → abs (fst (fst ϕ) h) ≡ 1 + lem₁ p = ⊎→abs _ _ (groupEquivPresGen _ ψ h p ϕ) + + lem₂ : ((fst (fst ϕ) h) ≡ 1) ⊎ (fst (fst ϕ) h ≡ -1) + → ((fst (fst ϕ) g) ≡ 1) ⊎ (fst (fst ϕ) g ≡ -1) + → (h ≡ g) ⊎ (h ≡ (-ₕ g)) + lem₂ (inl x) (inl x₁) = + inl (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) + ∙∙ retEq (fst ϕ) g) + lem₂ (inl x) (inr x₁) = + inr (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) x + ∙ IsGroupHom.presinv (snd (invGroupEquiv ϕ)) (negsuc zero) + ∙∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g))) + lem₂ (inr x) (inl x₁) = + inr (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) x + ∙∙ (IsGroupHom.presinv (snd (invGroupEquiv ϕ)) 1 + ∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g)))) + lem₂ (inr x) (inr x₁) = + inl (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) + ∙∙ retEq (fst ϕ) g) + + -ₕeq : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → Iso (coHom n A) (coHom n A) + Iso.fun (-ₕeq n) = -ₕ_ + Iso.inv (-ₕeq n) = -ₕ_ + Iso.rightInv (-ₕeq n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → -ₖ^2 (f x)) + Iso.leftInv (-ₕeq n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → -ₖ^2 (f x)) + + theEq : coHom 2 G ≃ coHom 4 G + theEq = compEquiv (_ , ⌣eq) (isoToEquiv (-ₕeq 4)) + + lem₃ : (h ≡ g) ⊎ (h ≡ (-ₕ g)) → isEquiv {A = coHom 2 G} (_⌣ h) + lem₃ (inl x) = subst isEquiv (λ i → _⌣ (x (~ i))) ⌣eq + lem₃ (inr x) = + subst isEquiv (funExt (λ x → -ₕDistᵣ 2 2 x g) ∙ (λ i → _⌣ (x (~ i)))) + (theEq .snd) + + main : GroupEquiv (coHomGr 2 G) (coHomGr 4 G) + fst main = _ , + (lem₃ (lem₂ (abs→⊎ _ _ (lem₁ (abs→⊎ _ _ idh))) (abs→⊎ _ _ idg))) + snd main = + makeIsGroupHom λ g1 g2 → rightDistr-⌣ _ _ g1 g2 h + +-- The hopf invariant is ±1 for both definitions of the hopf map +HopfInvariant-HopfMap' : + abs (HopfInvariant zero (HopfMap' , λ _ → HopfMap' (snd (S₊∙ 3)))) ≡ 1 +HopfInvariant-HopfMap' = + cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (HopfMap' , refl)))) + (transportRefl (⌣-α 0 (HopfMap' , refl)))) + ∙ ⌣equiv→pres1 (sym CP2≡CP²') + GysinS².e (Hopfα zero (HopfMap' , refl)) + (l isGenerator≃ℤ-e) + (GroupIso→GroupEquiv (Hopfα-Iso 0 (HopfMap' , refl)) , refl) + (snd (fst ⌣Equiv)) + (GroupIso→GroupEquiv (Hopfβ-Iso zero (HopfMap' , refl))) + where + l : Σ[ ϕ ∈ GroupIso (coHomGr 2 CP²) ℤGroup ] + (abs (Iso.fun (fst ϕ) GysinS².e) ≡ 1) + → Σ[ ϕ ∈ GroupEquiv (coHomGr 2 CP²) ℤGroup ] + (abs (fst (fst ϕ) GysinS².e) ≡ 1) + l p = (GroupIso→GroupEquiv (fst p)) , (snd p) + +HopfInvariant-HopfMap : abs (HopfInvariant zero HopfMap) ≡ 1 +HopfInvariant-HopfMap = cong abs (cong (HopfInvariant zero) hopfMap≡HopfMap') + ∙ HopfInvariant-HopfMap' diff --git a/Cubical/Homotopy/HopfInvariant/Whitehead.agda b/Cubical/Homotopy/HopfInvariant/Whitehead.agda new file mode 100644 index 0000000000..ff7dc1bd74 --- /dev/null +++ b/Cubical/Homotopy/HopfInvariant/Whitehead.agda @@ -0,0 +1,808 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.HopfInvariant.Whitehead where + +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.SuspensionMap +open import Cubical.Homotopy.HopfInvariant.HopfMap +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.Group.Pi4S3.Tricky +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso +open import Cubical.Homotopy.Loopspace + + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.GradedCommutativity +open import Cubical.ZCohomology.Gysin + +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Unit + +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact + +open import Cubical.Relation.Nullary + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.Truncation +open import Cubical.HITs.SetTruncation + renaming (elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation renaming (map to pMap ; rec to pRec) + + +-- abs (HopfInvariant-π' 0 [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π') ≡ 2 + + +open import Cubical.Homotopy.Whitehead +open import Cubical.HITs.S1 +open import Cubical.HITs.Join +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Unit + +¬lem : (n m : ℕ) → ¬ suc (n +ℕ m) ≡ m +¬lem n zero = snotz +¬lem n (suc m) p = ¬lem n m (sym (+-suc n m) ∙ cong predℕ p) + +H²-genₗ' : coHom 2 (S₊ 2 × S₊ 2) +H²-genₗ' = ∣ (λ x → ∣ fst x ∣) ∣₂ + +H²-genᵣ' : coHom 2 (S₊ 2 × S₊ 2) +H²-genᵣ' = ∣ (λ x → ∣ snd x ∣) ∣₂ + +abstract + H²-genₗabs : coHom 2 (S₊ 2 × S₊ 2) + H²-genₗabs = H²-genₗ' + + H²-genᵣabs : coHom 2 (S₊ 2 × S₊ 2) + H²-genᵣabs = H²-genᵣ' + + absInd : ∀ {ℓ} (A : coHom 2 (S₊ 2 × S₊ 2) → coHom 2 (S₊ 2 × S₊ 2) → Type ℓ) + → A H²-genₗabs H²-genᵣabs + → A H²-genₗ' H²-genᵣ' + absInd A p = p + + absInd⁻ : ∀ {ℓ} (A : coHom 2 (S₊ 2 × S₊ 2) → coHom 2 (S₊ 2 × S₊ 2) → Type ℓ) + → A H²-genₗ' H²-genᵣ' + → A H²-genₗabs H²-genᵣabs + absInd⁻ A p = p + + +∨map = joinTo⋁ {A = S₊∙ 1} {B = S₊∙ 1} +module mv = MV ((S₊∙ 2) ⋁ (S₊∙ 2)) Unit (join S¹ S¹) ∨map (λ _ → tt) + +downFun : (n m : ℕ) + → (S₊ (suc (suc n)) → S₊ (suc m) → coHomK (suc (suc ((suc n) +ℕ m)))) + → S₊ (suc n) → S₊ (suc m) → coHomK (suc (suc (n +ℕ m))) +downFun n m f x y = + ΩKn+1→Kn ((suc ((suc n) +ℕ m))) + (sym (rCancelₖ _ (f north y)) + ∙∙ cong (λ x → f x y -ₖ f north y) (merid x ∙ sym (merid (ptSn (suc n)))) + ∙∙ rCancelₖ _ (f north y)) + +upFun : (n m : ℕ) + → (S₊ (suc n) → S₊ (suc m) → coHomK (suc (suc (n +ℕ m)))) + → (S₊ (suc (suc n)) → S₊ (suc m) → coHomK (suc (suc ((suc n) +ℕ m)))) +upFun n m f north y = 0ₖ _ +upFun n m f south y = 0ₖ _ +upFun n m f (merid a i) y = + Kn→ΩKn+1 _ (f a y) i + +funLem : (n m : ℕ) → (f : S₊ (suc m) → coHomK (suc n)) + → ¬ (n ≡ m) + → ∥ f ≡ (λ _ → 0ₖ (suc n)) ∥ +funLem n m f p = + Iso.fun PathIdTrunc₀Iso + (isContr→isProp + (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Sᵐ≅0 n m p)) isContrUnit) + ∣ f ∣₂ (0ₕ _)) + +upDownIso : (n m : ℕ) + → GroupIso (coHomGr (suc (((suc n) +ℕ m))) + (S₊ (suc n) × S₊ (suc m))) + (coHomGr (suc (suc ((suc n) +ℕ m))) + (S₊ (suc (suc n)) × S₊ (suc m))) +Iso.fun (fst (upDownIso n m)) = sMap (uncurry ∘ upFun n m ∘ curry) +Iso.inv (fst (upDownIso n m)) = sMap ((uncurry ∘ downFun n m ∘ curry)) +Iso.rightInv (fst (upDownIso n m)) = + sElim (λ _ → isSetPathImplicit) + λ f → Iso.inv PathIdTrunc₀Iso + (pRec squash + (uncurry (λ g p + → pMap (λ gid → funExt λ {(x , y) + → (λ i → uncurry (upFun n m (curry (uncurry (downFun n m (curry (p (~ i))))))) (x , y)) + ∙∙ main g gid x y + ∙∙ funExt⁻ p (x , y)}) + (gIdLem g))) + (rewr f)) + where + lem : (n m : ℕ) → ¬ suc (suc (n +ℕ m)) ≡ m + lem n zero = snotz + lem n (suc m) p = lem n m (cong suc (sym (+-suc n m)) ∙ cong predℕ p) + + charac-fun : (S₊ (suc n) → S₊ (suc m) → typ (Ω (coHomK-ptd (suc (suc (suc n +ℕ m)))))) + → S₊ (suc (suc n)) × S₊ (suc m) → coHomK (suc (suc (suc n +ℕ m))) + charac-fun g (north , y) = 0ₖ _ + charac-fun g (south , y) = 0ₖ _ + charac-fun g (merid c i , y) = g c y i + + rewr : (f : S₊ (suc (suc n)) × S₊ (suc m) → coHomK (suc (suc (suc n +ℕ m)))) + → ∃[ g ∈ (S₊ (suc n) → S₊ (suc m) → typ (Ω (coHomK-ptd (suc (suc (suc n +ℕ m)))))) ] + charac-fun g ≡ f + rewr f = + pMap (λ p → (λ x y → sym (funExt⁻ p y) + ∙∙ (λ i → f ((merid x ∙ sym (merid (ptSn (suc n)))) i , y)) + ∙∙ funExt⁻ p y) + , funExt λ { (north , y) → sym (funExt⁻ p y) + ; (south , y) → sym (funExt⁻ p y) + ∙ λ i → f (merid (ptSn (suc n)) i , y) + ; (merid a i , y) j → + hcomp (λ k → λ { (i = i0) → p (~ j ∧ k) y + ; (i = i1) → compPath-filler' + (sym (funExt⁻ p y)) + (λ i → f (merid (ptSn (suc n)) i , y)) k j + ; (j = i1) → f (merid a i , y) }) + (f (compPath-filler (merid a) (sym (merid (ptSn (suc n)))) (~ j) i + , y))}) + (funLem _ _ (λ x → f (north , x)) + (lem n m)) + + gIdLem : (g : S₊ (suc n) → S₊ (suc m) → typ (Ω (coHomK-ptd (suc (suc (suc n +ℕ m)))))) + → ∥ (g (ptSn _)) ≡ (λ _ → refl) ∥ + gIdLem g = + Iso.fun PathIdTrunc₀Iso + (isContr→isProp + (isOfHLevelRetractFromIso 0 + ((invIso (fst (coHom≅coHomΩ _ (S₊ (suc m)))))) + (0ₕ _ , sElim (λ _ → isProp→isSet (squash₂ _ _)) + λ f → pRec (squash₂ _ _) (sym ∘ cong ∣_∣₂) + (funLem _ _ f (¬lem n m)))) + ∣ g (ptSn (suc n)) ∣₂ ∣ (λ _ → refl) ∣₂) + + main : (g : _) → (g (ptSn _)) ≡ (λ _ → refl) + → (x : S₊ (suc (suc n))) (y : S₊ (suc m)) + → uncurry + (upFun n m (curry (uncurry (downFun n m (curry (charac-fun g)))))) + (x , y) + ≡ charac-fun g (x , y) + main g gid north y = refl + main g gid south y = refl + main g gid (merid a i) y j = help j i + where + help : (λ i → uncurry + (upFun n m (curry (uncurry (downFun n m (curry (charac-fun g)))))) + (merid a i , y)) ≡ g a y + help = (λ i → Iso.rightInv (Iso-Kn-ΩKn+1 _) + ((sym (rCancel≡refl _ i) + ∙∙ cong-∙ (λ x → rUnitₖ _ (charac-fun g (x , y)) i) (merid a) (sym (merid (ptSn (suc n)))) i + ∙∙ rCancel≡refl _ i)) i) + ∙∙ sym (rUnit _) + ∙∙ (cong (g a y ∙_) (cong sym (funExt⁻ gid y)) + ∙ sym (rUnit (g a y))) +Iso.leftInv (fst (upDownIso n m)) = + sElim (λ _ → isSetPathImplicit) + λ f → pRec (squash₂ _ _) + (λ id + → cong ∣_∣₂ + (funExt (λ {(x , y) → (λ i → ΩKn+1→Kn _ (sym (rCancel≡refl _ i) + ∙∙ ((λ j → rUnitₖ _ + (upFun n m (curry f) + ((merid x ∙ sym (merid (ptSn (suc n)))) j) y) i)) + ∙∙ rCancel≡refl _ i)) + ∙ (λ i → ΩKn+1→Kn _ (rUnit + (cong-∙ (λ r → upFun n m (curry f) r y) + (merid x) (sym (merid (ptSn (suc n)))) i) (~ i))) + ∙ cong (ΩKn+1→Kn _) (cong (Kn→ΩKn+1 _ (f (x , y)) ∙_) + (cong sym (cong (Kn→ΩKn+1 _) (funExt⁻ id y) ∙ (Kn→ΩKn+10ₖ _))) + ∙ sym (rUnit _)) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (x , y))}))) + (funLem (suc n +ℕ m) m (λ x → f (ptSn _ , x)) + (¬lem n m)) +snd (upDownIso n m) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ + (funExt λ { (north , y) → refl + ; (south , y) → refl + ; (merid a i , y) j + → (sym (∙≡+₂ _ (Kn→ΩKn+1 _ (f (a , y))) (Kn→ΩKn+1 _ (g (a , y)))) + ∙ sym (Kn→ΩKn+1-hom _ (f (a , y)) (g (a , y)))) (~ j) i})) + +toBaseIso : (n m : ℕ) → GroupIso (coHomGr (suc (((suc n) +ℕ m))) + (S₊ (suc n) × S₊ (suc m))) + ((coHomGr (suc (suc m)) + (S₊ (suc zero) × S₊ (suc m)))) +toBaseIso zero m = idGroupIso +toBaseIso (suc n) m = + compGroupIso (invGroupIso (upDownIso n m)) (toBaseIso n m) + +BaseIso→ : (m : ℕ) + → (S₊ (suc m) → coHomK (suc m)) + → S₊ (suc zero) → S₊ (suc m) → coHomK (suc (suc m)) +BaseIso→ m f base y = 0ₖ _ +BaseIso→ m f (loop i) y = Kn→ΩKn+1 (suc m) (f y) i + +BaseIso← : (m : ℕ) + → (S₊ (suc zero) → S₊ (suc m) → coHomK (suc (suc m))) + → (S₊ (suc m) → coHomK (suc m)) +BaseIso← m f x = + ΩKn+1→Kn (suc m) + (sym (rCancelₖ _ (f base x)) ∙∙ ((λ i → f (loop i) x -ₖ f base x)) ∙∙ rCancelₖ _ (f base x)) + +BaseIso : (m : ℕ) → GroupIso (coHomGr (suc m) (S₊ (suc m))) + ((coHomGr (suc (suc m)) + (S₊ (suc zero) × S₊ (suc m)))) + +Iso.fun (fst (BaseIso m)) = sMap (uncurry ∘ BaseIso→ m) +Iso.inv (fst (BaseIso m)) = sMap (BaseIso← m ∘ curry) +Iso.rightInv (fst (BaseIso m)) = + sElim (λ _ → isSetPathImplicit) + λ f → Iso.inv PathIdTrunc₀Iso + (pMap (uncurry (λ g p + → funExt λ {(x , y) + → (λ i → uncurry (BaseIso→ m (BaseIso← m (curry (p i)))) (x , y)) + ∙∙ main g x y + ∙∙ sym (funExt⁻ p (x , y))})) + (ss f)) + where + characFun : (x : S₊ (suc m) → coHomK (suc m)) + → S₊ 1 × S₊ (suc m) → coHomK (suc (suc m)) + characFun f (base , y) = 0ₖ _ + characFun f (loop i , y) = Kn→ΩKn+1 _ (f y) i + + main : (g : _) → (x : S¹) (y : S₊ (suc m)) → uncurry (BaseIso→ m (BaseIso← m (curry (characFun g)))) + (x , y) + ≡ characFun g (x , y) + main g base y = refl + main g (loop i) y j = help j i + where + help : cong (λ x → BaseIso→ m (BaseIso← m (curry (characFun g))) x y) loop + ≡ Kn→ΩKn+1 _ (g y) + help = Iso.rightInv (Iso-Kn-ΩKn+1 (suc m)) + (sym (rCancelₖ _ (0ₖ _)) ∙∙ ((λ i → Kn→ΩKn+1 _ (g y) i -ₖ 0ₖ _)) ∙∙ rCancelₖ _ (0ₖ _)) + ∙∙ (λ i → sym (rCancel≡refl _ i) ∙∙ (λ j → rUnitₖ _ (Kn→ΩKn+1 _ (g y) j) i) ∙∙ rCancel≡refl _ i) + ∙∙ sym (rUnit _) + + lem : (m : ℕ) → ¬ suc m ≡ m + lem zero = snotz + lem (suc m) p = lem m (cong predℕ p) + + ss : (f : S₊ 1 × S₊ (suc m) → coHomK (suc (suc m))) + → ∃[ g ∈ (S₊ (suc m) → coHomK (suc m)) ] f ≡ characFun g + ss f = + pMap (λ p → (λ x → ΩKn+1→Kn _ (sym (funExt⁻ p x) ∙∙ (λ i → f (loop i , x)) ∙∙ funExt⁻ p x)) + , funExt λ { (base , y) → funExt⁻ p y + ; (loop i , x) j → + hcomp (λ k → λ {(i = i0) → funExt⁻ p x j + ; (i = i1) → funExt⁻ p x j + ; (j = i0) → f (loop i , x) + ; (j = i1) → + Iso.rightInv (Iso-Kn-ΩKn+1 (suc m)) + (sym (funExt⁻ p x) ∙∙ (λ i → f (loop i , x)) ∙∙ funExt⁻ p x) (~ k) i}) + (doubleCompPath-filler + (sym (funExt⁻ p x)) (λ i → f (loop i , x)) (funExt⁻ p x) j i)}) + (funLem (suc m) m (λ x → f (base , x)) (lem m)) +Iso.leftInv (fst (BaseIso m)) = + sElim (λ _ → isSetPathImplicit) + λ f + → cong ∣_∣₂ (funExt λ x + → cong (ΩKn+1→Kn (suc m)) + ((λ i → sym (rCancel≡refl _ i) + ∙∙ cong (λ z → rUnitₖ _ (BaseIso→ m f z x) i) loop + ∙∙ rCancel≡refl _ i) ∙ sym (rUnit (Kn→ΩKn+1 (suc m) (f x)))) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) +snd (BaseIso m) = + makeIsGroupHom + (sElim2 + (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ + (funExt λ { (base , y) → refl + ; (loop i , y) j → + (sym (∙≡+₂ _ (Kn→ΩKn+1 _ (f y)) (Kn→ΩKn+1 _ (g y))) + ∙ sym (Kn→ΩKn+1-hom _ (f y) (g y))) (~ j) i})) + +GroupIso2 : ∀ {ℓ} {A : Type ℓ} (n : ℕ) + → GroupIso (×coHomGr (suc n) A Unit) (coHomGr (suc n) A) +Iso.fun (fst (GroupIso2 n)) = fst +Iso.inv (fst (GroupIso2 n)) x = x , 0ₕ (suc n) +Iso.rightInv (fst (GroupIso2 n)) _ = refl +Iso.leftInv (fst (GroupIso2 n)) x = + ΣPathP (refl , isContr→isProp (isContrHⁿ-Unit n) (0ₕ (suc n)) (snd x)) +snd (GroupIso2 n) = makeIsGroupHom λ _ _ → refl + +theIso : GroupIso (coHomGr 2 (S₊ 2)) (coHomGr 4 (S₊ 2 × S₊ 2)) +theIso = (compGroupIso (BaseIso 1) (upDownIso 0 1)) + +Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ : (n m : ℕ) + → GroupIso (coHomGr ((suc n) +' (suc m)) + (S₊ (suc n) × S₊ (suc m))) + ℤGroup +Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ zero m = + compGroupIso (invGroupIso (BaseIso m)) (Hⁿ-Sⁿ≅ℤ m) +Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ (suc n) m = + compGroupIso + (invGroupIso (upDownIso n m)) + (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ n m) + + +theIso⌣ : Iso.fun (fst (theIso)) ∣ ∣_∣ₕ ∣₂ ≡ H²-genₗ' ⌣ H²-genᵣ' +theIso⌣ = + cong ∣_∣₂ (funExt (uncurry + (λ { north y → refl + ; south y → refl + ; (merid a i) y j → Kn→ΩKn+1 3 (lem₃ a y j) i}))) + where + lem₃ : (a : S¹) (y : S₊ 2) → BaseIso→ 1 ∣_∣ₕ a y ≡ _⌣ₖ_ {n = 1} {m = 2} ∣ a ∣ₕ ∣ y ∣ₕ + lem₃ base y = refl + lem₃ (loop i) y = refl + +abstract + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs : (n m : ℕ) + → GroupIso (coHomGr ((suc n) +' (suc m)) + (S₊ (suc n) × S₊ (suc m))) + ℤGroup + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs = Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ + + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs≡ : (n m : ℕ) → Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs n m ≡ Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ n m + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs≡ n m = refl + + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-⌣ : Iso.fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1)) (H²-genₗ' ⌣ H²-genᵣ') ≡ 1 + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-⌣ = + cong (Iso.fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1))) (sym theIso⌣) + ∙ speedUp ∣_∣ₕ + ∙ refl -- Computation! :-) + where + speedUp : (f : _) → Iso.fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1)) (Iso.fun (fst (theIso)) ∣ f ∣₂) + ≡ Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 1)) ∣ f ∣₂ + speedUp f i = (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 1))) (Iso.leftInv (fst theIso) ∣ f ∣₂ i) + + + +open import Cubical.Data.Sum renaming (rec to ⊎rec) +open Iso + + +Pushout≅Torus2 : GroupIso (coHomGr 2 (S₊ 2 × S₊ 2)) (coHomGr 2 (Pushout joinTo⋁ (λ _ → tt))) +Pushout≅Torus2 = coHomIso 2 (invIso (Iso-Susp×Susp-cofibJoinTo⋁ S¹ S¹ base base)) + +GroupHom2 : + GroupHom (coHomGr 2 (Pushout ∨map (λ _ → tt))) + (×coHomGr 2 (S₊∙ 2 ⋁ S₊∙ 2) Unit) +GroupHom2 = mv.i 2 + +GroupEquiv2 : GroupEquiv + (coHomGr 2 (Pushout ∨map (λ _ → tt))) + (×coHomGr 2 (S₊∙ 2 ⋁ S₊∙ 2) Unit) +fst (fst GroupEquiv2) = fst GroupHom2 +snd (fst GroupEquiv2) = + SES→isEquiv + (isContr→≡UnitGroup + (isOfHLevelRetractFromIso 0 + (compIso (coHomIso 1 (invIso (IsoSphereJoin 1 1)) .fst) + (fst (Hⁿ-Sᵐ≅0 0 2 λ p → snotz (sym p)))) + isContrUnit)) + ((isContr→≡UnitGroup + (isOfHLevelRetractFromIso 0 + (compIso (coHomIso 2 (invIso (IsoSphereJoin 1 1)) .fst) + (fst (Hⁿ-Sᵐ≅0 1 2 λ p → snotz (cong predℕ (sym p))))) + isContrUnit))) + (mv.d 1) (mv.i 2) + (mv.Δ 2) + (mv.Ker-i⊂Im-d 1) + (mv.Ker-Δ⊂Im-i 2) +snd GroupEquiv2 = snd GroupHom2 + +coHom2S²×S² : GroupIso (coHomGr 2 (S₊ 2 × S₊ 2)) (DirProd ℤGroup ℤGroup) +coHom2S²×S² = + compGroupIso Pushout≅Torus2 + (compGroupIso + (GroupEquiv→GroupIso GroupEquiv2) + (compGroupIso (GroupIso2 1) + (compGroupIso (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 1) + (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ 1) (Hⁿ-Sⁿ≅ℤ 1))))) + +zz : GroupIso (coHomGr 4 (S₊ 2 × S₊ 2)) (coHomGr 4 (Pushout joinTo⋁ (λ _ → tt))) +zz = coHomIso 4 (invIso (Iso-Susp×Susp-cofibJoinTo⋁ S¹ S¹ base base)) + +GroupEquiv1 : GroupEquiv + (coHomGr 3 (join S¹ S¹)) + (coHomGr 4 (Pushout ∨map (λ _ → tt))) +fst (fst GroupEquiv1) = fst (mv.d 3) +snd (fst GroupEquiv1) = ss + where + abstract + ss : isEquiv (fst (mv.d 3)) + ss = SES→isEquiv + (isContr→≡UnitGroup + (isOfHLevelΣ 0 + (isOfHLevelRetractFromIso 0 + (compIso + (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 2)) + (compIso + (prodIso (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) + (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p)))) + lUnit×Iso)) + isContrUnit) + (λ _ → isContrHⁿ-Unit 2))) + (isContr→≡UnitGroup + (isOfHLevelΣ 0 + (isOfHLevelRetractFromIso 0 + (compIso + (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 3)) + (compIso + (prodIso (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) + (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p)))) + lUnit×Iso)) + isContrUnit) + λ _ → isContrHⁿ-Unit 3)) + (mv.Δ 3) + (mv.d 3) + (mv.i 4) + (mv.Ker-d⊂Im-Δ 3) + (mv.Ker-i⊂Im-d 3) +snd GroupEquiv1 = snd (mv.d 3) + + +coHom4S²×S² : GroupIso (coHomGr 4 (S₊ 2 × S₊ 2)) ℤGroup +coHom4S²×S² = + compGroupIso zz + (compGroupIso + (invGroupIso (GroupEquiv→GroupIso GroupEquiv1)) + (compGroupIso + (coHomIso 3 (invIso (IsoSphereJoin 1 1))) + (Hⁿ-Sⁿ≅ℤ 2))) + + +CHopf : Type +CHopf = HopfInvariantPush 0 fold∘W + +Hopfαfold∘W = Hopfα 0 (fold∘W , refl) +Hopfβfold∘W = Hopfβ 0 (fold∘W , refl) + +lem₂ : Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ ≃ fst thePushout∙ +lem₂ = compEquiv + (compEquiv pushoutSwitchEquiv + (isoToEquiv (PushoutDistr.PushoutDistrIso fold⋁ W λ _ → tt))) + pushoutSwitchEquiv + +lem₁ : Pushout W (λ _ → tt) ≃ cofibW S¹ S¹ base base +lem₁ = pushoutEquiv W (λ _ → tt) joinTo⋁ (λ _ → tt) + (isoToEquiv (invIso (IsoSphereJoin 1 1))) + (idEquiv _) + (idEquiv _) + refl + refl + +lem : Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ + ≃ fst (Pushout⋁↪fold⋁∙ (S₊∙ 2)) +lem = pushoutEquiv inl _ ⋁↪ fold⋁ + (idEquiv _) + (compEquiv lem₁ (isoToEquiv (invIso (Iso-Susp×Susp-cofibJoinTo⋁ S¹ S¹ base base)))) + (idEquiv _) + (Susp×Susp→cofibW≡ S¹ S¹ base base) + refl + +lem∙ : (Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ , inr north) + ≃∙ (Pushout⋁↪fold⋁∙ (S₊∙ 2)) +fst lem∙ = lem +snd lem∙ = sym (push (inl north)) + + +CHopfIso : Iso CHopf (Pushout⋁↪fold⋁ (S₊∙ 2)) +CHopfIso = + compIso (invIso (equivToIso + (compEquiv (compEquiv pushoutSwitchEquiv + (isoToEquiv (PushoutDistr.PushoutDistrIso fold⋁ W λ _ → tt))) + pushoutSwitchEquiv))) + (equivToIso lem) + + + +S²×S² : Pointed₀ +fst S²×S² = S₊ 2 × S₊ 2 +snd S²×S² = north , north + +q : (S²×S² →∙ (Pushout⋁↪fold⋁ (S₊∙ 2) , inl (north , north))) +fst q = inl +snd q = refl + +qHom : GroupHom (coHomGr 4 (Pushout⋁↪fold⋁ (S₊∙ 2))) + (coHomGr 4 (S₊ 2 × S₊ 2)) +qHom = coHomMorph 4 (fst q) + +qHomGen : (n : ℕ) → + GroupHom (coHomGr n (Pushout⋁↪fold⋁ (S₊∙ 2))) + (coHomGr n (S₊ 2 × S₊ 2)) +qHomGen = λ n → coHomMorph n (fst q) + +module mv2 = MV _ _ (S₊∙ 2 ⋁ S₊∙ 2) ⋁↪ fold⋁ + +gaha : isEquiv (fst (mv2.i 4)) +gaha = + SES→isEquiv + (isContr→≡UnitGroup + (isOfHLevelRetractFromIso 0 + (compIso + (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 2)) + (compIso + (prodIso (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) + (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p)))) + rUnit×Iso)) + isContrUnit)) + ((isContr→≡UnitGroup + (isOfHLevelRetractFromIso 0 + (compIso + (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 3)) + (compIso + (prodIso (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) + (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p)))) + rUnit×Iso)) + isContrUnit))) + (mv2.d 3) + (mv2.i 4) + (mv2.Δ 4) + (mv2.Ker-i⊂Im-d 3) + (mv2.Ker-Δ⊂Im-i 4) + +qHomPP : {A B C : Type} + → Unit ≡ C + → (f : A → B × C) + → isEquiv f + → isEquiv (fst ∘ f) +qHomPP {A = A} {B = B} = + J (λ C _ → (f : A → B × C) → isEquiv f → isEquiv (fst ∘ f)) + λ f eq → record { equiv-proof = + λ b → ((fst (fst (equiv-proof eq (b , tt)))) + , cong fst (fst (equiv-proof eq (b , tt)) .snd)) + , λ y → ΣPathP ((cong fst (equiv-proof eq (b , tt) .snd ((fst y) + , ΣPathP ((snd y) , refl)))) + , λ i j → equiv-proof eq (b , tt) .snd ((fst y) + , ΣPathP ((snd y) , refl)) i .snd j .fst) } + + +isEquiv-qHom : GroupEquiv (coHomGr 4 (Pushout⋁↪fold⋁ (S₊∙ 2))) + (coHomGr 4 (S₊ 2 × S₊ 2)) +fst (fst isEquiv-qHom) = qHom .fst +snd (fst isEquiv-qHom) = + subst isEquiv + (funExt (sElim (λ _ → isSetPathImplicit) (λ _ → refl))) + (qHomPP (isoToPath + (invIso (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))))) _ gaha) +snd isEquiv-qHom = qHom .snd + +coHomCHopfIso : (n : ℕ) + → GroupIso (coHomGr n CHopf) (coHomGr n (Pushout⋁↪fold⋁ (S₊∙ 2))) +coHomCHopfIso n = invGroupIso (coHomIso n CHopfIso) + +H⁴-gen = Iso.inv (fst coHom4S²×S²) 1 +H²-genₗ = Iso.inv (fst coHom2S²×S²) (1 , 0) +H²-genᵣ = Iso.inv (fst coHom2S²×S²) (0 , 1) + + + +open import Cubical.ZCohomology.RingStructure.RingLaws + +module postul (is : GroupIso (coHomGr 4 (S₊ 2 × S₊ 2)) ℤGroup) + (isEq : (fun (fst is) (H²-genₗ' ⌣ H²-genᵣ') ≡ 1)) where + α' : coHom 2 (Pushout⋁↪fold⋁ (S₊∙ 2)) + α' = fun (fst (coHomCHopfIso 2)) Hopfαfold∘W + + β' : coHom 4 (Pushout⋁↪fold⋁ (S₊∙ 2)) + β' = fun (fst (coHomCHopfIso 4)) Hopfβfold∘W + + ⌣→≡ : (α' ⌣ α' ≡ β' +ₕ β') ⊎ (α' ⌣ α' ≡ -ₕ (β' +ₕ β')) + → (Hopfαfold∘W ⌣ Hopfαfold∘W ≡ Hopfβfold∘W +ₕ Hopfβfold∘W) + ⊎ (Hopfαfold∘W ⌣ Hopfαfold∘W ≡ -ₕ (Hopfβfold∘W +ₕ Hopfβfold∘W)) + ⌣→≡ (inl x) = inl ((λ i → leftInv (fst (coHomCHopfIso 2)) Hopfαfold∘W (~ i) + ⌣ leftInv (fst (coHomCHopfIso 2)) Hopfαfold∘W (~ i)) + ∙∙ cong (inv (fst (coHomCHopfIso 4))) x + ∙∙ leftInv (fst (coHomCHopfIso 4)) (Hopfβfold∘W +ₕ Hopfβfold∘W)) + ⌣→≡ (inr x) = inr ((λ i → leftInv (fst (coHomCHopfIso 2)) Hopfαfold∘W (~ i) + ⌣ leftInv (fst (coHomCHopfIso 2)) Hopfαfold∘W (~ i)) + ∙∙ cong (inv (fst (coHomCHopfIso 4))) x + ∙∙ leftInv (fst (coHomCHopfIso 4)) + (-ₕ (Hopfβfold∘W +ₕ Hopfβfold∘W))) + + Q : (qHom .fst β' ≡ H²-genₗ' ⌣ H²-genᵣ') + ⊎ (qHom .fst β' ≡ -ₕ (H²-genₗ' ⌣ H²-genᵣ')) + Q = + ⊎rec + (λ p → inl (sym (leftInv (fst is) (qHom .fst β')) + ∙∙ cong (inv (fst is)) (p ∙ sym isEq) + ∙∙ leftInv (fst is) (H²-genₗ' ⌣ H²-genᵣ'))) + (λ p → inr (sym (leftInv (fst is) (qHom .fst β')) + ∙∙ cong (inv (fst is)) + (p ∙ sym (cong (GroupStr.inv (snd ℤGroup)) isEq)) + ∙∙ (IsGroupHom.presinv (invGroupIso is .snd) (fun (fst is) (H²-genₗ' ⌣ H²-genᵣ')) + ∙ cong -ₕ_ (leftInv (fst is) (H²-genₗ' ⌣ H²-genᵣ'))))) + p2 + where + h : GroupEquiv (coHomGr 4 (HopfInvariantPush 0 fold∘W)) ℤGroup + h = compGroupEquiv (GroupIso→GroupEquiv (coHomCHopfIso 4)) + (compGroupEquiv + isEquiv-qHom + (GroupIso→GroupEquiv is)) + + p2 : (fst (fst h) Hopfβfold∘W ≡ 1) ⊎ (fst (fst h) Hopfβfold∘W ≡ -1) + p2 = groupEquivPresGen _ (GroupIso→GroupEquiv (Hopfβ-Iso 0 (fold∘W , refl))) + Hopfβfold∘W (inl (Hopfβ↦1 0 (fold∘W , refl))) h + + qpres⌣ : (x y : coHom 2 _) + → fst qHom (x ⌣ y) ≡ fst (qHomGen 2) x ⌣ fst (qHomGen 2) y + qpres⌣ = sElim2 (λ _ _ → isSetPathImplicit) λ _ _ → refl + +-- H¹(...) → H²(S² ∨ S²) → H²(S² × S²) × H²(S²) → H²(...) + + i* : coHom 2 (Pushout⋁↪fold⋁ (S₊∙ 2)) → coHom 2 (S₊ 2) + i* = coHomFun 2 inr + + kata : i* (coHomFun 2 (Iso.inv CHopfIso) Hopfαfold∘W) ≡ ∣ ∣_∣ ∣₂ + kata = refl + + incl∘q : fst q ∘ (λ x → x , north) ≡ inr + incl∘q = funExt λ x → (push (inl x)) + + incr∘q : fst q ∘ (λ x → north , x) ≡ inr + incr∘q = funExt λ x → (push (inr x)) + + q≡' : coHomFun 2 (λ x → x , north) ∘ fst (qHomGen 2) ≡ i* + q≡' = funExt (sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (funExt λ x → cong f (push (inl x)))) + + kzz : (x : _) → i* x ≡ ∣ ∣_∣ ∣₂ → fst (qHomGen 2) x + ≡ ∣ (λ x → ∣ fst x ∣) ∣₂ + +ₕ ∣ (λ x → ∣ snd x ∣) ∣₂ + kzz = sElim (λ _ → isSetΠ λ _ → isSetPathImplicit) + λ f p → Cubical.HITs.PropositionalTruncation.rec + (squash₂ _ _) + (λ r → cong ∣_∣₂ (funExt (uncurry + (wedgeconFun 1 1 (λ _ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + (λ x → cong f (push (inr x)) ∙∙ funExt⁻ r x ∙∙ refl) + ((λ x → cong f (push (inl x)) ∙∙ funExt⁻ r x ∙∙ sym (rUnitₖ 2 ∣ x ∣ₕ))) + (cong (_∙∙ funExt⁻ r north ∙∙ refl) + (cong (cong f) λ j i → push (push tt j) i)))))) + (Iso.fun PathIdTrunc₀Iso p) + + + q≡'2 : coHomFun 2 (λ x → north , x) ∘ fst (qHomGen 2) ≡ i* + q≡'2 = funExt (sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (funExt λ x → cong f (push (inr x)))) + + q₂ : fst (qHomGen 2) α' + ≡ ∣ (λ x → ∣ fst x ∣) ∣₂ +ₕ ∣ (λ x → ∣ snd x ∣) ∣₂ + q₂ = kzz ((coHomFun 2 (Iso.inv CHopfIso) Hopfαfold∘W)) refl + + q₂' : fst (qHomGen 2) α' ≡ H²-genₗ' +ₕ H²-genᵣ' + q₂' = q₂ + + triv⌣ : (a : S¹) → cong₂ (_⌣ₖ_ {n = 2} {m = 2}) (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (merid a)) ≡ λ _ → ∣ north ∣ₕ + triv⌣ a = cong₂Funct (_⌣ₖ_ {n = 2} {m = 2}) (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (merid a)) + ∙ sym (rUnit λ j → _⌣ₖ_ {n = 2} {m = 2} ∣ merid a j ∣ₕ ∣ north ∣) + ∙ (λ i j → ⌣ₖ-0ₖ 2 2 ∣ merid a j ∣ₕ i) + + distLem : (x : coHom 4 (S₊ 2 × S₊ 2)) → -ₕ ((-ₕ x) +ₕ (-ₕ x)) ≡ x +ₕ x + distLem = + sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (funExt λ x → cong -ₖ_ (sym (-distrₖ 4 (f x) (f x))) + ∙ -ₖ^2 (f x +ₖ f x)) + + MainId : ((fst qHom) (α' ⌣ α') ≡ qHom .fst (β' +ₕ β')) + ⊎ ((fst qHom) (α' ⌣ α') ≡ qHom .fst (-ₕ (β' +ₕ β'))) + MainId = + ⊎rec + (λ id → inl (h ∙ H + ∙ cong (λ x → x +ₕ x) (sym id) + ∙ sym (IsGroupHom.pres· (snd qHom) β' β'))) + (λ id → inr (h ∙ H + ∙ ((sym (distLem (x ⌣ y)) + ∙ cong -ₕ_ (cong (λ x → x +ₕ x) (sym id))) + ∙ cong (-ₕ_) (IsGroupHom.pres· (snd qHom) β' β')) + ∙ sym (IsGroupHom.presinv (snd qHom) (β' +ₕ β')))) + Q + where + H²-genₗ⌣H²-genₗ : H²-genₗ' ⌣ H²-genₗ' ≡ 0ₕ 4 + H²-genₗ⌣H²-genₗ = + cong ∣_∣₂ + (funExt (uncurry λ { north y → refl + ; south y → refl + ; (merid a i) y j → triv⌣ a j i})) + + H²-genᵣ⌣H²-genᵣ : H²-genᵣ' ⌣ H²-genᵣ' ≡ 0ₕ 4 + H²-genᵣ⌣H²-genᵣ = cong ∣_∣₂ + (funExt (uncurry λ { x north → refl + ; x south → refl + ; x (merid a i) j → triv⌣ a j i})) + + x = H²-genₗ' + y = H²-genᵣ' + + -ₕId : (x : coHom 4 (S₊ 2 × S₊ 2)) → (-ₕ^ 2 · 2) x ≡ x + -ₕId = sElim (λ _ → isSetPathImplicit) + λ a → cong ∣_∣₂ (funExt λ x → -ₖ-gen-inl-left 2 2 tt (inl tt) (a x)) + + H : (H²-genₗ' +ₕ H²-genᵣ') ⌣ (H²-genₗ' +ₕ H²-genᵣ') + ≡ (H²-genₗ' ⌣ H²-genᵣ') +ₕ (H²-genₗ' ⌣ H²-genᵣ') + H = (x +ₕ y) ⌣ (x +ₕ y) ≡⟨ leftDistr-⌣ 2 2 (x +ₕ y) x y ⟩ + ((x +ₕ y) ⌣ x) +ₕ ((x +ₕ y) ⌣ y) ≡⟨ cong₂ _+ₕ_ (rightDistr-⌣ 2 2 x y x) (rightDistr-⌣ 2 2 x y y) ⟩ + ((x ⌣ x +ₕ y ⌣ x)) +ₕ (x ⌣ y +ₕ y ⌣ y) ≡⟨ cong₂ _+ₕ_ (cong₂ _+ₕ_ H²-genₗ⌣H²-genₗ + ((gradedComm-⌣ 2 2 y x) + ∙ cong (-ₕ^ 2 · 2) + (transportRefl (x ⌣ y)) + ∙ -ₕId (x ⌣ y)) + ∙ lUnitₕ 4 (x ⌣ y)) + (cong (x ⌣ y +ₕ_) H²-genᵣ⌣H²-genᵣ ∙ rUnitₕ 4 (x ⌣ y)) ⟩ + ((x ⌣ y) +ₕ (x ⌣ y)) ∎ + + h : (fst qHom) (α' ⌣ α') ≡ (H²-genₗ' +ₕ H²-genᵣ') ⌣ (H²-genₗ' +ₕ H²-genᵣ') + h = (fst qHom) (α' ⌣ α') ≡⟨ refl ⟩ + fst (qHomGen 2) α' ⌣ fst (qHomGen 2) α' ≡⟨ cong (λ x → x ⌣ x) q₂ ⟩ + ((H²-genₗ' +ₕ H²-genᵣ') ⌣ (H²-genₗ' +ₕ H²-genᵣ')) ∎ + + HopfInvariantLem : (HopfInvariant 0 (fold∘W , refl) ≡ 2) + ⊎ (HopfInvariant 0 (fold∘W , refl) ≡ -2) + HopfInvariantLem = + ⊎rec (λ p → inl (p1 + ∙ cong (Iso.fun (fst (Hopfβ-Iso 0 (fold∘W , refl)))) p + ∙ IsGroupHom.pres· (Hopfβ-Iso 0 (fold∘W , refl) .snd) Hopfβfold∘W Hopfβfold∘W + ∙ cong (λ x → x + x) (Hopfβ↦1 0 (fold∘W , refl)))) + (λ p → inr (p1 + ∙ cong (Iso.fun (fst (Hopfβ-Iso 0 (fold∘W , refl)))) p + ∙ IsGroupHom.presinv (Hopfβ-Iso 0 (fold∘W , refl) .snd) (Hopfβfold∘W +ₕ Hopfβfold∘W) + ∙ cong (GroupStr.inv (snd ℤGroup)) + (IsGroupHom.pres· (Hopfβ-Iso 0 (fold∘W , refl) .snd) Hopfβfold∘W Hopfβfold∘W + ∙ cong (λ x → x + x) (Hopfβ↦1 0 (fold∘W , refl))))) + p2 + where + p1 : HopfInvariant 0 (fold∘W , refl) + ≡ Iso.fun (fst (Hopfβ-Iso 0 (fold∘W , refl))) (Hopfαfold∘W ⌣ Hopfαfold∘W) + p1 = cong (Iso.fun (fst (Hopfβ-Iso 0 (fold∘W , refl)))) (transportRefl (Hopfαfold∘W ⌣ Hopfαfold∘W)) + + p2 : (Hopfαfold∘W ⌣ Hopfαfold∘W ≡ Hopfβfold∘W +ₕ Hopfβfold∘W) + ⊎ (Hopfαfold∘W ⌣ Hopfαfold∘W ≡ -ₕ (Hopfβfold∘W +ₕ Hopfβfold∘W)) + p2 = ⌣→≡ (⊎rec (λ p → inl (sym (retEq (fst isEquiv-qHom) (α' ⌣ α')) + ∙∙ cong (invEq (fst isEquiv-qHom)) p + ∙∙ retEq (fst isEquiv-qHom) (β' +ₕ β'))) + (λ p → inr ((sym (retEq (fst isEquiv-qHom) (α' ⌣ α')) + ∙∙ cong (invEq (fst isEquiv-qHom)) p + ∙∙ retEq (fst isEquiv-qHom) (-ₕ (β' +ₕ β'))))) + MainId) + +HopfInvariantLemfold∘W : abs (HopfInvariant 0 (fold∘W , refl)) ≡ 2 +HopfInvariantLemfold∘W = + ⊎→abs (HopfInvariant 0 (fold∘W , refl)) 2 + (postul.HopfInvariantLem + (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1) Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-⌣) + +Whitehead≡ : [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π' ≡ ∣ fold∘W , refl ∣₂ +Whitehead≡ = + cong ∣_∣₂ ([]≡[]₂ (idfun∙ (S₊∙ 2)) (idfun∙ (S₊∙ 2)) ) + ∙ sym Fold∘W + ∙ cong ∣_∣₂ (∘∙-idˡ (fold∘W , refl)) + +HopfInvariantWhitehead : + abs (HopfInvariant-π' 0 [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π') ≡ 2 +HopfInvariantWhitehead = + cong abs (cong (HopfInvariant-π' 0) Whitehead≡) ∙ HopfInvariantLemfold∘W diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda new file mode 100644 index 0000000000..05f722d1e4 --- /dev/null +++ b/Cubical/Homotopy/Loopspace.agda @@ -0,0 +1,449 @@ +{-# OPTIONS --safe #-} + +module Cubical.Homotopy.Loopspace where + +open import Cubical.Core.Everything + +open import Cubical.Data.Nat + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Truncation hiding (elim2) renaming (rec to trRec) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Morphism +open import Cubical.Data.Sigma +open Iso + +{- loop space of a pointed type -} +Ω : {ℓ : Level} → Pointed ℓ → Pointed ℓ +Ω (_ , a) = ((a ≡ a) , refl) + +{- n-fold loop space of a pointed type -} +Ω^_ : ∀ {ℓ} → ℕ → Pointed ℓ → Pointed ℓ +(Ω^ 0) p = p +(Ω^ (suc n)) p = Ω ((Ω^ n) p) + +{- loop space map -} +Ω→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (A →∙ B) → (Ω A →∙ Ω B) +fst (Ω→ {A = A} {B = B} (f , p)) q = sym p ∙∙ cong f q ∙∙ p +snd (Ω→ {A = A} {B = B} (f , p)) = ∙∙lCancel p + +Ω^→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (A →∙ B) → ((Ω^ n) A →∙ (Ω^ n) B) +Ω^→ zero f = f +Ω^→ (suc n) f = Ω→ (Ω^→ n f) + +{- loop space map functoriality (missing pointedness proof) -} +Ω→∘ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (g : B →∙ C) (f : A →∙ B) + → ∀ p → Ω→ (g ∘∙ f) .fst p ≡ (Ω→ g ∘∙ Ω→ f) .fst p +Ω→∘ g f p k i = + hcomp + (λ j → λ + { (i = i0) → compPath-filler' (cong (g .fst) (f .snd)) (g .snd) (~ k) j + ; (i = i1) → compPath-filler' (cong (g .fst) (f .snd)) (g .snd) (~ k) j + }) + (g .fst (doubleCompPath-filler (sym (f .snd)) (cong (f .fst) p) (f .snd) k i)) + +Ω→∘∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (g : B →∙ C) (f : A →∙ B) + → Ω→ (g ∘∙ f) ≡ (Ω→ g ∘∙ Ω→ f) +Ω→∘∙ g f = →∙Homogeneous≡ (isHomogeneousPath _ _) (funExt (Ω→∘ g f)) + +Ω→const : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → Ω→ {A = A} {B = B} ((λ _ → pt B) , refl) ≡ ((λ _ → refl) , refl) +Ω→const = →∙Homogeneous≡ (isHomogeneousPath _ _) (funExt λ _ → sym (rUnit _)) + +{- Ω→ is a homomorphism -} +Ω→pres∙filler : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → (p q : typ (Ω A)) + → I → I → I → fst B +Ω→pres∙filler f p q i j k = + hfill + (λ k → λ + { (i = i0) → doubleCompPath-filler (sym (snd f)) (cong (fst f) (p ∙ q)) (snd f) k j + ; (i = i1) → + (doubleCompPath-filler + (sym (snd f)) (cong (fst f) p) (snd f) k + ∙ doubleCompPath-filler + (sym (snd f)) (cong (fst f) q) (snd f) k) j + ; (j = i0) → snd f k + ; (j = i1) → snd f k}) + (inS (cong-∙ (fst f) p q i j)) + k + +Ω→pres∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → (p q : typ (Ω A)) + → fst (Ω→ f) (p ∙ q) ≡ fst (Ω→ f) p ∙ fst (Ω→ f) q +Ω→pres∙ f p q i j = Ω→pres∙filler f p q i j i1 + +Ω→pres∙reflrefl : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Ω→pres∙ {A = A} {B = B} f refl refl + ≡ cong (fst (Ω→ f)) (sym (rUnit refl)) + ∙ snd (Ω→ f) + ∙ rUnit _ + ∙ cong₂ _∙_ (sym (snd (Ω→ f))) (sym (snd (Ω→ f))) +Ω→pres∙reflrefl {A = A} {B = B} = + →∙J (λ b₀ f → Ω→pres∙ {A = A} {B = (fst B , b₀)} f refl refl + ≡ cong (fst (Ω→ f)) (sym (rUnit refl)) + ∙ snd (Ω→ f) + ∙ rUnit _ + ∙ cong₂ _∙_ (sym (snd (Ω→ f))) (sym (snd (Ω→ f)))) + λ f → lem f + ∙ cong (cong (fst (Ω→ (f , refl))) (sym (rUnit refl)) ∙_) + (((lUnit (cong₂ _∙_ (sym (snd (Ω→ (f , refl)))) + (sym (snd (Ω→ (f , refl)))))) + ∙ cong (_∙ (cong₂ _∙_ (sym (snd (Ω→ (f , refl)))) + (sym (snd (Ω→ (f , refl)))))) + (sym (rCancel (snd (Ω→ (f , refl)))))) + ∙ sym (assoc (snd (Ω→ (f , refl))) + (sym (snd (Ω→ (f , refl)))) + (cong₂ _∙_ (sym (snd (Ω→ (f , refl)))) + (sym (snd (Ω→ (f , refl))))))) + where + lem : (f : fst A → fst B) → Ω→pres∙ (f , refl) (λ _ → snd A) (λ _ → snd A) ≡ + (λ i → fst (Ω→ (f , refl)) (rUnit (λ _ → snd A) (~ i))) ∙ + (λ i → snd (Ω→ (f , refl)) (~ i) ∙ snd (Ω→ (f , refl)) (~ i)) + lem f k i j = + hcomp (λ r → λ { (i = i0) → doubleCompPath-filler + refl (cong f ((λ _ → pt A) ∙ refl)) refl (r ∨ k) j + ; (i = i1) → (∙∙lCancel (λ _ → f (pt A)) (~ r) + ∙ ∙∙lCancel (λ _ → f (pt A)) (~ r)) j + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (k = i0) → Ω→pres∙filler {A = A} {B = fst B , f (pt A)} + (f , refl) refl refl i j r + ; (k = i1) → compPath-filler + ((λ i → fst (Ω→ (f , refl)) + (rUnit (λ _ → snd A) (~ i)))) + ((λ i → snd (Ω→ (f , refl)) (~ i) + ∙ snd (Ω→ (f , refl)) (~ i))) r i j}) + (hcomp (λ r → λ { (i = i0) → doubleCompPath-filler refl (cong f (rUnit (λ _ → pt A) r)) refl k j + ; (i = i1) → rUnit (λ _ → f (pt A)) (r ∨ k) j + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (k = i0) → cong-∙∙-filler f (λ _ → pt A) (λ _ → pt A) (λ _ → pt A) r i j + ; (k = i1) → fst (Ω→ (f , refl)) (rUnit (λ _ → snd A) (~ i ∧ r)) j}) + (rUnit (λ _ → f (pt A)) k j)) + +{- Ω^→ is homomorphism -} +Ω^→pres∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → (n : ℕ) + → (p q : typ ((Ω^ (suc n)) A)) + → fst (Ω^→ (suc n) f) (p ∙ q) + ≡ fst (Ω^→ (suc n) f) p ∙ fst (Ω^→ (suc n) f) q +Ω^→pres∙ {A = A} {B = B} f n p q = Ω→pres∙ (Ω^→ n f) p q + +Ω^→∘∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} (n : ℕ) + (g : B →∙ C) (f : A →∙ B) + → Ω^→ n (g ∘∙ f) ≡ (Ω^→ n g ∘∙ Ω^→ n f) +Ω^→∘∙ zero g f = refl +Ω^→∘∙ (suc n) g f = cong Ω→ (Ω^→∘∙ n g f) ∙ Ω→∘∙ (Ω^→ n g) (Ω^→ n f) + +Ω^→const : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → Ω^→ {A = A} {B = B} n ((λ _ → pt B) , refl) + ≡ ((λ _ → snd ((Ω^ n) B)) , refl) +Ω^→const zero = refl +Ω^→const (suc n) = cong Ω→ (Ω^→const n) ∙ Ω→const + +isEquivΩ→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (f : (A →∙ B)) + → isEquiv (fst f) → isEquiv (Ω→ f .fst) +isEquivΩ→ {B = (B , b)} = + uncurry λ f → + J (λ b y → isEquiv f + → isEquiv (λ q → (λ i → y (~ i)) ∙∙ (λ i → f (q i)) ∙∙ y)) + λ eqf → subst isEquiv (funExt (rUnit ∘ cong f)) + (isoToIsEquiv (congIso (equivToIso (f , eqf)))) + +isEquivΩ^→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (f : A →∙ B) + → isEquiv (fst f) + → isEquiv (Ω^→ n f .fst) +isEquivΩ^→ zero f iseq = iseq +isEquivΩ^→ (suc n) f iseq = isEquivΩ→ (Ω^→ n f) (isEquivΩ^→ n f iseq) + +Ω≃∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (e : A ≃∙ B) + → (Ω A) ≃∙ (Ω B) +fst (fst (Ω≃∙ e)) = fst (Ω→ (fst (fst e) , snd e)) +snd (fst (Ω≃∙ e)) = isEquivΩ→ (fst (fst e) , snd e) (snd (fst e)) +snd (Ω≃∙ e) = snd (Ω→ (fst (fst e) , snd e)) + +Ω≃∙pres∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (e : A ≃∙ B) + → (p q : typ (Ω A)) + → fst (fst (Ω≃∙ e)) (p ∙ q) + ≡ fst (fst (Ω≃∙ e)) p + ∙ fst (fst (Ω≃∙ e)) q +Ω≃∙pres∙ e p q = Ω→pres∙ (fst (fst e) , snd e) p q + +Ω^≃∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (e : A ≃∙ B) + → ((Ω^ n) A) ≃∙ ((Ω^ n) B) +Ω^≃∙ zero e = e +fst (fst (Ω^≃∙ (suc n) e)) = + fst (Ω→ (fst (fst (Ω^≃∙ n e)) , snd (Ω^≃∙ n e))) +snd (fst (Ω^≃∙ (suc n) e)) = + isEquivΩ→ (fst (fst (Ω^≃∙ n e)) , snd (Ω^≃∙ n e)) (snd (fst (Ω^≃∙ n e))) +snd (Ω^≃∙ (suc n) e) = + snd (Ω→ (fst (fst (Ω^≃∙ n e)) , snd (Ω^≃∙ n e))) + +ΩfunExtIso : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → Iso (typ (Ω (A →∙ B ∙))) (A →∙ Ω B) +fst (fun (ΩfunExtIso A B) p) x = funExt⁻ (cong fst p) x +snd (fun (ΩfunExtIso A B) p) i j = snd (p j) i +fst (inv (ΩfunExtIso A B) (f , p) i) x = f x i +snd (inv (ΩfunExtIso A B) (f , p) i) j = p j i +rightInv (ΩfunExtIso A B) _ = refl +leftInv (ΩfunExtIso A B) _ = refl + +{- Commutativity of loop spaces -} +isComm∙ : ∀ {ℓ} (A : Pointed ℓ) → Type ℓ +isComm∙ A = (p q : typ (Ω A)) → p ∙ q ≡ q ∙ p + +private + mainPath : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (α β : typ ((Ω^ (2 + n)) A)) + → (λ i → α i ∙ refl) ∙ (λ i → refl ∙ β i) + ≡ (λ i → refl ∙ β i) ∙ (λ i → α i ∙ refl) + mainPath n α β i = (λ j → α (j ∧ ~ i) ∙ β (j ∧ i)) ∙ λ j → α (~ i ∨ j) ∙ β (i ∨ j) + +EH-filler : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → typ ((Ω^ (2 + n)) A) + → typ ((Ω^ (2 + n)) A) → I → I → I → _ +EH-filler {A = A} n α β i j z = + hfill (λ k → λ { (i = i0) → ((cong (λ x → rUnit x (~ k)) α) + ∙ cong (λ x → lUnit x (~ k)) β) j + ; (i = i1) → ((cong (λ x → lUnit x (~ k)) β) + ∙ cong (λ x → rUnit x (~ k)) α) j + ; (j = i0) → rUnit refl (~ k) + ; (j = i1) → rUnit refl (~ k)}) + (inS (mainPath n α β i j)) z + +{- Eckmann-Hilton -} +EH : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → isComm∙ ((Ω^ (suc n)) A) +EH {A = A} n α β i j = EH-filler n α β i j i1 + +{- Lemmas for the syllepsis : EH α β ≡ (EH β α) ⁻¹ -} + +EH-refl-refl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → EH {A = A} n refl refl ≡ refl +EH-refl-refl {A = A} n k i j = + hcomp (λ r → λ { (k = i1) → (refl ∙ (λ _ → basep)) j + ; (j = i0) → rUnit basep (~ r ∧ ~ k) + ; (j = i1) → rUnit basep (~ r ∧ ~ k) + ; (i = i0) → (refl ∙ (λ _ → lUnit basep (~ r ∧ ~ k))) j + ; (i = i1) → (refl ∙ (λ _ → lUnit basep (~ r ∧ ~ k))) j}) + (((cong (λ x → rUnit x (~ k)) (λ _ → basep)) + ∙ cong (λ x → lUnit x (~ k)) (λ _ → basep)) j) + where + basep = snd (Ω ((Ω^ n) A)) + +{- Generalisations of EH α β when α or β is refl -} +EH-gen-l : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → {x y : typ ((Ω^ (suc n)) A)} (α : x ≡ y) + → α ∙ refl ≡ refl ∙ α +EH-gen-l {ℓ = ℓ} {A = A} n {x = x} {y = y} α i j z = + hcomp (λ k → λ { (i = i0) → ((cong (λ x → rUnit x (~ k)) α) ∙ refl) j z + ; (i = i1) → (refl ∙ cong (λ x → rUnit x (~ k)) α) j z + ; (j = i0) → rUnit (refl {x = x z}) (~ k) z + ; (j = i1) → rUnit (refl {x = y z}) (~ k) z + ; (z = i0) → x i1 + ; (z = i1) → y i1}) + (((λ j → α (j ∧ ~ i) ∙ refl) ∙ λ j → α (~ i ∨ j) ∙ refl) j z) + +EH-gen-r : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → {x y : typ ((Ω^ (suc n)) A)} (β : x ≡ y) + → refl ∙ β ≡ β ∙ refl +EH-gen-r {A = A} n {x = x} {y = y} β i j z = + hcomp (λ k → λ { (i = i0) → (refl ∙ cong (λ x → lUnit x (~ k)) β) j z + ; (i = i1) → ((cong (λ x → lUnit x (~ k)) β) ∙ refl) j z + ; (j = i0) → lUnit (λ k → x (k ∧ z)) (~ k) z + ; (j = i1) → lUnit (λ k → y (k ∧ z)) (~ k) z + ; (z = i0) → x i1 + ; (z = i1) → y i1}) + (((λ j → refl ∙ β (j ∧ i)) ∙ λ j → refl ∙ β (i ∨ j)) j z) + +{- characterisations of EH α β when α or β is refl -} +EH-α-refl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (α : typ ((Ω^ (2 + n)) A)) + → EH n α refl ≡ sym (rUnit α) ∙ lUnit α +EH-α-refl {A = A} n α i j k = + hcomp (λ r → λ { (i = i0) → EH-gen-l n (λ i → α (i ∧ r)) j k + ; (i = i1) → (sym (rUnit λ i → α (i ∧ r)) ∙ lUnit λ i → α (i ∧ r)) j k + ; (j = i0) → ((λ i → α (i ∧ r)) ∙ refl) k + ; (j = i1) → (refl ∙ (λ i → α (i ∧ r))) k + ; (k = i0) → refl + ; (k = i1) → α r}) + ((EH-refl-refl n ∙ sym (lCancel (rUnit refl))) i j k) + +EH-refl-β : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (β : typ ((Ω^ (2 + n)) A)) + → EH n refl β ≡ sym (lUnit β) ∙ rUnit β +EH-refl-β {A = A} n β i j k = + hcomp (λ r → λ { (i = i0) → EH-gen-r n (λ i → β (i ∧ r)) j k + ; (i = i1) → (sym (lUnit λ i → β (i ∧ r)) ∙ rUnit λ i → β (i ∧ r)) j k + ; (j = i0) → (refl ∙ (λ i → β (i ∧ r))) k + ; (j = i1) → ((λ i → β (i ∧ r)) ∙ refl) k + ; (k = i0) → refl + ; (k = i1) → β r}) + ((EH-refl-refl n ∙ sym (lCancel (rUnit refl))) i j k) + +syllepsis : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (α β : typ ((Ω^ 3) A)) + → EH 0 α β ≡ sym (EH 0 β α) +syllepsis {A = A} n α β k i j = + hcomp (λ r → λ { (i = i0) → i=i0 r j k + ; (i = i1) → i=i1 r j k + ; (j = i0) → j-filler r j k + ; (j = i1) → j-filler r j k + ; (k = i0) → EH-filler 1 α β i j r + ; (k = i1) → EH-filler 1 β α (~ i) j r}) + (btm-filler (~ k) i j) + where + guy = snd (Ω (Ω A)) + + btm-filler : I → I → I → typ (Ω (Ω A)) + btm-filler j i k = + hcomp (λ r + → λ {(j = i0) → mainPath 1 β α (~ i) k + ; (j = i1) → mainPath 1 α β i k + ; (i = i0) → (cong (λ x → EH-α-refl 0 x r (~ j)) α + ∙ cong (λ x → EH-refl-β 0 x r (~ j)) β) k + ; (i = i1) → (cong (λ x → EH-refl-β 0 x r (~ j)) β + ∙ cong (λ x → EH-α-refl 0 x r (~ j)) α) k + ; (k = i0) → EH-α-refl 0 guy r (~ j) + ; (k = i1) → EH-α-refl 0 guy r (~ j)}) + (((λ l → EH 0 (α (l ∧ ~ i)) (β (l ∧ i)) (~ j)) + ∙ λ l → EH 0 (α (l ∨ ~ i)) (β (l ∨ i)) (~ j)) k) + + link : I → I → I → _ + link z i j = + hfill (λ k → λ { (i = i1) → refl + ; (j = i0) → rUnit refl (~ i) + ; (j = i1) → lUnit guy (~ i ∧ k)}) + (inS (rUnit refl (~ i ∧ ~ j))) z + + i=i1 : I → I → I → typ (Ω (Ω A)) + i=i1 r j k = + hcomp (λ i → λ { (r = i0) → (cong (λ x → compPath-filler (sym (lUnit x)) (rUnit x) i k) β + ∙ cong (λ x → compPath-filler (sym (rUnit x)) (lUnit x) i k) α) j + ; (r = i1) → (β ∙ α) j + ; (k = i0) → (cong (λ x → lUnit x (~ r)) β ∙ + cong (λ x → rUnit x (~ r)) α) j + ; (k = i1) → (cong (λ x → rUnit x (~ r ∧ i)) β ∙ + cong (λ x → lUnit x (~ r ∧ i)) α) j + ; (j = i0) → link i r k + ; (j = i1) → link i r k}) + (((cong (λ x → lUnit x (~ r ∧ ~ k)) β + ∙ cong (λ x → rUnit x (~ r ∧ ~ k)) α)) j) + + i=i0 : I → I → I → typ (Ω (Ω A)) + i=i0 r j k = + hcomp (λ i → λ { (r = i0) → (cong (λ x → compPath-filler (sym (rUnit x)) (lUnit x) i k) α + ∙ cong (λ x → compPath-filler (sym (lUnit x)) (rUnit x) i k) β) j + ; (r = i1) → (α ∙ β) j + ; (k = i0) → (cong (λ x → rUnit x (~ r)) α ∙ + cong (λ x → lUnit x (~ r)) β) j + ; (k = i1) → (cong (λ x → lUnit x (~ r ∧ i)) α ∙ + cong (λ x → rUnit x (~ r ∧ i)) β) j + ; (j = i0) → link i r k + ; (j = i1) → link i r k}) + ((cong (λ x → rUnit x (~ r ∧ ~ k)) α + ∙ cong (λ x → lUnit x (~ r ∧ ~ k)) β) j) + + j-filler : I → I → I → typ (Ω (Ω A)) + j-filler r i k = + hcomp (λ j → λ { (i = i0) → link j r k + ; (i = i1) → link j r k + ; (r = i0) → compPath-filler (sym (rUnit guy)) + (lUnit guy) j k + ; (r = i1) → refl + ; (k = i0) → rUnit guy (~ r) + ; (k = i1) → rUnit guy (j ∧ ~ r)}) + (rUnit guy (~ r ∧ ~ k)) + +------ Ωⁿ⁺¹ A ≃ Ωⁿ(Ω A) ------ +flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) +flipΩPath {A = A} zero = refl +flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) + +flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) +flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) + +flipΩIso⁻pres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst ((Ω^ (suc n)) (Ω A))) + → inv (flipΩIso (suc n)) (f ∙ g) + ≡ (inv (flipΩIso (suc n)) f) + ∙ (inv (flipΩIso (suc n)) g) +flipΩIso⁻pres· {A = A} n f g i = + transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd + ≡ flipΩPath n (~ i ∧ ~ j) .snd) i + (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f + ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) + +flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst (Ω ((Ω^ (suc n)) A))) + → fun (flipΩIso (suc n)) (f ∙ g) + ≡ (fun (flipΩIso (suc n)) f) + ∙ (fun (flipΩIso (suc n)) g) +flipΩIsopres· n = + morphLemmas.isMorphInv _∙_ _∙_ + (inv (flipΩIso (suc n))) + (flipΩIso⁻pres· n) + (fun (flipΩIso (suc n))) + (Iso.leftInv (flipΩIso (suc n))) + (Iso.rightInv (flipΩIso (suc n))) + +flipΩrefl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → fun (flipΩIso {A = A} (suc n)) refl ≡ refl +flipΩrefl {A = A} n j = + transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j))))) + j (snd (Ω (flipΩPath n j))) + +---- Misc. ---- + +isCommA→isCommTrunc : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → isComm∙ A + → isOfHLevel (suc n) (typ A) + → isComm∙ (∥ typ A ∥ (suc n) , ∣ pt A ∣) +isCommA→isCommTrunc {A = (A , a)} n comm hlev p q = + ((λ i j → (leftInv (truncIdempotentIso (suc n) hlev) ((p ∙ q) j) (~ i))) + ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) + (cong (trRec hlev (λ x → x)) (p ∙ q))) + ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) + (congFunct {A = ∥ A ∥ (suc n)} {B = A} (trRec hlev (λ x → x)) p q i))) + ∙ ((λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) + (comm (cong (trRec hlev (λ x → x)) p) (cong (trRec hlev (λ x → x)) q) i)) + ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) + (congFunct {A = ∥ A ∥ (suc n)} {B = A} (trRec hlev (λ x → x)) q p (~ i))) + ∙∙ (λ i j → (leftInv (truncIdempotentIso (suc n) hlev) ((q ∙ p) j) i))) + +ptdIso→comm : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Type ℓ'} (e : Iso (typ A) B) + → isComm∙ A → isComm∙ (B , fun e (pt A)) +ptdIso→comm {A = (A , a)} {B = B} e comm p q = + sym (rightInv (congIso e) (p ∙ q)) + ∙∙ (cong (fun (congIso e)) ((invCongFunct e p q) + ∙∙ (comm (inv (congIso e) p) (inv (congIso e) q)) + ∙∙ (sym (invCongFunct e q p)))) + ∙∙ rightInv (congIso e) (q ∙ p) + +{- Homotopy group version -} +π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₂ + → ∥ typ ((Ω^ (suc n)) A) ∥₂ → ∥ typ ((Ω^ (suc n)) A) ∥₂ +π-comp n = elim2 (λ _ _ → isSetSetTrunc) λ p q → ∣ p ∙ q ∣₂ + +EH-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : ∥ typ ((Ω^ (2 + n)) A) ∥₂) + → π-comp (1 + n) p q ≡ π-comp (1 + n) q p +EH-π n = elim2 (λ x y → isOfHLevelPath 2 isSetSetTrunc _ _) + λ p q → cong ∣_∣₂ (EH n p q) diff --git a/Cubical/Homotopy/MayerVietorisCofiber.agda b/Cubical/Homotopy/MayerVietorisCofiber.agda new file mode 100644 index 0000000000..894352405b --- /dev/null +++ b/Cubical/Homotopy/MayerVietorisCofiber.agda @@ -0,0 +1,175 @@ +{- + +Mayer-Vietoris cofiber sequence: + + Let X be a pointed type, and let a span B ←[f]- X -[g]→ C be given. + Then the mapping cone of the canonical map (B ⋁ C) → B ⊔_X C is equivalent to Susp X. + +The sequence Susp X → (B ⋁ C) → B ⊔_X C therefore induces a long exact sequence in cohomology. +Proof is adapted from Evan Cavallo's master's thesis. + +-} +{-# OPTIONS --safe #-} +module Cubical.Homotopy.MayerVietorisCofiber where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Pointed +open import Cubical.Data.Unit +open import Cubical.HITs.MappingCones +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge + +module _ {ℓX ℓB ℓC} {X : Pointed ℓX} {B : Type ℓB} {C : Type ℓC} (f : X .fst → B) (g : X .fst → C) + where + + private + Y : Pointed _ + Y = (B , f (X .snd)) + + Z : Pointed _ + Z = (C , g (X .snd)) + + wedgeToPushout : Y ⋁ Z → Pushout f g + wedgeToPushout (inl y) = inl y + wedgeToPushout (inr z) = inr z + wedgeToPushout (push _ i) = push (pt X) i + + pushoutToSusp : Pushout f g → Susp (X .fst) + pushoutToSusp (inl y) = north + pushoutToSusp (inr z) = south + pushoutToSusp (push x i) = merid x i + + {- + Coherence lemma: + To construct a function f : (c : Cone wedgeToPushout) → D c, we can always + adjust the definition of (f (spoke (inr z) i)) so that there is a canonical + choice for (f (spoke (push tt j) i)) + -} + module Helper {ℓD} {D : Cone wedgeToPushout → Type ℓD} + (inj* : ∀ w → D (inj w)) + (hub* : D hub) + (inl* : ∀ y → PathP (λ i → D (spoke (inl y) i)) hub* (inj* (inl y))) + (inr* : ∀ z → PathP (λ i → D (spoke (inr z) i)) hub* (inj* (inr z))) + where + + cap : (i j : I) → D (spoke (push tt j) i) + cap i j = + fill + (λ i → D (spoke (push tt j) (~ i))) + (λ i → λ + { (j = i0) → inl* (Y .snd) (~ i) + ; (j = i1) → inr* (Z .snd) (~ i) + }) + (inS (inj* (push (X .snd) j))) + (~ i) + + cap0 : (j : I) → D hub + cap0 = cap i0 + + face-i≡0 : (k j : I) → D hub + face-i≡0 k j = + hfill + (λ j → λ + { (k = i0) → cap0 j + ; (k = i1) → hub* + }) + (inS hub*) + j + + inrFiller : ∀ z → (k i : I) → D (spoke (inr z) i) + inrFiller z k i = + hfill + (λ k → λ + { (i = i0) → face-i≡0 k i1 + ; (i = i1) → inj* (inr z) + }) + (inS (inr* z i)) + k + + fun : ∀ c → D c + fun (inj w) = inj* w + fun hub = hub* + fun (spoke (inl y) i) = inl* y i + fun (spoke (inr z) i) = inrFiller z i1 i + fun (spoke (push tt j) i) = + hcomp + (λ k → λ + { (i = i0) → face-i≡0 k j + ; (i = i1) → inj* (push (X .snd) j) + ; (j = i0) → inl* (Y .snd) i + }) + (cap i j) + + equiv : Cone wedgeToPushout ≃ Susp (X .fst) + equiv = isoToEquiv (iso fwd bwd fwdBwd bwdFwd) + where + fwd : Cone wedgeToPushout → Susp (X .fst) + fwd (inj w) = pushoutToSusp w + fwd hub = north + fwd (spoke (inl y) i) = north + fwd (spoke (inr z) i) = merid (X .snd) i + fwd (spoke (push tt j) i) = merid (X .snd) (i ∧ j) + + bwd : Susp (X .fst) → Cone wedgeToPushout + bwd north = hub + bwd south = hub + bwd (merid x i) = + hcomp + (λ k → λ + { (i = i0) → spoke (inl (f x)) (~ k) + ; (i = i1) → spoke (inr (g x)) (~ k) + }) + (inj (push x i)) + + bwdPushout : (w : Pushout f g) → bwd (pushoutToSusp w) ≡ inj w + bwdPushout (inl y) = spoke (inl y) + bwdPushout (inr z) = spoke (inr z) + bwdPushout (push x i) k = + hfill + (λ k → λ + { (i = i0) → spoke (inl (f x)) (~ k) + ; (i = i1) → spoke (inr (g x)) (~ k) + }) + (inS (inj (push x i))) + (~ k) + + bwdMeridPt : refl ≡ cong bwd (merid (X .snd)) + bwdMeridPt j i = + hcomp + (λ k → λ + { (i = i0) → spoke (inl (Y .snd)) (~ k) + ; (i = i1) → spoke (inr (Z .snd)) (~ k) + ; (j = i0) → spoke (push _ i) (~ k) + }) + (inj (push (X .snd) i)) + + bwdFwd : (c : Cone wedgeToPushout) → bwd (fwd c) ≡ c + bwdFwd = + Helper.fun + bwdPushout + refl + (λ y i k → spoke (inl y) (i ∧ k)) + (λ z i k → + hcomp + (λ l → λ + { (i = i0) → hub + ; (i = i1) → spoke (inr z) k + ; (k = i0) → bwdMeridPt l i + ; (k = i1) → spoke (inr z) i + }) + (spoke (inr z) (i ∧ k))) + + fwdBwd : (s : Susp (X .fst)) → fwd (bwd s) ≡ s + fwdBwd north = refl + fwdBwd south = merid (X .snd) + fwdBwd (merid a i) j = + fill + (λ _ → Susp (X .fst)) + (λ j → λ + { (i = i0) → north + ; (i = i1) → merid (X .snd) (~ j) + }) + (inS (merid a i)) + (~ j) diff --git a/Cubical/Homotopy/Prespectrum.agda b/Cubical/Homotopy/Prespectrum.agda new file mode 100644 index 0000000000..e6352c9b34 --- /dev/null +++ b/Cubical/Homotopy/Prespectrum.agda @@ -0,0 +1,54 @@ +{-# OPTIONS --safe #-} +{- + This uses ideas from Floris van Doorn's phd thesis and the code in + https://github.com/cmu-phil/Spectral/blob/master/spectrum/basic.hlean +-} +module Cubical.Homotopy.Prespectrum where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Pointed +open import Cubical.Data.Unit.Pointed + +open import Cubical.Structures.Successor + +open import Cubical.Data.Nat +open import Cubical.Data.Int + +open import Cubical.HITs.Susp + +open import Cubical.Homotopy.Loopspace + +private + variable + ℓ ℓ′ : Level + +record GenericPrespectrum (S : SuccStr ℓ) (ℓ′ : Level) : Type (ℓ-max (ℓ-suc ℓ′) ℓ) where + open SuccStr S + field + space : Index → Pointed ℓ′ + map : (i : Index) → (space i →∙ Ω (space (succ i))) + +Prespectrum = GenericPrespectrum ℤ+ + +Unit∙→ΩUnit∙ : {ℓ : Level} → (Unit∙ {ℓ = ℓ}) →∙ Ω (Unit∙ {ℓ = ℓ}) +Unit∙→ΩUnit∙ = (λ {tt* → refl}) , refl + +makeℤPrespectrum : (space : ℕ → Pointed ℓ) + (map : (i : ℕ) → (space i) →∙ Ω (space (suc i))) + → Prespectrum ℓ +GenericPrespectrum.space (makeℤPrespectrum space map) (pos n) = space n +GenericPrespectrum.space (makeℤPrespectrum space map) (negsuc n) = Unit∙ +GenericPrespectrum.map (makeℤPrespectrum space map) (pos n) = map n +GenericPrespectrum.map (makeℤPrespectrum space map) (negsuc zero) = (λ {tt* → refl}) , refl +GenericPrespectrum.map (makeℤPrespectrum space map) (negsuc (suc n)) = Unit∙→ΩUnit∙ + +SuspensionPrespectrum : Pointed ℓ → Prespectrum ℓ +SuspensionPrespectrum A = makeℤPrespectrum space map + where + space : ℕ → Pointed _ + space zero = A + space (suc n) = Susp∙ (typ (space n)) + + map : (n : ℕ) → _ + map n = toSuspPointed (space n) diff --git a/Cubical/Homotopy/Spectrum.agda b/Cubical/Homotopy/Spectrum.agda new file mode 100644 index 0000000000..c2a36d017b --- /dev/null +++ b/Cubical/Homotopy/Spectrum.agda @@ -0,0 +1,22 @@ +{-# OPTIONS --safe #-} +{- + This uses ideas from Floris van Doorn's phd thesis and the code in + https://github.com/cmu-phil/Spectral/blob/master/spectrum/basic.hlean +-} +module Cubical.Homotopy.Spectrum where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Unit.Pointed +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Int + +open import Cubical.Homotopy.Prespectrum + +private + variable + ℓ : Level + +Spectrum : (ℓ : Level) → Type (ℓ-suc ℓ) +Spectrum ℓ = let open GenericPrespectrum + in Σ[ P ∈ Prespectrum ℓ ] ((k : ℤ) → isEquiv (fst (map P k))) diff --git a/Cubical/Homotopy/WedgeConnectivity.agda b/Cubical/Homotopy/WedgeConnectivity.agda new file mode 100644 index 0000000000..2a8f451fb2 --- /dev/null +++ b/Cubical/Homotopy/WedgeConnectivity.agda @@ -0,0 +1,66 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.WedgeConnectivity where + +open import Cubical.Foundations.Everything +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.HITs.Nullification +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation as Trunc +open import Cubical.Homotopy.Connected + +module WedgeConnectivity {ℓ ℓ' ℓ''} (n m : ℕ) + (A : Pointed ℓ) (connA : isConnected (suc n) (typ A)) + (B : Pointed ℓ') (connB : isConnected (suc m) (typ B)) + (P : typ A → typ B → TypeOfHLevel ℓ'' (n + m)) + (f : (a : typ A) → P a (pt B) .fst) + (g : (b : typ B) → P (pt A) b .fst) + (p : f (pt A) ≡ g (pt B)) + where + + private + Q : typ A → TypeOfHLevel _ n + Q a = + ( (Σ[ k ∈ ((b : typ B) → P a b .fst) ] k (pt B) ≡ f a) + , isOfHLevelRetract n + (λ {(h , q) → h , funExt λ _ → q}) + (λ {(h , q) → h , funExt⁻ q _}) + (λ _ → refl) + (isOfHLevelPrecomposeConnected n m (P a) (λ _ → pt B) + (isConnectedPoint m connB (pt B)) (λ _ → f a)) + ) + + main : isContr (fiber (λ s _ → s (pt A)) (λ _ → g , p ⁻¹)) + main = + elim.isEquivPrecompose (λ _ → pt A) n Q + (isConnectedPoint n connA (pt A)) + .equiv-proof (λ _ → g , p ⁻¹) + + + extension : ∀ a b → P a b .fst + extension a b = main .fst .fst a .fst b + + left : ∀ a → extension a (pt B) ≡ f a + left a = main .fst .fst a .snd + + right : ∀ b → extension (pt A) b ≡ g b + right = funExt⁻ (cong fst (funExt⁻ (main .fst .snd) _)) + + hom : left (pt A) ⁻¹ ∙ right (pt B) ≡ p + hom i j = hcomp (λ k → λ { (i = i1) → p j + ; (j = i0) → (cong snd (funExt⁻ (main .fst .snd) tt)) i (~ j) + ; (j = i1) → right (pt B) (i ∨ k)}) + (cong snd (funExt⁻ (main .fst .snd) tt) i (~ j)) + + hom' : left (pt A) ≡ right (pt B) ∙ sym p + hom' = (lUnit (left _) ∙ cong (_∙ left (pt A)) (sym (rCancel (right (pt B))))) + ∙∙ sym (assoc _ _ _) + ∙∙ cong (right (pt B) ∙_) (sym (symDistr (left (pt A) ⁻¹) (right (pt B))) ∙ (cong sym hom)) + + homSquare : PathP (λ i → extension (pt A) (pt B) ≡ p i) (left (pt A)) (right (pt B)) + homSquare i j = hcomp (λ k → λ { (i = i0) → left (pt A) j + ; (i = i1) → compPath-filler (right (pt B)) (sym p) (~ k) j + ; (j = i0) → extension (pt A) (pt B) + ; (j = i1) → p (i ∧ k) }) + (hom' i j) diff --git a/Cubical/Homotopy/Whitehead.agda b/Cubical/Homotopy/Whitehead.agda new file mode 100644 index 0000000000..af7df6635a --- /dev/null +++ b/Cubical/Homotopy/Whitehead.agda @@ -0,0 +1,407 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Whitehead where + +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Join +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation + +open import Cubical.Homotopy.Group.Base + +open Iso +open 3x3-span + +joinTo⋁ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → join (typ A) (typ B) + → (Susp (typ A) , north) ⋁ (Susp (typ B) , north) +joinTo⋁ (inl x) = inr north +joinTo⋁ (inr x) = inl north +joinTo⋁ {A = A} {B = B} (push a b i) = + ((λ i → inr (σ B b i)) + ∙∙ sym (push tt) + ∙∙ λ i → inl (σ A a i)) i + +-- Whitehead product (main definition) +[_∣_] : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (S₊∙ (suc n) →∙ X) + → (S₊∙ (suc m) →∙ X) + → S₊∙ (suc (n + m)) →∙ X +fst ([_∣_] {X = X} {n = n} {m = m} f g) x = + _∨→_ (f ∘∙ (inv (IsoSucSphereSusp n) , IsoSucSphereSusp∙ n)) + (g ∘∙ (inv (IsoSucSphereSusp m) , IsoSucSphereSusp∙ m)) + (joinTo⋁ {A = S₊∙ n} {B = S₊∙ m} + (inv (IsoSphereJoin n m) x)) +snd ([_∣_] {n = n} {m = m} f g) = + cong (_∨→_ (f ∘∙ (inv (IsoSucSphereSusp n) , IsoSucSphereSusp∙ n)) + (g ∘∙ (inv (IsoSucSphereSusp m) , IsoSucSphereSusp∙ m))) + (cong (joinTo⋁ {A = S₊∙ n} {B = S₊∙ m}) (IsoSphereJoin⁻Pres∙ n m)) + ∙ cong (fst g) (IsoSucSphereSusp∙ m) + ∙ snd g + +-- For Sⁿ, Sᵐ with n, m ≥ 2, we can avoid some bureaucracy. We make +-- a separate definition and prove it equivalent. +[_∣_]-pre : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (S₊∙ (suc (suc n)) →∙ X) + → (S₊∙ (suc (suc m)) →∙ X) + → join (typ (S₊∙ (suc n))) (typ (S₊∙ (suc m))) → fst X +[_∣_]-pre {n = n} {m = m} f g x = + _∨→_ f g + (joinTo⋁ {A = S₊∙ (suc n)} {B = S₊∙ (suc m)} + x) + +[_∣_]₂ : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (S₊∙ (suc (suc n)) →∙ X) + → (S₊∙ (suc (suc m)) →∙ X) + → S₊∙ (suc ((suc n) + (suc m))) →∙ X +fst ([_∣_]₂ {n = n} {m = m} f g) x = + [ f ∣ g ]-pre (inv (IsoSphereJoin (suc n) (suc m)) x) +snd ([_∣_]₂ {n = n} {m = m} f g) = + cong ([ f ∣ g ]-pre) (IsoSphereJoin⁻Pres∙ (suc n) (suc m)) + ∙ snd g + +[]≡[]₂ : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (f : (S₊∙ (suc (suc n)) →∙ X)) + → (g : (S₊∙ (suc (suc m)) →∙ X)) + → [ f ∣ g ] ≡ [ f ∣ g ]₂ +[]≡[]₂ {n = n} {m = m} f g = + ΣPathP ( + (λ i x → _∨→_ (∘∙-idˡ f i) + (∘∙-idˡ g i) + (joinTo⋁ {A = S₊∙ (suc n)} {B = S₊∙ (suc m)} + (inv (IsoSphereJoin (suc n) (suc m)) x))) + , (cong (cong (_∨→_ (f ∘∙ idfun∙ _) + (g ∘∙ idfun∙ _)) + (cong (joinTo⋁ {A = S₊∙ (suc n)} {B = S₊∙ (suc m)}) + (IsoSphereJoin⁻Pres∙ (suc n) (suc m))) ∙_) + (sym (lUnit (snd g))) + ◁ λ j → (λ i → _∨→_ (∘∙-idˡ f j) + (∘∙-idˡ g j) + ( joinTo⋁ {A = S₊∙ (suc n)} {B = S₊∙ (suc m)} + ((IsoSphereJoin⁻Pres∙ (suc n) (suc m)) i))) ∙ snd g)) + +-- Homotopy group version +[_∣_]π' : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → π' (suc n) X → π' (suc m) X + → π' (suc (n + m)) X +[_∣_]π' = elim2 (λ _ _ → squash₂) λ f g → ∣ [ f ∣ g ] ∣₂ + +-- We prove that the function joinTo⋁ used in the definition of the whitehead +-- product gives an equivalence between (Susp A × Susp B) and the +-- appropriate cofibre of joinTo⋁ (last two theorems in the following +-- module). + +module _ (A B : Type) (a₀ : A) (b₀ : B) where + private + W = joinTo⋁ {A = (A , a₀)} {B = (B , b₀)} + + A∨B = (Susp A , north) ⋁ (Susp B , north) + + σB = σ (B , b₀) + σA = σ (A , a₀) + + cofibW = Pushout W λ _ → tt + + whitehead3x3 : 3x3-span + A00 whitehead3x3 = Susp A + A02 whitehead3x3 = B + A04 whitehead3x3 = Unit + A20 whitehead3x3 = B + A22 whitehead3x3 = A × B + A24 whitehead3x3 = A + A40 whitehead3x3 = B + A42 whitehead3x3 = B + A44 whitehead3x3 = Unit + f10 whitehead3x3 _ = south + f12 whitehead3x3 = snd + f14 whitehead3x3 _ = tt + f30 whitehead3x3 = idfun B + f32 whitehead3x3 = snd + f34 whitehead3x3 _ = tt + f01 whitehead3x3 _ = north + f21 whitehead3x3 = snd + f41 whitehead3x3 = idfun B + f03 whitehead3x3 _ = tt + f23 whitehead3x3 = fst + f43 whitehead3x3 _ = tt + H11 whitehead3x3 x = merid (fst x) + H13 whitehead3x3 _ = refl + H31 whitehead3x3 _ = refl + H33 whitehead3x3 _ = refl + + A0□→A∨B : A0□ whitehead3x3 → A∨B + A0□→A∨B (inl x) = inl x + A0□→A∨B (inr x) = inr north + A0□→A∨B (push a i) = (push tt ∙ λ i → inr (σB a (~ i))) i + + A∨B→A0□ : A∨B → A0□ whitehead3x3 + A∨B→A0□ (inl x) = inl x + A∨B→A0□ (inr north) = inl north + A∨B→A0□ (inr south) = inl north + A∨B→A0□ (inr (merid b i)) = (push b₀ ∙ sym (push b)) i + A∨B→A0□ (push a i) = inl north + + Iso-A0□-⋁ : Iso (A0□ whitehead3x3) A∨B + fun Iso-A0□-⋁ = A0□→A∨B + inv Iso-A0□-⋁ = A∨B→A0□ + rightInv Iso-A0□-⋁ (inl x) = refl + rightInv Iso-A0□-⋁ (inr north) = push tt + rightInv Iso-A0□-⋁ (inr south) = push tt ∙ λ i → inr (merid b₀ i) + rightInv Iso-A0□-⋁ (inr (merid a i)) j = lem j i + where + lem : PathP (λ i → push tt i ≡ (push tt ∙ (λ i → inr (merid b₀ i))) i) + (cong A0□→A∨B (cong A∨B→A0□ λ i → inr (merid a i))) + (λ i → inr (merid a i)) + lem = (cong-∙ A0□→A∨B (push b₀) (sym (push a)) + ∙ cong₂ _∙_ (cong (push tt ∙_) + (λ j i → inr (rCancel (merid b₀) j (~ i))) + ∙ sym (rUnit (push tt))) + (symDistr (push tt) (λ i → inr (σB a (~ i))))) + ◁ λ i j → hcomp (λ k → + λ { (i = i0) → compPath-filler' (push tt) + (compPath-filler (λ i → inr (σB a i)) + (sym (push tt)) k) k j + ; (i = i1) → inr (merid a j) + ; (j = i0) → push tt (i ∨ ~ k) + ; (j = i1) → compPath-filler' (push tt) + (λ i → inr (merid b₀ i)) k i}) + (inr (compPath-filler (merid a) + (sym (merid b₀)) (~ i) j)) + + rightInv Iso-A0□-⋁ (push a i) j = push tt (i ∧ j) + leftInv Iso-A0□-⋁ (inl x) = refl + leftInv Iso-A0□-⋁ (inr tt) = push b₀ + leftInv Iso-A0□-⋁ (push b i) j = help j i + where + help : PathP (λ i → inl north ≡ push b₀ i) + (cong A∨B→A0□ (cong A0□→A∨B (push b))) + (push b) + help = (cong-∙ A∨B→A0□ (push tt) (λ i → inr (σB b (~ i))) + ∙ (λ i → lUnit (sym (cong-∙ (A∨B→A0□ ∘ inr) + (merid b) (sym (merid b₀)) i)) (~ i)) + ∙ cong sym (cong ((push b₀ ∙ sym (push b)) ∙_) + (cong sym (rCancel (push b₀)))) + ∙ cong sym (sym (rUnit (push b₀ ∙ sym (push b))))) + ◁ λ i j → compPath-filler' (push b₀) (sym (push b)) (~ i) (~ j) + + Iso-A2□-join : Iso (A2□ whitehead3x3) (join A B) + fun Iso-A2□-join (inl x) = inr x + fun Iso-A2□-join (inr x) = inl x + fun Iso-A2□-join (push (a , b) i) = push a b (~ i) + inv Iso-A2□-join (inl x) = inr x + inv Iso-A2□-join (inr x) = inl x + inv Iso-A2□-join (push a b i) = push (a , b) (~ i) + rightInv Iso-A2□-join (inl x) = refl + rightInv Iso-A2□-join (inr x) = refl + rightInv Iso-A2□-join (push a b i) = refl + leftInv Iso-A2□-join (inl x) = refl + leftInv Iso-A2□-join (inr x) = refl + leftInv Iso-A2□-join (push a i) = refl + + isContrA4□ : isContr (A4□ whitehead3x3) + fst isContrA4□ = inr tt + snd isContrA4□ (inl x) = sym (push x) + snd isContrA4□ (inr x) = refl + snd isContrA4□ (push a i) j = push a (i ∨ ~ j) + + A4□≃Unit : A4□ whitehead3x3 ≃ Unit + A4□≃Unit = isContr→≃Unit isContrA4□ + + Iso-A□0-Susp : Iso (A□0 whitehead3x3) (Susp A) + fun Iso-A□0-Susp (inl x) = x + fun Iso-A□0-Susp (inr x) = north + fun Iso-A□0-Susp (push a i) = merid a₀ (~ i) + inv Iso-A□0-Susp x = inl x + rightInv Iso-A□0-Susp x = refl + leftInv Iso-A□0-Susp (inl x) = refl + leftInv Iso-A□0-Susp (inr x) = (λ i → inl (merid a₀ i)) ∙ push x + leftInv Iso-A□0-Susp (push a i) j = + hcomp (λ k → λ { (i = i0) → inl (merid a₀ (k ∨ j)) + ; (i = i1) → compPath-filler + (λ i₁ → inl (merid a₀ i₁)) + (push (idfun B a)) k j + ; (j = i0) → inl (merid a₀ (~ i ∧ k)) + ; (j = i1) → push a (i ∧ k)}) + (inl (merid a₀ j)) + + Iso-A□2-Susp× : Iso (A□2 whitehead3x3) (Susp A × B) + fun Iso-A□2-Susp× (inl x) = north , x + fun Iso-A□2-Susp× (inr x) = south , x + fun Iso-A□2-Susp× (push a i) = merid (fst a) i , (snd a) + inv Iso-A□2-Susp× (north , y) = inl y + inv Iso-A□2-Susp× (south , y) = inr y + inv Iso-A□2-Susp× (merid a i , y) = push (a , y) i + rightInv Iso-A□2-Susp× (north , snd₁) = refl + rightInv Iso-A□2-Susp× (south , snd₁) = refl + rightInv Iso-A□2-Susp× (merid a i , snd₁) = refl + leftInv Iso-A□2-Susp× (inl x) = refl + leftInv Iso-A□2-Susp× (inr x) = refl + leftInv Iso-A□2-Susp× (push a i) = refl + + Iso-A□4-Susp : Iso (A□4 whitehead3x3) (Susp A) + fun Iso-A□4-Susp (inl x) = north + fun Iso-A□4-Susp (inr x) = south + fun Iso-A□4-Susp (push a i) = merid a i + inv Iso-A□4-Susp north = inl tt + inv Iso-A□4-Susp south = inr tt + inv Iso-A□4-Susp (merid a i) = push a i + rightInv Iso-A□4-Susp north = refl + rightInv Iso-A□4-Susp south = refl + rightInv Iso-A□4-Susp (merid a i) = refl + leftInv Iso-A□4-Susp (inl x) = refl + leftInv Iso-A□4-Susp (inr x) = refl + leftInv Iso-A□4-Susp (push a i) = refl + + Iso-PushSusp×-Susp×Susp : + Iso (Pushout {A = Susp A × B} fst fst) (Susp A × Susp B) + Iso-PushSusp×-Susp×Susp = theIso + where + F : Pushout {A = Susp A × B} fst fst + → Susp A × Susp B + F (inl x) = x , north + F (inr x) = x , north + F (push (x , b) i) = x , σB b i + + G : Susp A × Susp B → Pushout {A = Susp A × B} fst fst + G (a , north) = inl a + G (a , south) = inr a + G (a , merid b i) = push (a , b) i + + retr : retract F G + retr (inl x) = refl + retr (inr x) = push (x , b₀) + retr (push (a , b) i) j = help j i + where + help : PathP (λ i → Path (Pushout fst fst) (inl a) (push (a , b₀) i)) + (cong G (λ i → a , σB b i)) + (push (a , b)) + help = cong-∙ (λ x → G (a , x)) (merid b) (sym (merid b₀)) + ◁ λ i j → compPath-filler + (push (a , b)) + (sym (push (a , b₀))) + (~ i) j + + theIso : Iso (Pushout fst fst) (Susp A × Susp B) + fun theIso = F + inv theIso = G + rightInv theIso (a , north) = refl + rightInv theIso (a , south) = ΣPathP (refl , merid b₀) + rightInv theIso (a , merid b i) j = + a , compPath-filler (merid b) (sym (merid b₀)) (~ j) i + leftInv theIso = retr + + Iso-A□○-PushSusp× : + Iso (A□○ whitehead3x3) (Pushout {A = Susp A × B} fst fst) + Iso-A□○-PushSusp× = + pushoutIso _ _ fst fst + (isoToEquiv Iso-A□2-Susp×) + (isoToEquiv Iso-A□0-Susp) + (isoToEquiv Iso-A□4-Susp) + (funExt (λ { (inl x) → refl + ; (inr x) → merid a₀ + ; (push a i) j → help₁ a j i})) + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j + → fun Iso-A□4-Susp (rUnit (push (fst a)) (~ j) i)}) + where + help₁ : (a : A × B) + → PathP (λ i → north ≡ merid a₀ i) + (cong (fun Iso-A□0-Susp) + (cong (f□1 whitehead3x3) (push a))) + (merid (fst a)) + help₁ a = + (cong-∙∙ (fun Iso-A□0-Susp) + (λ i → inl (merid (fst a) i)) + (push (snd a)) + refl) + ◁ (λ i j → hcomp (λ k → λ {(i = i1) → merid (fst a) (j ∨ ~ k) + ; (j = i0) → merid (fst a) (~ k) + ; (j = i1) → merid a₀ i}) + (merid a₀ (i ∨ ~ j))) + + Iso-A□○-Susp×Susp : Iso (A□○ whitehead3x3) (Susp A × Susp B) + Iso-A□○-Susp×Susp = compIso Iso-A□○-PushSusp× Iso-PushSusp×-Susp×Susp + + Iso-A○□-cofibW : Iso (A○□ whitehead3x3) cofibW + Iso-A○□-cofibW = + pushoutIso _ _ + W (λ _ → tt) + (isoToEquiv Iso-A2□-join) (isoToEquiv Iso-A0□-⋁) + A4□≃Unit + (funExt lem) + λ _ _ → tt + where + lem : (x : A2□ whitehead3x3) + → A0□→A∨B (f1□ whitehead3x3 x) ≡ W (fun Iso-A2□-join x) + lem (inl x) = (λ i → inl (merid a₀ (~ i))) + lem (inr x) = refl + lem (push (a , b) i) j = help j i + where + help : PathP (λ i → Path (Pushout (λ _ → north) (λ _ → north)) + ((inl (merid a₀ (~ i)))) + (inr north)) + (cong A0□→A∨B (cong (f1□ whitehead3x3) (push (a , b)))) + (cong W (cong (fun Iso-A2□-join) (push (a , b)))) + help = (cong-∙∙ A0□→A∨B (λ i → inl (merid a (~ i))) (push b) refl + ∙ λ j → (λ i₂ → inl (merid a (~ i₂))) + ∙∙ compPath-filler (push tt) (λ i → inr (σB b (~ i))) (~ j) + ∙∙ λ i → inr (σB b (~ i ∧ j))) + ◁ (λ j → (λ i → inl (sym (compPath-filler + (merid a) (sym (merid a₀)) j) i)) + ∙∙ push tt + ∙∙ λ i → inr (σB b (~ i))) + + Iso₁-Susp×Susp-cofibW : Iso (Susp A × Susp B) cofibW + Iso₁-Susp×Susp-cofibW = + compIso (invIso Iso-A□○-Susp×Susp) + (compIso (3x3-Iso whitehead3x3) Iso-A○□-cofibW) + + -- Main iso + Iso-Susp×Susp-cofibJoinTo⋁ : Iso (Susp A × Susp B) cofibW + Iso-Susp×Susp-cofibJoinTo⋁ = + compIso (Σ-cong-iso-snd (λ _ → invSuspIso)) + Iso₁-Susp×Susp-cofibW + + -- The induced function A ∨ B → Susp A × Susp B satisfies + -- the following identity + Susp×Susp→cofibW≡ : Path (A∨B → Susp A × Susp B) + (Iso.inv Iso-Susp×Susp-cofibJoinTo⋁ ∘ inl) + ⋁↪ + Susp×Susp→cofibW≡ = + funExt λ { (inl x) → ΣPathP (refl , (sym (merid b₀))) + ; (inr north) → ΣPathP (refl , (sym (merid b₀))) + ; (inr south) → refl + ; (inr (merid a i)) j → lem₂ a j i + ; (push a i) j → north , (merid b₀ (~ j))} + where + f₁ = fun Iso-PushSusp×-Susp×Susp + f₂ = fun Iso-A□○-PushSusp× + f₃ = backward-l whitehead3x3 + f₄ = fun (Σ-cong-iso-snd (λ _ → invSuspIso)) + + lem : (b : B) + → cong (f₁ ∘ f₂ ∘ f₃) (push b) + ≡ (λ i → north , σB b i) + lem b = cong (cong f₁) (sym (rUnit (push (north , b)))) + + lem₂ : (a : B) + → PathP (λ i → (north , merid b₀ (~ i)) ≡ (north , south)) + (cong (f₄ ∘ f₁ ∘ f₂ ∘ f₃ ∘ A∨B→A0□ ∘ inr) (merid a)) + λ i → north , merid a i + lem₂ a = + cong (cong f₄) (cong-∙ (f₁ ∘ f₂ ∘ f₃) (push b₀) (sym (push a)) + ∙∙ cong₂ _∙_ (lem b₀ ∙ (λ j i → north , rCancel (merid b₀) j i)) + (cong sym (lem a)) + ∙∙ sym (lUnit (λ i₁ → north , σB a (~ i₁)))) + ∙ (λ i j → north , cong-∙ invSusp (merid a) (sym (merid b₀)) i (~ j) ) + ◁ λ i j → north , compPath-filler (sym (merid a)) (merid b₀) (~ i) (~ j) diff --git a/Cubical/Induction/WellFounded.agda b/Cubical/Induction/WellFounded.agda new file mode 100644 index 0000000000..561cb3ca90 --- /dev/null +++ b/Cubical/Induction/WellFounded.agda @@ -0,0 +1,47 @@ +{-# OPTIONS --safe #-} + +module Cubical.Induction.WellFounded where + +open import Cubical.Foundations.Everything + +Rel : ∀{ℓ} → Type ℓ → ∀ ℓ' → Type _ +Rel A ℓ = A → A → Type ℓ + +module _ {ℓ ℓ'} {A : Type ℓ} (_<_ : A → A → Type ℓ') where + WFRec : ∀{ℓ''} → (A → Type ℓ'') → A → Type _ + WFRec P x = ∀ y → y < x → P y + + data Acc (x : A) : Type (ℓ-max ℓ ℓ') where + acc : WFRec Acc x → Acc x + + WellFounded : Type _ + WellFounded = ∀ x → Acc x + + +module _ {ℓ ℓ'} {A : Type ℓ} {_<_ : A → A → Type ℓ'} where + isPropAcc : ∀ x → isProp (Acc _<_ x) + isPropAcc x (acc p) (acc q) + = λ i → acc (λ y y>=_ ; return + ; fib ; fibTail) public +-- Computation +_ : fib 20 ≡ (6765 , PropositionalTruncation.∣ 21890 ∣) +_ = refl + +_ : fibTail 20 ≡ (6765 , PropositionalTruncation.∣ 19 ∣) +_ = refl + +-- Set Quotients +open SetQuotients using (_/_ ; setQuotUniversal) public +-- Rational Numbers +open SetQuoQ using (_∼_ ; ℚ) public +open SigmaQ renaming (ℚ to ℚ') public + + + +------------------------------------------------------------------------- +-- 3. The Structure Identity Principle +-- 3.1 Structures +-- 3.2 Building Strucutres +------------------------------------------------------------------------- + +-- 3.1 Structures +open SIP using (TypeWithStr ; StrEquiv ; _≃[_]_ + ; UnivalentStr ; SIP ; sip) public + +-- the last two terms above correspond to lemma 3.3 +-- and corollary 3.4 respectively +open Axioms using ( AxiomsStructure ; AxiomsEquivStr + ; axiomsUnivalentStr ; transferAxioms) public + +-- Monoids are defined using records and Σ-types in the library + +RawMonoidStructure : Type → Type +RawMonoidStructure X = X × (X → X → X) + +MonoidAxioms : (M : Type) → RawMonoidStructure M → Type +MonoidAxioms M (e , _·_) = Semigroup.IsSemigroup _·_ + × ((x : M) → (x · e ≡ x) × (e · x ≡ x)) + +MonoidStructure : Type → Type +MonoidStructure = AxiomsStructure RawMonoidStructure MonoidAxioms + +Monoid : Type₁ +Monoid = TypeWithStr ℓ-zero MonoidStructure + +MonoidEquiv : (M N : Monoid) → fst M ≃ fst N → Type +MonoidEquiv (_ , (εᴹ , _·_) , _) (_ , (εᴺ , _∗_) , _) (φ , _) = + (φ εᴹ ≡ εᴺ) × (∀ x y → φ (x · y) ≡ (φ x) ∗ (φ y)) + + +-- 3.2 Building Structures +-- Constant and Pointed Structures +open PointedStr using (PointedStructure ; PointedEquivStr + ; pointedUnivalentStr) public +open ConstantStr using (ConstantStructure ; ConstantEquivStr + ; constantUnivalentStr) public + +-- Product Structures +open ProductStr using (ProductStructure ; ProductEquivStr + ; productUnivalentStr) public + +-- Function Structures +open FunctionStr using (FunctionEquivStr) public + +-- Maybe Structures +open MaybeStr using (MaybeEquivStr) public + +-- Transport Structures +open Structure using (EquivAction) public +open SIP using (TransportStr ; TransportStr→UnivalentStr + ; UnivalentStr→TransportStr) public +open Structure using (EquivAction→StrEquiv) public +open FunctionStr using (FunctionEquivStr+) public + +-- Monoids Revisited + +RawMonoid : Type₁ +RawMonoid = TypeWithStr _ RawMonoidStructure + +Monoid→RawMonoid : Monoid → RawMonoid +Monoid→RawMonoid (A , r , _) = (A , r) + +RawMonoidEquivStr = Auto.AutoEquivStr RawMonoidStructure + +rawMonoidUnivalentStr : UnivalentStr _ RawMonoidEquivStr +rawMonoidUnivalentStr = Auto.autoUnivalentStr RawMonoidStructure + +isPropMonoidAxioms : (M : Type) (s : RawMonoidStructure M) → isProp (MonoidAxioms M s) +isPropMonoidAxioms M (e , _·_) = + HLevels.isPropΣ + (Semigroup.isPropIsSemigroup _·_) + (λ α → isPropΠ λ _ → + HLevels.isProp× + (Semigroup.IsSemigroup.is-set α _ _) + (Semigroup.IsSemigroup.is-set α _ _)) + +MonoidEquivStr : StrEquiv MonoidStructure ℓ-zero +MonoidEquivStr = AxiomsEquivStr RawMonoidEquivStr MonoidAxioms + +monoidUnivalentStr : UnivalentStr MonoidStructure MonoidEquivStr +monoidUnivalentStr = axiomsUnivalentStr _ isPropMonoidAxioms rawMonoidUnivalentStr + +MonoidΣPath : (M N : Monoid) → (M ≃[ MonoidEquivStr ] N) ≃ (M ≡ N) +MonoidΣPath = SIP monoidUnivalentStr + +InducedMonoid : (M : Monoid) (N : RawMonoid) (e : M .fst ≃ N .fst) + → RawMonoidEquivStr (Monoid→RawMonoid M) N e → Monoid +InducedMonoid M N e r = + Axioms.inducedStructure rawMonoidUnivalentStr M N (e , r) + +InducedMonoidPath : (M : Monoid) (N : RawMonoid) (e : M .fst ≃ N .fst) + (E : RawMonoidEquivStr (Monoid→RawMonoid M) N e) + → M ≡ InducedMonoid M N e E +InducedMonoidPath M N e E = + MonoidΣPath M (InducedMonoid M N e E) .fst (e , E) + +-- Automation +open Auto using (Transp[_] ; AutoEquivStr ; autoUnivalentStr) public +module _ (A : Type) (Aset : isSet A) where + RawQueueEquivStr = + AutoEquivStr (λ (X : Type) → X × (A → X → X) × (X → Transp[ Maybe (X × A) ])) + + + +------------------------------------------------------------------------- +-- 4. Representation Independence through the SIP +-- 4.1 Matrices +-- 4.2 Queues +------------------------------------------------------------------------- + + +-- 4.1 Matrices +open Vector using (Vec) public +open Finite using (Fin ; _==_) public +open Matrices using (VecMatrix ; FinMatrix ; FinMatrix≡VecMatrix + ; FinMatrix≃VecMatrix) public +open Matrices.FinMatrixAbGroup using (addFinMatrix ; addFinMatrixComm) public + +-- example (not in the library) +open import Cubical.Data.Int hiding (-_) + +ℤ-AbGroup : AbGroup ℓ-zero +ℤ-AbGroup = makeAbGroup {G = ℤ} 0 _+_ -_ isSetℤ +Assoc (λ x _ → x) rem +Comm + where + -_ : ℤ → ℤ + - x = 0 - x + rem : (x : ℤ) → x + (- x) ≡ 0 + rem x = +Comm x (pos 0 - x) Prelude.∙ minusPlus x 0 + +module experiment where + open Prelude + + M : FinMatrix ℤ 2 2 + M i j = if i == j then 1 else 0 + + N : FinMatrix ℤ 2 2 + N i j = if i == j then 0 else 1 + + replaceGoal : {A B : Type} {x y : A} → (e : A ≃ B) + (h : invEq e (equivFun e x) ≡ invEq e (equivFun e y)) → x ≡ y + replaceGoal e h = sym (retEq e _) ∙∙ h ∙∙ retEq e _ + + _ : addFinMatrix ℤ-AbGroup M N ≡ (λ _ _ → 1) + _ = replaceGoal (FinMatrix≃VecMatrix) refl + + +-- 4.2 Queues +open Queues.Queues-on using (RawQueueStructure ; QueueAxioms) public +open BatchedQueues.Truncated2List renaming (Q to BatchedQueueHIT) + using (Raw-1≡2 ; WithLaws) public + + + +------------------------------------------------------------------------- +-- 5. Structured Equivalences from Structured Relations +-- 5.1 Quasi-Equivalence Relations +-- 5.2 Structured Relations +------------------------------------------------------------------------- + + +-- 5.1 Quasi-Equivalence Relations +--Lemma (5.1) +open BinRel using (idPropRel ; invPropRel + ; compPropRel ; graphRel) public +-- Definitions (5.2) and (5.3) +open QER using (isZigZagComplete ; isQuasiEquivRel) public +-- Lemma (5.4) +open QER.QER→Equiv using (Thm ; bwd≡ToRel) public +-- Multisets +open MultiSets renaming (AList to AssocList) public +open MultiSets.Lists&ALists using (addIfEq ; R ; φ ; ψ + ; List/Rᴸ≃AList/Rᴬᴸ) public +open MultiSets.Lists&ALists.L using (insert ; union ; count) +open MultiSets.Lists&ALists.AL using (insert ; union ; count) + + +-- 5.2 Structured Relations +open RelStructure using (StrRel) public +-- Definition (5.6) +open RelStructure using (SuitableStrRel) public +-- Theorem (5.7) +open RelStructure using (structuredQER→structuredEquiv) public +-- Definition (5.9) +open RelStructure using (StrRelAction) public +-- Lemma (5.10) +open RelStructure using (strRelQuotientComparison) public +-- Definition (5.11) +open RelStructure using (PositiveStrRel) public +-- Theorem (5.12) +open RelFunction using (functionSuitableRel) public +-- Multisets +-- (main is applying 5.7 to the example) +open MultiSets.Lists&ALists using (multisetShape ; isStructuredR ; main ; List/Rᴸ≡AList/Rᴬᴸ) + renaming ( hasAssociativeUnion to unionAssocAxiom + ; LQassoc to LUnionAssoc + ; ALQassoc to ALUnionAssoc) public diff --git a/Cubical/Papers/Synthetic.agda b/Cubical/Papers/Synthetic.agda new file mode 100644 index 0000000000..d834d82707 --- /dev/null +++ b/Cubical/Papers/Synthetic.agda @@ -0,0 +1,214 @@ +{-# OPTIONS --safe #-} +module Cubical.Papers.Synthetic where + +-- Cubical synthetic homotopy theory +-- Mörtberg and Pujet, „Cubical synthetic homotopy theory“. +-- https://dl.acm.org/doi/abs/10.1145/3372885.3373825 + +-- 2.1 +import Agda.Builtin.Cubical.Path as Path +import Cubical.Foundations.Prelude as Prelude +-- 2.2 +import Agda.Primitive.Cubical as PrimitiveCubical +import Cubical.Data.Bool as Bool +import Cubical.Core.Primitives as CorePrimitives +-- 2.3 +import Cubical.HITs.S1 as S1 +-- 2.4 +import Cubical.Foundations.Equiv as Equiv +import Cubical.Core.Glue as CoreGlue +import Cubical.Foundations.Univalence as Univalence +-- 3. +import Cubical.HITs.Torus as T2 +-- 3.1 +import Cubical.Data.Int as Int +import Cubical.Data.Int.Properties as IntProp +import Cubical.Foundations.Isomorphism as Isomorphism +-- 4.1 +import Cubical.HITs.Susp as Suspension +import Cubical.HITs.Sn as Sn +import Agda.Builtin.Nat as BNat +import Agda.Builtin.Bool as BBool +import Cubical.Foundations.GroupoidLaws as GroupoidLaws +import Cubical.HITs.S2 as S2 +import Cubical.HITs.S3 as S3 +-- 4.2 +import Cubical.HITs.Pushout as Push +import Cubical.HITs.Pushout.Properties as PushProp +-- 4.3 +import Cubical.HITs.Join as Join +import Cubical.HITs.Join.Properties as JoinProp +-- 5. +import Cubical.Homotopy.Hopf as Hopf + +-------------------------------------------------------------------------------- +-- 2. Cubical Agda +-- 2.1 The Interval and Path Types +-- 2.2 Transport and Composition +-- 2.3 Higher Inductive Types +-- 2.4 Glue Types and Univalence +-------------------------------------------------------------------------------- + +-- 2.1 The Interval and Path Types +open Path using (PathP) public +open Prelude using (_≡_ ; refl ; funExt) public +open Prelude renaming (sym to _⁻¹) public + +-- 2.2 Transport and Composition +open Prelude using (transport ; subst ; J ; JRefl) public +open PrimitiveCubical using (Partial) public +open Bool using (Bool ; true ; false) public + +partialBool : ∀ i → Partial (i ∨ ~ i) Bool +partialBool i = λ {(i = i1) → true + ; (i = i0) → false} + +open CorePrimitives using (inS ; outS ; hcomp) public +open Prelude using (_∙_) public + +-- 2.3 Higher Inductive Types +open S1 using (S¹ ; base ; loop) public + +double : S¹ → S¹ +double base = base +double (loop i) = (loop ∙ loop) i + +-- 2.4 Glue Types and Univalence +open Equiv using (idEquiv) public +open CoreGlue using (Glue) public +open Univalence using (ua) public + +-------------------------------------------------------------------------------- +-- 3. The Circle and Torus +-- 3.1 The Loop Spaces of the Circle and Torus +-------------------------------------------------------------------------------- + +open T2 using ( Torus ; point ; line1 ; line2 ; square + ; t2c ; c2t ; c2t-t2c ; t2c-c2t ; Torus≡S¹×S¹) + +-- 3.1 The Loop Spaces of the Circle and Torus +open S1 using (ΩS¹) public +open T2 using (ΩTorus) public +open Int using (pos ; negsuc) renaming (ℤ to Int) public +open IntProp using (sucPathℤ) public +open S1 using (helix ; winding) public + +-- Examples computing the winding numbers of the circle +_ : winding (loop ∙ loop ∙ loop) ≡ pos 3 +_ = refl + +_ : winding ((loop ⁻¹) ∙ loop ∙ (loop ⁻¹)) ≡ negsuc 0 +_ = refl + +open S1 renaming (intLoop to loopn) public +open S1 renaming (windingℤLoop to winding-loopn) public +open S1 using (encode ; decode ; decodeEncode ; ΩS¹≡ℤ) public +open Isomorphism using (isoToPath ; iso) public + +-- Notation of the paper, current notation under ΩS¹≡Int +ΩS¹≡Int' : ΩS¹ ≡ Int +ΩS¹≡Int' = isoToPath (iso winding loopn + winding-loopn (decodeEncode base)) + +open T2 using (ΩTorus≡ℤ×ℤ ; windingTorus) public + +-- Examples at the end of section 3. +_ : windingTorus (line1 ∙ line2) ≡ (pos 1 , pos 1) +_ = refl + +_ : windingTorus ((line1 ⁻¹) ∙ line2 ∙ line1) ≡ (pos 0 , pos 1) +_ = refl + +-------------------------------------------------------------------------------- +-- 4. Suspension, Spheres and Pushouts +-- 4.1 Suspension +-- 4.2 Pushouts and the 3 × 3 Lemma +-- 4.3 The Join and S³ +-------------------------------------------------------------------------------- + +-- 4.1 Suspension +open Suspension using (Susp ; north ; south ; merid) public +open Sn using (S₊) public +open Suspension using ( SuspBool→S¹ ; S¹→SuspBool + ; SuspBool→S¹→SuspBool + ; S¹→SuspBool→S¹) public + +-- Deprecated version of S₊ +open BNat renaming (Nat to ℕ) hiding (_*_) public +open CorePrimitives renaming (Type to Set) public +open BBool using (Bool) public +-- At the time the paper was published, Set was used instead of Type +_-sphere : ℕ → Set +0 -sphere = Bool +(suc n) -sphere = Susp (n -sphere) + +-- Lemma 4.1. The (1)-sphere is equal to the circle S1. + +open BBool using (true ; false) public +-- Deprecated version of SuspBool→S¹ +s2c : 1 -sphere → S¹ +s2c north = base +s2c south = base +s2c (merid true i) = loop i +s2c (merid false i) = base -- (loop ⁻¹) i + +-- Deprecated version of S¹→SuspBool +c2s : S¹ → 1 -sphere +c2s base = north +c2s (loop i) = (merid true ∙ (merid false ⁻¹)) i + +open GroupoidLaws using (rUnit) public +-- Deprecated version of SuspBool→S¹→SuspBool +s2c-c2s : ∀ (p : S¹) → s2c (c2s p) ≡ p +s2c-c2s base = refl +s2c-c2s (loop i) j = rUnit loop (~ j) i + +h1 : I → I → 1 -sphere +h1 i j = merid false (i ∧ j) + +h2 : I → I → 1 -sphere +h2 i j = hcomp (λ k → λ { (i = i0) → north + ; (i = i1) → merid false (j ∨ ~ k) + ; (j = i1) → merid true i }) + (merid true i) + +-- Deprecated version of S¹→SuspBool→S¹ +c2s-s2c : ∀ (t : 1 -sphere) → c2s (s2c t) ≡ t +c2s-s2c north j = north +c2s-s2c south j = merid false j +c2s-s2c (merid true i) j = h2 i j +c2s-s2c (merid false i) j = merid false (i ∧ j) + +-- Notation of the paper, current notation under S¹≡SuspBool +-- Proof of Lemma 4.1 +1-sphere≡S¹ : 1 -sphere ≡ S¹ +1-sphere≡S¹ = isoToPath (iso s2c c2s s2c-c2s c2s-s2c) + +-- Definitions of S2 and S3 +open S2 using (S²) public +open S3 using (S³) public + +-- 4.2 Pushouts and the 3 × 3 Lemma +open Push using (Pushout) public +-- 3x3-span is implemented as a record +open PushProp using (3x3-span) public +open 3x3-span using (f□1) public +-- The rest of the definitions inside the 3x3-lemma +-- A□0-A○□ , A□○-A○□ ... +open 3x3-span using (3x3-lemma) public + +-- 4.3 The Join and S³ +open Join renaming (join to Join) using (S³≡joinS¹S¹) public +open JoinProp using (join-assoc) public + +-------------------------------------------------------------------------------- +-- 5. The Hopf Fibration +-------------------------------------------------------------------------------- + +-- rot (denoted by _·_ here) in the paper is substituted by a rot and rotLoop in S1 +open S1 using (_·_ ; rotLoop) public +open Hopf.S¹Hopf renaming ( HopfSuspS¹ to Hopf + ; JoinS¹S¹→TotalHopf to j2h + ; TotalHopf→JoinS¹S¹ to h2j) + using (HopfS²) public +open S1 renaming (rotInv-1 to lem-rot-inv) public diff --git a/Cubical/Papers/ZCohomology.agda b/Cubical/Papers/ZCohomology.agda new file mode 100644 index 0000000000..cc390dd426 --- /dev/null +++ b/Cubical/Papers/ZCohomology.agda @@ -0,0 +1,536 @@ +{- + +Please do not move this file. Changes should only be made if +necessary. + +This file contains pointers to the code examples and main results from +the paper: + +Synthetic Integral Cohomology in Cubical Agda +Guillaume Brunerie, Axel Ljungström, Anders Mörtberg +Computer Science Logic (CSL) 2022 + +-} + +-- The "--safe" flag ensures that there are no postulates or +-- unfinished goals +{-# OPTIONS --safe #-} +module Cubical.Papers.ZCohomology where + +-- Misc. +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.Foundations.Everything +open import Cubical.HITs.S1 +open import Cubical.Data.Sum +open import Cubical.Data.Sigma + +-- 2 +open import Cubical.Core.Glue as Glue +import Cubical.Foundations.Prelude as Prelude +import Cubical.Foundations.GroupoidLaws as GroupoidLaws +import Cubical.Foundations.Path as Path +import Cubical.Foundations.Pointed as Pointed + renaming (Pointed to Type∙) +open import Cubical.HITs.S1 as S1 +open import Cubical.HITs.Susp as Suspension +open import Cubical.HITs.Sn as Sn +open import Cubical.Homotopy.Loopspace as Loop +open import Cubical.Foundations.HLevels as n-types +open import Cubical.HITs.Truncation as Trunc +open import Cubical.Homotopy.Connected as Connected +import Cubical.HITs.Pushout as Push +import Cubical.HITs.Wedge as ⋁ +import Cubical.Foundations.Univalence as Unival +import Cubical.Foundations.SIP as StructIdPrinc +import Cubical.Algebra.Group as Gr +import Cubical.Algebra.Group.GroupPath as GrPath + +-- 3 +import Cubical.ZCohomology.Base as coHom + renaming (coHomK to K ; coHomK-ptd to K∙) +import Cubical.HITs.Sn.Properties as S +import Cubical.ZCohomology.GroupStructure as GroupStructure +import Cubical.ZCohomology.Properties as Properties + renaming (Kn→ΩKn+1 to σ ; ΩKn+1→Kn to σ⁻¹) +import Cubical.Experiments.ZCohomologyOld.Properties as oldCohom + +-- 4 +import Cubical.ZCohomology.RingStructure.CupProduct as Cup +import Cubical.ZCohomology.RingStructure.RingLaws as ⌣Ring +import Cubical.ZCohomology.RingStructure.GradedCommutativity as ⌣Comm +import Cubical.Foundations.Pointed.Homogeneous as Homogen + +-- 5 +import Cubical.HITs.Torus as 𝕋² + renaming (Torus to 𝕋²) +import Cubical.HITs.KleinBottle as 𝕂² + renaming (KleinBottle to 𝕂²) +import Cubical.HITs.RPn as ℝP + renaming (RP² to ℝP²) +import Cubical.ZCohomology.Groups.Sn as HⁿSⁿ + renaming (Hⁿ-Sᵐ≅0 to Hⁿ-Sᵐ≅1) +import Cubical.ZCohomology.Groups.Torus as HⁿT² +import Cubical.ZCohomology.Groups.Wedge as Hⁿ-wedge +import Cubical.ZCohomology.Groups.KleinBottle as Hⁿ𝕂² +import Cubical.ZCohomology.Groups.RP2 as HⁿℝP² + renaming (H¹-RP²≅0 to H¹-RP²≅1) +import Cubical.ZCohomology.Groups.CP2 as HⁿℂP² + renaming (CP² to ℂP² ; ℤ→HⁿCP²→ℤ to g) + {- Remark: ℂP² is defined as the pushout S² ← TotalHopf → 1 in + the formalisation. TotalHopf is just the total space from the Hopf + fibration. We have TotalHopf ≃ S³, and the map TotalHopf → S² + is given by taking the first projection. This is equivalent to the + description given in the paper, since h : S³ → S² is given by + S³ ≃ TotalHopf → S² -} + +-- Additional material +import Cubical.Homotopy.EilenbergSteenrod as ES-axioms +import Cubical.ZCohomology.EilenbergSteenrodZ as satisfies-ES-axioms + renaming (coHomFunctor to H^~ ; coHomFunctor' to Ĥ) +import Cubical.ZCohomology.MayerVietorisUnreduced as MayerVietoris + +----- 2. HOMOTOPY TYPE THEORY IN CUBICAL AGDA ----- + +-- 2.1 Important notions in Cubical Agda +open Prelude using ( PathP + ; _≡_ + ; refl + ; cong + ; cong₂ + ; funExt) + +open GroupoidLaws using (_⁻¹) + +open Prelude using ( transport + ; subst + ; hcomp) + +--- 2.2 Important concepts from HoTT/UF in Cubical Agda + +-- Pointed Types +open Pointed using (Type∙) + +-- The circle, 𝕊¹ +open S1 using (S¹) + +-- Suspensions +open Suspension using (Susp) + +-- (Pointed) n-spheres, 𝕊ⁿ +open Sn using (S₊∙) + +-- Loop spaces +open Loop using (Ω^_) + +-- Eckmann-Hilton argument +Eckmann-Hilton : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → isComm∙ ((Ω^ (suc n)) A) +Eckmann-Hilton n α β = + transport (λ i → cong (λ x → rUnit x (~ i)) α ∙ cong (λ x → lUnit x (~ i)) β + ≡ cong (λ x → lUnit x (~ i)) β ∙ cong (λ x → rUnit x (~ i)) α) + (λ i → (λ j → α (j ∧ ~ i) ∙ β (j ∧ i)) ∙ λ j → α (~ i ∨ j) ∙ β (i ∨ j)) + +-- n-types Note that we start indexing from 0 in the Cubical Library +-- (so (-2)-types as referred to as 0-types, (-1) as 1-types, and so +-- on) +open n-types using (isOfHLevel) + +-- truncations +open Trunc using (hLevelTrunc) + +-- elimination principle +open Trunc using (elim) + +-- elimination principle for paths +truncPathElim : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} (n : ℕ) + → {B : Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ y ∣ → Type ℓ'} + → ((q : _) → isOfHLevel n (B q)) + → ((p : x ≡ y) → B (cong ∣_∣ p)) + → (q : _) → B q +truncPathElim zero hlev ind q = hlev q .fst +truncPathElim (suc n) {B = B} hlev ind q = + subst B (Iso.leftInv (Trunc.PathIdTruncIso _) q) + (help (ΩTrunc.encode-fun ∣ _ ∣ ∣ _ ∣ q)) + where + help : (q : _) → B (ΩTrunc.decode-fun ∣ _ ∣ ∣ _ ∣ q) + help = Trunc.elim (λ _ → hlev _) ind + +-- Connectedness +open Connected using (isConnected) + +-- Pushouts +open Push using (Pushout) + +-- Wedge sum +open ⋁ using (_⋁_) + + +-- 2.3 Univalence + +-- Univalence and the ua function respectively +open Unival using (univalence ; ua) + +-- Glue types +open Glue using (Glue) + +-- The structure identity principle and the sip function +-- respectively +open StructIdPrinc using (SIP ; sip) + +-- Groups +open Gr using (Group) + +-- Isomorphic groups are path equal +open GrPath using (GroupPath) + + +----- 3. INTEGRAL COHOMOLOGY IN CUBICAL AGDA ----- + + +-- 3.1 Eilenberg-MacLane spaces + +-- Eilenberg-MacLane spaces Kₙ (unpointed and pointed respectively) +open coHom using (K ; K∙) + +-- Proposition 7 +open S using (sphereConnected) + +-- Lemma 8 +open S using (wedgeconFun; wedgeconLeft ; wedgeconRight) + +-- restated to match the formulation in the paper +wedgeConSn' : ∀ {ℓ} (n m : ℕ) {A : (S₊ (suc n)) → (S₊ (suc m)) → Type ℓ} + → ((x : S₊ (suc n)) (y : S₊ (suc m)) → isOfHLevel ((suc n) + (suc m)) (A x y)) + → (fₗ : (x : _) → A x (ptSn (suc m))) + → (fᵣ : (x : _) → A (ptSn (suc n)) x) + → (p : fₗ (ptSn (suc n)) ≡ fᵣ (ptSn (suc m))) + → Σ[ F ∈ ((x : S₊ (suc n)) (y : S₊ (suc m)) → A x y) ] + (Σ[ (left , right) ∈ ((x : S₊ (suc n)) → fₗ x ≡ F x (ptSn (suc m))) + × ((x : S₊ (suc m)) → fᵣ x ≡ F (ptSn (suc n)) x) ] + p ≡ left (ptSn (suc n)) ∙ (right (ptSn (suc m))) ⁻¹) +wedgeConSn' zero zero hlev fₗ fᵣ p = + (wedgeconFun 0 0 hlev fᵣ fₗ p) + , ((λ x → sym (wedgeconRight 0 0 hlev fᵣ fₗ p x)) + , λ _ → refl) -- right holds by refl + , rUnit _ +wedgeConSn' zero (suc m) hlev fₗ fᵣ p = + (wedgeconFun 0 (suc m) hlev fᵣ fₗ p) + , ((λ _ → refl) -- left holds by refl + , (λ x → sym (wedgeconLeft 0 (suc m) hlev fᵣ fₗ p x))) + , lUnit _ +wedgeConSn' (suc n) m hlev fₗ fᵣ p = + (wedgeconFun (suc n) m hlev fᵣ fₗ p) + , ((λ x → sym (wedgeconRight (suc n) m hlev fᵣ fₗ p x)) + , λ _ → refl) -- right holds by refl + , rUnit _ + +-- +ₖ (addition) and 0ₖ +open GroupStructure using (_+ₖ_ ; 0ₖ) + +-- -ₖ (subtraction) +open GroupStructure using (-ₖ_) + +-- The function σ : Kₙ → ΩKₙ₊₁ +open Properties using (σ) + +-- Group laws for +ₖ +open GroupStructure using ( rUnitₖ ; lUnitₖ + ; rCancelₖ ; lCancelₖ + ; commₖ + ; assocₖ) + +-- All group laws are equal to refl at 0ₖ +-- rUnitₖ (definitional) +0-rUnit≡refl : rUnitₖ 0 (0ₖ 0) ≡ refl +1-rUnit≡refl : rUnitₖ 1 (0ₖ 1) ≡ refl +n≥2-rUnit≡refl : {n : ℕ} → rUnitₖ (2 + n) (0ₖ (2 + n)) ≡ refl +0-rUnit≡refl = refl +1-rUnit≡refl = refl +n≥2-rUnit≡refl = refl + +-- lUnitₖ (definitional) +0-lUnit≡refl : lUnitₖ 0 (0ₖ 0) ≡ refl +1-lUnit≡refl : lUnitₖ 1 (0ₖ 1) ≡ refl +n≥2-lUnit≡refl : {n : ℕ} → lUnitₖ (2 + n) (0ₖ (2 + n)) ≡ refl +0-lUnit≡refl = refl +1-lUnit≡refl = refl +n≥2-lUnit≡refl = refl + +-- assocₖ (definitional) +0-assoc≡refl : assocₖ 0 (0ₖ 0) (0ₖ 0) (0ₖ 0) ≡ refl +1-assoc≡refl : assocₖ 1 (0ₖ 1) (0ₖ 1) (0ₖ 1) ≡ refl +n≥2-assoc≡refl : {n : ℕ} → assocₖ (2 + n) (0ₖ (2 + n)) (0ₖ (2 + n)) (0ₖ (2 + n)) ≡ refl +0-assoc≡refl = refl +1-assoc≡refl = refl +n≥2-assoc≡refl = refl + +-- commₖ (≡ refl ∙ refl for n ≥ 2) +0-comm≡refl : commₖ 0 (0ₖ 0) (0ₖ 0) ≡ refl +1-comm≡refl : commₖ 1 (0ₖ 1) (0ₖ 1) ≡ refl +n≥2-comm≡refl : {n : ℕ} → commₖ (2 + n) (0ₖ (2 + n)) (0ₖ (2 + n)) ≡ refl +0-comm≡refl = refl +1-comm≡refl = refl +n≥2-comm≡refl = sym (rUnit refl) + +-- lCancelₖ (definitional) +0-lCancel≡refl : lCancelₖ 0 (0ₖ 0) ≡ refl +1-lCancel≡refl : lCancelₖ 1 (0ₖ 1) ≡ refl +n≥2-lCancel≡refl : {n : ℕ} → lCancelₖ (2 + n) (0ₖ (2 + n)) ≡ refl +0-lCancel≡refl = refl +1-lCancel≡refl = refl +n≥2-lCancel≡refl = refl + +-- rCancelₖ (≡ (refl ∙ refl) ∙ refl for n ≥ 2) +0-rCancel≡refl : rCancelₖ 0 (0ₖ 0) ≡ refl +1-rCancel≡refl : rCancelₖ 1 (0ₖ 1) ≡ refl +n≥2-rCancel≡refl : {n : ℕ} → rCancelₖ (2 + n) (0ₖ (2 + n)) ≡ refl +0-rCancel≡refl = refl +1-rCancel≡refl = refl +n≥2-rCancel≡refl i = rUnit (rUnit refl (~ i)) (~ i) + +-- Proof that there is a unique h-structure on Kₙ +-- +ₖ defines an h-Structure on Kₙ +open GroupStructure using (_+ₖ_ ; 0ₖ ; rUnitₖ ; lUnitₖ ; lUnitₖ≡rUnitₖ) + +-- and so does Brunerie's addition +open oldCohom using (_+ₖ_ ; 0ₖ ; rUnitₖ ; lUnitₖ ; rUnitlUnit0) + +-- consequently both additions agree +open GroupStructure using (+ₖ-unique) +open oldCohom using (addLemma) +additionsAgree : (n : ℕ) → GroupStructure._+ₖ_ {n = n} ≡ oldCohom._+ₖ_ {n = n} +additionsAgree zero i x y = oldCohom.addLemma x y (~ i) +additionsAgree (suc n) i x y = + +ₖ-unique n (GroupStructure._+ₖ_) (oldCohom._+ₖ_) + (GroupStructure.rUnitₖ (suc n)) (GroupStructure.lUnitₖ (suc n)) + (oldCohom.rUnitₖ (suc n)) (oldCohom.lUnitₖ (suc n)) + (sym (lUnitₖ≡rUnitₖ (suc n))) + (rUnitlUnit0 (suc n)) x y i + +-- Theorem 9 (Kₙ ≃ ΩKₙ₊₁) +open Properties using (Kn≃ΩKn+1) + +-- σ and σ⁻¹ are morphisms +-- (for σ⁻¹ this is proved directly without using the fact that σ is a morphism) +open Properties using (Kn→ΩKn+1-hom ; ΩKn+1→Kn-hom) + +-- Lemma 10 (p ∙ q ≡ cong²₊(p,q)) for n = 1 and n ≥ 2 respectively +open GroupStructure using (∙≡+₁ ; ∙≡+₂) + +-- Lemma 11 (cong²₊ is commutative) and Theorem 12 respectively +open GroupStructure using (cong+ₖ-comm ; isCommΩK) + +-- 3.2 Group structure on Hⁿ(A) + +-- +ₕ (addition), -ₕ and 0ₕ +open GroupStructure using (_+ₕ_ ; -ₕ_ ; 0ₕ) + +-- Cohomology group structure +open GroupStructure using ( rUnitₕ ; lUnitₕ + ; rCancelₕ ; lCancelₕ + ; commₕ + ; assocₕ) + +--- Additional material ------------------------------------------- + +-- Reduced cohomology, group structure +open GroupStructure using (coHomRedGroupDir) + +-- Equality of unreduced and reduced cohmology +open Properties using (coHomGroup≡coHomRedGroup) +-------------------------------------------------------------------- + +----- 4. The Cup Product and Cohomology Ring ----- +-- 4.1 +-- Lemma 13 +open Properties using (isOfHLevel↑∙) + +-- ⌣ₖ +open Cup using (_⌣ₖ_) + +-- ⌣ₖ is pointed in both arguments +open ⌣Ring using (0ₖ-⌣ₖ ; ⌣ₖ-0ₖ) + +-- The cup product +open Cup using (_⌣_) + +-- 4.2 +-- Lemma 14 +Lem14 : ∀ {ℓ} {A : Type∙ ℓ} (n : ℕ) (f g : A →∙ K∙ n) → fst f ≡ fst g → f ≡ g +Lem14 n f g p = Homogen.→∙Homogeneous≡ (Properties.isHomogeneousKn n) p + +-- Proposition 15 +open ⌣Ring using (leftDistr-⌣ₖ ; rightDistr-⌣ₖ) + +-- Lemma 16 +open ⌣Ring using (assocer-helpFun≡) + +-- Proposition 17 +open ⌣Ring using (assoc-⌣ₖ) + +-- Proposition 18 +open ⌣Comm using (gradedComm-⌣ₖ) + +-- Ring structure on ⌣ +open ⌣Ring using (leftDistr-⌣ ; rightDistr-⌣ + ; assoc-⌣ ; 1⌣ + ; rUnit⌣ ; lUnit⌣ + ; ⌣0 ; 0⌣) +open ⌣Comm using (gradedComm-⌣) + +----- 5. CHARACTERIZING INTEGRAL COHOMOLOGY GROUPS ----- + +-- 5.1 +-- Proposition 19 +open HⁿSⁿ using (Hⁿ-Sⁿ≅ℤ) + +-- 5.2 +-- The torus +open 𝕋² using (𝕋²) + +-- Propositions 20 and 21 respectively +open HⁿT² using (H¹-T²≅ℤ×ℤ ; H²-T²≅ℤ) + +-- 5.3 +-- The Klein bottle +open 𝕂² using (𝕂²) + +-- The real projective plane +open ℝP using (ℝP²) + +-- Proposition 22 and 24 respectively +-- ℤ/2ℤ is represented by Bool with the unique group structure +-- Lemma 23 is used implicitly in H²-𝕂²≅Bool +open Hⁿ𝕂² using (H¹-𝕂²≅ℤ ; H²-𝕂²≅Bool) + +-- First and second cohomology groups of ℝP² respectively +open HⁿℝP² using (H¹-RP²≅1 ; H²-RP²≅Bool) + +-- 5.4 +-- The complex projective plane +open HⁿℂP² using (ℂP²) + +-- Second and fourth cohomology groups ℂP² respectively +open HⁿℂP² using (H²CP²≅ℤ ; H⁴CP²≅ℤ) + +-- 6 Proving by computations in Cubical Agda +-- Proof of m = n = 1 case of graded commutativity (post truncation elimination): +-- Uncomment and give it a minute. The proof is currently not running very fast. + +{- +open ⌣Comm using (-ₖ^_·_ ) +n=m=1 : (a b : S¹) + → _⌣ₖ_ {n = 1} {m = 1} ∣ a ∣ ∣ b ∣ + ≡ (-ₖ (_⌣ₖ_ {n = 1} {m = 1} ∣ b ∣ ∣ a ∣)) +n=m=1 base base = refl +n=m=1 base (loop i) k = -ₖ (Properties.Kn→ΩKn+10ₖ _ (~ k) i) +n=m=1 (loop i) base k = Properties.Kn→ΩKn+10ₖ _ k i +n=m=1 (loop i) (loop j) k = -- This hcomp is just a simple rewriting to get paths in Ω²K₂ + hcomp (λ r → λ { (i = i0) → -ₖ Properties.Kn→ΩKn+10ₖ _ (~ k ∨ ~ r) j + ; (i = i1) → -ₖ Properties.Kn→ΩKn+10ₖ _ (~ k ∨ ~ r) j + ; (j = i0) → Properties.Kn→ΩKn+10ₖ _ (k ∨ ~ r) i + ; (j = i1) → Properties.Kn→ΩKn+10ₖ _ (k ∨ ~ r) i + ; (k = i0) → + doubleCompPath-filler + (sym (Properties.Kn→ΩKn+10ₖ _)) + (λ j i → _⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣) + (Properties.Kn→ΩKn+10ₖ _) + (~ r) j i + ; (k = i1) → + -ₖ doubleCompPath-filler + (sym (Properties.Kn→ΩKn+10ₖ _)) + (λ j i → _⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣) + (Properties.Kn→ΩKn+10ₖ _) + (~ r) i j}) + ((main + ∙ sym (cong-∙∙ (cong (-ₖ_)) (sym (Properties.Kn→ΩKn+10ₖ _)) + (λ j i → (_⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣)) + (Properties.Kn→ΩKn+10ₖ _))) k i j) + where + open import Cubical.Foundations.Equiv.HalfAdjoint + t : Iso (typ ((Ω^ 2) (K∙ 2))) ℤ + t = compIso (congIso (invIso (Properties.Iso-Kn-ΩKn+1 1))) + (invIso (Properties.Iso-Kn-ΩKn+1 0)) + + p₁ = flipSquare ((sym (Properties.Kn→ΩKn+10ₖ _)) + ∙∙ (λ j i → _⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣) + ∙∙ (Properties.Kn→ΩKn+10ₖ _)) + p₂ = (cong (cong (-ₖ_)) + ((sym (Properties.Kn→ΩKn+10ₖ _)))) + ∙∙ (λ j i → -ₖ (_⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣)) + ∙∙ (cong (cong (-ₖ_)) (Properties.Kn→ΩKn+10ₖ _)) + + computation : Iso.fun t p₁ ≡ Iso.fun t p₂ + computation = refl + + main : p₁ ≡ p₂ + main = p₁ ≡⟨ sym (Iso.leftInv t p₁) ⟩ + (Iso.inv t (Iso.fun t p₁)) ≡⟨ cong (Iso.inv t) computation ⟩ + Iso.inv t (Iso.fun t p₂) ≡⟨ Iso.leftInv t p₂ ⟩ + p₂ ∎ +-} + +-- 𝕋² ≠ S² ∨ S¹ ∨ S¹ +open HⁿT² using (T²≠S²⋁S¹⋁S¹) + +-- Second "Brunerie number" +open HⁿℂP² using (g) +brunerie2 : ℤ +brunerie2 = g 1 + +-- Additional material (from the appendix of the preprint) +----- A. Proofs ----- + +-- A.2 Proofs for Section 4 + +-- Lemma 27 +open Homogen using (→∙Homogeneous≡) + +-- Lemma 28 +open Homogen using (isHomogeneous→∙) + +-- Lemma 29 +open Properties using (isHomogeneousKn) + +-- Lemma 30, parts 1-3 respectively +open Path using (sym≡flipSquare ; sym-cong-sym≡id ; sym≡cong-sym) + +-- Lemma 31 +open ⌣Comm using (cong-ₖ-gen-inr) + + +-- A.3 Proofs for Section 5 + +-- Proposition 32 +open HⁿSⁿ using (Hⁿ-Sᵐ≅1) + +----- B. THE EILENBERG-STEENROD AXIOMS ----- + +-- B.1 The axioms in HoTT/UF + +-- The axioms are defined as a record type +open ES-axioms.coHomTheory + +-- Proof of the claim that the alternative reduced cohomology functor +-- Ĥ is the same as the usual reduced cohomology functor +_ : ∀ {ℓ} → satisfies-ES-axioms.H^~ {ℓ} ≡ satisfies-ES-axioms.Ĥ +_ = satisfies-ES-axioms.coHomFunctor≡coHomFunctor' + +-- B.2 Verifying the axioms + +-- Propositions 35 and 36. +_ : ∀ {ℓ} → ES-axioms.coHomTheory {ℓ} satisfies-ES-axioms.H^~ +_ = satisfies-ES-axioms.isCohomTheoryZ + + +-- B.3 Characterizing Z-cohomology groups using the axioms + +-- Theorem 37 +open MayerVietoris.MV using ( Ker-i⊂Im-d ; Im-d⊂Ker-i + ; Ker-Δ⊂Im-i ; Im-i⊂Ker-Δ + ; Ker-d⊂Im-Δ ; Im-Δ⊂Ker-d) + + +----- C. BENCHMARKING COMPUTATIONS WITH THE COHOMOLOGY GROUPS ----- + +import Cubical.Experiments.ZCohomology.Benchmarks diff --git a/Cubical/README.agda b/Cubical/README.agda new file mode 100644 index 0000000000..c35afe47a3 --- /dev/null +++ b/Cubical/README.agda @@ -0,0 +1,73 @@ +{-# OPTIONS --guardedness #-} + +module Cubical.README where + +------------------------------------------------------------------------ +-- An experimental library for Cubical Agda +----------------------------------------------------------------------- + +-- The library comes with a .agda-lib file, for use with the library +-- management system. + +------------------------------------------------------------------------ +-- Module hierarchy +------------------------------------------------------------------------ + +-- The core library for Cubical Agda. +-- It contains basic primitives, equivalences, glue types. +import Cubical.Core.Everything + +-- The foundations for Cubical Agda. +-- The Prelude module is self-explanatory. +import Cubical.Foundations.Prelude +import Cubical.Foundations.Everything + +-- Kinds and properties of functions +import Cubical.Functions.Everything + +-- Data types and properties +import Cubical.Data.Everything + +-- Higher-inductive types +import Cubical.HITs.Everything + +-- Coinductive data types and properties +import Cubical.Codata.Everything + +-- Papers +import Cubical.Papers.Everything + +-- Properties and proofs about relations +import Cubical.Relation.Everything + +-- Category theory +import Cubical.Categories.Everything + +-- Homotopy theory +import Cubical.Homotopy.Everything + +-- Properties and kinds of Modalities +import Cubical.Modalities.Everything + +-- Various experiments using Cubical Agda +import Cubical.Experiments.Everything + +-- Other modules (TODO: add descriptions) +import Cubical.Induction.Everything +import Cubical.Structures.Everything +import Cubical.ZCohomology.Everything + +-- Algebra library (in development) +import Cubical.Algebra.Everything + +-- Various talks +import Cubical.Talks.Everything + +-- Reflection +import Cubical.Reflection.Everything + +-- Displayed univalent graphs +import Cubical.Displayed.Everything + +-- Syntax typeclasses +import Cubical.Syntax.Everything diff --git a/Cubical/Reflection/Base.agda b/Cubical/Reflection/Base.agda new file mode 100644 index 0000000000..b807d51e5c --- /dev/null +++ b/Cubical/Reflection/Base.agda @@ -0,0 +1,59 @@ +{- + +Some basic utilities for reflection + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Reflection.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.List.Base +open import Cubical.Data.Nat.Base + +import Agda.Builtin.Reflection as R +open import Agda.Builtin.String + +_>>=_ = R.bindTC +_<|>_ = R.catchTC + +_$_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → A → B +f $ a = f a + +_>>_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → R.TC A → R.TC B → R.TC B +f >> g = f >>= λ _ → g + +infixl 4 _>>=_ _>>_ _<|>_ +infixr 3 _$_ + +liftTC : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → R.TC A → R.TC B +liftTC f ta = ta >>= λ a → R.returnTC (f a) + +v : ℕ → R.Term +v n = R.var n [] + +pattern varg t = R.arg (R.arg-info R.visible (R.modality R.relevant R.quantity-ω)) t +pattern harg {q = q} t = R.arg (R.arg-info R.hidden (R.modality R.relevant q)) t +pattern _v∷_ a l = varg a ∷ l +pattern _h∷_ a l = harg a ∷ l + +infixr 5 _v∷_ _h∷_ + +vlam : String → R.Term → R.Term +vlam str t = R.lam R.visible (R.abs str t) + +hlam : String → R.Term → R.Term +hlam str t = R.lam R.hidden (R.abs str t) + +newMeta = R.checkType R.unknown + +extend*Context : ∀ {ℓ} {A : Type ℓ} → List (R.Arg R.Type) → R.TC A → R.TC A +extend*Context [] tac = tac +extend*Context (a ∷ as) tac = R.extendContext a (extend*Context as tac) + +makeAuxiliaryDef : String → R.Type → R.Term → R.TC R.Term +makeAuxiliaryDef s ty term = + R.freshName s >>= λ name → + R.declareDef (varg name) ty >> + R.defineFun name [ R.clause [] [] term ] >> + R.returnTC (R.def name []) + diff --git a/Cubical/Reflection/RecordEquiv.agda b/Cubical/Reflection/RecordEquiv.agda new file mode 100644 index 0000000000..9e76f9d284 --- /dev/null +++ b/Cubical/Reflection/RecordEquiv.agda @@ -0,0 +1,196 @@ +{- + + Reflection-based tools for converting between iterated record types, particularly between + record types and iterated Σ-types. + + See end of file for examples. + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Reflection.RecordEquiv where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Data.List as List +open import Cubical.Data.Nat +open import Cubical.Data.Maybe as Maybe +open import Cubical.Data.Sigma + +open import Agda.Builtin.String +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +Projections = Maybe (List R.Name) + +-- Describes a correspondence between two iterated record types +RecordAssoc = List (Projections × Projections) + +-- Describes a correspondence between a record type and an iterated Σ-types; +-- more convenient than RecordAssoc for this special case +data ΣFormat : Type where + leaf : R.Name → ΣFormat + _,_ : ΣFormat → ΣFormat → ΣFormat + unit : ΣFormat + +infixr 4 _,_ + +flipRecordAssoc : RecordAssoc → RecordAssoc +flipRecordAssoc = List.map λ {p .fst → p .snd; p .snd → p .fst} + +fstIdRecordAssoc : RecordAssoc → RecordAssoc +fstIdRecordAssoc = List.map λ {p .fst → p .fst; p .snd → p .fst} + +List→ΣFormat : List R.Name → ΣFormat +List→ΣFormat [] = unit +List→ΣFormat (x ∷ []) = leaf x +List→ΣFormat (x ∷ y ∷ xs) = leaf x , List→ΣFormat (y ∷ xs) + +ΣFormat→RecordAssoc : ΣFormat → RecordAssoc +ΣFormat→RecordAssoc = go [] + where + go : List R.Name → ΣFormat → RecordAssoc + go prefix unit = [ just prefix , nothing ] + go prefix (leaf fieldName) = [ just prefix , just [ fieldName ] ] + go prefix (sig₁ , sig₂) = + go (quote fst ∷ prefix) sig₁ ++ go (quote snd ∷ prefix) sig₂ + +-- Derive the shape of the compound Σ-type +ΣFormat→Ty : ΣFormat → R.Type +ΣFormat→Ty unit = R.def (quote Unit) [] +ΣFormat→Ty (leaf _) = R.unknown +ΣFormat→Ty (sig₁ , sig₂) = + R.def (quote Σ) (ΣFormat→Ty sig₁ v∷ R.lam R.visible (R.abs "_" (ΣFormat→Ty sig₂)) v∷ []) + +recordName→isoTy : R.Name → R.Term → R.TC R.Term +recordName→isoTy name σShape = + R.inferType (R.def name []) >>= R.normalise >>= go [] + where + go : List R.ArgInfo → R.Type → R.TC R.Term + go acc (R.pi (R.arg i argTy) (R.abs s ty)) = + liftTC (λ t → R.pi (R.arg i' argTy) (R.abs s t)) (go (i ∷ acc) ty) + where + i' = R.arg-info R.hidden (R.modality R.relevant R.quantity-ω) + go acc (R.agda-sort _) = + R.returnTC (R.def (quote Iso) (R.def name (makeArgs 0 [] acc) v∷ σShape v∷ [])) + where + makeArgs : ℕ → List (R.Arg R.Term) → List R.ArgInfo → List (R.Arg R.Term) + makeArgs n acc [] = acc + makeArgs n acc (i ∷ infos) = makeArgs (suc n) (R.arg i (v n) ∷ acc) infos + go _ _ = R.typeError (R.strErr "Not a record type name: " ∷ R.nameErr name ∷ []) + +projNames→Patterns : List R.Name → List (R.Arg R.Pattern) +projNames→Patterns = go [] + where + go : List (R.Arg R.Pattern) → List R.Name → List (R.Arg R.Pattern) + go acc [] = acc + go acc (π ∷ projs) = go (varg (R.proj π) ∷ acc) projs + +projNames→Term : R.Term → List R.Name → R.Term +projNames→Term term [] = term +projNames→Term term (π ∷ projs) = R.def π [ varg (projNames→Term term projs) ] + +convertClauses : RecordAssoc → R.Term → List R.Clause +convertClauses al term = fixIfEmpty (List.filterMap makeClause al) + where + makeClause : Projections × Projections → Maybe R.Clause + makeClause (projl , just projr) = + just (R.clause [] (goPat [] projr) (Maybe.rec R.unknown goTm projl)) + where + goPat : List (R.Arg R.Pattern) → List R.Name → List (R.Arg R.Pattern) + goPat acc [] = acc + goPat acc (π ∷ projs) = goPat (varg (R.proj π) ∷ acc) projs + + goTm : List R.Name → R.Term + goTm [] = term + goTm (π ∷ projs) = R.def π [ varg (goTm projs) ] + makeClause (_ , nothing) = nothing + + fixIfEmpty : List R.Clause → List R.Clause + fixIfEmpty [] = [ R.clause [] [] R.unknown ] + fixIfEmpty (c ∷ cs) = c ∷ cs + +mapClause : + (List (String × R.Arg R.Type) → List (String × R.Arg R.Type)) + → (List (R.Arg R.Pattern) → List (R.Arg R.Pattern)) + → (R.Clause → R.Clause) +mapClause f g (R.clause tel ps t) = R.clause (f tel) (g ps) t +mapClause f g (R.absurd-clause tel ps) = R.absurd-clause (f tel) (g ps) + +recordIsoΣClauses : ΣFormat → List R.Clause +recordIsoΣClauses σ = + funClauses (quote Iso.fun) Σ↔R ++ + funClauses (quote Iso.inv) R↔Σ ++ + pathClauses (quote Iso.rightInv) R↔Σ ++ + pathClauses (quote Iso.leftInv) Σ↔R + where + R↔Σ = ΣFormat→RecordAssoc σ + Σ↔R = flipRecordAssoc R↔Σ + + funClauses : R.Name → RecordAssoc → List R.Clause + funClauses name al = + List.map + (mapClause + (("_" , varg R.unknown) ∷_) + (λ ps → R.proj name v∷ R.var 0 v∷ ps)) + (convertClauses al (v 0)) + + pathClauses : R.Name → RecordAssoc → List R.Clause + pathClauses name al = + List.map + (mapClause + (λ vs → ("_" , varg R.unknown) ∷ ("_" , varg R.unknown) ∷ vs) + (λ ps → R.proj name v∷ R.var 1 v∷ R.var 0 v∷ ps)) + (convertClauses (fstIdRecordAssoc al) (v 1)) + +recordIsoΣTerm : ΣFormat → R.Term +recordIsoΣTerm σ = R.pat-lam (recordIsoΣClauses σ) [] + +-- with a provided ΣFormat for the record +declareRecordIsoΣ' : R.Name → ΣFormat → R.Name → R.TC Unit +declareRecordIsoΣ' idName σ recordName = + let σTy = ΣFormat→Ty σ in + recordName→isoTy recordName σTy >>= λ isoTy → + R.declareDef (varg idName) isoTy >> + R.defineFun idName (recordIsoΣClauses σ) + +-- using the right-associated Σ given by the record fields +declareRecordIsoΣ : R.Name → R.Name → R.TC Unit +declareRecordIsoΣ idName recordName = + R.getDefinition recordName >>= λ where + (R.record-type _ fs) → + let σ = List→ΣFormat (List.map (λ {(R.arg _ n) → n}) fs) in + declareRecordIsoΣ' idName σ recordName + _ → + R.typeError (R.strErr "Not a record type name:" ∷ R.nameErr recordName ∷ []) + +private + module Example where + variable + ℓ ℓ' : Level + A : Type ℓ + B : A → Type ℓ' + + record Example0 {A : Type ℓ} (B : A → Type ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality -- works with or without eta equality + field + cool : A + fun : A + wow : B cool + + -- Declares a function `Example0IsoΣ` that gives an isomorphism between the record type and a + -- right-associated nested Σ-type (with the parameters to Example0 as implict arguments). + unquoteDecl Example0IsoΣ = declareRecordIsoΣ Example0IsoΣ (quote Example0) + + -- `Example0IsoΣ` has the type we expect + test0 : Iso (Example0 B) (Σ[ a ∈ A ] (Σ[ _ ∈ A ] B a)) + test0 = Example0IsoΣ + + -- A record with no fields is isomorphic to Unit + + record Example1 : Type where + + unquoteDecl Example1IsoΣ = declareRecordIsoΣ Example1IsoΣ (quote Example1) + + test1 : Iso Example1 Unit + test1 = Example1IsoΣ diff --git a/Cubical/Reflection/StrictEquiv.agda b/Cubical/Reflection/StrictEquiv.agda new file mode 100644 index 0000000000..63a819696d --- /dev/null +++ b/Cubical/Reflection/StrictEquiv.agda @@ -0,0 +1,81 @@ +{-# OPTIONS --safe #-} +module Cubical.Reflection.StrictEquiv where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.Base +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.List.Base +open import Cubical.Data.Unit.Base + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +strictEquivClauses : R.Term → R.Term → List R.Clause +strictEquivClauses f g = + R.clause [] + (R.proj (quote fst) v∷ []) + f + ∷ R.clause [] + (R.proj (quote snd) v∷ R.proj (quote equiv-proof) v∷ []) + (R.def (quote strictContrFibers) (g v∷ [])) + ∷ [] + +strictEquivTerm : R.Term → R.Term → R.Term +strictEquivTerm f g = R.pat-lam (strictEquivClauses f g) [] + +strictEquivMacro : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (A → B) → (B → A) → R.Term → R.TC Unit +strictEquivMacro {A = A} {B} f g hole = + R.quoteTC (A ≃ B) >>= λ equivTy → + R.checkType hole equivTy >> + R.quoteTC f >>= λ `f` → + R.quoteTC g >>= λ `g` → + R.unify (strictEquivTerm `f` `g`) hole + +strictIsoToEquivMacro : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Iso A B → R.Term → R.TC Unit +strictIsoToEquivMacro isom = + strictEquivMacro (Iso.fun isom) (Iso.inv isom) + +-- For use with unquoteDef + +defStrictEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → R.Name → (A → B) → (B → A) → R.TC Unit +defStrictEquiv idName f g = + R.quoteTC f >>= λ `f` → + R.quoteTC g >>= λ `g` → + R.defineFun idName (strictEquivClauses `f` `g`) + +defStrictIsoToEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → R.Name → Iso A B → R.TC Unit +defStrictIsoToEquiv idName isom = + defStrictEquiv idName (Iso.fun isom) (Iso.inv isom) + +-- For use with unquoteDef + +declStrictEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → R.Name → (A → B) → (B → A) → R.TC Unit +declStrictEquiv {A = A} {B = B} idName f g = + R.quoteTC (A ≃ B) >>= λ ty → + R.declareDef (varg idName) ty >> + defStrictEquiv idName f g + +declStrictIsoToEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → R.Name → Iso A B → R.TC Unit +declStrictIsoToEquiv idName isom = + declStrictEquiv idName (Iso.fun isom) (Iso.inv isom) + +macro + -- (f : A → B) → (g : B → A) → (A ≃ B) + -- Assumes that `f` and `g` are inverse up to definitional equality + strictEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (A → B) → (B → A) → R.Term → R.TC Unit + strictEquiv = strictEquivMacro + + -- (isom : Iso A B) → (A ≃ B) + -- Assumes that the functions defining `isom` are inverse up to definitional equality + strictIsoToEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Iso A B → R.Term → R.TC Unit + strictIsoToEquiv = strictIsoToEquivMacro diff --git a/Cubical/Relation/Binary.agda b/Cubical/Relation/Binary.agda new file mode 100644 index 0000000000..ec2f3e272c --- /dev/null +++ b/Cubical/Relation/Binary.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary where + +open import Cubical.Relation.Binary.Base public +open import Cubical.Relation.Binary.Properties public + diff --git a/Cubical/Relation/Binary/Base.agda b/Cubical/Relation/Binary/Base.agda new file mode 100644 index 0000000000..e476fa5fe2 --- /dev/null +++ b/Cubical/Relation/Binary/Base.agda @@ -0,0 +1,148 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Base where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Fiberwise +open import Cubical.Data.Sigma +open import Cubical.HITs.SetQuotients.Base +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓA ℓ≅A ℓA' ℓ≅A' : Level + +Rel : ∀ {ℓ} (A B : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +Rel A B ℓ' = A → B → Type ℓ' + +PropRel : ∀ {ℓ} (A B : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +PropRel A B ℓ' = Σ[ R ∈ Rel A B ℓ' ] ∀ a b → isProp (R a b) + +idPropRel : ∀ {ℓ} (A : Type ℓ) → PropRel A A ℓ +idPropRel A .fst a a' = ∥ a ≡ a' ∥ +idPropRel A .snd _ _ = squash + +invPropRel : ∀ {ℓ ℓ'} {A B : Type ℓ} + → PropRel A B ℓ' → PropRel B A ℓ' +invPropRel R .fst b a = R .fst a b +invPropRel R .snd b a = R .snd a b + +compPropRel : ∀ {ℓ ℓ' ℓ''} {A B C : Type ℓ} + → PropRel A B ℓ' → PropRel B C ℓ'' → PropRel A C (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) +compPropRel R S .fst a c = ∥ Σ[ b ∈ _ ] (R .fst a b × S .fst b c) ∥ +compPropRel R S .snd _ _ = squash + +graphRel : ∀ {ℓ} {A B : Type ℓ} → (A → B) → Rel A B ℓ +graphRel f a b = f a ≡ b + +module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : Rel A A ℓ') where + isRefl : Type (ℓ-max ℓ ℓ') + isRefl = (a : A) → R a a + + isSym : Type (ℓ-max ℓ ℓ') + isSym = (a b : A) → R a b → R b a + + isAntisym : Type (ℓ-max ℓ ℓ') + isAntisym = (a b : A) → R a b → R b a → a ≡ b + + isTrans : Type (ℓ-max ℓ ℓ') + isTrans = (a b c : A) → R a b → R b c → R a c + + record isEquivRel : Type (ℓ-max ℓ ℓ') where + constructor equivRel + field + reflexive : isRefl + symmetric : isSym + transitive : isTrans + + isPropValued : Type (ℓ-max ℓ ℓ') + isPropValued = (a b : A) → isProp (R a b) + + isSetValued : Type (ℓ-max ℓ ℓ') + isSetValued = (a b : A) → isSet (R a b) + + isEffective : Type (ℓ-max ℓ ℓ') + isEffective = + (a b : A) → isEquiv (eq/ {R = R} a b) + + + impliesIdentity : Type _ + impliesIdentity = {a a' : A} → (R a a') → (a ≡ a') + + -- the total space corresponding to the binary relation w.r.t. a + relSinglAt : (a : A) → Type (ℓ-max ℓ ℓ') + relSinglAt a = Σ[ a' ∈ A ] (R a a') + + -- the statement that the total space is contractible at any a + contrRelSingl : Type (ℓ-max ℓ ℓ') + contrRelSingl = (a : A) → isContr (relSinglAt a) + + isUnivalent : Type (ℓ-max ℓ ℓ') + isUnivalent = (a a' : A) → (R a a') ≃ (a ≡ a') + + contrRelSingl→isUnivalent : isRefl → contrRelSingl → isUnivalent + contrRelSingl→isUnivalent ρ c a a' = isoToEquiv i + where + h : isProp (relSinglAt a) + h = isContr→isProp (c a) + aρa : relSinglAt a + aρa = a , ρ a + Q : (y : A) → a ≡ y → _ + Q y _ = R a y + i : Iso (R a a') (a ≡ a') + Iso.fun i r = cong fst (h aρa (a' , r)) + Iso.inv i = J Q (ρ a) + Iso.rightInv i = J (λ y p → cong fst (h aρa (y , J Q (ρ a) p)) ≡ p) + (J (λ q _ → cong fst (h aρa (a , q)) ≡ refl) + (J (λ α _ → cong fst α ≡ refl) refl + (isProp→isSet h _ _ refl (h _ _))) + (sym (JRefl Q (ρ a)))) + Iso.leftInv i r = J (λ w β → J Q (ρ a) (cong fst β) ≡ snd w) + (JRefl Q (ρ a)) (h aρa (a' , r)) + + isUnivalent→contrRelSingl : isUnivalent → contrRelSingl + isUnivalent→contrRelSingl u a = q + where + abstract + f : (x : A) → a ≡ x → R a x + f x p = invEq (u a x) p + + t : singl a → relSinglAt a + t (x , p) = x , f x p + + q : isContr (relSinglAt a) + q = isOfHLevelRespectEquiv 0 (t , totalEquiv _ _ f λ x → invEquiv (u a x) .snd) + (isContrSingl a) + +EquivRel : ∀ {ℓ} (A : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +EquivRel A ℓ' = Σ[ R ∈ Rel A A ℓ' ] BinaryRelation.isEquivRel R + +EquivPropRel : ∀ {ℓ} (A : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +EquivPropRel A ℓ' = Σ[ R ∈ PropRel A A ℓ' ] BinaryRelation.isEquivRel (R .fst) + +record RelIso {A : Type ℓA} (_≅_ : Rel A A ℓ≅A) + {A' : Type ℓA'} (_≅'_ : Rel A' A' ℓ≅A') : Type (ℓ-max (ℓ-max ℓA ℓA') (ℓ-max ℓ≅A ℓ≅A')) where + constructor reliso + field + fun : A → A' + inv : A' → A + rightInv : (a' : A') → fun (inv a') ≅' a' + leftInv : (a : A) → inv (fun a) ≅ a + +open BinaryRelation + +RelIso→Iso : {A : Type ℓA} {A' : Type ℓA'} + (_≅_ : Rel A A ℓ≅A) (_≅'_ : Rel A' A' ℓ≅A') + (uni : impliesIdentity _≅_) (uni' : impliesIdentity _≅'_) + (f : RelIso _≅_ _≅'_) + → Iso A A' +Iso.fun (RelIso→Iso _ _ _ _ f) = RelIso.fun f +Iso.inv (RelIso→Iso _ _ _ _ f) = RelIso.inv f +Iso.rightInv (RelIso→Iso _ _ uni uni' f) a' + = uni' (RelIso.rightInv f a') +Iso.leftInv (RelIso→Iso _ _ uni uni' f) a + = uni (RelIso.leftInv f a) diff --git a/Cubical/Relation/Binary/Poset.agda b/Cubical/Relation/Binary/Poset.agda new file mode 100644 index 0000000000..57d7560484 --- /dev/null +++ b/Cubical/Relation/Binary/Poset.agda @@ -0,0 +1,138 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Poset where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.RecordEquiv +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Relation.Binary.Base + +open Iso +open BinaryRelation + + +private + variable + ℓ ℓ' ℓ'' ℓ₀ ℓ₀' ℓ₁ ℓ₁' : Level + +record IsPoset {A : Type ℓ} (_≤_ : A → A → Type ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality + constructor isposet + + field + is-set : isSet A + is-prop-valued : isPropValued _≤_ + is-refl : isRefl _≤_ + is-trans : isTrans _≤_ + is-antisym : isAntisym _≤_ + +unquoteDecl IsPosetIsoΣ = declareRecordIsoΣ IsPosetIsoΣ (quote IsPoset) + + +record PosetStr (ℓ' : Level) (A : Type ℓ) : Type (ℓ-max ℓ (ℓ-suc ℓ')) where + + constructor posetstr + + field + _≤_ : A → A → Type ℓ' + isPoset : IsPoset _≤_ + + infixl 7 _≤_ + + open IsPoset isPoset public + +Poset : ∀ ℓ ℓ' → Type (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) +Poset ℓ ℓ' = TypeWithStr ℓ (PosetStr ℓ') + +poset : (A : Type ℓ) (_≤_ : A → A → Type ℓ') (h : IsPoset _≤_) → Poset ℓ ℓ' +poset A _≤_ h = A , posetstr _≤_ h + +record IsPosetEquiv {A : Type ℓ₀} {B : Type ℓ₁} + (M : PosetStr ℓ₀' A) (e : A ≃ B) (N : PosetStr ℓ₁' B) + : Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') ℓ₁') + where + constructor + isposetequiv + -- Shorter qualified names + private + module M = PosetStr M + module N = PosetStr N + + field + pres≤ : (x y : A) → x M.≤ y ≃ equivFun e x N.≤ equivFun e y + + +PosetEquiv : (M : Poset ℓ₀ ℓ₀') (M : Poset ℓ₁ ℓ₁') → Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') (ℓ-max ℓ₁ ℓ₁')) +PosetEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsPosetEquiv (M .snd) e (N .snd) + +isPropIsPoset : {A : Type ℓ} (_≤_ : A → A → Type ℓ') → isProp (IsPoset _≤_) +isPropIsPoset _≤_ = isOfHLevelRetractFromIso 1 IsPosetIsoΣ + (isPropΣ isPropIsSet + λ isSetA → isPropΣ (isPropΠ2 (λ _ _ → isPropIsProp)) + λ isPropValued≤ → isProp×2 + (isPropΠ (λ _ → isPropValued≤ _ _)) + (isPropΠ5 λ _ _ _ _ _ → isPropValued≤ _ _) + (isPropΠ4 λ _ _ _ _ → isSetA _ _)) + +𝒮ᴰ-Poset : DUARel (𝒮-Univ ℓ) (PosetStr ℓ') (ℓ-max ℓ ℓ') +𝒮ᴰ-Poset = + 𝒮ᴰ-Record (𝒮-Univ _) IsPosetEquiv + (fields: + data[ _≤_ ∣ autoDUARel _ _ ∣ pres≤ ] + prop[ isPoset ∣ (λ _ _ → isPropIsPoset _) ]) + where + open PosetStr + open IsPoset + open IsPosetEquiv + +PosetPath : (M N : Poset ℓ ℓ') → PosetEquiv M N ≃ (M ≡ N) +PosetPath = ∫ 𝒮ᴰ-Poset .UARel.ua + +-- an easier way of establishing an equivalence of posets +module _ {P : Poset ℓ₀ ℓ₀'} {S : Poset ℓ₁ ℓ₁'} (e : ⟨ P ⟩ ≃ ⟨ S ⟩) where + private + module P = PosetStr (P .snd) + module S = PosetStr (S .snd) + + module _ (isMon : ∀ x y → x P.≤ y → equivFun e x S.≤ equivFun e y) + (isMonInv : ∀ x y → x S.≤ y → invEq e x P.≤ invEq e y) where + open IsPosetEquiv + open IsPoset + + makeIsPosetEquiv : IsPosetEquiv (P .snd) e (S .snd) + pres≤ makeIsPosetEquiv x y = propBiimpl→Equiv (P.isPoset .is-prop-valued _ _) + (S.isPoset .is-prop-valued _ _) + (isMon _ _) (isMonInv' _ _) + where + isMonInv' : ∀ x y → equivFun e x S.≤ equivFun e y → x P.≤ y + isMonInv' x y ex≤ey = transport (λ i → retEq e x i P.≤ retEq e y i) (isMonInv _ _ ex≤ey) + + +module PosetReasoning (P' : Poset ℓ ℓ') where + private P = fst P' + open PosetStr (snd P') + open IsPoset + + _≤⟨_⟩_ : (x : P) {y z : P} → x ≤ y → y ≤ z → x ≤ z + x ≤⟨ p ⟩ q = isPoset .is-trans x _ _ p q + + _◾ : (x : P) → x ≤ x + x ◾ = isPoset .is-refl x + + infixr 0 _≤⟨_⟩_ + infix 1 _◾ diff --git a/Cubical/Relation/Binary/Properties.agda b/Cubical/Relation/Binary/Properties.agda new file mode 100644 index 0000000000..90aa324b4e --- /dev/null +++ b/Cubical/Relation/Binary/Properties.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels diff --git a/Cubical/Relation/Nullary.agda b/Cubical/Relation/Nullary.agda new file mode 100644 index 0000000000..ac3c578188 --- /dev/null +++ b/Cubical/Relation/Nullary.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Nullary where + +open import Cubical.Relation.Nullary.Base public +open import Cubical.Relation.Nullary.Properties public diff --git a/Cubical/Relation/Nullary/Base.agda b/Cubical/Relation/Nullary/Base.agda new file mode 100644 index 0000000000..fdd1b3fa6d --- /dev/null +++ b/Cubical/Relation/Nullary/Base.agda @@ -0,0 +1,66 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Nullary.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Functions.Fixpoint + +open import Cubical.Data.Empty as ⊥ +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +-- Negation +infix 3 ¬_ + +¬_ : Type ℓ → Type ℓ +¬ A = A → ⊥ + +-- Decidable types (inspired by standard library) +data Dec (P : Type ℓ) : Type ℓ where + yes : ( p : P) → Dec P + no : (¬p : ¬ P) → Dec P + +NonEmpty : Type ℓ → Type ℓ +NonEmpty A = ¬ ¬ A + +Stable : Type ℓ → Type ℓ +Stable A = NonEmpty A → A + +-- reexport propositional truncation for uniformity +open Cubical.HITs.PropositionalTruncation.Base + using (∥_∥) public + +SplitSupport : Type ℓ → Type ℓ +SplitSupport A = ∥ A ∥ → A + +Collapsible : Type ℓ → Type ℓ +Collapsible A = Σ[ f ∈ (A → A) ] 2-Constant f + +Populated ⟪_⟫ : Type ℓ → Type ℓ +Populated A = (f : Collapsible A) → Fixpoint (f .fst) +⟪_⟫ = Populated + +PStable : Type ℓ → Type ℓ +PStable A = ⟪ A ⟫ → A + +onAllPaths : (Type ℓ → Type ℓ) → Type ℓ → Type ℓ +onAllPaths S A = (x y : A) → S (x ≡ y) + +Separated : Type ℓ → Type ℓ +Separated = onAllPaths Stable + +HSeparated : Type ℓ → Type ℓ +HSeparated = onAllPaths SplitSupport + +Collapsible≡ : Type ℓ → Type ℓ +Collapsible≡ = onAllPaths Collapsible + +PStable≡ : Type ℓ → Type ℓ +PStable≡ = onAllPaths PStable + +Discrete : Type ℓ → Type ℓ +Discrete = onAllPaths Dec diff --git a/Cubical/Relation/Nullary/DecidableEq.agda b/Cubical/Relation/Nullary/DecidableEq.agda new file mode 100644 index 0000000000..7c66fe00a6 --- /dev/null +++ b/Cubical/Relation/Nullary/DecidableEq.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Nullary.DecidableEq where + +open import Cubical.Relation.Nullary.Properties + using (Dec→Stable; Discrete→isSet) public diff --git a/Cubical/Relation/Nullary/HLevels.agda b/Cubical/Relation/Nullary/HLevels.agda new file mode 100644 index 0000000000..e4f29c1ff1 --- /dev/null +++ b/Cubical/Relation/Nullary/HLevels.agda @@ -0,0 +1,38 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Nullary.HLevels where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Functions.Fixpoint + +open import Cubical.Relation.Nullary + +private + variable + ℓ : Level + A : Type ℓ + +isPropPopulated : isProp (Populated A) +isPropPopulated = isPropΠ λ x → 2-Constant→isPropFixpoint (x .fst) (x .snd) + +isPropHSeparated : isProp (HSeparated A) +isPropHSeparated f g i x y a = HSeparated→isSet f x y (f x y a) (g x y a) i + +isPropCollapsible≡ : isProp (Collapsible≡ A) +isPropCollapsible≡ {A = A} f = (isPropΠ2 λ x y → isPropCollapsiblePointwise) f where + sA : isSet A + sA = Collapsible≡→isSet f + gA : isGroupoid A + gA = isSet→isGroupoid sA + isPropCollapsiblePointwise : ∀ {x y} → isProp (Collapsible (x ≡ y)) + isPropCollapsiblePointwise {x} {y} (a , ca) (b , cb) = λ i → endoFunction i , endoFunctionIsConstant i where + endoFunction : a ≡ b + endoFunction = funExt λ p → sA _ _ (a p) (b p) + isProp2-Constant : (k : I) → isProp (2-Constant (endoFunction k)) + isProp2-Constant k = isPropΠ2 λ r s → gA x y (endoFunction k r) (endoFunction k s) + endoFunctionIsConstant : PathP (λ i → 2-Constant (endoFunction i)) ca cb + endoFunctionIsConstant = isProp→PathP isProp2-Constant ca cb + +isPropDiscrete : isProp (Discrete A) +isPropDiscrete p q = isPropΠ2 (λ x y → isPropDec (Discrete→isSet p x y)) p q diff --git a/Cubical/Relation/Nullary/Properties.agda b/Cubical/Relation/Nullary/Properties.agda new file mode 100644 index 0000000000..ab6343dcce --- /dev/null +++ b/Cubical/Relation/Nullary/Properties.agda @@ -0,0 +1,168 @@ +{-# OPTIONS --safe #-} +{- + +Properties of nullary relations, i.e. structures on types. + +Includes several results from [Notions of Anonymous Existence in Martin-Löf Type Theory](https://lmcs.episciences.org/3217) +by Altenkirch, Coquand, Escardo, Kraus. + +-} +module Cubical.Relation.Nullary.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Fixpoint + +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Relation.Nullary.Base +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +IsoPresDiscrete : ∀ {ℓ ℓ'}{A : Type ℓ} {B : Type ℓ'} → Iso A B + → Discrete A → Discrete B +IsoPresDiscrete e dA x y with dA (Iso.inv e x) (Iso.inv e y) +... | yes p = subst Dec (λ i → Iso.rightInv e x i ≡ Iso.rightInv e y i) + (yes (cong (Iso.fun e) p)) +... | no p = subst Dec (λ i → Iso.rightInv e x i ≡ Iso.rightInv e y i) + (no λ q → p (sym (Iso.leftInv e (Iso.inv e x)) + ∙∙ cong (Iso.inv e) q + ∙∙ Iso.leftInv e (Iso.inv e y))) + +EquivPresDiscrete : ∀ {ℓ ℓ'}{A : Type ℓ} {B : Type ℓ'} → A ≃ B + → Discrete A → Discrete B +EquivPresDiscrete e = IsoPresDiscrete (equivToIso e) + +isProp¬ : (A : Type ℓ) → isProp (¬ A) +isProp¬ A p q i x = isProp⊥ (p x) (q x) i + +Stable¬ : Stable (¬ A) +Stable¬ ¬¬¬a a = ¬¬¬a ¬¬a + where + ¬¬a = λ ¬a → ¬a a + +fromYes : A → Dec A → A +fromYes _ (yes a) = a +fromYes a (no _) = a + +discreteDec : (Adis : Discrete A) → Discrete (Dec A) +discreteDec Adis (yes p) (yes p') = decideYes (Adis p p') -- TODO: monad would simply stuff + where + decideYes : Dec (p ≡ p') → Dec (yes p ≡ yes p') + decideYes (yes eq) = yes (cong yes eq) + decideYes (no ¬eq) = no λ eq → ¬eq (cong (fromYes p) eq) +discreteDec Adis (yes p) (no ¬p) = ⊥.rec (¬p p) +discreteDec Adis (no ¬p) (yes p) = ⊥.rec (¬p p) +discreteDec {A = A} Adis (no ¬p) (no ¬p') = yes (cong no (isProp¬ A ¬p ¬p')) + +isPropDec : (Aprop : isProp A) → isProp (Dec A) +isPropDec Aprop (yes a) (yes a') = cong yes (Aprop a a') +isPropDec Aprop (yes a) (no ¬a) = ⊥.rec (¬a a) +isPropDec Aprop (no ¬a) (yes a) = ⊥.rec (¬a a) +isPropDec {A = A} Aprop (no ¬a) (no ¬a') = cong no (isProp¬ A ¬a ¬a') + +mapDec : ∀ {B : Type ℓ} → (A → B) → (¬ A → ¬ B) → Dec A → Dec B +mapDec f _ (yes p) = yes (f p) +mapDec _ f (no ¬p) = no (f ¬p) + +-- we have the following implications +-- X ── ∣_∣ ─→ ∥ X ∥ +-- ∥ X ∥ ── populatedBy ─→ ⟪ X ⟫ +-- ⟪ X ⟫ ── notEmptyPopulated ─→ NonEmpty X + +-- reexport propositional truncation for uniformity +open Cubical.HITs.PropositionalTruncation.Base + using (∣_∣) public + +populatedBy : ∥ A ∥ → ⟪ A ⟫ +populatedBy {A = A} a (f , fIsConst) = h a where + h : ∥ A ∥ → Fixpoint f + h ∣ a ∣ = f a , fIsConst (f a) a + h (squash a b i) = 2-Constant→isPropFixpoint f fIsConst (h a) (h b) i + +notEmptyPopulated : ⟪ A ⟫ → NonEmpty A +notEmptyPopulated {A = A} pop u = u (fixpoint (pop (h , hIsConst))) where + h : A → A + h a = ⊥.elim (u a) + hIsConst : ∀ x y → h x ≡ h y + hIsConst x y i = ⊥.elim (isProp⊥ (u x) (u y) i) + +-- these implications induce the following for different kinds of stability, gradually weakening the assumption +Dec→Stable : Dec A → Stable A +Dec→Stable (yes x) = λ _ → x +Dec→Stable (no x) = λ f → ⊥.elim (f x) + +Stable→PStable : Stable A → PStable A +Stable→PStable st = st ∘ notEmptyPopulated + +PStable→SplitSupport : PStable A → SplitSupport A +PStable→SplitSupport pst = pst ∘ populatedBy + +-- Although SplitSupport and Collapsible are not properties, their path versions, HSeparated and Collapsible≡, are. +-- Nevertheless they are logically equivalent +SplitSupport→Collapsible : SplitSupport A → Collapsible A +SplitSupport→Collapsible {A = A} hst = h , hIsConst where + h : A → A + h p = hst ∣ p ∣ + hIsConst : 2-Constant h + hIsConst p q i = hst (squash ∣ p ∣ ∣ q ∣ i) + +Collapsible→SplitSupport : Collapsible A → SplitSupport A +Collapsible→SplitSupport f x = fixpoint (populatedBy x f) + +HSeparated→Collapsible≡ : HSeparated A → Collapsible≡ A +HSeparated→Collapsible≡ hst x y = SplitSupport→Collapsible (hst x y) + +Collapsible≡→HSeparated : Collapsible≡ A → HSeparated A +Collapsible≡→HSeparated col x y = Collapsible→SplitSupport (col x y) + +-- stability of path space under truncation or path collapsability are necessary and sufficient properties +-- for a a type to be a set. +Collapsible≡→isSet : Collapsible≡ A → isSet A +Collapsible≡→isSet {A = A} col a b p q = p≡q where + f : (x : A) → a ≡ x → a ≡ x + f x = col a x .fst + fIsConst : (x : A) → (p q : a ≡ x) → f x p ≡ f x q + fIsConst x = col a x .snd + rem : (p : a ≡ b) → PathP (λ i → a ≡ p i) (f a refl) (f b p) + rem p j = f (p j) (λ i → p (i ∧ j)) + p≡q : p ≡ q + p≡q j i = hcomp (λ k → λ { (i = i0) → f a refl k + ; (i = i1) → fIsConst b p q j k + ; (j = i0) → rem p i k + ; (j = i1) → rem q i k }) a + +HSeparated→isSet : HSeparated A → isSet A +HSeparated→isSet = Collapsible≡→isSet ∘ HSeparated→Collapsible≡ + +isSet→HSeparated : isSet A → HSeparated A +isSet→HSeparated setA x y = extract where + extract : ∥ x ≡ y ∥ → x ≡ y + extract ∣ p ∣ = p + extract (squash p q i) = setA x y (extract p) (extract q) i + +-- by the above more sufficient conditions to inhibit isSet A are given +PStable≡→HSeparated : PStable≡ A → HSeparated A +PStable≡→HSeparated pst x y = PStable→SplitSupport (pst x y) + +PStable≡→isSet : PStable≡ A → isSet A +PStable≡→isSet = HSeparated→isSet ∘ PStable≡→HSeparated + +Separated→PStable≡ : Separated A → PStable≡ A +Separated→PStable≡ st x y = Stable→PStable (st x y) + +Separated→isSet : Separated A → isSet A +Separated→isSet = PStable≡→isSet ∘ Separated→PStable≡ + +-- Proof of Hedberg's theorem: a type with decidable equality is an h-set +Discrete→Separated : Discrete A → Separated A +Discrete→Separated d x y = Dec→Stable (d x y) + +Discrete→isSet : Discrete A → isSet A +Discrete→isSet = Separated→isSet ∘ Discrete→Separated diff --git a/Cubical/Relation/ZigZag/Applications/MultiSet.agda b/Cubical/Relation/ZigZag/Applications/MultiSet.agda new file mode 100644 index 0000000000..df50082472 --- /dev/null +++ b/Cubical/Relation/ZigZag/Applications/MultiSet.agda @@ -0,0 +1,358 @@ +-- We apply the theory of quasi equivalence relations (QERs) to finite multisets and association lists. +{-# OPTIONS --safe #-} +module Cubical.Relation.ZigZag.Applications.MultiSet where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.RelationalStructure +open import Cubical.Foundations.Structure +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Univalence +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat +open import Cubical.Data.List hiding ([_]) +open import Cubical.Data.Sigma +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.FiniteMultiset as FMS hiding ([_] ; _++_) +open import Cubical.HITs.FiniteMultiset.CountExtensionality +open import Cubical.HITs.PropositionalTruncation +open import Cubical.Relation.Nullary +open import Cubical.Relation.ZigZag.Base + +open import Cubical.Structures.MultiSet +open import Cubical.Structures.Relational.Auto +open import Cubical.Structures.Relational.Macro + +-- we define simple association lists without any higher constructors +data AList {ℓ} (A : Type ℓ) : Type ℓ where + ⟨⟩ : AList A + ⟨_,_⟩∷_ : A → ℕ → AList A → AList A + +infixr 5 ⟨_,_⟩∷_ + +private + variable + ℓ : Level + +-- We have a CountStructure on List and AList and use these to get a QER between the two +module Lists&ALists {A : Type ℓ} (discA : Discrete A) where + + multisetShape : Type ℓ → Type ℓ + multisetShape X = X × (A → X → X) × (X → X → X) × (A → X → Const[ ℕ , isSetℕ ]) + + module S = RelMacro ℓ (autoRelDesc multisetShape) + + addIfEq : (a x : A) → ℕ → ℕ → ℕ + addIfEq a x m n with discA a x + ... | yes _ = m + n + ... | no _ = n + + module _ {a x : A} {n : ℕ} where + + addIfEq≡ : {m : ℕ} → a ≡ x → addIfEq a x m n ≡ m + n + addIfEq≡ a≡x with discA a x + ... | yes _ = refl + ... | no a≢x = ⊥.rec (a≢x a≡x) + + addIfEq≢ : {m : ℕ} → ¬ (a ≡ x) → addIfEq a x m n ≡ n + addIfEq≢ a≢x with discA a x + ... | yes a≡x = ⊥.rec (a≢x a≡x) + ... | no _ = refl + + addIfEq0 : addIfEq a x 0 n ≡ n + addIfEq0 with discA a x + ... | yes _ = refl + ... | no _ = refl + + addIfEq+ : {m : ℕ} (n' : ℕ) → addIfEq a x m (n + n') ≡ addIfEq a x m n + n' + addIfEq+ n' with discA a x + ... | yes _ = +-assoc _ n n' + ... | no _ = refl + + module L where + emp : List A + emp = [] + + insert : A → List A → List A + insert x xs = x ∷ xs + + union : List A → List A → List A + union xs ys = xs ++ ys + + count : A → List A → ℕ + count a [] = zero + count a (x ∷ xs) = addIfEq a x 1 (count a xs) + + structure : S.structure (List A) + structure = emp , insert , union , count + + countUnion : ∀ a xs ys → count a (union xs ys) ≡ count a xs + count a ys + countUnion a [] ys = refl + countUnion a (x ∷ xs) ys = + cong (addIfEq a x 1) (countUnion a xs ys) + ∙ addIfEq+ (count a ys) + + module AL where + emp : AList A + emp = ⟨⟩ + + insert* : ℕ → A → AList A → AList A + insert* m a ⟨⟩ = ⟨ a , m ⟩∷ ⟨⟩ + insert* m a (⟨ y , n ⟩∷ ys) with (discA a y) + ... | yes _ = ⟨ y , m + n ⟩∷ ys + ... | no _ = ⟨ y , n ⟩∷ insert* m a ys + + insert : A → AList A → AList A + insert = insert* 1 + + union : AList A → AList A → AList A + union ⟨⟩ ys = ys + union (⟨ x , n ⟩∷ xs) ys = insert* n x (union xs ys) + + count : A → AList A → ℕ + count a ⟨⟩ = zero + count a (⟨ y , n ⟩∷ ys) = addIfEq a y n (count a ys) + + structure : S.structure (AList A) + structure = emp , insert , union , count + + countInsert* : ∀ m a x ys → count a (insert* m x ys) ≡ addIfEq a x m (count a ys) + countInsert* m a x ⟨⟩ = refl + countInsert* m a x (⟨ y , n ⟩∷ ys) with discA a x | discA a y | discA x y + ... | yes a≡x | yes a≡y | yes x≡y = addIfEq≡ a≡y ∙ sym (+-assoc m n _) + ... | yes a≡x | yes a≡y | no x≢y = ⊥.rec (x≢y (sym a≡x ∙ a≡y)) + ... | yes a≡x | no a≢y | yes x≡y = ⊥.rec (a≢y (a≡x ∙ x≡y)) + ... | yes a≡x | no a≢y | no x≢y = addIfEq≢ a≢y ∙ countInsert* m a x ys ∙ addIfEq≡ a≡x + ... | no a≢x | yes a≡y | yes x≡y = ⊥.rec (a≢x (a≡y ∙ sym x≡y)) + ... | no a≢x | yes a≡y | no x≢y = addIfEq≡ a≡y ∙ cong (n +_) (countInsert* m a x ys ∙ addIfEq≢ a≢x) + ... | no a≢x | no a≢y | yes x≡y = addIfEq≢ a≢y + ... | no a≢x | no a≢y | no x≢y = addIfEq≢ a≢y ∙ countInsert* m a x ys ∙ addIfEq≢ a≢x + + countInsert = countInsert* 1 + + countUnion : ∀ a xs ys → count a (union xs ys) ≡ count a xs + count a ys + countUnion a ⟨⟩ ys = refl + countUnion a (⟨ x , n ⟩∷ xs) ys = + countInsert* n a x (union xs ys) + ∙ cong (addIfEq a x n) (countUnion a xs ys) + ∙ addIfEq+ (count a ys) + + -- now for the QER between List and Alist + + R : List A → AList A → Type ℓ + R xs ys = ∀ a → L.count a xs ≡ AL.count a ys + + φ : List A → AList A + φ [] = ⟨⟩ + φ (x ∷ xs) = AL.insert x (φ xs) + + ψ : AList A → List A + ψ ⟨⟩ = [] + ψ (⟨ x , zero ⟩∷ xs) = ψ xs + ψ (⟨ x , suc n ⟩∷ xs) = x ∷ ψ (⟨ x , n ⟩∷ xs) + + η : ∀ xs → R xs (φ xs) + η [] a = refl + η (x ∷ xs) a = cong (addIfEq a x 1) (η xs a) ∙ sym (AL.countInsert a x (φ xs)) + + -- for the other direction we need a little helper function + ε : ∀ y → R (ψ y) y + ε' : (x : A) (n : ℕ) (xs : AList A) (a : A) + → L.count a (ψ (⟨ x , n ⟩∷ xs)) ≡ AL.count a (⟨ x , n ⟩∷ xs) + + ε ⟨⟩ a = refl + ε (⟨ x , n ⟩∷ xs) a = ε' x n xs a + + ε' x zero xs a = ε xs a ∙ sym addIfEq0 + ε' x (suc n) xs a with discA a x + ... | yes a≡x = cong suc (ε' x n xs a ∙ addIfEq≡ a≡x) + ... | no a≢x = ε' x n xs a ∙ addIfEq≢ a≢x + + -- Induced quotients and equivalence + + open isQuasiEquivRel + + -- R is a QER + QuasiR : QuasiEquivRel _ _ ℓ + QuasiR .fst .fst = R + QuasiR .fst .snd _ _ = isPropΠ λ _ → isSetℕ _ _ + QuasiR .snd .zigzag r r' r'' a = (r a) ∙∙ sym (r' a) ∙∙ (r'' a) + QuasiR .snd .fwd a = ∣ φ a , η a ∣ + QuasiR .snd .bwd b = ∣ ψ b , ε b ∣ + + isStructuredInsert : (x : A) {xs : List A} {ys : AList A} + → R xs ys → R (L.insert x xs) (AL.insert x ys) + isStructuredInsert x {xs} {ys} r a = + cong (addIfEq a x 1) (r a) ∙ sym (AL.countInsert a x ys) + + isStructuredUnion : + {xs : List A} {ys : AList A} (r : R xs ys) + {xs' : List A} {ys' : AList A} (r' : R xs' ys') + → R (L.union xs xs') (AL.union ys ys') + isStructuredUnion {xs} {ys} r {xs'} {ys'} r' a = + L.countUnion a xs xs' ∙ cong₂ _+_ (r a) (r' a) ∙ sym (AL.countUnion a ys ys') + + -- R is structured + isStructuredR : S.relation R L.structure AL.structure + isStructuredR .fst a = refl + isStructuredR .snd .fst = isStructuredInsert + isStructuredR .snd .snd .fst {xs} {ys} = isStructuredUnion {xs} {ys} + isStructuredR .snd .snd .snd a r = r a + + module E = QER→Equiv QuasiR + open E renaming (Rᴸ to Rᴸ; Rᴿ to Rᴬᴸ) + + List/Rᴸ = (List A) / Rᴸ + AList/Rᴬᴸ = (AList A) / Rᴬᴸ + + List/Rᴸ≃AList/Rᴬᴸ : List/Rᴸ ≃ AList/Rᴬᴸ + List/Rᴸ≃AList/Rᴬᴸ = E.Thm + + main : QERDescends _ S.relation (List A , L.structure) (AList A , AL.structure) QuasiR + main = structuredQER→structuredEquiv S.suitable _ _ QuasiR isStructuredR + + open QERDescends + + LQstructure : S.structure List/Rᴸ + LQstructure = main .quoᴸ .fst + + ALQstructure : S.structure AList/Rᴬᴸ + ALQstructure = main .quoᴿ .fst + + -- We get a path between structure over the equivalence from the fact that the QER is structured + List/Rᴸ≡AList/Rᴬᴸ : + Path (TypeWithStr ℓ S.structure) (List/Rᴸ , LQstructure) (AList/Rᴬᴸ , ALQstructure) + List/Rᴸ≡AList/Rᴬᴸ = + sip S.univalent _ _ + (E.Thm , S.matches (List/Rᴸ , LQstructure) (AList/Rᴬᴸ , ALQstructure) E.Thm .fst (main .rel)) + + -- Deriving associativity of union for association list multisets + + LQunion = LQstructure .snd .snd .fst + ALQunion = ALQstructure .snd .snd .fst + + hasAssociativeUnion : TypeWithStr ℓ S.structure → Type ℓ + hasAssociativeUnion (_ , _ , _ , _⊔_ , _) = + ∀ xs ys zs → (xs ⊔ ys) ⊔ zs ≡ xs ⊔ (ys ⊔ zs) + + LQassoc : hasAssociativeUnion (List/Rᴸ , LQstructure) + LQassoc = elimProp3 (λ _ _ _ → squash/ _ _) (λ xs ys zs i → [ ++-assoc xs ys zs i ]) + + ALQassoc : hasAssociativeUnion (AList/Rᴬᴸ , ALQstructure) + ALQassoc = subst hasAssociativeUnion List/Rᴸ≡AList/Rᴬᴸ LQassoc + + -- We now show that List/Rᴸ≃FMSet + + _∷/_ : A → List/Rᴸ → List/Rᴸ + _∷/_ = LQstructure .snd .fst + + multisetShape' : Type ℓ → Type ℓ + multisetShape' X = X × (A → X → X) × (A → X → Const[ ℕ , isSetℕ ]) + + FMSstructure : S.structure (FMSet A) + FMSstructure = [] , _∷_ , FMS._++_ , FMScount discA + + infixr 5 _∷/_ + + FMSet→List/Rᴸ : FMSet A → List/Rᴸ + FMSet→List/Rᴸ = FMS.Rec.f squash/ [ [] ] _∷/_ β + where + δ : ∀ c a b xs → L.count c (a ∷ b ∷ xs) ≡ L.count c (b ∷ a ∷ xs) + δ c a b xs with discA c a | discA c b + δ c a b xs | yes _ | yes _ = refl + δ c a b xs | yes _ | no _ = refl + δ c a b xs | no _ | yes _ = refl + δ c a b xs | no _ | no _ = refl + + γ : ∀ a b xs → Rᴸ (a ∷ b ∷ xs) (b ∷ a ∷ xs) + γ a b xs = + ∣ φ (a ∷ b ∷ xs) , η (a ∷ b ∷ xs) , (λ c → δ c b a xs ∙ η (a ∷ b ∷ xs) c) ∣ + + β : ∀ a b [xs] → a ∷/ b ∷/ [xs] ≡ b ∷/ a ∷/ [xs] + β a b = elimProp (λ _ → squash/ _ _) (λ xs → eq/ _ _ (γ a b xs)) + + -- The inverse is induced by the standard projection of lists into finite multisets, + -- which is a morphism of CountStructures + -- Moreover, we need 'count-extensionality' for finite multisets + List→FMSet : List A → FMSet A + List→FMSet [] = [] + List→FMSet (x ∷ xs) = x ∷ List→FMSet xs + + List→FMSet-count : ∀ a xs → L.count a xs ≡ FMScount discA a (List→FMSet xs) + List→FMSet-count a [] = refl + List→FMSet-count a (x ∷ xs) with discA a x + ... | yes _ = cong suc (List→FMSet-count a xs) + ... | no _ = List→FMSet-count a xs + + List/Rᴸ→FMSet : List/Rᴸ → FMSet A + List/Rᴸ→FMSet [ xs ] = List→FMSet xs + List/Rᴸ→FMSet (eq/ xs ys r i) = path i + where + countsAgree : ∀ a → L.count a xs ≡ L.count a ys + countsAgree a = cong (LQstructure .snd .snd .snd a) (eq/ xs ys r) + + θ : ∀ a → FMScount discA a (List→FMSet xs) ≡ FMScount discA a (List→FMSet ys) + θ a = sym (List→FMSet-count a xs) ∙∙ countsAgree a ∙∙ List→FMSet-count a ys + + path : List→FMSet xs ≡ List→FMSet ys + path = FMScountExt.Thm discA _ _ θ + List/Rᴸ→FMSet (squash/ xs/ xs/' p q i j) = + trunc (List/Rᴸ→FMSet xs/) (List/Rᴸ→FMSet xs/') (cong List/Rᴸ→FMSet p) (cong List/Rᴸ→FMSet q) i j + + List/Rᴸ→FMSet-insert : (x : A) (ys : List/Rᴸ) → List/Rᴸ→FMSet (x ∷/ ys) ≡ x ∷ List/Rᴸ→FMSet ys + List/Rᴸ→FMSet-insert x = elimProp (λ _ → FMS.trunc _ _) λ xs → refl + + List→FMSet-union : (xs ys : List A) + → List→FMSet (xs ++ ys) ≡ FMS._++_ (List→FMSet xs) (List→FMSet ys) + List→FMSet-union [] ys = refl + List→FMSet-union (x ∷ xs) ys = cong (x ∷_) (List→FMSet-union xs ys) + + List/Rᴸ≃FMSet : List/Rᴸ ≃ FMSet A + List/Rᴸ≃FMSet = isoToEquiv (iso List/Rᴸ→FMSet FMSet→List/Rᴸ τ σ) + where + σ' : (xs : List A) → FMSet→List/Rᴸ (List/Rᴸ→FMSet [ xs ]) ≡ [ xs ] + σ' [] = refl + σ' (x ∷ xs) = cong (x ∷/_) (σ' xs) + + σ : section FMSet→List/Rᴸ List/Rᴸ→FMSet + σ = elimProp (λ _ → squash/ _ _) σ' + + τ' : ∀ x {xs} → List/Rᴸ→FMSet (FMSet→List/Rᴸ xs) ≡ xs → List/Rᴸ→FMSet (FMSet→List/Rᴸ (x ∷ xs)) ≡ x ∷ xs + τ' x {xs} p = List/Rᴸ→FMSet-insert x (FMSet→List/Rᴸ xs) ∙ cong (x ∷_) p + + τ : retract FMSet→List/Rᴸ List/Rᴸ→FMSet + τ = FMS.ElimProp.f (FMS.trunc _ _) refl τ' + + List/Rᴸ≃FMSet-EquivStr : S.equiv (List/Rᴸ , LQstructure) (FMSet A , FMSstructure) List/Rᴸ≃FMSet + List/Rᴸ≃FMSet-EquivStr .fst = refl + List/Rᴸ≃FMSet-EquivStr .snd .fst a xs = List/Rᴸ→FMSet-insert a xs + List/Rᴸ≃FMSet-EquivStr .snd .snd .fst = elimProp2 (λ _ _ → trunc _ _) List→FMSet-union + List/Rᴸ≃FMSet-EquivStr .snd .snd .snd a = + elimProp (λ _ → isSetℕ _ _) (List→FMSet-count a) + + {- + Putting everything together we get: + ≃ + List/Rᴸ ------------> AList/Rᴬᴸ + | + |≃ + | + ∨ + ≃ + FMSet A ------------> AssocList A + We thus get that AList/Rᴬᴸ≃AssocList. + Constructing such an equivalence directly requires count extensionality for association lists, + which should be even harder to prove than for finite multisets. + This strategy should work for all implementations of multisets with HITs. + We just have to show that: + ∙ The HIT is equivalent to FMSet (like AssocList) + ∙ There is a QER between lists and the basic data type of the HIT + with the higher constructors removed (like AList) + Then we get that this HIT is equivalent to the corresponding set quotient that identifies elements + that give the same count on each a : A. + TODO: Show that all the equivalences are indeed isomorphisms of multisets not only of CountStructures! + -} diff --git a/Cubical/Relation/ZigZag/Base.agda b/Cubical/Relation/ZigZag/Base.agda new file mode 100644 index 0000000000..01cf6c1191 --- /dev/null +++ b/Cubical/Relation/ZigZag/Base.agda @@ -0,0 +1,161 @@ +-- We define ZigZag-complete relations and prove that quasi equivalence relations +-- give rise to equivalences on the set quotients. +{-# OPTIONS --safe #-} +module Cubical.Relation.ZigZag.Base where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.Relation.Binary.Base + +open BinaryRelation +open isEquivRel + +private + variable + ℓ ℓ' : Level + +isZigZagComplete : {A B : Type ℓ} (R : A → B → Type ℓ') → Type (ℓ-max ℓ ℓ') +isZigZagComplete R = ∀ {a b a' b'} → R a b → R a' b → R a' b' → R a b' + +ZigZagRel : (A B : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +ZigZagRel A B ℓ' = Σ[ R ∈ (A → B → Type ℓ') ] (isZigZagComplete R) + +record isQuasiEquivRel {A B : Type ℓ} (R : A → B → Type ℓ') : Type (ℓ-max ℓ ℓ') where + field + zigzag : isZigZagComplete R + fwd : (a : A) → ∃[ b ∈ B ] R a b + bwd : (b : B) → ∃[ a ∈ A ] R a b + +open isQuasiEquivRel + +QuasiEquivRel : (A B : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +QuasiEquivRel A B ℓ' = + Σ[ R ∈ PropRel A B ℓ' ] isQuasiEquivRel (R .fst) + +invQER : {A B : Type ℓ} {ℓ' : Level} → QuasiEquivRel A B ℓ' → QuasiEquivRel B A ℓ' +invQER (R , qer) .fst = invPropRel R +invQER (R , qer) .snd .zigzag aRb aRb' a'Rb' = qer .zigzag a'Rb' aRb' aRb +invQER (R , qer) .snd .fwd = qer .bwd +invQER (R , qer) .snd .bwd = qer .fwd + +QER→EquivRel : {A B : Type ℓ} + → QuasiEquivRel A B ℓ' → EquivPropRel A (ℓ-max ℓ ℓ') +QER→EquivRel (R , sim) .fst = compPropRel R (invPropRel R) +QER→EquivRel (R , sim) .snd .reflexive a = Trunc.map (λ {(b , r) → b , r , r}) (sim .fwd a) +QER→EquivRel (R , sim) .snd .symmetric _ _ = Trunc.map (λ {(b , r₀ , r₁) → b , r₁ , r₀}) +QER→EquivRel (R , sim) .snd .transitive _ _ _ = + Trunc.map2 (λ {(b , r₀ , r₁) (b' , r₀' , r₁') → b , r₀ , sim .zigzag r₁' r₀' r₁}) + +-- The following result is due to Carlo Angiuli +module QER→Equiv {A B : Type ℓ} (R : QuasiEquivRel A B ℓ') where + + Rᴸ = QER→EquivRel R .fst .fst + Rᴿ = QER→EquivRel (invQER R) .fst .fst + + private + sim = R .snd + + private + f : (a : A) → ∃[ b ∈ B ] R .fst .fst a b → B / Rᴿ + f a = + Trunc.rec→Set squash/ + ([_] ∘ fst) + (λ {(b , r) (b' , r') → eq/ b b' ∣ a , r , r' ∣}) + + fPath : + (a₀ : A) (s₀ : ∃[ b ∈ B ] R .fst .fst a₀ b) + (a₁ : A) (s₁ : ∃[ b ∈ B ] R .fst .fst a₁ b) + → Rᴸ a₀ a₁ + → f a₀ s₀ ≡ f a₁ s₁ + fPath a₀ = + Trunc.elim (λ _ → isPropΠ3 λ _ _ _ → squash/ _ _) + (λ {(b₀ , r₀) a₁ → + Trunc.elim (λ _ → isPropΠ λ _ → squash/ _ _) + (λ {(b₁ , r₁) → + Trunc.elim (λ _ → squash/ _ _) + (λ {(b' , r₀' , r₁') → eq/ b₀ b₁ ∣ a₀ , r₀ , sim .zigzag r₀' r₁' r₁ ∣})})}) + + φ : A / Rᴸ → B / Rᴿ + φ [ a ] = f a (sim .fwd a) + φ (eq/ a₀ a₁ r i) = fPath a₀ (sim .fwd a₀) a₁ (sim .fwd a₁) r i + φ (squash/ _ _ p q j i) = squash/ _ _ (cong φ p) (cong φ q) j i + + + relToFwd≡ : ∀ {a b} → R .fst .fst a b → φ [ a ] ≡ [ b ] + relToFwd≡ {a} {b} r = + Trunc.elim {P = λ s → f a s ≡ [ b ]} (λ _ → squash/ _ _) + (λ {(b' , r') → eq/ b' b ∣ a , r' , r ∣}) + (sim .fwd a) + + fwd≡ToRel : ∀ {a b} → φ [ a ] ≡ [ b ] → R .fst .fst a b + fwd≡ToRel {a} {b} = + Trunc.elim {P = λ s → f a s ≡ [ b ] → R .fst .fst a b} + (λ _ → isPropΠ λ _ → R .fst .snd _ _) + (λ {(b' , r') p → + Trunc.rec (R .fst .snd _ _) + (λ {(a' , s' , s) → R .snd .zigzag r' s' s}) + (effective + (QER→EquivRel (invQER R) .fst .snd) + (QER→EquivRel (invQER R) .snd) + b' b p)}) + (sim .fwd a) + + private + g : (b : B) → ∃[ a ∈ A ] R .fst .fst a b → A / Rᴸ + g b = + Trunc.rec→Set squash/ + ([_] ∘ fst) + (λ {(a , r) (a' , r') → eq/ a a' ∣ b , r , r' ∣}) + + gPath : + (b₀ : B) (s₀ : ∃[ a ∈ A ] R .fst .fst a b₀) + (b₁ : B) (s₁ : ∃[ a ∈ A ] R .fst .fst a b₁) + → Rᴿ b₀ b₁ + → g b₀ s₀ ≡ g b₁ s₁ + gPath b₀ = + Trunc.elim (λ _ → isPropΠ3 λ _ _ _ → squash/ _ _) + (λ {(a₀ , r₀) b₁ → + Trunc.elim (λ _ → isPropΠ λ _ → squash/ _ _) + (λ {(a₁ , r₁) → + Trunc.elim (λ _ → squash/ _ _) + (λ {(a' , r₀' , r₁') → eq/ a₀ a₁ ∣ b₀ , r₀ , sim .zigzag r₁ r₁' r₀' ∣})})}) + + ψ : B / Rᴿ → A / Rᴸ + ψ [ b ] = g b (sim .bwd b) + ψ (eq/ b₀ b₁ r i) = gPath b₀ (sim .bwd b₀) b₁ (sim .bwd b₁) r i + ψ (squash/ _ _ p q j i) = squash/ _ _ (cong ψ p) (cong ψ q) j i + + relToBwd≡ : ∀ {a b} → R .fst .fst a b → ψ [ b ] ≡ [ a ] + relToBwd≡ {a} {b} r = + Trunc.elim {P = λ s → g b s ≡ [ a ]} (λ _ → squash/ _ _) + (λ {(a' , r') → eq/ a' a ∣ b , r' , r ∣}) + (sim .bwd b) + + private + η : ∀ qb → φ (ψ qb) ≡ qb + η = + elimProp (λ _ → squash/ _ _) + (λ b → + Trunc.elim {P = λ s → φ (g b s) ≡ [ b ]} (λ _ → squash/ _ _) + (λ {(a , r) → relToFwd≡ r}) + (sim .bwd b)) + + ε : ∀ qa → ψ (φ qa) ≡ qa + ε = + elimProp (λ _ → squash/ _ _) + (λ a → + Trunc.elim {P = λ s → ψ (f a s) ≡ [ a ]} (λ _ → squash/ _ _) + (λ {(b , r) → relToBwd≡ r}) + (sim .fwd a)) + + bwd≡ToRel : ∀ {a b} → ψ [ b ] ≡ [ a ] → R .fst .fst a b + bwd≡ToRel {a} {b} p = fwd≡ToRel (cong φ (sym p) ∙ η [ b ]) + + Thm : (A / Rᴸ) ≃ (B / Rᴿ) + Thm = isoToEquiv (iso φ ψ η ε) diff --git a/Cubical/Structures/Auto.agda b/Cubical/Structures/Auto.agda new file mode 100644 index 0000000000..26145e992e --- /dev/null +++ b/Cubical/Structures/Auto.agda @@ -0,0 +1,250 @@ +{- + +Macros (autoDesc, AutoStructure, AutoEquivStr, autoUnivalentStr) for automatically generating structure definitions. + +For example: + + autoDesc (λ (X : Type₀) → X → X × ℕ) ↦ function+ var (var , constant ℕ) + +We prefer to use the constant structure whenever possible, e.g., [autoDesc (λ (X : Type₀) → ℕ → ℕ)] +is [constant (ℕ → ℕ)] rather than [function (constant ℕ) (constant ℕ)]. + +Writing [auto* (λ X → ⋯)] doesn't seem to work, but [auto* (λ (X : Type ℓ) → ⋯)] does. + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Auto where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.List +open import Cubical.Data.Bool +open import Cubical.Data.Maybe +open import Cubical.Structures.Macro as Macro + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +-- Magic number +private + FUEL = 10000 + +-- Mark part of a structure definition to use transport-based structured equivalences +abstract + Transp[_] : ∀ {ℓ} → Type ℓ → Type ℓ + Transp[ A ] = A + +-- Some reflection utilities +private + tLevel = R.def (quote Level) [] + + tType : R.Term → R.Term + tType ℓ = R.def (quote Type) [ varg ℓ ] + + tTranspDesc : R.Term → R.Term → R.Term + tTranspDesc ℓ ℓ' = R.def (quote TranspDesc) (ℓ v∷ ℓ' v∷ []) + + tDesc : R.Term → R.Term → R.Term → R.Term + tDesc ℓ ℓ₁ ℓ₁' = R.def (quote Desc) (ℓ v∷ ℓ₁ v∷ ℓ₁' v∷ []) + + func : (ℓ ℓ' : Level) → Type (ℓ-suc (ℓ-max ℓ ℓ')) + func ℓ ℓ' = Type ℓ → Type ℓ' + + tStruct : R.Term → R.Term → R.Term + tStruct ℓ ℓ' = R.def (quote func) (varg ℓ ∷ varg ℓ' ∷ []) + +-- We try to build a descriptor by unifying the provided structure with combinators we're aware of. We +-- redefine the structure combinators as the *Shape terms below so that we don't depend on the specific way +-- these are defined in other files (order of implicit arguments and so on); the syntactic analysis that goes +-- on here means that we would get mysterious errors if those changed. +private + constantShape : ∀ {ℓ'} (ℓ : Level) (A : Type ℓ') → (Type ℓ → Type ℓ') + constantShape _ A _ = A + + pointedShape : (ℓ : Level) → Type ℓ → Type ℓ + pointedShape _ X = X + + productShape : ∀ {ℓ₀ ℓ₁} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → (Type ℓ → Type ℓ₁) → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + productShape _ A₀ A₁ X = A₀ X × A₁ X + + functionShape : ∀ {ℓ₀ ℓ₁} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → (Type ℓ → Type ℓ₁) → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + functionShape _ A₀ A₁ X = A₀ X → A₁ X + + maybeShape : ∀ {ℓ₀} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → Type ℓ → Type ℓ₀ + maybeShape _ A₀ X = Maybe (A₀ X) + + transpShape : ∀ {ℓ₀} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → Type ℓ → Type ℓ₀ + transpShape _ A₀ X = Transp[ A₀ X ] + +private + -- Build transport structure descriptor from a function [t : Type ℓ → Type ℓ'] + buildTranspDesc : ℕ → R.Term → R.Term → R.Term → R.TC R.Term + buildTranspDesc zero ℓ ℓ' t = R.typeError (R.strErr "Ran out of fuel! at \n" ∷ R.termErr t ∷ []) + buildTranspDesc (suc fuel) ℓ ℓ' t = + tryConstant t <|> tryPointed t <|> tryProduct t <|> tryFunction t <|> tryMaybe t <|> + R.typeError (R.strErr "Can't automatically generate a transp structure for\n" ∷ R.termErr t ∷ []) + where + tryConstant : R.Term → R.TC R.Term + tryConstant t = + newMeta (tType ℓ') >>= λ A → + R.unify t (R.def (quote constantShape) (ℓ v∷ A v∷ [])) >> + R.returnTC (R.con (quote TranspDesc.constant) (A v∷ [])) + + tryPointed : R.Term → R.TC R.Term + tryPointed t = + R.unify t (R.def (quote pointedShape) (ℓ v∷ [])) >> + R.returnTC (R.con (quote TranspDesc.var) []) + + tryFunction : R.Term → R.TC R.Term + tryFunction t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote functionShape) (ℓ v∷ A₀ v∷ A₁ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildTranspDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote TranspDesc.function) (d₀ v∷ d₁ v∷ [])) + + tryProduct : R.Term → R.TC R.Term + tryProduct t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote productShape) (ℓ v∷ A₀ v∷ A₁ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildTranspDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote TranspDesc._,_) (d₀ v∷ d₁ v∷ [])) + + tryMaybe : R.Term → R.TC R.Term + tryMaybe t = + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote maybeShape) (ℓ v∷ A₀ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote TranspDesc.maybe) (d₀ v∷ [])) + + autoTranspDesc' : R.Term → R.Term → R.TC Unit + autoTranspDesc' t hole = + R.inferType hole >>= λ H → + newMeta tLevel >>= λ ℓ → + newMeta tLevel >>= λ ℓ' → + R.unify (tTranspDesc ℓ ℓ') H >> + R.checkType t (tStruct ℓ ℓ') >> + buildTranspDesc FUEL ℓ ℓ' t >>= R.unify hole + +-- Build structure descriptor from a function [t : Type ℓ → Type ℓ'] +buildDesc : ℕ → R.Term → R.Term → R.Term → R.TC R.Term +buildDesc zero ℓ ℓ' t = R.typeError (R.strErr "Ran out of fuel! at \n" ∷ R.termErr t ∷ []) +buildDesc (suc fuel) ℓ ℓ' t = + tryConstant t <|> tryPointed t <|> tryProduct t <|> tryFunction t <|> + tryMaybe t <|> tryTransp t <|> + R.typeError (R.strErr "Can't automatically generate a structure for\n" ∷ R.termErr t ∷ []) + where + tryConstant : R.Term → R.TC R.Term + tryConstant t = + newMeta (tType ℓ') >>= λ A → + R.unify t (R.def (quote constantShape) (ℓ v∷ A v∷ [])) >> + R.returnTC (R.con (quote Desc.constant) (A v∷ [])) + + tryPointed : R.Term → R.TC R.Term + tryPointed t = + R.unify t (R.def (quote pointedShape) (ℓ v∷ [])) >> + R.returnTC (R.con (quote Desc.var) []) + + tryProduct : R.Term → R.TC R.Term + tryProduct t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote productShape) (ℓ v∷ A₀ v∷ A₁ v∷ [])) >> + buildDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote Desc._,_) (d₀ v∷ d₁ v∷ [])) + + tryFunction : R.Term → R.TC R.Term + tryFunction t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote functionShape) (ℓ v∷ A₀ v∷ A₁ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote Desc.function+) (d₀ v∷ d₁ v∷ [])) + + tryMaybe : R.Term → R.TC R.Term + tryMaybe t = + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote maybeShape) (ℓ v∷ A₀ v∷ [])) >> + buildDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote Desc.maybe) (d₀ v∷ [])) + + tryTransp : R.Term → R.TC R.Term + tryTransp t = + newMeta (tStruct ℓ ℓ') >>= λ A₀ → + R.unify t (R.def (quote transpShape) (ℓ v∷ A₀ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ' A₀ >>= λ d₀ → + R.returnTC (R.con (quote Desc.transpDesc) (d₀ v∷ [])) + +autoDesc' : R.Term → R.Term → R.TC Unit +autoDesc' t hole = + R.inferType hole >>= λ H → + newMeta tLevel >>= λ ℓ → + newMeta tLevel >>= λ ℓ' → + R.unify (tDesc ℓ ℓ' R.unknown) H >> + R.checkType t (tStruct ℓ ℓ') >> + buildDesc FUEL ℓ ℓ' t >>= R.unify hole + +macro + -- (Type ℓ → Type ℓ₁) → TranspDesc ℓ + autoTranspDesc : R.Term → R.Term → R.TC Unit + autoTranspDesc = autoTranspDesc' + + -- (S : Type ℓ → Type ℓ₁) → EquivAction (AutoStructure S) + autoEquivAction : R.Term → R.Term → R.TC Unit + autoEquivAction t hole = + newMeta (tTranspDesc R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote transpMacroAction) [ varg d ]) >> + autoTranspDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → TransportStr (autoEquivAction S) + autoTransportStr : R.Term → R.Term → R.TC Unit + autoTransportStr t hole = + newMeta (tTranspDesc R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote transpMacroTransportStr) [ varg d ]) >> + autoTranspDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → Desc ℓ + autoDesc : R.Term → R.Term → R.TC Unit + autoDesc = autoDesc' + + -- (S : Type ℓ → Type ℓ₁) → (Type ℓ → Type ℓ₁) + -- Removes Transp[_] annotations + AutoStructure : R.Term → R.Term → R.TC Unit + AutoStructure t hole = + newMeta (tDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote MacroStructure) [ varg d ]) >> + autoDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → StrEquiv (AutoStructure S) _ + AutoEquivStr : R.Term → R.Term → R.TC Unit + AutoEquivStr t hole = + newMeta (tDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote MacroEquivStr) [ varg d ]) >> + autoDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → UnivalentStr (AutoStructure S) (AutoEquivStr S) + autoUnivalentStr : R.Term → R.Term → R.TC Unit + autoUnivalentStr t hole = + newMeta (tDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote MacroUnivalentStr) [ varg d ]) >> + autoDesc' t d diff --git a/Cubical/Structures/Axioms.agda b/Cubical/Structures/Axioms.agda new file mode 100644 index 0000000000..e3d2a225d3 --- /dev/null +++ b/Cubical/Structures/Axioms.agda @@ -0,0 +1,69 @@ +{- + +Add axioms (i.e., propositions) to a structure S without changing the definition of structured equivalence. + +X ↦ Σ[ s ∈ S X ] (P X s) where (P X s) is a proposition for all X and s. + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Axioms where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path +open import Cubical.Foundations.SIP +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₂ : Level + +AxiomsStructure : (S : Type ℓ → Type ℓ₁) + (axioms : (X : Type ℓ) → S X → Type ℓ₂) + → Type ℓ → Type (ℓ-max ℓ₁ ℓ₂) +AxiomsStructure S axioms X = Σ[ s ∈ S X ] (axioms X s) + +AxiomsEquivStr : {S : Type ℓ → Type ℓ₁} (ι : StrEquiv S ℓ₁') + (axioms : (X : Type ℓ) → S X → Type ℓ₂) + → StrEquiv (AxiomsStructure S axioms) ℓ₁' +AxiomsEquivStr ι axioms (X , (s , a)) (Y , (t , b)) e = ι (X , s) (Y , t) e + +axiomsUnivalentStr : {S : Type ℓ → Type ℓ₁} + (ι : (A B : TypeWithStr ℓ S) → A .fst ≃ B .fst → Type ℓ₁') + {axioms : (X : Type ℓ) → S X → Type ℓ₂} + (axioms-are-Props : (X : Type ℓ) (s : S X) → isProp (axioms X s)) + (θ : UnivalentStr S ι) + → UnivalentStr (AxiomsStructure S axioms) (AxiomsEquivStr ι axioms) +axiomsUnivalentStr {S = S} ι {axioms = axioms} axioms-are-Props θ {X , s , a} {Y , t , b} e = + ι (X , s) (Y , t) e + ≃⟨ θ e ⟩ + PathP (λ i → S (ua e i)) s t + ≃⟨ invEquiv (Σ-contractSnd λ _ → isOfHLevelPathP' 0 (axioms-are-Props _ _) _ _) ⟩ + Σ[ p ∈ PathP (λ i → S (ua e i)) s t ] PathP (λ i → axioms (ua e i) (p i)) a b + ≃⟨ ΣPath≃PathΣ ⟩ + PathP (λ i → AxiomsStructure S axioms (ua e i)) (s , a) (t , b) + ■ + +inducedStructure : {S : Type ℓ → Type ℓ₁} + {ι : (A B : TypeWithStr ℓ S) → A .fst ≃ B .fst → Type ℓ₁'} + (θ : UnivalentStr S ι) + {axioms : (X : Type ℓ) → S X → Type ℓ₂} + (A : TypeWithStr ℓ (AxiomsStructure S axioms)) (B : TypeWithStr ℓ S) + → (typ A , str A .fst) ≃[ ι ] B + → TypeWithStr ℓ (AxiomsStructure S axioms) +inducedStructure θ {axioms} A B eqv = + B .fst , B .snd , subst (uncurry axioms) (sip θ _ _ eqv) (A .snd .snd) + +transferAxioms : {S : Type ℓ → Type ℓ₁} + {ι : (A B : TypeWithStr ℓ S) → A .fst ≃ B .fst → Type ℓ₁'} + (θ : UnivalentStr S ι) + {axioms : (X : Type ℓ) → S X → Type ℓ₂} + (A : TypeWithStr ℓ (AxiomsStructure S axioms)) (B : TypeWithStr ℓ S) + → (typ A , str A .fst) ≃[ ι ] B + → axioms (fst B) (snd B) +transferAxioms θ {axioms} A B eqv = + subst (uncurry axioms) (sip θ _ _ eqv) (A .snd .snd) diff --git a/Cubical/Structures/Constant.agda b/Cubical/Structures/Constant.agda new file mode 100644 index 0000000000..ec605bba17 --- /dev/null +++ b/Cubical/Structures/Constant.agda @@ -0,0 +1,35 @@ +{- + +Constant structure: _ ↦ A + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Constant where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Foundations.SIP + +private + variable + ℓ ℓ' : Level + +-- Structured isomorphisms + +module _ (A : Type ℓ') where + + ConstantStructure : Type ℓ → Type ℓ' + ConstantStructure _ = A + + ConstantEquivStr : StrEquiv {ℓ} ConstantStructure ℓ' + ConstantEquivStr (_ , a) (_ , a') _ = a ≡ a' + + constantUnivalentStr : UnivalentStr {ℓ} ConstantStructure ConstantEquivStr + constantUnivalentStr e = idEquiv _ + + constantEquivAction : EquivAction {ℓ} ConstantStructure + constantEquivAction e = idEquiv _ + + constantTransportStr : TransportStr {ℓ} constantEquivAction + constantTransportStr e _ = sym (transportRefl _) diff --git a/Cubical/Structures/Function.agda b/Cubical/Structures/Function.agda new file mode 100644 index 0000000000..4371e4c016 --- /dev/null +++ b/Cubical/Structures/Function.agda @@ -0,0 +1,82 @@ +{- + +Functions between structures S and T: X ↦ S X → T X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Function where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence + +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Data.Nat +open import Cubical.Data.Vec + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₂ ℓ₂' : Level + +-- General function structure + +FunctionStructure : (S : Type ℓ → Type ℓ₁) (T : Type ℓ → Type ℓ₂) + → Type ℓ → Type (ℓ-max ℓ₁ ℓ₂) +FunctionStructure S T X = S X → T X + +FunctionEquivStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + → StrEquiv S ℓ₁' → StrEquiv T ℓ₂' + → StrEquiv (FunctionStructure S T) (ℓ-max ℓ₁ (ℓ-max ℓ₁' ℓ₂')) +FunctionEquivStr {S = S} {T} ι₁ ι₂ (X , f) (Y , g) e = + {s : S X} {t : S Y} → ι₁ (X , s) (Y , t) e → ι₂ (X , f s) (Y , g t) e + +functionUnivalentStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (ι₁ : StrEquiv S ℓ₁') (θ₁ : UnivalentStr S ι₁) + (ι₂ : StrEquiv T ℓ₂') (θ₂ : UnivalentStr T ι₂) + → UnivalentStr (FunctionStructure S T) (FunctionEquivStr ι₁ ι₂) +functionUnivalentStr ι₁ θ₁ ι₂ θ₂ e = + compEquiv + (equivImplicitΠCod (equivImplicitΠCod (equiv→ (θ₁ e) (θ₂ e)))) + funExtDepEquiv + +functionEquivAction : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + → EquivAction S → EquivAction T + → EquivAction (FunctionStructure S T) +functionEquivAction α₁ α₂ e = equiv→ (α₁ e) (α₂ e) + +functionTransportStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (α₁ : EquivAction S) (τ₁ : TransportStr α₁) + (α₂ : EquivAction T) (τ₂ : TransportStr α₂) + → TransportStr (functionEquivAction α₁ α₂) +functionTransportStr {S = S} α₁ τ₁ α₂ τ₂ e f = + funExt λ t → + cong (equivFun (α₂ e) ∘ f) (invTransportStr α₁ τ₁ e t) + ∙ τ₂ e (f (subst⁻ S (ua e) t)) + +-- Definition of structured equivalence using an action in the domain + +FunctionEquivStr+ : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + → EquivAction S → StrEquiv T ℓ₂' + → StrEquiv (FunctionStructure S T) (ℓ-max ℓ₁ ℓ₂') +FunctionEquivStr+ {S = S} {T} α₁ ι₂ (X , f) (Y , g) e = + (s : S X) → ι₂ (X , f s) (Y , g (equivFun (α₁ e) s)) e + +functionUnivalentStr+ : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (α₁ : EquivAction S) (τ₁ : TransportStr α₁) + (ι₂ : StrEquiv T ℓ₂') (θ₂ : UnivalentStr T ι₂) + → UnivalentStr (FunctionStructure S T) (FunctionEquivStr+ α₁ ι₂) +functionUnivalentStr+ {S = S} {T} α₁ τ₁ ι₂ θ₂ {X , f} {Y , g} e = + compEquiv + (compEquiv + (equivΠCod λ s → + compEquiv + (θ₂ e) + (pathToEquiv (cong (PathP (λ i → T (ua e i)) (f s) ∘ g) (τ₁ e s)))) + (invEquiv heteroHomotopy≃Homotopy)) + funExtDepEquiv diff --git a/Cubical/Structures/LeftAction.agda b/Cubical/Structures/LeftAction.agda new file mode 100644 index 0000000000..59cf68aa8c --- /dev/null +++ b/Cubical/Structures/LeftAction.agda @@ -0,0 +1,17 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.LeftAction where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.SIP + +open import Cubical.Structures.Auto + +module _ {ℓ ℓ' : Level} (A : Type ℓ') where + + LeftActionStructure : Type ℓ → Type (ℓ-max ℓ ℓ') + LeftActionStructure X = A → X → X + + LeftActionEquivStr = AutoEquivStr LeftActionStructure + + leftActionUnivalentStr : UnivalentStr _ LeftActionEquivStr + leftActionUnivalentStr = autoUnivalentStr LeftActionStructure diff --git a/Cubical/Structures/Macro.agda b/Cubical/Structures/Macro.agda new file mode 100644 index 0000000000..d80e785a70 --- /dev/null +++ b/Cubical/Structures/Macro.agda @@ -0,0 +1,157 @@ +{- + +Descriptor language for easily defining structures + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Macro where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.SIP +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.Data.Maybe + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Constant +open import Cubical.Structures.Function +open import Cubical.Structures.Maybe +open import Cubical.Structures.Parameterized +open import Cubical.Structures.Pointed +open import Cubical.Structures.Product + +{- Transport structures -} + +data TranspDesc (ℓ : Level) : Level → Typeω where + -- constant structure: X ↦ A + constant : ∀ {ℓ₁} (A : Type ℓ₁) → TranspDesc ℓ ℓ₁ + -- pointed structure: X ↦ X + var : TranspDesc ℓ ℓ + -- product of structures S,T : X ↦ (S X × T X) + _,_ : ∀ {ℓ₁ ℓ₂} (d₀ : TranspDesc ℓ ℓ₁) (d₁ : TranspDesc ℓ ℓ₂) → TranspDesc ℓ (ℓ-max ℓ₁ ℓ₂) + -- functions between structures S,T: X ↦ (S X → T X) + function : ∀ {ℓ₁ ℓ₂} (d₀ : TranspDesc ℓ ℓ₁) (d₁ : TranspDesc ℓ ℓ₂) → TranspDesc ℓ (ℓ-max ℓ₁ ℓ₂) + -- Maybe on a structure S: X ↦ Maybe (S X) + maybe : ∀ {ℓ₁} → TranspDesc ℓ ℓ₁ → TranspDesc ℓ ℓ₁ + -- arbitrary transport structure + foreign : ∀ {ℓ₁} {S : Type ℓ → Type ℓ₁} (α : EquivAction S) → TransportStr α → TranspDesc ℓ ℓ₁ + +-- Structure defined by a transport descriptor +TranspMacroStructure : ∀ {ℓ ℓ₁} → TranspDesc ℓ ℓ₁ → Type ℓ → Type ℓ₁ +TranspMacroStructure (constant A) X = A +TranspMacroStructure var X = X +TranspMacroStructure (d₀ , d₁) X = TranspMacroStructure d₀ X × TranspMacroStructure d₁ X +TranspMacroStructure (function d₀ d₁) X = TranspMacroStructure d₀ X → TranspMacroStructure d₁ X +TranspMacroStructure (maybe d) = MaybeStructure (TranspMacroStructure d) +TranspMacroStructure (foreign {S = S} α τ) = S + +-- Action defined by a transport descriptor +transpMacroAction : ∀ {ℓ ℓ₁} (d : TranspDesc ℓ ℓ₁) → EquivAction (TranspMacroStructure d) +transpMacroAction (constant A) = constantEquivAction A +transpMacroAction var = pointedEquivAction +transpMacroAction (d₀ , d₁) = productEquivAction (transpMacroAction d₀) (transpMacroAction d₁) +transpMacroAction (function d₀ d₁) = + functionEquivAction (transpMacroAction d₀) (transpMacroAction d₁) +transpMacroAction (maybe d) = maybeEquivAction (transpMacroAction d) +transpMacroAction (foreign α _) = α + +-- Action defines a transport structure +transpMacroTransportStr : ∀ {ℓ ℓ₁} (d : TranspDesc ℓ ℓ₁) → TransportStr (transpMacroAction d) +transpMacroTransportStr (constant A) = constantTransportStr A +transpMacroTransportStr var = pointedTransportStr +transpMacroTransportStr (d₀ , d₁) = + productTransportStr + (transpMacroAction d₀) (transpMacroTransportStr d₀) + (transpMacroAction d₁) (transpMacroTransportStr d₁) +transpMacroTransportStr (function d₀ d₁) = + functionTransportStr + (transpMacroAction d₀) (transpMacroTransportStr d₀) + (transpMacroAction d₁) (transpMacroTransportStr d₁) +transpMacroTransportStr (maybe d) = + maybeTransportStr (transpMacroAction d) (transpMacroTransportStr d) +transpMacroTransportStr (foreign α τ) = τ + +{- General structures -} + +mutual + data Desc (ℓ : Level) : Level → Level → Typeω where + -- constant structure: X ↦ A + constant : ∀ {ℓ₁} (A : Type ℓ₁) → Desc ℓ ℓ₁ ℓ₁ + -- pointed structure: X ↦ X + var : Desc ℓ ℓ ℓ + -- product of structures S,T : X ↦ (S X × T X) + _,_ : ∀ {ℓ₁ ℓ₁' ℓ₂ ℓ₂'} (d₀ : Desc ℓ ℓ₁ ℓ₁') (d₁ : Desc ℓ ℓ₂ ℓ₂') + → Desc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁' ℓ₂') + -- functions between structures S,T : X ↦ (S X → T X) + function : ∀ {ℓ₁ ℓ₁' ℓ₂ ℓ₂'} (d₀ : Desc ℓ ℓ₁ ℓ₁') (d₁ : Desc ℓ ℓ₂ ℓ₂') + → Desc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁ (ℓ-max ℓ₁' ℓ₂')) + -- functions between structures S,T where S is functorial : X ↦ (S X → T X) + function+ : ∀ {ℓ₁ ℓ₂ ℓ₂'} (d₀ : TranspDesc ℓ ℓ₁) (d₁ : Desc ℓ ℓ₂ ℓ₂') → Desc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁ ℓ₂') + -- Maybe on a structure S: X ↦ Maybe (S X) + maybe : ∀ {ℓ₁ ℓ₁'} → Desc ℓ ℓ₁ ℓ₁' → Desc ℓ ℓ₁ ℓ₁' + -- add axioms to a structure + axioms : ∀ {ℓ₁ ℓ₁' ℓ₂} (d : Desc ℓ ℓ₁ ℓ₁') (ax : ∀ X → MacroStructure d X → Type ℓ₂) + → (∀ X s → isProp (ax X s)) + → Desc ℓ (ℓ-max ℓ₁ ℓ₂) ℓ₁' + -- univalent structure from transport structure + transpDesc : ∀ {ℓ₁} → TranspDesc ℓ ℓ₁ → Desc ℓ ℓ₁ ℓ₁ + -- arbitrary univalent notion of structure + foreign : ∀ {ℓ₁ ℓ₁'} {S : Type ℓ → Type ℓ₁} (ι : StrEquiv S ℓ₁') → UnivalentStr S ι → Desc ℓ ℓ₁ ℓ₁' + + infixr 4 _,_ + + -- Structure defined by a descriptor + MacroStructure : ∀ {ℓ ℓ₁ ℓ₁'} → Desc ℓ ℓ₁ ℓ₁' → Type ℓ → Type ℓ₁ + MacroStructure (constant A) X = A + MacroStructure var X = X + MacroStructure (d₀ , d₁) X = MacroStructure d₀ X × MacroStructure d₁ X + MacroStructure (function+ d₀ d₁) X = TranspMacroStructure d₀ X → MacroStructure d₁ X + MacroStructure (function d₀ d₁) X = MacroStructure d₀ X → MacroStructure d₁ X + MacroStructure (maybe d) = MaybeStructure (MacroStructure d) + MacroStructure (axioms d ax _) = AxiomsStructure (MacroStructure d) ax + MacroStructure (transpDesc d) = TranspMacroStructure d + MacroStructure (foreign {S = S} _ _) = S + +-- Notion of structured equivalence defined by a descriptor +MacroEquivStr : ∀ {ℓ ℓ₁ ℓ₁'} → (d : Desc ℓ ℓ₁ ℓ₁') → StrEquiv (MacroStructure d) ℓ₁' +MacroEquivStr (constant A) = ConstantEquivStr A +MacroEquivStr var = PointedEquivStr +MacroEquivStr (d₀ , d₁) = ProductEquivStr (MacroEquivStr d₀) (MacroEquivStr d₁) +MacroEquivStr (function+ d₀ d₁) = FunctionEquivStr+ (transpMacroAction d₀) (MacroEquivStr d₁) +MacroEquivStr (function d₀ d₁) = FunctionEquivStr (MacroEquivStr d₀) (MacroEquivStr d₁) +MacroEquivStr (maybe d) = MaybeEquivStr (MacroEquivStr d) +MacroEquivStr (axioms d ax _) = AxiomsEquivStr (MacroEquivStr d) ax +MacroEquivStr (transpDesc d) = EquivAction→StrEquiv (transpMacroAction d) +MacroEquivStr (foreign ι _) = ι + +-- Proof that structure induced by descriptor is univalent +MacroUnivalentStr : ∀ {ℓ ℓ₁ ℓ₁'} → (d : Desc ℓ ℓ₁ ℓ₁') → UnivalentStr (MacroStructure d) (MacroEquivStr d) +MacroUnivalentStr (constant A) = constantUnivalentStr A +MacroUnivalentStr var = pointedUnivalentStr +MacroUnivalentStr (d₀ , d₁) = + productUnivalentStr + (MacroEquivStr d₀) (MacroUnivalentStr d₀) + (MacroEquivStr d₁) (MacroUnivalentStr d₁) +MacroUnivalentStr (function+ d₀ d₁) = + functionUnivalentStr+ + (transpMacroAction d₀) (transpMacroTransportStr d₀) + (MacroEquivStr d₁) (MacroUnivalentStr d₁) +MacroUnivalentStr (function d₀ d₁) = + functionUnivalentStr + (MacroEquivStr d₀) (MacroUnivalentStr d₀) + (MacroEquivStr d₁) (MacroUnivalentStr d₁) +MacroUnivalentStr (maybe d) = maybeUnivalentStr (MacroEquivStr d) (MacroUnivalentStr d) +MacroUnivalentStr (axioms d _ isPropAx) = axiomsUnivalentStr (MacroEquivStr d) isPropAx (MacroUnivalentStr d) +MacroUnivalentStr (transpDesc d) = + TransportStr→UnivalentStr (transpMacroAction d) (transpMacroTransportStr d) +MacroUnivalentStr (foreign _ θ) = θ + +-- Module for easy importing +module Macro ℓ {ℓ₁ ℓ₁'} (d : Desc ℓ ℓ₁ ℓ₁') where + + structure = MacroStructure d + equiv = MacroEquivStr d + univalent = MacroUnivalentStr d diff --git a/Cubical/Structures/Maybe.agda b/Cubical/Structures/Maybe.agda new file mode 100644 index 0000000000..7198a3a124 --- /dev/null +++ b/Cubical/Structures/Maybe.agda @@ -0,0 +1,108 @@ +{- + + Maybe structure: X ↦ Maybe (S X) + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Maybe where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.SIP +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Maybe + +private + variable + ℓ ℓ₁ ℓ₁' : Level + +MaybeRel : {A B : Type ℓ} (R : A → B → Type ℓ₁) → Maybe A → Maybe B → Type ℓ₁ +MaybeRel R nothing nothing = Lift Unit +MaybeRel R nothing (just _) = Lift ⊥ +MaybeRel R (just _) nothing = Lift ⊥ +MaybeRel R (just x) (just y) = R x y + +congMaybeRel : {A B : Type ℓ} {R : A → B → Type ℓ₁} {S : A → B → Type ℓ₁'} + → (∀ x y → R x y ≃ S x y) + → ∀ ox oy → MaybeRel R ox oy ≃ MaybeRel S ox oy +congMaybeRel e nothing nothing = Lift≃Lift (idEquiv _) +congMaybeRel e nothing (just _) = Lift≃Lift (idEquiv _) +congMaybeRel e (just _) nothing = Lift≃Lift (idEquiv _) +congMaybeRel e (just x) (just y) = e x y + +module MaybePathP where + + Code : (A : I → Type ℓ) → Maybe (A i0) → Maybe (A i1) → Type ℓ + Code A = MaybeRel (PathP A) + + encodeRefl : {A : Type ℓ} → ∀ ox → Code (λ _ → A) ox ox + encodeRefl nothing = lift tt + encodeRefl (just _) = refl + + encode : (A : I → Type ℓ) → ∀ ox oy → PathP (λ i → Maybe (A i)) ox oy → Code A ox oy + encode A ox oy p = transport (λ j → Code (λ i → A (i ∧ j)) ox (p j)) (encodeRefl ox) + + decode : {A : I → Type ℓ} → ∀ ox oy → Code A ox oy → PathP (λ i → Maybe (A i)) ox oy + decode nothing nothing p i = nothing + decode (just _) (just _) p i = just (p i) + + decodeEncodeRefl : {A : Type ℓ} (ox : Maybe A) → decode ox ox (encodeRefl ox) ≡ refl + decodeEncodeRefl nothing = refl + decodeEncodeRefl (just _) = refl + + decodeEncode : {A : I → Type ℓ} → ∀ ox oy p → decode ox oy (encode A ox oy p) ≡ p + decodeEncode {A = A} ox oy p = + transport + (λ k → + decode _ _ (transp (λ j → Code (λ i → A (i ∧ j ∧ k)) ox (p (j ∧ k))) (~ k) (encodeRefl ox)) + ≡ (λ i → p (i ∧ k))) + (decodeEncodeRefl ox) + + encodeDecode : (A : I → Type ℓ) → ∀ ox oy c → encode A ox oy (decode ox oy c) ≡ c + encodeDecode A nothing nothing c = refl + encodeDecode A (just x) (just y) c = + transport + (λ k → + encode (λ i → A (i ∧ k)) _ _ (decode (just x) (just (c k)) (λ i → c (i ∧ k))) + ≡ (λ i → c (i ∧ k))) + (transportRefl _) + + Code≃PathP : {A : I → Type ℓ} → ∀ ox oy → Code A ox oy ≃ PathP (λ i → Maybe (A i)) ox oy + Code≃PathP {A = A} ox oy = isoToEquiv isom + where + isom : Iso _ _ + isom .Iso.fun = decode ox oy + isom .Iso.inv = encode _ ox oy + isom .Iso.rightInv = decodeEncode ox oy + isom .Iso.leftInv = encodeDecode A ox oy + +-- Structured isomorphisms + +MaybeStructure : (S : Type ℓ → Type ℓ₁) → Type ℓ → Type ℓ₁ +MaybeStructure S X = Maybe (S X) + +MaybeEquivStr : {S : Type ℓ → Type ℓ₁} + → StrEquiv S ℓ₁' → StrEquiv (MaybeStructure S) ℓ₁' +MaybeEquivStr ι (X , ox) (Y , oy) e = MaybeRel (λ x y → ι (X , x) (Y , y) e) ox oy + +maybeUnivalentStr : {S : Type ℓ → Type ℓ₁} (ι : StrEquiv S ℓ₁') + → UnivalentStr S ι → UnivalentStr (MaybeStructure S) (MaybeEquivStr ι) +maybeUnivalentStr ι θ {X , ox} {Y , oy} e = + compEquiv + (congMaybeRel (λ x y → θ {X , x} {Y , y} e) ox oy) + (MaybePathP.Code≃PathP ox oy) + +maybeEquivAction : {S : Type ℓ → Type ℓ₁} + → EquivAction S → EquivAction (MaybeStructure S) +maybeEquivAction α e = congMaybeEquiv (α e) + +maybeTransportStr : {S : Type ℓ → Type ℓ₁} (α : EquivAction S) + → TransportStr α → TransportStr (maybeEquivAction α) +maybeTransportStr _ τ e nothing = refl +maybeTransportStr _ τ e (just x) = cong just (τ e x) diff --git a/Cubical/Structures/MultiSet.agda b/Cubical/Structures/MultiSet.agda new file mode 100644 index 0000000000..6352a2f58f --- /dev/null +++ b/Cubical/Structures/MultiSet.agda @@ -0,0 +1,42 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.MultiSet where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.SIP +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Structures.Auto + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +private + variable + ℓ : Level + +module _ (A : Type ℓ) (Aset : isSet A) where + + CountStructure : Type ℓ → Type ℓ + CountStructure X = A → X → ℕ + + CountEquivStr = AutoEquivStr CountStructure + + countUnivalentStr : UnivalentStr _ CountEquivStr + countUnivalentStr = autoUnivalentStr CountStructure + + Count : Type (ℓ-suc ℓ) + Count = TypeWithStr ℓ CountStructure + + MultiSetStructure : Type ℓ → Type ℓ + MultiSetStructure X = X × (A → X → X) × (A → X → ℕ) + + MultiSetEquivStr = AutoEquivStr MultiSetStructure + + multiSetUnivalentStr : UnivalentStr _ MultiSetEquivStr + multiSetUnivalentStr = autoUnivalentStr MultiSetStructure + + MultiSet : Type (ℓ-suc ℓ) + MultiSet = TypeWithStr ℓ MultiSetStructure diff --git a/Cubical/Structures/Parameterized.agda b/Cubical/Structures/Parameterized.agda new file mode 100644 index 0000000000..04a189d03a --- /dev/null +++ b/Cubical/Structures/Parameterized.agda @@ -0,0 +1,48 @@ +{- + +A parameterized family of structures S can be combined into a single structure: +X ↦ (a : A) → S a X + +This is more general than Structures.Function in that S can vary in A. + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Parameterized where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Functions.FunExtEquiv +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Univalence + +private + variable + ℓ ℓ₁ ℓ₁' : Level + +module _ {ℓ₀} (A : Type ℓ₀) where + + ParamStructure : (S : A → Type ℓ → Type ℓ₁) + → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + ParamStructure S X = (a : A) → S a X + + ParamEquivStr : {S : A → Type ℓ → Type ℓ₁} + → (∀ a → StrEquiv (S a) ℓ₁') → StrEquiv (ParamStructure S) (ℓ-max ℓ₀ ℓ₁') + ParamEquivStr ι (X , l) (Y , m) e = ∀ a → ι a (X , l a) (Y , m a) e + + paramUnivalentStr : {S : A → Type ℓ → Type ℓ₁} + (ι : ∀ a → StrEquiv (S a) ℓ₁') (θ : ∀ a → UnivalentStr (S a) (ι a)) + → UnivalentStr (ParamStructure S) (ParamEquivStr ι) + paramUnivalentStr ι θ e = compEquiv (equivΠCod λ a → θ a e) funExtEquiv + + paramEquivAction : {S : A → Type ℓ → Type ℓ₁} + → (∀ a → EquivAction (S a)) → EquivAction (ParamStructure S) + paramEquivAction α e = equivΠCod (λ a → α a e) + + paramTransportStr : {S : A → Type ℓ → Type ℓ₁} + (α : ∀ a → EquivAction (S a)) (τ : ∀ a → TransportStr (α a)) + → TransportStr (paramEquivAction α) + paramTransportStr {S = S} α τ e f = + funExt λ a → + τ a e (f a) + ∙ cong (λ fib → transport (λ i → S (fib .snd (~ i)) (ua e i)) (f (fib .snd i1))) + (isContrSingl a .snd (_ , sym (transportRefl a))) diff --git a/Cubical/Structures/Pointed.agda b/Cubical/Structures/Pointed.agda new file mode 100644 index 0000000000..ea6c76b1ea --- /dev/null +++ b/Cubical/Structures/Pointed.agda @@ -0,0 +1,41 @@ +{- + +Pointed structure: X ↦ X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Pointed where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.SIP + +open import Cubical.Foundations.Pointed.Base + +private + variable + ℓ : Level + +-- Structured isomorphisms + +PointedStructure : Type ℓ → Type ℓ +PointedStructure X = X + +PointedEquivStr : StrEquiv PointedStructure ℓ +PointedEquivStr A B f = equivFun f (pt A) ≡ pt B + +pointedUnivalentStr : UnivalentStr {ℓ} PointedStructure PointedEquivStr +pointedUnivalentStr f = invEquiv (ua-ungluePath-Equiv f) + +pointedSIP : (A B : Pointed ℓ) → A ≃[ PointedEquivStr ] B ≃ (A ≡ B) +pointedSIP = SIP pointedUnivalentStr + +pointed-sip : (A B : Pointed ℓ) → A ≃[ PointedEquivStr ] B → (A ≡ B) +pointed-sip A B = equivFun (pointedSIP A B) -- ≡ λ (e , p) i → ua e i , ua-gluePath e p i + +pointedEquivAction : EquivAction {ℓ} PointedStructure +pointedEquivAction e = e + +pointedTransportStr : TransportStr {ℓ} pointedEquivAction +pointedTransportStr e s = sym (transportRefl _) diff --git a/Cubical/Structures/Product.agda b/Cubical/Structures/Product.agda new file mode 100644 index 0000000000..de2ddba3d3 --- /dev/null +++ b/Cubical/Structures/Product.agda @@ -0,0 +1,50 @@ +{- + +Product of structures S and T: X ↦ S X × T X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Product where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path +open import Cubical.Foundations.SIP +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₂ ℓ₂' : Level + +ProductStructure : (S₁ : Type ℓ → Type ℓ₁) (S₂ : Type ℓ → Type ℓ₂) + → Type ℓ → Type (ℓ-max ℓ₁ ℓ₂) +ProductStructure S₁ S₂ X = S₁ X × S₂ X + +ProductEquivStr : + {S₁ : Type ℓ → Type ℓ₁} (ι₁ : StrEquiv S₁ ℓ₁') + {S₂ : Type ℓ → Type ℓ₂} (ι₂ : StrEquiv S₂ ℓ₂') + → StrEquiv (ProductStructure S₁ S₂) (ℓ-max ℓ₁' ℓ₂') +ProductEquivStr ι₁ ι₂ (X , s₁ , s₂) (Y , t₁ , t₂) f = + (ι₁ (X , s₁) (Y , t₁) f) × (ι₂ (X , s₂) (Y , t₂) f) + +productUnivalentStr : + {S₁ : Type ℓ → Type ℓ₁} (ι₁ : StrEquiv S₁ ℓ₁') (θ₁ : UnivalentStr S₁ ι₁) + {S₂ : Type ℓ → Type ℓ₂} (ι₂ : StrEquiv S₂ ℓ₂') (θ₂ : UnivalentStr S₂ ι₂) + → UnivalentStr (ProductStructure S₁ S₂) (ProductEquivStr ι₁ ι₂) +productUnivalentStr {S₁ = S₁} ι₁ θ₁ {S₂} ι₂ θ₂ {X , s₁ , s₂} {Y , t₁ , t₂} e = + compEquiv (Σ-cong-equiv (θ₁ e) (λ _ → θ₂ e)) ΣPath≃PathΣ + +productEquivAction : + {S₁ : Type ℓ → Type ℓ₁} (α₁ : EquivAction S₁) + {S₂ : Type ℓ → Type ℓ₂} (α₂ : EquivAction S₂) + → EquivAction (ProductStructure S₁ S₂) +productEquivAction α₁ α₂ e = Σ-cong-equiv (α₁ e) (λ _ → α₂ e) + +productTransportStr : + {S₁ : Type ℓ → Type ℓ₁} (α₁ : EquivAction S₁) (τ₁ : TransportStr α₁) + {S₂ : Type ℓ → Type ℓ₂} (α₂ : EquivAction S₂) (τ₂ : TransportStr α₂) + → TransportStr (productEquivAction α₁ α₂) +productTransportStr _ τ₁ _ τ₂ e (s₁ , s₂) = ΣPathP (τ₁ e s₁ , τ₂ e s₂) diff --git a/Cubical/Structures/Queue.agda b/Cubical/Structures/Queue.agda new file mode 100644 index 0000000000..567cf70dac --- /dev/null +++ b/Cubical/Structures/Queue.agda @@ -0,0 +1,115 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Queue where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Functions.FunExtEquiv +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.SIP + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Macro +open import Cubical.Structures.Auto + +open import Cubical.Data.Unit +open import Cubical.Data.Maybe as Maybe +open import Cubical.Data.Sigma +open import Cubical.Data.List + + +-- Developing Queues as a standard notion of structure, see +-- https://github.com/ecavallo/cubical/blob/queue/Cubical/Experiments/Queue.agda +-- for the original development + +private + variable + ℓ ℓ' : Level + +-- We start fixing a set A on which we define what it means for a type Q to have +-- a Queue structure (wrt. A) +module Queues-on (A : Type ℓ) (Aset : isSet A) where + -- A Queue structure has three components, the empty Queue, an enqueue function and a dequeue function + -- We first deal with enq and deq as separate structures + + -- deq as a structure + -- First, a few preliminary results that we will need later + deqMap : {X Y : Type ℓ} → (X → Y) → Maybe (X × A) → Maybe (Y × A) + deqMap = map-Maybe ∘ map-fst + + deqMapId : {X : Type ℓ} → ∀ r → deqMap (idfun X) r ≡ r + deqMapId = map-Maybe-id + + deqMap-∘ :{B C D : Type ℓ} + (g : C → D) (f : B → C) + → ∀ r → deqMap {X = C} g (deqMap f r) ≡ deqMap (λ b → g (f b)) r + deqMap-∘ g f nothing = refl + deqMap-∘ g f (just (b , a)) = refl + + -- Now we can do Queues: + rawQueueDesc = + autoDesc (λ (X : Type ℓ) → X × (A → X → X) × (X → Transp[ Maybe (X × A) ])) + + open Macro ℓ rawQueueDesc public renaming + ( structure to RawQueueStructure + ; equiv to RawQueueEquivStr + ; univalent to rawQueueUnivalentStr + ) + + RawQueue : Type (ℓ-suc ℓ) + RawQueue = TypeWithStr ℓ RawQueueStructure + + returnOrEnq : {Q : Type ℓ} + → RawQueueStructure Q → A → Maybe (Q × A) → Q × A + returnOrEnq (emp , enq , _) a qr = + Maybe.rec (emp , a) (λ {(q , b) → enq a q , b}) qr + + QueueAxioms : (Q : Type ℓ) → RawQueueStructure Q → Type ℓ + QueueAxioms Q S@(emp , enq , deq) = + (isSet Q) + × (deq emp ≡ nothing) + × (∀ a q → deq (enq a q) ≡ just (returnOrEnq S a (deq q))) + × (∀ a a' q q' → enq a q ≡ enq a' q' → (a ≡ a') × (q ≡ q')) + × (∀ q q' → deq q ≡ deq q' → q ≡ q') + + isPropQueueAxioms : ∀ Q S → isProp (QueueAxioms Q S) + isPropQueueAxioms Q S = + isPropΣ isPropIsSet + (λ Qset → isProp×3 (isOfHLevelDeq Qset _ _) + (isPropΠ2 λ _ _ → isOfHLevelDeq Qset _ _) + (isPropΠ3 λ _ _ _ → isPropΠ2 λ _ _ → isProp× (Aset _ _) (Qset _ _)) + (isPropΠ3 λ _ _ _ → Qset _ _)) + where + isOfHLevelDeq : isSet Q → isOfHLevel 2 (Maybe (Q × A)) + isOfHLevelDeq Qset = isOfHLevelMaybe 0 (isSet× Qset Aset) + + QueueStructure : Type ℓ → Type ℓ + QueueStructure = AxiomsStructure RawQueueStructure QueueAxioms + + Queue : Type (ℓ-suc ℓ) + Queue = TypeWithStr ℓ QueueStructure + + QueueEquivStr : StrEquiv QueueStructure ℓ + QueueEquivStr = AxiomsEquivStr RawQueueEquivStr QueueAxioms + + queueUnivalentStr : UnivalentStr QueueStructure QueueEquivStr + queueUnivalentStr = axiomsUnivalentStr RawQueueEquivStr isPropQueueAxioms rawQueueUnivalentStr + + + FiniteQueueAxioms : (Q : Type ℓ) → QueueStructure Q → Type ℓ + FiniteQueueAxioms Q ((emp , enq , _) , _) = isEquiv (foldr enq emp) + + isPropFiniteQueueAxioms : ∀ Q S → isProp (FiniteQueueAxioms Q S) + isPropFiniteQueueAxioms Q S = isPropIsEquiv _ + + FiniteQueueStructure : Type ℓ → Type ℓ + FiniteQueueStructure = AxiomsStructure QueueStructure FiniteQueueAxioms + + FiniteQueue : Type (ℓ-suc ℓ) + FiniteQueue = TypeWithStr ℓ FiniteQueueStructure + + FiniteQueueEquivStr : StrEquiv FiniteQueueStructure ℓ + FiniteQueueEquivStr = AxiomsEquivStr QueueEquivStr FiniteQueueAxioms + + finiteQueueUnivalentStr : UnivalentStr FiniteQueueStructure FiniteQueueEquivStr + finiteQueueUnivalentStr = axiomsUnivalentStr QueueEquivStr isPropFiniteQueueAxioms queueUnivalentStr diff --git a/Cubical/Structures/Record.agda b/Cubical/Structures/Record.agda new file mode 100644 index 0000000000..ce26b89961 --- /dev/null +++ b/Cubical/Structures/Record.agda @@ -0,0 +1,453 @@ +{- + +Automatically generating proofs of UnivalentStr for records + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Record where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Univalence +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.List as List +open import Cubical.Data.Vec as Vec +open import Cubical.Data.Bool +open import Cubical.Data.Maybe +open import Cubical.Data.Sum +open import Cubical.Structures.Auto +import Cubical.Structures.Macro as M +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +-- Magic number +private + FUEL = 10000 + +-- Types for specifying inputs to the tactics + +data AutoFieldSpec : Typeω where + autoFieldSpec : ∀ {ℓ ℓ₁ ℓ₂} (R : Type ℓ → Type ℓ₁) {S : Type ℓ → Type ℓ₂} + → ({X : Type ℓ} → R X → S X) + → AutoFieldSpec + +module _ {ℓ ℓ₁ ℓ₁'} where + mutual + data AutoFields (R : Type ℓ → Type ℓ₁) (ι : StrEquiv R ℓ₁') : Typeω + where + fields: : AutoFields R ι + _data[_∣_] : (fs : AutoFields R ι) + → ∀ {ℓ₂ ℓ₂'} {S : Type ℓ → Type ℓ₂} {ι' : StrEquiv S ℓ₂'} + → (f : {X : Type ℓ} → R X → S X) + → ({A B : TypeWithStr ℓ R} {e : typ A ≃ typ B} → ι A B e → ι' (map-snd f A) (map-snd f B) e) + → AutoFields R ι + _prop[_∣_] : (fs : AutoFields R ι) + → ∀ {ℓ₂} {P : (X : Type ℓ) → GatherFields fs X → Type ℓ₂} + → ({X : Type ℓ} (r : R X) → P X (projectFields fs r)) + → isPropProperty R ι fs P + → AutoFields R ι + + GatherFieldsLevel : {R : Type ℓ → Type ℓ₁} {ι : StrEquiv R ℓ₁'} + → AutoFields R ι + → Level + GatherFieldsLevel fields: = ℓ-zero + GatherFieldsLevel (_data[_∣_] fs {ℓ₂ = ℓ₂} _ _) = ℓ-max (GatherFieldsLevel fs) ℓ₂ + GatherFieldsLevel (_prop[_∣_] fs {ℓ₂ = ℓ₂} _ _) = ℓ-max (GatherFieldsLevel fs) ℓ₂ + + GatherFields : {R : Type ℓ → Type ℓ₁} {ι : StrEquiv R ℓ₁'} + (dat : AutoFields R ι) + → Type ℓ → Type (GatherFieldsLevel dat) + GatherFields fields: X = Unit + GatherFields (_data[_∣_] fs {S = S} _ _) X = GatherFields fs X × S X + GatherFields (_prop[_∣_] fs {P = P} _ _) X = + Σ[ s ∈ GatherFields fs X ] (P X s) + + projectFields : {R : Type ℓ → Type ℓ₁} {ι : StrEquiv R ℓ₁'} + (fs : AutoFields R ι) + → {X : Type ℓ} → R X → GatherFields fs X + projectFields fields: = _ + projectFields (fs data[ f ∣ _ ]) r = projectFields fs r , f r + projectFields (fs prop[ f ∣ _ ]) r = projectFields fs r , f r + + isPropProperty : ∀ {ℓ₂} (R : Type ℓ → Type ℓ₁) + (ι : StrEquiv R ℓ₁') + (fs : AutoFields R ι) + (P : (X : Type ℓ) → GatherFields fs X → Type ℓ₂) + → Type (ℓ-max (ℓ-suc ℓ) (ℓ-max ℓ₁ ℓ₂)) + isPropProperty R ι fs P = + {X : Type ℓ} (r : R X) → isProp (P X (projectFields fs r)) + + data AutoRecordSpec : Typeω where + autoRecordSpec : (R : Type ℓ → Type ℓ₁) (ι : StrEquiv R ℓ₁') + → AutoFields R ι + → AutoRecordSpec + +-- Some reflection utilities +private + tApply : R.Term → List (R.Arg R.Term) → R.Term + tApply t l = R.def (quote idfun) (R.unknown v∷ t v∷ l) + + tStrMap : R.Term → R.Term → R.Term + tStrMap A f = R.def (quote map-snd) (f v∷ A v∷ []) + + tStrProj : R.Term → R.Name → R.Term + tStrProj A sfield = tStrMap A (R.def sfield []) + + Fun : ∀ {ℓ ℓ'} → Type ℓ → Type ℓ' → Type (ℓ-max ℓ ℓ') + Fun A B = A → B + +-- Helper functions used in the generated univalence proof +private + pathMap : ∀ {ℓ ℓ'} {S : I → Type ℓ} {T : I → Type ℓ'} (f : {i : I} → S i → T i) + {x : S i0} {y : S i1} → PathP S x y → PathP T (f x) (f y) + pathMap f p i = f (p i) + + -- Property field helper functions + module _ + {ℓ ℓ₁ ℓ₁' ℓ₂} + (R : Type ℓ → Type ℓ₁) -- Structure record + (ι : StrEquiv R ℓ₁') -- Equivalence record + (fs : AutoFields R ι) -- Prior fields + (P : (X : Type ℓ) → GatherFields fs X → Type ℓ₂) -- Property type + (f : {X : Type ℓ} (r : R X) → P X (projectFields fs r)) -- Property projection + where + + prev = projectFields fs + Prev = GatherFields fs + + PropHelperCenterType : Type _ + PropHelperCenterType = + (A B : TypeWithStr ℓ R) (e : A .fst ≃ B .fst) + (p : PathP (λ i → Prev (ua e i)) (prev (A .snd)) (prev (B .snd))) + → PathP (λ i → P (ua e i) (p i)) (f (A .snd)) (f (B .snd)) + + PropHelperContractType : PropHelperCenterType → Type _ + PropHelperContractType c = + (A B : TypeWithStr ℓ R) (e : A .fst ≃ B .fst) + {p₀ : PathP (λ i → Prev (ua e i)) (prev (A .snd)) (prev (B .snd))} + (q : PathP (λ i → R (ua e i)) (A .snd) (B .snd)) + (p : p₀ ≡ (λ i → prev (q i))) + → PathP (λ k → PathP (λ i → P (ua e i) (p k i)) (f (A .snd)) (f (B .snd))) + (c A B e p₀) + (λ i → f (q i)) + + PropHelperType : Type _ + PropHelperType = + Σ PropHelperCenterType PropHelperContractType + + derivePropHelper : isPropProperty R ι fs P → PropHelperType + derivePropHelper propP .fst A B e p = + isOfHLevelPathP' 0 (propP _) (f (A .snd)) (f (B .snd)) .fst + derivePropHelper propP .snd A B e q p = + isOfHLevelPathP' 0 (isOfHLevelPathP 1 (propP _) _ _) _ _ .fst + + -- Build proof of univalence from an isomorphism + module _ {ℓ ℓ₁ ℓ₁'} (S : Type ℓ → Type ℓ₁) (ι : StrEquiv S ℓ₁') where + + fwdShape : Type _ + fwdShape = + (A B : TypeWithStr ℓ S) (e : typ A ≃ typ B) → ι A B e → PathP (λ i → S (ua e i)) (str A) (str B) + + bwdShape : Type _ + bwdShape = + (A B : TypeWithStr ℓ S) (e : typ A ≃ typ B) → PathP (λ i → S (ua e i)) (str A) (str B) → ι A B e + + fwdBwdShape : fwdShape → bwdShape → Type _ + fwdBwdShape fwd bwd = + (A B : TypeWithStr ℓ S) (e : typ A ≃ typ B) → ∀ p → fwd A B e (bwd A B e p) ≡ p + + bwdFwdShape : fwdShape → bwdShape → Type _ + bwdFwdShape fwd bwd = + (A B : TypeWithStr ℓ S) (e : typ A ≃ typ B) → ∀ r → bwd A B e (fwd A B e r) ≡ r + + -- The implicit arguments A,B in UnivalentStr make some things annoying so let's avoid them + ExplicitUnivalentStr : Type _ + ExplicitUnivalentStr = + (A B : TypeWithStr _ S) (e : typ A ≃ typ B) → ι A B e ≃ PathP (λ i → S (ua e i)) (str A) (str B) + + explicitUnivalentStr : (fwd : fwdShape) (bwd : bwdShape) + → fwdBwdShape fwd bwd → bwdFwdShape fwd bwd + → ExplicitUnivalentStr + explicitUnivalentStr fwd bwd fwdBwd bwdFwd A B e = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = fwd A B e + isom .inv = bwd A B e + isom .rightInv = fwdBwd A B e + isom .leftInv = bwdFwd A B e + + ExplicitUnivalentDesc : ∀ ℓ {ℓ₁ ℓ₁'} → (d : M.Desc ℓ ℓ₁ ℓ₁') → Type _ + ExplicitUnivalentDesc _ d = + ExplicitUnivalentStr (M.MacroStructure d) (M.MacroEquivStr d) + + explicitUnivalentDesc : ∀ ℓ {ℓ₁ ℓ₁'} → (d : M.Desc ℓ ℓ₁ ℓ₁') → ExplicitUnivalentDesc ℓ d + explicitUnivalentDesc _ d A B e = M.MacroUnivalentStr d e + +-- Internal record specification type +private + record TypedTerm : Type where + field + type : R.Term + term : R.Term + + record InternalDatumField : Type where + field + sfield : R.Name -- name of structure field + efield : R.Name -- name of equivalence field + + record InternalPropField : Type where + field + sfield : R.Name -- name of structure field + + InternalField : Type + InternalField = InternalDatumField ⊎ InternalPropField + + record InternalSpec (A : Type) : Type where + field + srec : R.Term -- structure record type + erec : R.Term -- equivalence record type + fields : List (InternalField × A) -- in reverse order + + open TypedTerm + open InternalDatumField + open InternalPropField + +-- Parse a field and record specifications +private + findName : R.Term → R.TC R.Name + findName (R.def name _) = R.returnTC name + findName (R.lam R.hidden (R.abs _ t)) = findName t + findName t = R.typeError (R.strErr "Not a name + spine: " ∷ R.termErr t ∷ []) + + parseFieldSpec : R.Term → R.TC (R.Term × R.Term × R.Term × R.Term) + parseFieldSpec (R.con (quote autoFieldSpec) (ℓ h∷ ℓ₁ h∷ ℓ₂ h∷ R v∷ S h∷ f v∷ [])) = + R.reduce ℓ >>= λ ℓ → + R.returnTC (ℓ , ℓ₂ , S , f) + parseFieldSpec t = + R.typeError (R.strErr "Malformed field specification: " ∷ R.termErr t ∷ []) + + parseSpec : R.Term → R.TC (InternalSpec TypedTerm) + parseSpec (R.con (quote autoRecordSpec) (ℓ h∷ ℓ₁ h∷ ℓ₁' h∷ srecTerm v∷ erecTerm v∷ fs v∷ [])) = + parseFields fs >>= λ fs' → + R.returnTC λ { .srec → srecTerm ; .erec → erecTerm ; .fields → fs'} + where + open InternalSpec + + parseFields : R.Term → R.TC (List (InternalField × TypedTerm)) + parseFields (R.con (quote fields:) _) = R.returnTC [] + parseFields (R.con (quote _data[_∣_]) + (ℓ h∷ ℓ₁ h∷ ℓ₁' h∷ R h∷ ι h∷ fs v∷ ℓ₂ h∷ ℓ₂' h∷ S h∷ ι' h∷ sfieldTerm v∷ efieldTerm v∷ [])) + = + R.reduce ℓ >>= λ ℓ → + findName sfieldTerm >>= λ sfieldName → + findName efieldTerm >>= λ efieldName → + buildDesc FUEL ℓ ℓ₂ S >>= λ d → + let + f : InternalField × TypedTerm + f = λ + { .fst → inl λ { .sfield → sfieldName ; .efield → efieldName } + ; .snd .type → R.def (quote ExplicitUnivalentDesc) (ℓ v∷ d v∷ []) + ; .snd .term → R.def (quote explicitUnivalentDesc) (ℓ v∷ d v∷ []) + } + in + liftTC (f ∷_) (parseFields fs) + parseFields (R.con (quote _prop[_∣_]) + (ℓ h∷ ℓ₁ h∷ ℓ₁' h∷ R h∷ ι h∷ fs v∷ ℓ₂ h∷ P h∷ fieldTerm v∷ prop v∷ [])) + = + findName fieldTerm >>= λ fieldName → + let + p : InternalField × TypedTerm + p = λ + { .fst → inr λ { .sfield → fieldName } + ; .snd .type → + R.def (quote PropHelperType) (srecTerm v∷ erecTerm v∷ fs v∷ P v∷ fieldTerm v∷ []) + ; .snd .term → + R.def (quote derivePropHelper) (srecTerm v∷ erecTerm v∷ fs v∷ P v∷ fieldTerm v∷ prop v∷ []) + } + in + liftTC (p ∷_) (parseFields fs) + + parseFields t = R.typeError (R.strErr "Malformed autoRecord specification (1): " ∷ R.termErr t ∷ []) + + parseSpec t = R.typeError (R.strErr "Malformed autoRecord specification (2): " ∷ R.termErr t ∷ []) + +-- Build a proof of univalence from an InternalSpec +module _ (spec : InternalSpec ℕ) where + open InternalSpec spec + private + + fwdDatum : Vec R.Term 4 → R.Term → InternalDatumField × ℕ → R.Term + fwdDatum (A ∷ B ∷ e ∷ streq ∷ _) i (dat , n) = + R.def (quote equivFun) + (tApply (v n) (tStrProj A (dat .sfield) v∷ tStrProj B (dat .sfield) v∷ e v∷ []) + v∷ R.def (dat .efield) (streq v∷ []) + v∷ i + v∷ []) + + fwdProperty : Vec R.Term 4 → R.Term → R.Term → InternalPropField × ℕ → R.Term + fwdProperty (A ∷ B ∷ e ∷ streq ∷ _) i prevPath prop = + R.def (quote fst) (v (prop .snd) v∷ A v∷ B v∷ e v∷ prevPath v∷ i v∷ []) + + bwdClause : Vec R.Term 4 → InternalDatumField × ℕ → R.Clause + bwdClause (A ∷ B ∷ e ∷ q ∷ _) (dat , n) = + R.clause [] (R.proj (dat .efield) v∷ []) + (R.def (quote invEq) + (tApply + (v n) + (tStrProj A (dat .sfield) v∷ tStrProj B (dat .sfield) v∷ e v∷ []) + v∷ R.def (quote pathMap) (R.def (dat .sfield) [] v∷ q v∷ []) + v∷ [])) + + fwdBwdDatum : Vec R.Term 4 → R.Term → R.Term → InternalDatumField × ℕ → R.Term + fwdBwdDatum (A ∷ B ∷ e ∷ q ∷ _) j i (dat , n) = + R.def (quote secEq) + (tApply + (v n) + (tStrProj A (dat .sfield) v∷ tStrProj B (dat .sfield) v∷ e v∷ []) + v∷ R.def (quote pathMap) (R.def (dat .sfield) [] v∷ q v∷ []) + v∷ j v∷ i + v∷ []) + + fwdBwdProperty : Vec R.Term 4 → (j i prevPath : R.Term) → InternalPropField × ℕ → R.Term + fwdBwdProperty (A ∷ B ∷ e ∷ q ∷ _) j i prevPath prop = + R.def (quote snd) (v (prop .snd) v∷ A v∷ B v∷ e v∷ q v∷ prevPath v∷ j v∷ i v∷ []) + + bwdFwdClause : Vec R.Term 4 → R.Term → InternalDatumField × ℕ → R.Clause + bwdFwdClause (A ∷ B ∷ e ∷ streq ∷ _) j (dat , n) = + R.clause [] (R.proj (dat .efield) v∷ []) + (R.def (quote retEq) + (tApply + (v n) + (tStrProj A (dat .sfield) v∷ tStrProj B (dat .sfield) v∷ e v∷ []) + v∷ R.def (dat .efield) (streq v∷ []) + v∷ j + v∷ [])) + + makeVarsFrom : {n : ℕ} → ℕ → Vec R.Term n + makeVarsFrom {zero} k = [] + makeVarsFrom {suc n} k = v (n + k) ∷ (makeVarsFrom k) + + fwd : R.Term + fwd = + vlam "A" (vlam "B" (vlam "e" (vlam "streq" (vlam "i" (R.pat-lam body []))))) + where + -- input list is in reverse order + fwdClauses : ℕ → List (InternalField × ℕ) → List (R.Name × R.Term) + fwdClauses k [] = [] + fwdClauses k ((inl f , n) ∷ fs) = + fwdClauses k fs + ∷ʳ (f .sfield , fwdDatum (makeVarsFrom k) (v 0) (map-snd (4 + k +_) (f , n))) + fwdClauses k ((inr p , n) ∷ fs) = + fwdClauses k fs + ∷ʳ (p .sfield , fwdProperty (makeVarsFrom k) (v 0) prevPath (map-snd (4 + k +_) (p , n))) + where + prevPath = + vlam "i" + (List.foldl + (λ t (_ , t') → R.con (quote _,_) (t v∷ t' v∷ [])) + (R.con (quote tt) []) + (fwdClauses (suc k) fs)) + + body = + List.map (λ (n , t) → R.clause [] [ varg (R.proj n) ] t) (fwdClauses 1 fields) + + bwd : R.Term + bwd = + vlam "A" (vlam "B" (vlam "e" (vlam "q" (R.pat-lam (bwdClauses fields) [])))) + where + -- input is in reverse order + bwdClauses : List (InternalField × ℕ) → List R.Clause + bwdClauses [] = [] + bwdClauses ((inl f , n) ∷ fs) = + bwdClauses fs + ∷ʳ bwdClause (makeVarsFrom 0) (map-snd (4 +_) (f , n)) + bwdClauses ((inr p , n) ∷ fs) = bwdClauses fs + + fwdBwd : R.Term + fwdBwd = + vlam "A" (vlam "B" (vlam "e" (vlam "q" (vlam "j" (vlam "i" (R.pat-lam body [])))))) + where + -- input is in reverse order + fwdBwdClauses : ℕ → List (InternalField × ℕ) → List (R.Name × R.Term) + fwdBwdClauses k [] = [] + fwdBwdClauses k ((inl f , n) ∷ fs) = + fwdBwdClauses k fs + ∷ʳ (f .sfield , fwdBwdDatum (makeVarsFrom k) (v 1) (v 0) (map-snd (4 + k +_) (f , n))) + fwdBwdClauses k ((inr p , n) ∷ fs) = + fwdBwdClauses k fs + ∷ʳ ((p .sfield , fwdBwdProperty (makeVarsFrom k) (v 1) (v 0) prevPath (map-snd (4 + k +_) (p , n)))) + where + prevPath = + vlam "j" + (vlam "i" + (List.foldl + (λ t (_ , t') → R.con (quote _,_) (t v∷ t' v∷ [])) + (R.con (quote tt) []) + (fwdBwdClauses (2 + k) fs))) + + body = List.map (λ (n , t) → R.clause [] [ varg (R.proj n) ] t) (fwdBwdClauses 2 fields) + + bwdFwd : R.Term + bwdFwd = + vlam "A" (vlam "B" (vlam "e" (vlam "streq" (vlam "j" (R.pat-lam (bwdFwdClauses fields) []))))) + where + bwdFwdClauses : List (InternalField × ℕ) → List R.Clause + bwdFwdClauses [] = [] + bwdFwdClauses ((inl f , n) ∷ fs) = + bwdFwdClauses fs + ∷ʳ bwdFwdClause (makeVarsFrom 1) (v 0) (map-snd (5 +_) (f , n)) + bwdFwdClauses ((inr _ , n) ∷ fs) = bwdFwdClauses fs + + univalentRecord : R.Term + univalentRecord = + R.def (quote explicitUnivalentStr) + (R.unknown v∷ R.unknown v∷ fwd v∷ bwd v∷ fwdBwd v∷ bwdFwd v∷ []) + +macro + autoFieldEquiv : R.Term → R.Term → R.Term → R.Term → R.TC Unit + autoFieldEquiv spec A B hole = + (R.reduce spec >>= parseFieldSpec) >>= λ (ℓ , ℓ₂ , S , f) → + buildDesc FUEL ℓ ℓ₂ S >>= λ d → + R.unify hole (R.def (quote M.MacroEquivStr) (d v∷ tStrMap A f v∷ tStrMap B f v∷ [])) + + autoUnivalentRecord : R.Term → R.Term → R.TC Unit + autoUnivalentRecord t hole = + (R.reduce t >>= parseSpec) >>= λ spec → + -- R.typeError (R.strErr "WOW: " ∷ R.termErr (main spec) ∷ []) + R.unify (main spec) hole + where + module _ (spec : InternalSpec TypedTerm) where + open InternalSpec spec + + mapUp : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (ℕ → A → B) → ℕ → List A → List B + mapUp f _ [] = [] + mapUp f n (x ∷ xs) = f n x ∷ mapUp f (suc n) xs + + closureSpec : InternalSpec ℕ + closureSpec .InternalSpec.srec = srec + closureSpec .InternalSpec.erec = erec + closureSpec .InternalSpec.fields = mapUp (λ n → map-snd (λ _ → n)) 0 fields + + closure : R.Term + closure = + iter (List.length fields) (vlam "") (univalentRecord closureSpec) + + env : List (R.Arg R.Term) + env = List.map (varg ∘ term ∘ snd) (List.rev fields) + + closureTy : R.Term + closureTy = + List.foldr + (λ ty cod → R.def (quote Fun) (ty v∷ cod v∷ [])) + (R.def (quote ExplicitUnivalentStr) (srec v∷ erec v∷ [])) + (List.map (type ∘ snd) (List.rev fields)) + + main : R.Term + main = R.def (quote idfun) (closureTy v∷ closure v∷ env) diff --git a/Cubical/Structures/Relational/Auto.agda b/Cubical/Structures/Relational/Auto.agda new file mode 100644 index 0000000000..9901bdb145 --- /dev/null +++ b/Cubical/Structures/Relational/Auto.agda @@ -0,0 +1,239 @@ +{- + +Macros (autoDesc, AutoStructure, AutoEquivStr, autoUnivalentStr) for automatically generating structure definitions. + +For example: + + autoDesc (λ (X : Type₀) → X → X × ℕ) ↦ function+ var (var , constant ℕ) + +We prefer to use the constant structure whenever possible, e.g., [autoDesc (λ (X : Type₀) → ℕ → ℕ)] +is [constant (ℕ → ℕ)] rather than [function (constant ℕ) (constant ℕ)]. + +Writing [auto* (λ X → ⋯)] doesn't seem to work, but [auto* (λ (X : Type ℓ) → ⋯)] does. + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Relational.Auto where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.List +open import Cubical.Data.Bool +open import Cubical.Data.Maybe + +open import Cubical.Structures.Relational.Macro as Macro + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +-- Magic number +private + FUEL = 10000 + +-- Mark a constant type with a proof it is a set +abstract + Const[_] : ∀ {ℓ} → hSet ℓ → Type ℓ + Const[ A ] = A .fst + +-- Some reflection utilities +private + tLevel = R.def (quote Level) [] + + tType : R.Term → R.Term + tType ℓ = R.def (quote Type) [ varg ℓ ] + + thSet : R.Term → R.Term + thSet ℓ = R.def (quote hSet) [ varg ℓ ] + + tPosRelDesc : R.Term → R.Term → R.Term + tPosRelDesc ℓ ℓ₁ = R.def (quote PosRelDesc) (ℓ v∷ ℓ₁ v∷ []) + + tRelDesc : R.Term → R.Term → R.Term → R.Term + tRelDesc ℓ ℓ₁ ℓ₁' = R.def (quote RelDesc) (ℓ v∷ ℓ₁ v∷ ℓ₁' v∷ []) + + func : (ℓ ℓ' : Level) → Type (ℓ-suc (ℓ-max ℓ ℓ')) + func ℓ ℓ' = Type ℓ → Type ℓ' + + tStruct : R.Term → R.Term → R.Term + tStruct ℓ ℓ' = R.def (quote func) (varg ℓ ∷ varg ℓ' ∷ []) + +-- We try to build a descriptor by unifying the provided structure with combinators we're aware of. We +-- redefine the structure combinators as the *Shape terms below so that we don't depend on the specific way +-- these are defined in other files (order of implicit arguments and so on); the syntactic analysis that goes +-- on here means that we would get mysterious errors if those changed. +private + constantShape : ∀ {ℓ'} (ℓ : Level) (A : hSet ℓ') → (Type ℓ → Type ℓ') + constantShape _ A _ = Const[ A ] + + pointedShape : (ℓ : Level) → Type ℓ → Type ℓ + pointedShape _ X = X + + productShape : ∀ {ℓ₀ ℓ₁} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → (Type ℓ → Type ℓ₁) → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + productShape _ A₀ A₁ X = A₀ X × A₁ X + + paramShape : ∀ {ℓ₀ ℓ'} (ℓ : Level) + → Type ℓ' → (Type ℓ → Type ℓ₀) → Type ℓ → Type (ℓ-max ℓ' ℓ₀) + paramShape _ A A₀ X = A → A₀ X + + functionShape : ∀ {ℓ₀ ℓ₁} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → (Type ℓ → Type ℓ₁) → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + functionShape _ A₀ A₁ X = A₀ X → A₁ X + + maybeShape : ∀ {ℓ₀} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → Type ℓ → Type ℓ₀ + maybeShape _ A₀ X = Maybe (A₀ X) + +private + -- Build transport structure descriptor from a function [t : Type ℓ → Type ℓ'] + buildPosRelDesc : ℕ → R.Term → R.Term → R.Term → R.TC R.Term + buildPosRelDesc zero ℓ ℓ' t = R.typeError (R.strErr "Ran out of fuel! at \n" ∷ R.termErr t ∷ []) + buildPosRelDesc (suc fuel) ℓ ℓ' t = + tryConstant t <|> tryPointed t <|> tryProduct t <|> tryMaybe t <|> + R.typeError (R.strErr "Can't automatically generate a positive structure for\n" ∷ R.termErr t ∷ []) + where + tryConstant : R.Term → R.TC R.Term + tryConstant t = + newMeta (thSet ℓ') >>= λ A → + R.unify t (R.def (quote constantShape) (varg ℓ ∷ varg A ∷ [])) >> + R.returnTC (R.con (quote PosRelDesc.constant) (varg A ∷ [])) + + tryPointed : R.Term → R.TC R.Term + tryPointed t = + R.unify t (R.def (quote pointedShape) (varg ℓ ∷ [])) >> + R.returnTC (R.con (quote PosRelDesc.var) []) + + tryProduct : R.Term → R.TC R.Term + tryProduct t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote productShape) (varg ℓ ∷ varg A₀ ∷ varg A₁ ∷ [])) >> + buildPosRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildPosRelDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote PosRelDesc._,_) (varg d₀ ∷ varg d₁ ∷ [])) + + tryMaybe : R.Term → R.TC R.Term + tryMaybe t = + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote maybeShape) (varg ℓ ∷ varg A₀ ∷ [])) >> + buildPosRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote PosRelDesc.maybe) (varg d₀ ∷ [])) + + autoPosRelDesc' : R.Term → R.Term → R.TC Unit + autoPosRelDesc' t hole = + R.inferType hole >>= λ H → + newMeta tLevel >>= λ ℓ → + newMeta tLevel >>= λ ℓ' → + R.unify (tPosRelDesc ℓ ℓ') H >> + R.checkType t (tStruct ℓ ℓ') >> + buildPosRelDesc FUEL ℓ ℓ' t >>= R.unify hole + + -- Build structure descriptor from a function [t : Type ℓ → Type ℓ'] + buildRelDesc : ℕ → R.Term → R.Term → R.Term → R.TC R.Term + buildRelDesc zero ℓ ℓ' t = R.typeError (R.strErr "Ran out of fuel! at \n" ∷ R.termErr t ∷ []) + buildRelDesc (suc fuel) ℓ ℓ' t = + tryConstant t <|> tryPointed t <|> tryProduct t <|> tryParam t <|> tryFunction t <|> + tryMaybe t <|> + R.typeError (R.strErr "Can't automatically generate a structure for\n" ∷ R.termErr t ∷ []) + where + tryConstant : R.Term → R.TC R.Term + tryConstant t = + newMeta (thSet ℓ') >>= λ A → + R.unify t (R.def (quote constantShape) (varg ℓ ∷ varg A ∷ [])) >> + R.returnTC (R.con (quote RelDesc.constant) (varg A ∷ [])) + + tryPointed : R.Term → R.TC R.Term + tryPointed t = + R.unify t (R.def (quote pointedShape) (varg ℓ ∷ [])) >> + R.returnTC (R.con (quote RelDesc.var) []) + + tryProduct : R.Term → R.TC R.Term + tryProduct t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote productShape) (varg ℓ ∷ varg A₀ ∷ varg A₁ ∷ [])) >> + buildRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildRelDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote RelDesc._,_) (varg d₀ ∷ varg d₁ ∷ [])) + + tryParam : R.Term → R.TC R.Term + tryParam t = + newMeta (tType R.unknown) >>= λ A → + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote paramShape) (varg ℓ ∷ varg A ∷ varg A₀ ∷ [])) >> + buildRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote RelDesc.param) (varg A ∷ varg d₀ ∷ [])) + + tryFunction : R.Term → R.TC R.Term + tryFunction t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote functionShape) (varg ℓ ∷ varg A₀ ∷ varg A₁ ∷ [])) >> + buildPosRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildRelDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote RelDesc.function+) (varg d₀ ∷ varg d₁ ∷ [])) + + tryMaybe : R.Term → R.TC R.Term + tryMaybe t = + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote maybeShape) (varg ℓ ∷ varg A₀ ∷ [])) >> + buildRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote RelDesc.maybe) (varg d₀ ∷ [])) + + autoRelDesc' : R.Term → R.Term → R.TC Unit + autoRelDesc' t hole = + R.inferType hole >>= λ H → + newMeta tLevel >>= λ ℓ → + newMeta tLevel >>= λ ℓ' → + R.unify (tRelDesc ℓ ℓ' R.unknown) H >> + R.checkType t (tStruct ℓ ℓ') >> + buildRelDesc FUEL ℓ ℓ' t >>= R.unify hole + +macro + -- (Type ℓ → Type ℓ₁) → PosRelDesc ℓ + autoPosRelDesc : R.Term → R.Term → R.TC Unit + autoPosRelDesc = autoPosRelDesc' + + -- (S : Type ℓ → Type ℓ₁) → RelDesc ℓ + autoRelDesc : R.Term → R.Term → R.TC Unit + autoRelDesc = autoRelDesc' + + -- (S : Type ℓ → Type ℓ₁) → (Type ℓ → Type ℓ₁) + -- Sanity check: should be the identity + AutoStructure : R.Term → R.Term → R.TC Unit + AutoStructure t hole = + newMeta (tRelDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote RelMacroStructure) [ varg d ]) >> + autoRelDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → StrRel S _ + AutoRelStr : R.Term → R.Term → R.TC Unit + AutoRelStr t hole = + newMeta (tRelDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote RelMacroRelStr) [ varg d ]) >> + autoRelDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → SuitableStrRel S (AutoStrRel S) + autoSuitableRel : R.Term → R.Term → R.TC Unit + autoSuitableRel t hole = + newMeta (tRelDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote relMacroSuitableRel) [ varg d ]) >> + autoRelDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → StrRelMatchesEquiv (AutoRelStr S) (AutoEquivStr S) + autoMatchesEquiv : R.Term → R.Term → R.TC Unit + autoMatchesEquiv t hole = + newMeta (tRelDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote relMacroMatchesEquiv) [ varg d ]) >> + autoRelDesc' t d diff --git a/Cubical/Structures/Relational/Constant.agda b/Cubical/Structures/Relational/Constant.agda new file mode 100644 index 0000000000..78394cef36 --- /dev/null +++ b/Cubical/Structures/Relational/Constant.agda @@ -0,0 +1,57 @@ +{- + +Constant structure: _ ↦ A + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Constant where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.SetQuotients + +open import Cubical.Structures.Constant + +private + variable + ℓ ℓ' : Level + +-- Structured relations + +module _ (A : hSet ℓ') where + + ConstantRelStr : StrRel {ℓ = ℓ} (ConstantStructure (A .fst)) ℓ' + ConstantRelStr _ a₀ a₁ = a₀ ≡ a₁ + + constantSuitableRel : SuitableStrRel {ℓ = ℓ} (ConstantStructure (A .fst)) ConstantRelStr + constantSuitableRel .quo _ _ _ = isContrSingl _ + constantSuitableRel .symmetric _ = sym + constantSuitableRel .transitive _ _ = _∙_ + constantSuitableRel .set _ = A .snd + constantSuitableRel .prop _ = A .snd + + constantRelMatchesEquiv : StrRelMatchesEquiv {ℓ = ℓ} ConstantRelStr (ConstantEquivStr (A .fst)) + constantRelMatchesEquiv _ _ _ = idEquiv _ + + constantRelAction : StrRelAction {ℓ = ℓ} ConstantRelStr + constantRelAction .actStr _ a = a + constantRelAction .actStrId _ = refl + constantRelAction .actRel _ a a' p = p + + constantPositiveRel : PositiveStrRel {ℓ = ℓ} constantSuitableRel + constantPositiveRel .act = constantRelAction + constantPositiveRel .reflexive a = refl + constantPositiveRel .detransitive R R' p = ∣ _ , p , refl ∣ + constantPositiveRel .quo R = isoToIsEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = _ + isom .inv = [_] + isom .rightInv _ = refl + isom .leftInv = elimProp (λ _ → squash/ _ _) (λ a → refl) diff --git a/Cubical/Structures/Relational/Equalizer.agda b/Cubical/Structures/Relational/Equalizer.agda new file mode 100644 index 0000000000..faeeaa38d0 --- /dev/null +++ b/Cubical/Structures/Relational/Equalizer.agda @@ -0,0 +1,67 @@ +{- + +Equalizer of functions f g : S X ⇉ T X such that f and g act on relation structures + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Equalizer where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.RelationalStructure +open import Cubical.Data.Sigma +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.Relation.Binary.Base + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₂ ℓ₂' : Level + +EqualizerStructure : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (f g : ∀ X → S X → T X) + → Type ℓ → Type (ℓ-max ℓ₁ ℓ₂) +EqualizerStructure {S = S} f g X = Σ[ s ∈ S X ] f X s ≡ g X s + +EqualizerRelStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + {f g : ∀ X → S X → T X} + → StrRel S ℓ₁' → StrRel T ℓ₂' → StrRel (EqualizerStructure f g) ℓ₁' +EqualizerRelStr ρ₁ ρ₂ R (s , _) (s' , _) = ρ₁ R s s' + +equalizerSuitableRel : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + {f g : ∀ X → S X → T X} + {ρ₁ : StrRel S ℓ₁'} {ρ₂ : StrRel T ℓ₂'} + (αf : ∀ {X Y} {R : X → Y → Type ℓ} {s s'} → ρ₁ R s s' → ρ₂ R (f X s) (f Y s')) + (αg : ∀ {X Y} {R : X → Y → Type ℓ} {s s'} → ρ₁ R s s' → ρ₂ R (g X s) (g Y s')) + → SuitableStrRel S ρ₁ + → SuitableStrRel T ρ₂ + → SuitableStrRel (EqualizerStructure f g) (EqualizerRelStr ρ₁ ρ₂) +equalizerSuitableRel {f = f} {g} {ρ₁} {ρ₂} αf αg θ₁ θ₂ .quo (X , s , p) (R , qer) r = + ( ((quo₁ .fst .fst , sym step₀ ∙ step₁) , quo₁ .fst .snd) + , λ {((s' , _) , r') → + Σ≡Prop (λ _ → θ₁ .prop (λ _ _ → squash/ _ _) _ _) + (Σ≡Prop (λ _ → θ₂ .set squash/ _ _) + (cong fst (quo₁ .snd (s' , r'))))} + ) + where + quo₁ = θ₁ .quo (X , s) (R , qer) r + quo₂ = θ₂ .quo (X , f X s) (R , qer) (αf r) + + step₀ : quo₂ .fst .fst ≡ f (X / R .fst) (quo₁ .fst .fst) + step₀ = + cong fst + (quo₂ .snd + (f _ (quo₁ .fst .fst) , αf (quo₁ .fst .snd))) + + step₁ : quo₂ .fst .fst ≡ g (X / R .fst) (quo₁ .fst .fst) + step₁ = + cong fst + (quo₂ .snd + ( g _ (quo₁ .fst .fst) + , subst (λ t → ρ₂ (graphRel [_]) t (g _ (quo₁ .fst .fst))) (sym p) (αg (quo₁ .fst .snd)) + )) +equalizerSuitableRel _ _ θ₁ _ .symmetric R = θ₁ .symmetric R +equalizerSuitableRel _ _ θ₁ _ .transitive R R' = θ₁ .transitive R R' +equalizerSuitableRel _ _ θ₁ θ₂ .set setX = + isSetΣ (θ₁ .set setX) λ _ → isOfHLevelPath 2 (θ₂ .set setX) _ _ +equalizerSuitableRel _ _ θ₁ _ .prop propR _ _ = θ₁ .prop propR _ _ diff --git a/Cubical/Structures/Relational/Function.agda b/Cubical/Structures/Relational/Function.agda new file mode 100644 index 0000000000..e618bd14bb --- /dev/null +++ b/Cubical/Structures/Relational/Function.agda @@ -0,0 +1,191 @@ +{- + +Index a structure T a positive structure S: X ↦ S X → T X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Function where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Foundations.Univalence +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.Relation.Binary.Base +open import Cubical.Relation.ZigZag.Base +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.PropositionalTruncation as Trunc + +open import Cubical.Structures.Function + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₁'' ℓ₂ ℓ₂' ℓ₂'' : Level + +FunctionRelStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + → StrRel S ℓ₁' → StrRel T ℓ₂' → StrRel (FunctionStructure S T) (ℓ-max ℓ₁ (ℓ-max ℓ₁' ℓ₂')) +FunctionRelStr ρ₁ ρ₂ R f g = + ∀ {x y} → ρ₁ R x y → ρ₂ R (f x) (g y) + +open BinaryRelation +open isEquivRel + +private + composeWith[_] : {A : Type ℓ} (R : EquivPropRel A ℓ) + → compPropRel (R .fst) (quotientPropRel (R .fst .fst)) .fst ≡ graphRel [_] + composeWith[_] R = + funExt₂ λ a t → + hPropExt squash (squash/ _ _) + (Trunc.rec (squash/ _ _) (λ {(b , r , p) → eq/ a b r ∙ p })) + (λ p → ∣ a , R .snd .reflexive a , p ∣) + + [_]∙[_]⁻¹ : {A : Type ℓ} (R : EquivPropRel A ℓ) + → compPropRel (quotientPropRel (R .fst .fst)) (invPropRel (quotientPropRel (R .fst .fst))) .fst + ≡ R .fst .fst + [_]∙[_]⁻¹ R = + funExt₂ λ a b → + hPropExt squash (R .fst .snd a b) + (Trunc.rec (R .fst .snd a b) + (λ {(c , p , q) → effective (R .fst .snd) (R .snd) a b (p ∙ sym q)})) + (λ r → ∣ _ , eq/ a b r , refl ∣) + +functionSuitableRel : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + {ρ₁ : StrRel S ℓ₁'} {ρ₂ : StrRel T ℓ₂'} + (θ₁ : SuitableStrRel S ρ₁) + → PositiveStrRel θ₁ + → SuitableStrRel T ρ₂ + → SuitableStrRel (FunctionStructure S T) (FunctionRelStr ρ₁ ρ₂) +functionSuitableRel {S = S} {T = T} {ρ₁ = ρ₁} {ρ₂} θ₁ σ₁ θ₂ .quo (X , f) R h = + final + where + ref : (s : S X) → ρ₁ (R .fst .fst) s s + ref = posRelReflexive σ₁ R + + [f] : S X / ρ₁ (R .fst .fst) → T (X / R .fst .fst) + [f] [ s ] = θ₂ .quo (X , f s) R (h (ref s)) .fst .fst + [f] (eq/ s₀ s₁ r i) = + cong fst + (θ₂ .quo (X , f s₀) R (h (ref s₀)) .snd + ( [f] [ s₁ ] + , subst (λ R' → ρ₂ R' (f s₀) ([f] [ s₁ ])) + (composeWith[_] R) + (θ₂ .transitive (R .fst) (quotientPropRel (R .fst .fst)) + (h r) + (θ₂ .quo (X , f s₁) R (h (ref s₁)) .fst .snd)) + )) + i + [f] (squash/ _ _ p q j i) = + θ₂ .set squash/ _ _ (cong [f] p) (cong [f] q) j i + + relLemma : (s : S X) (t : S X) + → ρ₁ (graphRel [_]) s (funIsEq (σ₁ .quo R) [ t ]) + → ρ₂ (graphRel [_]) (f s) ([f] [ t ]) + relLemma s t r = + subst (λ R' → ρ₂ R' (f s) ([f] [ t ])) + (composeWith[_] R) + (θ₂ .transitive (R .fst) (quotientPropRel (R .fst .fst)) + (h r') + (θ₂ .quo (X , f t) R (h (ref t)) .fst .snd)) + where + r' : ρ₁ (R .fst .fst) s t + r' = + subst (λ R' → ρ₁ R' s t) ([_]∙[_]⁻¹ R) + (θ₁ .transitive + (quotientPropRel (R .fst .fst)) + (invPropRel (quotientPropRel (R .fst .fst))) + r + (θ₁ .symmetric (quotientPropRel (R .fst .fst)) + (subst + (λ t' → ρ₁ (graphRel [_]) t' (funIsEq (σ₁ .quo R) [ t ])) + (σ₁ .act .actStrId t) + (σ₁ .act .actRel eq/ t t (ref t))))) + + quoRelLemma : (s : S X) (t : S X / ρ₁ (R .fst .fst)) + → ρ₁ (graphRel [_]) s (funIsEq (σ₁ .quo R) t) + → ρ₂ (graphRel [_]) (f s) ([f] t) + quoRelLemma s = + elimProp (λ _ → isPropΠ λ _ → θ₂ .prop (λ _ _ → squash/ _ _) _ _) + (relLemma s) + + final : Σ (Σ _ _) _ + final .fst .fst = [f] ∘ invIsEq (σ₁ .quo R) + final .fst .snd {s} {t} r = + quoRelLemma s + (invIsEq (σ₁ .quo R) t) + (subst (ρ₁ (graphRel [_]) s) (sym (secIsEq (σ₁ .quo R) t)) r) + final .snd (f' , c) = + Σ≡Prop + (λ _ → isPropImplicitΠ λ s → + isPropImplicitΠ λ t → + isPropΠ λ _ → θ₂ .prop (λ _ _ → squash/ _ _) _ _) + (funExt λ s → contractorLemma (invIsEq (σ₁ .quo R) s) ∙ cong f' (secIsEq (σ₁ .quo R) s)) + where + contractorLemma : (s : S X / ρ₁ (R .fst .fst)) + → [f] s ≡ f' (funIsEq (σ₁ .quo R) s) + contractorLemma = + elimProp + (λ _ → θ₂ .set squash/ _ _) + (λ s → + cong fst + (θ₂ .quo (X , f s) R (h (ref s)) .snd + ( f' (funIsEq (σ₁ .quo R) [ s ]) + , c + (subst + (λ s' → ρ₁ (graphRel [_]) s' (funIsEq (σ₁ .quo R) [ s ])) + (σ₁ .act .actStrId s) + (σ₁ .act .actRel eq/ s s (ref s))) + ))) +functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .symmetric R h r = + θ₂ .symmetric R (h (θ₁ .symmetric (invPropRel R) r)) +functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .transitive R R' h h' rr' = + Trunc.rec + (θ₂ .prop (λ _ _ → squash) _ _) + (λ {(_ , r , r') → θ₂ .transitive R R' (h r) (h' r')}) + (σ .detransitive R R' rr') +functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .set setX = + isSetΠ λ _ → θ₂ .set setX +functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .prop propR f g = + isPropImplicitΠ λ _ → + isPropImplicitΠ λ _ → + isPropΠ λ _ → + θ₂ .prop propR _ _ + +functionRelMatchesEquiv : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (ρ₁ : StrRel S ℓ₁') {ι₁ : StrEquiv S ℓ₁''} + (ρ₂ : StrRel T ℓ₂') {ι₂ : StrEquiv T ℓ₂''} + → StrRelMatchesEquiv ρ₁ ι₁ + → StrRelMatchesEquiv ρ₂ ι₂ + → StrRelMatchesEquiv (FunctionRelStr ρ₁ ρ₂) (FunctionEquivStr ι₁ ι₂) +functionRelMatchesEquiv ρ₁ ρ₂ μ₁ μ₂ (X , f) (Y , g) e = + equivImplicitΠCod (equivImplicitΠCod (equiv→ (μ₁ _ _ e) (μ₂ _ _ e))) + +functionRelMatchesEquiv+ : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (ρ₁ : StrRel S ℓ₁') (α₁ : EquivAction S) + (ρ₂ : StrRel T ℓ₂') (ι₂ : StrEquiv T ℓ₂'') + → StrRelMatchesEquiv ρ₁ (EquivAction→StrEquiv α₁) + → StrRelMatchesEquiv ρ₂ ι₂ + → StrRelMatchesEquiv (FunctionRelStr ρ₁ ρ₂) (FunctionEquivStr+ α₁ ι₂) +functionRelMatchesEquiv+ ρ₁ α₁ ρ₂ ι₂ μ₁ μ₂ (X , f) (Y , g) e = + compEquiv + (functionRelMatchesEquiv ρ₁ ρ₂ μ₁ μ₂ (X , f) (Y , g) e) + (isoToEquiv isom) + where + open Iso + isom : Iso + (FunctionEquivStr (EquivAction→StrEquiv α₁) ι₂ (X , f) (Y , g) e) + (FunctionEquivStr+ α₁ ι₂ (X , f) (Y , g) e) + isom .fun h s = h refl + isom .inv k {x} = J (λ y _ → ι₂ (X , f x) (Y , g y) e) (k x) + isom .rightInv k i x = JRefl (λ y _ → ι₂ (X , f x) (Y , g y) e) (k x) i + isom .leftInv h = + implicitFunExt λ {x} → + implicitFunExt λ {y} → + funExt λ p → + J (λ y p → isom .inv (isom .fun h) p ≡ h p) + (funExt⁻ (isom .rightInv (isom .fun h)) x) + p diff --git a/Cubical/Structures/Relational/Macro.agda b/Cubical/Structures/Relational/Macro.agda new file mode 100644 index 0000000000..61228a91b6 --- /dev/null +++ b/Cubical/Structures/Relational/Macro.agda @@ -0,0 +1,161 @@ +{- + +Descriptor language for easily defining relational structures + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Relational.Macro where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Data.Sigma + +open import Cubical.Structures.Relational.Constant +open import Cubical.Structures.Relational.Function +open import Cubical.Structures.Relational.Maybe +open import Cubical.Structures.Relational.Parameterized +open import Cubical.Structures.Relational.Pointed +open import Cubical.Structures.Relational.Product + +open import Cubical.Structures.Macro +open import Cubical.Structures.Maybe + +data PosRelDesc (ℓ : Level) : Level → Typeω where + -- constant structure: X ↦ A + constant : ∀ {ℓ₁} → hSet ℓ₁ → PosRelDesc ℓ ℓ₁ + -- pointed structure: X ↦ X + var : PosRelDesc ℓ ℓ + -- product of structures S,T : X ↦ (S X × T X) + _,_ : ∀ {ℓ₁ ℓ₂} → PosRelDesc ℓ ℓ₁ → PosRelDesc ℓ ℓ₂ → PosRelDesc ℓ (ℓ-max ℓ₁ ℓ₂) + -- Maybe on a structure S: X ↦ Maybe (S X) + maybe : ∀ {ℓ₁} → PosRelDesc ℓ ℓ₁ → PosRelDesc ℓ ℓ₁ + +data RelDesc (ℓ : Level) : Level → Level → Typeω where + -- constant structure: X ↦ A + constant : ∀ {ℓ₁} → hSet ℓ₁ → RelDesc ℓ ℓ₁ ℓ₁ + -- pointed structure: X ↦ X + var : RelDesc ℓ ℓ ℓ + -- product of structures S,T : X ↦ (S X × T X) + _,_ : ∀ {ℓ₁ ℓ₁' ℓ₂ ℓ₂'} → RelDesc ℓ ℓ₁ ℓ₁' → RelDesc ℓ ℓ₂ ℓ₂' → RelDesc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁' ℓ₂') + -- structure S parameterized by constant A : X ↦ (A → S X) + param : ∀ {ℓ₁ ℓ₂ ℓ₂'} → (A : Type ℓ₁) → RelDesc ℓ ℓ₂ ℓ₂' → RelDesc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁ ℓ₂') + -- function from positive structure S to T: X ↦ (S X → T X) + function+ : ∀ {ℓ₁ ℓ₂ ℓ₂'} → PosRelDesc ℓ ℓ₁ → RelDesc ℓ ℓ₂ ℓ₂' → RelDesc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁ ℓ₂') + -- Maybe on a structure S: X ↦ Maybe (S X) + maybe : ∀ {ℓ₁ ℓ₁'} → RelDesc ℓ ℓ₁ ℓ₁' → RelDesc ℓ ℓ₁ ℓ₁' + +infixr 4 _,_ + +posRelDesc→TranspDesc : ∀ {ℓ ℓ₁} → PosRelDesc ℓ ℓ₁ → TranspDesc ℓ ℓ₁ +posRelDesc→TranspDesc (constant A) = constant (A .fst) +posRelDesc→TranspDesc var = var +posRelDesc→TranspDesc (d₀ , d₁) = posRelDesc→TranspDesc d₀ , posRelDesc→TranspDesc d₁ +posRelDesc→TranspDesc (maybe d) = maybe (posRelDesc→TranspDesc d) + +posRelDesc→RelDesc : ∀ {ℓ ℓ₁} → PosRelDesc ℓ ℓ₁ → RelDesc ℓ ℓ₁ ℓ₁ +posRelDesc→RelDesc (constant A) = constant A +posRelDesc→RelDesc var = var +posRelDesc→RelDesc (d₀ , d₁) = posRelDesc→RelDesc d₀ , posRelDesc→RelDesc d₁ +posRelDesc→RelDesc (maybe d) = maybe (posRelDesc→RelDesc d) + +relDesc→Desc : ∀ {ℓ ℓ₁ ℓ₁'} → RelDesc ℓ ℓ₁ ℓ₁' → Desc ℓ ℓ₁ ℓ₁' +relDesc→Desc (constant A) = constant (A .fst) +relDesc→Desc var = var +relDesc→Desc (d₀ , d₁) = relDesc→Desc d₀ , relDesc→Desc d₁ +relDesc→Desc (param A d) = function+ (constant A) (relDesc→Desc d) +relDesc→Desc (function+ d₀ d₁) = function+ (posRelDesc→TranspDesc d₀) (relDesc→Desc d₁) +relDesc→Desc (maybe d) = maybe (relDesc→Desc d) + +{- Definition of structure -} + +PosRelMacroStructure : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) → Type ℓ → Type ℓ₁ +PosRelMacroStructure d = TranspMacroStructure (posRelDesc→TranspDesc d) + +RelMacroStructure : ∀ {ℓ ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') → Type ℓ → Type ℓ₁ +RelMacroStructure d = MacroStructure (relDesc→Desc d) + +{- Notion of structured relation defined by a descriptor -} + +PosRelMacroRelStr : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) → StrRel (PosRelMacroStructure d) ℓ₁ +PosRelMacroRelStr (constant A) = ConstantRelStr A +PosRelMacroRelStr var = PointedRelStr +PosRelMacroRelStr (d₀ , d₁) = ProductRelStr (PosRelMacroRelStr d₀) (PosRelMacroRelStr d₁) +PosRelMacroRelStr (maybe d) = MaybeRelStr (PosRelMacroRelStr d) + +RelMacroRelStr : ∀ {ℓ ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') → StrRel (RelMacroStructure d) ℓ₁' +RelMacroRelStr (constant A) = ConstantRelStr A +RelMacroRelStr var = PointedRelStr +RelMacroRelStr (d₀ , d₁) = ProductRelStr (RelMacroRelStr d₀) (RelMacroRelStr d₁) +RelMacroRelStr (param A d) = ParamRelStr A (λ _ → RelMacroRelStr d) +RelMacroRelStr (function+ d₀ d₁) = + FunctionRelStr (PosRelMacroRelStr d₀) (RelMacroRelStr d₁) +RelMacroRelStr (maybe d) = MaybeRelStr (RelMacroRelStr d) + +{- Proof that structure induced by descriptor is suitable or positive -} + +posRelMacroSuitableRel : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) → SuitableStrRel _ (PosRelMacroRelStr d) +posRelMacroSuitableRel (constant A) = constantSuitableRel A +posRelMacroSuitableRel var = pointedSuitableRel +posRelMacroSuitableRel (d₀ , d₁) = + productSuitableRel (posRelMacroSuitableRel d₀) (posRelMacroSuitableRel d₁) +posRelMacroSuitableRel (maybe d) = maybeSuitableRel (posRelMacroSuitableRel d) + +posRelMacroPositiveRel : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) → PositiveStrRel (posRelMacroSuitableRel d) +posRelMacroPositiveRel (constant A) = constantPositiveRel A +posRelMacroPositiveRel var = pointedPositiveRel +posRelMacroPositiveRel (d₀ , d₁) = + productPositiveRel (posRelMacroPositiveRel d₀) (posRelMacroPositiveRel d₁) +posRelMacroPositiveRel (maybe d) = maybePositiveRel (posRelMacroPositiveRel d) + +relMacroSuitableRel : ∀ {ℓ ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') → SuitableStrRel _ (RelMacroRelStr d) +relMacroSuitableRel (constant A) = constantSuitableRel A +relMacroSuitableRel var = pointedSuitableRel +relMacroSuitableRel (d₀ , d₁) = productSuitableRel (relMacroSuitableRel d₀) (relMacroSuitableRel d₁) +relMacroSuitableRel (param A d) = paramSuitableRel A (λ _ → relMacroSuitableRel d) +relMacroSuitableRel (function+ d₀ d₁) = + functionSuitableRel (posRelMacroSuitableRel d₀) (posRelMacroPositiveRel d₀) (relMacroSuitableRel d₁) +relMacroSuitableRel (maybe d) = maybeSuitableRel (relMacroSuitableRel d) + +{- Proof that structured relations and equivalences agree -} + +posRelMacroMatchesEquiv : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) + → StrRelMatchesEquiv (PosRelMacroRelStr d) (EquivAction→StrEquiv (transpMacroAction (posRelDesc→TranspDesc d))) +posRelMacroMatchesEquiv (constant A) _ _ _ = idEquiv _ +posRelMacroMatchesEquiv var _ _ _ = idEquiv _ +posRelMacroMatchesEquiv (d₀ , d₁) = + productRelMatchesTransp + (PosRelMacroRelStr d₀) (transpMacroAction (posRelDesc→TranspDesc d₀)) + (PosRelMacroRelStr d₁) (transpMacroAction (posRelDesc→TranspDesc d₁)) + (posRelMacroMatchesEquiv d₀) (posRelMacroMatchesEquiv d₁) +posRelMacroMatchesEquiv (maybe d) = + maybeRelMatchesTransp + (PosRelMacroRelStr d) (transpMacroAction (posRelDesc→TranspDesc d)) + (posRelMacroMatchesEquiv d) + +relMacroMatchesEquiv : ∀ {ℓ ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') + → StrRelMatchesEquiv (RelMacroRelStr d) (MacroEquivStr (relDesc→Desc d)) +relMacroMatchesEquiv (constant A) = constantRelMatchesEquiv A +relMacroMatchesEquiv var = pointedRelMatchesEquiv +relMacroMatchesEquiv (d₁ , d₂) = + productRelMatchesEquiv + (RelMacroRelStr d₁) (RelMacroRelStr d₂) + (relMacroMatchesEquiv d₁) (relMacroMatchesEquiv d₂) +relMacroMatchesEquiv (param A d) = + paramRelMatchesEquiv A (λ _ → RelMacroRelStr d) (λ _ → relMacroMatchesEquiv d) +relMacroMatchesEquiv (function+ d₀ d₁) = + functionRelMatchesEquiv+ + (PosRelMacroRelStr d₀) (transpMacroAction (posRelDesc→TranspDesc d₀)) + (RelMacroRelStr d₁) (MacroEquivStr (relDesc→Desc d₁)) + (posRelMacroMatchesEquiv d₀) (relMacroMatchesEquiv d₁) +relMacroMatchesEquiv (maybe d) = + maybeRelMatchesEquiv (RelMacroRelStr d) (relMacroMatchesEquiv d) + +-- Module for easy importing +module RelMacro ℓ {ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') where + relation = RelMacroRelStr d + suitable = relMacroSuitableRel d + matches = relMacroMatchesEquiv d + open Macro ℓ (relDesc→Desc d) public diff --git a/Cubical/Structures/Relational/Maybe.agda b/Cubical/Structures/Relational/Maybe.agda new file mode 100644 index 0000000000..aa641c3378 --- /dev/null +++ b/Cubical/Structures/Relational/Maybe.agda @@ -0,0 +1,127 @@ +{- + +Maybe structure: X ↦ Maybe (S X) + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Relational.Maybe where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Maybe +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.HITs.SetQuotients + +open import Cubical.Structures.Maybe + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₁'' : Level + +-- Structured relations + +MaybeRelStr : {S : Type ℓ → Type ℓ₁} {ℓ₁' : Level} + → StrRel S ℓ₁' → StrRel (λ X → Maybe (S X)) ℓ₁' +MaybeRelStr ρ R = MaybeRel (ρ R) + +maybeSuitableRel : {S : Type ℓ → Type ℓ₁} {ρ : StrRel S ℓ₁'} + → SuitableStrRel S ρ + → SuitableStrRel (MaybeStructure S) (MaybeRelStr ρ) +maybeSuitableRel θ .quo (X , nothing) R _ .fst = nothing , _ +maybeSuitableRel θ .quo (X , nothing) R _ .snd (nothing , _) = refl +maybeSuitableRel θ .quo (X , just s) R c .fst = + just (θ .quo (X , s) R c .fst .fst) , θ .quo (X , s) R c .fst .snd +maybeSuitableRel θ .quo (X , just s) R c .snd (just s' , r) = + cong (λ {(t , r') → just t , r'}) (θ .quo (X , s) R c .snd (s' , r)) +maybeSuitableRel θ .symmetric R {nothing} {nothing} r = _ +maybeSuitableRel θ .symmetric R {just s} {just t} r = θ .symmetric R r +maybeSuitableRel θ .transitive R R' {nothing} {nothing} {nothing} r r' = _ +maybeSuitableRel θ .transitive R R' {just s} {just t} {just u} r r' = θ .transitive R R' r r' +maybeSuitableRel θ .set setX = isOfHLevelMaybe 0 (θ .set setX) +maybeSuitableRel θ .prop propR nothing nothing = isOfHLevelLift 1 isPropUnit +maybeSuitableRel θ .prop propR nothing (just y) = isOfHLevelLift 1 isProp⊥ +maybeSuitableRel θ .prop propR (just x) nothing = isOfHLevelLift 1 isProp⊥ +maybeSuitableRel θ .prop propR (just x) (just y) = θ .prop propR x y + +maybeRelMatchesEquiv : {S : Type ℓ → Type ℓ₁} (ρ : StrRel S ℓ₁') {ι : StrEquiv S ℓ₁''} + → StrRelMatchesEquiv ρ ι + → StrRelMatchesEquiv (MaybeRelStr ρ) (MaybeEquivStr ι) +maybeRelMatchesEquiv ρ μ (X , nothing) (Y , nothing) _ = Lift≃Lift (idEquiv _) +maybeRelMatchesEquiv ρ μ (X , nothing) (Y , just y) _ = Lift≃Lift (idEquiv _) +maybeRelMatchesEquiv ρ μ (X , just x) (Y , nothing) _ = Lift≃Lift (idEquiv _) +maybeRelMatchesEquiv ρ μ (X , just x) (Y , just y) = μ (X , x) (Y , y) + +maybeRelAction : + {S : Type ℓ → Type ℓ₁} {ρ : StrRel S ℓ₁'} + → StrRelAction ρ + → StrRelAction (MaybeRelStr ρ) +maybeRelAction α .actStr f = map-Maybe (α .actStr f) +maybeRelAction α .actStrId s = + funExt⁻ (cong map-Maybe (funExt (α .actStrId))) s ∙ map-Maybe-id s +maybeRelAction α .actRel h nothing nothing = _ +maybeRelAction α .actRel h (just s) (just t) r = α .actRel h s t r + +maybePositiveRel : + {S : Type ℓ → Type ℓ₁} {ρ : StrRel S ℓ₁'} {θ : SuitableStrRel S ρ} + → PositiveStrRel θ + → PositiveStrRel (maybeSuitableRel θ) +maybePositiveRel σ .act = maybeRelAction (σ .act) +maybePositiveRel σ .reflexive nothing = _ +maybePositiveRel σ .reflexive (just s) = σ .reflexive s +maybePositiveRel σ .detransitive R R' {nothing} {nothing} r = ∣ nothing , _ , _ ∣ +maybePositiveRel σ .detransitive R R' {just s} {just u} rr' = + Trunc.map (λ {(t , r , r') → just t , r , r'}) (σ .detransitive R R' rr') +maybePositiveRel {S = S} {ρ = ρ} {θ = θ} σ .quo {X} R = + subst isEquiv + (funExt + (elimProp (λ _ → maybeSuitableRel θ .set squash/ _ _) + (λ {nothing → refl; (just _) → refl}))) + (compEquiv (isoToEquiv isom) (congMaybeEquiv (_ , σ .quo R)) .snd) + where + fwd : Maybe (S X) / MaybeRel (ρ (R .fst .fst)) → Maybe (S X / ρ (R .fst .fst)) + fwd [ nothing ] = nothing + fwd [ just s ] = just [ s ] + fwd (eq/ nothing nothing r i) = nothing + fwd (eq/ (just s) (just t) r i) = just (eq/ s t r i) + fwd (squash/ _ _ p q i j) = + isOfHLevelMaybe 0 squash/ _ _ (cong fwd p) (cong fwd q) i j + + bwd : Maybe (S X / ρ (R .fst .fst)) → Maybe (S X) / MaybeRel (ρ (R .fst .fst)) + bwd nothing = [ nothing ] + bwd (just [ s ]) = [ just s ] + bwd (just (eq/ s t r i)) = eq/ (just s) (just t) r i + bwd (just (squash/ _ _ p q i j)) = + squash/ _ _ (cong (bwd ∘ just) p) (cong (bwd ∘ just) q) i j + + open Iso + isom : Iso (Maybe (S X) / MaybeRel (ρ (R .fst .fst))) (Maybe (S X / ρ (R .fst .fst))) + isom .fun = fwd + isom .inv = bwd + isom .rightInv nothing = refl + isom .rightInv (just x) = + elimProp {P = λ x → fwd (bwd (just x)) ≡ just x} + (λ _ → isOfHLevelMaybe 0 squash/ _ _) + (λ _ → refl) + x + isom .leftInv = elimProp (λ _ → squash/ _ _) (λ {nothing → refl; (just _) → refl}) + +maybeRelMatchesTransp : {S : Type ℓ → Type ℓ₁} + (ρ : StrRel S ℓ₁') (α : EquivAction S) + → StrRelMatchesEquiv ρ (EquivAction→StrEquiv α) + → StrRelMatchesEquiv (MaybeRelStr ρ) (EquivAction→StrEquiv (maybeEquivAction α)) +maybeRelMatchesTransp _ _ μ (X , nothing) (Y , nothing) _ = + isContr→Equiv (isOfHLevelLift 0 isContrUnit) isContr-nothing≡nothing +maybeRelMatchesTransp _ _ μ (X , nothing) (Y , just y) _ = + uninhabEquiv lower ¬nothing≡just +maybeRelMatchesTransp _ _ μ (X , just x) (Y , nothing) _ = + uninhabEquiv lower ¬just≡nothing +maybeRelMatchesTransp _ _ μ (X , just x) (Y , just y) e = + compEquiv (μ (X , x) (Y , y) e) (_ , isEmbedding-just _ _) diff --git a/Cubical/Structures/Relational/Parameterized.agda b/Cubical/Structures/Relational/Parameterized.agda new file mode 100644 index 0000000000..e051ddbbfd --- /dev/null +++ b/Cubical/Structures/Relational/Parameterized.agda @@ -0,0 +1,70 @@ +{- + +A parameterized family of structures S can be combined into a single structure: +X ↦ (a : A) → S a X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Parameterized where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.SetQuotients + +open import Cubical.Structures.Parameterized + +private + variable + ℓ ℓ₀ ℓ₁ ℓ₁' ℓ₁'' : Level + +-- Structured relations + +module _ (A : Type ℓ₀) where + + ParamRelStr : {S : A → Type ℓ → Type ℓ₁} + → (∀ a → StrRel (S a) ℓ₁') + → StrRel (ParamStructure A S) (ℓ-max ℓ₀ ℓ₁') + ParamRelStr ρ R s t = + (a : A) → ρ a R (s a) (t a) + + paramSuitableRel : {S : A → Type ℓ → Type ℓ₁} {ρ : ∀ a → StrRel (S a) ℓ₁'} + → (∀ a → SuitableStrRel (S a) (ρ a)) + → SuitableStrRel (ParamStructure A S) (ParamRelStr ρ) + paramSuitableRel {ρ = ρ} θ .quo (X , f) R r .fst .fst a = + θ a .quo (X , f a) R (r a) .fst .fst + paramSuitableRel {ρ = ρ} θ .quo (X , f) R r .fst .snd a = + θ a .quo (X , f a) R (r a) .fst .snd + paramSuitableRel {ρ = ρ} θ .quo (X , f) R r .snd (q , c) i .fst a = + θ a .quo (X , f a) R (r a) .snd (q a , c a) i .fst + paramSuitableRel {ρ = ρ} θ .quo (X , f) R r .snd (q , c) i .snd a = + θ a .quo (X , f a) R (r a) .snd (q a , c a) i .snd + paramSuitableRel {ρ = ρ} θ .symmetric R r a = + θ a .symmetric R (r a) + paramSuitableRel {ρ = ρ} θ .transitive R R' r r' a = + θ a .transitive R R' (r a) (r' a) + paramSuitableRel {ρ = ρ} θ .set setX = + isSetΠ λ a → θ a .set setX + paramSuitableRel {ρ = ρ} θ .prop propR s t = + isPropΠ λ a → θ a .prop propR (s a) (t a) + + paramRelMatchesEquiv : {S : A → Type ℓ → Type ℓ₁} + (ρ : ∀ a → StrRel (S a) ℓ₁') {ι : ∀ a → StrEquiv (S a) ℓ₁''} + → (∀ a → StrRelMatchesEquiv (ρ a) (ι a)) + → StrRelMatchesEquiv (ParamRelStr ρ) (ParamEquivStr A ι) + paramRelMatchesEquiv ρ μ _ _ e = equivΠCod λ a → μ a _ _ e + + paramRelAction : {S : A → Type ℓ → Type ℓ₁} {ρ : ∀ a → StrRel (S a) ℓ₁'} + → (∀ a → StrRelAction (ρ a)) + → StrRelAction (ParamRelStr ρ) + paramRelAction α .actStr f s a = α a .actStr f (s a) + paramRelAction α .actStrId s = funExt λ a → α a .actStrId (s a) + paramRelAction α .actRel h _ _ r a = α a .actRel h _ _ (r a) + + -- Detransitivity of ParamRelStr would depend on choice in general, so + -- we don't get positivity diff --git a/Cubical/Structures/Relational/Pointed.agda b/Cubical/Structures/Relational/Pointed.agda new file mode 100644 index 0000000000..1531eef109 --- /dev/null +++ b/Cubical/Structures/Relational/Pointed.agda @@ -0,0 +1,56 @@ +{- + +Pointed structure: X ↦ X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Pointed where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Foundations.Univalence +open import Cubical.Relation.ZigZag.Base +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Structures.Pointed + +private + variable + ℓ : Level + +-- Structured relations + +PointedRelStr : StrRel PointedStructure ℓ +PointedRelStr R = R + +pointedSuitableRel : SuitableStrRel {ℓ = ℓ} PointedStructure PointedRelStr +pointedSuitableRel .quo _ _ _ = isContrSingl _ +pointedSuitableRel .symmetric _ r = r +pointedSuitableRel .transitive _ _ r r' = ∣ _ , r , r' ∣ +pointedSuitableRel .set setX = setX +pointedSuitableRel .prop propR = propR + +pointedRelMatchesEquiv : StrRelMatchesEquiv {ℓ = ℓ} PointedRelStr PointedEquivStr +pointedRelMatchesEquiv _ _ _ = idEquiv _ + +pointedRelAction : StrRelAction {ℓ = ℓ} PointedRelStr +pointedRelAction .actStr f = f +pointedRelAction .actStrId _ = refl +pointedRelAction .actRel α = α + +pointedPositiveRel : PositiveStrRel {ℓ = ℓ} pointedSuitableRel +pointedPositiveRel .act = pointedRelAction +pointedPositiveRel .reflexive x = ∣ refl ∣ +pointedPositiveRel .detransitive R R' rr' = rr' +pointedPositiveRel .quo R = isoToIsEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = _ + isom .inv q = q + isom .rightInv = elimProp (λ _ → squash/ _ _) (λ _ → refl) + isom .leftInv = elimProp (λ _ → squash/ _ _) (λ _ → refl) diff --git a/Cubical/Structures/Relational/Product.agda b/Cubical/Structures/Relational/Product.agda new file mode 100644 index 0000000000..b2148f36ed --- /dev/null +++ b/Cubical/Structures/Relational/Product.agda @@ -0,0 +1,140 @@ +{- + +Product of structures S and T: X ↦ S X × T X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Product where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Univalence +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.HITs.SetQuotients + +open import Cubical.Structures.Product + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₁'' ℓ₂ ℓ₂' ℓ₂'' : Level + +-- Structured relations + +ProductRelStr : + {S₁ : Type ℓ → Type ℓ₁} (ρ₁ : StrRel S₁ ℓ₁') + {S₂ : Type ℓ → Type ℓ₂} (ρ₂ : StrRel S₂ ℓ₂') + → StrRel (ProductStructure S₁ S₂) (ℓ-max ℓ₁' ℓ₂') +ProductRelStr ρ₁ ρ₂ R (s₁ , s₂) (t₁ , t₂) = + ρ₁ R s₁ t₁ × ρ₂ R s₂ t₂ + +productSuitableRel : + {S₁ : Type ℓ → Type ℓ₁} {ρ₁ : StrRel S₁ ℓ₁'} + {S₂ : Type ℓ → Type ℓ₂} {ρ₂ : StrRel S₂ ℓ₂'} + → SuitableStrRel S₁ ρ₁ → SuitableStrRel S₂ ρ₂ + → SuitableStrRel (ProductStructure S₁ S₂) (ProductRelStr ρ₁ ρ₂) +productSuitableRel θ₁ θ₂ .quo (X , s₁ , s₂) R (r₁ , r₂) .fst .fst = + θ₁ .quo (X , s₁) R r₁ .fst .fst , θ₂ .quo (X , s₂) R r₂ .fst .fst +productSuitableRel θ₁ θ₂ .quo (X , s₁ , s₂) R (r₁ , r₂) .fst .snd = + θ₁ .quo (X , s₁) R r₁ .fst .snd , θ₂ .quo (X , s₂) R r₂ .fst .snd +productSuitableRel θ₁ θ₂ .quo (X , s₁ , s₂) R (r₁ , r₂) .snd ((q₁ , q₂) , (c₁ , c₂)) i .fst = + θ₁ .quo (X , s₁) R r₁ .snd (q₁ , c₁) i .fst , θ₂ .quo (X , s₂) R r₂ .snd (q₂ , c₂) i .fst +productSuitableRel θ₁ θ₂ .quo (X , s₁ , s₂) R (r₁ , r₂) .snd ((q₁ , q₂) , (c₁ , c₂)) i .snd = + θ₁ .quo (X , s₁) R r₁ .snd (q₁ , c₁) i .snd , θ₂ .quo (X , s₂) R r₂ .snd (q₂ , c₂) i .snd +productSuitableRel θ₁ θ₂ .symmetric R (r₁ , r₂) = + θ₁ .symmetric R r₁ , θ₂ .symmetric R r₂ +productSuitableRel θ₁ θ₂ .transitive R R' (r₁ , r₂) (r₁' , r₂') = + θ₁ .transitive R R' r₁ r₁' , θ₂ .transitive R R' r₂ r₂' +productSuitableRel θ₁ θ₂ .set setX = + isSet× (θ₁ .set setX) (θ₂ .set setX) +productSuitableRel θ₁ θ₂ .prop propR (s₁ , s₂) (t₁ , t₂) = + isProp× (θ₁ .prop propR s₁ t₁) (θ₂ .prop propR s₂ t₂) + +productRelMatchesEquiv : + {S₁ : Type ℓ → Type ℓ₁} (ρ₁ : StrRel S₁ ℓ₁') {ι₁ : StrEquiv S₁ ℓ₁''} + {S₂ : Type ℓ → Type ℓ₂} (ρ₂ : StrRel S₂ ℓ₂') {ι₂ : StrEquiv S₂ ℓ₂''} + → StrRelMatchesEquiv ρ₁ ι₁ → StrRelMatchesEquiv ρ₂ ι₂ + → StrRelMatchesEquiv (ProductRelStr ρ₁ ρ₂) (ProductEquivStr ι₁ ι₂) +productRelMatchesEquiv ρ₁ ρ₂ μ₁ μ₂ A B e = + Σ-cong-equiv (μ₁ _ _ e) (λ _ → μ₂ _ _ e) + +productRelAction : + {S₁ : Type ℓ → Type ℓ₁} {ρ₁ : StrRel S₁ ℓ₁'} (α₁ : StrRelAction ρ₁) + {S₂ : Type ℓ → Type ℓ₂} {ρ₂ : StrRel S₂ ℓ₂'} (α₂ : StrRelAction ρ₂) + → StrRelAction (ProductRelStr ρ₁ ρ₂) +productRelAction α₁ α₂ .actStr f (s₁ , s₂) = α₁ .actStr f s₁ , α₂ .actStr f s₂ +productRelAction α₁ α₂ .actStrId (s₁ , s₂) = ΣPathP (α₁ .actStrId s₁ , α₂ .actStrId s₂) +productRelAction α₁ α₂ .actRel h _ _ (r₁ , r₂) = α₁ .actRel h _ _ r₁ , α₂ .actRel h _ _ r₂ + +productPositiveRel : + {S₁ : Type ℓ → Type ℓ₁} {ρ₁ : StrRel S₁ ℓ₁'} {θ₁ : SuitableStrRel S₁ ρ₁} + {S₂ : Type ℓ → Type ℓ₂} {ρ₂ : StrRel S₂ ℓ₂'} {θ₂ : SuitableStrRel S₂ ρ₂} + → PositiveStrRel θ₁ + → PositiveStrRel θ₂ + → PositiveStrRel (productSuitableRel θ₁ θ₂) +productPositiveRel σ₁ σ₂ .act = productRelAction (σ₁ .act) (σ₂ .act) +productPositiveRel σ₁ σ₂ .reflexive (s₁ , s₂) = σ₁ .reflexive s₁ , σ₂ .reflexive s₂ +productPositiveRel σ₁ σ₂ .detransitive R R' (rr'₁ , rr'₂) = + Trunc.rec squash + (λ {(s₁ , r₁ , r₁') → + Trunc.rec squash + (λ {(s₂ , r₂ , r₂') → ∣ (s₁ , s₂) , (r₁ , r₂) , (r₁' , r₂') ∣}) + (σ₂ .detransitive R R' rr'₂)}) + (σ₁ .detransitive R R' rr'₁) +productPositiveRel {S₁ = S₁} {ρ₁} {θ₁} {S₂} {ρ₂} {θ₂} σ₁ σ₂ .quo {X} R = + subst isEquiv + (funExt (elimProp (λ _ → productSuitableRel θ₁ θ₂ .set squash/ _ _) (λ _ → refl))) + (compEquiv + (isoToEquiv isom) + (Σ-cong-equiv (_ , σ₁ .quo R) (λ _ → _ , σ₂ .quo R)) .snd) + where + fwd : + ProductStructure S₁ S₂ X / ProductRelStr ρ₁ ρ₂ (R .fst .fst) + → (S₁ X / ρ₁ (R .fst .fst)) × (S₂ X / ρ₂ (R .fst .fst)) + fwd [ s₁ , s₂ ] = [ s₁ ] , [ s₂ ] + fwd (eq/ (s₁ , s₂) (t₁ , t₂) (r₁ , r₂) i) = eq/ s₁ t₁ r₁ i , eq/ s₂ t₂ r₂ i + fwd (squash/ _ _ p q i j) = + isSet× squash/ squash/ _ _ (cong fwd p) (cong fwd q) i j + + bwd[] : S₁ X → S₂ X / ρ₂ (R .fst .fst) + → ProductStructure S₁ S₂ X / ProductRelStr ρ₁ ρ₂ (R .fst .fst) + bwd[] s₁ [ s₂ ] = [ s₁ , s₂ ] + bwd[] s₁ (eq/ s₂ t₂ r₂ i) = + eq/ (s₁ , s₂) (s₁ , t₂) (posRelReflexive σ₁ R s₁ , r₂) i + bwd[] s₁ (squash/ _ _ p q i j) = + squash/ _ _ (λ j → bwd[] s₁ (p j)) (λ j → bwd[] s₁ (q j)) i j + + bwd : S₁ X / ρ₁ (R .fst .fst) → S₂ X / ρ₂ (R .fst .fst) + → ProductStructure S₁ S₂ X / ProductRelStr ρ₁ ρ₂ (R .fst .fst) + bwd [ s₁ ] u = bwd[] s₁ u + bwd (eq/ s₁ t₁ r₁ i) u = path u i + where + path : ∀ u → bwd [ s₁ ] u ≡ bwd [ t₁ ] u + path = elimProp (λ _ → squash/ _ _) (λ s₂ → eq/ (s₁ , s₂) (t₁ , s₂) (r₁ , posRelReflexive σ₂ R s₂)) + bwd (squash/ _ _ p q i j) = + isSetΠ (λ _ → squash/) _ _ (cong bwd p) (cong bwd q) i j + + open Iso + isom : Iso _ _ + isom .fun = fwd + isom .inv = uncurry bwd + isom .rightInv = + uncurry + (elimProp (λ _ → isPropΠ λ _ → isSet× squash/ squash/ _ _) + (λ _ → elimProp (λ _ → isSet× squash/ squash/ _ _) (λ _ → refl))) + isom .leftInv = elimProp (λ _ → squash/ _ _) (λ _ → refl) + +productRelMatchesTransp : + {S₁ : Type ℓ → Type ℓ₁} (ρ₁ : StrRel S₁ ℓ₁') (α₁ : EquivAction S₁) + {S₂ : Type ℓ → Type ℓ₂} (ρ₂ : StrRel S₂ ℓ₂') (α₂ : EquivAction S₂) + → StrRelMatchesEquiv ρ₁ (EquivAction→StrEquiv α₁) + → StrRelMatchesEquiv ρ₂ (EquivAction→StrEquiv α₂) + → StrRelMatchesEquiv (ProductRelStr ρ₁ ρ₂) (EquivAction→StrEquiv (productEquivAction α₁ α₂)) +productRelMatchesTransp _ _ _ _ μ₁ μ₂ _ _ e = + compEquiv (Σ-cong-equiv (μ₁ _ _ e) (λ _ → μ₂ _ _ e)) ΣPath≃PathΣ diff --git a/Cubical/Structures/Successor.agda b/Cubical/Structures/Successor.agda new file mode 100644 index 0000000000..ce72551c60 --- /dev/null +++ b/Cubical/Structures/Successor.agda @@ -0,0 +1,29 @@ +{-# OPTIONS --safe #-} +{- + Successor structures for spectra, chain complexes and fiber sequences. + This is an idea from Floris van Doorn's phd thesis. +-} +module Cubical.Structures.Successor where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Int +open import Cubical.Data.Nat + +private + variable + ℓ : Level + +record SuccStr (ℓ : Level) : Type (ℓ-suc ℓ) where + field + Index : Type ℓ + succ : Index → Index + +open SuccStr + +ℤ+ : SuccStr ℓ-zero +ℤ+ .Index = ℤ +ℤ+ .succ = sucℤ + +ℕ+ : SuccStr ℓ-zero +ℕ+ .Index = ℕ +ℕ+ .succ = suc diff --git a/Cubical/Structures/Transfer.agda b/Cubical/Structures/Transfer.agda new file mode 100644 index 0000000000..809fbff536 --- /dev/null +++ b/Cubical/Structures/Transfer.agda @@ -0,0 +1,47 @@ +{- + +Transferring properties of terms between equivalent structures + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Transfer where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Transport +open import Cubical.Structures.Product + +private + variable + ℓ ℓ₀ ℓ₁ ℓ₂ ℓ₂' : Level + +transfer : {ℓ₂' ℓ₀ : Level} {S : Type ℓ → Type ℓ₁} {H : Type ℓ → Type ℓ₂} + (P : ∀ X → S X → H X → Type ℓ₀) + (α : EquivAction H) (τ : TransportStr α) + (ι : StrEquiv S ℓ₂') (θ : UnivalentStr S ι) + {X Y : Type ℓ} {s : S X} {t : S Y} + (e : (X , s) ≃[ ι ] (Y , t)) + → (h : H Y) → P X s (invEq (α (e .fst)) h) → P Y t h +transfer P α τ ι θ {X} {Y} {s} {t} e h = + subst (λ {(Z , u , h) → P Z u h}) + (sip + (productUnivalentStr ι θ (EquivAction→StrEquiv α) (TransportStr→UnivalentStr α τ)) + (X , s , invEq (α (e .fst)) h) + (Y , t , h) + (e .fst , e .snd , secEq (α (e .fst)) h)) + +transfer⁻ : {ℓ₂' ℓ₀ : Level} {S : Type ℓ → Type ℓ₁} {H : Type ℓ → Type ℓ₂} + (P : ∀ X → S X → H X → Type ℓ₀) + (α : EquivAction H) (τ : TransportStr α) + (ι : StrEquiv S ℓ₂') (θ : UnivalentStr S ι) + {X Y : Type ℓ} {s : S X} {t : S Y} + (e : (X , s) ≃[ ι ] (Y , t)) + → (h : H X) → P Y t (equivFun (α (e .fst)) h) → P X s h +transfer⁻ P α τ ι θ {X} {Y} {s} {t} e h = + subst⁻ (λ {(Z , u , h) → P Z u h}) + (sip + (productUnivalentStr ι θ (EquivAction→StrEquiv α) (TransportStr→UnivalentStr α τ)) + (X , s , h) + (Y , t , equivFun (α (e .fst)) h) + (e .fst , e .snd , refl)) diff --git a/Cubical/Structures/TypeEqvTo.agda b/Cubical/Structures/TypeEqvTo.agda new file mode 100644 index 0000000000..e142bb9135 --- /dev/null +++ b/Cubical/Structures/TypeEqvTo.agda @@ -0,0 +1,41 @@ +{-# OPTIONS --safe #-} +module Cubical.Structures.TypeEqvTo where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.HITs.PropositionalTruncation +open import Cubical.Data.Sigma + +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Pointed +open import Cubical.Structures.Axioms +open import Cubical.Structures.Pointed + +private + variable + ℓ ℓ' ℓ'' : Level + +TypeEqvTo : (ℓ : Level) (X : Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +TypeEqvTo ℓ X = TypeWithStr ℓ (λ Y → ∥ Y ≃ X ∥) + +PointedEqvTo : (ℓ : Level) (X : Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +PointedEqvTo ℓ X = TypeWithStr ℓ (λ Y → Y × ∥ Y ≃ X ∥) + +module _ (X : Type ℓ') where + + PointedEqvToStructure : Type ℓ → Type (ℓ-max ℓ ℓ') + PointedEqvToStructure = AxiomsStructure PointedStructure (λ Y _ → ∥ Y ≃ X ∥) + + PointedEqvToEquivStr : StrEquiv PointedEqvToStructure ℓ'' + PointedEqvToEquivStr = AxiomsEquivStr PointedEquivStr (λ Y _ → ∥ Y ≃ X ∥) + + pointedEqvToUnivalentStr : UnivalentStr {ℓ} PointedEqvToStructure PointedEqvToEquivStr + pointedEqvToUnivalentStr = axiomsUnivalentStr PointedEquivStr {axioms = λ Y _ → ∥ Y ≃ X ∥} + (λ _ _ → squash) pointedUnivalentStr + + PointedEqvToSIP : (A B : PointedEqvTo ℓ X) → A ≃[ PointedEqvToEquivStr ] B ≃ (A ≡ B) + PointedEqvToSIP = SIP pointedEqvToUnivalentStr + + PointedEqvTo-sip : (A B : PointedEqvTo ℓ X) → A ≃[ PointedEqvToEquivStr ] B → (A ≡ B) + PointedEqvTo-sip A B = equivFun (PointedEqvToSIP A B) diff --git "a/Cubical/Syntax/\342\237\250\342\237\251.agda" "b/Cubical/Syntax/\342\237\250\342\237\251.agda" new file mode 100644 index 0000000000..03496011b7 --- /dev/null +++ "b/Cubical/Syntax/\342\237\250\342\237\251.agda" @@ -0,0 +1,12 @@ +{-# OPTIONS --safe #-} +module Cubical.Syntax.⟨⟩ where + +open import Cubical.Core.Primitives + +-- A syntax typeclass for types which contain a "carrier" type in the +-- sence of an algebraic structure. +record has-⟨⟩ {ℓᵢ ℓc} (Instance : Type ℓᵢ) : Type (ℓ-max ℓᵢ (ℓ-suc ℓc)) where + field + ⟨_⟩ : Instance → Type ℓc + +open has-⟨⟩ ⦃ … ⦄ public diff --git a/Cubical/Talks/EPA2020.agda b/Cubical/Talks/EPA2020.agda new file mode 100644 index 0000000000..dc1f2e0b54 --- /dev/null +++ b/Cubical/Talks/EPA2020.agda @@ -0,0 +1,343 @@ +{- + +Cubical Agda - A Dependently Typed PL with Univalence and HITs +============================================================== + Anders Mörtberg + Every Proof Assistant - September 17, 2020 + + +Link to slides: https://staff.math.su.se/anders.mortberg/slides/EPA2020.pdf + +Link to video: https://vimeo.com/459020971 + +-} + +-- To make Agda cubical add the --cubical option. +-- This is implicitly added for files in the cubical library via the cubical.agda-lib file. +{-# OPTIONS --safe #-} +module Cubical.Talks.EPA2020 where + +-- The "Foundations" package contain a lot of important results (in +-- particular the univalence theorem) +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Int +open import Cubical.Data.Prod.Base + + +-- The interval in Cubical Agda is written I and the endpoints i0 and i1. + +apply0 : (A : Type) (p : I → A) → A +apply0 A p = p i0 + +-- We omit the universe level ℓ for simplicity in this talk. In the +-- library everything is maximally universe polymorphic. + + +-- We can write functions out of the interval using λ-abstraction: +refl' : {A : Type} (x : A) → x ≡ x +refl' x = λ i → x +-- In fact, x ≡ y is short for PathP (λ _ → A) x y + +refl'' : {A : Type} (x : A) → PathP (λ _ → A) x x +refl'' x = λ _ → x + +-- In general PathP A x y has A : I → Type, x : A i0 and y : A i1 +-- PathP = Dependent paths (Path over a Path) + +-- We cannot pattern-match on interval variables as I is not inductively defined +-- foo : {A : Type} → I → A +-- foo r = {!r!} -- Try typing C-c C-c + +-- cong has a direct proof +cong' : {A B : Type} (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +cong' f p i = f (p i) + +-- function extensionality also has a direct proof. +-- It also has computational content: swap the arguments. +funExt' : {A B : Type} {f g : A → B} (p : (x : A) → f x ≡ g x) → f ≡ g +funExt' p i x = p x i + +-- Transport is more complex as ≡ isn't inductively defined (so we +-- can't define it by pattern-matching on p) +transport' : {A B : Type} → A ≡ B → A → B +transport' p a = transp (λ i → p i) i0 a + +-- This lets us define subst (which is called "transport" in the HoTT book) +subst' : {A : Type} (P : A → Type) {x y : A} (p : x ≡ y) → P x → P y +subst' P p pa = transport (λ i → P (p i)) pa + +-- The transp operation reduces differently for different types +-- formers. For paths it reduces to another primitive operation called +-- hcomp. + +-- We can also define the J eliminator (aka path induction) +J' : {A : Type} {B : A → Type} {x : A} + (P : (z : A) → x ≡ z → Type) + (d : P x refl) {y : A} (p : x ≡ y) → P y p +J' P d p = transport (λ i → P (p i) (λ j → p (i ∧ j))) d + +-- So J is provable, but it doesn't satisfy computation rule +-- definitionally. This is almost never a problem in practice as the +-- cubical primitives satisfy many new definitional equalities. + + + +-- Another key concept in HoTT/UF is the Univalence Axiom. In Cubical +-- Agda this is provable, we hence refer to it as the Univalence +-- Theorem. + +-- The univalence theorem: equivalences of types give paths of types +ua' : {A B : Type} → A ≃ B → A ≡ B +ua' = ua + +-- Any isomorphism of types gives rise to an equivalence +isoToEquiv' : {A B : Type} → Iso A B → A ≃ B +isoToEquiv' = isoToEquiv + +-- And hence to a path +isoToPath' : {A B : Type} → Iso A B → A ≡ B +isoToPath' e = ua' (isoToEquiv' e) + +-- ua satisfies the following computation rule +-- This suffices to be able to prove the standard formulation of univalence. +uaβ' : {A B : Type} (e : A ≃ B) (x : A) + → transport (ua' e) x ≡ fst e x +uaβ' e x = transportRefl (equivFun e x) + + + +-- Time for an example! + +-- Booleans +data Bool : Type where + false true : Bool + +not : Bool → Bool +not false = true +not true = false + +notPath : Bool ≡ Bool +notPath = isoToPath' (iso not not rem rem) + where + rem : (b : Bool) → not (not b) ≡ b + rem false = refl + rem true = refl + +_ : transport notPath true ≡ false +_ = refl + + +-- Another example, integers: + +sucPath : ℤ ≡ ℤ +sucPath = isoToPath' (iso sucℤ predℤ sucPred predSuc) + +_ : transport sucPath (pos 0) ≡ pos 1 +_ = refl + +_ : transport (sucPath ∙ sucPath) (pos 0) ≡ pos 2 +_ = refl + +_ : transport (sym sucPath) (pos 0) ≡ negsuc 0 +_ = refl + + + +----------------------------------------------------------------------------- +-- Higher inductive types + +-- The following definition of finite multisets is due to Vikraman +-- Choudhury and Marcelo Fiore. + +infixr 5 _∷_ + +data FMSet (A : Type) : Type where + [] : FMSet A + _∷_ : (x : A) → (xs : FMSet A) → FMSet A + comm : (x y : A) (xs : FMSet A) → x ∷ y ∷ xs ≡ y ∷ x ∷ xs +-- trunc : (xs ys : FMSet A) (p q : xs ≡ ys) → p ≡ q + +-- We need to add the trunc constructor for FMSets to be sets, omitted +-- here for simplicity. + +_++_ : ∀ {A : Type} (xs ys : FMSet A) → FMSet A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ xs ++ ys +comm x y xs i ++ ys = comm x y (xs ++ ys) i +-- trunc xs zs p q i j ++ ys = +-- trunc (xs ++ ys) (zs ++ ys) (cong (_++ ys) p) (cong (_++ ys) q) i j + +unitr-++ : {A : Type} (xs : FMSet A) → xs ++ [] ≡ xs +unitr-++ [] = refl +unitr-++ (x ∷ xs) = cong (x ∷_) (unitr-++ xs) +unitr-++ (comm x y xs i) j = comm x y (unitr-++ xs j) i +-- unitr-++ (trunc xs ys x y i j) = {!!} + + +-- This is a special case of set quotients! Very useful for +-- programming and set level mathematics + +data _/_ (A : Type) (R : A → A → Type) : Type where + [_] : A → A / R + eq/ : (a b : A) → R a b → [ a ] ≡ [ b ] + trunc : (a b : A / R) (p q : a ≡ b) → p ≡ q + +-- Proving that they are effective ((a b : A) → [ a ] ≡ [ b ] → R a b) +-- requires univalence for propositions. + + +------------------------------------------------------------------------- +-- Topological examples of things that are not sets + +-- We can define the circle as the following simple data declaration: +data S¹ : Type where + base : S¹ + loop : base ≡ base + +-- We can write functions on S¹ using pattern-matching equations: +double : S¹ → S¹ +double base = base +double (loop i) = (loop ∙ loop) i + +helix : S¹ → Type +helix base = ℤ +helix (loop i) = sucPathℤ i + +ΩS¹ : Type +ΩS¹ = base ≡ base + +winding : ΩS¹ → ℤ +winding p = subst helix p (pos 0) + +_ : winding (λ i → double ((loop ∙ loop) i)) ≡ pos 4 +_ = refl + + +-- We can define the Torus as: +data Torus : Type where + point : Torus + line1 : point ≡ point + line2 : point ≡ point + square : PathP (λ i → line1 i ≡ line1 i) line2 line2 + +-- And prove that it is equivalent to two circle: +t2c : Torus → S¹ × S¹ +t2c point = (base , base) +t2c (line1 i) = (loop i , base) +t2c (line2 j) = (base , loop j) +t2c (square i j) = (loop i , loop j) + +c2t : S¹ × S¹ → Torus +c2t (base , base) = point +c2t (loop i , base) = line1 i +c2t (base , loop j) = line2 j +c2t (loop i , loop j) = square i j + +c2t-t2c : (t : Torus) → c2t (t2c t) ≡ t +c2t-t2c point = refl +c2t-t2c (line1 _) = refl +c2t-t2c (line2 _) = refl +c2t-t2c (square _ _) = refl + +t2c-c2t : (p : S¹ × S¹) → t2c (c2t p) ≡ p +t2c-c2t (base , base) = refl +t2c-c2t (base , loop _) = refl +t2c-c2t (loop _ , base) = refl +t2c-c2t (loop _ , loop _) = refl + +-- Using univalence we get the following equality: +Torus≡S¹×S¹ : Torus ≡ S¹ × S¹ +Torus≡S¹×S¹ = isoToPath' (iso t2c c2t t2c-c2t c2t-t2c) + + +windingTorus : point ≡ point → ℤ × ℤ +windingTorus l = ( winding (λ i → proj₁ (t2c (l i))) + , winding (λ i → proj₂ (t2c (l i)))) + +_ : windingTorus (line1 ∙ sym line2) ≡ (pos 1 , negsuc 0) +_ = refl + +-- We have many more topological examples, including Klein bottle, RP^n, +-- higher spheres, suspensions, join, wedges, smash product: +open import Cubical.HITs.KleinBottle +open import Cubical.HITs.RPn +open import Cubical.HITs.S2 +open import Cubical.HITs.S3 +open import Cubical.HITs.Susp +open import Cubical.HITs.Join +open import Cubical.HITs.Wedge +open import Cubical.HITs.SmashProduct + +-- There's also a proof of the "3x3 lemma" for pushouts in less than +-- 200LOC. In HoTT-Agda this took about 3000LOC. For details see: +-- https://github.com/HoTT/HoTT-Agda/tree/master/theorems/homotopy/3x3 +open import Cubical.HITs.Pushout + +-- We also defined the Hopf fibration and proved that its total space +-- is S³ in about 300LOC: +open import Cubical.Homotopy.Hopf +open S¹Hopf + +-- There is also some integer cohomology: +open import Cubical.ZCohomology.Everything +-- To compute cohomology groups of various spaces we need a bunch of +-- interesting theorems: Freudenthal suspension theorem, +-- Mayer-Vietoris sequence... +open import Cubical.Homotopy.Freudenthal +open import Cubical.ZCohomology.MayerVietorisUnreduced + + +------------------------------------------------------------------------- +-- The structure identity principle + +-- A more efficient version of finite multisets based on association lists +open import Cubical.HITs.AssocList.Base + +-- data AssocList (A : Type) : Type where +-- ⟨⟩ : AssocList A +-- ⟨_,_⟩∷_ : (a : A) (n : ℕ) (xs : AssocList A) → AssocList A +-- per : (a b : A) (m n : ℕ) (xs : AssocList A) +-- → ⟨ a , m ⟩∷ ⟨ b , n ⟩∷ xs ≡ ⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs +-- agg : (a : A) (m n : ℕ) (xs : AssocList A) +-- → ⟨ a , m ⟩∷ ⟨ a , n ⟩∷ xs ≡ ⟨ a , m + n ⟩∷ xs +-- del : (a : A) (xs : AssocList A) → ⟨ a , 0 ⟩∷ xs ≡ xs +-- trunc : (xs ys : AssocList A) (p q : xs ≡ ys) → p ≡ q + + +-- Programming and proving is more complicated with AssocList compared +-- to FMSet. This kind of example occurs everywhere in programming and +-- mathematics: one representation is easier to work with, but not +-- efficient, while another is efficient but difficult to work with. + +-- Solution: substitute using univalence +substIso : {A B : Type} (P : Type → Type) (e : Iso A B) → P A → P B +substIso P e = subst P (isoToPath e) + +-- Can transport for example Monoid structure from FMSet to AssocList +-- this way, but the achieved Monoid structure is not very efficient +-- to work with. A better solution is to prove that FMSet and +-- AssocList are equal *as monoids*, but how to do this? + +-- Solution: structure identity principle (SIP) +-- This is a very useful consequence of univalence +open import Cubical.Foundations.SIP + +sip' : {ℓ : Level} {S : Type ℓ → Type ℓ} {ι : StrEquiv S ℓ} + (θ : UnivalentStr S ι) (A B : TypeWithStr ℓ S) → A ≃[ ι ] B → A ≡ B +sip' = sip + +-- The tricky thing is to prove that (S,ι) is a univalent structure. +-- Luckily we provide automation for this in the library, see for example: +open import Cubical.Algebra.Monoid.Base + +-- Another cool application of the SIP: matrices represented as +-- functions out of pairs of Fin's and vectors are equal as abelian +-- groups: +open import Cubical.Algebra.Matrix + + +-- The end, back to slides! diff --git a/Cubical/WithK.agda b/Cubical/WithK.agda new file mode 100644 index 0000000000..b0140999a7 --- /dev/null +++ b/Cubical/WithK.agda @@ -0,0 +1,48 @@ +{- Cubical Agda with K + +This file demonstrates the incompatibility of the --cubical +and --with-K flags, relying on the well-known incosistency of K with +univalence. + +The --safe flag can be used to prevent accidentally mixing such +incompatible flags. + +-} + +{-# OPTIONS --with-K #-} + +module Cubical.WithK where + +open import Cubical.Data.Equality +open import Cubical.Data.Bool +open import Cubical.Data.Empty + + +private + variable + ℓ : Level + A : Type ℓ + x y : A + +uip : (prf : x ≡p x) → prf ≡c reflp +uip reflp i = reflp + +transport-uip : (prf : A ≡p A) → transport (ptoc prf) x ≡c x +transport-uip {x = x} prf = + cong (λ m → transport (ptoc m) x) (uip prf) ∙ transportRefl x + +transport-not : transport (ptoc (ctop notEq)) true ≡c false +transport-not = cong (λ a → transport a true) (ptoc-ctop notEq) + +false-true : false ≡c true +false-true = sym transport-not ∙ transport-uip (ctop notEq) + +absurd : (X : Type) → X +absurd X = transport (cong sel false-true) true + where + sel : Bool → Type + sel false = Bool + sel true = X + +inconsistency : ⊥ +inconsistency = absurd ⊥ diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda new file mode 100644 index 0000000000..da01224fad --- /dev/null +++ b/Cubical/ZCohomology/Base.agda @@ -0,0 +1,51 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Base where + +open import Cubical.Data.Int.Base hiding (_+_) +open import Cubical.Data.Nat.Base +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed.Base + +open import Cubical.HITs.Nullification.Base +open import Cubical.HITs.SetTruncation.Base +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.Susp.Base +open import Cubical.HITs.Truncation.Base +open import Cubical.Homotopy.Loopspace + +private + variable + ℓ : Level + A : Type ℓ + +--- Cohomology --- + +{- EM-spaces Kₙ from Brunerie 2016 -} +coHomK : (n : ℕ) → Type₀ +coHomK zero = ℤ +coHomK (suc n) = ∥ S₊ (suc n) ∥ (2 + suc n) + +{- Cohomology -} +coHom : (n : ℕ) → Type ℓ → Type ℓ +coHom n A = ∥ (A → coHomK n) ∥₂ + +--- Reduced cohomology --- + +coHom-pt : (n : ℕ) → coHomK n +coHom-pt 0 = 0 +coHom-pt (suc n) = ∣ (ptSn (suc n)) ∣ + +{- Pointed version of Kₙ -} +coHomK-ptd : (n : ℕ) → Pointed (ℓ-zero) +coHomK-ptd n = coHomK n , coHom-pt n + +{- Reduced cohomology -} +coHomRed : (n : ℕ) → (A : Pointed ℓ) → Type ℓ +coHomRed n A = ∥ A →∙ coHomK-ptd n ∥₂ + +{- Kₙ, untruncated version -} +coHomKType : (n : ℕ) → Type +coHomKType zero = ℤ +coHomKType (suc n) = S₊ (suc n) diff --git a/Cubical/ZCohomology/EilenbergSteenrodZ.agda b/Cubical/ZCohomology/EilenbergSteenrodZ.agda new file mode 100644 index 0000000000..26b58b7788 --- /dev/null +++ b/Cubical/ZCohomology/EilenbergSteenrodZ.agda @@ -0,0 +1,379 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} + +module Cubical.ZCohomology.EilenbergSteenrodZ where + +open import Cubical.Homotopy.EilenbergSteenrod + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Int +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.HITs.SetTruncation + renaming (map to sMap ; rec to sRec ; elim2 to sElim2; elim to sElim ; rec2 to sRec2) +open import Cubical.HITs.PropositionalTruncation + renaming (map to pMap ; rec to pRec ; elim to pElim ; elim2 to pElim2) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; map to trMap ; elim to trEilim ; elim2 to trElim2) +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.Wedge +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 + +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup) +open import Cubical.Algebra.AbGroup + +open coHomTheory +open Iso +open IsGroup +open GroupStr +open IsGroupHom + +private + suspΩFun' : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (f : A → Path _ (0ₖ _) (0ₖ _)) + → Susp A → coHomK (suc n) + suspΩFun' n f north = 0ₖ _ + suspΩFun' n f south = 0ₖ _ + suspΩFun' n f (merid a i) = f a i + + suspΩFun : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (f : A → coHomK n) → Susp A → coHomK (suc n) + suspΩFun n f = suspΩFun' n λ a → Kn→ΩKn+1 n (f a) + + ≡suspΩFun : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) (f : Susp (typ A) → coHomK (suc n)) (p : f north ≡ 0ₖ _) + → f ≡ suspΩFun' n λ a → sym p ∙∙ (cong f (merid a)) ∙∙ (cong f (sym (merid (pt A))) ∙ p) + ≡suspΩFun n s f p i north = p i + ≡suspΩFun n s f p i south = (cong f (sym (merid (pt s))) ∙ p) i + ≡suspΩFun n s f p i (merid a j) = + doubleCompPath-filler (sym p) (cong f (merid a)) (cong f (sym (merid (pt s))) ∙ p) i j + +-- induction principle for Hⁿ(Susp A) +SuspCohomElim : ∀ {ℓ ℓ'} {A : Pointed ℓ} (n : ℕ) {B : coHom (suc n) (Susp (typ A)) → Type ℓ'} + → ((x : coHom (suc n) (Susp (typ A))) → isProp (B x)) + → ((f : typ A → Path _ (0ₖ _) (0ₖ _)) → f (pt A) ≡ refl → B ∣ suspΩFun' n f ∣₂) + → (x : _) → B x +SuspCohomElim {A = A} n {B = B} isprop f = + coHomPointedElim _ north isprop λ g gid + → subst (B ∘ ∣_∣₂) (sym (≡suspΩFun n A g gid)) + (f _ ((λ i → (sym gid ∙∙ (λ j → g (merid (pt A) (~ i ∧ j))) ∙∙ ((λ j → g (merid (pt A) (~ i ∧ ~ j))) ∙ gid))) + ∙∙ (λ i → (λ j → gid (i ∨ ~ j)) ∙∙ refl ∙∙ lUnit (λ j → gid (i ∨ j)) (~ i)) + ∙∙ sym (rUnit refl))) + +-- (Reduced) cohomology functor +coHomFunctor : {ℓ : Level} (n : ℤ) → Pointed ℓ → AbGroup ℓ +coHomFunctor (pos n) = coHomRedGroup n +coHomFunctor (negsuc n) _ = trivialAbGroup + +-- Alternative definition with reduced groups replaced by unreduced one for n ≥ 1 +coHomFunctor' : {ℓ : Level} (n : ℤ) → Pointed ℓ → AbGroup ℓ +coHomFunctor' (pos zero) = coHomFunctor 0 +coHomFunctor' (pos (suc n)) A = coHomGroup (suc n) (typ A) +coHomFunctor' (negsuc n) = coHomFunctor (negsuc n) + +coHomFunctor≡coHomFunctor' : ∀ {ℓ} → coHomFunctor {ℓ} ≡ coHomFunctor' +coHomFunctor≡coHomFunctor' = funExt λ {(pos zero) → refl + ; (pos (suc n)) → funExt λ A → sym (coHomGroup≡coHomRedGroup n A) + ; (negsuc n) → refl} + +-- Ĥ⁰(Susp A) is contractible +H0-susp : ∀ {ℓ} {A : Pointed ℓ} → isContr (coHomRed 0 (Susp (typ A) , north)) +fst H0-susp = 0ₕ∙ _ +snd (H0-susp {A = A}) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , p) + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (funExt λ {north → sym p + ; south → sym p ∙ cong f (merid (pt A)) + ; (merid a i) j → isSet→isSet' (isSetℤ) + (sym p) + (sym p ∙ cong f (merid (pt A))) + refl (cong f (merid a)) i j}))} +-- We need that Hⁿ⁺¹(Susp A) ≃ Hⁿ(A) + +private + suspFunCharacFun : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → ((Susp (typ A)) → coHomK (suc n)) + → (typ A → (coHomK n)) + suspFunCharacFun {A = A} n f x = ΩKn+1→Kn n (sym (rCancelₖ (suc n) (f north)) + ∙∙ cong (λ x → f x -[ (suc n) ]ₖ f north) ((merid x) ∙ sym (merid (pt A))) + ∙∙ rCancelₖ (suc n) (f north)) + + linvLem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f : typ A → Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) + (fId : f (pt A) ≡ refl) + → (x : _) → suspΩFun n (suspFunCharacFun {A = A} n (suspΩFun' n f)) x + ≡ suspΩFun' n f x + linvLem n f fId north = refl + linvLem n f fId south = refl + linvLem {A = A} n f fId (merid x i) j = helper n x f fId j i + where + helper : (n : ℕ) (a : typ A) (f : typ A → Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) + (fId : f (pt A) ≡ refl) + → Kn→ΩKn+1 n (suspFunCharacFun {A = A} n (suspΩFun' n f) a) ≡ f a + helper zero a f fId = + Iso.rightInv (Iso-Kn-ΩKn+1 0) (sym (rCancelₖ 1 (0ₖ 1)) + ∙∙ cong (λ x → suspΩFun' 0 f x +ₖ 0ₖ 1) (merid a ∙ (sym (merid (pt A)))) + ∙∙ (rCancelₖ 1 (0ₖ 1))) + ∙∙ sym (rUnit _) + ∙∙ (λ j i → rUnitₖ 1 (congFunct (suspΩFun' 0 f) (merid a) (sym (merid (pt A))) j i) j) + ∙∙ cong (λ p → f a ∙ sym p) fId + ∙∙ sym (rUnit (f a)) + helper (suc n) a f fId = + Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) + ((sym (rCancelₖ _ (0ₖ (suc (suc n)))) + ∙∙ cong (λ x → suspΩFun' (suc n) f x +ₖ 0ₖ (suc (suc n))) (merid a ∙ (sym (merid (pt A)))) + ∙∙ rCancelₖ _ (0ₖ (suc (suc n))))) + ∙∙ (((λ i → sym (rCancel≡refl n i) + ∙∙ cong (λ x → rUnitₖ (suc (suc n)) (suspΩFun' (suc n) f x) i) (merid a ∙ (sym (merid (pt A)))) + ∙∙ rCancel≡refl n i))) + ∙∙ ((λ i → rUnit (congFunct (suspΩFun' (suc n) f) (merid a) (sym (merid (pt A))) i) (~ i))) + ∙∙ cong (λ p → f a ∙ sym p) fId + ∙∙ sym (rUnit (f a)) + + +suspFunCharac : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (coHom (suc (suc n)) (Susp (typ A))) (coHom (suc n) (typ A)) +fun (suspFunCharac {A = A} n) = + sMap λ f → suspFunCharacFun {A = A} (suc n) f +inv (suspFunCharac {A = A} n) = sMap (suspΩFun (suc n)) +rightInv (suspFunCharac {A = A} n) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ fId → cong ∣_∣₂ + (funExt (λ x → cong (ΩKn+1→Kn (suc n)) + ((λ i → sym (rCancel≡refl n i) ∙∙ cong (λ x → suspΩFun (suc n) f x +ₖ 0ₖ _) + (merid x ∙ sym (merid (pt A))) ∙∙ rCancel≡refl n i) + ∙∙ sym (rUnit (cong (λ x → suspΩFun (suc n) f x +ₖ 0ₖ _) + ((merid x) ∙ sym (merid (pt A))))) + ∙∙ λ i → congFunct (λ x → rUnitₖ _ (suspΩFun (suc n) f x) i) + (merid x) (sym (merid (pt A))) i) + ∙∙ ΩKn+1→Kn-hom (suc n) (Kn→ΩKn+1 (suc n) (f x)) + (sym (Kn→ΩKn+1 (suc n) (f (pt A)))) + ∙∙ cong₂ _+ₖ_ (Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) (f x)) + (cong (λ x → ΩKn+1→Kn (suc n) (sym (Kn→ΩKn+1 (suc n) x))) fId) + ∙∙ cong (λ y → f x +ₖ ΩKn+1→Kn (suc n) (sym y)) (Kn→ΩKn+10ₖ (suc n)) + ∙∙ cong (f x +ₖ_) (ΩKn+1→Kn-refl (suc n)) + ∙ rUnitₖ _ (f x)))) + (fst (isConnectedPathKn n (f (pt A)) (0ₖ _))) +leftInv (suspFunCharac {A = A} n) = + SuspCohomElim {A = A} _ (λ _ → isSetSetTrunc _ _) + λ f fId → cong ∣_∣₂ (funExt (linvLem (suc n) f fId)) + +-- We also need that H¹(Susp A) ≃ Ĥ⁰(A) +suspFunCharac0 : ∀ {ℓ} {A : Pointed ℓ} → Iso (∥ ((Susp (typ A)) → coHomK 1) ∥₂) ∥ A →∙ (ℤ , 0) ∥₂ +fun (suspFunCharac0 {A = A}) = + sMap λ f → suspFunCharacFun {A = A} 0 f + , (cong (ΩKn+1→Kn 0) ((λ i → sym (rCancelₖ _ (f north)) + ∙∙ cong (λ x → f x -ₖ f north) (rCancel (merid (pt A)) i) + ∙∙ rCancelₖ _ (f north)) + ∙∙ (doubleCompPath-elim (sym (rCancelₖ _ (f north))) refl (rCancelₖ _ (f north))) + ∙∙ (cong (_∙ (rCancelₖ _ (f north))) (sym (rUnit (sym (rCancelₖ _ (f north)))))) + ∙ (lCancel (rCancelₖ _ (f north))))) +inv suspFunCharac0 = sMap λ f → suspΩFun 0 (fst f) +rightInv (suspFunCharac0 {A = A}) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , p) + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (funExt (λ x → (λ j → transp (λ i → helix (wedgeMapS¹ (intLoop (p j) (~ i)) base)) j + ((transport (λ i → helix (trRec isGroupoidS¹ (λ x → x) + (rUnitₖ 1 ∣ intLoop (f x) i ∣ j))) + (pos 0)))) + ∙ windingℤLoop (f x))))} +leftInv (suspFunCharac0 {A = A}) = + SuspCohomElim {A = A} _ (λ _ → isSetSetTrunc _ _) + λ f fId → cong ∣_∣₂ (funExt (linvLem 0 f fId)) + +-- We now prove that the alternative definition of cohomology is a cohomology theory. +private + -- First, we need to that coHomFunctor' is contravariant + theMorph : ∀ {ℓ} (n : ℤ) {A B : Pointed ℓ} (f : A →∙ B) + → AbGroupHom (coHomFunctor' n B) (coHomFunctor' n A) + fst (theMorph (pos zero) f) = sMap λ g → (λ x → fst g (fst f x)) , cong (fst g) (snd f) ∙ snd g + snd (theMorph (pos zero) f) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl)) + theMorph (pos (suc n)) f = coHomMorph _ (fst f) + fst (theMorph (negsuc n) f) = idfun _ + snd (theMorph (negsuc n) f) = makeIsGroupHom λ _ _ → refl + + open coHomTheory + isCohomTheoryZ' : ∀ {ℓ} → coHomTheory {ℓ} coHomFunctor' + Hmap isCohomTheoryZ' = theMorph + + -------------------------- Suspension -------------------------- + -- existence of suspension isomorphism + fst (Suspension isCohomTheoryZ') (pos zero) {A = A} = + invGroupEquiv + (GroupIso→GroupEquiv + ( invIso suspFunCharac0 + , makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g → cong ∣_∣₂ (funExt λ { north → refl + ; south → refl + ; (merid a i) j → helper a (fst f) (fst g) j i})))) + where + helper : (a : typ A) (f g : typ A → coHomK 0) + → Kn→ΩKn+1 0 (f a +[ 0 ]ₖ g a) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 0 (f a)) (Kn→ΩKn+1 0 (g a)) + helper a f g = Kn→ΩKn+1-hom 0 (f a) (g a) + ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a)) + fst (Suspension isCohomTheoryZ') (pos (suc n)) {A = A} = + invGroupEquiv + (GroupIso→GroupEquiv + ( invIso (suspFunCharac {A = A} n) + , makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g → cong ∣_∣₂ (funExt λ { north → refl + ; south → refl + ; (merid a i) j → helper a f g j i})))) + where + helper : (a : typ A) (f g : typ A → coHomK (suc n)) + → Kn→ΩKn+1 (suc n) (f a +ₖ g a) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc n) (f a)) (Kn→ΩKn+1 (suc n) (g a)) + helper a f g = Kn→ΩKn+1-hom (suc n) (f a) (g a) + ∙ ∙≡+₂ n (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a)) + fst (Suspension isCohomTheoryZ') (negsuc zero) {A = A} = + GroupIso→GroupEquiv + ( isContr→Iso (H0-susp {A = _ , pt A}) isContrUnit* + , makeIsGroupHom λ _ _ → refl) + fst (Suspension isCohomTheoryZ') (negsuc (suc n)) = idGroupEquiv + + -- naturality of the suspension isomorphism + snd (Suspension (isCohomTheoryZ' {ℓ})) (f , p) (pos zero) = + funExt (sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , _) → cong ∣_∣₂ (funExt λ {north → refl + ; south → refl + ; (merid a i) → refl})}) + snd (Suspension (isCohomTheoryZ' {ℓ})) (f , p) (pos (suc n)) = + funExt (sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ∣_∣₂ (funExt λ {north → refl + ; south → refl + ; (merid a i) → refl})) + snd (Suspension (isCohomTheoryZ' {ℓ})) (f , p) (negsuc zero) = refl + snd (Suspension (isCohomTheoryZ' {ℓ})) (f , p) (negsuc (suc n)) = refl + + -------------------------- Exactness --------------------------- + Exactness isCohomTheoryZ' {A = A} {B = B} f n = isoToPath (exactnessIso n f) + where + exactnessIso : (n : ℤ) (f : A →∙ B) + → Iso (Ker (theMorph n f)) (Im (theMorph n (cfcod (fst f) , refl))) + fun (exactnessIso (pos zero) (f , p)) = + uncurry (sElim (λ _ → isSetΠ λ _ → isSetΣ isSetSetTrunc λ _ → isProp→isSet isPropPropTrunc) + λ {(g , q) inker → ∣ g , q ∣₂ + , pRec isPropPropTrunc + (λ gId → ∣ ∣ (λ { (inl tt) → 0 + ; (inr b) → g b + ; (push a i) → funExt⁻ (cong fst gId) a (~ i)}) , q ∣₂ + , cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl) ∣) + (Iso.fun PathIdTrunc₀Iso inker)}) + inv (exactnessIso (pos zero) (f , p)) = + uncurry (sElim (λ _ → isSetΠ λ _ → isSetΣ isSetSetTrunc λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(g , q) inim' + → ∣ g , q ∣₂ + , pRec (isSetSetTrunc _ _) + (uncurry + (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _)) + (λ pushmap pushId' + → pRec (isSetSetTrunc _ _) + (λ pushId + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (funExt λ x → sym (funExt⁻ (cong fst pushId) (f x)) + ∙∙ cong (fst pushmap) (sym (push x) ∙ push (pt A)) + ∙∙ (cong (fst pushmap ∘ inr) p + ∙ snd pushmap)))) + (Iso.fun PathIdTrunc₀Iso pushId')))) + inim'}) + rightInv (exactnessIso (pos zero) (f , p)) = + uncurry (sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 + (isSetΣ isSetSetTrunc + (λ _ → isProp→isSet isPropPropTrunc)) _ _) + λ {(p , q) _ → Σ≡Prop (λ _ → isPropPropTrunc) refl}) + leftInv (exactnessIso (pos zero) (f , p)) = + uncurry (sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 + (isSetΣ isSetSetTrunc + (λ _ → isProp→isSet (isSetSetTrunc _ _))) _ _) + λ {(p , q) _ → Σ≡Prop (λ _ → isSetSetTrunc _ _) refl}) + fun (exactnessIso (pos (suc n)) f) ker = (fst ker) , inIm-helper (fst ker) (snd ker) + where + inIm-helper : (x : coHom (suc n) (typ B)) + → isInKer (theMorph (pos (suc n)) {A = A} {B = B} f) x + → isInIm (theMorph (pos (suc n)) {A = B} {B = _ , inr (pt B)} (cfcod (fst f) , refl)) x + inIm-helper = + coHomPointedElim _ (pt B) (λ _ → isPropΠ λ _ → isPropPropTrunc) + λ g gId inker → pRec isPropPropTrunc + (λ gIdTot → ∣ ∣ (λ { (inl tt) → 0ₖ _ + ; (inr b) → g b + ; (push a i) → funExt⁻ gIdTot a (~ i)}) ∣₂ + , cong ∣_∣₂ (funExt λ b → refl) ∣) + (Iso.fun PathIdTrunc₀Iso inker) + inv (exactnessIso (pos (suc n)) f) im = fst im , inKer-helper (fst im) (snd im) + where + inKer-helper : (x : coHom (suc n) (typ B)) + → isInIm (theMorph (pos (suc n)) {A = B} {B = _ , inr (pt B)} (cfcod (fst f) , refl)) x + → isInKer (theMorph (pos (suc n)) {A = A} {B = B} f) x + inKer-helper = + coHomPointedElim _ (pt B) (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ g gId → pRec (isSetSetTrunc _ _) + (uncurry λ cg p + → subst (isInKer (coHomMorph (suc n) (fst f))) + p + (helper cg)) + where + helper : (cg : _) → coHomFun (suc n) (fst f) (coHomFun (suc n) (cfcod (fst f)) cg) ≡ 0ₕ _ + helper = sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ cg → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ p → (cong ∣_∣₂ (funExt λ x → cong cg (sym (push x)) + ∙ p))) + (isConnectedPathKn _ (cg (inl tt)) (0ₖ (suc n)) .fst) + rightInv (exactnessIso (pos (suc n)) f) _ = Σ≡Prop (λ _ → isPropPropTrunc) refl + leftInv (exactnessIso (pos (suc n)) f) _ = Σ≡Prop (λ _ → isSetSetTrunc _ _) refl + exactnessIso (negsuc n) (f , p) = + isContr→Iso ((tt* , refl) + , λ {(tt* , p) → Σ≡Prop (λ _ → isOfHLevelPath 1 isPropUnit* _ _) + refl}) + ((tt* , ∣ tt* , refl ∣) + , λ {(tt* , p) → Σ≡Prop (λ _ → isPropPropTrunc) + refl}) + + -------------------------- Dimension --------------------------- + Dimension isCohomTheoryZ' (pos zero) p = ⊥-rec (p refl) + fst (Dimension isCohomTheoryZ' (pos (suc n)) _) = 0ₕ _ + snd (Dimension isCohomTheoryZ' (pos (suc n)) _) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ f → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ f-true → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ f-false → cong ∣_∣₂ (funExt (λ {(lift true) → f-true + ; (lift false) → f-false}))) + (isConnectedPathKn n (0ₖ _) (f (lift false)) .fst)) + (isConnectedPathKn n (0ₖ _) (f (lift true)) .fst)) + Dimension isCohomTheoryZ' (negsuc n) _ = isContrUnit* + + ------------------------ Binary wedges ------------------------- + BinaryWedge isCohomTheoryZ' (pos zero) = GroupIso→GroupEquiv (H⁰Red-⋁ _ _) + BinaryWedge isCohomTheoryZ' (pos (suc n)) = GroupIso→GroupEquiv (Hⁿ-⋁ _ _ n) + BinaryWedge isCohomTheoryZ' (negsuc n) = + GroupIso→GroupEquiv + (compGroupIso (contrGroupIsoUnit isContrUnit*) + (invGroupIso (contrGroupIsoUnit (isOfHLevel× 0 isContrUnit* isContrUnit*)))) + +-- Substituting back for our original theory, we are done +isCohomTheoryZ : ∀ {ℓ} → coHomTheory {ℓ} coHomFunctor +isCohomTheoryZ = subst coHomTheory (sym coHomFunctor≡coHomFunctor') + isCohomTheoryZ' diff --git a/Cubical/ZCohomology/GroupStructure.agda b/Cubical/ZCohomology/GroupStructure.agda new file mode 100644 index 0000000000..43aff7ccfd --- /dev/null +++ b/Cubical/ZCohomology/GroupStructure.agda @@ -0,0 +1,930 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.GroupStructure where + +open import Cubical.ZCohomology.Base + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed hiding (id) +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) +open import Cubical.Data.Sigma +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; isSetSetTrunc to §) +open import Cubical.Data.Int renaming (_+_ to _ℤ+_ ; -_ to -ℤ_) +open import Cubical.Data.Nat renaming (+-assoc to +-assocℕ ; +-comm to +-commℕ) +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3 ; map2 to trMap2) +open import Cubical.Homotopy.Loopspace +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup) +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + +open Iso renaming (inv to inv') + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + A' : Pointed ℓ + +infixr 34 _+ₖ_ +infixr 34 _+ₕ_ +infixr 34 _+ₕ∙_ + +-- Addition in the Eilenberg-Maclane spaces is uniquely determined if we require it to have left- and right-unit laws, +-- such that these agree on 0. In particular, any h-structure (see http://ericfinster.github.io/files/emhott.pdf) is unique. ++ₖ-unique : (n : ℕ) → (comp1 comp2 : coHomK (suc n) → coHomK (suc n) → coHomK (suc n)) + → (rUnit1 : (x : _) → comp1 x (coHom-pt (suc n)) ≡ x) + → (lUnit1 : (x : _) → comp1 (coHom-pt (suc n)) x ≡ x) + → (rUnit2 : (x : _) → comp2 x (coHom-pt (suc n)) ≡ x) + → (lUnit2 : (x : _) → comp2 (coHom-pt (suc n)) x ≡ x) + → (unId1 : rUnit1 (coHom-pt (suc n)) ≡ lUnit1 (coHom-pt (suc n))) + → (unId2 : rUnit2 (coHom-pt (suc n)) ≡ lUnit2 (coHom-pt (suc n))) + → (x y : _) → comp1 x y ≡ comp2 x y ++ₖ-unique n comp1 comp2 rUnit1 lUnit1 rUnit2 lUnit2 unId1 unId2 = + elim2 (λ _ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + (wedgeconFun _ _ + (λ _ _ → help _ _) + (λ x → lUnit1 ∣ x ∣ ∙ sym (lUnit2 ∣ x ∣)) + (λ x → rUnit1 ∣ x ∣ ∙ sym (rUnit2 ∣ x ∣)) + (cong₂ _∙_ unId1 (cong sym unId2))) + where + help : isOfHLevel (2 + (n + suc n)) (coHomK (suc n)) + help = subst (λ x → isOfHLevel x (coHomK (suc n))) (+-suc n (2 + n) ∙ +-suc (suc n) (suc n)) + (isOfHLevelPlus n (isOfHLevelTrunc (3 + n))) + +wedgeConHLev : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (coHomK (2 + n)) +wedgeConHLev n = subst (λ x → isOfHLevel x (coHomK (2 + n))) + (sym (+-suc (2 + n) (suc n) ∙ +-suc (3 + n) n)) + (isOfHLevelPlus' {n = n} (4 + n) (isOfHLevelTrunc (4 + n))) +wedgeConHLev' : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (typ (Ω (coHomK-ptd (3 + n)))) +wedgeConHLev' n = subst (λ x → isOfHLevel x (typ (Ω (coHomK-ptd (3 + n))))) + (sym (+-suc (2 + n) (suc n) ∙ +-suc (3 + n) n)) + (isOfHLevelPlus' {n = n} (4 + n) (isOfHLevelTrunc (5 + n) _ _)) + +wedgeConHLevPath : (n : ℕ) → (x y : coHomK (suc n)) → isOfHLevel ((suc n) + (suc n)) (x ≡ y) +wedgeConHLevPath zero x y = isOfHLevelTrunc 3 _ _ +wedgeConHLevPath (suc n) x y = isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _ + +-- addition for n ≥ 2 together with the left- and right-unit laws (modulo truncations) +preAdd : (n : ℕ) → (S₊ (2 + n) → S₊ (2 + n) → coHomK (2 + n)) +preAdd n = + wedgeconFun _ _ + (λ _ _ → wedgeConHLev n) + ∣_∣ + ∣_∣ + refl + +preAdd-l : (n : ℕ) → (x : (S₊ (2 + n))) → preAdd n north x ≡ ∣ x ∣ +preAdd-l n _ = refl + +preAdd-r : (n : ℕ) → (x : (S₊ (2 + n))) → preAdd n x north ≡ ∣ x ∣ +preAdd-r n = + wedgeconRight _ (suc n) + (λ _ _ → wedgeConHLev n) + ∣_∣ + ∣_∣ + refl + +-- addition for n = 1 +wedgeMapS¹ : S¹ → S¹ → S¹ +wedgeMapS¹ base y = y +wedgeMapS¹ (loop i) base = loop i +wedgeMapS¹ (loop i) (loop j) = + hcomp (λ k → λ { (i = i0) → loop j + ; (i = i1) → loop (j ∧ k) + ; (j = i0) → loop i + ; (j = i1) → loop (i ∧ k)}) + (loop (i ∨ j)) + +---------- Algebra/Group stuff -------- +0ₖ : (n : ℕ) → coHomK n +0ₖ = coHom-pt + +_+ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n +_+ₖ_ {n = zero} x y = x ℤ+ y +_+ₖ_ {n = suc zero} = trMap2 wedgeMapS¹ +_+ₖ_ {n = suc (suc n)} = trRec (isOfHLevelΠ (4 + n) λ _ → isOfHLevelTrunc (4 + n)) + λ x → trRec (isOfHLevelTrunc (4 + n)) (preAdd n x) + +private + isEquiv+ : (n : ℕ) → (x : coHomK (suc n)) → isEquiv (_+ₖ_ {n = (suc n)} x) + isEquiv+ zero = + trElim (λ _ → isProp→isOfHLevelSuc 2 (isPropIsEquiv _)) + (toPropElim (λ _ → isPropIsEquiv _) + (subst isEquiv (sym help) (idIsEquiv _))) + where + help : _+ₖ_ {n = 1} (coHom-pt 1) ≡ idfun _ + help = funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ _ → refl) + isEquiv+ (suc n) = + trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsEquiv _)) + (suspToPropElim (ptSn (suc n)) (λ _ → isPropIsEquiv _) + (subst isEquiv (sym help) (idIsEquiv _))) + where + help : _+ₖ_ {n = (2 + n)} (coHom-pt (2 + n)) ≡ idfun _ + help = funExt (trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) λ _ → refl) + + + Kₙ≃Kₙ : (n : ℕ) (x : coHomK (suc n)) → coHomK (suc n) ≃ coHomK (suc n) + Kₙ≃Kₙ n x = _ , isEquiv+ n x + +-ₖ_ : {n : ℕ} → coHomK n → coHomK n +-ₖ_ {n = zero} x = 0 - x +-ₖ_ {n = suc zero} = trMap λ {base → base ; (loop i) → (loop (~ i))} +-ₖ_ {n = suc (suc n)} = trMap λ {north → north + ; south → north + ; (merid a i) → ((merid (ptSn (suc n)) ∙ sym (merid a))) i} + +_-ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n +_-ₖ_ {n = n} x y = _+ₖ_ {n = n} x (-ₖ_ {n = n} y) + ++ₖ-syntax : (n : ℕ) → coHomK n → coHomK n → coHomK n ++ₖ-syntax n = _+ₖ_ {n = n} + +-ₖ-syntax : (n : ℕ) → coHomK n → coHomK n +-ₖ-syntax n = -ₖ_ {n = n} + +-'ₖ-syntax : (n : ℕ) → coHomK n → coHomK n → coHomK n +-'ₖ-syntax n = _-ₖ_ {n = n} + +syntax +ₖ-syntax n x y = x +[ n ]ₖ y +syntax -ₖ-syntax n x = -[ n ]ₖ x +syntax -'ₖ-syntax n x y = x -[ n ]ₖ y + +-ₖ^2 : {n : ℕ} → (x : coHomK n) → (-ₖ (-ₖ x)) ≡ x +-ₖ^2 {n = zero} x = + +Comm (pos zero) (-ℤ (pos zero ℤ+ (-ℤ x))) ∙∙ -Dist+ (pos zero) (-ℤ x) + ∙∙ (+Comm (pos zero) (-ℤ (-ℤ x)) ∙ -Involutive x) +-ₖ^2 {n = suc zero} = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ { base → refl ; (loop i) → refl} +-ₖ^2 {n = suc (suc n)} = + trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ { north → refl + ; south j → ∣ merid (ptSn _) j ∣ₕ + ; (merid a i) j + → hcomp (λ k → λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ compPath-filler' (merid a) (sym (merid (ptSn _))) (~ k) (~ j) ∣ₕ + ; (j = i0) → help a (~ k) i + ; (j = i1) → ∣ merid a (i ∧ k) ∣}) + ∣ (merid a ∙ sym (merid (ptSn _))) (i ∧ ~ j) ∣ₕ} + where + help : (a : _) → cong (-ₖ_ ∘ (-ₖ_ {n = suc (suc n)})) (cong ∣_∣ₕ (merid a)) + ≡ cong ∣_∣ₕ (merid a ∙ sym (merid (ptSn _))) + help a = cong (cong ((-ₖ_ {n = suc (suc n)}))) (cong-∙ ∣_∣ₕ (merid (ptSn (suc n))) (sym (merid a))) + ∙∙ cong-∙ (-ₖ_ {n = suc (suc n)}) (cong ∣_∣ₕ (merid (ptSn (suc n)))) (cong ∣_∣ₕ (sym (merid a))) + ∙∙ (λ i → (λ j → ∣ rCancel (merid (ptSn (suc n))) i j ∣ₕ) + ∙ λ j → ∣ symDistr (merid (ptSn (suc n))) (sym (merid a)) i j ∣ₕ) + ∙ sym (lUnit _) + + +------- Groupoid Laws for Kₙ --------- +commₖ : (n : ℕ) → (x y : coHomK n) → x +[ n ]ₖ y ≡ y +[ n ]ₖ x +commₖ zero = +Comm +commₖ (suc zero) = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (wedgeconFun _ _ + (λ _ _ → isOfHLevelTrunc 3 _ _) + (λ {base → refl ; (loop i) → refl}) + (λ {base → refl ; (loop i) → refl}) + refl) +commₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeconFun _ _ + (λ x y → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _) + (λ x → preAdd-l n x ∙ sym (preAdd-r n x)) + (λ x → preAdd-r n x ∙ sym (preAdd-l n x)) + refl) + +commₖ-base : (n : ℕ) → commₖ n (coHom-pt n) (coHom-pt n) ≡ refl +commₖ-base zero = refl +commₖ-base (suc zero) = refl +commₖ-base (suc (suc n)) = sym (rUnit _) + +rUnitₖ : (n : ℕ) → (x : coHomK n) → x +[ n ]ₖ coHom-pt n ≡ x +rUnitₖ zero x = refl +rUnitₖ (suc zero) = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ {base → refl + ; (loop i) → refl} +rUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (preAdd-r n) + +lUnitₖ : (n : ℕ) → (x : coHomK n) → coHom-pt n +[ n ]ₖ x ≡ x +lUnitₖ zero x = sym (pos0+ x) +lUnitₖ (suc zero) = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ {base → refl + ; (loop i) → refl} +lUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ x → refl + +∙≡+₁ : (p q : typ (Ω (coHomK-ptd 1))) → p ∙ q ≡ cong₂ _+ₖ_ p q +∙≡+₁ p q = (λ i → (λ j → rUnitₖ 1 (p j) (~ i)) ∙ λ j → lUnitₖ 1 (q j) (~ i)) ∙ sym (cong₂Funct _+ₖ_ p q) + +∙≡+₂ : (n : ℕ) (p q : typ (Ω (coHomK-ptd (suc (suc n))))) → p ∙ q ≡ cong₂ _+ₖ_ p q +∙≡+₂ n p q = (λ i → (λ j → rUnitₖ (2 + n) (p j) (~ i)) ∙ λ j → lUnitₖ (2 + n) (q j) (~ i)) ∙ sym (cong₂Funct _+ₖ_ p q) + +lCancelₖ : (n : ℕ) → (x : coHomK n) → (-ₖ_ {n = n} x) +ₖ x ≡ coHom-pt n +lCancelₖ zero x = minusPlus x 0 +lCancelₖ (suc zero) = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ {base → refl ; (loop i) j → help j i} + where + help : cong₂ _+ₖ_ (sym (cong ∣_∣ loop)) (cong ∣_∣ loop) ≡ refl + help = sym (∙≡+₁ (sym (cong ∣_∣ loop)) (cong ∣_∣ loop)) ∙ lCancel _ +lCancelₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ {north → refl ; south → cong ∣_∣ (sym (merid (ptSn (suc n)))) ; (merid a i) → help a i } + where + s : (a : _) → _ ≡ _ + s x = cong₂ _+ₖ_ (sym (cong ∣_∣ (merid (ptSn (suc n)) ∙ sym (merid x)))) (cong ∣_∣ (sym (merid x))) + + help : (a : _) → PathP (λ i → (preAdd n ((merid (ptSn (suc n)) ∙ (λ i₁ → merid a (~ i₁))) i) + (merid a i)) ≡ ∣ north ∣) refl λ i₁ → ∣ merid (ptSn (suc n)) (~ i₁) ∣ + help x = + compPathR→PathP + ((sym (lCancel _) + ∙∙ (λ i → ∙≡+₂ _ (cong ∣_∣ (symDistr (merid x) (sym (merid (ptSn (suc n)))) i)) (cong ∣_∣ ((merid x) ∙ sym (merid (ptSn (suc n))))) i) + ∙∙ rUnit _) + ∙∙ (λ j → cong₂ _+ₖ_ ((cong ∣_∣ (merid (ptSn (suc n)) ∙ sym (merid x)))) + (λ i → ∣ compPath-filler ((merid x)) ((sym (merid (ptSn (suc n))))) (~ j) i ∣) + ∙ λ i → ∣ merid (ptSn (suc n)) (~ i ∧ j) ∣) + ∙∙ λ i → sym (s x) ∙ rUnit (cong ∣_∣ (sym (merid (ptSn (suc n))))) i) + +rCancelₖ : (n : ℕ) → (x : coHomK n) → x +ₖ (-ₖ_ {n = n} x) ≡ coHom-pt n +rCancelₖ zero x = +Comm x (pos 0 - x) ∙ minusPlus x 0 -- +-comm x (pos 0 - x) ∙ minusPlus x 0 +rCancelₖ (suc zero) = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ {base → refl ; (loop i) j → help j i} + where + help : (λ i → ∣ loop i ∣ₕ +ₖ (-ₖ ∣ loop i ∣ₕ)) ≡ refl + help = sym (∙≡+₁ (cong ∣_∣ₕ loop) (sym (cong ∣_∣ₕ loop))) ∙ rCancel _ +rCancelₖ (suc (suc n)) x = commₖ _ x (-ₖ x) ∙ lCancelₖ _ x + +rCancel≡refl : (n : ℕ) → rCancelₖ (2 + n) (0ₖ _) ≡ refl +rCancel≡refl n i = rUnit (rUnit refl (~ i)) (~ i) + +assocₖ : (n : ℕ) → (x y z : coHomK n) → x +[ n ]ₖ (y +[ n ]ₖ z) ≡ (x +[ n ]ₖ y) +[ n ]ₖ z +assocₖ zero = +Assoc +assocₖ (suc zero) = + trElim3 (λ _ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ x → wedgeconFun _ _ + (λ _ _ → isOfHLevelTrunc 3 _ _) + (λ y i → rUnitₖ 1 ∣ x ∣ (~ i) +ₖ ∣ y ∣) + (λ z → cong (∣ x ∣ +ₖ_) (rUnitₖ 1 ∣ z ∣) ∙ sym (rUnitₖ 1 (∣ x ∣ +ₖ ∣ z ∣))) + (helper x) + where + helper : (x : S¹) → cong (∣ x ∣ +ₖ_) (rUnitₖ 1 ∣ base ∣) ∙ sym (rUnitₖ 1 (∣ x ∣ +ₖ ∣ base ∣)) + ≡ (cong (_+ₖ ∣ base ∣) (sym (rUnitₖ 1 ∣ x ∣))) + helper = toPropElim (λ _ → isOfHLevelTrunc 3 _ _ _ _) + (sym (lUnit refl)) + +assocₖ (suc (suc n)) = + trElim3 (λ _ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeConSn-×3 _ + (λ x z i → preAdd-r n x (~ i) +ₖ ∣ z ∣) + (λ x y → cong (∣ x ∣ +ₖ_) (rUnitₖ (2 + n) ∣ y ∣) ∙ sym (rUnitₖ (2 + n) (∣ x ∣ +ₖ ∣ y ∣))) + (lUnit (sym (rUnitₖ (2 + n) (∣ north ∣ +ₖ ∣ north ∣))))) + where + wedgeConSn-×3 : (n : ℕ) + → (f : (x z : S₊ (2 + n))→ ∣ x ∣ +ₖ ((0ₖ _) +ₖ ∣ z ∣) ≡ (∣ x ∣ +ₖ (0ₖ _)) +ₖ ∣ z ∣) + → (g : (x y : S₊ (2 + n)) → ∣ x ∣ +ₖ (∣ y ∣ +ₖ 0ₖ _) ≡ (∣ x ∣ +ₖ ∣ y ∣) +ₖ 0ₖ _) + → (f (ptSn _) (ptSn _) ≡ g (ptSn _) (ptSn _)) + → (x y z : S₊ (2 + n)) → ∣ x ∣ +ₖ (∣ y ∣ +ₖ ∣ z ∣) ≡ (∣ x ∣ +ₖ ∣ y ∣) +ₖ ∣ z ∣ + wedgeConSn-×3 n f g d x = + wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _) + (f x) + (g x) + (sphereElim _ {A = λ x → g x (ptSn (suc (suc n))) ≡ f x (ptSn (suc (suc n))) } + (λ _ → isOfHLevelTrunc (4 + n) _ _ _ _) + (sym d) x) +{- +This was the original proof for the case n ≥ 2: +For some reason it doesn't check in reasonable time anymore: +assocₖ (suc (suc n)) = + trElim3 (λ _ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ x → wedgeConSn _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _) + (λ z i → preAdd n .snd .snd x (~ i) +ₖ ∣ z ∣) + (λ y → cong (∣ x ∣ +ₖ_) (rUnitₖ (2 + n) ∣ y ∣) ∙ sym (rUnitₖ (2 + n) (∣ x ∣ +ₖ ∣ y ∣))) + (helper x) .fst + where + helper : (x : S₊ (2 + n)) → cong (∣ x ∣ +ₖ_) (rUnitₖ (2 + n) ∣ north ∣) ∙ sym (rUnitₖ (2 + n) (∣ x ∣ +ₖ ∣ north ∣)) + ≡ cong (_+ₖ ∣ north ∣) (sym (preAdd n .snd .snd x)) + helper = sphereElim (suc n) (λ _ → isOfHLevelTrunc (4 + n) _ _ _ _) + (sym (lUnit (sym (rUnitₖ (2 + n) (∣ north ∣ +ₖ ∣ north ∣))))) +-} + + +lUnitₖ≡rUnitₖ : (n : ℕ) → lUnitₖ n (coHom-pt n) ≡ rUnitₖ n (coHom-pt n) +lUnitₖ≡rUnitₖ zero = isSetℤ _ _ _ _ +lUnitₖ≡rUnitₖ (suc zero) = refl +lUnitₖ≡rUnitₖ (suc (suc n)) = refl + +------ Commutativity of ΩKₙ +-- We show that p ∙ q ≡ (λ i → (p i) +ₖ (q i)) for any p q : ΩKₙ₊₁. This allows us to prove that p ∙ q ≡ q ∙ p +-- without having to use the equivalence Kₙ ≃ ΩKₙ₊₁ + + +cong+ₖ-comm : (n : ℕ) (p q : typ (Ω (coHomK-ptd (suc n)))) → cong₂ _+ₖ_ p q ≡ cong₂ _+ₖ_ q p +cong+ₖ-comm zero p q = + rUnit (cong₂ _+ₖ_ p q) + ∙∙ (λ i → (λ j → commₖ 1 ∣ base ∣ ∣ base ∣ (i ∧ j)) + ∙∙ (λ j → commₖ 1 (p j) (q j) i) + ∙∙ λ j → commₖ 1 ∣ base ∣ ∣ base ∣ (i ∧ ~ j)) + ∙∙ ((λ i → commₖ-base 1 i ∙∙ cong₂ _+ₖ_ q p ∙∙ sym (commₖ-base 1 i)) + ∙ sym (rUnit (cong₂ _+ₖ_ q p))) +cong+ₖ-comm (suc n) p q = + rUnit (cong₂ _+ₖ_ p q) + ∙∙ (λ i → (λ j → commₖ (2 + n) ∣ north ∣ ∣ north ∣ (i ∧ j)) + ∙∙ (λ j → commₖ (2 + n) (p j) (q j) i ) + ∙∙ λ j → commₖ (2 + n) ∣ north ∣ ∣ north ∣ (i ∧ ~ j)) + ∙∙ ((λ i → commₖ-base (2 + n) i ∙∙ cong₂ _+ₖ_ q p ∙∙ sym (commₖ-base (2 + n) i)) + ∙ sym (rUnit (cong₂ _+ₖ_ q p))) + +isCommΩK : (n : ℕ) → isComm∙ (coHomK-ptd n) +isCommΩK zero p q = isSetℤ _ _ (p ∙ q) (q ∙ p) +isCommΩK (suc zero) p q = ∙≡+₁ p q ∙∙ cong+ₖ-comm 0 p q ∙∙ sym (∙≡+₁ q p) +isCommΩK (suc (suc n)) p q = ∙≡+₂ n p q ∙∙ cong+ₖ-comm (suc n) p q ∙∙ sym (∙≡+₂ n q p) + +----- some other useful lemmas about algebra in Kₙ +-0ₖ : {n : ℕ} → -[ n ]ₖ (0ₖ n) ≡ (0ₖ n) +-0ₖ {n = zero} = refl +-0ₖ {n = suc zero} = refl +-0ₖ {n = suc (suc n)} = refl + +-distrₖ : (n : ℕ) (x y : coHomK n) → -[ n ]ₖ (x +[ n ]ₖ y) ≡ (-[ n ]ₖ x) +[ n ]ₖ (-[ n ]ₖ y) +-distrₖ zero x y = GroupTheory.invDistr ℤGroup x y ∙ +Comm (0 - y) (0 - x) +-distrₖ (suc zero) = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 3 _ _) + (λ x → sym (lUnitₖ 1 (-[ 1 ]ₖ ∣ x ∣))) + (λ x → cong (λ x → -[ 1 ]ₖ x) (rUnitₖ 1 ∣ x ∣) ∙ sym (rUnitₖ 1 (-[ 1 ]ₖ ∣ x ∣))) + (sym (rUnit refl))) +-distrₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _) + (λ x → sym (lUnitₖ (2 + n) (-[ (2 + n) ]ₖ ∣ x ∣))) + (λ x → cong (λ x → -[ (2 + n) ]ₖ x) (rUnitₖ (2 + n) ∣ x ∣ ) ∙ sym (rUnitₖ (2 + n) (-[ (2 + n) ]ₖ ∣ x ∣))) + (sym (rUnit refl))) + +-cancelRₖ : (n : ℕ) (x y : coHomK n) → (y +[ n ]ₖ x) -[ n ]ₖ x ≡ y +-cancelRₖ zero x y = sym (+Assoc y x (0 - x)) + ∙∙ cong (y ℤ+_) (+Comm x (0 - x)) + ∙∙ cong (y ℤ+_) (minusPlus x (pos 0)) +-cancelRₖ (suc zero) = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (wedgeconFun _ _ (λ _ _ → wedgeConHLevPath 0 _ _) + (λ x → cong (_+ₖ ∣ base ∣) (rUnitₖ 1 ∣ x ∣) ∙ rUnitₖ 1 ∣ x ∣) + (λ x → rCancelₖ 1 ∣ x ∣) + (rUnit refl)) +-cancelRₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeconFun _ _ (λ _ _ → wedgeConHLevPath (suc n) _ _) + (λ x → cong (_+ₖ ∣ north ∣) (rUnitₖ (2 + n) ∣ x ∣) ∙ rUnitₖ (2 + n) ∣ x ∣) + (λ x → rCancelₖ (2 + n) ∣ x ∣) + (sym (rUnit _))) + +-cancelLₖ : (n : ℕ) (x y : coHomK n) → (x +[ n ]ₖ y) -[ n ]ₖ x ≡ y +-cancelLₖ n x y = cong (λ z → z -[ n ]ₖ x) (commₖ n x y) ∙ -cancelRₖ n x y + +-+cancelₖ : (n : ℕ) (x y : coHomK n) → (x -[ n ]ₖ y) +[ n ]ₖ y ≡ x +-+cancelₖ zero x y = sym (+Assoc x (0 - y) y) ∙ cong (x ℤ+_) (minusPlus y (pos 0)) +-+cancelₖ (suc zero) = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (wedgeconFun _ _ (λ _ _ → wedgeConHLevPath 0 _ _) + (λ x → cong (_+ₖ ∣ x ∣) (lUnitₖ 1 (-ₖ ∣ x ∣)) ∙ lCancelₖ 1 ∣ x ∣) + (λ x → cong (_+ₖ ∣ base ∣) (rUnitₖ 1 ∣ x ∣) ∙ rUnitₖ 1 ∣ x ∣) + refl) +-+cancelₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeconFun _ _ (λ _ _ → wedgeConHLevPath (suc n) _ _) + (λ x → cong (_+ₖ ∣ x ∣) (lUnitₖ (2 + n) (-ₖ ∣ x ∣)) ∙ lCancelₖ (2 + n) ∣ x ∣) + (λ x → cong (_+ₖ ∣ north ∣) (rUnitₖ (2 + n) ∣ x ∣) ∙ rUnitₖ (2 + n) ∣ x ∣) + refl) + +---- Group structure of cohomology groups +_+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_+ₕ_ {n = n} = sRec2 § λ a b → ∣ (λ x → a x +[ n ]ₖ b x) ∣₂ + +-ₕ_ : {n : ℕ} → coHom n A → coHom n A +-ₕ_ {n = n} = sRec § λ a → ∣ (λ x → -[ n ]ₖ a x) ∣₂ + +_-ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_-ₕ_ {n = n} = sRec2 § λ a b → ∣ (λ x → a x -[ n ]ₖ b x) ∣₂ + ++ₕ-syntax : (n : ℕ) → coHom n A → coHom n A → coHom n A ++ₕ-syntax n = _+ₕ_ {n = n} + +-ₕ-syntax : (n : ℕ) → coHom n A → coHom n A +-ₕ-syntax n = -ₕ_ {n = n} + +-ₕ'-syntax : (n : ℕ) → coHom n A → coHom n A → coHom n A +-ₕ'-syntax n = _-ₕ_ {n = n} + +syntax +ₕ-syntax n x y = x +[ n ]ₕ y +syntax -ₕ-syntax n x = -[ n ]ₕ x +syntax -ₕ'-syntax n x y = x -[ n ]ₕ y + +0ₕ : (n : ℕ) → coHom n A +0ₕ n = ∣ (λ _ → (0ₖ n)) ∣₂ + +_+'ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_+'ₕ_ {n = n} x y = (x +ₕ 0ₕ _) +ₕ y +ₕ 0ₕ _ + +rUnitₕ : (n : ℕ) (x : coHom n A) → x +[ n ]ₕ (0ₕ n) ≡ x +rUnitₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitₖ n (a x)) i ∣₂ + +lUnitₕ : (n : ℕ) (x : coHom n A) → (0ₕ n) +[ n ]ₕ x ≡ x +lUnitₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitₖ n (a x)) i ∣₂ + +rCancelₕ : (n : ℕ) (x : coHom n A) → x +[ n ]ₕ (-[ n ]ₕ x) ≡ 0ₕ n +rCancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelₖ n (a x)) i ∣₂ + +lCancelₕ : (n : ℕ) (x : coHom n A) → (-[ n ]ₕ x) +[ n ]ₕ x ≡ 0ₕ n +lCancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelₖ n (a x)) i ∣₂ + +assocₕ : (n : ℕ) (x y z : coHom n A) → (x +[ n ]ₕ (y +[ n ]ₕ z)) ≡ ((x +[ n ]ₕ y) +[ n ]ₕ z) +assocₕ n = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocₖ n (a x) (b x) (c x)) i ∣₂ + +commₕ : (n : ℕ) (x y : coHom n A) → (x +[ n ]ₕ y) ≡ (y +[ n ]ₕ x) +commₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commₖ n (a x) (b x)) i ∣₂ + +-cancelLₕ : (n : ℕ) (x y : coHom n A) → (x +[ n ]ₕ y) -[ n ]ₕ x ≡ y +-cancelLₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelLₖ n (a x) (b x) i) ∣₂ + +-cancelRₕ : (n : ℕ) (x y : coHom n A) → (y +[ n ]ₕ x) -[ n ]ₕ x ≡ y +-cancelRₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelRₖ n (a x) (b x) i) ∣₂ + +-+cancelₕ : (n : ℕ) (x y : coHom n A) → (x -[ n ]ₕ y) +[ n ]ₕ y ≡ x +-+cancelₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -+cancelₖ n (a x) (b x) i) ∣₂ + +-- Group structure of reduced cohomology groups (in progress - might need K to compute properly first) +_+ₕ∙_ : {A : Pointed ℓ} {n : ℕ} → coHomRed n A → coHomRed n A → coHomRed n A +_+ₕ∙_ {n = zero} = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ zero ]ₖ b x) + , (λ i → (pa i +[ zero ]ₖ pb i)) ∣₂ } +_+ₕ∙_ {n = (suc zero)} = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ 1 ]ₖ b x) + , (λ i → pa i +[ 1 ]ₖ pb i) ∣₂ } +_+ₕ∙_ {n = (suc (suc n))} = + sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ (2 + n) ]ₖ b x) + , (λ i → pa i +[ (2 + n) ]ₖ pb i) ∣₂ } + +-ₕ∙_ : {A : Pointed ℓ} {n : ℕ} → coHomRed n A → coHomRed n A +-ₕ∙_ {n = zero} = sRec § λ {(f , p) → ∣ (λ x → -[ 0 ]ₖ (f x)) + , cong (λ x → -[ 0 ]ₖ x) p ∣₂} +-ₕ∙_ {n = suc zero} = sRec § λ {(f , p) → ∣ (λ x → -ₖ (f x)) + , cong -ₖ_ p ∣₂} +-ₕ∙_ {n = suc (suc n)} = sRec § λ {(f , p) → ∣ (λ x → -ₖ (f x)) + , cong -ₖ_ p ∣₂} + +0ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A +0ₕ∙ n = ∣ (λ _ → 0ₖ n) , refl ∣₂ + ++ₕ∙-syntax : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A ++ₕ∙-syntax n = _+ₕ∙_ {n = n} + +-ₕ∙-syntax : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A +-ₕ∙-syntax n = -ₕ∙_ {n = n} + +-'ₕ∙-syntax : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A +-'ₕ∙-syntax n x y = _+ₕ∙_ {n = n} x (-ₕ∙_ {n = n} y) + +syntax +ₕ∙-syntax n x y = x +[ n ]ₕ∙ y +syntax -ₕ∙-syntax n x = -[ n ]ₕ∙ x +syntax -'ₕ∙-syntax n x y = x -[ n ]ₕ∙ y + +commₕ∙ : {A : Pointed ℓ} (n : ℕ) (x y : coHomRed n A) → x +[ n ]ₕ∙ y ≡ y +[ n ]ₕ∙ x +commₕ∙ zero = + sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → commₖ 0 (f x) (g x) i)} +commₕ∙ (suc zero) = + sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) + → cong ∣_∣₂ (ΣPathP ((λ i x → commₖ 1 (f x) (g x) i) + , λ i j → commₖ 1 (p j) (q j) i))} +commₕ∙ {A = A} (suc (suc n)) = + sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) + → cong ∣_∣₂ (ΣPathP ((λ i x → commₖ (2 + n) (f x) (g x) i) + , λ i j → hcomp (λ k → λ {(i = i0) → p j +ₖ q j + ; (i = i1) → q j +ₖ p j + ; (j = i0) → commₖ (2 + n) (f (pt A)) (g (pt A)) i + ; (j = i1) → rUnit (refl {x = 0ₖ (2 + n)}) (~ k) i}) + (commₖ (2 + n) (p j) (q j) i)))} + +rUnitₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → x +[ n ]ₕ∙ 0ₕ∙ n ≡ x +rUnitₕ∙ zero = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → rUnitₖ zero (f x) i)} +rUnitₕ∙ (suc zero) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → rUnitₖ 1 (f x) i) , λ i j → rUnitₖ 1 (p j) i))} +rUnitₕ∙ (suc (suc n)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → rUnitₖ (2 + n) (f x) i) , λ i j → rUnitₖ (2 + n) (p j) i))} + +lUnitₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → 0ₕ∙ n +[ n ]ₕ∙ x ≡ x +lUnitₕ∙ zero = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → lUnitₖ zero (f x) i)} +lUnitₕ∙ (suc zero) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → lUnitₖ 1 (f x) i) , λ i j → lUnitₖ 1 (p j) i))} +lUnitₕ∙ (suc (suc n)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → lUnitₖ (2 + n) (f x) i) , λ i j → lUnitₖ (2 + n) (p j) i))} + + +private + pp : {A : Pointed ℓ} (n : ℕ) → (f : fst A → coHomK (suc (suc n))) + → (p : f (snd A) ≡ snd (coHomK-ptd (suc (suc n)))) + → PathP (λ i → rCancelₖ (2 + n) (f (snd A)) i ≡ 0ₖ (suc (suc n))) + (λ i → (p i) +ₖ (-ₖ p i)) (λ _ → 0ₖ (suc (suc n))) + pp {A = A} n f p i j = + hcomp (λ k → λ {(i = i0) → rCancelₖ (suc (suc n)) (p j) (~ k) + ; (i = i1) → 0ₖ (suc (suc n)) + ; (j = i0) → rCancelₖ (2 + n) (f (snd A)) (i ∨ ~ k) + ; (j = i1) → rUnit (rUnit (λ _ → 0ₖ (suc (suc n))) (~ i)) (~ i) k}) + (0ₖ (suc (suc n))) + +rCancelₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → x +[ n ]ₕ∙ (-[ n ]ₕ∙ x) ≡ 0ₕ∙ n +rCancelₕ∙ zero = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → rCancelₖ zero (f x) i)} +rCancelₕ∙ {A = A} (suc zero) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → rCancelₖ 1 (f x) i) , λ i j → rCancelₖ 1 (p j) i))} +rCancelₕ∙ {A = A} (suc (suc n)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) + → cong ∣_∣₂ (ΣPathP ((λ i x → rCancelₖ (2 + n) (f x) i) + , pp n f p))} + +lCancelₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → (-[ n ]ₕ∙ x) +[ n ]ₕ∙ x ≡ 0ₕ∙ n +lCancelₕ∙ zero = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → lCancelₖ zero (f x) i)} +lCancelₕ∙ {A = A} (suc zero) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) + → cong ∣_∣₂ (ΣPathP ((λ i x → lCancelₖ 1 (f x) i) + , λ i j → (lCancelₖ 1 (p j) i)))} +lCancelₕ∙ {A = A} (suc (suc n)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) + → cong ∣_∣₂ (ΣPathP ((λ i x → lCancelₖ (2 + n) (f x) i) + , λ i j → lCancelₖ (2 + n) (p j) i))} + +assocₕ∙ : {A : Pointed ℓ} (n : ℕ) (x y z : coHomRed n A) + → (x +[ n ]ₕ∙ (y +[ n ]ₕ∙ z)) ≡ ((x +[ n ]ₕ∙ y) +[ n ]ₕ∙ z) +assocₕ∙ zero = + elim3 (λ _ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) (h , r) + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (λ i x → assocₖ zero (f x) (g x) (h x) i))} +assocₕ∙ (suc zero) = + elim3 (λ _ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) (h , r) + → cong ∣_∣₂ (ΣPathP ((λ i x → assocₖ 1 (f x) (g x) (h x) i) + , λ i j → assocₖ 1 (p j) (q j) (r j) i))} +assocₕ∙ (suc (suc n)) = + elim3 (λ _ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) (h , r) + → cong ∣_∣₂ (ΣPathP ((λ i x → assocₖ (2 + n) (f x) (g x) (h x) i) + , λ i j → assocₖ (2 + n) (p j) (q j) (r j) i))} + +open IsSemigroup +open IsMonoid +open GroupStr +open IsGroupHom + +coHomGr : (n : ℕ) (A : Type ℓ) → Group ℓ +coHomGr n A = coHom n A , coHomGrnA + where + coHomGrnA : GroupStr (coHom n A) + 1g coHomGrnA = 0ₕ n + GroupStr._·_ coHomGrnA = λ x y → x +[ n ]ₕ y + inv coHomGrnA = λ x → -[ n ]ₕ x + isGroup coHomGrnA = helper + where + abstract + helper : IsGroup {G = coHom n A} (0ₕ n) (λ x y → x +[ n ]ₕ y) (λ x → -[ n ]ₕ x) + helper = makeIsGroup § (assocₕ n) (rUnitₕ n) (lUnitₕ n) (rCancelₕ n) (lCancelₕ n) + +×coHomGr : (n : ℕ) (A : Type ℓ) (B : Type ℓ') → Group _ +×coHomGr n A B = DirProd (coHomGr n A) (coHomGr n B) + +coHomGroup : (n : ℕ) (A : Type ℓ) → AbGroup ℓ +fst (coHomGroup n A) = coHom n A +AbGroupStr.0g (snd (coHomGroup n A)) = 0ₕ n +AbGroupStr._+_ (snd (coHomGroup n A)) = _+ₕ_ {n = n} +AbGroupStr.- snd (coHomGroup n A) = -ₕ_ {n = n} +IsAbGroup.isGroup (AbGroupStr.isAbGroup (snd (coHomGroup n A))) = isGroup (snd (coHomGr n A)) +IsAbGroup.comm (AbGroupStr.isAbGroup (snd (coHomGroup n A))) = commₕ n + +-- Reduced cohomology group (direct def) + +coHomRedGroupDir : (n : ℕ) (A : Pointed ℓ) → AbGroup ℓ +fst (coHomRedGroupDir n A) = coHomRed n A +AbGroupStr.0g (snd (coHomRedGroupDir n A)) = 0ₕ∙ n +AbGroupStr._+_ (snd (coHomRedGroupDir n A)) = _+ₕ∙_ {n = n} +AbGroupStr.- snd (coHomRedGroupDir n A) = -ₕ∙_ {n = n} +IsAbGroup.isGroup (AbGroupStr.isAbGroup (snd (coHomRedGroupDir n A))) = helper + where + abstract + helper : IsGroup {G = coHomRed n A} (0ₕ∙ n) (_+ₕ∙_ {n = n}) (-ₕ∙_ {n = n}) + helper = makeIsGroup § (assocₕ∙ n) (rUnitₕ∙ n) (lUnitₕ∙ n) (rCancelₕ∙ n) (lCancelₕ∙ n) +IsAbGroup.comm (AbGroupStr.isAbGroup (snd (coHomRedGroupDir n A))) = commₕ∙ n + +coHomRedGrDir : (n : ℕ) (A : Pointed ℓ) → Group ℓ +coHomRedGrDir n A = AbGroup→Group (coHomRedGroupDir n A) + +-- Induced map +coHomFun : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) (f : A → B) → coHom n B → coHom n A +coHomFun n f = sRec § λ β → ∣ β ∘ f ∣₂ + +coHomMorph : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) (f : A → B) → GroupHom (coHomGr n B) (coHomGr n A) +fst (coHomMorph n f) = coHomFun n f +snd (coHomMorph n f) = makeIsGroupHom (helper n) + where + helper : ℕ → _ + helper zero = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + helper (suc zero) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + helper (suc (suc n)) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + +coHomIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) → Iso A B + → GroupIso (coHomGr n B) (coHomGr n A) +fun (fst (coHomIso n is)) = fst (coHomMorph n (fun is)) +inv' (fst (coHomIso n is)) = fst (coHomMorph n (inv' is)) +rightInv (fst (coHomIso n is)) = + sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (leftInv is x)) +leftInv (fst (coHomIso n is)) = + sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (rightInv is x)) +snd (coHomIso n is) = snd (coHomMorph n (fun is)) + +-- Alternative definition of cohomology using ΩKₙ instead. Useful for breaking proofs of group isos +-- up into smaller parts +coHomGrΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Group ℓ +coHomGrΩ n A = ∥ (A → typ (Ω (coHomK-ptd (suc n)))) ∥₂ , coHomGrnA + where + coHomGrnA : GroupStr ∥ (A → typ (Ω (coHomK-ptd (suc n)))) ∥₂ + 1g coHomGrnA = ∣ (λ _ → refl) ∣₂ + GroupStr._·_ coHomGrnA = sRec2 § λ p q → ∣ (λ x → p x ∙ q x) ∣₂ + inv coHomGrnA = map λ f x → sym (f x) + isGroup coHomGrnA = helper + where + abstract + helper : + IsGroup {G = ∥ (A → typ (Ω (coHomK-ptd (suc n)))) ∥₂} + (∣ (λ _ → refl) ∣₂) (sRec2 § λ p q → ∣ (λ x → p x ∙ q x) ∣₂) (map λ f x → sym (f x)) + helper = makeIsGroup § (elim3 (λ _ _ _ → isOfHLevelPath 2 § _ _) + (λ p q r → cong ∣_∣₂ (funExt λ x → assoc∙ (p x) (q x) (r x)))) + (sElim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → sym (rUnit (p x)))) + (sElim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → sym (lUnit (p x)))) + (sElim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → rCancel (p x))) + (sElim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → lCancel (p x))) + +--- the loopspace of Kₙ is commutative regardless of base +addIso : (n : ℕ) (x : coHomK n) → Iso (coHomK n) (coHomK n) +fun (addIso n x) y = y +[ n ]ₖ x +inv' (addIso n x) y = y -[ n ]ₖ x +rightInv (addIso n x) y = -+cancelₖ n y x +leftInv (addIso n x) y = -cancelRₖ n x y + +baseChange : (n : ℕ) (x : coHomK (suc n)) → (0ₖ (suc n) ≡ 0ₖ (suc n)) ≃ (x ≡ x) +baseChange n x = isoToEquiv is + where + f : (n : ℕ) (x : coHomK (suc n)) → (0ₖ (suc n) ≡ 0ₖ (suc n)) → (x ≡ x) + f n x p = sym (rUnitₖ _ x) ∙∙ cong (x +ₖ_) p ∙∙ rUnitₖ _ x + + g : (n : ℕ) (x : coHomK (suc n)) → (x ≡ x) → (0ₖ (suc n) ≡ 0ₖ (suc n)) + g n x p = sym (rCancelₖ _ x) ∙∙ cong (λ y → y -ₖ x) p ∙∙ rCancelₖ _ x + + f-g : (n : ℕ) (x : coHomK (suc n)) (p : x ≡ x) → f n x (g n x p) ≡ p + f-g n = + trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelPath (3 + n) + (isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) _ _) + (ind n) + where + ind : (n : ℕ) (a : S₊ (suc n)) (p : ∣ a ∣ₕ ≡ ∣ a ∣ₕ) → f n ∣ a ∣ₕ (g n ∣ a ∣ₕ p) ≡ p + ind zero = + toPropElim (λ _ → isPropΠ λ _ → isOfHLevelTrunc 3 _ _ _ _) + λ p → cong (f zero (0ₖ 1)) (sym (rUnit _) ∙ (λ k i → rUnitₖ _ (p i) k)) + ∙∙ sym (rUnit _) + ∙∙ λ k i → lUnitₖ _ (p i) k + ind (suc n) = + sphereElim (suc n) (λ _ → isOfHLevelΠ (2 + n) λ _ → isOfHLevelTrunc (4 + n) _ _ _ _) + λ p → cong (f (suc n) (0ₖ (2 + n))) + ((λ k → (sym (rUnit (refl ∙ refl)) ∙ sym (rUnit refl)) k + ∙∙ (λ i → p i +ₖ 0ₖ (2 + n)) ∙∙ (sym (rUnit (refl ∙ refl)) ∙ sym (rUnit refl)) k) + ∙ (λ k → rUnit (λ i → rUnitₖ _ (p i) k) (~ k))) + ∙ λ k → rUnit (λ i → lUnitₖ _ (p i) k) (~ k) + + g-f : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ 0ₖ _) → g n x (f n x p) ≡ p + g-f n = + trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelPath (3 + n) + (isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) _ _) + (ind n) + where + ind : (n : ℕ) (a : S₊ (suc n)) (p : 0ₖ (suc n) ≡ 0ₖ (suc n)) → g n ∣ a ∣ₕ (f n ∣ a ∣ₕ p) ≡ p + ind zero = + toPropElim (λ _ → isPropΠ λ _ → isOfHLevelTrunc 3 _ _ _ _) + λ p → cong (g zero (0ₖ 1)) (λ k → rUnit (λ i → lUnitₖ _ (p i) k) (~ k)) + ∙ (λ k → rUnit (λ i → rUnitₖ _ (p i) k) (~ k)) + ind (suc n) = + sphereElim (suc n) (λ _ → isOfHLevelΠ (2 + n) λ _ → isOfHLevelTrunc (4 + n) _ _ _ _) + λ p → cong (g (suc n) (0ₖ (2 + n))) + (λ k → rUnit (λ i → lUnitₖ _ (p i) k) (~ k)) + ∙∙ (λ k → (sym (rUnit (refl ∙ refl)) ∙ sym (rUnit refl)) k + ∙∙ (λ i → p i +ₖ 0ₖ (2 + n)) + ∙∙ (sym (rUnit (refl ∙ refl)) ∙ sym (rUnit refl)) k) + ∙∙ λ k → rUnit (λ i → rUnitₖ _ (p i) k) (~ k) + + is : Iso _ _ + fun is = f n x + inv' is = g n x + rightInv is = f-g n x + leftInv is = g-f n x + +isCommΩK-based : (n : ℕ) (x : coHomK n) → isComm∙ (coHomK n , x) +isCommΩK-based zero x p q = isSetℤ _ _ (p ∙ q) (q ∙ p) +isCommΩK-based (suc zero) x = + subst isComm∙ (λ i → coHomK 1 , lUnitₖ 1 x i) + (ptdIso→comm {A = (_ , 0ₖ 1)} (addIso 1 x) + (isCommΩK 1)) +isCommΩK-based (suc (suc n)) x = + subst isComm∙ (λ i → coHomK (suc (suc n)) , lUnitₖ (suc (suc n)) x i) + (ptdIso→comm {A = (_ , 0ₖ (suc (suc n)))} (addIso (suc (suc n)) x) + (isCommΩK (suc (suc n)))) + +-- hidden versions of cohom stuff using the "lock" hack. The locked versions can be used when proving things. +-- Swapping "key" for "tt*" will then give computing functions. +Unit' : Type₀ +Unit' = lockUnit {ℓ-zero} + +lock : ∀ {ℓ} {A : Type ℓ} → Unit' → A → A +lock unlock = λ x → x + +module lockedCohom (key : Unit') where + +K : (n : ℕ) → coHomK n → coHomK n → coHomK n + +K n = lock key (_+ₖ_ {n = n}) + + -K : (n : ℕ) → coHomK n → coHomK n + -K n = lock key (-ₖ_ {n = n}) + + -Kbin : (n : ℕ) → coHomK n → coHomK n → coHomK n + -Kbin n x y = +K n x (-K n y) + + rUnitK : (n : ℕ) (x : coHomK n) → +K n x (0ₖ n) ≡ x + rUnitK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (0ₖ n) ≡ x + pm unlock = rUnitₖ n x + + lUnitK : (n : ℕ) (x : coHomK n) → +K n (0ₖ n) x ≡ x + lUnitK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (0ₖ n) x ≡ x + pm unlock = lUnitₖ n x + + rCancelK : (n : ℕ) (x : coHomK n) → +K n x (-K n x) ≡ 0ₖ n + rCancelK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (lock t (-ₖ_ {n = n}) x) ≡ 0ₖ n + pm unlock = rCancelₖ n x + + lCancelK : (n : ℕ) (x : coHomK n) → +K n (-K n x) x ≡ 0ₖ n + lCancelK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (-ₖ_ {n = n}) x) x ≡ 0ₖ n + pm unlock = lCancelₖ n x + + -cancelRK : (n : ℕ) (x y : coHomK n) → -Kbin n (+K n y x) x ≡ y + -cancelRK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) y x) (lock t (-ₖ_ {n = n}) x) ≡ y + pm unlock = -cancelRₖ n x y + + -cancelLK : (n : ℕ) (x y : coHomK n) → -Kbin n (+K n x y) x ≡ y + -cancelLK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x y) (lock t (-ₖ_ {n = n}) x) ≡ y + pm unlock = -cancelLₖ n x y + + -+cancelK : (n : ℕ) (x y : coHomK n) → +K n (-Kbin n x y) y ≡ x + -+cancelK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x (lock t (-ₖ_ {n = n}) y)) y ≡ x + pm unlock = -+cancelₖ n x y + + assocK : (n : ℕ) (x y z : coHomK n) → +K n x (+K n y z) ≡ +K n (+K n x y) z + assocK n x y z = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (lock t (_+ₖ_ {n = n}) y z) + ≡ lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x y) z + pm unlock = assocₖ n x y z + + commK : (n : ℕ) (x y : coHomK n) → +K n x y ≡ +K n y x + commK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x y ≡ lock t (_+ₖ_ {n = n}) y x + pm unlock = commₖ n x y + + -- cohom + + +H : (n : ℕ) (x y : coHom n A) → coHom n A + +H n = sRec2 § λ a b → ∣ (λ x → +K n (a x) (b x)) ∣₂ + + -H : (n : ℕ) (x : coHom n A) → coHom n A + -H n = sRec § λ a → ∣ (λ x → -K n (a x)) ∣₂ + + -Hbin : (n : ℕ) → coHom n A → coHom n A → coHom n A + -Hbin n = sRec2 § λ a b → ∣ (λ x → -Kbin n (a x) (b x)) ∣₂ + + rUnitH : (n : ℕ) (x : coHom n A) → +H n x (0ₕ n) ≡ x + rUnitH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitK n (a x)) i ∣₂ + + lUnitH : (n : ℕ) (x : coHom n A) → +H n (0ₕ n) x ≡ x + lUnitH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitK n (a x)) i ∣₂ + + rCancelH : (n : ℕ) (x : coHom n A) → +H n x (-H n x) ≡ 0ₕ n + rCancelH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelK n (a x)) i ∣₂ + + lCancelH : (n : ℕ) (x : coHom n A) → +H n (-H n x) x ≡ 0ₕ n + lCancelH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelK n (a x)) i ∣₂ + + assocH : (n : ℕ) (x y z : coHom n A) → (+H n x (+H n y z)) ≡ (+H n (+H n x y) z) + assocH n = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocK n (a x) (b x) (c x)) i ∣₂ + + commH : (n : ℕ) (x y : coHom n A) → (+H n x y) ≡ (+H n y x) + commH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commK n (a x) (b x)) i ∣₂ + + -cancelRH : (n : ℕ) (x y : coHom n A) → -Hbin n (+H n y x) x ≡ y + -cancelRH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelRK n (a x) (b x) i) ∣₂ + + -cancelLH : (n : ℕ) (x y : coHom n A) → -Hbin n (+H n x y) x ≡ y + -cancelLH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelLK n (a x) (b x) i) ∣₂ + + -+cancelH : (n : ℕ) (x y : coHom n A) → +H n (-Hbin n x y) y ≡ x + -+cancelH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -+cancelK n (a x) (b x) i) ∣₂ + +lUnitK≡rUnitK : (key : Unit') (n : ℕ) → lockedCohom.lUnitK key n (0ₖ n) ≡ lockedCohom.rUnitK key n (0ₖ n) +lUnitK≡rUnitK unlock = lUnitₖ≡rUnitₖ + +open GroupStr renaming (_·_ to _+gr_) +open IsGroupHom + +-- inducedCoHom : ∀ {ℓ ℓ'} {A : Type ℓ} {G : Group {ℓ'}} {n : ℕ} +-- → GroupIso (coHomGr n A) G +-- → Group +-- inducedCoHom {A = A} {G = G} {n = n} e = +-- InducedGroup (coHomGr n A) +-- (coHom n A , λ x y → Iso.inv (isom e) (_+gr_ (snd G) (fun (isom e) x) +-- (fun (isom e) y))) +-- (idEquiv _) +-- λ x y → sym (leftInv (isom e) _) +-- ∙ cong (Iso.inv (isom e)) (isHom e x y) + +-- induced+ : ∀ {ℓ ℓ'} {A : Type ℓ} {G : Group {ℓ'}} {n : ℕ} +-- → (e : GroupIso (coHomGr n A) G) +-- → fst (inducedCoHom e) → fst (inducedCoHom e) → fst (inducedCoHom e) +-- induced+ e = _+gr_ (snd (inducedCoHom e)) + +-- inducedCoHomIso : ∀ {ℓ ℓ'} {A : Type ℓ} {G : Group {ℓ'}} {n : ℕ} +-- → (e : GroupIso (coHomGr n A) G) +-- → GroupIso (coHomGr n A) (inducedCoHom e) +-- isom (inducedCoHomIso e) = idIso +-- isHom (inducedCoHomIso e) x y = sym (leftInv (isom e) _) +-- ∙ cong (Iso.inv (isom e)) (isHom e x y) + +-- inducedCoHomPath : ∀ {ℓ ℓ'} {A : Type ℓ} {G : Group {ℓ'}} {n : ℕ} +-- → (e : GroupIso (coHomGr n A) G) +-- → coHomGr n A ≡ inducedCoHom e +-- inducedCoHomPath e = InducedGroupPath _ _ _ _ diff --git a/Cubical/ZCohomology/Groups/CP2.agda b/Cubical/ZCohomology/Groups/CP2.agda new file mode 100644 index 0000000000..0d70660920 --- /dev/null +++ b/Cubical/ZCohomology/Groups/CP2.agda @@ -0,0 +1,249 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.CP2 where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sigma +open import Cubical.Data.Int +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) + +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Join +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim2 to pElim2 ; ∣_∣ to ∣_∣₁ ; map to pMap) +open import Cubical.HITs.Truncation + +open import Cubical.Relation.Nullary + +open import Cubical.Homotopy.Hopf +open S¹Hopf + +open IsGroupHom +open Iso + +CP² : Type +CP² = Pushout {A = TotalHopf} fst λ _ → tt + +characFunSpaceCP² : ∀ {ℓ} {A : Type ℓ} + → Iso (CP² → A) (Σ[ x ∈ A ] Σ[ f ∈ (S₊ (suc (suc zero)) → A) ] + ((y : TotalHopf) → f (fst y) ≡ x)) +fun characFunSpaceCP² f = (f (inr tt)) , ((f ∘ inl ) , (λ a → cong f (push a))) +inv characFunSpaceCP² (a , f , p) (inl x) = f x +inv characFunSpaceCP² (a , f , p) (inr x) = a +inv characFunSpaceCP² (a , f , p) (push b i) = p b i +rightInv characFunSpaceCP² _ = refl +leftInv characFunSpaceCP² _ = + funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl} + +H⁰CP²≅ℤ : GroupIso (coHomGr 0 CP²) ℤGroup +H⁰CP²≅ℤ = + H⁰-connected (inr tt) + (PushoutToProp (λ _ → squash) + (sphereElim _ (λ _ → isOfHLevelSuc 1 squash) + ∣ sym (push (north , base)) ∣₁) + λ _ → ∣ refl ∣₁) + +module M = MV (S₊ 2) Unit TotalHopf fst (λ _ → tt) + +H²CP²≅ℤ : GroupIso (coHomGr 2 CP²) ℤGroup +H²CP²≅ℤ = compGroupIso (BijectionIso→GroupIso bij) + (compGroupIso (invGroupIso trivIso) (Hⁿ-Sⁿ≅ℤ 1)) + where + isContrH¹TotalHopf : isContr (coHom 1 TotalHopf) + isContrH¹TotalHopf = + isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) + (isOfHLevelRetractFromIso 0 ((fst (H¹-Sⁿ≅0 1))) isContrUnit) + + isContrH²TotalHopf : isContr (coHom 2 TotalHopf) + isContrH²TotalHopf = + isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) + ((isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 1 2 λ p → snotz (sym (cong predℕ p)))) isContrUnit)) + + trivIso : GroupIso (coHomGr 2 (Susp S¹)) (×coHomGr 2 (Susp S¹) Unit) + fun (fst trivIso) x = x , 0ₕ _ + inv (fst trivIso) = fst + rightInv (fst trivIso) (x , p) = + ΣPathP (refl , isContr→isProp (isContrHⁿ-Unit 1) _ _) + leftInv (fst trivIso) x = refl + snd trivIso = makeIsGroupHom λ _ _ → refl + + bij : BijectionIso (coHomGr 2 CP²) (×coHomGr 2 (Susp S¹) Unit) + BijectionIso.fun bij = M.i 2 + BijectionIso.inj bij x p = + pRec (squash₂ _ _) + (uncurry (λ z q + → sym q + ∙∙ cong (fst (M.d 1)) (isContr→isProp isContrH¹TotalHopf z (0ₕ _)) + ∙∙ (M.d 1) .snd .pres1)) + (M.Ker-i⊂Im-d 1 x p) + where + help : isInIm (M.d 1) x + help = M.Ker-i⊂Im-d 1 x p + BijectionIso.surj bij y = + M.Ker-Δ⊂Im-i 2 y (isContr→isProp isContrH²TotalHopf _ _) + +H⁴CP²≅ℤ : GroupIso (coHomGr 4 CP²) ℤGroup +H⁴CP²≅ℤ = compGroupIso (invGroupIso (BijectionIso→GroupIso bij)) + (compGroupIso help (Hⁿ-Sⁿ≅ℤ 2)) + where + help : GroupIso (coHomGr 3 TotalHopf) (coHomGr 3 (S₊ 3)) + help = isoType→isoCohom 3 (invIso IsoS³TotalHopf) + + bij : BijectionIso (coHomGr 3 TotalHopf) (coHomGr 4 CP²) + BijectionIso.fun bij = M.d 3 + BijectionIso.inj bij x p = + pRec (squash₂ _ _) + (uncurry (λ z q → + sym q + ∙∙ cong (M.Δ 3 .fst) + (isOfHLevelΣ 1 (isContr→isProp + (isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) isContrUnit)) + (λ _ → isContr→isProp (isContrHⁿ-Unit 2)) + z (0ₕ _ , 0ₕ _)) + ∙∙ M.Δ 3 .snd .pres1)) + (M.Ker-d⊂Im-Δ _ x p) + BijectionIso.surj bij y = + M.Ker-i⊂Im-d 3 y (isOfHLevelΣ 1 + (isContr→isProp (isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) isContrUnit)) + (λ _ → isContr→isProp (isContrHⁿ-Unit _)) _ _) + + +-- Characterisations of the trivial groups + +private + elim-TotalHopf : (B : TotalHopf → Type) + → ((x : _) → isOfHLevel 3 (B x)) → B (north , base) + → (x : _) → B x + elim-TotalHopf = + transport (λ i → (B : isoToPath IsoS³TotalHopf i → Type) + → ((x : _) → isOfHLevel 3 (B x)) + → B (transp (λ j → isoToPath IsoS³TotalHopf (i ∨ ~ j)) i (north , base)) → (x : _) → B x) + λ B hLev elim-TotalHopf → sphereElim _ (λ _ → hLev _) elim-TotalHopf + +H¹-CP²≅0 : GroupIso (coHomGr 1 CP²) UnitGroup +H¹-CP²≅0 = + contrGroupIsoUnit + (isOfHLevelRetractFromIso 0 (setTruncIso characFunSpaceCP²) + (isOfHLevelRetractFromIso 0 lem₂ lem₃)) + where + lem₁ : (f : (Susp S¹ → coHomK 1)) → ∥ (λ _ → 0ₖ _) ≡ f ∥ + lem₁ f = pMap (λ p → p) + (Iso.fun PathIdTrunc₀Iso (isOfHLevelRetractFromIso 1 + (fst (Hⁿ-Sᵐ≅0 0 1 (λ p → snotz (sym p)))) isPropUnit (0ₕ _) ∣ f ∣₂)) + + lem₂ : Iso ∥ (Σ[ x ∈ coHomK 1 ] ( Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ x))) ∥₂ + ∥ (Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ 0ₖ 1)) ∥₂ + fun lem₂ = sMap (uncurry λ x → uncurry λ f p → (λ y → (-ₖ x) +ₖ f y) , λ y → cong ((-ₖ x) +ₖ_) (p y) ∙ lCancelₖ _ x) + inv lem₂ = sMap λ p → 0ₖ _ , p + rightInv lem₂ = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) + , (funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) + ◁ λ j y i → lUnitₖ _ (p y i) j)))} + leftInv lem₂ = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + (uncurry (coHomK-elim _ (λ _ → isPropΠ (λ _ → squash₂ _ _)) + (uncurry λ f p → cong ∣_∣₂ (ΣPathP (refl , (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) + , ((funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) + ◁ λ j y i → lUnitₖ _ (p y i) j))))))))) + + lem₃ : isContr _ + fst lem₃ = ∣ (λ _ → 0ₖ 1) , (λ _ → refl) ∣₂ + snd lem₃ = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + (uncurry λ f → pRec (isPropΠ (λ _ → squash₂ _ _)) + (J (λ f _ → (y : (y₁ : TotalHopf) → f (fst y₁) ≡ 0ₖ 1) → + ∣ (λ _ → 0ₖ 1) , (λ _ _ → 0ₖ 1) ∣₂ ≡ ∣ f , y ∣₂) + (λ y → cong ∣_∣₂ (ΣPathP ((funExt (λ z → sym (y (north , base)))) , toPathP (s y))))) + (lem₁ f)) + + where + s : (y : TotalHopf → 0ₖ 1 ≡ 0ₖ 1) + → transport (λ i → (_ : TotalHopf) → y (north , base) (~ i) ≡ ∣ base ∣) + (λ _ _ → 0ₖ 1) ≡ y + s y = funExt (elim-TotalHopf _ (λ _ → isOfHLevelPath 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) _ _) + λ k → transp (λ i → y (north , base) (~ i ∧ ~ k) ≡ ∣ base ∣) k + λ j → y (north , base) (~ k ∨ j)) + +Hⁿ-CP²≅0-higher : (n : ℕ) → ¬ (n ≡ 1) → GroupIso (coHomGr (3 +ℕ n) CP²) UnitGroup +Hⁿ-CP²≅0-higher n p = contrGroupIsoUnit ((0ₕ _) , (λ x → sym (main x))) + where + h : GroupHom (coHomGr (2 +ℕ n) TotalHopf) (coHomGr (3 +ℕ n) CP²) + h = M.d (2 +ℕ n) + + propᵣ : isProp (fst (×coHomGr (3 +ℕ n) (Susp S¹) Unit)) + propᵣ = + isPropΣ + (isOfHLevelRetractFromIso 1 + (fst (Hⁿ-Sᵐ≅0 (2 +ℕ n) 1 λ p → ⊥-rec (snotz (cong predℕ p)))) isPropUnit) + λ _ → isContr→isProp (isContrHⁿ-Unit _) + + propₗ : isProp (coHom (2 +ℕ n) TotalHopf) + propₗ = subst (λ x → isProp (coHom (2 +ℕ n) x)) (isoToPath IsoS³TotalHopf) + (isOfHLevelRetractFromIso 1 + (fst (Hⁿ-Sᵐ≅0 (suc n) 2 λ q → p (cong predℕ q))) isPropUnit) + + inIm : (x : coHom (3 +ℕ n) CP²) → isInIm h x + inIm x = M.Ker-i⊂Im-d (2 +ℕ n) x (propᵣ _ _) + + main : (x : coHom (3 +ℕ n) CP²) → x ≡ 0ₕ _ + main x = + pRec (squash₂ _ _) + (uncurry (λ f p → sym p ∙∙ cong (h .fst) (propₗ f (0ₕ _)) ∙∙ pres1 (snd h))) + (inIm x) + +-- All trivial groups: +Hⁿ-CP²≅0 : (n : ℕ) → ¬ suc n ≡ 2 → ¬ suc n ≡ 4 + → GroupIso (coHomGr (suc n) CP²) UnitGroup +Hⁿ-CP²≅0 zero p q = H¹-CP²≅0 +Hⁿ-CP²≅0 (suc zero) p q = ⊥-rec (p refl) +Hⁿ-CP²≅0 (suc (suc zero)) p q = Hⁿ-CP²≅0-higher 0 λ p → snotz (sym p) +Hⁿ-CP²≅0 (suc (suc (suc zero))) p q = ⊥-rec (q refl) +Hⁿ-CP²≅0 (suc (suc (suc (suc n)))) p q = + Hⁿ-CP²≅0-higher (suc (suc n)) + λ p → snotz (cong predℕ p) + +-- Another Brunerie number +ℤ→HⁿCP²→ℤ : ℤ → ℤ +ℤ→HⁿCP²→ℤ x = + Iso.fun (fst H⁴CP²≅ℤ) + (Iso.inv (fst H²CP²≅ℤ) x ⌣ Iso.inv (fst H²CP²≅ℤ) x) + +brunerie2 : ℤ +brunerie2 = ℤ→HⁿCP²→ℤ 1 + +{- +|brunerie2|≡1 : abs (ℤ→HⁿCP²→ℤ 1) ≡ 1 +|brunerie2|≡1 = refl +-} diff --git a/Cubical/ZCohomology/Groups/Connected.agda b/Cubical/ZCohomology/Groups/Connected.agda new file mode 100644 index 0000000000..b134d252b3 --- /dev/null +++ b/Cubical/ZCohomology/Groups/Connected.agda @@ -0,0 +1,45 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Groups.Connected where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Unit + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) +open import Cubical.HITs.Nullification + +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.Data.Int hiding (ℤ) renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (rec₊ to trRec) +open import Cubical.Algebra.Group + +open import Cubical.Homotopy.Connected +open import Cubical.Foundations.Equiv + +private + H⁰-connected-type : ∀ {ℓ} {A : Type ℓ} (a : A) → isConnected 2 A → Iso (coHom 0 A) (fst ℤ) + Iso.fun (H⁰-connected-type a con) = sRec isSetℤ λ f → f a + Iso.inv (H⁰-connected-type a con) b = ∣ (λ x → b) ∣₂ + Iso.rightInv (H⁰-connected-type a con) b = refl + Iso.leftInv (H⁰-connected-type a con) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ∣_∣₂ (funExt λ x → trRec (isSetℤ _ _) (cong f) (isConnectedPath 1 con a x .fst)) + +open IsGroupHom +open Iso + +H⁰-connected : ∀ {ℓ} {A : Type ℓ} (a : A) → ((x : A) → ∥ a ≡ x ∥₁) → GroupIso (coHomGr 0 A) ℤ +fun (fst (H⁰-connected a con)) = sRec isSetℤ (λ f → f a) +inv (fst (H⁰-connected a con)) b = ∣ (λ _ → b) ∣₂ +rightInv (fst (H⁰-connected a con)) _ = refl +leftInv (fst (H⁰-connected a con)) = + sElim (λ _ → isProp→isSet (isSetSetTrunc _ _)) + (λ f → cong ∣_∣₂ (funExt λ x → pRec (isSetℤ _ _) (cong f) (con x))) +snd (H⁰-connected a con) = makeIsGroupHom (sElim2 (λ _ _ → isProp→isSet (isSetℤ _ _)) λ x y → refl) diff --git a/Cubical/ZCohomology/Groups/KleinBottle.agda b/Cubical/ZCohomology/Groups/KleinBottle.agda new file mode 100644 index 0000000000..15945515cd --- /dev/null +++ b/Cubical/ZCohomology/Groups/KleinBottle.agda @@ -0,0 +1,454 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.KleinBottle where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to pRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∣_∣ to ∣_∣₁) +open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; elim2 to trElim2) +open import Cubical.Data.Nat hiding (isEven) +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) + +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Transport + +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn + +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.Equiv +open import Cubical.Homotopy.Connected + +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Bool +open import Cubical.Data.Int renaming (+Comm to +-commℤ ; _+_ to _+ℤ_) + +open import Cubical.HITs.KleinBottle +open import Cubical.Data.Empty +open import Cubical.Foundations.Path + +open import Cubical.Homotopy.Loopspace + +open IsGroupHom +open Iso + +characFunSpace𝕂² : ∀ {ℓ} (A : Type ℓ) → + Iso (KleinBottle → A) + (Σ[ x ∈ A ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) +fun (characFunSpace𝕂² A) f = + (f point) , + ((cong f line1) , + (cong f line2 , + fst (Square≃doubleComp + (cong f line2) (cong f line2) + (sym (cong f line1)) (cong f line1)) + (λ i j → f (square i j)))) +inv (characFunSpace𝕂² A) (x , p , q , sq) point = x +inv (characFunSpace𝕂² A) (x , p , q , sq) (line1 i) = p i +inv (characFunSpace𝕂² A) (x , p , q , sq) (line2 i) = q i +inv (characFunSpace𝕂² A) (x , p , q , sq) (square i j) = + invEq (Square≃doubleComp q q (sym p) p) sq i j +rightInv (characFunSpace𝕂² A) (x , (p , (q , sq))) = + ΣPathP (refl , (ΣPathP (refl , (ΣPathP (refl , secEq (Square≃doubleComp q q (sym p) p) sq))))) +leftInv (characFunSpace𝕂² A) f _ point = f point +leftInv (characFunSpace𝕂² A) f _ (line1 i) = f (line1 i) +leftInv (characFunSpace𝕂² A) f _ (line2 i) = f (line2 i) +leftInv (characFunSpace𝕂² A) f z (square i j) = + retEq (Square≃doubleComp + (cong f line2) (cong f line2) + (sym (cong f line1)) (cong f line1)) + (λ i j → f (square i j)) z i j +private + movePathLem : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) → isComm∙ (A , x) + → (p ∙∙ q ∙∙ p ≡ q) ≡ ((p ∙ p) ∙ q ≡ q) + movePathLem p q comm = + cong (_≡ q) (doubleCompPath-elim' p q p ∙∙ cong (p ∙_) (comm q p) ∙∙ assoc _ _ _) + + movePathLem2 : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) + → (((p ∙ p) ∙ q) ∙ sym q ≡ q ∙ sym q) ≡ (p ∙ p ≡ refl) + movePathLem2 p q = + cong₂ _≡_ (sym (assoc (p ∙ p) q (sym q)) ∙∙ cong ((p ∙ p) ∙_) (rCancel q) ∙∙ sym (rUnit (p ∙ p))) + (rCancel q) + + movePathIso : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) → isComm∙ (A , x) + → Iso (p ∙∙ q ∙∙ p ≡ q) (p ∙ p ≡ refl) + movePathIso {x = x} p q comm = + compIso (pathToIso (movePathLem p q comm)) + (compIso (helper (p ∙ p)) + (pathToIso (movePathLem2 p q))) + where + helper : (p : x ≡ x) → Iso (p ∙ q ≡ q) ((p ∙ q) ∙ sym q ≡ q ∙ sym q) + helper p = congIso (equivToIso (_ , compPathr-isEquiv (sym q))) + +------ H¹(𝕂²) ≅ 0 -------------- +H⁰-𝕂² : GroupIso (coHomGr 0 KleinBottle) ℤGroup +fun (fst H⁰-𝕂²) = sRec isSetℤ λ f → f point +inv (fst H⁰-𝕂²) x = ∣ (λ _ → x) ∣₂ +rightInv (fst H⁰-𝕂²) _ = refl +leftInv (fst H⁰-𝕂²) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ∣_∣₂ (funExt (λ {point → refl + ; (line1 i) j → isSetℤ (f point) (f point) refl (cong f line1) j i + ; (line2 i) j → isSetℤ (f point) (f point) refl (cong f line2) j i + ; (square i j) z → helper f i j z})) + where + helper : (f : KleinBottle → ℤ) + → Cube (λ j z → isSetℤ (f point) (f point) refl (cong f line2) z j) + (λ j z → isSetℤ (f point) (f point) refl (cong f line2) z j) + (λ i z → isSetℤ (f point) (f point) refl (cong f line1) z (~ i)) + (λ i z → isSetℤ (f point) (f point) refl (cong f line1) z i) + refl + λ i j → f (square i j) + helper f = isGroupoid→isGroupoid' (isOfHLevelSuc 2 isSetℤ) _ _ _ _ _ _ +snd H⁰-𝕂² = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) λ _ _ → refl) + +------ H¹(𝕂¹) ≅ ℤ ------------ +{- +Step one : +H¹(𝕂²) := ∥ 𝕂² → K₁ ∥₂ + ≡ ∥ Σ[ x ∈ K₁ ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] (p ∙∙ q ∙∙ p ≡ q) ∥₂ (characFunSpace𝕂²) + ≡ ∥ Σ[ x ∈ K₁ ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ (movePathIso, using commutativity of ΩK₂) + ≡ ∥ Σ[ x ∈ K₁ ] (x ≡ x) ∥₂ (p ∙ p ≡ refl forces p ≡ refl. Also, p ∙ p ≡ refl is an hProp) +-} + +nilpotent→≡0 : (x : ℤ) → x +ℤ x ≡ 0 → x ≡ 0 +nilpotent→≡0 (pos zero) p = refl +nilpotent→≡0 (pos (suc n)) p = + ⊥-rec (negsucNotpos _ _ + (sym (cong (_- 1) (cong sucℤ (sym (helper2 n)) ∙ p)))) + where + helper2 : (n : ℕ) → pos (suc n) +pos n ≡ pos (suc (n + n)) + helper2 zero = refl + helper2 (suc n) = cong sucℤ (sym (sucℤ+pos n (pos (suc n)))) + ∙∙ cong (sucℤ ∘ sucℤ) (helper2 n) + ∙∙ cong (pos ∘ suc ∘ suc) (sym (+-suc n n)) +nilpotent→≡0 (negsuc n) p = ⊥-rec (negsucNotpos _ _ (helper2 n p)) + where + helper2 : (n : ℕ) → (negsuc n +negsuc n) ≡ pos 0 → negsuc n ≡ pos (suc n) + helper2 n p = cong (negsuc n +ℤ_) (sym (helper3 n)) + ∙ +Assoc (negsuc n) (negsuc n) (pos (suc n)) + ∙∙ cong (_+ℤ (pos (suc n))) p + ∙∙ cong sucℤ (+-commℤ (pos 0) (pos n)) + where + helper3 : (n : ℕ) → negsuc n +pos (suc n) ≡ 0 + helper3 zero = refl + helper3 (suc n) = cong sucℤ (sucℤ+pos n (negsuc (suc n))) ∙ helper3 n + +nilpotent→≡refl : (x : coHomK 1) (p : x ≡ x) → p ∙ p ≡ refl → p ≡ refl +nilpotent→≡refl = + trElim (λ _ → isGroupoidΠ2 λ _ _ → isOfHLevelPlus {n = 1} 2 (isOfHLevelTrunc 3 _ _ _ _)) + (toPropElim (λ _ → isPropΠ2 λ _ _ → isOfHLevelTrunc 3 _ _ _ _) + λ p pId → sym (rightInv (Iso-Kn-ΩKn+1 0) p) + ∙∙ cong (Kn→ΩKn+1 0) (nilpotent→≡0 (ΩKn+1→Kn 0 p) + (sym (ΩKn+1→Kn-hom 0 p p) + ∙ cong (ΩKn+1→Kn 0) pId)) + ∙∙ Kn→ΩKn+10ₖ 0) + +Iso-H¹-𝕂²₁ : Iso (Σ[ x ∈ coHomK 1 ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙ p ≡ refl) + (Σ[ x ∈ coHomK 1 ] x ≡ x) +fun Iso-H¹-𝕂²₁ (x , (_ , (q , _))) = x , q +inv Iso-H¹-𝕂²₁ (x , q) = x , (refl , (q , (sym (rUnit refl)))) +rightInv Iso-H¹-𝕂²₁ _ = refl +leftInv Iso-H¹-𝕂²₁ (x , (p , (q , P))) = + ΣPathP (refl , + (ΣPathP (sym (nilpotent→≡refl x p P) + , toPathP (Σ≡Prop (λ _ → isOfHLevelTrunc 3 _ _ _ _) + (transportRefl q))))) + +{- But this is precisely the type (minus set-truncation) of H¹(S¹) -} +Iso-H¹-𝕂²₂ : Iso (Σ[ x ∈ coHomK 1 ] x ≡ x) (S¹ → coHomK 1) +Iso-H¹-𝕂²₂ = invIso IsoFunSpaceS¹ + +H¹-𝕂²≅ℤ : GroupIso (coHomGr 1 KleinBottle) ℤGroup +H¹-𝕂²≅ℤ = compGroupIso theGroupIso (Hⁿ-Sⁿ≅ℤ 0) + where + theIso : Iso (coHom 1 KleinBottle) (coHom 1 S¹) + theIso = + setTruncIso ( + compIso (characFunSpace𝕂² (coHomK 1)) + (compIso + (Σ-cong-iso-snd (λ x → Σ-cong-iso-snd + λ p → Σ-cong-iso-snd + λ q → movePathIso p q (isCommΩK-based 1 x))) + (compIso Iso-H¹-𝕂²₁ + Iso-H¹-𝕂²₂))) + + is-hom : IsGroupHom (coHomGr 1 KleinBottle .snd) (fun theIso) (coHomGr 1 S¹ .snd) + is-hom = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g → cong ∣_∣₂ (funExt λ {base → refl ; (loop i) → refl})) + + theGroupIso : GroupIso (coHomGr 1 KleinBottle) (coHomGr 1 S¹) + theGroupIso = (theIso , is-hom) + +------ H²(𝕂²) ≅ ℤ/2ℤ (represented here by BoolGroup) ------- +-- It suffices to show that H²(Klein) is equivalent to Bool as types + +{- +Step one : +H²(𝕂²) := ∥ 𝕂² → K₂ ∥₂ + ≡ ∥ Σ[ x ∈ K₂ ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] (p ∙∙ q ∙∙ p ≡ q) ∥₂ (characFunSpace𝕂²) + ≡ ∥ Σ[ x ∈ K₂ ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ (movePathIso, using commutativity of ΩK₂) + ≡ ∥ Σ[ p ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ (connectedness of K₂) +-} + + +Iso-H²-𝕂²₁ : Iso ∥ Σ[ x ∈ coHomK 2 ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ + ∥ Σ[ p ∈ 0ₖ 2 ≡ 0ₖ 2 ] p ∙ p ≡ refl ∥₂ +fun Iso-H²-𝕂²₁ = + sRec isSetSetTrunc + (uncurry (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 2} 2 isSetSetTrunc) + (sphereElim _ (λ _ → isSetΠ λ _ → isSetSetTrunc) + λ y → ∣ fst y , snd (snd y) ∣₂))) +inv Iso-H²-𝕂²₁ = + sMap λ p → (0ₖ 2) , ((fst p) , (refl , (snd p))) +rightInv Iso-H²-𝕂²₁ = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ p → refl +leftInv Iso-H²-𝕂²₁ = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 1} 3 (isSetSetTrunc _ _)) + (sphereToPropElim _ + (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ {(p , (q , sq)) + → trRec (isSetSetTrunc _ _) + (λ qid → cong ∣_∣₂ (ΣPathP (refl , (ΣPathP (refl , (ΣPathP (sym qid , refl))))))) + (fun (PathIdTruncIso _) + (isContr→isProp (isConnectedPathKn 1 (0ₖ 2) (0ₖ 2)) ∣ q ∣ ∣ refl ∣))}))) + +{- Step two : ∥ Σ[ p ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ ≡ ∥ Σ[ x ∈ K₁ ] x + x ≡ 0 ∥₂ -} +Iso-H²-𝕂²₂ : Iso ∥ (Σ[ p ∈ 0ₖ 2 ≡ 0ₖ 2 ] p ∙ p ≡ refl) ∥₂ ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ +Iso-H²-𝕂²₂ = setTruncIso (Σ-cong-iso {B' = λ x → x +ₖ x ≡ 0ₖ 1} (invIso (Iso-Kn-ΩKn+1 1)) + λ p → compIso (congIso (invIso (Iso-Kn-ΩKn+1 1))) + (pathToIso λ i → ΩKn+1→Kn-hom 1 p p i ≡ 0ₖ 1)) + +{- Step three : +∥ Σ[ x ∈ K₁ ] x + x ≡ 0 ∥₂ ≡ Bool +We begin by defining the a map Σ[ x ∈ K₁ ] x + x ≡ 0 → Bool. For a point +(0 , p) we map it to true if winding(p) is even and false if winding(p) is odd. +We also have to show that this map respects the loop +-} + +ΣKₙNilpot→Bool : Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 → Bool +ΣKₙNilpot→Bool = uncurry (trElim (λ _ → isGroupoidΠ λ _ → isOfHLevelSuc 2 isSetBool) + λ {base p → isEven (ΩKn+1→Kn 0 p) + ; (loop i) p → hcomp (λ k → λ { (i = i0) → respectsLoop p k + ; (i = i1) → isEven (ΩKn+1→Kn 0 p)}) + (isEven (ΩKn+1→Kn 0 (transp (λ j → ∣ (loop ∙ loop) (i ∨ j) ∣ ≡ 0ₖ 1) i + p)))}) + where + isEven-2 : (x : ℤ) → isEven (-2 +ℤ x) ≡ isEven x + isEven-2 (pos zero) = refl + isEven-2 (pos (suc zero)) = refl + isEven-2 (pos (suc (suc n))) = + cong isEven (cong sucℤ (sucℤ+pos _ _) + ∙∙ sucℤ+pos _ _ + ∙∙ +-commℤ 0 (pos n)) + ∙ lossy n + where + lossy : (n : ℕ) → isEven (pos n) ≡ isEven (pos n) + lossy n = refl + isEven-2 (negsuc zero) = refl + isEven-2 (negsuc (suc n)) = + cong isEven (predℤ+negsuc n _ + ∙ +-commℤ -3 (negsuc n)) + ∙ lossy2 n + where + lossy2 : (n : ℕ) → isEven (negsuc (suc (suc (suc n)))) ≡ isEven (pos n) + lossy2 n = refl + respectsLoop : (p : 0ₖ 1 ≡ 0ₖ 1) + → isEven (ΩKn+1→Kn 0 (transport (λ i → ∣ (loop ∙ loop) i ∣ ≡ 0ₖ 1) p)) + ≡ isEven (ΩKn+1→Kn 0 p) + respectsLoop p = + cong isEven (cong (ΩKn+1→Kn 0) (cong (transport (λ i → ∣ (loop ∙ loop) i ∣ ≡ 0ₖ 1)) + (lUnit p))) + ∙∙ cong isEven (cong (ΩKn+1→Kn 0) + λ j → transp (λ i → ∣ (loop ∙ loop) (i ∨ j) ∣ ≡ 0ₖ 1) j + ((λ i → ∣ (loop ∙ loop) (~ i ∧ j) ∣) ∙ p)) + ∙∙ cong isEven (ΩKn+1→Kn-hom 0 (sym (cong ∣_∣ (loop ∙ loop))) p) + ∙ isEven-2 (ΩKn+1→Kn 0 p) + +{- +We show that for any x : ℤ we have ∣ (0ₖ 1 , Kn→ΩKn+1 0 x) ∣₂ ≡ ∣ (0ₖ 1 , refl) ∣₂ when x is even +and ∣ (0ₖ 1 , Kn→ΩKn+1 0 x) ∣₂ ≡ ∣ (0ₖ 1 , cong ∣_∣ loop) ∣₂ when x is odd + +This is done by induction on x. For the inductive step we define a multiplication _*_ on ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ +which is just ∣ (0 , p) ∣₂ * ∣ (0 , q) ∣₂ ≡ ∣ (0 , p ∙ q) ∣₂ when x is 0 +-} + +private + _*_ : ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ → ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ → ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ + _*_ = sRec (isSetΠ (λ _ → isSetSetTrunc)) λ a → sRec isSetSetTrunc λ b → *' (fst a) (fst b) (snd a) (snd b) + where + *' : (x y : coHomK 1) (p : x +ₖ x ≡ 0ₖ 1) (q : y +ₖ y ≡ 0ₖ 1) → ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ + *' = + trElim2 (λ _ _ → isGroupoidΠ2 λ _ _ → isOfHLevelSuc 2 isSetSetTrunc) + (wedgeconFun _ _ + (λ _ _ → isSetΠ2 λ _ _ → isSetSetTrunc) + (λ x p q → ∣ ∣ x ∣ , cong₂ _+ₖ_ p q ∣₂) + (λ y p q → ∣ ∣ y ∣ , sym (rUnitₖ 1 (∣ y ∣ +ₖ ∣ y ∣)) ∙ cong₂ _+ₖ_ p q ∣₂) + (funExt λ p → funExt λ q → cong ∣_∣₂ (ΣPathP (refl , (sym (lUnit _)))))) + + *=∙ : (p q : 0ₖ 1 ≡ 0ₖ 1) → ∣ 0ₖ 1 , p ∣₂ * ∣ 0ₖ 1 , q ∣₂ ≡ ∣ 0ₖ 1 , p ∙ q ∣₂ + *=∙ p q = cong ∣_∣₂ (ΣPathP (refl , sym (∙≡+₁ p q))) + +isEvenNegsuc : (n : ℕ) → isEven (pos (suc n)) ≡ true → isEven (negsuc n) ≡ true +isEvenNegsuc zero p = ⊥-rec (true≢false (sym p)) +isEvenNegsuc (suc n) p = p + +¬isEvenNegSuc : (n : ℕ) → isEven (pos (suc n)) ≡ false → isEven (negsuc n) ≡ false +¬isEvenNegSuc zero p = refl +¬isEvenNegSuc (suc n) p = p + +evenCharac : (x : ℤ) → isEven x ≡ true + → Path ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ + ∣ (0ₖ 1 , Kn→ΩKn+1 0 x) ∣₂ + ∣ (0ₖ 1 , refl) ∣₂ +evenCharac (pos zero) isisEven i = ∣ (0ₖ 1) , (rUnit refl (~ i)) ∣₂ +evenCharac (pos (suc zero)) isisEven = ⊥-rec (true≢false (sym isisEven)) +evenCharac (pos (suc (suc zero))) isisEven = + cong ∣_∣₂ ((λ i → 0ₖ 1 , rUnit (cong ∣_∣ ((lUnit loop (~ i)) ∙ loop)) (~ i)) + ∙ (ΣPathP (cong ∣_∣ loop , λ i j → ∣ (loop ∙ loop) (i ∨ j) ∣))) +evenCharac (pos (suc (suc (suc n)))) isisEven = + (λ i → ∣ 0ₖ 1 , Kn→ΩKn+1-hom 0 (pos (suc n)) 2 i ∣₂) + ∙∙ sym (*=∙ (Kn→ΩKn+1 0 (pos (suc n))) (Kn→ΩKn+1 0 (pos 2))) + ∙∙ (cong₂ _*_ (evenCharac (pos (suc n)) isisEven) (evenCharac 2 refl)) + +evenCharac (negsuc zero) isisEven = ⊥-rec (true≢false (sym isisEven)) +evenCharac (negsuc (suc zero)) isisEven = + cong ∣_∣₂ ((λ i → 0ₖ 1 + , λ i₁ → hfill (doubleComp-faces (λ i₂ → ∣ base ∣) (λ _ → ∣ base ∣) i₁) + (inS ∣ compPath≡compPath' (sym loop) (sym loop) i i₁ ∣) (~ i)) + ∙ ΣPathP ((cong ∣_∣ (sym loop)) , λ i j → ∣ (sym loop ∙' sym loop) (i ∨ j) ∣)) +evenCharac (negsuc (suc (suc n))) isisEven = + cong ∣_∣₂ (λ i → 0ₖ 1 , Kn→ΩKn+1-hom 0 (negsuc n) -2 i) + ∙∙ sym (*=∙ (Kn→ΩKn+1 0 (negsuc n)) (Kn→ΩKn+1 0 -2)) + ∙∙ cong₂ _*_ (evenCharac (negsuc n) (isEvenNegsuc n isisEven)) (evenCharac -2 refl) + +oddCharac : (x : ℤ) → isEven x ≡ false + → Path ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ + ∣ (0ₖ 1 , Kn→ΩKn+1 0 x) ∣₂ + ∣ (0ₖ 1 , cong ∣_∣ loop) ∣₂ +oddCharac (pos zero) isOdd = ⊥-rec (true≢false isOdd) +oddCharac (pos (suc zero)) isOdd i = + ∣ (0ₖ 1 , λ j → hfill (doubleComp-faces (λ i₂ → ∣ base ∣) (λ _ → ∣ base ∣) j) + (inS ∣ lUnit loop (~ i) j ∣) (~ i)) ∣₂ +oddCharac (pos (suc (suc n))) isOdd = + (λ i → ∣ 0ₖ 1 , Kn→ΩKn+1-hom 0 (pos n) 2 i ∣₂) + ∙∙ sym (*=∙ (Kn→ΩKn+1 0 (pos n)) (Kn→ΩKn+1 0 2)) + ∙∙ cong₂ _*_ (oddCharac (pos n) isOdd) (evenCharac 2 refl) +oddCharac (negsuc zero) isOdd = + cong ∣_∣₂ ((λ i → 0ₖ 1 , rUnit (sym (cong ∣_∣ loop)) (~ i)) + ∙ ΣPathP (cong ∣_∣ (sym loop) , λ i j → ∣ hcomp (λ k → λ { (i = i0) → loop (~ j ∧ k) + ; (i = i1) → loop j + ; (j = i1) → base}) + (loop (j ∨ ~ i)) ∣)) +oddCharac (negsuc (suc zero)) isOdd = ⊥-rec (true≢false isOdd) +oddCharac (negsuc (suc (suc n))) isOdd = + cong ∣_∣₂ (λ i → 0ₖ 1 , Kn→ΩKn+1-hom 0 (negsuc n) -2 i) + ∙∙ sym (*=∙ (Kn→ΩKn+1 0 (negsuc n)) (Kn→ΩKn+1 0 -2)) + ∙∙ cong₂ _*_ (oddCharac (negsuc n) (¬isEvenNegSuc n isOdd)) (evenCharac (negsuc 1) refl) + +{- We now have all we need to establish the Iso -} +Bool→ΣKₙNilpot : Bool → ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ +Bool→ΣKₙNilpot false = ∣ 0ₖ 1 , cong ∣_∣ loop ∣₂ +Bool→ΣKₙNilpot true = ∣ 0ₖ 1 , refl ∣₂ + +testIso : Iso ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ Bool +fun testIso = sRec isSetBool ΣKₙNilpot→Bool +inv testIso = Bool→ΣKₙNilpot +rightInv testIso false = refl +rightInv testIso true = refl +leftInv testIso = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry (trElim + (λ _ → isGroupoidΠ λ _ → isOfHLevelPlus {n = 1} 2 (isSetSetTrunc _ _)) + (toPropElim (λ _ → isPropΠ (λ _ → isSetSetTrunc _ _)) + (λ p → path p (isEven (ΩKn+1→Kn 0 p)) refl)))) + where + path : (p : 0ₖ 1 ≡ 0ₖ 1) (b : Bool) → (isEven (ΩKn+1→Kn 0 p) ≡ b) + → Bool→ΣKₙNilpot (ΣKₙNilpot→Bool (∣ base ∣ , p)) ≡ ∣ ∣ base ∣ , p ∣₂ + path p false q = + (cong Bool→ΣKₙNilpot q) + ∙∙ sym (oddCharac (ΩKn+1→Kn 0 p) q) + ∙∙ cong ∣_∣₂ λ i → 0ₖ 1 , rightInv (Iso-Kn-ΩKn+1 0) p i + path p true q = + cong Bool→ΣKₙNilpot q + ∙∙ sym (evenCharac (ΩKn+1→Kn 0 p) q) + ∙∙ cong ∣_∣₂ λ i → 0ₖ 1 , rightInv (Iso-Kn-ΩKn+1 0) p i + + +H²-𝕂²≅Bool : GroupIso (coHomGr 2 KleinBottle) BoolGroup +H²-𝕂²≅Bool = invGroupIso (≅Bool theIso) + where + theIso : Iso _ _ + theIso = + compIso (setTruncIso + (compIso (characFunSpace𝕂² (coHomK 2)) + (Σ-cong-iso-snd + λ x → Σ-cong-iso-snd + λ p → Σ-cong-iso-snd + λ q → (movePathIso p q (isCommΩK-based 2 x))))) + (compIso Iso-H²-𝕂²₁ + (compIso + Iso-H²-𝕂²₂ + testIso)) + +------ Hⁿ(𝕂²) ≅ 0 , n ≥ 3 ------ +isContrHⁿ-𝕂² : (n : ℕ) → isContr (coHom (3 + n) KleinBottle) +isContrHⁿ-𝕂² n = + isOfHLevelRetractFromIso 0 + (setTruncIso (characFunSpace𝕂² (coHomK _))) + isContrΣ-help + where + helper : (x : coHomK (3 + n))(p : x ≡ x) → (refl ≡ p) → (q : x ≡ x) → (refl ≡ q) + → (P : p ∙∙ q ∙∙ p ≡ q) + → Path ∥ (Σ[ x ∈ coHomK (3 + n) ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) ∥₂ + ∣ x , p , q , P ∣₂ + ∣ 0ₖ _ , refl , refl , sym (rUnit refl) ∣₂ + helper = + trElim (λ _ → isProp→isOfHLevelSuc (4 + n) (isPropΠ4 λ _ _ _ _ → isPropΠ λ _ → isSetSetTrunc _ _)) + (sphereToPropElim _ (λ _ → isPropΠ4 λ _ _ _ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ p → J (λ p _ → (q : 0ₖ _ ≡ 0ₖ _) → (refl ≡ q) + → (P : p ∙∙ q ∙∙ p ≡ q) + → Path ∥ (Σ[ x ∈ coHomK (3 + n) ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) ∥₂ + ∣ 0ₖ _ , p , q , P ∣₂ + ∣ 0ₖ _ , refl , refl , sym (rUnit refl) ∣₂) + λ q → J (λ q _ → (P : refl ∙∙ q ∙∙ refl ≡ q) + → Path ∥ (Σ[ x ∈ coHomK (3 + n) ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) ∥₂ + ∣ 0ₖ _ , refl , q , P ∣₂ + ∣ 0ₖ _ , refl , refl , sym (rUnit refl) ∣₂) + λ P → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ P≡rUnitrefl i → ∣ 0ₖ (3 + n) , refl , refl , P≡rUnitrefl i ∣₂) + (fun (PathIdTruncIso _) + (isContr→isProp (isConnectedPath _ (isConnectedPathKn (2 + n) _ _) + (refl ∙∙ refl ∙∙ refl) refl) + ∣ P ∣ ∣ sym (rUnit refl) ∣))) + + isContrΣ-help : isContr ∥ (Σ[ x ∈ coHomK (3 + n) ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) ∥₂ + fst isContrΣ-help = ∣ 0ₖ _ , refl , refl , sym (rUnit refl) ∣₂ + snd isContrΣ-help = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(x , p , q , P) + → trRec (isProp→isOfHLevelSuc (suc n) (isSetSetTrunc _ _)) + (λ pId → trRec (isProp→isOfHLevelSuc (suc n) (isSetSetTrunc _ _)) + (λ qId → sym (helper x p pId q qId P)) + (fun (PathIdTruncIso (2 + n)) + (isContr→isProp (isConnectedPathKn (2 + n) _ _) ∣ refl ∣ ∣ q ∣))) + (fun (PathIdTruncIso (2 + n)) + (isContr→isProp (isConnectedPathKn (2 + n) _ _) ∣ refl ∣ ∣ p ∣))} + +Hⁿ⁺³-𝕂²≅0 : (n : ℕ) → GroupIso (coHomGr (3 + n) KleinBottle) UnitGroup +Hⁿ⁺³-𝕂²≅0 n = contrGroupIsoUnit (isContrHⁿ-𝕂² n) diff --git a/Cubical/ZCohomology/Groups/Prelims.agda b/Cubical/ZCohomology/Groups/Prelims.agda new file mode 100644 index 0000000000..8db18ee7f2 --- /dev/null +++ b/Cubical/ZCohomology/Groups/Prelims.agda @@ -0,0 +1,280 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.Prelims where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 + +open import Cubical.Homotopy.Loopspace +open import Cubical.Data.Sigma +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) +open import Cubical.HITs.SetTruncation renaming (elim to sElim ; map to sMap ; rec to sRec) + +infixr 33 _⋄_ + +_⋄_ : _ +_⋄_ = compIso + +-- We strengthen the elimination rule for Hⁿ(S¹). We show that we only need to work with elements ∣ f ∣₂ (definitionally) sending loop to some loop p +-- and sending base to 0 +elimFunS¹ : (n : ℕ) → (p : typ (Ω (coHomK-ptd (suc n)))) → S¹ → coHomK (suc n) +elimFunS¹ n p base = ∣ ptSn (suc n) ∣ +elimFunS¹ n p (loop i) = p i + +coHomPointedElimS¹ : ∀ {ℓ} (n : ℕ) {B : coHom (suc n) S¹ → Type ℓ} + → ((x : coHom (suc n) S¹) → isProp (B x)) + → ((p : typ (Ω (coHomK-ptd (suc n)))) → B ∣ elimFunS¹ n p ∣₂) + → (x : coHom (suc n) S¹) → B x +coHomPointedElimS¹ n {B = B} x p = + coHomPointedElim n base x + λ f Id → subst B + (cong ∣_∣₂ (funExt (λ {base → sym Id ; (loop i) j → doubleCompPath-filler (sym Id) (cong f loop) Id (~ j) i}))) + (p (sym Id ∙∙ (cong f loop) ∙∙ Id)) + +coHomPointedElimS¹2 : ∀ {ℓ} (n : ℕ) {B : (x y : coHom (suc n) S¹) → Type ℓ} + → ((x y : coHom (suc n) S¹) → isProp (B x y)) + → ((p q : typ (Ω (coHomK-ptd (suc n)))) → B ∣ elimFunS¹ n p ∣₂ ∣ elimFunS¹ n q ∣₂) + → (x y : coHom (suc n) S¹) → B x y +coHomPointedElimS¹2 n {B = B} x p = + coHomPointedElim2 _ base x λ f g fId gId + → subst2 B (cong ∣_∣₂ (funExt (λ {base → sym fId ; (loop i) j → doubleCompPath-filler (sym (fId)) (cong f loop) fId (~ j) i}))) + (cong ∣_∣₂ (funExt (λ {base → sym gId ; (loop i) j → doubleCompPath-filler (sym (gId)) (cong g loop) gId (~ j) i}))) + (p (sym fId ∙∙ cong f loop ∙∙ fId) (sym gId ∙∙ cong g loop ∙∙ gId)) + +-- We do the same thing for Sⁿ, n ≥ 2. +elimFunSⁿ : (n m : ℕ) (p : S₊ (suc m) → typ (Ω (coHomK-ptd (suc n)))) + → (S₊ (2 + m)) → coHomK (suc n) +elimFunSⁿ n m p north = ∣ ptSn (suc n) ∣ +elimFunSⁿ n m p south = ∣ ptSn (suc n) ∣ +elimFunSⁿ n m p (merid a i) = p a i + +coHomPointedElimSⁿ : ∀ {ℓ} (n m : ℕ) {B : (x : coHom (suc n) (S₊ (2 + m))) → Type ℓ} + → ((x : coHom (suc n) (S₊ (2 + m))) → isProp (B x)) + → ((p : _) → B ∣ elimFunSⁿ n m p ∣₂) + → (x : coHom (suc n) (S₊ (2 + m))) → B x +coHomPointedElimSⁿ n m {B = B} isprop ind = + coHomPointedElim n north isprop + λ f fId → subst B (cong ∣_∣₂ (funExt (λ {north → sym fId + ; south → sym fId ∙' cong f (merid (ptSn (suc m))) + ; (merid a i) j → hcomp (λ k → λ {(i = i0) → fId (~ j ∧ k) + ; (i = i1) → compPath'-filler (sym fId) + (cong f (merid (ptSn (suc m)))) k j + ; (j = i1) → f (merid a i)}) + (hcomp (λ k → λ {(i = i0) → f north ; + (i = i1) → f (merid (ptSn (suc m)) (j ∨ ~ k)) ; + (j = i1) → f (merid a i)}) + (f (merid a i)))}))) + (ind λ a → sym fId ∙∙ cong f (merid a) ∙ cong f (sym (merid (ptSn (suc m)))) ∙∙ fId) + +0₀ = 0ₖ 0 +0₁ = 0ₖ 1 +0₂ = 0ₖ 2 +0₃ = 0ₖ 3 +0₄ = 0ₖ 4 + +S¹map : hLevelTrunc 3 S¹ → S¹ +S¹map = trRec isGroupoidS¹ (idfun _) + +S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x +S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → refl + + +-- We prove that (S¹ → ∥ S¹ ∥) ≃ S¹ × ℤ (Needed for H¹(S¹)). Note that the truncation doesn't really matter, since S¹ is a groupoid. +-- Given a map f : S¹ → S¹, the idea is to send this to (f(base) , winding (f(loop))). For this to be +-- well-typed, we need to translate f(loop) into an element in Ω(S¹,base). + +S¹→S¹≡S¹×ℤ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × ℤ) +Iso.fun S¹→S¹≡S¹×ℤ f = S¹map (f base) + , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) +Iso.inv S¹→S¹≡S¹×ℤ (s , int) base = ∣ s ∣ +Iso.inv S¹→S¹≡S¹×ℤ (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ +Iso.rightInv S¹→S¹≡S¹×ℤ (s , int) = ΣPathP (refl , ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) + ∙ windingℤLoop int)) +Iso.leftInv S¹→S¹≡S¹×ℤ f = funExt λ { base → S¹map-id (f base) + ; (loop i) j → helper j i} + where + helper : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) + (λ i → ∣ basechange2 (S¹map (f base)) + (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) + (cong f loop) + helper i j = + hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j + ; (i = i1) → S¹map-id (f (loop j)) k + ; (j = i0) → S¹map-id (f base) (i ∧ k) + ; (j = i1) → S¹map-id (f base) (i ∧ k)}) + (helper2 i j) + where + helper2 : Path (Path (hLevelTrunc 3 _) _ _) + (cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop + (winding + (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁)))))))) + λ i → ∣ S¹map (f (loop i)) ∣ + helper2 i j = + ∣ ((cong (basechange2 (S¹map (f base))) + (decodeEncode base (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁))))) + ∙ basechange2-sect (S¹map (f base)) + (λ i → S¹map (f (loop i)))) i) j ∣ + +{- Proof that (S¹ → K₁) ≃ K₁ × ℤ. Needed for H¹(T²) -} +S1→K₁≡S1×ℤ : Iso ((S₊ 1) → coHomK 1) (coHomK 1 × ℤ) +S1→K₁≡S1×ℤ = S¹→S¹≡S¹×ℤ ⋄ prodIso (invIso (truncIdempotentIso 3 (isGroupoidS¹))) idIso + +module _ (key : Unit') where + module P = lockedCohom key + private + _+K_ : {n : ℕ} → coHomK n → coHomK n → coHomK n + _+K_ {n = n} = P.+K n + + -K_ : {n : ℕ} → coHomK n → coHomK n + -K_ {n = n} = P.-K n + + _-K_ : {n : ℕ} → coHomK n → coHomK n → coHomK n + _-K_ {n = n} = P.-Kbin n + + infixr 55 _+K_ + infixr 55 -K_ + infixr 56 _-K_ + + {- Proof that S¹→K2 is isomorphic to K2×K1 (as types). Needed for H²(T²) -} + S1→K2≡K2×K1' : Iso (S₊ 1 → coHomK 2) (coHomK 2 × coHomK 1) + Iso.fun S1→K2≡K2×K1' f = f base , ΩKn+1→Kn 1 (sym (P.rCancelK 2 (f base)) + ∙∙ cong (λ x → (f x) -K f base) loop + ∙∙ P.rCancelK 2 (f base)) + Iso.inv S1→K2≡K2×K1' = invmap + where + invmap : (∥ Susp S¹ ∥ 4) × (∥ S¹ ∥ 3) → S¹ → ∥ Susp S¹ ∥ 4 + invmap (a , b) base = a +K 0₂ + invmap (a , b) (loop i) = a +K Kn→ΩKn+1 1 b i + Iso.rightInv S1→K2≡K2×K1' (a , b) = ΣPathP ((P.rUnitK 2 a) + , (cong (ΩKn+1→Kn 1) (doubleCompPath-elim' (sym (P.rCancelK 2 (a +K 0₂))) + (λ i → (a +K Kn→ΩKn+1 1 b i) -K (a +K 0₂)) + (P.rCancelK 2 (a +K 0₂))) + ∙∙ cong (ΩKn+1→Kn 1) (congHelper2 (Kn→ΩKn+1 1 b) (λ x → (a +K x) -K (a +K 0₂)) + (funExt (λ x → sym (cancelHelper a x))) + (P.rCancelK 2 (a +K 0₂))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 1) b)) + + module _ where + cancelHelper : (a b : coHomK 2) → (a +K b) -K (a +K 0₂) ≡ b + cancelHelper a b = cong (λ x → (a +K b) -K x) (P.rUnitK 2 a) + ∙ P.-cancelLK 2 a b + + congHelper2 : (p : 0₂ ≡ 0₂) (f : coHomK 2 → coHomK 2) (id : (λ x → x) ≡ f) → (q : (f 0₂) ≡ 0₂) + → (sym q) ∙ cong f p ∙ q ≡ p + congHelper2 p f = J (λ f _ → (q : (f 0₂) ≡ 0₂) → (sym q) ∙ cong f p ∙ q ≡ p) + λ q → (cong (sym q ∙_) (isCommΩK 2 p _) ∙∙ assoc _ _ _ ∙∙ cong (_∙ p) (lCancel q)) + ∙ sym (lUnit p) + + conghelper3 : (x : coHomK 2) (p : x ≡ x) (f : coHomK 2 → coHomK 2) (id : (λ x → x) ≡ f) (q : f x ≡ x) + → (sym q) ∙ cong f p ∙ q ≡ p + conghelper3 x p f = J (λ f _ → (q : (f x) ≡ x) → (sym q) ∙ cong f p ∙ q ≡ p) + λ q → (cong (sym q ∙_) (isCommΩK-based 2 x p _) ∙∙ assoc _ _ _ ∙∙ cong (_∙ p) (lCancel q)) + ∙ sym (lUnit p) + Iso.leftInv S1→K2≡K2×K1' a i base = P.rUnitK _ (a base) i + Iso.leftInv S1→K2≡K2×K1' a i (loop j) = loop-helper i j + where + loop-helper : PathP (λ i → P.rUnitK _ (a base) i ≡ P.rUnitK _ (a base) i) + (cong (a base +K_) (Kn→ΩKn+1 1 (ΩKn+1→Kn 1 ((sym (P.rCancelK 2 (a base)) + ∙∙ (λ i → a (loop i) -K (a (base))) + ∙∙ P.rCancelK 2 (a base)))))) + (cong a loop) + loop-helper i j = + hcomp (λ k → λ { (i = i0) → (G (doubleCompPath-elim' + (sym rP) (λ i₁ → a (loop i₁) -K a base) rP (~ k))) j + ; (i = i1) → cong a loop j + ; (j = i0) → P.rUnitK 2 (a base) i + ; (j = i1) → P.rUnitK 2 (a base) i}) + (loop-helper2 i j) + + where + F : typ (Ω (coHomK-ptd 2)) → a base +K snd (coHomK-ptd 2) ≡ a base +K snd (coHomK-ptd 2) + F = cong (_+K_ {n = 2} (a base)) + + G : 0ₖ 2 ≡ 0ₖ 2 → a base +K snd (coHomK-ptd 2) ≡ a base +K snd (coHomK-ptd 2) + G p = F (Kn→ΩKn+1 1 (ΩKn+1→Kn 1 p)) + + rP : P.+K 2 (a base) (P.-K 2 (a base)) ≡ 0ₖ 2 + rP = P.rCancelK 2 (a base) + + lem : (x : coHomK 2) (p : x ≡ x) (q : 0₂ ≡ x) + → Kn→ΩKn+1 1 (ΩKn+1→Kn 1 (q ∙ p ∙ sym q)) ≡ q ∙ p ∙ sym q + lem x p q = Iso.rightInv (Iso-Kn-ΩKn+1 1) (q ∙ p ∙ sym q) + + subtr-lem : (a b : hLevelTrunc 4 (S₊ 2)) → a +K (b -K a) ≡ b + subtr-lem a b = P.commK 2 a (b -K a) ∙ P.-+cancelK 2 b a + + subtr-lem-coher : PathP (λ i → subtr-lem (a base) (a base) i ≡ subtr-lem (a base) (a base) i) + (cong (a base +K_) (λ i₁ → a (loop i₁) -K a base)) + λ i → a (loop i) + subtr-lem-coher i j = subtr-lem (a base) (a (loop j)) i + + abstract + helperFun2 : {A : Type₀} {0A a b : A} (main : 0A ≡ 0A) (start : b ≡ b) (p : a ≡ a) (q : a ≡ b) (r : b ≡ 0A) (Q : a ≡ 0A) + (R : PathP (λ i → Q i ≡ Q i) p main) + → start ≡ sym q ∙ p ∙ q + → isComm∙ (A , 0A) + → sym r ∙ start ∙ r ≡ main + helperFun2 main start p q r Q R startId comm = + sym r ∙ start ∙ r ≡[ i ]⟨ sym r ∙ startId i ∙ r ⟩ + sym r ∙ (sym q ∙ p ∙ q) ∙ r ≡[ i ]⟨ sym r ∙ assoc (sym q) (p ∙ q) r (~ i) ⟩ + sym r ∙ sym q ∙ (p ∙ q) ∙ r ≡[ i ]⟨ sym r ∙ sym q ∙ assoc p q r (~ i) ⟩ + sym r ∙ sym q ∙ p ∙ q ∙ r ≡[ i ]⟨ assoc (sym r) (rUnit (sym q) i) (p ∙ lUnit (q ∙ r) i) i ⟩ + (sym r ∙ sym q ∙ refl) ∙ p ∙ refl ∙ q ∙ r ≡[ i ]⟨ (sym r ∙ sym q ∙ λ j → Q (i ∧ j)) ∙ R i ∙ (λ j → Q ( i ∧ (~ j))) ∙ q ∙ r ⟩ + (sym r ∙ sym q ∙ Q) ∙ main ∙ sym Q ∙ q ∙ r ≡[ i ]⟨ (sym r ∙ sym q ∙ Q) ∙ main ∙ sym Q ∙ symDistr (sym r) (sym q) (~ i) ⟩ + (sym r ∙ sym q ∙ Q) ∙ main ∙ sym Q ∙ sym (sym r ∙ sym q) ≡[ i ]⟨ (assoc (sym r) (sym q) Q i) ∙ main ∙ symDistr (sym r ∙ sym q) Q (~ i) ⟩ + ((sym r ∙ sym q) ∙ Q) ∙ main ∙ sym ((sym r ∙ sym q) ∙ Q) ≡[ i ]⟨ ((sym r ∙ sym q) ∙ Q) ∙ comm main (sym ((sym r ∙ sym q) ∙ Q)) i ⟩ + ((sym r ∙ sym q) ∙ Q) ∙ sym ((sym r ∙ sym q) ∙ Q) ∙ main ≡⟨ assoc ((sym r ∙ sym q) ∙ Q) (sym ((sym r ∙ sym q) ∙ Q)) main ⟩ + (((sym r ∙ sym q) ∙ Q) ∙ sym ((sym r ∙ sym q) ∙ Q)) ∙ main ≡[ i ]⟨ rCancel (((sym r ∙ sym q) ∙ Q)) i ∙ main ⟩ + refl ∙ main ≡⟨ sym (lUnit main) ⟩ + main ∎ + + congFunct₃ : ∀ {A B : Type₀} {a b c d : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) + → cong f (p ∙ q ∙ r) ≡ cong f p ∙ cong f q ∙ cong f r + congFunct₃ f p q r = congFunct f p (q ∙ r) ∙ cong (cong f p ∙_) (congFunct f q r) + + lem₀ : G (sym rP ∙ (λ i₁ → a (loop i₁) -K a base) ∙ rP) + ≡ F (sym rP ∙ (λ i₁ → a (loop i₁) -K a base) ∙ rP) + lem₀ = cong F (lem (a base -K (a base)) + (λ i₁ → a (loop i₁) -K a base) + (sym (P.rCancelK 2 (a base)))) + + lem₁ : G (sym rP ∙ (λ i₁ → a (loop i₁) -K a base) ∙ rP) + ≡ sym (λ i₁ → a base +K P.rCancelK 2 (a base) i₁) ∙ + cong (a base +K_) (λ i₁ → a (loop i₁) -K a base) ∙ + (λ i₁ → a base +K P.rCancelK 2 (a base) i₁) + lem₁ = lem₀ ∙ congFunct₃ (a base +K_) (sym rP) + (λ i₁ → a (loop i₁) -K a base) + rP + + loop-helper2 : PathP (λ i → P.rUnitK _ (a base) i ≡ P.rUnitK _ (a base) i) + (cong (a base +K_) (Kn→ΩKn+1 1 (ΩKn+1→Kn 1 ((sym (P.rCancelK 2 (a base)) + ∙ (λ i → a (loop i) -K (a (base))) + ∙ P.rCancelK 2 (a base)))))) + (cong a loop) + loop-helper2 = compPathL→PathP (helperFun2 (cong a loop) _ _ + (cong (a base +K_) (P.rCancelK 2 (a base))) _ _ + subtr-lem-coher + lem₁ + (isCommΩK-based 2 (a base))) + +S1→K2≡K2×K1 : Iso (S₊ 1 → coHomK 2) (coHomK 2 × coHomK 1) +S1→K2≡K2×K1 = S1→K2≡K2×K1' unlock diff --git a/Cubical/ZCohomology/Groups/RP2.agda b/Cubical/ZCohomology/Groups/RP2.agda new file mode 100644 index 0000000000..0154df3147 --- /dev/null +++ b/Cubical/ZCohomology/Groups/RP2.agda @@ -0,0 +1,128 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.RP2 where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.Groups.KleinBottle +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to pRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim) hiding (map) +open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; elim2 to trElim2) +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) + +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Transport + +open import Cubical.ZCohomology.Groups.Connected + +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.Equiv +open import Cubical.Homotopy.Connected +open import Cubical.HITs.RPn.Base + +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Bool +open import Cubical.Data.Int + +open import Cubical.Foundations.Path + +private + variable + ℓ : Level + A : Type ℓ + +funSpaceIso-RP² : Iso (RP² → A) (Σ[ x ∈ A ] Σ[ p ∈ x ≡ x ] p ≡ sym p) +Iso.fun funSpaceIso-RP² f = f point , (cong f line , λ i j → f (square i j)) +Iso.inv funSpaceIso-RP² (x , p , P) point = x +Iso.inv funSpaceIso-RP² (x , p , P) (line i) = p i +Iso.inv funSpaceIso-RP² (x , p , P) (square i j) = P i j +Iso.rightInv funSpaceIso-RP² (x , p , P) i = x , p , P +Iso.leftInv funSpaceIso-RP² f _ point = f point +Iso.leftInv funSpaceIso-RP² f _ (line i) = f (line i) +Iso.leftInv funSpaceIso-RP² f _ (square i j) = f (square i j) + +private + pathIso : {x : A} {p : x ≡ x} → Iso (p ≡ sym p) (p ∙ p ≡ refl) + pathIso {p = p} = compIso (congIso (equivToIso (_ , compPathr-isEquiv p))) + (pathToIso (cong (p ∙ p ≡_) (lCancel p))) + +--- H⁰(RP²) ≅ ℤ ---- +H⁰-RP²≅ℤ : GroupIso (coHomGr 0 RP²) ℤGroup +H⁰-RP²≅ℤ = H⁰-connected point connectedRP¹ + where + connectedRP¹ : (x : RP²) → ∥ point ≡ x ∥ + connectedRP¹ point = ∣ refl ∣ + connectedRP¹ (line i) = + isOfHLevel→isOfHLevelDep 1 {B = λ x → ∥ point ≡ x ∥} + (λ _ → isPropPropTrunc) ∣ refl ∣ ∣ refl ∣ line i + connectedRP¹ (square i j) = helper i j + where + helper : SquareP (λ i j → ∥ point ≡ square i j ∥) + (isOfHLevel→isOfHLevelDep 1 {B = λ x → ∥ point ≡ x ∥} + (λ _ → isPropPropTrunc) ∣ refl ∣ ∣ refl ∣ line) + (symP (isOfHLevel→isOfHLevelDep 1 {B = λ x → ∥ point ≡ x ∥} + (λ _ → isPropPropTrunc) ∣ refl ∣ ∣ refl ∣ line)) + refl refl + helper = toPathP (isOfHLevelPathP 1 isPropPropTrunc _ _ _ _) + +--- H¹(RP²) ≅ 0 ---- +isContr-H¹-RP²-helper : isContr ∥ Σ[ x ∈ coHomK 1 ] Σ[ p ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ +fst isContr-H¹-RP²-helper = ∣ 0ₖ 1 , refl , sym (rUnit refl) ∣₂ +snd isContr-H¹-RP²-helper = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry + (trElim (λ _ → isGroupoidΠ λ _ → isOfHLevelPlus {n = 1} 2 (isSetSetTrunc _ _)) + (toPropElim (λ _ → isPropΠ (λ _ → isSetSetTrunc _ _)) + λ {(p , nilp) + → cong ∣_∣₂ (ΣPathP (refl , Σ≡Prop (λ _ → isOfHLevelTrunc 3 _ _ _ _) + (rUnit refl + ∙∙ cong (Kn→ΩKn+1 0) (sym (nilpotent→≡0 (ΩKn+1→Kn 0 p) + (sym (ΩKn+1→Kn-hom 0 p p) + ∙ cong (ΩKn+1→Kn 0) nilp))) + ∙∙ Iso.rightInv (Iso-Kn-ΩKn+1 0) p)))}))) + +H¹-RP²≅0 : GroupIso (coHomGr 1 RP²) UnitGroup +H¹-RP²≅0 = + contrGroupIsoUnit + (isOfHLevelRetractFromIso 0 + (setTruncIso (compIso funSpaceIso-RP² + (Σ-cong-iso-snd (λ _ → Σ-cong-iso-snd λ _ → pathIso)))) + isContr-H¹-RP²-helper) + +--- H²(RP²) ≅ ℤ/2ℤ ---- + +Iso-H²-RP²₁ : Iso ∥ Σ[ x ∈ coHomK 2 ] Σ[ p ∈ x ≡ x ] p ≡ sym p ∥₂ + ∥ Σ[ p ∈ 0ₖ 2 ≡ 0ₖ 2 ] p ≡ sym p ∥₂ +Iso.fun Iso-H²-RP²₁ = + sRec isSetSetTrunc + (uncurry + (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 2} 2 isSetSetTrunc) + (sphereElim _ (λ _ → isSetΠ (λ _ → isSetSetTrunc)) + λ p → ∣ fst p , snd p ∣₂))) +Iso.inv Iso-H²-RP²₁ = sMap λ p → (0ₖ 2) , p +Iso.rightInv Iso-H²-RP²₁ = sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ _ → refl +Iso.leftInv Iso-H²-RP²₁ = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 1} 3 (isSetSetTrunc _ _)) + (sphereToPropElim _ (λ _ → isPropΠ (λ _ → isSetSetTrunc _ _)) + λ p → refl))) + +Iso-H²-RP²₂ : Iso ∥ Σ[ p ∈ 0ₖ 2 ≡ 0ₖ 2 ] p ≡ sym p ∥₂ Bool +Iso-H²-RP²₂ = compIso (setTruncIso (Σ-cong-iso-snd λ _ → pathIso)) + (compIso Iso-H²-𝕂²₂ testIso) + + +H²-RP²≅Bool : GroupIso (coHomGr 2 RP²) BoolGroup +H²-RP²≅Bool = invGroupIso (≅Bool (compIso + (compIso (setTruncIso funSpaceIso-RP²) + Iso-H²-RP²₁) + Iso-H²-RP²₂)) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda new file mode 100644 index 0000000000..d2116749e6 --- /dev/null +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -0,0 +1,330 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Groups.Sn where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Prelims + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) + hiding (map) + +open import Cubical.Relation.Nullary +open import Cubical.Data.Sum hiding (map) +open import Cubical.Data.Empty renaming (rec to ⊥-rec) + +open import Cubical.Data.Bool +open import Cubical.Data.Sigma +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) + +open import Cubical.Homotopy.Connected + +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) + +infixr 31 _□_ +_□_ : _ +_□_ = compGroupIso + +open IsGroupHom +open BijectionIso +open Iso + +Sn-connected : (n : ℕ) (x : typ (S₊∙ (suc n))) → ∥ pt (S₊∙ (suc n)) ≡ x ∥₁ +Sn-connected zero = toPropElim (λ _ → isPropPropTrunc) ∣ refl ∣₁ +Sn-connected (suc zero) = suspToPropElim base (λ _ → isPropPropTrunc) ∣ refl ∣₁ +Sn-connected (suc (suc n)) = suspToPropElim north (λ _ → isPropPropTrunc) ∣ refl ∣₁ + +suspensionAx-Sn : (n m : ℕ) → GroupIso (coHomGr (2 + n) (S₊ (2 + m))) + (coHomGr (suc n) (S₊ (suc m))) +suspensionAx-Sn n m = + compIso (setTruncIso (invIso funSpaceSuspIso)) helperIso , + makeIsGroupHom funIsHom + where + helperIso : Iso ∥ (Σ[ x ∈ coHomK (2 + n) ] + Σ[ y ∈ coHomK (2 + n) ] + (S₊ (suc m) → x ≡ y)) ∥₂ + (coHom (suc n) (S₊ (suc m))) + Iso.fun helperIso = + sRec isSetSetTrunc + (uncurry + (coHomK-elim _ + (λ _ → isOfHLevelΠ (2 + n) + λ _ → isOfHLevelPlus' {n = n} 2 isSetSetTrunc) + (uncurry + (coHomK-elim _ + (λ _ → isOfHLevelΠ (2 + n) + λ _ → isOfHLevelPlus' {n = n} 2 isSetSetTrunc) + λ f → ∣ (λ x → ΩKn+1→Kn (suc n) (f x)) ∣₂)))) + Iso.inv helperIso = + sMap λ f → (0ₖ _) , (0ₖ _ , λ x → Kn→ΩKn+1 (suc n) (f x)) + Iso.rightInv helperIso = + coHomPointedElim _ (ptSn (suc m)) (λ _ → isSetSetTrunc _ _) + λ f fId → cong ∣_∣₂ (funExt (λ x → Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x))) + Iso.leftInv helperIso = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry + (coHomK-elim _ + (λ _ → isProp→isOfHLevelSuc (suc n) (isPropΠ λ _ → isSetSetTrunc _ _)) + (uncurry + (coHomK-elim _ + (λ _ → isProp→isOfHLevelSuc (suc n) (isPropΠ λ _ → isSetSetTrunc _ _)) + λ f → cong ∣_∣₂ + (ΣPathP (refl , + ΣPathP (refl , + (λ i x → Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) (f x) i)))))))) + + theFun : coHom (2 + n) (S₊ (2 + m)) → coHom (suc n) (S₊ (suc m)) + theFun = Iso.fun (compIso (setTruncIso (invIso funSpaceSuspIso)) + helperIso) + + funIsHom : (x y : coHom (2 + n) (S₊ (2 + m))) + → theFun (x +ₕ y) ≡ theFun x +ₕ theFun y + funIsHom = + coHomPointedElimSⁿ _ _ (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ f → coHomPointedElimSⁿ _ _ (λ _ → isSetSetTrunc _ _) + λ g → cong ∣_∣₂ (funExt λ x → cong (ΩKn+1→Kn (suc n)) (sym (∙≡+₂ n (f x) (g x))) + ∙ ΩKn+1→Kn-hom (suc n) (f x) (g x)) + + +H⁰-Sⁿ≅ℤ : (n : ℕ) → GroupIso (coHomGr 0 (S₊ (suc n))) ℤGroup +H⁰-Sⁿ≅ℤ zero = H⁰-connected base (Sn-connected 0) +H⁰-Sⁿ≅ℤ (suc n) = H⁰-connected north (Sn-connected (suc n)) + +-- -- ---------------------------------------------------------------------- + +--- We will need to switch between Sⁿ defined using suspensions and using pushouts +--- in order to apply Mayer Vietoris. + + +S1Iso : Iso S¹ (Pushout {A = Bool} (λ _ → tt) λ _ → tt) +S1Iso = S¹IsoSuspBool ⋄ invIso PushoutSuspIsoSusp + +coHomPushout≅coHomSn : (n m : ℕ) → GroupIso (coHomGr m (S₊ (suc n))) + (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) +coHomPushout≅coHomSn zero m = + setTruncIso (domIso S1Iso) , + makeIsGroupHom (sElim2 (λ _ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ _ _ → refl)) +coHomPushout≅coHomSn (suc n) m = + setTruncIso (domIso (invIso PushoutSuspIsoSusp)) , + makeIsGroupHom (sElim2 (λ _ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ _ _ → refl)) + +-------------------------- H⁰(S⁰) ----------------------------- +S0→ℤ : (a : ℤ × ℤ) → S₊ 0 → ℤ +S0→ℤ a true = fst a +S0→ℤ a false = snd a + +H⁰-S⁰≅ℤ×ℤ : GroupIso (coHomGr 0 (S₊ 0)) (DirProd ℤGroup ℤGroup) +fun (fst H⁰-S⁰≅ℤ×ℤ) = sRec (isSet× isSetℤ isSetℤ) λ f → (f true) , (f false) +inv (fst H⁰-S⁰≅ℤ×ℤ) a = ∣ S0→ℤ a ∣₂ +rightInv (fst H⁰-S⁰≅ℤ×ℤ) _ = refl +leftInv (fst H⁰-S⁰≅ℤ×ℤ) = + sElim (λ _ → isSet→isGroupoid isSetSetTrunc _ _) + (λ f → cong ∣_∣₂ (funExt (λ {true → refl ; false → refl}))) +snd H⁰-S⁰≅ℤ×ℤ = + makeIsGroupHom + (sElim2 (λ _ _ → isSet→isGroupoid (isSet× isSetℤ isSetℤ) _ _) λ a b → refl) + + +------------------------- H¹(S⁰) ≅ 0 ------------------------------- + + +private + Hⁿ-S0≃Kₙ×Kₙ : (n : ℕ) → Iso (S₊ 0 → coHomK (suc n)) (coHomK (suc n) × coHomK (suc n)) + Iso.fun (Hⁿ-S0≃Kₙ×Kₙ n) f = (f true) , (f false) + Iso.inv (Hⁿ-S0≃Kₙ×Kₙ n) (a , b) true = a + Iso.inv (Hⁿ-S0≃Kₙ×Kₙ n) (a , b) false = b + Iso.rightInv (Hⁿ-S0≃Kₙ×Kₙ n) a = refl + Iso.leftInv (Hⁿ-S0≃Kₙ×Kₙ n) b = funExt λ {true → refl ; false → refl} + + isContrHⁿ-S0 : (n : ℕ) → isContr (coHom (suc n) (S₊ 0)) + isContrHⁿ-S0 n = isContrRetract (Iso.fun (setTruncIso (Hⁿ-S0≃Kₙ×Kₙ n))) + (Iso.inv (setTruncIso (Hⁿ-S0≃Kₙ×Kₙ n))) + (Iso.leftInv (setTruncIso (Hⁿ-S0≃Kₙ×Kₙ n))) + (isContrHelper n) + where + isContrHelper : (n : ℕ) → isContr (∥ (coHomK (suc n) × coHomK (suc n)) ∥₂) + fst (isContrHelper zero) = ∣ (0₁ , 0₁) ∣₂ + snd (isContrHelper zero) = sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ y → elim2 {B = λ x y → ∣ (0₁ , 0₁) ∣₂ ≡ ∣(x , y) ∣₂ } + (λ _ _ → isOfHLevelPlus {n = 2} 2 isSetSetTrunc _ _) + (toPropElim2 (λ _ _ → isSetSetTrunc _ _) refl) (fst y) (snd y) + isContrHelper (suc zero) = ∣ (0₂ , 0₂) ∣₂ + , sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ y → elim2 {B = λ x y → ∣ (0₂ , 0₂) ∣₂ ≡ ∣(x , y) ∣₂ } + (λ _ _ → isOfHLevelPlus {n = 2} 3 isSetSetTrunc _ _) + (suspToPropElim2 base (λ _ _ → isSetSetTrunc _ _) refl) (fst y) (snd y) + isContrHelper (suc (suc n)) = ∣ (0ₖ (3 + n) , 0ₖ (3 + n)) ∣₂ + , sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ y → elim2 {B = λ x y → ∣ (0ₖ (3 + n) , 0ₖ (3 + n)) ∣₂ ≡ ∣(x , y) ∣₂ } + (λ _ _ → isProp→isOfHLevelSuc (4 + n) (isSetSetTrunc _ _)) + (suspToPropElim2 north (λ _ _ → isSetSetTrunc _ _) refl) (fst y) (snd y) + +H¹-S⁰≅0 : (n : ℕ) → GroupIso (coHomGr (suc n) (S₊ 0)) UnitGroup +H¹-S⁰≅0 n = contrGroupIsoUnit (isContrHⁿ-S0 n) + +------------------------- H²(S¹) ≅ 0 ------------------------------- + +Hⁿ-S¹≅0 : (n : ℕ) → GroupIso (coHomGr (2 + n) (S₊ 1)) UnitGroup +Hⁿ-S¹≅0 n = contrGroupIsoUnit + (isOfHLevelRetractFromIso 0 helper + (_ , helper2)) + where + helper : Iso ⟨ coHomGr (2 + n) (S₊ 1)⟩ ∥ Σ (hLevelTrunc (4 + n) (S₊ (2 + n))) (λ x → ∥ x ≡ x ∥₂) ∥₂ + helper = compIso (setTruncIso IsoFunSpaceS¹) IsoSetTruncateSndΣ + + helper2 : (x : ∥ Σ (hLevelTrunc (4 + n) (S₊ (2 + n))) (λ x → ∥ x ≡ x ∥₂) ∥₂) → ∣ ∣ north ∣ , ∣ refl ∣₂ ∣₂ ≡ x + helper2 = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry + (trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isProp→isOfHLevelSuc (3 + n) (isSetSetTrunc _ _)) + (suspToPropElim (ptSn (suc n)) (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + (sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ p + → cong ∣_∣₂ (ΣPathP (refl , isContr→isProp helper3 _ _)))))) + where + helper4 : isConnected (n + 3) (hLevelTrunc (4 + n) (S₊ (2 + n))) + helper4 = subst (λ m → isConnected m (hLevelTrunc (4 + n) (S₊ (2 + n)))) (+-comm 3 n) + (isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (3 + n) 1)) (sphereConnected (2 + n))) + + helper3 : isContr ∥ ∣ north ∣ ≡ ∣ north ∣ ∥₂ + helper3 = isOfHLevelRetractFromIso 0 setTruncTrunc2Iso + (isConnectedPath 2 (isConnectedSubtr 3 n helper4) _ _) + +-- --------------- H¹(Sⁿ), n ≥ 1 -------------------------------------------- + +H¹-Sⁿ≅0 : (n : ℕ) → GroupIso (coHomGr 1 (S₊ (2 + n))) UnitGroup +H¹-Sⁿ≅0 zero = contrGroupIsoUnit isContrH¹S² + where + isContrH¹S² : isContr ⟨ coHomGr 1 (S₊ 2) ⟩ + isContrH¹S² = ∣ (λ _ → ∣ base ∣) ∣₂ + , coHomPointedElim 0 north (λ _ → isSetSetTrunc _ _) + λ f p → cong ∣_∣₂ (funExt λ x → sym p ∙∙ sym (spoke f north) ∙∙ spoke f x) +H¹-Sⁿ≅0 (suc n) = contrGroupIsoUnit isContrH¹S³⁺ⁿ + where + anIso : Iso ⟨ coHomGr 1 (S₊ (3 + n)) ⟩ ∥ (S₊ (3 + n) → hLevelTrunc (4 + n) (coHomK 1)) ∥₂ + anIso = + setTruncIso + (codomainIso + (invIso (truncIdempotentIso (4 + n) (isOfHLevelPlus' {n = 1 + n} 3 (isOfHLevelTrunc 3))))) + + isContrH¹S³⁺ⁿ-ish : (f : (S₊ (3 + n) → hLevelTrunc (4 + n) (coHomK 1))) + → ∣ (λ _ → ∣ ∣ base ∣ ∣) ∣₂ ≡ ∣ f ∣₂ + isContrH¹S³⁺ⁿ-ish f = ind (f north) refl + where + ind : (x : hLevelTrunc (4 + n) (coHomK 1)) + → x ≡ f north + → ∣ (λ _ → ∣ ∣ base ∣ ∣) ∣₂ ≡ ∣ f ∣₂ + ind = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPlus' {n = (3 + n)} 1 (isSetSetTrunc _ _)) + (trElim (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 (isSetSetTrunc _ _)) + (toPropElim (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ p → cong ∣_∣₂ (funExt λ x → p ∙∙ sym (spoke f north) ∙∙ spoke f x))) + isContrH¹S³⁺ⁿ : isContr ⟨ coHomGr 1 (S₊ (3 + n)) ⟩ + isContrH¹S³⁺ⁿ = + isOfHLevelRetractFromIso 0 + anIso + (∣ (λ _ → ∣ ∣ base ∣ ∣) ∣₂ + , sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) isContrH¹S³⁺ⁿ-ish) + +--------- H¹(S¹) ≅ ℤ ------- +{- +Idea : +H¹(S¹) := ∥ S¹ → K₁ ∥₂ + ≃ ∥ S¹ → S¹ ∥₂ + ≃ ∥ S¹ × ℤ ∥₂ + ≃ ∥ S¹ ∥₂ × ∥ ℤ ∥₂ + ≃ ℤ +-} +coHom1S1≃ℤ : GroupIso (coHomGr 1 (S₊ 1)) ℤGroup +coHom1S1≃ℤ = theIso + where + F = Iso.fun S¹→S¹≡S¹×ℤ + F⁻ = Iso.inv S¹→S¹≡S¹×ℤ + + theIso : GroupIso (coHomGr 1 (S₊ 1)) ℤGroup + fun (fst theIso) = sRec isSetℤ (λ f → snd (F f)) + inv (fst theIso) a = ∣ (F⁻ (base , a)) ∣₂ + rightInv (fst theIso) a = cong snd (Iso.rightInv S¹→S¹≡S¹×ℤ (base , a)) + leftInv (fst theIso) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ((sRec isSetSetTrunc ∣_∣₂) + ∘ sRec isSetSetTrunc λ x → ∣ F⁻ (x , (snd (F f))) ∣₂) + (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (fst (F f)))) + ∙ cong ∣_∣₂ (Iso.leftInv S¹→S¹≡S¹×ℤ f) + snd theIso = + makeIsGroupHom + (coHomPointedElimS¹2 _ (λ _ _ → isSetℤ _ _) + λ p q → (λ i → winding (guy ∣ base ∣ (cong S¹map (help p q i)))) + ∙∙ (λ i → winding (guy ∣ base ∣ (congFunct S¹map p q i))) + ∙∙ winding-hom (guy ∣ base ∣ (cong S¹map p)) + (guy ∣ base ∣ (cong S¹map q))) + where + guy = basechange2⁻ ∘ S¹map + help : (p q : Path (coHomK 1) ∣ base ∣ ∣ base ∣) → cong₂ _+ₖ_ p q ≡ p ∙ q + help p q = cong₂Funct _+ₖ_ p q ∙ (λ i → cong (λ x → rUnitₖ 1 x i) p ∙ cong (λ x → lUnitₖ 1 x i) q) + +---------------------------- Hⁿ(Sⁿ) ≅ ℤ , n ≥ 1 ------------------- +Hⁿ-Sⁿ≅ℤ : (n : ℕ) → GroupIso (coHomGr (suc n) (S₊ (suc n))) ℤGroup +Hⁿ-Sⁿ≅ℤ zero = coHom1S1≃ℤ +Hⁿ-Sⁿ≅ℤ (suc n) = suspensionAx-Sn n n □ Hⁿ-Sⁿ≅ℤ n + +-------------- Hⁿ(Sᵐ) ≅ ℤ for n , m ≥ 1 with n ≠ m ---------------- +Hⁿ-Sᵐ≅0 : (n m : ℕ) → ¬ (n ≡ m) → GroupIso (coHomGr (suc n) (S₊ (suc m))) UnitGroup +Hⁿ-Sᵐ≅0 zero zero pf = ⊥-rec (pf refl) +Hⁿ-Sᵐ≅0 zero (suc m) pf = H¹-Sⁿ≅0 m +Hⁿ-Sᵐ≅0 (suc n) zero pf = Hⁿ-S¹≅0 n +Hⁿ-Sᵐ≅0 (suc n) (suc m) pf = suspensionAx-Sn n m + □ Hⁿ-Sᵐ≅0 n m λ p → pf (cong suc p) + + +-- Test functions +private + to₁ : coHom 1 S¹ → ℤ + to₁ = Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 0)) + + to₂ : coHom 2 (S₊ 2) → ℤ + to₂ = Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 1)) + + to₃ : coHom 3 (S₊ 3) → ℤ + to₃ = Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 2)) + + + from₁ : ℤ → coHom 1 S¹ + from₁ = Iso.inv (fst (Hⁿ-Sⁿ≅ℤ 0)) + + from₂ : ℤ → coHom 2 (S₊ 2) + from₂ = Iso.inv (fst (Hⁿ-Sⁿ≅ℤ 1)) + + from₃ : ℤ → coHom 3 (S₊ 3) + from₃ = Iso.inv (fst (Hⁿ-Sⁿ≅ℤ 2)) + +{- +Strangely, the following won't compute +test₀ : to₂ (from₂ 1 +ₕ from₂ 1) ≡ 2 +test₀ = refl +However, the same example works for S¹ ∨ S¹ ∨ S², where the functions are essentially the same as here +(although somewhat more complicated). + +But with our new strange addition +'ₕ, it computes just fine: + +test₀ : to₂ (from₂ 1 +'ₕ from₂ 1) ≡ 2 +test₀ = refl + +-} diff --git a/Cubical/ZCohomology/Groups/Torus.agda b/Cubical/ZCohomology/Groups/Torus.agda new file mode 100644 index 0000000000..cf9a73bca3 --- /dev/null +++ b/Cubical/ZCohomology/Groups/Torus.agda @@ -0,0 +1,309 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.Torus where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Prelims +open import Cubical.ZCohomology.RingStructure.CupProduct + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.Data.Unit +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) + +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) hiding (map) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim2 to pElim2 ; ∣_∣ to ∣_∣₁) hiding (map) +open import Cubical.HITs.Nullification +open import Cubical.HITs.Truncation renaming (elim to trElim ; elim2 to trElim2 ; map to trMap ; rec to trRec) +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Loopspace + +open import Cubical.ZCohomology.Groups.WedgeOfSpheres + renaming (to₂ to to₂-∨ ; from₂ to from₂-∨ ; from₁ to from₁-∨ ; to₁ to to₁-∨) hiding (to₀ ; from₀) +open import Cubical.Data.Empty +open import Cubical.HITs.Wedge + +open import Cubical.Relation.Nullary + +open IsGroupHom +open Iso + +-- The following section contains stengthened induction principles for cohomology groups of T². They are particularly useful for showing that +-- that some Isos are morphisms. They make things type-check faster, but should probably not be used for computations. + +-- We first need some functions +elimFunT² : (n : ℕ) (p q : typ (Ω (coHomK-ptd (suc n)))) + → Square q q p p + → S¹ × S¹ → coHomK (suc n) +elimFunT² n p q P (base , base) = ∣ ptSn (suc n) ∣ +elimFunT² n p q P (base , loop i) = q i +elimFunT² n p q P (loop i , base) = p i +elimFunT² n p q P (loop i , loop j) = P i j + +elimFunT²' : (n : ℕ) → Square (refl {ℓ-zero} {coHomK (suc n)} {∣ ptSn (suc n) ∣}) refl refl refl + → S¹ × S¹ → coHomK (suc n) +elimFunT²' n P (x , base) = ∣ ptSn (suc n) ∣ +elimFunT²' n P (base , loop j) = ∣ ptSn (suc n) ∣ +elimFunT²' n P (loop i , loop j) = P i j + +elimFunT²'≡elimFunT² : (n : ℕ) → (P : _) → elimFunT²' n P ≡ elimFunT² n refl refl P +elimFunT²'≡elimFunT² n P i (base , base) = ∣ ptSn (suc n) ∣ +elimFunT²'≡elimFunT² n P i (base , loop k) = ∣ ptSn (suc n) ∣ +elimFunT²'≡elimFunT² n P i (loop j , base) = ∣ ptSn (suc n) ∣ +elimFunT²'≡elimFunT² n P i (loop j , loop k) = P j k + +{- +The first induction principle says that when proving a proposition for some x : Hⁿ(T²), n ≥ 1, it suffices to show that it holds for +(elimFunT² p q P) for any paths p q : ΩKₙ, and square P : Square q q p p. This is useful because elimFunT² p q P (base , base) recudes to 0 +-} + +coHomPointedElimT² : ∀ {ℓ} (n : ℕ) {B : coHom (suc n) (S¹ × S¹) → Type ℓ} + → ((x : coHom (suc n) (S¹ × S¹)) → isProp (B x)) + → ((p q : _) (P : _) → B ∣ elimFunT² n p q P ∣₂) + → (x : coHom (suc n) (S¹ × S¹)) → B x +coHomPointedElimT² n {B = B} isprop indp = + coHomPointedElim _ (base , base) isprop + λ f fId → subst B (cong ∣_∣₂ (funExt (λ { (base , base) → sym fId + ; (base , loop i) j → helper f fId i1 i (~ j) + ; (loop i , base) j → helper f fId i i1 (~ j) + ; (loop i , loop j) k → helper f fId i j (~ k)}))) + (indp (λ i → helper f fId i i1 i1) + (λ i → helper f fId i1 i i1) + λ i j → helper f fId i j i1) + where + helper : (f : S¹ × S¹ → coHomK (suc n)) → f (base , base) ≡ ∣ ptSn (suc n) ∣ + → I → I → I → coHomK (suc n) + helper f fId i j k = + hfill (λ k → λ {(i = i0) → doubleCompPath-filler (sym fId) (cong f (λ i → (base , loop i))) fId k j + ; (i = i1) → doubleCompPath-filler (sym fId) (cong f (λ i → (base , loop i))) fId k j + ; (j = i0) → doubleCompPath-filler (sym fId) (cong f (λ i → (loop i , base))) fId k i + ; (j = i1) → doubleCompPath-filler (sym fId) (cong f (λ i → (loop i , base))) fId k i}) + (inS (f ((loop i) , (loop j)))) + k + +private + lem : ∀ {ℓ} (n : ℕ) {B : coHom (2 + n) (S¹ × S¹) → Type ℓ} + → ((P : _) → B ∣ elimFunT² (suc n) refl refl P ∣₂) + → (p : _) → (refl ≡ p) + → (q : _) → (refl ≡ q) + → (P : _) + → B ∣ elimFunT² (suc n) p q P ∣₂ + lem n {B = B} elimP p = + J (λ p _ → (q : _) → (refl ≡ q) + → (P : _) + → B ∣ elimFunT² (suc n) p q P ∣₂) + λ q → + J (λ q _ → (P : _) → B ∣ elimFunT² (suc n) refl q P ∣₂) + elimP + +{- When working with Hⁿ(T²) , n ≥ 2, we are, in the case described above, allowed to assume that any f : Hⁿ(T²) is + elimFunT² n refl refl P -} +coHomPointedElimT²' : ∀ {ℓ} (n : ℕ) {B : coHom (2 + n) (S¹ × S¹) → Type ℓ} + → ((x : coHom (2 + n) (S¹ × S¹)) → isProp (B x)) + → ((P : _) → B ∣ elimFunT² (suc n) refl refl P ∣₂) + → (x : coHom (2 + n) (S¹ × S¹)) → B x +coHomPointedElimT²' n {B = B} prop ind = + coHomPointedElimT² (suc n) prop + λ p q P → trRec (isProp→isOfHLevelSuc n (prop _)) + (λ p-refl → trRec (isProp→isOfHLevelSuc n (prop _)) + (λ q-refl → lem n {B = B} ind p (sym p-refl) q (sym q-refl) P) + (isConnectedPath _ (isConnectedPathKn (suc n) _ _) q refl .fst)) + (isConnectedPath _ (isConnectedPathKn (suc n) _ _) p refl .fst) + +{- A slight variation of the above which gives definitional equalities for all points (x , base) -} +private + coHomPointedElimT²'' : ∀ {ℓ} (n : ℕ) {B : coHom (2 + n) (S¹ × S¹) → Type ℓ} + → ((x : coHom (2 + n) (S¹ × S¹)) → isProp (B x)) + → ((P : _) → B ∣ elimFunT²' (suc n) P ∣₂) + → (x : coHom (2 + n) (S¹ × S¹)) → B x + coHomPointedElimT²'' n {B = B} prop ind = + coHomPointedElimT²' n prop λ P → subst (λ x → B ∣ x ∣₂) + (elimFunT²'≡elimFunT² (suc n) P) (ind P) + +--------- H⁰(T²) ------------ +H⁰-T²≅ℤ : GroupIso (coHomGr 0 (S₊ 1 × S₊ 1)) ℤGroup +H⁰-T²≅ℤ = + H⁰-connected (base , base) + λ (a , b) → pRec isPropPropTrunc + (λ id1 → pRec isPropPropTrunc + (λ id2 → ∣ ΣPathP (id1 , id2) ∣₁) + (Sn-connected 0 b) ) + (Sn-connected 0 a) + +--------- H¹(T²) ------------------------------- + +H¹-T²≅ℤ×ℤ : GroupIso (coHomGr 1 ((S₊ 1) × (S₊ 1))) (DirProd ℤGroup ℤGroup) +H¹-T²≅ℤ×ℤ = theIso □ GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ 0) (H⁰-Sⁿ≅ℤ 0) + where + typIso : Iso _ _ + typIso = setTruncIso (curryIso ⋄ codomainIso S1→K₁≡S1×ℤ ⋄ toProdIso) + ⋄ setTruncOfProdIso + + theIso : GroupIso _ _ + fst theIso = typIso + snd theIso = + makeIsGroupHom + (coHomPointedElimT² _ (λ _ → isPropΠ λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ pf qf Pf → + coHomPointedElimT² _ (λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ pg qg Pg i → ∣ funExt (helperFst pf qf pg qg Pg Pf) i ∣₂ + , ∣ funExt (helperSnd pf qf pg qg Pg Pf) i ∣₂) + where + module _ (pf qf pg qg : 0ₖ 1 ≡ 0ₖ 1) (Pg : Square qg qg pg pg) (Pf : Square qf qf pf pf) where + helperFst : (x : S¹) + → fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y) +ₖ elimFunT² 0 pg qg Pg (x , y)) .fst + ≡ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y)) .fst + +ₖ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pg qg Pg (x , y)) .fst + helperFst base = refl + helperFst (loop i) j = loopLem j i + where + loopLem : cong (λ x → fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y) +ₖ elimFunT² 0 pg qg Pg (x , y)) .fst) loop + ≡ cong (λ x → fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y)) .fst + +ₖ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pg qg Pg (x , y)) .fst) loop + loopLem = (λ i j → S¹map-id (pf j +ₖ pg j) i) + ∙ (λ i j → S¹map-id (pf j) (~ i) +ₖ S¹map-id (pg j) (~ i)) + + helperSnd : (x : S¹) + → fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y) +ₖ elimFunT² 0 pg qg Pg (x , y)) .snd + ≡ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y)) .snd +ℤ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pg qg Pg (x , y)) .snd + helperSnd = + toPropElim (λ _ → isSetℤ _ _) + ((λ i → winding (basechange2⁻ base λ j → S¹map (∙≡+₁ qf qg (~ i) j))) + ∙∙ cong (winding ∘ basechange2⁻ base) (congFunct S¹map qf qg) + ∙∙ (cong winding (basechange2⁻-morph base (cong S¹map qf) (cong S¹map qg)) + ∙ winding-hom (basechange2⁻ base (cong S¹map qf)) (basechange2⁻ base (cong S¹map qg)))) + +----------------------- H²(T²) ------------------------------ + +H²-T²≅ℤ : GroupIso (coHomGr 2 (S₊ 1 × S₊ 1)) ℤGroup +H²-T²≅ℤ = compGroupIso helper2 (Hⁿ-Sⁿ≅ℤ 0) + where + helper : Iso (∥ ((a : S¹) → coHomK 2) ∥₂ × ∥ ((a : S¹) → coHomK 1) ∥₂) (coHom 1 S¹) + inv helper s = 0ₕ _ , s + fun helper = snd + leftInv helper _ = + ΣPathP (isOfHLevelSuc 0 (isOfHLevelRetractFromIso 0 (fst (Hⁿ-S¹≅0 0)) (isContrUnit)) _ _ + , refl) + rightInv helper _ = refl + theIso : Iso (coHom 2 (S¹ × S¹)) (coHom 1 S¹) + theIso = setTruncIso (curryIso ⋄ codomainIso S1→K2≡K2×K1 ⋄ toProdIso) + ⋄ setTruncOfProdIso + ⋄ helper + + helper2 : GroupIso (coHomGr 2 (S¹ × S¹)) (coHomGr 1 S¹) + helper2 .fst = theIso + helper2 .snd = makeIsGroupHom ( + coHomPointedElimT²'' 0 (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ P → coHomPointedElimT²'' 0 (λ _ → isSetSetTrunc _ _) + λ Q → ((λ i → ∣ (λ a → ΩKn+1→Kn 1 (sym (rCancel≡refl 0 i) + ∙∙ cong (λ x → (elimFunT²' 1 P (a , x) +ₖ elimFunT²' 1 Q (a , x)) -ₖ ∣ north ∣) loop + ∙∙ rCancel≡refl 0 i)) ∣₂)) + ∙∙ (λ i → ∣ (λ a → ΩKn+1→Kn 1 (rUnit (cong (λ x → rUnitₖ 2 (elimFunT²' 1 P (a , x) +ₖ elimFunT²' 1 Q (a , x)) i) loop) (~ i))) ∣₂) + ∙∙ (λ i → ∣ (λ a → ΩKn+1→Kn 1 (∙≡+₂ 0 (cong (λ x → elimFunT²' 1 P (a , x)) loop) (cong (λ x → elimFunT²' 1 Q (a , x)) loop) (~ i))) ∣₂) + ∙∙ (λ i → ∣ (λ a → ΩKn+1→Kn-hom 1 (cong (λ x → elimFunT²' 1 P (a , x)) loop) (cong (λ x → elimFunT²' 1 Q (a , x)) loop) i) ∣₂) + ∙∙ (λ i → ∣ ((λ a → ΩKn+1→Kn 1 (rUnit (cong (λ x → rUnitₖ 2 (elimFunT²' 1 P (a , x)) (~ i)) loop) i) + +ₖ ΩKn+1→Kn 1 (rUnit (cong (λ x → rUnitₖ 2 (elimFunT²' 1 Q (a , x)) (~ i)) loop) i))) ∣₂) + ∙ (λ i → ∣ ((λ a → ΩKn+1→Kn 1 (sym (rCancel≡refl 0 (~ i)) + ∙∙ cong (λ x → elimFunT²' 1 P (a , x) +ₖ ∣ north ∣) loop + ∙∙ rCancel≡refl 0 (~ i)) + +ₖ ΩKn+1→Kn 1 (sym (rCancel≡refl 0 (~ i)) + ∙∙ cong (λ x → elimFunT²' 1 Q (a , x) +ₖ ∣ north ∣) loop + ∙∙ rCancel≡refl 0 (~ i)))) ∣₂)) + +private + to₂ : coHom 2 (S₊ 1 × S₊ 1) → ℤ + to₂ = fun (fst H²-T²≅ℤ) + + from₂ : ℤ → coHom 2 (S₊ 1 × S₊ 1) + from₂ = inv (fst H²-T²≅ℤ) + + to₁ : coHom 1 (S₊ 1 × S₊ 1) → ℤ × ℤ + to₁ = fun (fst H¹-T²≅ℤ×ℤ) + + from₁ : ℤ × ℤ → coHom 1 (S₊ 1 × S₊ 1) + from₁ = inv (fst H¹-T²≅ℤ×ℤ) + + to₀ : coHom 0 (S₊ 1 × S₊ 1) → ℤ + to₀ = fun (fst H⁰-T²≅ℤ) + + from₀ : ℤ → coHom 0 (S₊ 1 × S₊ 1) + from₀ = inv (fst H⁰-T²≅ℤ) + +{- +-- Compute fast: +test : to₁ (from₁ (0 , 1) +ₕ from₁ (1 , 0)) ≡ (1 , 1) +test = refl + +test2 : to₁ (from₁ (5 , 1) +ₕ from₁ (-2 , 3)) ≡ (3 , 4) +test2 = refl + +-- Compute pretty fast +test3 : to₂ (from₂ 1) ≡ 1 +test3 = refl + +test4 : to₂ (from₂ 2) ≡ 2 +test4 = refl + +test5 : to₂ (from₂ 3) ≡ 3 +test5 = refl + +-- Compute, but slower +test6 : to₂ (from₂ 0 +ₕ from₂ 0) ≡ 0 +test6 = refl + +test6 : to₂ (from₂ 0 +ₕ from₂ 1) ≡ 1 +test6 = refl + +-- Does not compute +test7 : to₂ (from₂ 1 +ₕ from₂ 0) ≡ 1 +test7 = refl +-} + + +-- Proof (by computation) that T² ≠ S² ∨ S¹ ∨ S¹ +private + hasTrivial⌣₁ : ∀ {ℓ} (A : Type ℓ) → Type ℓ + hasTrivial⌣₁ A = (x y : coHom 1 A) → x ⌣ y ≡ 0ₕ 2 + + hasTrivial⌣₁S²∨S¹∨S¹ : hasTrivial⌣₁ S²⋁S¹⋁S¹ + hasTrivial⌣₁S²∨S¹∨S¹ x y = + x ⌣ y ≡⟨ cong₂ _⌣_ (sym (leftInv (fst (H¹-S²⋁S¹⋁S¹)) x)) (sym (leftInv (fst (H¹-S²⋁S¹⋁S¹)) y)) ⟩ + from₁-∨ (to₁-∨ x) ⌣ from₁-∨ (to₁-∨ y) ≡⟨ sym (leftInv (fst (H²-S²⋁S¹⋁S¹)) (from₁-∨ (to₁-∨ x) ⌣ from₁-∨ (to₁-∨ y))) ⟩ + from₂-∨ (to₂-∨ (from₁-∨ (to₁-∨ x) ⌣ from₁-∨ (to₁-∨ y))) ≡⟨ refl ⟩ -- holds by computation (even with open terms in the context)! + from₂-∨ 0 ≡⟨ hom1g (snd ℤGroup) from₂-∨ (snd (coHomGr 2 S²⋁S¹⋁S¹)) + ((invGroupEquiv (GroupIso→GroupEquiv H²-S²⋁S¹⋁S¹)) .snd .pres·) ⟩ + 0ₕ 2 ∎ + + 1≠0 : ¬ (Path ℤ 1 0) + 1≠0 p = posNotnegsuc _ _ (cong predℤ p) + + ¬hasTrivial⌣₁T² : ¬ (hasTrivial⌣₁ (S¹ × S¹)) + ¬hasTrivial⌣₁T² p = 1≠0 1=0 + where + 1=0 : pos 1 ≡ pos 0 + 1=0 = + 1 ≡⟨ refl ⟩ -- holds by computation! + to₂ (from₁ (0 , 1) ⌣ from₁ (1 , 0)) ≡⟨ cong to₂ (p (from₁ (0 , 1)) (from₁ (1 , 0))) ⟩ + 0 ∎ + +T²≠S²⋁S¹⋁S¹ : ¬ S¹ × S¹ ≡ S²⋁S¹⋁S¹ +T²≠S²⋁S¹⋁S¹ p = ¬hasTrivial⌣₁T² (subst hasTrivial⌣₁ (sym p) hasTrivial⌣₁S²∨S¹∨S¹) diff --git a/Cubical/ZCohomology/Groups/Unit.agda b/Cubical/ZCohomology/Groups/Unit.agda new file mode 100644 index 0000000000..6c0350f737 --- /dev/null +++ b/Cubical/ZCohomology/Groups/Unit.agda @@ -0,0 +1,85 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Groups.Unit where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (ℤ ; _+_ ; +Comm) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation +open import Cubical.Homotopy.Connected +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Algebra.Group renaming (Unit to UnitGroup) + +-- H⁰(Unit) +open IsGroupHom +open Iso + +H⁰-Unit≅ℤ : GroupIso (coHomGr 0 Unit) ℤ +fun (fst H⁰-Unit≅ℤ) = sRec isSetℤ (λ f → f tt) +inv (fst H⁰-Unit≅ℤ) a = ∣ (λ _ → a) ∣₂ +rightInv (fst H⁰-Unit≅ℤ) _ = refl +leftInv (fst H⁰-Unit≅ℤ) = sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ a → refl +snd H⁰-Unit≅ℤ = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) λ a b → refl) + +{- Hⁿ(Unit) for n ≥ 1 -} +isContrHⁿ-Unit : (n : ℕ) → isContr (coHom (suc n) Unit) +isContrHⁿ-Unit n = subst isContr (λ i → ∥ UnitToTypePath (coHomK (suc n)) (~ i) ∥₂) (helper' n) + where + helper' : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₂) + helper' n = + subst isContr + ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) + ∙∙ sym propTrunc≡Trunc2 + ∙∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₂) + (isConnectedSubtr 2 (helper2 n .fst) + (subst (λ x → isConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) + where + helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n + helper2 zero = 0 , refl + helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) + +Hⁿ-Unit≅0 : (n : ℕ) → GroupIso (coHomGr (suc n) Unit) UnitGroup +fun (fst (Hⁿ-Unit≅0 n)) _ = _ +inv (fst (Hⁿ-Unit≅0 n)) _ = 0ₕ (suc n) +rightInv (fst (Hⁿ-Unit≅0 n)) _ = refl +leftInv (fst (Hⁿ-Unit≅0 n)) _ = isOfHLevelSuc 0 (isContrHⁿ-Unit n) _ _ +snd (Hⁿ-Unit≅0 n) = makeIsGroupHom λ _ _ → refl + + +{- Hⁿ for arbitrary contractible types -} +private + Hⁿ-contrTypeIso : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A + → Iso (coHom (suc n) A) (coHom (suc n) Unit) + Hⁿ-contrTypeIso n contr = compIso (setTruncIso (isContr→Iso2 contr)) + (setTruncIso (invIso (isContr→Iso2 isContrUnit))) + +Hⁿ-contrType≅0 : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A + → GroupIso (coHomGr (suc n) A) UnitGroup +fun (fst (Hⁿ-contrType≅0 n contr)) _ = _ +inv (fst (Hⁿ-contrType≅0 n contr)) _ = 0ₕ (suc n) +rightInv (fst (Hⁿ-contrType≅0 n contr)) _ = refl +leftInv (fst (Hⁿ-contrType≅0 {A = A} n contr)) _ = isOfHLevelSuc 0 helper _ _ + where + helper : isContr (coHom (suc n) A) + helper = (Iso.inv (Hⁿ-contrTypeIso n contr) (0ₕ (suc n))) + , λ y → cong (Iso.inv (Hⁿ-contrTypeIso n contr)) + (isOfHLevelSuc 0 (isContrHⁿ-Unit n) (0ₕ (suc n)) (Iso.fun (Hⁿ-contrTypeIso n contr) y)) + ∙ Iso.leftInv (Hⁿ-contrTypeIso n contr) y +snd (Hⁿ-contrType≅0 n contr) = makeIsGroupHom λ _ _ → refl + +isContr-HⁿRed-Unit : (n : ℕ) → isContr (coHomRed n (Unit , tt)) +fst (isContr-HⁿRed-Unit n) = 0ₕ∙ _ +snd (isContr-HⁿRed-Unit n) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP (funExt (λ _ → sym p) + , λ i j → p (~ i ∨ j)))} diff --git a/Cubical/ZCohomology/Groups/Wedge.agda b/Cubical/ZCohomology/Groups/Wedge.agda new file mode 100644 index 0000000000..afc92655ce --- /dev/null +++ b/Cubical/ZCohomology/Groups/Wedge.agda @@ -0,0 +1,243 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.Wedge where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) +open import Cubical.HITs.Wedge +open import Cubical.Data.Int hiding (_+_) +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∣_∣ to ∣_∣₁) +open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; elim2 to trElim2) +open import Cubical.Data.Nat +open import Cubical.Algebra.Group + +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Isomorphism +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.Equiv + +open IsGroupHom +open Iso +{- +This module proves that Hⁿ(A ⋁ B) ≅ Hⁿ(A) × Hⁿ(B) for n ≥ 1 directly (rather than by means of Mayer-Vietoris). +It also proves that Ĥⁿ(A ⋁ B) ≅ Ĥ⁰(A) × Ĥ⁰(B) (reduced groups) + +Proof sketch for n ≥ 1: + +Any ∣ f ∣₂ ∈ Hⁿ(A ⋁ B) is uniquely characterised by a pair of functions + f₁ : A → Kₙ + f₂ : B → Kₙ +together with a path + p : f₁ (pt A) ≡ f₂ (pt B) + +The map F : Hⁿ(A ⋁ B) → Hⁿ(A) × Hⁿ(B) simply forgets about p, i.e.: + F(∣ f ∣₂) := (∣ f₁ ∣₂ , ∣ f₂ ∣₂) + +The construction of its inverse is defined by + F⁻(∣ f₁ ∣₂ , ∣ f₂ ∣₂) := ∣ f₁∨f₂ ∣₂ + where + f₁∨f₂ : A ⋁ B → Kₙ is defined inductively by + f₁∨f₂ (inl x) := f₁ x + f₂ (pt B) + f₁∨f₂ (inr x) := f₁ (pt B) + f₂ x + cong f₁∨f₂ (push tt) := refl + (this is the map wedgeFun⁻ below) + +The fact that F and F⁻ cancel out is a proposition and we may thus assume for any + ∣ f ∣₂ ∈ Hⁿ(A ⋁ B) and its corresponding f₁ that + f₁ (pt A) = f₂ (pt B) = 0 (*) + and + f (inl (pt A)) = 0 (**) + +The fact that F(F⁻(∣ f₁ ∣₂ , ∣ f₂ ∣₂)) = ∣ f₁ ∣₂ , ∣ f₂ ∣₂) follows immediately from (*) + +The other way is slightly trickier. We need to construct paths + Pₗ(x) : f (inl (x)) + f (inr (pt B)) ---> f (inl (x)) + Pᵣ(x) : f (inl (pt A)) + f (inr (x)) ---> f (inr (x)) + +Together with a filler of the following square + + cong f (push tt) + -----------------> + ^ ^ + | | + | | +Pₗ(pr A) | | Pᵣ(pt B) + | | + | | + | | + -----------------> + refl + +The square is filled by first constructing Pₗ by + f (inl (x)) + f (inr (pt B)) ---[cong f (push tt)⁻¹]---> + f (inl (x)) + f (inl (pt A)) ---[(**)]---> + f (inl (x)) + 0 ---[right-unit]---> + f (inl (x)) + +and then Pᵣ by + f (inl (pt A)) + f (inr (x)) ---[(**)⁻¹]---> + 0 + f (inr (x)) ---[left-unit]---> + f (inr (x)) + +and finally by using the fact that the group laws for Kₙ are refl at its base point. +-} + +module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where + + private + wedgeFun⁻ : ∀ n → (f : typ A → coHomK (suc n)) (g : typ B → coHomK (suc n)) + → ((A ⋁ B) → coHomK (suc n)) + wedgeFun⁻ n f g (inl x) = f x +ₖ g (pt B) + wedgeFun⁻ n f g (inr x) = f (pt A) +ₖ g x + wedgeFun⁻ n f g (push a i) = f (pt A) +ₖ g (pt B) + + Hⁿ-⋁ : (n : ℕ) → GroupIso (coHomGr (suc n) (A ⋁ B)) (×coHomGr (suc n) (typ A) (typ B)) + fun (fst (Hⁿ-⋁ zero)) = + sElim (λ _ → isSet× isSetSetTrunc isSetSetTrunc) + λ f → ∣ (λ x → f (inl x)) ∣₂ , ∣ (λ x → f (inr x)) ∣₂ + inv (fst (Hⁿ-⋁ zero)) = + uncurry (sElim2 (λ _ _ → isSetSetTrunc) + λ f g → ∣ wedgeFun⁻ 0 f g ∣₂) + rightInv (fst (Hⁿ-⋁ zero)) = + uncurry + (coHomPointedElim _ (pt A) (λ _ → isPropΠ λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ f fId → coHomPointedElim _ (pt B) (λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ g gId → ΣPathP ((cong ∣_∣₂ (funExt (λ x → cong (f x +ₖ_) gId ∙ rUnitₖ 1 (f x)))) + , cong ∣_∣₂ (funExt (λ x → cong (_+ₖ g x) fId ∙ lUnitₖ 1 (g x))))) + leftInv (fst (Hⁿ-⋁ zero)) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ f → pRec (isSetSetTrunc _ _) + (λ fId → cong ∣_∣₂ (sym fId)) + (helper f _ refl)) + where + helper : (f : A ⋁ B → coHomK 1) (x : coHomK 1) + → f (inl (pt A)) ≡ x + → ∥ f ≡ wedgeFun⁻ 0 (λ x → f (inl x)) (λ x → f (inr x)) ∥ + helper f = + trElim (λ _ → isProp→isOfHLevelSuc 2 (isPropΠ λ _ → isPropPropTrunc)) + (sphereElim 0 (λ _ → isPropΠ λ _ → isPropPropTrunc) + λ inlId → ∣ funExt (λ { (inl x) → sym (rUnitₖ 1 (f (inl x))) + ∙∙ cong ((f (inl x)) +ₖ_) (sym inlId) + ∙∙ cong ((f (inl x)) +ₖ_) (cong f (push tt)) + ; (inr x) → sym (lUnitₖ 1 (f (inr x))) + ∙ cong (_+ₖ (f (inr x))) (sym inlId) + ; (push tt i) j → helper2 (f (inl (pt A))) (sym (inlId)) + (f (inr (pt B))) (cong f (push tt)) j i} ) ∣₁) + where + helper2 : (x : coHomK 1) (r : ∣ base ∣ ≡ x) (y : coHomK 1) (p : x ≡ y) + → PathP (λ j → ((sym (rUnitₖ 1 x) ∙∙ cong (x +ₖ_) r ∙∙ cong (x +ₖ_) p)) j + ≡ (sym (lUnitₖ 1 y) ∙ cong (_+ₖ y) r) j) + p refl + helper2 x = J (λ x r → (y : coHomK 1) (p : x ≡ y) + → PathP (λ j → ((sym (rUnitₖ 1 x) ∙∙ cong (x +ₖ_) r ∙∙ cong (x +ₖ_) p)) j + ≡ (sym (lUnitₖ 1 y) ∙ cong (_+ₖ y) r) j) + p refl) + λ y → J (λ y p → PathP (λ j → ((sym (rUnitₖ 1 ∣ base ∣) ∙∙ refl ∙∙ cong (∣ base ∣ +ₖ_) p)) j + ≡ (sym (lUnitₖ 1 y) ∙ refl) j) + p refl) + λ i _ → (refl ∙ (λ _ → 0ₖ 1)) i + snd (Hⁿ-⋁ zero) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ _ _ → refl) + fun (fst (Hⁿ-⋁ (suc n))) = + sElim (λ _ → isSet× isSetSetTrunc isSetSetTrunc) + λ f → ∣ (λ x → f (inl x)) ∣₂ , ∣ (λ x → f (inr x)) ∣₂ + inv (fst (Hⁿ-⋁ (suc n))) = + uncurry (sElim2 (λ _ _ → isSetSetTrunc) + λ f g → ∣ wedgeFun⁻ (suc n) f g ∣₂) + rightInv (fst (Hⁿ-⋁ (suc n))) = + uncurry + (coHomPointedElim _ (pt A) (λ _ → isPropΠ λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ f fId → coHomPointedElim _ (pt B) (λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ g gId → ΣPathP ((cong ∣_∣₂ (funExt (λ x → cong (f x +ₖ_) gId ∙ rUnitₖ (2 + n) (f x)))) + , cong ∣_∣₂ (funExt (λ x → cong (_+ₖ g x) fId ∙ lUnitₖ (2 + n) (g x))))) + leftInv (fst (Hⁿ-⋁ (suc n))) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ f → pRec (isSetSetTrunc _ _) + (λ fId → cong ∣_∣₂ (sym fId)) + (helper f _ refl)) + where + helper : (f : A ⋁ B → coHomK (2 + n)) (x : coHomK (2 + n)) + → f (inl (pt A)) ≡ x + → ∥ f ≡ wedgeFun⁻ (suc n) (λ x → f (inl x)) (λ x → f (inr x)) ∥ + helper f = + trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropΠ λ _ → isPropPropTrunc)) + (sphereToPropElim (suc n) (λ _ → isPropΠ λ _ → isPropPropTrunc) + λ inlId → (∣ funExt (λ { (inl x) → sym (rUnitₖ (2 + n) (f (inl x))) + ∙∙ cong ((f (inl x)) +ₖ_) (sym inlId) + ∙∙ cong ((f (inl x)) +ₖ_) (cong f (push tt)) + ; (inr x) → sym (lUnitₖ (2 + n) (f (inr x))) + ∙ cong (_+ₖ (f (inr x))) (sym inlId) + ; (push tt i) j → helper2 (f (inl (pt A))) (sym (inlId)) + (f (inr (pt B))) (cong f (push tt)) j i}) ∣₁)) + where + helper2 : (x : coHomK (2 + n)) (r : ∣ north ∣ ≡ x) (y : coHomK (2 + n)) (p : x ≡ y) + → PathP (λ j → ((sym (rUnitₖ (2 + n) x) ∙∙ cong (x +ₖ_) r ∙∙ cong (x +ₖ_) p)) j + ≡ (sym (lUnitₖ (2 + n) y) ∙ cong (_+ₖ y) r) j) + p refl + helper2 x = J (λ x r → (y : coHomK (2 + n)) (p : x ≡ y) + → PathP (λ j → ((sym (rUnitₖ (2 + n) x) ∙∙ cong (x +ₖ_) r ∙∙ cong (x +ₖ_) p)) j + ≡ (sym (lUnitₖ (2 + n) y) ∙ cong (_+ₖ y) r) j) + p refl) + λ y → J (λ y p → PathP (λ j → ((sym (rUnitₖ (2 + n) ∣ north ∣) ∙∙ refl ∙∙ cong (∣ north ∣ +ₖ_) p)) j + ≡ (sym (lUnitₖ (2 + n) y) ∙ refl) j) + p refl) + λ i j → ((λ _ → ∣ north ∣) ∙ refl) i + snd (Hⁿ-⋁ (suc n)) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ _ _ → refl) + + H⁰Red-⋁ : GroupIso (coHomRedGrDir 0 (A ⋁ B , inl (pt A))) + (DirProd (coHomRedGrDir 0 A) (coHomRedGrDir 0 B)) + fun (fst H⁰Red-⋁) = + sRec (isSet× isSetSetTrunc isSetSetTrunc) + λ {(f , p) → ∣ (f ∘ inl) , p ∣₂ + , ∣ (f ∘ inr) , cong f (sym (push tt)) ∙ p ∣₂} + inv (fst H⁰Red-⋁) = + uncurry (sRec2 isSetSetTrunc + λ {(f , p) (g , q) → ∣ (λ {(inl a) → f a + ; (inr b) → g b + ; (push tt i) → (p ∙ sym q) i}) + , p ∣₂}) + rightInv (fst H⁰Red-⋁) = + uncurry + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ {(_ , _) (_ , _) → ΣPathP (cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl) + , cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl))}) + leftInv (fst H⁰Red-⋁) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (funExt λ {(inl a) → refl + ; (inr b) → refl + ; (push tt i) j → (cong (p ∙_) (symDistr (cong f (sym (push tt))) p) + ∙∙ assoc∙ p (sym p) (cong f (push tt)) + ∙∙ cong (_∙ (cong f (push tt))) (rCancel p) + ∙ sym (lUnit (cong f (push tt)))) j i}))} + -- Alt. use isOfHLevel→isOfHLevelDep + snd H⁰Red-⋁ = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ {(f , p) (g , q) → ΣPathP (cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl) + , cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl))}) + + wedgeConnected : ((x : typ A) → ∥ pt A ≡ x ∥) → ((x : typ B) → ∥ pt B ≡ x ∥) → (x : A ⋁ B) → ∥ inl (pt A) ≡ x ∥ + wedgeConnected conA conB = + PushoutToProp (λ _ → isPropPropTrunc) + (λ a → pRec isPropPropTrunc (λ p → ∣ cong inl p ∣₁) (conA a)) + λ b → pRec isPropPropTrunc (λ p → ∣ push tt ∙ cong inr p ∣₁) (conB b) diff --git a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda new file mode 100644 index 0000000000..94e21aede2 --- /dev/null +++ b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda @@ -0,0 +1,134 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.WedgeOfSpheres where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Int renaming (_+_ to _ℤ+_) + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.RingStructure.CupProduct + +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.Pushout +open import Cubical.HITs.Truncation renaming (elim to trElim) hiding (map ; elim2) +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim) + +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) + + +S¹⋁S¹ : Type₀ +S¹⋁S¹ = S₊∙ 1 ⋁ S₊∙ 1 + +S²⋁S¹⋁S¹ : Type₀ +S²⋁S¹⋁S¹ = S₊∙ 2 ⋁ (S¹⋁S¹ , inl base) + +------------- H⁰(S¹⋁S¹) ------------ +H⁰-S¹⋁S¹ : GroupIso (coHomGr 0 S¹⋁S¹) ℤGroup +H⁰-S¹⋁S¹ = H⁰-connected (inl base) (wedgeConnected _ _ (Sn-connected 0) (Sn-connected 0)) + +------------- H¹(S¹⋁S¹) ------------ +H¹-S¹⋁S¹ : GroupIso (coHomGr 1 S¹⋁S¹) (DirProd ℤGroup ℤGroup) +H¹-S¹⋁S¹ = (Hⁿ-⋁ _ _ 0) □ GroupIsoDirProd coHom1S1≃ℤ coHom1S1≃ℤ + +------------- H⁰(S²⋁S¹⋁S¹) --------- +H⁰-S²⋁S¹⋁S¹ : GroupIso (coHomGr 0 S²⋁S¹⋁S¹) ℤGroup +H⁰-S²⋁S¹⋁S¹ = H⁰-connected (inl north) + (wedgeConnected _ _ + (Sn-connected 1) + (wedgeConnected _ _ + (Sn-connected 0) + (Sn-connected 0))) + +------------- H¹(S²⋁S¹⋁S¹) --------- +H¹-S²⋁S¹⋁S¹ : GroupIso (coHomGr 1 S²⋁S¹⋁S¹) (DirProd ℤGroup ℤGroup) +H¹-S²⋁S¹⋁S¹ = + Hⁿ-⋁ (S₊∙ 2) (S¹⋁S¹ , inl base) 0 + □ GroupIsoDirProd (H¹-Sⁿ≅0 0) H¹-S¹⋁S¹ + □ lUnitGroupIso + +------------- H²(S²⋁S¹⋁S¹) --------- + +H²-S²⋁S¹⋁S¹ : GroupIso (coHomGr 2 S²⋁S¹⋁S¹) ℤGroup +H²-S²⋁S¹⋁S¹ = + compGroupIso + (Hⁿ-⋁ _ _ 1) + (GroupIsoDirProd {B = UnitGroup} + (Hⁿ-Sⁿ≅ℤ 1) + ((Hⁿ-⋁ _ _ 1) □ GroupIsoDirProd (Hⁿ-S¹≅0 0) (Hⁿ-S¹≅0 0) □ rUnitGroupIso) + □ rUnitGroupIso) + +open Iso + +to₂ : coHom 2 S²⋁S¹⋁S¹ → ℤ +to₂ = fun (fst H²-S²⋁S¹⋁S¹) +from₂ : ℤ → coHom 2 S²⋁S¹⋁S¹ +from₂ = inv (fst H²-S²⋁S¹⋁S¹) + +to₁ : coHom 1 S²⋁S¹⋁S¹ → ℤ × ℤ +to₁ = fun (fst H¹-S²⋁S¹⋁S¹) +from₁ : ℤ × ℤ → coHom 1 S²⋁S¹⋁S¹ +from₁ = inv (fst H¹-S²⋁S¹⋁S¹) + +to₀ : coHom 0 S²⋁S¹⋁S¹ → ℤ +to₀ = fun (fst H⁰-S²⋁S¹⋁S¹) +from₀ : ℤ → coHom 0 S²⋁S¹⋁S¹ +from₀ = inv (fst H⁰-S²⋁S¹⋁S¹) + +{- + +-- Compute pretty fast +test1 : to₁ (from₁ (1 , 0) +ₕ from₁ (0 , 1)) ≡ (1 , 1) +test1 = refl + +-- Computes, but only when computing some smaller numbers first +test2 : to₁ (from₁ (50 , 3) +ₕ from₁ (2 , -2)) ≡ (52 , 1) +test2 = refl + +test3 : to₂ (from₂ 0) ≡ 0 +test3 = refl + +test4 : to₂ (from₂ 3) ≡ 3 +test4 = refl + +test5 : to₂ (from₂ 1 +ₕ from₂ 1) ≡ 2 +test5 = refl +-} +{- + g : S²⋁S¹⋁S¹ → ∥ Susp S¹ ∥ 4 + g (inl x) = ∣ x ∣ + g (inr x) = 0ₖ _ + g (push a i) = 0ₖ _ + + G = ∣ g ∣₂ + +-- Does not compute: +test₀ : to₂ (G +ₕ G) ≡ 2 +test₀ = refl + +but this does: +test₀ : to₂ (G +'ₕ G) ≡ 2 +test₀ = refl +-} + + +-- Cup product is trivial +⌣-gen₁ : to₂ (from₁ (1 , 0) ⌣ from₁ (0 , 1)) ≡ 0 +⌣-gen₁ = refl + +-- Even better: +⌣-gen : (x y : ℤ × ℤ) → to₂ (from₁ x ⌣ from₁ y) ≡ 0 +⌣-gen x y = refl diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda new file mode 100644 index 0000000000..ee91d7f1ab --- /dev/null +++ b/Cubical/ZCohomology/Gysin.agda @@ -0,0 +1,1137 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Gysin where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.RingStructure.GradedCommutativity + +open import Cubical.Relation.Nullary + +open import Cubical.Functions.Embedding + +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Data.Sum +open import Cubical.Data.Fin +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.AbGroup + +open import Cubical.HITs.Pushout.Flattening +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) +open import Cubical.HITs.Join + +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Hopf +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Group.Base + +-- There seems to be some problems with the termination checker. +-- Spelling out integer induction with 3 base cases like this +-- solves the issue. +private + Int-ind : ∀ {ℓ} (P : ℤ → Type ℓ) + → P (pos zero) → P (pos 1) + → P (negsuc zero) + → ((x y : ℤ) → P x → P y → P (x + y)) → (x : ℤ) → P x + Int-ind P z one min ind (pos zero) = z + Int-ind P z one min ind (pos (suc zero)) = one + Int-ind P z one min ind (pos (suc (suc n))) = + ind (pos (suc n)) 1 (Int-ind P z one min ind (pos (suc n))) one + Int-ind P z one min ind (negsuc zero) = min + Int-ind P z one min ind (negsuc (suc n)) = + ind (negsuc n) (negsuc zero) (Int-ind P z one min ind (negsuc n)) min + +-- The untruncated version (coHomRed n (S₊ n)), i.e. +-- (S₊∙ n →∙ coHomK-ptd n) is in fact equal to +-- (coHomRed n (S₊ n)). Its useful to formulate +-- (S₊∙ n →∙ coHomK-ptd n) as a group in its own right +-- and prove it equivalent to (coHomRed n (S₊ n)). + +-- We start with the addition +private + _++_ : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → + (A →∙ coHomK-ptd n) → (A →∙ coHomK-ptd n) → (A →∙ coHomK-ptd n) + fst (f ++ g) x = fst f x +ₖ fst g x + snd (f ++ g) = cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ _ (0ₖ _) + +-- If we truncate the addition, we get back our usual +-- addition on (coHomRed n (S₊ n)) +addAgree : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x y : A →∙ coHomK-ptd n) + → Path (fst (coHomRedGrDir n A)) + ∣ x ++ y ∣₂ + (∣ x ∣₂ +ₕ∙ ∣ y ∣₂) +addAgree {A = A} zero f g = + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) +addAgree {A = A} (suc zero) f g = + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) +addAgree {A = A} (suc (suc n)) f g = + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) + +-- We formulate (S₊∙ n →∙ coHomK-ptd n) as a group, πS. +module _ where + open import Cubical.Algebra.Monoid + open import Cubical.Algebra.Semigroup + open GroupStr + open IsGroup + open IsMonoid + open IsSemigroup + πS : (n : ℕ) → Group ℓ-zero + fst (πS n) = S₊∙ n →∙ coHomK-ptd n + 1g (snd (πS n)) = (λ _ → 0ₖ n) , refl + GroupStr._·_ (snd (πS n)) = + λ f g → (λ x → fst f x +ₖ fst g x) + , cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ n (0ₖ n) + inv (snd (πS n)) f = (λ x → -ₖ fst f x) , cong -ₖ_ (snd f) ∙ -0ₖ {n = n} + is-set (isSemigroup (isMonoid (isGroup (snd (πS zero))))) = + isOfHLevelΣ 2 (isSetΠ (λ _ → isSetℤ)) + λ _ → isOfHLevelPath 2 isSetℤ _ _ + is-set (isSemigroup (isMonoid (isGroup (snd (πS (suc n)))))) = + isOfHLevel↑∙' 0 n + IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (snd (πS n))))) x y z = + →∙Homogeneous≡ (isHomogeneousKn n) + (funExt λ w → assocₖ n (fst x w) (fst y w) (fst z w)) + fst (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rUnitₖ n (f x)) + snd (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lUnitₖ n (f x)) + fst (inverse (isGroup (snd (πS n))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rCancelₖ n (f x)) + snd (inverse (isGroup (snd (πS n))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lCancelₖ n (f x)) + + isSetπS : (n : ℕ) → isSet (S₊∙ n →∙ coHomK-ptd n) + isSetπS = λ n → is-set (snd (πS n)) + +-- πS is equivalent to (coHomRed n (S₊∙ n)) +K : (n : ℕ) → GroupIso (coHomRedGrDir n (S₊∙ n)) (πS n) +fst (K n) = setTruncIdempotentIso (isSetπS n) +snd (K zero) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 0) _ _) + λ f g → →∙Homogeneous≡ (isHomogeneousKn 0) refl) +snd (K (suc zero)) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 1) _ _) + λ f g → →∙Homogeneous≡ (isHomogeneousKn 1) refl) +snd (K (suc (suc n))) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS _) _ _) + λ f g → →∙Homogeneous≡ (isHomogeneousKn _) refl) + +-- πS has the following generator +genFunSpace : (n : ℕ) → S₊∙ n →∙ coHomK-ptd n +fst (genFunSpace zero) false = 1 +fst (genFunSpace zero) true = 0 +snd (genFunSpace zero) = refl +fst (genFunSpace (suc n)) = ∣_∣ +snd (genFunSpace (suc zero)) = refl +snd (genFunSpace (suc (suc n))) = refl + +πS≅ℤ : (n : ℕ) → GroupIso (πS n) ℤGroup +fst (πS≅ℤ zero) = IsoBool→∙ +snd (πS≅ℤ zero) = makeIsGroupHom λ _ _ → refl +πS≅ℤ (suc n) = + compGroupIso + (invGroupIso (K (suc n))) + (compGroupIso + (GroupEquiv→GroupIso (coHomGr≅coHomRedGr n (S₊∙ (suc n)))) + (Hⁿ-Sⁿ≅ℤ n)) + +Iso-πS-ℤ : (n : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc n)) ℤ +Iso-πS-ℤ n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) + (compIso (equivToIso (fst (coHomGr≅coHomRedGr n (S₊∙ (suc n))))) + (fst (Hⁿ-Sⁿ≅ℤ n))) + +Iso-πS-ℤ' : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ +Iso-πS-ℤ' n = fst (πS≅ℤ n) + +Iso-πS-ℤPres1 : (n : ℕ) → Iso.fun (fst (πS≅ℤ (suc n))) (∣_∣ , refl) ≡ pos 1 +Iso-πS-ℤPres1 zero = refl +Iso-πS-ℤPres1 (suc n) = + cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ Iso-πS-ℤPres1 n + where + lem : (n : ℕ) → Iso.fun (fst (suspensionAx-Sn n n)) ∣ ∣_∣ ∣₂ ≡ ∣ ∣_∣ ∣₂ + lem zero = + cong ∣_∣₂ (funExt λ x → transportRefl (∣ x ∣ +ₖ (0ₖ 1)) ∙ rUnitₖ 1 ∣ x ∣) + lem (suc n) = cong ∣_∣₂ + (funExt λ x → (λ i → transportRefl ((ΩKn+1→Kn (suc (suc n)) + (transp (λ j → 0ₖ (suc (suc (suc n))) ≡ ∣ merid north (~ j ∧ ~ i) ∣) i + (λ z → ∣ compPath-filler + (merid (transportRefl (transportRefl x i) i)) + (sym (merid north)) i z + ∣)))) i) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc (suc n))) ∣ x ∣) + +-- The first step of the Gysin sequence is to formulate +-- an equivalence g : Kᵢ ≃ (Sⁿ →∙ Kᵢ₊ₙ) +module g-base where + g : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (i +' n)) + fst (g n i x) y = x ⌣ₖ (genFunSpace n) .fst y + snd (g n i x) = cong (x ⌣ₖ_) ((genFunSpace n) .snd) ∙ ⌣ₖ-0ₖ i n x + + -- We give another version of g which will be easier to work with now + G : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (n +' i)) + fst (G n i x) y = (genFunSpace n) .fst y ⌣ₖ x + snd (G n i x) = cong (_⌣ₖ x) ((genFunSpace n) .snd) ∙ 0ₖ-⌣ₖ n i x + + -ₖ^-Iso : (n : ℕ) (i : ℕ) + → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + -ₖ^-Iso n i = isoToEquiv (iso F F FF FF) + where + lem : (i n : ℕ) → (-ₖ^ i · n) (snd (coHomK-ptd (i +' n))) ≡ 0ₖ _ + lem zero zero = refl + lem zero (suc zero) = refl + lem zero (suc (suc n)) = refl + lem (suc zero) zero = refl + lem (suc (suc i)) zero = refl + lem (suc i) (suc n) = refl + + F : S₊∙ n →∙ coHomK-ptd (i +' n) → S₊∙ n →∙ coHomK-ptd (i +' n) + fst (F f) x = (-ₖ^ i · n) (fst f x) + snd (F f) = cong (-ₖ^ i · n) (snd f) ∙ lem i n + + FF : (x : _) → F (F x) ≡ x + FF x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ y → -ₖ-gen² i n _ _ (fst x y)) + + transpPres0ₖ : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m + transpPres0ₖ {k = k} = + J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) + + -- There will be some index swapping going on. We state this explicitly, since we will + -- need to trace the maps later + indexSwap : (n : ℕ) (i : ℕ) + → (S₊∙ n →∙ coHomK-ptd (n +' i)) + ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + indexSwap n i = + isoToEquiv (iso (λ f → (λ x → subst coHomK (+'-comm n i) (fst f x)) , + cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0ₖ (+'-comm n i)) + (λ f → (λ x → subst coHomK (sym (+'-comm n i)) (fst f x)) + , (cong (subst coHomK (sym (+'-comm n i))) (snd f) + ∙ transpPres0ₖ (sym (+'-comm n i)))) + (λ f → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → transportTransport⁻ _ (f .fst x))) + λ f → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → transportTransport⁻ _ (f .fst x))) + + -- g is a composition of G and our two previous equivs. + g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x + → fst (compEquiv (indexSwap n i) (-ₖ^-Iso n i)) ((G n i) x) + g≡ n i = funExt (λ f → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ y → gradedComm-⌣ₖ _ _ f (genFunSpace n .fst y))) + + -- We need a third Iso. + + suspKn-Iso-fun : (n m : ℕ) → + (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + → (S₊ n → coHomK m) + suspKn-Iso-fun zero m (f , p) false = ΩKn+1→Kn _ (sym p ∙∙ cong f loop ∙∙ p) + suspKn-Iso-fun zero m (f , p) true = 0ₖ _ + suspKn-Iso-fun (suc n) m (f , p) x = + ΩKn+1→Kn _ (sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) + + suspKn-Iso-fun∙ : (n m : ℕ) → (f : _) → suspKn-Iso-fun n m f (ptSn _) ≡ 0ₖ m + suspKn-Iso-fun∙ zero m = λ _ → refl + suspKn-Iso-fun∙ (suc n) m (f , p) = + cong (ΩKn+1→Kn m) + (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) + ∙∙ cong (ΩKn+1→Kn m) (∙∙lCancel p) + ∙∙ ΩKn+1→Kn-refl m + + suspKn-Iso-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) + → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + fst (suspKn-Iso-inv zero m (f , p)) base = 0ₖ _ + fst (suspKn-Iso-inv zero m (f , p)) (loop i) = Kn→ΩKn+1 m (f false) i + snd (suspKn-Iso-inv zero m (f , p)) = refl + fst (suspKn-Iso-inv (suc n) m (f , p)) north = 0ₖ _ + fst (suspKn-Iso-inv (suc n) m (f , p)) south = 0ₖ _ + fst (suspKn-Iso-inv (suc n) m (f , p)) (merid a i) = Kn→ΩKn+1 m (f a) i + snd (suspKn-Iso-inv (suc n) m (f , p)) = refl + + suspKn-Iso : (n m : ℕ) → + Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + (S₊∙ n →∙ coHomK-ptd m) + Iso.fun (suspKn-Iso n m) f = (suspKn-Iso-fun n m f) + , (suspKn-Iso-fun∙ n m f) + Iso.inv (suspKn-Iso n m) = suspKn-Iso-inv n m + Iso.rightInv (suspKn-Iso zero m) (f , p) = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { false → cong (ΩKn+1→Kn m) (sym (rUnit _)) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f false) + ; true → sym p}) + Iso.rightInv (suspKn-Iso (suc n) m) (f , p) = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → + (λ i → ΩKn+1→Kn m + (sym (rUnit + (cong-∙ + (suspKn-Iso-inv (suc n) m (f , p) .fst) + (merid x) (sym (merid (ptSn _))) i)) i)) + ∙∙ cong (ΩKn+1→Kn m) + (cong (Kn→ΩKn+1 m (f x) ∙_) + (cong sym (cong (Kn→ΩKn+1 m) p ∙ Kn→ΩKn+10ₖ m)) + ∙ sym (rUnit _)) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) + Iso.leftInv (suspKn-Iso zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { base → sym p + ; (loop i) j → lem j i}) + where + lem : PathP (λ i → p (~ i) ≡ p (~ i)) + (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f loop ∙∙ p))) + (cong f loop) + lem = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ + ◁ λ i → doubleCompPath-filler (sym p) (cong f loop) p (~ i) + Iso.leftInv (suspKn-Iso (suc n) m) (f , p) = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { north → sym p + ; south → sym p ∙ cong f (merid (ptSn _)) + ; (merid a i) j → lem a j i}) + where + lem : (a : S₊ (suc n)) + → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) + (Kn→ΩKn+1 m + (ΩKn+1→Kn m + (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p))) + (cong f (merid a)) + lem a = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ + ◁ λ i j → hcomp (λ k → + λ { (i = i1) → (f (merid a j)) + ; (j = i0) → p (k ∧ ~ i) + ; (j = i1) → compPath-filler' + (sym p) (cong f (merid (ptSn _))) k i }) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + + glIsoInvHom : (n m : ℕ) (x y : coHomK n) (z : S₊ (suc m)) + → Iso.inv (suspKn-Iso _ _) (G m n (x +ₖ y)) .fst z + ≡ Iso.inv (suspKn-Iso _ _) (G m n x) .fst z + +ₖ Iso.inv (suspKn-Iso _ _) (G m n y) .fst z + glIsoInvHom zero zero x y base = refl + glIsoInvHom (suc n) zero x y base = refl + glIsoInvHom zero zero x y (loop i) j = lem j i + where + lem : (cong (Iso.inv (suspKn-Iso _ _) (G zero zero (x + y)) .fst) loop) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (suspKn-Iso _ _) (G zero zero x) .fst) loop) + (cong (Iso.inv (suspKn-Iso _ _) (G zero zero y) .fst) loop) + lem = Kn→ΩKn+1-hom 0 x y + ∙ ∙≡+₁ (cong (Iso.inv (suspKn-Iso _ _) (G zero zero x) .fst) loop) + (cong (Iso.inv (suspKn-Iso _ _) (G zero zero y) .fst) loop) + glIsoInvHom (suc n) zero x y (loop i) j = lem j i + where + lem : Kn→ΩKn+1 (suc n) ((pos (suc zero)) ·₀ (x +ₖ y)) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (suspKn-Iso zero (zero +' suc n)) + (G zero (suc n) x) .fst) loop) + (cong (Iso.inv (suspKn-Iso zero (zero +' suc n)) + (G zero (suc n) y) .fst) loop) + lem = cong (Kn→ΩKn+1 (suc n)) (lUnit⌣ₖ _ (x +ₖ y)) + ∙∙ Kn→ΩKn+1-hom (suc n) x y + ∙∙ (λ i → ∙≡+₂ n (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ x (~ i))) + (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ y (~ i))) i) + glIsoInvHom zero (suc m) x y north = refl + glIsoInvHom zero (suc m) x y south = refl + glIsoInvHom zero (suc m) x y (merid a i) j = lem j i + where + lem : Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ (x + y)) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ x)) + (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ y)) + lem = cong (Kn→ΩKn+1 (suc m)) (leftDistr-⌣ₖ (suc m) 0 ∣ a ∣ x y) + ∙∙ Kn→ΩKn+1-hom (suc m) _ _ + ∙∙ ∙≡+₂ _ _ _ + glIsoInvHom (suc n) (suc m) x y north = refl + glIsoInvHom (suc n) (suc m) x y south = refl + glIsoInvHom (suc n) (suc m) x y (merid a i) j = lem j i + where + lem : Kn→ΩKn+1 (suc (suc (m +ℕ n))) + (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ (x +ₖ y)) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc (suc (m +ℕ n))) + (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ x)) + (Kn→ΩKn+1 (suc (suc (m +ℕ n))) + (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ y)) + lem = cong (Kn→ΩKn+1 (suc (suc (m +ℕ n)))) + (leftDistr-⌣ₖ (suc m) (suc n) ∣ a ∣ x y) + ∙∙ Kn→ΩKn+1-hom _ _ _ + ∙∙ ∙≡+₂ _ _ _ + + private + +'-suc : (n m : ℕ) → suc n +' m ≡ suc (n +' m) + +'-suc zero zero = refl + +'-suc (suc n) zero = refl + +'-suc zero (suc m) = refl + +'-suc (suc n) (suc m) = refl + + decomposeG : (i n : ℕ) (x : coHomK i) + → G (suc n) i x + ≡ subst (λ x → S₊∙ (suc n) →∙ coHomK-ptd x) + (sym (+'-suc n i)) + (Iso.inv (suspKn-Iso n _) (G n i x)) + decomposeG zero zero x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → (λ i → x ·₀ ∣ z ∣) ∙ h3 x z ∙ sym (transportRefl _)) + where + h3 : (x : ℤ) (z : S¹) + → _·₀_ {n = 1} x ∣ z ∣ + ≡ fst (Iso.inv (suspKn-Iso 0 zero) (G zero zero x)) z + h3 = + Int-ind _ + (λ { base → refl ; (loop i) j → Kn→ΩKn+10ₖ zero (~ j) i}) + (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (lUnit loop j)) j i}) + (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (sym loop)) j i}) + λ x y inx iny z + → rightDistr-⌣ₖ 0 1 x y ∣ z ∣ + ∙∙ cong₂ _+ₖ_ (inx z) (iny z) + ∙∙ sym (glIsoInvHom zero zero x y z) + decomposeG (suc i) zero x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → lem z + ∙ sym (transportRefl + ((Iso.inv (suspKn-Iso zero (suc i)) (G zero (suc i) x)) .fst z))) + where + lem : (z : S₊ 1) + → _ ≡ Iso.inv (suspKn-Iso zero (suc i)) (G zero (suc i) x) .fst z + lem base = refl + lem (loop k) j = Kn→ΩKn+1 (suc i) (lUnit⌣ₖ _ x (~ j)) k + decomposeG zero (suc n) x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → main x z + ∙ λ i → transportRefl + (Iso.inv (suspKn-Iso (suc n) (suc n)) + (G (suc n) zero x)) (~ i) .fst z) + where + +merid : rUnitₖ (suc (suc n)) ∣ south ∣ ≡ cong ∣_∣ₕ (merid (ptSn _)) + +merid = sym (lUnit _) + ∙ cong (cong ∣_∣ₕ) + λ j i → transp (λ _ → S₊ (suc (suc n))) (i ∨ j) + (merid (ptSn (suc n)) i) + + lem : (a : S₊ (suc n)) + → PathP (λ i → ∣ merid a i ∣ₕ + ≡ Kn→ΩKn+1 (suc n) (∣ a ∣ +ₖ (0ₖ _)) i) + refl (sym (rUnitₖ (suc (suc n)) ∣ south ∣)) + lem a = + flipSquare ((λ i j → ∣ compPath-filler (merid a) + (sym (merid (ptSn _))) i j ∣ₕ) + ▷ cong (Kn→ΩKn+1 (suc n)) (sym (rUnitₖ (suc n) ∣ a ∣ₕ))) + ▷ sym (cong sym +merid) + + lem₂ : (a : S₊ (suc n)) → (λ i → -ₖ ∣ merid a i ∣) + ≡ Kn→ΩKn+1 (suc n) (-ₖ ∣ a ∣) + lem₂ a = + cong (cong ∣_∣ₕ) (sym (symDistr (merid a) + (sym (merid (ptSn (suc n)))))) + ∙ sym (Kn→ΩKn+1-ₖ (suc n) ∣ a ∣) + + main : (x : ℤ) (z : S₊ (suc (suc n))) + → _⌣ₖ_ {n = suc (suc n)} {m = 0} ∣ z ∣ x + ≡ Iso.inv (suspKn-Iso (suc n) (suc n)) (G (suc n) zero x) .fst z + main = Int-ind _ + (λ { north → refl ; south → refl + ; (merid a i) j → Kn→ΩKn+10ₖ (suc n) (~ j) i}) + (λ { north → refl ; south → refl + ; (merid a i) j → + hcomp (λ k → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → rUnitₖ (suc (suc n)) ∣ south ∣ (~ k) + ; (j = i0) → rUnitₖ (suc (suc n)) ∣ merid a i ∣ (~ k) + ; (j = i1) → lem a i k}) + ∣ merid a i ∣ₕ}) + (λ { north → refl + ; south → refl + ; (merid a i) j → lem₂ a j i}) + λ x y indx indy z → leftDistr-⌣ₖ _ _ ∣ z ∣ x y + ∙ cong₂ _+ₖ_ (indx z) (indy z) + ∙ sym (glIsoInvHom _ _ _ _ _) + decomposeG (suc i) (suc n) x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → lem z + ∙ λ j → + transportRefl ((Iso.inv (suspKn-Iso (suc n) (suc (suc (n +ℕ i)))) + (G (suc n) (suc i) x))) (~ j) .fst z) + where + lem : (z : S₊ (suc (suc n))) → _ + lem north = refl + lem south = refl + lem (merid a i) = refl + + isEquivGzero : (i : ℕ) → isEquiv (G zero i) + isEquivGzero i = + isoToIsEquiv + (iso _ (λ f → fst f false) + (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) + (lUnit⌣ₖ _)) + + isEquivG : (n i : ℕ) → isEquiv (G n i) + isEquivG zero i = + isoToIsEquiv + (iso _ (λ f → fst f false) + (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) + (lUnit⌣ₖ _)) + isEquivG (suc n) i = + subst isEquiv (sym (funExt (decomposeG i n))) + (compEquiv (compEquiv (G n i , isEquivG n i) + (isoToEquiv (invIso (suspKn-Iso n (n +' i))))) + (pathToEquiv (λ j → S₊∙ (suc n) →∙ coHomK-ptd (+'-suc n i (~ j)))) .snd) + + isEquiv-g : (n i : ℕ) → isEquiv (g n i) + isEquiv-g n i = + subst isEquiv (sym (g≡ n i)) + (compEquiv (G n i , isEquivG n i) + (compEquiv (indexSwap n i) (-ₖ^-Iso n i)) .snd) + +-- We now generealise the equivalence g to also apply to arbitrary fibrations (Q : B → Type) +-- satisfying (Q * ≃∙ Sⁿ) +module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥) + (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) + (c : (b : typ B) → (Q b →∙ coHomK-ptd n)) + (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) + where + + g : (b : typ B) (i : ℕ) → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) + fst (g b i x) y = x ⌣ₖ c b .fst y + snd (g b i x) = cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x + + g-hom : (b : typ B) (i : ℕ) → (x y : coHomK i) + → g b i (x +ₖ y) ≡ ((g b i x) ++ (g b i y)) + g-hom b i x y = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → rightDistr-⌣ₖ i n x y (c b .fst z)) + + gPathP' : (i : ℕ) + → PathP (λ j → coHomK i → + (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) + →∙ coHomK-ptd (i +' n)) + (g (pt B) i) (g-base.g n i) + gPathP' i = + toPathP + (funExt + λ x → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ y + → (λ i → transportRefl (transportRefl x i ⌣ₖ c (pt B) .fst + (Iso.inv Q-is (transportRefl y i))) i) + ∙ cong (x ⌣ₖ_) + (funExt⁻ c-pt (Iso.inv Q-is y) + ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) + + g-base : (i : ℕ) → isEquiv (g (pt B) i) + g-base i = transport (λ j → isEquiv (gPathP' i (~ j))) (g-base.isEquiv-g n i) + + g-equiv : (b : typ B) (i : ℕ) → isEquiv (g b i) + g-equiv b i = + pRec (isPropIsEquiv _) + (J (λ b _ → isEquiv (g b i)) + (g-base i)) + (conB (pt B) b) + +module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥₂) + (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) where + + is-setQ→K : (b : typ B) → isSet (Q b →∙ coHomK-ptd n) + is-setQ→K b = + sRec (isProp→isOfHLevelSuc 1 isPropIsSet) + (J (λ b _ → isSet (Q b →∙ coHomK-ptd n)) + (subst isSet (cong (_→∙ coHomK-ptd n) + (ua∙ (isoToEquiv (invIso Q-is)) + (cong (Iso.inv Q-is) (sym Q-is-ptd) ∙ Iso.leftInv Q-is _))) + (isOfHLevelRetractFromIso 2 (fst (πS≅ℤ n)) isSetℤ))) + (conB (pt B) b) + + + isConnB : isConnected 3 (typ B) + fst isConnB = ∣ pt B ∣ + snd isConnB = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) + + isPropPath : isProp (∥ pt B ≡ pt B ∥₂) + isPropPath = + isOfHLevelRetractFromIso 1 setTruncTrunc2Iso + (isContr→isProp (isConnectedPath _ isConnB (pt B) (pt B))) + + -- We construct a term in c* : (b : B) → (Q b →∙ Kₙ) + -- Which is equal to the generator of (Sⁿ →∙ Kₙ) over the base point. + c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] + (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) + fst c* b = + sRec (is-setQ→K b) + (J (λ b _ → Q b →∙ coHomK-ptd n) + ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) + , cong (genFunSpace n .fst) Q-is-ptd ∙ genFunSpace n .snd)) + (conB (pt B) b) + snd c* = + funExt λ x → (λ i → sRec (is-setQ→K (pt B)) + (J (λ b _ → Q b →∙ coHomK-ptd n) + ((λ x₁ → genFunSpace n .fst (Iso.fun Q-is x₁)) , + (λ i → genFunSpace n .fst (Q-is-ptd i)) + ∙ genFunSpace n .snd)) + (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i) .fst x) + ∙ (λ i → transportRefl (genFunSpace n .fst + (Iso.fun Q-is (transportRefl x i))) i) + + p-help : {b : fst B} (p : pt B ≡ b) + → (subst (fst ∘ Q) (sym p) (snd (Q b))) ≡ (snd (Q (pt B))) + p-help {b = b} = + J (λ b p → subst (fst ∘ Q) (sym p) (snd (Q b)) ≡ snd (Q (pt B))) + (transportRefl _) + + -- This form of c* will make things somewhat easier to work with later on. + c≡ : (b : fst B) (p : ∥ pt B ≡ b ∥₂) + → (c* .fst b) + ≡ sRec (is-setQ→K b) + (λ pp → (λ qb → + genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) + , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) + ∙ ((λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) p + c≡ b = + sElim (λ _ → isOfHLevelPath 2 (is-setQ→K b) _ _) + (J (λ b a → c* .fst b ≡ + sRec (is-setQ→K b) (λ pp → + (λ qb → + genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) + , + cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ + (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd) ∣ a ∣₂) + ((λ i → sRec (is-setQ→K (pt B)) + (J (λ b₁ _ → Q b₁ →∙ coHomK-ptd n) + ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) , + (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) + (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i)) + ∙ →∙Homogeneous≡ (isHomogeneousKn n) + (transportRefl ((λ x → genFunSpace n .fst (Iso.fun Q-is x))) + ∙ funExt λ x → cong (genFunSpace n .fst ∘ Iso.fun Q-is) + (sym (transportRefl x))))) + +-- We are now almost ready to define the Thom isomorphism. +-- The following module contains the types and functions occuring in it. +module preThom {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where + E : Type _ + E = Σ (typ B) P + + Ẽ : Type _ + Ẽ = Pushout {A = E} (λ _ → tt) fst + + i : (n : ℕ) → (P-base : Iso (P (pt B)) (S₊ n)) → S₊ (suc n) → Ẽ + i zero P-base base = inr (pt B) + i zero P-base (loop j) = (sym (push (pt B , Iso.inv P-base false)) + ∙ push ((pt B) , Iso.inv P-base true)) j + i (suc n) P-base north = inl tt + i (suc n) P-base south = inr (pt B) + i (suc n) P-base (merid a i₁) = push (pt B , Iso.inv P-base a) i₁ + + Q : typ B → Pointed ℓ' + Q x = Susp (P x) , north + + F : Type _ + F = Σ (typ B) λ x → Q x .fst + + F̃ : Type _ + F̃ = Pushout {A = typ B} {C = F} (λ _ → tt) λ b → b , north + + invFE : Ẽ → F̃ + invFE (inl x) = inl tt + invFE (inr x) = inr (x , south) + invFE (push (x , a) i₁) = ((push x) ∙ λ i → inr (x , merid a i)) i₁ + + funFE : F̃ → Ẽ + funFE (inl x) = inl tt + funFE (inr (x , north)) = inl tt + funFE (inr (x , south)) = inr x + funFE (inr (x , merid a i₁)) = push (x , a) i₁ + funFE (push a i₁) = inl tt + + IsoFE : Iso F̃ Ẽ + Iso.fun IsoFE = funFE + Iso.inv IsoFE = invFE + Iso.rightInv IsoFE (inl x) = refl + Iso.rightInv IsoFE (inr x) = refl + Iso.rightInv IsoFE (push (x , a) i₁) k = lem k i₁ + where + lem : cong funFE (((push x) ∙ λ i → inr (x , merid a i))) + ≡ push (x , a) + lem = congFunct funFE (push x) (λ i → inr (x , merid a i)) + ∙ sym (lUnit (push (x , a))) + Iso.leftInv IsoFE (inl x) = refl + Iso.leftInv IsoFE (inr (x , north)) = push x + Iso.leftInv IsoFE (inr (x , south)) = refl + Iso.leftInv IsoFE (inr (x , merid a i)) j = + compPath-filler' (push x) (λ i₁ → inr (x , merid a i₁)) (~ j) i + Iso.leftInv IsoFE (push a i₁) k = push a (i₁ ∧ k) + + + F̃→Q : ∀ {ℓ} {A : Pointed ℓ} → (F̃ , inl tt) →∙ A → (b : typ B) → Q b →∙ A + fst (F̃→Q {A = A , a} (f , p) b) north = f (inr (b , north)) + fst (F̃→Q {A = A , a} (f , p) b) south = f (inr (b , south)) + fst (F̃→Q {A = A , a} (f , p) b) (merid a₁ i₁) = f (inr (b , merid a₁ i₁)) + snd (F̃→Q {A = A , a} (f , p) b) = sym (cong f (push b)) ∙ p + + Q→F̃ : ∀ {ℓ} {A : Pointed ℓ} → ((b : typ B) → Q b →∙ A) → (F̃ , inl tt) →∙ A + fst (Q→F̃ {A = A , a} f) (inl x) = a + fst (Q→F̃ {A = A , a} f) (inr (x , north)) = f x .fst north + fst (Q→F̃ {A = A , a} f) (inr (x , south)) = f x .fst south + fst (Q→F̃ {A = A , a} f) (inr (x , merid a₁ i₁)) = f x .fst (merid a₁ i₁) + fst (Q→F̃ {A = A , a} f) (push a₁ i₁) = snd (f a₁) (~ i₁) + snd (Q→F̃ {A = A , a} f) = refl + + Q→F̃-hom : (n : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd n)) + → Q→F̃ (λ b → f b ++ g b) ≡ (Q→F̃ f ++ Q→F̃ g) + Q→F̃-hom n f g = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { (inl x) → sym (rUnitₖ _ (0ₖ _)) + ; (inr (x , north)) → refl + ; (inr (x , south)) → refl + ; (inr (x , merid a i₁)) → refl + ; (push a j) i → + compPath-filler (cong₂ _+ₖ_ (snd (f a)) (snd (g a))) + (rUnitₖ n (0ₖ n)) (~ i) (~ j)}) + + Q→F̃→Q : ∀ {ℓ} {A : Pointed ℓ} + → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) + → F̃→Q (Q→F̃ x) b .fst q ≡ x b .fst q + Q→F̃→Q f b north = refl + Q→F̃→Q f b south = refl + Q→F̃→Q f b (merid a i₁) = refl + + F̃→Q→F̃ : ∀ {ℓ} {A : Pointed ℓ} (f : F̃ → typ A) (p : _) + → (x : F̃) → fst (Q→F̃ {A = A} (F̃→Q (f , p))) x ≡ f x + F̃→Q→F̃ f p (inl x) = sym p + F̃→Q→F̃ f p (inr (x , north)) = refl + F̃→Q→F̃ f p (inr (x , south)) = refl + F̃→Q→F̃ f p (inr (x , merid a i₁)) = refl + F̃→Q→F̃ f p (push a i₁) j = + compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) + + IsoF̃Q : ∀ {ℓ} {A : Pointed ℓ} + → Iso ((F̃ , inl tt) →∙ A) + ((b : typ B) → Q b →∙ A) + Iso.fun (IsoF̃Q {A = A , a}) = F̃→Q + Iso.inv (IsoF̃Q {A = A , a}) = Q→F̃ + Iso.rightInv (IsoF̃Q {A = A , a}) f = + funExt λ b → ΣPathP (funExt (Q→F̃→Q f b) + , sym (rUnit (snd (f b)))) + Iso.leftInv (IsoF̃Q {A = A , a}) (f , p) = + ΣPathP ((funExt (F̃→Q→F̃ f p)) + , λ i j → p (~ i ∨ j)) + + -- The main result + ι : (k : ℕ) → Iso ((b : typ B) → Q b →∙ coHomK-ptd k) + ((Ẽ , inl tt) →∙ coHomK-ptd k) + ι k = compIso (invIso IsoF̃Q) IsoFE-extend + where + + IsoFE-extend : Iso ((F̃ , inl tt) →∙ coHomK-ptd k) + ((Ẽ , inl tt) →∙ coHomK-ptd k) + Iso.fun IsoFE-extend G = (λ x → G .fst (Iso.inv IsoFE x)) + , (snd G) + Iso.inv IsoFE-extend G = (λ x → G .fst (Iso.fun IsoFE x)) + , (snd G) + Iso.rightInv IsoFE-extend G = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → cong (G .fst) (Iso.rightInv IsoFE x)) + Iso.leftInv IsoFE-extend G = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → cong (G .fst) (Iso.leftInv IsoFE x)) + + ι-hom : (k : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd k)) + → Iso.fun (ι k) (λ b → f b ++ g b) + ≡ (Iso.fun (ι k) f ++ Iso.fun (ι k) g) + ι-hom k f g = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → funExt⁻ (cong fst (Q→F̃-hom _ f g)) (invFE x)) + +-- Packing everything up gives us the Thom Isomorphism between +-- the nᵗʰ cohomology of B and the (n+i)ᵗʰ reduced cohomology of Ẽ, +-- as defined above +module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥) + (n : ℕ) (Q-is : Iso (typ (preThom.Q B P (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (preThom.Q B P (pt B))) ≡ snd (S₊∙ n)) + (c : (b : typ B) → (preThom.Q B P b →∙ coHomK-ptd n)) + (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) + where + + ϕ : (i : ℕ) + → GroupEquiv (coHomGr i (typ B)) + (coHomRedGrDir (i +' n) (preThom.Ẽ B P , inl tt)) + fst (ϕ i) = + isoToEquiv + (setTruncIso + (compIso + (codomainIsoDep + λ b → + equivToIso ((g B (preThom.Q B P) conB n Q-is Q-is-ptd c c-pt b i) + , g-equiv B (preThom.Q B P) conB n Q-is Q-is-ptd c c-pt b i)) + (preThom.ι B P (i +' n)))) + snd (ϕ i) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ F G → + cong ∣_∣₂ (cong (Iso.fun (preThom.ι B P (i +' n))) + (funExt (λ a → + g-hom B (preThom.Q B P) + conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) + ∙ preThom.ι-hom B P (i +' n) _ _) + ∙ addAgree (i +' n) _ _) + +-- We finally get the Gysin sequence +module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥₂) + (n : ℕ) (Q-is : Iso (typ (preThom.Q B P (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (preThom.Q B P (pt B))) ≡ snd (S₊∙ n)) + where + + 0-connB : (x y : typ B) → ∥ x ≡ y ∥ + 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) + + c = (c* B (preThom.Q B P) conB n Q-is Q-is-ptd .fst) + c-ptd = (c* B (preThom.Q B P) conB n Q-is Q-is-ptd .snd) + cEq = c≡ B (preThom.Q B P) conB n Q-is Q-is-ptd + + module w = Thom B P 0-connB n Q-is Q-is-ptd c c-ptd + + ϕ = w.ϕ + + E' = preThom.E B P + + E'̃ = preThom.Ẽ B P + + -- The generator of coHom n (typ B) + e : coHom n (typ B) + e = ∣ (λ b → c b .fst south) ∣₂ + + -- The maps of interest are ⌣, p, E-susp and j*. In reality, we are interested + -- in the composition of ϕ and j* (which is just the cup product), + -- but it's easier to first give an exact sequence involving p, E-susp and j* + ⌣-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) + fst (⌣-hom i) x = x ⌣ e + snd (⌣-hom i) = + makeIsGroupHom λ f g → rightDistr-⌣ _ _ f g e + + p : E' → typ B + p = fst + + p-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr i E') + p-hom i = coHomMorph _ p + + E-susp : (i : ℕ) → + GroupHom (coHomGr i E') (coHomRedGrDir (suc i) (E'̃ , inl tt)) + fst (E-susp i) = + sMap λ f → (λ { (inl x) → 0ₖ _ + ; (inr x) → 0ₖ _ + ; (push a j) → Kn→ΩKn+1 _ (f a) j}) , refl + snd (E-susp zero) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn 1) + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a j) k → + (Kn→ΩKn+1-hom zero (f a) (g a) + ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) + snd (E-susp (suc i)) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a j) k → (Kn→ΩKn+1-hom _ (f a) (g a) + ∙ ∙≡+₂ _ (Kn→ΩKn+1 _ (f a)) + (Kn→ΩKn+1 _ (g a))) k j}))) + + module cofibSeq where + j* : (i : ℕ) → + GroupHom (coHomRedGrDir i (E'̃ , (inl tt))) (coHomGr i (typ B)) + fst (j* i) = sMap λ f → λ x → fst f (inr x) + snd (j* zero) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + snd (j* (suc zero)) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + snd (j* (suc (suc i₁))) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + + Im-j⊂Ker-p : (i : ℕ) (x : _) → isInIm (j* i) x → isInKer (p-hom i) x + Im-j⊂Ker-p i = + sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f → pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g P → subst (isInKer (p-hom i)) P + (cong ∣_∣₂ (funExt λ x → + cong (g .fst) (sym (push x)) ∙ g .snd)))) + + Ker-p⊂Im-j : (i : ℕ) (x : _) → isInKer (p-hom i) x → isInIm (j* i) x + Ker-p⊂Im-j i = + sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + λ f ker → pRec squash + (λ id → ∣ ∣ (λ { (inl x) → 0ₖ _ + ; (inr x) → f x + ; (push a i₁) → funExt⁻ id a (~ i₁)}) , refl ∣₂ + , refl ∣) + (Iso.fun PathIdTrunc₀Iso ker) + + Im-p⊂Ker-Susp : (i : ℕ) (x : _) + → isInIm (p-hom i) x → isInKer (E-susp i) x + Im-p⊂Ker-Susp i = + sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f → pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g y → subst (isInKer (E-susp i)) y + (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { (inl x) → refl + ; (inr x) → sym (Kn→ΩKn+1 _ (g x)) + ; (push a j) k → Kn→ΩKn+1 i (g (fst a)) (~ k ∧ j)}))))) + + Ker-Susp⊂Im-p : (i : ℕ) (x : _) + → isInKer (E-susp i) x → isInIm (p-hom i) x + Ker-Susp⊂Im-p i = + sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + λ f ker → pRec squash + (λ id → ∣ ∣ (λ x → ΩKn+1→Kn i (sym (funExt⁻ (cong fst id) (inr x)))) ∣₂ + , cong ∣_∣₂ (funExt (λ { (a , b) → + cong (ΩKn+1→Kn i) + (lUnit _ + ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) + (sym (flipSquare (cong snd id)) + ∙∙ (PathP→compPathR + (cong (funExt⁻ (cong fst id)) (push (a , b)))) + ∙∙ assoc _ _ _ + ∙ sym (rUnit _)) + ∙ (sym (assoc _ _ _) + ∙∙ cong (Kn→ΩKn+1 i (f (a , b)) ∙_) (rCancel _) + ∙∙ sym (rUnit _))) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (a , b))})) ∣) + (Iso.fun PathIdTrunc₀Iso ker) + + Im-Susp⊂Ker-j : (i : ℕ) (x : _) + → isInIm (E-susp i) x → isInKer (cofibSeq.j* (suc i)) x + Im-Susp⊂Ker-j i = + sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g → pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f id → pRec (squash₂ _ _) + (λ P → subst (isInKer (cofibSeq.j* (suc i))) (cong ∣_∣₂ P) + (cong ∣_∣₂ refl)) + (Iso.fun PathIdTrunc₀Iso id))) + + Ker-j⊂Im-Susp : (i : ℕ) (x : _) + → isInKer (cofibSeq.j* (suc i)) x → isInIm (E-susp i) x + Ker-j⊂Im-Susp i = + sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + λ f ker + → pRec squash + (λ p → ∣ ∣ (λ x → ΩKn+1→Kn _ (sym (snd f) + ∙∙ cong (fst f) (push x) + ∙∙ funExt⁻ p (fst x))) ∣₂ + , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) + (funExt (λ { (inl x) → sym (snd f) + ; (inr x) → sym (funExt⁻ p x) + ; (push a j) k → h3 f p a k j}))) ∣) + (Iso.fun PathIdTrunc₀Iso ker) + where + h3 : (f : (E'̃ , inl tt) →∙ coHomK-ptd (suc i)) + → (p : (fst f ∘ inr) ≡ (λ _ → 0ₖ (suc i))) + → (a : preThom.E B P) + → PathP (λ i → snd f (~ i) ≡ p (~ i) (fst a)) + (Kn→ΩKn+1 i (ΩKn+1→Kn i (sym (snd f) + ∙∙ cong (fst f) (push a) + ∙∙ funExt⁻ p (fst a)))) + (cong (fst f) (push a)) + h3 f p a = + Iso.rightInv (Iso-Kn-ΩKn+1 i) _ + ◁ λ i j → + doubleCompPath-filler (sym (snd f)) (cong (fst f) (push a)) + (funExt⁻ p (fst a)) (~ i) j + + -- We compose ϕ and j*. The above exact sequence will induce another one for + -- this composition. In fact, this group hom is definitionally equal to + -- (λ x → x ⌣ e) modulo truncation elimination and snd fields. + ϕ∘j : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) + ϕ∘j i = compGroupHom (fst (fst (ϕ i)) , snd (ϕ i)) (cofibSeq.j* (i +' n)) + + private + +'-suc : (i n : ℕ) → (suc i +' n) ≡ suc (i +' n) + +'-suc zero zero = refl + +'-suc (suc i₁) zero = refl + +'-suc zero (suc n) = refl + +'-suc (suc i₁) (suc n) = refl + + Path→GroupPath : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') + (p : n ≡ m) + → GroupEquiv (G n) H + → GroupEquiv (G m) H + Path→GroupPath {n = n} G H = + J (λ m _ → GroupEquiv (G n) H → GroupEquiv (G m) H) + λ p → p + + h-ret : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') + → (e : GroupEquiv (G n) H) + → (p : n ≡ m) + → (x : G m .fst) + → invEq (fst e) (fst (fst (Path→GroupPath G H p e)) x) + ≡ subst (λ x → G x .fst) (sym p) x + h-ret G H e = + J (λ m p → ((x : G m .fst) + → invEq (fst e) (fst (fst (Path→GroupPath G H p e)) x) + ≡ subst (λ x → G x .fst) (sym p) x)) + λ x → cong (invEq (fst e) ) + (λ i → transportRefl (transportRefl (fst (fst e) + (transportRefl (transportRefl x i) i)) i) i) + ∙∙ retEq (fst e) x + ∙∙ sym (transportRefl _) + + thomIso' : (i : ℕ) → GroupEquiv (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) + (coHomGr (suc i) (typ B)) + thomIso' i = (Path→GroupPath + (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) + (invGroupEquiv (ϕ (suc i)))) + + ϕ' : (i : ℕ) → GroupHom (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) + (coHomGr (suc i) (typ B)) + ϕ' i = fst (fst (thomIso' i)) , snd (thomIso' i) + + susp∘ϕ : (i : ℕ) → GroupHom (coHomGr (i +' n) E') (coHomGr (suc i) (typ B)) + susp∘ϕ i = compGroupHom (E-susp (i +' n)) (ϕ' i) + + Im-ϕ∘j⊂Ker-p : (i : ℕ) (x : _) → isInIm (ϕ∘j i) x → isInKer (p-hom _) x + Im-ϕ∘j⊂Ker-p i x p = + cofibSeq.Im-j⊂Ker-p _ x + (pRec squash (uncurry (λ f p → ∣ fst (fst (ϕ _)) f , p ∣)) p) + + Ker-p⊂Im-ϕ∘j : (i : ℕ) (x : _) → isInKer (p-hom _) x → isInIm (ϕ∘j i) x + Ker-p⊂Im-ϕ∘j i x p = + pRec squash + (uncurry (λ f p → + ∣ (invEq (fst (ϕ _)) f) + , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣)) + (cofibSeq.Ker-p⊂Im-j _ x p) + + Im-p⊂KerSusp∘ϕ : (i : ℕ) (x : _) + → isInIm (p-hom _) x → isInKer (susp∘ϕ i) x + Im-p⊂KerSusp∘ϕ i x p = + cong (fst (ϕ' _)) (Im-p⊂Ker-Susp _ x p) ∙ IsGroupHom.pres1 (snd (ϕ' _)) + + KerSusp∘ϕ⊂Im-p : (i : ℕ) (x : _) + → isInKer (susp∘ϕ i) x → isInIm (p-hom _) x + KerSusp∘ϕ⊂Im-p i x p = + Ker-Susp⊂Im-p _ x (sym (retEq (fst (thomIso' _)) _) + ∙ (cong (invEq (fst (thomIso' _))) p + ∙ IsGroupHom.pres1 (snd (invGroupEquiv (thomIso' _))))) + + Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) + → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x + Im-Susp∘ϕ⊂Ker-ϕ∘j i x = + pRec (squash₂ _ _) + (uncurry λ f → J (λ x p → isInKer (ϕ∘j (suc i)) x) + ((λ i → fst (cofibSeq.j* _) (fst (fst (ϕ _)) (fst (ϕ' _) (fst (E-susp _) f)))) + ∙∙ cong (fst (cofibSeq.j* _)) + ((h-ret (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ + (invGroupEquiv (ϕ (suc i))) (+'-suc i n)) (fst (E-susp _) f)) + ∙∙ (natTranspLem _ (λ n → fst (cofibSeq.j* n)) (sym (+'-suc i n)) + ∙ cong (subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n))) + (Im-Susp⊂Ker-j _ (fst (E-susp (i +' n)) f) ∣ f , refl ∣) + ∙ tLem i n))) + where + tLem : (i n : ℕ) → subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n)) + (0ₕ _) ≡ 0ₕ _ + tLem zero zero = refl + tLem zero (suc n) = refl + tLem (suc i₁) zero = refl + tLem (suc i₁) (suc n) = refl + + Ker-ϕ∘j⊂Im-Susp∘ϕ : (i : ℕ) (x : _) + → isInKer (ϕ∘j (suc i)) x → isInIm (susp∘ϕ i) x + Ker-ϕ∘j⊂Im-Susp∘ϕ i x p = + pRec squash + (uncurry (λ f p → ∣ f , cong (fst (fst (thomIso' i))) p + ∙ secEq (fst (thomIso' _)) x ∣)) + (Ker-j⊂Im-Susp _ (invEq (fst (thomIso' _)) x) + ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem₁ + ∙ natTranspLem _ (λ n → cofibSeq.j* n .fst) (+'-suc i n)) + ∙∙ cong (transport (λ j → (coHomGr (+'-suc i n j) (typ B) .fst))) p + ∙∙ lem₂ i n)) + where + lem₁ : invEq (fst (thomIso' i)) x + ≡ transport (λ j → coHomRed (+'-suc i n j) + (E'̃ , inl tt)) (fst (fst (ϕ _)) x) + lem₁ = cong (transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt))) + (transportRefl _ ∙ cong (fst (fst (ϕ _))) + λ i → transportRefl (transportRefl x i) i) + + lem₂ : (i n : ℕ) + → transport (λ j → coHomGr (+'-suc i n j) (typ B) .fst) + (GroupStr.1g (coHomGr (suc i +' n) (typ B) .snd)) ≡ 0ₕ _ + lem₂ zero zero = refl + lem₂ zero (suc n) = refl + lem₂ (suc i₁) zero = refl + lem₂ (suc i₁) (suc n) = refl + + -- Finally, we have that ϕ∘j is just the cup product, and we have arrived + -- at an exact sequence involving it. + ϕ∘j≡ : (i : ℕ) → ϕ∘j i ≡ ⌣-hom i + ϕ∘j≡ i = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ _ → refl)) + + -- We can now restate the previous resluts for (λ x → x ⌣ e) + Im-⌣e⊂Ker-p : (i : ℕ) (x : _) + → isInIm (⌣-hom i) x → isInKer (p-hom _) x + Im-⌣e⊂Ker-p i x p = + Im-ϕ∘j⊂Ker-p i x (subst (λ p → isInIm p x) (sym (ϕ∘j≡ i)) p) + + Ker-p⊂Im-⌣e : (i : ℕ) (x : _) + → isInKer (p-hom _) x → isInIm (⌣-hom i) x + Ker-p⊂Im-⌣e i x p = + subst (λ p → isInIm p x) (ϕ∘j≡ i) (Ker-p⊂Im-ϕ∘j i x p) + + Im-Susp∘ϕ⊂Ker-⌣e : (i : ℕ) (x : _) + → isInIm (susp∘ϕ i) x → isInKer (⌣-hom (suc i)) x + Im-Susp∘ϕ⊂Ker-⌣e i x p = + subst (λ p → isInKer p x) (ϕ∘j≡ (suc i)) (Im-Susp∘ϕ⊂Ker-ϕ∘j i x p) + + Ker-⌣e⊂Im-Susp∘ϕ : (i : ℕ) (x : _) + → isInKer (⌣-hom (suc i)) x → isInIm (susp∘ϕ i) x + Ker-⌣e⊂Im-Susp∘ϕ i x p = + Ker-ϕ∘j⊂Im-Susp∘ϕ i x (subst (λ p → isInKer p x) (sym (ϕ∘j≡ (suc i))) p) diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda new file mode 100644 index 0000000000..22631d8622 --- /dev/null +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -0,0 +1,250 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.MayerVietorisUnreduced where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Sigma +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) +open import Cubical.Data.Nat +open import Cubical.Algebra.Group +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) + +open IsGroupHom + +module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where + -- Proof from Brunerie 2016. + -- We first define the three morphisms involved: i, Δ and d. + + private + i* : (n : ℕ) → coHom n (Pushout f g) → coHom n A × coHom n B + i* n = sRec (isSet× isSetSetTrunc isSetSetTrunc) λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ + + iIsHom : (n : ℕ) → IsGroupHom (coHomGr n (Pushout f g) .snd) (i* n) (×coHomGr n A B .snd) + iIsHom n = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ → refl) + + i : (n : ℕ) → GroupHom (coHomGr n (Pushout f g)) (×coHomGr n A B) + fst (i n) = i* n + snd (i n) = iIsHom n + + private + distrLem : (n : ℕ) (x y z w : coHomK n) → (x +[ n ]ₖ y) -[ n ]ₖ (z +[ n ]ₖ w) ≡ (x -[ n ]ₖ z) +[ n ]ₖ (y -[ n ]ₖ w) + distrLem n x y z w = cong (λ z → (x +[ n ]ₖ y) +[ n ]ₖ z) (-distrₖ n z w) + ∙∙ sym (assocₖ n x y ((-[ n ]ₖ z) +[ n ]ₖ (-[ n ]ₖ w))) + ∙∙ cong (λ y → x +[ n ]ₖ y) (commₖ n y ((-[ n ]ₖ z) +[ n ]ₖ (-[ n ]ₖ w)) ∙ sym (assocₖ n _ _ _)) + ∙∙ assocₖ n _ _ _ + ∙∙ cong (λ y → (x -[ n ]ₖ z) +[ n ]ₖ y) (commₖ n (-[ n ]ₖ w) y) + + Δ' : (n : ℕ) → coHom n A × coHom n B → coHom n C + Δ' n (α , β) = coHomFun n f α -[ n ]ₕ coHomFun n g β + + Δ'-isMorph : (n : ℕ) → IsGroupHom (×coHomGr n A B .snd) (Δ' n) (coHomGr n C .snd) + Δ'-isMorph n = + makeIsGroupHom + (prodElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _ ) + λ f' x1 g' x2 i → ∣ (λ x → distrLem n (f' (f x)) (g' (f x)) (x1 (g x)) (x2 (g x)) i) ∣₂) + + Δ : (n : ℕ) → GroupHom (×coHomGr n A B) (coHomGr n C) + fst (Δ n) = Δ' n + snd (Δ n) = Δ'-isMorph n + + d-pre : (n : ℕ) → (C → coHomK n) → Pushout f g → coHomK (suc n) + d-pre n γ (inl x) = 0ₖ (suc n) + d-pre n γ (inr x) = 0ₖ (suc n) + d-pre n γ (push a i) = Kn→ΩKn+1 n (γ a) i + + dHomHelper : (n : ℕ) (h l : C → coHomK n) (x : Pushout f g) + → d-pre n (λ x → h x +[ n ]ₖ l x) x ≡ d-pre n h x +[ suc n ]ₖ d-pre n l x + dHomHelper n h l (inl x) = sym (rUnitₖ (suc n) (0ₖ (suc n))) + dHomHelper n h l (inr x) = sym (lUnitₖ (suc n) (0ₖ (suc n))) + dHomHelper n h l (push a i) j = + hcomp (λ k → λ { (i = i0) → rUnitₖ (suc n) (0ₖ (suc n)) (~ j) + ; (i = i1) → lUnitₖ (suc n) (0ₖ (suc n)) (~ j) + ; (j = i0) → Kn→ΩKn+1-hom n (h a) (l a) (~ k) i + ; (j = i1) → cong₂Funct (λ x y → x +[ (suc n) ]ₖ y) (Kn→ΩKn+1 n (h a)) (Kn→ΩKn+1 n (l a)) (~ k) i }) + (hcomp (λ k → λ { (i = i0) → rUnitₖ (suc n) (0ₖ (suc n)) (~ j) + ; (i = i1) → lUnitₖ (suc n) (Kn→ΩKn+1 n (l a) k) (~ j)}) + (hcomp (λ k → λ { (i = i0) → rUnitₖ (suc n) (0ₖ (suc n)) (~ j) + ; (i = i1) → lUnitₖ≡rUnitₖ (suc n) (~ k) (~ j) + ; (j = i0) → Kn→ΩKn+1 n (h a) i + ; (j = i1) → (Kn→ΩKn+1 n (h a) i) +[ (suc n) ]ₖ coHom-pt (suc n)}) + (rUnitₖ (suc n) (Kn→ΩKn+1 n (h a) i) (~ j)))) + + dIsHom : (n : ℕ) → IsGroupHom (coHomGr n C .snd) (sRec isSetSetTrunc λ a → ∣ d-pre n a ∣₂) (coHomGr (suc n) (Pushout f g) .snd) + dIsHom n = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g i → ∣ funExt (λ x → dHomHelper n f g x) i ∣₂) + + d : (n : ℕ) → GroupHom (coHomGr n C) (coHomGr (suc n) (Pushout f g)) + fst (d n) = sRec isSetSetTrunc λ a → ∣ d-pre n a ∣₂ + snd (d n) = dIsHom n + + -- The long exact sequence + Im-d⊂Ker-i : (n : ℕ) (x : ⟨ (coHomGr (suc n) (Pushout f g)) ⟩) + → isInIm (d n) x + → isInKer (i (suc n)) x + Im-d⊂Ker-i n = sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ a → pRec (isOfHLevelPath' 1 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ δ b i → sRec (isSet× isSetSetTrunc isSetSetTrunc) + (λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ ) (b (~ i))) + + + Ker-i⊂Im-d : (n : ℕ) (x : ⟨ coHomGr (suc n) (Pushout f g) ⟩) + → isInKer (i (suc n)) x + → isInIm (d n) x + Ker-i⊂Im-d n = + sElim (λ _ → isSetΠ λ _ → isProp→isSet isPropPropTrunc) + λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ (suc n)} + (isProp→ isPropPropTrunc) + (λ p1 → pRec isPropPropTrunc λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn n (sym (cong (λ F → F (f c)) p1) + ∙∙ cong a (push c) + ∙∙ cong (λ F → F (g c)) p2)) ∣₂ + , cong ∣_∣₂ (funExt (λ δ → helper a p1 p2 δ)) ∣₁) + (Iso.fun PathIdTrunc₀Iso (cong fst p)) + (Iso.fun PathIdTrunc₀Iso (cong snd p)) + + where + helper : (F : (Pushout f g) → coHomK (suc n)) + (p1 : Path (_ → coHomK (suc n)) (λ a₁ → F (inl a₁)) (λ _ → coHom-pt (suc n))) + (p2 : Path (_ → coHomK (suc n)) (λ a₁ → F (inr a₁)) (λ _ → coHom-pt (suc n))) + → (δ : Pushout f g) + → d-pre n (λ c → ΩKn+1→Kn n ((λ i₁ → p1 (~ i₁) (f c)) + ∙∙ cong F (push c) + ∙∙ cong (λ F → F (g c)) p2)) δ + ≡ F δ + helper F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + helper F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso-Kn-ΩKn+1 n) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (doubleCompPath-filler (sym (cong (λ F → F (f a)) p1)) (cong F (push a)) (cong (λ F → F (g a)) p2) (~ j) i) + + Im-i⊂Ker-Δ : (n : ℕ) (x : ⟨ ×coHomGr n A B ⟩) + → isInIm (i n) x + → isInKer (Δ n) x + Im-i⊂Ker-Δ n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInIm (i n) (Fa , Fb) + → isInKer (Δ n) (Fa , Fb)} + (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ Fa → sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fb → pRec (isSetSetTrunc _ _) + (sigmaElim (λ x → isProp→isSet (isSetSetTrunc _ _)) + λ Fd p → helper n Fa Fb Fd p)) + Fa + Fb + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : (Pushout f g) → coHomK n) + → (fst (i n) ∣ Fd ∣₂ ≡ (∣ Fa ∣₂ , ∣ Fb ∣₂)) + → (fst (Δ n)) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ n + helper n Fa Fb Fd p = cong (fst (Δ n)) (sym p) + ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ -[ n ]ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂ ) + ∙∙ rCancelₕ n ∣ (λ x → Fd (inl (f x))) ∣₂ + + Ker-Δ⊂Im-i : (n : ℕ) (a : ⟨ ×coHomGr n A B ⟩) + → isInKer (Δ n) a + → isInIm (i n) a + Ker-Δ⊂Im-i n = prodElim (λ _ → isSetΠ (λ _ → isProp→isSet isPropPropTrunc)) + (λ Fa Fb p → pRec isPropPropTrunc + (λ q → ∣ ∣ helpFun Fa Fb q ∣₂ , refl ∣₁) + (helper Fa Fb p)) + where + helper : (Fa : A → coHomK n) (Fb : B → coHomK n) + → fst (Δ n) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ n + → ∥ Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c)) ∥₁ + helper Fa Fb p = Iso.fun PathIdTrunc₀Iso + (sym (cong ∣_∣₂ (funExt (λ x → sym (assocₖ n _ _ _) + ∙∙ cong (λ y → Fa (f x) +[ n ]ₖ y) (lCancelₖ n (Fb (g x))) + ∙∙ rUnitₖ n (Fa (f x))))) + ∙∙ cong (λ x → x +[ n ]ₕ ∣ (λ c → Fb (g c)) ∣₂) p + ∙∙ lUnitₕ n _) + + helpFun : (Fa : A → coHomK n) (Fb : B → coHomK n) + → ((λ c → Fa (f c)) ≡ (λ c → Fb (g c))) + → Pushout f g → coHomK n + helpFun Fa Fb p (inl x) = Fa x + helpFun Fa Fb p (inr x) = Fb x + helpFun Fa Fb p (push a i) = p i a + + private + distrHelper : (n : ℕ) (p q : _) + → ΩKn+1→Kn n p +[ n ]ₖ (-[ n ]ₖ ΩKn+1→Kn n q) ≡ ΩKn+1→Kn n (p ∙ sym q) + distrHelper n p q = cong (λ x → ΩKn+1→Kn n p +[ n ]ₖ x) helper ∙ sym (ΩKn+1→Kn-hom n _ _) + where + helper : -[ n ]ₖ ΩKn+1→Kn n q ≡ ΩKn+1→Kn n (sym q) + helper = + sym (rUnitₖ n _) + ∙∙ cong (λ x → (-[ n ]ₖ (ΩKn+1→Kn n q)) +[ n ]ₖ x) (sym helper2) + ∙∙ (assocₖ n _ _ _ ∙∙ cong (λ x → x +[ n ]ₖ (ΩKn+1→Kn n (sym q))) (lCancelₖ n _) ∙∙ lUnitₖ n _) + where + helper2 : ΩKn+1→Kn n q +[ n ]ₖ (ΩKn+1→Kn n (sym q)) ≡ coHom-pt n + helper2 = sym (ΩKn+1→Kn-hom n q (sym q)) ∙∙ cong (ΩKn+1→Kn n) (rCancel q) ∙∙ ΩKn+1→Kn-refl n + + Ker-d⊂Im-Δ : (n : ℕ) (a : coHom n C) + → isInKer (d n) a + → isInIm (Δ n) a + Ker-d⊂Im-Δ n = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + λ Fc p → pRec isPropPropTrunc (λ p → ∣ (∣ (λ a → ΩKn+1→Kn n (cong (λ f → f (inl a)) p)) ∣₂ , + ∣ (λ b → ΩKn+1→Kn n (cong (λ f → f (inr b)) p)) ∣₂) , + Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 Fc p c) ∣₁ ∣₁) + (Iso.fun (PathIdTrunc₀Iso) p) + + where + + helper2 : (Fc : C → coHomK n) + (p : d-pre n Fc ≡ (λ _ → coHom-pt (suc n))) (c : C) + → ΩKn+1→Kn n (λ i₁ → p i₁ (inl (f c))) -[ n ]ₖ (ΩKn+1→Kn n (λ i₁ → p i₁ (inr (g c)))) ≡ Fc c + helper2 Fc p c = distrHelper n _ _ ∙∙ cong (ΩKn+1→Kn n) helper3 ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) (Fc c) + where + helper3 : (λ i₁ → p i₁ (inl (f c))) ∙ sym (λ i₁ → p i₁ (inr (g c))) ≡ Kn→ΩKn+1 n (Fc c) + helper3 = cong ((λ i₁ → p i₁ (inl (f c))) ∙_) (lUnit _) ∙ sym (PathP→compPathR (cong (λ f → cong f (push c)) p)) + + Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) + → isInIm (Δ n) a + → isInKer (d n) a + Im-Δ⊂Ker-d n = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fc → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + (sigmaProdElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fa Fb p → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + (λ q → ((λ i → fst (d n) ∣ (q (~ i)) ∣₂) ∙ dΔ-Id n Fa Fb)) + (Iso.fun (PathIdTrunc₀Iso) p)) + + where + d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : (Pushout f g)) + → d-pre n (Fa ∘ f) d ≡ 0ₖ (suc n) + d-preLeftId n Fa (inl x) = Kn→ΩKn+1 n (Fa x) + d-preLeftId n Fa (inr x) = refl + d-preLeftId n Fa (push a i) j = Kn→ΩKn+1 n (Fa (f a)) (j ∨ i) + + d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : (Pushout f g)) + → d-pre n (Fb ∘ g) d ≡ 0ₖ (suc n) + d-preRightId n Fb (inl x) = refl + d-preRightId n Fb (inr x) = sym (Kn→ΩKn+1 n (Fb x)) + d-preRightId n Fb (push a i) j = Kn→ΩKn+1 n (Fb (g a)) (~ j ∧ i) + + dΔ-Id : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → fst (d n) (fst (Δ n) (∣ Fa ∣₂ , ∣ Fb ∣₂)) ≡ 0ₕ (suc n) + dΔ-Id n Fa Fb = -distrLemma n (suc n) (d n) ∣ Fa ∘ f ∣₂ ∣ Fb ∘ g ∣₂ + ∙∙ (λ i → ∣ (λ x → d-preLeftId n Fa x i) ∣₂ -[ (suc n) ]ₕ ∣ (λ x → d-preRightId n Fb x i) ∣₂) + ∙∙ rCancelₕ (suc n) (0ₕ (suc n)) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda new file mode 100644 index 0000000000..60918431a2 --- /dev/null +++ b/Cubical/ZCohomology/Properties.agda @@ -0,0 +1,676 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Properties where + +{- +This module contains: +1. direct proofs of connectedness of Kn and ΩKn +2. Induction principles for cohomology groups of pointed types +3. Equivalence between cohomology of A and reduced cohomology of (A + 1) +4. Equivalence between cohomology and reduced cohomology for dimension ≥ 1 +5. Encode-decode proof of Kₙ ≃ ΩKₙ₊₁ and proofs that this equivalence + and its inverse are morphisms +6. A proof of coHomGr ≅ coHomGrΩ +7. A locked (non-reducing) version of Kₙ ≃ ΩKₙ₊₁ +8. Some HLevel lemmas used for the cup product +-} + + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure + +open import Cubical.HITs.S1 hiding (encode ; decode ; _·_) +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) +open import Cubical.Foundations.Univalence +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; isSetSetTrunc to §) +open import Cubical.Data.Int renaming (_+_ to _ℤ+_) hiding (-_) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; map2 to trMap2; rec to trRec ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Algebra.Group hiding (Unit ; ℤ) +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Data.Sum.Base hiding (map) +open import Cubical.Functions.Morphism +open import Cubical.Data.Sigma + +open Iso renaming (inv to inv') + +private + variable + ℓ ℓ' : Level + +------------------- Connectedness --------------------- +is2ConnectedKn : (n : ℕ) → isConnected 2 (coHomK (suc n)) +is2ConnectedKn zero = ∣ ∣ base ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) + (toPropElim (λ _ → isOfHLevelTrunc 2 _ _) refl)) +is2ConnectedKn (suc n) = ∣ ∣ north ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isOfHLevelTrunc 2 _ _)) + (suspToPropElim (ptSn (suc n)) (λ _ → isOfHLevelTrunc 2 _ _) refl)) + +isConnectedKn : (n : ℕ) → isConnected (2 + n) (coHomK (suc n)) +isConnectedKn n = isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (2 + n) 1)) (sphereConnected (suc n)) + +-- direct proof of connectedness of ΩKₙ₊₁ not relying on the equivalence ∥ a ≡ b ∥ₙ ≃ (∣ a ∣ₙ₊₁ ≡ ∣ b ∣ₙ₊₁) +isConnectedPathKn : (n : ℕ) (x y : (coHomK (suc n))) → isConnected (suc n) (x ≡ y) +isConnectedPathKn n = + trElim (λ _ → isProp→isOfHLevelSuc (2 + n) (isPropΠ λ _ → isPropIsContr)) + (sphereElim _ (λ _ → isProp→isOfHLevelSuc n (isPropΠ λ _ → isPropIsContr)) + λ y → isContrRetractOfConstFun + {B = (hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)))} ∣ refl ∣ + (fun⁻ n y + , trElim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (J (λ y p → fun⁻ n y _ ≡ _) (funExt⁻ (fun⁻Id n) ∣ refl ∣)))) + where + fun⁻ : (n : ℕ) → (y : coHomK (suc n)) → + hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)) + → hLevelTrunc (suc n) (∣ ptSn (suc n) ∣ ≡ y) + fun⁻ n = + trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelSuc (2 + n) (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n)))) + (sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) λ _ → ∣ refl ∣) + + fun⁻Id : (n : ℕ) → fun⁻ n ∣ ptSn (suc n) ∣ ≡ λ _ → ∣ refl ∣ + fun⁻Id zero = refl + fun⁻Id (suc n) = refl + +------------------- +-- Induction principles for cohomology groups (n ≥ 1) +-- If we want to show a proposition about some x : Hⁿ(A), it suffices to show it under the +-- assumption that x = ∣ f ∣₂ for some f : A → Kₙ and that f is pointed + +coHomPointedElim : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (x : coHom (suc n) A) → B x +coHomPointedElim {ℓ' = ℓ'} {A = A} n a isprop indp = + sElim (λ _ → isOfHLevelSuc 1 (isprop _)) + λ f → helper n isprop indp f (f a) refl + where + helper : (n : ℕ) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (f : A → coHomK (suc n)) + → (x : coHomK (suc n)) + → f a ≡ x → B ∣ f ∣₂ + -- pattern matching a bit extra to avoid isOfHLevelPlus' + helper zero isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ λ _ → isprop _)) + (toPropElim (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc zero) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ λ _ → isprop _)) + (suspToPropElim base (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc zero)) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc (suc n))) isprop ind f = + trElim (λ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + + +coHomPointedElim2 : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (x y : coHom (suc n) A) → B x y +coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = sElim2 (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) + λ f g → helper n a isprop indp f g (f a) (g a) refl refl + where + helper : (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (f g : A → coHomK (suc n)) + → (x y : coHomK (suc n)) + → f a ≡ x → g a ≡ y + → B ∣ f ∣₂ ∣ g ∣₂ + helper zero a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ2 λ _ _ → isprop _ _)) + (toPropElim2 (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc zero) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 base (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc zero)) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc (suc n))) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + +coHomK-elim : ∀ {ℓ} (n : ℕ) {B : coHomK (suc n) → Type ℓ} + → ((x : _) → isOfHLevel (suc n) (B x)) + → B (0ₖ (suc n)) + → (x : _) → B x +coHomK-elim n {B = B } hlev b = + trElim (λ _ → isOfHLevelPlus {n = (suc n)} 2 (hlev _)) + (sphereElim _ (hlev ∘ ∣_∣) b) + +{- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} +coHomRed+1Equiv : (n : ℕ) → + (A : Type ℓ) → + (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHomRed+1Equiv zero A i = ∥ helpLemma {C = (ℤ , pos 0)} i ∥₂ + module coHomRed+1 where + helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →∙ C))) + helpLemma {C = C} = isoToPath (iso map1 + map2 + (λ b → linvPf b) + (λ _ → refl)) + where + map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) + map1 f = map1' , refl + module helpmap where + map1' : A ⊎ Unit → fst C + map1' (inl x) = f x + map1' (inr x) = pt C + + map2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) + map2 (g , pf) x = g (inl x) + + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) + where + helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x + helper (inl x) = refl + helper (inr tt) = sym snd +coHomRed+1Equiv (suc zero) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK 1 , ∣ base ∣)} i ∥₂ +coHomRed+1Equiv (suc (suc n)) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK (2 + n) , ∣ north ∣)} i ∥₂ + +Iso-coHom-coHomRed : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (coHomRed (suc n) A) (coHom (suc n) (typ A)) +fun (Iso-coHom-coHomRed {A = A , a} n) = map fst +inv' (Iso-coHom-coHomRed {A = A , a} n) = map λ f → (λ x → f x -ₖ f a) , rCancelₖ _ _ +rightInv (Iso-coHom-coHomRed {A = A , a} n) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ f → trRec (isProp→isOfHLevelSuc _ (§ _ _)) + (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → f x +ₖ y) (cong -ₖ_ p ∙ -0ₖ) ∙ rUnitₖ _ (f x))) + (Iso.fun (PathIdTruncIso (suc n)) (isContr→isProp (isConnectedKn n) ∣ f a ∣ ∣ 0ₖ _ ∣)) +leftInv (Iso-coHom-coHomRed {A = A , a} n) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP (((funExt λ x → (cong (λ y → f x -ₖ y) p + ∙∙ cong (λ y → f x +ₖ y) -0ₖ + ∙∙ rUnitₖ _ (f x)) ∙ refl)) + , helper n (f a) (sym p)))} + where + path : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) → _ + path n x p = (cong (λ y → x -ₖ y) (sym p) ∙∙ cong (λ y → x +ₖ y) -0ₖ ∙∙ rUnitₖ _ x) ∙ refl + + helper : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) + → PathP (λ i → path n x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p) + helper zero x = + J (λ x p → PathP (λ i → path 0 x p i ≡ 0ₖ _) + (rCancelₖ _ x) (sym p)) + λ i j → rUnit (rUnit (λ _ → 0ₖ 1) (~ j)) (~ j) i + helper (suc n) x = + J (λ x p → PathP (λ i → path (suc n) x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p)) + λ i j → rCancelₖ (suc (suc n)) (0ₖ (suc (suc n))) (~ i ∧ ~ j) + ++∙≡+ : (n : ℕ) {A : Pointed ℓ} (x y : coHomRed (suc n) A) + → Iso.fun (Iso-coHom-coHomRed n) (x +ₕ∙ y) + ≡ Iso.fun (Iso-coHom-coHomRed n) x +ₕ Iso.fun (Iso-coHom-coHomRed n) y + ++∙≡+ zero = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl ++∙≡+ (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + +private + homhelp : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) (x y : coHom (suc n) (typ A)) + → Iso.inv (Iso-coHom-coHomRed {A = A} n) (x +ₕ y) + ≡ Iso.inv (Iso-coHom-coHomRed n) x +ₕ∙ Iso.inv (Iso-coHom-coHomRed n) y + homhelp n A = morphLemmas.isMorphInv _+ₕ∙_ _+ₕ_ + (Iso.fun (Iso-coHom-coHomRed n)) (+∙≡+ n) _ + (Iso.rightInv (Iso-coHom-coHomRed n)) (Iso.leftInv (Iso-coHom-coHomRed n)) + +coHomGr≅coHomRedGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) + → GroupEquiv (coHomRedGrDir (suc n) A) (coHomGr (suc n) (typ A)) +fst (coHomGr≅coHomRedGr n A) = isoToEquiv (Iso-coHom-coHomRed n) +snd (coHomGr≅coHomRedGr n A) = makeIsGroupHom (+∙≡+ n) + +coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → AbGroup ℓ +coHomRedGroup zero A = coHomRedGroupDir zero A +coHomRedGroup (suc n) A = + InducedAbGroup (coHomGroup (suc n) (typ A)) + _+ₕ∙_ + (isoToEquiv (invIso (Iso-coHom-coHomRed n))) + (homhelp n A) + +abstract + coHomGroup≡coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) + → coHomGroup (suc n) (typ A) ≡ coHomRedGroup (suc n) A + coHomGroup≡coHomRedGroup n A = + InducedAbGroupPath (coHomGroup (suc n) (typ A)) + _+ₕ∙_ + (isoToEquiv (invIso (Iso-coHom-coHomRed n))) + (homhelp n A) + +------------------- Kₙ ≃ ΩKₙ₊₁ --------------------- +-- This proof uses the encode-decode method rather than Freudenthal + +-- We define the map σ : Kₙ → ΩKₙ₊₁ and prove that it is a morphism +private + module _ (n : ℕ) where + σ : {n : ℕ} → coHomK (suc n) → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ + σ {n = n} = trRec (isOfHLevelTrunc (4 + n) _ _) + λ a → cong ∣_∣ (merid a ∙ sym (merid (ptSn (suc n)))) + + σ-hom-helper : ∀ {ℓ} {A : Type ℓ} {a : A} (p : a ≡ a) (r : refl ≡ p) + → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r + σ-hom-helper p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl + + σ-hom : {n : ℕ} (x y : coHomK (suc n)) → σ (x +ₖ y) ≡ σ x ∙ σ y + σ-hom {n = zero} = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 4 _ _) _ _) + (wedgeconFun _ _ + (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) + (λ x → lUnit _ + ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))) + (λ y → cong σ (rUnitₖ 1 ∣ y ∣) + ∙∙ rUnit _ + ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid base))))) + (sym (σ-hom-helper (σ ∣ base ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))))) + σ-hom {n = suc n} = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) + (wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev' n) _ _) + (λ x → lUnit _ + ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))) + (λ y → cong σ (rUnitₖ (2 + n) ∣ y ∣) + ∙∙ rUnit _ + ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid north))))) + (sym (σ-hom-helper (σ ∣ north ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))))) + + -- We will need to following lemma + σ-minusDistr : {n : ℕ} (x y : coHomK (suc n)) → σ (x -ₖ y) ≡ σ x ∙ sym (σ y) + σ-minusDistr {n = n} = + morphLemmas.distrMinus' + _+ₖ_ _∙_ + σ σ-hom ∣ (ptSn (suc n)) ∣ refl + -ₖ_ sym + (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) + (rUnitₖ (suc n)) + (lCancelₖ (suc n)) rCancel + (assocₖ (suc n)) assoc∙ + (cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) + + -- we define the code using addIso + Code : (n : ℕ) → coHomK (2 + n) → Type₀ + Code n x = (trRec {B = TypeOfHLevel ℓ-zero (3 + n)} (isOfHLevelTypeOfHLevel (3 + n)) + λ a → Code' a , hLevCode' a) x .fst + where + Code' : (S₊ (2 + n)) → Type₀ + Code' north = coHomK (suc n) + Code' south = coHomK (suc n) + Code' (merid a i) = isoToPath (addIso (suc n) ∣ a ∣) i + + hLevCode' : (x : S₊ (2 + n)) → isOfHLevel (3 + n) (Code' x) + hLevCode' = suspToPropElim (ptSn (suc n)) (λ _ → isPropIsOfHLevel (3 + n)) (isOfHLevelTrunc (3 + n)) + + symMeridLem : (n : ℕ) → (x : S₊ (suc n)) (y : coHomK (suc n)) + → subst (Code n) (cong ∣_∣ (sym (merid x))) y ≡ y -ₖ ∣ x ∣ + symMeridLem n x = trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + (λ y → cong (_-ₖ ∣ x ∣) (transportRefl ∣ y ∣)) + + decode : {n : ℕ} (x : coHomK (2 + n)) → Code n x → ∣ north ∣ ≡ x + decode {n = n} = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + decode-elim + where + north≡merid : (a : S₊ (suc n)) + → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ + ≡ (Path (coHomK (2 + n)) ∣ north ∣ ∣ south ∣) + north≡merid a i = Path (coHomK (2 + n)) ∣ north ∣ ∣ merid a i ∣ + + decode-elim : (a : S₊ (2 + n)) → Code n ∣ a ∣ → Path (coHomK (2 + n)) ∣ north ∣ ∣ a ∣ + decode-elim north = σ + decode-elim south = trRec (isOfHLevelTrunc (4 + n) _ _) + λ a → cong ∣_∣ (merid a) + decode-elim (merid a i) = + hcomp (λ k → λ { (i = i0) → σ + ; (i = i1) → mainPath a k}) + (funTypeTransp (Code n) (λ x → ∣ north ∣ ≡ x) (cong ∣_∣ (merid a)) σ i) + where + mainPath : (a : (S₊ (suc n))) → + transport (north≡merid a) ∘ σ ∘ transport (λ i → Code n ∣ merid a (~ i) ∣) + ≡ trRec (isOfHLevelTrunc (4 + n) _ _) λ a → cong ∣_∣ (merid a) + mainPath a = funExt (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) + (λ x → (λ i → transport (north≡merid a) (σ (symMeridLem n a ∣ x ∣ i))) + ∙∙ cong (transport (north≡merid a)) (-distrHelp x) + ∙∙ (substAbove x))) + where + -distrHelp : (x : S₊ (suc n)) → σ (∣ x ∣ -ₖ ∣ a ∣) ≡ cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a)) + -distrHelp x = + σ-minusDistr ∣ x ∣ ∣ a ∣ + ∙ (λ i → (cong ∣_∣ (compPath-filler (merid x) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))) + ∙ (cong ∣_∣ (sym (compPath-filler (merid a) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))))) + + substAbove : (x : S₊ (suc n)) → transport (north≡merid a) (cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a))) + ≡ cong ∣_∣ (merid x) + substAbove x i = transp (λ j → north≡merid a (i ∨ j)) i + (compPath-filler (cong ∣_∣ (merid x)) (λ j → ∣ merid a (~ j ∨ i) ∣) (~ i)) + + encode : {n : ℕ} {x : coHomK (2 + n)} → Path (coHomK (2 + n)) ∣ north ∣ x → Code n x + encode {n = n} p = transport (cong (Code n) p) ∣ (ptSn (suc n)) ∣ + + decode-encode : {n : ℕ} {x : coHomK (2 + n)} (p : Path (coHomK (2 + n)) ∣ north ∣ x) + → decode _ (encode p) ≡ p + decode-encode {n = n} = + J (λ y p → decode _ (encode p) ≡ p) + (cong (decode ∣ north ∣) (transportRefl ∣ ptSn (suc n) ∣) + ∙ cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) + +-- We define an addition operation on Code which we can use in order to show that encode is a +-- morphism (in a very loose sense) + hLevCode : {n : ℕ} (x : coHomK (2 + n)) → isOfHLevel (3 + n) (Code n x) + hLevCode {n = n} = + trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsOfHLevel (3 + n))) + (sphereToPropElim _ + (λ _ → (isPropIsOfHLevel (3 + n))) (isOfHLevelTrunc (3 + n))) + + Code-add' : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n ∣ x ∣ → Code n ∣ x ∣ + Code-add' {n = n} north = _+ₖ_ + Code-add' {n = n} south = _+ₖ_ + Code-add' {n = n} (merid a i) = helper n a i + where + help : (n : ℕ) → (x y a : S₊ (suc n)) + → transport (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) + (_+ₖ_) ∣ x ∣ ∣ y ∣ + ≡ ∣ x ∣ +ₖ ∣ y ∣ + help n x y a = + (λ i → transportRefl ((∣ transportRefl x i ∣ +ₖ (∣ transportRefl y i ∣ -ₖ ∣ a ∣)) +ₖ ∣ a ∣) i) + ∙∙ cong (_+ₖ ∣ a ∣) (assocₖ _ ∣ x ∣ ∣ y ∣ (-ₖ ∣ a ∣)) + ∙∙ sym (assocₖ _ (∣ x ∣ +ₖ ∣ y ∣) (-ₖ ∣ a ∣) ∣ a ∣) + ∙∙ cong ((∣ x ∣ +ₖ ∣ y ∣) +ₖ_) (lCancelₖ _ ∣ a ∣) + ∙∙ rUnitₖ _ _ + + helper : (n : ℕ) (a : S₊ (suc n)) + → PathP (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) _+ₖ_ _+ₖ_ + helper n a = + toPathP (funExt + (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelΠ (3 + n) (λ _ → isOfHLevelTrunc (3 + n))) _ _) + λ x → funExt + (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + λ y → help n x y a))) + + Code-add : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n x → Code n x + Code-add {n = n} = + trElim (λ x → isOfHLevelΠ (4 + n) λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelSuc (3 + n) (hLevCode {n = n} x)) + Code-add' + + encode-hom : {n : ℕ} {x : _} (q : 0ₖ _ ≡ 0ₖ _) (p : 0ₖ _ ≡ x) + → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p) + encode-hom {n = n} q = J (λ x p → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p)) + (cong encode (sym (rUnit q)) + ∙∙ sym (rUnitₖ _ (encode q)) + ∙∙ cong (encode q +ₖ_) (cong ∣_∣ (sym (transportRefl _)))) + +stabSpheres : (n : ℕ) → Iso (coHomK (suc n)) (typ (Ω (coHomK-ptd (2 + n)))) +fun (stabSpheres n) = decode _ +inv' (stabSpheres n) = encode +rightInv (stabSpheres n) p = decode-encode p +leftInv (stabSpheres n) = + trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + λ a → cong encode (congFunct ∣_∣ (merid a) (sym (merid (ptSn (suc n))))) + ∙∙ (λ i → transport (congFunct (Code n) (cong ∣_∣ (merid a)) + (cong ∣_∣ (sym (merid (ptSn (suc n))))) i) ∣ ptSn (suc n) ∣) + ∙∙ (substComposite (λ x → x) + (cong (Code n) (cong ∣_∣ (merid a))) + (cong (Code n) (cong ∣_∣ (sym (merid (ptSn (suc n)))))) ∣ ptSn (suc n) ∣ + ∙∙ cong (transport (λ i → Code n ∣ merid (ptSn (suc n)) (~ i) ∣)) + (transportRefl (∣ (ptSn (suc n)) ∣ +ₖ ∣ a ∣) ∙ lUnitₖ (suc n) ∣ a ∣) + ∙∙ symMeridLem n (ptSn (suc n)) ∣ a ∣ + ∙∙ cong (∣ a ∣ +ₖ_) -0ₖ + ∙∙ rUnitₖ (suc n) ∣ a ∣) + +Iso-Kn-ΩKn+1 : (n : HLevel) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +Iso-Kn-ΩKn+1 zero = invIso (compIso (congIso (truncIdempotentIso _ isGroupoidS¹)) ΩS¹Isoℤ) +Iso-Kn-ΩKn+1 (suc n) = stabSpheres n + +Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) + +-- Some properties of the Iso +Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +Kn→ΩKn+1 n = Iso.fun (Iso-Kn-ΩKn+1 n) + +ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n +ΩKn+1→Kn n = Iso.inv (Iso-Kn-ΩKn+1 n) + +Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n (0ₖ n) ≡ refl +Kn→ΩKn+10ₖ zero = sym (rUnit refl) +Kn→ΩKn+10ₖ (suc n) i j = ∣ (rCancel (merid (ptSn (suc n))) i j) ∣ + +ΩKn+1→Kn-refl : (n : ℕ) → ΩKn+1→Kn n refl ≡ 0ₖ n +ΩKn+1→Kn-refl zero = refl +ΩKn+1→Kn-refl (suc zero) = refl +ΩKn+1→Kn-refl (suc (suc n)) = refl + +Kn≃ΩKn+1∙ : {n : ℕ} → coHomK-ptd n ≡ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1∙ {n = n} = ua∙ Kn≃ΩKn+1 (Kn→ΩKn+10ₖ n) + +Kn→ΩKn+1-hom : (n : ℕ) (x y : coHomK n) → Kn→ΩKn+1 n (x +[ n ]ₖ y) ≡ Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y +Kn→ΩKn+1-hom zero x y = (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) + (inS (∣ intLoop (x ℤ+ y) i ∣)) (~ j)) + ∙∙ (λ j i → ∣ intLoop-hom x y (~ j) i ∣) + ∙∙ (congFunct ∣_∣ (intLoop x) (intLoop y) + ∙ cong₂ _∙_ (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) + (inS (∣ intLoop x i ∣)) j) + λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) + (inS (∣ intLoop y i ∣)) j) +Kn→ΩKn+1-hom (suc n) = σ-hom + +ΩKn+1→Kn-hom : (n : ℕ) (x y : Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) + → ΩKn+1→Kn n (x ∙ y) ≡ ΩKn+1→Kn n x +[ n ]ₖ ΩKn+1→Kn n y +ΩKn+1→Kn-hom zero p q = + cong winding (congFunct (trRec isGroupoidS¹ (λ x → x)) p q) + ∙ winding-hom (cong (trRec isGroupoidS¹ (λ x → x)) p) (cong (trRec isGroupoidS¹ (λ x → x)) q) +ΩKn+1→Kn-hom (suc n) = encode-hom + +Kn→ΩKn+1-ₖ : (n : ℕ) (x : coHomK n) → Kn→ΩKn+1 n (-ₖ x) ≡ sym (Kn→ΩKn+1 n x) +Kn→ΩKn+1-ₖ n x = + lUnit _ + ∙∙ cong (_∙ Kn→ΩKn+1 n (-ₖ x)) (sym (lCancel _)) + ∙∙ sym (assoc∙ _ _ _) + ∙∙ cong (sym (Kn→ΩKn+1 n x) ∙_) help + ∙∙ sym (rUnit _) + where + help : Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n (-ₖ x) ≡ refl + help = sym (Kn→ΩKn+1-hom n x (-ₖ x)) ∙∙ cong (Kn→ΩKn+1 n) (rCancelₖ n x) ∙∙ Kn→ΩKn+10ₖ n + +isHomogeneousKn : (n : HLevel) → isHomogeneous (coHomK-ptd n) +isHomogeneousKn n = + subst isHomogeneous + (sym (ΣPathP (ua Kn≃ΩKn+1 , ua-gluePath _ (Kn→ΩKn+10ₖ n)))) + (isHomogeneousPath _ _) + +-- With the equivalence Kn≃ΩKn+1, we get that the two definitions of cohomology groups agree +open IsGroupHom + +coHom≅coHomΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → GroupIso (coHomGr n A) (coHomGrΩ n A) +fun (fst (coHom≅coHomΩ n A)) = map λ f a → Kn→ΩKn+1 n (f a) +inv' (fst (coHom≅coHomΩ n A)) = map λ f a → ΩKn+1→Kn n (f a) +rightInv (fst (coHom≅coHomΩ n A)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ f → cong ∣_∣₂ (funExt λ x → rightInv (Iso-Kn-ΩKn+1 n) (f x)) +leftInv (fst (coHom≅coHomΩ n A)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ f → cong ∣_∣₂ (funExt λ x → leftInv (Iso-Kn-ΩKn+1 n) (f x)) +snd (coHom≅coHomΩ n A) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + λ f g → cong ∣_∣₂ (funExt λ x → Kn→ΩKn+1-hom n (f x) (g x))) + +module lockedKnIso (key : Unit') where + Kn→ΩKn+1' : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) + Kn→ΩKn+1' n = lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) + + ΩKn+1→Kn' : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n + ΩKn+1→Kn' n = lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) + + ΩKn+1→Kn→ΩKn+1 : (n : ℕ) → (x : typ (Ω (coHomK-ptd (suc n)))) → Kn→ΩKn+1' n (ΩKn+1→Kn' n x) ≡ x + ΩKn+1→Kn→ΩKn+1 n x = pm key + where + pm : (key : Unit') → lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) (lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) x) ≡ x + pm unlock = Iso.rightInv (Iso-Kn-ΩKn+1 n) x + + Kn→ΩKn+1→Kn : (n : ℕ) → (x : coHomK n) → ΩKn+1→Kn' n (Kn→ΩKn+1' n x) ≡ x + Kn→ΩKn+1→Kn n x = pm key + where + pm : (key : Unit') → lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) (lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) x) ≡ x + pm unlock = Iso.leftInv (Iso-Kn-ΩKn+1 n) x + +-distrLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : GroupHom (coHomGr n A) (coHomGr m B)) + (x y : coHom n A) + → fst f (x -[ n ]ₕ y) ≡ fst f x -[ m ]ₕ fst f y +-distrLemma n m f' x y = sym (-cancelRₕ m (f y) (f (x -[ n ]ₕ y))) + ∙∙ cong (λ x → x -[ m ]ₕ f y) (sym (f' .snd .pres· (x -[ n ]ₕ y) y)) + ∙∙ cong (λ x → x -[ m ]ₕ f y) ( cong f (-+cancelₕ n _ _)) + where + f = fst f' + +-- HLevel stuff for cup product +isContr-↓∙ : (n : ℕ) → isContr (coHomK-ptd (suc n) →∙ coHomK-ptd n) +fst (isContr-↓∙ zero) = (λ _ → 0) , refl +snd (isContr-↓∙ zero) (f , p) = + Σ≡Prop (λ f → isSetℤ _ _) + (funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 isSetℤ) _ _) + (toPropElim (λ _ → isSetℤ _ _) (sym p)))) +fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ _) , refl +fst (snd (isContr-↓∙ (suc n)) f i) x = + trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (isOfHLevelTrunc (3 + n))) _ _) + (sphereElim _ (λ _ → isOfHLevelTrunc (3 + n) _ _) + (sym (snd f))) x i +snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) + +isContr-↓∙' : (n : ℕ) → isContr (S₊∙ (suc n) →∙ coHomK-ptd n) +fst (isContr-↓∙' zero) = (λ _ → 0) , refl +snd (isContr-↓∙' zero) (f , p) = + Σ≡Prop (λ f → isSetℤ _ _) + (funExt (toPropElim (λ _ → isSetℤ _ _) (sym p))) +fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ _) , refl +fst (snd (isContr-↓∙' (suc n)) f i) x = + sphereElim _ {A = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelTrunc (3 + n) _ _) + (sym (snd f)) x i +snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) + +isOfHLevel→∙Kn : {A : Pointed₀} (n m : ℕ) + → isOfHLevel (suc m) (A →∙ coHomK-ptd n) → isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) +isOfHLevel→∙Kn {A = A} n m hlev = step₃ + where + step₁ : isOfHLevel (suc m) (A →∙ Ω (coHomK-ptd (suc n))) + step₁ = + subst (isOfHLevel (suc m)) + (λ i → A →∙ ua∙ {A = Ω (coHomK-ptd (suc n))} {B = coHomK-ptd n} + (invEquiv Kn≃ΩKn+1) + (ΩKn+1→Kn-refl n) (~ i)) hlev + + step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ coHomK-ptd (suc n) ∙))) + step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ + + step₃ : isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) + step₃ = + isOfHLevelΩ→isOfHLevel m + λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) + (isHomogeneous→∙ (isHomogeneousKn (suc n)) f) + step₂ + +isOfHLevel↑∙ : ∀ n m → isOfHLevel (2 + n) (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m))) +isOfHLevel↑∙ zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙ m)) +isOfHLevel↑∙ (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) + +isOfHLevel↑∙' : ∀ n m → isOfHLevel (2 + n) (S₊∙ (suc m) →∙ coHomK-ptd (suc (n + m))) +isOfHLevel↑∙' zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙' m)) +isOfHLevel↑∙' (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) + +→∙KnPath : (A : Pointed₀) (n : ℕ) → Ω (A →∙ coHomK-ptd (suc n) ∙) ≡ (A →∙ coHomK-ptd n ∙) +→∙KnPath A n = + ua∙ (isoToEquiv (ΩfunExtIso A (coHomK-ptd (suc n)))) refl + ∙ (λ i → (A →∙ Kn≃ΩKn+1∙ {n = n} (~ i) ∙)) + +contr∙-lem : (n m : ℕ) + → isContr (coHomK-ptd (suc n) + →∙ (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (n + m)) ∙)) +fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl +snd (contr∙-lem n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) (sym (funExt help)) + where + help : (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) + help = + trElim (λ _ → isOfHLevelPath (3 + n) + (subst (λ x → isOfHLevel x (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m)))) (λ i → suc (suc (+-comm n 1 i))) + (isOfHLevelPlus' {n = 1} (2 + n) (isOfHLevel↑∙ n m))) _ _) + (sphereElim _ (λ _ → isOfHLevel↑∙ n m _ _) p) + +isOfHLevel↑∙∙ : ∀ n m l + → isOfHLevel (2 + l) (coHomK-ptd (suc n) + →∙ (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (l + n + m))) ∙)) +isOfHLevel↑∙∙ n m zero = + isOfHLevelΩ→isOfHLevel 0 λ f + → subst + isProp (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) + (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) + where + h : isProp (coHomK-ptd (suc n) + →∙ (Ω (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (n + m))) ∙))) + h = + subst isProp (λ i → coHomK-ptd (suc n) →∙ (→∙KnPath (coHomK-ptd (suc m)) (suc (n + m)) (~ i))) + (isContr→isProp (contr∙-lem n m)) +isOfHLevel↑∙∙ n m (suc l) = + isOfHLevelΩ→isOfHLevel (suc l) + λ f → + subst + (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) + (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) + where + h : isOfHLevel (2 + l) + (coHomK-ptd (suc n) + →∙ (Ω (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (suc (l + n + m)))) ∙))) + h = + subst (isOfHLevel (2 + l)) + (λ i → coHomK-ptd (suc n) →∙ →∙KnPath (coHomK-ptd (suc m)) (suc (suc (l + n + m))) (~ i)) + (isOfHLevel↑∙∙ n m l) + +----------- Misc. ------------ +isoType→isoCohom : {A : Type ℓ} {B : Type ℓ'} (n : ℕ) + → Iso A B + → GroupIso (coHomGr n A) (coHomGr n B) +fst (isoType→isoCohom n is) = setTruncIso (domIso is) +snd (isoType→isoCohom n is) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + (λ _ _ → refl)) + +-- Explicit index swapping for cohomology groups +transportCohomIso : {A : Type ℓ} {n m : ℕ} + → (p : n ≡ m) + → GroupIso (coHomGr n A) (coHomGr m A) +Iso.fun (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) p +Iso.inv (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) (sym p) +Iso.rightInv (fst (transportCohomIso p)) = + transportTransport⁻ (cong (\ n → coHomGr n _ .fst) p) +Iso.leftInv (fst (transportCohomIso p)) = + transportTransport⁻ (cong (\ n → coHomGr n _ .fst) (sym p)) +snd (transportCohomIso {A = A} {n = n} {m = m} p) = + makeIsGroupHom (λ x y → help x y p) + where + help : (x y : coHom n A) (p : n ≡ m) + → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y + help x y = + J (λ m p → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y) + (transportRefl (x +ₕ y) + ∙ cong₂ _+ₕ_ (sym (transportRefl x)) (sym (transportRefl y))) diff --git a/Cubical/ZCohomology/RingStructure/CupProduct.agda b/Cubical/ZCohomology/RingStructure/CupProduct.agda new file mode 100644 index 0000000000..dd9da81808 --- /dev/null +++ b/Cubical/ZCohomology/RingStructure/CupProduct.agda @@ -0,0 +1,87 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.RingStructure.CupProduct where +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties + +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec2 to sRec2) +open import Cubical.HITs.Truncation renaming (rec to trRec) + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Data.Int hiding (_+'_ ; +'≡+ ; _+_) +open import Cubical.Data.Nat + +infixl 30 _·₀_ +infixr 35 _⌣ₖ_ +infixr 35 _⌣_ + +--- This definition of ℕ-addition removes some unnecessary transports. +_+'_ : ℕ → ℕ → ℕ +zero +' b = b +suc a +' zero = suc a +suc a +' suc b = 2 + (a + b) + ++'≡+ : (n m : ℕ) → n +' m ≡ n + m ++'≡+ zero m = refl ++'≡+ (suc n) zero = cong suc (sym (+-comm n zero)) ++'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) + +-- Cup product with one integer (K₀) argument +_·₀_ : {n : ℕ} (m : ℤ) → coHomK n → coHomK n +_·₀_ {n = n} (pos zero) x = 0ₖ _ +_·₀_ {n = n} (pos (suc m)) x = x +ₖ (pos m ·₀ x) +_·₀_ {n = n} (negsuc zero) x = -ₖ x +_·₀_ {n = n} (negsuc (suc m)) x = (negsuc m ·₀ x) -ₖ x + +·₀-0ₖ : {n : ℕ} (m : ℤ) → _·₀_ m (0ₖ n) ≡ 0ₖ n +·₀-0ₖ (pos zero) = refl +·₀-0ₖ (pos (suc n)) = cong (0ₖ _ +ₖ_) (·₀-0ₖ (pos n)) ∙ rUnitₖ _ (0ₖ _) +·₀-0ₖ (negsuc zero) = -0ₖ +·₀-0ₖ (negsuc (suc n)) = cong (λ x → x -ₖ (0ₖ _)) (·₀-0ₖ (negsuc n)) ∙ rCancelₖ _ (0ₖ _) + +-- Pointed version first (enables truncation elimination) +⌣ₖ∙ : (n m : ℕ) → coHomK n → coHomK-ptd m →∙ coHomK-ptd (n +' m) +fst (⌣ₖ∙ zero m a) b = a ·₀ b +snd (⌣ₖ∙ zero m a) = ·₀-0ₖ a +fst (⌣ₖ∙ (suc n) zero a) b = b ·₀ a +snd (⌣ₖ∙ (suc n) zero a) = refl +⌣ₖ∙ (suc n) (suc m) = trRec (isOfHLevel↑∙ (suc n) m) (cup n m) + where + cup : (n m : ℕ) → S₊ (suc n) → coHomK-ptd (suc m) →∙ coHomK-ptd (suc (suc (n + m))) + fst (cup zero m base) _ = 0ₖ _ + fst (cup zero m (loop i)) x = Kn→ΩKn+1 _ x i + fst (cup (suc n) m north) _ = 0ₖ _ + fst (cup (suc n) m south) _ = 0ₖ _ + fst (cup (suc n) m (merid a i)) x = Kn→ΩKn+1 _ (fst (cup n m a) x) i + snd (cup zero m base) = refl + snd (cup zero m (loop i)) k = Kn→ΩKn+10ₖ _ k i + snd (cup (suc n) m north) = refl + snd (cup (suc n) m south) = refl + snd (cup (suc n) m (merid a i)) k = (cong (Kn→ΩKn+1 _) (snd (cup n m a)) ∙ Kn→ΩKn+10ₖ _) k i + +-- Non pointed version +_⌣ₖ_ : {n m : ℕ} → coHomK n → coHomK m → coHomK (n +' m) +_⌣ₖ_ {n = n} {m = m} x y = fst (⌣ₖ∙ n m x) y + +-- Doubly pointed version +⌣ₖ∙∙ : (n m : ℕ) → coHomK-ptd n →∙ (coHomK-ptd m →∙ coHomK-ptd (n +' m) ∙) +fst (⌣ₖ∙∙ n m) = ⌣ₖ∙ n m +fst (snd (⌣ₖ∙∙ zero zero) i) x = 0 +fst (snd (⌣ₖ∙∙ zero (suc m)) i) x = 0ₖ _ +fst (snd (⌣ₖ∙∙ (suc n) zero) i) x = ·₀-0ₖ x i +fst (snd (⌣ₖ∙∙ (suc zero) (suc m)) i) x = 0ₖ _ +fst (snd (⌣ₖ∙∙ (suc (suc n)) (suc m)) i) x = 0ₖ _ +snd (snd (⌣ₖ∙∙ zero zero) i) = refl +snd (snd (⌣ₖ∙∙ zero (suc m)) i) = refl +snd (snd (⌣ₖ∙∙ (suc n) zero) i) = refl +snd (snd (⌣ₖ∙∙ (suc zero) (suc m)) i) = refl +snd (snd (⌣ₖ∙∙ (suc (suc n)) (suc m)) i) = refl + +-- Cup product +_⌣_ : ∀ {ℓ} {A : Type ℓ} {n m : ℕ} → coHom n A → coHom m A → coHom (n +' m) A +_⌣_ = sRec2 squash₂ λ f g → ∣ (λ x → f x ⌣ₖ g x) ∣₂ diff --git a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda new file mode 100644 index 0000000000..68ea383b7f --- /dev/null +++ b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda @@ -0,0 +1,1105 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.RingStructure.GradedCommutativity where +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.Properties + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.Truncation + renaming (elim to trElim ; map to trMap ; map2 to trMap2 + ; rec to trRec ; elim3 to trElim3) + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.GroupoidLaws hiding (assoc) +open import Cubical.Foundations.Path +open import Cubical.Data.Int + renaming (_+_ to _ℤ+_ ; _·_ to _ℤ∙_ ; +Comm to +ℤ-comm + ; ·Comm to ∙-comm ; +Assoc to ℤ+-assoc ; -_ to -ℤ_) + hiding (_+'_ ; +'≡+) +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sum + +private + variable + ℓ : Level + +natTranspLem : ∀ {ℓ} {A B : ℕ → Type ℓ} {n m : ℕ} (a : A n) + (f : (n : ℕ) → (a : A n) → B n) (p : n ≡ m) + → f m (subst A p a) ≡ subst B p (f n a) +natTranspLem {A = A} {B = B} a f p = sym (substCommSlice A B f p a) + ++'-comm : (n m : ℕ) → n +' m ≡ m +' n ++'-comm n m = +'≡+ n m ∙∙ +-comm n m ∙∙ sym (+'≡+ m n) + +private + transp0₁ : (n : ℕ) → subst coHomK (+'-comm 1 (suc n)) (0ₖ _) ≡ 0ₖ _ + transp0₁ zero = refl + transp0₁ (suc n) = refl + + transp0₂ : (n m : ℕ) → subst coHomK (+'-comm (suc (suc n)) (suc m)) (0ₖ _) ≡ 0ₖ _ + transp0₂ n zero = refl + transp0₂ n (suc m) = refl + +-- Recurring expressions +private + ΩKn+1→Ω²Kn+2 : {k : ℕ} → typ (Ω (coHomK-ptd k)) → typ ((Ω^ 2) (coHomK-ptd (suc k))) + ΩKn+1→Ω²Kn+2 x = sym (Kn→ΩKn+10ₖ _) ∙∙ cong (Kn→ΩKn+1 _) x ∙∙ Kn→ΩKn+10ₖ _ + + ΩKn+1→Ω²Kn+2' : {k : ℕ} → Kn→ΩKn+1 k (0ₖ k) ≡ Kn→ΩKn+1 k (0ₖ k) → typ ((Ω^ 2) (coHomK-ptd (suc k))) + ΩKn+1→Ω²Kn+2' p = sym (Kn→ΩKn+10ₖ _) ∙∙ p ∙∙ Kn→ΩKn+10ₖ _ + + Kn→Ω²Kn+2 : {k : ℕ} → coHomK k → typ ((Ω^ 2) (coHomK-ptd (2 + k))) + Kn→Ω²Kn+2 x = ΩKn+1→Ω²Kn+2 (Kn→ΩKn+1 _ x) + +-- Definition of of -ₖⁿ̇*ᵐ +-ₖ-helper : {k : ℕ} (n m : ℕ) + → isEvenT n ⊎ isOddT n → isEvenT m ⊎ isOddT m + → coHomKType k → coHomKType k +-ₖ-helper {k = zero} n m (inl x₁) q x = x +-ₖ-helper {k = zero} n m (inr x₁) (inl x₂) x = x +-ₖ-helper {k = zero} n m (inr x₁) (inr x₂) x = 0 - x +-ₖ-helper {k = suc zero} n m p q base = base +-ₖ-helper {k = suc zero} n m (inl x) q (loop i) = loop i +-ₖ-helper {k = suc zero} n m (inr x) (inl x₁) (loop i) = loop i +-ₖ-helper {k = suc zero} n m (inr x) (inr x₁) (loop i) = loop (~ i) +-ₖ-helper {k = suc (suc k)} n m p q north = north +-ₖ-helper {k = suc (suc k)} n m p q south = north +-ₖ-helper {k = suc (suc k)} n m (inl x) q (merid a i) = + (merid a ∙ sym (merid (ptSn (suc k)))) i +-ₖ-helper {k = suc (suc k)} n m (inr x) (inl x₁) (merid a i) = + (merid a ∙ sym (merid (ptSn (suc k)))) i +-ₖ-helper {k = suc (suc k)} n m (inr x) (inr x₁) (merid a i) = + (merid a ∙ sym (merid (ptSn (suc k)))) (~ i) + +-ₖ-gen : {k : ℕ} (n m : ℕ) + (p : isEvenT n ⊎ isOddT n) + (q : isEvenT m ⊎ isOddT m) + → coHomK k → coHomK k +-ₖ-gen {k = zero} = -ₖ-helper {k = zero} +-ₖ-gen {k = suc k} n m p q = trMap (-ₖ-helper {k = suc k} n m p q) + +-- -ₖⁿ̇*ᵐ +-ₖ^_·_ : {k : ℕ} (n m : ℕ) → coHomK k → coHomK k +-ₖ^_·_ {k = k} n m = -ₖ-gen n m (evenOrOdd n) (evenOrOdd m) + +-- cohomology version +-ₕ^_·_ : {k : ℕ} {A : Type ℓ} (n m : ℕ) → coHom k A → coHom k A +-ₕ^_·_ n m = sMap λ f x → (-ₖ^ n · m) (f x) + +-- -ₖⁿ̇*ᵐ = -ₖ for n m odd +-ₖ-gen-inr≡-ₖ : {k : ℕ} (n m : ℕ) (p : _) (q : _) (x : coHomK k) + → -ₖ-gen n m (inr p) (inr q) x ≡ (-ₖ x) +-ₖ-gen-inr≡-ₖ {k = zero} n m p q _ = refl +-ₖ-gen-inr≡-ₖ {k = suc zero} n m p q = + trElim ((λ _ → isOfHLevelTruncPath)) + λ { base → refl + ; (loop i) → refl} +-ₖ-gen-inr≡-ₖ {k = suc (suc k)} n m p q = + trElim ((λ _ → isOfHLevelTruncPath)) + λ { north → refl + ; south → refl + ; (merid a i) k → ∣ symDistr (merid (ptSn _)) (sym (merid a)) (~ k) (~ i) ∣ₕ} + +-- -ₖⁿ̇*ᵐ x = x for n even +-ₖ-gen-inl-left : {k : ℕ} (n m : ℕ) (p : _) (q : _) (x : coHomK k) + → -ₖ-gen n m (inl p) q x ≡ x +-ₖ-gen-inl-left {k = zero} n m p q x = refl +-ₖ-gen-inl-left {k = suc zero} n m p q = + trElim (λ _ → isOfHLevelTruncPath) + λ { base → refl ; (loop i) → refl} +-ₖ-gen-inl-left {k = suc (suc k)} n m p q = + trElim (λ _ → isOfHLevelPath (4 + k) (isOfHLevelTrunc (4 + k)) _ _) + λ { north → refl + ; south → cong ∣_∣ₕ (merid (ptSn _)) + ; (merid a i) k → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) (~ k) i ∣ₕ} + +-- -ₖⁿ̇*ᵐ x = x for m even +-ₖ-gen-inl-right : {k : ℕ} (n m : ℕ) (p : _) (q : _) (x : coHomK k) + → -ₖ-gen n m p (inl q) x ≡ x +-ₖ-gen-inl-right {k = zero} n m (inl x₁) q x = refl +-ₖ-gen-inl-right {k = zero} n m (inr x₁) q x = refl +-ₖ-gen-inl-right {k = suc zero} n m (inl x₁) q = + trElim (λ _ → isOfHLevelTruncPath) + λ { base → refl ; (loop i) → refl} +-ₖ-gen-inl-right {k = suc zero} n m (inr x₁) q = + trElim (λ _ → isOfHLevelTruncPath) + λ { base → refl ; (loop i) → refl} +-ₖ-gen-inl-right {k = suc (suc k)} n m (inl x₁) q = + trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → cong ∣_∣ₕ (merid (ptSn _)) + ; (merid a i) k → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) (~ k) i ∣ₕ} +-ₖ-gen-inl-right {k = suc (suc k)} n m (inr x₁) q = + trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → cong ∣_∣ₕ (merid (ptSn _)) + ; (merid a i) k → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) (~ k) i ∣ₕ} + +-ₖ-gen² : {k : ℕ} (n m : ℕ) + (p : isEvenT n ⊎ isOddT n) + (q : isEvenT m ⊎ isOddT m) + → (x : coHomK k) → -ₖ-gen n m p q (-ₖ-gen n m p q x) ≡ x +-ₖ-gen² {k = zero} n m (inl x₁) q x = refl +-ₖ-gen² {k = zero} n m (inr x₁) (inl x₂) x = refl +-ₖ-gen² {k = zero} n m (inr x₁) (inr x₂) x = + cong (pos 0 -_) (-AntiComm (pos 0) x) + ∙∙ -AntiComm (pos 0) (-ℤ (x - pos 0)) + ∙∙ h x + where + h : (x : _) → -ℤ (-ℤ (x - pos 0) - pos 0) ≡ x + h (pos zero) = refl + h (pos (suc n)) = refl + h (negsuc n) = refl +-ₖ-gen² {k = suc k} n m (inl x₁) q x i = + -ₖ-gen-inl-left n m x₁ q (-ₖ-gen-inl-left n m x₁ q x i) i +-ₖ-gen² {k = suc k} n m (inr x₁) (inl x₂) x i = + -ₖ-gen-inl-right n m (inr x₁) x₂ (-ₖ-gen-inl-right n m (inr x₁) x₂ x i) i +-ₖ-gen² {k = suc k} n m (inr x₁) (inr x₂) x = + (λ i → -ₖ-gen-inr≡-ₖ n m x₁ x₂ (-ₖ-gen-inr≡-ₖ n m x₁ x₂ x i) i) ∙ -ₖ^2 x + +-ₖ-genIso : {k : ℕ} (n m : ℕ) + (p : isEvenT n ⊎ isOddT n) + (q : isEvenT m ⊎ isOddT m) + → Iso (coHomK k) (coHomK k) +Iso.fun (-ₖ-genIso {k = k} n m p q) = -ₖ-gen n m p q +Iso.inv (-ₖ-genIso {k = k} n m p q) = -ₖ-gen n m p q +Iso.rightInv (-ₖ-genIso {k = k} n m p q) = -ₖ-gen² n m p q +Iso.leftInv (-ₖ-genIso {k = k} n m p q) = -ₖ-gen² n m p q + +-- action of cong on -ₖⁿ̇*ᵐ +cong-ₖ-gen-inr : {k : ℕ} (n m : ℕ) (p : _) (q : _) (P : Path (coHomK (2 + k)) (0ₖ _) (0ₖ _)) + → cong (-ₖ-gen n m (inr p) (inr q)) P ≡ sym P +cong-ₖ-gen-inr {k = k} n m p q P = code≡sym (0ₖ _) P + where + code : (x : coHomK (2 + k)) → 0ₖ _ ≡ x → x ≡ 0ₖ _ + code = trElim (λ _ → isOfHLevelΠ (4 + k) λ _ → isOfHLevelTruncPath) + λ { north → cong (-ₖ-gen n m (inr p) (inr q)) + ; south P → cong ∣_∣ₕ (sym (merid (ptSn _))) ∙ (cong (-ₖ-gen n m (inr p) (inr q)) P) + ; (merid a i) → t a i} + where + t : (a : S₊ (suc k)) → PathP (λ i → 0ₖ (2 + k) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + k)) + (cong (-ₖ-gen n m (inr p) (inr q))) + (λ P → cong ∣_∣ₕ (sym (merid (ptSn _))) ∙ (cong (-ₖ-gen n m (inr p) (inr q)) P)) + t a = toPathP (funExt λ P → cong (transport (λ i → ∣ merid a i ∣ ≡ 0ₖ (suc (suc k)))) + (cong (cong (-ₖ-gen n m (inr p) (inr q))) (λ i → (transp (λ j → 0ₖ (suc (suc k)) ≡ ∣ merid a (~ j ∧ ~ i) ∣) i + (compPath-filler P (λ j → ∣ merid a (~ j) ∣ₕ) i)))) + ∙∙ cong (transport (λ i → ∣ merid a i ∣ ≡ 0ₖ (suc (suc k)))) (congFunct (-ₖ-gen n m (inr p) (inr q)) P (sym (cong ∣_∣ₕ (merid a)))) + ∙∙ (λ j → transp (λ i → ∣ merid a (i ∨ j) ∣ ≡ 0ₖ (suc (suc k))) j + (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong (-ₖ-gen n m (inr p) (inr q)) P + ∙ cong (-ₖ-gen n m (inr p) (inr q)) (sym (cong ∣_∣ₕ (merid a)))) j)) + ∙∙ (λ i → sym (cong ∣_∣ₕ (merid a)) + ∙ isCommΩK (2 + k) (cong (-ₖ-gen n m (inr p) (inr q)) P) + (cong (-ₖ-gen n m (inr p) (inr q)) (sym (cong ∣_∣ₕ (merid a)))) i) + ∙∙ (λ j → (λ i → ∣ merid a (~ i ∨ j) ∣) + ∙ (cong ∣_∣ₕ (compPath-filler' (merid a) (sym (merid (ptSn _))) (~ j)) ∙ (λ i → -ₖ-gen n m (inr p) (inr q) (P i)))) + ∙ sym (lUnit _)) + + code≡sym : (x : coHomK (2 + k)) → (p : 0ₖ _ ≡ x) → code x p ≡ sym p + code≡sym x = J (λ x p → code x p ≡ sym p) refl + +cong-cong-ₖ-gen-inr : {k : ℕ} (n m : ℕ) (p : _) (q : _) + (P : Square (refl {x = 0ₖ (suc (suc k))}) refl refl refl) + → cong (cong (-ₖ-gen n m (inr p) (inr q))) P ≡ sym P +cong-cong-ₖ-gen-inr n m p q P = + rUnit _ + ∙∙ (λ k → (λ i → cong-ₖ-gen-inr n m p q refl (i ∧ k)) + ∙∙ (λ i → cong-ₖ-gen-inr n m p q (P i) k) + ∙∙ λ i → cong-ₖ-gen-inr n m p q refl (~ i ∧ k)) + ∙∙ (λ k → transportRefl refl k + ∙∙ cong sym P + ∙∙ transportRefl refl k) + ∙∙ sym (rUnit (cong sym P)) + ∙∙ sym (sym≡cong-sym P) + +Kn→ΩKn+1-ₖ' : {k : ℕ} (n m : ℕ) (p : _) (q : _) (x : coHomK k) + → Kn→ΩKn+1 k (-ₖ-gen n m (inr p) (inr q) x) ≡ sym (Kn→ΩKn+1 k x) +Kn→ΩKn+1-ₖ' n m p q x = cong (Kn→ΩKn+1 _) (-ₖ-gen-inr≡-ₖ n m p q x) ∙ Kn→ΩKn+1-ₖ _ x + +transpΩ² : {n m : ℕ} (p q : n ≡ m) → (P : _) + → transport (λ i → refl {x = 0ₖ (p i)} ≡ refl {x = 0ₖ (p i)}) P + ≡ transport (λ i → refl {x = 0ₖ (q i)} ≡ refl {x = 0ₖ (q i)}) P +transpΩ² p q P k = subst (λ n → refl {x = 0ₖ n} ≡ refl {x = 0ₖ n}) (isSetℕ _ _ p q k) P + +-- Some technical lemmas about Kn→Ω²Kn+2 and its interaction with -ₖⁿ̇*ᵐ and transports +-- TODO : Check if this can be cleaned up more by having more general lemmas +private + lem₁ : (n : ℕ) (a : _) + → (cong (cong (subst coHomK (+'-comm (suc zero) (suc (suc n))))) + (Kn→Ω²Kn+2 ∣ a ∣ₕ)) + ≡ ΩKn+1→Ω²Kn+2 + (sym (transp0₁ n) ∙∙ cong (subst coHomK (+'-comm (suc zero) (suc n))) (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ) ∙∙ transp0₁ n) + lem₁ zero a = + (λ k i j → transportRefl (Kn→Ω²Kn+2 ∣ a ∣ₕ i j) k) + ∙ cong ΩKn+1→Ω²Kn+2 λ k → rUnit (λ i → transportRefl (Kn→ΩKn+1 1 ∣ a ∣ i) (~ k)) k + lem₁ (suc n) a = + (λ k → transp (λ i → refl {x = 0ₖ (+'-comm 1 (suc (suc (suc n))) (i ∨ ~ k))} + ≡ refl {x = 0ₖ (+'-comm 1 (suc (suc (suc n))) (i ∨ ~ k))}) (~ k) + (λ i j → transp (λ i → coHomK (+'-comm 1 (suc (suc (suc n))) (i ∧ ~ k))) k + (Kn→Ω²Kn+2 ∣ a ∣ₕ i j))) + ∙∙ transpΩ² (+'-comm 1 (suc (suc (suc n)))) + (cong suc (+'-comm (suc zero) (suc (suc n)))) + (Kn→Ω²Kn+2 ∣ a ∣ₕ) + ∙∙ sym (natTranspLem {A = λ n → 0ₖ n ≡ 0ₖ n} + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣) + (λ _ → ΩKn+1→Ω²Kn+2) + (+'-comm 1 (suc (suc n)))) + ∙∙ cong ΩKn+1→Ω²Kn+2 + (λ k → transp (λ i → 0ₖ (+'-comm (suc zero) (suc (suc n)) (i ∨ k)) + ≡ 0ₖ (+'-comm (suc zero) (suc (suc n)) (i ∨ k))) k + (λ i → transp (λ i → coHomK (+'-comm (suc zero) (suc (suc n)) (i ∧ k))) (~ k) + (Kn→ΩKn+1 _ ∣ a ∣ₕ i))) + ∙∙ cong ΩKn+1→Ω²Kn+2 (rUnit (cong (subst coHomK (+'-comm (suc zero) (suc (suc n)))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ))) + + lem₂ : (n : ℕ) (a : _) (p : _) (q : _) + → (cong (cong (-ₖ-gen (suc (suc n)) (suc zero) p q + ∘ (subst coHomK (+'-comm 1 (suc (suc n)))))) + (Kn→Ω²Kn+2 (∣ a ∣ₕ))) + ≡ ΩKn+1→Ω²Kn+2 + (sym (transp0₁ n) + ∙∙ cong (subst coHomK (+'-comm (suc zero) (suc n))) + (cong (-ₖ-gen (suc (suc n)) (suc zero) p q) + (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ)) + ∙∙ transp0₁ n) + lem₂ n a (inl x) (inr y) = + (λ k i j → (-ₖ-gen-inl-left (suc (suc n)) 1 x (inr y) ( + subst coHomK (+'-comm 1 (suc (suc n))) + (Kn→Ω²Kn+2 ∣ a ∣ₕ i j))) k) + ∙∙ lem₁ n a + ∙∙ cong ΩKn+1→Ω²Kn+2 (cong (sym (transp0₁ n) ∙∙_∙∙ transp0₁ n) + λ k i → subst coHomK (+'-comm 1 (suc n)) + (-ₖ-gen-inl-left (suc (suc n)) 1 x (inr y) (Kn→ΩKn+1 (suc n) ∣ a ∣ i) (~ k))) + lem₂ n a (inr x) (inr y) = + cong-cong-ₖ-gen-inr (suc (suc n)) 1 x y + (cong + (cong + (subst coHomK (+'-comm 1 (suc (suc n))))) + (Kn→Ω²Kn+2 ∣ a ∣ₕ)) + ∙∙ cong sym (lem₁ n a) + ∙∙ λ k → ΩKn+1→Ω²Kn+2 + (sym (transp0₁ n) ∙∙ + cong (subst coHomK (+'-comm 1 (suc n))) + (cong-ₖ-gen-inr (suc (suc n)) 1 x y + (Kn→ΩKn+1 (suc n) ∣ a ∣) (~ k)) + ∙∙ transp0₁ n) + + lem₃ : (n m : ℕ) (q : _) (p : isEvenT (suc (suc n)) ⊎ isOddT (suc (suc n))) (x : _) + → (((sym (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) + ∙∙ (λ j → -ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (+'-comm (suc (suc m)) (suc n)) (Kn→ΩKn+1 _ x j))) + ∙∙ cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)))) + ≡ (Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) x))) + lem₃ n m q p x = + sym (cong-∙∙ (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (sym (transp0₂ m n)) + (λ j → subst coHomK (+'-comm (suc (suc m)) (suc n)) (Kn→ΩKn+1 _ x j)) + (transp0₂ m n)) + ∙ h n m p q x + where + help : (n m : ℕ) (x : _) + → ((sym (transp0₂ m n)) + ∙∙ (λ j → subst coHomK (+'-comm (suc (suc m)) (suc n)) + (Kn→ΩKn+1 (suc (suc (m + n))) x j)) + ∙∙ transp0₂ m n) + ≡ Kn→ΩKn+1 (suc (n + suc m)) + (subst coHomK (cong suc (+-comm (suc m) n)) x) + help zero m x = + sym (rUnit _) + ∙∙ (λ k i → transp (λ i → coHomK (+'-comm (suc (suc m)) 1 (i ∨ k))) k + (Kn→ΩKn+1 _ + (transp (λ i → coHomK (predℕ (+'-comm (suc (suc m)) 1 (i ∧ k)))) (~ k) x) i)) + ∙∙ cong (Kn→ΩKn+1 _) + λ k → subst coHomK (isSetℕ _ _ (cong predℕ (+'-comm (suc (suc m)) 1)) + (cong suc (+-comm (suc m) zero)) k) x + help (suc n) m x = + sym (rUnit _) + ∙∙ ((λ k i → transp (λ i → coHomK (+'-comm (suc (suc m)) (suc (suc n)) (i ∨ k))) k + (Kn→ΩKn+1 _ + (transp (λ i → coHomK (predℕ (+'-comm (suc (suc m)) (suc (suc n)) (i ∧ k)))) (~ k) x) i))) + ∙∙ cong (Kn→ΩKn+1 _) + (λ k → subst coHomK (isSetℕ _ _ (cong predℕ (+'-comm (suc (suc m)) (suc (suc n)))) + (cong suc (+-comm (suc m) (suc n))) k) x) + + h : (n m : ℕ) (p : isEvenT (suc (suc n)) ⊎ isOddT (suc (suc n))) (q : _) (x : _) + → cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) + (sym (transp0₂ m n) + ∙∙ (λ j → subst coHomK (+'-comm (suc (suc m)) (suc n)) + (Kn→ΩKn+1 (suc (suc (m + n))) x j)) + ∙∙ transp0₂ m n) + ≡ Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) x)) + h n m (inl p) (inl q) x = + (λ k → cong (-ₖ-gen (suc n) (suc (suc m)) + (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inr p) k) (inl q)) + (help n m x k)) + ∙∙ ((λ k i → -ₖ-gen-inl-right (suc n) (suc (suc m)) (inr p) q (help n m x i1 i) k)) + ∙∙ λ i → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-right (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (inr p) (evenOrOdd (suc n)) i) q + (subst coHomK (cong suc (+-comm (suc m) n)) x) (~ i)) + h n m (inl p) (inr q) x = + (λ k → cong (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inr p) k) (inr q)) + (help n m x k)) + ∙∙ cong-ₖ-gen-inr (suc n) (suc (suc m)) p q (help n m x i1) + ∙∙ sym (Kn→ΩKn+1-ₖ' (suc n) (suc (suc m)) p q + (subst coHomK (λ i → suc (+-comm (suc m) n i)) x)) + ∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (inr p) (evenOrOdd (suc n)) k) (inr q) + (subst coHomK (cong suc (+-comm (suc m) n)) x)) + h n m (inr p) (inl q) x = + (λ k → cong (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inl p) k) (inl q)) + (help n m x k)) + ∙∙ (λ k i → -ₖ-gen-inl-left (suc n) (suc (suc m)) p (inl q) (help n m x i1 i) k) + ∙∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-right (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) x) (~ k)) + h n m (inr p) (inr q) x = + (λ k → cong (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inl p) k) (inr q)) + (help n m x k)) + ∙∙ (λ k i → -ₖ-gen-inl-left (suc n) (suc (suc m)) p (inr q) (help n m x i1 i) k) + ∙∙ cong (Kn→ΩKn+1 (suc (n + suc m))) + (sym (-ₖ-gen-inl-left (suc n) (suc (suc m)) p (inr q) + (subst coHomK (λ i → suc (+-comm (suc m) n i)) x))) + ∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (inl p) (evenOrOdd (suc n)) k) (inr q) + (subst coHomK (cong suc (+-comm (suc m) n)) x)) + + lem₄ : (n m : ℕ) (q : _) (p : isEvenT (suc (suc n)) ⊎ isOddT (suc (suc n))) (a : _) (b : _) + → cong (Kn→ΩKn+1 _) (((sym (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) + ∙∙ (λ j → -ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (+'-comm (suc (suc m)) (suc n)) + (_⌣ₖ_ {n = suc (suc m)} {m = (suc n)} ∣ merid b j ∣ₕ ∣ a ∣))) + ∙∙ cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)))) + ≡ cong (Kn→ΩKn+1 _) (Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) + (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ₕ ∣ a ∣)))) + lem₄ n m q p a b = cong (cong (Kn→ΩKn+1 _)) (lem₃ n m q p (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ₕ ∣ a ∣)) + + lem₅ : (n m : ℕ) (p : _) (q : _) (a : _) (b : _) + → cong (cong (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + ∘ (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))))) + (ΩKn+1→Ω²Kn+2 (sym (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m)) + ∙∙ (λ i → -ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (+'-comm (suc (suc n)) (suc m)) + (_⌣ₖ_ {n = suc (suc n)} {m = suc m} ∣ merid a i ∣ₕ ∣ b ∣ₕ))) + ∙∙ cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m))) + ≡ Kn→Ω²Kn+2 (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) (_⌣ₖ_ {n = suc n} {m = suc m} ∣ a ∣ₕ ∣ b ∣ₕ)))) + lem₅ n m p q a b = + cong (cong (cong (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + ∘ (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))))) + (cong (sym (Kn→ΩKn+10ₖ _) ∙∙_∙∙ Kn→ΩKn+10ₖ _) + (lem₄ m n p q b a)) + ∙ help p q (_⌣ₖ_ {n = suc n} {m = suc m} ∣ a ∣ ∣ b ∣) + where + annoying : (x : _) + → cong (cong (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))) + (Kn→Ω²Kn+2 (subst coHomK (cong suc (+-comm (suc n) m)) x)) + ≡ Kn→Ω²Kn+2 (subst coHomK (cong suc (sym (+-suc n m))) x) + annoying x = + ((λ k → transp (λ i → refl {x = 0ₖ ((+'-comm (suc (suc m)) (suc (suc n))) (i ∨ ~ k))} + ≡ refl {x = 0ₖ ((+'-comm (suc (suc m)) (suc (suc n))) (i ∨ ~ k))}) (~ k) + + λ i j → transp (λ i → coHomK (+'-comm (suc (suc m)) (suc (suc n)) (i ∧ ~ k))) k + (Kn→Ω²Kn+2 (subst coHomK (cong suc (+-comm (suc n) m)) x) i j))) + ∙∙ cong (transport (λ i → refl {x = 0ₖ ((+'-comm (suc (suc m)) (suc (suc n))) i)} + ≡ refl {x = 0ₖ ((+'-comm (suc (suc m)) (suc (suc n))) i)})) + (natTranspLem {A = coHomK} x (λ _ → Kn→Ω²Kn+2) (cong suc (+-comm (suc n) m))) + ∙∙ sym (substComposite (λ n → refl {x = 0ₖ n} ≡ refl {x = 0ₖ n}) + (cong (suc ∘ suc ∘ suc) (+-comm (suc n) m)) (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→Ω²Kn+2 x)) + ∙∙ (λ k → subst (λ n → refl {x = 0ₖ n} ≡ refl {x = 0ₖ n}) + (isSetℕ _ _ + (cong (suc ∘ suc ∘ suc) (+-comm (suc n) m) ∙ (+'-comm (suc (suc m)) (suc (suc n)))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n m))) k) + (Kn→Ω²Kn+2 x)) + ∙∙ sym (natTranspLem {A = coHomK} x (λ _ → Kn→Ω²Kn+2) (cong suc (sym (+-suc n m)))) + + help : (p : _) (q : _) (x : _) → + cong (cong (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + ∘ subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))) + (Kn→Ω²Kn+2 (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (+-comm (suc n) m)) x))) + ≡ Kn→Ω²Kn+2 + (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) x))) + help (inl x) (inl y) z = + (λ k i j → + -ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inl x) y + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + ((ΩKn+1→Ω²Kn+2 + (Kn→ΩKn+1 _ (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (subst coHomK (cong suc (+-comm (suc n) m)) z) k))) i j)) k) + ∙∙ annoying z + ∙∙ cong Kn→Ω²Kn+2 + λ k → (-ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inl x) y + (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (subst coHomK (cong suc (sym (+-suc n m))) z) (~ k)) (~ k)) + help (inl x) (inr y) z = + (λ k i j → + -ₖ-gen-inl-left (suc (suc n)) (suc (suc m)) x (inr y) + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→Ω²Kn+2 (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (subst coHomK (cong suc (+-comm (suc n) m)) z) k) i j)) k) + ∙∙ annoying z + ∙∙ cong Kn→Ω²Kn+2 + (λ k → (-ₖ-gen-inl-left (suc (suc n)) (suc (suc m)) x (inr y) + (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (subst coHomK (cong suc (sym (+-suc n m))) z) (~ k)) (~ k))) + help (inr x) (inl y) z = + (λ k i j → -ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inr x) y + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→Ω²Kn+2 + (-ₖ-gen (suc m) (suc (suc n)) (isPropEvenOrOdd (suc m) (evenOrOdd (suc m)) (inr y) k) (inr x) + (subst coHomK (cong suc (+-comm (suc n) m)) z)) i j)) k) + ∙∙ cong (cong (cong (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))) ∘ ΩKn+1→Ω²Kn+2) + (Kn→ΩKn+1-ₖ' (suc m) (suc (suc n)) y x + (subst coHomK (cong suc (+-comm (suc n) m)) z)) + ∙∙ cong sym (annoying z) + ∙∙ cong ΩKn+1→Ω²Kn+2 (sym (Kn→ΩKn+1-ₖ' (suc m) (suc (suc n)) y x + (subst coHomK (cong suc (sym (+-suc n m))) z))) + ∙∙ cong Kn→Ω²Kn+2 λ k → (-ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inr x) y + (-ₖ-gen (suc m) (suc (suc n)) (isPropEvenOrOdd (suc m) (evenOrOdd (suc m)) (inr y) (~ k)) (inr x) + (subst coHomK (cong suc (sym (+-suc n m))) z)) (~ k)) + help (inr x) (inr y) z = + (λ k → cong-cong-ₖ-gen-inr (suc (suc n)) (suc (suc m)) x y + (λ i j → subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→Ω²Kn+2 + (-ₖ-gen (suc m) (suc (suc n)) + (isPropEvenOrOdd (suc m) (evenOrOdd (suc m)) (inl y) k) (inr x) + (subst coHomK (cong suc (+-comm (suc n) m)) z)) i j)) k) + ∙∙ cong (sym ∘ cong (cong (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))) ∘ Kn→Ω²Kn+2) + (-ₖ-gen-inl-left (suc m) (suc (suc n)) y (inr x) + (subst coHomK (cong suc (+-comm (suc n) m)) z)) + ∙∙ cong sym (annoying z) + ∙∙ cong (sym ∘ Kn→Ω²Kn+2) + (sym (-ₖ-gen-inl-left (suc m) (suc (suc n)) y (inr x) + (subst coHomK (cong suc (sym (+-suc n m))) z))) + ∙∙ cong ΩKn+1→Ω²Kn+2 + λ k → (Kn→ΩKn+1-ₖ' (suc (suc n)) (suc (suc m)) x y + (-ₖ-gen (suc m) (suc (suc n)) ( + isPropEvenOrOdd (suc m) (evenOrOdd (suc m)) (inl y) (~ k)) (inr x) + (subst coHomK (cong suc (sym (+-suc n m))) z))) (~ k) + + lem₆ : (n m : ℕ) (p : _) (q : _) (a : _) (b : _) + → flipSquare + (Kn→Ω²Kn+2 (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) + (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ₕ ∣ a ∣)))) + ≡ Kn→Ω²Kn+2 + (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) + (subst coHomK (+'-comm (suc m) (suc n)) (∣ b ∣ₕ ⌣ₖ ∣ a ∣ₕ)))))) + lem₆ n m p q a b = + sym (sym≡flipSquare (Kn→Ω²Kn+2 (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) + (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ₕ ∣ a ∣))))) + ∙ cong ΩKn+1→Ω²Kn+2 + (help₁ + (subst coHomK (cong suc (+-comm (suc m) n)) + (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ ∣ a ∣)) p q + ∙ cong (Kn→ΩKn+1 _) (sym (help₂ (∣ b ∣ ⌣ₖ ∣ a ∣)))) + where + help₁ : (x : _) (p : _) (q : _) + → sym (Kn→ΩKn+1 (suc (n + suc m)) (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q x)) + ≡ Kn→ΩKn+1 (suc (n + suc m)) ((-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) x)))) + help₁ z (inl x) (inl y) = + cong (λ x → sym (Kn→ΩKn+1 (suc (n + suc m)) x)) + (-ₖ-gen-inl-right (suc n) (suc (suc m)) (evenOrOdd (suc n)) y z) + ∙∙ sym (Kn→ΩKn+1-ₖ' (suc n) (suc m) x y z) + ∙∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inl x) y + (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (-ₖ-gen (suc n) (suc m) + (isPropEvenOrOdd (suc n) (inr x) (evenOrOdd (suc n)) k) + (isPropEvenOrOdd (suc m) (inr y) (evenOrOdd (suc m)) k) z) (~ k)) (~ k)) + help₁ z (inl x) (inr y) = + (λ k → sym (Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) + (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inr x) k) (inr y) z))) + ∙∙ cong sym (Kn→ΩKn+1-ₖ' (suc n) (suc (suc m)) x y z) + ∙∙ cong (Kn→ΩKn+1 (suc (n + suc m))) (sym (-ₖ-gen-inl-right (suc n) (suc m) (inr x) y z)) + ∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-left (suc (suc n)) (suc (suc m)) x (inr y) + (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (-ₖ-gen (suc n) (suc m) + (isPropEvenOrOdd (suc n) (inr x) (evenOrOdd (suc n)) k) + (isPropEvenOrOdd (suc m) (inl y) (evenOrOdd (suc m)) k) z) (~ k)) (~ k)) + help₁ z (inr x) (inl y) = + cong (λ x → sym (Kn→ΩKn+1 (suc (n + suc m)) x)) + (-ₖ-gen-inl-right (suc n) (suc (suc m)) (evenOrOdd (suc n)) y z) + ∙∙ (λ k → Kn→ΩKn+1-ₖ' (suc m) (suc (suc n)) y x + (-ₖ-gen-inl-left (suc n) (suc m) x (inr y) z (~ k)) (~ k)) + ∙∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inr x) y + (-ₖ-gen (suc m) (suc (suc n)) + (isPropEvenOrOdd (suc m) (inr y) (evenOrOdd (suc m)) k) (inr x) + (-ₖ-gen (suc n) (suc m) + (isPropEvenOrOdd (suc n) (inl x) (evenOrOdd (suc n)) k) + (isPropEvenOrOdd (suc m) (inr y) (evenOrOdd (suc m)) k) z)) (~ k)) + help₁ z (inr x) (inr y) = + ((λ k → sym (Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) + (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inl x) k) (inr y) z)))) + ∙∙ cong sym (cong (Kn→ΩKn+1 (suc (n + suc m))) (-ₖ-gen-inl-left (suc n) (suc (suc m)) x (inr y) z)) + ∙∙ (λ k → sym (Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-left (suc m) (suc (suc n)) y (inr x) + (-ₖ-gen-inl-right (suc n) (suc m) (inl x) y z (~ k)) (~ k)))) + ∙ λ k → Kn→ΩKn+1-ₖ' (suc (suc n)) (suc (suc m)) x y + (-ₖ-gen (suc m) (suc (suc n)) (isPropEvenOrOdd (suc m) (inl y) (evenOrOdd (suc m)) k) (inr x) + (-ₖ-gen (suc n) (suc m) + (isPropEvenOrOdd (suc n) (inl x) (evenOrOdd (suc n)) k) + (isPropEvenOrOdd (suc m) (inl y) (evenOrOdd (suc m)) k) z)) (~ k) + + help₂ : (x : _) → + (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) + (subst coHomK (+'-comm (suc m) (suc n)) x))))) + ≡ -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) + (subst coHomK (cong suc (+-comm (suc m) n)) x))) + help₂ x = + (λ k → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (transp (λ i → coHomK ((cong suc (sym (+-suc n m))) (i ∨ k))) k + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) + (transp (λ i → coHomK ((cong suc (sym (+-suc n m))) (i ∧ k))) (~ k) + (subst coHomK (+'-comm (suc m) (suc n)) x)))))) + ∙ cong (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + ∘ -ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + ∘ -ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m))) + (sym (substComposite coHomK (+'-comm (suc m) (suc n)) ((cong suc (sym (+-suc n m)))) x) + ∙ λ k → subst coHomK (isSetℕ _ _ (+'-comm (suc m) (suc n) ∙ cong suc (sym (+-suc n m))) + ((cong suc (+-comm (suc m) n))) k) x) + + lem₇ : (n : ℕ) (a : _) (p : _) (q : _) + → ((λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc zero) (evenOrOdd (suc n)) (inr tt) + (transp0₁ n (~ i)))) + ∙∙ (λ i j → Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc zero) (evenOrOdd (suc n)) (inr tt) + (subst coHomK (+'-comm (suc zero) (suc n)) (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ i))) j) + ∙∙ (λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc zero) (evenOrOdd (suc n)) (inr tt) + (transp0₁ n i)))) + ≡ (cong (Kn→ΩKn+1 (suc (suc (n + zero)))) + (sym (transp0₁ n) + ∙∙ sym (cong (subst coHomK (+'-comm (suc zero) (suc n))) + (cong (-ₖ-gen (suc (suc n)) (suc zero) p q) (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ))) + ∙∙ transp0₁ n)) + lem₇ zero a (inl x) (inr tt) = + (λ k → rUnit ((cong (Kn→ΩKn+1 _) (cong-ₖ-gen-inr (suc zero) (suc zero) tt tt + (λ i → (subst coHomK (+'-comm (suc zero) (suc zero)) + (Kn→ΩKn+1 (suc zero) ∣ a ∣ₕ i))) k))) (~ k)) + ∙ λ k → ((cong (Kn→ΩKn+1 (suc (suc zero))) + (rUnit (λ i → subst coHomK (+'-comm (suc zero) (suc zero)) + (-ₖ-gen-inl-left (suc (suc zero)) (suc zero) tt (inr tt) + (Kn→ΩKn+1 (suc zero) ∣ a ∣ₕ (~ i)) (~ k))) k))) + lem₇ (suc n) a (inl x) (inr tt) = + ((λ k → rUnit (cong (Kn→ΩKn+1 _) + (λ i → -ₖ-gen (suc (suc n)) (suc zero) + (isPropEvenOrOdd n (evenOrOdd (suc (suc n))) (inr x) k) (inr tt) + (subst coHomK (+'-comm (suc zero) (suc (suc n))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ i)))) (~ k))) + ∙∙ (((λ k → ((cong (Kn→ΩKn+1 _) (cong-ₖ-gen-inr (suc (suc n)) (suc zero) x tt + (λ i → (subst coHomK (+'-comm (suc zero) (suc (suc n))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ i))) k)))))) + ∙∙ λ k → ((cong (Kn→ΩKn+1 (suc (suc (suc n + zero)))) + (rUnit (λ i → subst coHomK (+'-comm (suc zero) (suc (suc n))) + (-ₖ-gen-inl-left (suc (suc (suc n))) (suc zero) x (inr tt) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ (~ i)) (~ k))) k))) + lem₇ (suc n) a (inr x) (inr tt) = + (λ k → rUnit (λ i j → Kn→ΩKn+1 _ + (-ₖ-gen (suc (suc n)) (suc zero) + (isPropEvenOrOdd (suc (suc n)) (evenOrOdd (suc (suc n))) (inl x) k) (inr tt) + (subst coHomK (+'-comm (suc zero) (suc (suc n))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ i))) j) (~ k)) + ∙∙ (λ k i j → Kn→ΩKn+1 _ (-ₖ-gen-inl-left (suc (suc n)) (suc zero) x (inr tt) + (subst coHomK (+'-comm (suc zero) (suc (suc n))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ i)) k) j) + ∙∙ λ k → cong (Kn→ΩKn+1 _) + (rUnit (sym (cong (subst coHomK (+'-comm (suc zero) (suc (suc n)))) + (cong-ₖ-gen-inr (suc (suc (suc n))) (suc zero) x tt (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ) (~ k)))) k) + +-- ∣ a ∣ ⌣ₖ ∣ b ∣ ≡ -ₖⁿ*ᵐ (∣ b ∣ ⌣ₖ ∣ a ∣) for n ≥ 1, m = 1 +gradedComm-elimCase-left : (n : ℕ) (p : _) (q : _) (a : S₊ (suc n)) (b : S¹) → + (_⌣ₖ_ {n = suc n} {m = (suc zero)} ∣ a ∣ₕ ∣ b ∣ₕ) + ≡ (-ₖ-gen (suc n) (suc zero) p q) + (subst coHomK (+'-comm (suc zero) (suc n)) + (_⌣ₖ_ {n = suc zero} {m = suc n} ∣ b ∣ₕ ∣ a ∣ₕ)) +gradedComm-elimCase-left zero (inr tt) (inr tt) a b = + proof a b + ∙ cong (-ₖ-gen 1 1 (inr tt) (inr tt)) + (sym (transportRefl ((_⌣ₖ_ {n = suc zero} {m = suc zero} ∣ b ∣ ∣ a ∣)))) + where + help : flipSquare (ΩKn+1→Ω²Kn+2' (λ j i → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop i ∣ₕ ∣ loop j ∣ₕ)) ≡ + cong (cong (-ₖ-gen 1 1 (inr tt) (inr tt))) + (ΩKn+1→Ω²Kn+2' (λ i j → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop j ∣ₕ ∣ loop i ∣ₕ)) + help = sym (sym≡flipSquare _) + ∙ sym (cong-cong-ₖ-gen-inr 1 1 tt tt + (ΩKn+1→Ω²Kn+2' (λ i j → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop j ∣ ∣ loop i ∣))) + + proof : (a b : S¹) → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ a ∣ₕ ∣ b ∣ₕ ≡ + -ₖ-gen 1 1 (inr tt) (inr tt) (_⌣ₖ_ {n = suc zero} {m = suc zero} ∣ b ∣ ∣ a ∣) + proof base base = refl + proof base (loop i) k = -ₖ-gen 1 1 (inr tt) (inr tt) (Kn→ΩKn+10ₖ _ (~ k) i) + proof (loop i) base k = Kn→ΩKn+10ₖ _ k i + proof (loop i) (loop j) k = + hcomp (λ r → λ { (i = i0) → -ₖ-gen 1 1 (inr tt) (inr tt) (Kn→ΩKn+10ₖ _ (~ k ∨ ~ r) j) + ; (i = i1) → -ₖ-gen 1 1 (inr tt) (inr tt) (Kn→ΩKn+10ₖ _ (~ k ∨ ~ r) j) + ; (j = i0) → Kn→ΩKn+10ₖ _ (k ∨ ~ r) i + ; (j = i1) → Kn→ΩKn+10ₖ _ (k ∨ ~ r) i + ; (k = i0) → doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (λ j i → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop i ∣ₕ ∣ loop j ∣ₕ) + (Kn→ΩKn+10ₖ _) (~ r) j i + ; (k = i1) → (-ₖ-gen 1 1 (inr tt) (inr tt) + (doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (λ i j → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop j ∣ₕ ∣ loop i ∣ₕ) + (Kn→ΩKn+10ₖ _) (~ r) i j))}) + (help k i j) +gradedComm-elimCase-left (suc n) p q north b = + cong (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ b ∣ₕ)) +gradedComm-elimCase-left (suc n) p q south b = + cong (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + ((sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ b ∣ₕ)) ∙ λ i → ∣ b ∣ ⌣ₖ ∣ merid (ptSn (suc n)) i ∣ₕ) +gradedComm-elimCase-left (suc n) p q (merid a i) base k = + hcomp (λ j → λ {(i = i0) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (0ₖ _) + ; (i = i1) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (compPath-filler (sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ base ∣ₕ)) + (λ i → ∣ base ∣ ⌣ₖ ∣ merid a i ∣ₕ) j k) + ; (k = i0) → _⌣ₖ_ {n = suc (suc n)} {m = suc zero} ∣ merid a i ∣ₕ ∣ base ∣ₕ + ; (k = i1) → -ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (∣ base ∣ₕ ⌣ₖ ∣ merid a i ∣ₕ))}) + (hcomp (λ j → λ {(i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (k = i0) → (sym (Kn→ΩKn+10ₖ _) + ∙ (λ j → Kn→ΩKn+1 _ + (sym (gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a base + ∙ cong (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt)) (transp0₁ n)) j))) j i + ; (k = i1) → ∣ north ∣}) + ∣ north ∣) +gradedComm-elimCase-left (suc n) p q (merid a i) (loop j) k = + hcomp (λ r → + λ { (i = i0) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ (loop j) ∣ₕ) k) + ; (i = i1) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (compPath-filler (sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ (loop j) ∣ₕ)) + (λ i → ∣ loop j ∣ ⌣ₖ ∣ merid (ptSn (suc n)) i ∣ₕ) r k) + ; (k = i0) → _⌣ₖ_ {n = suc (suc n)} {m = suc zero} ∣ merid a i ∣ₕ ∣ loop j ∣ₕ + ; (k = i1) → -ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (_⌣ₖ_ {n = suc zero} {m = suc (suc n)} ∣ loop j ∣ₕ + ∣ compPath-filler (merid a) (sym (merid (ptSn (suc n)))) (~ r) i ∣ₕ))}) + (hcomp (λ r → + λ { (i = i0) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + ∣ rCancel (merid (ptSn (suc (suc n)))) (~ k ∧ r) j ∣ + ; (i = i1) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + ∣ rCancel (merid (ptSn (suc (suc n)))) (~ k ∧ r) j ∣ + ; (k = i0) → help₂ r i j + ; (k = i1) → -ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (_⌣ₖ_ {n = suc zero} {m = suc (suc n)} ∣ loop j ∣ₕ + (Kn→ΩKn+1 _ ∣ a ∣ₕ i)))}) + (-ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (_⌣ₖ_ {n = suc zero} {m = suc (suc n)} ∣ loop j ∣ₕ + (Kn→ΩKn+1 _ ∣ a ∣ₕ i))))) + where + P : Path _ (Kn→ΩKn+1 (suc (suc (n + 0))) (0ₖ _)) + (Kn→ΩKn+1 (suc (suc (n + 0))) (_⌣ₖ_ {n = (suc n)} {m = suc zero} ∣ a ∣ ∣ base ∣)) + P i = Kn→ΩKn+1 (suc (suc (n + 0))) + ((sym (gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a base + ∙ cong (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt)) (transp0₁ n)) i)) + + help₁ : (P ∙∙ ((λ i j → _⌣ₖ_ {n = suc (suc n)} {m = suc zero} ∣ merid a j ∣ₕ ∣ loop i ∣ₕ)) ∙∙ sym P) + ≡ ((λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (transp0₁ n (~ i)))) + ∙∙ (λ i j → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (subst coHomK (+'-comm 1 (suc n)) (∣ loop i ∣ₕ ⌣ₖ ∣ a ∣ₕ))) j) + ∙∙ (λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (transp0₁ n i)))) + help₁ k i j = + ((λ i → (Kn→ΩKn+1 (suc (suc (n + 0)))) + (compPath-filler' + ((gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a base)) + (cong (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt)) + (transp0₁ n)) (~ k) (~ i))) + ∙∙ (λ i j → (Kn→ΩKn+1 _ + (gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a (loop i) k) j)) + ∙∙ λ i → (Kn→ΩKn+1 (suc (suc (n + 0)))) + (compPath-filler' + ((gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a base)) + (cong (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt)) + (transp0₁ n)) (~ k) i)) i j + + help₂ : I → I → I → coHomK _ + help₂ r i j = + hcomp (λ k → + λ { (i = i0) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + subst coHomK (+'-comm 1 (suc (suc n)))) + ∣ rCancel (merid (ptSn (suc (suc n)))) (~ k ∨ r) j ∣ + ; (i = i1) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + subst coHomK (+'-comm 1 (suc (suc n)))) + ∣ rCancel (merid (ptSn (suc (suc n)))) (~ k ∨ r) j ∣ + ; (j = i0) → compPath-filler (sym (Kn→ΩKn+10ₖ (suc (suc (n + 0))))) + P k r i + ; (j = i1) → compPath-filler (sym (Kn→ΩKn+10ₖ (suc (suc (n + 0))))) + P k r i + ; (r = i0) → -ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (doubleCompPath-filler (sym (Kn→ΩKn+10ₖ _)) + (λ i j → _⌣ₖ_ {n = suc zero} {m = suc (suc n)} ∣ loop j ∣ₕ + (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ i)) (Kn→ΩKn+10ₖ _) (~ k) i j)) + ; (r = i1) → doubleCompPath-filler P + (λ i j → _⌣ₖ_ {n = suc (suc n)} {m = suc zero} ∣ merid a j ∣ₕ ∣ loop i ∣ₕ) + (sym P) (~ k) j i}) + (hcomp (λ k → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → (Kn→ΩKn+10ₖ (suc (suc (n + 0)))) (~ r) i + ; (j = i1) → (Kn→ΩKn+10ₖ (suc (suc (n + 0)))) (~ r) i + ; (r = i0) → lem₂ n a p q (~ k) i j + ; (r = i1) → help₁ (~ k) j i}) + (hcomp (λ k → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+10ₖ (suc (suc (n + 0))) (~ r) i + ; (j = i1) → Kn→ΩKn+10ₖ (suc (suc (n + 0))) (~ r) i + ; (r = i0) → flipSquare≡cong-sym (flipSquare (ΩKn+1→Ω²Kn+2 + (sym (transp0₁ n) + ∙∙ cong (subst coHomK (+'-comm 1 (suc n))) + (cong (-ₖ-gen (suc (suc n)) 1 p q) + (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ)) + ∙∙ transp0₁ n))) (~ k) i j + ; (r = i1) → ((λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (transp0₁ n (~ i)))) + ∙∙ (λ i j → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (subst coHomK (+'-comm 1 (suc n)) (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ i))) j) + ∙∙ (λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (transp0₁ n i)))) j i}) + (hcomp (λ k → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+10ₖ (suc (suc (n + 0))) (~ r ∧ k) i + ; (j = i1) → Kn→ΩKn+10ₖ (suc (suc (n + 0))) (~ r ∧ k) i + ; (r = i0) → doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (cong (Kn→ΩKn+1 (suc (suc (n + 0)))) + (sym (transp0₁ n) + ∙∙ sym (cong (subst coHomK (+'-comm 1 (suc n))) + (cong (-ₖ-gen (suc (suc n)) 1 p q) + (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ))) + ∙∙ transp0₁ n)) + (Kn→ΩKn+10ₖ _) k j i + ; (r = i1) → lem₇ n a p q (~ k) j i}) + (lem₇ n a p q i1 j i)))) + +-- ∣ a ∣ ⌣ₖ ∣ b ∣ ≡ -ₖⁿ*ᵐ (∣ b ∣ ⌣ₖ ∣ a ∣) for all n, m ≥ 1 +gradedComm-elimCase : (k n m : ℕ) (term : n + m ≡ k) (p : _) (q : _) (a : _) (b : _) → + (_⌣ₖ_ {n = suc n} {m = (suc m)} ∣ a ∣ₕ ∣ b ∣ₕ) + ≡ (-ₖ-gen (suc n) (suc m) p q) + (subst coHomK (+'-comm (suc m) (suc n)) + (_⌣ₖ_ {n = suc m} {m = suc n} ∣ b ∣ₕ ∣ a ∣ₕ)) +gradedComm-elimCase k zero zero term p q a b = gradedComm-elimCase-left zero p q a b +gradedComm-elimCase k zero (suc m) term (inr tt) q a b = + help q + ∙ sym (cong (-ₖ-gen 1 (suc (suc m)) (inr tt) q + ∘ (subst coHomK (+'-comm (suc (suc m)) 1))) + (gradedComm-elimCase-left (suc m) q (inr tt) b a)) + where + help : (q : _) → ∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ ≡ + -ₖ-gen 1 (suc (suc m)) (inr tt) q + (subst coHomK (+'-comm (suc (suc m)) 1) + (-ₖ-gen (suc (suc m)) 1 q (inr tt) + (subst coHomK (+'-comm 1 (suc (suc m))) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)))) + help (inl x) = + (sym (transportRefl _) + ∙ (λ i → subst coHomK (isSetℕ _ _ refl (+'-comm 1 (suc (suc m)) ∙ +'-comm (suc (suc m)) 1) i) + (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ))) + ∙∙ substComposite coHomK + (+'-comm 1 (suc (suc m))) + (+'-comm (suc (suc m)) 1) + ((∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)) + ∙∙ λ i → -ₖ-gen-inl-right (suc zero) (suc (suc m)) (inr tt) x + ((subst coHomK (+'-comm (suc (suc m)) 1) + (-ₖ-gen-inl-left (suc (suc m)) 1 x (inr tt) + (subst coHomK (+'-comm 1 (suc (suc m))) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)) (~ i)))) (~ i) + help (inr x) = + (sym (transportRefl _) + ∙∙ (λ k → subst coHomK (isSetℕ _ _ refl (+'-comm 1 (suc (suc m)) ∙ +'-comm (suc (suc m)) 1) k) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)) + ∙∙ sym (-ₖ^2 (subst coHomK (+'-comm 1 (suc (suc m)) ∙ +'-comm (suc (suc m)) 1) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)))) + ∙∙ (λ i → -ₖ-gen-inr≡-ₖ 1 (suc (suc m)) tt x + (-ₖ-gen-inr≡-ₖ (suc (suc m)) 1 x tt + (substComposite coHomK (+'-comm 1 (suc (suc m))) (+'-comm (suc (suc m)) 1) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ) i) + (~ i)) (~ i)) + ∙∙ λ i → (-ₖ-gen 1 (suc (suc m)) (inr tt) (inr x) + (transp (λ j → coHomK ((+'-comm (suc (suc m)) 1) (j ∨ ~ i))) (~ i) + (-ₖ-gen (suc (suc m)) 1 (inr x) (inr tt) + (transp (λ j → coHomK ((+'-comm (suc (suc m)) 1) (j ∧ ~ i))) i + ((subst coHomK (+'-comm 1 (suc (suc m))) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ))))))) +gradedComm-elimCase k (suc n) zero term p q a b = + gradedComm-elimCase-left (suc n) p q a b +gradedComm-elimCase zero (suc n) (suc m) term p q a b = + ⊥-rec (snotz (sym (+-suc n m) ∙ cong predℕ term)) +gradedComm-elimCase (suc zero) (suc n) (suc m) term p q a b = + ⊥-rec (snotz (sym (+-suc n m) ∙ cong predℕ term)) +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q north north = refl +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q north south = refl +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q north (merid a i) r = + -ₖ-gen (suc (suc n)) (suc (suc m)) p q ( + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))) + ((sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 _) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (sym (transp0₂ n m)) + ∙ sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p a north))) r i)) +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q south north = refl +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q south south = refl +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q south (merid a i) r = + -ₖ-gen (suc (suc n)) (suc (suc m)) p q ( + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))) + ((sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 _) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (sym (transp0₂ n m)) + ∙ sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p a south))) r i)) +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q (merid a i) north r = + (cong (Kn→ΩKn+1 (suc (suc (n + suc m)))) + (gradedComm-elimCase (suc k) n (suc m) (cong predℕ term) (evenOrOdd (suc n)) q a north + ∙ cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) + ∙' Kn→ΩKn+10ₖ _) r i +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q (merid a i) south r = + (cong (Kn→ΩKn+1 (suc (suc (n + suc m)))) + (gradedComm-elimCase (suc k) n (suc m) (cong predℕ term) (evenOrOdd (suc n)) q a south + ∙ cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) + ∙' Kn→ΩKn+10ₖ _) r i +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q (merid a i) (merid b j) r = + hcomp (λ l → + λ { (i = i0) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q ( + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))) + ((compPath-filler (sym (Kn→ΩKn+10ₖ _)) + (cong (Kn→ΩKn+1 _) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (sym (transp0₂ n m)) + ∙ sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b north))) l r j))) + ; (i = i1) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q ( + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))) + ((compPath-filler (sym (Kn→ΩKn+10ₖ _)) + (cong (Kn→ΩKn+1 _) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (sym (transp0₂ n m)) + ∙ sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b south))) l r j))) + ; (r = i0) → help₂ l i j + ; (r = i1) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (help₁ l i j))}) + (hcomp (λ l → + λ { (i = i0) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→ΩKn+10ₖ _ (~ r ∨ ~ l) j)) + ; (i = i1) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→ΩKn+10ₖ _ (~ r ∨ ~ l) j)) + ; (j = i0) → Kn→ΩKn+10ₖ _ r i + ; (j = i1) → Kn→ΩKn+10ₖ _ r i + ; (r = i0) → lem₄ n m q p a b (~ l) j i + ; (r = i1) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (λ i j → Kn→ΩKn+1 _ ((sym (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m)) + ∙∙ (λ i → -ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (+'-comm (suc (suc n)) (suc m)) + (_⌣ₖ_ {n = suc (suc n)} {m = suc m} ∣ merid a i ∣ₕ ∣ b ∣ₕ))) + ∙∙ cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m)) i) j) + (Kn→ΩKn+10ₖ _) (~ l) i j))}) + (hcomp (λ l → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+10ₖ _ r i + ; (j = i1) → Kn→ΩKn+10ₖ _ r i + ; (r = i0) → lem₄ n m q p a b i1 j i + ; (r = i1) → lem₅ n m p q a b (~ l) i j}) + (hcomp (λ l → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+10ₖ _ (r ∨ ~ l) i + ; (j = i1) → Kn→ΩKn+10ₖ _ (r ∨ ~ l) i + ; (r = i0) → doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (lem₄ n m q p a b i1) + (Kn→ΩKn+10ₖ _) (~ l) j i + ; (r = i1) → Kn→Ω²Kn+2 (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) + (gradedComm-elimCase k n m + (+-comm n m ∙∙ cong predℕ (+-comm (suc m) n) ∙∙ cong (predℕ ∘ predℕ) term) + (evenOrOdd (suc n)) (evenOrOdd (suc m)) a b (~ l))))) i j}) + (lem₆ n m p q a b r i j)))) + where + help₁ : I → I → I → coHomK _ + help₁ l i j = + Kn→ΩKn+1 _ + (hcomp (λ r + → λ { (i = i0) → compPath-filler' (cong ((-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p)) (sym (transp0₂ n m))) + (sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b north)) r l + ; (i = i1) → compPath-filler' (cong ((-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p)) (sym (transp0₂ n m))) + (sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b south)) r l + ; (l = i0) → doubleCompPath-filler (sym (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m))) + (λ i → -ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (+'-comm (suc (suc n)) (suc m)) + (_⌣ₖ_ {n = suc (suc n)} {m = suc m} ∣ merid a i ∣ₕ ∣ b ∣ₕ))) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m)) r i + ; (l = i1) → _⌣ₖ_ {n = suc m} {m = suc (suc n)} ∣ b ∣ₕ ∣ merid a i ∣ₕ}) + (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b (merid a i) (~ l))) j + + help₂ : I → I → I → coHomK _ + help₂ l i j = + hcomp (λ r → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → + Kn→ΩKn+1 (suc (suc (n + suc m))) + (compPath-filler (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a north) + (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) r (~ l)) i + ; (j = i1) → + Kn→ΩKn+1 (suc (suc (n + suc m))) + (compPath-filler (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a south) + (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) r (~ l)) i + ; (l = i0) → + Kn→ΩKn+1 _ + (doubleCompPath-filler (sym (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n))) + (λ j → -ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (+'-comm (suc (suc m)) (suc n)) + (_⌣ₖ_ {n = suc (suc m)} {m = (suc n)} ∣ merid b j ∣ₕ ∣ a ∣))) + (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) r j) i + ; (l = i1) → Kn→ΩKn+1 _ (_⌣ₖ_ {n = (suc n)} {m = suc (suc m)} ∣ a ∣ ∣ merid b j ∣ₕ) i}) + (hcomp (λ r → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+1 (suc (suc (n + suc m))) + (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a north (~ l ∨ ~ r)) i + ; (j = i1) → Kn→ΩKn+1 (suc (suc (n + suc m))) + (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a south (~ l ∨ ~ r)) i + ; (l = i0) → Kn→ΩKn+1 (suc (suc (n + suc m))) + (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a (merid b j) i1) i + ; (l = i1) → Kn→ΩKn+1 _ (gradedComm-elimCase (suc k) n (suc m) (cong predℕ term) + (evenOrOdd (suc n)) q a (merid b j) (~ r)) i}) + (Kn→ΩKn+1 (suc (suc (n + suc m))) + (gradedComm-elimCase (suc k) n (suc m) (cong predℕ term) + (evenOrOdd (suc n)) q a (merid b j) i1) i)) + +private + coherence-transp : (n m : ℕ) (p : _) (q : _) + → -ₖ-gen (suc n) (suc m) p q + (subst coHomK (+'-comm (suc m) (suc n)) (0ₖ (suc m +' suc n))) ≡ 0ₖ _ + coherence-transp zero zero p q = refl + coherence-transp zero (suc m) p q = refl + coherence-transp (suc n) zero p q = refl + coherence-transp (suc n) (suc m) p q = refl + +gradedComm-⌣ₖ∙ : (n m : ℕ) (p : _) (q : _) (a : _) + → ⌣ₖ∙ (suc n) (suc m) a + ≡ ((λ b → -ₖ-gen (suc n) (suc m) p q (subst coHomK (+'-comm (suc m) (suc n)) (b ⌣ₖ a))) + , (cong (-ₖ-gen (suc n) (suc m) p q) + (cong (subst coHomK (+'-comm (suc m) (suc n))) + (0ₖ-⌣ₖ (suc m) (suc n) a)) + ∙ coherence-transp n m p q)) +gradedComm-⌣ₖ∙ n m p q = + trElim (λ _ → isOfHLevelPath (3 + n) ((isOfHLevel↑∙ (suc n) m)) _ _) + λ a → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ b → funExt⁻ (cong fst (f₁≡f₂ b)) a) + where + f₁ : coHomK (suc m) → S₊∙ (suc n) →∙ coHomK-ptd (suc n +' suc m) + fst (f₁ b) a = _⌣ₖ_ {n = suc n} {m = suc m} ∣ a ∣ₕ b + snd (f₁ b) = 0ₖ-⌣ₖ (suc n) (suc m) b + + f₂ : coHomK (suc m) → S₊∙ (suc n) →∙ coHomK-ptd (suc n +' suc m) + fst (f₂ b) a = + -ₖ-gen (suc n) (suc m) p q (subst coHomK (+'-comm (suc m) (suc n)) + (_⌣ₖ_ {n = suc m} {m = suc n} b ∣ a ∣ₕ)) + snd (f₂ b) = + (cong (-ₖ-gen (suc n) (suc m) p q) + (cong (subst coHomK (+'-comm (suc m) (suc n))) + (⌣ₖ-0ₖ (suc m) (suc n) b)) + ∙ coherence-transp n m p q) + + f₁≡f₂ : (b : _) → f₁ b ≡ f₂ b + f₁≡f₂ = + trElim (λ _ → isOfHLevelPath (3 + m) + (subst (isOfHLevel (3 + m)) + (λ i → S₊∙ (suc n) →∙ coHomK-ptd (+'-comm (suc n) (suc m) (~ i))) + (isOfHLevel↑∙' (suc m) n)) _ _) + λ b → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ a → gradedComm-elimCase (n + m) n m refl p q a b) + +-- Finally, graded commutativity: +gradedComm-⌣ₖ : (n m : ℕ) (a : coHomK n) (b : coHomK m) + → a ⌣ₖ b ≡ (-ₖ^ n · m) (subst coHomK (+'-comm m n) (b ⌣ₖ a)) +gradedComm-⌣ₖ zero zero a b = sym (transportRefl _) ∙ cong (transport refl) (comm-·₀ a b) +gradedComm-⌣ₖ zero (suc m) a b = + sym (transportRefl _) + ∙∙ (λ k → subst coHomK (isSetℕ _ _ refl (+'-comm (suc m) zero) k) (b ⌣ₖ a)) + ∙∙ sym (-ₖ-gen-inl-left zero (suc m) tt (evenOrOdd (suc m)) + (subst coHomK (+'-comm (suc m) zero) (b ⌣ₖ a))) +gradedComm-⌣ₖ (suc n) zero a b = + sym (transportRefl _) + ∙∙ ((λ k → subst coHomK (isSetℕ _ _ refl (+'-comm zero (suc n)) k) (b ⌣ₖ a))) + ∙∙ sym (-ₖ-gen-inl-right (suc n) zero (evenOrOdd (suc n)) tt + (subst coHomK (+'-comm zero (suc n)) (b ⌣ₖ a))) +gradedComm-⌣ₖ (suc n) (suc m) a b = + funExt⁻ (cong fst (gradedComm-⌣ₖ∙ n m (evenOrOdd (suc n)) (evenOrOdd (suc m)) a)) b + +gradedComm-⌣ : {A : Type ℓ} (n m : ℕ) (a : coHom n A) (b : coHom m A) + → a ⌣ b ≡ (-ₕ^ n · m) (subst (λ n → coHom n A) (+'-comm m n) (b ⌣ a)) +gradedComm-⌣ n m = + sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → + cong ∣_∣₂ (funExt (λ x → + gradedComm-⌣ₖ n m (f x) (g x) + ∙ cong ((-ₖ^ n · m) ∘ (subst coHomK (+'-comm m n))) + λ i → g (transportRefl x (~ i)) ⌣ₖ f (transportRefl x (~ i)))) diff --git a/Cubical/ZCohomology/RingStructure/RingLaws.agda b/Cubical/ZCohomology/RingStructure/RingLaws.agda new file mode 100644 index 0000000000..bbe67c027c --- /dev/null +++ b/Cubical/ZCohomology/RingStructure/RingLaws.agda @@ -0,0 +1,686 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.RingStructure.RingLaws where +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.RingStructure.CupProduct + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.Truncation renaming (elim to trElim) + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.GroupoidLaws hiding (assoc) +open import Cubical.Data.Int + renaming (_+_ to _ℤ+_ ; _·_ to _ℤ∙_ ; +Comm to +ℤ-comm + ; ·Comm to ∙-comm ; +Assoc to ℤ+-assoc ; -_ to -ℤ_) + hiding (_+'_ ; +'≡+) +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +private + variable + ℓ : Level + +-- Some boring lemmas +·₀≡·ℤ : (x y : ℤ) → _·₀_ {n = zero} x y ≡ x ℤ∙ y +·₀≡·ℤ (pos zero) y = refl +·₀≡·ℤ (pos (suc n)) y = cong (y ℤ+_) (·₀≡·ℤ (pos n) y) +·₀≡·ℤ (negsuc zero) (pos zero) = refl +·₀≡·ℤ (negsuc zero) (pos (suc n)) = -AntiComm 0 (pos (suc n)) +·₀≡·ℤ (negsuc zero) (negsuc n) = -AntiComm 0 (negsuc n) +·₀≡·ℤ (negsuc (suc n)) (pos m) = + (cong ((_ℤ+ (pos 0 - pos m))) (·₀≡·ℤ (negsuc n) (pos m)) + ∙ cong ((negsuc n ℤ∙ (pos m)) ℤ+_) + (-AntiComm 0 (pos m) ∙ cong (-ℤ_) (help m))) + ∙ sym (+ℤ-comm (-ℤ (pos m)) (negsuc n ℤ∙ (pos m))) + where + help : (m : ℕ) → (pos m - 0) ≡ pos m + help zero = refl + help (suc m) = refl +·₀≡·ℤ (negsuc (suc n)) (negsuc m) = + +ℤ-comm (_·₀_{n = zero} (negsuc n) (negsuc m)) (pos 0 - (negsuc m)) + ∙∙ cong ((pos 0 - negsuc m) ℤ+_) (·₀≡·ℤ (negsuc n) (negsuc m)) + ∙∙ cong (_ℤ+ (negsuc n ℤ∙ negsuc m)) (help m) + where + help : (m : ℕ) → (pos 0 - negsuc m) ≡ pos (suc m) + help zero = refl + help (suc n) = cong sucℤ (help n) + +comm-·₀ : (x y : ℤ) → _·₀_ {n = 0} x y ≡ y ·₀ x +comm-·₀ x y = ·₀≡·ℤ x y ∙∙ ∙-comm x y ∙∙ sym (·₀≡·ℤ y x) + ++'-assoc : (n m l : ℕ) → (n +' (m +' l)) ≡ ((n +' m) +' l) ++'-assoc n m l = + (λ i → +'≡+ n (+'≡+ m l i) i) + ∙∙ +-assoc n m l + ∙∙ (λ i → +'≡+ (+'≡+ n m (~ i)) l (~ i)) + +-- Zero multiplication + +⌣ₖ-0ₖ : (n m : ℕ) → (x : coHomK n) + → x ⌣ₖ (0ₖ _) ≡ 0ₖ _ +⌣ₖ-0ₖ n m x = snd (⌣ₖ∙ n m x) + +0ₖ-⌣ₖ : (n m : ℕ) → (x : coHomK m) → + (0ₖ _) ⌣ₖ x ≡ 0ₖ _ +0ₖ-⌣ₖ n m = funExt⁻ (cong fst (snd (⌣ₖ∙∙ n m))) + +⌣ₖ-0ₖ≡0ₖ-⌣ₖ : (n m : ℕ) → ⌣ₖ-0ₖ n m (0ₖ _) ≡ 0ₖ-⌣ₖ n m (0ₖ _) +⌣ₖ-0ₖ≡0ₖ-⌣ₖ zero zero = refl +⌣ₖ-0ₖ≡0ₖ-⌣ₖ zero (suc m) = refl +⌣ₖ-0ₖ≡0ₖ-⌣ₖ (suc n) zero = refl +⌣ₖ-0ₖ≡0ₖ-⌣ₖ (suc zero) (suc m) = refl +⌣ₖ-0ₖ≡0ₖ-⌣ₖ (suc (suc n)) (suc m) = refl + +-- Left distributivity + +private + ⌣ₖ-distrFun : -- z ⌣ₖ (x +ₖ y) + (n m : ℕ) → (x y : coHomK n) + → coHomK-ptd m →∙ coHomK-ptd (m +' n) + fst (⌣ₖ-distrFun n m x y) z = + z ⌣ₖ (x +ₖ y) + snd (⌣ₖ-distrFun n m x y) = + 0ₖ-⌣ₖ m n (x +ₖ y) + + ⌣ₖ-distrFun2 : -- z ⌣ₖ x +ₖ z ⌣ₖ y + (n m : ℕ) → (x y : coHomK n) + → coHomK-ptd m →∙ coHomK-ptd (m +' n) + fst (⌣ₖ-distrFun2 n m x y) z = + z ⌣ₖ x +ₖ z ⌣ₖ y + snd (⌣ₖ-distrFun2 n m x y) = + cong₂ _+ₖ_ (0ₖ-⌣ₖ m n x) (0ₖ-⌣ₖ m n y) ∙ rUnitₖ _ _ + + leftDistr-⌣ₖ· : (n m : ℕ) (x y : coHomK (suc n)) → ⌣ₖ-distrFun (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2 (suc n) (suc m) x y + leftDistr-⌣ₖ· n m = + elim2 (λ _ _ → isOfHLevelSuc (2 + n) (hLevHelp n m _ _)) + main + where + hLevHelp : (n m : ℕ) (x y : _) → isOfHLevel (2 + n) ( ⌣ₖ-distrFun (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2 (suc n) (suc m) x y) + hLevHelp n m x y = + (subst (isOfHLevel (3 + n)) + (λ i → (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (suc (+-comm n m i))))) + (isOfHLevel↑∙ (suc n) m)) _ _ + + left-fst : (n : ℕ) (x : S₊ (suc n)) → + fst (⌣ₖ-distrFun (suc n) (suc m) ∣ ptSn (suc n) ∣ ∣ x ∣) ≡ + fst (⌣ₖ-distrFun2 (suc n) (suc m) ∣ ptSn (suc n) ∣ ∣ x ∣) + left-fst zero y = + (funExt (λ z → sym (lUnitₖ _ (z ⌣ₖ ∣ y ∣)) ∙ λ i → ⌣ₖ-0ₖ _ _ z (~ i) +ₖ z ⌣ₖ ∣ y ∣)) + left-fst (suc n) y = + (funExt (λ z → sym (lUnitₖ _ (z ⌣ₖ ∣ y ∣)) ∙ λ i → ⌣ₖ-0ₖ _ _ z (~ i) +ₖ z ⌣ₖ ∣ y ∣)) + + right-fst : (n : ℕ) (x : S₊ (suc n)) → + fst (⌣ₖ-distrFun (suc n) (suc m) ∣ x ∣ (0ₖ _)) ≡ + fst (⌣ₖ-distrFun2 (suc n) (suc m) ∣ x ∣ (0ₖ _)) + right-fst n x = funExt (λ z → (cong (z ⌣ₖ_) (rUnitₖ _ ∣ x ∣) ∙∙ sym (rUnitₖ _ _) ∙∙ λ i → z ⌣ₖ ∣ x ∣ +ₖ ⌣ₖ-0ₖ _ (suc n) z (~ i))) + + helper : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) + → sym (lUnitₖ _ x) ∙ cong (_+ₖ x) p ≡ sym (rUnitₖ _ x) ∙ cong (x +ₖ_) p + helper n x = J (λ x p → sym (lUnitₖ _ x) ∙ cong (_+ₖ x) p ≡ sym (rUnitₖ _ x) ∙ cong (x +ₖ_) p) + (helper' n) + where + helper' : (n : ℕ) → (λ i → lUnitₖ (suc n) (0ₖ (suc n)) (~ i)) ∙ + (λ i → 0ₖ (suc n) +ₖ 0ₖ (suc n)) + ≡ + (λ i → rUnitₖ (suc n) (0ₖ (suc n)) (~ i)) ∙ + (λ i → 0ₖ (suc n) +ₖ 0ₖ (suc n)) + helper' zero = refl + helper' (suc n) = refl + + left-fst≡right-fst : (n : ℕ) → left-fst n (ptSn _) ≡ right-fst n (ptSn _) + left-fst≡right-fst zero i j z = helper _ (z ⌣ₖ ∣ base ∣) (sym (⌣ₖ-0ₖ _ _ z)) i j + left-fst≡right-fst (suc n) i j z = helper _ (z ⌣ₖ ∣ north ∣) (sym (⌣ₖ-0ₖ _ (suc (suc n)) z)) i j + + main : (a b : S₊ (suc n)) → + ⌣ₖ-distrFun (suc n) (suc m) ∣ a ∣ₕ ∣ b ∣ₕ ≡ + ⌣ₖ-distrFun2 (suc n) (suc m) ∣ a ∣ₕ ∣ b ∣ₕ + main = + wedgeconFun n n + (λ x y → subst (λ l → isOfHLevel l ((⌣ₖ-distrFun (suc n) (suc m) ∣ x ∣ ∣ y ∣) + ≡ ⌣ₖ-distrFun2 (suc n) (suc m) ∣ x ∣ ∣ y ∣)) + (+-suc n (suc n)) + (isOfHLevelPlus {n = 2 + n} n (hLevHelp n m ∣ x ∣ ∣ y ∣))) + (λ x → →∙Homogeneous≡ (isHomogeneousKn _) (left-fst n x)) + (λ x → →∙Homogeneous≡ (isHomogeneousKn _) (right-fst n x)) + (cong (→∙Homogeneous≡ (isHomogeneousKn _)) (sym (left-fst≡right-fst n))) + + -- Distributivity for 0 dimensional cases + leftDistr₀n : (n : ℕ) → (z : coHomK 0) (x y : coHomK n) + → z ·₀ (x +[ n ]ₖ y) ≡ z ·₀ x +[ n ]ₖ (z ·₀ y) + leftDistr₀n n (pos zero) x y = sym (rUnitₖ n (0ₖ _)) + leftDistr₀n n (pos (suc z)) x y = + cong ((x +ₖ y) +ₖ_) (leftDistr₀n n (pos z) x y) + ∙∙ sym (assocₖ n x y (pos z ·₀ x +[ n ]ₖ (pos z ·₀ y))) + ∙∙ cong (x +ₖ_) (assocₖ n y (pos z ·₀ x) (pos z ·₀ y) + ∙∙ cong (_+ₖ (pos z ·₀ y)) (commₖ n y (pos z ·₀ x)) + ∙∙ sym (assocₖ n (pos z ·₀ x) y (pos z ·₀ y))) + ∙ assocₖ n x _ _ + leftDistr₀n n (negsuc zero) x y = -distrₖ n x y + leftDistr₀n n (negsuc (suc z)) x y = + cong₂ (_+ₖ_) (leftDistr₀n n (negsuc z) x y) (-distrₖ n x y) + ∙∙ assocₖ n ((negsuc z ·₀ x) +ₖ (negsuc z ·₀ y)) (-ₖ x) (-ₖ y) + ∙∙ cong (_-ₖ y) (sym (assocₖ n (negsuc z ·₀ x) (negsuc z ·₀ y) (-ₖ x)) + ∙∙ cong ((negsuc z ·₀ x) +ₖ_) (commₖ n (negsuc z ·₀ y) (-ₖ x)) + ∙∙ assocₖ n (negsuc z ·₀ x) (-ₖ x) (negsuc z ·₀ y)) + ∙ sym (assocₖ n (negsuc (suc z) ·₀ x) _ _) + + leftDistrn₀ : (n : ℕ) → (z : coHomK n) (x y : coHomK 0) + → (x ℤ+ y) ·₀ z ≡ x ·₀ z +[ n ]ₖ (y ·₀ z) + leftDistrn₀ n z x (pos zero) = sym (rUnitₖ n (x ·₀ z)) + leftDistrn₀ n z x (pos (suc y)) = + lem (x +pos y) + ∙∙ cong (z +ₖ_) (leftDistrn₀ n z x (pos y) ∙ commₖ n _ _) + ∙∙ assocₖ n z _ _ + ∙ commₖ n _ _ + where + lem : (a : ℤ) → (sucℤ a) ·₀ z ≡ z +ₖ (a ·₀ z) + lem (pos zero) = refl + lem (pos (suc a)) = cong (z +ₖ_) (lem (pos a)) + lem (negsuc zero) = sym (rCancelₖ n z) + lem (negsuc (suc a)) = sym (-cancelLₖ n z (negsuc a ·₀ z)) ∙ sym (assocₖ n z (negsuc a ·₀ z) (-ₖ z)) + leftDistrn₀ n z x (negsuc y) = main y + where + help : (x : ℤ) → predℤ x ·₀ z ≡ (x ·₀ z -ₖ z) + help (pos zero) = sym (lUnitₖ n (-ₖ z)) + help (pos (suc x)) = sym (-cancelLₖ n z (pos x ·₀ z)) + help (negsuc x) = refl + + main : (y : _) → (x ℤ+ negsuc y) ·₀ z ≡ (x ·₀ z) +ₖ (negsuc y ·₀ z) + main zero = help x + main (suc y) = + help (x +negsuc y) + ∙∙ cong (_-ₖ z) (main y) + ∙∙ sym (assocₖ n _ _ _) + +leftDistr-⌣ₖ : (n m : ℕ) (z : coHomK n) (x y : coHomK m) + → z ⌣ₖ (x +ₖ y) ≡ (z ⌣ₖ x +ₖ z ⌣ₖ y) +leftDistr-⌣ₖ zero m z x y = leftDistr₀n m z x y +leftDistr-⌣ₖ (suc n) zero z x y = leftDistrn₀ (suc n) z x y +leftDistr-⌣ₖ (suc n) (suc m) z x y = funExt⁻ (cong fst (leftDistr-⌣ₖ· m n x y)) z + + +-- Right distributivity + +private + ⌣ₖ-distrFun-r : -- (x +ₖ y) ⌣ₖ z + (n m : ℕ) → (x y : coHomK n) + → coHomK-ptd m →∙ coHomK-ptd (n +' m) + fst (⌣ₖ-distrFun-r n m x y) z = + (x +ₖ y) ⌣ₖ z + snd (⌣ₖ-distrFun-r n m x y) = + ⌣ₖ-0ₖ n m (x +ₖ y) -- ⌣ₖ-0ₖ m n (x +ₖ y) + + ⌣ₖ-distrFun2-r : + (n m : ℕ) → (x y : coHomK n) + → coHomK-ptd m →∙ coHomK-ptd (n +' m) + fst (⌣ₖ-distrFun2-r n m x y) z = + x ⌣ₖ z +ₖ y ⌣ₖ z + snd (⌣ₖ-distrFun2-r n m x y) = + cong₂ _+ₖ_ (⌣ₖ-0ₖ n m x) (⌣ₖ-0ₖ n m y) ∙ rUnitₖ _ _ + + rightDistr-⌣ₖ· : (n m : ℕ) (x y : coHomK (suc n)) → ⌣ₖ-distrFun-r (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2-r (suc n) (suc m) x y + rightDistr-⌣ₖ· n m = + elim2 (λ _ _ → isOfHLevelPath (3 + n) (isOfHLevel↑∙ (suc n) m) _ _) + main + where + fst-left : (n : ℕ) (y : S₊ (suc n)) → + fst (⌣ₖ-distrFun-r (suc n) (suc m) ∣ ptSn (suc n) ∣ ∣ y ∣) ≡ + fst (⌣ₖ-distrFun2-r (suc n) (suc m) ∣ ptSn (suc n) ∣ ∣ y ∣) + fst-left n y = + funExt (λ z + → cong (_⌣ₖ z) (lUnitₖ _ ∣ y ∣) + ∙∙ sym (lUnitₖ _ (∣ y ∣ ⌣ₖ z)) + ∙∙ cong (_+ₖ (∣ y ∣ ⌣ₖ z)) (sym (0ₖ-⌣ₖ _ _ z))) + + fst-right : (n : ℕ) (x : S₊ (suc n)) → + fst (⌣ₖ-distrFun-r (suc n) (suc m) ∣ x ∣ ∣ ptSn (suc n) ∣) ≡ + fst (⌣ₖ-distrFun2-r (suc n) (suc m) ∣ x ∣ ∣ ptSn (suc n) ∣) + fst-right n x = + funExt λ z + → cong (_⌣ₖ z) (rUnitₖ _ ∣ x ∣) + ∙∙ sym (rUnitₖ _ _) + ∙∙ cong (∣ x ∣ ⌣ₖ z +ₖ_) (sym (0ₖ-⌣ₖ _ _ z)) + + left≡right : (n : ℕ) → fst-left n (ptSn (suc n)) ≡ fst-right n (ptSn (suc n)) + left≡right zero = refl + left≡right (suc n) = refl + + main : (a b : S₊ (suc n)) → + ⌣ₖ-distrFun-r (suc n) (suc m) ∣ a ∣ₕ ∣ b ∣ₕ ≡ + ⌣ₖ-distrFun2-r (suc n) (suc m) ∣ a ∣ₕ ∣ b ∣ₕ + main = + wedgeconFun n n + (λ x y → subst (λ l → isOfHLevel l ((⌣ₖ-distrFun-r (suc n) (suc m) ∣ x ∣ ∣ y ∣) + ≡ ⌣ₖ-distrFun2-r (suc n) (suc m) ∣ x ∣ ∣ y ∣)) + (+-suc n (suc n)) + (isOfHLevelPlus {n = 2 + n} n (isOfHLevel↑∙ (suc n) m _ _))) + (λ x → →∙Homogeneous≡ (isHomogeneousKn _) (fst-left n x)) + (λ x → →∙Homogeneous≡ (isHomogeneousKn _) (fst-right n x)) + (cong (→∙Homogeneous≡ (isHomogeneousKn _)) (sym (left≡right n))) + +rightDistr-⌣ₖ : (n m : ℕ) (x y : coHomK n) (z : coHomK m) + → (x +ₖ y) ⌣ₖ z ≡ (x ⌣ₖ z +ₖ y ⌣ₖ z) +rightDistr-⌣ₖ zero zero x y z = + comm-·₀ (x ℤ+ y) z + ∙∙ leftDistr-⌣ₖ zero zero z x y + ∙∙ cong₂ _+ₖ_ (sym (comm-·₀ x z)) (sym (comm-·₀ y z)) +rightDistr-⌣ₖ zero (suc m) x y z = leftDistr-⌣ₖ _ zero z x y +rightDistr-⌣ₖ (suc n) zero x y z = leftDistr-⌣ₖ zero (suc n) z x y +rightDistr-⌣ₖ (suc n) (suc m) x y z = (funExt⁻ (cong fst (rightDistr-⌣ₖ· n m x y))) z + +-- Associativity + +private +-- We need to give the two associators as (doubly) pointed functions + assocer : (n m k : ℕ) → coHomK (suc n) + → coHomK-ptd (suc m) + →∙ (coHomK-ptd (suc k) + →∙ coHomK-ptd ((suc n) +' ((suc m) +' (suc k))) ∙) + fst (fst (assocer n m k x) y) z = x ⌣ₖ (y ⌣ₖ z) + snd (fst (assocer n m k x) y) = cong (x ⌣ₖ_) (⌣ₖ-0ₖ _ (suc k) y) ∙ ⌣ₖ-0ₖ _ _ x + snd (assocer n m k x) = + ΣPathP (funExt (λ z → cong (x ⌣ₖ_) (0ₖ-⌣ₖ (suc m) (suc k) z) ∙ ⌣ₖ-0ₖ (suc n) _ x) + , help) + where + h : (n m k : ℕ) (x : coHomK (suc n)) → cong (_⌣ₖ_ x) (⌣ₖ-0ₖ (suc m) (suc k) (0ₖ _)) ≡ + (λ i → x ⌣ₖ 0ₖ-⌣ₖ (suc m) (suc k) (0ₖ _) i) + h zero zero k x = refl + h zero (suc m) k x = refl + h (suc n) zero k x = refl + h (suc n) (suc m) k x = refl + + help : PathP (λ i → (cong (x ⌣ₖ_) (0ₖ-⌣ₖ (suc m) (suc k) (0ₖ _)) ∙ ⌣ₖ-0ₖ (suc n) _ x) i ≡ 0ₖ _) + (cong (x ⌣ₖ_) (⌣ₖ-0ₖ (suc m) (suc k) (0ₖ _)) ∙ ⌣ₖ-0ₖ (suc n) _ x) refl + help = compPathR→PathP + (cong (_∙ ⌣ₖ-0ₖ (suc n) ((suc m) +' (suc k)) x) (h n m k x) + ∙∙ rUnit _ + ∙∙ cong ((cong (x ⌣ₖ_) (0ₖ-⌣ₖ (suc m) (suc k) (0ₖ _)) ∙ ⌣ₖ-0ₖ (suc n) _ x) ∙_) (rUnit refl)) + + assoc2-sub : (n m k : ℕ) → _ → _ + assoc2-sub n m k = subst coHomK (sym (+'-assoc n m k)) + + assoc2-sub-0 : (n m k : ℕ) → assoc2-sub n m k (0ₖ _) ≡ 0ₖ _ + assoc2-sub-0 zero zero zero = refl + assoc2-sub-0 zero zero (suc zero) = refl + assoc2-sub-0 zero zero (suc (suc k)) = refl + assoc2-sub-0 zero (suc zero) zero = refl + assoc2-sub-0 zero (suc zero) (suc k) = refl + assoc2-sub-0 zero (suc (suc m)) zero = refl + assoc2-sub-0 zero (suc (suc m)) (suc k) = refl + assoc2-sub-0 (suc zero) zero zero = refl + assoc2-sub-0 (suc zero) zero (suc k) = refl + assoc2-sub-0 (suc (suc n)) zero zero = refl + assoc2-sub-0 (suc (suc n)) zero (suc k) = refl + assoc2-sub-0 (suc zero) (suc m) zero = refl + assoc2-sub-0 (suc (suc n)) (suc m) zero = refl + assoc2-sub-0 (suc zero) (suc m) (suc k) = refl + assoc2-sub-0 (suc (suc n)) (suc m) (suc k) = refl + + assocer2 : (n m k : ℕ) + → coHomK (suc n) + → coHomK-ptd (suc m) →∙ (coHomK-ptd (suc k) →∙ coHomK-ptd ((suc n) +' ((suc m) +' (suc k))) ∙) + fst (fst (assocer2 n m k x) y) z = + subst coHomK (sym (+'-assoc (suc n) (suc m) (suc k))) ((x ⌣ₖ y) ⌣ₖ z) -- + snd (fst (assocer2 zero m k x) y) = + cong (subst coHomK (sym (+'-assoc 1 (suc m) (suc k)))) (⌣ₖ-0ₖ _ _ (x ⌣ₖ y)) + snd (fst (assocer2 (suc n) m k x) y) = + cong (subst coHomK (sym (+'-assoc (2 + n) (suc m) (suc k)))) (⌣ₖ-0ₖ _ _ (x ⌣ₖ y)) + fst (snd (assocer2 zero m k x) i) z = + subst coHomK (sym (+'-assoc (suc zero) (suc m) (suc k))) (⌣ₖ-0ₖ _ _ x i ⌣ₖ z) + snd (snd (assocer2 zero m k x) i) j = + subst coHomK (sym (+'-assoc (suc zero) (suc m) (suc k))) (⌣ₖ-0ₖ _ _ (⌣ₖ-0ₖ _ _ x i) j) + fst (snd (assocer2 (suc n) m k x) i) z = + subst coHomK (sym (+'-assoc (2 + n) (suc m) (suc k))) (⌣ₖ-0ₖ _ _ x i ⌣ₖ z) + snd (snd (assocer2 (suc n) m k x) i) j = + subst coHomK (sym (+'-assoc (2 + n) (suc m) (suc k))) (⌣ₖ-0ₖ _ _ (⌣ₖ-0ₖ _ _ x i) j) + + assocer-helpFun : (n m : ℕ) → coHomK (suc n) → coHomK-ptd (suc m) →∙ Ω (coHomK-ptd (3 + (n + m))) + fst (assocer-helpFun n m a) b = Kn→ΩKn+1 _ (a ⌣ₖ b) + snd (assocer-helpFun n m a) = cong (Kn→ΩKn+1 _) (⌣ₖ-0ₖ (suc n) (suc m) a) ∙ Kn→ΩKn+10ₖ _ + + assocer-helpFun2 : (n m : ℕ) → coHomK (suc n) → coHomK-ptd (suc m) →∙ Ω (coHomK-ptd (3 + (n + m))) + fst (assocer-helpFun2 n m a) b i = (Kn→ΩKn+1 _ a i) ⌣ₖ b + snd (assocer-helpFun2 n m a) i j = ⌣ₖ-0ₖ _ (suc m) (Kn→ΩKn+1 _ a j) i + +-- Key lemma for associativity +assocer-helpFun≡ : (n m : ℕ) → (x : coHomK (suc n)) → assocer-helpFun n m x ≡ assocer-helpFun2 n m x +assocer-helpFun≡ n m = + trElim (λ _ → isOfHLevelPath (3 + n) (hLev-assocer-helpFun n m) _ _) + λ a → →∙Homogeneous≡ (subst isHomogeneous Kn≃ΩKn+1∙ (isHomogeneousKn _)) + (funExt (main n a)) + where + hLev-assocer-helpFun : (n m : ℕ) → isOfHLevel (3 + n) (coHomK-ptd (suc m) →∙ Ω (coHomK-ptd (3 + (n + m)))) + hLev-assocer-helpFun n m = + subst (isOfHLevel (3 + n)) (cong (coHomK-ptd (suc m) →∙_) + (Kn≃ΩKn+1∙)) + (isOfHLevel↑∙ (suc n) m) + + main : (n : ℕ) (a : S₊ (suc n)) (b : _) + → fst (assocer-helpFun n m ∣ a ∣) b ≡ fst (assocer-helpFun2 n m ∣ a ∣) b + main zero a b k i = + hcomp + (λ r → λ {(i = i0) → 0ₖ _ + ; (i = i1) → ∣ rCancel (merid north) r (~ k) ∣ + ; (k = i0) → Kn→ΩKn+1 _ (∣ a ∣ ⌣ₖ b) i + ; (k = i1) → (Kn→ΩKn+1 _ ∣ a ∣ i) ⌣ₖ b}) + (∣ compPath-filler (merid a) (sym (merid base)) k i ∣ ⌣ₖ b) + main (suc n) a b k i = + hcomp + (λ r → λ {(i = i0) → 0ₖ _ + ; (i = i1) → ∣ rCancel (merid north) r (~ k) ∣ + ; (k = i0) → Kn→ΩKn+1 _ (∣ a ∣ ⌣ₖ b) i + ; (k = i1) → (Kn→ΩKn+1 _ ∣ a ∣ i) ⌣ₖ b}) + (∣ compPath-filler (merid a) (sym (merid north)) k i ∣ ⌣ₖ b) + +assoc-helper : (n m : ℕ) (x : coHomK (suc n)) (y : coHomK (suc m)) + → Kn→ΩKn+1 _ (x ⌣ₖ y) ≡ λ i → (Kn→ΩKn+1 _ x i) ⌣ₖ y +assoc-helper n m x y = funExt⁻ (cong fst (assocer-helpFun≡ n m x)) y + +assoc-⌣ₖ· : (n m k : ℕ) → (x : coHomK (suc n)) → assocer n m k x ≡ assocer2 n m k x +assoc-⌣ₖ· n m k = + trElim (λ _ → isOfHLevelPath (3 + n) + (transport (λ i → isOfHLevel (3 + n) + (coHomK-ptd (suc m) →∙ (coHomK-ptd (suc k) →∙ coHomK-ptd (h (~ i)) ∙))) + (isOfHLevel↑∙∙ m k (suc n))) _ _) + λ a → →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) + (funExt λ b → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt (main n m k a b))) + where + + h : (suc n) +' ((suc m) +' (suc k)) ≡ suc (suc (suc n + m + k)) + h = cong (2 +_) (+-suc n (m + k)) ∙ λ i → suc (suc (suc (+-assoc n m k i))) + + + main : (n m k : ℕ) (a : S₊ (suc n)) (b : coHomK (suc m)) (c : coHomK (suc k)) + → ∣ a ∣ ⌣ₖ b ⌣ₖ c ≡ + (subst coHomK (λ i → +'-assoc (suc n) (suc m) (suc k) (~ i)) + ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + main zero m k a b c = goal a ∙ sym (funExt⁻ t ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + where + t : subst coHomK (λ i → +'-assoc 1 (suc m) (suc k) (~ i)) ≡ idfun _ + t = cong (subst coHomK) (isSetℕ _ _ (+'-assoc 1 (suc m) (suc k)) refl) ∙ funExt transportRefl + goal : (a : _) → ∣ a ∣ ⌣ₖ b ⌣ₖ c ≡ (∣ a ∣ ⌣ₖ b) ⌣ₖ c + goal base = refl + goal (loop i) j = assoc-helper m k b c j i + main (suc n) m k north b c = refl + main (suc n) m k south b c = refl + main (suc n) m k (merid a i) b c j = help2 j i + where + transpLem : (n m : ℕ) (p : n ≡ m) → subst coHomK p (0ₖ _) ≡ 0ₖ _ + transpLem zero m = J (λ m p → subst coHomK p (0ₖ _) ≡ 0ₖ _) refl + transpLem (suc zero) m = J (λ m p → subst coHomK p (0ₖ _) ≡ 0ₖ _) refl + transpLem (suc (suc n)) m = J (λ m p → subst coHomK p (0ₖ _) ≡ 0ₖ _) refl + + transpLem-refl : transpLem (suc (suc (suc (suc (n + m + k))))) + (suc (suc n +' (suc m +' suc k))) + (λ i₃ → +'-assoc (2 + n) (suc m) (suc k) (~ i₃)) ≡ refl + transpLem-refl = transportRefl refl + moveTransports : (n m : ℕ) (x : coHomK n) (p : n ≡ m) (q : (suc n) ≡ suc m) + → PathP (λ i → transpLem _ _ q (~ i) ≡ transpLem _ _ q (~ i)) + (Kn→ΩKn+1 _ (subst coHomK p x)) + (cong (subst coHomK q) (Kn→ΩKn+1 _ x)) + moveTransports n m x = + J (λ m p → (q : (suc n) ≡ suc m) + → PathP (λ i → transpLem _ _ q (~ i) ≡ transpLem _ _ q (~ i)) + (Kn→ΩKn+1 _ (subst coHomK p x)) + (cong (subst coHomK q) (Kn→ΩKn+1 _ x))) + λ q → transport (λ j → PathP (λ i → transpLem (suc n) (suc n) (isSetℕ _ _ refl q j) (~ i) + ≡ transpLem (suc n) (suc n) (isSetℕ _ _ refl q j) (~ i)) + (Kn→ΩKn+1 _ (subst coHomK refl x)) + (cong (subst coHomK (isSetℕ _ _ refl q j)) (Kn→ΩKn+1 _ x))) + (h2 n x) + where + h2 : (n : ℕ) (x : _) + → PathP (λ i₁ → + transpLem (suc n) (suc n) (λ _ → suc n) (~ i₁) ≡ + transpLem (suc n) (suc n) (λ _ → suc n) (~ i₁)) + (Kn→ΩKn+1 n (subst coHomK (λ _ → n) x)) + (λ i₁ → subst coHomK (λ _ → suc n) (Kn→ΩKn+1 n x i₁)) + h2 zero x k i = + hcomp (λ j → λ {(i = i0) → transportRefl (refl {x = 0ₖ _}) (~ j) k + ; (i = i1) → transportRefl (refl {x = 0ₖ _}) (~ j) k + ; (k = i0) → Kn→ΩKn+1 _ (transportRefl x j) i + ; (k = i1) → transportRefl (Kn→ΩKn+1 _ x i) (~ j)}) + (Kn→ΩKn+1 _ x i) + h2 (suc n) x k i = + hcomp (λ j → λ {(i = i0) → transportRefl (refl {x = 0ₖ _}) (~ j) k + ; (i = i1) → transportRefl (refl {x = 0ₖ _}) (~ j) k + ; (k = i0) → Kn→ΩKn+1 _ (transport refl x) i + ; (k = i1) → transportRefl (Kn→ΩKn+1 _ x i) (~ j)}) + (Kn→ΩKn+1 _ (transportRefl x k) i) + + finalTransp : (n : ℕ) (a : _) → Kn→ΩKn+1 _ (subst coHomK (λ i₁ → +'-assoc (suc n) (suc m) (suc k) (~ i₁)) ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + ≡ cong (subst coHomK (λ i₁ → +'-assoc (2 + n) (suc m) (suc k) (~ i₁))) (Kn→ΩKn+1 _ ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + finalTransp n a = + rUnit _ + ∙∙ (λ i → (λ j → transpLem (suc (suc (suc (suc (n + m + k))))) + (suc (suc n +' (suc m +' suc k))) + (λ i₃ → +'-assoc (2 + n) (suc m) (suc k) (~ i₃)) (i ∧ j)) + ∙∙ moveTransports _ _ ((∣ a ∣ ⌣ₖ b) ⌣ₖ c) + (sym (+'-assoc (suc n) (suc m) (suc k))) + (sym (+'-assoc (2 + n) (suc m) (suc k))) i + ∙∙ λ j → transpLem (suc (suc (suc (suc (n + m + k))))) + (suc (suc n +' (suc m +' suc k))) + (λ i₃ → +'-assoc (2 + n) (suc m) (suc k) (~ i₃)) (i ∧ ~ j)) + ∙∙ (λ i → transportRefl refl i + ∙∙ cong (subst coHomK (λ i₁ → +'-assoc (2 + n) (suc m) (suc k) (~ i₁))) (Kn→ΩKn+1 _ ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + ∙∙ transportRefl refl i) + ∙ sym (rUnit _) + + help2 : cong (λ x → (∣ x ∣) ⌣ₖ (b ⌣ₖ c)) (merid a) ≡ cong (assoc2-sub (2 + n) (suc m) (suc k)) (cong (λ x → (∣ x ∣ ⌣ₖ b) ⌣ₖ c) (merid a)) + help2 = ((λ r → Kn→ΩKn+1 _ (main n m k a b c r))) + ∙∙ finalTransp n a + ∙∙ λ r i → subst coHomK (sym (+'-assoc (2 + n) (suc m) (suc k))) (assoc-helper _ _ (∣ a ∣ ⌣ₖ b) c r i) + + +-- Some key distributivity lemmas +-Distₗ : (n m : ℕ) (x : coHomK n) (y : coHomK m) → (-ₖ (x ⌣ₖ y)) ≡ (-ₖ x) ⌣ₖ y +-Distₗ n m x y = + sym (rUnitₖ _ (-ₖ (x ⌣ₖ y))) + ∙∙ cong ((-ₖ (x ⌣ₖ y)) +ₖ_) (sym (rCancelₖ _ (x ⌣ₖ y))) + ∙∙ assocₖ _ (-ₖ (x ⌣ₖ y)) (x ⌣ₖ y) (-ₖ (x ⌣ₖ y)) + ∙∙ cong (_-ₖ (x ⌣ₖ y)) help + ∙∙ sym (assocₖ _ ((-ₖ x) ⌣ₖ y) (x ⌣ₖ y) (-ₖ (x ⌣ₖ y))) + ∙∙ cong ((-ₖ x) ⌣ₖ y +ₖ_) (rCancelₖ _ (x ⌣ₖ y)) + ∙∙ rUnitₖ _ ((-ₖ x) ⌣ₖ y) + where + help : (-ₖ (x ⌣ₖ y)) +ₖ (x ⌣ₖ y) ≡ (-ₖ x) ⌣ₖ y +ₖ (x ⌣ₖ y) + help = lCancelₖ _ _ + ∙∙ sym (0ₖ-⌣ₖ _ _ y) + ∙ cong (_⌣ₖ y) (sym (lCancelₖ _ x)) + ∙∙ rightDistr-⌣ₖ _ _ (-ₖ x) x y + +-Distᵣ : (n m : ℕ) (x : coHomK n) (y : coHomK m) → (-ₖ (x ⌣ₖ y)) ≡ x ⌣ₖ (-ₖ y) +-Distᵣ n m x y = + sym (rUnitₖ _ (-ₖ (x ⌣ₖ y))) + ∙∙ cong ((-ₖ (x ⌣ₖ y)) +ₖ_) (sym (rCancelₖ _ (x ⌣ₖ y))) + ∙∙ assocₖ _ (-ₖ (x ⌣ₖ y)) (x ⌣ₖ y) (-ₖ (x ⌣ₖ y)) + ∙∙ cong (_-ₖ (x ⌣ₖ y)) help + ∙∙ sym (assocₖ _ (x ⌣ₖ (-ₖ y)) (x ⌣ₖ y) (-ₖ (x ⌣ₖ y))) + ∙∙ cong (x ⌣ₖ (-ₖ y) +ₖ_) (rCancelₖ _ (x ⌣ₖ y)) + ∙∙ rUnitₖ _ (x ⌣ₖ (-ₖ y)) + where + help : (-ₖ (x ⌣ₖ y)) +ₖ (x ⌣ₖ y) ≡ x ⌣ₖ (-ₖ y) +ₖ (x ⌣ₖ y) + help = lCancelₖ _ _ + ∙∙ sym (⌣ₖ-0ₖ _ _ x) + ∙∙ cong (x ⌣ₖ_) (sym (lCancelₖ _ y)) + ∙ leftDistr-⌣ₖ _ _ x (-ₖ y) y + +assoc₀ : (m k : ℕ) (x : ℤ) (y : coHomK m) (z : coHomK k) → _⌣ₖ_{n = zero} x (y ⌣ₖ z) ≡ (_⌣ₖ_{n = zero} {m = m} x y) ⌣ₖ z +assoc₀ m k x y z = main x + where + h : subst coHomK (sym (+'-assoc zero m k)) ≡ idfun _ + h = cong (subst coHomK) (isSetℕ _ _ _ refl) ∙ funExt transportRefl + mainPos : (x : ℕ) → _⌣ₖ_ {n = zero} (pos x) (y ⌣ₖ z) ≡ ((_⌣ₖ_ {n = zero} {m = m} (pos x) y) ⌣ₖ z) + mainPos zero = sym (0ₖ-⌣ₖ _ _ z) ∙ cong (_⌣ₖ z) (sym (0ₖ-⌣ₖ _ _ y)) + mainPos (suc n) = + cong (y ⌣ₖ z +ₖ_) (mainPos n) + ∙∙ sym (rightDistr-⌣ₖ _ _ y (_⌣ₖ_ {n = zero} {m = m} (pos n) y) z) + ∙∙ (λ i → (y +ₖ (_⌣ₖ_ {n = zero} {m = m} (pos n) y)) ⌣ₖ z) + + main : (x : ℤ) → x ⌣ₖ y ⌣ₖ z ≡ ((_⌣ₖ_ {n = zero} {m = m} x y) ⌣ₖ z) + main (pos n) = mainPos n + main (negsuc n) = + (λ i → _⌣ₖ_ {n = zero} (+ℤ-comm (negsuc n) 0 i) (y ⌣ₖ z)) + ∙∙ sym (-Distₗ zero _ (pos (suc n)) (y ⌣ₖ z)) + ∙ cong (-ₖ_) (mainPos (suc n)) + ∙∙ -Distₗ _ _ (pos (suc n) ⌣ₖ y) z + ∙∙ cong (_⌣ₖ z) ((-Distₗ zero _ (pos (suc n)) y)) + ∙∙ λ i → (_⌣ₖ_ {n = zero} (+ℤ-comm (negsuc n) 0 (~ i)) y) ⌣ₖ z + +assoc-⌣ₖ : (n m k : ℕ) (x : coHomK n) (y : coHomK m) (z : coHomK k) + → x ⌣ₖ y ⌣ₖ z ≡ subst coHomK (sym (+'-assoc n m k)) ((x ⌣ₖ y) ⌣ₖ z) +assoc-⌣ₖ zero m k x y z = assoc₀ _ _ x y z ∙ sym (funExt⁻ h ((x ⌣ₖ y) ⌣ₖ z)) + where + h : subst coHomK (sym (+'-assoc zero m k)) ≡ idfun _ + h = cong (subst coHomK) (isSetℕ _ _ _ refl) ∙ funExt transportRefl + +assoc-⌣ₖ (suc n) zero k x y z = + help y + ∙∙ sym (transportRefl ((x ⌣ₖ y) ⌣ₖ z)) + ∙∙ λ i → transport (λ j → coHomK ((isSetℕ _ _ ((sym (+'-assoc (suc n) zero k))) refl (~ i) j))) + ((_⌣ₖ_ {m = zero} x y) ⌣ₖ z) + where + helpPos : (y : ℕ) → x ⌣ₖ (_⌣ₖ_ {n = zero} (pos y) z) ≡ + ((_⌣ₖ_ {m = zero} x (pos y)) ⌣ₖ z) + helpPos zero = + (⌣ₖ-0ₖ _ _ x) + ∙∙ (sym (0ₖ-⌣ₖ _ _ z)) + ∙∙ cong (_⌣ₖ z) (sym (⌣ₖ-0ₖ _ _ x)) + helpPos (suc y) = + leftDistr-⌣ₖ _ _ x z (_⌣ₖ_{n = zero} (pos y) z) + ∙∙ cong ((x ⌣ₖ z) +ₖ_) (helpPos y) + ∙∙ sym (rightDistr-⌣ₖ _ _ x (_⌣ₖ_{n = zero} (pos y) x) z) + + help : (y : ℤ) → x ⌣ₖ (_⌣ₖ_ {n = zero} y z) ≡ + ((_⌣ₖ_ {m = zero} x y) ⌣ₖ z) + help (pos n) = helpPos n + help (negsuc n) = + (λ i → x ⌣ₖ (_⌣ₖ_ {n = zero} (+ℤ-comm (negsuc n) 0 i) z)) + ∙∙ cong (x ⌣ₖ_) (sym (-Distₗ zero _ (pos (suc n)) z)) + ∙∙ sym (-Distᵣ _ _ x _) + ∙∙ cong -ₖ_ (helpPos (suc n)) + ∙∙ -Distₗ _ _ _ z + ∙∙ cong (_⌣ₖ z) (-Distᵣ _ zero x (pos (suc n))) + ∙∙ λ i → (_⌣ₖ_ {m = zero} x (+ℤ-comm (negsuc n) 0 (~ i))) ⌣ₖ z +assoc-⌣ₖ (suc n) (suc m) zero x y z = + (assoc-⌣ₖ (suc n) zero (suc m) x z y) + ∙∙ cong (subst coHomK (sym (+'-assoc (suc n) zero (suc m)))) + (sym (assoc₀ _ _ z x y)) + ∙∙ (funExt⁻ (sym h) (_⌣ₖ_ {n = zero} z (x ⌣ₖ y))) + where + h : subst coHomK (sym (+'-assoc (suc n) (suc m) zero)) + ≡ subst coHomK (λ i → +'-assoc (suc n) zero (suc m) (~ i)) + h = cong (subst coHomK) (isSetℕ _ _ _ _) +assoc-⌣ₖ (suc n) (suc m) (suc k) x y z = + funExt⁻ (cong fst (funExt⁻ (cong fst (assoc-⌣ₖ· n m k x)) y)) z + +-- Ring laws for ⌣ +module _ {A : Type ℓ} (n m : ℕ) where + ⌣-0ₕ : (f : coHom n A) → (f ⌣ 0ₕ m) ≡ 0ₕ _ + ⌣-0ₕ = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → ⌣ₖ-0ₖ n m (f x)) + + 0ₕ-⌣ : (f : coHom m A) → (0ₕ n ⌣ f) ≡ 0ₕ _ + 0ₕ-⌣ = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → 0ₖ-⌣ₖ n m (f x)) + + + leftDistr-⌣ : (f : coHom n A) (g h : coHom m A) + → f ⌣ (g +ₕ h) ≡ f ⌣ g +ₕ f ⌣ h + leftDistr-⌣ = + sElim (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ g h → cong ∣_∣₂ (funExt λ x → leftDistr-⌣ₖ n m (f x) (g x) (h x)) + + rightDistr-⌣ : (g h : coHom n A) (f : coHom m A) → (g +ₕ h) ⌣ f ≡ g ⌣ f +ₕ h ⌣ f + rightDistr-⌣ = + sElim2 (λ _ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g h → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → rightDistr-⌣ₖ n m (g x) (h x) (f x)) + + assoc-⌣ : (l : ℕ) (f : coHom n A) (g : coHom m A) (h : coHom l A) + → f ⌣ g ⌣ h ≡ subst (λ x → coHom x A) (sym (+'-assoc n m l)) ((f ⌣ g) ⌣ h) + assoc-⌣ l = + sElim (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 squash₂ _ _) + λ g → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ h → + cong ∣_∣₂ ((funExt (λ x → assoc-⌣ₖ n m l (f x) (g x) (h x) + ∙ cong (subst coHomK (sym (+'-assoc n m l))) + λ i → (f (transportRefl x (~ i)) ⌣ₖ g (transportRefl x (~ i))) + ⌣ₖ (h (transportRefl x (~ i)))))) + +-- Additive unit(s) +0⌣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (x : coHom n A) + → x ⌣ (0ₕ m) ≡ 0ₕ _ +0⌣ n m = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → ⌣ₖ-0ₖ n m (f x)) + +⌣0 : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (x : coHom n A) + → (0ₕ m) ⌣ x ≡ 0ₕ _ +⌣0 n m = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → 0ₖ-⌣ₖ m n (f x)) + +-- Multiplicative unit +1⌣ : ∀ {ℓ} {A : Type ℓ} → coHom 0 A +1⌣ = ∣ (λ _ → 1) ∣₂ + +private + n+'0 : (n : ℕ) → n +' 0 ≡ n + n+'0 zero = refl + n+'0 (suc n) = refl + +lUnit⌣ₖ : (n : ℕ) (x : coHomK n) → _⌣ₖ_ {n = 0} (pos 1) x ≡ x +lUnit⌣ₖ zero = λ _ → refl +lUnit⌣ₖ (suc n) x = rUnitₖ _ x + +lUnit⌣ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (x : coHom n A) + → x ⌣ 1⌣ ≡ subst (λ n → coHom n A) (sym (n+'0 n)) x +lUnit⌣ zero = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt (λ x → comm-·₀ (f x) (pos 1))) ∙ sym (transportRefl ∣ f ∣₂) +lUnit⌣ (suc n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt (λ x → cong (f x +ₖ_) (0ₖ-⌣ₖ zero _ (f x)) ∙ rUnitₖ _ (f x))) + ∙ sym (transportRefl ∣ f ∣₂) + +rUnit⌣ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (x : coHom n A) + → 1⌣ ⌣ x ≡ x +rUnit⌣ zero = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → refl +rUnit⌣ (suc n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → rUnitₖ _ (f x)) + +-ₕDistᵣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) + (x : coHom n A) (y : coHom m A) → (-ₕ (x ⌣ y)) ≡ x ⌣ (-ₕ y) +-ₕDistᵣ n m = + sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → cong ∣_∣₂ (funExt λ x → -Distᵣ n m (f x) (g x)) + +-ₕDistₗ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) + (x : coHom n A) (y : coHom m A) → (-ₕ (x ⌣ y)) ≡ (-ₕ x) ⌣ y +-ₕDistₗ n m = + sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → cong ∣_∣₂ (funExt λ x → -Distₗ n m (f x) (g x)) + +-ₕDistₗᵣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) + (x : coHom n A) (y : coHom m A) → (-ₕ x) ⌣ (-ₕ y) ≡ x ⌣ y +-ₕDistₗᵣ n m x y = + sym (-ₕDistₗ n m x (-ₕ y)) + ∙∙ cong -ₕ_ (sym (-ₕDistᵣ n m x y)) + ∙∙ sElim2 {C = λ x y → (-ₕ (-ₕ (x ⌣ y))) ≡ x ⌣ y} + (λ _ _ → isSetPathImplicit) + (λ f g → cong ∣_∣₂ (funExt λ _ → -ₖ^2 _)) x y + +-- TODO : Graded ring structure diff --git a/Everythings.hs b/Everythings.hs new file mode 100644 index 0000000000..e1e85e61f2 --- /dev/null +++ b/Everythings.hs @@ -0,0 +1,124 @@ +import System.Environment +import System.Directory +import System.Exit +import Control.Monad +import Data.Maybe +import Data.List + +stripSuffix :: (Eq a) => [a] -> [a] -> Maybe [a] +stripSuffix sfx = fmap reverse . stripPrefix (reverse sfx) . reverse + +splitOn :: Char -> String -> [String] +splitOn d = foldr go [] + where go :: Char -> [String] -> [String] + go x acc | d == x = []:acc + go x (a:acc) = (x:a):acc + go x [] = [[x]] + + +type SplitFilePath = [String] + +showFP :: Char -> SplitFilePath -> String +showFP c fp = intercalate [c] (reverse fp) + +addToFP :: SplitFilePath -> String -> SplitFilePath +addToFP fp dir = dir : fp + +-- Given a path to a directory, returns a pair containing the list of all its +-- subdirectories and the list of all agda files it contains +getSubDirsFiles :: SplitFilePath -> IO ([String], [String]) +getSubDirsFiles fp = do + ls <- getDirectoryContents ("./" ++ showFP '/' fp) + let sub_dirs = filter ('.' `notElem`) ls + files = mapMaybe (stripSuffix ".agda") ls + pure (sub_dirs, files) + +-- Given a path to an agda file, returns the list of all files it imports +getImported :: SplitFilePath -> IO [SplitFilePath] +getImported fp = do + ls <- fmap words . lines <$> readFile ("./" ++ showFP '/' fp ++ ".agda") + pure $ fmap (reverse . splitOn '.') (mapMaybe f ls) + where f :: [String] -> Maybe String + f ("open":"import":x:_) = Just x + f ("import":x:_) = Just x + f _ = Nothing + +-- Given a path to a directory $fp and a path to an agda file $fileToCheck.agda, +-- returns the list of all files in* $fp not imported in $fileToCheck.agda +-- * recursively +getMissingFiles :: SplitFilePath -> Maybe SplitFilePath -> IO [SplitFilePath] +getMissingFiles fp fpToCheck = do + (sub_dirs, sub_files) <- getSubDirsFiles fp + -- recursively get all files in $fp/X not imported in $fp/X.agda (if it exists) + missing_files <- concat <$> forM sub_dirs (\sub_dir -> + getMissingFiles (addToFP fp sub_dir) + (addToFP fp <$> find (== sub_dir) sub_files)) + -- return all of these files, plus all agda files in the current directory, + -- except for those which are imported in $fpToCheck.agda (if it exists) or + -- which are $fpToCheck.agda itself + imported <- maybe (pure []) getImported fpToCheck + pure $ ((addToFP fp <$> sub_files) ++ missing_files) + \\ (maybeToList fpToCheck ++ imported) + + +checkEverythings :: [String] -> IO () +checkEverythings dirs = do + missing_files <- concat <$> forM dirs (\dir -> + getMissingFiles [dir,"Cubical"] (Just ["Everything",dir,"Cubical"])) + if null missing_files then pure () + else do putStrLn "Found some files which are not imported in any Everything.agda:" + forM_ missing_files (putStrLn . (" " ++) . showFP '.') + exitFailure + +checkREADME :: IO () +checkREADME = do + (sub_dirs, _) <- getSubDirsFiles ["Cubical"] + imported <- getImported ["README","Cubical"] + let missing_files = fmap (\dir -> ["Everything",dir,"Cubical"]) sub_dirs \\ imported + if null missing_files then pure () + else do putStrLn "Found some Everything.agda's which are not imported in README.agda:" + forM_ missing_files (putStrLn . (" " ++) . showFP '.') + exitFailure + +genEverythings :: [String] -> IO () +genEverythings = + mapM_ (\dir -> do + let fp = addToFP ["Cubical"] dir + files <- getMissingFiles fp Nothing + let ls = ["{-# OPTIONS --safe #-}", + "module " ++ showFP '.' (addToFP fp "Everything") ++ " where",[]] + ++ sort (fmap (\file -> "import " ++ showFP '.' file) + (delete (addToFP fp "Everything") files)) + writeFile ("./" ++ showFP '/' (addToFP fp "Everything") ++ ".agda") + (unlines ls)) + + +helpText :: String +helpText = unlines [ + "Accepted arguments: ", + " check d1 d2 ... dn checks imports in the Everything files in the", + " given directories", + " check-except d1 d2 ... dn checks imports in all Everything files except those", + " in the given directories", + " gen d1 d2 ... dn generates Everything files in the given directories", + " gen-except d1 d2 ... dn generates Everything files in all directories", + " except in those given", + " check-README checks all Everything files are imported in README", + " get-imports-README gets the list of all Everything files imported in README"] + +main :: IO () +main = do + all_dirs <- filter ('.' `notElem`) <$> getDirectoryContents "./Cubical" + args <- getArgs + case args of + "check":dirs -> checkEverythings dirs + "gen" :dirs -> genEverythings dirs + "check-except":ex_dirs -> checkEverythings (all_dirs \\ ex_dirs) + "gen-except" :ex_dirs -> genEverythings (all_dirs \\ ex_dirs) + ["check-README"] -> checkREADME + ["get-imports-README"] -> do + imported <- filter (\fp -> head fp == "Everything") + <$> getImported ["README","Cubical"] + putStrLn . unwords $ map (\fp -> showFP '/' fp ++ ".agda") imported + "help":_ -> putStrLn helpText + _ -> putStrLn "argument(s) not recognized, try 'help'" diff --git a/GNUmakefile b/GNUmakefile new file mode 100644 index 0000000000..3d446d2033 --- /dev/null +++ b/GNUmakefile @@ -0,0 +1,67 @@ +AGDA_BIN?=agda +AGDA_FLAGS?=-W error +AGDA_EXEC?=$(AGDA_BIN) $(AGDA_FLAGS) +FIX_WHITESPACE?=fix-whitespace +RTS_OPTIONS=+RTS -H3G -RTS +AGDA=$(AGDA_EXEC) $(RTS_OPTIONS) +RUNHASKELL?=runhaskell +EVERYTHINGS=$(RUNHASKELL) ./Everythings.hs + +.PHONY : all +all : build + +.PHONY : build +build : + $(MAKE) AGDA_EXEC=$(AGDA_BIN) gen-everythings check + +.PHONY : test +test : check-whitespace gen-and-check-everythings check-README check + +# checking and fixing whitespace + +.PHONY : fix-whitespace +fix-whitespace: + $(FIX_WHITESPACE) + +.PHONY : check-whitespace +check-whitespace: + $(FIX_WHITESPACE) --check + +# checking and generating Everything files + +.PHONY : check-everythings +check-everythings: + $(EVERYTHINGS) check-except Experiments + +.PHONY : gen-everythings +gen-everythings: + $(EVERYTHINGS) gen-except Core Foundations Codata Experiments + +.PHONY : gen-and-check-everythings +gen-and-check-everythings: + $(EVERYTHINGS) gen-except Core Foundations Codata Experiments + $(EVERYTHINGS) check Core Foundations Codata + +.PHONY : check-README +check-README: + $(EVERYTHINGS) check-README + +# typechecking and generating listings for all files imported in README + +.PHONY : check +check: gen-everythings + $(AGDA) Cubical/README.agda + $(AGDA) Cubical/WithK.agda + +.PHONY : timings +timings: clean gen-everythings + $(AGDA) -v profile.modules:10 Cubical/README.agda + +.PHONY : listings +listings: $(wildcard Cubical/**/*.agda) + $(AGDA) -i. -isrc --html Cubical/README.agda -v0 + +.PHONY : clean +clean: + find . -type f -name '*.agdai' -delete + diff --git a/INSTALL.md b/INSTALL.md new file mode 100644 index 0000000000..bfcdb881ec --- /dev/null +++ b/INSTALL.md @@ -0,0 +1,281 @@ +Installation of agda/cubical +============================ + +The cubical library should compile on the latest development version +of Agda: + +https://github.com/agda/agda + +This file contains detailed instruction for how to install this on +Linux and Mac. Windows users might be able to use Cygwin to mimic +these instructions, but this hasn't been tested. + +There are three main ways of installing the development version of Agda: + +1. Using cabal `v2-build`: https://www.haskell.org/cabal/ +2. Using cabal sandboxes: https://www.haskell.org/cabal/ +3. Using stack: https://docs.haskellstack.org/ + +We recommend whichever approach that works for you. This INSTALL file +contains some detailed instructions for installing Agda using either +cabal and stack. For more instructions and general documentation of +Agda see: + +https://agda.readthedocs.io/ + +cabal v2-install instructions +============================= + +cabal `v2-build` is a new operating mode, which makes projects +not interfere with each other. To download and compile the development +version Agda with a `v2-build`, you need to download recent (2.4 or later) +version of [`cabal-install`](https://www.haskell.org/cabal/download.html). +Then, execute following: + +``` +> cabal v2-update +> git clone https://github.com/agda/agda +> cd agda +> touch doc/user-manual.pdf +> cabal v2-install agda agda-mode +``` + +This should put the agda and agda-mode executables in the folder +`~/.cabal/bin` (the location can be configured with `--symlink-bindir` flag). + +In order to be able to access these on your system you need to add +them to your `$PATH` environment variable. On a typical Linux/Mac +installation this can be done by adding + +``` +export PATH=$HOME/.cabal/bin:$PATH +``` + +in your `~/.bashrc` or `~/.bash_profile`. Here `path/to/agda` is the +absolute path to where you cloned the agda repository. In order for +this change to take effect you then have to run + +``` +> source ~/.bashrc +``` + +or + +``` +> source ~/.bash_profile +``` + +or restart the terminal. You should now be able to run: + +``` +> agda --version +``` + +to see that agda has been properly installed and is available in your +`$PATH`. You then also want to setup the agda-mode for emacs: + +``` +> agda-mode setup +``` + +Once this works go to a suitable directory and run + +``` +> git clone https://github.com/agda/cubical +> cd cubical +> make +``` + +This should compile all of the agda/cubical files. To test that it +works in emacs run + +``` +> emacs Cubical/Core/Primitives.agda +``` + +and then type `C-c C-l`. This should now load the file and you can +start developing your own cubical files. + +You can also register cubical as a library to depend on it in your own +Agda developments: + +https://agda.readthedocs.io/en/latest/tools/package-system.html + + +cabal sandbox install instructions +================================== + +In order to have you Agda installation not interfere with any other +Haskell packages that you might have installed it is recommended to +use a sandbox (that installs agda locally in a folder and not globally +on your system). To download and compile the development version Agda +in a cabal sandbox do the following: + +``` +> git clone https://github.com/agda/agda +> cd agda +> cabal sandbox init +> cabal update +> make +``` + +If you have cabal v2 installed the sandbox command should be replaced +by `cabal v1-sandbox init`. + +If the above commands succeed this the agda and agda-mode executables +will be in the folder `agda/.cabal-sandbox/bin`. In order to be able +to access these on your system you need to add them to your `$PATH` +environment variable. On a typical Linux/Mac installation (using Bash +or Zsh as the shell) this can be done by adding + +``` +export PATH=/path/to/agda/.cabal-sandbox/bin:$PATH +``` + +in your `~/.bashrc` or `~/.bash_profile`. Here `path/to/agda` is the +absolute path to where you cloned the agda repository. In order for +this change to take effect you then have to run + +``` +> source ~/.bashrc +``` + +or + +``` +> source ~/.bash_profile +``` + +or restart the terminal. You should now be able to run: + +``` +> agda --version +``` + +to see that agda has been properly installed and is available in your +`$PATH`. You then also want to setup the agda-mode for emacs: + +``` +> agda-mode setup +``` + +Once this works go to a suitable directory and run + +``` +> git clone https://github.com/agda/cubical +> cd cubical +> make +``` + +This should compile all of the agda/cubical files. To test that it +works in emacs run + +``` +> emacs Cubical/Core/Primitives.agda +``` + +and then type `C-c C-l`. This should now load the file and you can +start developing your own cubical files. + +You can also register cubical as a library to depend on it in your own +Agda developments: + +https://agda.readthedocs.io/en/latest/tools/package-system.html + + +stack install instructions +========================== + +In order to install Agda using stack do the following: + +``` +> git clone https://github.com/agda/agda +> cd agda +> stack build --stack-yaml stack-"version".yaml +``` + +Where "version" is a suitable version of ghc (for example 8.6.3). This +should put the agda and agda-mode executables in the folder +`agda/.stack-work/install/.../.../.../bin`. + +where `...` are system and setup dependent. In order to be able to +access these on your system you need to add them to your `$PATH` +environment variable. On a typical Linux/Mac installation this can be +done by adding + +``` +export PATH=/path/to/agda/.stack-work/install/.../.../.../bin:$PATH +``` + +in your `~/.bashrc` or `~/.bash_profile`. Here `path/to/agda` is the +absolute path to where you cloned the agda repository. In order for +this change to take effect you then have to run + +``` +> source ~/.bashrc +``` + +or + +``` +> source ~/.bash_profile +``` + +or restart the terminal. + +You can also run + +``` +> stack install +``` + +which will copy `agda` and `agda-mode` to your `~/.local/bin` folder. +Once these executables are in your PATH you should be able to run: + +``` +> agda --version +``` + +to see that agda has been properly installed and is available in your +`$PATH`. You then also want to setup the agda-mode for emacs: + +``` +> agda-mode setup +``` + +Once this works go to a suitable directory and run + +``` +> git clone https://github.com/agda/cubical +> cd cubical +> make +``` + +This should compile all of the agda/cubical files. To test that it +works in emacs run + +``` +> emacs Cubical/Core/Primitives.agda +``` + +and then type `C-c C-l`. This should now load the file and you can +start developing your own cubical files. + +Registering the cubical library +=============================== + +You can also register cubical as a library to depend on it in your own +Agda developments: + +https://agda.readthedocs.io/en/latest/tools/package-system.html + +On a Linux/Mac installation you need the following files (where +`/path/to/cubical.agda-lib` has been replaced by the path to the +`cubical.agda-lib` file): + +``` +$ cat .agda/defaults +cubical +$ cat .agda/libraries +/path/to/cubical.agda-lib +``` diff --git a/LICENSE b/LICENSE new file mode 100644 index 0000000000..6c4ffd3530 --- /dev/null +++ b/LICENSE @@ -0,0 +1,21 @@ +MIT License + +Copyright (c) 2018 github.com/agda/cubical contributors + +Permission is hereby granted, free of charge, to any person obtaining a copy +of this software and associated documentation files (the "Software"), to deal +in the Software without restriction, including without limitation the rights +to use, copy, modify, merge, publish, distribute, sublicense, and/or sell +copies of the Software, and to permit persons to whom the Software is +furnished to do so, subject to the following conditions: + +The above copyright notice and this permission notice shall be included in all +copies or substantial portions of the Software. + +THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR +IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, +FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE +AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER +LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, +OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE +SOFTWARE. diff --git a/MAKEWINDOWS.md b/MAKEWINDOWS.md new file mode 100644 index 0000000000..296ef7fdc1 --- /dev/null +++ b/MAKEWINDOWS.md @@ -0,0 +1,32 @@ +# Instructions for getting `make` to work on Windows + +The GNU command `make` does not work out of the box on Windows. +Here are some steps you could take to get `make` working on a Windows machine: + +1. You probably already installed the Chocolatey package manager when + installing Agda (check by running `choco` in terminal). If not, install it + from [here](https://chocolatey.org/install). + +2. You need to install the `make` command by running + ``` + choco install make + ``` + +3. The [makefile](GNUmakefile) runs the `runhaskell` command, which might not + work on your system. Test this by running in terminal + ``` + runhaskell + ``` + If that doesn't work, create a file called `runhaskell.bat` with the contents + ``` + runghc %* + ``` + and put this somewhere in your PATH. + +4. Give your user account `Read & Execute` access to `C:\Windows\System32\pthread.dll` + +5. Now, **each time before you run `make`,** run the following command in your terminal: + ``` + chcp.com 65001 + ``` + so that Haskell can parse UTF-8 characters. diff --git a/README.md b/README.md new file mode 100644 index 0000000000..03cc0acbea --- /dev/null +++ b/README.md @@ -0,0 +1,62 @@ +A standard library for Cubical Agda +=================================== + +This library compiles with the master branch of the development +version of [Agda](https://github.com/agda/agda/). For detailed install +instructions see the +[INSTALL](https://github.com/agda/cubical/blob/master/INSTALL.md) +file. + +The source code has a glorious clickable [rendered version](https://agda.github.io/cubical/Cubical.README.html). + +If you want to use Agda 2.6.2 instead of the latest development version, you +can check out the tag `v0.3` of this library. + +If you want to use Agda 2.6.1.3 instead of the latest development version, you +can check out the tag `v0.2` of this library. + +If you want to use Agda 2.6.0.1 instead of the latest development version, you +can check out the tag `v0.1` of this library. + +For some introductory lecture notes see the material for the Cubical Agda course +of the [EPIT 2021 spring school](https://github.com/HoTT/EPIT-2020/blob/main/04-cubical-type-theory/). + +For a paper on with details about Cubical Agda, see [Cubical Agda: a dependently typed +programming language with univalence and higher inductive +types](https://dl.acm.org/doi/10.1145/3341691) by Andrea Vezzosi, Anders +Mörtberg, and Andreas Abel. + +For an introduction to this library, see this [blog +post](https://homotopytypetheory.org/2018/12/06/cubical-agda/). Note that many +files and results have moved since this blog post was written. + +The type theory that Cubical Agda implements is a variation of the +cubical type theory of: + +[Cubical Type Theory: a constructive interpretation of the univalence +axiom](https://arxiv.org/abs/1611.02108) - Cyril Cohen, Thierry +Coquand, Simon Huber, Anders Mörtberg. + + +The key difference is that the Kan composition operations are +decomposed into homogeneous composition and generalized transport as +in: + +[On Higher Inductive Types in Cubical Type +Theory](https://arxiv.org/abs/1802.01170) - Thierry Coquand, Simon +Huber, Anders Mörtberg. + +This makes it possible to directly represent higher inductive types. + +Maintainers +----------- + +* [Anders Mörtberg](https://staff.math.su.se/anders.mortberg/) + +* [Andrea Vezzosi](http://saizan.github.io/) + +* [Evan Cavallo](https://staff.math.su.se/evan.cavallo/) + +[![Build Status](https://travis-ci.org/agda/cubical.svg?branch=master)](https://travis-ci.org/agda/cubical) + +[![Gitter](https://badges.gitter.im/agda/cubical.svg)](https://gitter.im/agda/cubical?utm_source=badge&utm_medium=badge&utm_campaign=pr-badge) diff --git a/RELEASE.md b/RELEASE.md new file mode 100644 index 0000000000..43100ccd61 --- /dev/null +++ b/RELEASE.md @@ -0,0 +1,16 @@ +When releasing a new version of +[agda/cubical](https://github.com/agda/cubical), the following +procedure should be followed: + +Post-release +============ + +* Increment the version number in the name field in the library file + [cubical.agda-lib](cubical.agda-lib): + + ```diff + -name: cubical-0.2 + +name: cubical-0.3 + include: . + depend: + ``` diff --git a/cubical.agda-lib b/cubical.agda-lib new file mode 100644 index 0000000000..4b7b690180 --- /dev/null +++ b/cubical.agda-lib @@ -0,0 +1,4 @@ +name: cubical-0.4 +include: . +depend: +flags: --cubical --no-import-sorts diff --git a/fix-whitespace.yaml b/fix-whitespace.yaml new file mode 100644 index 0000000000..fafbb42c3e --- /dev/null +++ b/fix-whitespace.yaml @@ -0,0 +1,40 @@ +# This file contains the project-specific settings for `fix-whitespace` a tiny +# but useful tool to +# +# * Removes trailing whitespace. +# * Removes trailing lines containing nothing but whitespace. +# * Ensures that the file ends in a newline character. +# +# By default, fix-whitespace checks every directory under the current working +# directory but no files. This program should be placed under a text-based +# project. +# +# For directories, +# +# 1) excluded-dirs is a black-list of directories, +# 2) included-dirs is a white-list of excluded-dirs +# +# For files, +# +# 3) included-files is a white-list of files, +# 4) excluded-files is a black-list of included-files. +# +# The extended glob pattern can be used to specify file/direcotory names. +# For details, see http://hackage.haskell.org/package/filemanip-0.3.6.3/docs/System-FilePath-GlobPattern.html +# + +# Every matched filename is included unless it is matched by excluded-files. +included-files: + - "**.agda" + - "**.cabal" + - "**.el" + - "**.hs" + - "**.hs-boot" + - "**.lagda" + - "**.lhs" + - "**.md" + - "**.rst" + - "**.x" + - "**.y" + - "**.yaml" + - "**.yml" From e8244ac1f3364abe7adb794f27d805d466437524 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 24 Apr 2023 17:10:45 +0200 Subject: [PATCH 02/73] duplicate --- Cubical/ZCohomology/GroupStructure.agda | 10 ---------- 1 file changed, 10 deletions(-) diff --git a/Cubical/ZCohomology/GroupStructure.agda b/Cubical/ZCohomology/GroupStructure.agda index 2af104028f..7944f1d9b2 100644 --- a/Cubical/ZCohomology/GroupStructure.agda +++ b/Cubical/ZCohomology/GroupStructure.agda @@ -696,16 +696,6 @@ leftInv (fst (coHomIso n is)) = ST.elim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (rightInv is x)) snd (coHomIso n is) = snd (coHomMorph n (fun is)) -coHomIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) → Iso A B - → GroupIso (coHomGr n B) (coHomGr n A) -fun (fst (coHomIso n is)) = fst (coHomMorph n (fun is)) -inv' (fst (coHomIso n is)) = fst (coHomMorph n (inv' is)) -rightInv (fst (coHomIso n is)) = - sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (leftInv is x)) -leftInv (fst (coHomIso n is)) = - sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (rightInv is x)) -snd (coHomIso n is) = snd (coHomMorph n (fun is)) - -- Alternative definition of cohomology using ΩKₙ instead. Useful for breaking proofs of group isos -- up into smaller parts coHomGrΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Group ℓ From aadde33a93f44d140376682ec774a7d460ac183f Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 24 Apr 2023 17:12:41 +0200 Subject: [PATCH 03/73] wups --- Cubical/Homotopy/Group/Pi4S3/Summary.agda | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Cubical/Homotopy/Group/Pi4S3/Summary.agda b/Cubical/Homotopy/Group/Pi4S3/Summary.agda index e6a5877520..ad0238341e 100644 --- a/Cubical/Homotopy/Group/Pi4S3/Summary.agda +++ b/Cubical/Homotopy/Group/Pi4S3/Summary.agda @@ -66,6 +66,7 @@ open import Cubical.Algebra.Group.ZAction β : ℕ β = Brunerie + -- The connection to π₄(S³) is then also proved in the BrunerieNumber -- file following Corollary 3.4.5 in Guillaume Brunerie's PhD thesis. βSpec : GroupEquiv (π 4 𝕊³) (ℤGroup/ β) @@ -94,6 +95,7 @@ _ = π₂S³-gen-by-HopfMap -- etc. For more details see the proof of "Brunerie≡2". + -- Combining all of this gives us the desired equivalence of groups: π₄S³≃ℤ/2ℤ : GroupEquiv (π 4 𝕊³) (ℤGroup/ 2) π₄S³≃ℤ/2ℤ = subst (GroupEquiv (π 4 𝕊³)) (cong ℤGroup/_ β≡2) βSpec @@ -135,4 +137,3 @@ _ = π₂S³-gen-by-HopfMap -- computation as conjectured in Brunerie's thesis: π₄S³≃ℤ/2ℤ-computation : GroupEquiv (π 4 𝕊³) (ℤGroup/ 2) π₄S³≃ℤ/2ℤ-computation = DirectProof.BrunerieGroupEquiv'' ->>>>>>> agda/master From 17f3fc1f5dbed5cf98bb5f8f58d9d4ed56aa81b5 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 24 Apr 2023 17:13:52 +0200 Subject: [PATCH 04/73] rme --- Cubical/Homotopy/Group/Pi4S3/Tricky.agda | 772 ----------------- Cubical/Homotopy/HopfInvariant/Whitehead.agda | 808 ------------------ 2 files changed, 1580 deletions(-) delete mode 100644 Cubical/Homotopy/Group/Pi4S3/Tricky.agda delete mode 100644 Cubical/Homotopy/HopfInvariant/Whitehead.agda diff --git a/Cubical/Homotopy/Group/Pi4S3/Tricky.agda b/Cubical/Homotopy/Group/Pi4S3/Tricky.agda deleted file mode 100644 index c9ca909f0d..0000000000 --- a/Cubical/Homotopy/Group/Pi4S3/Tricky.agda +++ /dev/null @@ -1,772 +0,0 @@ -{-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Homotopy.Group.Pi4S3.Tricky where - -open import Cubical.Homotopy.Loopspace - -open import Cubical.Homotopy.Group.Base -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Pointed.Homogeneous -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) -open import Cubical.Foundations.Path -open import Cubical.Foundations.Isomorphism -open Iso -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Univalence -open import Cubical.Functions.Morphism - -open import Cubical.Data.Unit -open import Cubical.Data.Empty renaming (rec to ⊥-rec) -open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 - ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 - ; map to sMap) -open import Cubical.HITs.Sn -open import Cubical.HITs.Susp renaming (toSusp to σ) -open import Cubical.HITs.S1 hiding (decode ; encode) - -open import Cubical.Data.Sigma -open import Cubical.Data.Nat -open import Cubical.Data.Bool - -open import Cubical.Algebra.Group renaming (Unit to UnitGr ; Bool to BoolGr) -open import Cubical.Algebra.Semigroup -open import Cubical.Algebra.Monoid -open import Cubical.Algebra.Group.ZAction - -open import Cubical.HITs.Join -open import Cubical.HITs.Pushout -open import Cubical.HITs.Wedge -open import Cubical.Homotopy.Freudenthal hiding (Code ; encode) -open import Cubical.Homotopy.Connected -open import Cubical.HITs.Truncation renaming - (rec to trRec ; elim to trElim ; elim2 to trElim2 ; map to trMap) -open import Cubical.Foundations.Function -open import Cubical.HITs.S2 - -open import Cubical.Homotopy.BlakersMassey -open import Cubical.Homotopy.Whitehead - -open import Cubical.HITs.PropositionalTruncation - renaming (rec to pRec ; elim to pElim ; map to pMap) - - -W' = joinTo⋁ {A = S₊∙ 1} {B = S₊∙ 1} - -W : S₊ 3 → (S₊∙ 2 ⋁ S₊∙ 2) -W = W' ∘ Iso.inv (IsoSphereJoin 1 1) - -fold∘W : S₊ 3 → S₊ 2 -fold∘W = fold⋁ ∘ W - -thePushout : Type -thePushout = Pushout (λ _ → tt) fold∘W - -isConnectedS3→S2 : (f : S₊ 3 → S₊ 2) → isConnectedFun 2 f -isConnectedS3→S2 f p = - trRec (isProp→isOfHLevelSuc 1 isPropIsContr) - (J (λ p _ → isConnected 2 (fiber f p)) - (∣ north , refl ∣ - , (trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) - (uncurry (sphereElim 2 (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) - λ p → trRec (isOfHLevelTrunc 2 _ _) (λ r → cong ∣_∣ₕ (ΣPathP (refl , r))) (fun (PathIdTruncIso 1) (isContr→isProp (isConnectedPath 2 (sphereConnected 2) (f north) (f north)) ∣ refl ∣ ∣ p ∣))))))) - (fun (PathIdTruncIso 2) - (isContr→isProp (sphereConnected 2) ∣ f north ∣ ∣ p ∣)) - -fibreUnitMap : ∀ {ℓ} {A : Type ℓ} → Iso (fiber (λ (a : A) → tt) tt) A -fun fibreUnitMap (x , p) = x -inv fibreUnitMap a = a , refl -rightInv fibreUnitMap _ = refl -leftInv fibreUnitMap _ = refl - -isConnectedUnitmap : ∀ {ℓ} {A : Type ℓ} (n : ℕ) - → isConnected n A - → isConnectedFun n (λ (a : A) → tt) -isConnectedUnitmap n con tt = - subst (isConnected n) (isoToPath (invIso fibreUnitMap)) con - -module BM-inst = - BlakersMassey□ (λ _ → tt) fold∘W 3 - 1 - (isConnectedUnitmap 4 (sphereConnected 3)) - (isConnectedS3→S2 fold∘W) - -open BM-inst - -Grr : S₊ 3 → Σ (Unit × S₊ 2) PushoutPath× -Grr = toPullback - -thePullback∙ : Pointed₀ -fst thePullback∙ = Σ (Unit × S₊ 2) PushoutPath× -snd thePullback∙ = (tt , north) , (push north) - -isConGrr : isConnectedFun 4 Grr -isConGrr = isConnected-toPullback - -open import Cubical.Homotopy.Group.LES - -thePushout∙ : Pointed₀ -thePushout∙ = thePushout , inl tt - -inr' : S₊ 2 → thePushout -inr' = inr - -fibreinr' : Iso (fiber inr' (inl tt)) (Σ (Unit × S₊ 2) PushoutPath×) -fun fibreinr' (x , p) = (tt , x) , (sym p) -inv fibreinr' ((tt , x) , p) = x , (sym p) -rightInv fibreinr' _ = refl -leftInv fibreinr' _ = refl - -TotalPushoutPath×∙ : Pointed ℓ-zero -fst TotalPushoutPath×∙ = Σ (Unit × S₊ 2) PushoutPath× -snd TotalPushoutPath×∙ = (tt , north) , push north - -P : Pointed₀ -P = (fiber inr' (inl tt) , north , (sym (push north))) - - -π'P→π'TotalPath× : (n : ℕ) → GroupEquiv (π'Gr n TotalPushoutPath×∙) (π'Gr n P) -fst (fst (π'P→π'TotalPath× n)) = - π'eqFun n ((invEquiv (isoToEquiv fibreinr')) , refl) -snd (fst (π'P→π'TotalPath× n)) = π'eqFunIsEquiv n _ -snd (π'P→π'TotalPath× n) = π'eqFunIsHom n _ - -inr∙ : S₊∙ 2 →∙ thePushout∙ -fst inr∙ = inr' -snd inr∙ = sym (push north) - -module LESinst = πLES {A = S₊∙ 2} {B = thePushout∙} inr∙ - - -≃∙→πHom : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} - → (A ≃∙ B) - → (n : ℕ) → GroupEquiv (πGr n A) (πGr n B) -fst (fst (≃∙→πHom e n)) = fst (πHom n (≃∙map e)) -snd (fst (≃∙→πHom e n)) = - isoToIsEquiv (setTruncIso (equivToIso (_ , isEquivΩ^→ (suc n) (≃∙map e) (snd (fst e))))) -snd (≃∙→πHom e n) = snd (πHom n (≃∙map e)) - -isConnectedΩ→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) - → (f : A →∙ B) - → isConnectedFun (suc n) (fst f) - → isConnectedFun n (fst (Ω→ f)) -isConnectedΩ→ n f g = - transport (λ i → isConnectedFun n λ b - → doubleCompPath-filler (sym (snd f)) (cong (fst f) b) (snd f) i) - (isConnectedCong n _ g) - -isConnectedΩ^→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n m : ℕ) - → (f : A →∙ B) - → isConnectedFun (suc n) (fst f) - → isConnectedFun ((suc n) ∸ m) (fst (Ω^→ m f)) -isConnectedΩ^→ n zero f conf = conf -isConnectedΩ^→ n (suc zero) f conf = isConnectedΩ→ n f conf -isConnectedΩ^→ {A = A} {B = B} n (suc (suc m)) f conf = - transport (λ i → isConnectedFun (suc n ∸ suc (suc m)) - λ q → doubleCompPath-filler - (sym (snd (Ω^→ (suc m) f))) - (cong (fst (Ω^→ (suc m) f)) q) - (snd (Ω^→ (suc m) f)) i) - (help n m (help2 n (suc m)) (isConnectedΩ^→ n (suc m) f conf)) - where - open import Cubical.Data.Sum - help2 : (n m : ℕ) → ((suc n ∸ m ≡ suc (n ∸ m))) ⊎ (suc n ∸ suc m ≡ 0) - help2 n zero = inl refl - help2 zero (suc m) = inr refl - help2 (suc n) (suc m) = help2 n m - - help : (n m : ℕ) → ((suc n ∸ (suc m) ≡ suc (n ∸ suc m))) ⊎ (suc n ∸ suc (suc m) ≡ 0) - → isConnectedFun (suc n ∸ suc m) (fst (Ω^→ (suc m) f)) - → isConnectedFun (suc n ∸ suc (suc m)) - {A = Path ((Ω^ (suc m)) (_ , pt A) .fst) - refl refl} - (cong (fst (Ω^→ (suc m) f))) - help n m (inl x) conf = - isConnectedCong (n ∸ suc m) (fst (Ω^→ (suc m) f)) - (subst (λ x → isConnectedFun x (fst (Ω^→ (suc m) f))) x conf) - help n m (inr x) conf = - subst (λ x → isConnectedFun x (cong {x = refl} {y = refl} (fst (Ω^→ (suc m) f)))) - (sym x) - λ b → tt* , λ _ → refl - -π₂P≅0 : GroupEquiv (πGr 1 P) UnitGr -π₂P≅0 = compGroupEquiv (≃∙→πHom (isoToEquiv fibreinr' , refl) 1) - (GroupIso→GroupEquiv - (contrGroupIsoUnit (isOfHLevelRetractFromIso 0 (invIso map23) isContrπ₂S³))) - where - mapp : Iso (hLevelTrunc 4 (S₊ 3)) (hLevelTrunc 4 (Σ (Unit × S₊ 2) PushoutPath×)) - mapp = connectedTruncIso 4 Grr isConnected-toPullback - - map23 : Iso (π 2 (hLevelTrunc∙ 4 (S₊∙ 3))) (π 2 thePullback∙) - map23 = - (compIso (setTruncIso - (equivToIso (_ , - (isEquivΩ^→ 2 (fun mapp , refl) (isoToIsEquiv mapp))))) - (invIso (πTruncIso 2))) - - zz : isContr (π 2 (Unit , tt)) - fst zz = ∣ refl ∣₂ - snd zz = sElim (λ _ → isSetPathImplicit) λ p → refl - - isContrπ₂S³ : isContr (π 2 (hLevelTrunc∙ 4 (S₊∙ 3))) - isContrπ₂S³ = - subst (λ x → isContr (π 2 x)) - (λ i → ((sym (isContr→≡Unit (sphereConnected 3))) i) - , transp (λ j → isContr→≡Unit - (sphereConnected 3) (~ i ∧ j)) i ∣ north ∣) - zz - -testComm : (λ x → snd (fst x)) ∘ Grr ≡ fold∘W -testComm = refl --- nice - --- G → H → L → R -record exact {ℓ ℓ' ℓ'' ℓ''' : Level} (G : Group ℓ) (H : Group ℓ') (L : Group ℓ'') (R : Group ℓ''') - (G→H : GroupHom G H) (H→L : GroupHom H L) (L→R : GroupHom L R) - : Type (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ'''))) where - field - ImG→H⊂KerH→L : (x : fst H) → isInIm G→H x → isInKer H→L x - KerH→L⊂ImG→H : (x : fst H) → isInKer H→L x → isInIm G→H x - - ImH→L⊂KerL→R : (x : fst L) → isInIm H→L x → isInKer L→R x - KerL→R⊂ImH→L : (x : fst L) → isInKer L→R x → isInIm H→L x - -P→S²→Pushout : - exact (πGr 2 P) (πGr 2 (S₊∙ 2)) (πGr 2 thePushout∙) (πGr 1 P) - (LESinst.fib→A 2) - (LESinst.A→B 2) - (LESinst.B→fib 1) -exact.ImG→H⊂KerH→L P→S²→Pushout = LESinst.Im-fib→A⊂Ker-A→B 2 -exact.KerH→L⊂ImG→H P→S²→Pushout = LESinst.Ker-A→B⊂Im-fib→A 2 -exact.ImH→L⊂KerL→R P→S²→Pushout = LESinst.Im-A→B⊂Ker-B→fib 1 -exact.KerL→R⊂ImH→L P→S²→Pushout = LESinst.Ker-B→fib⊂Im-A→B 1 - -transportExact : {ℓ ℓ' ℓ'' ℓ''' : Level} - (G G₂ : Group ℓ) (H H₂ : Group ℓ') (L L₂ : Group ℓ'') (R R₂ : Group ℓ''') - → (G≡G₂ : G ≡ G₂) → (H≡H₂ : H ≡ H₂) (L≡L₂ : L ≡ L₂) (R≡R₂ : R ≡ R₂) - → (G→H : GroupHom G H) (G₂→H₂ : GroupHom G₂ H₂) - → (H→L : GroupHom H L) (H₂→L₂ : GroupHom H₂ L₂) - → (L→R : GroupHom L R) - → exact G H L R G→H H→L L→R - → PathP (λ i → GroupHom (G≡G₂ i) (H≡H₂ i)) G→H G₂→H₂ - → PathP (λ i → GroupHom (H≡H₂ i) (L≡L₂ i)) H→L H₂→L₂ - → Σ[ L₂→R₂ ∈ GroupHom L₂ R₂ ] - exact G₂ H₂ L₂ R₂ G₂→H₂ H₂→L₂ L₂→R₂ -transportExact G G₂ H H₂ L L₂ R R₂ = - J4 (λ G₂ H₂ L₂ R₂ G≡G₂ H≡H₂ L≡L₂ R≡R₂ - → (G→H : GroupHom G H) (G₂→H₂ : GroupHom G₂ H₂) - → (H→L : GroupHom H L) (H₂→L₂ : GroupHom H₂ L₂) - → (L→R : GroupHom L R) - → exact G H L R G→H H→L L→R - → PathP (λ i → GroupHom (G≡G₂ i) (H≡H₂ i)) G→H G₂→H₂ - → PathP (λ i → GroupHom (H≡H₂ i) (L≡L₂ i)) H→L H₂→L₂ - → Σ[ L₂→R₂ ∈ GroupHom L₂ R₂ ] - exact G₂ H₂ L₂ R₂ G₂→H₂ H₂→L₂ L₂→R₂) - (λ G→H G₂→H₂ H→L H₂→L₂ L→R ex - → J (λ G₂→H₂ _ - → H→L ≡ H₂→L₂ - → Σ[ L→R ∈ GroupHom L R ] - exact G H L R G₂→H₂ H₂→L₂ L→R) - (J (λ H₂→L₂ _ → Σ[ L→R ∈ GroupHom L R ] - exact G H L R G→H H₂→L₂ L→R) - (L→R , ex))) - G₂ H₂ L₂ R₂ - where - abstract - J4 : ∀ {ℓ ℓ₂ ℓ₃ ℓ₄ ℓ'} {A : Type ℓ} {A₂ : Type ℓ₂} {A₃ : Type ℓ₃} {A₄ : Type ℓ₄} - {x : A} {x₂ : A₂} {x₃ : A₃} {x₄ : A₄} - (B : (y : A) (z : A₂) (w : A₃) (u : A₄) → x ≡ y → x₂ ≡ z → x₃ ≡ w → x₄ ≡ u → Type ℓ') - → B x x₂ x₃ x₄ refl refl refl refl - → (y : A) (z : A₂) (w : A₃) (u : A₄) (p : x ≡ y) (q : x₂ ≡ z) (r : x₃ ≡ w) (s : x₄ ≡ u) - → B y z w u p q r s - J4 {x = x} {x₂ = x₂} {x₃ = x₃} {x₄ = x₄} B b y z w u = - J (λ y p → (q : x₂ ≡ z) (r : x₃ ≡ w) (s : x₄ ≡ u) → - B y z w u p q r s) - (J (λ z q → (r : x₃ ≡ w) (s : x₄ ≡ u) → B x z w u refl q r s) - (J (λ w r → (s : x₄ ≡ u) → B x x₂ w u refl refl r s) - (J (λ u s → B x x₂ x₃ u refl refl refl s) b))) - -ΣP→S²→Pushout' : - Σ[ F ∈ GroupHom (π'Gr 2 thePushout∙) (π'Gr 1 P) ] - (exact (π'Gr 2 P) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) (π'Gr 1 P) - (π'∘∙Hom 2 (fst , refl)) - (π'∘∙Hom 2 inr∙) - F) -ΣP→S²→Pushout' = - transportExact _ _ _ _ _ _ _ _ - (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 2 P))))) - (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 2 (S₊∙ 2)))))) - (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 2 thePushout∙))))) - (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 1 P))))) - _ _ _ _ _ - P→S²→Pushout - (toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (fromPathP λ i → fst (π∘∙fib→A-PathP 2 inr∙ i)))) - ((toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (fromPathP λ i → fst (π∘∙A→B-PathP 2 inr∙ i))))) - -abstract - π₂thePushout→π₁P : GroupHom (π'Gr 2 thePushout∙) (π'Gr 1 P) - π₂thePushout→π₁P = fst ΣP→S²→Pushout' - - P→S²→Pushout→P' : - exact (π'Gr 2 P) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) (π'Gr 1 P) - (π'∘∙Hom 2 (fst , refl)) - (π'∘∙Hom 2 inr∙) - π₂thePushout→π₁P - P→S²→Pushout→P' = ΣP→S²→Pushout' .snd - -→UnitHom : ∀ {ℓ} (G : Group ℓ) → GroupHom G UnitGr -fst (→UnitHom G) _ = tt -snd (→UnitHom G) = makeIsGroupHom λ _ _ → refl - -compSurj : ∀ {ℓ ℓ' ℓ''} {G : Group ℓ} {H : Group ℓ'} {L : Group ℓ''} - → (G→H : GroupHom G H) (H→L : GroupHom H L) - → isSurjective G→H → isSurjective H→L - → isSurjective (compGroupHom G→H H→L) -compSurj G→H H→L surj1 surj2 l = - pRec squash - (λ {(h , p) - → pMap (λ {(g , q) → g , (cong (fst H→L) q ∙ p)}) - (surj1 h)}) - (surj2 l) - -π₃S³→π₃PΩ : GroupHom (πGr 2 (S₊∙ 3)) (πGr 2 TotalPushoutPath×∙) -π₃S³→π₃PΩ = πHom 2 (Grr , refl) - -TotalPushoutPath×∙→P : TotalPushoutPath×∙ →∙ P -fst TotalPushoutPath×∙→P x = (snd (fst x)) , (sym (snd x)) -snd TotalPushoutPath×∙→P = refl - -isSurjective-π₃S³→π₃PΩ : isSurjective π₃S³→π₃PΩ -isSurjective-π₃S³→π₃PΩ = - sElim (λ _ → isProp→isSet squash) - λ p → trRec squash - (λ s → ∣ ∣ fst s ∣₂ , (cong ∣_∣₂ (snd s)) ∣) - (((isConnectedΩ^→ 3 3 (Grr , refl) isConGrr) p) .fst) - - -π₃S³→π₃P : GroupHom (π'Gr 2 (S₊∙ 3)) (π'Gr 2 TotalPushoutPath×∙) -π₃S³→π₃P = π'∘∙Hom 2 (Grr , refl) - -transportLem : PathP (λ i → GroupHomπ≅π'PathP (S₊∙ 3) TotalPushoutPath×∙ 2 i) - π₃S³→π₃PΩ π₃S³→π₃P -transportLem = - toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (π'∘∙Hom'≡π'∘∙fun {A = S₊∙ 3} {B = TotalPushoutPath×∙} 2 (Grr , refl))) - -isSurjective-π₃S³→π₃P : isSurjective π₃S³→π₃P -isSurjective-π₃S³→π₃P = - transport (λ i → isSurjective (transportLem i)) - isSurjective-π₃S³→π₃PΩ - -π₃S³→π₃S²' : GroupHom (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) -π₃S³→π₃S²' = compGroupHom π₃S³→π₃P (π'∘∙Hom 2 ((λ x → snd (fst x)) , refl)) - -test : π₃S³→π₃S²' ≡ π'∘∙Hom 2 (fold∘W , refl) -test = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (sElim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (ΣPathP (refl , - (λ j → (λ i → snd (fst ((rUnit (λ k → Grr (snd f k)) (~ j)) i))) - ∙ (λ _ → north)))))) - -amap : TotalPushoutPath×∙ →∙ P -fst amap ((tt , s), p) = s , (sym p) -snd amap = refl - -π₃S³→π₃P' : GroupHom (π'Gr 2 (S₊∙ 3)) (π'Gr 2 P) -π₃S³→π₃P' = compGroupHom π₃S³→π₃P (π'∘∙Hom 2 amap) - -isSurjective-π₃S³→π₃P' : isSurjective π₃S³→π₃P' -isSurjective-π₃S³→π₃P' = - compSurj π₃S³→π₃P (π'∘∙Hom 2 amap) - isSurjective-π₃S³→π₃P - (sElim (λ _ → isProp→isSet squash) - λ {(s , p) → ∣ ∣ (λ x → (tt , s x .fst) , sym (s x .snd)) - , ΣPathP ((ΣPathP (refl , (cong fst p))) - , (λ i j → snd (p i) (~ j))) ∣₂ - , cong ∣_∣₂ (ΣPathP (refl , sym (rUnit p))) ∣}) - --- π₃P → π₃S² → π₃ Pushout → Unit -P→S²→Pushout→Unit : - exact (π'Gr 2 P) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr - (π'∘∙Hom 2 (fst , refl)) - (π'∘∙Hom 2 inr∙) - (→UnitHom (π'Gr 2 thePushout∙)) -P→S²→Pushout→Unit = - transport (λ i → - exact (π'Gr 2 P) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) - (GroupPath _ _ .fst - (compGroupEquiv (GroupIso→GroupEquiv (π'Gr≅πGr 1 P)) π₂P≅0) i) - (π'∘∙Hom 2 (fst , refl)) - (π'∘∙Hom 2 inr∙) - (r i)) - P→S²→Pushout→P' - where - r : PathP (λ i → GroupHom (π'Gr 2 thePushout∙) - ((GroupPath _ _ .fst - (compGroupEquiv (GroupIso→GroupEquiv (π'Gr≅πGr 1 P)) π₂P≅0) i))) - π₂thePushout→π₁P - (→UnitHom (π'Gr 2 thePushout∙)) - r = toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt λ _ → refl)) -open exact - -extendExactSurj : {ℓ ℓ' ℓ'' ℓ''' ℓ'''' : Level} - (G : Group ℓ) (H : Group ℓ') (L : Group ℓ'') (R : Group ℓ''') (S : Group ℓ'''') - (G→H : GroupHom G H) (H→L : GroupHom H L) (L→R : GroupHom L R) (R→S : GroupHom R S) - → isSurjective G→H - → exact H L R S H→L L→R R→S - → exact G L R S (compGroupHom G→H H→L) L→R R→S -ImG→H⊂KerH→L (extendExactSurj G H L R S G→H H→L L→R R→S surj ex) x = - pRec (GroupStr.is-set (snd R) _ _) - (uncurry λ g → J (λ x _ → isInKer L→R x) - (ImG→H⊂KerH→L ex (fst H→L (fst G→H g)) - ∣ (fst G→H g) , refl ∣)) -KerH→L⊂ImG→H (extendExactSurj G H L R S G→H H→L L→R R→S surj ex) x ker = - pRec squash - (uncurry λ y → J (λ x _ → isInIm (compGroupHom G→H H→L) x) - (pMap (uncurry - (λ y → J (λ y _ → Σ[ g ∈ fst G ] fst H→L (fst G→H g) ≡ H→L .fst y) - (y , refl))) (surj y))) - (KerH→L⊂ImG→H ex x ker) -ImH→L⊂KerL→R (extendExactSurj G H L R S G→H H→L L→R R→S surj ex) = - ImH→L⊂KerL→R ex -KerL→R⊂ImH→L (extendExactSurj G H L R S G→H H→L L→R R→S surj ex) = - KerL→R⊂ImH→L ex - -P'→S²→Pushout→Unit' : - exact (π'Gr 2 TotalPushoutPath×∙) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr - (compGroupHom (π'∘∙Hom 2 TotalPushoutPath×∙→P) (π'∘∙Hom 2 (fst , refl))) - (π'∘∙Hom 2 inr∙) - (→UnitHom (π'Gr 2 thePushout∙)) -P'→S²→Pushout→Unit' = - extendExactSurj _ _ _ _ _ _ _ _ _ - (sElim (λ _ → isProp→isSet squash) - (λ f → ∣ ∣ (λ x → (tt , fst f x .fst) , sym (fst f x .snd)) - , ΣPathP ((ΣPathP (refl , cong fst (snd f))) , λ j i → snd f j .snd (~ i)) ∣₂ - , cong ∣_∣₂ (ΣPathP (refl , sym (rUnit _))) ∣)) - P→S²→Pushout→Unit - - -S³→S²→Pushout→Unit'' : - exact (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr - (compGroupHom π₃S³→π₃P - (compGroupHom (π'∘∙Hom 2 TotalPushoutPath×∙→P) (π'∘∙Hom 2 (fst , refl)))) - (π'∘∙Hom 2 inr∙) - (→UnitHom (π'Gr 2 thePushout∙)) -S³→S²→Pushout→Unit'' = - extendExactSurj _ _ _ _ _ _ _ _ _ - isSurjective-π₃S³→π₃P P'→S²→Pushout→Unit' - -tripleComp≡ : - (compGroupHom π₃S³→π₃P - (compGroupHom - (π'∘∙Hom 2 TotalPushoutPath×∙→P) (π'∘∙Hom 2 (fst , refl)))) - ≡ π'∘∙Hom 2 (fold∘W , refl) -tripleComp≡ = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (sElim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (ΣPathP (refl , (cong (_∙ refl) - (λ j → cong fst (rUnit (cong (fst TotalPushoutPath×∙→P) - (rUnit (cong Grr (snd f)) (~ j))) (~ j)))))))) - -S³→S²→Pushout→Unit : - exact (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr - (π'∘∙Hom 2 (fold∘W , refl)) - (π'∘∙Hom 2 inr∙) - (→UnitHom (π'Gr 2 thePushout∙)) -S³→S²→Pushout→Unit = - subst - (λ F → exact (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) UnitGr - F (π'∘∙Hom 2 inr∙) - (→UnitHom (π'Gr 2 thePushout∙))) - tripleComp≡ - S³→S²→Pushout→Unit'' - -indΠ₃S₂ : ∀ {ℓ} {A : Pointed ℓ} → (f g : A →∙ S₊∙ 2) → fst f ≡ fst g → ∥ f ≡ g ∥ -indΠ₃S₂ {A = A} f g p = - trRec squash - (λ r → ∣ ΣPathP (p , r) ∣) - (isConnectedPathP 1 {A = (λ i → p i (snd A) ≡ north)} - (isConnectedPathSⁿ 1 (fst g (pt A)) north) (snd f) (snd g) .fst ) - -Fold∘W : fst (π'∘∙Hom 2 (fold∘W , refl)) ∣ idfun∙ (S₊∙ 3) ∣₂ - ≡ ∣ [ idfun∙ (S₊∙ 2) ∣ idfun∙ (S₊∙ 2) ]₂ ∣₂ -Fold∘W = - pRec (squash₂ _ _) - (cong ∣_∣₂) - (indΠ₃S₂ _ _ - (funExt (λ x → funExt⁻ (sym (cong fst (id∨→∙id {A = S₊∙ 2}))) (W x)))) - --- Todo: generalise this - --- ℤ --f--> ℤ --g--> ℤ ---> 0 - -open import Cubical.Data.Int renaming (ℤ to Z ; _·_ to _·Z_ ; _+_ to _+Z_) -module exact→BoolIso (G : Group₀) (F : GroupHom ℤ ℤ) (H : GroupHom ℤ G) (P : fst F 1 ≡ 2) - (ex : exact ℤ ℤ G UnitGr F H (→UnitHom G)) where - - strG = snd G - strH = snd H - open GroupStr - - Bool→G : Bool → fst G - Bool→G false = fst H 1 - Bool→G true = 1g strG - - isSurjectiveH : isSurjective H - isSurjectiveH b = exact.KerL→R⊂ImH→L ex b refl - - Bool→GHom : IsGroupHom (snd BoolGr) Bool→G strG - Bool→GHom = - makeIsGroupHom - λ { false false → sym (exact.ImG→H⊂KerH→L ex 2 ∣ 1 , P ∣) - ∙ IsGroupHom.pres· (snd H) 1 1 - ; false true → sym (rid (snd G) _) - ; true y → sym (lid (snd G) _)} - - BoolGHom : GroupHom BoolGr G - fst BoolGHom = Bool→G - snd BoolGHom = Bool→GHom - - genG : fst G - genG = fst H 1 - - open IsGroupHom - - _·G_ = GroupStr._·_ (snd G) - - isSurjectiveBoolGHomPos : (n : ℕ) → Σ[ b ∈ Bool ] fst H (pos n) ≡ Bool→G b - isSurjectiveBoolGHomPos zero = true , pres1 strH - isSurjectiveBoolGHomPos (suc zero) = false , refl - isSurjectiveBoolGHomPos (suc (suc n)) = - isSurjectiveBoolGHomPos n .fst - , (cong (fst H) (+Comm (pos n) 2) - ∙ (pres· (snd H) 2 (pos n))) - ∙ cong (_·G (fst H (pos n))) - (sym (sym (exact.ImG→H⊂KerH→L ex 2 ∣ 1 , P ∣))) - ∙ lid (snd G) _ - ∙ isSurjectiveBoolGHomPos n .snd - - isSurjectiveBoolGHomPre : (x : Z) → Σ[ b ∈ Bool ] fst H x ≡ Bool→G b - isSurjectiveBoolGHomPre (pos n) = isSurjectiveBoolGHomPos n - isSurjectiveBoolGHomPre (negsuc n) = - (isSurjectiveBoolGHomPos (suc n) .fst) - , cong (fst H) (+Comm (- (pos (suc n))) 0) - ∙ presinv (snd H) (pos (suc n)) - ∙ cong (inv strG) (isSurjectiveBoolGHomPos (suc n) .snd) - ∙ sym (presinv Bool→GHom (isSurjectiveBoolGHomPos (suc n) .fst)) - - isSurjectiveBoolGHom : isSurjective BoolGHom - isSurjectiveBoolGHom g = - pMap (λ {(x , p) → (isSurjectiveBoolGHomPre x .fst) - , (sym (isSurjectiveBoolGHomPre x .snd) ∙ p)}) - (isSurjectiveH g) - - isInjectiveBoolGHom : isInjective BoolGHom - isInjectiveBoolGHom false p = - ⊥-rec (pRec isProp⊥ (uncurry contr) ∃1≡2) - where - H≡0 : isInKer H 1 - H≡0 = p - - H≡ : isInIm F 1 - H≡ = exact.KerH→L⊂ImG→H ex 1 p - - F≡pos : (n : ℕ) → fst F (pos n) ≡ 2 ·Z (pos n) - F≡pos zero = pres1 (snd F) - F≡pos (suc zero) = P - F≡pos (suc (suc n)) = - cong (fst F) (+Comm (pos (suc n)) (pos 1)) - ∙∙ pres· (snd F) (pos 1) (pos (suc n)) - ∙∙ cong₂ _+Z_ P (F≡pos (suc n)) - ∙∙ +Comm 2 (2 ·Z (pos (suc n))) - ∙∙ (sym (·DistR+ 2 (pos (suc n)) 1)) - - F≡ : (x : Z) → fst F x ≡ 2 ·Z x - F≡ (pos n) = F≡pos n - F≡ (negsuc n) = - cong (fst F) (+Comm (negsuc n) 0) - ∙ presinv (snd F) (pos (suc n)) - ∙ (sym (+Comm (- (fst F (pos (suc n)))) 0) - ∙ cong -_ (F≡pos (suc n))) - ∙ sym (ℤ·negsuc 2 n) - - ∃1≡2 : ∃[ x ∈ Z ] 1 ≡ 2 ·Z x - ∃1≡2 = - pMap (λ {(x , p) → x , (sym p ∙ F≡ x)}) H≡ - - contr : (x : Z) → 1 ≡ 2 ·Z x → ⊥ - contr (pos zero) p = snotz (injPos p) - contr (pos (suc n)) p = snotz (injPos (sym (cong predℤ p ∙ predSuc _ ∙ help1 n))) - where - help1 : (n : ℕ) → pos (suc n) +pos n ≡ pos (suc (n + n)) - help1 zero = refl - help1 (suc n) = cong sucℤ (sym (sucℤ+ (pos (suc n)) (pos n))) - ∙ cong (sucℤ ∘ sucℤ) (help1 n) - ∙ cong pos (cong (2 +_) (sym (+-suc n n))) - contr (negsuc n) p = negsucNotpos (suc (n + n)) 1 (sym (p ∙ negsuclem n)) - where - negsuclem : (n : ℕ) → (negsuc n +negsuc n) ≡ negsuc (suc (n + n)) - negsuclem zero = refl - negsuclem (suc n) = - cong predℤ (+Comm (negsuc (suc n)) (negsuc n) - ∙ cong predℤ (negsuclem n)) - ∙ cong (negsuc ∘ (2 +_)) (sym (+-suc n n)) - isInjectiveBoolGHom true p = refl - - BijectionIsoBoolG : BijectionIso BoolGr G - BijectionIso.fun BijectionIsoBoolG = BoolGHom - BijectionIso.inj BijectionIsoBoolG = isInjectiveBoolGHom - BijectionIso.surj BijectionIsoBoolG = isSurjectiveBoolGHom - - Bool≅G : GroupIso BoolGr G - Bool≅G = BijectionIso→GroupIso BijectionIsoBoolG - -open import Cubical.Data.Sum -exact→Bool≅G± : (G : Group₀) (F : GroupHom ℤ ℤ) (H : GroupHom ℤ G) - (P : (fst F 1 ≡ 2) ⊎ (fst F 1 ≡ - 2)) - (ex : exact ℤ ℤ G UnitGr F H (→UnitHom G)) - → GroupIso BoolGr G -exact→Bool≅G± G F H (inl x) ex = exact→BoolIso.Bool≅G G F H x ex -exact→Bool≅G± G F H (inr x) ex = - exact→BoolIso.Bool≅G _ _ _ (cong (fst flip') x) exact' - where - flip' : GroupHom ℤ ℤ - fst flip' x = GroupStr.inv (snd ℤ) x - snd flip' = makeIsGroupHom (λ x y → - (+Comm (pos 0) (- (x +Z y)) - ∙ -Dist+ x y) - ∙ λ i → GroupStr.lid (snd ℤ) (- x) (~ i) +Z GroupStr.lid (snd ℤ) (- y) (~ i)) - - exact' : exact ℤ ℤ G UnitGr (compGroupHom F flip') H (→UnitHom G) - ImG→H⊂KerH→L exact' x inim = - (cong (fst H) (sym (GroupTheory.invInv ℤ x)) - ∙ (IsGroupHom.presinv (snd H) (fst flip' x))) - ∙∙ cong (GroupStr.inv (snd G)) (ImG→H⊂KerH→L ex _ (help inim)) - ∙∙ GroupTheory.inv1g G - where - help : isInIm (compGroupHom F flip') x → isInIm F (fst flip' x) - help = - pMap λ {(x , p) → x , - ((sym (-Involutive (fst F x)) - ∙ +Comm (- (- fst F x)) (pos 0)) - ∙ (λ i → pos 0 - (+Comm (- fst F x) (pos 0) i))) - ∙ λ i → pos 0 - p i } - KerH→L⊂ImG→H exact' x inker = - pMap (λ {(x , p) → (fst flip' x) , (GroupStr.lid (snd ℤ) (- (fst F (pos 0 - x))) - ∙ cong -_ (IsGroupHom.presinv (snd F) x - ∙ GroupStr.lid (snd ℤ) (- (fst F x))) - ∙ -Involutive (fst F x) ∙ p)}) - (KerH→L⊂ImG→H ex x inker) - ImH→L⊂KerL→R exact' = ImH→L⊂KerL→R ex - KerL→R⊂ImH→L exact' = KerL→R⊂ImH→L ex - - -exact→boolIsoGenPre : - (G H L : Group₀) (Z≅G : GroupEquiv ℤ G) (Z≅H : GroupEquiv ℤ H) - → (G→H : GroupHom G H) (H→L : GroupHom H L) - → ((invEq (fst Z≅H) (fst G→H (fst (fst Z≅G) 1)) ≡ 2) - ⊎ (invEq (fst Z≅H) (fst G→H (fst (fst Z≅G) 1)) ≡ - 2)) - → exact G H L UnitGr G→H H→L (→UnitHom L) - → GroupIso BoolGr L -exact→boolIsoGenPre G H L = - GroupEquivJ (λ G Z≅G - → (Z≅H : GroupEquiv ℤ H) - → (G→H : GroupHom G H) (H→L : GroupHom H L) - → ((invEq (fst Z≅H) (fst G→H (fst (fst Z≅G) 1)) ≡ 2) - ⊎ (invEq (fst Z≅H) (fst G→H (fst (fst Z≅G) 1)) ≡ - 2)) - → exact G H L UnitGr G→H H→L (→UnitHom L) - → GroupIso BoolGr L) - (GroupEquivJ (λ H Z≅H → - (G→H : GroupHom ℤ H) - (H→L : GroupHom H L) → - (invEq (fst Z≅H) (fst G→H (idfun (typ ℤ) (pos 1))) ≡ pos 2) ⊎ - (invEq (fst Z≅H) (fst G→H (idfun (typ ℤ) (pos 1))) ≡ negsuc 1) → - exact ℤ H L UnitGr G→H H→L (→UnitHom L) → GroupIso BoolGr L) - λ Z→Z L P ex → exact→Bool≅G± _ _ L P ex) - -open import Cubical.Homotopy.HopfInvariant.Base -open import Cubical.Homotopy.HopfInvariant.HopfMap -open import Cubical.Homotopy.HopfInvariant.Homomorphism -open import Cubical.Homotopy.Group.Pi3S2 -open import Cubical.Homotopy.Group.PinSn - - --- π₂S³-gen-by-HopfMap - --- abs (HopfInvariant-π' 0 ([ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×)) ≡ 2 - -hopfInvariantEquiv : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤ -fst (fst hopfInvariantEquiv) = HopfInvariant-π' 0 -snd (fst hopfInvariantEquiv) = - GroupEquivℤ-isEquiv (invGroupEquiv π₃S²≅ℤ) ∣ HopfMap ∣₂ - π₂S³-gen-by-HopfMap (GroupHom-HopfInvariant-π' 0) - (abs→⊎ _ _ HopfInvariant-HopfMap) -snd hopfInvariantEquiv = snd (GroupHom-HopfInvariant-π' 0) - -π₃PushoutCharac : - abs (HopfInvariant-π' 0 [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π') ≡ 2 - → GroupIso BoolGr (π'Gr 2 thePushout∙) -π₃PushoutCharac p = - exact→boolIsoGenPre (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 thePushout∙) - (GroupIso→GroupEquiv (invGroupIso (πₙ'Sⁿ≅ℤ 2))) - (invGroupEquiv hopfInvariantEquiv) - (π'∘∙Hom 2 (fold∘W , refl)) - (π'∘∙Hom 2 inr∙) - (abs→⊎ _ _ (cong abs - (cong (invEq (fst (invGroupEquiv hopfInvariantEquiv))) - (cong (fst (π'∘∙Hom 2 (fold∘W , refl))) - help) - ∙ cong (HopfInvariant-π' 0) - (Fold∘W ∙ sym (cong ∣_∣₂ ([]≡[]₂ (idfun∙ (S₊∙ 2)) (idfun∙ (S₊∙ 2)))))) ∙ p)) - S³→S²→Pushout→Unit - where - help : inv (fst (πₙ'Sⁿ≅ℤ 2)) 1 ≡ ∣ idfun∙ _ ∣₂ - help = cong (inv (fst (πₙ'Sⁿ≅ℤ 2))) (sym (πₙ'Sⁿ≅ℤ-idfun∙ 2)) - ∙ leftInv (fst (πₙ'Sⁿ≅ℤ 2)) ∣ idfun∙ _ ∣₂ - -open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 -open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso - -∣HopfWhitehead∣≡2→π₄S³≅Bool : - abs (HopfInvariant-π' 0 [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π') ≡ 2 - → GroupEquiv (πGr 3 (S₊∙ 3)) BoolGr -∣HopfWhitehead∣≡2→π₄S³≅Bool p = - compGroupEquiv - (compGroupEquiv - (GroupIso→GroupEquiv - (compGroupIso π₄S³≅π₃PushS² - (invGroupIso (π'Gr≅πGr 2 (Pushout⋁↪fold⋁∙ (S₊∙ 2)))))) - (compGroupEquiv (invGroupEquiv (π'Iso 2 lem∙)) - (π'Iso 2 (lem₂ , sym (push north))))) - (invGroupEquiv (GroupIso→GroupEquiv (π₃PushoutCharac p))) - where - lem₂ : Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ ≃ fst thePushout∙ - lem₂ = compEquiv - (compEquiv pushoutSwitchEquiv (isoToEquiv (PushoutDistr.PushoutDistrIso fold⋁ W λ _ → tt))) - pushoutSwitchEquiv - - lem₁ : Pushout W (λ _ → tt) ≃ cofibW S¹ S¹ base base - lem₁ = pushoutEquiv W (λ _ → tt) joinTo⋁ (λ _ → tt) - (isoToEquiv (invIso (IsoSphereJoin 1 1))) - (idEquiv _) - (idEquiv _) - refl - refl - - lem : Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ - ≃ fst (Pushout⋁↪fold⋁∙ (S₊∙ 2)) - lem = pushoutEquiv inl _ ⋁↪ fold⋁ - (idEquiv _) - (compEquiv lem₁ (isoToEquiv (invIso (Iso-Susp×Susp-cofibJoinTo⋁ S¹ S¹ base base)))) - (idEquiv _) - (Susp×Susp→cofibW≡ S¹ S¹ base base) - refl - - lem∙ : (Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ , inr north) - ≃∙ (Pushout⋁↪fold⋁∙ (S₊∙ 2)) - fst lem∙ = lem - snd lem∙ = sym (push (inl north)) diff --git a/Cubical/Homotopy/HopfInvariant/Whitehead.agda b/Cubical/Homotopy/HopfInvariant/Whitehead.agda deleted file mode 100644 index ff7dc1bd74..0000000000 --- a/Cubical/Homotopy/HopfInvariant/Whitehead.agda +++ /dev/null @@ -1,808 +0,0 @@ -{-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Homotopy.HopfInvariant.Whitehead where - -open import Cubical.Homotopy.HopfInvariant.Base -open import Cubical.Homotopy.Group.Base -open import Cubical.Homotopy.Group.SuspensionMap -open import Cubical.Homotopy.HopfInvariant.HopfMap -open import Cubical.Homotopy.HopfInvariant.Base -open import Cubical.Homotopy.Group.Pi4S3.Tricky -open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso -open import Cubical.Homotopy.Loopspace - - -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.GroupStructure -open import Cubical.ZCohomology.Properties -open import Cubical.ZCohomology.MayerVietorisUnreduced -open import Cubical.ZCohomology.Groups.Unit -open import Cubical.ZCohomology.Groups.Wedge -open import Cubical.ZCohomology.Groups.Sn -open import Cubical.ZCohomology.RingStructure.CupProduct -open import Cubical.ZCohomology.RingStructure.GradedCommutativity -open import Cubical.ZCohomology.Gysin - -open import Cubical.Foundations.Path -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Function -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Equiv - -open import Cubical.Data.Sigma -open import Cubical.Data.Int hiding (_+'_) -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) -open import Cubical.Data.Unit - -open import Cubical.Algebra.Group - renaming (ℤ to ℤGroup ; Unit to UnitGroup) -open import Cubical.Algebra.Group.ZAction -open import Cubical.Algebra.Group.Exact - -open import Cubical.Relation.Nullary - -open import Cubical.HITs.Pushout -open import Cubical.HITs.Sn -open import Cubical.HITs.Susp -open import Cubical.HITs.Wedge -open import Cubical.HITs.Truncation -open import Cubical.HITs.SetTruncation - renaming (elim to sElim ; elim2 to sElim2 ; map to sMap) -open import Cubical.HITs.PropositionalTruncation renaming (map to pMap ; rec to pRec) - - --- abs (HopfInvariant-π' 0 [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π') ≡ 2 - - -open import Cubical.Homotopy.Whitehead -open import Cubical.HITs.S1 -open import Cubical.HITs.Join -open import Cubical.ZCohomology.Groups.Wedge -open import Cubical.ZCohomology.Groups.Sn -open import Cubical.ZCohomology.Groups.Unit - -¬lem : (n m : ℕ) → ¬ suc (n +ℕ m) ≡ m -¬lem n zero = snotz -¬lem n (suc m) p = ¬lem n m (sym (+-suc n m) ∙ cong predℕ p) - -H²-genₗ' : coHom 2 (S₊ 2 × S₊ 2) -H²-genₗ' = ∣ (λ x → ∣ fst x ∣) ∣₂ - -H²-genᵣ' : coHom 2 (S₊ 2 × S₊ 2) -H²-genᵣ' = ∣ (λ x → ∣ snd x ∣) ∣₂ - -abstract - H²-genₗabs : coHom 2 (S₊ 2 × S₊ 2) - H²-genₗabs = H²-genₗ' - - H²-genᵣabs : coHom 2 (S₊ 2 × S₊ 2) - H²-genᵣabs = H²-genᵣ' - - absInd : ∀ {ℓ} (A : coHom 2 (S₊ 2 × S₊ 2) → coHom 2 (S₊ 2 × S₊ 2) → Type ℓ) - → A H²-genₗabs H²-genᵣabs - → A H²-genₗ' H²-genᵣ' - absInd A p = p - - absInd⁻ : ∀ {ℓ} (A : coHom 2 (S₊ 2 × S₊ 2) → coHom 2 (S₊ 2 × S₊ 2) → Type ℓ) - → A H²-genₗ' H²-genᵣ' - → A H²-genₗabs H²-genᵣabs - absInd⁻ A p = p - - -∨map = joinTo⋁ {A = S₊∙ 1} {B = S₊∙ 1} -module mv = MV ((S₊∙ 2) ⋁ (S₊∙ 2)) Unit (join S¹ S¹) ∨map (λ _ → tt) - -downFun : (n m : ℕ) - → (S₊ (suc (suc n)) → S₊ (suc m) → coHomK (suc (suc ((suc n) +ℕ m)))) - → S₊ (suc n) → S₊ (suc m) → coHomK (suc (suc (n +ℕ m))) -downFun n m f x y = - ΩKn+1→Kn ((suc ((suc n) +ℕ m))) - (sym (rCancelₖ _ (f north y)) - ∙∙ cong (λ x → f x y -ₖ f north y) (merid x ∙ sym (merid (ptSn (suc n)))) - ∙∙ rCancelₖ _ (f north y)) - -upFun : (n m : ℕ) - → (S₊ (suc n) → S₊ (suc m) → coHomK (suc (suc (n +ℕ m)))) - → (S₊ (suc (suc n)) → S₊ (suc m) → coHomK (suc (suc ((suc n) +ℕ m)))) -upFun n m f north y = 0ₖ _ -upFun n m f south y = 0ₖ _ -upFun n m f (merid a i) y = - Kn→ΩKn+1 _ (f a y) i - -funLem : (n m : ℕ) → (f : S₊ (suc m) → coHomK (suc n)) - → ¬ (n ≡ m) - → ∥ f ≡ (λ _ → 0ₖ (suc n)) ∥ -funLem n m f p = - Iso.fun PathIdTrunc₀Iso - (isContr→isProp - (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Sᵐ≅0 n m p)) isContrUnit) - ∣ f ∣₂ (0ₕ _)) - -upDownIso : (n m : ℕ) - → GroupIso (coHomGr (suc (((suc n) +ℕ m))) - (S₊ (suc n) × S₊ (suc m))) - (coHomGr (suc (suc ((suc n) +ℕ m))) - (S₊ (suc (suc n)) × S₊ (suc m))) -Iso.fun (fst (upDownIso n m)) = sMap (uncurry ∘ upFun n m ∘ curry) -Iso.inv (fst (upDownIso n m)) = sMap ((uncurry ∘ downFun n m ∘ curry)) -Iso.rightInv (fst (upDownIso n m)) = - sElim (λ _ → isSetPathImplicit) - λ f → Iso.inv PathIdTrunc₀Iso - (pRec squash - (uncurry (λ g p - → pMap (λ gid → funExt λ {(x , y) - → (λ i → uncurry (upFun n m (curry (uncurry (downFun n m (curry (p (~ i))))))) (x , y)) - ∙∙ main g gid x y - ∙∙ funExt⁻ p (x , y)}) - (gIdLem g))) - (rewr f)) - where - lem : (n m : ℕ) → ¬ suc (suc (n +ℕ m)) ≡ m - lem n zero = snotz - lem n (suc m) p = lem n m (cong suc (sym (+-suc n m)) ∙ cong predℕ p) - - charac-fun : (S₊ (suc n) → S₊ (suc m) → typ (Ω (coHomK-ptd (suc (suc (suc n +ℕ m)))))) - → S₊ (suc (suc n)) × S₊ (suc m) → coHomK (suc (suc (suc n +ℕ m))) - charac-fun g (north , y) = 0ₖ _ - charac-fun g (south , y) = 0ₖ _ - charac-fun g (merid c i , y) = g c y i - - rewr : (f : S₊ (suc (suc n)) × S₊ (suc m) → coHomK (suc (suc (suc n +ℕ m)))) - → ∃[ g ∈ (S₊ (suc n) → S₊ (suc m) → typ (Ω (coHomK-ptd (suc (suc (suc n +ℕ m)))))) ] - charac-fun g ≡ f - rewr f = - pMap (λ p → (λ x y → sym (funExt⁻ p y) - ∙∙ (λ i → f ((merid x ∙ sym (merid (ptSn (suc n)))) i , y)) - ∙∙ funExt⁻ p y) - , funExt λ { (north , y) → sym (funExt⁻ p y) - ; (south , y) → sym (funExt⁻ p y) - ∙ λ i → f (merid (ptSn (suc n)) i , y) - ; (merid a i , y) j → - hcomp (λ k → λ { (i = i0) → p (~ j ∧ k) y - ; (i = i1) → compPath-filler' - (sym (funExt⁻ p y)) - (λ i → f (merid (ptSn (suc n)) i , y)) k j - ; (j = i1) → f (merid a i , y) }) - (f (compPath-filler (merid a) (sym (merid (ptSn (suc n)))) (~ j) i - , y))}) - (funLem _ _ (λ x → f (north , x)) - (lem n m)) - - gIdLem : (g : S₊ (suc n) → S₊ (suc m) → typ (Ω (coHomK-ptd (suc (suc (suc n +ℕ m)))))) - → ∥ (g (ptSn _)) ≡ (λ _ → refl) ∥ - gIdLem g = - Iso.fun PathIdTrunc₀Iso - (isContr→isProp - (isOfHLevelRetractFromIso 0 - ((invIso (fst (coHom≅coHomΩ _ (S₊ (suc m)))))) - (0ₕ _ , sElim (λ _ → isProp→isSet (squash₂ _ _)) - λ f → pRec (squash₂ _ _) (sym ∘ cong ∣_∣₂) - (funLem _ _ f (¬lem n m)))) - ∣ g (ptSn (suc n)) ∣₂ ∣ (λ _ → refl) ∣₂) - - main : (g : _) → (g (ptSn _)) ≡ (λ _ → refl) - → (x : S₊ (suc (suc n))) (y : S₊ (suc m)) - → uncurry - (upFun n m (curry (uncurry (downFun n m (curry (charac-fun g)))))) - (x , y) - ≡ charac-fun g (x , y) - main g gid north y = refl - main g gid south y = refl - main g gid (merid a i) y j = help j i - where - help : (λ i → uncurry - (upFun n m (curry (uncurry (downFun n m (curry (charac-fun g)))))) - (merid a i , y)) ≡ g a y - help = (λ i → Iso.rightInv (Iso-Kn-ΩKn+1 _) - ((sym (rCancel≡refl _ i) - ∙∙ cong-∙ (λ x → rUnitₖ _ (charac-fun g (x , y)) i) (merid a) (sym (merid (ptSn (suc n)))) i - ∙∙ rCancel≡refl _ i)) i) - ∙∙ sym (rUnit _) - ∙∙ (cong (g a y ∙_) (cong sym (funExt⁻ gid y)) - ∙ sym (rUnit (g a y))) -Iso.leftInv (fst (upDownIso n m)) = - sElim (λ _ → isSetPathImplicit) - λ f → pRec (squash₂ _ _) - (λ id - → cong ∣_∣₂ - (funExt (λ {(x , y) → (λ i → ΩKn+1→Kn _ (sym (rCancel≡refl _ i) - ∙∙ ((λ j → rUnitₖ _ - (upFun n m (curry f) - ((merid x ∙ sym (merid (ptSn (suc n)))) j) y) i)) - ∙∙ rCancel≡refl _ i)) - ∙ (λ i → ΩKn+1→Kn _ (rUnit - (cong-∙ (λ r → upFun n m (curry f) r y) - (merid x) (sym (merid (ptSn (suc n)))) i) (~ i))) - ∙ cong (ΩKn+1→Kn _) (cong (Kn→ΩKn+1 _ (f (x , y)) ∙_) - (cong sym (cong (Kn→ΩKn+1 _) (funExt⁻ id y) ∙ (Kn→ΩKn+10ₖ _))) - ∙ sym (rUnit _)) - ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (x , y))}))) - (funLem (suc n +ℕ m) m (λ x → f (ptSn _ , x)) - (¬lem n m)) -snd (upDownIso n m) = - makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) - λ f g → cong ∣_∣₂ - (funExt λ { (north , y) → refl - ; (south , y) → refl - ; (merid a i , y) j - → (sym (∙≡+₂ _ (Kn→ΩKn+1 _ (f (a , y))) (Kn→ΩKn+1 _ (g (a , y)))) - ∙ sym (Kn→ΩKn+1-hom _ (f (a , y)) (g (a , y)))) (~ j) i})) - -toBaseIso : (n m : ℕ) → GroupIso (coHomGr (suc (((suc n) +ℕ m))) - (S₊ (suc n) × S₊ (suc m))) - ((coHomGr (suc (suc m)) - (S₊ (suc zero) × S₊ (suc m)))) -toBaseIso zero m = idGroupIso -toBaseIso (suc n) m = - compGroupIso (invGroupIso (upDownIso n m)) (toBaseIso n m) - -BaseIso→ : (m : ℕ) - → (S₊ (suc m) → coHomK (suc m)) - → S₊ (suc zero) → S₊ (suc m) → coHomK (suc (suc m)) -BaseIso→ m f base y = 0ₖ _ -BaseIso→ m f (loop i) y = Kn→ΩKn+1 (suc m) (f y) i - -BaseIso← : (m : ℕ) - → (S₊ (suc zero) → S₊ (suc m) → coHomK (suc (suc m))) - → (S₊ (suc m) → coHomK (suc m)) -BaseIso← m f x = - ΩKn+1→Kn (suc m) - (sym (rCancelₖ _ (f base x)) ∙∙ ((λ i → f (loop i) x -ₖ f base x)) ∙∙ rCancelₖ _ (f base x)) - -BaseIso : (m : ℕ) → GroupIso (coHomGr (suc m) (S₊ (suc m))) - ((coHomGr (suc (suc m)) - (S₊ (suc zero) × S₊ (suc m)))) - -Iso.fun (fst (BaseIso m)) = sMap (uncurry ∘ BaseIso→ m) -Iso.inv (fst (BaseIso m)) = sMap (BaseIso← m ∘ curry) -Iso.rightInv (fst (BaseIso m)) = - sElim (λ _ → isSetPathImplicit) - λ f → Iso.inv PathIdTrunc₀Iso - (pMap (uncurry (λ g p - → funExt λ {(x , y) - → (λ i → uncurry (BaseIso→ m (BaseIso← m (curry (p i)))) (x , y)) - ∙∙ main g x y - ∙∙ sym (funExt⁻ p (x , y))})) - (ss f)) - where - characFun : (x : S₊ (suc m) → coHomK (suc m)) - → S₊ 1 × S₊ (suc m) → coHomK (suc (suc m)) - characFun f (base , y) = 0ₖ _ - characFun f (loop i , y) = Kn→ΩKn+1 _ (f y) i - - main : (g : _) → (x : S¹) (y : S₊ (suc m)) → uncurry (BaseIso→ m (BaseIso← m (curry (characFun g)))) - (x , y) - ≡ characFun g (x , y) - main g base y = refl - main g (loop i) y j = help j i - where - help : cong (λ x → BaseIso→ m (BaseIso← m (curry (characFun g))) x y) loop - ≡ Kn→ΩKn+1 _ (g y) - help = Iso.rightInv (Iso-Kn-ΩKn+1 (suc m)) - (sym (rCancelₖ _ (0ₖ _)) ∙∙ ((λ i → Kn→ΩKn+1 _ (g y) i -ₖ 0ₖ _)) ∙∙ rCancelₖ _ (0ₖ _)) - ∙∙ (λ i → sym (rCancel≡refl _ i) ∙∙ (λ j → rUnitₖ _ (Kn→ΩKn+1 _ (g y) j) i) ∙∙ rCancel≡refl _ i) - ∙∙ sym (rUnit _) - - lem : (m : ℕ) → ¬ suc m ≡ m - lem zero = snotz - lem (suc m) p = lem m (cong predℕ p) - - ss : (f : S₊ 1 × S₊ (suc m) → coHomK (suc (suc m))) - → ∃[ g ∈ (S₊ (suc m) → coHomK (suc m)) ] f ≡ characFun g - ss f = - pMap (λ p → (λ x → ΩKn+1→Kn _ (sym (funExt⁻ p x) ∙∙ (λ i → f (loop i , x)) ∙∙ funExt⁻ p x)) - , funExt λ { (base , y) → funExt⁻ p y - ; (loop i , x) j → - hcomp (λ k → λ {(i = i0) → funExt⁻ p x j - ; (i = i1) → funExt⁻ p x j - ; (j = i0) → f (loop i , x) - ; (j = i1) → - Iso.rightInv (Iso-Kn-ΩKn+1 (suc m)) - (sym (funExt⁻ p x) ∙∙ (λ i → f (loop i , x)) ∙∙ funExt⁻ p x) (~ k) i}) - (doubleCompPath-filler - (sym (funExt⁻ p x)) (λ i → f (loop i , x)) (funExt⁻ p x) j i)}) - (funLem (suc m) m (λ x → f (base , x)) (lem m)) -Iso.leftInv (fst (BaseIso m)) = - sElim (λ _ → isSetPathImplicit) - λ f - → cong ∣_∣₂ (funExt λ x - → cong (ΩKn+1→Kn (suc m)) - ((λ i → sym (rCancel≡refl _ i) - ∙∙ cong (λ z → rUnitₖ _ (BaseIso→ m f z x) i) loop - ∙∙ rCancel≡refl _ i) ∙ sym (rUnit (Kn→ΩKn+1 (suc m) (f x)))) - ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) -snd (BaseIso m) = - makeIsGroupHom - (sElim2 - (λ _ _ → isSetPathImplicit) - λ f g → cong ∣_∣₂ - (funExt λ { (base , y) → refl - ; (loop i , y) j → - (sym (∙≡+₂ _ (Kn→ΩKn+1 _ (f y)) (Kn→ΩKn+1 _ (g y))) - ∙ sym (Kn→ΩKn+1-hom _ (f y) (g y))) (~ j) i})) - -GroupIso2 : ∀ {ℓ} {A : Type ℓ} (n : ℕ) - → GroupIso (×coHomGr (suc n) A Unit) (coHomGr (suc n) A) -Iso.fun (fst (GroupIso2 n)) = fst -Iso.inv (fst (GroupIso2 n)) x = x , 0ₕ (suc n) -Iso.rightInv (fst (GroupIso2 n)) _ = refl -Iso.leftInv (fst (GroupIso2 n)) x = - ΣPathP (refl , isContr→isProp (isContrHⁿ-Unit n) (0ₕ (suc n)) (snd x)) -snd (GroupIso2 n) = makeIsGroupHom λ _ _ → refl - -theIso : GroupIso (coHomGr 2 (S₊ 2)) (coHomGr 4 (S₊ 2 × S₊ 2)) -theIso = (compGroupIso (BaseIso 1) (upDownIso 0 1)) - -Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ : (n m : ℕ) - → GroupIso (coHomGr ((suc n) +' (suc m)) - (S₊ (suc n) × S₊ (suc m))) - ℤGroup -Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ zero m = - compGroupIso (invGroupIso (BaseIso m)) (Hⁿ-Sⁿ≅ℤ m) -Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ (suc n) m = - compGroupIso - (invGroupIso (upDownIso n m)) - (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ n m) - - -theIso⌣ : Iso.fun (fst (theIso)) ∣ ∣_∣ₕ ∣₂ ≡ H²-genₗ' ⌣ H²-genᵣ' -theIso⌣ = - cong ∣_∣₂ (funExt (uncurry - (λ { north y → refl - ; south y → refl - ; (merid a i) y j → Kn→ΩKn+1 3 (lem₃ a y j) i}))) - where - lem₃ : (a : S¹) (y : S₊ 2) → BaseIso→ 1 ∣_∣ₕ a y ≡ _⌣ₖ_ {n = 1} {m = 2} ∣ a ∣ₕ ∣ y ∣ₕ - lem₃ base y = refl - lem₃ (loop i) y = refl - -abstract - Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs : (n m : ℕ) - → GroupIso (coHomGr ((suc n) +' (suc m)) - (S₊ (suc n) × S₊ (suc m))) - ℤGroup - Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs = Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ - - Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs≡ : (n m : ℕ) → Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs n m ≡ Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ n m - Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs≡ n m = refl - - Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-⌣ : Iso.fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1)) (H²-genₗ' ⌣ H²-genᵣ') ≡ 1 - Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-⌣ = - cong (Iso.fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1))) (sym theIso⌣) - ∙ speedUp ∣_∣ₕ - ∙ refl -- Computation! :-) - where - speedUp : (f : _) → Iso.fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1)) (Iso.fun (fst (theIso)) ∣ f ∣₂) - ≡ Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 1)) ∣ f ∣₂ - speedUp f i = (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 1))) (Iso.leftInv (fst theIso) ∣ f ∣₂ i) - - - -open import Cubical.Data.Sum renaming (rec to ⊎rec) -open Iso - - -Pushout≅Torus2 : GroupIso (coHomGr 2 (S₊ 2 × S₊ 2)) (coHomGr 2 (Pushout joinTo⋁ (λ _ → tt))) -Pushout≅Torus2 = coHomIso 2 (invIso (Iso-Susp×Susp-cofibJoinTo⋁ S¹ S¹ base base)) - -GroupHom2 : - GroupHom (coHomGr 2 (Pushout ∨map (λ _ → tt))) - (×coHomGr 2 (S₊∙ 2 ⋁ S₊∙ 2) Unit) -GroupHom2 = mv.i 2 - -GroupEquiv2 : GroupEquiv - (coHomGr 2 (Pushout ∨map (λ _ → tt))) - (×coHomGr 2 (S₊∙ 2 ⋁ S₊∙ 2) Unit) -fst (fst GroupEquiv2) = fst GroupHom2 -snd (fst GroupEquiv2) = - SES→isEquiv - (isContr→≡UnitGroup - (isOfHLevelRetractFromIso 0 - (compIso (coHomIso 1 (invIso (IsoSphereJoin 1 1)) .fst) - (fst (Hⁿ-Sᵐ≅0 0 2 λ p → snotz (sym p)))) - isContrUnit)) - ((isContr→≡UnitGroup - (isOfHLevelRetractFromIso 0 - (compIso (coHomIso 2 (invIso (IsoSphereJoin 1 1)) .fst) - (fst (Hⁿ-Sᵐ≅0 1 2 λ p → snotz (cong predℕ (sym p))))) - isContrUnit))) - (mv.d 1) (mv.i 2) - (mv.Δ 2) - (mv.Ker-i⊂Im-d 1) - (mv.Ker-Δ⊂Im-i 2) -snd GroupEquiv2 = snd GroupHom2 - -coHom2S²×S² : GroupIso (coHomGr 2 (S₊ 2 × S₊ 2)) (DirProd ℤGroup ℤGroup) -coHom2S²×S² = - compGroupIso Pushout≅Torus2 - (compGroupIso - (GroupEquiv→GroupIso GroupEquiv2) - (compGroupIso (GroupIso2 1) - (compGroupIso (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 1) - (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ 1) (Hⁿ-Sⁿ≅ℤ 1))))) - -zz : GroupIso (coHomGr 4 (S₊ 2 × S₊ 2)) (coHomGr 4 (Pushout joinTo⋁ (λ _ → tt))) -zz = coHomIso 4 (invIso (Iso-Susp×Susp-cofibJoinTo⋁ S¹ S¹ base base)) - -GroupEquiv1 : GroupEquiv - (coHomGr 3 (join S¹ S¹)) - (coHomGr 4 (Pushout ∨map (λ _ → tt))) -fst (fst GroupEquiv1) = fst (mv.d 3) -snd (fst GroupEquiv1) = ss - where - abstract - ss : isEquiv (fst (mv.d 3)) - ss = SES→isEquiv - (isContr→≡UnitGroup - (isOfHLevelΣ 0 - (isOfHLevelRetractFromIso 0 - (compIso - (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 2)) - (compIso - (prodIso (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) - (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p)))) - lUnit×Iso)) - isContrUnit) - (λ _ → isContrHⁿ-Unit 2))) - (isContr→≡UnitGroup - (isOfHLevelΣ 0 - (isOfHLevelRetractFromIso 0 - (compIso - (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 3)) - (compIso - (prodIso (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) - (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p)))) - lUnit×Iso)) - isContrUnit) - λ _ → isContrHⁿ-Unit 3)) - (mv.Δ 3) - (mv.d 3) - (mv.i 4) - (mv.Ker-d⊂Im-Δ 3) - (mv.Ker-i⊂Im-d 3) -snd GroupEquiv1 = snd (mv.d 3) - - -coHom4S²×S² : GroupIso (coHomGr 4 (S₊ 2 × S₊ 2)) ℤGroup -coHom4S²×S² = - compGroupIso zz - (compGroupIso - (invGroupIso (GroupEquiv→GroupIso GroupEquiv1)) - (compGroupIso - (coHomIso 3 (invIso (IsoSphereJoin 1 1))) - (Hⁿ-Sⁿ≅ℤ 2))) - - -CHopf : Type -CHopf = HopfInvariantPush 0 fold∘W - -Hopfαfold∘W = Hopfα 0 (fold∘W , refl) -Hopfβfold∘W = Hopfβ 0 (fold∘W , refl) - -lem₂ : Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ ≃ fst thePushout∙ -lem₂ = compEquiv - (compEquiv pushoutSwitchEquiv - (isoToEquiv (PushoutDistr.PushoutDistrIso fold⋁ W λ _ → tt))) - pushoutSwitchEquiv - -lem₁ : Pushout W (λ _ → tt) ≃ cofibW S¹ S¹ base base -lem₁ = pushoutEquiv W (λ _ → tt) joinTo⋁ (λ _ → tt) - (isoToEquiv (invIso (IsoSphereJoin 1 1))) - (idEquiv _) - (idEquiv _) - refl - refl - -lem : Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ - ≃ fst (Pushout⋁↪fold⋁∙ (S₊∙ 2)) -lem = pushoutEquiv inl _ ⋁↪ fold⋁ - (idEquiv _) - (compEquiv lem₁ (isoToEquiv (invIso (Iso-Susp×Susp-cofibJoinTo⋁ S¹ S¹ base base)))) - (idEquiv _) - (Susp×Susp→cofibW≡ S¹ S¹ base base) - refl - -lem∙ : (Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ , inr north) - ≃∙ (Pushout⋁↪fold⋁∙ (S₊∙ 2)) -fst lem∙ = lem -snd lem∙ = sym (push (inl north)) - - -CHopfIso : Iso CHopf (Pushout⋁↪fold⋁ (S₊∙ 2)) -CHopfIso = - compIso (invIso (equivToIso - (compEquiv (compEquiv pushoutSwitchEquiv - (isoToEquiv (PushoutDistr.PushoutDistrIso fold⋁ W λ _ → tt))) - pushoutSwitchEquiv))) - (equivToIso lem) - - - -S²×S² : Pointed₀ -fst S²×S² = S₊ 2 × S₊ 2 -snd S²×S² = north , north - -q : (S²×S² →∙ (Pushout⋁↪fold⋁ (S₊∙ 2) , inl (north , north))) -fst q = inl -snd q = refl - -qHom : GroupHom (coHomGr 4 (Pushout⋁↪fold⋁ (S₊∙ 2))) - (coHomGr 4 (S₊ 2 × S₊ 2)) -qHom = coHomMorph 4 (fst q) - -qHomGen : (n : ℕ) → - GroupHom (coHomGr n (Pushout⋁↪fold⋁ (S₊∙ 2))) - (coHomGr n (S₊ 2 × S₊ 2)) -qHomGen = λ n → coHomMorph n (fst q) - -module mv2 = MV _ _ (S₊∙ 2 ⋁ S₊∙ 2) ⋁↪ fold⋁ - -gaha : isEquiv (fst (mv2.i 4)) -gaha = - SES→isEquiv - (isContr→≡UnitGroup - (isOfHLevelRetractFromIso 0 - (compIso - (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 2)) - (compIso - (prodIso (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) - (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p)))) - rUnit×Iso)) - isContrUnit)) - ((isContr→≡UnitGroup - (isOfHLevelRetractFromIso 0 - (compIso - (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 3)) - (compIso - (prodIso (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) - (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p)))) - rUnit×Iso)) - isContrUnit))) - (mv2.d 3) - (mv2.i 4) - (mv2.Δ 4) - (mv2.Ker-i⊂Im-d 3) - (mv2.Ker-Δ⊂Im-i 4) - -qHomPP : {A B C : Type} - → Unit ≡ C - → (f : A → B × C) - → isEquiv f - → isEquiv (fst ∘ f) -qHomPP {A = A} {B = B} = - J (λ C _ → (f : A → B × C) → isEquiv f → isEquiv (fst ∘ f)) - λ f eq → record { equiv-proof = - λ b → ((fst (fst (equiv-proof eq (b , tt)))) - , cong fst (fst (equiv-proof eq (b , tt)) .snd)) - , λ y → ΣPathP ((cong fst (equiv-proof eq (b , tt) .snd ((fst y) - , ΣPathP ((snd y) , refl)))) - , λ i j → equiv-proof eq (b , tt) .snd ((fst y) - , ΣPathP ((snd y) , refl)) i .snd j .fst) } - - -isEquiv-qHom : GroupEquiv (coHomGr 4 (Pushout⋁↪fold⋁ (S₊∙ 2))) - (coHomGr 4 (S₊ 2 × S₊ 2)) -fst (fst isEquiv-qHom) = qHom .fst -snd (fst isEquiv-qHom) = - subst isEquiv - (funExt (sElim (λ _ → isSetPathImplicit) (λ _ → refl))) - (qHomPP (isoToPath - (invIso (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))))) _ gaha) -snd isEquiv-qHom = qHom .snd - -coHomCHopfIso : (n : ℕ) - → GroupIso (coHomGr n CHopf) (coHomGr n (Pushout⋁↪fold⋁ (S₊∙ 2))) -coHomCHopfIso n = invGroupIso (coHomIso n CHopfIso) - -H⁴-gen = Iso.inv (fst coHom4S²×S²) 1 -H²-genₗ = Iso.inv (fst coHom2S²×S²) (1 , 0) -H²-genᵣ = Iso.inv (fst coHom2S²×S²) (0 , 1) - - - -open import Cubical.ZCohomology.RingStructure.RingLaws - -module postul (is : GroupIso (coHomGr 4 (S₊ 2 × S₊ 2)) ℤGroup) - (isEq : (fun (fst is) (H²-genₗ' ⌣ H²-genᵣ') ≡ 1)) where - α' : coHom 2 (Pushout⋁↪fold⋁ (S₊∙ 2)) - α' = fun (fst (coHomCHopfIso 2)) Hopfαfold∘W - - β' : coHom 4 (Pushout⋁↪fold⋁ (S₊∙ 2)) - β' = fun (fst (coHomCHopfIso 4)) Hopfβfold∘W - - ⌣→≡ : (α' ⌣ α' ≡ β' +ₕ β') ⊎ (α' ⌣ α' ≡ -ₕ (β' +ₕ β')) - → (Hopfαfold∘W ⌣ Hopfαfold∘W ≡ Hopfβfold∘W +ₕ Hopfβfold∘W) - ⊎ (Hopfαfold∘W ⌣ Hopfαfold∘W ≡ -ₕ (Hopfβfold∘W +ₕ Hopfβfold∘W)) - ⌣→≡ (inl x) = inl ((λ i → leftInv (fst (coHomCHopfIso 2)) Hopfαfold∘W (~ i) - ⌣ leftInv (fst (coHomCHopfIso 2)) Hopfαfold∘W (~ i)) - ∙∙ cong (inv (fst (coHomCHopfIso 4))) x - ∙∙ leftInv (fst (coHomCHopfIso 4)) (Hopfβfold∘W +ₕ Hopfβfold∘W)) - ⌣→≡ (inr x) = inr ((λ i → leftInv (fst (coHomCHopfIso 2)) Hopfαfold∘W (~ i) - ⌣ leftInv (fst (coHomCHopfIso 2)) Hopfαfold∘W (~ i)) - ∙∙ cong (inv (fst (coHomCHopfIso 4))) x - ∙∙ leftInv (fst (coHomCHopfIso 4)) - (-ₕ (Hopfβfold∘W +ₕ Hopfβfold∘W))) - - Q : (qHom .fst β' ≡ H²-genₗ' ⌣ H²-genᵣ') - ⊎ (qHom .fst β' ≡ -ₕ (H²-genₗ' ⌣ H²-genᵣ')) - Q = - ⊎rec - (λ p → inl (sym (leftInv (fst is) (qHom .fst β')) - ∙∙ cong (inv (fst is)) (p ∙ sym isEq) - ∙∙ leftInv (fst is) (H²-genₗ' ⌣ H²-genᵣ'))) - (λ p → inr (sym (leftInv (fst is) (qHom .fst β')) - ∙∙ cong (inv (fst is)) - (p ∙ sym (cong (GroupStr.inv (snd ℤGroup)) isEq)) - ∙∙ (IsGroupHom.presinv (invGroupIso is .snd) (fun (fst is) (H²-genₗ' ⌣ H²-genᵣ')) - ∙ cong -ₕ_ (leftInv (fst is) (H²-genₗ' ⌣ H²-genᵣ'))))) - p2 - where - h : GroupEquiv (coHomGr 4 (HopfInvariantPush 0 fold∘W)) ℤGroup - h = compGroupEquiv (GroupIso→GroupEquiv (coHomCHopfIso 4)) - (compGroupEquiv - isEquiv-qHom - (GroupIso→GroupEquiv is)) - - p2 : (fst (fst h) Hopfβfold∘W ≡ 1) ⊎ (fst (fst h) Hopfβfold∘W ≡ -1) - p2 = groupEquivPresGen _ (GroupIso→GroupEquiv (Hopfβ-Iso 0 (fold∘W , refl))) - Hopfβfold∘W (inl (Hopfβ↦1 0 (fold∘W , refl))) h - - qpres⌣ : (x y : coHom 2 _) - → fst qHom (x ⌣ y) ≡ fst (qHomGen 2) x ⌣ fst (qHomGen 2) y - qpres⌣ = sElim2 (λ _ _ → isSetPathImplicit) λ _ _ → refl - --- H¹(...) → H²(S² ∨ S²) → H²(S² × S²) × H²(S²) → H²(...) - - i* : coHom 2 (Pushout⋁↪fold⋁ (S₊∙ 2)) → coHom 2 (S₊ 2) - i* = coHomFun 2 inr - - kata : i* (coHomFun 2 (Iso.inv CHopfIso) Hopfαfold∘W) ≡ ∣ ∣_∣ ∣₂ - kata = refl - - incl∘q : fst q ∘ (λ x → x , north) ≡ inr - incl∘q = funExt λ x → (push (inl x)) - - incr∘q : fst q ∘ (λ x → north , x) ≡ inr - incr∘q = funExt λ x → (push (inr x)) - - q≡' : coHomFun 2 (λ x → x , north) ∘ fst (qHomGen 2) ≡ i* - q≡' = funExt (sElim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (funExt λ x → cong f (push (inl x)))) - - kzz : (x : _) → i* x ≡ ∣ ∣_∣ ∣₂ → fst (qHomGen 2) x - ≡ ∣ (λ x → ∣ fst x ∣) ∣₂ - +ₕ ∣ (λ x → ∣ snd x ∣) ∣₂ - kzz = sElim (λ _ → isSetΠ λ _ → isSetPathImplicit) - λ f p → Cubical.HITs.PropositionalTruncation.rec - (squash₂ _ _) - (λ r → cong ∣_∣₂ (funExt (uncurry - (wedgeconFun 1 1 (λ _ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) - (λ x → cong f (push (inr x)) ∙∙ funExt⁻ r x ∙∙ refl) - ((λ x → cong f (push (inl x)) ∙∙ funExt⁻ r x ∙∙ sym (rUnitₖ 2 ∣ x ∣ₕ))) - (cong (_∙∙ funExt⁻ r north ∙∙ refl) - (cong (cong f) λ j i → push (push tt j) i)))))) - (Iso.fun PathIdTrunc₀Iso p) - - - q≡'2 : coHomFun 2 (λ x → north , x) ∘ fst (qHomGen 2) ≡ i* - q≡'2 = funExt (sElim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (funExt λ x → cong f (push (inr x)))) - - q₂ : fst (qHomGen 2) α' - ≡ ∣ (λ x → ∣ fst x ∣) ∣₂ +ₕ ∣ (λ x → ∣ snd x ∣) ∣₂ - q₂ = kzz ((coHomFun 2 (Iso.inv CHopfIso) Hopfαfold∘W)) refl - - q₂' : fst (qHomGen 2) α' ≡ H²-genₗ' +ₕ H²-genᵣ' - q₂' = q₂ - - triv⌣ : (a : S¹) → cong₂ (_⌣ₖ_ {n = 2} {m = 2}) (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (merid a)) ≡ λ _ → ∣ north ∣ₕ - triv⌣ a = cong₂Funct (_⌣ₖ_ {n = 2} {m = 2}) (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (merid a)) - ∙ sym (rUnit λ j → _⌣ₖ_ {n = 2} {m = 2} ∣ merid a j ∣ₕ ∣ north ∣) - ∙ (λ i j → ⌣ₖ-0ₖ 2 2 ∣ merid a j ∣ₕ i) - - distLem : (x : coHom 4 (S₊ 2 × S₊ 2)) → -ₕ ((-ₕ x) +ₕ (-ₕ x)) ≡ x +ₕ x - distLem = - sElim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (funExt λ x → cong -ₖ_ (sym (-distrₖ 4 (f x) (f x))) - ∙ -ₖ^2 (f x +ₖ f x)) - - MainId : ((fst qHom) (α' ⌣ α') ≡ qHom .fst (β' +ₕ β')) - ⊎ ((fst qHom) (α' ⌣ α') ≡ qHom .fst (-ₕ (β' +ₕ β'))) - MainId = - ⊎rec - (λ id → inl (h ∙ H - ∙ cong (λ x → x +ₕ x) (sym id) - ∙ sym (IsGroupHom.pres· (snd qHom) β' β'))) - (λ id → inr (h ∙ H - ∙ ((sym (distLem (x ⌣ y)) - ∙ cong -ₕ_ (cong (λ x → x +ₕ x) (sym id))) - ∙ cong (-ₕ_) (IsGroupHom.pres· (snd qHom) β' β')) - ∙ sym (IsGroupHom.presinv (snd qHom) (β' +ₕ β')))) - Q - where - H²-genₗ⌣H²-genₗ : H²-genₗ' ⌣ H²-genₗ' ≡ 0ₕ 4 - H²-genₗ⌣H²-genₗ = - cong ∣_∣₂ - (funExt (uncurry λ { north y → refl - ; south y → refl - ; (merid a i) y j → triv⌣ a j i})) - - H²-genᵣ⌣H²-genᵣ : H²-genᵣ' ⌣ H²-genᵣ' ≡ 0ₕ 4 - H²-genᵣ⌣H²-genᵣ = cong ∣_∣₂ - (funExt (uncurry λ { x north → refl - ; x south → refl - ; x (merid a i) j → triv⌣ a j i})) - - x = H²-genₗ' - y = H²-genᵣ' - - -ₕId : (x : coHom 4 (S₊ 2 × S₊ 2)) → (-ₕ^ 2 · 2) x ≡ x - -ₕId = sElim (λ _ → isSetPathImplicit) - λ a → cong ∣_∣₂ (funExt λ x → -ₖ-gen-inl-left 2 2 tt (inl tt) (a x)) - - H : (H²-genₗ' +ₕ H²-genᵣ') ⌣ (H²-genₗ' +ₕ H²-genᵣ') - ≡ (H²-genₗ' ⌣ H²-genᵣ') +ₕ (H²-genₗ' ⌣ H²-genᵣ') - H = (x +ₕ y) ⌣ (x +ₕ y) ≡⟨ leftDistr-⌣ 2 2 (x +ₕ y) x y ⟩ - ((x +ₕ y) ⌣ x) +ₕ ((x +ₕ y) ⌣ y) ≡⟨ cong₂ _+ₕ_ (rightDistr-⌣ 2 2 x y x) (rightDistr-⌣ 2 2 x y y) ⟩ - ((x ⌣ x +ₕ y ⌣ x)) +ₕ (x ⌣ y +ₕ y ⌣ y) ≡⟨ cong₂ _+ₕ_ (cong₂ _+ₕ_ H²-genₗ⌣H²-genₗ - ((gradedComm-⌣ 2 2 y x) - ∙ cong (-ₕ^ 2 · 2) - (transportRefl (x ⌣ y)) - ∙ -ₕId (x ⌣ y)) - ∙ lUnitₕ 4 (x ⌣ y)) - (cong (x ⌣ y +ₕ_) H²-genᵣ⌣H²-genᵣ ∙ rUnitₕ 4 (x ⌣ y)) ⟩ - ((x ⌣ y) +ₕ (x ⌣ y)) ∎ - - h : (fst qHom) (α' ⌣ α') ≡ (H²-genₗ' +ₕ H²-genᵣ') ⌣ (H²-genₗ' +ₕ H²-genᵣ') - h = (fst qHom) (α' ⌣ α') ≡⟨ refl ⟩ - fst (qHomGen 2) α' ⌣ fst (qHomGen 2) α' ≡⟨ cong (λ x → x ⌣ x) q₂ ⟩ - ((H²-genₗ' +ₕ H²-genᵣ') ⌣ (H²-genₗ' +ₕ H²-genᵣ')) ∎ - - HopfInvariantLem : (HopfInvariant 0 (fold∘W , refl) ≡ 2) - ⊎ (HopfInvariant 0 (fold∘W , refl) ≡ -2) - HopfInvariantLem = - ⊎rec (λ p → inl (p1 - ∙ cong (Iso.fun (fst (Hopfβ-Iso 0 (fold∘W , refl)))) p - ∙ IsGroupHom.pres· (Hopfβ-Iso 0 (fold∘W , refl) .snd) Hopfβfold∘W Hopfβfold∘W - ∙ cong (λ x → x + x) (Hopfβ↦1 0 (fold∘W , refl)))) - (λ p → inr (p1 - ∙ cong (Iso.fun (fst (Hopfβ-Iso 0 (fold∘W , refl)))) p - ∙ IsGroupHom.presinv (Hopfβ-Iso 0 (fold∘W , refl) .snd) (Hopfβfold∘W +ₕ Hopfβfold∘W) - ∙ cong (GroupStr.inv (snd ℤGroup)) - (IsGroupHom.pres· (Hopfβ-Iso 0 (fold∘W , refl) .snd) Hopfβfold∘W Hopfβfold∘W - ∙ cong (λ x → x + x) (Hopfβ↦1 0 (fold∘W , refl))))) - p2 - where - p1 : HopfInvariant 0 (fold∘W , refl) - ≡ Iso.fun (fst (Hopfβ-Iso 0 (fold∘W , refl))) (Hopfαfold∘W ⌣ Hopfαfold∘W) - p1 = cong (Iso.fun (fst (Hopfβ-Iso 0 (fold∘W , refl)))) (transportRefl (Hopfαfold∘W ⌣ Hopfαfold∘W)) - - p2 : (Hopfαfold∘W ⌣ Hopfαfold∘W ≡ Hopfβfold∘W +ₕ Hopfβfold∘W) - ⊎ (Hopfαfold∘W ⌣ Hopfαfold∘W ≡ -ₕ (Hopfβfold∘W +ₕ Hopfβfold∘W)) - p2 = ⌣→≡ (⊎rec (λ p → inl (sym (retEq (fst isEquiv-qHom) (α' ⌣ α')) - ∙∙ cong (invEq (fst isEquiv-qHom)) p - ∙∙ retEq (fst isEquiv-qHom) (β' +ₕ β'))) - (λ p → inr ((sym (retEq (fst isEquiv-qHom) (α' ⌣ α')) - ∙∙ cong (invEq (fst isEquiv-qHom)) p - ∙∙ retEq (fst isEquiv-qHom) (-ₕ (β' +ₕ β'))))) - MainId) - -HopfInvariantLemfold∘W : abs (HopfInvariant 0 (fold∘W , refl)) ≡ 2 -HopfInvariantLemfold∘W = - ⊎→abs (HopfInvariant 0 (fold∘W , refl)) 2 - (postul.HopfInvariantLem - (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1) Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-⌣) - -Whitehead≡ : [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π' ≡ ∣ fold∘W , refl ∣₂ -Whitehead≡ = - cong ∣_∣₂ ([]≡[]₂ (idfun∙ (S₊∙ 2)) (idfun∙ (S₊∙ 2)) ) - ∙ sym Fold∘W - ∙ cong ∣_∣₂ (∘∙-idˡ (fold∘W , refl)) - -HopfInvariantWhitehead : - abs (HopfInvariant-π' 0 [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π') ≡ 2 -HopfInvariantWhitehead = - cong abs (cong (HopfInvariant-π' 0) Whitehead≡) ∙ HopfInvariantLemfold∘W From 934171020cab0a9fc68f298d4c34a9958c621ca1 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 2 Jun 2023 15:04:54 +0200 Subject: [PATCH 05/73] ganea stuff --- Cubical/Foundations/Prelude.agda | 8 + Cubical/HITs/Join/Properties.agda | 209 +++++++++++ Cubical/HITs/SmashProduct/Ganea.agda | 14 + .../Group/Pi4S3/NewBrunerieNumbers.agda | 326 ++++++++++++++++++ 4 files changed, 557 insertions(+) create mode 100644 Cubical/HITs/SmashProduct/Ganea.agda create mode 100644 Cubical/Homotopy/Group/Pi4S3/NewBrunerieNumbers.agda diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index a402935890..21b1675a2f 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -171,6 +171,14 @@ compPath-filler' {z = z} p q j i = ; (i = i1) → q k ; (j = i0) → q (i ∧ k) }) (p (i ∨ ~ j)) + +compPath-filler'-filler : (p : x ≡ y) (q : y ≡ z) → (i j k : I) → _ +compPath-filler'-filler {z = z} p q j i k = + hfill (λ k → λ { (i = i0) → p (~ j) + ; (i = i1) → q k + ; (j = i0) → q (i ∧ k) }) + (inS (p (i ∨ ~ j))) + k -- Note: We can omit a (j = i1) case here since when (j = i1), the whole expression is -- definitionally equal to `p ∙ q`. (Notice that `p ∙ q` is also an hcomp.) Nevertheless, -- we could have given `compPath-filler p q k i` as the (j = i1) case. diff --git a/Cubical/HITs/Join/Properties.agda b/Cubical/HITs/Join/Properties.agda index e465083060..76b4a88249 100644 --- a/Cubical/HITs/Join/Properties.agda +++ b/Cubical/HITs/Join/Properties.agda @@ -482,3 +482,212 @@ fst joinAnnihilL = inl tt* snd joinAnnihilL (inl tt*) = refl snd joinAnnihilL (inr a) = push tt* a snd joinAnnihilL (push tt* a i) j = push tt* a (i ∧ j) + +open import Cubical.Foundations.Pointed +open import Cubical.Homotopy.Loopspace + +module _ {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where + fib⋆ : Pointed _ + proj₁ fib⋆ = fiber (fst f) (pt B) + snd fib⋆ = (pt A) , (snd f) + + fib-gen : (b : fst B) → Type _ + fib-gen b = fiber (fst f) b + + fib-proj : fib⋆ →∙ A + proj₁ fib-proj = fst + snd fib-proj = refl + + fib-cofib : Type _ + fib-cofib = cofib (fst fib-proj) + + GaneaMap : fib-cofib → fst B + GaneaMap (inl x) = pt B + GaneaMap (inr x) = fst f x + GaneaMap (push a i) = a .snd (~ i) + + GaneaMap* : (b : fst B) (f : A →∙ (fst B , b)) → cofib {A = fiber (fst f) b} {B = fst A} fst → fst B + GaneaMap* b f (inl x) = b + GaneaMap* b f (inr x) = fst f x + GaneaMap* b f (push a i) = a .snd (~ i) + + GaneaFib : Type _ + GaneaFib = fiber GaneaMap (pt B) + + GaneaFib→join-unc' : (x : fib-cofib) → join (fib-gen (GaneaMap x)) (GaneaMap x ≡ GaneaMap x) + GaneaFib→join-unc' (inl x) = inr refl + GaneaFib→join-unc' (inr x) = inl (x , refl) + GaneaFib→join-unc' (push (a , p) i) = push (a , (λ t → p (~ i ∧ t))) refl (~ i) + + GaneaFib→join-unc : (x : fib-cofib) (b : fst B) → GaneaMap x ≡ b → join (fib-gen b) (b ≡ b) + GaneaFib→join-unc x = J> GaneaFib→join-unc' x + + other-gen : (b : fst B) → join (fiber (fst f) b) (b ≡ b) → GaneaFib + other-gen b (inl x) = (inr (fst x)) , {!snd x!} + other-gen b (inr x) = {!!} + other-gen b (push a b₁ i) = {!!} + + other-fill : (a : fib⋆ .fst) (p : Ω B .fst) → (i j k : I) → fst B + other-fill a b i j k = + hfill (λ k → λ {(i = i0) → snd a j + ; (i = i1) → b (j ∨ ~ k) + ; (j = i0) → compPath-filler (snd a) (sym b) k i + ; (j = i1) → snd B}) + (inS (snd a (j ∨ i))) + k + + other : join (fib⋆ .fst) (Ω B .fst) → GaneaFib + other (inl x) = inr (fst x) , snd x + other (inr x) = (inl tt) , x + proj₁ (other (push a b i)) = (push (fst a , snd a ∙ sym b)) (~ i) -- + snd (other (push a b i)) j = other-fill a b i j i1 + + c1 : (i j k : I) → (a : fst A) (q : fst f a ≡ pt B) (p : q (~ i) ≡ pt B) + → fst B + c1 i j k a q p = + hfill (λ k + → λ {(i = i0) → p j -- p (j ∧ k) + ; (i = i1) → compPath-filler' (sym q) p k j -- p j + ; (j = i0) → q (k ∨ ~ i) -- pt B + ; (j = i1) → pt B}) -- p k}) + (inS (p j)) -- (inS (q (~ i ∨ ~ j))) + k + + c2 : (i k : I) → (a : fst A) (q : fst f a ≡ pt B) (p : q (~ i) ≡ pt B) + → join (fib⋆ .fst) (Ω B .fst) + c2 i k a q p = + hfill (λ k → λ {(i = i0) → inr p + ; (i = i1) → push (a , p) (sym q ∙ p) (~ k)}) + (inS (inr λ j → c1 i j i1 a q p)) k + + GaneaFib→join : GaneaFib → join (fib⋆ .fst) (Ω B .fst) + GaneaFib→join (inl x , p) = inr p + GaneaFib→join (inr x , p) = inl (x , p) + GaneaFib→join (push (a , q) i , p) = c2 i i1 a q p + + cancel : (x : GaneaFib) → other (GaneaFib→join x) ≡ x + cancel (inl x , y) = refl + cancel (inr x , y) = refl + cancel (push (a , q) i , p) j = + hcomp (λ k → λ {(i = i0) → inl tt , p + ; (i = i1) → mg p k j + ; (j = i0) → other (c2 i k a q p) + ; (j = i1) → push (a , q) i , p}) + ((push (a , q) (i ∧ j)) + , λ k → hcomp (λ r → λ {(i = i0) → p k + ; (i = i1) → compPath-filler' (sym q) p (r ∧ (~ j)) k + ; (j = i0) → c1 i k r a q p + ; (j = i1) → p k + ; (k = i0) → q ((r ∧ ~ j) ∨ ~ i) + ; (k = i1) → snd B}) + (p k)) + where + fill1 : (i j k : I) (p : fst f a ≡ pt B) + → fst B + fill1 i j k p = + hfill (λ k → λ {(i = i0) → compPath-filler p (sym (sym q ∙ p)) k j + ; (i = i1) → q (j ∨ ~ k) + ; (j = i0) → q (~ k ∧ i) + ; (j = i1) → ((λ i₂ → q (~ i₂)) ∙ p) (~ k ∧ ~ i)}) + (inS (compPath-filler (sym q) p j (~ i))) k + + lazy : ∀ {ℓ} {A : Type ℓ} {x : A} (y : A) (q : x ≡ y) (z : A) (p : x ≡ z) + → PathP (λ k → Square (λ j → q (j ∨ ~ k)) refl + (compPath-filler' (λ i₂ → q (~ i₂)) p (~ k)) (sym q ∙ p)) + (λ i _ → (sym q ∙ p) i) λ i j → compPath-filler' (sym q) p j i + lazy {x = x} = J> (J> h x refl (refl ∙ refl) (compPath-filler' (sym refl) refl)) + where + h : ∀ {ℓ} {A : Type ℓ} {x : A} (y : A) (p : x ≡ y) (q : x ≡ y) (r : p ≡ q) + → PathP (λ k → Square refl refl (r (~ k)) q) (λ i _ → q i) λ i j → r j i + h = J> (J> refl) + + lazy2 : ∀ {ℓ} {A : Type ℓ} {x : A} (y : A) (q : x ≡ y) (z : A) (p : x ≡ z) + → PathP (λ r → Square (λ i → compPath-filler' (sym q) p i r) (λ i → (sym q ∙ p) (r ∨ ~ i)) (λ j → p (r ∨ j)) refl) + (λ j i → compPath-filler (sym q) p j (~ i)) + refl + lazy2 {x = x} = J> (J> λ i j k → {!lazy x refl _ refl (~ i) k (~ j)!}) + + help : (p : fst f a ≡ pt B) + → cong other (push (a , p) (sym q ∙ p)) + ≡ λ i → (push (a , q) (~ i)) , (compPath-filler' (sym q) p i) + proj₁ (help p i j) = push (a , λ j → fill1 i j i1 p) (~ j) + snd (help p i j) r = + hcomp (λ k → λ {(i = i0) → other-fill (a , p) (sym q ∙ p) j r k + ; (i = i1) → lazy _ q _ p k r j -- compPath-filler'-filler (sym q) p j r k + ; (j = i0) → compPath-filler' (sym q) p (~ k ∧ i) r + ; (j = i1) → (sym q ∙ p) (r ∨ ~ (k ∨ i)) + ; (r = i0) → fill1 i j k p + ; (r = i1) → snd B}) + (lazy2 _ q _ p r j i) + {- + r = i0 ⊢ +r = i1 ⊢ snd B +j = i0 ⊢ p r +j = i1 ⊢ hcomp (doubleComp-faces (λ _ → snd B) p r) (q (~ r)) +i = i0 ⊢ hcomp + (λ { k (j = i0) → p r + ; k (j = i1) + → hcomp (doubleComp-faces (λ _ → snd B) p (r ∨ ~ k)) + (q (~ (r ∨ ~ k))) + ; k (r = i0) + → compPath-filler p + (λ i₂ → + hcomp (doubleComp-faces (λ _ → snd B) p (~ i₂)) (q (~ (~ i₂)))) + k j + ; k (r = i1) → snd B + }) + (p (r ∨ j)) +i = i1 ⊢ hcomp + (λ { k (r = i0) → q (~ (~ j)) + ; k (r = i1) → p k + ; k (j = i0) → p (r ∧ k) + }) + (q (~ r ∧ j)) + -} + {- + hcomp (λ k → λ {(i = i0) → {!c2 j i1 a p (sym q ∙ p)!} + ; (i = i1) → {!!} + ; (j = i0) → {!!} + ; (j = i1) → {!!}}) + {!!} -} + -- other (push (a , p) (sym q ∙ p) j) + open import Cubical.Foundations.Path + + mg : (p : fst f a ≡ pt B) + → PathP (λ k → other (push (a , p) (sym q ∙ p) (~ k)) ≡ (inr a , p)) + (λ i → push (a , q) i , compPath-filler' (sym q) p (~ i)) + refl + mg p = flipSquare (cong sym (help p) + ◁ λ j i → push (a , q) (j ∨ i) + , compPath-filler' (sym q) p (~ (j ∨ i))) + + cool : PathP (λ j → other {!c2 i i1 a q !} ≡ {!!}) {!!} {!!} + cool = {!j = i0 ⊢ inr a , p +j = i1 ⊢ inl tt , sym q ∙ p +i = i0 ⊢ other (push (a , p) (sym q ∙ p) j) +i = i1 ⊢ push (a , q) (~ j) , compPath-filler' (sym q) p j!} + + + {- + Goal: join GaneaFib (Ω B .proj₁) +———— Boundary (wanted) ————————————————————————————————————— +i = i0 ⊢ inr b +i = i1 ⊢ inl (inr a , b) + -} + {- help a p i b + where + help : (a : fst A) (p : fst f a ≡ pt B) + → PathP (λ i → (b : (p (~ i) ≡ pt B)) → join GaneaFib (Ω B .fst)) + inr + λ b → inl (inr a , b) + help a p = {!!} -} + + -- (sym (push {!!} {!!}) ∙∙ {!!} ∙∙ {!!}) i -- push ((inr (fst a)) , {!a!}) {!!} (~ i) + + lem : join GaneaFib (Ω B .fst) → GaneaFib + lem (inl x) = x + lem (inr x) = inl tt , x + lem (push (inl x , p) b i) = (push ((pt A) , (snd f ∙ p)) ∙ sym (push ((pt A) , (snd f ∙ b)))) i , {!!} + lem (push (inr x , p) b i) = {!i = i0 ⊢ inr b +i = i1 ⊢ inl (inr a , b)!} + lem (push (push a i₁ , p) b i) = {!!} diff --git a/Cubical/HITs/SmashProduct/Ganea.agda b/Cubical/HITs/SmashProduct/Ganea.agda new file mode 100644 index 0000000000..339b350f0d --- /dev/null +++ b/Cubical/HITs/SmashProduct/Ganea.agda @@ -0,0 +1,14 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.SmashProduct.Ganea where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.Pushout.Base +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.HITs.Wedge +open import Cubical.Foundations.GroupoidLaws +open import Cubical.HITs.SmashProduct.Base + + diff --git a/Cubical/Homotopy/Group/Pi4S3/NewBrunerieNumbers.agda b/Cubical/Homotopy/Group/Pi4S3/NewBrunerieNumbers.agda new file mode 100644 index 0000000000..1f221f2518 --- /dev/null +++ b/Cubical/Homotopy/Group/Pi4S3/NewBrunerieNumbers.agda @@ -0,0 +1,326 @@ +-- Mostly self contained definitions of the numbers from: https://arxiv.org/abs/2302.00151 +{-# OPTIONS --safe #-} +module Cubical.Homotopy.Group.Pi4S3.NewBrunerieNumbers where + +open import Cubical.Core.Primitives + +open import Cubical.Foundations.Prelude using + (_,_ ; fst ; snd ; _≡_ ; cong ; _∙_ ; refl ; isProp→isSet ; J ; sym ; isProp ; isSet ; isGroupoid ; is2Groupoid) +open import Cubical.Foundations.Equiv using + (_≃_ ; isEquiv ; isPropIsEquiv ; idIsEquiv ; idEquiv) +open import Cubical.Foundations.HLevels using + (hSet ; hGroupoid ; isOfHLevelTypeOfHLevel ; isOfHLevelPath ; isOfHLevelΠ ; isOfHLevel→isOfHLevelDep ; is2GroupoidΠ) +open import Cubical.Foundations.Univalence using + (Glue ; ua) +open import Cubical.Data.Int using + (ℤ ; pos ; neg ; isSetℤ ; sucPathℤ) +open import Cubical.Foundations.CartesianKanOps + +private variable ℓ ℓ' ℓ'' : Level + +-- S1 +data S¹ : Type₀ where + base : S¹ + loop : base ≡ base + +helix : S¹ → Type₀ +helix base = ℤ +helix (loop i) = sucPathℤ i + +rotLoop : (a : S¹) → a ≡ a +rotLoop base = loop +rotLoop (loop i) j = + hcomp (λ k → λ { (i = i0) → loop (j ∨ ~ k) + ; (i = i1) → loop (j ∧ k) + ; (j = i0) → loop (i ∨ ~ k) + ; (j = i1) → loop (i ∧ k)}) base + +_·_ : S¹ → S¹ → S¹ +base · x = x +(loop i) · x = rotLoop x i + +isPropFamS¹ : (P : S¹ → Type ℓ) (pP : (x : S¹) → isProp (P x)) (b0 : P base) → + PathP (λ i → P (loop i)) b0 b0 +isPropFamS¹ P pP b0 i = pP (loop i) (coe0→i (λ k → P (loop k)) i b0) + (coe1→i (λ k → P (loop k)) i b0) i + +rotIsEquiv : (a : S¹) → isEquiv (a ·_) +rotIsEquiv base = idIsEquiv S¹ +rotIsEquiv (loop i) = isPropFamS¹ (λ x → isEquiv (x ·_)) + (λ x → isPropIsEquiv (x ·_)) + (idIsEquiv _) i + +-- S2 + +-- Direct S², useful for later part of the definition (or is it?) +data S²' : Type₀ where + base : S²' + surf : PathP (λ i → base ≡ base) refl refl + +S²ToSetElim' : {A : S²' → Type ℓ} + → ((x : S²') → isSet (A x)) + → A base + → (x : S²') → A x +S²ToSetElim' set b base = b +S²ToSetElim' set b (surf i j) = + isOfHLevel→isOfHLevelDep 2 set b b {a0 = refl} {a1 = refl} refl refl surf i j + +-- Join +data join (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') where + inl : A → join A B + inr : B → join A B + push : ∀ a b → inl a ≡ inr b + + +-- SetTruncation +data ∥_∥₀ (A : Type ℓ) : Type ℓ where + ∣_∣₀ : A → ∥ A ∥₀ + squash₀ : ∀ (x y : ∥ A ∥₀) (p q : x ≡ y) → p ≡ q + +rec₀ : {A : Type ℓ} {B : Type ℓ'} → isSet B → (A → B) → ∥ A ∥₀ → B +rec₀ Bset f ∣ x ∣₀ = f x +rec₀ Bset f (squash₀ x y p q i j) = + Bset _ _ (cong (rec₀ Bset f) p) (cong (rec₀ Bset f) q) i j + + +-- GroupoidTruncation +data ∥_∥₁ (A : Type ℓ) : Type ℓ where + ∣_∣₁ : A → ∥ A ∥₁ + squash₁ : ∀ (x y : ∥ A ∥₁) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s + +rec₁ : {A : Type ℓ} {B : Type ℓ'} → isGroupoid B → (A → B) → ∥ A ∥₁ → B +rec₁ gB f ∣ x ∣₁ = f x +rec₁ gB f (squash₁ x y p q r s i j k) = + gB _ _ _ _ (λ m n → rec₁ gB f (r m n)) (λ m n → rec₁ gB f (s m n)) i j k + + +-- 2GroupoidTruncation +data ∥_∥₂ (A : Type ℓ) : Type ℓ where + ∣_∣₂ : A → ∥ A ∥₂ + squash₂ : ∀ (x y : ∥ A ∥₂) (p q : x ≡ y) (r s : p ≡ q) (t u : r ≡ s) → t ≡ u + +rec₂ : ∀ {A : Type ℓ} {B : Type ℓ'} → is2Groupoid B → (A → B) → ∥ A ∥₂ → B +rec₂ gB f ∣ x ∣₂ = f x +rec₂ gB f (squash₂ _ _ _ _ _ _ t u i j k l) = + gB _ _ _ _ _ _ (λ m n o → rec₂ gB f (t m n o)) (λ m n o → rec₂ gB f (u m n o)) + i j k l + +-- rec₂Bin : ∀ {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → is2Groupoid C → (A → B → C) → ∥ A ∥₂ → ∥ B ∥₂ → C +-- rec₂Bin gB f = rec₂ (is2GroupoidΠ λ _ → gB) λ x → rec₂ gB (f x) + +elim₂ : {A : Type ℓ} {B : ∥ A ∥₂ → Type ℓ} + (bG : (x : ∥ A ∥₂) → is2Groupoid (B x)) + (f : (x : A) → B ∣ x ∣₂) (x : ∥ A ∥₂) → B x +elim₂ bG f ∣ x ∣₂ = f x +elim₂ bG f (squash₂ x y p q r s u v i j k l) = + isOfHLevel→isOfHLevelDep 4 bG _ _ _ _ _ _ + (λ j k l → elim₂ bG f (u j k l)) (λ j k l → elim₂ bG f (v j k l)) + (squash₂ x y p q r s u v) + i j k l + + +-- Susp +data Susp (A : Type ℓ) : Type ℓ where + north : Susp A + south : Susp A + merid : (a : A) → north ≡ south + +S² S³ : Type₀ +S² = Susp S¹ +S³ = Susp S² + +-- Pointed +Pointed₀ : Type₁ +Pointed₀ = Σ[ X ∈ Type₀ ] X + +S¹∙ S²'∙ : Pointed₀ +S¹∙ = (S¹ , base) +S²'∙ = (S²' , base) + +Susp∙ : Type₀ → Pointed₀ +Susp∙ A = Susp A , north + +∥_∥₁∙ ∥_∥₂∙ : Pointed₀ → Pointed₀ +∥ A , a ∥₁∙ = ∥ A ∥₁ , ∣ a ∣₁ +∥ A , a ∥₂∙ = ∥ A ∥₂ , ∣ a ∣₂ + +Ω Ω² : Pointed₀ → Pointed₀ +Ω (_ , a) = ((a ≡ a) , refl) +Ω² p = Ω (Ω p) + +-- The maps +σ' : S²' → Ω (Susp∙ S²') .fst +σ' x = merid x ∙ sym (merid base) + +S¹×S¹→S²' : S¹ → S¹ → S²' +S¹×S¹→S²' base y = base +S¹×S¹→S²' (loop i) base = base +S¹×S¹→S²' (loop i) (loop j) = surf i j + +f7 : Ω (Susp∙ S²') .fst → ∥ S²' ∥₂ +f7 p = coe0→1 (λ i → Code (p i)) ∣ base ∣₂ + where + _+₂_ : ∥ S²' ∥₂ → ∥ S²' ∥₂ → ∥ S²' ∥₂ + _+₂_ = rec₂ (is2GroupoidΠ λ _ → squash₂) + λ { base x → x + ; (surf i j) x → surfc x i j} + where + surfc : (x : ∥ S²' ∥₂) → refl {x = x} ≡ refl {x = x} + surfc = elim₂ (λ x → isOfHLevelPath 4 (isOfHLevelPath 4 squash₂ _ _) refl refl) + (S²ToSetElim' (λ _ → squash₂ _ _ _ _) λ i j → ∣ surf i j ∣₂) + + ∥S²'∥₂≃∥S²'∥₂ : (x : S²') → ∥ S²' ∥₂ ≃ ∥ S²' ∥₂ + fst (∥S²'∥₂≃∥S²'∥₂ x) = ∣ x ∣₂ +₂_ + snd (∥S²'∥₂≃∥S²'∥₂ x) = help x + where + help : (x : _) → isEquiv (λ y → ∣ x ∣₂ +₂ y) + help = S²ToSetElim' (λ _ → isProp→isSet (isPropIsEquiv _)) (idIsEquiv _) + + Code : Susp S²' → Type₀ + Code north = ∥ S²' ∥₂ + Code south = ∥ S²' ∥₂ + Code (merid a i) = ua (∥S²'∥₂≃∥S²'∥₂ a) i + +g8 : Ω² ∥ S²'∙ ∥₂∙ .fst → Ω ∥ S¹∙ ∥₁∙ .fst +g8 p i = coe0→1 (λ j → codeTruncS²' (p i j) .fst) ∣ base ∣₁ + where + HopfS²' : S²' → Type₀ + HopfS²' base = S¹ + HopfS²' (surf i j) = Glue S¹ (λ { (i = i0) → S¹ , idEquiv S¹ + ; (i = i1) → S¹ , idEquiv S¹ + ; (j = i0) → S¹ , idEquiv S¹ + ; (j = i1) → S¹ , (loop i) ·_ , rotIsEquiv (loop i) } ) + + codeTruncS²' : ∥ S²' ∥₂ → hGroupoid _ + codeTruncS²' = rec₂ (isOfHLevelTypeOfHLevel 3) (λ s → ∥ HopfS²' s ∥₁ , squash₁) + +g9 : Ω ∥ S¹∙ ∥₁∙ .fst → ∥ ℤ ∥₀ +g9 p = coe0→1 (λ i → codeTruncS¹ (p i) .fst) ∣ pos 0 ∣₀ + where + codeTruncS¹ : ∥ S¹ ∥₁ → hSet _ + codeTruncS¹ = rec₁ (isOfHLevelTypeOfHLevel 2) (λ s → ∥ helix s ∥₀ , squash₀) + +-- Use trick to eliminate away the truncation last +g10 : ∥ ℤ ∥₀ → ℤ +g10 = rec₀ isSetℤ (λ x → x) + +-- Original η₃ as in the paper +η₃ : join S¹ S¹ → Susp S²' +η₃ (inl x) = north +η₃ (inr x) = north +η₃ (push a b i) = + (sym (σ' (S¹×S¹→S²' a b)) ∙ sym (σ' (S¹×S¹→S²' a b))) i + +-- Dropping the syms makes it compute fast (why?! maybe some slow inverse map gets applied with the sym?) +η₃' : join S¹ S¹ → Susp S²' +η₃' (inl x) = north +η₃' (inr x) = north +η₃' (push a b i) = + (σ' (S¹×S¹→S²' a b) ∙ σ' (S¹×S¹→S²' a b)) i + +-- Remark: dropping only one sym does not seem to make it fast enough + +-- Brunerie numbers + +β₃ : ℤ +β₃ = g10 (g9 (g8 λ i j → f7 λ k → η₃ (push (loop i) (loop j) k))) + +-- This does not terminate fast +-- β₃≡-2 : β₃ ≡ -2 +-- β₃≡-2 = refl + +β₃' : ℤ +β₃' = g10 (g9 (g8 λ i j → f7 λ k → η₃' (push (loop i) (loop j) k))) + +-- This terminates fast +β₃'≡-2 : β₃' ≡ -2 +β₃'≡-2 = refl + +β₃'-pos : ℤ +β₃'-pos = g10 (g9 (g8 λ i j → f7 λ k → η₃' (push (loop (~ i)) (loop j) k))) + +β₃'≡2 : β₃'-pos ≡ pos 2 +β₃'≡2 = refl + + +-- We now define β₂ + +invLooper : S¹ → S¹ +invLooper base = base +invLooper (loop i) = loop (~ i) + +η₂ : join S¹ S¹ → join S¹ S¹ +η₂ (inl x) = inr (invLooper x) +η₂ (inr x) = inr x +η₂ (push a b i) = + (sym (push (b · invLooper a) (invLooper a)) ∙ push (b · invLooper a) b) i + +rCancel : {A : Type ℓ} {x y : A} (p : x ≡ y) → p ∙ sym p ≡ refl +rCancel {x = x} p j i = + hcomp (λ k → λ { (i = i0) → x + ; (i = i1) → p (~ k ∧ ~ j) + ; (j = i1) → x + }) (p (i ∧ ~ j)) + +_∙∙_∙∙_ : {A : Type ℓ} {w x y z : A} → w ≡ x → x ≡ y → y ≡ z → w ≡ z +(p ∙∙ q ∙∙ r) i = + hcomp (λ j → λ { (i = i0) → p (~ j) ; (i = i1) → r j}) (q i) + +S¹×S¹→S² : S¹ → S¹ → S² +S¹×S¹→S² base y = north +S¹×S¹→S² (loop i) base = north +S¹×S¹→S² (loop i) (loop j) = + (sym (rCancel (merid base)) + ∙∙ (λ i → merid (loop i) ∙ sym (merid base)) + ∙∙ rCancel (merid base)) i j + +joinS¹S¹→S³ : join S¹ S¹ → S³ +joinS¹S¹→S³ (inl x) = north +joinS¹S¹→S³ (inr x) = north +joinS¹S¹→S³ (push a b i) = (merid (S¹×S¹→S² a b) ∙ sym (merid north)) i + +S²→S²' : S² → S²' +S²→S²' north = base +S²→S²' south = base +S²→S²' (merid base i) = base +S²→S²' (merid (loop i) j) = surf i j + +S³→SuspS²' : S³ → Susp S²' +S³→SuspS²' north = north +S³→SuspS²' south = north +S³→SuspS²' (merid x i) = (merid (S²→S²' x) ∙ sym (merid base)) i + +joinS¹S¹→SuspS²' : join S¹ S¹ → Susp S²' +joinS¹S¹→SuspS²' x = S³→SuspS²' (joinS¹S¹→S³ x) + +β₂ : ℤ +β₂ = g10 (g9 (g8 {!!})) -- λ i j → {!f7 -- joinS¹S¹→SuspS²' (η₂ (push (loop i0) (loop j) ?))!})) + where + _→∙_ : Type + _→∙_ = {!Σ[ f ∈ (Pointed₀ → Pointed₀) ] ?!} + + m : Ω (Susp∙ S²') .fst → ∥ S²'∙ ∥₂∙ .fst + m f7 = {!!} + + t : Ω (Ω² (Susp∙ S²')) .fst + t i j k = + hcomp (λ r → λ { (i = i0) → {!!} + ; (i = i1) → {!η₂ (push (loop i) (loop j) k)!} + ; (j = i0) → {!north!} + ; (j = i1) → {!north!} + ; (k = i0) → north + ; (k = i1) → north + }) + (joinS¹S¹→SuspS²' + (η₂ (push (loop i) (loop j) k))) + + h : Ω² ∥ S²'∙ ∥₂∙ .fst -- Ω² ∥ S²' ∥₂ .fst + h i j = f7 λ k → t i j k + + +-- β₂ : ℤ +-- β₂ = g10 (g9 (g8 λ i j → f7 (λ k → joinS¹S¹→SuspS²' (η₂ (push (loop i) (loop j) k))))) + + + +-- -- TODO: define β₁ From ea75092e20dc74fd00e788088ad7a35cbe5dfc97 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 8 Jun 2023 09:44:37 +0200 Subject: [PATCH 06/73] w --- Cubical/HITs/Join/Properties.agda | 209 ----------- Cubical/HITs/SmashProduct/Ganea.agda | 14 - .../Group/Pi4S3/NewBrunerieNumbers.agda | 326 ------------------ 3 files changed, 549 deletions(-) delete mode 100644 Cubical/HITs/SmashProduct/Ganea.agda delete mode 100644 Cubical/Homotopy/Group/Pi4S3/NewBrunerieNumbers.agda diff --git a/Cubical/HITs/Join/Properties.agda b/Cubical/HITs/Join/Properties.agda index 76b4a88249..e465083060 100644 --- a/Cubical/HITs/Join/Properties.agda +++ b/Cubical/HITs/Join/Properties.agda @@ -482,212 +482,3 @@ fst joinAnnihilL = inl tt* snd joinAnnihilL (inl tt*) = refl snd joinAnnihilL (inr a) = push tt* a snd joinAnnihilL (push tt* a i) j = push tt* a (i ∧ j) - -open import Cubical.Foundations.Pointed -open import Cubical.Homotopy.Loopspace - -module _ {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where - fib⋆ : Pointed _ - proj₁ fib⋆ = fiber (fst f) (pt B) - snd fib⋆ = (pt A) , (snd f) - - fib-gen : (b : fst B) → Type _ - fib-gen b = fiber (fst f) b - - fib-proj : fib⋆ →∙ A - proj₁ fib-proj = fst - snd fib-proj = refl - - fib-cofib : Type _ - fib-cofib = cofib (fst fib-proj) - - GaneaMap : fib-cofib → fst B - GaneaMap (inl x) = pt B - GaneaMap (inr x) = fst f x - GaneaMap (push a i) = a .snd (~ i) - - GaneaMap* : (b : fst B) (f : A →∙ (fst B , b)) → cofib {A = fiber (fst f) b} {B = fst A} fst → fst B - GaneaMap* b f (inl x) = b - GaneaMap* b f (inr x) = fst f x - GaneaMap* b f (push a i) = a .snd (~ i) - - GaneaFib : Type _ - GaneaFib = fiber GaneaMap (pt B) - - GaneaFib→join-unc' : (x : fib-cofib) → join (fib-gen (GaneaMap x)) (GaneaMap x ≡ GaneaMap x) - GaneaFib→join-unc' (inl x) = inr refl - GaneaFib→join-unc' (inr x) = inl (x , refl) - GaneaFib→join-unc' (push (a , p) i) = push (a , (λ t → p (~ i ∧ t))) refl (~ i) - - GaneaFib→join-unc : (x : fib-cofib) (b : fst B) → GaneaMap x ≡ b → join (fib-gen b) (b ≡ b) - GaneaFib→join-unc x = J> GaneaFib→join-unc' x - - other-gen : (b : fst B) → join (fiber (fst f) b) (b ≡ b) → GaneaFib - other-gen b (inl x) = (inr (fst x)) , {!snd x!} - other-gen b (inr x) = {!!} - other-gen b (push a b₁ i) = {!!} - - other-fill : (a : fib⋆ .fst) (p : Ω B .fst) → (i j k : I) → fst B - other-fill a b i j k = - hfill (λ k → λ {(i = i0) → snd a j - ; (i = i1) → b (j ∨ ~ k) - ; (j = i0) → compPath-filler (snd a) (sym b) k i - ; (j = i1) → snd B}) - (inS (snd a (j ∨ i))) - k - - other : join (fib⋆ .fst) (Ω B .fst) → GaneaFib - other (inl x) = inr (fst x) , snd x - other (inr x) = (inl tt) , x - proj₁ (other (push a b i)) = (push (fst a , snd a ∙ sym b)) (~ i) -- - snd (other (push a b i)) j = other-fill a b i j i1 - - c1 : (i j k : I) → (a : fst A) (q : fst f a ≡ pt B) (p : q (~ i) ≡ pt B) - → fst B - c1 i j k a q p = - hfill (λ k - → λ {(i = i0) → p j -- p (j ∧ k) - ; (i = i1) → compPath-filler' (sym q) p k j -- p j - ; (j = i0) → q (k ∨ ~ i) -- pt B - ; (j = i1) → pt B}) -- p k}) - (inS (p j)) -- (inS (q (~ i ∨ ~ j))) - k - - c2 : (i k : I) → (a : fst A) (q : fst f a ≡ pt B) (p : q (~ i) ≡ pt B) - → join (fib⋆ .fst) (Ω B .fst) - c2 i k a q p = - hfill (λ k → λ {(i = i0) → inr p - ; (i = i1) → push (a , p) (sym q ∙ p) (~ k)}) - (inS (inr λ j → c1 i j i1 a q p)) k - - GaneaFib→join : GaneaFib → join (fib⋆ .fst) (Ω B .fst) - GaneaFib→join (inl x , p) = inr p - GaneaFib→join (inr x , p) = inl (x , p) - GaneaFib→join (push (a , q) i , p) = c2 i i1 a q p - - cancel : (x : GaneaFib) → other (GaneaFib→join x) ≡ x - cancel (inl x , y) = refl - cancel (inr x , y) = refl - cancel (push (a , q) i , p) j = - hcomp (λ k → λ {(i = i0) → inl tt , p - ; (i = i1) → mg p k j - ; (j = i0) → other (c2 i k a q p) - ; (j = i1) → push (a , q) i , p}) - ((push (a , q) (i ∧ j)) - , λ k → hcomp (λ r → λ {(i = i0) → p k - ; (i = i1) → compPath-filler' (sym q) p (r ∧ (~ j)) k - ; (j = i0) → c1 i k r a q p - ; (j = i1) → p k - ; (k = i0) → q ((r ∧ ~ j) ∨ ~ i) - ; (k = i1) → snd B}) - (p k)) - where - fill1 : (i j k : I) (p : fst f a ≡ pt B) - → fst B - fill1 i j k p = - hfill (λ k → λ {(i = i0) → compPath-filler p (sym (sym q ∙ p)) k j - ; (i = i1) → q (j ∨ ~ k) - ; (j = i0) → q (~ k ∧ i) - ; (j = i1) → ((λ i₂ → q (~ i₂)) ∙ p) (~ k ∧ ~ i)}) - (inS (compPath-filler (sym q) p j (~ i))) k - - lazy : ∀ {ℓ} {A : Type ℓ} {x : A} (y : A) (q : x ≡ y) (z : A) (p : x ≡ z) - → PathP (λ k → Square (λ j → q (j ∨ ~ k)) refl - (compPath-filler' (λ i₂ → q (~ i₂)) p (~ k)) (sym q ∙ p)) - (λ i _ → (sym q ∙ p) i) λ i j → compPath-filler' (sym q) p j i - lazy {x = x} = J> (J> h x refl (refl ∙ refl) (compPath-filler' (sym refl) refl)) - where - h : ∀ {ℓ} {A : Type ℓ} {x : A} (y : A) (p : x ≡ y) (q : x ≡ y) (r : p ≡ q) - → PathP (λ k → Square refl refl (r (~ k)) q) (λ i _ → q i) λ i j → r j i - h = J> (J> refl) - - lazy2 : ∀ {ℓ} {A : Type ℓ} {x : A} (y : A) (q : x ≡ y) (z : A) (p : x ≡ z) - → PathP (λ r → Square (λ i → compPath-filler' (sym q) p i r) (λ i → (sym q ∙ p) (r ∨ ~ i)) (λ j → p (r ∨ j)) refl) - (λ j i → compPath-filler (sym q) p j (~ i)) - refl - lazy2 {x = x} = J> (J> λ i j k → {!lazy x refl _ refl (~ i) k (~ j)!}) - - help : (p : fst f a ≡ pt B) - → cong other (push (a , p) (sym q ∙ p)) - ≡ λ i → (push (a , q) (~ i)) , (compPath-filler' (sym q) p i) - proj₁ (help p i j) = push (a , λ j → fill1 i j i1 p) (~ j) - snd (help p i j) r = - hcomp (λ k → λ {(i = i0) → other-fill (a , p) (sym q ∙ p) j r k - ; (i = i1) → lazy _ q _ p k r j -- compPath-filler'-filler (sym q) p j r k - ; (j = i0) → compPath-filler' (sym q) p (~ k ∧ i) r - ; (j = i1) → (sym q ∙ p) (r ∨ ~ (k ∨ i)) - ; (r = i0) → fill1 i j k p - ; (r = i1) → snd B}) - (lazy2 _ q _ p r j i) - {- - r = i0 ⊢ -r = i1 ⊢ snd B -j = i0 ⊢ p r -j = i1 ⊢ hcomp (doubleComp-faces (λ _ → snd B) p r) (q (~ r)) -i = i0 ⊢ hcomp - (λ { k (j = i0) → p r - ; k (j = i1) - → hcomp (doubleComp-faces (λ _ → snd B) p (r ∨ ~ k)) - (q (~ (r ∨ ~ k))) - ; k (r = i0) - → compPath-filler p - (λ i₂ → - hcomp (doubleComp-faces (λ _ → snd B) p (~ i₂)) (q (~ (~ i₂)))) - k j - ; k (r = i1) → snd B - }) - (p (r ∨ j)) -i = i1 ⊢ hcomp - (λ { k (r = i0) → q (~ (~ j)) - ; k (r = i1) → p k - ; k (j = i0) → p (r ∧ k) - }) - (q (~ r ∧ j)) - -} - {- - hcomp (λ k → λ {(i = i0) → {!c2 j i1 a p (sym q ∙ p)!} - ; (i = i1) → {!!} - ; (j = i0) → {!!} - ; (j = i1) → {!!}}) - {!!} -} - -- other (push (a , p) (sym q ∙ p) j) - open import Cubical.Foundations.Path - - mg : (p : fst f a ≡ pt B) - → PathP (λ k → other (push (a , p) (sym q ∙ p) (~ k)) ≡ (inr a , p)) - (λ i → push (a , q) i , compPath-filler' (sym q) p (~ i)) - refl - mg p = flipSquare (cong sym (help p) - ◁ λ j i → push (a , q) (j ∨ i) - , compPath-filler' (sym q) p (~ (j ∨ i))) - - cool : PathP (λ j → other {!c2 i i1 a q !} ≡ {!!}) {!!} {!!} - cool = {!j = i0 ⊢ inr a , p -j = i1 ⊢ inl tt , sym q ∙ p -i = i0 ⊢ other (push (a , p) (sym q ∙ p) j) -i = i1 ⊢ push (a , q) (~ j) , compPath-filler' (sym q) p j!} - - - {- - Goal: join GaneaFib (Ω B .proj₁) -———— Boundary (wanted) ————————————————————————————————————— -i = i0 ⊢ inr b -i = i1 ⊢ inl (inr a , b) - -} - {- help a p i b - where - help : (a : fst A) (p : fst f a ≡ pt B) - → PathP (λ i → (b : (p (~ i) ≡ pt B)) → join GaneaFib (Ω B .fst)) - inr - λ b → inl (inr a , b) - help a p = {!!} -} - - -- (sym (push {!!} {!!}) ∙∙ {!!} ∙∙ {!!}) i -- push ((inr (fst a)) , {!a!}) {!!} (~ i) - - lem : join GaneaFib (Ω B .fst) → GaneaFib - lem (inl x) = x - lem (inr x) = inl tt , x - lem (push (inl x , p) b i) = (push ((pt A) , (snd f ∙ p)) ∙ sym (push ((pt A) , (snd f ∙ b)))) i , {!!} - lem (push (inr x , p) b i) = {!i = i0 ⊢ inr b -i = i1 ⊢ inl (inr a , b)!} - lem (push (push a i₁ , p) b i) = {!!} diff --git a/Cubical/HITs/SmashProduct/Ganea.agda b/Cubical/HITs/SmashProduct/Ganea.agda deleted file mode 100644 index 339b350f0d..0000000000 --- a/Cubical/HITs/SmashProduct/Ganea.agda +++ /dev/null @@ -1,14 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.HITs.SmashProduct.Ganea where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.HITs.Pushout.Base -open import Cubical.Data.Unit -open import Cubical.Data.Sigma -open import Cubical.HITs.Wedge -open import Cubical.Foundations.GroupoidLaws -open import Cubical.HITs.SmashProduct.Base - - diff --git a/Cubical/Homotopy/Group/Pi4S3/NewBrunerieNumbers.agda b/Cubical/Homotopy/Group/Pi4S3/NewBrunerieNumbers.agda deleted file mode 100644 index 1f221f2518..0000000000 --- a/Cubical/Homotopy/Group/Pi4S3/NewBrunerieNumbers.agda +++ /dev/null @@ -1,326 +0,0 @@ --- Mostly self contained definitions of the numbers from: https://arxiv.org/abs/2302.00151 -{-# OPTIONS --safe #-} -module Cubical.Homotopy.Group.Pi4S3.NewBrunerieNumbers where - -open import Cubical.Core.Primitives - -open import Cubical.Foundations.Prelude using - (_,_ ; fst ; snd ; _≡_ ; cong ; _∙_ ; refl ; isProp→isSet ; J ; sym ; isProp ; isSet ; isGroupoid ; is2Groupoid) -open import Cubical.Foundations.Equiv using - (_≃_ ; isEquiv ; isPropIsEquiv ; idIsEquiv ; idEquiv) -open import Cubical.Foundations.HLevels using - (hSet ; hGroupoid ; isOfHLevelTypeOfHLevel ; isOfHLevelPath ; isOfHLevelΠ ; isOfHLevel→isOfHLevelDep ; is2GroupoidΠ) -open import Cubical.Foundations.Univalence using - (Glue ; ua) -open import Cubical.Data.Int using - (ℤ ; pos ; neg ; isSetℤ ; sucPathℤ) -open import Cubical.Foundations.CartesianKanOps - -private variable ℓ ℓ' ℓ'' : Level - --- S1 -data S¹ : Type₀ where - base : S¹ - loop : base ≡ base - -helix : S¹ → Type₀ -helix base = ℤ -helix (loop i) = sucPathℤ i - -rotLoop : (a : S¹) → a ≡ a -rotLoop base = loop -rotLoop (loop i) j = - hcomp (λ k → λ { (i = i0) → loop (j ∨ ~ k) - ; (i = i1) → loop (j ∧ k) - ; (j = i0) → loop (i ∨ ~ k) - ; (j = i1) → loop (i ∧ k)}) base - -_·_ : S¹ → S¹ → S¹ -base · x = x -(loop i) · x = rotLoop x i - -isPropFamS¹ : (P : S¹ → Type ℓ) (pP : (x : S¹) → isProp (P x)) (b0 : P base) → - PathP (λ i → P (loop i)) b0 b0 -isPropFamS¹ P pP b0 i = pP (loop i) (coe0→i (λ k → P (loop k)) i b0) - (coe1→i (λ k → P (loop k)) i b0) i - -rotIsEquiv : (a : S¹) → isEquiv (a ·_) -rotIsEquiv base = idIsEquiv S¹ -rotIsEquiv (loop i) = isPropFamS¹ (λ x → isEquiv (x ·_)) - (λ x → isPropIsEquiv (x ·_)) - (idIsEquiv _) i - --- S2 - --- Direct S², useful for later part of the definition (or is it?) -data S²' : Type₀ where - base : S²' - surf : PathP (λ i → base ≡ base) refl refl - -S²ToSetElim' : {A : S²' → Type ℓ} - → ((x : S²') → isSet (A x)) - → A base - → (x : S²') → A x -S²ToSetElim' set b base = b -S²ToSetElim' set b (surf i j) = - isOfHLevel→isOfHLevelDep 2 set b b {a0 = refl} {a1 = refl} refl refl surf i j - --- Join -data join (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') where - inl : A → join A B - inr : B → join A B - push : ∀ a b → inl a ≡ inr b - - --- SetTruncation -data ∥_∥₀ (A : Type ℓ) : Type ℓ where - ∣_∣₀ : A → ∥ A ∥₀ - squash₀ : ∀ (x y : ∥ A ∥₀) (p q : x ≡ y) → p ≡ q - -rec₀ : {A : Type ℓ} {B : Type ℓ'} → isSet B → (A → B) → ∥ A ∥₀ → B -rec₀ Bset f ∣ x ∣₀ = f x -rec₀ Bset f (squash₀ x y p q i j) = - Bset _ _ (cong (rec₀ Bset f) p) (cong (rec₀ Bset f) q) i j - - --- GroupoidTruncation -data ∥_∥₁ (A : Type ℓ) : Type ℓ where - ∣_∣₁ : A → ∥ A ∥₁ - squash₁ : ∀ (x y : ∥ A ∥₁) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s - -rec₁ : {A : Type ℓ} {B : Type ℓ'} → isGroupoid B → (A → B) → ∥ A ∥₁ → B -rec₁ gB f ∣ x ∣₁ = f x -rec₁ gB f (squash₁ x y p q r s i j k) = - gB _ _ _ _ (λ m n → rec₁ gB f (r m n)) (λ m n → rec₁ gB f (s m n)) i j k - - --- 2GroupoidTruncation -data ∥_∥₂ (A : Type ℓ) : Type ℓ where - ∣_∣₂ : A → ∥ A ∥₂ - squash₂ : ∀ (x y : ∥ A ∥₂) (p q : x ≡ y) (r s : p ≡ q) (t u : r ≡ s) → t ≡ u - -rec₂ : ∀ {A : Type ℓ} {B : Type ℓ'} → is2Groupoid B → (A → B) → ∥ A ∥₂ → B -rec₂ gB f ∣ x ∣₂ = f x -rec₂ gB f (squash₂ _ _ _ _ _ _ t u i j k l) = - gB _ _ _ _ _ _ (λ m n o → rec₂ gB f (t m n o)) (λ m n o → rec₂ gB f (u m n o)) - i j k l - --- rec₂Bin : ∀ {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → is2Groupoid C → (A → B → C) → ∥ A ∥₂ → ∥ B ∥₂ → C --- rec₂Bin gB f = rec₂ (is2GroupoidΠ λ _ → gB) λ x → rec₂ gB (f x) - -elim₂ : {A : Type ℓ} {B : ∥ A ∥₂ → Type ℓ} - (bG : (x : ∥ A ∥₂) → is2Groupoid (B x)) - (f : (x : A) → B ∣ x ∣₂) (x : ∥ A ∥₂) → B x -elim₂ bG f ∣ x ∣₂ = f x -elim₂ bG f (squash₂ x y p q r s u v i j k l) = - isOfHLevel→isOfHLevelDep 4 bG _ _ _ _ _ _ - (λ j k l → elim₂ bG f (u j k l)) (λ j k l → elim₂ bG f (v j k l)) - (squash₂ x y p q r s u v) - i j k l - - --- Susp -data Susp (A : Type ℓ) : Type ℓ where - north : Susp A - south : Susp A - merid : (a : A) → north ≡ south - -S² S³ : Type₀ -S² = Susp S¹ -S³ = Susp S² - --- Pointed -Pointed₀ : Type₁ -Pointed₀ = Σ[ X ∈ Type₀ ] X - -S¹∙ S²'∙ : Pointed₀ -S¹∙ = (S¹ , base) -S²'∙ = (S²' , base) - -Susp∙ : Type₀ → Pointed₀ -Susp∙ A = Susp A , north - -∥_∥₁∙ ∥_∥₂∙ : Pointed₀ → Pointed₀ -∥ A , a ∥₁∙ = ∥ A ∥₁ , ∣ a ∣₁ -∥ A , a ∥₂∙ = ∥ A ∥₂ , ∣ a ∣₂ - -Ω Ω² : Pointed₀ → Pointed₀ -Ω (_ , a) = ((a ≡ a) , refl) -Ω² p = Ω (Ω p) - --- The maps -σ' : S²' → Ω (Susp∙ S²') .fst -σ' x = merid x ∙ sym (merid base) - -S¹×S¹→S²' : S¹ → S¹ → S²' -S¹×S¹→S²' base y = base -S¹×S¹→S²' (loop i) base = base -S¹×S¹→S²' (loop i) (loop j) = surf i j - -f7 : Ω (Susp∙ S²') .fst → ∥ S²' ∥₂ -f7 p = coe0→1 (λ i → Code (p i)) ∣ base ∣₂ - where - _+₂_ : ∥ S²' ∥₂ → ∥ S²' ∥₂ → ∥ S²' ∥₂ - _+₂_ = rec₂ (is2GroupoidΠ λ _ → squash₂) - λ { base x → x - ; (surf i j) x → surfc x i j} - where - surfc : (x : ∥ S²' ∥₂) → refl {x = x} ≡ refl {x = x} - surfc = elim₂ (λ x → isOfHLevelPath 4 (isOfHLevelPath 4 squash₂ _ _) refl refl) - (S²ToSetElim' (λ _ → squash₂ _ _ _ _) λ i j → ∣ surf i j ∣₂) - - ∥S²'∥₂≃∥S²'∥₂ : (x : S²') → ∥ S²' ∥₂ ≃ ∥ S²' ∥₂ - fst (∥S²'∥₂≃∥S²'∥₂ x) = ∣ x ∣₂ +₂_ - snd (∥S²'∥₂≃∥S²'∥₂ x) = help x - where - help : (x : _) → isEquiv (λ y → ∣ x ∣₂ +₂ y) - help = S²ToSetElim' (λ _ → isProp→isSet (isPropIsEquiv _)) (idIsEquiv _) - - Code : Susp S²' → Type₀ - Code north = ∥ S²' ∥₂ - Code south = ∥ S²' ∥₂ - Code (merid a i) = ua (∥S²'∥₂≃∥S²'∥₂ a) i - -g8 : Ω² ∥ S²'∙ ∥₂∙ .fst → Ω ∥ S¹∙ ∥₁∙ .fst -g8 p i = coe0→1 (λ j → codeTruncS²' (p i j) .fst) ∣ base ∣₁ - where - HopfS²' : S²' → Type₀ - HopfS²' base = S¹ - HopfS²' (surf i j) = Glue S¹ (λ { (i = i0) → S¹ , idEquiv S¹ - ; (i = i1) → S¹ , idEquiv S¹ - ; (j = i0) → S¹ , idEquiv S¹ - ; (j = i1) → S¹ , (loop i) ·_ , rotIsEquiv (loop i) } ) - - codeTruncS²' : ∥ S²' ∥₂ → hGroupoid _ - codeTruncS²' = rec₂ (isOfHLevelTypeOfHLevel 3) (λ s → ∥ HopfS²' s ∥₁ , squash₁) - -g9 : Ω ∥ S¹∙ ∥₁∙ .fst → ∥ ℤ ∥₀ -g9 p = coe0→1 (λ i → codeTruncS¹ (p i) .fst) ∣ pos 0 ∣₀ - where - codeTruncS¹ : ∥ S¹ ∥₁ → hSet _ - codeTruncS¹ = rec₁ (isOfHLevelTypeOfHLevel 2) (λ s → ∥ helix s ∥₀ , squash₀) - --- Use trick to eliminate away the truncation last -g10 : ∥ ℤ ∥₀ → ℤ -g10 = rec₀ isSetℤ (λ x → x) - --- Original η₃ as in the paper -η₃ : join S¹ S¹ → Susp S²' -η₃ (inl x) = north -η₃ (inr x) = north -η₃ (push a b i) = - (sym (σ' (S¹×S¹→S²' a b)) ∙ sym (σ' (S¹×S¹→S²' a b))) i - --- Dropping the syms makes it compute fast (why?! maybe some slow inverse map gets applied with the sym?) -η₃' : join S¹ S¹ → Susp S²' -η₃' (inl x) = north -η₃' (inr x) = north -η₃' (push a b i) = - (σ' (S¹×S¹→S²' a b) ∙ σ' (S¹×S¹→S²' a b)) i - --- Remark: dropping only one sym does not seem to make it fast enough - --- Brunerie numbers - -β₃ : ℤ -β₃ = g10 (g9 (g8 λ i j → f7 λ k → η₃ (push (loop i) (loop j) k))) - --- This does not terminate fast --- β₃≡-2 : β₃ ≡ -2 --- β₃≡-2 = refl - -β₃' : ℤ -β₃' = g10 (g9 (g8 λ i j → f7 λ k → η₃' (push (loop i) (loop j) k))) - --- This terminates fast -β₃'≡-2 : β₃' ≡ -2 -β₃'≡-2 = refl - -β₃'-pos : ℤ -β₃'-pos = g10 (g9 (g8 λ i j → f7 λ k → η₃' (push (loop (~ i)) (loop j) k))) - -β₃'≡2 : β₃'-pos ≡ pos 2 -β₃'≡2 = refl - - --- We now define β₂ - -invLooper : S¹ → S¹ -invLooper base = base -invLooper (loop i) = loop (~ i) - -η₂ : join S¹ S¹ → join S¹ S¹ -η₂ (inl x) = inr (invLooper x) -η₂ (inr x) = inr x -η₂ (push a b i) = - (sym (push (b · invLooper a) (invLooper a)) ∙ push (b · invLooper a) b) i - -rCancel : {A : Type ℓ} {x y : A} (p : x ≡ y) → p ∙ sym p ≡ refl -rCancel {x = x} p j i = - hcomp (λ k → λ { (i = i0) → x - ; (i = i1) → p (~ k ∧ ~ j) - ; (j = i1) → x - }) (p (i ∧ ~ j)) - -_∙∙_∙∙_ : {A : Type ℓ} {w x y z : A} → w ≡ x → x ≡ y → y ≡ z → w ≡ z -(p ∙∙ q ∙∙ r) i = - hcomp (λ j → λ { (i = i0) → p (~ j) ; (i = i1) → r j}) (q i) - -S¹×S¹→S² : S¹ → S¹ → S² -S¹×S¹→S² base y = north -S¹×S¹→S² (loop i) base = north -S¹×S¹→S² (loop i) (loop j) = - (sym (rCancel (merid base)) - ∙∙ (λ i → merid (loop i) ∙ sym (merid base)) - ∙∙ rCancel (merid base)) i j - -joinS¹S¹→S³ : join S¹ S¹ → S³ -joinS¹S¹→S³ (inl x) = north -joinS¹S¹→S³ (inr x) = north -joinS¹S¹→S³ (push a b i) = (merid (S¹×S¹→S² a b) ∙ sym (merid north)) i - -S²→S²' : S² → S²' -S²→S²' north = base -S²→S²' south = base -S²→S²' (merid base i) = base -S²→S²' (merid (loop i) j) = surf i j - -S³→SuspS²' : S³ → Susp S²' -S³→SuspS²' north = north -S³→SuspS²' south = north -S³→SuspS²' (merid x i) = (merid (S²→S²' x) ∙ sym (merid base)) i - -joinS¹S¹→SuspS²' : join S¹ S¹ → Susp S²' -joinS¹S¹→SuspS²' x = S³→SuspS²' (joinS¹S¹→S³ x) - -β₂ : ℤ -β₂ = g10 (g9 (g8 {!!})) -- λ i j → {!f7 -- joinS¹S¹→SuspS²' (η₂ (push (loop i0) (loop j) ?))!})) - where - _→∙_ : Type - _→∙_ = {!Σ[ f ∈ (Pointed₀ → Pointed₀) ] ?!} - - m : Ω (Susp∙ S²') .fst → ∥ S²'∙ ∥₂∙ .fst - m f7 = {!!} - - t : Ω (Ω² (Susp∙ S²')) .fst - t i j k = - hcomp (λ r → λ { (i = i0) → {!!} - ; (i = i1) → {!η₂ (push (loop i) (loop j) k)!} - ; (j = i0) → {!north!} - ; (j = i1) → {!north!} - ; (k = i0) → north - ; (k = i1) → north - }) - (joinS¹S¹→SuspS²' - (η₂ (push (loop i) (loop j) k))) - - h : Ω² ∥ S²'∙ ∥₂∙ .fst -- Ω² ∥ S²' ∥₂ .fst - h i j = f7 λ k → t i j k - - --- β₂ : ℤ --- β₂ = g10 (g9 (g8 λ i j → f7 (λ k → joinS¹S¹→SuspS²' (η₂ (push (loop i) (loop j) k))))) - - - --- -- TODO: define β₁ From 3acc6b862989c2d0e5d0f58137d1b9c2a144a48a Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 8 Jun 2023 09:45:28 +0200 Subject: [PATCH 07/73] w --- Cubical/Foundations/Prelude.agda | 7 ------- 1 file changed, 7 deletions(-) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 21b1675a2f..08ee95f74c 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -172,13 +172,6 @@ compPath-filler' {z = z} p q j i = ; (j = i0) → q (i ∧ k) }) (p (i ∨ ~ j)) -compPath-filler'-filler : (p : x ≡ y) (q : y ≡ z) → (i j k : I) → _ -compPath-filler'-filler {z = z} p q j i k = - hfill (λ k → λ { (i = i0) → p (~ j) - ; (i = i1) → q k - ; (j = i0) → q (i ∧ k) }) - (inS (p (i ∨ ~ j))) - k -- Note: We can omit a (j = i1) case here since when (j = i1), the whole expression is -- definitionally equal to `p ∙ q`. (Notice that `p ∙ q` is also an hcomp.) Nevertheless, -- we could have given `compPath-filler p q k i` as the (j = i1) case. From 227b10b9d16e0bdb866941a1eb9195e481c13999 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 8 Jun 2023 09:45:56 +0200 Subject: [PATCH 08/73] w --- Cubical/Foundations/Prelude.agda | 1 - 1 file changed, 1 deletion(-) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 08ee95f74c..a402935890 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -171,7 +171,6 @@ compPath-filler' {z = z} p q j i = ; (i = i1) → q k ; (j = i0) → q (i ∧ k) }) (p (i ∨ ~ j)) - -- Note: We can omit a (j = i1) case here since when (j = i1), the whole expression is -- definitionally equal to `p ∙ q`. (Notice that `p ∙ q` is also an hcomp.) Nevertheless, -- we could have given `compPath-filler p q k i` as the (j = i1) case. From 9d1915dd861102c72f1a497bf2d7e9b9e0295edf Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 17 Jan 2024 02:53:26 +0100 Subject: [PATCH 09/73] fix --- Cubical/Algebra/AbGroup/Base.agda | 4 - Cubical/Algebra/AbGroup/Properties.agda | 25 +++ Cubical/Algebra/Group/GroupPath.agda | 43 +++++- Cubical/Algebra/Group/MorphismProperties.agda | 4 + Cubical/Data/Nat/Order.agda | 24 +++ Cubical/Data/Unit/Properties.agda | 3 + Cubical/Foundations/Path.agda | 6 + Cubical/Foundations/Pointed/Properties.agda | 10 ++ Cubical/Foundations/Prelude.agda | 10 ++ .../PropositionalTruncation/Properties.agda | 11 ++ Cubical/HITs/Pushout/Properties.agda | 24 +++ Cubical/HITs/SequentialColimit/Base.agda | 7 + .../HITs/SequentialColimit/Properties.agda | 146 ++++++++++++++++++ Cubical/HITs/Sn/Base.agda | 4 + Cubical/HITs/Sn/Properties.agda | 109 +++++++++---- Cubical/HITs/Susp/Properties.agda | 18 +++ Cubical/HITs/Truncation/Properties.agda | 8 + Cubical/Homotopy/Connected.agda | 32 ++++ 18 files changed, 456 insertions(+), 32 deletions(-) diff --git a/Cubical/Algebra/AbGroup/Base.agda b/Cubical/Algebra/AbGroup/Base.agda index eb86373396..74605f2388 100644 --- a/Cubical/Algebra/AbGroup/Base.agda +++ b/Cubical/Algebra/AbGroup/Base.agda @@ -315,10 +315,6 @@ module _ {ℓ ℓ' : Level} (AGr : Group ℓ) (BGr : AbGroup ℓ') where renaming (_·_ to _∙A_ ; inv to -A_ ; 1g to 1A ; ·IdR to ·IdRA) - trivGroupHom : GroupHom AGr (BGr *) - fst trivGroupHom x = 0B - snd trivGroupHom = makeIsGroupHom λ _ _ → sym (+IdRB 0B) - compHom : GroupHom AGr (BGr *) → GroupHom AGr (BGr *) → GroupHom AGr (BGr *) fst (compHom f g) x = fst f x +B fst g x snd (compHom f g) = diff --git a/Cubical/Algebra/AbGroup/Properties.agda b/Cubical/Algebra/AbGroup/Properties.agda index 2f16f05be3..2baeeef678 100644 --- a/Cubical/Algebra/AbGroup/Properties.agda +++ b/Cubical/Algebra/AbGroup/Properties.agda @@ -5,6 +5,8 @@ open import Cubical.Foundations.Prelude open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties private variable ℓ : Level @@ -24,3 +26,26 @@ module AbGroupTheory (A : AbGroup ℓ) where implicitInverse : ∀ {a b} → a + b ≡ 0g → b ≡ - a implicitInverse b+a≡0 = invUniqueR b+a≡0 + +addGroupHom : (A B : AbGroup ℓ) (ϕ ψ : AbGroupHom A B) → AbGroupHom A B +fst (addGroupHom A B ϕ ψ) x = AbGroupStr._+_ (snd B) (ϕ .fst x) (ψ .fst x) +snd (addGroupHom A B ϕ ψ) = makeIsGroupHom + λ x y → cong₂ (AbGroupStr._+_ (snd B)) + (IsGroupHom.pres· (snd ϕ) x y) + (IsGroupHom.pres· (snd ψ) x y) + ∙ AbGroupTheory.comm-4 B (fst ϕ x) (fst ϕ y) (fst ψ x) (fst ψ y) + +negGroupHom : (A B : AbGroup ℓ) (ϕ : AbGroupHom A B) → AbGroupHom A B +fst (negGroupHom A B ϕ) x = AbGroupStr.-_ (snd B) (ϕ .fst x) +snd (negGroupHom A B ϕ) = + makeIsGroupHom λ x y + → sym (IsGroupHom.presinv (snd ϕ) (AbGroupStr._+_ (snd A) x y)) + ∙ cong (fst ϕ) (GroupTheory.invDistr (AbGroup→Group A) x y + ∙ AbGroupStr.+Comm (snd A) _ _) + ∙ IsGroupHom.pres· (snd ϕ) _ _ + ∙ cong₂ (AbGroupStr._+_ (snd B)) + (IsGroupHom.presinv (snd ϕ) x) + (IsGroupHom.presinv (snd ϕ) y) + +subtrGroupHom : (A B : AbGroup ℓ) (ϕ ψ : AbGroupHom A B) → AbGroupHom A B +subtrGroupHom A B ϕ ψ = addGroupHom A B ϕ (negGroupHom A B ψ) diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index 84c33a481e..4159140591 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -1,5 +1,5 @@ -- The SIP applied to groups -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.Group.GroupPath where open import Cubical.Foundations.Prelude @@ -24,6 +24,8 @@ open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.HITs.PropositionalTruncation + private variable ℓ ℓ' ℓ'' : Level @@ -195,3 +197,42 @@ GroupEquivJ {G = G} P p {H} e = (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof (transportRefl (fst (fst e) (transportRefl x i)) i)))) ∙ retEq (fst e) x)) + + +isGroupoidGroup : ∀ {ℓ} → isGroupoid (Group ℓ) +isGroupoidGroup G H = + isOfHLevelRespectEquiv 2 (GroupPath _ _) + (isOfHLevelΣ 2 (isOfHLevel≃ 2 (GroupStr.is-set (snd G)) (GroupStr.is-set (snd H))) + λ _ → isProp→isSet (isPropIsGroupHom _ _)) + +module _ {ℓ ℓ'} {A : Type ℓ} + (G : A → Group ℓ') + (G-coh : (x y : A) → GroupEquiv (G x) (G y)) + (G-coh-coh : (x y z : A) (g : fst (G x)) + → fst (fst (G-coh y z)) ((fst (fst (G-coh x y)) g)) + ≡ fst (fst (G-coh x z)) g ) where + + PropTrunc→Group-coh : (x y : A) → G x ≡ G y + PropTrunc→Group-coh x y = uaGroup (G-coh x y) + + PropTrunc→Group-coh-coh : (x y z : A) → compGroupEquiv (G-coh x y) (G-coh y z) ≡ G-coh x z + PropTrunc→Group-coh-coh x y z = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (Σ≡Prop (λ _ → isPropIsEquiv _) + (funExt (G-coh-coh x y z))) + + open import Cubical.Foundations.Path + PropTrunc→Group : ∥ A ∥₁ → Group ℓ' + PropTrunc→Group = rec→Gpd isGroupoidGroup + G + (record { link = PropTrunc→Group-coh + ; coh₁ = coh-coh }) + where + coh-coh : (x y z : A) + → Square (PropTrunc→Group-coh x y) (PropTrunc→Group-coh x z) + refl (PropTrunc→Group-coh y z) + coh-coh x y z = + compPathL→PathP + (sym (lUnit _) + ∙∙ sym (uaCompGroupEquiv (G-coh x y) (G-coh y z)) + ∙∙ cong uaGroup (PropTrunc→Group-coh-coh x y z)) diff --git a/Cubical/Algebra/Group/MorphismProperties.agda b/Cubical/Algebra/Group/MorphismProperties.agda index 071230f91e..bc5f285f7d 100644 --- a/Cubical/Algebra/Group/MorphismProperties.agda +++ b/Cubical/Algebra/Group/MorphismProperties.agda @@ -177,6 +177,10 @@ compGroupHom : GroupHom F G → GroupHom G H → GroupHom F H fst (compGroupHom f g) = fst g ∘ fst f snd (compGroupHom f g) = isGroupHomComp f g +trivGroupHom : GroupHom G H +fst (trivGroupHom {H = H}) _ = 1g (snd H) +snd (trivGroupHom {H = H}) = makeIsGroupHom λ _ _ → sym (·IdR (snd H) _) + GroupHomDirProd : {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → GroupHom A C → GroupHom B D → GroupHom (DirProd A B) (DirProd C D) fst (GroupHomDirProd mf1 mf2) = map-× (fst mf1) (fst mf2) diff --git a/Cubical/Data/Nat/Order.agda b/Cubical/Data/Nat/Order.agda index 48e68c9d1c..2698bd7eb8 100644 --- a/Cubical/Data/Nat/Order.agda +++ b/Cubical/Data/Nat/Order.agda @@ -191,6 +191,24 @@ predℕ-≤-predℕ {suc m} {suc n} ineq = pred-≤-pred ineq <-k+-trans : (k + m < n) → m < n <-k+-trans {m} {k} {n} p = ≤<-trans (m , refl) p +¬-suc-n m) +Dichotomyℕ n m with (n ≟ m) +... | lt x = inl (<-weaken x) +... | Trichotomy.eq x = inl (0 , x) +... | gt x = inr x + splitℕ-≤ : (m n : ℕ) → (m ≤ n) ⊎ (n < m) splitℕ-≤ m n with m ≟ n ... | lt x = inl (<-weaken x) diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index b2938ee014..2f8e5cf5dc 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -25,6 +25,9 @@ private variable ℓ ℓ' : Level +terminal : (A : Type ℓ) → A → Unit +terminal A x = tt + isContrUnit : isContr Unit isContrUnit = tt , λ {tt → refl} diff --git a/Cubical/Foundations/Path.agda b/Cubical/Foundations/Path.agda index f8a3270618..a738bdf499 100644 --- a/Cubical/Foundations/Path.agda +++ b/Cubical/Foundations/Path.agda @@ -437,3 +437,9 @@ Square→compPathΩ² {a = a} sq k i j = ; (j = i1) → a ; (k = i1) → cong (λ x → rUnit x r) (flipSquare sq) i j}) (sq j i) + +pathFiber : {B : Type ℓ} (f : A → B) + (b : B) {a a' : A} {t : f a ≡ b} {t' : f a' ≡ b} → + ((a , t) ≡ (a' , t' )) → Σ[ e ∈ a ≡ a' ] (t ≡ cong f e ∙ t') +pathFiber {A} {B} f b {a} {a'} {t} {t'} e = + J (λ X _ → Σ[ e ∈ a ≡ fst X ] (t ≡ cong f e ∙ (snd X))) (refl , lUnit t) e diff --git a/Cubical/Foundations/Pointed/Properties.agda b/Cubical/Foundations/Pointed/Properties.agda index 014d605213..75b88c221d 100644 --- a/Cubical/Foundations/Pointed/Properties.agda +++ b/Cubical/Foundations/Pointed/Properties.agda @@ -244,3 +244,13 @@ compPathlEquiv∙ : {A : Type ℓ} {a b c : A} {q : b ≡ c} (p : a ≡ b) → ((b ≡ c) , q) ≃∙ ((a ≡ c) , p ∙ q) fst (compPathlEquiv∙ p) = compPathlEquiv p snd (compPathlEquiv∙ p) = refl + +-- a pointed map is constant iff its non-pointed part is constant +constantPointed≡ : {A B : Pointed ℓ} (f : A →∙ B) + → fst f ≡ fst (const∙ A B) → f ≡ const∙ A B +constantPointed≡ {A = A} {B = B , b} f Hf i = + (λ x → ((λ j → Hf j x) ∙∙ (λ j → Hf (~ j) (pt A)) ∙∙ (snd f)) i) + , λ j → hcomp (λ k → λ { (i = i0) → invSides-filler (λ i → Hf i (pt A)) (λ i → snd f i) k (~ j) + ; (i = i1) → snd f k + ; (j = i1) → snd f k }) + (Hf ((~ i) ∧ (~ j)) (pt A)) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index d5796056f3..6bdd34e7bf 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -181,6 +181,16 @@ compPath-filler' {z = z} p q j i = ; (i = i1) → q k ; (j = i0) → q (i ∧ k) }) (p (i ∨ ~ j)) + +compPath-filler'' : (p : x ≡ y) (q : y ≡ z) + → PathP (λ i → p (~ i) ≡ q i) refl (p ∙ q) +compPath-filler'' p q i j = + hcomp (λ k → λ {(i = i0) → p i1 + ; (i = i1) → compPath-filler p q k j + ; (j = i0) → p (~ i) + ; (j = i1) → q (i ∧ k)}) + (p (~ i ∨ j)) + -- Note: We can omit a (j = i1) case here since when (j = i1), the whole expression is -- definitionally equal to `p ∙ q`. (Notice that `p ∙ q` is also an hcomp.) Nevertheless, -- we could have given `compPath-filler p q k i` as the (j = i1) case. diff --git a/Cubical/HITs/PropositionalTruncation/Properties.agda b/Cubical/HITs/PropositionalTruncation/Properties.agda index 69b13578bb..92f6ea0b08 100644 --- a/Cubical/HITs/PropositionalTruncation/Properties.agda +++ b/Cubical/HITs/PropositionalTruncation/Properties.agda @@ -297,6 +297,17 @@ elim2→Set {A = A} {B = B} {P = P} Pset f kf₁ kf₂ sf = → PathP (λ i → (u : ∥ B ∥₁) → P (squash₁ ∣ x ∣₁ ∣ y ∣₁ i) u) (mapHelper x) (mapHelper y) squareHelper x y i = elim→Set (λ _ → Pset _ _) (λ v → kf₁ x y v i) λ v w → sf x y v w i +elim→Setβ : ∀ {ℓ ℓ'} {A : Type ℓ} + {P : ∥ A ∥₁ → Type ℓ'} + → (is-set : ∀ t → isSet (P t)) + → (f : (x : A) → P ∣ x ∣₁) + → (kf : ∀ x y → PathP (λ i → P (squash₁ ∣ x ∣₁ ∣ y ∣₁ i)) (f x) (f y)) + → (t : A) → elim→Set is-set f kf ∣ t ∣₁ ≡ f t +elim→Setβ {A = A} {P = P} Pset f kf t = + cong (λ p → subst P p (f t)) + (isProp→isSet squash₁ _ _ (squash₁ ∣ t ∣₁ ∣ t ∣₁) refl) + ∙ transportRefl (f t) + RecHProp : (P : A → hProp ℓ) (kP : ∀ x y → P x ≡ P y) → ∥ A ∥₁ → hProp ℓ RecHProp P kP = rec→Set isSetHProp P kP diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index c01e0b4b44..2e832226b5 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -29,8 +29,11 @@ open import Cubical.Foundations.Univalence open import Cubical.Foundations.Transport open import Cubical.Foundations.Function +open import Cubical.Relation.Nullary + open import Cubical.Data.Sigma open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ open import Cubical.HITs.Pushout.Base @@ -527,3 +530,24 @@ module PushoutDistr {ℓ ℓ' ℓ'' ℓ''' : Level} leftInv PushoutDistrIso (inr (push a i)) j = compPath-filler' (push (g a)) (λ j → inr (push a j)) (~ j) i leftInv PushoutDistrIso (push a i) j = push a (i ∧ j) + +PushoutEmptyFam : ∀ {ℓ} {A B C : Type ℓ} + → ¬ A → ¬ C + → {f : A → B} {g : A → C} + → Iso B (Pushout {A = A} {B = B} {C = C} f g) +PushoutEmptyFam {A = A} {B = B} {C = C} ¬A ¬C = + compIso is + (pushoutIso _ _ _ _ + (uninhabEquiv (λ {()}) ¬A) + (idEquiv B) + (uninhabEquiv (λ {()}) ¬C) + (funExt (λ{()})) (funExt λ{()})) + where + lift⊥ : ∀ ℓ → Type ℓ + lift⊥ ℓ = Lift ⊥ + + is : Iso B (Pushout {A = lift⊥ ℓ} {B = B} {C = lift⊥ ℓ''} (λ{()}) λ{()}) + Iso.fun is = inl + Iso.inv is (inl x) = x + Iso.rightInv is (inl x) = refl + Iso.leftInv is x = refl diff --git a/Cubical/HITs/SequentialColimit/Base.agda b/Cubical/HITs/SequentialColimit/Base.agda index af80b4b5f7..27ca5d7fbc 100644 --- a/Cubical/HITs/SequentialColimit/Base.agda +++ b/Cubical/HITs/SequentialColimit/Base.agda @@ -3,6 +3,7 @@ module Cubical.HITs.SequentialColimit.Base where open import Cubical.Foundations.Prelude open import Cubical.Data.Nat +open import Cubical.Foundations.Equiv private variable @@ -18,3 +19,9 @@ open Sequence data SeqColim (X : Sequence ℓ) : Type ℓ where incl : {n : ℕ} → X .obj n → SeqColim X push : {n : ℕ} (x : X .obj n) → incl x ≡ incl (X .map x) + +converges : ∀ {ℓ} → Sequence ℓ → (n : ℕ) → Type ℓ +converges seq n = (k : ℕ) → isEquiv (Sequence.map seq {n = k + n}) + +finiteSequence : (ℓ : Level) (m : ℕ) → Type (ℓ-suc ℓ) +finiteSequence ℓ m = Σ[ S ∈ Sequence ℓ ] converges S m diff --git a/Cubical/HITs/SequentialColimit/Properties.agda b/Cubical/HITs/SequentialColimit/Properties.agda index c0321c9737..02cfad2f80 100644 --- a/Cubical/HITs/SequentialColimit/Properties.agda +++ b/Cubical/HITs/SequentialColimit/Properties.agda @@ -13,6 +13,8 @@ open import Cubical.Foundations.Equiv.Properties open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence open import Cubical.Data.Nat hiding (elim) open import Cubical.HITs.SequentialColimit.Base @@ -235,3 +237,147 @@ module _ isConnectedIncl∞ : isConnectedFun d (incl∞ n) isConnectedIncl∞ = elim.isConnectedPrecompose _ _ hasSectionIncl∘ + +open Sequence +open ElimDataShifted +-- Proof that colim Aₘ ≃ Aₙ for a converging sequence +-- A₀ → A₁ → ... → Aₙ ≃ Aₙ₊₁ ≃ ... +module _ {ℓ : Level} (seq : Sequence ℓ) (n : ℕ) (term : converges seq n) + where + shiftEq : ∀ {k} → seq .obj n ≃ (seq .obj (k + n)) + shiftEq {k = zero} = idEquiv _ + shiftEq {k = suc k} = compEquiv (shiftEq {k}) (_ , term k) + + shiftEqCoh : (k : ℕ) (x : _) + → invEq shiftEq x + ≡ invEq (compEquiv shiftEq (Sequence.map seq , term k)) (seq .Sequence.map x) + shiftEqCoh zero x = sym (retEq (_ , term 0) x) + shiftEqCoh (suc k) x = + cong (invEq shiftEq) (sym (retEq (_ , term (suc k)) x)) + + shiftEqShifted : ElimDataShifted seq n (λ _ → obj seq n) + inclP shiftEqShifted = invEq shiftEq + pushP shiftEqShifted = shiftEqCoh _ + + -- induction principle for terminating sequences + module _ {P : (k : ℕ) → seq .obj (k + n) → Type ℓ} + (bs : (s : _) → P zero s) + (indr : (k : ℕ) → ((y : _) → P k y) → (x : _) → P (suc k) (Sequence.map seq x)) + where + terminaates-seq-ind : (k : ℕ) (x : _) → P k x + terminaates-seq-ind zero = bs + terminaates-seq-ind (suc k) x = + subst (P (suc k)) (secEq (_ , term k) x) + (indr k (terminaates-seq-ind k) (invEq (_ , term k) x)) + + terminaates-seq-indβ : (k : ℕ) (s : ((y : seq .obj (k + n)) → P k y)) (x : _) + → terminaates-seq-ind (suc k) (Sequence.map seq x) + ≡ indr k (terminaates-seq-ind k) x + terminaates-seq-indβ k s x = + lem (Sequence.map seq , term k) x (P (suc k)) + (indr k (terminaates-seq-ind k)) + where + lem : ∀ {ℓ ℓ'} {A B : Type ℓ} (e : A ≃ B) (x : A) + (P : B → Type ℓ') (πP : (x : A) → P (fst e x)) + → subst P (secEq e (fst e x)) (πP (invEq e (fst e x))) ≡ πP x + lem {ℓ' = ℓ'} {B = B} = + EquivJ (λ A e → (x : A) (P : B → Type ℓ') (πP : (x : A) → P (fst e x)) + → subst P (secEq e (fst e x)) (πP (invEq e (fst e x))) + ≡ πP x) + λ b P πP → transportRefl _ + + -- a special case + module _ + (0case : (x : seq .obj n) + → Path (SeqColim seq) + (incl (elimShifted seq n (λ _ → obj seq n) shiftEqShifted (incl x))) + (incl x)) where + + converges→ColimIso-main-lem : (k : ℕ) (x : seq .obj (k + n)) + → Path (SeqColim seq) + (incl (elimShifted seq n (λ _ → obj seq n) + shiftEqShifted (incl x))) + (incl x) + converges→ColimIso-main-lem = + terminaates-seq-ind + 0case + (λ k id x + → cong incl + (sym (cong (elimShifted seq n (λ _ → obj seq n) shiftEqShifted) (push x))) + ∙∙ id x + ∙∙ push x) + + converges→ColimIso-main-lemβ : (k : ℕ) (x : seq .obj (k + n)) + → converges→ColimIso-main-lem (suc k) (Sequence.map seq x) + ≡ (cong incl + (sym (cong (elimShifted seq n (λ _ → obj seq n) + shiftEqShifted) (push x))) + ∙∙ converges→ColimIso-main-lem k x + ∙∙ push x) + converges→ColimIso-main-lemβ k x = + terminaates-seq-indβ + 0case + (λ k id x + → cong incl + (sym (cong (elimShifted seq n (λ _ → obj seq n) shiftEqShifted) (push x))) + ∙∙ id x + ∙∙ push x) + k (converges→ColimIso-main-lem k) x + +-- main result +-- (todo: add more thy about colimits to get nicer proof) +converges→ColimIso : ∀ {ℓ} {seq : Sequence ℓ} (n : ℕ) + → converges seq n + → Iso (obj seq n) (SeqColim seq) +Iso.fun (converges→ColimIso {seq = seq} n e) = incl +Iso.inv (converges→ColimIso {seq = seq} n e) = elimShifted seq n _ (shiftEqShifted seq n e) +Iso.rightInv (converges→ColimIso {seq = seq} n e) = elimShifted seq n _ + (record { inclP = λ {k} → paths k + ; pushP = λ {k} → cohs k}) + where + zero-case : (x : seq .obj n) + → Path (SeqColim seq) + (incl (elimShifted seq n (λ _ → obj seq n) (shiftEqShifted seq n e) (incl x))) + (incl x) + zero-case x i = incl (elimShiftedβ seq n 0 (λ _ → obj seq n) (shiftEqShifted seq n e) i x) + + paths : (k : ℕ) → (x : seq .obj (k + n)) + → Path (SeqColim seq) + (incl (elimShifted seq n (λ _ → obj seq n) (shiftEqShifted seq n e) (incl x))) + (incl x) + paths = converges→ColimIso-main-lem seq n e zero-case + + cohs : (k : ℕ) (x : seq .obj (k + n)) + → PathP (λ i → Path (SeqColim seq) + (incl (elimShifted seq n (λ _ → obj seq n) + (shiftEqShifted seq n e) + (push x i))) + (push x i)) + (paths k x) (paths (suc k) (seq .Sequence.map x)) + cohs k x = + doubleCompPath-filler + (λ i → incl (elimShifted seq n (λ _ → obj seq n) + (shiftEqShifted seq n e) (push x (~ i)))) + (converges→ColimIso-main-lem seq n e zero-case k x) + (push x) + ▷ sym (converges→ColimIso-main-lemβ seq n e (zero-case) k x) +Iso.leftInv (converges→ColimIso {seq = seq} n e) a = + funExt⁻ (elimShiftedβ seq n _ (λ _ → obj seq n) + (shiftEqShifted seq n e)) a + +-- different form +converges→isEquivIncl : ∀ {ℓ} {seq : Sequence ℓ} (n : ℕ) + → converges seq n + → isEquiv {A = obj seq n} {B = SeqColim seq} (incl {n = n}) +converges→isEquivIncl n x = isoToEquiv (converges→ColimIso n x) .snd + +-- elimination from colimit into prop +SeqColim→Prop : ∀ {ℓ ℓ'} {C : Sequence ℓ} {B : SeqColim C → Type ℓ'} + → ((x : _) → isProp (B x)) + → (s : (n : ℕ) (x : obj C n) → B (incl x)) + → (x : _) → B x +SeqColim→Prop {C = C} pr ind (incl x) = ind _ x +SeqColim→Prop {C = C} {B = B} pr ind (push x i) = + isProp→PathP {B = λ i → B (push x i)} + (λ i → pr _) + (ind _ x) (ind (suc _) (C .Sequence.map x)) i diff --git a/Cubical/HITs/Sn/Base.agda b/Cubical/HITs/Sn/Base.agda index 9234d84ae1..f0474e3698 100644 --- a/Cubical/HITs/Sn/Base.agda +++ b/Cubical/HITs/Sn/Base.agda @@ -19,6 +19,10 @@ S₊ 0 = Bool S₊ 1 = S¹ S₊ (suc (suc n)) = Susp (S₊ (suc n)) +S⁻ : ℕ → Type +S⁻ zero = ⊥ +S⁻ (suc n) = S₊ n + ptSn : (n : ℕ) → S₊ n ptSn zero = true ptSn (suc zero) = base diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index bf127bdd3d..50b9689589 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -17,8 +17,9 @@ open import Cubical.HITs.S3 open import Cubical.Data.Nat hiding (elim) open import Cubical.Data.Sigma open import Cubical.HITs.Sn.Base -open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Susp open import Cubical.HITs.Truncation +open import Cubical.HITs.PropositionalTruncation as PT hiding (rec ; elim) open import Cubical.Homotopy.Connected open import Cubical.HITs.Join renaming (joinS¹S¹→S³ to joinS¹S¹→S3) open import Cubical.Data.Bool hiding (elim) @@ -29,6 +30,11 @@ private open Iso +σSn : (n : ℕ) → S₊ n → Path (S₊ (suc n)) (ptSn (suc n)) (ptSn (suc n)) +σSn zero false = loop +σSn zero true = refl +σSn (suc n) x = toSusp (S₊∙ (suc n)) x + IsoSucSphereSusp : (n : ℕ) → Iso (S₊ (suc n)) (Susp (S₊ n)) IsoSucSphereSusp zero = S¹IsoSuspBool IsoSucSphereSusp (suc n) = idIso @@ -67,6 +73,12 @@ private ; (j = i1) → y }) (p (j ∨ i)) +sphereElim' : (n : ℕ) {A : S₊ n → Type ℓ} → + ((x : S₊ n) → isOfHLevel n (A x)) → + A (ptSn n) → (x : S₊ n) → A x +sphereElim' zero st _ x = st x .fst +sphereElim' (suc n) = sphereElim n + sphereToPropElim : (n : ℕ) {A : (S₊ (suc n)) → Type ℓ} → ((x : S₊ (suc n)) → isProp (A x)) → A (ptSn (suc n)) → (x : S₊ (suc n)) → A x @@ -288,6 +300,49 @@ sphereConnected : (n : HLevel) → isConnected (suc n) (S₊ n) sphereConnected n = ∣ ptSn n ∣ , elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) (λ a → sym (spoke ∣_∣ (ptSn n)) ∙ spoke ∣_∣ a) +sphereToTrunc : (n : ℕ) {A : S₊ n → Type} + → ((x : S₊ n) → hLevelTrunc (suc n) (A x)) + → ∥ ((x : _) → A x) ∥₁ +sphereToTrunc zero {A = A} indr = + rec squash₁ (λ p → rec squash₁ + (λ q → ∣ (λ { false → q ; true → p}) ∣₁) + (indr false)) (indr true) +sphereToTrunc (suc zero) {A = A} indr = + lem (indr base) (cong indr loop) + where + lem : (x : hLevelTrunc 2 (A base)) + → PathP (λ i → hLevelTrunc 2 (A (loop i))) x x + → ∥ ((x : S¹) → A x) ∥₁ + lem = elim (λ _ → isSetΠ λ _ → isProp→isSet squash₁) λ a p + → rec squash₁ (λ q → ∣ (λ { base → a + ; (loop i) → toPathP {A = λ i → A (loop i)} q i}) ∣₁) + (PathIdTruncIso 1 .Iso.fun + (fromPathP p)) +sphereToTrunc (suc (suc n)) {A = A} indr = + lem (sphereToTrunc (suc n)) (indr north) (indr south) + λ a → cong indr (merid a) + where + lem : ({A : S₊ (suc n) → Type} → + ((i : S₊ (suc n)) → hLevelTrunc (suc (suc n)) (A i)) → + ∥ ((x : S₊ (suc n)) → A x) ∥₁) + → (x : hLevelTrunc (3 + n) (A north)) + (y : hLevelTrunc (3 + n) (A south)) + → ((a : _) → PathP (λ i → hLevelTrunc (3 + n) (A (merid a i))) x y) + → ∥ ((x : S₊ (2 + n)) → A x) ∥₁ + lem indr = + elim (λ _ → isOfHLevelΠ2 (3 + n) + λ _ _ → isProp→isOfHLevelSuc (2 + n) squash₁) + λ a → elim (λ _ → isOfHLevelΠ (3 + n) + λ _ → isProp→isOfHLevelSuc (2 + n) squash₁) + λ b → λ f → + PT.map (λ ma → λ { north → a + ; south → b + ; (merid a i) → ma a i}) + (indr {A = λ x → PathP (λ i → A (merid x i)) a b} + λ x → rec (isOfHLevelTrunc (2 + n)) + (λ p → ∣ toPathP p ∣) + (Iso.fun (PathIdTruncIso _) (fromPathP (f x)))) + -- The fact that path spaces of Sn are connected can be proved directly for Sⁿ. -- (Unfortunately, this does not work for higher paths) pathIdTruncSⁿ : (n : ℕ) (x y : S₊ (suc n)) @@ -348,18 +403,18 @@ invLooperDistr = SuspS¹-hom : (a x : S¹) → Path (Path (hLevelTrunc 4 (S₊ 2)) _ _) - (cong ∣_∣ₕ (σ (S₊∙ 1) (a * x))) - (cong ∣_∣ₕ (σ (S₊∙ 1) a) - ∙ (cong ∣_∣ₕ (σ (S₊∙ 1) x))) + (cong ∣_∣ₕ (σSn 1 (a * x))) + (cong ∣_∣ₕ (σSn 1 a) + ∙ (cong ∣_∣ₕ (σSn 1 x))) SuspS¹-hom = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) (λ x → lUnit _ - ∙ cong (_∙ cong ∣_∣ₕ (σ (S₊∙ 1) x)) + ∙ cong (_∙ cong ∣_∣ₕ (σSn 1 x)) (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) - (λ x → (λ i → cong ∣_∣ₕ (σ (S₊∙ 1) (rUnitS¹ x i))) + (λ x → (λ i → cong ∣_∣ₕ (σSn 1 (rUnitS¹ x i))) ∙∙ rUnit _ - ∙∙ cong (cong ∣_∣ₕ (σ (S₊∙ 1) x) ∙_) + ∙∙ cong (cong ∣_∣ₕ (σSn 1 x) ∙_) (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) - (sym (l (cong ∣_∣ₕ (σ (S₊∙ 1) base)) + (sym (l (cong ∣_∣ₕ (σSn 1 base)) (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))))) where l : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (P : refl ≡ p) @@ -373,24 +428,24 @@ rCancelS¹ (loop i) j = base SuspS¹-inv : (x : S¹) → Path (Path (hLevelTrunc 4 (S₊ 2)) _ _) - (cong ∣_∣ₕ (σ (S₊∙ 1) (invLooper x))) - (cong ∣_∣ₕ (sym (σ (S₊∙ 1) x))) + (cong ∣_∣ₕ (σSn 1 (invLooper x))) + (cong ∣_∣ₕ (sym (σSn 1 x))) SuspS¹-inv x = (lUnit _ - ∙∙ cong (_∙ cong ∣_∣ₕ (σ (S₊∙ 1) (invLooper x))) - (sym (lCancel (cong ∣_∣ₕ (σ (S₊∙ 1) x)))) + ∙∙ cong (_∙ cong ∣_∣ₕ (σSn 1 (invLooper x))) + (sym (lCancel (cong ∣_∣ₕ (σSn 1 x)))) ∙∙ sym (assoc _ _ _)) - ∙∙ cong (sym (cong ∣_∣ₕ (σ (S₊∙ 1) x)) ∙_) lem + ∙∙ cong (sym (cong ∣_∣ₕ (σSn 1 x)) ∙_) lem ∙∙ (assoc _ _ _ - ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (σ (S₊∙ 1) x)))) - (lCancel (cong ∣_∣ₕ (σ (S₊∙ 1) x))) + ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (σSn 1 x)))) + (lCancel (cong ∣_∣ₕ (σSn 1 x))) ∙∙ sym (lUnit _)) where - lem : cong ∣_∣ₕ (σ (S₊∙ 1) x) - ∙ cong ∣_∣ₕ (σ (S₊∙ 1) (invLooper x)) - ≡ cong ∣_∣ₕ (σ (S₊∙ 1) x) - ∙ cong ∣_∣ₕ (sym (σ (S₊∙ 1) x)) + lem : cong ∣_∣ₕ (σSn 1 x) + ∙ cong ∣_∣ₕ (σSn 1 (invLooper x)) + ≡ cong ∣_∣ₕ (σSn 1 x) + ∙ cong ∣_∣ₕ (sym (σSn 1 x)) lem = sym (SuspS¹-hom x (invLooper x)) - ∙ ((λ i → cong ∣_∣ₕ (σ (S₊∙ 1) (rCancelS¹ x (~ i)))) + ∙ ((λ i → cong ∣_∣ₕ (σSn 1 (rCancelS¹ x (~ i)))) ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) -------------------- join Sⁿ Sᵐ ≃ Sⁿ⁺¹⁺ᵐ ------------------------- @@ -460,7 +515,7 @@ joinS¹S¹→S³'≡joinS¹S¹→S³' (push (loop i) (loop j) k) l = ; (k = i1) → merid (sym (rCancel (merid base)) (~ r) j) (~ l) ; (l = i0) → merid (doubleCompPath-filler (sym (rCancel (merid base))) - (cong (σ (S₊∙ 1)) loop) + (cong (σSn 1) loop) (rCancel (merid base)) r i j) k ; (l = i1) → 3cell i1 i j k}) (hcomp (λ r → λ {(i = i0) → merid (cp-fill base r j) (k ∧ ~ l) @@ -581,21 +636,21 @@ invSphere² (suc (suc n)) = invSusp² -- Interaction between σ and invSphere σ-invSphere : (n : ℕ) (x : S₊ (suc n)) - → σ (S₊∙ (suc n)) (invSphere x) - ≡ sym (σ (S₊∙ (suc n)) x) + → σSn (suc n) (invSphere x) + ≡ sym (σSn (suc n) x) σ-invSphere zero base = rCancel (merid base) ∙∙ refl ∙∙ cong sym (sym (rCancel (merid base))) σ-invSphere zero (loop i) j = hcomp (λ k → λ { (j = i0) → doubleCompPath-filler (sym (rCancel (merid base))) - (λ i → (σ (S₊∙ 1) (loop (~ i)))) + (λ i → (σSn 1 (loop (~ i)))) (rCancel (merid base)) (~ k) i ; (j = i1) → doubleCompPath-filler (sym (cong sym (rCancel (merid base)))) - (λ i → sym (σ (S₊∙ 1) (loop i))) + (λ i → sym (σSn 1 (loop i))) (cong sym (rCancel (merid base))) (~ k) i}) (sym≡cong-sym (sym (rCancel (merid base)) - ∙∙ (λ i → (σ (S₊∙ 1) (loop i))) + ∙∙ (λ i → (σSn 1 (loop i))) ∙∙ (rCancel (merid base))) j i) σ-invSphere (suc n) x = toSusp-invSusp (S₊∙ (suc n)) x @@ -658,7 +713,7 @@ invSusp∘S¹×S¹→S² (loop i) (loop j) k = ; (k = i1) → invSusp (cp-filler (loop i) r j)}) (merid (loop i) (~ j))) where - σ₁ = σ (S₊∙ 1) + σ₁ = σSn 1 m-b = merid base rCancel-mb = rCancel m-b rCancel-mb⁻¹ = sym (rCancel m-b) diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index 4a9358a86d..db54691e4a 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -9,6 +9,8 @@ open import Cubical.Foundations.Path open import Cubical.Foundations.Pointed open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function + open import Cubical.Data.Bool open import Cubical.Data.Sigma @@ -23,6 +25,22 @@ private open Iso +suspFunComp : {A B C : Type ℓ} (f : B → C) (g : A → B) + → suspFun (f ∘ g) ≡ (suspFun f) ∘ (suspFun g) +suspFunComp f g i north = north +suspFunComp f g i south = south +suspFunComp f g i (merid a i₁) = merid (f (g a)) i₁ + +suspFunConst : {A B : Type ℓ} (b : B) → suspFun (λ (_ : A) → b) ≡ λ _ → north +suspFunConst b i north = north +suspFunConst b i south = merid b (~ i) +suspFunConst b i (merid a j) = merid b (~ i ∧ j) + +suspFunIdFun : {A : Type ℓ} → suspFun (λ (a : A) → a) ≡ λ x → x +suspFunIdFun i north = north +suspFunIdFun i south = south +suspFunIdFun i (merid a j) = merid a j + Susp-iso-joinBool : ∀ {ℓ} {A : Type ℓ} → Iso (Susp A) (join A Bool) fun Susp-iso-joinBool north = inr true fun Susp-iso-joinBool south = inr false diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 591d74f1e5..7f89a76036 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -496,6 +496,14 @@ PathIdTruncIso zero = (isOfHLevelUnit* 0) PathIdTruncIso (suc n) = ΩTrunc.IsoFinal ∣ _ ∣ ∣ _ ∣ +PathPIdTruncIso : {A : I → Type ℓ} {a : A i0} {b : A i1} (n : HLevel) + → Iso (PathP (λ i → ∥ A i ∥ suc n) ∣ a ∣ ∣ b ∣) (∥ PathP (λ i → A i) a b ∥ n) +PathPIdTruncIso {A = A} n = lem (A i0) (A i1) (λ i → A i) n + where + lem : ∀ {ℓ} (A B : Type ℓ) (A' : A ≡ B) {a : A} {b : B} (n : HLevel) + → Iso (PathP (λ i → ∥ A' i ∥ suc n) ∣ a ∣ ∣ b ∣) (∥ PathP (λ i → A' i) a b ∥ n) + lem A = J> PathIdTruncIso + PathIdTrunc : {a b : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) PathIdTrunc n = isoToPath (PathIdTruncIso n) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 5640935b83..b2a5adcb26 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -332,6 +332,26 @@ isConnectedCong n f cf {a₀} {a₁} q = (sym (fiberCong f q)) (isConnectedPath n (cf (f a₁)) (a₀ , q) (a₁ , refl)) +isConnectedCong² : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) + → isConnectedFun (suc (suc n)) f + → ∀ {a₀ a₁ a₂ a₃} {p : a₀ ≡ a₁} {q : a₂ ≡ a₃} + {r : a₀ ≡ a₂} {s : a₁ ≡ a₃} + → isConnectedFun n + {A = Square p q r s} + {B = Square (cong f p) (cong f q) (cong f r) (cong f s)} + (λ p i j → f (p i j)) +isConnectedCong² n {A = A} f cf {a₀} {a₁} {r = r} {s = s} + = isConnectedCong²' _ r _ s + where + isConnectedCong²' : (a₂ : A) (r : a₀ ≡ a₂) (a₃ : A) (s : a₁ ≡ a₃) + {p : a₀ ≡ a₁} {q : a₂ ≡ a₃} + → isConnectedFun n + {A = Square p q r s} + {B = Square (cong f p) (cong f q) (cong f r) (cong f s)} + (λ p i j → f (p i j)) + isConnectedCong²' = + J> (J> isConnectedCong n (cong f) (isConnectedCong (suc n) f cf)) + isConnectedCong' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x : A} {y : B} (n : ℕ) (f : A → B) → isConnectedFun (suc n) f @@ -596,6 +616,18 @@ inrConnected {A = A} {B = B} {C = C} n f g iscon = (equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .snd i a)) +inlConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : ℕ) + → (f : C → A) (g : C → B) + → isConnectedFun n g + → isConnectedFun n {A = A} {B = Pushout f g} inl +inlConnected {A = A} {B = B} {C = C} n f g iscon = + transport (λ i → isConnectedFun n (lem i)) + (inrConnected n g f iscon) + where + lem : PathP (λ i → A → ua (symPushout g f) i) inr inl + lem = toPathP (λ j x → inl (transportRefl (transportRefl x j) j)) + + isConnectedPushout→ : (f₁ : X₀ → X₁) (f₂ : X₀ → X₂) (g₁ : Y₀ → Y₁) (g₂ : Y₀ → Y₂) (h₀ : X₀ → Y₀) (h₁ : X₁ → Y₁) (h₂ : X₂ → Y₂) From c0345787321d5730595763e4d0f1dc706eebd6af Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 18 Jan 2024 17:17:24 +0100 Subject: [PATCH 10/73] stuff --- .../AbGroup/Instances/DirectProduct.agda | 15 + Cubical/Algebra/AbGroup/Instances/Pi.agda | 12 + Cubical/Algebra/AbGroup/Properties.agda | 71 ++- Cubical/Algebra/CommRing/Instances/Int.agda | 27 + Cubical/Algebra/Group/Instances/Pi.agda | 33 ++ Cubical/Axiom/Choice.agda | 102 ++++ Cubical/Cohomology/EilenbergMacLane/Base.agda | 81 ++- .../EilenbergMacLane/EilenbergSteenrod.agda | 521 ++++++++++-------- .../EilenbergMacLane/Groups/KleinBottle.agda | 238 +++++++- .../EilenbergMacLane/Groups/RP2.agda | 286 +++++++++- .../EilenbergMacLane/Groups/Torus.agda | 238 ++++++++ .../EilenbergMacLane/Groups/Unit.agda | 49 ++ .../EilenbergMacLane/Rings/KleinBottle.agda | 1 + Cubical/Data/Fin/Properties.agda | 38 ++ Cubical/HITs/KleinBottle/Properties.agda | 18 + Cubical/HITs/RPn/Base.agda | 14 + Cubical/HITs/Wedge/Base.agda | 7 + .../Homotopy/EilenbergMacLane/Properties.agda | 18 + Cubical/Homotopy/EilenbergSteenrod.agda | 32 ++ .../ComputationalSyntheticCohomology.agda | 334 +++++++++++ 20 files changed, 1876 insertions(+), 259 deletions(-) create mode 100644 Cubical/Algebra/AbGroup/Instances/DirectProduct.agda create mode 100644 Cubical/Algebra/AbGroup/Instances/Pi.agda create mode 100644 Cubical/Algebra/Group/Instances/Pi.agda create mode 100644 Cubical/Axiom/Choice.agda create mode 100644 Cubical/Cohomology/EilenbergMacLane/Groups/Torus.agda create mode 100644 Cubical/Cohomology/EilenbergMacLane/Groups/Unit.agda create mode 100644 Cubical/Papers/ComputationalSyntheticCohomology.agda diff --git a/Cubical/Algebra/AbGroup/Instances/DirectProduct.agda b/Cubical/Algebra/AbGroup/Instances/DirectProduct.agda new file mode 100644 index 0000000000..0848ae9b75 --- /dev/null +++ b/Cubical/Algebra/AbGroup/Instances/DirectProduct.agda @@ -0,0 +1,15 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.AbGroup.Instances.DirectProduct where + +open import Cubical.Data.Sigma +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.Group.DirProd + +AbDirProd : ∀ {ℓ ℓ'} → AbGroup ℓ → AbGroup ℓ' → AbGroup (ℓ-max ℓ ℓ') +AbDirProd G H = + Group→AbGroup (DirProd (AbGroup→Group G) (AbGroup→Group H)) comm + where + comm : (x y : fst G × fst H) → _ ≡ _ + comm (g1 , h1) (g2 , h2) = + ΣPathP (AbGroupStr.+Comm (snd G) g1 g2 + , AbGroupStr.+Comm (snd H) h1 h2) diff --git a/Cubical/Algebra/AbGroup/Instances/Pi.agda b/Cubical/Algebra/AbGroup/Instances/Pi.agda new file mode 100644 index 0000000000..5e61f2928e --- /dev/null +++ b/Cubical/Algebra/AbGroup/Instances/Pi.agda @@ -0,0 +1,12 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.AbGroup.Instances.Pi where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group.Instances.Pi + +module _ {ℓ ℓ' : Level} {X : Type ℓ} (G : X → AbGroup ℓ') where + ΠAbGroup : AbGroup (ℓ-max ℓ ℓ') + ΠAbGroup = Group→AbGroup (ΠGroup (λ x → AbGroup→Group (G x))) + λ f g i x → IsAbGroup.+Comm (AbGroupStr.isAbGroup (G x .snd)) (f x) (g x) i diff --git a/Cubical/Algebra/AbGroup/Properties.agda b/Cubical/Algebra/AbGroup/Properties.agda index 2baeeef678..618f5f7025 100644 --- a/Cubical/Algebra/AbGroup/Properties.agda +++ b/Cubical/Algebra/AbGroup/Properties.agda @@ -7,6 +7,15 @@ open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.QuotientGroup +open import Cubical.Algebra.Group.Subgroup +open import Cubical.Algebra.Group.ZAction + +open import Cubical.HITs.SetQuotients as SQ hiding (_/_) + +open import Cubical.Data.Int using (ℤ ; pos ; negsuc) +open import Cubical.Data.Nat hiding (_+_) +open import Cubical.Data.Sigma private variable ℓ : Level @@ -28,7 +37,7 @@ module AbGroupTheory (A : AbGroup ℓ) where implicitInverse b+a≡0 = invUniqueR b+a≡0 addGroupHom : (A B : AbGroup ℓ) (ϕ ψ : AbGroupHom A B) → AbGroupHom A B -fst (addGroupHom A B ϕ ψ) x = AbGroupStr._+_ (snd B) (ϕ .fst x) (ψ .fst x) +fst (addGroupHom A B ϕ ψ) x = AbGroupStr._+_ (snd B) (ϕ .fst x) (ψ .fst x) snd (addGroupHom A B ϕ ψ) = makeIsGroupHom λ x y → cong₂ (AbGroupStr._+_ (snd B)) (IsGroupHom.pres· (snd ϕ) x y) @@ -49,3 +58,63 @@ snd (negGroupHom A B ϕ) = subtrGroupHom : (A B : AbGroup ℓ) (ϕ ψ : AbGroupHom A B) → AbGroupHom A B subtrGroupHom A B ϕ ψ = addGroupHom A B ϕ (negGroupHom A B ψ) + + + +-- ℤ-multiplication defines a homomorphism for abelian groups +private module _ (G : AbGroup ℓ) where + ℤ·isHom-pos : (n : ℕ) (x y : fst G) + → (pos n ℤ[ (AbGroup→Group G) ]· (AbGroupStr._+_ (snd G) x y)) + ≡ AbGroupStr._+_ (snd G) ((pos n) ℤ[ (AbGroup→Group G) ]· x) + ((pos n) ℤ[ (AbGroup→Group G) ]· y) + ℤ·isHom-pos zero x y = + sym (AbGroupStr.+IdR (snd G) (AbGroupStr.0g (snd G))) + ℤ·isHom-pos (suc n) x y = + cong₂ (AbGroupStr._+_ (snd G)) + refl + (ℤ·isHom-pos n x y) + ∙ AbGroupTheory.comm-4 G _ _ _ _ + + ℤ·isHom-neg : (n : ℕ) (x y : fst G) + → (negsuc n ℤ[ (AbGroup→Group G) ]· (AbGroupStr._+_ (snd G) x y)) + ≡ AbGroupStr._+_ (snd G) ((negsuc n) ℤ[ (AbGroup→Group G) ]· x) + ((negsuc n) ℤ[ (AbGroup→Group G) ]· y) + ℤ·isHom-neg zero x y = + GroupTheory.invDistr (AbGroup→Group G) _ _ + ∙ AbGroupStr.+Comm (snd G) _ _ + ℤ·isHom-neg (suc n) x y = + cong₂ (AbGroupStr._+_ (snd G)) + (GroupTheory.invDistr (AbGroup→Group G) _ _ + ∙ AbGroupStr.+Comm (snd G) _ _) + (ℤ·isHom-neg n x y) + ∙ AbGroupTheory.comm-4 G _ _ _ _ + +ℤ·isHom : (n : ℤ) (G : AbGroup ℓ) (x y : fst G) + → (n ℤ[ (AbGroup→Group G) ]· (AbGroupStr._+_ (snd G) x y)) + ≡ AbGroupStr._+_ (snd G) (n ℤ[ (AbGroup→Group G) ]· x) + (n ℤ[ (AbGroup→Group G) ]· y) +ℤ·isHom (pos n) G = ℤ·isHom-pos G n +ℤ·isHom (negsuc n) G = ℤ·isHom-neg G n + +-- left multiplication as a group homomorphism +multₗHom : (G : AbGroup ℓ) (n : ℤ) → AbGroupHom G G +fst (multₗHom G n) g = n ℤ[ (AbGroup→Group G) ]· g +snd (multₗHom G n) = makeIsGroupHom (ℤ·isHom n G) + +-- Abelian groups quotiented by a natural number +_/^_ : (G : AbGroup ℓ) (n : ℕ) → AbGroup ℓ +G /^ n = + Group→AbGroup + ((AbGroup→Group G) + / (imSubgroup (multₗHom G (pos n)) + , isNormalIm (multₗHom G (pos n)) (AbGroupStr.+Comm (snd G)))) + (SQ.elimProp2 (λ _ _ → squash/ _ _) + λ a b → cong [_] (AbGroupStr.+Comm (snd G) a b)) + +-- Torsion subgrous +_[_] : (G : AbGroup ℓ) (n : ℕ) → AbGroup ℓ +G [ n ] = + Group→AbGroup (Subgroup→Group (AbGroup→Group G) + (kerSubgroup (multₗHom G (pos n)))) + λ {(x , p) (y , q) → Σ≡Prop (λ _ → isPropIsInKer (multₗHom G (pos n)) _) + (AbGroupStr.+Comm (snd G) _ _)} diff --git a/Cubical/Algebra/CommRing/Instances/Int.agda b/Cubical/Algebra/CommRing/Instances/Int.agda index 2e6a6b5360..b5a9e710a9 100644 --- a/Cubical/Algebra/CommRing/Instances/Int.agda +++ b/Cubical/Algebra/CommRing/Instances/Int.agda @@ -4,8 +4,16 @@ module Cubical.Algebra.CommRing.Instances.Int where open import Cubical.Foundations.Prelude open import Cubical.Algebra.CommRing +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Instances.Int + open import Cubical.Data.Int as Int renaming ( ℤ to ℤ ; _+_ to _+ℤ_; _·_ to _·ℤ_; -_ to -ℤ_) +open import Cubical.Data.Nat + renaming (_+_ to _+ℕ_; _·_ to _·ℕ_) +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ open CommRingStr using (0r ; 1r ; _+_ ; _·_ ; -_ ; isCommRing) @@ -23,3 +31,22 @@ isCommRing (snd ℤCommRing) = isCommRingℤ isCommRingℤ = makeIsCommRing isSetℤ Int.+Assoc (λ _ → refl) -Cancel Int.+Comm Int.·Assoc Int.·IdR ·DistR+ Int.·Comm + +ℤAbGroup : AbGroup ℓ-zero +ℤAbGroup = Group→AbGroup ℤGroup +Comm + +ℤTorsion : (n : ℕ) → isContr (fst (ℤAbGroup [ (suc n) ])) +fst (ℤTorsion n) = AbGroupStr.0g (snd (ℤAbGroup [ (suc n) ])) +snd (ℤTorsion n) (a , p) = Σ≡Prop (λ _ → isSetℤ _ _) + (sym (help a (ℤ·≡· (pos (suc n)) a ∙ p))) + where + help : (a : ℤ) → a +ℤ pos n ·ℤ a ≡ 0 → a ≡ 0 + help (pos zero) p = refl + help (pos (suc a)) p = + ⊥.rec (snotz (injPos (pos+ (suc a) (n ·ℕ suc a) + ∙ cong (pos (suc a) +ℤ_) (pos·pos n (suc a)) ∙ p))) + help (negsuc a) p = ⊥.rec + (snotz (injPos (cong -ℤ_ (negsuc+ a (n ·ℕ suc a) + ∙ (cong (negsuc a +ℤ_) + (cong (-ℤ_) (pos·pos n (suc a))) + ∙ sym (cong (negsuc a +ℤ_) (pos·negsuc n a)) ∙ p))))) diff --git a/Cubical/Algebra/Group/Instances/Pi.agda b/Cubical/Algebra/Group/Instances/Pi.agda new file mode 100644 index 0000000000..0d1483652b --- /dev/null +++ b/Cubical/Algebra/Group/Instances/Pi.agda @@ -0,0 +1,33 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Group.Instances.Pi where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + +open IsGroup +open GroupStr +open IsMonoid +open IsSemigroup + +module _ {ℓ ℓ' : Level} {X : Type ℓ} (G : X → Group ℓ') where + -- TODO: upstream + ΠGroup : Group (ℓ-max ℓ ℓ') + fst ΠGroup = (x : X) → fst (G x) + 1g (snd ΠGroup) x = 1g (G x .snd) + GroupStr._·_ (snd ΠGroup) f g x = GroupStr._·_ (G x .snd) (f x) (g x) + inv (snd ΠGroup) f x = inv (G x .snd) (f x) + is-set (isSemigroup (isMonoid (isGroup (snd ΠGroup)))) = + isSetΠ λ x → is-set (G x .snd) + IsSemigroup.·Assoc (isSemigroup (isMonoid (isGroup (snd ΠGroup)))) f g h = + funExt λ x → IsSemigroup.·Assoc (isSemigroup (isMonoid (G x .snd))) (f x) (g x) (h x) + IsMonoid.·IdR (isMonoid (isGroup (snd ΠGroup))) f = + funExt λ x → IsMonoid.·IdR (isMonoid (isGroup (snd (G x)))) (f x) + IsMonoid.·IdL (isMonoid (isGroup (snd ΠGroup))) f = + funExt λ x → IsMonoid.·IdL (isMonoid (isGroup (snd (G x)))) (f x) + ·InvR (isGroup (snd ΠGroup)) f = + funExt λ x → ·InvR (isGroup (snd (G x))) (f x) + ·InvL (isGroup (snd ΠGroup)) f = + funExt λ x → ·InvL (isGroup (snd (G x))) (f x) diff --git a/Cubical/Axiom/Choice.agda b/Cubical/Axiom/Choice.agda new file mode 100644 index 0000000000..ba2cafa007 --- /dev/null +++ b/Cubical/Axiom/Choice.agda @@ -0,0 +1,102 @@ +{-# OPTIONS --safe #-} + +module Cubical.Axiom.Choice where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Fin as FN +open import Cubical.Data.Nat.Order + +open import Cubical.HITs.Truncation as TR +open import Cubical.HITs.PropositionalTruncation as PT + +private + variable + ℓ ℓ' ℓ'' : Level + +-- Stuff about choice: +choiceMap : {A : Type ℓ} {B : A → Type ℓ'} (n : ℕ) + → hLevelTrunc n ((a : A) → B a) + → (a : A) → hLevelTrunc n (B a) +choiceMap n = + TR.rec (isOfHLevelΠ n (λ _ → isOfHLevelTrunc n)) + λ f a → ∣ f a ∣ₕ + +-- n-level axiom of choice +satAC : (ℓ' : Level) (n : ℕ) (A : Type ℓ) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +satAC ℓ' n A = (B : A → Type ℓ') → isEquiv (choiceMap {A = A} {B} n) + +-- Version of (propositional) AC with ∃ +AC∃-map : {A : Type ℓ} {B : A → Type ℓ'} + {C : (a : A) → B a → Type ℓ''} + → ∃[ f ∈ ((a : A) → B a) ] ((a : _) → C a (f a)) + → ((a : A) → ∃ (B a) (C a)) +AC∃-map = + PT.rec (isPropΠ (λ _ → squash₁)) + λ f → λ a → ∣ f .fst a , f .snd a ∣₁ + +satAC∃ : ∀ {ℓ} (ℓ' ℓ'' : Level) (A : Type ℓ) → Type _ +satAC∃ ℓ' ℓ'' A = (B : A → Type ℓ') (C : (a : A) → B a → Type ℓ'') + → isEquiv (AC∃-map {B = B} {C = C}) + +satAC→satAC∃ : {A : Type ℓ} + → satAC (ℓ-max ℓ' ℓ'') (suc zero) A → satAC∃ ℓ' ℓ'' A +satAC→satAC∃ F B C = propBiimpl→Equiv squash₁ (isPropΠ (λ _ → squash₁)) + _ + (λ f → invEq propTrunc≃Trunc1 (TR.map (λ f → fst ∘ f , λ a → f a .snd ) + (invEq (_ , F (λ x → Σ (B x) (C x))) λ a → fst propTrunc≃Trunc1 (f a)))) .snd + +-- All types satisfy (-2) level axiom of choice +satAC₀ : {A : Type ℓ} → satAC ℓ' 0 A +satAC₀ B = isoToIsEquiv (iso (λ _ _ → tt*) (λ _ → tt*) (λ _ → refl) λ _ → refl) + + + +-- Main theorem Fin m satisfies AC for any level n. +FinSatAC : (n m : ℕ) → ∀ {ℓ} → satAC ℓ n (Fin m) +FinSatAC n zero B = + isoToIsEquiv (iso _ + (λ f → ∣ (λ x → ⊥.rec (¬Fin0 x)) ∣ₕ) + (λ f → funExt λ x → ⊥.rec (¬Fin0 x)) + (TR.elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) + λ a → cong ∣_∣ₕ (funExt λ x → ⊥.rec (¬Fin0 x)))) +FinSatAC n (suc m) B = + subst isEquiv (ac-eq n m {B} (FinSatAC n m)) + (isoToIsEquiv (ac-map' n m B (FinSatAC n m))) + where + ac-map' : ∀ {ℓ} (n m : ℕ) (B : Fin (suc m) → Type ℓ) → (satAC ℓ n (Fin m)) + → Iso (hLevelTrunc n ((x : _) → B x)) ((x : _) → hLevelTrunc n (B x)) + ac-map' n m B ise = + compIso (mapCompIso (CharacΠFinIso m)) + (compIso (truncOfProdIso n) + (compIso (Σ-cong-iso-snd λ _ → equivToIso (_ , ise (λ x → B (fsuc x)))) + (invIso (CharacΠFinIso m)))) + + ac-eq : (n m : ℕ) {B : _} → (eq : satAC ℓ n (Fin m)) + → Iso.fun (ac-map' n m B eq) ≡ choiceMap {B = B} n + ac-eq zero m {B = B} x = refl + ac-eq (suc n) m {B = B} x = + funExt (TR.elim (λ _ → isOfHLevelPath (suc n) + (isOfHLevelΠ (suc n) + (λ _ → isOfHLevelTrunc (suc n))) _ _) + λ F → funExt + λ { (zero , p) j → ∣ transp (λ i → B (p1 p (j ∨ i))) j (F (p1 p j)) ∣ₕ + ; (suc x , p) j → ∣ transp (λ i → B (p2 x p (j ∨ i))) j (F (p2 x p j)) ∣ₕ}) + where + p1 : (p : _ ) → fzero ≡ (zero , p) + p1 p = Fin-fst-≡ refl + + p2 : (x : _) (p : suc x < suc m) + → Path (Fin _) (fsuc (x , pred-≤-pred p)) (suc x , p) + p2 x p = Fin-fst-≡ refl + +-- Key result for construction of cw-approx at lvl 0 +satAC∃Fin : (n : ℕ) → satAC∃ ℓ ℓ' (Fin n) +satAC∃Fin n = satAC→satAC∃ (FinSatAC 1 n) diff --git a/Cubical/Cohomology/EilenbergMacLane/Base.agda b/Cubical/Cohomology/EilenbergMacLane/Base.agda index f0c192a445..57cc24968e 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Base.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Base.agda @@ -6,8 +6,10 @@ open import Cubical.Homotopy.EilenbergMacLane.GroupStructure open import Cubical.Homotopy.EilenbergMacLane.Base open import Cubical.Homotopy.EilenbergMacLane.Properties open import Cubical.Homotopy.EilenbergMacLane.Order2 +open import Cubical.Homotopy.Connected open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.Transport open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws @@ -17,6 +19,7 @@ open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.Isomorphism open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as TR open import Cubical.Data.Nat open import Cubical.Data.Sigma @@ -28,6 +31,8 @@ open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.Monoid open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.AbGroup.Instances.DirectProduct +open import Cubical.Algebra.AbGroup.Properties open import Cubical.HITs.SetTruncation as ST hiding (rec ; map ; elim ; elim2 ; elim3) @@ -270,6 +275,44 @@ coHom≅coHomRed n G A = makeIsGroupHom (ST.elim2 (λ _ _ → isSetPathImplicit) λ _ _ → refl) +coHom⁰≅coHomRed⁰ : (G : AbGroup ℓ) (A : Pointed ℓ) + → AbGroupEquiv (AbDirProd (coHomRedGr 0 G A) G) (coHomGr 0 G (typ A)) +fst (coHom⁰≅coHomRed⁰ G A) = isoToEquiv is + where + is : Iso _ _ + Iso.fun is = uncurry (ST.rec (isSetΠ (λ _ → squash₂)) + λ f g → ∣ (λ x → AbGroupStr._+_ (snd G) (fst f x) g) ∣₂) + Iso.inv is = ST.rec (isSet× squash₂ (is-set (snd G))) + λ f → ∣ (λ x → AbGroupStr._-_ (snd G) (f x) (f (pt A))) + , +InvR (snd G) (f (pt A)) ∣₂ , f (pt A) + Iso.rightInv is = ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (funExt λ x + → sym (+Assoc (snd G) _ _ _) + ∙∙ cong (AbGroupStr._+_ (snd G) (f x)) + (+InvL (snd G) (f (pt A))) + ∙∙ +IdR (snd G) (f x)) + Iso.leftInv is = + uncurry (ST.elim + (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 + (isSet× squash₂ (is-set (snd G))) _ _)) + λ f → λ g + → ΣPathP (cong ∣_∣₂ (Σ≡Prop (λ _ → is-set (snd G) _ _) + (funExt (λ x → cong₂ (AbGroupStr._+_ (snd G)) + refl + (cong (- (snd G)) + (cong₂ (AbGroupStr._+_ (snd G)) (snd f) refl + ∙ +IdL (snd G) g)) + ∙ sym (+Assoc (snd G) _ _ _) + ∙ cong (AbGroupStr._+_ (snd G) (fst f x)) + (+InvR (snd G) g) + ∙ +IdR (snd G) (f .fst x)))) + , (cong (λ x → AbGroupStr._+_ (snd G) x g) (snd f) + ∙ +IdL (snd G) g))) +snd (coHom⁰≅coHomRed⁰ G A) = + makeIsGroupHom (uncurry (ST.elim (λ _ → isSetΠ2 λ _ _ → isSetPathImplicit) + λ f1 g1 → uncurry (ST.elim (λ _ → isSetΠ λ _ → isSetPathImplicit) + λ f2 g2 → cong ∣_∣₂ (funExt λ a → AbGroupTheory.comm-4 G _ _ _ _)))) + -- ℤ/2 lemmas +ₕ≡id-ℤ/2 : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (x : coHom n ℤ/2 A) → x +ₕ x ≡ 0ₕ n +ₕ≡id-ℤ/2 n = @@ -297,6 +340,23 @@ snd (coHomHom n f) = makeIsGroupHom (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) +coHomEquiv : ∀ {ℓ''} {A : Type ℓ} {B : Type ℓ'} {G : AbGroup ℓ''} + (n : ℕ) + → (Iso A B) + → AbGroupEquiv (coHomGr n G B) (coHomGr n G A) +fst (coHomEquiv n f) = isoToEquiv is + where + is : Iso _ _ + Iso.fun is = coHomFun n (Iso.fun f) + Iso.inv is = coHomFun n (Iso.inv f) + Iso.rightInv is = + ST.elim (λ _ → isSetPathImplicit) + λ g → cong ∣_∣₂ (funExt λ x → cong g (Iso.leftInv f x)) + Iso.leftInv is = + ST.elim (λ _ → isSetPathImplicit) + λ g → cong ∣_∣₂ (funExt λ x → cong g (Iso.rightInv f x)) +snd (coHomEquiv n f) = snd (coHomHom n (Iso.fun f)) + coHomFun∙ : ∀ {ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {G : AbGroup ℓ''} (n : ℕ) (f : A →∙ B) → coHomRed n G B → coHomRed n G A @@ -343,7 +403,26 @@ subst-EM-0ₖ : ∀{ℓ} {G : AbGroup ℓ} {n m : ℕ} (p : n ≡ m) subst-EM-0ₖ {G = G} {n = n} = J (λ m p → subst (EM G) p (0ₖ n) ≡ 0ₖ m) (transportRefl _) -subst-EM∙ : ∀{ℓ} {G : AbGroup ℓ} {n m : ℕ} (p : n ≡ m) +subst-EM∙ : ∀ {ℓ} {G : AbGroup ℓ} {n m : ℕ} (p : n ≡ m) → EM∙ G n →∙ EM∙ G m fst (subst-EM∙ {G = G} p) = subst (EM G) p snd (subst-EM∙ p) = subst-EM-0ₖ p + +coHomPointedElim : ∀ {ℓ ℓ' ℓ''} {G : AbGroup ℓ} + {A : Type ℓ'} (n : ℕ) (a : A) {B : coHom (suc n) G A → Type ℓ''} + → ((x : coHom (suc n) G A) → isProp (B x)) + → ((f : A → EM G (suc n)) → f a ≡ 0ₖ (suc n) → B ∣ f ∣₂) + → (x : coHom (suc n) G A) → B x +coHomPointedElim {ℓ'' = ℓ''} {G = G} {A = A} n a isprop indp = + ST.elim (λ _ → isOfHLevelSuc 1 (isprop _)) + λ f → helper n isprop indp f + where + helper : (n : ℕ) {B : coHom (suc n) G A → Type ℓ''} + → ((x : coHom (suc n) G A) → isProp (B x)) + → ((f : A → EM G (suc n)) → f a ≡ 0ₖ (suc n) → B ∣ f ∣₂) + → (f : A → EM G (suc n)) + → B ∣ f ∣₂ + helper n isprop ind f = + TR.rec (isProp→isOfHLevelSuc n (isprop _)) + (ind f) + (isConnectedPath (suc n) (isConnectedEM (suc n)) (f a) (0ₖ (suc n)) .fst) diff --git a/Cubical/Cohomology/EilenbergMacLane/EilenbergSteenrod.agda b/Cubical/Cohomology/EilenbergMacLane/EilenbergSteenrod.agda index 0cd001a289..b6cdb7aaa0 100644 --- a/Cubical/Cohomology/EilenbergMacLane/EilenbergSteenrod.agda +++ b/Cubical/Cohomology/EilenbergMacLane/EilenbergSteenrod.agda @@ -25,6 +25,7 @@ open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.Unit open import Cubical.Algebra.Group.Instances.Unit +open import Cubical.Algebra.AbGroup.Instances.Pi open import Cubical.HITs.Wedge open import Cubical.HITs.Pushout @@ -40,243 +41,309 @@ open import Cubical.Data.Sigma open import Cubical.Data.Int open import Cubical.Data.Unit +open import Cubical.Axiom.Choice + open Iso -open coHomTheory +module _ where + open coHomTheory + + -- This module verifies the (reduced) Eilenberg-Steenrod axioms for + -- the following ℤ-indexed cohomology theory + + coHomRedℤ : ∀ {ℓ ℓ'} → AbGroup ℓ' → ℤ → Pointed ℓ → AbGroup (ℓ-max ℓ ℓ') + coHomRedℤ G (pos n) A = coHomRedGr n G A + coHomRedℤ G (negsuc n) A = UnitAbGroup --- This module verifies the (reduced) Eilenberg-Steenrod axioms for --- the following ℤ-indexed cohomology theory + module coHomRedℤ {ℓ ℓ'} {G : AbGroup ℓ} where + Hmap∙ : (n : ℤ) {A B : Pointed ℓ'} + → A →∙ B + → AbGroupHom (coHomRedℤ G n B) (coHomRedℤ G n A) + Hmap∙ (pos n) = coHomHom∙ n + Hmap∙ (negsuc n) f = idGroupHom -coHomRedℤ : ∀ {ℓ ℓ'} → AbGroup ℓ' → ℤ → Pointed ℓ → AbGroup (ℓ-max ℓ ℓ') -coHomRedℤ G (pos n) A = coHomRedGr n G A -coHomRedℤ G (negsuc n) A = UnitAbGroup + suspMap : (n : ℤ) {A : Pointed ℓ'} → + AbGroupHom (coHomRedℤ G (sucℤ n) (Susp (typ A) , north)) + (coHomRedℤ G n A) + fst (suspMap (pos n) {A = A}) = + ST.map λ f → (λ x → ΩEM+1→EM n + (sym (snd f) + ∙∙ cong (fst f) (merid x ∙ sym (merid (pt A))) + ∙∙ snd f)) + , cong (ΩEM+1→EM n) + (cong (sym (snd f) ∙∙_∙∙ snd f) + (cong (cong (fst f)) + (rCancel (merid (pt A)))) + ∙ ∙∙lCancel (snd f)) + ∙ ΩEM+1→EM-refl n + snd (suspMap (pos n) {A = A}) = + makeIsGroupHom + (ST.elim2 (λ _ _ → isSetPathImplicit) + (λ f g → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) + (funExt λ x → main n _ (sym (snd f)) _ (sym (snd g)) + (cong (fst f) (merid x ∙ sym (merid (pt A)))) + (cong (fst g) (merid x ∙ sym (merid (pt A)))))))) + where + main : (n : ℕ) (x : EM G (suc n)) (f0 : 0ₖ (suc n) ≡ x) + (y : EM G (suc n)) (g0 : 0ₖ (suc n) ≡ y) + (f1 : x ≡ x) (g1 : y ≡ y) + → ΩEM+1→EM n (sym (cong₂ _+ₖ_ (sym f0) (sym g0) + ∙ rUnitₖ (suc n) (0ₖ (suc n))) + ∙∙ cong₂ _+ₖ_ f1 g1 + ∙∙ (cong₂ _+ₖ_ (sym f0) (sym g0) + ∙ rUnitₖ (suc n) (0ₖ (suc n)))) + ≡ ΩEM+1→EM n (f0 ∙∙ f1 ∙∙ sym f0) + +[ n ]ₖ ΩEM+1→EM n (g0 ∙∙ g1 ∙∙ sym g0) + main zero = J> (J> λ f g → + cong (ΩEM+1→EM {G = G} zero) + (cong (λ x → sym x ∙∙ cong₂ _+ₖ_ f g ∙∙ x) + (sym (rUnit refl)) + ∙∙ sym (rUnit _) + ∙∙ cong₂+₁ f g) + ∙∙ ΩEM+1→EM-hom zero f g + ∙∙ cong₂ _+ₖ_ (cong (ΩEM+1→EM zero) (rUnit f)) + (cong (ΩEM+1→EM zero) (rUnit g))) + main (suc n) = + J> (J> λ f g → + cong (ΩEM+1→EM {G = G} (suc n)) + (cong (λ x → sym x ∙∙ cong₂ _+ₖ_ f g ∙∙ x) + (sym (rUnit refl)) + ∙∙ sym (rUnit _) + ∙∙ cong₂+₂ n f g) + ∙∙ ΩEM+1→EM-hom (suc n) f g + ∙∙ cong₂ _+ₖ_ (cong (ΩEM+1→EM (suc n)) (rUnit f)) + (cong (ΩEM+1→EM (suc n)) (rUnit g))) + fst (suspMap (negsuc n)) _ = tt* + snd (suspMap (negsuc n)) = makeIsGroupHom λ _ _ → refl -module coHomRedℤ {ℓ ℓ'} {G : AbGroup ℓ} where - Hmap∙ : (n : ℤ) {A B : Pointed ℓ'} - → A →∙ B - → AbGroupHom (coHomRedℤ G n B) (coHomRedℤ G n A) - Hmap∙ (pos n) = coHomHom∙ n - Hmap∙ (negsuc n) f = idGroupHom + toSusp-coHomRed : (n : ℕ) {A : Pointed ℓ'} + → A →∙ EM∙ G n → (Susp (typ A) , north) →∙ EM∙ G (suc n) + fst (toSusp-coHomRed n f) north = 0ₖ (suc n) + fst (toSusp-coHomRed n f) south = 0ₖ (suc n) + fst (toSusp-coHomRed n f) (merid a i) = EM→ΩEM+1 n (fst f a) i + snd (toSusp-coHomRed n f) = refl - suspMap : (n : ℤ) {A : Pointed ℓ'} → - AbGroupHom (coHomRedℤ G (sucℤ n) (Susp (typ A) , north)) - (coHomRedℤ G n A) - fst (suspMap (pos n) {A = A}) = - ST.map λ f → (λ x → ΩEM+1→EM n - (sym (snd f) - ∙∙ cong (fst f) (merid x ∙ sym (merid (pt A))) - ∙∙ snd f)) - , cong (ΩEM+1→EM n) - (cong (sym (snd f) ∙∙_∙∙ snd f) - (cong (cong (fst f)) - (rCancel (merid (pt A)))) - ∙ ∙∙lCancel (snd f)) - ∙ ΩEM+1→EM-refl n - snd (suspMap (pos n) {A = A}) = - makeIsGroupHom - (ST.elim2 (λ _ _ → isSetPathImplicit) - (λ f g → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) - (funExt λ x → main n _ (sym (snd f)) _ (sym (snd g)) - (cong (fst f) (merid x ∙ sym (merid (pt A)))) - (cong (fst g) (merid x ∙ sym (merid (pt A)))))))) + suspMapIso : (n : ℤ) {A : Pointed ℓ'} → + AbGroupIso (coHomRedℤ G (sucℤ n) (Susp (typ A) , north)) + (coHomRedℤ G n A) + fun (fst (suspMapIso n)) = suspMap n .fst + inv (fst (suspMapIso (pos n))) = ST.map (toSusp-coHomRed n) + inv (fst (suspMapIso (negsuc zero))) _ = 0ₕ∙ zero + inv (fst (suspMapIso (negsuc (suc n)))) _ = tt* + rightInv (fst (suspMapIso (pos n) {A = A})) = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) + (funExt λ x → cong (ΩEM+1→EM n) + (sym (rUnit _) + ∙∙ cong-∙ (toSusp-coHomRed n f .fst) + (merid x) (sym (merid (pt A))) + ∙∙ cong (EM→ΩEM+1 n (fst f x) ∙_) + (cong sym (cong (EM→ΩEM+1 n) (snd f) + ∙ EM→ΩEM+1-0ₖ n)) + ∙ sym (rUnit _)) + ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) (fst f x))) + rightInv (fst (suspMapIso (negsuc zero))) tt* = refl + rightInv (fst (suspMapIso (negsuc (suc n)))) tt* = refl + leftInv (fst (suspMapIso (pos n) {A = A})) = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM (suc n)) + (funExt λ { north → sym (snd f) + ; south → sym (snd f) ∙ cong (fst f) (merid (pt A)) + ; (merid a i) j → lem a f j i})) + where + lem : (a : typ A) (f : Susp∙ (typ A) →∙ EM∙ G (suc n)) + → PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (pt A))) i) + (EM→ΩEM+1 n (ΩEM+1→EM n (sym (snd f) + ∙∙ cong (fst f) (toSusp A a) + ∙∙ snd f))) + (cong (fst f) (merid a)) + lem a f = Iso.rightInv (Iso-EM-ΩEM+1 n) _ + ◁ λ i j → hcomp (λ k + → λ { (i = i1) → fst f (merid a j) + ; (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' + (sym (snd f)) (cong (fst f) (merid (pt A))) k i}) + (fst f (compPath-filler (merid a) + (sym (merid (pt A))) (~ i) j)) + leftInv (fst (suspMapIso (negsuc zero) {A = A})) = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (Σ≡Prop (λ _ → hLevelEM G 0 _ _) + (funExt (suspToPropElim (pt A) + (λ _ → hLevelEM G 0 _ _) + (sym (snd f))))) + leftInv (fst (suspMapIso (negsuc (suc n)))) tt* = refl + snd (suspMapIso n) = suspMap n .snd + + satisfies-ES : ∀ {ℓ ℓ'} (G : AbGroup ℓ) → coHomTheory {ℓ'} (coHomRedℤ G) + Hmap (satisfies-ES G) = coHomRedℤ.Hmap∙ + fst (Suspension (satisfies-ES G)) n = GroupIso→GroupEquiv (coHomRedℤ.suspMapIso n) + snd (Suspension (satisfies-ES G)) f (pos n) = + funExt (ST.elim (λ _ → isSetPathImplicit) λ f + → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM (suc n)) + (funExt λ { north → refl + ; south → refl + ; (merid a i) → refl}))) + snd (Suspension (satisfies-ES G)) f (negsuc zero) = + funExt λ {tt* → cong ∣_∣₂ (Σ≡Prop (λ _ → hLevelEM G 0 _ _) refl)} + snd (Suspension (satisfies-ES G)) f (negsuc (suc n)) = refl + Exactness (satisfies-ES G) {A = A} {B = B} f (pos n) = isoToPath help where - main : (n : ℕ) (x : EM G (suc n)) (f0 : 0ₖ (suc n) ≡ x) - (y : EM G (suc n)) (g0 : 0ₖ (suc n) ≡ y) - (f1 : x ≡ x) (g1 : y ≡ y) - → ΩEM+1→EM n (sym (cong₂ _+ₖ_ (sym f0) (sym g0) - ∙ rUnitₖ (suc n) (0ₖ (suc n))) - ∙∙ cong₂ _+ₖ_ f1 g1 - ∙∙ (cong₂ _+ₖ_ (sym f0) (sym g0) - ∙ rUnitₖ (suc n) (0ₖ (suc n)))) - ≡ ΩEM+1→EM n (f0 ∙∙ f1 ∙∙ sym f0) - +[ n ]ₖ ΩEM+1→EM n (g0 ∙∙ g1 ∙∙ sym g0) - main zero = J> (J> λ f g → - cong (ΩEM+1→EM {G = G} zero) - (cong (λ x → sym x ∙∙ cong₂ _+ₖ_ f g ∙∙ x) - (sym (rUnit refl)) - ∙∙ sym (rUnit _) - ∙∙ cong₂+₁ f g) - ∙∙ ΩEM+1→EM-hom zero f g - ∙∙ cong₂ _+ₖ_ (cong (ΩEM+1→EM zero) (rUnit f)) - (cong (ΩEM+1→EM zero) (rUnit g))) - main (suc n) = - J> (J> λ f g → - cong (ΩEM+1→EM {G = G} (suc n)) - (cong (λ x → sym x ∙∙ cong₂ _+ₖ_ f g ∙∙ x) - (sym (rUnit refl)) - ∙∙ sym (rUnit _) - ∙∙ cong₂+₂ n f g) - ∙∙ ΩEM+1→EM-hom (suc n) f g - ∙∙ cong₂ _+ₖ_ (cong (ΩEM+1→EM (suc n)) (rUnit f)) - (cong (ΩEM+1→EM (suc n)) (rUnit g))) - fst (suspMap (negsuc n)) _ = tt* - snd (suspMap (negsuc n)) = makeIsGroupHom λ _ _ → refl + To : (x : coHomRed n G B) + → isInKer (coHomHom∙ n f) x + → isInIm (coHomHom∙ n (cfcod (fst f) , refl)) x + To = ST.elim (λ _ → isSetΠ λ _ → isProp→isSet (isPropIsInIm _ _)) + λ g p → PT.map (λ id → ∣ (λ { (inl x) → 0ₖ n + ; (inr x) → fst g x + ; (push a i) → id (~ i) .fst a}) + , snd g ∣₂ + , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) + refl)) + (Iso.fun ST.PathIdTrunc₀Iso p) - toSusp-coHomRed : (n : ℕ) {A : Pointed ℓ'} - → A →∙ EM∙ G n → (Susp (typ A) , north) →∙ EM∙ G (suc n) - fst (toSusp-coHomRed n f) north = 0ₖ (suc n) - fst (toSusp-coHomRed n f) south = 0ₖ (suc n) - fst (toSusp-coHomRed n f) (merid a i) = EM→ΩEM+1 n (fst f a) i - snd (toSusp-coHomRed n f) = refl + From : (x : coHomRed n G B) + → isInIm (coHomHom∙ n (cfcod (fst f) , refl)) x + → isInKer (coHomHom∙ n f) x + From = + ST.elim (λ _ → isSetΠ λ _ → isProp→isSet (isPropIsInKer _ _)) + λ g → PT.rec (isPropIsInKer _ _) + (uncurry (ST.elim (λ _ → isSetΠ λ _ → isProp→isSet (isPropIsInKer _ _)) + λ h p → PT.rec (squash₂ _ _) + (λ id → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) + (funExt λ x → sym (funExt⁻ (cong fst id) (fst f x)) + ∙ cong (fst h) (sym (push x) + ∙ push (pt A) + ∙ λ i → inr (snd f i)) + ∙ snd h))) + (Iso.fun ST.PathIdTrunc₀Iso p))) - suspMapIso : (n : ℤ) {A : Pointed ℓ'} → - AbGroupIso (coHomRedℤ G (sucℤ n) (Susp (typ A) , north)) - (coHomRedℤ G n A) - fun (fst (suspMapIso n)) = suspMap n .fst - inv (fst (suspMapIso (pos n))) = ST.map (toSusp-coHomRed n) - inv (fst (suspMapIso (negsuc zero))) _ = 0ₕ∙ zero - inv (fst (suspMapIso (negsuc (suc n)))) _ = tt* - rightInv (fst (suspMapIso (pos n) {A = A})) = - ST.elim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) - (funExt λ x → cong (ΩEM+1→EM n) - (sym (rUnit _) - ∙∙ cong-∙ (toSusp-coHomRed n f .fst) - (merid x) (sym (merid (pt A))) - ∙∙ cong (EM→ΩEM+1 n (fst f x) ∙_) - (cong sym (cong (EM→ΩEM+1 n) (snd f) - ∙ EM→ΩEM+1-0ₖ n)) - ∙ sym (rUnit _)) - ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) (fst f x))) - rightInv (fst (suspMapIso (negsuc zero))) tt* = refl - rightInv (fst (suspMapIso (negsuc (suc n)))) tt* = refl - leftInv (fst (suspMapIso (pos n) {A = A})) = - ST.elim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM (suc n)) - (funExt λ { north → sym (snd f) - ; south → sym (snd f) ∙ cong (fst f) (merid (pt A)) - ; (merid a i) j → lem a f j i})) + help : Iso (Ker (coHomHom∙ n f)) + (Im (coHomHom∙ n (cfcod (fst f) , refl))) + fun help (x , p) = x , To x p + inv help (x , p) = x , From x p + rightInv help (x , p) = Σ≡Prop (λ _ → isPropIsInIm _ _) refl + leftInv help (x , p) = Σ≡Prop (λ _ → isPropIsInKer _ _) refl + Exactness (satisfies-ES {ℓ} {ℓ'} G) {A = A} {B = B} f (negsuc n) = + isoToPath help + where + help : Iso (Ker (coHomRedℤ.Hmap∙ {G = G} (negsuc n) f)) + (Im (coHomRedℤ.Hmap∙ {G = G} (negsuc n) {A = B} + (cfcod (fst f) , refl))) + fun help (tt* , p) = tt* , ∣ tt* , refl ∣₁ + inv help (tt* , p) = tt* , refl + rightInv help (tt* , p) = Σ≡Prop (λ _ → isPropIsInIm _ _) refl + leftInv help (tt* , p) = Σ≡Prop (λ _ → isPropIsInKer _ _) refl + Dimension (satisfies-ES G) (pos zero) p = ⊥.rec (p refl) + fst (Dimension (satisfies-ES G) (pos (suc n)) _) = 0ₕ∙ (suc n) + snd (Dimension (satisfies-ES G) (pos (suc n)) _) = + ST.elim (λ _ → isSetPathImplicit) + λ f → T.rec (isProp→isOfHLevelSuc n (squash₂ _ _)) + (λ p → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM (suc n)) + (funExt λ { (lift false) → sym p + ; (lift true) → sym (snd f)}))) + (Iso.fun (PathIdTruncIso (suc n)) + (isContr→isProp (isConnectedEM {G' = G} (suc n)) + ∣ fst f (lift false) ∣ ∣ 0ₖ (suc n) ∣)) + Dimension (satisfies-ES G) (negsuc n) _ = isContrUnit* + BinaryWedge (satisfies-ES G) (pos n) {A = A} {B = B} = + GroupIso→GroupEquiv main where - lem : (a : typ A) (f : Susp∙ (typ A) →∙ EM∙ G (suc n)) - → PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (pt A))) i) - (EM→ΩEM+1 n (ΩEM+1→EM n (sym (snd f) - ∙∙ cong (fst f) (toSusp A a) - ∙∙ snd f))) - (cong (fst f) (merid a)) - lem a f = Iso.rightInv (Iso-EM-ΩEM+1 n) _ - ◁ λ i j → hcomp (λ k - → λ { (i = i1) → fst f (merid a j) - ; (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' - (sym (snd f)) (cong (fst f) (merid (pt A))) k i}) - (fst f (compPath-filler (merid a) - (sym (merid (pt A))) (~ i) j)) - leftInv (fst (suspMapIso (negsuc zero) {A = A})) = - ST.elim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (Σ≡Prop (λ _ → hLevelEM G 0 _ _) - (funExt (suspToPropElim (pt A) - (λ _ → hLevelEM G 0 _ _) - (sym (snd f))))) - leftInv (fst (suspMapIso (negsuc (suc n)))) tt* = refl - snd (suspMapIso n) = suspMap n .snd + main : GroupIso _ _ + fun (fst main) = + ST.rec (isSet× squash₂ squash₂) + λ f → ∣ f ∘∙ (inl , refl) ∣₂ , ∣ f ∘∙ (inr , sym (push tt)) ∣₂ + inv (fst main) = + uncurry (ST.rec2 squash₂ + λ f g → ∣ (λ { (inl x) → fst f x + ; (inr x) → fst g x + ; (push a i) → (snd f ∙ sym (snd g)) i}) + , snd f ∣₂) + rightInv (fst main) = + uncurry + (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSet× squash₂ squash₂) _ _) + λ f g → ΣPathP + ((cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl)) + , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl))) + leftInv (fst main) = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j → lem (snd f) (cong (fst f) (push tt)) j i})) + where + lem : ∀ {ℓ} {A : Type ℓ} {x y z : A} + → (p : x ≡ y) (q : x ≡ z) + → (refl ∙ p) ∙ sym (sym q ∙ p) ≡ q + lem p q = cong₂ _∙_ (sym (lUnit p)) (symDistr (sym q) p) + ∙ assoc p (sym p) q + ∙ cong (_∙ q) (rCancel p) + ∙ sym (lUnit q) + snd main = + makeIsGroupHom + (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSet× squash₂ squash₂) _ _) + λ f g → ΣPathP + (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl) + , (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl)))) + BinaryWedge (satisfies-ES G) (negsuc n) {A = A} {B = B} = + invGroupEquiv (GroupIso→GroupEquiv lUnitGroupIso^) -satisfies-ES : ∀ {ℓ ℓ'} (G : AbGroup ℓ) → coHomTheory {ℓ'} (coHomRedℤ G) -Hmap (satisfies-ES G) = coHomRedℤ.Hmap∙ -fst (Suspension (satisfies-ES G)) n = GroupIso→GroupEquiv (coHomRedℤ.suspMapIso n) -snd (Suspension (satisfies-ES G)) f (pos n) = - funExt (ST.elim (λ _ → isSetPathImplicit) λ f - → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM (suc n)) - (funExt λ { north → refl - ; south → refl - ; (merid a i) → refl}))) -snd (Suspension (satisfies-ES G)) f (negsuc zero) = - funExt λ {tt* → cong ∣_∣₂ (Σ≡Prop (λ _ → hLevelEM G 0 _ _) refl)} -snd (Suspension (satisfies-ES G)) f (negsuc (suc n)) = refl -Exactness (satisfies-ES G) {A = A} {B = B} f (pos n) = isoToPath help - where - To : (x : coHomRed n G B) - → isInKer (coHomHom∙ n f) x - → isInIm (coHomHom∙ n (cfcod (fst f) , refl)) x - To = ST.elim (λ _ → isSetΠ λ _ → isProp→isSet (isPropIsInIm _ _)) - λ g p → PT.map (λ id → ∣ (λ { (inl x) → 0ₖ n - ; (inr x) → fst g x - ; (push a i) → id (~ i) .fst a}) - , snd g ∣₂ - , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) - refl)) - (Iso.fun ST.PathIdTrunc₀Iso p) +open coHomTheoryGen +private + compAx : ∀ {ℓ ℓ'} (G : AbGroup ℓ) (n : ℤ) {A B C : Pointed ℓ'} + (g : B →∙ C) (f : A →∙ B) → + compGroupHom (coHomRedℤ.Hmap∙ {G = G} n g) (coHomRedℤ.Hmap∙ n f) + ≡ coHomRedℤ.Hmap∙ n (g ∘∙ f) + compAx G (pos n) g f = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (ST.elim (λ _ → isSetPathImplicit) + λ a → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl))) + compAx G (negsuc n) g f = Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl - From : (x : coHomRed n G B) - → isInIm (coHomHom∙ n (cfcod (fst f) , refl)) x - → isInKer (coHomHom∙ n f) x - From = - ST.elim (λ _ → isSetΠ λ _ → isProp→isSet (isPropIsInKer _ _)) - λ g → PT.rec (isPropIsInKer _ _) - (uncurry (ST.elim (λ _ → isSetΠ λ _ → isProp→isSet (isPropIsInKer _ _)) - λ h p → PT.rec (squash₂ _ _) - (λ id → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) - (funExt λ x → sym (funExt⁻ (cong fst id) (fst f x)) - ∙ cong (fst h) (sym (push x) - ∙ push (pt A) - ∙ λ i → inr (snd f i)) - ∙ snd h))) - (Iso.fun ST.PathIdTrunc₀Iso p))) + idAx : ∀ {ℓ ℓ'} (G : AbGroup ℓ) (n : ℤ) {A : Pointed ℓ'} → + coHomRedℤ.Hmap∙ {G = G} n (idfun∙ A) ≡ idGroupHom + idAx G (pos n) = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (ΣPathP (refl , sym (lUnit (snd f)))))) + idAx G (negsuc n) = refl - help : Iso (Ker (coHomHom∙ n f)) - (Im (coHomHom∙ n (cfcod (fst f) , refl))) - fun help (x , p) = x , To x p - inv help (x , p) = x , From x p - rightInv help (x , p) = Σ≡Prop (λ _ → isPropIsInIm _ _) refl - leftInv help (x , p) = Σ≡Prop (λ _ → isPropIsInKer _ _) refl -Exactness (satisfies-ES {ℓ} {ℓ'} G) {A = A} {B = B} f (negsuc n) = - isoToPath help +satisfies-ES-gen : ∀ {ℓ ℓ'} (G : AbGroup ℓ) → coHomTheoryGen {ℓ'} (coHomRedℤ G) +Hmap (satisfies-ES-gen G) = coHomTheory.Hmap (satisfies-ES G) +HMapComp (satisfies-ES-gen G) = compAx G +HMapId (satisfies-ES-gen G) = idAx G +Suspension (satisfies-ES-gen G) = coHomTheory.Suspension (satisfies-ES G) +Exactness (satisfies-ES-gen G) = coHomTheory.Exactness (satisfies-ES G) +Dimension (satisfies-ES-gen G) = coHomTheory.Dimension (satisfies-ES G) +Wedge (satisfies-ES-gen G) (pos n) {I = I} satAC {A = A} = + subst isEquiv altEquiv≡ (snd altEquiv) where - help : Iso (Ker (coHomRedℤ.Hmap∙ {G = G} (negsuc n) f)) - (Im (coHomRedℤ.Hmap∙ {G = G} (negsuc n) {A = B} - (cfcod (fst f) , refl))) - fun help (tt* , p) = tt* , ∣ tt* , refl ∣₁ - inv help (tt* , p) = tt* , refl - rightInv help (tt* , p) = Σ≡Prop (λ _ → isPropIsInIm _ _) refl - leftInv help (tt* , p) = Σ≡Prop (λ _ → isPropIsInKer _ _) refl -Dimension (satisfies-ES G) (pos zero) p = ⊥.rec (p refl) -fst (Dimension (satisfies-ES G) (pos (suc n)) _) = 0ₕ∙ (suc n) -snd (Dimension (satisfies-ES G) (pos (suc n)) _) = - ST.elim (λ _ → isSetPathImplicit) - λ f → T.rec (isProp→isOfHLevelSuc n (squash₂ _ _)) - (λ p → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM (suc n)) - (funExt λ { (lift false) → sym p - ; (lift true) → sym (snd f)}))) - (Iso.fun (PathIdTruncIso (suc n)) - (isContr→isProp (isConnectedEM {G' = G} (suc n)) - ∣ fst f (lift false) ∣ ∣ 0ₖ (suc n) ∣)) -Dimension (satisfies-ES G) (negsuc n) _ = isContrUnit* -BinaryWedge (satisfies-ES G) (pos n) {A = A} {B = B} = - GroupIso→GroupEquiv main - where - main : GroupIso _ _ - fun (fst main) = - ST.rec (isSet× squash₂ squash₂) - λ f → ∣ f ∘∙ (inl , refl) ∣₂ , ∣ f ∘∙ (inr , sym (push tt)) ∣₂ - inv (fst main) = - uncurry (ST.rec2 squash₂ - λ f g → ∣ (λ { (inl x) → fst f x - ; (inr x) → fst g x - ; (push a i) → (snd f ∙ sym (snd g)) i}) - , snd f ∣₂) - rightInv (fst main) = - uncurry - (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSet× squash₂ squash₂) _ _) - λ f g → ΣPathP - ((cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl)) - , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl))) - leftInv (fst main) = - ST.elim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) - (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push a i) j → lem (snd f) (cong (fst f) (push tt)) j i})) - where - lem : ∀ {ℓ} {A : Type ℓ} {x y z : A} - → (p : x ≡ y) (q : x ≡ z) - → (refl ∙ p) ∙ sym (sym q ∙ p) ≡ q - lem p q = cong₂ _∙_ (sym (lUnit p)) (symDistr (sym q) p) - ∙ assoc p (sym p) q - ∙ cong (_∙ q) (rCancel p) - ∙ sym (lUnit q) - snd main = - makeIsGroupHom - (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSet× squash₂ squash₂) _ _) - λ f g → ΣPathP - (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl) - , (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl)))) -BinaryWedge (satisfies-ES G) (negsuc n) {A = A} {B = B} = - invGroupEquiv (GroupIso→GroupEquiv lUnitGroupIso^) + eq : _ ≃ _ + fst eq = _ + snd eq = satAC _ + + mainIso : Iso _ _ + fun mainIso = ST.map λ f i + → (λ x → f .fst (inr (i , x))) , cong (fst f) (sym (push i)) ∙ snd f + inv mainIso = ST.map λ f → (λ { (inl x) → 0ₖ n + ; (inr x) → f (fst x) .fst (snd x) + ; (push a i) → f a .snd (~ i)}) + , refl + rightInv mainIso = ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (funExt (λ i → ΣPathP (refl , (sym (rUnit (snd (f i))))))) + leftInv mainIso = ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (ΣPathP ((funExt + (λ { (inl x) → sym (snd f) + ; (inr x) → refl + ; (push a i) j + → compPath-filler (cong (fst f) (sym (push a))) (snd f) (~ j) (~ i)})) + , λ i j → snd f (~ i ∨ j))) + + altEquiv : (coHomRedℤ G (pos n) (⋁gen∙ I A) .fst) + ≃ (ΠAbGroup (λ i → coHomRedℤ G (pos n) (A i)) .fst) + altEquiv = isoToEquiv + (compIso + (compIso (compIso mainIso setTruncTrunc2Iso) + (equivToIso eq)) + (codomainIsoDep λ i → invIso setTruncTrunc2Iso)) + + altEquiv≡ : fst altEquiv ≡ _ + altEquiv≡ = funExt (ST.elim (λ _ → isOfHLevelPath 2 (isSetΠ (λ _ → squash₂)) _ _) + λ _ → refl) +Wedge (satisfies-ES-gen G) (negsuc n) {I = I} satAC {A = A} = + isoToIsEquiv (iso _ (λ _ → tt*) (λ _ → refl) λ _ → refl) diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda index 9f76f417e0..c0042296c5 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda @@ -4,6 +4,7 @@ module Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle where open import Cubical.Cohomology.EilenbergMacLane.Base open import Cubical.Cohomology.EilenbergMacLane.Groups.Connected +open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2 open import Cubical.Homotopy.EilenbergMacLane.GroupStructure open import Cubical.Homotopy.EilenbergMacLane.Order2 @@ -20,6 +21,9 @@ open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Univalence open import Cubical.Foundations.Path open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function + +open import Cubical.Relation.Nullary open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order @@ -31,14 +35,16 @@ open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Instances.IntMod open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.DirectProduct open import Cubical.HITs.KleinBottle renaming (rec to KleinFun) open import Cubical.HITs.SetTruncation as ST open import Cubical.HITs.Truncation as TR open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.HITs.EilenbergMacLane1 hiding (elimProp ; elimSet) +open import Cubical.HITs.EilenbergMacLane1 renaming (elimProp to elimPropEM1 ; elimSet to elimSetEM1) open import Cubical.HITs.Susp +open import Cubical.HITs.RPn open AbGroupStr @@ -203,11 +209,6 @@ H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 = invGroupEquiv ℤ/2×ℤ/2≅H¹[K²,ℤ/2] ------ H²(K²,ℤ/2) ------ -- The equivalence Ω²K₂ ≃ ℤ/2 will be needed -Iso-Ω²K₂-ℤ/2 : Iso (fst ((Ω^ 2) (EM∙ ℤ/2 2))) (fst ℤ/2) -Iso-Ω²K₂-ℤ/2 = - compIso (congIso (invIso (Iso-EM-ΩEM+1 {G = ℤ/2} 1))) - (invIso (Iso-EM-ΩEM+1 {G = ℤ/2} 0)) - Ω²K₂→ℤ/2 : fst ((Ω^ 2) (EM∙ ℤ/2 2)) → fst ℤ/2 Ω²K₂→ℤ/2 = Iso.fun Iso-Ω²K₂-ℤ/2 @@ -303,24 +304,6 @@ H²K²→ℤ/2-pres0 = cong H²K²→ℤ/2 KleinFun-trivₕ ∙ H²K²→ℤ/2-rewrite (λ _ _ → ∣ north ∣ₕ) ∙ refl -Isoℤ/2-morph : {A : Type} (f : A ≃ fst ℤ/2) (0A : A) - → 0 ≡ fst f 0A → (_+'_ : A → A → A) (-m : A → A) - → (λ x → x) ≡ -m - → (e : IsAbGroup 0A _+'_ -m) - → IsGroupHom (AbGroup→Group (A , abgroupstr 0A _+'_ (λ x → -m x) e) .snd) - (fst f) ((ℤGroup/ 2) .snd) -Isoℤ/2-morph = - EquivJ (λ A f → (0A : A) → 0 ≡ fst f 0A → (_+'_ : A → A → A) (-m : A → A) - → (λ x → x) ≡ -m - → (e : IsAbGroup 0A _+'_ -m) - → IsGroupHom (AbGroup→Group (A , abgroupstr 0A _+'_ (λ x → -m x) e) .snd) - (fst f) ((ℤGroup/ 2) .snd)) - (J> λ _+'_ → J> - λ e → makeIsGroupHom (ℤ/2-elim (ℤ/2-elim (IsAbGroup.+IdR e fzero) - (IsAbGroup.+IdL e 1)) - (ℤ/2-elim (IsAbGroup.+IdR e 1) - (IsAbGroup.+InvR e 1)))) - H²[K²,ℤ/2]≅ℤ/2 : AbGroupEquiv (coHomGr 2 ℤ/2 KleinBottle) ℤ/2 fst H²[K²,ℤ/2]≅ℤ/2 = isoToEquiv is where @@ -355,7 +338,206 @@ snd (isContr-HⁿKleinBottle n G) = ST.elim (λ _ → isSetPathImplicit) (isConnectedEM (3 +ℕ n))))) _ _) _ _) refl p .fst)) -H³⁺ⁿK²≅0 : (n : ℕ) (G : AbGroup ℓ) + +---- With general coefficients + +-- H⁰(K²,G) : ? +H⁰[K²,G]≅G : (G : AbGroup ℓ) → AbGroupEquiv (coHomGr 0 G KleinBottle) G +H⁰[K²,G]≅G G = H⁰conn (∣ point ∣ + , (TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (elimProp (λ _ → isOfHLevelTrunc 2 _ _) refl))) G + +-- H¹(K²,G) +H¹[K²,G]≅G×H¹[RP²,G] : (G : AbGroup ℓ) + → AbGroupEquiv (coHomGr 1 G KleinBottle) (AbDirProd G (coHomGr 1 G RP²)) +fst (H¹[K²,G]≅G×H¹[RP²,G] G) = isoToEquiv mainIso + where + F→ : (KleinBottle → EM G 1) → (RP² → EM G 1) + F→ f point = f point + F→ f (line i) = f (line1 i) + F→ f (square i j) = help i (~ j) + where + help : cong f (sym line1) ≡ cong f line1 + help = + lUnit (cong f (sym line1)) + ∙ cong (_∙ cong f (sym line1)) + (sym (lCancel _)) + ∙ sym (assoc _ _ _) + ∙ cong (sym (cong f line2) ∙_) (isCommΩEM-base 0 _ _ _) + ∙ PathP→compPathL (λ i j → f (square j i)) + + F← : (g : fst G) → (RP² → EM G 1) → KleinBottle → EM G 1 + F← g f point = f point + F← g f (line1 i) = f (line i) + F← g f (line2 i) = EM→ΩEM+1-gen 0 (f point) g i + F← g f (square i j) = help j i + where + l1 = cong f line + l2 = EM→ΩEM+1-gen 0 (f point) g + + help : Square (sym l1) l1 l2 l2 + help = compPathL→PathP + (cong₂ _∙_ refl (isCommΩEM-base 0 _ _ _) + ∙∙ assoc _ _ _ + ∙∙ cong (_∙ sym l1) (lCancel l2) + ∙∙ sym (lUnit _) + ∙∙ sym (cong (cong f) square)) + + mainIso : Iso (coHom 1 G KleinBottle) (fst G × coHom 1 G RP²) + Iso.fun mainIso = + ST.rec (isSet× (is-set (snd G)) squash₂) + λ f → (ΩEM+1→EM-gen 0 _ (cong f line2)) , ∣ F→ f ∣₂ + Iso.inv mainIso = uncurry λ g → ST.map (F← g) + Iso.rightInv mainIso = uncurry λ g + → ST.elim (λ _ → isOfHLevelPath 2 (isSet× (is-set (snd G)) squash₂) _ _) + λ f → ΣPathP ((Iso.leftInv (Iso-EM-ΩEM+1-gen 0 (f point)) g) + , cong ∣_∣₂ (funExt (elimSetRP² (λ _ → hLevelEM G 1 _ _) + refl λ i j → f (line i)))) + Iso.leftInv mainIso = ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (funExt (elimSet (λ _ → hLevelEM G 1 _ _) + refl + (λ i j → f (line1 i)) + (flipSquare (Iso.rightInv (Iso-EM-ΩEM+1-gen 0 (f point)) (cong f line2))))) +snd (H¹[K²,G]≅G×H¹[RP²,G] G) = + makeIsGroupHom (ST.elim2 + (λ _ _ → isOfHLevelPath 2 lem _ _) + (ConnK².elim₁ (isConnectedEM 1) (λ _ → isPropΠ λ _ → lem _ _) embase + λ p1 q1 r1 + → ConnK².elim₁ (isConnectedEM 1) (λ _ → lem _ _) embase + λ p2 q2 r2 → ΣPathP (cong (ΩEM+1→EM 0) (cong₂+₁ q1 q2) + ∙ ΩEM+1→EM-hom 0 q1 q2 + , (cong ∣_∣₂ (funExt (elimSetRP² (λ _ → hLevelEM G 1 _ _) + refl + λ i j → (p1 i) +[ 1 ]ₖ (p2 i))))))) + where + lem = isSet× (is-set (snd G)) squash₂ + +H¹[K²,G]≅G×G[2] : (G : AbGroup ℓ) + → AbGroupEquiv (coHomGr 1 G KleinBottle) (AbDirProd G (G [ 2 ])) +H¹[K²,G]≅G×G[2] G = + compGroupEquiv (H¹[K²,G]≅G×H¹[RP²,G] G) + (GroupEquivDirProd idGroupEquiv (H¹[RP²,G]≅G[2] G)) + +-- H²(K²,G) +private + H²K²-helperFun₁ : (G : AbGroup ℓ) (n : ℕ) + → (Σ[ p ∈ Ω (EM∙ G (2 +ℕ n)) .fst ] + Σ[ q ∈ Ω (EM∙ G (2 +ℕ n)) .fst ] Square (sym p) p q q) + → ((Σ[ p ∈ Ω (EM∙ G (2 +ℕ n)) .fst ] + p ≡ sym p)) + H²K²-helperFun₁ G n (p , q , r) = (p , λ i j + → hcomp (λ k → λ {(i = i0) → rUnitₖ (suc (suc n)) (p j) k + ; (i = i1) → rUnitₖ (suc (suc n)) (p (~ j)) k + ; (j = i0) → rCancelₖ (suc (suc n)) (q (~ i)) k + ; (j = i1) → rCancelₖ (suc (suc n)) (q (~ i)) k}) + (r (~ i) j -[ suc (suc n) ]ₖ q (~ i) )) + + H²K²-helperIso₁ : (G : AbGroup ℓ) (n : ℕ) → + Iso ∥ (Σ[ p ∈ Ω (EM∙ G (2 +ℕ n)) .fst ] + Σ[ q ∈ Ω (EM∙ G (2 +ℕ n)) .fst ] Square (sym p) p q q) ∥₂ + ∥ (Σ[ x ∈ EM G (2 +ℕ n) ] (Σ[ p ∈ Ω (EM G (2 +ℕ n) , x) .fst ] + Σ[ q ∈ Ω (EM G (2 +ℕ n) , x) .fst ] Square (sym p) p q q)) ∥₂ + Iso.fun (H²K²-helperIso₁ G n) = ST.map (0ₖ _ ,_) + Iso.inv (H²K²-helperIso₁ G n) = ST.rec squash₂ + (uncurry (EM-raw'-elim G (2 +ℕ n) + (λ _ → isOfHLevelΠ (3 +ℕ n) λ _ → isOfHLevelPlus' {n = suc n} 2 squash₂) + (EM-raw'-trivElim G (suc n) (λ _ + → isOfHLevelΠ (2 +ℕ n) λ _ → isOfHLevelPlus' {n = n} 2 squash₂) + ∣_∣₂))) + Iso.rightInv (H²K²-helperIso₁ G n) = ST.elim (λ _ → isSetPathImplicit) + (uncurry (EM-raw'-elim G (2 +ℕ n) + (λ _ → isOfHLevelΠ (3 +ℕ n) + λ _ → isOfHLevelPlus' {n = suc (suc n)} 1 (squash₂ _ _)) + (EM-raw'-trivElim G (suc n) (λ _ → isOfHLevelΠ (2 +ℕ n) + λ _ → isOfHLevelPlus' {n = suc n} 1 (squash₂ _ _)) + λ _ → refl))) + Iso.leftInv (H²K²-helperIso₁ G n) = ST.elim (λ _ → isSetPathImplicit) λ _ → refl + + H²K²-helperIso₂ : (G : AbGroup ℓ) (n : ℕ) → + Iso ∥ (Σ[ p ∈ Ω (EM∙ G (2 +ℕ n)) .fst ] + Σ[ q ∈ Ω (EM∙ G (2 +ℕ n)) .fst ] Square (sym p) p q q) ∥₂ + ∥ ((Σ[ p ∈ Ω (EM∙ G (2 +ℕ n)) .fst ] + p ≡ sym p)) ∥₂ + Iso.fun (H²K²-helperIso₂ G n) = ST.map (H²K²-helperFun₁ G n) + Iso.inv (H²K²-helperIso₂ G n) = ST.map (λ p → fst p , refl , sym (snd p)) + Iso.rightInv (H²K²-helperIso₂ G n) = ST.elim (λ _ → isSetPathImplicit) λ {(p , r) + → TR.rec (isProp→isOfHLevelSuc n (squash₂ _ _)) + (λ { P → cong ∣_∣₂ (main p P r)}) + ((F p .fst))} + where + F = isConnectedPath (suc n) (isConnectedPath (suc (suc n)) + (isConnectedEM (suc (suc n))) _ _) refl + main : (p : Ω (EM∙ G (2 +ℕ n)) .fst) → refl ≡ p + → (r : p ≡ sym p) + → H²K²-helperFun₁ G n (p , refl , sym r) ≡ (p , r) + main = (J> λ r → (ΣPathP (refl , sym (rUnit _) + ∙ λ i j k → rUnitₖ (suc (suc n)) (r j k) i))) + Iso.leftInv (H²K²-helperIso₂ G n) = ST.elim (λ _ → isSetPathImplicit) + (uncurry λ p → TR.rec (isProp→isOfHLevelSuc n (isPropΠ (λ _ → squash₂ _ _ ))) + (λ P → uncurry λ q → TR.rec (isProp→isOfHLevelSuc n (isPropΠ (λ _ → squash₂ _ _ ))) + (λ Q x → cong ∣_∣₂ (main p P q Q x)) + (F q .fst)) + (F p .fst)) + where + F = isConnectedPath (suc n) (isConnectedPath (suc (suc n)) + (isConnectedEM (suc (suc n))) _ _) refl + main : (p : _) → refl ≡ p → + (q : _) → refl ≡ q → + (x : Square (sym p) p q q) → + (fst (H²K²-helperFun₁ G n (p , q , x)) + , refl + , (λ i → snd (H²K²-helperFun₁ G n (p , q , x)) (~ i))) + ≡ (p , q , x) + main = J> (J> λ p → ΣPathP (refl , (ΣPathP (refl + , cong sym (sym (rUnit _)) + ∙ λ i j k → rUnitₖ (suc (suc n)) (p j k) i)))) + + Iso-H²⁺ⁿ[K²,G]-H²⁺ⁿ[RP²,G] : (G : AbGroup ℓ) (n : ℕ) + → Iso (coHom (2 +ℕ n) G KleinBottle) + (coHom (2 +ℕ n) G RP²) + Iso-H²⁺ⁿ[K²,G]-H²⁺ⁿ[RP²,G] G n = + compIso + (setTruncIso + (compIso K²FunCharacIso + (Σ-cong-iso-snd λ p → Σ-cong-iso-snd λ q + → Σ-cong-iso-snd λ r → equivToIso flipSquareEquiv))) + (compIso (invIso (H²K²-helperIso₁ G n)) + (compIso (H²K²-helperIso₂ G n) + (compIso (invIso (killFirstCompIsoGen G n)) + (setTruncIso (invIso RP²FunCharacIso))))) + +H²[RP²,G]≅H²[K²,G] : (G : AbGroup ℓ) + → AbGroupEquiv (coHomGr 2 G RP²) (coHomGr 2 G KleinBottle) +fst (H²[RP²,G]≅H²[K²,G] G) = isoToEquiv (invIso (Iso-H²⁺ⁿ[K²,G]-H²⁺ⁿ[RP²,G] G 0)) +snd (H²[RP²,G]≅H²[K²,G] G) = + makeIsGroupHom (ST.elim (λ _ → isSetΠ λ _ → isSetPathImplicit) + (RP²Conn.elim₂ (isConnectedEM 2) + (λ _ → isPropΠ λ _ → squash₂ _ _) + (0ₖ 2) + λ p → ST.elim (λ _ → isSetPathImplicit) + (RP²Conn.elim₂ (isConnectedEM 2) + (λ _ → squash₂ _ _) + (0ₖ 2) λ q + → cong ∣_∣₂ (funExt λ { point → refl + ; (line1 i) → refl + ; (line2 i) → refl + ; (square i i₁) → refl})))) + +H²⁺ⁿ[K²,G]≅H²⁺ⁿ[RP²,G] : (G : AbGroup ℓ) (n : ℕ) + → AbGroupEquiv (coHomGr (2 +ℕ n) G KleinBottle) + (coHomGr (2 +ℕ n) G RP²) +fst (H²⁺ⁿ[K²,G]≅H²⁺ⁿ[RP²,G] G n) = isoToEquiv (Iso-H²⁺ⁿ[K²,G]-H²⁺ⁿ[RP²,G] G n) +snd (H²⁺ⁿ[K²,G]≅H²⁺ⁿ[RP²,G] G zero) = invGroupEquiv (H²[RP²,G]≅H²[K²,G] G) .snd +snd (H²⁺ⁿ[K²,G]≅H²⁺ⁿ[RP²,G] G (suc n)) = + makeIsGroupHom λ _ _ → isContr→isProp + (isOfHLevelRetractFromIso 0 (equivToIso (fst (H³⁺ⁿ[RP²,G]≅0 G n))) + isContrUnit*) _ _ + +H²[K²,G]≅G/2 : (G : AbGroup ℓ) + → AbGroupEquiv (coHomGr 2 G KleinBottle) + (G /^ 2) +H²[K²,G]≅G/2 G = compGroupEquiv (H²⁺ⁿ[K²,G]≅H²⁺ⁿ[RP²,G] G 0) (H²[RP²,G]≅G/2 G) + +H³⁺ⁿ[K²,G]≅0 : (n : ℕ) (G : AbGroup ℓ) → AbGroupEquiv (coHomGr (3 +ℕ n) G KleinBottle) (trivialAbGroup {ℓ}) -fst (H³⁺ⁿK²≅0 n G) = isContr→Equiv (isContr-HⁿKleinBottle n G) isContrUnit* -snd (H³⁺ⁿK²≅0 n G) = makeIsGroupHom λ _ _ → refl +H³⁺ⁿ[K²,G]≅0 n G = compGroupEquiv (H²⁺ⁿ[K²,G]≅H²⁺ⁿ[RP²,G] G (suc n)) (H³⁺ⁿ[RP²,G]≅0 G n) diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda index c8f381b49a..06c2069d69 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda @@ -3,7 +3,7 @@ module Cubical.Cohomology.EilenbergMacLane.Groups.RP2 where open import Cubical.Cohomology.EilenbergMacLane.Base -open import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle +open import Cubical.Cohomology.EilenbergMacLane.Groups.Connected open import Cubical.Homotopy.EilenbergMacLane.GroupStructure open import Cubical.Homotopy.EilenbergMacLane.Order2 @@ -17,21 +17,32 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Path open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Unit +open import Cubical.Data.Sigma open import Cubical.Data.Fin open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties +open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Instances.IntMod open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group.Subgroup +open import Cubical.Algebra.Group.QuotientGroup +open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.SetTruncation as ST open import Cubical.HITs.Truncation as TR -open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.HITs.EilenbergMacLane1 as EM1 open import Cubical.HITs.RPn +open import Cubical.HITs.SetQuotients as SQ renaming ([_] to ⟦_⟧) open AbGroupStr @@ -88,6 +99,31 @@ module RP²Conn {B : (RP² → A) → Type ℓ} where (isConnectedPath 2 (isConnectedPath 3 conA _ _) _ _) refl sq .fst) +--- lemmas +Isoℤ/2-morph : {A : Type} (f : A ≃ fst ℤ/2) (0A : A) + → 0 ≡ fst f 0A → (_+'_ : A → A → A) (-m : A → A) + → (λ x → x) ≡ -m + → (e : IsAbGroup 0A _+'_ -m) + → IsGroupHom (AbGroup→Group (A , abgroupstr 0A _+'_ (λ x → -m x) e) .snd) + (fst f) ((ℤGroup/ 2) .snd) +Isoℤ/2-morph = + EquivJ (λ A f → (0A : A) → 0 ≡ fst f 0A → (_+'_ : A → A → A) (-m : A → A) + → (λ x → x) ≡ -m + → (e : IsAbGroup 0A _+'_ -m) + → IsGroupHom (AbGroup→Group (A , abgroupstr 0A _+'_ (λ x → -m x) e) .snd) + (fst f) ((ℤGroup/ 2) .snd)) + (J> λ _+'_ → J> + λ e → makeIsGroupHom (ℤ/2-elim (ℤ/2-elim (IsAbGroup.+IdR e fzero) + (IsAbGroup.+IdL e 1)) + (ℤ/2-elim (IsAbGroup.+IdR e 1) + (IsAbGroup.+InvR e 1)))) + +Iso-Ω²K₂-ℤ/2 : Iso (fst ((Ω^ 2) (EM∙ ℤ/2 2))) (fst ℤ/2) +Iso-Ω²K₂-ℤ/2 = + compIso (congIso (invIso (Iso-EM-ΩEM+1 {G = ℤ/2} 1))) + (invIso (Iso-EM-ΩEM+1 {G = ℤ/2} 0)) + + ----- H¹(RP²,ℤ/2) ------ H¹[RP²,ℤ/2]→ℤ/2 : coHom 1 ℤ/2 RP² → fst ℤ/2 @@ -210,3 +246,249 @@ fst (H³⁺ⁿ[RP²,G]≅G n G) = (cong ∣_∣₂ (characRP²Fun (λ _ → 0ₖ (3 +ℕ n))))))) isContrUnit*) snd (H³⁺ⁿ[RP²,G]≅G n G) = makeIsGroupHom λ _ _ → refl + + +---- H¹(RP², G) ----- +H⁰[RP²,G]≅G : ∀ {ℓ} (G : AbGroup ℓ) + → AbGroupEquiv (coHomGr 0 G RP²) G +H⁰[RP²,G]≅G G = + H⁰conn (∣ point ∣ₕ , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (elimPropRP² (λ _ → isOfHLevelTrunc 2 _ _) refl)) G + +----- H¹(RP², G) ------ +module _ {ℓ : Level} (G : AbGroup ℓ) where + private + G[2]path : (x : (G [ 2 ]) .fst) → _+_ (snd G) (fst x) (fst x) ≡ 0g (snd G) + G[2]path (x , p) = cong (_+_ (snd G) x) (sym (+IdR (snd G) x)) ∙ p + + G[2]path' : (x : (G [ 2 ]) .fst) → fst x ≡ -_ (snd G) (fst x) + G[2]path' x = (sym (+IdR (snd G) _) + ∙ cong (_+_ (snd G) (fst x)) (sym (+InvR (snd G) (fst x))) + ∙ (+Assoc (snd G) _ _ _)) + ∙∙ cong (λ w → _+_ (snd G) w (-_ (snd G) (fst x))) (G[2]path x) + ∙∙ +IdL (snd G) _ + + H¹[RP²,G]→G[2] : coHom 1 G RP² → (G [ 2 ]) .fst + H¹[RP²,G]→G[2] = + ST.rec (is-set (snd (G [ 2 ]))) + λ f → ΩEM+1→EM-gen {G = G} 0 (f point) (cong f line) + , cong (_+_ (snd G) (ΩEM+1→EM-gen 0 (f point) (cong f line))) + (+IdR (snd G) _) + ∙ sym (ΩEM+1→EM-gen-hom {G = G} 0 (f point) (cong f line) (cong f line)) + ∙ cong (ΩEM+1→EM-gen 0 (f point)) + (sym (cong-∙ f line line) + ∙ cong (cong f) (cong (line ∙_) square ∙ rCancel line) ) + ∙ ΩEM+1→EM-gen-refl 0 (f point) + + G[2]→H¹[RP²,G] : (G [ 2 ]) .fst → coHom 1 G RP² + G[2]→H¹[RP²,G] g = + ∣ (λ { point → embase + ; (line i) → emloop (fst g) i + ; (square i j) → + (cong emloop (G[2]path' g) + ∙ emloop-inv (AbGroup→Group G) (fst g)) i j }) ∣₂ + + private + G[2]≅H¹[RP²,G] : AbGroupEquiv (G [ 2 ]) (coHomGr 1 G RP²) + fst G[2]≅H¹[RP²,G] = isoToEquiv (invIso is) + where + is : Iso _ _ + Iso.fun is = H¹[RP²,G]→G[2] + Iso.inv is = G[2]→H¹[RP²,G] + Iso.rightInv is (g , p) = + Σ≡Prop (λ _ → is-set (snd G) _ _) + (Iso.leftInv (Iso-EM-ΩEM+1 0) g) + Iso.leftInv is = ST.elim (λ _ → isSetPathImplicit) + (RP²Conn.elim₁ (isConnectedEM 1) (λ _ → squash₂ _ _) + embase + λ p q → cong ∣_∣₂ + (funExt (elimSetRP² (λ _ → emsquash _ _) + refl + (flipSquare (Iso.rightInv (Iso-EM-ΩEM+1 {G = G} 0) p))))) + snd G[2]≅H¹[RP²,G] = makeIsGroupHom + λ x y → cong ∣_∣₂ + (funExt (elimSetRP² (λ _ → emsquash _ _) refl + (flipSquare + (EM→ΩEM+1-hom {G = G} 0 (fst x) (fst y) + ∙ sym (cong₂+₁ (emloop (fst x)) (emloop (fst y))))))) + + + H¹[RP²,G]≅G[2] : AbGroupEquiv (coHomGr 1 G RP²) (G [ 2 ]) + H¹[RP²,G]≅G[2] = invGroupEquiv G[2]≅H¹[RP²,G] + +----- H²(RP², G) ------ +module _ (G : AbGroup ℓ) where + private + ΩEM+1→EM-sym' : (p : fst (Ω (EM∙ G 2))) + → ΩEM+1→EM 1 (sym p) ≡ (-ₖ ΩEM+1→EM 1 p) + ΩEM+1→EM-sym' p = + ΩEM+1→EM-sym (suc zero) p + ∙ (sym (rUnitₖ 1 ((-ₖ ΩEM+1→EM 1 p))) + ∙∙ cong ((-ₖ ΩEM+1→EM 1 p) +ₖ_) (sym (ΩEM+1→EM-sym (suc zero) refl)) + ∙∙ rUnitₖ 1 ((-ₖ ΩEM+1→EM 1 p))) + + ΩEM+1→EM-sym'-refl : ΩEM+1→EM-sym' refl ≡ refl + ΩEM+1→EM-sym'-refl = + (λ i → ΩEM+1→EM-sym (suc zero) refl ∙ rUnit (sym (ΩEM+1→EM-sym (suc zero) refl)) (~ i)) + ∙ rCancel _ + + + G/2-ord2 : (x : fst (G /^ 2)) → (- snd (G /^ 2)) x ≡ x + G/2-ord2 = SQ.elimProp (λ _ → SQ.squash/ _ _) + λ a → eq/ _ _ ∣ snd G .-_ a + , cong (snd G ._+_ (snd G .-_ a)) (+IdR (snd G) _) ∣₁ + + open EM2 (G /^ 2) G/2-ord2 + killFirstCompIsoGen : (n : ℕ) + → Iso ∥ (Σ[ x ∈ EM G (2 +ℕ n) ] (Σ[ p ∈ x ≡ x ] p ≡ sym p)) ∥₂ + ∥ (Σ[ p ∈ Ω (EM∙ G (2 +ℕ n)) .fst ] p ≡ sym p) ∥₂ + Iso.fun (killFirstCompIsoGen n) = + ST.rec squash₂ + (uncurry (TR.elim (λ _ → isOfHLevelPlus {n = 2 +ℕ n} 2 + (isOfHLevelPlus' {n = n} 2 (isSetΠ λ _ → squash₂))) + (EM-raw'-trivElim G (suc n) + (λ _ → isOfHLevelΠ (suc (suc n)) λ _ → isOfHLevelPlus' {n = n} 2 squash₂) ∣_∣₂))) + Iso.inv (killFirstCompIsoGen n) = ST.map (0ₖ (2 +ℕ n) ,_) + Iso.rightInv (killFirstCompIsoGen n) = ST.elim (λ _ → isSetPathImplicit) λ _ → refl + Iso.leftInv (killFirstCompIsoGen n) = ST.elim (λ _ → isSetPathImplicit) + (uncurry (TR.elim (λ _ → isProp→isOfHLevelSuc (3 +ℕ n) (isPropΠ (λ _ → squash₂ _ _))) + (EM-raw'-trivElim G (suc n) (λ _ → isProp→isOfHLevelSuc (suc n) + (isPropΠ λ _ → squash₂ _ _ )) + λ _ → refl))) + + Iso₁ : Iso (coHom 2 G RP²) ∥ (Σ[ p ∈ Ω (EM∙ G 2) .fst ] p ≡ sym p) ∥₂ + Iso₁ = compIso (setTruncIso RP²FunCharacIso) (killFirstCompIsoGen 0) + + Iso₂ : Iso ((Σ[ p ∈ Ω (EM∙ G 2) .fst ] p ≡ sym p)) (Σ[ p ∈ EM G 1 ] p ≡ -ₖ p) + Iso₂ = (Σ-cong-iso (invIso (Iso-EM-ΩEM+1 1)) + λ x → compIso (congIso (invIso (Iso-EM-ΩEM+1 1))) + (equivToIso (compPathrEquiv (ΩEM+1→EM-sym' x)))) + + RP²Fun+ : (p q : (Ω^ 2) (EM∙ G 2) .fst) (x : _) + → RP²Fun _ _ p x +[ 2 ]ₖ RP²Fun _ _ q x ≡ RP²Fun _ _ (p ∙ q) x + RP²Fun+ p q point = refl + RP²Fun+ p q (line i) j = 0ₖ 2 + RP²Fun+ p q (square i j) k = + hcomp (λ r → λ {(i = i0) → lem3 k j r + ; (i = i1) → lem2 k (~ r) j + ; (j = i0) → p (i ∨ ~ k) r + ; (j = i1) → p (~ i ∧ k) (~ r) + ; (k = i0) → cong₂+₂ 0 (p i) (q i) (~ r) j + ; (k = i1) → lem1 _ refl p (sym q) r i j}) + ((p i ∙ q i) j) + where + lem1 : ∀ {ℓ} {A : Type ℓ} (x : A) (p : x ≡ x) (P Q : refl ≡ p) + → Cube (λ i j → (P i ∙ Q (~ i)) j) (P ∙ sym Q) + (λ k → compPath-filler refl p (~ k)) + (λ k → compPath-filler' p refl (~ k)) + (λ i j → P j i) λ i j → P (~ j) (~ i) + lem1 x = J> λ Q → (λ i j k → lUnit (Q (~ j)) (~ i) k) ▷ lUnit (sym Q) + + lem2 : cong₂+₂ {G = G} 0 refl refl ≡ rUnit refl + lem2 = sym (rUnit _) + + lem3 : Cube (λ j r → cong₂+₂ {G = G} 0 refl refl (~ r) j) + (λ j r → compPath-filler (refl {x = 0ₖ 2}) refl (~ r) j) + (λ k r → p (~ k) r) (λ k r → p k (~ r)) + (λ j i → (rUnit (refl {x = 0ₖ 2}) i1) i) + (λ _ _ → 0ₖ 2) + lem3 = (λ i j r → lem2 i (~ r) j) + ◁ (λ k j r → + hcomp (λ i → λ {(j = i0) → p (~ k) r + ; (j = i1) → p k (~ r) + ; (r = i0) → rUnit (λ _ → 0ₖ 2) i j + ; (r = i1) → 0ₖ 2 + ; (k = i0) → rUnit (λ _ → 0ₖ 2) (i ∧ ~ r) j + ; (k = i1) → rUnit (λ _ → 0ₖ 2) (i ∧ ~ r) j}) + (((sym≡flipSquare p) ∙ flipSquare≡cong-sym p) j k r)) + + Iso₂Iso₂' : (p : (Ω^ 2) (EM∙ G 2) .fst) + → ST.map (Iso.fun Iso₂) (Iso.fun Iso₁ ∣ RP²Fun _ _ p ∣₂) + ≡ ∣ embase , (cong (ΩEM+1→EM 1) p) ∣₂ + Iso₂Iso₂' p = cong ∣_∣₂ (ΣPathP (refl + , (λ i → cong (ΩEM+1→EM {G = G} 1) p ∙ ΩEM+1→EM-sym'-refl i) + ∙ sym (rUnit (cong (ΩEM+1→EM 1) p)))) + + Iso₂Iso₂+ : (p q : (Ω^ 2) (EM∙ G 2) .fst) + → ST.map (Iso.fun Iso₂) (Iso.fun Iso₁ (∣ RP²Fun _ _ p ∣₂ +ₕ ∣ RP²Fun _ _ q ∣₂)) + ≡ ∣ embase , cong (ΩEM+1→EM 1) p ∙ cong (ΩEM+1→EM 1) q ∣₂ + Iso₂Iso₂+ p q = cong (ST.map (Iso.fun Iso₂) ∘ Iso.fun Iso₁) + (cong ∣_∣₂ (funExt (RP²Fun+ p q))) + ∙ Iso₂Iso₂' (p ∙ q) + ∙ cong ∣_∣₂ (ΣPathP (refl , cong-∙ (ΩEM+1→EM 1) p q)) + + Iso₃→ : (p : EM G 1) → (p ≡ -[ 1 ]ₖ p) → fst (G /^ 2) + Iso₃→ = EM1.elimSet (AbGroup→Group G) (λ _ → isSetΠ λ _ → squash/) + (λ p → ⟦ ΩEM+1→EM 0 p ⟧) + λ g → toPathP (funExt λ q + → cong ⟦_⟧ (transportRefl _ + ∙ cong (ΩEM+1→EM 0) + ((λ j → transp (λ i → Path (EM G 1) (emloop g (~ i ∧ ~ j)) (emloop g (i ∨ j))) j + (doubleCompPath-filler (emloop g) q (emloop g) j)) + ∙ doubleCompPath≡compPath _ _ _ + ∙ cong (emloop g ∙_) (isCommΩEM {G = G} 0 q (emloop g)) + ∙ assoc (emloop g) (emloop g) q) + ∙ ΩEM+1→EM-hom 0 (emloop g ∙ emloop g) q) + ∙ eq/ _ _ ∣ g + , cong (snd G ._+_ g) (+IdR (snd G) _) + ∙ sym (+IdR (snd G) (snd G ._+_ g g)) + ∙ cong₂ (snd G ._+_) + (sym (Iso.leftInv (Iso-EM-ΩEM+1 0) _) + ∙ cong (ΩEM+1→EM 0) (emloop-comp _ g g)) + (sym (+InvR (snd G) (ΩEM+1→EM 0 q))) + ∙ +Assoc (snd G) _ _ _ ∣₁) + + Iso₃ : Iso ∥ (Σ[ p ∈ EM G 1 ] p ≡ -ₖ p) ∥₂ (fst (G /^ 2)) + Iso.fun Iso₃ = ST.rec squash/ (uncurry Iso₃→) + Iso.inv Iso₃ = + SQ.elim (λ _ → squash₂) (λ a → ∣ embase , emloop a ∣₂) + λ a b → PT.rec (squash₂ _ _) + (uncurry λ x y → cong ∣_∣₂ (ΣPathP ((emloop x) + , compPathL→PathP + (cong₂ _∙_ (sym (emloop-inv _ x)) + (cong (emloop a ∙_) (sym (emloop-inv _ x)) + ∙ sym (EM→ΩEM+1-hom {G = G} 0 a (AbGroupStr.-_ (snd G) x))) + ∙ sym (EM→ΩEM+1-hom {G = G} 0 _ _) + ∙ cong emloop (cong (_+_ (snd G) (AbGroupStr.-_ (snd G) x)) + (+Comm (snd G) _ _) + ∙ +Assoc (snd G) _ _ _ + ∙ cong (λ w → _+_ (snd G) w a) + (sym (GroupTheory.invDistr (AbGroup→Group G) x x) + ∙ cong (AbGroupStr.-_ (snd G)) + (cong (snd G ._+_ x ) (sym (+IdR (snd G) _)) ∙ y) + ∙ GroupTheory.invDistr (AbGroup→Group G) a _ + ∙ cong (λ w → _+_ (snd G) w (AbGroupStr.-_ (snd G) a)) + (GroupTheory.invInv (AbGroup→Group G) b)) + ∙ (sym (+Assoc (snd G) _ _ _) + ∙ cong (_+_ (snd G) b) (+InvL (snd G) a)) + ∙ +IdR (snd G) b))))) + Iso.rightInv Iso₃ = SQ.elimProp (λ _ → squash/ _ _) + λ a → cong ⟦_⟧ (Iso.leftInv (Iso-EM-ΩEM+1 0) a) + Iso.leftInv Iso₃ = ST.elim (λ _ → isSetPathImplicit) + (uncurry (EM1.elimProp (AbGroup→Group G) (λ _ → isPropΠ λ _ → squash₂ _ _) + λ p → cong ∣_∣₂ (ΣPathP (refl , (Iso.rightInv (Iso-EM-ΩEM+1 0) p))))) + + H²[RP²,G]≅G/2 : AbGroupEquiv (coHomGr 2 G RP²) (G /^ 2) + fst H²[RP²,G]≅G/2 = isoToEquiv is + where + is : Iso _ _ + is = compIso (compIso Iso₁ (setTruncIso Iso₂)) Iso₃ + snd H²[RP²,G]≅G/2 = + makeIsGroupHom (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash/ _ _) + (RP²Conn.elim₂ (isConnectedEM 2) (λ _ → isPropΠ λ _ → squash/ _ _) + (0ₖ 2) λ p → RP²Conn.elim₂ (isConnectedEM 2) (λ _ → squash/ _ _) + (0ₖ 2) λ q → cong (Iso.fun Iso₃) (Iso₂Iso₂+ p q) + ∙ cong ⟦_⟧ (ΩEM+1→EM-hom 0 (cong (ΩEM+1→EM 1) p) + (cong (ΩEM+1→EM 1) q)) + ∙ cong₂ ((G /^ 2) .snd ._+_) + (cong (Iso.fun Iso₃) (sym (Iso₂Iso₂' p))) + (cong (Iso.fun Iso₃) (sym (Iso₂Iso₂' q))))) + + H³⁺ⁿ[RP²,G]≅0 : (n : ℕ) → AbGroupEquiv (coHomGr (3 +ℕ n) G RP²) (trivialAbGroup {ℓ}) + fst (H³⁺ⁿ[RP²,G]≅0 n) = isContr→≃Unit* (0ₕ _ + , ST.elim (λ _ → isSetPathImplicit) + (RP²Conn.elim₃ (isConnectedSubtr 4 n + (subst (λ k → isConnected k (EM G (3 +ℕ n))) + (+-comm 4 n) (isConnectedEM (3 +ℕ n)))) + (λ _ → squash₂ _ _) (0ₖ _) refl)) + snd (H³⁺ⁿ[RP²,G]≅0 n) = makeIsGroupHom λ _ _ → refl diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/Torus.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/Torus.agda new file mode 100644 index 0000000000..1788faa995 --- /dev/null +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/Torus.agda @@ -0,0 +1,238 @@ +{-# OPTIONS --safe --lossy-unification #-} + +module Cubical.Cohomology.EilenbergMacLane.Groups.Torus where + +open import Cubical.Cohomology.EilenbergMacLane.Base +open import Cubical.Cohomology.EilenbergMacLane.Groups.Connected +open import Cubical.Cohomology.EilenbergMacLane.Groups.Sn + +open import Cubical.Homotopy.EilenbergMacLane.Base +open import Cubical.Homotopy.EilenbergMacLane.GroupStructure +open import Cubical.Homotopy.EilenbergMacLane.Properties +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Loopspace + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv + + + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.AbGroup.Instances.DirectProduct + + +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as Trunc +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 renaming (rec to S¹fun ; elim to S¹elim) + + +open import Cubical.Data.Unit + +-- Preliminary lemmas +elimFunT² : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) + (p q : typ (Ω (EM∙ G (suc n)))) + → Square q q p p + → S¹ × S¹ → EM G (suc n) +elimFunT² n p q P (base , base) = 0ₖ (suc n) +elimFunT² n p q P (base , loop i) = q i +elimFunT² n p q P (loop i , base) = p i +elimFunT² n p q P (loop i , loop j) = P i j + +elimFunT²' : ∀ {ℓ} {G : AbGroup ℓ} + (n : ℕ) → Square (refl {ℓ} {EM G (suc n)} {0ₖ (suc n)}) refl refl refl + → S¹ × S¹ → EM G (suc n) +elimFunT²' n P (x , base) = 0ₖ (suc n) +elimFunT²' n P (base , loop j) = 0ₖ (suc n) +elimFunT²' n P (loop i , loop j) = P i j + +elimFunT²'≡elimFunT² : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) + → (P : _) → elimFunT²' {G = G} n P ≡ elimFunT² n refl refl P +elimFunT²'≡elimFunT² n P i (base , base) = 0ₖ (suc n) +elimFunT²'≡elimFunT² n P i (base , loop k) = 0ₖ (suc n) +elimFunT²'≡elimFunT² n P i (loop j , base) = 0ₖ (suc n) +elimFunT²'≡elimFunT² n P i (loop j , loop k) = P j k + +coHomPointedElimT² : ∀ {ℓ ℓ'} {G : AbGroup ℓ} (n : ℕ) {B : coHom (suc n) G (S¹ × S¹) → Type ℓ'} + → ((x : coHom (suc n) G (S¹ × S¹)) → isProp (B x)) + → ((p q : _) (P : _) → B ∣ elimFunT² n p q P ∣₂) + → (x : coHom (suc n) G (S¹ × S¹)) → B x +coHomPointedElimT² {G = G} n {B = B} isprop indp = + coHomPointedElim n (base , base) isprop + λ f fId → subst B (cong ∣_∣₂ (funExt + (λ { (base , base) → sym fId + ; (base , loop i) j → helper f fId i1 i (~ j) + ; (loop i , base) j → helper f fId i i1 (~ j) + ; (loop i , loop j) k → helper f fId i j (~ k)}))) + (indp (λ i → helper f fId i i1 i1) + (λ i → helper f fId i1 i i1) + λ i j → helper f fId i j i1) + where + helper : (f : S¹ × S¹ → EM G (suc n)) → f (base , base) ≡ 0ₖ (suc n) + → I → I → I → EM G (suc n) + helper f fId i j k = + hfill (λ k → λ {(i = i0) → doubleCompPath-filler (sym fId) (cong f (λ i → (base , loop i))) fId k j + ; (i = i1) → doubleCompPath-filler (sym fId) (cong f (λ i → (base , loop i))) fId k j + ; (j = i0) → doubleCompPath-filler (sym fId) (cong f (λ i → (loop i , base))) fId k i + ; (j = i1) → doubleCompPath-filler (sym fId) (cong f (λ i → (loop i , base))) fId k i}) + (inS (f ((loop i) , (loop j)))) + k + +private + lem : ∀ {ℓ ℓ'} {G : AbGroup ℓ} (n : ℕ) {B : coHom (2 + n) G (S¹ × S¹) → Type ℓ'} + → ((P : _) → B ∣ elimFunT² (suc n) refl refl P ∣₂) + → (p : _) → (refl ≡ p) + → (q : _) → (refl ≡ q) + → (P : _) + → B ∣ elimFunT² (suc n) p q P ∣₂ + lem n {B = B} elimP p = + J (λ p _ → (q : _) → (refl ≡ q) + → (P : _) + → B ∣ elimFunT² (suc n) p q P ∣₂) + λ q → + J (λ q _ → (P : _) → B ∣ elimFunT² (suc n) refl q P ∣₂) + elimP + +{- When working with Hⁿ(T²) , n ≥ 2, we are, in the case described above, allowed to assume that any f : Hⁿ(T²) is + elimFunT² n refl refl P -} +coHomPointedElimT²' : ∀ {ℓ ℓ'} {G : AbGroup ℓ} (n : ℕ) {B : coHom (2 + n) G (S¹ × S¹) → Type ℓ'} + → ((x : coHom (2 + n) G (S¹ × S¹)) → isProp (B x)) + → ((P : _) → B ∣ elimFunT² (suc n) refl refl P ∣₂) + → (x : coHom (2 + n) G (S¹ × S¹)) → B x +coHomPointedElimT²' n {B = B} prop ind = + coHomPointedElimT² (suc n) prop + λ p q P → Trunc.rec (isProp→isOfHLevelSuc n (prop _)) + (λ p-refl → Trunc.rec (isProp→isOfHLevelSuc n (prop _)) + (λ q-refl → lem n {B = B} ind p (sym p-refl) q (sym q-refl) P) + (isConnectedPath (suc n) + (isConnectedPath (suc (suc n)) (isConnectedEM (suc (suc n))) _ _) q refl .fst)) + (isConnectedPath (suc n) + (isConnectedPath (suc (suc n)) (isConnectedEM (suc (suc n))) _ _) p refl .fst) + +{- A slight variation of the above which gives definitional equalities for all points (x , base) -} +private + coHomPointedElimT²'' : ∀ {ℓ ℓ'} {G : AbGroup ℓ} (n : ℕ) {B : coHom (2 + n) G (S¹ × S¹) → Type ℓ'} + → ((x : coHom (2 + n) G (S¹ × S¹)) → isProp (B x)) + → ((P : _) → B ∣ elimFunT²' (suc n) P ∣₂) + → (x : coHom (2 + n) G (S¹ × S¹)) → B x + coHomPointedElimT²'' n {B = B} prop ind = + coHomPointedElimT²' n prop λ P → subst (λ x → B ∣ x ∣₂) + (elimFunT²'≡elimFunT² (suc n) P) (ind P) + +--------- H⁰(T²) ------------ +H⁰[T²,G]≅G : ∀ {ℓ} (G : AbGroup ℓ) → AbGroupEquiv (coHomGr 0 G (S₊ 1 × S₊ 1)) G +H⁰[T²,G]≅G = H⁰conn (∣ base , base ∣ + , Trunc.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + λ ab → PT.rec2 (isOfHLevelTrunc 2 _ _) + (λ p q i → ∣ p i , q i ∣ₕ ) + (isConnectedS¹ (fst ab)) + (isConnectedS¹ (snd ab))) + +--------- H¹(T²) ------------------------------- + +H¹[T²,G]≅G×G : ∀ {ℓ} (G : AbGroup ℓ) + → AbGroupEquiv (coHomGr 1 G (S¹ × S¹)) (AbDirProd G G) +H¹[T²,G]≅G×G G = + (compGroupEquiv + (invGroupEquiv (GroupIso→GroupEquiv theIso)) + (GroupEquivDirProd + (H¹[S¹,G]≅G G) + (H¹[S¹,G]≅G G))) + where + _+1_ : EM G 1 → EM G 1 → EM G 1 + _+1_ = _+ₖ_ {G = G} {n = 1} + + typIso : Iso _ _ + Iso.fun typIso = ST.rec (isSet× squash₂ squash₂) + λ f → ∣ (λ x → f (x , base)) ∣₂ , ∣ (λ x → f (base , x)) ∣₂ + Iso.inv typIso = uncurry (ST.rec2 squash₂ + λ f g → ∣ (λ x → f (fst x) +1 g (snd x)) ∣₂) + Iso.rightInv typIso = + uncurry (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSet× squash₂ squash₂) _ _) + (S¹-connElim (isConnectedEM {G' = G} 1) + (λ _ → isPropΠ λ _ → isSet× squash₂ squash₂ _ _) + (0ₖ {G = G} 1) + λ p → S¹-connElim (isConnectedEM {G' = G} 1) + (λ _ → isSet× squash₂ squash₂ _ _) + (0ₖ {G = G} 1) + λ q i → ∣ (λ l → rUnitₖ {G = G} 1 (S¹fun (0ₖ {G = G} 1) p l) i) ∣₂ + , ∣ S¹fun (0ₖ {G = G} 1) q ∣₂)) + Iso.leftInv typIso = coHomPointedElimT² {G = G} 0 (λ _ → squash₂ _ _) + λ p q r → cong ∣_∣₂ + (funExt (uncurry (S¹elim _ (λ _ → refl) + (funExt (toPropElim (λ _ → isOfHLevelPathP' 1 (hLevelEM G 1 _ _) _ _ ) + λ i j → rUnitₖ {G = G} 1 (p i) j))))) + + theIso : AbGroupIso (AbDirProd (coHomGr 1 G S¹) (coHomGr 1 G S¹)) (coHomGr 1 G (S¹ × S¹)) + fst theIso = invIso typIso + snd theIso = makeIsGroupHom + (uncurry (ST.elim2 (λ _ _ → isSetΠ λ _ → isSetPathImplicit) + λ f1 g1 → uncurry (ST.elim2 (λ _ _ → isSetPathImplicit) + λ f2 g2 → cong ∣_∣₂ (funExt λ {(x , y) + → move4 (f1 x) (f2 x) (g1 y) (g2 y) _+1_ (assocₖ {G = G} 1) (commₖ {G = G} 1)})))) + +----------------------- H²(T²) ------------------------------ + +H²[T²,G]≅G : ∀ {ℓ} (G : AbGroup ℓ) → AbGroupEquiv (coHomGr 2 G (S¹ × S¹)) G +H²[T²,G]≅G G = compGroupEquiv (GroupIso→GroupEquiv theIso) (Hⁿ[Sⁿ,G]≅G G 0) + where + _+1_ : EM G 1 → EM G 1 → EM G 1 + _+1_ = _+ₖ_ {G = G} {n = 1} + + _+2_ : EM G 2 → EM G 2 → EM G 2 + _+2_ = _+ₖ_ {G = G} {n = 2} + + + is1 : Iso (S¹ → EM G 2) (EM G 2 × EM G 1) + is1 = compIso IsoFunSpaceS¹ + (Σ-cong-iso-snd λ t → invIso (Iso-EM-ΩEM+1-gen {G = G} 1 t)) + + is : Iso _ _ + is = compIso (setTruncIso (compIso curryIso + (compIso (codomainIso is1) toProdIso))) + (compIso setTruncOfProdIso + (compIso (prodIso + (equivToIso (fst (Hᵐ⁺ⁿ[Sⁿ,G]≅0 G 0 0))) idIso) + lUnit*×Iso)) + + theIso : AbGroupIso (coHomGr 2 G (S¹ × S¹)) (coHomGr 1 G S¹) + fst theIso = is + snd theIso = makeIsGroupHom + (coHomPointedElimT²'' {G = G} 0 (λ _ → isPropΠ λ _ → squash₂ _ _) + λ p → coHomPointedElimT²'' {G = G} 0 (λ _ → squash₂ _ _) + λ q → cong ∣_∣₂ (funExt λ x + → cong (ΩEM+1→EM {G = G} 1) (help p q x) + ∙ ΩEM+1→EM-hom {G = G} 1 (λ i → elimFunT²' {G = G} 1 p (x , loop i)) + (λ i → elimFunT²' {G = G} 1 q (x , loop i)))) + where + help : (p q : Path (0ₖ {G = G} 2 ≡ 0ₖ {G = G} 2) refl refl) (x : S¹) + → (λ i → elimFunT²' 1 p (x , loop i) +2 elimFunT²' 1 q (x , loop i)) + ≡ (λ i → elimFunT²' 1 p (x , loop i)) ∙ (λ i → elimFunT²' 1 q (x , loop i)) + help p q x = cong₂+₂ {G = G} 0 (cong (λ y → elimFunT²' {G = G} 1 p (x , y)) loop) + (cong (λ y → elimFunT²' {G = G} 1 q (x , y)) loop) + +H³⁺ⁿ[T²,G]≅0 : ∀ {ℓ} (G : AbGroup ℓ) (n : ℕ) + → AbGroupEquiv (coHomGr (3 + n) G (S¹ × S¹)) (trivialAbGroup {ℓ}) +fst (H³⁺ⁿ[T²,G]≅0 G n) = + isoToEquiv (isContr→Iso + (0ₕ (3 + n) + , coHomPointedElimT²'' {G = G} (suc n) + (λ _ → squash₂ _ _) + λ P → Trunc.rec (isProp→isOfHLevelSuc n (squash₂ _ _)) + (λ P → cong ∣_∣₂ (funExt λ { (base , base) → refl + ; (base , loop i) → refl + ; (loop i , base) → refl + ; (loop i , loop j) k → P k i j})) + (isConnectedPath (suc n) (isConnectedPath (2 + n) + (isConnectedPath (3 + n) (isConnectedEM (3 + n)) _ _) _ _) refl P .fst)) + isContrUnit*) +snd (H³⁺ⁿ[T²,G]≅0 G n) = makeIsGroupHom λ _ _ → refl diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/Unit.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/Unit.agda new file mode 100644 index 0000000000..c013143348 --- /dev/null +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/Unit.agda @@ -0,0 +1,49 @@ +{-# OPTIONS --safe #-} + +module Cubical.Cohomology.EilenbergMacLane.Groups.Unit where + +open import Cubical.Cohomology.EilenbergMacLane.Base +open import Cubical.Cohomology.EilenbergMacLane.Groups.Connected + +open import Cubical.Homotopy.EilenbergMacLane.Base +open import Cubical.Homotopy.EilenbergMacLane.Properties +open import Cubical.Homotopy.Connected + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup.Base + +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as Trunc + +open import Cubical.Data.Unit + +H⁰[Unit,G]≅G : ∀ {ℓ} {G : AbGroup ℓ} + → AbGroupEquiv (coHomGr 0 G Unit) G +H⁰[Unit,G]≅G {G = G} = H⁰conn isConnectedUnit G + where + isConnectedUnit : isConnected 2 Unit + fst isConnectedUnit = ∣ tt ∣ + snd isConnectedUnit = Trunc.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + λ _ → refl + +isContr-Hⁿ⁺¹[Unit,G] : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) + → isContr (coHom (suc n) G Unit) +fst (isContr-Hⁿ⁺¹[Unit,G] {G = G} n) = 0ₕ (suc n) +snd (isContr-Hⁿ⁺¹[Unit,G] {G = G} n) = + ST.elim (λ _ → isSetPathImplicit) + λ f → Trunc.rec + (isProp→isOfHLevelSuc n (squash₂ _ _)) + (λ p → cong ∣_∣₂ (funExt λ _ → p)) + (isConnectedPath (suc n) + (isConnectedEM (suc n)) (0ₖ (suc n)) (f tt) .fst) + +Hⁿ⁺¹[Unit,G]≅0 : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) + → AbGroupEquiv (coHomGr (suc n) G Unit) (trivialAbGroup {ℓ = ℓ}) +fst (Hⁿ⁺¹[Unit,G]≅0 {G = G} n) = isoToEquiv (isContr→Iso (isContr-Hⁿ⁺¹[Unit,G] n) isContrUnit*) +snd (Hⁿ⁺¹[Unit,G]≅0 n) = makeIsGroupHom λ _ _ → refl diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda index 8ad9613002..bcc000cba0 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda @@ -3,6 +3,7 @@ module Cubical.Cohomology.EilenbergMacLane.Rings.KleinBottle where open import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle +open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2 open import Cubical.Cohomology.EilenbergMacLane.Base open import Cubical.Cohomology.EilenbergMacLane.CupProduct open import Cubical.Cohomology.EilenbergMacLane.RingStructure diff --git a/Cubical/Data/Fin/Properties.agda b/Cubical/Data/Fin/Properties.agda index b57b4ec979..d48638d7e9 100644 --- a/Cubical/Data/Fin/Properties.agda +++ b/Cubical/Data/Fin/Properties.agda @@ -680,3 +680,41 @@ isProp→Fin≤1 : (n : ℕ) → isProp (Fin n) → n ≤ 1 isProp→Fin≤1 0 _ = ≤-solver 0 1 isProp→Fin≤1 1 _ = ≤-solver 1 1 isProp→Fin≤1 (suc (suc n)) p = Empty.rec (fzero≠fone (p fzero fone)) + +-- Characterisation of Π over (Fin (suc n)) +CharacΠFinIso : ∀ {ℓ} (n : ℕ) {B : Fin (suc n) → Type ℓ} + → Iso ((x : _) → B x) (B fzero × ((x : _) → B (fsuc x))) +Iso.fun (CharacΠFinIso n {B = B}) f = f fzero , f ∘ fsuc +Iso.inv (CharacΠFinIso n {B = B}) (x , f) (zero , p) = + subst B (Fin-fst-≡ {i = fzero} {j = zero , p} refl) x +Iso.inv (CharacΠFinIso n {B = B}) (x , f) (suc y , p) = + subst B (Fin-fst-≡ refl) (f (y , pred-≤-pred p)) +Iso.rightInv (CharacΠFinIso n {B = B}) (x , f) = + ΣPathP ((λ j → + transp (λ i → B (isSetFin fzero fzero (Fin-fst-≡ (λ _ → zero)) refl j i)) j x) + , funExt λ x → (λ j → subst B (pathid x j) + (f (fst x , pred-≤-pred (suc-≤-suc (snd x))))) + ∙ (λ i → transp (λ j → B (path₃ x (i ∨ j))) i (f (path₂ x i)))) + where + module _ (x : Fin n) where + path₁ : _ + path₁ = Fin-fst-≡ {i = (fsuc (fst x , pred-≤-pred (snd (fsuc x))))} + {j = fsuc x} refl + + path₂ : _ + path₂ = Fin-fst-≡ refl + + path₃ : _ + path₃ = cong fsuc path₂ + + pathid : path₁ ≡ path₃ + pathid = isSetFin _ _ _ _ + +Iso.leftInv (CharacΠFinIso n {B = B}) f = + funExt λ { (zero , p) j + → transp (λ i → B (Fin-fst-≡ {i = fzero} {j = zero , p} refl (i ∨ j))) + j (f (Fin-fst-≡ {i = fzero} {j = zero , p} refl j)) + ; (suc x , p) j → transp (λ i → B (q x p (i ∨ j))) j (f (q x p j))} + where + q : (x : _) (p : _) → _ + q x p = Fin-fst-≡ {i = (fsuc (x , pred-≤-pred p))} {j = suc x , p} refl diff --git a/Cubical/HITs/KleinBottle/Properties.agda b/Cubical/HITs/KleinBottle/Properties.agda index d9c9628ca8..8b4a30ce21 100644 --- a/Cubical/HITs/KleinBottle/Properties.agda +++ b/Cubical/HITs/KleinBottle/Properties.agda @@ -55,6 +55,24 @@ elimProp prop a₀ = (isProp→PathP (λ _ → prop _) _ _) (isProp→PathP (λ _ → prop _) _ _) +K²FunCharacIso : ∀ {ℓ} {A : KleinBottle → Type ℓ} + → Iso ((x : KleinBottle) → A x) + (Σ[ x ∈ A point ] + Σ[ p ∈ PathP (λ i → A (line1 i)) x x ] + (Σ[ q ∈ PathP (λ i → A (line2 i)) x x ] + SquareP (λ i j → A (square i j)) + q q (λ i → p (~ i)) p)) +Iso.fun K²FunCharacIso f = + f point , cong f line1 , cong f line2 , λ i j → f (square i j) +Iso.inv K²FunCharacIso (a , p1 , p2 , sq) point = a +Iso.inv K²FunCharacIso (a , p1 , p2 , sq) (line1 i) = p1 i +Iso.inv K²FunCharacIso (a , p1 , p2 , sq) (line2 i) = p2 i +Iso.inv K²FunCharacIso (a , p1 , p2 , sq) (square i j) = sq i j +Iso.rightInv K²FunCharacIso _ = refl +Iso.leftInv K²FunCharacIso f = + funExt λ { point → refl ; (line1 i) → refl + ; (line2 i) → refl ; (square i i₁) → refl} + loop1 : S¹ → KleinBottle loop1 base = point loop1 (loop i) = line1 i diff --git a/Cubical/HITs/RPn/Base.agda b/Cubical/HITs/RPn/Base.agda index db0a62202c..5f4c20c7ea 100644 --- a/Cubical/HITs/RPn/Base.agda +++ b/Cubical/HITs/RPn/Base.agda @@ -348,3 +348,17 @@ characRP²Fun : ∀ {ℓ} {A : Type ℓ} (f : RP² → A) → RP²Fun (f point) (cong f line) (λ i j → f (square i j)) ≡ f characRP²Fun f = funExt λ { point → refl ; (line i) → refl ; (square i i₁) → refl} + +RP²FunCharacIso : {A : RP² → Type ℓ} + → Iso ((x : RP²) → A x) + (Σ[ x ∈ A point ] + Σ[ p ∈ PathP (λ i → A (line i)) x x ] + SquareP (λ i j → A (square i j)) + p (λ i → p (~ i)) refl refl) +Iso.fun RP²FunCharacIso f = f point , cong f line , cong (cong f) square +Iso.inv RP²FunCharacIso (x , p , q) point = x +Iso.inv RP²FunCharacIso (x , p , q) (line i) = p i +Iso.inv RP²FunCharacIso (x , p , q) (square i j) = q i j +Iso.rightInv RP²FunCharacIso _ = refl +Iso.leftInv RP²FunCharacIso f = + funExt λ { point → refl ; (line i) → refl ; (square i i₁) → refl} diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda index 887dc002f4..192c7aedcc 100644 --- a/Cubical/HITs/Wedge/Base.agda +++ b/Cubical/HITs/Wedge/Base.agda @@ -11,6 +11,10 @@ open import Cubical.Foundations.GroupoidLaws _⋁_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') _⋁_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ _ → ptA) (λ _ → ptB) +-- Arbitrary wedges +⋁gen : ∀ {ℓ ℓ'} (A : Type ℓ) (B : A → Pointed ℓ') → Type (ℓ-max ℓ ℓ') +⋁gen A B = cofib {A = A} {B = Σ A λ a → fst (B a)} + (λ a → a , snd (B a)) -- Pointed versions _⋁∙ₗ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') @@ -19,6 +23,9 @@ A ⋁∙ₗ B = (A ⋁ B) , (inl (snd A)) _⋁∙ᵣ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') A ⋁∙ᵣ B = (A ⋁ B) , (inr (snd B)) +⋁gen∙ : ∀ {ℓ ℓ'} (A : Type ℓ) (B : A → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +⋁gen∙ A B = ⋁gen A B , inl tt + -- Wedge sums of functions _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → (f : A →∙ C) (g : B →∙ C) diff --git a/Cubical/Homotopy/EilenbergMacLane/Properties.agda b/Cubical/Homotopy/EilenbergMacLane/Properties.agda index 76a5092bec..953ce48309 100644 --- a/Cubical/Homotopy/EilenbergMacLane/Properties.agda +++ b/Cubical/Homotopy/EilenbergMacLane/Properties.agda @@ -512,6 +512,24 @@ module _ {G : AbGroup ℓ} where lem zero = ΩEM+1→EM-refl 0 lem (suc n) = ΩEM+1→EM-refl (suc n) + ΩEM+1→EM-gen-hom : (n : ℕ) (x : EM G (suc n)) (p q : x ≡ x) + → ΩEM+1→EM-gen n x (p ∙ q) ≡ ΩEM+1→EM-gen n x p +ₖ ΩEM+1→EM-gen n x q + ΩEM+1→EM-gen-hom n = + EM-raw'-elim _ (suc n) + (λ _ → isOfHLevelΠ2 (suc (suc n)) + (λ _ _ → isOfHLevelPath (suc (suc n)) (hLevelEM _ n) _ _)) + (EM-raw'-trivElim _ n + (λ _ → isOfHLevelΠ2 (suc n) (λ _ _ → hLevelEM _ n _ _)) + (lem n)) + where + lem : (n : ℕ) (p q : EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n))) + ≡ EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n)))) + → ΩEM+1→EM-gen n _ (p ∙ q) + ≡ ΩEM+1→EM-gen n _ p + +ₖ ΩEM+1→EM-gen n _ q + lem zero p q = ΩEM+1→EM-hom 0 p q + lem (suc n) p q = ΩEM+1→EM-hom (suc n) p q + -- Some HLevel lemmas about function spaces (EM∙ G n →∙ EM∙ H m), mainly used for -- the cup product module _ where diff --git a/Cubical/Homotopy/EilenbergSteenrod.agda b/Cubical/Homotopy/EilenbergSteenrod.agda index 26502a1c09..bf49cd26aa 100644 --- a/Cubical/Homotopy/EilenbergSteenrod.agda +++ b/Cubical/Homotopy/EilenbergSteenrod.agda @@ -31,6 +31,9 @@ open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.Pi + +open import Cubical.Axiom.Choice record coHomTheory {ℓ ℓ' : Level} (H : (n : ℤ) → Pointed ℓ → AbGroup ℓ') : Type (ℓ-suc (ℓ-max ℓ ℓ')) where @@ -47,3 +50,32 @@ record coHomTheory {ℓ ℓ' : Level} (H : (n : ℤ) → Pointed ℓ → AbGroup ≡ Im (Hmap n {B = _ , inr (pt B)} (cfcod (fst f) , refl)) Dimension : (n : ℤ) → ¬ n ≡ 0 → isContr (fst (H n Boolℓ)) BinaryWedge : (n : ℤ) {A B : Pointed ℓ} → AbGroupEquiv (H n (A ⋁ B , (inl (pt A)))) (dirProdAb (H n A) (H n B)) + +record coHomTheoryGen {ℓ ℓ' : Level} (H : (n : ℤ) → Pointed ℓ → AbGroup ℓ') : Type (ℓ-suc (ℓ-max ℓ ℓ')) + where + Boolℓ : Pointed ℓ + Boolℓ = Lift Bool , lift true + field + Hmap : (n : ℤ) → {A B : Pointed ℓ} (f : A →∙ B) → AbGroupHom (H n B) (H n A) + HMapComp : (n : ℤ) → {A B C : Pointed ℓ} (g : B →∙ C) (f : A →∙ B) → compGroupHom (Hmap n g) (Hmap n f) ≡ Hmap n (g ∘∙ f) + HMapId : (n : ℤ) → {A : Pointed ℓ} → Hmap n (idfun∙ A) ≡ idGroupHom + + Suspension : Σ[ F ∈ ((n : ℤ) {A : Pointed ℓ} → AbGroupEquiv (H (sucℤ n) (Susp (typ A) , north)) (H n A)) ] + ({A B : Pointed ℓ} (f : A →∙ B) (n : ℤ) + → fst (Hmap (sucℤ n) (suspFun (fst f) , refl)) ∘ invEq (fst (F n {A = B})) + ≡ invEq (fst (F n {A = A})) ∘ fst (Hmap n f)) + Exactness : {A B : Pointed ℓ} (f : A →∙ B) (n : ℤ) + → Ker (Hmap n f) + ≡ Im (Hmap n {B = _ , inr (pt B)} (cfcod (fst f) , refl)) + Dimension : (n : ℤ) → ¬ n ≡ 0 → isContr (fst (H n Boolℓ)) + Wedge : (n : ℤ) {I : Type ℓ} (satAC : satAC (ℓ-max ℓ ℓ') 2 I) {A : I → Pointed ℓ} + → isEquiv {A = H n (⋁gen∙ I A) .fst} + {B = ΠAbGroup (λ i → H n (A i)) .fst} + λ x i → Hmap n ((λ x → inr (i , x)) , sym (push i)) .fst x + + Wedge→AbGroupEquiv : (n : ℤ) {I : Type ℓ} (satAC : satAC (ℓ-max ℓ ℓ') 2 I) {A : I → Pointed ℓ} + → AbGroupEquiv (H n (⋁gen∙ I A)) (ΠAbGroup (λ i → H n (A i)) ) + fst (fst (Wedge→AbGroupEquiv n satAC)) = _ + snd (fst (Wedge→AbGroupEquiv n satAC)) = Wedge n satAC + snd (Wedge→AbGroupEquiv n satAC) = makeIsGroupHom λ f g + → funExt λ i → IsGroupHom.pres· (Hmap n ((λ x → inr (i , x)) , sym (push i)) .snd) _ _ diff --git a/Cubical/Papers/ComputationalSyntheticCohomology.agda b/Cubical/Papers/ComputationalSyntheticCohomology.agda new file mode 100644 index 0000000000..25c768039b --- /dev/null +++ b/Cubical/Papers/ComputationalSyntheticCohomology.agda @@ -0,0 +1,334 @@ +{- + +Please do not move this file. Changes should only be made if necessary. + +This file contains pointers to the code examples and main results from +the paper: + +Computational Synthetic Cohomology Theory in Homotopy Type Theory + +-} + +-- The "--safe" flag ensures that there are no postulates or unfinished goals +{-# OPTIONS --safe #-} +module Cubical.Papers.ComputationalSyntheticCohomology where + +open import Cubical.Data.Nat +open import Cubical.Algebra.AbGroup +open import Cubical.Foundations.GroupoidLaws + +-- 1: Introduction + +import Cubical.ZCohomology.Groups.S2wedgeS1wedgeS1 as HⁿS²∨S¹∨S¹ + +-- 2: Background + +open import Cubical.Foundations.Prelude as Prelude +import Cubical.Homotopy.HSpace as hSpace +import Cubical.Foundations.Pointed.Homogeneous as Homogeneous +import Cubical.HITs.Susp as Suspensions +import Cubical.HITs.Pushout as Pushouts +import Cubical.Foundations.Path as Paths + +-- 3: Stuff +import Cubical.Homotopy.EilenbergMacLane.Base as EMSpace +import Cubical.Homotopy.EilenbergMacLane.Properties as EMProps +import Cubical.Homotopy.EilenbergMacLane.WedgeConnectivity as WC +import Cubical.Homotopy.EilenbergMacLane.GroupStructure as EMGr +import Cubical.Algebra.AbGroup.TensorProduct as Tensor +import Cubical.Homotopy.EilenbergMacLane.CupProductTensor as Cup⊗ +import Cubical.Homotopy.EilenbergMacLane.CupProduct as Cupₖ +import Cubical.Homotopy.EilenbergMacLane.GradedCommTensor as CupComm + +-- 4: Cohomology +import Cubical.Cohomology.EilenbergMacLane.Base as Cohom +import Cubical.Cohomology.EilenbergMacLane.EilenbergSteenrod as Axioms +import Cubical.Cohomology.EilenbergMacLane.CupProduct as CohomCup +import Cubical.Cohomology.EilenbergMacLane.MayerVietoris as MV +import Cubical.Axiom.Choice as Choice +import Cubical.HITs.Wedge as ⋁ + +-- 5: Computations of cohomology groups and rings +import Cubical.Cohomology.EilenbergMacLane.Groups.Unit as HⁿUnit +import Cubical.Cohomology.EilenbergMacLane.Groups.Connected as CohomConnected +import Cubical.Cohomology.EilenbergMacLane.Groups.Sn as CohomSn +import Cubical.Cohomology.EilenbergMacLane.Groups.Torus as CohomT² +import Cubical.HITs.Torus as T² +import Cubical.HITs.RPn.Base as RP² +import Cubical.HITs.KleinBottle as K² +import Cubical.Cohomology.EilenbergMacLane.Groups.RP2 as CohomRP² +import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle as CohomK² +import Cubical.Cohomology.EilenbergMacLane.Rings.KleinBottle as K²Ring +-- import Cubical.Cohomology.Groups.KleinBottle as HⁿK² +-- import Cubical.ZCohomology.Groups.KleinBottle as HⁿK² + +-- import Cubical.Cohomology.EilenbergMacLane.RSnings as CohomRingSn + +-- import Cubical.ZCohomology.RingStructure.CohomologyRing as ℤCohomologyRing +-- import Cubical.HITs.S1 as S1 + +-- import Cubical.HITs.Sn as Sn +-- import Cubical.ZCohomology.Groups.Sn as HⁿSᵐ +-- import Cubical.ZCohomology.CohomologyRings.S1 as H*S¹ +-- import Cubical.ZCohomology.CohomologyRings.Sn as H*Sᵐ +-- open import Cubical.Homotopy.Hopf as HopfFibration +-- import Cubical.ZCohomology.Groups.CP2 as HⁿℂP² +-- import Cubical.ZCohomology.CohomologyRings.CP2 as H*ℂP² +-- import Cubical.HITs.Wedge as ⋁ +-- import Cubical.ZCohomology.Groups.S2wedgeS4 as HⁿS²∨S⁴ +-- import Cubical.ZCohomology.CohomologyRings.S2wedgeS4 as H*S²∨S⁴ +-- import Cubical.Cohomology.EilenbergMacLane.RingStructure as GCohomologyRing +-- import Cubical.HITs.KleinBottle as 𝕂² + +-- import Cubical.ZCohomology.CohomologyRings.KleinBottle as H*𝕂² +-- import Cubical.ZCohomology.Groups.RP2wedgeS1 as HⁿℝP²∨S¹ +-- import Cubical.ZCohomology.CohomologyRings.RP2wedgeS1 as H*ℝP²∨S¹ +-- import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle as ℤ/2-Hⁿ𝕂² +-- open import Cubical.Cohomology.EilenbergMacLane.Rings.KleinBottle as ℤ/2-H*𝕂² +-- import Cubical.Cohomology.EilenbergMacLane.Groups.RP2wedgeS1 as ℤ/2-HⁿℝP²∨S¹ +-- import Cubical.Cohomology.EilenbergMacLane.Rings.RP2wedgeS1 as ℤ/2-H*ℝP²∨S¹ +-- renaming (RP²∨S¹-CohomologyRing to H*RP²∨S¹≅ℤ/2[X,Y]/) + + +----- 1. INNTRODUCTION ----- + +----- 2. Background on HoTT/UF and notations ----- + +-- 2.1 Homotopy Type Theory in Cubical Agda + +-- Definition 1. Binary ap +open Prelude renaming (cong₂ to ap²) + +-- Definition 2. H-spaces +open hSpace using (HSpace) + +-- Definition 3. Homogeneous types +open Homogeneous using (isHomogeneous) + +-- Definition 4. Suspension +open Suspensions using (Susp) + +-- Definition 5. Pushouts +open Pushouts using (Pushout) + +-- Lemma 6. Flip +open Paths using (sym≡flipSquare) + +----- 3. Eilenberg-MacLane spaces ----- +-- Definition 7. +-- Included for readability: not used explicitly in formalisation. + +-- Definition 8. (raw Eilenbreg-MacLane spaces) EM-raw +open EMSpace using (EM-raw) + +-- Definition 9. (Eilenbreg-MacLane spaces) EM-raw +open EMSpace using (EM) + +-- Proposition 10. Connectivity of K(G,n) +open EMProps using (isConnectedEM) + +-- Lemma 11. +open EMProps using (EMFun-EM→ΩEM+1) + +--- 3.1 Group Structure +-- Proposition 12. (Wedge connectivity for K(G,n)) +open WC.wedgeConEM using (fun ; left ; right) + +-- Proposition 13. (unit laws for +ₖ) +open EMGr using (rUnitₖ ; lUnitₖ) +lUnit≡rUnit : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) + → rUnitₖ {G = G} n (EMSpace.0ₖ n) ≡ lUnitₖ {G = G} n (EMSpace.0ₖ n) +lUnit≡rUnit {G = G} zero = AbGroupStr.is-set (snd G) _ _ _ _ +lUnit≡rUnit (suc zero) = refl +lUnit≡rUnit (suc (suc n)) = refl + +-- Proposition 14. (commutativity of +ₖ) +open EMGr using (commₖ) + +-- Proposition 15. (associativity of +ₖ) +open EMGr using (assocₖ) + +-- Proposition 16. (ap² on +ₖ) +open EMGr using (cong₂+₁ ; cong₂+₂) + +-- Proposition 17. (cancellation laws of +ₖ) +open EMGr using (rCancelₖ ; lCancelₖ) + +--- 3.2 K(Gn) vs ΩK(G,n+1) +-- Proposition 18. Commutativity of Ω(K(G,n),x) +open EMProps using (isCommΩEM-base) + +-- Proposition 19. (ap -ₖ = path inversion) +open EMGr using (cong-₁ ; cong-₂) + +-- Proposition 20. (σ preserves +ₖ) +open EMProps using (EM→ΩEM+1-hom) + +-- Corollary 21. (σ preserves -ₖ) +open EMProps using (EM→ΩEM+1-sym) + +-- Theorem 22. K(G,n) ≃ Ω(K(G,n+1)) +open EMProps using (EM≃ΩEM+1) + +--- 3.3 The cup product with group coefficients +-- Definition 23. Tensor products of abelian groups +open Tensor using (_⨂_) + +-- Proposition 24. Truncation of K(G,n) →* K(H;n+m) +open EMProps using (isOfHLevel↑∙) + +-- The general cup product +open Cup⊗ renaming (_⌣ₖ_ to _⌣ₖ⊗_) + +-- Proposition 25 (anihilation laws) +open Cup⊗ using (⌣ₖ-0ₖ ; 0ₖ-⌣ₖ) + +-- Lemma 26 (Evan's trick) +open Homogeneous using (→∙Homogeneous≡) + +-- Propositions 27 and 28 (distributivity of ⌣ₖ) +open Cup⊗.LeftDistributivity using (main) +open Cup⊗.RightDistributivity using (main) + +-- Lemma 29 +open Cup⊗ using (EM→ΩEM+1-distr₀ₙ ; EM→ΩEM+1-distrₙsuc ; EM→ΩEM+1-distrₙ₀) + +-- Proposition 30 (associativity of ⌣ₖ) +open Cup⊗.Assoc using (main) + +-- Proposition 31 (graded commutativity of CupComm) +open CupComm using (⌣ₖ-comm) + +--- 3.4 The cup product with ring coefficients +-- Proposition 32 (ring structure on ⌣ₖ with ring coefficents) +open Cupₖ using (assoc⌣ₖ ; distrL⌣ₖ ; distrR⌣ₖ ; ⌣ₖ-0ₖ ; 0ₖ-⌣ₖ) + +-- Proposition 33 (graded commutativity) +open Cupₖ using (⌣ₖ-comm) + +-- Proposition 34 (neutral element) +open Cupₖ using (⌣ₖ-1ₖ ; 1ₖ-⌣ₖ) + + +----- 4. Cohomology ----- +-- Defininition of cohomology groups +open Cohom using (coHomGr) + +-- Cup products on cohomology +open CohomCup using (_⌣_) + +--- 4.1 Reduced cohomology +-- Definition of reduced cohomology groups +open Cohom using (coHomRedGr) + +-- Proposition 35. +open Cohom using (coHom≅coHomRed) + +-- Proposition 36. +open Cohom using (coHom⁰≅coHomRed⁰) + +--- 4.2 Eilenberg-Steenrod axioms for cohomology +-- Definition 37. Cofibres +open Pushouts using (cofib) + +-- Definition 38. Arbitrary wedges +open ⋁ using (⋁gen) + +-- Definition 39. Choice +open Choice using (satAC) + +-- Definition 40. Eilenberg-Steenrod axioms. Todo!! + +-- Theorem 41. Eilenberg-Steenrod axioms. Todo!! + +-- Theorem 42 (The mayer-Vietoris sequence) +open MV.MV using ( Ker-i⊂Im-d ; Im-d⊂Ker-i + ; Ker-Δ⊂Im-i ; Im-i⊂Ker-Δ + ; Ker-d⊂Im-Δ ; Im-Δ⊂Ker-d) + +--- 4.4 The Thom isomorphism and the Gysin sequence +-- Theorem 43 (Gysin sequence) +-- TODO! Port from Steenrod branch. + +-- + +-- Proposition 44. Cohom of 1 +open HⁿUnit using (H⁰[Unit,G]≅G ; Hⁿ⁺¹[Unit,G]≅0) + +-- Lemma 45. Cohom of truncation +-- Todo! + +-- Proposition 46. Cohomology of connected types +open CohomConnected using (H⁰conn) + +--- 5.1 Speres +-- Proposition 47. H¹(S¹,G) +open CohomSn using (H¹[S¹,G]≅G) + +-- Proposition 48. Hⁿ(Sⁿ,G) +open CohomSn using (Hⁿ[Sⁿ,G]≅G) + +-- Proposition 49. Hⁿ(Sᵐ,G) , m ≠ n +open CohomSn using (Hⁿ[Sᵐ⁺ⁿ,G]≅0 ; Hᵐ⁺ⁿ[Sⁿ,G]≅0) + +-- Proposition 50. H*(Sᵐ,G) +-- Todo (just transalte from ℤ version) + +--- 5.2 The Torus +-- Definition 51. Torus +open T² using (Torus) + +-- Proposition 52. Hⁿ(T²,G) +open CohomT² using (H⁰[T²,G]≅G ; H¹[T²,G]≅G×G ; H²[T²,G]≅G ; H³⁺ⁿ[T²,G]≅0) + +--- 5.3 The Real Projective Plane and The Klein Bottle +-- Definition 53. Real Projective Plane +open RP² using (RP²) +-- Definition 54. The Klein Bottle +open K² using (KleinBottle) + +-- Proposition 55 H¹(RP²,G) ≅ G[2] +open CohomRP² using (H¹[RP²,G]≅G[2]) + +-- Proposition 56 H²(RP²,G) ≅ G/2 +open CohomRP² using (H²[RP²,G]≅G/2) + +-- Proposition 57 Hᵐ(RP²,G) ≅ 0, m ≥ 3 +open CohomRP² using (H³⁺ⁿ[RP²,G]≅0) + +-- Proposition 58. Hᵐ(K²,G) +open CohomK² using (H⁰[K²,G]≅G ; H¹[K²,G]≅G×H¹[RP²,G]; H²⁺ⁿ[K²,G]≅H²⁺ⁿ[RP²,G]) + +-- Proposition 59. H*(RP²,\bZ) +-- Todo + +-- Defnition 60 kill-Δ +-- Omitted (implicit in formalisation) + +-- Proposition 61 Hᵐ(K²,G) +-- Thomas proof + +-- Definition 62. Res (only formliased for relevant special case of F = ⌣ on K(ℤ/2,1)) +open K²Ring using (Res⌣) + +-- Propsoiton 63. Res refl = refl +-- omitted (used implicitly in formalisation) + +-- Lemma 64/65. ap2-funct ocherence (this is roughly the special case of lemma 61) +open K²Ring using (sym-cong₂-⌣≡) + +-- Proposition 66. Implicitly formalised via the following theorem +open K²Ring using (α²↦1) + +-- Proposition 67. H*(RP²,Z/2) +-- Todo + +-- Proposition 68 (Roughly) +open K²Ring using (α²↦1 ; βα↦1 ; β²↦0) + +-- Proposition 69 H*(K²,ℤ/2) +open K²Ring using (H*KleinBottle≅ℤ/2[X,Y]/) + +-- Lemma 70 (implicitly used) + +-- Proposition 71 (do!) From 2fef4ef8ccaf61ec6105bb9c6201281bd305eb47 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 22 Jan 2024 18:36:10 +0100 Subject: [PATCH 11/73] stuff --- Cubical/Algebra/AbGroup/Instances/IntMod.agda | 68 +++ Cubical/Cohomology/EilenbergMacLane/Base.agda | 9 + .../EilenbergMacLane/CupProduct.agda | 1 + .../EilenbergMacLane/Groups/KleinBottle.agda | 1 + .../EilenbergMacLane/Groups/RP2.agda | 1 + .../EilenbergMacLane/Groups/RP2wedgeS1.agda | 1 + .../EilenbergMacLane/Groups/Sn.agda | 41 +- .../EilenbergMacLane/Rings/KleinBottle.agda | 1 + .../EilenbergMacLane/Rings/RP2.agda | 521 ++++++++++++++++++ .../EilenbergMacLane/Rings/RP2wedgeS1.agda | 3 +- .../Cohomology/EilenbergMacLane/Rings/Sn.agda | 428 ++++++++++++++ .../EilenbergMacLane/Rings/Z2-properties.agda | 1 + .../Homotopy/EilenbergMacLane/CupProduct.agda | 47 +- Cubical/Homotopy/EilenbergMacLane/Order2.agda | 4 +- .../ComputationalSyntheticCohomology.agda | 36 +- .../CohomologyRings/KleinBottle.agda | 2 +- Cubical/ZCohomology/CohomologyRings/RP2.agda | 2 +- 17 files changed, 1145 insertions(+), 22 deletions(-) create mode 100644 Cubical/Algebra/AbGroup/Instances/IntMod.agda create mode 100644 Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda create mode 100644 Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda diff --git a/Cubical/Algebra/AbGroup/Instances/IntMod.agda b/Cubical/Algebra/AbGroup/Instances/IntMod.agda new file mode 100644 index 0000000000..cb6bfcf652 --- /dev/null +++ b/Cubical/Algebra/AbGroup/Instances/IntMod.agda @@ -0,0 +1,68 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.AbGroup.Instances.IntMod where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.MorphismProperties + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Int renaming (_+_ to _+ℤ_) +open import Cubical.Data.Fin +open import Cubical.Data.Fin.Arithmetic +open import Cubical.Data.Sigma + +open import Cubical.HITs.SetQuotients as SQ renaming ([_] to ⟦_⟧) +open import Cubical.HITs.PropositionalTruncation as PT + +ℤAbGroup/_ : ℕ → AbGroup ℓ-zero +ℤAbGroup/ n = Group→AbGroup (ℤGroup/ n) (comm n) + where + comm : (n : ℕ) (x y : fst (ℤGroup/ n)) + → GroupStr._·_ (snd (ℤGroup/ n)) x y + ≡ GroupStr._·_ (snd (ℤGroup/ n)) y x + comm zero = +Comm + comm (suc n) = +ₘ-comm + +ℤ/2 : AbGroup ℓ-zero +ℤ/2 = ℤAbGroup/ 2 + +ℤ/2[2]≅ℤ/2 : AbGroupIso (ℤ/2 [ 2 ]) ℤ/2 +Iso.fun (fst ℤ/2[2]≅ℤ/2) = fst +Iso.inv (fst ℤ/2[2]≅ℤ/2) x = x , cong (x +ₘ_) (+ₘ-rUnit x) ∙ x+x x + where + x+x : (x : ℤ/2 .fst) → x +ₘ x ≡ fzero + x+x = ℤ/2-elim refl refl +Iso.rightInv (fst ℤ/2[2]≅ℤ/2) x = Σ≡Prop (λ _ → isProp≤) refl +Iso.leftInv (fst ℤ/2[2]≅ℤ/2) x = Σ≡Prop (λ _ → isSetFin _ _) refl +snd ℤ/2[2]≅ℤ/2 = makeIsGroupHom λ _ _ → refl + +ℤ/2/2≅ℤ/2 : AbGroupIso (ℤ/2 /^ 2) ℤ/2 +Iso.fun (fst ℤ/2/2≅ℤ/2) = + SQ.rec isSetFin (λ x → x) lem + where + lem : _ + lem = ℤ/2-elim (ℤ/2-elim (λ _ → refl) + (PT.rec (isSetFin _ _) + (uncurry (ℤ/2-elim + (λ p → ⊥.rec (snotz (sym (cong fst p)))) + λ p → ⊥.rec (snotz (sym (cong fst p))))))) + (ℤ/2-elim (PT.rec (isSetFin _ _) + (uncurry (ℤ/2-elim + (λ p → ⊥.rec (snotz (sym (cong fst p)))) + λ p → ⊥.rec (snotz (sym (cong fst p)))))) + λ _ → refl) +Iso.inv (fst ℤ/2/2≅ℤ/2) = ⟦_⟧ +Iso.rightInv (fst ℤ/2/2≅ℤ/2) _ = refl +Iso.leftInv (fst ℤ/2/2≅ℤ/2) = + SQ.elimProp (λ _ → squash/ _ _) λ _ → refl +snd ℤ/2/2≅ℤ/2 = makeIsGroupHom + (SQ.elimProp (λ _ → isPropΠ λ _ → isSetFin _ _) + λ a → SQ.elimProp (λ _ → isSetFin _ _) λ b → refl) diff --git a/Cubical/Cohomology/EilenbergMacLane/Base.agda b/Cubical/Cohomology/EilenbergMacLane/Base.agda index 57cc24968e..5b87f5e509 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Base.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Base.agda @@ -31,6 +31,7 @@ open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.Monoid open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.AbGroup.Instances.IntMod open import Cubical.Algebra.AbGroup.Instances.DirectProduct open import Cubical.Algebra.AbGroup.Properties @@ -426,3 +427,11 @@ coHomPointedElim {ℓ'' = ℓ''} {G = G} {A = A} n a isprop indp = TR.rec (isProp→isOfHLevelSuc n (isprop _)) (ind f) (isConnectedPath (suc n) (isConnectedEM (suc n)) (f a) (0ₖ (suc n)) .fst) + +coHomTruncEquiv : {A : Type ℓ} (G : AbGroup ℓ) (n : ℕ) + → AbGroupEquiv (coHomGr n G (∥ A ∥ (suc (suc n)))) (coHomGr n G A) +fst (coHomTruncEquiv G n) = + isoToEquiv (setTruncIso (univTrunc (suc (suc n)) {B = _ , hLevelEM G n})) +snd (coHomTruncEquiv G n) = + makeIsGroupHom (ST.elim2 (λ _ _ → isSetPathImplicit) + λ _ _ → refl) diff --git a/Cubical/Cohomology/EilenbergMacLane/CupProduct.agda b/Cubical/Cohomology/EilenbergMacLane/CupProduct.agda index 74d3f9f61e..65a9ad0eda 100644 --- a/Cubical/Cohomology/EilenbergMacLane/CupProduct.agda +++ b/Cubical/Cohomology/EilenbergMacLane/CupProduct.agda @@ -17,6 +17,7 @@ open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.IntMod open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.AbGroup.Instances.IntMod open import Cubical.Foundations.Prelude diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda index c0042296c5..0e9f39f8ad 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda @@ -37,6 +37,7 @@ open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.DirectProduct +open import Cubical.Algebra.AbGroup.Instances.IntMod open import Cubical.HITs.KleinBottle renaming (rec to KleinFun) open import Cubical.HITs.SetTruncation as ST diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda index 06c2069d69..8f79b44abe 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda @@ -34,6 +34,7 @@ open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.IntMod open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.IntMod open import Cubical.Algebra.Group.Subgroup open import Cubical.Algebra.Group.QuotientGroup diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/RP2wedgeS1.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/RP2wedgeS1.agda index 3d31c7e09c..cc854b7092 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/RP2wedgeS1.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/RP2wedgeS1.agda @@ -29,6 +29,7 @@ open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.AbGroup.Instances.Unit +open import Cubical.Algebra.AbGroup.Instances.IntMod open import Cubical.HITs.SetTruncation as ST open import Cubical.HITs.Truncation as TR diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/Sn.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/Sn.agda index e71c8db909..92dea22524 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/Sn.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/Sn.agda @@ -22,8 +22,9 @@ open import Cubical.Foundations.Pointed open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order open import Cubical.Data.Unit -open import Cubical.Data.Bool +open import Cubical.Data.Bool hiding (_≟_) open import Cubical.Data.Sigma open import Cubical.Algebra.Group.MorphismProperties @@ -286,6 +287,44 @@ module _ (G : AbGroup ℓ) where (isConnectedEM (suc (suc n))) (0ₖ _) (f north) .fst))) isContrUnit*) snd (Hⁿ[Sᵐ⁺ⁿ,G]≅0 (suc n) m) = makeIsGroupHom λ _ _ → refl + open import Cubical.Data.Sum as ⊎ + open import Cubical.Data.Empty as ⊥ + open import Cubical.Relation.Nullary + open import Cubical.Cohomology.EilenbergMacLane.Groups.Connected + + H⁰[Sⁿ,G]≅G : (n : ℕ) → AbGroupEquiv (coHomGr 0 G (S₊ (suc n))) G + H⁰[Sⁿ,G]≅G n = H⁰conn + (∣ ptSn (suc n) ∣ₕ + , (TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (sphereElim n (λ _ → isProp→isOfHLevelSuc n (isOfHLevelTrunc 2 _ _)) + refl))) G + + Hⁿ[Sᵐ,G]Full : (n m : ℕ) + → (((n ≡ 0) ⊎ (n ≡ suc m)) + → AbGroupEquiv (coHomGr n G (S₊ (suc m))) G) + × ((¬ n ≡ 0) × (¬ (n ≡ suc m)) + → AbGroupEquiv (coHomGr n G (S₊ (suc m))) (trivialAbGroup {ℓ-zero})) + fst (Hⁿ[Sᵐ,G]Full zero m) _ = H⁰[Sⁿ,G]≅G m + snd (Hⁿ[Sᵐ,G]Full zero m) p = ⊥.rec (fst p refl) + Hⁿ[Sᵐ,G]Full (suc n) m with (n ≟ m) + ... | lt x = (λ { (inl x) → ⊥.rec (snotz x) + ; (inr p) → ⊥.rec (¬m = A[X1,···,Xn]/ ℤ/2CommRing 1 0 3 + + open GradedRing-⊕HIT-index + NatMonoid + (Poly ℤ/2CommRing) + (λ n → CommRing→AbGroup (PolyCommRing ℤ/2CommRing n) .snd) + + open RingStr (snd (H*R ℤ/2Ring RP²)) + renaming + ( 0r to 0H*RP + ; 1r to 1H*RP + ; _+_ to _+H*RP_ + ; -_ to -H*RP_ + ; _·_ to _cupS_ + ; +Assoc to +H*RPAssoc + ; +IdR to +H*RPIdR + ; +Comm to +H*RPComm + ; ·Assoc to ·H*RPAssoc + ; is-set to isSetH*RP + ; ·DistR+ to ·DistR+H* + ; ·DistL+ to ·DistL+H*) + + open RingStr (snd ℤ₂[X]) + renaming + ( 0r to 0Z₂X + ; 1r to 1Z₂X + ; _+_ to _+Z₂X_ + ; -_ to -Z₂X_ + ; _·_ to _·Z₂X_ + ; +Assoc to +Z₂XAssoc + ; +IdR to +Z₂XIdR + ; +Comm to +Z₂XComm + ; ·Assoc to ·Z₂XAssoc + ; is-set to isSetZ₂X + ; ·DistR+ to ·DistR+Z₂ + ; ·DistL+ to ·DistL+Z₂) + + private + H⁰[RP²,ℤ/2]≅ℤ/2 = H⁰[RP²,G]≅G ℤ/2 + + α* : fst (H*R ℤ/2Ring RP²) + α* = base 1 α + + γ* : fst (H*R ℤ/2Ring RP²) + γ* = base 2 γ + + α*²≡γ* : α* cupS α* ≡ γ* + α*²≡γ* = cong (base 2) α²=γ + + α-isGen : base 1 (invEq (fst (H¹[RP²,ℤ/2]≅ℤ/2)) fone) ≡ α* + α-isGen = cong (base 1) + (cong (invEq (fst (H¹[RP²,ℤ/2]≅ℤ/2))) (sym α↦1) + ∙ retEq (fst (H¹[RP²,ℤ/2]≅ℤ/2)) α) + + γ-isGen : base 2 (invEq (fst (H²[RP²,ℤ/2]≅ℤ/2)) fone) ≡ γ* + γ-isGen = cong (base 2) + (cong (invEq (fst (H²[RP²,ℤ/2]≅ℤ/2))) (sym γ↦1') + ∙ retEq (fst (H²[RP²,ℤ/2]≅ℤ/2)) γ) + + ℤ₂[X]→H*[RP²,ℤ₂]-main' : (n : ℕ) (g : ℤ/2 .fst) → coHom n ℤ/2 RP² + ℤ₂[X]→H*[RP²,ℤ₂]-main' zero g = invEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) g + ℤ₂[X]→H*[RP²,ℤ₂]-main' one g = invEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) g + ℤ₂[X]→H*[RP²,ℤ₂]-main' two g = invEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) g + ℤ₂[X]→H*[RP²,ℤ₂]-main' (suc (suc (suc n))) g = 0ₕ (3 +ℕ n) + + ℤ₂[X]→H*[RP²,ℤ₂]-main : Vec ℕ 1 → ℤ/2 .fst → fst (H*R ℤ/2Ring RP²) + ℤ₂[X]→H*[RP²,ℤ₂]-main (n ∷ []) g = base n (ℤ₂[X]→H*[RP²,ℤ₂]-main' n g) + + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₁ : (r : Vec ℕ 1) → + ℤ₂[X]→H*[RP²,ℤ₂]-main r fzero ≡ neutral + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₁ (zero ∷ []) = base-neutral 0 + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₁ (one ∷ []) = + cong (base 1) (IsGroupHom.pres1 (invGroupEquiv (H¹[RP²,ℤ/2]≅ℤ/2) .snd)) + ∙ base-neutral 1 + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₁ (two ∷ []) = + cong (base 2) (IsGroupHom.pres1 (invGroupEquiv (H²[RP²,ℤ/2]≅ℤ/2) .snd)) + ∙ base-neutral 2 + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₁ (suc (suc (suc s)) ∷ []) = base-neutral _ + + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₂ : (r : Vec ℕ 1) (a b : Fin 2) + → (ℤ₂[X]→H*[RP²,ℤ₂]-main r a add ℤ₂[X]→H*[RP²,ℤ₂]-main r b) + ≡ ℤ₂[X]→H*[RP²,ℤ₂]-main r (a +ₘ b) + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₂ (zero ∷ []) a b = base-add 0 _ _ + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₂ (one ∷ []) a b = base-add 1 _ _ + ∙ cong (base 1) (sym (IsGroupHom.pres· (invGroupEquiv (H¹[RP²,ℤ/2]≅ℤ/2) .snd) a b)) + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₂ (two ∷ []) a b = base-add 2 _ _ + ∙ cong (base 2) (sym (IsGroupHom.pres· (invGroupEquiv (H²[RP²,ℤ/2]≅ℤ/2) .snd) a b)) + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₂ (suc (suc (suc x)) ∷ []) a b = + cong (base (suc (suc (suc x))) (0ₕ (suc (suc (suc x)))) add_) (base-neutral _) + ∙ addRid _ + + ℤ₂[X]→H*[RP²,ℤ₂]-fun : ℤ₂[x] → fst (H*R ℤ/2Ring RP²) + ℤ₂[X]→H*[RP²,ℤ₂]-fun = + DS-Rec-Set.f _ _ _ _ isSetH*RP + neutral + ℤ₂[X]→H*[RP²,ℤ₂]-main + _add_ + addAssoc + addRid + addComm + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₁ + ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₂ + + open import Cubical.Foundations.Equiv.Properties + open import Cubical.Foundations.Transport +{- + Hⁿ[Sⁿ,G]≅G-pres⌣ : (n : ℕ) (G : Ring ℓ) (g : fst G) + (x : coHom (suc n) (Ring→AbGroup G) (S₊ (suc n))) + → Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n .fst .fst + (_⌣_ {G'' = G} {n = 0} {suc n} ∣ (λ _ → g) ∣₂ x) + ≡ RingStr._·_ (snd G) g (Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n .fst .fst x) + Hⁿ[Sⁿ,G]≅G-pres⌣ zero G g = + ST.elim (λ _ → isOfHLevelPath 2 (RingStr.is-set (snd G)) _ _) + (S¹-connElim (isConnectedEM 1) + (λ _ → RingStr.is-set (snd G) _ _) + embase + λ p → cong (H¹S¹→G (Ring→AbGroup G)) + (cong (_⌣_ {G'' = G} {n = 0} {1} ∣ (λ _ → g) ∣₂) + (cong (∣_∣₂ ∘ S¹Fun embase) + (sym (Iso.rightInv (Iso-EM-ΩEM+1 0) p)))) + ∙∙ help (ΩEM+1→EM 0 p) + ∙∙ cong (RingStr._·_ (snd G) g) (λ _ → ΩEM+1→EM 0 p)) + where + help : (h : fst G) + → H¹S¹→G (Ring→AbGroup G) + (_⌣_ {G'' = G} {n = 0} {1} ∣ (λ _ → g) ∣₂ + ∣ S¹Fun embase (emloop h) ∣₂) + ≡ RingStr._·_ (snd G) g h + help h = Iso.leftInv (Iso-EM-ΩEM+1 0) (RingStr._·_ (snd G) g h) + Hⁿ[Sⁿ,G]≅G-pres⌣ (suc n) G g = + ST.elim (λ _ → isOfHLevelPath 2 (RingStr.is-set (snd G)) _ _) + (Sⁿ-connElim n (isConnectedSubtr 2 (suc n) + (subst (λ k → isConnected (suc k) (EM (Ring→AbGroup G) (suc (suc n)))) + (+-comm 2 n) (isConnectedEM (2 +ℕ n)))) + (λ _ → RingStr.is-set (snd G) _ _) ∣ north ∣ₕ + λ f → cong (fst (Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n) .fst) + (cong ∣_∣₂ (funExt (λ x → sym (ΩEM+1→EM-distr⌣ₖ0n {G'' = G} (suc n) g _)))) + ∙ Hⁿ[Sⁿ,G]≅G-pres⌣ n G g _) +-} + + isContrH³⁺ⁿ[RP²,G] : ∀ {ℓ} (n : ℕ) (G : AbGroup ℓ) → isContr (coHom (3 +ℕ n) G RP²) + isContrH³⁺ⁿ[RP²,G] n G = + isOfHLevelRetractFromIso 0 + (equivToIso (fst (H³⁺ⁿ[RP²,G]≅0 G n))) isContrUnit* + + ℤ/2[X]→H*[RP²,ℤ/2]-pres· : (x y : ℤ₂[x]) + → ℤ₂[X]→H*[RP²,ℤ₂]-fun (RingStr._·_ (snd ℤ₂[X]) x y) + ≡ (ℤ₂[X]→H*[RP²,ℤ₂]-fun x cupS ℤ₂[X]→H*[RP²,ℤ₂]-fun y) + ℤ/2[X]→H*[RP²,ℤ/2]-pres· = + DS-Ind-Prop.f _ _ _ _ + (λ _ → isPropΠ λ _ → trunc _ _) + (λ _ → refl) + (λ v a → + DS-Ind-Prop.f _ _ _ _ + (λ _ → trunc _ _) + (cong (ℤ₂[X]→H*[RP²,ℤ₂]-fun) + (0RightAnnihilates ℤ₂[X] (base v a)) + ∙ sym (0RightAnnihilates (H*R ℤ/2Ring RP²) + (ℤ₂[X]→H*[RP²,ℤ₂]-fun (base v a)))) + (λ w b → main v w a b) + λ {x} {y} p q → cong ℤ₂[X]→H*[RP²,ℤ₂]-fun (·DistR+Z₂ (base v a) x y) + ∙ cong₂ _add_ p q + ∙ sym (·DistR+H* (ℤ₂[X]→H*[RP²,ℤ₂]-fun (base v a)) + (ℤ₂[X]→H*[RP²,ℤ₂]-fun x) + (ℤ₂[X]→H*[RP²,ℤ₂]-fun y))) + λ {x} {y} p q z + → cong (ℤ₂[X]→H*[RP²,ℤ₂]-fun) + (·DistL+Z₂ x y z) + ∙ cong₂ _add_ (p z) (q z) + where + main-main : (v w : Vec ℕ 1) + → ℤ₂[X]→H*[RP²,ℤ₂]-fun (base v fone ·Z₂X base w fone) + ≡ ℤ₂[X]→H*[RP²,ℤ₂]-fun (base v fone) + cupS (ℤ₂[X]→H*[RP²,ℤ₂]-fun (base w fone)) + main-main (zero ∷ []) (w ∷ []) = + cong (ℤ₂[X]→H*[RP²,ℤ₂]-fun) + (RingStr.·IdL (snd ℤ₂[X]) (base (w ∷ []) fone)) + ∙ sym (RingStr.·IdL (snd (H*R ℤ/2Ring RP²)) + (ℤ₂[X]→H*[RP²,ℤ₂]-fun (base (w ∷ []) fone))) + main-main (suc v ∷ []) (zero ∷ []) = + cong (ℤ₂[X]→H*[RP²,ℤ₂]-fun) + (RingStr.·IdR (snd ℤ₂[X]) (base (suc v ∷ []) fone)) + ∙ sym (RingStr.·IdR (snd (H*R ℤ/2Ring RP²)) + (ℤ₂[X]→H*[RP²,ℤ₂]-fun (base (suc v ∷ []) fone))) + main-main (one ∷ []) (one ∷ []) = + γ-isGen ∙ sym α*²≡γ* ∙ cong (λ x → x cupS x) (sym (α-isGen)) + main-main (one ∷ []) (suc (suc w₁) ∷ []) = + cong (base (suc (suc (suc w₁)))) + (isContr→isProp (isContrH³⁺ⁿ[RP²,G] w₁ ℤ/2) _ _) + main-main (suc (suc v) ∷ []) (suc w ∷ []) = + cong (base (suc (suc v) +ℕ suc w)) + (isContr→isProp + (subst (λ k → isContr (coHom (suc (suc k)) ℤ/2 RP²)) (sym (+-suc v w)) + (isContrH³⁺ⁿ[RP²,G] (v +ℕ w) ℤ/2)) _ (0ₕ _)) + ∙ (base-neutral _ + ∙ sym (base-neutral _)) + ∙ cong (base (suc (suc v) +' suc w)) + (isContr→isProp (isContrH³⁺ⁿ[RP²,G] (v +ℕ w) ℤ/2) (0ₕ _) _) + + main : (v w : Vec ℕ 1) (a b : ℤ/2 .fst) + → ℤ₂[X]→H*[RP²,ℤ₂]-fun (base v a ·Z₂X base w b) + ≡ ℤ₂[X]→H*[RP²,ℤ₂]-fun (base v a) cupS ℤ₂[X]→H*[RP²,ℤ₂]-fun (base w b) + main v w = ℤ/2-elim + (λ b → + (cong ℤ₂[X]→H*[RP²,ℤ₂]-fun + (cong (_·Z₂X base w b) (base-neutral v) + ∙ 0LeftAnnihilates ℤ₂[X] (base w b)) + ∙ cong₂ _cupS_ (cong (ℤ₂[X]→H*[RP²,ℤ₂]-fun) + (sym (base-neutral v))) refl)) + (ℤ/2-elim + (cong ℤ₂[X]→H*[RP²,ℤ₂]-fun + (cong (base v fone ·Z₂X_) (base-neutral w) + ∙ 0RightAnnihilates ℤ₂[X] (base w fone)) + ∙ sym (0RightAnnihilates (H*R ℤ/2Ring RP²) + (ℤ₂[X]→H*[RP²,ℤ₂]-fun (base v fone))) + ∙ cong₂ _cupS_ + (λ _ → ℤ₂[X]→H*[RP²,ℤ₂]-fun (base v fone)) + (cong (ℤ₂[X]→H*[RP²,ℤ₂]-fun) (sym (base-neutral w)))) + (main-main v w)) + + ℤ₂[X]→H*[RP²,ℤ₂] : RingHom ℤ₂[X] (H*R ℤ/2Ring RP²) + fst ℤ₂[X]→H*[RP²,ℤ₂] = ℤ₂[X]→H*[RP²,ℤ₂]-fun + snd ℤ₂[X]→H*[RP²,ℤ₂] = makeIsRingHom refl (λ _ _ → refl) ℤ/2[X]→H*[RP²,ℤ/2]-pres· + + open Quotient-FGideal-CommRing-Ring (PolyCommRing ℤ/2CommRing 1) (H*R ℤ/2Ring RP²) + ℤ₂[X]→H*[RP²,ℤ₂] ( ℤ/2CommRing 1 0 3) + (λ { zero → base-neutral _}) + + ℤ₂[X]/→H*[RP²,ℤ₂] : RingHom (CommRing→Ring ℤ₂[X]/) (H*R ℤ/2Ring RP²) + ℤ₂[X]/→H*[RP²,ℤ₂] = inducedHom + + H*[RP²,ℤ₂]→ℤ₂[X]/-fun' : (r : ℕ) → coHom r ℤ/2 RP² → ℤ/2 .fst + H*[RP²,ℤ₂]→ℤ₂[X]/-fun' zero x = H⁰[RP²,ℤ/2]≅ℤ/2 .fst .fst x + H*[RP²,ℤ₂]→ℤ₂[X]/-fun' one x = H¹[RP²,ℤ/2]≅ℤ/2 .fst .fst x + H*[RP²,ℤ₂]→ℤ₂[X]/-fun' two x = H²[RP²,ℤ/2]≅ℤ/2 .fst .fst x + H*[RP²,ℤ₂]→ℤ₂[X]/-fun' (suc (suc (suc r))) x = fzero + + + H*[RP²,ℤ₂]→ℤ₂[X]/-fun : (r : ℕ) → coHom r ℤ/2 RP² → ℤ₂[X]/ .fst + H*[RP²,ℤ₂]→ℤ₂[X]/-fun r x = [ base (r ∷ []) (H*[RP²,ℤ₂]→ℤ₂[X]/-fun' r x) ] + + H*[RP²,ℤ₂]→ℤ₂[X]/ : H*R ℤ/2Ring RP² .fst → ℤ₂[X]/ .fst + H*[RP²,ℤ₂]→ℤ₂[X]/ = DS-Rec-Set.f _ _ _ _ squash/ + [ neutral ] + H*[RP²,ℤ₂]→ℤ₂[X]/-fun + (CommRingStr._+_ (snd ℤ₂[X]/)) + (CommRingStr.+Assoc (snd ℤ₂[X]/)) + (CommRingStr.+IdR (snd ℤ₂[X]/)) + (CommRingStr.+Comm (snd ℤ₂[X]/)) + (λ { zero → cong [_] (base-neutral (0 ∷ [])) + ; one → cong [_] (cong (base (1 ∷ [])) (IsGroupHom.pres1 (snd (H¹[RP²,ℤ/2]≅ℤ/2))) + ∙ base-neutral (1 ∷ [])) + ; two → cong [_] (cong (base (2 ∷ [])) (IsGroupHom.pres1 (snd (H²[RP²,ℤ/2]≅ℤ/2))) + ∙ base-neutral (2 ∷ [])) + ; (suc (suc (suc r))) → cong [_] (base-neutral _)}) + (λ { zero a b → cong [_] (base-add (0 ∷ []) _ _ + ∙ cong (base (0 ∷ [])) + (sym (IsGroupHom.pres· (snd (H⁰[RP²,ℤ/2]≅ℤ/2)) a b))) + ; one a b → cong [_] (base-add (1 ∷ []) _ _ + ∙ cong (base (1 ∷ [])) + (sym (IsGroupHom.pres· (snd (H¹[RP²,ℤ/2]≅ℤ/2)) a b))) + ; two a b → cong [_] (base-add (2 ∷ []) _ _ + ∙ cong (base (2 ∷ [])) + (sym (IsGroupHom.pres· (snd (H²[RP²,ℤ/2]≅ℤ/2)) a b))) + ; (suc (suc (suc r))) a b + → cong [_] (cong₂ _add_ refl (base-neutral (3 +ℕ r ∷ [])) + ∙ addRid _)}) + + ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ : (x : _) + → H*[RP²,ℤ₂]→ℤ₂[X]/ (ℤ₂[X]/→H*[RP²,ℤ₂] .fst x) ≡ x + ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ = SQ.elimProp (λ _ → squash/ _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) + refl + (λ { (zero ∷ []) a → cong [_] (cong (base (zero ∷ [])) + (secEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) a)) + ; (one ∷ []) a → cong [_] (cong (base (one ∷ [])) + (secEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) a)) + ; (two ∷ []) a → cong [_] (cong (base (two ∷ [])) + (secEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) a)) + ; (suc (suc (suc r)) ∷ []) → ℤ/2-elim refl + (sym (eq/ _ _ ∣ (λ {zero → base (r ∷ []) fone}) + , ((cong₂ _add_ refl (base-neutral _) + ∙ addRid _ + ∙ λ i → base (+-comm 3 r i ∷ []) fone)) + ∙ sym (addRid _) ∣₁))}) + (λ p q → cong₂ (CommRingStr._+_ (snd ℤ₂[X]/)) p q)) + + H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] : (x : H*R ℤ/2Ring RP² .fst) + → ℤ₂[X]/→H*[RP²,ℤ₂] .fst (H*[RP²,ℤ₂]→ℤ₂[X]/ x) ≡ x + H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] = DS-Ind-Prop.f _ _ _ _ + (λ _ → trunc _ _) + refl + (λ { zero x → cong (base zero) (retEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) x) + ; one x → cong (base one) (retEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) x) + ; two x → cong (base two) (retEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) x) + ; (suc (suc (suc r))) x → cong (base (3 +ℕ r)) + (isContr→isProp (isContrH³⁺ⁿ[RP²,G] r ℤ/2) _ _)}) + λ {x} {y} p q + → (IsRingHom.pres+ (ℤ₂[X]/→H*[RP²,ℤ₂] .snd) _ _ ∙ cong₂ _add_ p q) + + ℤ₂[X]/≅H*[RP²,ℤ₂] : RingEquiv (CommRing→Ring ℤ₂[X]/) (H*R ℤ/2Ring RP²) + fst ℤ₂[X]/≅H*[RP²,ℤ₂] = + isoToEquiv (iso (ℤ₂[X]/→H*[RP²,ℤ₂] .fst) H*[RP²,ℤ₂]→ℤ₂[X]/ + H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/) + snd ℤ₂[X]/≅H*[RP²,ℤ₂] = ℤ₂[X]/→H*[RP²,ℤ₂] .snd + + H*[RP²,ℤ₂]≅ℤ₂[X]/ : RingEquiv (H*R ℤ/2Ring RP²) (CommRing→Ring ℤ₂[X]/) + H*[RP²,ℤ₂]≅ℤ₂[X]/ = RingEquivs.invRingEquiv ℤ₂[X]/≅H*[RP²,ℤ₂] diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda index 3c6011cfd9..0e632396f1 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda @@ -57,8 +57,9 @@ open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.IntMod -open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup hiding (_[_]) open import Cubical.Algebra.AbGroup.Instances.Unit +open import Cubical.Algebra.AbGroup.Instances.IntMod open import Cubical.Algebra.DirectSum.DirectSumHIT.Base open import Cubical.Algebra.Ring open import Cubical.Algebra.AbGroup.TensorProduct diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda new file mode 100644 index 0000000000..f5d2bc9a2e --- /dev/null +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda @@ -0,0 +1,428 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Cohomology.EilenbergMacLane.Rings.Sn where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Path +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2wedgeS1 +open import Cubical.Cohomology.EilenbergMacLane.Base +open import Cubical.Cohomology.EilenbergMacLane.Groups.Connected +open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2 +open import Cubical.Cohomology.EilenbergMacLane.Groups.Sn +open import Cubical.Cohomology.EilenbergMacLane.Groups.Wedge +open import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle +open import Cubical.Cohomology.EilenbergMacLane.CupProduct +open import Cubical.Cohomology.EilenbergMacLane.Rings.KleinBottle hiding (α↦1) + +open import Cubical.Homotopy.EilenbergMacLane.GroupStructure +open import Cubical.Homotopy.EilenbergMacLane.Order2 +open import Cubical.Homotopy.EilenbergMacLane.Properties +open import Cubical.Homotopy.EilenbergMacLane.Base +open import Cubical.Homotopy.EilenbergMacLane.CupProduct +open import Cubical.Homotopy.EilenbergMacLane.CupProductTensor + renaming (_⌣ₖ_ to _⌣ₖ⊗_ ; ⌣ₖ-0ₖ to ⌣ₖ-0ₖ⊗ ; 0ₖ-⌣ₖ to 0ₖ-⌣ₖ⊗) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Cohomology.EilenbergMacLane.RingStructure +open import Cubical.Cohomology.EilenbergMacLane.Rings.Z2-properties + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Fin +open import Cubical.Data.Sum +open import Cubical.Data.Fin.Arithmetic +open import Cubical.Data.Sigma +open import Cubical.Data.Vec +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.FinData renaming (Fin to FinI) + +open import Cubical.HITs.KleinBottle renaming (rec to KleinFun) +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as TR +open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _sq/_) +open import Cubical.HITs.EilenbergMacLane1 hiding (elimProp ; elimSet) +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 renaming (rec to S¹Fun) +open import Cubical.HITs.Sn +open import Cubical.HITs.RPn +open import Cubical.HITs.Wedge +open import Cubical.Algebra.GradedRing.DirectSumHIT + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.AbGroup hiding (_[_]) +open import Cubical.Algebra.AbGroup.Instances.Unit +open import Cubical.Algebra.DirectSum.DirectSumHIT.Base +open import Cubical.Algebra.Ring +open import Cubical.Algebra.AbGroup.TensorProduct +open import Cubical.Algebra.Monoid + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.Quotient +open import Cubical.Algebra.CommRing.Instances.IntMod +open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly +open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly-Quotient +open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly-notationZ2 + +open import Cubical.Algebra.Monoid.Instances.Nat + +open Iso +open PlusBis +open RingTheory + +open AbGroupStr + +-- open import Cubical.Agebra.Monoids.Instance.NatVec + +module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where + G[x] = Poly G 1 + + G[X] = CommRing→Ring (PolyCommRing G 1) + + G[X]/ = A[X1,···,Xn]/ G 1 0 2 + + private + GRing = CommRing→Ring G + GAb = CommRing→AbGroup G + + open GradedRing-⊕HIT-index + NatMonoid + (Poly G) + (λ n → CommRing→AbGroup (PolyCommRing G n) .snd) + + open RingStr (snd (H*R GRing (S₊ (suc n)))) + renaming + ( 0r to 0H*S + ; 1r to 1H*S + ; _+_ to _+H*S_ + ; -_ to -H*S_ + ; _·_ to _cupS_ + ; +Assoc to +H*SAssoc + ; +IdR to +H*SIdR + ; +Comm to +H*SComm + ; ·Assoc to ·H*SAssoc + ; is-set to isSetH*S + ; ·DistR+ to ·DistR+H* + ; ·DistL+ to ·DistL+H*) + + open RingStr (snd G[X]) + renaming + ( 0r to 0GX + ; 1r to 1GX + ; _+_ to _+GX_ + ; -_ to -GX_ + ; _·_ to _·GX_ + ; +Assoc to +GXAssoc + ; +IdR to +GXIdR + ; +Comm to +GXComm + ; ·Assoc to ·GXAssoc + ; is-set to isSetGX + ; ·DistR+ to ·DistR+G + ; ·DistL+ to ·DistL+G) + + G[X]→H*[Sⁿ,G]-main : Vec ℕ 1 → fst G → fst (H*R GRing (S₊ (suc n))) + G[X]→H*[Sⁿ,G]-main (zero ∷ []) g = base 0 (invEquiv (H⁰[Sⁿ,G]≅G GAb n .fst) .fst g) + G[X]→H*[Sⁿ,G]-main (one ∷ []) g = base (suc n) (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g) + G[X]→H*[Sⁿ,G]-main (suc (suc x) ∷ []) g = ⊕HIT.neutral + + G[X]→H*[Sⁿ,G]-main-coh₁ : (r : Vec ℕ 1) + → G[X]→H*[Sⁿ,G]-main r (CommRingStr.0r (snd G)) ≡ neutral + G[X]→H*[Sⁿ,G]-main-coh₁ (zero ∷ []) = base-neutral 0 + G[X]→H*[Sⁿ,G]-main-coh₁ (one ∷ []) = + cong (base (suc n)) (IsGroupHom.pres1 (invGroupEquiv (Hⁿ[Sⁿ,G]≅G GAb n) .snd) ) + ∙ base-neutral (suc n) + G[X]→H*[Sⁿ,G]-main-coh₁ (suc (suc x) ∷ []) = refl + + G[X]→H*[Sⁿ,G]-main-coh₂ : (r : Vec ℕ 1) (a b : fst G) + → (G[X]→H*[Sⁿ,G]-main r a add G[X]→H*[Sⁿ,G]-main r b) + ≡ G[X]→H*[Sⁿ,G]-main r ((snd (CommRing→Ring G) RingStr.+ a) b) + G[X]→H*[Sⁿ,G]-main-coh₂ (zero ∷ []) a b = base-add 0 ∣ (λ _ → a) ∣₂ ∣ (λ _ → b) ∣₂ + G[X]→H*[Sⁿ,G]-main-coh₂ (one ∷ []) a b = + base-add (suc n) _ _ + ∙ cong (base (suc n)) + (sym (IsGroupHom.pres· (invGroupEquiv (Hⁿ[Sⁿ,G]≅G GAb n) .snd) a b)) + G[X]→H*[Sⁿ,G]-main-coh₂ (suc (suc x) ∷ []) a b = addRid neutral + + G[X]→H*[Sⁿ,G]-fun : G[x] → fst (H*R GRing (S₊ (suc n))) + G[X]→H*[Sⁿ,G]-fun = + DS-Rec-Set.f _ _ _ _ isSetH*S + neutral + G[X]→H*[Sⁿ,G]-main + _add_ + addAssoc + addRid + addComm + G[X]→H*[Sⁿ,G]-main-coh₁ + G[X]→H*[Sⁿ,G]-main-coh₂ + + open import Cubical.Foundations.Equiv.Properties + open import Cubical.Foundations.Transport + + Hⁿ[Sⁿ,G]≅G-pres⌣ : (n : ℕ) (G : Ring ℓ) (g : fst G) + (x : coHom (suc n) (Ring→AbGroup G) (S₊ (suc n))) + → Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n .fst .fst + (_⌣_ {G'' = G} {n = 0} {suc n} ∣ (λ _ → g) ∣₂ x) + ≡ RingStr._·_ (snd G) g (Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n .fst .fst x) + Hⁿ[Sⁿ,G]≅G-pres⌣ zero G g = + ST.elim (λ _ → isOfHLevelPath 2 (RingStr.is-set (snd G)) _ _) + (S¹-connElim (isConnectedEM 1) + (λ _ → RingStr.is-set (snd G) _ _) + embase + λ p → cong (H¹S¹→G (Ring→AbGroup G)) + (cong (_⌣_ {G'' = G} {n = 0} {1} ∣ (λ _ → g) ∣₂) + (cong (∣_∣₂ ∘ S¹Fun embase) + (sym (Iso.rightInv (Iso-EM-ΩEM+1 0) p)))) + ∙∙ help (ΩEM+1→EM 0 p) + ∙∙ cong (RingStr._·_ (snd G) g) (λ _ → ΩEM+1→EM 0 p)) + where + help : (h : fst G) + → H¹S¹→G (Ring→AbGroup G) + (_⌣_ {G'' = G} {n = 0} {1} ∣ (λ _ → g) ∣₂ + ∣ S¹Fun embase (emloop h) ∣₂) + ≡ RingStr._·_ (snd G) g h + help h = Iso.leftInv (Iso-EM-ΩEM+1 0) (RingStr._·_ (snd G) g h) + Hⁿ[Sⁿ,G]≅G-pres⌣ (suc n) G g = + ST.elim (λ _ → isOfHLevelPath 2 (RingStr.is-set (snd G)) _ _) + (Sⁿ-connElim n (isConnectedSubtr 2 (suc n) + (subst (λ k → isConnected (suc k) (EM (Ring→AbGroup G) (suc (suc n)))) + (+-comm 2 n) (isConnectedEM (2 +ℕ n)))) + (λ _ → RingStr.is-set (snd G) _ _) ∣ north ∣ₕ + λ f → cong (fst (Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n) .fst) + (cong ∣_∣₂ (funExt (λ x → sym (ΩEM+1→EM-distr⌣ₖ0n {G'' = G} (suc n) g _)))) + ∙ Hⁿ[Sⁿ,G]≅G-pres⌣ n G g _) + + G[X]→H*[Sⁿ,G]-pres· : (x y : G[x]) + → G[X]→H*[Sⁿ,G]-fun (x ·GX y) + ≡ (G[X]→H*[Sⁿ,G]-fun x cupS G[X]→H*[Sⁿ,G]-fun y) + G[X]→H*[Sⁿ,G]-pres· = DS-Ind-Prop.f _ _ _ _ + (λ _ → isPropΠ λ _ → isSetH*S _ _) + (λ _ → refl) + G[X]→H*[Sⁿ,G]-pres·+ + λ {x} {y} p q s → cong (G[X]→H*[Sⁿ,G]-fun) (·DistL+G x y s) + ∙ cong₂ _add_ (p s) (q s) + ∙ sym (·DistL+H* (G[X]→H*[Sⁿ,G]-fun x) + (G[X]→H*[Sⁿ,G]-fun y) (G[X]→H*[Sⁿ,G]-fun s)) + where + main₀₁ : (g g' : fst G) → _ ≡ _ + main₀₁ g g' = + (cong (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst)) + (cong₂ (CommRingStr._·_ (snd G)) refl + (sym (secEq (Hⁿ[Sⁿ,G]≅G (Ring→AbGroup GRing) n .fst) g'))) + ∙∙ cong (invEq (Hⁿ[Sⁿ,G]≅G (Ring→AbGroup GRing) n .fst)) + (sym (Hⁿ[Sⁿ,G]≅G-pres⌣ n GRing g (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g'))) + ∙∙ retEq (Hⁿ[Sⁿ,G]≅G (Ring→AbGroup GRing) n .fst) _) + + main : (v v' : Vec ℕ 1) (g g' : fst G) + → G[X]→H*[Sⁿ,G]-fun (base v g ·GX base v' g') + ≡ (G[X]→H*[Sⁿ,G]-fun (base v g) cupS G[X]→H*[Sⁿ,G]-fun (base v' g')) + main (zero ∷ []) (zero ∷ []) g g' = refl + main (zero ∷ []) (one ∷ []) g g' = cong (base (suc n)) (main₀₁ g g') + main (zero ∷ []) (suc (suc w) ∷ []) g g' = refl + main (one ∷ []) (zero ∷ []) g g' = + (λ j → base (+'-comm 0 (suc n) j) + (transp (λ k → coHom (+'-comm 0 (suc n) (k ∧ j)) + GAb (S₊ (suc n))) (~ j) + (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) (CommRingStr._·_ (snd G) g g')))) + ∙ cong (base (suc n +' zero)) + (cong (subst (λ k → coHom k GAb (S₊ (suc n))) (+'-comm 0 (suc n))) + (cong (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst)) (CommRingStr.·Comm (snd G) g g') + ∙ main₀₁ g' g + ∙ sym (-ₕ^[_·_]-even (suc n) 0 (inr tt) _)) + ∙ sym (comm⌣ (suc n) 0 (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g) + (invEquiv (H⁰[Sⁿ,G]≅G GAb n .fst) .fst g'))) + main (one ∷ []) (one ∷ []) g g' = + sym (base-neutral (suc n +' suc n)) + ∙ cong (base (suc n +' suc n)) + (isContr→isProp + (subst (λ k → isContr (coHom (suc k) GAb (S₊ (suc n)))) + (+-suc n n) (isContr-Hᵐ⁺ⁿ[Sⁿ,G] GAb n n)) _ _) + main (one ∷ []) (suc (suc w) ∷ []) g g' = refl + main (suc (suc v) ∷ []) (w ∷ []) g g' = refl + + G[X]→H*[Sⁿ,G]-pres·+ : (r : Vec ℕ 1) (a : fst G) (y : G[x]) + → G[X]→H*[Sⁿ,G]-fun (base r a ·GX y) + ≡ G[X]→H*[Sⁿ,G]-fun (base r a) cupS G[X]→H*[Sⁿ,G]-fun y + G[X]→H*[Sⁿ,G]-pres·+ v a = DS-Ind-Prop.f _ _ _ _ (λ _ → isSetH*S _ _) + (cong (G[X]→H*[Sⁿ,G]-fun) + (0RightAnnihilates G[X] (base v a)) + ∙ sym (0RightAnnihilates + (H*R GRing (S₊ (suc n))) (G[X]→H*[Sⁿ,G]-fun (base v a) ))) + (λ r t → main v r a t) + λ {x} {y} p q → cong (G[X]→H*[Sⁿ,G]-fun) (·DistR+G (base v a) x y) + ∙ cong₂ _add_ p q + ∙ sym (·DistR+H* (G[X]→H*[Sⁿ,G]-fun (base v a)) + (G[X]→H*[Sⁿ,G]-fun x) (G[X]→H*[Sⁿ,G]-fun y)) + + G[X]→H*[Sⁿ,G] : RingHom G[X] (H*R GRing (S₊ (suc n))) + fst G[X]→H*[Sⁿ,G] = G[X]→H*[Sⁿ,G]-fun + snd G[X]→H*[Sⁿ,G] = makeIsRingHom refl (λ _ _ → refl) G[X]→H*[Sⁿ,G]-pres· + + open Quotient-FGideal-CommRing-Ring (PolyCommRing G 1) (H*R GRing (S₊ (suc n))) + G[X]→H*[Sⁿ,G] ( G 1 0 2) + (λ { zero → refl}) + + G[X]/→H*[Sⁿ,G] : RingHom (CommRing→Ring G[X]/) (H*R GRing (S₊ (suc n))) + G[X]/→H*[Sⁿ,G] = inducedHom + + H*[Sⁿ,G]→G[X]/-fun^ : (m : ℕ) → Trichotomy n m + → (x : coHom (suc m) GAb (S₊ (suc n))) → G[X]/ .fst + H*[Sⁿ,G]→G[X]/-fun^ m (lt p) x = [ neutral ] + H*[Sⁿ,G]→G[X]/-fun^ m (eq e) x = + [ base (1 ∷ []) (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst (subst (λ k → coHom (suc k) GAb (S₊ (suc n))) (sym e) x)) ] + H*[Sⁿ,G]→G[X]/-fun^ m (gt q) x = [ neutral ] + + H*[Sⁿ,G]→G[X]/-fun : (m : ℕ) (x : coHom m GAb (S₊ (suc n))) + → G[X]/ .fst + H*[Sⁿ,G]→G[X]/-fun zero x = [ base (0 ∷ []) (H⁰[Sⁿ,G]≅G GAb n .fst .fst x) ] + H*[Sⁿ,G]→G[X]/-fun (suc m) x = H*[Sⁿ,G]→G[X]/-fun^ m (n ≟ m) x + + H*[Sⁿ,G]→G[X]/-fun-coh₁ : (r : ℕ) → H*[Sⁿ,G]→G[X]/-fun r (0ₕ r) ≡ [ neutral ] + H*[Sⁿ,G]→G[X]/-fun-coh₁ zero = + cong [_] (cong (base (0 ∷ [])) (IsGroupHom.pres1 (H⁰[Sⁿ,G]≅G GAb n .snd)) + ∙ base-neutral (0 ∷ [])) + H*[Sⁿ,G]→G[X]/-fun-coh₁ (suc r) = help (n ≟ r) + where + help : (e : _) → H*[Sⁿ,G]→G[X]/-fun^ r e (0ₕ (suc r)) ≡ [ neutral ] + help (lt x) = refl + help (eq x) = cong [_] + (cong (base (1 ∷ [])) + (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) + (λ j → transp (λ i → coHom (suc (x (~ i ∧ ~ j))) GAb (S₊ (suc n))) + j (0ₕ (suc (x (~ j))))) + ∙ IsGroupHom.pres1 (Hⁿ[Sⁿ,G]≅G GAb n .snd)) + ∙ base-neutral (1 ∷ [])) + help (gt x) = refl + + H*[Sⁿ,G]→G[X]/-fun-coh₂ : (r : ℕ) (x y : coHom r GAb (S₊ (suc n))) + → CommRingStr._+_ (snd (G[X]/)) + (H*[Sⁿ,G]→G[X]/-fun r x) (H*[Sⁿ,G]→G[X]/-fun r y) + ≡ H*[Sⁿ,G]→G[X]/-fun r (x +ₕ y) + H*[Sⁿ,G]→G[X]/-fun-coh₂ zero x y = + cong [_] (base-add (0 ∷ []) _ _ + ∙ cong (base (0 ∷ [])) (sym (IsGroupHom.pres· (H⁰[Sⁿ,G]≅G GAb n .snd) _ _))) + H*[Sⁿ,G]→G[X]/-fun-coh₂ (suc r) x y = help (n ≟ r) x y + where + help : (p : _) (x y : coHom (suc r) GAb (S₊ (suc n))) + → CommRingStr._+_ (snd (G[X]/)) + (H*[Sⁿ,G]→G[X]/-fun^ r p x) + (H*[Sⁿ,G]→G[X]/-fun^ r p y) + ≡ H*[Sⁿ,G]→G[X]/-fun^ r p (x +ₕ y) + help (lt p) x y = cong [_] (addRid neutral) + help (eq p) = help' r p + where + help' : (r : ℕ) (p : n ≡ r) (x y : _) + → CommRingStr._+_ (snd G[X]/) + (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) x) + (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) y) + ≡ H*[Sⁿ,G]→G[X]/-fun^ r (eq p) (x +ₕ y) + help' = J> λ a b → cong [_] + (cong₂ (λ x y → (base (1 ∷ []) x) add (base (1 ∷ []) y)) + (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl a)) + (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl b)) + ∙ base-add _ _ _ + ∙ cong (base (1 ∷ [])) + (sym (IsGroupHom.pres· (Hⁿ[Sⁿ,G]≅G GAb n .snd) _ _) + ∙ cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (sym (transportRefl (a +ₕ b))))) + help (gt p) x y = cong [_] (addRid neutral) + + H*[Sⁿ,G]→G[X]/ : fst (H*R GRing (S₊ (suc n))) → G[X]/ .fst + H*[Sⁿ,G]→G[X]/ = DS-Rec-Set.f _ _ _ _ squash/ + [ neutral ] + H*[Sⁿ,G]→G[X]/-fun + (CommRingStr._+_ (snd (G[X]/))) + (λ _ _ _ → CommRingStr.+Assoc (snd (G[X]/)) _ _ _) + (λ _ → CommRingStr.+IdR (snd (G[X]/)) _) + (λ _ _ → CommRingStr.+Comm (snd (G[X]/)) _ _) + H*[Sⁿ,G]→G[X]/-fun-coh₁ + H*[Sⁿ,G]→G[X]/-fun-coh₂ + + H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] : (x : fst (H*R GRing (S₊ (suc n)))) + → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/ x) ≡ x + H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] = + DS-Ind-Prop.f _ _ _ _ (λ _ → trunc _ _) + refl + (λ { zero a → cong (base zero) (retEq (H⁰[Sⁿ,G]≅G GAb n .fst) a) + ; (suc r) → help r (n ≟ r)}) + λ p q → IsRingHom.pres+ (G[X]/→H*[Sⁿ,G] .snd) _ _ + ∙ cong₂ _add_ p q + where + help : (r : ℕ) (p : _) (a : _) + → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r p a) + ≡ base (suc r) a + help r (lt x) a = + sym (base-neutral (suc r)) + ∙ cong (base (suc r)) + (isContr→isProp + (subst (λ k → isContr (coHom k GAb (S₊ (suc n)))) + (cong suc (snd x)) + (isContr-Hᵐ⁺ⁿ[Sⁿ,G] GAb n (fst x))) _ _) + help r (eq x) a = + J (λ r x → (a : coHom (suc r) GAb (S₊ (suc n))) + → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r (eq x) a) + ≡ base (suc r) a) + (λ a → cong (base (suc n)) + (retEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) _ ∙ transportRefl a)) x a + help r (gt x) a = + sym (base-neutral (suc r)) + ∙ cong (base (suc r)) + (isContr→isProp + (subst (λ k → isContr (coHom (suc r) GAb (S₊ k))) + (cong suc (snd x)) + (isOfHLevelRetractFromIso 0 + (equivToIso (fst (Hⁿ[Sᵐ⁺ⁿ,G]≅0 GAb r (fst x)))) + isContrUnit*)) _ _) + + G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] : (x : G[X]/ .fst) + → H*[Sⁿ,G]→G[X]/ (G[X]/→H*[Sⁿ,G] .fst x) ≡ x + G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] = SQ.elimProp (λ _ → squash/ _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) + refl + (λ { (zero ∷ []) g → cong [_] (cong (base (0 ∷ [])) + (secEq (H⁰[Sⁿ,G]≅G GAb n .fst) g)) + ; (one ∷ []) g → h2 g _ + ; (suc (suc x) ∷ []) g + → sym (eq/ _ neutral ∣ hh x g + , (λ i → base (+-comm 2 x i ∷ []) (CommRingStr.·IdR (snd G) g (~ i)) + add neutral) ∣₁)}) + λ p q → cong₂ (CommRingStr._+_ (snd (G[X]/))) p q) + where + hh : (x : ℕ) (g : fst G) → FinVec (fst (PolyCommRing G 1)) 1 + hh x g zero = base (x ∷ []) g + + h2 : (g : fst G) (p : _) + → H*[Sⁿ,G]→G[X]/-fun^ n p + (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g) + ≡ [ base (one ∷ []) g ] + h2 g (lt x) = ⊥.rec (¬m≅H*[Sⁿ,G] : RingEquiv (CommRing→Ring G[X]/) (H*R GRing (S₊ (suc n))) + fst G[X]/≅H*[Sⁿ,G] = + isoToEquiv (iso (G[X]/→H*[Sⁿ,G] .fst) H*[Sⁿ,G]→G[X]/ + H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G]) + snd G[X]/≅H*[Sⁿ,G] = G[X]/→H*[Sⁿ,G] .snd + + H*[Sⁿ,G]≅G[X]/ : RingEquiv (H*R GRing (S₊ (suc n))) (CommRing→Ring G[X]/) + H*[Sⁿ,G]≅G[X]/ = RingEquivs.invRingEquiv G[X]/≅H*[Sⁿ,G] diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/Z2-properties.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/Z2-properties.agda index 326bddac69..b34c8179b9 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/Z2-properties.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/Z2-properties.agda @@ -17,6 +17,7 @@ open import Cubical.Algebra.DirectSum.DirectSumHIT.Base open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.IntMod +open import Cubical.Algebra.AbGroup.Instances.IntMod open import Cubical.Homotopy.EilenbergMacLane.Order2 diff --git a/Cubical/Homotopy/EilenbergMacLane/CupProduct.agda b/Cubical/Homotopy/EilenbergMacLane/CupProduct.agda index 74733bfdb0..7253f8ec0b 100644 --- a/Cubical/Homotopy/EilenbergMacLane/CupProduct.agda +++ b/Cubical/Homotopy/EilenbergMacLane/CupProduct.agda @@ -20,10 +20,12 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function open import Cubical.HITs.EilenbergMacLane1 open import Cubical.HITs.Susp -open import Cubical.HITs.Truncation +open import Cubical.HITs.Truncation as TR open import Cubical.Algebra.AbGroup.Base open import Cubical.Data.Nat hiding (_·_) renaming (elim to ℕelim ; _+_ to _+ℕ_) @@ -168,6 +170,48 @@ module _ {G'' : Ring ℓ} where ∙ cong (cong (inducedFun-EM TensorMultHom (suc (suc n) +' suc m))) (EM→ΩEM+1-distrₙsuc n m x y) + ΩEM+1→EM-distr⌣ₖ0n : (n : ℕ) + → (x : EM G' zero) (y : Path (EM G' (suc n)) _ _) + → _⌣ₖ_ {n = zero} {m = n} x (ΩEM+1→EM n y) + ≡ ΩEM+1→EM-gen n _ λ i → _⌣ₖ_ {n = zero} {suc n} x (y i) + ΩEM+1→EM-distr⌣ₖ0n zero x y = + sym (Iso.leftInv (Iso-EM-ΩEM+1 0) _) + ∙ cong (ΩEM+1→EM 0) + λ j i → _⌣ₖ_ {n = zero} {1} x (Iso.rightInv (Iso-EM-ΩEM+1 0) y j i) + ΩEM+1→EM-distr⌣ₖ0n (suc n) x y = + sym (Iso.leftInv (Iso-EM-ΩEM+1 (suc n)) _) + ∙ cong (ΩEM+1→EM (suc n)) (EM→ΩEM+1-distr⌣ₖ0n n x (ΩEM+1→EM (suc n) y)) + ∙ cong (ΩEM+1→EM (suc n)) (help n (ΩEM+1→EM (suc n) y) + ∙ cong (cong (_⌣ₖ_ {n = zero} {suc (suc n)} x)) + (Iso.rightInv (Iso-EM-ΩEM+1 (suc n)) y)) + where + help : (n : ℕ) (y : EM G' (suc n)) + → (λ i → _⌣ₖ_ {n = suc zero} {suc n} (EM→ΩEM+1 0 x i) y) + ≡ cong (_⌣ₖ_ {n = zero} {suc (suc n)} x) (EM→ΩEM+1 (suc n) y) + help zero a = + (rUnit _ + ∙ cong₂ _∙_ refl + (cong (cong (EMTensorMult (suc (suc zero)))) + (cong sym (sym (EM→ΩEM+1-0ₖ (suc zero))) + ∙ cong (sym ∘ EM→ΩEM+1 (suc zero)) + (sym (⌣ₖ-0ₖ⊗ {G' = G'} {H' = G'} zero (suc zero) x))))) + ∙ sym (cong-∙ ((λ y → _⌣ₖ_ {n = zero} {suc (suc zero)} x ∣ y ∣ₕ)) + (merid a) + (sym (merid embase))) + help (suc n) = TR.elim + (λ _ → isOfHLevelPath' (4 +ℕ n) + (isOfHLevelPath (5 +ℕ n) (hLevelEM G' (3 +ℕ n)) _ _) _ _) + λ a → rUnit _ + ∙ cong₂ _∙_ refl + (cong (cong (EMTensorMult (suc (suc (suc n))))) + (cong sym (sym (EM→ΩEM+1-0ₖ (suc (suc n)))) + ∙ cong (sym ∘ EM→ΩEM+1 (suc (suc n))) + (sym (⌣ₖ-0ₖ⊗ {G' = G'} {H' = G'} zero (suc (suc n)) x)))) + ∙ sym (cong-∙ ((λ y → _⌣ₖ_ {n = zero} {suc (suc (suc n))} x ∣ y ∣ₕ)) + (merid a) + (sym (merid north))) + + -- graded commutativity module _ {G'' : CommRing ℓ} where private @@ -227,3 +271,4 @@ syntax ⌣[]ₖ-syntax R x y = x ⌣[ R R]ₖ y syntax ⌣[]Cₖ-syntax R x y = x ⌣[ R R]Cₖ y syntax ⌣[,,]ₖ-syntax n m R x y = x ⌣[ R , n , m ]ₖ y syntax ⌣[,,]Cₖ-syntax n m R x y = x ⌣[ R , n , m ]Cₖ y + diff --git a/Cubical/Homotopy/EilenbergMacLane/Order2.agda b/Cubical/Homotopy/EilenbergMacLane/Order2.agda index bd09f17d52..3de91de139 100644 --- a/Cubical/Homotopy/EilenbergMacLane/Order2.agda +++ b/Cubical/Homotopy/EilenbergMacLane/Order2.agda @@ -37,6 +37,7 @@ open import Cubical.HITs.Truncation as TR open import Cubical.Algebra.CommRing.Base open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.AbGroup.Instances.IntMod open import Cubical.Algebra.CommRing.Instances.IntMod open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.AbGroup.TensorProduct @@ -212,9 +213,6 @@ module EM2 {ℓ : Level} (G : AbGroup ℓ) symConstEM-refl {n = suc zero} = transportRefl refl symConstEM-refl {n = suc (suc n)} = transportRefl refl -ℤ/2 : AbGroup ℓ-zero -ℤ/2 = Group→AbGroup (ℤGroup/ 2) +ₘ-comm - private module EMZ/2 = EM2 ℤ/2 -Const-ℤ/2 diff --git a/Cubical/Papers/ComputationalSyntheticCohomology.agda b/Cubical/Papers/ComputationalSyntheticCohomology.agda index 25c768039b..65ec06c95e 100644 --- a/Cubical/Papers/ComputationalSyntheticCohomology.agda +++ b/Cubical/Papers/ComputationalSyntheticCohomology.agda @@ -42,28 +42,33 @@ import Cubical.Homotopy.EilenbergMacLane.GradedCommTensor as CupComm -- 4: Cohomology import Cubical.Cohomology.EilenbergMacLane.Base as Cohom -import Cubical.Cohomology.EilenbergMacLane.EilenbergSteenrod as Axioms import Cubical.Cohomology.EilenbergMacLane.CupProduct as CohomCup import Cubical.Cohomology.EilenbergMacLane.MayerVietoris as MV import Cubical.Axiom.Choice as Choice import Cubical.HITs.Wedge as ⋁ +import Cubical.Homotopy.EilenbergSteenrod as Axioms +import Cubical.Cohomology.EilenbergMacLane.EilenbergSteenrod as SatAxioms -- 5: Computations of cohomology groups and rings import Cubical.Cohomology.EilenbergMacLane.Groups.Unit as HⁿUnit import Cubical.Cohomology.EilenbergMacLane.Groups.Connected as CohomConnected import Cubical.Cohomology.EilenbergMacLane.Groups.Sn as CohomSn +import Cubical.Cohomology.EilenbergMacLane.Rings.Sn as CohomRingSn import Cubical.Cohomology.EilenbergMacLane.Groups.Torus as CohomT² import Cubical.HITs.Torus as T² import Cubical.HITs.RPn.Base as RP² import Cubical.HITs.KleinBottle as K² import Cubical.Cohomology.EilenbergMacLane.Groups.RP2 as CohomRP² +import Cubical.ZCohomology.CohomologyRings.RP2 as ZCohomRingRP² import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle as CohomK² +import Cubical.ZCohomology.CohomologyRings.KleinBottle as ZCohomRingK² +import Cubical.Cohomology.EilenbergMacLane.Rings.RP2 as RP²Ring import Cubical.Cohomology.EilenbergMacLane.Rings.KleinBottle as K²Ring -- import Cubical.Cohomology.Groups.KleinBottle as HⁿK² -- import Cubical.ZCohomology.Groups.KleinBottle as HⁿK² -- import Cubical.Cohomology.EilenbergMacLane.RSnings as CohomRingSn - + -- import Cubical.ZCohomology.RingStructure.CohomologyRing as ℤCohomologyRing -- import Cubical.HITs.S1 as S1 @@ -161,7 +166,7 @@ open EMProps using (isCommΩEM-base) -- Proposition 19. (ap -ₖ = path inversion) open EMGr using (cong-₁ ; cong-₂) --- Proposition 20. (σ preserves +ₖ) +-- Proposition 20. (σ preserves +ₖ) open EMProps using (EM→ΩEM+1-hom) -- Corollary 21. (σ preserves -ₖ) @@ -237,9 +242,11 @@ open ⋁ using (⋁gen) -- Definition 39. Choice open Choice using (satAC) --- Definition 40. Eilenberg-Steenrod axioms. Todo!! +-- Definition 40. Eilenberg-Steenrod axioms +open Axioms using (coHomTheoryGen) --- Theorem 41. Eilenberg-Steenrod axioms. Todo!! +-- Theorem 41. Eilenberg-Steenrod axioms are satified +open SatAxioms using (satisfies-ES-gen) -- Theorem 42 (The mayer-Vietoris sequence) open MV.MV using ( Ker-i⊂Im-d ; Im-d⊂Ker-i @@ -256,7 +263,7 @@ open MV.MV using ( Ker-i⊂Im-d ; Im-d⊂Ker-i open HⁿUnit using (H⁰[Unit,G]≅G ; Hⁿ⁺¹[Unit,G]≅0) -- Lemma 45. Cohom of truncation --- Todo! +open Cohom using (coHomTruncEquiv) -- Proposition 46. Cohomology of connected types open CohomConnected using (H⁰conn) @@ -271,8 +278,8 @@ open CohomSn using (Hⁿ[Sⁿ,G]≅G) -- Proposition 49. Hⁿ(Sᵐ,G) , m ≠ n open CohomSn using (Hⁿ[Sᵐ⁺ⁿ,G]≅0 ; Hᵐ⁺ⁿ[Sⁿ,G]≅0) --- Proposition 50. H*(Sᵐ,G) --- Todo (just transalte from ℤ version) +-- Proposition 50. H*(Sᵐ,G) +open CohomRingSn using (H*[Sⁿ,G]≅G[X]/) --- 5.2 The Torus -- Definition 51. Torus @@ -299,14 +306,14 @@ open CohomRP² using (H³⁺ⁿ[RP²,G]≅0) -- Proposition 58. Hᵐ(K²,G) open CohomK² using (H⁰[K²,G]≅G ; H¹[K²,G]≅G×H¹[RP²,G]; H²⁺ⁿ[K²,G]≅H²⁺ⁿ[RP²,G]) --- Proposition 59. H*(RP²,\bZ) --- Todo +-- Proposition 59. H*(RP²,\bZ) -- Formalisation taken from `Computing Cohomology Rings in Cubical Agda' +open ZCohomRingRP² using (RP²-CohomologyRing) -- Defnition 60 kill-Δ -- Omitted (implicit in formalisation) --- Proposition 61 Hᵐ(K²,G) --- Thomas proof +-- Proposition 61 H*(K²,ℤ) -- Formalisation taken from `Computing Cohomology Rings in Cubical Agda' +open ZCohomRingK² using (CohomologyRing-𝕂²) -- Definition 62. Res (only formliased for relevant special case of F = ⌣ on K(ℤ/2,1)) open K²Ring using (Res⌣) @@ -321,7 +328,7 @@ open K²Ring using (sym-cong₂-⌣≡) open K²Ring using (α²↦1) -- Proposition 67. H*(RP²,Z/2) --- Todo +open RP²Ring using (H*[RP²,ℤ₂]≅ℤ₂[X]/) -- Proposition 68 (Roughly) open K²Ring using (α²↦1 ; βα↦1 ; β²↦0) @@ -331,4 +338,5 @@ open K²Ring using (H*KleinBottle≅ℤ/2[X,Y]/) -- Lemma 70 (implicitly used) --- Proposition 71 (do!) +-- Proposition 71 +-- Todo diff --git a/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda b/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda index 0454d41bbd..4c882195f3 100644 --- a/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda +++ b/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda @@ -24,7 +24,7 @@ open import Cubical.Data.FinData open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup hiding (_[_]) open import Cubical.Algebra.Group.Instances.Unit open import Cubical.Algebra.Group.Instances.Bool open import Cubical.Algebra.Group.Instances.Int diff --git a/Cubical/ZCohomology/CohomologyRings/RP2.agda b/Cubical/ZCohomology/CohomologyRings/RP2.agda index b17c6f6ab4..a56751a623 100644 --- a/Cubical/ZCohomology/CohomologyRings/RP2.agda +++ b/Cubical/ZCohomology/CohomologyRings/RP2.agda @@ -24,7 +24,7 @@ open import Cubical.Data.FinData open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup hiding (_[_]) open import Cubical.Algebra.Group.Instances.Unit open import Cubical.Algebra.Group.Instances.Bool open import Cubical.Algebra.Group.Instances.Int renaming (ℤGroup to ℤG) From 2b8b5fda7e43eb304284c54383710752de3b2f10 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 22 Jan 2024 18:50:46 +0100 Subject: [PATCH 12/73] more --- .../EilenbergMacLane/Rings/RPinf.agda | 1985 +++++++++++++++++ 1 file changed, 1985 insertions(+) create mode 100644 Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda new file mode 100644 index 0000000000..2672655c75 --- /dev/null +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda @@ -0,0 +1,1985 @@ +{-# OPTIONS --safe --lossy-unification #-} + +{- +This file contains +1. The Thom isomorphism (various related forms of it) +2. The Gysin sequence +-} + +module Cubical.Cohomology.EilenbergMacLane.Rings.RPinf where + +open import Cubical.Cohomology.EilenbergMacLane.Base +open import Cubical.Cohomology.EilenbergMacLane.Groups.Sn +open import Cubical.Cohomology.EilenbergMacLane.CupProduct +open import Cubical.Cohomology.EilenbergMacLane.Gysin + +open import Cubical.Homotopy.EilenbergMacLane.CupProduct +open import Cubical.Homotopy.EilenbergMacLane.CupProductTensor + renaming (_⌣ₖ_ to _⌣ₖ⊗_ ; ⌣ₖ-0ₖ to ⌣ₖ-0ₖ⊗ ; 0ₖ-⌣ₖ to 0ₖ-⌣ₖ⊗) +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.EilenbergMacLane.GradedCommTensor + hiding (⌣ₖ-comm) +open import Cubical.Homotopy.EilenbergMacLane.GroupStructure +open import Cubical.Homotopy.EilenbergMacLane.Base +open import Cubical.Homotopy.EilenbergMacLane.Properties +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.EilenbergMacLane.Order2 + +open import Cubical.Functions.Morphism +open import Cubical.Functions.Embedding +open import Cubical.Functions.Surjection + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Transport +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Univalence + +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as TR +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.HITs.RPn + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Relation.Nullary + + +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order hiding (eq) +open import Cubical.Data.Sigma +open import Cubical.Data.Bool hiding (_≤_) + +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing hiding (_ˣ) +open import Cubical.Algebra.CommRing.Instances.IntMod +open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.AbGroup.Instances.IntMod + +open import Cubical.Data.Fin.Arithmetic +open import Cubical.Data.Fin.Base + + + +open RingStr renaming (_+_ to _+r_) +open PlusBis + +open Iso + +-- move to Bool? +Bool→Bool→∙Bool : Bool → (Bool , true) →∙ (Bool , true) +Bool→Bool→∙Bool false = idfun∙ _ +Bool→Bool→∙Bool true = const∙ _ _ + +Iso-Bool→∙Bool-Bool : Iso ((Bool , true) →∙ (Bool , true)) Bool +Iso.fun Iso-Bool→∙Bool-Bool f = fst f false +Iso.inv Iso-Bool→∙Bool-Bool = Bool→Bool→∙Bool +Iso.rightInv Iso-Bool→∙Bool-Bool false = refl +Iso.rightInv Iso-Bool→∙Bool-Bool true = refl +Iso.leftInv Iso-Bool→∙Bool-Bool f = Σ≡Prop (λ _ → isSetBool _ _) (help _ refl) + where + help : (x : Bool) → fst f false ≡ x + → Bool→Bool→∙Bool (fst f false) .fst ≡ f .fst + help false p = funExt + λ { false → (λ j → Bool→Bool→∙Bool (p j) .fst false) ∙ sym p + ; true → (λ j → Bool→Bool→∙Bool (p j) .fst true) ∙ sym (snd f)} + help true p = (λ j → Bool→Bool→∙Bool (p j) .fst) + ∙ funExt λ { false → sym p ; true → sym (snd f)} + +-- pres0→hom +pres0→GroupIso : ∀ {ℓ} {G : Group ℓ} (f : fst G ≃ ℤ/2 .fst) + → fst f (GroupStr.1g (snd G)) ≡ fzero + → IsGroupHom (snd G) (fst f) ((ℤGroup/ 2) .snd) +pres0→GroupIso {G = G} f fp = isGroupHomInv ((invEquiv f) , main) + where + one = invEq f fone + + f⁻∙ : invEq f fzero ≡ GroupStr.1g (snd G) + f⁻∙ = sym (cong (invEq f) fp) ∙ retEq f _ + + f⁻1 : GroupStr._·_ (snd G) (invEq f fone) (invEq f fone) + ≡ GroupStr.1g (snd G) + f⁻1 = sym (retEq f (GroupStr._·_ (snd G) (invEq f fone) (invEq f fone))) + ∙∙ cong (invEq f) (mainlem _ refl ∙ sym fp) + ∙∙ retEq f (GroupStr.1g (snd G)) + where + l : ¬ (fst f (GroupStr._·_ (snd G) (invEq f fone) (invEq f fone)) + ≡ fone) + l p = snotz (cong fst q) + where + q : fone ≡ fzero + q = (sym (secEq f fone) + ∙ cong (fst f) + ((sym (GroupStr.·IdL (snd G) one) + ∙ cong (λ x → GroupStr._·_ (snd G) x one) (sym (GroupStr.·InvL (snd G) one))) + ∙ sym (GroupStr.·Assoc (snd G) (GroupStr.inv (snd G) one) one one))) + ∙ cong (fst f) (cong (GroupStr._·_ (snd G) (GroupStr.inv (snd G) (invEq f fone))) + ((sym (retEq f _) ∙ cong (invEq f) p))) + ∙ cong (fst f) (GroupStr.·InvL (snd G) _) + ∙ fp + + + mainlem : (x : _) + → fst f (GroupStr._·_ (snd G) (invEq f fone) (invEq f fone)) ≡ x + → f .fst ((snd G GroupStr.· invEq f fone) (invEq f fone)) ≡ fzero + mainlem = ℤ/2-elim + (λ p → p) + λ p → ⊥.rec (l p) + + + main : IsGroupHom ((ℤGroup/ 2) .snd) (invEq f) (snd G) + main = + makeIsGroupHom + (ℤ/2-elim + (ℤ/2-elim (f⁻∙ ∙ sym (GroupStr.·IdR (snd G) (GroupStr.1g (snd G))) + ∙ cong (λ x → snd G .GroupStr._·_ x x) (sym f⁻∙)) + (sym (GroupStr.·IdL (snd G) one) + ∙ cong (λ x → snd G .GroupStr._·_ x one) (sym f⁻∙))) + (ℤ/2-elim (sym (GroupStr.·IdR (snd G) one) + ∙ cong (snd G .GroupStr._·_ (invEq f fone)) (sym f⁻∙)) + (f⁻∙ ∙ sym f⁻1))) + +----------------------------------------------------------- + +-- move to Cohomology GroupStr or somethign +EM→-charac : ∀ {ℓ ℓ'} {A : Pointed ℓ} {G : AbGroup ℓ'} (n : ℕ) + → Iso (fst A → EM G n) ((A →∙ EM∙ G n) × EM G n) +Iso.fun (EM→-charac {A = A} n) f = + ((λ x → f x -ₖ f (pt A)) , rCancelₖ n (f (pt A))) , f (pt A) +Iso.inv (EM→-charac n) (f , a) x = fst f x +ₖ a +Iso.rightInv (EM→-charac {A = A} n) ((f , p) , a) = + ΣPathP (→∙Homogeneous≡ (isHomogeneousEM _) + (funExt (λ x → (λ i → (f x +ₖ a) -ₖ (cong (_+ₖ a) p ∙ lUnitₖ n a) i) + ∙ sym (assocₖ n (f x) a (-ₖ a)) + ∙ cong (f x +ₖ_) (rCancelₖ n a) + ∙ rUnitₖ n (f x))) + , cong (_+ₖ a) p ∙ lUnitₖ n a) +Iso.leftInv (EM→-charac {A = A} n) f = + funExt λ x → sym (assocₖ n (f x) (-ₖ f (pt A)) (f (pt A))) + ∙∙ cong (f x +ₖ_) (lCancelₖ n (f (pt A))) + ∙∙ rUnitₖ n (f x) + +0ˣ : ∀ {ℓ} (A : ℕ → Type ℓ) (0A : (n : ℕ) → A n) → (n : ℕ) → A ˣ n +0ˣ A 0A zero = 0A zero +0ˣ A 0A (suc n) = (0ˣ A 0A n) , (0A (suc n)) + +EM-ℤ/2ˣ∙ : (n : ℕ) → EM ℤ/2 ˣ n +EM-ℤ/2ˣ∙ = 0ˣ (EM ℤ/2) 0ₖ + +private + nonConst→∙ : (b : EM ℤ/2 1) → Type _ + nonConst→∙ b = + Σ[ F ∈ Susp∙ (embase ≡ b) →∙ EM∙ (CommRing→AbGroup ℤ/2CommRing) 1 ] + (¬ F ≡ const∙ _ _) + + Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool : + Iso (Susp∙ (Ω (EM∙ ℤ/2 1) .fst) + →∙ EM∙ ℤ/2 1) + ((Bool , true) →∙ (Bool , true)) + Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool = compIso + (invIso (ΩSuspAdjointIso {A = Ω (EM∙ ℤ/2 1)})) + (compIso (post∘∙equiv (lem , refl)) (pre∘∙equiv (lem , refl))) + where + lem = isoToEquiv + (compIso (invIso (Iso-EM-ΩEM+1 {G = ℤ/2} 0)) + (invIso (Bool≅ℤGroup/2 .fst))) + + Σ¬Iso : ∀ {ℓ} {B A : Type ℓ} + → (e : A ≃ B) + → {x : A} + → Iso (Σ[ y ∈ A ] ¬ y ≡ x) + (Σ[ y ∈ B ] ¬ y ≡ fst e x) + Σ¬Iso {B = B} = + EquivJ (λ A e → {x : A} + → Iso (Σ[ y ∈ A ] ¬ y ≡ x) + (Σ[ y ∈ B ] ¬ y ≡ fst e x)) idIso + + nonConst→∙Iso : Iso (nonConst→∙ embase) (Σ[ F ∈ Bool ] ¬ F ≡ true) + nonConst→∙Iso = + Σ¬Iso (isoToEquiv + (compIso Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool + Iso-Bool→∙Bool-Bool)) + + isProp-nonConst→∙ : (b : EM ℤ/2 1) → isProp (nonConst→∙ b) + isProp-nonConst→∙ = EM→Prop _ 0 (λ _ → isPropIsProp) + (isOfHLevelRetractFromIso 1 nonConst→∙Iso + (isContr→isProp ((false , true≢false ∘ sym) + , λ { (false , p) → Σ≡Prop (λ _ → isProp¬ _) refl + ; (true , p) → ⊥.rec (p refl)}))) + +eRP∞∙ : Susp∙ (Ω (EM∙ ℤ/2 1) .fst) →∙ EM∙ ℤ/2 1 +fst eRP∞∙ north = embase +fst eRP∞∙ south = embase +fst eRP∞∙ (merid a i) = a i +snd eRP∞∙ = refl + +eRP∞-nonConst : nonConst→∙ embase +fst eRP∞-nonConst = eRP∞∙ +snd eRP∞-nonConst p = true≢false true≡false + where + true≡false : true ≡ false + true≡false i = + Iso.fun (compIso Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool Iso-Bool→∙Bool-Bool) + (p (~ i)) + +eRP∞ : (b : EM ℤ/2 1) → nonConst→∙ b +eRP∞ = EM→Prop ℤ/2 0 isProp-nonConst→∙ eRP∞-nonConst + +module ThomRP∞ = Thom (EM∙ ℤ/2 1) (0ₖ 1 ≡_) refl + (isConnectedEM 1) + ℤ/2CommRing + +open ThomRP∞ +isContrE : isContr E +isContrE = isContrSingl _ + +module conRP∞ = + con 0 (((compEquiv (isoToEquiv (invIso (Iso-EM-ΩEM+1 0))) + (isoToEquiv (invIso (fst Bool≅ℤGroup/2))))) , refl) + (λ b → eRP∞ b .fst) + (Iso.inv (congIso (compIso Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool + Iso-Bool→∙Bool-Bool)) λ i → false) + +open conRP∞ +ϕ-raw-contrRP∞ : (n : ℕ) + → ((EM ℤ/2 1 → EM ℤ/2 n)) + ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (n +' 1)) +ϕ-raw-contrRP∞ n = ϕ-raw-contr n isContrE + +⌣RP∞ : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + → EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (n +' 1) +fst (⌣RP∞ n f) x = (f x) ⌣[ ℤ/2Ring R]ₖ x +snd (⌣RP∞ n f) = ⌣ₖ-0ₖ _ _ (f (0ₖ 1)) + +⌣RP∞' : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + → EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (1 +' n) +fst (⌣RP∞' n f) x = x ⌣[ ℤ/2Ring R]ₖ (f x) +snd (⌣RP∞' n f) = 0ₖ-⌣ₖ _ _ (f (0ₖ 1)) + +⌣RP∞≡⌣RP∞' : (n : ℕ) → + PathP (λ i → + (EM ℤ/2 1 → EM ℤ/2 n) → + EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (+'-comm n 1 i)) + (⌣RP∞ n) + (⌣RP∞' n) +⌣RP∞≡⌣RP∞' n = funExt λ f → + →∙HomogeneousPathP refl (cong (EM∙ ℤ/2) (+'-comm n 1)) + (isHomogeneousEM _) + (funExt λ x → lem f x) + where + lem : (f : EM ℤ/2 1 → EM ℤ/2 n) + (x : EM ℤ/2 1) + → PathP (λ i → EM ℤ/2 (+'-comm n 1 i)) + (f x ⌣[ ℤ/2Ring R]ₖ x) (x ⌣[ ℤ/2Ring R]ₖ f x) + lem f x = toPathP (sym (⌣ₖ-commℤ/2 1 n x (f x))) + +⌣RP∞IsEq : (n : ℕ) → isEquiv (⌣RP∞ n) +⌣RP∞IsEq n = + subst isEquiv + (funExt (λ f → →∙Homogeneous≡ (isHomogeneousEM _) + (λ i x → f x ⌣[ ℤ/2Ring R]ₖ (eRP∞-lem x i)))) + (ϕ-raw-contrRP∞ n .snd) + where + help : (g : _) → (λ i → eRP∞ (emloop g i) .fst .fst south) + ≡ emloop g + help g j i = + hcomp (λ k → λ {(i = i0) → embase + ; (i = i1) → emloop g k + ; (j = i0) → eRP∞ (emloop g i) .fst .fst + (merid (λ w → emloop g (i ∧ w)) k) + ; (j = i1) → emloop g (i ∧ k)}) + (eRP∞ (emloop g i) .fst .snd j) + + eRP∞-lem : (x : _) → eRP∞ x .fst .fst south ≡ x + eRP∞-lem = EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) + λ { embase-raw → refl ; (emloop-raw g i) j → help g j i } + +abstract + ⌣RP∞'IsEq : (n : ℕ) → isEquiv (⌣RP∞' n) + ⌣RP∞'IsEq n = transport (λ i → isEquiv (⌣RP∞≡⌣RP∞' n i)) (⌣RP∞IsEq n) + +⌣RP∞Equiv : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (n +' 1)) +fst (⌣RP∞Equiv n) = ⌣RP∞ n +snd (⌣RP∞Equiv n) = ⌣RP∞IsEq n + +⌣RP∞'Equiv : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (1 +' n)) +fst (⌣RP∞'Equiv n) = ⌣RP∞' n +snd (⌣RP∞'Equiv n) = ⌣RP∞'IsEq n + + ++'-suc₁ : (n : ℕ) → 1 +' n ≡ suc n ++'-suc₁ zero = refl ++'-suc₁ (suc n) = refl + +⌣RP∞''Equiv : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (suc n)) +⌣RP∞''Equiv n = + compEquiv (⌣RP∞'Equiv n) + (isoToEquiv (pre∘∙equiv + ((pathToEquiv (cong (EM ℤ/2) (+'-suc₁ n))) + , (subst-EM-0ₖ (+'-suc₁ n))))) + +RP→Charac₀ : Iso (EM ℤ/2 1 → ℤ/2 .fst) + (ℤ/2 .fst) +Iso.fun RP→Charac₀ f = f embase +Iso.inv RP→Charac₀ a = λ _ → a +Iso.rightInv RP→Charac₀ a = refl +Iso.leftInv RP→Charac₀ f = + funExt (EM→Prop _ 0 (λ _ → is-set (snd ℤ/2Ring) _ _) refl) + +RP→EM-ℤ/2-CharacIso : (n : ℕ) + → Iso (EM ℤ/2 1 → EM ℤ/2 n) + ((EM ℤ/2) ˣ n) +RP→EM-ℤ/2-CharacIso zero = RP→Charac₀ -- RP→Charac₀ +RP→EM-ℤ/2-CharacIso (suc n) = + compIso (EM→-charac {A = EM∙ ℤ/2 1} (suc n)) + (prodIso + (compIso help (RP→EM-ℤ/2-CharacIso n)) + idIso) + where + help : Iso (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (suc n)) + (EM ℤ/2 1 → EM ℤ/2 n) + help = equivToIso (invEquiv (⌣RP∞''Equiv n)) + +∥EM-ℤ/2ˣ∥₀-Iso : (n : ℕ) → Iso ∥ EM ℤ/2 ˣ n ∥₂ (fst ℤ/2) +∥EM-ℤ/2ˣ∥₀-Iso zero = setTruncIdempotentIso (is-set (snd ℤ/2Ring)) +∥EM-ℤ/2ˣ∥₀-Iso (suc n) = + compIso + (compIso + setTruncOfProdIso + (compIso (Σ-cong-iso-snd λ _ + → equivToIso + (isContr→≃Unit (∣ 0ₖ (suc n) ∣₂ + , ST.elim (λ _ → isSetPathImplicit) + (EM→Prop _ _ (λ _ → squash₂ _ _) refl)))) + rUnit×Iso)) + (∥EM-ℤ/2ˣ∥₀-Iso n) + +⌣RP∞''Equiv∙ : (n : ℕ) + → ⌣RP∞''Equiv n .fst (λ _ → 0ₖ n) ≡ const∙ _ _ +⌣RP∞''Equiv∙ n = →∙Homogeneous≡ (isHomogeneousEM _) + (funExt λ x → cong (subst (EM ℤ/2) (+'-suc₁ n)) (⌣ₖ-0ₖ 1 n x) + ∙ subst-EM∙ (+'-suc₁ n) .snd) + +cohomRP∞Iso : (n : ℕ) → Iso (coHom n ℤ/2 (EM ℤ/2 1)) (ℤ/2 .fst) +cohomRP∞Iso n = compIso (setTruncIso (RP→EM-ℤ/2-CharacIso n)) (∥EM-ℤ/2ˣ∥₀-Iso n) + +RP→EM-ℤ/2-CharacIso∙ : (n : ℕ) + → Iso.fun (RP→EM-ℤ/2-CharacIso n) (λ _ → 0ₖ n) ≡ EM-ℤ/2ˣ∙ n +RP→EM-ℤ/2-CharacIso∙ zero = refl +RP→EM-ℤ/2-CharacIso∙ (suc n) = + ΣPathP (((cong (fun (RP→EM-ℤ/2-CharacIso n)) + (funExt λ x + → cong (λ f → invEq (⌣RP∞''Equiv n) f x) + (→∙Homogeneous≡ (isHomogeneousEM _) + (funExt (λ x → rCancelₖ (suc n) (0ₖ (suc n))))) + ∙ funExt⁻ (sym (cong (invEq (⌣RP∞''Equiv n)) (⌣RP∞''Equiv∙ n))) x + ∙ λ i → retEq (⌣RP∞''Equiv n) (λ _ → 0ₖ n) i x)) + ∙ RP→EM-ℤ/2-CharacIso∙ n) + , refl) + +HⁿRP∞≅ℤ/2 : (n : ℕ) → AbGroupIso (coHomGr n ℤ/2 (EM ℤ/2 1)) ℤ/2 +fst (HⁿRP∞≅ℤ/2 n) = cohomRP∞Iso n +snd (HⁿRP∞≅ℤ/2 n) = pres0→GroupIso (isoToEquiv (cohomRP∞Iso n)) (cohomRP∞Iso∙ n) + where + ∥EM-ℤ/2ˣ∥₀-Iso∙ : (n : ℕ) → Iso.fun (∥EM-ℤ/2ˣ∥₀-Iso n) ∣ EM-ℤ/2ˣ∙ n ∣₂ ≡ fzero + ∥EM-ℤ/2ˣ∥₀-Iso∙ zero = refl + ∥EM-ℤ/2ˣ∥₀-Iso∙ (suc n) = ∥EM-ℤ/2ˣ∥₀-Iso∙ n + + cohomRP∞Iso∙ : (n : ℕ) → cohomRP∞Iso n .fun (0ₕ n) ≡ fzero + cohomRP∞Iso∙ n = cong (Iso.fun (∥EM-ℤ/2ˣ∥₀-Iso n) ∘ ∣_∣₂) (RP→EM-ℤ/2-CharacIso∙ n) + ∙ ∥EM-ℤ/2ˣ∥₀-Iso∙ n + +-- makeSquare : (n : ℕ) +-- → (EM ℤ/2 1 +-- → EM ℤ/2 n +-- → EM ℤ/2 (n +' n)) +-- → (EM ℤ/2 n → (EM ℤ/2) ˣ (n +' n)) +-- makeSquare n f y = Iso.fun (RP→EM-ℤ/2-CharacIso _) λ x → f x y + +-- module _ {ℓ} (X Y : RP∞' ℓ) (A : fst X → fst Y → Pointed ℓ) where +-- UnordSmash∙² : Pointed ℓ +-- UnordSmash∙² = UnordSmash∙ X (λ x → UnordSmash∙ Y (A x)) + +-- UnordSmash∙²-fun : ((x : fst X) (y : fst Y) → A x y .fst) +-- → UnordSmash∙² .fst +-- UnordSmash∙²-fun f = inr (λ x → inr (f x)) + +-- -- open import Cubical.HITs.Join +-- -- ×⁴ : ∀ {ℓ} (A : Bool → Bool → Pointed ℓ) → Type ℓ +-- -- ×⁴ A = A true true .fst × A true false .fst +-- -- × A false true .fst × A false false .fst + +-- -- module _ {ℓ} {A : Type ℓ} {B C : A → Type ℓ} (contr : isContr (Σ A B)) where +-- -- private +-- -- push-c : (a : A) (p : contr .fst .fst ≡ a) +-- -- → (c : C a) +-- -- → Path (cofib {A = C (contr .fst .fst)} {B = Σ A C} λ x → _ , x) (inl tt) (inr (a , c)) +-- -- push-c = J> push + +-- -- cofib→join : (ptA : A) (ptB : B ptA) → (cofib {A = C ptA} {B = Σ A C} λ x → _ , x) → Σ[ a ∈ A ] join (B a) (C a) +-- -- cofib→join ptA ptB (inl x) = ptA , (inl ptB) -- contr .fst .fst , inl (contr .fst .snd) +-- -- cofib→join ptA ptB (inr (a , c)) = a , inr c +-- -- cofib→join ptA ptB (push a i) = ptA , push ptB a i + +-- -- push-c-coh : (a : A) (p : contr .fst .fst ≡ a) (d : B a) (pp : PathP (λ i → B (p i)) (contr .fst .snd) d) (c : C a) +-- -- → PathP (λ i → cofib→join (contr .fst .fst) (contr .fst .snd) (push-c a p c i) ≡ (a , push d c i)) +-- -- (ΣPathP (p , (λ j → inl (pp j)))) +-- -- refl -- refl +-- -- push-c-coh = +-- -- J> (J> λ c → flipSquare ((λ j i → cofib→join (contr .fst .fst) (contr .fst .snd) (help c j i)) +-- -- ◁ λ i j → (contr .fst .fst) , (push (contr .fst .snd) c j))) +-- -- where +-- -- help : (c : _) → push-c (contr .fst .fst) refl c ≡ push c +-- -- help = funExt⁻ (transportRefl push) + +-- -- joinC : Iso (Σ[ a ∈ A ] join (B a) (C a)) +-- -- (cofib {A = C (contr .fst .fst)} {B = Σ A C} λ x → _ , x) +-- -- Iso.fun joinC (a , inl x) = inl tt +-- -- Iso.fun joinC (a , inr x) = inr (a , x) +-- -- Iso.fun joinC (a , push b c i) = push-c a (cong fst (contr .snd (a , b))) c i +-- -- Iso.inv joinC = cofib→join (contr .fst .fst) (contr .fst .snd) +-- -- Iso.rightInv joinC (inl x) = refl +-- -- Iso.rightInv joinC (inr x) = refl +-- -- Iso.rightInv joinC (push a i) j = +-- -- ((λ j → push-c (contr .fst .fst) +-- -- (cong fst (isProp→isSet (isContr→isProp contr) _ _ (contr .snd (contr .fst)) refl j)) a) +-- -- ∙ lem a) j i +-- -- where +-- -- lem : (a : _) → Path (Path (cofib (λ x → contr .fst .fst , x)) _ _) (push-c (contr .fst .fst) refl a) (push a) +-- -- lem = funExt⁻ (transportRefl push) +-- -- Iso.leftInv joinC (a , inl x) = ΣPathP ((cong fst (contr .snd (a , x))) , (λ j → inl (contr .snd (a , x) j .snd))) +-- -- Iso.leftInv joinC (a , inr x) = refl +-- -- Iso.leftInv joinC (a , push c d i) j = +-- -- push-c-coh a (cong fst (contr .snd (a , c))) c (cong snd (contr .snd (a , c))) d i j + +-- -- JoinStructureBool : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) +-- -- (f : A true true .fst × A true false .fst +-- -- × A false true .fst × A false false .fst +-- -- → fst B) +-- -- → Type +-- -- JoinStructureBool A B f = +-- -- (g : A true true .fst × A true false .fst +-- -- × A false true .fst × A false false .fst) +-- -- → join (join (fst g ≡ A true true .snd) +-- -- (snd g .fst ≡ A true false .snd)) +-- -- (join (snd (snd g) .fst ≡ A false true .snd) +-- -- (snd (snd g) .snd ≡ A false false .snd)) +-- -- → f g ≡ B .snd + +-- -- module DEP (A B : Pointed ℓ-zero) (T : fst A → fst B → Pointed ℓ-zero) +-- -- (f : (a : fst A) (b : fst B) → fst (T a b)) where +-- -- JS₁ : Type +-- -- JS₁ = (a : A .fst) (b : B .fst) +-- -- → join (a ≡ A .snd) (b ≡ B .snd) +-- -- → f a b ≡ T a b .snd + +-- -- JS'l : singl (pt A) → Type +-- -- JS'l (a , p) = (b : fst B) → f a b ≡ T a b .snd + +-- -- JS'r : singl (pt B) → Type +-- -- JS'r (b , p) = (a : fst A) → f a b ≡ T a b .snd + +-- -- JS₁' : Type +-- -- JS₁' = Σ[ l ∈ ((a : _) → JS'l a) ] +-- -- Σ[ r ∈ ((a : _) → JS'r a) ] +-- -- ((a : _) (b : _) → l a (fst b) ≡ r b (fst a)) + + +-- -- JS₁'' : Type +-- -- JS₁'' = Σ[ l ∈ JS'l (pt A , refl) ] +-- -- Σ[ r ∈ JS'r (pt B , refl) ] +-- -- l (pt B) ≡ r (pt A) + +-- -- IS1 : Iso JS₁ JS₁' +-- -- fst (Iso.fun IS1 F) (a , p) b = F a b (inl (sym p)) +-- -- fst (snd (Iso.fun IS1 F)) (b , p) a = F a b (inr (sym p)) +-- -- snd (snd (Iso.fun IS1 F)) (a , p) (b , q) = cong (F a b) (push (sym p) (sym q)) +-- -- Iso.inv IS1 (l , r , lr) a b (inl x) = l (a , sym x) b +-- -- Iso.inv IS1 (l , r , lr) a b (inr x) = r (b , sym x) a +-- -- Iso.inv IS1 (l , r , lr) a b (push p q i) = lr (a , sym p) (b , sym q) i +-- -- Iso.rightInv IS1 _ = refl +-- -- Iso.leftInv IS1 F = funExt λ a → funExt λ b → funExt +-- -- λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} + +-- -- IS2 : Iso JS₁' JS₁'' +-- -- IS2 = compIso +-- -- (Σ-cong-iso +-- -- {B = λ l → Σ[ r ∈ ((a : _) → JS'r a) ] +-- -- ((a : _) (b : _) → l a (fst b) ≡ r b (fst a))} +-- -- {B' = λ l → Σ[ r ∈ ((a : _) → JS'r a) ] +-- -- ((b : _) → l (fst b) ≡ r b (pt A))} +-- -- (DomainContrIso (isContrSingl (pt A))) +-- -- λ x → Σ-cong-iso-snd λ r → DomainContrIso (isContrSingl (pt A))) +-- -- (Σ-cong-iso-snd +-- -- λ l → Σ-cong-iso {B = λ r → ((b : _) → l (fst b) ≡ r b (pt A))} +-- -- {B' = λ r → (l (pt B) ≡ r (pt A))} +-- -- (DomainContrIso (isContrSingl (pt B))) +-- -- λ _ → DomainContrIso (isContrSingl (pt B))) + +-- -- FullIso : Iso JS₁ JS₁'' +-- -- FullIso = compIso IS1 IS2 + +-- -- move later +-- -- RP∞'→SetElim : +-- -- ∀ {ℓ} {A : RP∞' ℓ → Type ℓ} (s : (X : _) → isSet (A X)) +-- -- → (f : (X : RP∞' ℓ) → fst X → A X) +-- -- → ((X : RP∞' ℓ) → (x : fst X) → f X x ≡ f X (RP∞'-fields.notRP∞' X x)) +-- -- → ((X : RP∞' ℓ) → A X) +-- -- RP∞'→SetElim {A = A} s f f-comm = +-- -- uncurry λ X → uncurry λ 2tx +-- -- → elim→Set (λ _ → s _) +-- -- (λ x → f (X , 2tx , ∣ x ∣₁) x) +-- -- λ x → RP∞'-fields.elimRP∞' (X , 2tx , ∣ x ∣₁) x +-- -- (λ i → f (X , 2tx , squash₁ ∣ x ∣₁ ∣ x ∣₁ i) x) +-- -- (f-comm (X , 2tx , ∣ x ∣₁) x +-- -- ◁ λ i → f (X , 2tx , squash₁ ∣ x ∣₁ ∣ fst 2tx x ∣₁ i) (fst 2tx x)) + +-- -- RP∞'→SetElimEq : +-- -- ∀ {ℓ} {A : RP∞' ℓ → Type ℓ} (s : (X : _) → isSet (A X)) +-- -- (f : (X : RP∞' ℓ) → fst X → A X) +-- -- (h : (X : RP∞' ℓ) → (x : fst X) → f X x ≡ f X (RP∞'-fields.notRP∞' X x)) +-- -- (X : RP∞' ℓ) (x : _) +-- -- → RP∞'→SetElim s f h X ≡ f X x +-- -- RP∞'→SetElimEq {A = A} s f h = uncurry λ X → uncurry +-- -- λ 2x → PT.elim (λ _ → isPropΠ λ _ → s _ _ _) +-- -- λ x → RP∞'-fields.elimRP∞' (X , 2x , ∣ x ∣₁) x +-- -- (fromPathP λ i → (f (X , 2x , squash₁ ∣ x ∣₁ ∣ x ∣₁ i) x)) +-- -- (fromPathP (h (X , 2x , ∣ x ∣₁) x +-- -- ◁ λ i → f (X , 2x , squash₁ ∣ x ∣₁ ∣ x ∣₁ i) (fst 2x x))) + +-- RP∞'→SetRec : ∀ {ℓ} {A : Type ℓ} (s : isSet A) (X : RP∞' ℓ) +-- → (f : fst X → A) +-- → ((x : _) → f x ≡ f (RP∞'-fields.notRP∞' X x)) +-- → A +-- RP∞'→SetRec s = uncurry λ X +-- → uncurry λ 2tx +-- → elim→Set (λ _ → isSetΠ2 λ _ _ → s) +-- (λ x f coh → f x) +-- λ x → RP∞'-fields.elimRP∞' (X , 2tx , ∣ x ∣₁) x +-- (λ i f coh → f x) +-- λ i f coh → coh x i + +-- RP∞'→SetRecEq : +-- ∀ {ℓ} {A : Type ℓ} (s : isSet A) (X : _) +-- (f : fst X → A) +-- (h : (x : fst X) → f x ≡ f (RP∞'-fields.notRP∞' X x)) +-- (x : fst X) +-- → RP∞'→SetRec s X f h ≡ f x +-- RP∞'→SetRecEq {A = A} s = uncurry λ X → uncurry +-- λ 2x → PT.elim (λ _ → isPropΠ3 λ _ _ _ → s _ _) +-- λ x f h → RP∞'-fields.elimRP∞' (X , 2x , ∣ x ∣₁) x +-- (λ i → transportRefl (transportRefl (f (transportRefl x i)) i) i) +-- ((λ i → transportRefl (transportRefl (f (transportRefl x i)) i) i) ∙ h x) + + +-- abstract +-- notNotRP∞' : ∀ {ℓ} (X : RP∞' ℓ) (x : fst X) +-- → RP∞'-fields.notRP∞' X (RP∞'-fields.notRP∞' X x) ≡ x +-- notNotRP∞' = JRP∞' refl + +-- RP∞'→GroupoidRec : ∀ {ℓ} {A : Type ℓ} (s : isGroupoid A) (X : RP∞' ℓ) +-- → (f : fst X → A) +-- → (f-coh : (x : _) → f x ≡ f (RP∞'-fields.notRP∞' X x)) +-- → (p : (x : fst X) +-- → PathP (λ i → f (notNotRP∞' X x (~ i)) ≡ f (RP∞'-fields.notRP∞' X x)) +-- (f-coh x) +-- (sym (f-coh (RP∞'-fields.notRP∞' X x)))) +-- → A +-- RP∞'→GroupoidRec {ℓ = ℓ} {A = A} grA = uncurry λ X +-- → uncurry λ 2tx +-- → elim→Gpd _ (λ _ → isGroupoidΠ3 λ _ _ _ → grA) +-- (F1 X 2tx) +-- (F2 X 2tx) +-- λ x → F3 (X , 2tx , ∣ x ∣₁) x +-- where +-- F1 : (X : Type) (2tx : is2Type ℓ X) (x : X) (f : X → A) +-- (f-coh : (x' : X) → f x' ≡ f (RP∞'-fields.notRP∞' (X , 2tx , ∣ x ∣₁) x')) → +-- ((x' : X) → +-- PathP (λ i → f (notNotRP∞' (X , 2tx , ∣ x ∣₁) x' (~ i)) +-- ≡ f (RP∞'-fields.notRP∞' (X , 2tx , ∣ x ∣₁) x')) +-- (f-coh x') +-- (sym (f-coh (RP∞'-fields.notRP∞' (X , 2tx , ∣ x ∣₁) x')))) +-- → A +-- F1 X 2tx x f coh p = f x + +-- F2 : (X : Type) (2tx : is2Type ℓ X) (x y : X) → +-- PathP (λ i → +-- (f : X → A) +-- (f-coh : (x' : X) +-- → f x' ≡ f (RP∞'-fields.notRP∞' (X , 2tx , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) x')) +-- → ((x' : X) → +-- PathP (λ i₁ → +-- f (notNotRP∞' (X , 2tx , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) x' (~ i₁)) +-- ≡ f (RP∞'-fields.notRP∞' (X , 2tx , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) x')) +-- (f-coh x') +-- (sym (f-coh (RP∞'-fields.notRP∞' (X , 2tx , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) x')))) → A) +-- (F1 X 2tx x) +-- (F1 X 2tx y) +-- F2 X 2tx x = +-- RP∞'-fields.elimRP∞' (X , 2tx , ∣ x ∣₁) x +-- (λ i f coh p → f x) +-- λ i f coh p → coh x i + +-- F3 : (X : RP∞' ℓ) (x y z : fst X) → +-- SquareP +-- (λ i j → (f : fst X → A) +-- (f-coh : (x' : fst X) +-- → f x' +-- ≡ f (RP∞'-fields.notRP∞' (fst X , fst (snd X) , squash₁ᵗ x y z i j) x')) +-- → ((x' : fst X) → +-- PathP +-- (λ i₁ → +-- f (notNotRP∞' (fst X , fst (snd X) , squash₁ᵗ x y z i j) x' (~ i₁)) +-- ≡ f (RP∞'-fields.notRP∞' (fst X , fst (snd X) , squash₁ᵗ x y z i j) x')) +-- (f-coh x') +-- (sym (f-coh (RP∞'-fields.notRP∞' (fst X , fst (snd X) , squash₁ᵗ x y z i j) x')))) → A) +-- (F2 (fst X) (fst (snd X)) x y) (F2 (fst X) (fst (snd X)) x z) +-- (λ i f f-coh p → f x) (F2 (fst X) (fst (snd X)) y z) +-- F3 = JRP∞' (CasesBool true +-- (CasesBool true (λ i j f f-coh _ → f true) +-- λ i j f f-coh p → f-coh true (i ∧ j)) +-- (CasesBool true +-- (λ i j f f-coh p +-- → hcomp (λ k → λ {(i = i0) → p true (~ k) j +-- ; (i = i1) → f true -- f-coh false (k ∧ j) -- f true +-- ; (j = i0) → f (notNotRP∞' (RP∞'∙ ℓ) true (k ∨ i)) +-- ; (j = i1) → f-coh false i}) +-- (hcomp (λ k → λ {(i = i0) → f-coh false (~ j) +-- ; (i = i1) → f true +-- ; (j = i0) → f (isSetBool _ _ refl (notNotRP∞' (RP∞'∙ ℓ) true) k i) +-- ; (j = i1) → f-coh false i}) +-- (f-coh false (i ∨ ~ j)))) +-- λ i j f f-coh p → f-coh true j)) + +-- RP∞'→GroupoidRecId : ∀ {ℓ} {A : Type ℓ} (s : isGroupoid A) (X : RP∞' ℓ) +-- → (f : fst X → A) +-- → (f-coh : (x : _) → f x ≡ f (RP∞'-fields.notRP∞' X x)) +-- → (p : (x : fst X) +-- → PathP (λ i → f (notNotRP∞' X x (~ i)) ≡ f (RP∞'-fields.notRP∞' X x)) +-- (f-coh x) +-- (sym (f-coh (RP∞'-fields.notRP∞' X x)))) +-- → Σ[ f-id ∈ ((x : fst X) → RP∞'→GroupoidRec s X f f-coh p ≡ f x) ] +-- ((x : fst X) → PathP (λ i → {!!}) {!!} {!!}) -- (RP∞'→GroupoidRec s X f f-coh p) (RP∞'→GroupoidRec s X f f-coh p)) +-- RP∞'→GroupoidRecId = {!!} + +-- isGroupoidRP∞' : isGroupoid (RP∞' ℓ-zero) +-- isGroupoidRP∞' = {!!} + +-- EM₁-ℤ/2→RP∞'-emloop : (g : ℤ/2 .fst) → RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero +-- EM₁-ℤ/2→RP∞'-emloop (zero , p) = refl +-- EM₁-ℤ/2→RP∞'-emloop (suc zero , p) = Σ≡Prop (isPropIsRP∞ ℓ-zero) (ua notEquiv) +-- EM₁-ℤ/2→RP∞'-emloop (suc (suc g) , p) = +-- ⊥.rec (¬-<-zero (<-k+-cancel {k = 2} p)) + +-- EM₁-ℤ/2→RP∞' : EM ℤ/2 1 → RP∞' ℓ-zero +-- EM₁-ℤ/2→RP∞' = +-- elimGroupoid _ (λ _ → isGroupoidRP∞') (RP∞'∙ ℓ-zero) +-- EM₁-ℤ/2→RP∞'-emloop +-- λ g h → ΣSquareSet (λ _ → isProp→isSet (isPropIsRP∞ ℓ-zero _)) +-- (coh g h) +-- where +-- coh : (g h : ℤ/2 .fst) +-- → Square refl (cong fst (EM₁-ℤ/2→RP∞'-emloop h)) +-- (cong fst (EM₁-ℤ/2→RP∞'-emloop g)) +-- (cong fst (EM₁-ℤ/2→RP∞'-emloop (g +ₘ h))) +-- coh = ℤ/2-elim (ℤ/2-elim refl λ i j → ua notEquiv (i ∧ j)) +-- (ℤ/2-elim (λ i j → ua notEquiv i) +-- ((λ i j → ua notEquiv (~ j ∧ i)) ▷ sym help)) +-- where +-- help : ua notEquiv ≡ sym (ua notEquiv) +-- help = cong ua (Σ≡Prop isPropIsEquiv (funExt (CasesBool true refl refl))) +-- ∙ uaInvEquiv notEquiv + +-- RP∞'→EM₁-ℤ/2' : EM ℤ/2 1 ≃ RP∞' ℓ-zero +-- RP∞'→EM₁-ℤ/2' = EM₁-ℤ/2→RP∞' +-- , record { equiv-proof = RP∞'pt→Prop (λ _ → isPropIsContr) +-- ((embase , refl) +-- , (isEmbedding→hasPropFibers′ {f = EM₁-ℤ/2→RP∞'} +-- (Cubical.HITs.EilenbergMacLane1.elimProp _ (λ _ → isPropΠ λ _ → isPropIsEquiv _) +-- (Cubical.HITs.EilenbergMacLane1.elimProp _ (λ _ → isPropIsEquiv _) +-- (subst isEquiv (sym main) (isoToEquiv (invIso (compIso p2 p1)) .snd )))) _ _)) } +-- where -- +-- p1 : Iso Bool (Ω (EM∙ ℤ/2 1) .fst) +-- p1 = compIso (fst Bool≅ℤGroup/2) (Iso-EM-ΩEM+1 0) + +-- is1 : Iso (RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero) (Bool ≡ Bool) +-- fun is1 = cong fst +-- inv is1 p = Σ≡Prop (isPropIsRP∞ ℓ-zero) p +-- rightInv is1 p = refl +-- leftInv is1 p = +-- ΣSquareSet (λ _ → isProp→isSet (isPropIsRP∞ ℓ-zero _)) +-- λ i j → fst (p j) + +-- p2 : Iso (Ω (RP∞' ℓ-zero , RP∞'∙ ℓ-zero) .fst) Bool +-- p2 = compIso is1 (compIso (equivToIso univalence) Bool≃Charac) + +-- main : cong EM₁-ℤ/2→RP∞' ≡ Iso.inv p2 ∘ Iso.inv p1 +-- main = funExt λ x → cong (cong EM₁-ℤ/2→RP∞') (sym (Iso.rightInv p1 x)) +-- ∙ sym (Iso.leftInv p2 _) +-- ∙ cong (Iso.inv p2) (is3 (Iso.inv p1 x)) +-- where +-- is3 : (x : Bool) → Iso.fun p2 (cong EM₁-ℤ/2→RP∞' (Iso.fun p1 x)) ≡ x +-- is3 false = refl +-- is3 true = refl + + +-- {- +-- help : (X : RP∞' ℓ-zero) → isContr (Σ[ X' ∈ (EM ℤ/2 1) ] EM₁-ℤ/2→RP∞' X' ≡ X) +-- help = +-- RP∞'pt→Prop (λ _ → isPropIsContr) ((embase , refl) +-- , (uncurry (elimSet _ (λ _ → isSetΠ λ _ +-- → isOfHLevelPath' 2 +-- (isOfHLevelΣ 3 emsquash (λ _ → isOfHLevelPath 3 isGroupoidRP∞' _ _)) _ _) +-- (λ y → p1 {!!}) +-- {!!}))) +-- where +-- fibr = Σ[ X' ∈ (EM ℤ/2 1) ] EM₁-ℤ/2→RP∞' X' ≡ RP∞'∙ ℓ-zero + + +-- uap' : Bool → RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero +-- uap' false = Σ≡Prop (isPropIsRP∞ ℓ-zero) (ua notEquiv) +-- uap' true = refl + +-- uapEq : Bool ≃ (RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero) +-- fst uapEq = uap' +-- snd uapEq = subst isEquiv (funExt l) (isoToEquiv (invIso uap) .snd) +-- where + + +-- l : (x : _) → inv uap x ≡ uap' x +-- l false = refl +-- l true = ΣSquareSet (λ _ → isProp→isSet (isPropIsRP∞ ℓ-zero _)) +-- (uaIdEquiv) + +-- p1 : (e : Bool) → Path fibr (embase , refl) (embase , uap' e) +-- p1 false = ΣPathP ((emloop fone) , +-- (ΣSquareSet (λ _ → isProp→isSet (isPropIsRP∞ ℓ-zero _)) +-- ((λ i j → ua (notEquiv) (i ∧ ~ j)) +-- ▷ (sym (uaInvEquiv notEquiv) ∙ cong ua (Σ≡Prop isPropIsEquiv refl))))) +-- p1 true = refl + +-- r : (e : RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero) +-- → Σ[ f ∈ Path fibr (embase , refl) (embase , e) ] +-- {!(g : ℤ/2) → PathP !} +-- r = {!!} +-- -} + + +-- RP∞'→EM₁-ℤ/2 : RP∞' ℓ-zero → EM ℤ/2 1 +-- RP∞'→EM₁-ℤ/2 X = +-- RP∞'→GroupoidRec emsquash X (λ _ → embase) +-- (λ x → {!!}) +-- {!!} + +-- -- ∑RP∞' : (X : RP∞' ℓ-zero) (n : fst X → ℕ) → ℕ +-- -- ∑RP∞' X n = +-- -- RP∞'→SetRec isSetℕ X +-- -- (λ x → n x +' n (RP∞'-fields.notRP∞' X x)) +-- -- λ x → +'-comm (n x) _ ∙ cong (n (RP∞'-fields.notRP∞' X x) +'_) +-- -- (cong n (sym (notNotRP∞' X x))) + + +-- -- ∑RP∞'≡ : (X : RP∞' ℓ-zero) (x : fst X) (n : fst X → ℕ) +-- -- → ∑RP∞' X n ≡ n x +' n (RP∞'-fields.notRP∞' X x) +-- -- ∑RP∞'≡ = RP∞'pt→Prop (λ _ → isPropΠ2 λ _ _ → isSetℕ _ _) +-- -- λ { false n → +'-comm (n true) (n false) +-- -- ; true n → refl} + +-- -- ∑RP∞'Fubini : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) +-- -- → (∑RP∞' Y (λ y → ∑RP∞' X (λ x → n x y))) +-- -- ≡ (∑RP∞' X (λ x → ∑RP∞' Y (n x))) +-- -- ∑RP∞'Fubini = +-- -- RP∞'pt→Prop (λ X → isPropΠ2 λ _ _ → isSetℕ _ _) +-- -- (RP∞'pt→Prop ((λ _ → isPropΠ λ _ → isSetℕ _ _)) +-- -- λ n → move4 (n true true) (n false true) (n true false) (n false false) _+'_ +-- -- +'-assoc +-- -- +'-comm) + +-- -- module _ {ℓ} (X : RP∞' ℓ) (A : fst X → Pointed ℓ) (B : Pointed ℓ) where +-- -- BipointedUnordJoin : (f : ((x : fst X) → A x .fst) → fst B) → Type ℓ +-- -- BipointedUnordJoin f = +-- -- (g : (x : fst X) → A x .fst) +-- -- → UnordJoinR X (λ x → g x ≡ A x .snd) +-- -- → f g ≡ B .snd + + +-- -- module _ {ℓ} (X Y : RP∞' ℓ) (A : fst X → fst Y → Pointed ℓ) (B : Pointed ℓ) where +-- -- QuadpointedUnordJoin : (f : ((x : fst X) (y : fst Y) → A x y .fst) → fst B) → Type ℓ +-- -- QuadpointedUnordJoin f = (g : (x : fst X) (y : fst Y) → A x y .fst) +-- -- → UnordJoinR X (λ x → UnordJoinR Y λ y → g x y ≡ A x y .snd) +-- -- → f g ≡ B .snd + +-- -- open import Cubical.HITs.Join +-- -- module _ {ℓ} (A B T : Pointed ℓ) +-- -- (f : fst A → fst B → fst T) where +-- -- BipointedJoinBool : Type ℓ +-- -- BipointedJoinBool = (a : A .fst) (b : B .fst) +-- -- → join (a ≡ A .snd) (b ≡ B .snd) +-- -- → f a b ≡ T .snd + +-- -- JS'l : singl (pt A) → Type ℓ +-- -- JS'l (a , p) = (b : fst B) → f a b ≡ T .snd + +-- -- JS'r : singl (pt B) → Type ℓ +-- -- JS'r (b , p) = (a : fst A) → f a b ≡ T .snd + +-- -- BipointedJoinBool' : Type ℓ +-- -- BipointedJoinBool' = Σ[ l ∈ ((a : _) → JS'l a) ] +-- -- Σ[ r ∈ ((a : _) → JS'r a) ] +-- -- ((a : _) (b : _) → l a (fst b) ≡ r b (fst a)) + + +-- -- BipointedJoinBool'' : Type ℓ +-- -- BipointedJoinBool'' = Σ[ l ∈ JS'l (pt A , refl) ] +-- -- Σ[ r ∈ JS'r (pt B , refl) ] +-- -- l (pt B) ≡ r (pt A) + +-- -- IS1 : Iso BipointedJoinBool BipointedJoinBool' +-- -- fst (Iso.fun IS1 F) (a , p) b = F a b (inl (sym p)) +-- -- fst (snd (Iso.fun IS1 F)) (b , p) a = F a b (inr (sym p)) +-- -- snd (snd (Iso.fun IS1 F)) (a , p) (b , q) = cong (F a b) (push (sym p) (sym q)) +-- -- Iso.inv IS1 (l , r , lr) a b (inl x) = l (a , sym x) b +-- -- Iso.inv IS1 (l , r , lr) a b (inr x) = r (b , sym x) a +-- -- Iso.inv IS1 (l , r , lr) a b (push p q i) = lr (a , sym p) (b , sym q) i +-- -- Iso.rightInv IS1 _ = refl +-- -- Iso.leftInv IS1 F = funExt λ a → funExt λ b → funExt +-- -- λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} + +-- -- IS2 : Iso BipointedJoinBool' BipointedJoinBool'' +-- -- IS2 = compIso +-- -- (Σ-cong-iso +-- -- {B = λ l → Σ[ r ∈ ((a : _) → JS'r a) ] +-- -- ((a : _) (b : _) → l a (fst b) ≡ r b (fst a))} +-- -- {B' = λ l → Σ[ r ∈ ((a : _) → JS'r a) ] +-- -- ((b : _) → l (fst b) ≡ r b (pt A))} +-- -- (DomainContrIso (isContrSingl (pt A))) +-- -- λ x → Σ-cong-iso-snd λ r → DomainContrIso (isContrSingl (pt A))) +-- -- (Σ-cong-iso-snd +-- -- λ l → Σ-cong-iso {B = λ r → ((b : _) → l (fst b) ≡ r b (pt A))} +-- -- {B' = λ r → (l (pt B) ≡ r (pt A))} +-- -- (DomainContrIso (isContrSingl (pt B))) +-- -- λ _ → DomainContrIso (isContrSingl (pt B))) + +-- -- FullIso : Iso BipointedJoinBool BipointedJoinBool'' +-- -- FullIso = compIso IS1 IS2 + +-- -- module _ (A B C D T : Pointed ℓ-zero) +-- -- (f : fst A → fst B → fst C → fst D → fst T) where +-- -- JS4 : Type +-- -- JS4 = (a : fst A) (b : fst B) (c : fst C) (d : fst D) +-- -- → join (join (a ≡ snd A) (b ≡ snd B)) (join (c ≡ snd C) (d ≡ snd D)) +-- -- → f a b c d ≡ pt T + +-- -- open import Cubical.HITs.SmashProduct + +-- -- isOfHLevelEM→ : ∀ {ℓ} {G H : AbGroup ℓ} (n : ℕ) +-- -- → isOfHLevel 3 (EM∙ G n →∙ EM∙ H (suc n)) +-- -- isOfHLevelEM→ {G = G} {H = H} n = +-- -- isConnected→∙ (suc n) 2 (isConnectedEM n) +-- -- (subst (λ m → isOfHLevel m (EM H (suc n))) (+-comm 2 (suc n)) +-- -- (hLevelEM _ (suc n))) + +-- -- mainLem' : (A B T : Pointed ℓ-zero) +-- -- → Iso (Σ[ f ∈ (fst A → fst B → fst T) ] BipointedJoinBool A B T f) +-- -- (A ⋀∙ B →∙ T) +-- -- mainLem' A B T = compIso (Σ-cong-iso-snd (λ f → FullIso A B T f)) is2 +-- -- where +-- -- F→ : (T : Type) (t : T) → Σ[ f ∈ (fst A → fst B → T) ] BipointedJoinBool'' A B (T , t) f → (A ⋀∙ B →∙ (T , t)) +-- -- fst (F→ T t (f , fl , fr , flr)) (inl x) = t +-- -- fst (F→ T t (f , fl , fr , flr)) (inr x) = f (fst x) (snd x) +-- -- fst (F→ T t (f , fl , fr , flr)) (push (inl x) i) = fr x (~ i) +-- -- fst (F→ T t (f , fl , fr , flr)) (push (inr x) i) = fl x (~ i) +-- -- fst (F→ T t (f , fl , fr , flr)) (push (push a i₁) i) = flr (~ i₁) (~ i) +-- -- snd (F→ T t (f , fl , fr , flr)) = refl + +-- -- mmB : (T : Type) (f : A ⋀ B → T) +-- -- → Σ[ g ∈ (fst A → fst B → T) ] BipointedJoinBool'' A B (T , f (inl tt)) g +-- -- fst (mmB T f) a b = f (inr (a , b)) +-- -- fst (snd (mmB T f)) b = cong f (sym (push (inr b))) +-- -- fst (snd (snd (mmB T f))) c = cong f (sym (push (inl c))) +-- -- snd (snd (snd (mmB T f))) j i = f (push (push tt (~ j)) (~ i)) + +-- -- mm' : (T : Type) (f : A ⋀ B → T) (t : T) +-- -- → (f (inl tt) ≡ t) +-- -- → BipointedJoinBool'' A B (T , t) (λ a b → f (inr (a , b))) +-- -- mm' T f = J> mmB T f .snd + +-- -- mm : (T : Type) (f : A ⋀ B → T) (t : T) +-- -- → (f (inl tt) ≡ t) +-- -- → Σ[ f ∈ (fst A → fst B → T) ] BipointedJoinBool'' A B (T , t) f +-- -- fst (mm T f t x) a b = f (inr (a , b)) +-- -- snd (mm T f t x) = mm' T f t x + +-- -- c1 : (T : Type) (f : A ⋀ B → T) (t : T) (p : f (inl tt) ≡ t) +-- -- → F→ T t (mm T f t p) ≡ (f , p) +-- -- c1 T f = +-- -- J> cong (F→ T (f (inl tt))) +-- -- (ΣPathP (refl , (transportRefl (mmB T f .snd)))) +-- -- ∙ ΣPathP ((funExt ( +-- -- λ { (inl x) → refl +-- -- ; (inr x) → refl +-- -- ; (push (inl x) i) → refl +-- -- ; (push (inr x) i) → refl +-- -- ; (push (push a i₁) i) → refl})) , refl) + +-- -- is2 : Iso (Σ[ f ∈ (fst A → fst B → fst T) ] BipointedJoinBool'' A B T f) +-- -- (A ⋀∙ B →∙ T) +-- -- Iso.fun is2 f = F→ (fst T) (snd T) f +-- -- Iso.inv is2 f = mm (fst T) (fst f) (snd T) (snd f) +-- -- Iso.rightInv is2 f = c1 (fst T) (fst f) _ (snd f) +-- -- Iso.leftInv is2 f = ΣPathP (refl , transportRefl (snd f)) + +-- -- mainLem : (A B T : Pointed ℓ-zero) +-- -- → Iso (Σ[ f ∈ (fst A → fst B → fst T) ] BipointedJoinBool A B T f) +-- -- (A →∙ (B →∙ T ∙)) +-- -- mainLem A B T = compIso (mainLem' A B T) SmashAdjIso + +-- -- open Iso +-- -- open import Cubical.Foundations.Equiv.HalfAdjoint + + +-- -- Iso-BipointedUnordJoin-BipointedJoinBool : +-- -- ∀ {ℓ} (A : Bool → Pointed ℓ) (B : Pointed ℓ) +-- -- → (f : ((x : Bool) → A x .fst) → fst B) +-- -- → Iso (BipointedUnordJoin (RP∞'∙ ℓ) A B f) +-- -- (BipointedJoinBool (A true) (A false) B +-- -- λ a b → f (CasesBool true a b)) +-- -- Iso-BipointedUnordJoin-BipointedJoinBool A B f = +-- -- compIso (codomainIsoDep (λ g → domIso join-UnordJoinR-iso)) +-- -- (compIso (ΠIso ΠBool×Iso λ g +-- -- → codomainIso (pathToIso (cong (_≡ B .snd) (cong f (sym (CasesBoolη g)))))) curryIso) + +-- -- SteenrodFunType : (X : RP∞' ℓ-zero) (n : fst X → ℕ) → Type +-- -- SteenrodFunType X n = +-- -- Σ[ f ∈ (((x : fst X) → EM ℤ/2 (n x)) → EM ℤ/2 (∑RP∞' X n)) ] +-- -- BipointedUnordJoin X (λ x → EM∙ ℤ/2 (n x)) (EM∙ ℤ/2 (∑RP∞' X n)) f + +-- -- isSetSteenrodFunTypeBoolIso : (n : Bool → ℕ) +-- -- → Iso (SteenrodFunType (RP∞'∙ ℓ-zero) n) +-- -- (EM∙ ℤ/2 (n true) →∙ (EM∙ ℤ/2 (n false) →∙ EM∙ ℤ/2 (n true +' n false) ∙)) +-- -- isSetSteenrodFunTypeBoolIso n = +-- -- (compIso (Σ-cong-iso-snd (λ f → Iso-BipointedUnordJoin-BipointedJoinBool _ _ _)) +-- -- (compIso (invIso (Σ-cong-iso (compIso (invIso curryIso) (invIso (ΠIso ΠBool×Iso λ f → idIso))) +-- -- λ g → idIso)) +-- -- (mainLem (EM∙ ℤ/2 (n true)) (EM∙ ℤ/2 (n false)) +-- -- (EM∙ ℤ/2 (n true +' n false))))) + +-- -- isSetSteenrodFunTypeBoolIsoId : (n : Bool → ℕ) (f : _) (x : _) (y : _) +-- -- → Iso.fun (isSetSteenrodFunTypeBoolIso n) f .fst x .fst y ≡ fst f (×→ΠBool (x , y)) +-- -- isSetSteenrodFunTypeBoolIsoId n f x y = transportRefl _ + +-- -- isSetSteenrodFunType : (X : RP∞' ℓ-zero) (n : fst X → ℕ) → isSet (SteenrodFunType X n) +-- -- isSetSteenrodFunType = RP∞'pt→Prop (λ _ → isPropΠ λ _ → isPropIsOfHLevel 2) +-- -- λ n → isOfHLevelRetractFromIso 2 +-- -- (isSetSteenrodFunTypeBoolIso n) +-- -- (isConnected→∙ (suc (n true)) 1 +-- -- (isConnectedEM (n true)) +-- -- (subst (λ m → isOfHLevel m (EM∙ ℤ/2 (n false) →∙ EM∙ ℤ/2 (n true +' n false))) +-- -- (cong suc (+-comm 1 (n true)) ) +-- -- (isConnected→∙ (suc (n false)) (suc (n true)) +-- -- (isConnectedEM (n false)) +-- -- (subst (λ m → isOfHLevel (suc m) (EM ℤ/2 (n true +' n false))) +-- -- (cong suc (+'≡+ (n true) (n false)) ∙ (+-comm (suc (n true)) (n false))) +-- -- (hLevelEM ℤ/2 (n true +' n false)))))) + +-- -- SteenrodFunType≡ : (X : RP∞' ℓ-zero) (n : fst X → ℕ) +-- -- (f g : SteenrodFunType X n) +-- -- → fst f ≡ fst g +-- -- → f ≡ g +-- -- SteenrodFunType≡ = +-- -- RP∞'pt→Prop (λ X → isPropΠ4 λ n _ _ _ → isSetSteenrodFunType X n _ _) +-- -- λ n f g p → sym (Iso.leftInv (isSetSteenrodFunTypeBoolIso n) f) +-- -- ∙∙ cong (inv (isSetSteenrodFunTypeBoolIso n)) +-- -- (→∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) +-- -- (funExt (λ x → →∙Homogeneous≡ (isHomogeneousEM _) +-- -- (funExt λ y → transportRefl _ +-- -- ∙ funExt⁻ p (×→ΠBool (x , y)) +-- -- ∙ sym (transportRefl _))))) +-- -- ∙∙ Iso.leftInv (isSetSteenrodFunTypeBoolIso n) g + +-- -- cp : {n m : ℕ} → EM ℤ/2 n → EM ℤ/2 m → EM ℤ/2 (n +' m) +-- -- cp = _⌣ₖ_ {G'' = ℤ/2Ring} + +-- -- Sq : (X : RP∞' ℓ-zero) (x : fst X) (n : fst X → ℕ) +-- -- (f : ((x₁ : fst X) → EM ℤ/2 (n x₁))) → EM ℤ/2 (∑RP∞' X n) +-- -- Sq X x n f = +-- -- subst (EM ℤ/2) (sym (∑RP∞'≡ X x n)) +-- -- (cp (f x) (f (RP∞'-fields.notRP∞' X x))) + +-- -- Sq-bool : (n : Bool → ℕ) (f : ((x₁ : Bool) → EM ℤ/2 (n x₁))) +-- -- → EM ℤ/2 (∑RP∞' (RP∞'∙ ℓ-zero) n) +-- -- Sq-bool n f = cp (f true) (f false) + +-- -- sq≡ : (n : Bool → ℕ) (f : ((x₁ : Bool) → EM ℤ/2 (n x₁))) +-- -- → Sq (RP∞'∙ ℓ-zero) true n f ≡ Sq-bool n f +-- -- sq≡ n f = (λ j → subst (EM ℤ/2) +-- -- (isSetℕ _ _ (sym (∑RP∞'≡ (RP∞'∙ ℓ-zero) true n)) refl j) +-- -- (Sq-bool n f)) +-- -- ∙ transportRefl _ + + +-- -- sq' : (n : Bool → ℕ) +-- -- → SmashPt (EM∙ ℤ/2 (n true)) (EM∙ ℤ/2 (n false)) +-- -- →∙ EM∙ ℤ/2 (n true +' n false) +-- -- fst (sq' n) basel = 0ₖ (n true +' n false) +-- -- fst (sq' n) baser = 0ₖ (n true +' n false) +-- -- fst (sq' n) (proj x y) = cp x y +-- -- fst (sq' n) (gluel a i) = ⌣ₖ-0ₖ {G'' = ℤ/2Ring} (n true) (n false) a i +-- -- fst (sq' n) (gluer b i) = 0ₖ-⌣ₖ {G'' = ℤ/2Ring} (n true) (n false) b i +-- -- snd (sq' n) = refl + +-- -- sq-coh-bool : (n : Bool → ℕ) +-- -- → BipointedJoinBool (EM∙ ℤ/2 (n true)) (EM∙ ℤ/2 (n false)) +-- -- (EM∙ ℤ/2 (∑RP∞' (RP∞'∙ ℓ-zero) n)) +-- -- (λ a b → Sq-bool n (CasesBool true a b)) +-- -- sq-coh-bool n = +-- -- mainLem' (EM∙ ℤ/2 (n true)) (EM∙ ℤ/2 (n false)) +-- -- (EM∙ ℤ/2 (∑RP∞' (RP∞'∙ ℓ-zero) n)) .Iso.inv +-- -- (sq' n ∘∙ (⋀→Smash , refl)) .snd + +-- -- sq-coh : (X : RP∞' ℓ-zero) (x : fst X) (n : fst X → ℕ) +-- -- → BipointedUnordJoin X (λ x₁ → EM∙ ℤ/2 (n x₁)) +-- -- (EM∙ ℤ/2 (∑RP∞' X n)) (Sq X x n) +-- -- sq-coh = JRP∞' λ n → Iso-BipointedUnordJoin-BipointedJoinBool +-- -- (λ x₁ → EM∙ ℤ/2 (n x₁)) +-- -- (EM∙ ℤ/2 (∑RP∞' (RP∞'∙ ℓ-zero) n)) +-- -- (Sq (RP∞'∙ ℓ-zero) true n) .Iso.inv +-- -- λ g x r → sq≡ n _ ∙ sq-coh-bool n g x r + +-- -- comm-field : (X : RP∞' ℓ-zero) (x : fst X) (n : fst X → ℕ) +-- -- → Path (SteenrodFunType X n) +-- -- (Sq X x n +-- -- , sq-coh X x n) +-- -- (Sq X (RP∞'-fields.notRP∞' X x) n +-- -- , sq-coh X (RP∞'-fields.notRP∞' X x) n) +-- -- comm-field = JRP∞' λ n → SteenrodFunType≡ (RP∞'∙ ℓ-zero) n _ _ +-- -- (funExt λ f → cong (subst (EM ℤ/2) (sym (∑RP∞'≡ (RP∞'∙ ℓ-zero) true n))) +-- -- (⌣ₖ-commℤ/2 (n true) (n false) (f true) (f false)) +-- -- ∙ help _ (+'-comm (n false) (n true)) _ +-- -- (sym (∑RP∞'≡ (RP∞'∙ ℓ-zero) true n)) +-- -- (sym (∑RP∞'≡ (RP∞'∙ ℓ-zero) false n)) +-- -- (cp (f false) (f true))) +-- -- where +-- -- help : {n : ℕ} (m : ℕ) (p : n ≡ m) (l : ℕ) (q : m ≡ l) (r : n ≡ l) +-- -- → (x : EM ℤ/2 n) +-- -- → subst (EM ℤ/2) q (subst (EM ℤ/2) p x) ≡ subst (EM ℤ/2) r x +-- -- help = J> (J> λ r x +-- -- → transportRefl _ +-- -- ∙ λ j → subst (EM ℤ/2) (isSetℕ _ _ refl r j) x) + + +-- -- SQ : (X : RP∞' ℓ-zero) (n : fst X → ℕ) → SteenrodFunType X n +-- -- SQ X n = RP∞'→SetRec (isSetSteenrodFunType X n) X +-- -- (λ x → Sq X x n , sq-coh X x n) +-- -- (λ x → comm-field X x n) + +-- -- STSQ : (X : RP∞' ℓ-zero) (n : fst X → ℕ) +-- -- (f : ((x₁ : fst X) → EM ℤ/2 (n x₁))) → EM ℤ/2 (∑RP∞' X n) +-- -- STSQ X n f = SQ X n .fst f + +-- -- SQId : (X : RP∞' ℓ-zero) (n : fst X → ℕ) +-- -- (x : fst X) (g : (x₁ : fst X) → EM ℤ/2 (n x₁)) +-- -- → SQ X n .fst g ≡ Sq X x n g +-- -- SQId X n x g i = +-- -- RP∞'→SetRecEq (isSetSteenrodFunType X n) X +-- -- (λ x → Sq X x n , sq-coh X x n) +-- -- (λ x → comm-field X x n) x i .fst g + +-- -- SQBool : (n : Bool → ℕ) (g : _) → SQ (RP∞'∙ ℓ-zero) n .fst g ≡ cp (g true) (g false) +-- -- SQBool n g = SQId (RP∞'∙ ℓ-zero) n true g ∙ sq≡ n g + +-- -- ΣQuadpointTy : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) +-- -- → Type +-- -- ΣQuadpointTy X Y n = +-- -- Σ[ f ∈ (((x : fst X) (y : fst Y) → EM ℤ/2 (n x y)) +-- -- → EM ℤ/2 (∑RP∞' X λ x → ∑RP∞' Y λ y → n x y)) ] +-- -- QuadpointedUnordJoin X Y +-- -- (λ x y → EM∙ ℤ/2 (n x y)) +-- -- (EM∙ ℤ/2 (∑RP∞' X λ x → ∑RP∞' Y λ y → n x y)) f + +-- -- SQ4 : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) +-- -- → ΣQuadpointTy X Y n +-- -- fst (SQ4 X Y n) f = SQ X (λ x → ∑RP∞' Y (n x)) .fst λ x → SQ Y (n x) .fst (f x) +-- -- snd (SQ4 X Y n) g (inlR p) = SQ X (λ x → ∑RP∞' Y (n x)) .snd _ (inlR (fst p , SQ Y (n (fst p)) .snd _ (snd p))) +-- -- snd (SQ4 X Y n) g (inrR f) = SQ X (λ x → ∑RP∞' Y (n x)) .snd _ (inrR λ x → SQ Y (n x) .snd _ (f x)) +-- -- snd (SQ4 X Y n) g (pushR p f r i) = +-- -- SQ X (λ x → ∑RP∞' Y (n x)) .snd _ +-- -- (pushR (fst p , SQ Y (n (fst p)) .snd _ (snd p)) +-- -- (λ x → SQ Y (n x) .snd _ (f x)) +-- -- (cong (SQ Y (n (fst p)) .snd (λ x₂ → g (fst p) x₂)) r) i) + +-- -- SQ4comm' : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) +-- -- → Σ[ f ∈ (((x : fst X) (y : fst Y) → EM ℤ/2 (n x y)) +-- -- → EM ℤ/2 (∑RP∞' Y λ y → ∑RP∞' X λ x → n x y)) ] +-- -- QuadpointedUnordJoin Y X +-- -- (λ y x → EM∙ ℤ/2 (n x y)) +-- -- (EM∙ ℤ/2 (∑RP∞' Y λ y → ∑RP∞' X λ x → n x y)) +-- -- (λ g → f (λ x y → g y x)) +-- -- fst (SQ4comm' X Y n) f = +-- -- SQ Y (λ z → ∑RP∞' X (λ x → n x z)) .fst +-- -- λ y → SQ X (λ x → n x y) .fst λ x → f x y +-- -- snd (SQ4comm' X Y n) g (inlR x) = +-- -- SQ Y (λ z → ∑RP∞' X (λ x₁ → n x₁ z)) .snd +-- -- (λ y → SQ X (λ x₁ → n x₁ y) .fst (λ x₁ → g y x₁)) +-- -- (inlR (fst x , SQ X (λ x₁ → n x₁ (fst x)) .snd (g (fst x)) (snd x))) +-- -- snd (SQ4comm' X Y n) g (inrR f) = +-- -- SQ Y (λ y → ∑RP∞' X (λ x → n x y)) .snd +-- -- (λ y → SQ X (λ x → n x y) .fst (g y)) +-- -- (inrR λ y → SQ X (λ x → n x y) .snd (g y) (f y)) +-- -- snd (SQ4comm' X Y n) g (pushR a b x i) = +-- -- SQ Y (λ y → ∑RP∞' X (λ x → n x y)) .snd +-- -- (λ y → SQ X (λ x₁ → n x₁ y) .fst (g y)) +-- -- (pushR (fst a , SQ X (λ x → n x (fst a)) .snd (g (fst a)) (snd a)) +-- -- (λ y → SQ X (λ x → n x y) .snd (g y) (b y)) +-- -- (cong (SQ X (λ x₁ → n x₁ (fst a)) .snd (g (fst a))) x) i) + +-- -- SQ4comm : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) +-- -- → ΣQuadpointTy X Y n +-- -- fst (SQ4comm X Y n) f = +-- -- subst (EM ℤ/2) (∑RP∞'Fubini X Y n) +-- -- (SQ Y (λ z → ∑RP∞' X (λ x → n x z)) .fst +-- -- λ y → SQ X (λ x → n x y) .fst λ x → f x y) +-- -- snd (SQ4comm X Y n) f p = +-- -- cong (subst (EM ℤ/2) (∑RP∞'Fubini X Y n)) +-- -- (SQ4comm' X Y n .snd (λ y x → f x y) +-- -- (UnordJoinFubiniFun X Y _ p)) +-- -- ∙ subst-EM∙ (∑RP∞'Fubini X Y n) .snd + +-- -- mainA : (A B C D T : Pointed ℓ-zero) +-- -- → Iso (Σ[ f ∈ (fst A → fst B → fst C → fst D → fst T) ] +-- -- JS4 A B C D T f) +-- -- (A →∙ (B →∙ C →∙ D →∙ T ∙ ∙ ∙)) +-- -- mainA A B C D T = +-- -- compIso +-- -- (compIso IsMain +-- -- (Σ-cong-iso +-- -- (codomainIso (codomainIso (mainLem C D T))) +-- -- λ f → codomainIsoDep λ a +-- -- → codomainIsoDep λ b +-- -- → codomainIso +-- -- (compIso (congIso {x = f a b} (mainLem C D T)) +-- -- (pathToIso (cong (fun (mainLem C D T) (f a b) ≡_) +-- -- (mainIs .snd)) )))) +-- -- (mainLem A B (C →∙ D →∙ T ∙ ∙)) +-- -- where +-- -- T1 = (Σ[ f ∈ (fst A → fst B → fst C → fst D → fst T) ] JS4 A B C D T f) +-- -- T2 = (Σ (fst A → fst B → Σ (fst C → fst D → fst T) (BipointedJoinBool C D T)) +-- -- (BipointedJoinBool A B (Σ (fst C → fst D → fst T) (BipointedJoinBool C D T) +-- -- , (λ _ _ → snd T) , (λ _ _ _ → refl) ))) + +-- -- module _ (a : fst A) (b : fst B) (c : fst C) (d : fst D) +-- -- (p : join (a ≡ snd A) (b ≡ snd B)) +-- -- (q : join (c ≡ snd C) (d ≡ snd D)) (i j k : I) where +-- -- fill₁ : T1 → fst T +-- -- fill₁ (f , g) = +-- -- hfill (λ k → λ {(i = i0) → g a b c d (push p q k) j +-- -- ; (i = i1) → snd T +-- -- ; (j = i0) → g a b c d (inl p) i +-- -- ; (j = i1) → snd T}) +-- -- (inS (g a b c d (inl p) (i ∨ j))) k + +-- -- fill₂ : T2 → fst T +-- -- fill₂ (f , g) = +-- -- hfill (λ k → λ {(i = i0) → g a b p j .snd c d q (~ k) +-- -- ; (i = i1) → f a b .snd c d q (~ k ∨ j) +-- -- ; (j = i0) → f a b .snd c d q (~ k) +-- -- ; (j = i1) → snd T}) +-- -- (inS (snd T)) k + +-- -- T1→T2 : T1 → T2 +-- -- fst (fst (T1→T2 (f , p)) a b) c d = f a b c d +-- -- snd (fst (T1→T2 (f , p)) a b) c d t = p a b c d (inr t) +-- -- fst (snd (T1→T2 (f , p)) a b t i) c d = p a b c d (inl t) i +-- -- snd (snd (T1→T2 (f , p)) a b t i) c d q j = fill₁ a b c d t q i j i1 (f , p) + +-- -- T2→T1 : T2 → T1 +-- -- fst (T2→T1 (f , p)) a b c d = f a b .fst c d +-- -- snd (T2→T1 (f , p)) a b c d (inl x) i = p a b x i .fst c d +-- -- snd (T2→T1 (f , p)) a b c d (inr x) = f a b .snd c d x +-- -- snd (T2→T1 (f , g)) a b c d (push p q i) j = fill₂ a b c d p q i j i1 (f , g) + +-- -- IsMain : Iso T1 T2 +-- -- Iso.fun IsMain = T1→T2 +-- -- Iso.inv IsMain = T2→T1 +-- -- fst (Iso.rightInv IsMain (f , p) i) = fst (T1→T2 (T2→T1 (f , p))) +-- -- fst (snd (Iso.rightInv IsMain (f , p) i) a b t j) = p a b t j .fst +-- -- snd (snd (Iso.rightInv IsMain (f , g) i) a b p j) c d q k = +-- -- hcomp (λ r → λ {(i = i0) → fill₁ a b c d p q j k r ((λ a b c d → f a b .fst c d) , (snd (T2→T1 (f , g)))) +-- -- ; (i = i1) → snd (g a b p j) c d q k +-- -- ; (j = i0) → sd r i k +-- -- ; (j = i1) → snd T +-- -- ; (k = i0) → g a b p j .fst c d +-- -- ; (k = i1) → snd T}) +-- -- (cb k j i) +-- -- where +-- -- sq : (k i r : I) → fst T +-- -- sq k i r = +-- -- hfill (λ r → λ {(i = i0) → g a b p k .snd c d q (~ r) +-- -- ; (i = i1) → f a b .snd c d q (~ r ∨ k) +-- -- ; (k = i0) → f a b .snd c d q (~ r) +-- -- ; (k = i1) → snd T}) +-- -- (inS (snd T)) +-- -- r + +-- -- sd : Cube (λ i j → sq j i i1) +-- -- (λ i k → f a b .snd c d q k) +-- -- (λ r k → fill₂ a b c d p q r k i1 +-- -- (f , λ a b p → g a b p)) +-- -- (λ r k → snd (f a b) c d q k) +-- -- (λ r i → f a b .fst c d) (λ _ _ → pt T) +-- -- sd r i k = +-- -- hcomp (λ w → λ {(r = i0) → sq k i w +-- -- ; (r = i1) → f a b .snd c d q (~ w ∨ k) +-- -- ; (i = i0) → fill₂ a b c d p q r k w (f , λ a b p → g a b p) +-- -- ; (i = i1) → f a b .snd c d q (~ w ∨ k) +-- -- ; (k = i0) → f a b .snd c d q (~ w) +-- -- ; (k = i1) → snd T}) +-- -- (snd T) + +-- -- cb : Cube (λ j i → g a b p j .fst c d) (λ _ _ → pt T) +-- -- (λ i j → sq i j i1) (λ _ _ → pt T) +-- -- (λ k j → g a b p (j ∨ k) .fst c d) +-- -- λ k j → snd (g a b p j) c d q k +-- -- cb k j i = +-- -- hcomp (λ r → λ {(i = i0) → g a b p (k ∨ j) .snd c d q (~ r) +-- -- ; (i = i1) → snd (g a b p j) c d q (~ r ∨ k) +-- -- ; (j = i1) → snd T +-- -- ; (k = i0) → g a b p j .snd c d q (~ r) +-- -- ; (k = i1) → snd T}) +-- -- (snd T) +-- -- fst (Iso.leftInv IsMain (f , g) i) = f +-- -- snd (Iso.leftInv IsMain (f , g) i) a b c d (inl p) = g a b c d (inl p) +-- -- snd (Iso.leftInv IsMain (f , g) i) a b c d (inr p) = g a b c d (inr p) +-- -- snd (Iso.leftInv IsMain (f , g) i) a b c d (push p q j) k = +-- -- hcomp (λ r → λ {(i = i0) → fill₂ a b c d p q j k r +-- -- (fst (T1→T2 (f , g)) , snd (T1→T2 (f , g))) +-- -- ; (i = i1) → ss r j k +-- -- ; (j = i0) → s2 r k i +-- -- ; (j = i1) → g a b c d (inr q) (~ r ∨ k) +-- -- ; (k = i0) → g a b c d (inr q) (~ r) +-- -- ; (k = i1) → pt T}) +-- -- (snd T) +-- -- where +-- -- PP : (i j k : I) → fst T +-- -- PP i j k = doubleWhiskFiller {A = λ i → g a b c d (inr q) (~ i) ≡ pt T} refl +-- -- (λ i j → g a b c d (inr q) (~ i ∨ j)) +-- -- (λ j k → g a b c d (push p q (~ j)) k) +-- -- k i j + +-- -- s2 : Cube (λ _ _ → pt T) (λ k i → g a b c d (inl p) k) +-- -- (λ r i → g a b c d (inr q) (~ r)) (λ _ _ → pt T) +-- -- (λ r k → fill₁ a b c d p q k (~ r) i1 (f , g)) +-- -- λ r k → PP r k i1 +-- -- s2 r k i = +-- -- hcomp (λ j → λ {(r = i0) → pt T +-- -- ; (r = i1) → g a b c d (push p q (~ j ∧ i)) k +-- -- ; (k = i0) → g a b c d (push p q (j ∨ i)) (~ r) +-- -- ; (k = i1) → pt T +-- -- ; (i = i0) → fill₁ a b c d p q k (~ r) j (f , g) +-- -- ; (i = i1) → PP r k j}) +-- -- (g a b c d (push p q i) (k ∨ ~ r)) + +-- -- ss : Cube (λ _ _ → pt T) (λ j k → g a b c d (push p q j) k) +-- -- (λ i j → PP i j i1) +-- -- (λ r k → g a b c d (inr q) (~ r ∨ k)) +-- -- (λ r j → g a b c d (inr q) (~ r)) +-- -- (λ r j → pt T) +-- -- ss r j k = +-- -- hcomp (λ w → λ {(r = i0) → pt T +-- -- ; (r = i1) → g a b c d (push p q (~ w ∨ j)) k +-- -- ; (j = i1) → g a b c d (inr q) (~ r ∨ k) +-- -- ; (k = i0) → g a b c d (inr q) (~ r) +-- -- ; (k = i1) → pt T}) +-- -- (g a b c d (inr q) (~ r ∨ k)) + +-- -- mainIs : (isoToPath (mainLem C D T) i0 , (λ c d → pt T) , λ _ _ _ → refl) +-- -- ≃∙ (C →∙ (D →∙ T ∙) ∙) +-- -- fst mainIs = isoToEquiv (mainLem C D T) +-- -- snd mainIs = cong (Iso.fun SmashAdjIso) +-- -- (ΣPathP ((funExt ( +-- -- λ { (inl x) → refl +-- -- ; (inr x) → refl +-- -- ; (push (inl x) i) → refl +-- -- ; (push (inr x) i) → refl +-- -- ; (push (push a i₁) i) → refl})) , refl)) +-- -- ∙ SmashAdj≃∙ .snd +-- -- {- +-- -- ((isoToEquiv (mainLem C D T))) +-- -- (cong (Iso.fun SmashAdjIso) +-- -- (ΣPathP ((funExt ( +-- -- λ { (inl x) → refl +-- -- ; (inr x) → refl +-- -- ; (push (inl x) i) → refl +-- -- ; (push (inr x) i) → refl +-- -- ; (push (push a i₁) i) → refl})) , refl)) +-- -- ∙ SmashAdj≃∙ .snd) +-- -- -} + +-- -- ΣQuadpointTyBool : (n : Bool → Bool → ℕ) +-- -- → Iso (ΣQuadpointTy (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n) +-- -- (EM∙ ℤ/2 (n true true) +-- -- →∙ (EM∙ ℤ/2 (n true false) +-- -- →∙ EM∙ ℤ/2 (n false true) +-- -- →∙ EM∙ ℤ/2 (n false false) +-- -- →∙ EM∙ ℤ/2 ((n true true +' n true false) +-- -- +' (n false true +' n false false)) ∙ ∙ ∙)) +-- -- ΣQuadpointTyBool n = +-- -- (compIso +-- -- (Σ-cong-iso +-- -- (invIso (compIso (invIso curryIso) +-- -- (compIso (invIso curryIso) +-- -- (compIso (invIso curryIso) +-- -- (domIso (invIso help)))))) +-- -- λ f → invIso (compIso +-- -- (compIso (invIso curryIso) +-- -- (compIso (invIso curryIso) +-- -- (compIso (invIso curryIso) +-- -- (invIso +-- -- (ΠIso idIso +-- -- λ {(((x , y) , z) , w) +-- -- → domIso (compIso join-UnordJoinR-iso +-- -- (Iso→joinIso +-- -- join-UnordJoinR-iso +-- -- join-UnordJoinR-iso))}))))) +-- -- (ΠIso (invIso help) λ _ → idIso))) +-- -- (mainA (EM∙ ℤ/2 (n true true)) +-- -- (EM∙ ℤ/2 (n true false)) +-- -- (EM∙ ℤ/2 (n false true)) +-- -- (EM∙ ℤ/2 (n false false)) +-- -- (EM∙ ℤ/2 ((n true true +' n true false) +-- -- +' (n false true +' n false false))))) +-- -- where +-- -- help : Iso ((x y : fst (RP∞'∙ ℓ-zero)) → EM∙ ℤ/2 (n x y) .fst) +-- -- (((EM ℤ/2 (n true true) +-- -- × EM ℤ/2 (n true false)) +-- -- × EM ℤ/2 (n false true)) +-- -- × EM ℤ/2 (n false false)) +-- -- help = compIso (compIso ΠBool×Iso +-- -- (prodIso ΠBool×Iso ΠBool×Iso)) +-- -- (invIso Σ-assoc-Iso) + +-- -- {- +-- -- ΣQuadpointTyPres : (n : Bool → Bool → ℕ) +-- -- (f : ΣQuadpointTy (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n) +-- -- → Path (EM ℤ/2 (n true true) → +-- -- (EM ℤ/2 (n true false) → +-- -- EM ℤ/2 (n false true) → +-- -- EM ℤ/2 (n false false) → +-- -- EM ℤ/2 +-- -- ((n true true +' n true false) +-- -- +' (n false true +' n false false)))) +-- -- (λ x y z w → Iso.fun (ΣQuadpointTyBool n) f .fst x .fst y .fst z .fst w) +-- -- λ x y z w → f .fst (CasesBool true +-- -- (CasesBool true x y) +-- -- (CasesBool true z w)) +-- -- ΣQuadpointTyPres n f = refl +-- -- -} + +-- -- isSetΣQuadPoint : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) +-- -- → isSet (ΣQuadpointTy X Y n) +-- -- isSetΣQuadPoint = +-- -- RP∞'pt→Prop (λ Y → isPropΠ2 (λ _ _ → isPropIsSet)) +-- -- (RP∞'pt→Prop (λ Y → isPropΠ (λ _ → isPropIsSet)) +-- -- λ n → isOfHLevelRetractFromIso 2 +-- -- (ΣQuadpointTyBool n) +-- -- (isConnected→∙ (suc (n true true)) 1 +-- -- (isConnectedEM (n true true)) +-- -- (isConnected→∙ (suc (n true false)) (n true true + 1) +-- -- (isConnectedEM (n true false)) +-- -- (isConnected→∙ (suc (n false true)) +-- -- ((n true false) + (n true true + 1)) +-- -- (isConnectedEM (n false true)) +-- -- (isConnected→∙ +-- -- (suc (n false false)) +-- -- (n false true + (n true false + (n true true + 1))) +-- -- (isConnectedEM (n false false)) +-- -- (subst (λ k → isOfHLevel (suc k) (EM ℤ/2 (N' n))) +-- -- (lem n) +-- -- (hLevelEM ℤ/2 (N' n)))))))) +-- -- where +-- -- N' : (n : Bool → Bool → ℕ) → ℕ +-- -- N' n = ((n true true +' n true false) +' (n false true +' n false false)) + +-- -- lem : (n : _) → suc (N' n) +-- -- ≡ (n false false + (n false true + (n true false + (n true true + 1)))) +-- -- lem n = cong suc ((cong₂ _+'_ (+'≡+ (n true true) (n true false)) +-- -- (+'≡+ (n false true) (n false false)) +-- -- ∙ +'≡+ _ _) +-- -- ∙ +-assoc (n true true + n true false ) (n false true) (n false false)) +-- -- ∙ cong (_+ n false false) +-- -- (cong (_+ n false true) +-- -- ((λ i → +-comm (+-comm 1 (n true true) i) (n true false) i)) +-- -- ∙ +-comm _ (n false true)) +-- -- ∙ +-comm _ (n false false) + +-- -- ΣQuadPoint≡ : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) +-- -- (f g : ΣQuadpointTy X Y n) +-- -- → ((t : _) → f .fst t ≡ g .fst t) +-- -- → f ≡ g +-- -- ΣQuadPoint≡ = +-- -- RP∞'pt→Prop (λ X → isPropΠ5 λ Y n _ _ _ → isSetΣQuadPoint X Y n _ _) +-- -- (RP∞'pt→Prop (λ Y → isPropΠ4 λ n _ _ _ +-- -- → isSetΣQuadPoint (RP∞'∙ ℓ-zero) Y n _ _) +-- -- λ n f g p +-- -- → sym (Iso.leftInv (ΣQuadpointTyBool n) f) +-- -- ∙∙ cong (inv (ΣQuadpointTyBool n)) +-- -- (main n f g p) +-- -- ∙∙ Iso.leftInv (ΣQuadpointTyBool n) g) +-- -- where +-- -- module _ (n : Bool → Bool → ℕ) +-- -- (f g : ΣQuadpointTy (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n) +-- -- (p : (t : (x y : fst (RP∞'∙ ℓ-zero)) → EM ℤ/2 (n x y)) +-- -- → f .fst t ≡ g .fst t) where +-- -- p' : (x : _) (y : _) (z : _) (w : _) +-- -- → fun (ΣQuadpointTyBool n) f .fst x .fst y .fst z .fst w +-- -- ≡ fun (ΣQuadpointTyBool n) g .fst x .fst y .fst z .fst w +-- -- p' x y z w = p (CasesBool true +-- -- (CasesBool true x y) +-- -- (CasesBool true z w)) + +-- -- module _ {ℓ ℓ' ℓT} {C : Pointed ℓ} {D : Pointed ℓ'} {T : Pointed ℓT} +-- -- (hom : isHomogeneous T) where +-- -- isHomogeneous→∙₂ : isHomogeneous (C →∙ (D →∙ T ∙) ∙) +-- -- isHomogeneous→∙₂ = isHomogeneous→∙ (isHomogeneous→∙ hom) + +-- -- module _ {ℓ'' : Level} {B : Pointed ℓ''} where +-- -- isHomogeneous→∙₃ : isHomogeneous (B →∙ (C →∙ (D →∙ T ∙) ∙) ∙) +-- -- isHomogeneous→∙₃ = isHomogeneous→∙ isHomogeneous→∙₂ + +-- -- isHomogeneous→∙₄ : ∀ {ℓ'''} {A : Pointed ℓ'''} +-- -- → isHomogeneous (A →∙ (B →∙ (C →∙ (D →∙ T ∙) ∙) ∙) ∙) +-- -- isHomogeneous→∙₄ = isHomogeneous→∙ isHomogeneous→∙₃ + + + +-- -- T = (n true true +' n true false) +' (n false true +' n false false) + +-- -- m4 : (x : EM ℤ/2 (n true true)) (y : EM ℤ/2 (n true false)) +-- -- (z : EM ℤ/2 (n false true)) +-- -- → fun (ΣQuadpointTyBool n) f .fst x .fst y .fst z +-- -- ≡ fun (ΣQuadpointTyBool n) g .fst x .fst y .fst z +-- -- m4 x y z = →∙Homogeneous≡ (isHomogeneousEM T) (funExt (p' x y z)) + +-- -- m3 : (x : EM ℤ/2 (n true true)) (y : EM ℤ/2 (n true false)) +-- -- → fun (ΣQuadpointTyBool n) f .fst x .fst y +-- -- ≡ fun (ΣQuadpointTyBool n) g .fst x .fst y +-- -- m3 x y = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM T)) +-- -- (funExt (m4 x y)) + +-- -- m2 : (x : EM ℤ/2 (n true true)) +-- -- → fun (ΣQuadpointTyBool n) f .fst x +-- -- ≡ fun (ΣQuadpointTyBool n) g .fst x +-- -- m2 x = →∙Homogeneous≡ (isHomogeneous→∙₂ (isHomogeneousEM T)) +-- -- (funExt (m3 x)) + +-- -- main : fun (ΣQuadpointTyBool n) f ≡ fun (ΣQuadpointTyBool n) g +-- -- main = →∙Homogeneous≡ (isHomogeneous→∙₃ (isHomogeneousEM T)) +-- -- (funExt m2) + +-- -- SQ4≡SQ4comm : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) +-- -- → SQ4 X Y n ≡ SQ4comm X Y n +-- -- SQ4≡SQ4comm = +-- -- RP∞'pt→Prop (λ _ → isPropΠ2 λ Y n → isSetΣQuadPoint _ Y n _ _) +-- -- (RP∞'pt→Prop (λ Y → isPropΠ λ n → isSetΣQuadPoint _ Y n _ _) +-- -- λ n → ΣQuadPoint≡ _ _ _ _ _ +-- -- λ f → SQBool (λ x → ∑RP∞' (RP∞'∙ ℓ-zero) (n x)) +-- -- (λ x → SQ (RP∞'∙ ℓ-zero) (n x) .fst (f x)) +-- -- ∙ cong₂ cp (SQBool (n true) (f true)) +-- -- (SQBool (n false) (f false)) +-- -- ∙ help (EM ℤ/2) (λ n m x y → cp {n = n} {m = m} x y) +-- -- ⌣ₖ-commℤ/2 assoc⌣ₖ +-- -- (n true true) (n true false) +-- -- (n false true) (n false false) +-- -- (f true true) (f true false) +-- -- (f false true) (f false false) +-- -- (∑RP∞'Fubini (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n) +-- -- ∙ cong (subst (EM ℤ/2) (∑RP∞'Fubini (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n)) +-- -- (sym (SQBool (λ z → ∑RP∞' (RP∞'∙ ℓ-zero) (λ x → n x z)) +-- -- (λ y → SQ (RP∞'∙ ℓ-zero) (λ x → n x y) .fst (λ x → f x y)) +-- -- ∙ cong₂ cp (SQBool (λ x → n x true) (λ x → f x true)) +-- -- (SQBool (λ x → n x false) (λ x → f x false))))) +-- -- where +-- -- help : ∀ {ℓ} (A : ℕ → Type ℓ) (compA : (n m : ℕ) (x : A n) (y : A m) → A (n +' m)) +-- -- → (⌣comm : (n m : ℕ) (x : A n) (y : A m) +-- -- → compA n m x y +-- -- ≡ subst A (+'-comm m n) (compA m n y x)) +-- -- → (⌣assoc : (n m l : ℕ) (x : A n) (y : A m) (z : A l) +-- -- → compA (n +' m) l (compA n m x y) z +-- -- ≡ subst A (+'-assoc n m l) +-- -- (compA n (m +' l) x (compA m l y z))) +-- -- → (n m k l : ℕ) (x : A n) (y : A m) (z : A k) (w : A l) +-- -- → (p : ((n +' k) +' (m +' l)) ≡ ((n +' m) +' (k +' l))) +-- -- → compA (n +' m) (k +' l) (compA n m x y) (compA k l z w) +-- -- ≡ subst A p (compA (n +' k) (m +' l) (compA n k x z) (compA m l y w)) +-- -- help A compA ⌣comm ⌣assoc n m k l x y z w p = +-- -- (sym (transportRefl _) +-- -- ∙ (λ i → subst A (isSetℕ _ _ refl (((sym (+'-assoc n m (k +' l))) ∙ p5 ∙ p4) ∙ p) i) +-- -- (compA (n +' m) (k +' l) (compA n m x y) (compA k l z w)))) +-- -- ∙ substComposite A ((sym (+'-assoc n m (k +' l)) ∙ p5 ∙ p4)) p _ +-- -- ∙ cong (subst A p) +-- -- ((substComposite A (sym (+'-assoc n m (k +' l))) (p5 ∙ p4) _ +-- -- ∙ cong (subst A (p5 ∙ p4)) +-- -- (cong (subst A (sym (+'-assoc n m (k +' l)))) +-- -- (⌣assoc _ _ _ x y (compA k l z w)) +-- -- ∙ subst⁻Subst A (+'-assoc n m (k +' l)) _)) +-- -- ∙ substComposite A (cong (_+'_ n) ((p1 ∙ p2) ∙ p3)) p4 +-- -- (compA n (m +' (k +' l)) x (compA m (k +' l) y (compA k l z w))) +-- -- ∙ cong (subst A (+'-assoc n k (m +' l))) +-- -- (sym (substLems _ _ ((p1 ∙ p2) ∙ p3) _ .snd +-- -- x (compA m (k +' l) y (compA k l z w))) +-- -- ∙ cong (compA n (k +' (m +' l)) x) +-- -- (substComposite A (p1 ∙ p2) p3 (compA m (k +' l) y (compA k l z w)) +-- -- ∙ cong (subst A p3) +-- -- ((substComposite A p1 p2 (compA m (k +' l) y (compA k l z w)) +-- -- ∙ cong (subst A (cong (_+' l) (+'-comm m k))) +-- -- (sym (⌣assoc m k l y z w))) +-- -- ∙ sym (substLems _ _ (+'-comm m k) _ .fst (compA m k y z) w) +-- -- ∙ cong (λ z → compA (k +' m) l z w) +-- -- (sym (⌣comm k m z y))) +-- -- ∙ cong (subst A p3) +-- -- (⌣assoc _ _ _ z y w) +-- -- ∙ subst⁻Subst A (+'-assoc k m l) _)) +-- -- ∙ sym (⌣assoc _ _ _ x z (compA m l y w))) +-- -- where +-- -- p1 = +'-assoc m k l +-- -- p2 = cong (_+' l) (+'-comm m k) +-- -- p3 = sym (+'-assoc k m l) +-- -- p4 = +'-assoc n k (m +' l) +-- -- p5 = cong (n +'_) ((p1 ∙ p2) ∙ p3) + +-- -- substLems : (n n' : ℕ) (p : n ≡ n') (m : ℕ) +-- -- → ((x : A n) (y : A m) +-- -- → compA n' m (subst A p x) y ≡ subst A (cong (_+' m) p) (compA n m x y)) +-- -- × ((x : A m) (y : A n) +-- -- → compA m n' x (subst A p y) ≡ subst A (cong (m +'_) p) (compA m n x y)) +-- -- substLems n = J> λ m +-- -- → (λ x y → cong (λ x → compA n m x y) (transportRefl x) +-- -- ∙ sym (transportRefl _)) +-- -- , ((λ x y → cong (λ y → compA m n x y) (transportRefl y) +-- -- ∙ sym (transportRefl _))) + +-- -- PreCartan : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) +-- -- → (f : (x : fst X) (y : fst Y) → EM ℤ/2 (n x y)) +-- -- → STSQ X (λ x → ∑RP∞' Y (n x)) (λ x → STSQ Y (n x) (f x)) +-- -- ≡ subst (EM ℤ/2) (∑RP∞'Fubini X Y n) +-- -- (STSQ Y (λ y → ∑RP∞' X (λ x → n x y)) +-- -- (λ y → STSQ X (λ x → n x y) (λ x → f x y))) +-- -- PreCartan X Y n f i = SQ4≡SQ4comm X Y n i .fst f + +-- -- -- module _ (A B C D : Pointed ℓ-zero) (T : Pointed ℓ-zero) +-- -- -- (f : A .fst × B .fst × C .fst × D .fst +-- -- -- → fst T) where +-- -- -- BipointedJoinBool : Type +-- -- -- BipointedJoinBool = (a : A .fst) (b : B .fst) (c : C .fst) (d : D .fst) +-- -- -- → join (a ≡ A .snd) +-- -- -- (join (b ≡ B .snd) +-- -- -- (join (c ≡ C .snd) +-- -- -- (d ≡ D .snd))) +-- -- -- → f (a , b , c , d) ≡ T .snd + +-- -- -- BipointedJoinBool* : Type +-- -- -- BipointedJoinBool* = (b : B .fst) (c : C .fst) (d : D .fst) +-- -- -- → Σ[ fr ∈ ((a : A .fst) → join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd)) +-- -- -- → f (a , b , c , d) ≡ T .snd) ] +-- -- -- ((x : singl (A .snd)) → +-- -- -- (Σ[ fl ∈ (f (x .fst , b , c , d) ≡ T .snd) ] +-- -- -- ((t : join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd))) +-- -- -- → fl ≡ fr (x .fst) t))) + +-- -- -- BipointedJoinBool** : Type +-- -- -- BipointedJoinBool** = (b : B .fst) (c : C .fst) (d : D .fst) +-- -- -- → Σ[ fr ∈ ((a : A .fst) → join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd)) +-- -- -- → f (a , b , c , d) ≡ T .snd) ] +-- -- -- Σ[ fl ∈ (f (pt A , b , c , d) ≡ T .snd) ] +-- -- -- ((t : join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd))) +-- -- -- → fl ≡ fr (pt A) t) + +-- -- -- JS₂ : Type +-- -- -- JS₂ = (c : C .fst) (d : D .fst) +-- -- -- → Σ[ fr ∈ ((a : A .fst) (b : fst B) +-- -- -- → join (c ≡ C .snd) (d ≡ D .snd) +-- -- -- → f (a , b , c , d) ≡ T .snd) ] +-- -- -- Σ[ fl ∈ ((b : fst B) → f (pt A , b , c , d) ≡ T .snd) ] +-- -- -- Σ[ flp ∈ ((b : fst B) → (t : join (c ≡ C .snd) (d ≡ D .snd)) +-- -- -- → fl b ≡ fr (pt A) b t) ] +-- -- -- ((x : singl (B .snd)) +-- -- -- → Σ[ frl ∈ ((a : fst A) → f (a , fst x , c , d) ≡ T .snd) ] +-- -- -- Σ[ frp ∈ ((a : fst A) (t : join (c ≡ C .snd) (d ≡ D .snd)) → frl a ≡ fr a (fst x) t) ] +-- -- -- Σ[ r ∈ fl (fst x) ≡ frl (pt A) ] +-- -- -- ((t : join (c ≡ C .snd) (d ≡ D .snd)) +-- -- -- → Square r (flp (fst x) t) refl (frp (pt A) t))) + +-- -- -- JS₂* : Type +-- -- -- JS₂* = (c : C .fst) (d : D .fst) +-- -- -- → Σ[ fr ∈ ((a : A .fst) (b : fst B) +-- -- -- → join (c ≡ C .snd) (d ≡ D .snd) +-- -- -- → f (a , b , c , d) ≡ T .snd) ] +-- -- -- Σ[ fl ∈ ((b : fst B) → f (pt A , b , c , d) ≡ T .snd) ] +-- -- -- Σ[ flp ∈ ((b : fst B) → (t : join (c ≡ C .snd) (d ≡ D .snd)) +-- -- -- → fl b ≡ fr (pt A) b t) ] +-- -- -- (Σ[ frl ∈ ((a : fst A) → f (a , pt B , c , d) ≡ T .snd) ] +-- -- -- Σ[ frp ∈ ((a : fst A) (t : join (c ≡ C .snd) (d ≡ D .snd)) → frl a ≡ fr a (pt B) t) ] +-- -- -- Σ[ r ∈ fl (pt B) ≡ frl (pt A) ] +-- -- -- ((t : join (c ≡ C .snd) (d ≡ D .snd)) +-- -- -- → Square r (flp (pt B) t) refl (frp (pt A) t))) + +-- -- -- module _ (fl₁ : ((b : fst B) (c : C .fst) (d : D .fst) → f (pt A , b , c , d) ≡ T .snd)) +-- -- -- (fl₂ : ((a : fst A) (c : C .fst) (d : D .fst) → f (a , pt B , c , d) ≡ T .snd)) +-- -- -- (fl₁₂ : (c : fst C) (d : fst D) → fl₁ (pt B) c d ≡ fl₂ (pt A) c d) +-- -- -- where +-- -- -- TL : singl (snd C) → Type +-- -- -- TL (c , p) = +-- -- -- Σ[ fr ∈ ((a : fst A) (b : fst B) (d : fst D) → f (a , b , c , d) ≡ T .snd) ] +-- -- -- Σ[ flp ∈ ((b : fst B) (d : fst D) → fl₁ b c d ≡ fr (pt A) b d) ] +-- -- -- Σ[ frp ∈ ((a : fst A) (d : fst D) → fl₂ a c d ≡ fr a (pt B) d) ] +-- -- -- ((d : fst D) → Square (fl₁₂ c d) (flp (pt B) d) refl (frp (pt A) d)) +-- -- -- TR : singl (snd D) → Type +-- -- -- TR (d , p) = +-- -- -- Σ[ fr ∈ ((a : fst A) (b : fst B) (c : fst C) → f (a , b , c , d) ≡ T .snd) ] +-- -- -- Σ[ flp ∈ ((b : fst B) (c : fst C) → fl₁ b c d ≡ fr (pt A) b c) ] +-- -- -- Σ[ frp ∈ ((a : fst A) (c : fst C) → fl₂ a c d ≡ fr a (pt B) c) ] +-- -- -- ((c : fst C) → Square (fl₁₂ c d) (flp (pt B) c) refl (frp (pt A) c)) + +-- -- -- TLR : (c : singl (snd C)) (d : singl (snd D)) → TL c → TR d → Type +-- -- -- TLR (c , p) (d , q) (frₗ , flpₗ , frpₗ , sqₗ) (frᵣ , flpᵣ , frpᵣ , sqᵣ) = +-- -- -- Σ[ frₗᵣ ∈ (((a : fst A) (b : fst B) → frₗ a b d ≡ frᵣ a b c)) ] +-- -- -- Σ[ flpₗᵣ ∈ ((b : fst B) → PathP (λ i → fl₁ b c d ≡ frₗᵣ (pt A) b i) (flpₗ b d) (flpᵣ b c)) ] +-- -- -- Σ[ frpₗᵣ ∈ ((a : fst A) → PathP (λ i → fl₂ a c d ≡ frₗᵣ a (pt B) i) (frpₗ a d) (frpᵣ a c)) ] +-- -- -- Cube (sqₗ d) (sqᵣ c) +-- -- -- (λ i j → fl₁₂ c d j) (flpₗᵣ (pt B)) +-- -- -- (λ j i → fl₁ (pt B) c d) (frpₗᵣ (pt A)) + + +-- -- -- JS₃* : Type +-- -- -- JS₃* = Σ[ fl₁ ∈ ((b : fst B) (c : C .fst) (d : D .fst) → f (pt A , b , c , d) ≡ T .snd) ] +-- -- -- Σ[ fl₂ ∈ ((a : fst A) (c : C .fst) (d : D .fst) → f (a , pt B , c , d) ≡ T .snd) ] +-- -- -- Σ[ fl₁₂ ∈ ((c : fst C) (d : fst D) → fl₁ (pt B) c d ≡ fl₂ (pt A) c d) ] +-- -- -- Σ[ l ∈ ((c : _) → TL fl₁ fl₂ fl₁₂ c) ] +-- -- -- Σ[ r ∈ ((c : _) → TR fl₁ fl₂ fl₁₂ c) ] +-- -- -- ((c : singl (snd C)) (d : singl (snd D)) → TLR fl₁ fl₂ fl₁₂ c d (l c) (r d)) + +-- -- -- JS₃** : Type +-- -- -- JS₃** = Σ[ fl₁ ∈ ((b : fst B) (c : C .fst) (d : D .fst) → f (pt A , b , c , d) ≡ T .snd) ] +-- -- -- Σ[ fl₂ ∈ ((a : fst A) (c : C .fst) (d : D .fst) → f (a , pt B , c , d) ≡ T .snd) ] +-- -- -- Σ[ fl₁₂ ∈ ((c : fst C) (d : fst D) → fl₁ (pt B) c d ≡ fl₂ (pt A) c d) ] +-- -- -- Σ[ l ∈ (TL fl₁ fl₂ fl₁₂ (pt C , refl)) ] +-- -- -- Σ[ r ∈ (TR fl₁ fl₂ fl₁₂ (pt D , refl)) ] +-- -- -- (TLR fl₁ fl₂ fl₁₂ (pt C , refl) (pt D , refl) l r) + +-- -- -- module _ (js : JS₃**) where +-- -- -- open import Cubical.HITs.SmashProduct +-- -- -- JS₃**→' : (A ⋀∙ (B ⋀∙ (C ⋀∙ D))) →∙ T +-- -- -- fst JS₃**→' (inl x) = pt T +-- -- -- fst JS₃**→' (inr (a , inl x)) = {!f (a , ?)!} -- pt T +-- -- -- fst JS₃**→' (inr (a , inr (b , inl x))) = {!!} -- pt T +-- -- -- fst JS₃**→' (inr (a , inr (b , inr (c , d)))) = f (a , b , c , d) +-- -- -- fst JS₃**→' (inr (a , inr (b , push (inl x) i))) = snd js .snd .snd .snd .fst .fst a b x (~ i) +-- -- -- fst JS₃**→' (inr (a , inr (b , push (inr x) i))) = snd js .snd .snd .fst .fst a b x (~ i) +-- -- -- fst JS₃**→' (inr (a , inr (b , push (push tt i₁) i))) = snd js .snd .snd .snd .snd .fst a b (~ i₁) (~ i) +-- -- -- fst JS₃**→' (inr (a , push (inl x) i)) = {!f (a , pt B , pt C , ?)!} -- pt T +-- -- -- fst JS₃**→' (inr (a , push (inr (inl x)) i)) = {!f (a , pt B , pt C , x)!} -- snd T +-- -- -- fst JS₃**→' (inr (a , push (inr (inr (c , d))) i)) = snd js .fst a c d (~ i) +-- -- -- fst JS₃**→' (inr (a , push (inr (push (inl x) i₁)) i)) = {!!} +-- -- -- fst JS₃**→' (inr (a , push (inr (push (inr x) i₁)) i)) = {!snd js .snd .snd .fst .snd .snd .fst a x (~ i₁) (~ i)!} +-- -- -- fst JS₃**→' (inr (a , push (inr (push (push a₁ i₂) i₁)) i)) = {!!} +-- -- -- fst JS₃**→' (inr (a , push (push a₁ i₁) i)) = {!!} +-- -- -- fst JS₃**→' (push a i) = {!!} +-- -- -- snd JS₃**→' = {!!} + +-- -- -- JS₃**→ : A →∙ (B →∙ (C →∙ (D →∙ T ∙) ∙) ∙) +-- -- -- fst (fst (fst (fst JS₃**→ a) b) c) d = f (a , b , c , d) +-- -- -- snd (fst (fst (fst JS₃**→ a) b) c) = snd js .snd .snd .snd .fst .fst a b c +-- -- -- fst (snd (fst (fst JS₃**→ a) b) i) d = snd js .snd .snd .fst .fst a b d i +-- -- -- snd (snd (fst (fst JS₃**→ a) b) i) = {!snd js .snd .snd .snd .snd .fst a b i!} +-- -- -- fst (fst (snd (fst JS₃**→ a) i) c) d = snd js .fst a c d i +-- -- -- snd (fst (snd (fst JS₃**→ a) i) c) j = {!!} +-- -- -- fst (snd (snd (fst JS₃**→ a) i) j) d = {!!} +-- -- -- snd (snd (snd (fst JS₃**→ a) i) j) k = {!!} +-- -- -- fst (fst (fst (snd JS₃**→ i) b) c) d = fst js b c d i +-- -- -- snd (fst (fst (snd JS₃**→ i) b) c) j = {!!} +-- -- -- fst (snd (fst (snd JS₃**→ i) b) j) d = {!!} +-- -- -- snd (snd (fst (snd JS₃**→ i) b) j) k = {!!} +-- -- -- fst (fst (snd (snd JS₃**→ i) j) c) d = {!!} +-- -- -- snd (fst (snd (snd JS₃**→ i) j) c) k = {!!} +-- -- -- snd (snd (snd JS₃**→ i) j) = {!!} + +-- -- -- Iso-JS₃*-JS₃** : Iso JS₃* JS₃** +-- -- -- Iso-JS₃*-JS₃** = +-- -- -- Σ-cong-iso-snd λ f' → Σ-cong-iso-snd λ g → Σ-cong-iso-snd λ fg +-- -- -- → compIso (Σ-cong-iso-snd (λ r → Σ-cong-iso-snd λ s +-- -- -- → compIso (DomainContrIso (isContrSingl (snd C))) +-- -- -- (DomainContrIso (isContrSingl (snd D))))) +-- -- -- (compIso (ΣDomainContrIso {C = λ c l → Σ ((d : _) → TR f' g fg d) +-- -- -- λ r → TLR f' g fg c (pt D , refl) l (r (pt D , refl))} (isContrSingl (snd C))) +-- -- -- (Σ-cong-iso-snd λ l → +-- -- -- ΣDomainContrIso {C = λ d r → TLR f' g fg (pt C , refl) d l r} (isContrSingl (snd D)))) + +-- -- -- Iso₂ : Iso BipointedJoinBool** JS₂ +-- -- -- fst (Iso.fun Iso₂ F c d) a b t = F b c d .fst a (inr t) +-- -- -- fst (snd (Iso.fun Iso₂ F c d)) b = F b c d .snd .fst +-- -- -- fst (snd (snd (Iso.fun Iso₂ F c d))) b t = F b c d .snd .snd (inr t) +-- -- -- fst (snd (snd (snd (Iso.fun Iso₂ F c d))) (b , p)) a = +-- -- -- F b c d .fst a (inl (sym p)) +-- -- -- fst (snd (snd (snd (snd (Iso.fun Iso₂ F c d))) (b , p))) a t = +-- -- -- cong (F b c d .fst a) (push (sym p) t) +-- -- -- fst (snd (snd (snd (snd (snd (Iso.fun Iso₂ F c d))) (b , p)))) = +-- -- -- F b c d .snd .snd (inl (sym p)) +-- -- -- snd (snd (snd (snd (snd (snd (Iso.fun Iso₂ F c d))) (b , p)))) t = +-- -- -- cong (F b c d .snd .snd) (push (sym p) t) +-- -- -- fst (Iso.inv Iso₂ F b c d) a (inl x) = F c d .snd .snd .snd (b , sym x) .fst a +-- -- -- fst (Iso.inv Iso₂ F b c d) a (inr t) = F c d .fst a b t +-- -- -- fst (Iso.inv Iso₂ F b c d) a (push x t i) = +-- -- -- F c d .snd .snd .snd (b , sym x) .snd .fst a t i +-- -- -- fst (snd (Iso.inv Iso₂ F b c d)) = F c d .snd .fst b +-- -- -- snd (snd (Iso.inv Iso₂ F b c d)) (inl x) = +-- -- -- F c d .snd .snd .snd (b , sym x) .snd .snd .fst +-- -- -- snd (snd (Iso.inv Iso₂ F b c d)) (inr t) = F c d .snd .snd .fst b t +-- -- -- snd (snd (Iso.inv Iso₂ F b c d)) (push a t i) = +-- -- -- F c d .snd .snd .snd (b , sym a) .snd .snd .snd t i +-- -- -- Iso.rightInv Iso₂ = λ _ → refl +-- -- -- Iso.leftInv Iso₂ F = funExt λ b → funExt λ c → funExt λ x → +-- -- -- ΣPathP (funExt (λ a → funExt λ { (inl x) → refl +-- -- -- ; (inr x) → refl +-- -- -- ; (push a x i) → refl}) +-- -- -- , ΣPathP (refl , (funExt (λ { (inl x) → refl +-- -- -- ; (inr x) → refl +-- -- -- ; (push a x i) → refl})))) + +-- -- -- Iso₂* : Iso BipointedJoinBool** JS₂* +-- -- -- Iso₂* = +-- -- -- compIso Iso₂ +-- -- -- (codomainIsoDep λ c → codomainIsoDep λ d → +-- -- -- Σ-cong-iso-snd λ f → Σ-cong-iso-snd λ g → Σ-cong-iso-snd +-- -- -- λ h → DomainContrIso (isContrSingl (B .snd))) + +-- -- -- Iso₃ : Iso JS₂* JS₃* +-- -- -- fst (Iso.fun Iso₃ F) b c d = F c d .snd .fst b +-- -- -- fst (snd (Iso.fun Iso₃ F)) a c d = F c d .snd .snd .snd .fst a +-- -- -- fst (snd (snd (Iso.fun Iso₃ F))) c d = F c d .snd .snd .snd .snd .snd .fst +-- -- -- fst (fst (snd (snd (snd (Iso.fun Iso₃ F)))) (c , p)) a b d = +-- -- -- F c d .fst a b (inl (sym p)) +-- -- -- fst (snd (fst (snd (snd (snd (Iso.fun Iso₃ F)))) (c , p))) b d = +-- -- -- F c d .snd .snd .fst b (inl (sym p)) +-- -- -- fst (snd (snd (fst (snd (snd (snd (Iso.fun Iso₃ F)))) (c , p)))) a d = +-- -- -- F c d .snd .snd .snd .snd .fst a (inl (sym p)) +-- -- -- snd (snd (snd (fst (snd (snd (snd (Iso.fun Iso₃ F)))) (c , p)))) d = +-- -- -- F c d .snd .snd .snd .snd .snd .snd (inl (sym p)) +-- -- -- fst (fst (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (d , p)) a b c = +-- -- -- F c d .fst a b (inr (sym p)) +-- -- -- fst (snd (fst (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (d , p))) b c = +-- -- -- F c d .snd .snd .fst b (inr (sym p)) +-- -- -- fst (snd (snd (fst (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (d , p)))) a c = +-- -- -- F c d .snd .snd .snd .snd .fst a (inr (sym p)) +-- -- -- snd (snd (snd (fst (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (d , p)))) c = +-- -- -- F c d .snd .snd .snd .snd .snd .snd (inr (sym p)) +-- -- -- fst (snd (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (c , p) (d , q)) a b = +-- -- -- cong (F c d .fst a b) (push (sym p) (sym q)) +-- -- -- fst (snd (snd (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (c , p) (d , q))) b i = +-- -- -- F c d .snd .snd .fst b (push (sym p) (sym q) i) +-- -- -- fst (snd (snd (snd (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (c , p) (d , q)))) a i = +-- -- -- F c d .snd .snd .snd .snd .fst a (push (sym p) (sym q) i) +-- -- -- snd (snd (snd (snd (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (c , p) (d , q)))) i = +-- -- -- F c d .snd .snd .snd .snd .snd .snd (push (sym p) (sym q) i) +-- -- -- fst (Iso.inv Iso₃ F c d) a b (inl x) = +-- -- -- F .snd .snd .snd .fst (c , sym x) .fst a b d +-- -- -- fst (Iso.inv Iso₃ F c d) a b (inr x) = +-- -- -- F .snd .snd .snd .snd .fst (d , sym x) .fst a b c +-- -- -- fst (Iso.inv Iso₃ F c d) a b (push p q i) = +-- -- -- F .snd .snd .snd .snd .snd (c , sym p) (d , sym q) .fst a b i +-- -- -- fst (snd (Iso.inv Iso₃ F c d)) b = F .fst b c d +-- -- -- fst (snd (snd (Iso.inv Iso₃ F c d))) b (inl x) = F .snd .snd .snd .fst (c , sym x) .snd .fst b d +-- -- -- fst (snd (snd (Iso.inv Iso₃ F c d))) b (inr x) = F .snd .snd .snd .snd .fst (d , sym x) .snd .fst b c +-- -- -- fst (snd (snd (Iso.inv Iso₃ F c d))) b (push p q i) = +-- -- -- F .snd .snd .snd .snd .snd (c , sym p) (d , sym q) .snd .fst b i +-- -- -- fst (snd (snd (snd (Iso.inv Iso₃ F c d)))) a = F .snd .fst a c d +-- -- -- fst (snd (snd (snd (snd (Iso.inv Iso₃ F c d))))) a (inl x) = +-- -- -- F .snd .snd .snd .fst (c , sym x) .snd .snd .fst a d +-- -- -- fst (snd (snd (snd (snd (Iso.inv Iso₃ F c d))))) a (inr x) = +-- -- -- F .snd .snd .snd .snd .fst (d , sym x) .snd .snd .fst a c +-- -- -- fst (snd (snd (snd (snd (Iso.inv Iso₃ F c d))))) a (push p q i) = +-- -- -- F .snd .snd .snd .snd .snd (c , sym p) (d , sym q) .snd .snd .fst a i +-- -- -- fst (snd (snd (snd (snd (snd (Iso.inv Iso₃ F c d)))))) = F .snd .snd .fst c d +-- -- -- snd (snd (snd (snd (snd (snd (Iso.inv Iso₃ F c d)))))) (inl x) = +-- -- -- F .snd .snd .snd .fst (c , sym x) .snd .snd .snd d +-- -- -- snd (snd (snd (snd (snd (snd (Iso.inv Iso₃ F c d)))))) (inr x) = +-- -- -- F .snd .snd .snd .snd .fst (d , sym x) .snd .snd .snd c +-- -- -- snd (snd (snd (snd (snd (snd (Iso.inv Iso₃ F c d)))))) (push p q i) = +-- -- -- F .snd .snd .snd .snd .snd (c , sym p) (d , sym q) .snd .snd .snd i +-- -- -- Iso.rightInv Iso₃ _ = refl +-- -- -- Iso.leftInv Iso₃ F = funExt λ c → funExt λ d → +-- -- -- ΣPathP ((funExt (λ a → funExt λ b → funExt +-- -- -- λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl})) +-- -- -- , ΣPathP (refl , (ΣPathP ((funExt (λ b +-- -- -- → funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl})) +-- -- -- , (ΣPathP (refl , (ΣPathP ((funExt (λ a → +-- -- -- funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl})) +-- -- -- , (ΣPathP (refl , (funExt (λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl})))))))))))) + + +-- -- -- Σ[ fl ∈ (f (pt A , b , c , d) ≡ T .snd) ] +-- -- -- ((t : join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd))) +-- -- -- → fl ≡ fr (pt A) t) + +-- -- -- Iso₁ : Iso BipointedJoinBool BipointedJoinBool* +-- -- -- fst (Iso.fun Iso₁ F b c d) a x = F a b c d (inr x) +-- -- -- fst (snd (Iso.fun Iso₁ F b c d) (a , p)) = F a b c d (inl (sym p)) +-- -- -- snd (snd (Iso.fun Iso₁ F b c d) (a , p)) t = cong (F a b c d) (push (sym p) t) +-- -- -- Iso.inv Iso₁ F a b c d (inl x) = F b c d .snd (a , sym x) .fst +-- -- -- Iso.inv Iso₁ F a b c d (inr t) = F b c d .fst a t +-- -- -- Iso.inv Iso₁ F a b c d (push p t i) = F b c d .snd (a , sym p) .snd t i +-- -- -- Iso.rightInv Iso₁ = λ _ → refl +-- -- -- Iso.leftInv Iso₁ F = funExt λ a → funExt λ b → funExt λ c → funExt λ d +-- -- -- → funExt λ { (inl x) → refl ; (inr x) → refl ; (push a x i) → refl} + +-- -- -- Iso₁' : Iso BipointedJoinBool BipointedJoinBool** +-- -- -- Iso₁' = compIso Iso₁ (codomainIsoDep λ b → codomainIsoDep λ c → codomainIsoDep λ d +-- -- -- → Σ-cong-iso-snd λ f → DomainContrIso (isContrSingl (A .snd))) + +-- -- -- JoinStructureBool* : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) +-- -- -- (f : A true true .fst × A true false .fst +-- -- -- × A false true .fst × A false false .fst +-- -- -- → fst B) +-- -- -- → Type +-- -- -- JoinStructureBool* A B f = +-- -- -- (g : A true true .fst × A true false .fst +-- -- -- × A false true .fst × A false false .fst) +-- -- -- → join (fst g ≡ A true true .snd) +-- -- -- (join (snd g .fst ≡ A true false .snd) +-- -- -- (join (snd (snd g) .fst ≡ A false true .snd) +-- -- -- (snd (snd g) .snd ≡ A false false .snd))) +-- -- -- → f g ≡ B .snd + +-- -- -- 4→∙ : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) → Type ℓ-zero +-- -- -- 4→∙ A B = A true true →∙ (A true false →∙ (A false true →∙ (A false false →∙ B ∙) ∙) ∙) + +-- -- -- 4→∙' : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) +-- -- -- → Type ℓ-zero +-- -- -- 4→∙' A B = +-- -- -- Σ[ f ∈ (A true true .fst → A true false .fst +-- -- -- → A false true .fst → A false false .fst → fst B) ] +-- -- -- Σ[ f-inl-inl ∈ ((a : singl (A true true .snd)) (b : _) (c : _) (d : _) → f (fst a) b c d ≡ pt B) ] +-- -- -- Σ[ f-inl-inr ∈ ((b : singl (A true false .snd)) (a : _) (c : _) (d : _) → f a (fst b) c d ≡ pt B) ] +-- -- -- Σ[ f-inl-push ∈ (((a : singl (A true true .snd)) (b : singl (A true false .snd)) (c : _) (d : _) +-- -- -- → f-inl-inl a (fst b) c d ≡ f-inl-inr b (fst a) c d)) ] +-- -- -- Σ[ f-inr-inl ∈ ((c : singl (A false true .snd)) (a : _) (b : _) (d : _) → f a b (fst c) d ≡ pt B) ] +-- -- -- Σ[ f-inr-inr ∈ ((d : singl (A false false .snd)) (a : _) (b : _) (c : _) → f a b c (fst d) ≡ pt B) ] +-- -- -- Σ[ f-inl-push ∈ ((c : singl (A false true .snd)) (d : singl (A false false .snd)) (a : _) (b : _) +-- -- -- → f-inr-inl c a b (fst d) ≡ f-inr-inr d a b (fst c)) ] +-- -- -- {!Σ[ f-inl-push ∈ ((c : singl (A false true .snd)) (d : singl (A false false .snd)) (a : _) (b : _) +-- -- -- → f-inr-inl c a b (fst d) ≡ f-inr-inr d a b (fst c)) ]!} + +-- -- -- pss : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) (f : _) → JoinStructureBool A B f +-- -- -- pss A B f (x , y , z , w) (inl (inl p)) = {!p!} +-- -- -- pss A B f (x , y , z , w) (inl (inr q)) = {!q!} +-- -- -- pss A B f (x , y , z , w) (inl (push p q i)) = {!!} +-- -- -- pss A B f (x , y , z , w) (inr (inl p)) = {!!} +-- -- -- pss A B f (x , y , z , w) (inr (inr q)) = {!!} +-- -- -- pss A B f (x , y , z , w) (inr (push p q i)) = {!!} +-- -- -- pss A B f (x , y , z , w) (push (inl p) (inl q) i) = {!!} +-- -- -- pss A B f (x , y , z , w) (push (inr p) (inl q) i) = {!!} +-- -- -- pss A B f (x , y , z , w) (push (push p q i₁) (inl r) i) = {!!} +-- -- -- pss A B f (x , y , z , w) (push p (inr q) i) = {!!} +-- -- -- pss A B f (x , y , z , w) (push p (push q r i₁) i) = {!!} + + +-- -- -- JoinStructureBoolD : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) +-- -- -- → Σ _ (JoinStructureBool A B) +-- -- -- → A true true →∙ (A true false →∙ (A false true →∙ (A false false →∙ B ∙) ∙) ∙) +-- -- -- fst (fst (fst (fst (JoinStructureBoolD A B (f , p)) x) y) z) w = +-- -- -- f (x , y , z , w) +-- -- -- snd (fst (fst (fst (JoinStructureBoolD A B (f , p)) x) y) z) = +-- -- -- p (x , y , z , snd (A false false)) (inr (inr refl)) +-- -- -- fst (snd (fst (fst (JoinStructureBoolD A B (f , p)) x) y) i) w = +-- -- -- p (x , y , snd (A false true) , w) (inr (inl refl)) i +-- -- -- snd (snd (fst (fst (JoinStructureBoolD A B (f , p)) x) y) i) = {!!} +-- -- -- fst (fst (snd (fst (JoinStructureBoolD A B (f , p)) x) i) z) w = +-- -- -- p (x , snd (A true false) , z , w) (inl (inr refl)) i +-- -- -- snd (fst (snd (fst (JoinStructureBoolD A B (f , p)) x) i) z) = {!!} +-- -- -- fst (snd (snd (fst (JoinStructureBoolD A B (f , p)) x) i) j) w = {!!} +-- -- -- snd (snd (snd (fst (JoinStructureBoolD A B (f , p)) x) i) j) = {!!} +-- -- -- fst (fst (fst (snd (JoinStructureBoolD A B (f , p)) i) y) z) w = +-- -- -- p (snd (A true true) , y , z , w) (inl (inl refl)) i +-- -- -- snd (fst (fst (snd (JoinStructureBoolD A B (f , p)) i) y) z) j = {!!} +-- -- -- snd (fst (snd (JoinStructureBoolD A B (f , p)) i) y) = {!!} +-- -- -- snd (snd (JoinStructureBoolD A B (f , p)) i) = {!!} + +-- -- -- module _ {ℓ} (X Y : RP∞' ℓ) (A : fst X → fst Y → Pointed ℓ) (B : Type ℓ) where + +-- -- -- JoinStructure : ((f : (x : fst X) (y : fst Y) → A x y .fst) → B) → Type ℓ +-- -- -- JoinStructure f = +-- -- -- (g : (x : fst X) (y : fst Y) → A x y .fst) +-- -- -- → UnordJoinR² X Y (λ x y → g x y ≡ pt (A x y)) +-- -- -- → f g ≡ f (λ x y → pt (A x y)) + +-- -- -- module _ {ℓ} (X Y : RP∞' ℓ) (A : fst X → fst Y → Pointed ℓ) (B : Type ℓ) where + +-- -- -- JoinStructure→ : (t : (f : (x : fst X) (y : fst Y) → A x y .fst) → B) +-- -- -- → JoinStructure X Y A B t +-- -- -- → JoinStructure Y X (λ y x → A x y) B +-- -- -- λ f → t λ x y → f y x +-- -- -- JoinStructure→ f st g pr = +-- -- -- st (λ x y → g y x) +-- -- -- (UnordJoinFubiniFun Y X (λ x y → g x y ≡ pt (A y x)) pr) + + +-- -- -- inr* : _ → UnordSmash∙² X Y A .fst +-- -- -- inr* = inr + +-- -- -- inr-case : (g : (x₁ : fst X) (y : fst Y) → A x₁ y .fst) (x : fst X) +-- -- -- → UnordJoinR Y (λ y → g x y ≡ pt (A x y)) +-- -- -- → Path (UnordSmash Y (A x)) (inr (g x)) (inr (λ y → pt (A x y))) +-- -- -- inr-case g x (inlR (y , z)) = {!z!} +-- -- -- inr-case g x (inrR z) = {!!} +-- -- -- inr-case g x (pushR a b x₁ i) = {!!} + +-- -- -- m2 : Σ _ (JoinStructure X Y A B) → UnordSmash∙² X Y A .fst → B +-- -- -- m2 (f , p) (inl x) = {!!} +-- -- -- m2 (f , p) (inr t) = {!!} +-- -- -- m2 (f , p) (push a i) = {!!} + +-- -- -- m1 : (f : UnordSmash∙² X Y A .fst → B) +-- -- -- → Σ _ (JoinStructure X Y A B) +-- -- -- fst (m1 f) g = f (inr λ x → inr (g x)) +-- -- -- snd (m1 f) g (inlR (x , inlR r)) = {!cong !} +-- -- -- snd (m1 f) g (inlR (x , inrR r)) = {!!} +-- -- -- snd (m1 f) g (inlR (x , pushR a b x₁ i)) = {!!} +-- -- -- snd (m1 f) g (inrR h) = cong f (cong inr* (funExt λ x → inr-case g x (h x))) +-- -- -- snd (m1 f) g (pushR a b x i) = {!!} + +-- -- -- -- Smash→ : (Σ[ h ∈ (UnordΠ X (fst ∘ A) → fst B) ] +-- -- -- -- ((f : UnordΠ X ((λ r → fst r) ∘ A)) +-- -- -- -- → UnordJoinR X (λ x → f x ≡ pt (A x)) +-- -- -- -- → h f ≡ pt B)) +-- -- -- -- → (UnordSmash X A → fst B) +-- -- -- -- Smash→ (h , f) (inl x) = pt B +-- -- -- -- Smash→ (h , f) (inr x) = h x +-- -- -- -- Smash→ (h , f) (push (x , a) i) = +-- -- -- -- f (UnordΣ→UnordΠ X A (x , a)) +-- -- -- -- (inlR (RP∞'-fields.notRP∞' X x +-- -- -- -- , RP∞'-fields.elimRP∞'β X x a +-- -- -- -- (pt (A (RP∞'-fields.notRP∞' X x))) .snd)) (~ i) From 74596b53b1c3b6d96b13ad5f468386ea76769155 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 23 Jan 2024 03:52:06 +0100 Subject: [PATCH 13/73] .. --- Cubical/Algebra/Group/Instances/IntMod.agda | 55 + Cubical/Cohomology/EilenbergMacLane/Base.agda | 17 + .../EilenbergMacLane/Groups/RPinf.agda | 282 ++ .../EilenbergMacLane/Rings/RPinf.agda | 2337 +++-------------- Cubical/Data/Bool/Properties.agda | 20 + Cubical/Data/Nat/Base.agda | 11 + Cubical/Foundations/Pointed/Homogeneous.agda | 8 + 7 files changed, 776 insertions(+), 1954 deletions(-) create mode 100644 Cubical/Cohomology/EilenbergMacLane/Groups/RPinf.agda diff --git a/Cubical/Algebra/Group/Instances/IntMod.agda b/Cubical/Algebra/Group/Instances/IntMod.agda index b00e675516..d3de7be2be 100644 --- a/Cubical/Algebra/Group/Instances/IntMod.agda +++ b/Cubical/Algebra/Group/Instances/IntMod.agda @@ -4,6 +4,9 @@ module Cubical.Algebra.Group.Instances.IntMod where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv + +open import Cubical.Relation.Nullary open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Unit @@ -215,3 +218,55 @@ isHomℤ→Fin n = -Const-ℤ/2 : (x : fst (ℤGroup/ 2)) → -ₘ x ≡ x -Const-ℤ/2 = ℤ/2-elim refl refl + +pres0→GroupIsoℤ/2 : ∀ {ℓ} {G : Group ℓ} (f : fst G ≃ (ℤGroup/ 2) .fst) + → fst f (GroupStr.1g (snd G)) ≡ fzero + → IsGroupHom (snd G) (fst f) ((ℤGroup/ 2) .snd) +pres0→GroupIsoℤ/2 {G = G} f fp = isGroupHomInv ((invEquiv f) , main) + where + one = invEq f fone + + f⁻∙ : invEq f fzero ≡ GroupStr.1g (snd G) + f⁻∙ = sym (cong (invEq f) fp) ∙ retEq f _ + + f⁻1 : GroupStr._·_ (snd G) (invEq f fone) (invEq f fone) + ≡ GroupStr.1g (snd G) + f⁻1 = sym (retEq f (GroupStr._·_ (snd G) (invEq f fone) (invEq f fone))) + ∙∙ cong (invEq f) (mainlem _ refl ∙ sym fp) + ∙∙ retEq f (GroupStr.1g (snd G)) + where + l : ¬ (fst f (GroupStr._·_ (snd G) (invEq f fone) (invEq f fone)) + ≡ fone) + l p = snotz (cong fst q) + where + q : fone ≡ fzero + q = (sym (secEq f fone) + ∙ cong (fst f) + ((sym (GroupStr.·IdL (snd G) one) + ∙ cong (λ x → GroupStr._·_ (snd G) x one) (sym (GroupStr.·InvL (snd G) one))) + ∙ sym (GroupStr.·Assoc (snd G) (GroupStr.inv (snd G) one) one one))) + ∙ cong (fst f) (cong (GroupStr._·_ (snd G) (GroupStr.inv (snd G) (invEq f fone))) + ((sym (retEq f _) ∙ cong (invEq f) p))) + ∙ cong (fst f) (GroupStr.·InvL (snd G) _) + ∙ fp + + + mainlem : (x : _) + → fst f (GroupStr._·_ (snd G) (invEq f fone) (invEq f fone)) ≡ x + → f .fst ((snd G GroupStr.· invEq f fone) (invEq f fone)) ≡ fzero + mainlem = ℤ/2-elim + (λ p → p) + λ p → ⊥.rec (l p) + + + main : IsGroupHom ((ℤGroup/ 2) .snd) (invEq f) (snd G) + main = + makeIsGroupHom + (ℤ/2-elim + (ℤ/2-elim (f⁻∙ ∙ sym (GroupStr.·IdR (snd G) (GroupStr.1g (snd G))) + ∙ cong (λ x → snd G .GroupStr._·_ x x) (sym f⁻∙)) + (sym (GroupStr.·IdL (snd G) one) + ∙ cong (λ x → snd G .GroupStr._·_ x one) (sym f⁻∙))) + (ℤ/2-elim (sym (GroupStr.·IdR (snd G) one) + ∙ cong (snd G .GroupStr._·_ (invEq f fone)) (sym f⁻∙)) + (f⁻∙ ∙ sym f⁻1))) diff --git a/Cubical/Cohomology/EilenbergMacLane/Base.agda b/Cubical/Cohomology/EilenbergMacLane/Base.agda index 5b87f5e509..12b470cc27 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Base.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Base.agda @@ -435,3 +435,20 @@ fst (coHomTruncEquiv G n) = snd (coHomTruncEquiv G n) = makeIsGroupHom (ST.elim2 (λ _ _ → isSetPathImplicit) λ _ _ → refl) + +EM→-charac : ∀ {ℓ ℓ'} {A : Pointed ℓ} {G : AbGroup ℓ'} (n : ℕ) + → Iso (fst A → EM G n) ((A →∙ EM∙ G n) × EM G n) +Iso.fun (EM→-charac {A = A} n) f = + ((λ x → f x -ₖ f (pt A)) , rCancelₖ n (f (pt A))) , f (pt A) +Iso.inv (EM→-charac n) (f , a) x = fst f x +ₖ a +Iso.rightInv (EM→-charac {A = A} n) ((f , p) , a) = + ΣPathP (→∙Homogeneous≡ (isHomogeneousEM _) + (funExt (λ x → (λ i → (f x +ₖ a) -ₖ (cong (_+ₖ a) p ∙ lUnitₖ n a) i) + ∙ sym (assocₖ n (f x) a (-ₖ a)) + ∙ cong (f x +ₖ_) (rCancelₖ n a) + ∙ rUnitₖ n (f x))) + , cong (_+ₖ a) p ∙ lUnitₖ n a) +Iso.leftInv (EM→-charac {A = A} n) f = + funExt λ x → sym (assocₖ n (f x) (-ₖ f (pt A)) (f (pt A))) + ∙∙ cong (f x +ₖ_) (lCancelₖ n (f (pt A))) + ∙∙ rUnitₖ n (f x) diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/RPinf.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/RPinf.agda new file mode 100644 index 0000000000..ab38f77ec5 --- /dev/null +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/RPinf.agda @@ -0,0 +1,282 @@ +{-# OPTIONS --safe --lossy-unification #-} + +module Cubical.Cohomology.EilenbergMacLane.Groups.RPinf where + +open import Cubical.Cohomology.EilenbergMacLane.Base +open import Cubical.Cohomology.EilenbergMacLane.Gysin + +open import Cubical.Homotopy.EilenbergMacLane.CupProduct +open import Cubical.Homotopy.EilenbergMacLane.GroupStructure +open import Cubical.Homotopy.EilenbergMacLane.Base +open import Cubical.Homotopy.EilenbergMacLane.Properties +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.EilenbergMacLane.Order2 + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Univalence + +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.HITs.Susp + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Bool hiding (_≤_) +open import Cubical.Data.Fin.Base + +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Instances.IntMod +open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.AbGroup.Instances.IntMod + +open RingStr renaming (_+_ to _+r_) +open PlusBis + +open Iso + +-- move to Cohomology GroupStr or somethign +EM-ℤ/2ˣ∙ : (n : ℕ) → EM ℤ/2 ˣ n +EM-ℤ/2ˣ∙ = 0ˣ (EM ℤ/2) 0ₖ + +private + nonConst→∙ : (b : EM ℤ/2 1) → Type _ + nonConst→∙ b = + Σ[ F ∈ Susp∙ (embase ≡ b) →∙ EM∙ (CommRing→AbGroup ℤ/2CommRing) 1 ] + (¬ F ≡ const∙ _ _) + + Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool : + Iso (Susp∙ (Ω (EM∙ ℤ/2 1) .fst) + →∙ EM∙ ℤ/2 1) + ((Bool , true) →∙ (Bool , true)) + Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool = compIso + (invIso (ΩSuspAdjointIso {A = Ω (EM∙ ℤ/2 1)})) + (compIso (post∘∙equiv (lem , refl)) (pre∘∙equiv (lem , refl))) + where + lem = isoToEquiv + (compIso (invIso (Iso-EM-ΩEM+1 {G = ℤ/2} 0)) + (invIso (Bool≅ℤGroup/2 .fst))) + + Σ¬Iso : ∀ {ℓ} {B A : Type ℓ} + → (e : A ≃ B) + → {x : A} + → Iso (Σ[ y ∈ A ] ¬ y ≡ x) + (Σ[ y ∈ B ] ¬ y ≡ fst e x) + Σ¬Iso {B = B} = + EquivJ (λ A e → {x : A} + → Iso (Σ[ y ∈ A ] ¬ y ≡ x) + (Σ[ y ∈ B ] ¬ y ≡ fst e x)) idIso + + nonConst→∙Iso : Iso (nonConst→∙ embase) (Σ[ F ∈ Bool ] ¬ F ≡ true) + nonConst→∙Iso = + Σ¬Iso (isoToEquiv + (compIso Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool + Iso-Bool→∙Bool-Bool)) + + isProp-nonConst→∙ : (b : EM ℤ/2 1) → isProp (nonConst→∙ b) + isProp-nonConst→∙ = EM→Prop _ 0 (λ _ → isPropIsProp) + (isOfHLevelRetractFromIso 1 nonConst→∙Iso + (isContr→isProp ((false , true≢false ∘ sym) + , λ { (false , p) → Σ≡Prop (λ _ → isProp¬ _) refl + ; (true , p) → ⊥.rec (p refl)}))) + +eRP∞∙ : Susp∙ (Ω (EM∙ ℤ/2 1) .fst) →∙ EM∙ ℤ/2 1 +fst eRP∞∙ north = embase +fst eRP∞∙ south = embase +fst eRP∞∙ (merid a i) = a i +snd eRP∞∙ = refl + +eRP∞-nonConst : nonConst→∙ embase +fst eRP∞-nonConst = eRP∞∙ +snd eRP∞-nonConst p = true≢false true≡false + where + true≡false : true ≡ false + true≡false i = + Iso.fun (compIso Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool Iso-Bool→∙Bool-Bool) + (p (~ i)) + +eRP∞ : (b : EM ℤ/2 1) → nonConst→∙ b +eRP∞ = EM→Prop ℤ/2 0 isProp-nonConst→∙ eRP∞-nonConst + +module ThomRP∞ = Thom (EM∙ ℤ/2 1) (0ₖ 1 ≡_) refl + (isConnectedEM 1) + ℤ/2CommRing + +open ThomRP∞ +isContrE : isContr E +isContrE = isContrSingl _ + +module conRP∞ = + con 0 (((compEquiv (isoToEquiv (invIso (Iso-EM-ΩEM+1 0))) + (isoToEquiv (invIso (fst Bool≅ℤGroup/2))))) , refl) + (λ b → eRP∞ b .fst) + (Iso.inv (congIso (compIso Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool + Iso-Bool→∙Bool-Bool)) λ i → false) + +open conRP∞ +ϕ-raw-contrRP∞ : (n : ℕ) + → ((EM ℤ/2 1 → EM ℤ/2 n)) + ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (n +' 1)) +ϕ-raw-contrRP∞ n = ϕ-raw-contr n isContrE + +⌣RP∞ : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + → EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (n +' 1) +fst (⌣RP∞ n f) x = (f x) ⌣[ ℤ/2Ring R]ₖ x +snd (⌣RP∞ n f) = ⌣ₖ-0ₖ _ _ (f (0ₖ 1)) + +⌣RP∞' : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + → EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (1 +' n) +fst (⌣RP∞' n f) x = x ⌣[ ℤ/2Ring R]ₖ (f x) +snd (⌣RP∞' n f) = 0ₖ-⌣ₖ _ _ (f (0ₖ 1)) + +⌣RP∞≡⌣RP∞' : (n : ℕ) → + PathP (λ i → + (EM ℤ/2 1 → EM ℤ/2 n) → + EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (+'-comm n 1 i)) + (⌣RP∞ n) + (⌣RP∞' n) +⌣RP∞≡⌣RP∞' n = funExt λ f → + →∙HomogeneousPathP refl (cong (EM∙ ℤ/2) (+'-comm n 1)) + (isHomogeneousEM _) + (funExt λ x → lem f x) + where + lem : (f : EM ℤ/2 1 → EM ℤ/2 n) + (x : EM ℤ/2 1) + → PathP (λ i → EM ℤ/2 (+'-comm n 1 i)) + (f x ⌣[ ℤ/2Ring R]ₖ x) (x ⌣[ ℤ/2Ring R]ₖ f x) + lem f x = toPathP (sym (⌣ₖ-commℤ/2 1 n x (f x))) + +⌣RP∞IsEq : (n : ℕ) → isEquiv (⌣RP∞ n) +⌣RP∞IsEq n = + subst isEquiv + (funExt (λ f → →∙Homogeneous≡ (isHomogeneousEM _) + (λ i x → f x ⌣[ ℤ/2Ring R]ₖ (eRP∞-lem x i)))) + (ϕ-raw-contrRP∞ n .snd) + where + help : (g : _) → (λ i → eRP∞ (emloop g i) .fst .fst south) + ≡ emloop g + help g j i = + hcomp (λ k → λ {(i = i0) → embase + ; (i = i1) → emloop g k + ; (j = i0) → eRP∞ (emloop g i) .fst .fst + (merid (λ w → emloop g (i ∧ w)) k) + ; (j = i1) → emloop g (i ∧ k)}) + (eRP∞ (emloop g i) .fst .snd j) + + eRP∞-lem : (x : _) → eRP∞ x .fst .fst south ≡ x + eRP∞-lem = EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) + λ { embase-raw → refl ; (emloop-raw g i) j → help g j i } + +abstract + ⌣RP∞'IsEq : (n : ℕ) → isEquiv (⌣RP∞' n) + ⌣RP∞'IsEq n = transport (λ i → isEquiv (⌣RP∞≡⌣RP∞' n i)) (⌣RP∞IsEq n) + +⌣RP∞Equiv : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (n +' 1)) +fst (⌣RP∞Equiv n) = ⌣RP∞ n +snd (⌣RP∞Equiv n) = ⌣RP∞IsEq n + +⌣RP∞'Equiv : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (1 +' n)) +fst (⌣RP∞'Equiv n) = ⌣RP∞' n +snd (⌣RP∞'Equiv n) = ⌣RP∞'IsEq n + ++'-suc₁ : (n : ℕ) → 1 +' n ≡ suc n ++'-suc₁ zero = refl ++'-suc₁ (suc n) = refl + +⌣RP∞''Equiv : (n : ℕ) + → (EM ℤ/2 1 → EM ℤ/2 n) + ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (suc n)) +⌣RP∞''Equiv n = + compEquiv (⌣RP∞'Equiv n) + (isoToEquiv (pre∘∙equiv + ((pathToEquiv (cong (EM ℤ/2) (+'-suc₁ n))) + , (subst-EM-0ₖ (+'-suc₁ n))))) + +RP→Charac₀ : Iso (EM ℤ/2 1 → ℤ/2 .fst) + (ℤ/2 .fst) +Iso.fun RP→Charac₀ f = f embase +Iso.inv RP→Charac₀ a = λ _ → a +Iso.rightInv RP→Charac₀ a = refl +Iso.leftInv RP→Charac₀ f = + funExt (EM→Prop _ 0 (λ _ → is-set (snd ℤ/2Ring) _ _) refl) + +RP→EM-ℤ/2-CharacIso : (n : ℕ) + → Iso (EM ℤ/2 1 → EM ℤ/2 n) + ((EM ℤ/2) ˣ n) +RP→EM-ℤ/2-CharacIso zero = RP→Charac₀ -- RP→Charac₀ +RP→EM-ℤ/2-CharacIso (suc n) = + compIso (EM→-charac {A = EM∙ ℤ/2 1} (suc n)) + (prodIso + (compIso help (RP→EM-ℤ/2-CharacIso n)) + idIso) + where + help : Iso (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (suc n)) + (EM ℤ/2 1 → EM ℤ/2 n) + help = equivToIso (invEquiv (⌣RP∞''Equiv n)) + +∥EM-ℤ/2ˣ∥₀-Iso : (n : ℕ) → Iso ∥ EM ℤ/2 ˣ n ∥₂ (fst ℤ/2) +∥EM-ℤ/2ˣ∥₀-Iso zero = setTruncIdempotentIso (is-set (snd ℤ/2Ring)) +∥EM-ℤ/2ˣ∥₀-Iso (suc n) = + compIso + (compIso + setTruncOfProdIso + (compIso (Σ-cong-iso-snd λ _ + → equivToIso + (isContr→≃Unit (∣ 0ₖ (suc n) ∣₂ + , ST.elim (λ _ → isSetPathImplicit) + (EM→Prop _ _ (λ _ → squash₂ _ _) refl)))) + rUnit×Iso)) + (∥EM-ℤ/2ˣ∥₀-Iso n) + +⌣RP∞''Equiv∙ : (n : ℕ) + → ⌣RP∞''Equiv n .fst (λ _ → 0ₖ n) ≡ const∙ _ _ +⌣RP∞''Equiv∙ n = →∙Homogeneous≡ (isHomogeneousEM _) + (funExt λ x → cong (subst (EM ℤ/2) (+'-suc₁ n)) (⌣ₖ-0ₖ 1 n x) + ∙ subst-EM∙ (+'-suc₁ n) .snd) + +cohomRP∞Iso : (n : ℕ) → Iso (coHom n ℤ/2 (EM ℤ/2 1)) (ℤ/2 .fst) +cohomRP∞Iso n = compIso (setTruncIso (RP→EM-ℤ/2-CharacIso n)) (∥EM-ℤ/2ˣ∥₀-Iso n) + +RP→EM-ℤ/2-CharacIso∙ : (n : ℕ) + → Iso.fun (RP→EM-ℤ/2-CharacIso n) (λ _ → 0ₖ n) ≡ EM-ℤ/2ˣ∙ n +RP→EM-ℤ/2-CharacIso∙ zero = refl +RP→EM-ℤ/2-CharacIso∙ (suc n) = + ΣPathP (((cong (fun (RP→EM-ℤ/2-CharacIso n)) + (funExt λ x + → cong (λ f → invEq (⌣RP∞''Equiv n) f x) + (→∙Homogeneous≡ (isHomogeneousEM _) + (funExt (λ x → rCancelₖ (suc n) (0ₖ (suc n))))) + ∙ funExt⁻ (sym (cong (invEq (⌣RP∞''Equiv n)) (⌣RP∞''Equiv∙ n))) x + ∙ λ i → retEq (⌣RP∞''Equiv n) (λ _ → 0ₖ n) i x)) + ∙ RP→EM-ℤ/2-CharacIso∙ n) + , refl) + +HⁿRP∞≅ℤ/2 : (n : ℕ) → AbGroupIso (coHomGr n ℤ/2 (EM ℤ/2 1)) ℤ/2 +fst (HⁿRP∞≅ℤ/2 n) = cohomRP∞Iso n +snd (HⁿRP∞≅ℤ/2 n) = pres0→GroupIsoℤ/2 (isoToEquiv (cohomRP∞Iso n)) (cohomRP∞Iso∙ n) + where + ∥EM-ℤ/2ˣ∥₀-Iso∙ : (n : ℕ) → Iso.fun (∥EM-ℤ/2ˣ∥₀-Iso n) ∣ EM-ℤ/2ˣ∙ n ∣₂ ≡ fzero + ∥EM-ℤ/2ˣ∥₀-Iso∙ zero = refl + ∥EM-ℤ/2ˣ∥₀-Iso∙ (suc n) = ∥EM-ℤ/2ˣ∥₀-Iso∙ n + + cohomRP∞Iso∙ : (n : ℕ) → cohomRP∞Iso n .fun (0ₕ n) ≡ fzero + cohomRP∞Iso∙ n = cong (Iso.fun (∥EM-ℤ/2ˣ∥₀-Iso n) ∘ ∣_∣₂) (RP→EM-ℤ/2-CharacIso∙ n) + ∙ ∥EM-ℤ/2ˣ∥₀-Iso∙ n diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda index 2672655c75..12f3287d2c 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda @@ -1,1985 +1,414 @@ {-# OPTIONS --safe --lossy-unification #-} - -{- -This file contains -1. The Thom isomorphism (various related forms of it) -2. The Gysin sequence --} - module Cubical.Cohomology.EilenbergMacLane.Rings.RPinf where +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Path +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2wedgeS1 open import Cubical.Cohomology.EilenbergMacLane.Base -open import Cubical.Cohomology.EilenbergMacLane.Groups.Sn +open import Cubical.Cohomology.EilenbergMacLane.Groups.RPinf open import Cubical.Cohomology.EilenbergMacLane.CupProduct -open import Cubical.Cohomology.EilenbergMacLane.Gysin +open import Cubical.Homotopy.EilenbergMacLane.GroupStructure +open import Cubical.Homotopy.EilenbergMacLane.Order2 +open import Cubical.Homotopy.EilenbergMacLane.Properties +open import Cubical.Homotopy.EilenbergMacLane.Base open import Cubical.Homotopy.EilenbergMacLane.CupProduct open import Cubical.Homotopy.EilenbergMacLane.CupProductTensor renaming (_⌣ₖ_ to _⌣ₖ⊗_ ; ⌣ₖ-0ₖ to ⌣ₖ-0ₖ⊗ ; 0ₖ-⌣ₖ to 0ₖ-⌣ₖ⊗) -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.EilenbergMacLane.GradedCommTensor - hiding (⌣ₖ-comm) -open import Cubical.Homotopy.EilenbergMacLane.GroupStructure -open import Cubical.Homotopy.EilenbergMacLane.Base -open import Cubical.Homotopy.EilenbergMacLane.Properties open import Cubical.Homotopy.Loopspace -open import Cubical.Homotopy.Group.Base -open import Cubical.Homotopy.EilenbergMacLane.Order2 - -open import Cubical.Functions.Morphism -open import Cubical.Functions.Embedding -open import Cubical.Functions.Surjection +open import Cubical.Homotopy.Connected +open import Cubical.Cohomology.EilenbergMacLane.RingStructure +open import Cubical.Cohomology.EilenbergMacLane.Rings.Z2-properties -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Transport -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Function -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Pointed.Homogeneous -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Foundations.Univalence +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Fin +open import Cubical.Data.Sum +open import Cubical.Data.Fin.Arithmetic +open import Cubical.Data.Sigma +open import Cubical.Data.Vec +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.FinData renaming (Fin to FinI) +open import Cubical.HITs.KleinBottle renaming (rec to KleinFun) open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.SetTruncation as ST open import Cubical.HITs.Truncation as TR -open import Cubical.HITs.Sn -open import Cubical.HITs.Pushout -open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _sq/_) +open import Cubical.HITs.EilenbergMacLane1 hiding (elimProp ; elimSet) open import Cubical.HITs.Susp -open import Cubical.HITs.S1 +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 renaming (rec to S¹Fun) +open import Cubical.HITs.Sn open import Cubical.HITs.RPn +open import Cubical.HITs.Wedge +open import Cubical.Algebra.GradedRing.DirectSumHIT -open import Cubical.Data.Empty as ⊥ -open import Cubical.Relation.Nullary - - -open import Cubical.Data.Unit -open import Cubical.Data.Nat -open import Cubical.Data.Nat.Order hiding (eq) -open import Cubical.Data.Sigma -open import Cubical.Data.Bool hiding (_≤_) - -open import Cubical.Algebra.Group.Base -open import Cubical.Algebra.AbGroup.Base -open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Ring -open import Cubical.Algebra.CommRing hiding (_ˣ) -open import Cubical.Algebra.CommRing.Instances.IntMod +open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.AbGroup hiding (_[_]) +open import Cubical.Algebra.AbGroup.Instances.Unit open import Cubical.Algebra.AbGroup.Instances.IntMod -open import Cubical.Data.Fin.Arithmetic -open import Cubical.Data.Fin.Base - - - -open RingStr renaming (_+_ to _+r_) -open PlusBis - -open Iso - --- move to Bool? -Bool→Bool→∙Bool : Bool → (Bool , true) →∙ (Bool , true) -Bool→Bool→∙Bool false = idfun∙ _ -Bool→Bool→∙Bool true = const∙ _ _ - -Iso-Bool→∙Bool-Bool : Iso ((Bool , true) →∙ (Bool , true)) Bool -Iso.fun Iso-Bool→∙Bool-Bool f = fst f false -Iso.inv Iso-Bool→∙Bool-Bool = Bool→Bool→∙Bool -Iso.rightInv Iso-Bool→∙Bool-Bool false = refl -Iso.rightInv Iso-Bool→∙Bool-Bool true = refl -Iso.leftInv Iso-Bool→∙Bool-Bool f = Σ≡Prop (λ _ → isSetBool _ _) (help _ refl) - where - help : (x : Bool) → fst f false ≡ x - → Bool→Bool→∙Bool (fst f false) .fst ≡ f .fst - help false p = funExt - λ { false → (λ j → Bool→Bool→∙Bool (p j) .fst false) ∙ sym p - ; true → (λ j → Bool→Bool→∙Bool (p j) .fst true) ∙ sym (snd f)} - help true p = (λ j → Bool→Bool→∙Bool (p j) .fst) - ∙ funExt λ { false → sym p ; true → sym (snd f)} - --- pres0→hom -pres0→GroupIso : ∀ {ℓ} {G : Group ℓ} (f : fst G ≃ ℤ/2 .fst) - → fst f (GroupStr.1g (snd G)) ≡ fzero - → IsGroupHom (snd G) (fst f) ((ℤGroup/ 2) .snd) -pres0→GroupIso {G = G} f fp = isGroupHomInv ((invEquiv f) , main) - where - one = invEq f fone - - f⁻∙ : invEq f fzero ≡ GroupStr.1g (snd G) - f⁻∙ = sym (cong (invEq f) fp) ∙ retEq f _ - - f⁻1 : GroupStr._·_ (snd G) (invEq f fone) (invEq f fone) - ≡ GroupStr.1g (snd G) - f⁻1 = sym (retEq f (GroupStr._·_ (snd G) (invEq f fone) (invEq f fone))) - ∙∙ cong (invEq f) (mainlem _ refl ∙ sym fp) - ∙∙ retEq f (GroupStr.1g (snd G)) - where - l : ¬ (fst f (GroupStr._·_ (snd G) (invEq f fone) (invEq f fone)) - ≡ fone) - l p = snotz (cong fst q) - where - q : fone ≡ fzero - q = (sym (secEq f fone) - ∙ cong (fst f) - ((sym (GroupStr.·IdL (snd G) one) - ∙ cong (λ x → GroupStr._·_ (snd G) x one) (sym (GroupStr.·InvL (snd G) one))) - ∙ sym (GroupStr.·Assoc (snd G) (GroupStr.inv (snd G) one) one one))) - ∙ cong (fst f) (cong (GroupStr._·_ (snd G) (GroupStr.inv (snd G) (invEq f fone))) - ((sym (retEq f _) ∙ cong (invEq f) p))) - ∙ cong (fst f) (GroupStr.·InvL (snd G) _) - ∙ fp - - - mainlem : (x : _) - → fst f (GroupStr._·_ (snd G) (invEq f fone) (invEq f fone)) ≡ x - → f .fst ((snd G GroupStr.· invEq f fone) (invEq f fone)) ≡ fzero - mainlem = ℤ/2-elim - (λ p → p) - λ p → ⊥.rec (l p) - - - main : IsGroupHom ((ℤGroup/ 2) .snd) (invEq f) (snd G) - main = - makeIsGroupHom - (ℤ/2-elim - (ℤ/2-elim (f⁻∙ ∙ sym (GroupStr.·IdR (snd G) (GroupStr.1g (snd G))) - ∙ cong (λ x → snd G .GroupStr._·_ x x) (sym f⁻∙)) - (sym (GroupStr.·IdL (snd G) one) - ∙ cong (λ x → snd G .GroupStr._·_ x one) (sym f⁻∙))) - (ℤ/2-elim (sym (GroupStr.·IdR (snd G) one) - ∙ cong (snd G .GroupStr._·_ (invEq f fone)) (sym f⁻∙)) - (f⁻∙ ∙ sym f⁻1))) - ------------------------------------------------------------ - --- move to Cohomology GroupStr or somethign -EM→-charac : ∀ {ℓ ℓ'} {A : Pointed ℓ} {G : AbGroup ℓ'} (n : ℕ) - → Iso (fst A → EM G n) ((A →∙ EM∙ G n) × EM G n) -Iso.fun (EM→-charac {A = A} n) f = - ((λ x → f x -ₖ f (pt A)) , rCancelₖ n (f (pt A))) , f (pt A) -Iso.inv (EM→-charac n) (f , a) x = fst f x +ₖ a -Iso.rightInv (EM→-charac {A = A} n) ((f , p) , a) = - ΣPathP (→∙Homogeneous≡ (isHomogeneousEM _) - (funExt (λ x → (λ i → (f x +ₖ a) -ₖ (cong (_+ₖ a) p ∙ lUnitₖ n a) i) - ∙ sym (assocₖ n (f x) a (-ₖ a)) - ∙ cong (f x +ₖ_) (rCancelₖ n a) - ∙ rUnitₖ n (f x))) - , cong (_+ₖ a) p ∙ lUnitₖ n a) -Iso.leftInv (EM→-charac {A = A} n) f = - funExt λ x → sym (assocₖ n (f x) (-ₖ f (pt A)) (f (pt A))) - ∙∙ cong (f x +ₖ_) (lCancelₖ n (f (pt A))) - ∙∙ rUnitₖ n (f x) - -0ˣ : ∀ {ℓ} (A : ℕ → Type ℓ) (0A : (n : ℕ) → A n) → (n : ℕ) → A ˣ n -0ˣ A 0A zero = 0A zero -0ˣ A 0A (suc n) = (0ˣ A 0A n) , (0A (suc n)) - -EM-ℤ/2ˣ∙ : (n : ℕ) → EM ℤ/2 ˣ n -EM-ℤ/2ˣ∙ = 0ˣ (EM ℤ/2) 0ₖ - -private - nonConst→∙ : (b : EM ℤ/2 1) → Type _ - nonConst→∙ b = - Σ[ F ∈ Susp∙ (embase ≡ b) →∙ EM∙ (CommRing→AbGroup ℤ/2CommRing) 1 ] - (¬ F ≡ const∙ _ _) +open import Cubical.Algebra.DirectSum.DirectSumHIT.Base +open import Cubical.Algebra.Ring +open import Cubical.Algebra.AbGroup.TensorProduct +open import Cubical.Algebra.Monoid - Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool : - Iso (Susp∙ (Ω (EM∙ ℤ/2 1) .fst) - →∙ EM∙ ℤ/2 1) - ((Bool , true) →∙ (Bool , true)) - Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool = compIso - (invIso (ΩSuspAdjointIso {A = Ω (EM∙ ℤ/2 1)})) - (compIso (post∘∙equiv (lem , refl)) (pre∘∙equiv (lem , refl))) - where - lem = isoToEquiv - (compIso (invIso (Iso-EM-ΩEM+1 {G = ℤ/2} 0)) - (invIso (Bool≅ℤGroup/2 .fst))) +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.Quotient +open import Cubical.Algebra.CommRing.Instances.IntMod +open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly +open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly-notationZ2 - Σ¬Iso : ∀ {ℓ} {B A : Type ℓ} - → (e : A ≃ B) - → {x : A} - → Iso (Σ[ y ∈ A ] ¬ y ≡ x) - (Σ[ y ∈ B ] ¬ y ≡ fst e x) - Σ¬Iso {B = B} = - EquivJ (λ A e → {x : A} - → Iso (Σ[ y ∈ A ] ¬ y ≡ x) - (Σ[ y ∈ B ] ¬ y ≡ fst e x)) idIso +open import Cubical.Algebra.Monoid.Instances.Nat - nonConst→∙Iso : Iso (nonConst→∙ embase) (Σ[ F ∈ Bool ] ¬ F ≡ true) - nonConst→∙Iso = - Σ¬Iso (isoToEquiv - (compIso Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool - Iso-Bool→∙Bool-Bool)) +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties - isProp-nonConst→∙ : (b : EM ℤ/2 1) → isProp (nonConst→∙ b) - isProp-nonConst→∙ = EM→Prop _ 0 (λ _ → isPropIsProp) - (isOfHLevelRetractFromIso 1 nonConst→∙Iso - (isContr→isProp ((false , true≢false ∘ sym) - , λ { (false , p) → Σ≡Prop (λ _ → isProp¬ _) refl - ; (true , p) → ⊥.rec (p refl)}))) +open import Cubical.Data.Fin.Arithmetic +open import Cubical.Cohomology.EilenbergMacLane.CupProduct -eRP∞∙ : Susp∙ (Ω (EM∙ ℤ/2 1) .fst) →∙ EM∙ ℤ/2 1 -fst eRP∞∙ north = embase -fst eRP∞∙ south = embase -fst eRP∞∙ (merid a i) = a i -snd eRP∞∙ = refl +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation as TR -eRP∞-nonConst : nonConst→∙ embase -fst eRP∞-nonConst = eRP∞∙ -snd eRP∞-nonConst p = true≢false true≡false +open Iso +open PlusBis +open RingTheory + +open AbGroupStr + +eRP∞^ : (n : ℕ) → coHom n ℤ/2 (EM ℤ/2 1) +eRP∞^ zero = 1ₕ {G'' = ℤ/2Ring} +eRP∞^ (suc n) = + subst (λ k → coHom k ℤ/2 (EM ℤ/2 1)) (+'-suc₁ n) + (_⌣_ {G'' = ℤ/2Ring} {n = 1} {n} ∣ idfun _ ∣₂ (eRP∞^ n)) + +isConnected₀EM : ∀ {ℓ} (G : AbGroup ℓ) (n : ℕ) → isContr ∥ EM G (suc n) ∥₂ +isConnected₀EM G n = + isOfHLevelRetractFromIso 0 setTruncTrunc2Iso + (isConnectedSubtr 2 n + (subst (λ x → isConnected x (EM G (suc n))) + (+-comm 2 n) + (isConnectedEM (suc n)))) + +EMOne : (n : ℕ) → EM ℤ/2 ˣ n +EMOne zero = fone +EMOne (suc n) = EMOne n , 0ₖ (suc n) + +eRP∞^-isGen : (n : ℕ) → Iso.inv (HⁿRP∞≅ℤ/2 n .fst) fone ≡ eRP∞^ n +eRP∞^-isGen zero = refl +eRP∞^-isGen (suc n) = + (cong₂ (λ v w → ST.rec isSetSetTrunc + (λ x → ∣ inv (RP→EM-ℤ/2-CharacIso (suc n)) x ∣₂) + (ST.rec2 isSetSetTrunc (λ a b → ∣ a , b ∣₂) + v w)) + (inv-∥EM-ℤ/2ˣ∥₀-Iso-fone n) + (isContr→isProp (isConnected₀EM ℤ/2 n) _ ∣ (0ₖ (suc n)) ∣₂) + ∙ cong ∣_∣₂ (funExt (λ x → rUnitₖ (suc n) _ + ∙ cong (subst (EM ℤ/2) (+'-suc₁ n)) + (cong₂ (_⌣ₖ_ {G'' = ℤ/2Ring} {n = 1} {n}) + (sym (transportRefl x)) + (cong (inv (RP→EM-ℤ/2-CharacIso n) (EMOne n)) + (sym (transportRefl x)))))) + ∙ sym (substCommSlice (λ k → EM ℤ/2 1 → EM ℤ/2 k) + (λ k → coHom k ℤ/2 (EM ℤ/2 1)) + (λ _ → ∣_∣₂) (+'-suc₁ n) + (λ x → _⌣ₖ_ {G'' = ℤ/2Ring} {n = 1} {n} x + (inv (RP→EM-ℤ/2-CharacIso n) (EMOne n) x)))) + ∙ cong (subst (λ k → coHom k ℤ/2 (EM ℤ/2 1)) (+'-suc₁ n)) + (cong (_⌣_ {G'' = ℤ/2Ring} {n = 1} {n} ∣ idfun _ ∣₂) + (cong (ST.rec isSetSetTrunc (λ x → ∣ inv (RP→EM-ℤ/2-CharacIso n) x ∣₂)) + (sym (inv-∥EM-ℤ/2ˣ∥₀-Iso-fone n)) + ∙ eRP∞^-isGen n)) where - true≡false : true ≡ false - true≡false i = - Iso.fun (compIso Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool Iso-Bool→∙Bool-Bool) - (p (~ i)) - -eRP∞ : (b : EM ℤ/2 1) → nonConst→∙ b -eRP∞ = EM→Prop ℤ/2 0 isProp-nonConst→∙ eRP∞-nonConst - -module ThomRP∞ = Thom (EM∙ ℤ/2 1) (0ₖ 1 ≡_) refl - (isConnectedEM 1) - ℤ/2CommRing - -open ThomRP∞ -isContrE : isContr E -isContrE = isContrSingl _ - -module conRP∞ = - con 0 (((compEquiv (isoToEquiv (invIso (Iso-EM-ΩEM+1 0))) - (isoToEquiv (invIso (fst Bool≅ℤGroup/2))))) , refl) - (λ b → eRP∞ b .fst) - (Iso.inv (congIso (compIso Iso-SuspΩEM₁→∙EM₁-Bool→∙Bool - Iso-Bool→∙Bool-Bool)) λ i → false) - -open conRP∞ -ϕ-raw-contrRP∞ : (n : ℕ) - → ((EM ℤ/2 1 → EM ℤ/2 n)) - ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (n +' 1)) -ϕ-raw-contrRP∞ n = ϕ-raw-contr n isContrE - -⌣RP∞ : (n : ℕ) - → (EM ℤ/2 1 → EM ℤ/2 n) - → EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (n +' 1) -fst (⌣RP∞ n f) x = (f x) ⌣[ ℤ/2Ring R]ₖ x -snd (⌣RP∞ n f) = ⌣ₖ-0ₖ _ _ (f (0ₖ 1)) - -⌣RP∞' : (n : ℕ) - → (EM ℤ/2 1 → EM ℤ/2 n) - → EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (1 +' n) -fst (⌣RP∞' n f) x = x ⌣[ ℤ/2Ring R]ₖ (f x) -snd (⌣RP∞' n f) = 0ₖ-⌣ₖ _ _ (f (0ₖ 1)) - -⌣RP∞≡⌣RP∞' : (n : ℕ) → - PathP (λ i → - (EM ℤ/2 1 → EM ℤ/2 n) → - EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (+'-comm n 1 i)) - (⌣RP∞ n) - (⌣RP∞' n) -⌣RP∞≡⌣RP∞' n = funExt λ f → - →∙HomogeneousPathP refl (cong (EM∙ ℤ/2) (+'-comm n 1)) - (isHomogeneousEM _) - (funExt λ x → lem f x) + inv-∥EM-ℤ/2ˣ∥₀-Iso-fone : (n : ℕ) + → inv (∥EM-ℤ/2ˣ∥₀-Iso n) fone ≡ ∣ EMOne n ∣₂ + inv-∥EM-ℤ/2ˣ∥₀-Iso-fone zero = refl + inv-∥EM-ℤ/2ˣ∥₀-Iso-fone (suc n) = + cong₂ (λ v w → prodRec isSetSetTrunc (λ a b → ∣ a , b ∣₂) + (v , w)) (inv-∥EM-ℤ/2ˣ∥₀-Iso-fone n) + (isContr→isProp (isConnected₀EM ℤ/2 n) _ ∣ (0ₖ (suc n)) ∣₂) + +eRP∞^-comp : (n m : ℕ) → _⌣_ {G'' = ℤ/2Ring} (eRP∞^ n) (eRP∞^ m) ≡ eRP∞^ (n +' m) +eRP∞^-comp zero m = 1ₕ-⌣ m (eRP∞^ m) +eRP∞^-comp (suc n) m = + transport⌣ _ (+'-suc₁ n) _ _ + ∙ sym (compSubstℕ (sym (+'-assoc 1 n m) ∙ +'-suc₁ (n +' m)) (+'-suc n m) + (cong₂ _+'_ (+'-suc₁ n) refl)) + ∙ cong (subst (λ k → coHom k ℤ/2 (EM ℤ/2 1)) (+'-suc n m)) + (sym (compSubstℕ (sym (+'-assoc 1 n m)) (+'-suc₁ (n +' m)) + (sym (+'-assoc 1 n m) ∙ +'-suc₁ (n +' m))) + ∙ cong (subst (λ k → coHom k ℤ/2 (EM ℤ/2 1)) (+'-suc₁ (n +' m))) + ((cong (subst (λ k → coHom k ℤ/2 (EM ℤ/2 1)) (sym (+'-assoc 1 n m))) + (assoc⌣ 1 n m ∣ idfun _ ∣₂ (eRP∞^ n) (eRP∞^ m)) + ∙ subst⁻Subst (λ k → coHom k ℤ/2 (EM ℤ/2 1)) (+'-assoc 1 n m) _) + ∙ cong₂ (_⌣_ {G'' = ℤ/2Ring} {n = 1} {n +' m}) refl (eRP∞^-comp n m))) + ∙ sym (help _ (sym (+'-suc n m))) where - lem : (f : EM ℤ/2 1 → EM ℤ/2 n) - (x : EM ℤ/2 1) - → PathP (λ i → EM ℤ/2 (+'-comm n 1 i)) - (f x ⌣[ ℤ/2Ring R]ₖ x) (x ⌣[ ℤ/2Ring R]ₖ f x) - lem f x = toPathP (sym (⌣ₖ-commℤ/2 1 n x (f x))) - -⌣RP∞IsEq : (n : ℕ) → isEquiv (⌣RP∞ n) -⌣RP∞IsEq n = - subst isEquiv - (funExt (λ f → →∙Homogeneous≡ (isHomogeneousEM _) - (λ i x → f x ⌣[ ℤ/2Ring R]ₖ (eRP∞-lem x i)))) - (ϕ-raw-contrRP∞ n .snd) + term = _⌣_ {G'' = ℤ/2Ring} {n = 1 +' n} {m} + (_⌣_ {G'' = ℤ/2Ring} {n = 1} {n} ∣ idfun _ ∣₂ (eRP∞^ n)) (eRP∞^ m) + + help : {w : ℕ} (k : ℕ) (p : w ≡ k) + → eRP∞^ w ≡ subst (λ k → coHom k ℤ/2 (EM ℤ/2 1)) (sym p) (eRP∞^ k) + help = J> sym (transportRefl _) + + transport⌣ : {w v : ℕ} (k : ℕ) (p : w ≡ k) + (x : coHom w ℤ/2 (EM ℤ/2 1)) (y : coHom v ℤ/2 (EM ℤ/2 1)) + → _⌣_ {G'' = ℤ/2Ring} (subst (λ k → coHom k ℤ/2 (EM ℤ/2 (suc zero))) p x) y + ≡ subst (λ k → coHom k ℤ/2 (EM ℤ/2 (suc zero))) + (cong₂ _+'_ p refl) (_⌣_ {G'' = ℤ/2Ring} x y) + transport⌣ = J> λ x y + → cong₂ (_⌣_ {G'' = ℤ/2Ring}) (transportRefl x) refl + ∙ sym (transportRefl _) + +HⁿRP∞≅ℤ/2⌣ : (n m : ℕ) (a b : ℤ/2 .fst) + → Iso.inv (HⁿRP∞≅ℤ/2 (n +' m) .fst) (a ·ₘ b) + ≡ _⌣_ {G'' = ℤ/2Ring} (Iso.inv (HⁿRP∞≅ℤ/2 n .fst) a) + (Iso.inv (HⁿRP∞≅ℤ/2 m .fst) b) +HⁿRP∞≅ℤ/2⌣ n m = + ℤ/2-elim (λ b + → IsGroupHom.pres1 (invGroupIso (HⁿRP∞≅ℤ/2 (n +' m)) .snd) + ∙ sym (0ₕ-⌣ {G'' = ℤ/2Ring} n m (inv (cohomRP∞Iso m) b)) + ∙ cong₂ (_⌣_ {G'' = ℤ/2Ring}) + (sym (IsGroupHom.pres1 (invGroupIso (HⁿRP∞≅ℤ/2 n) .snd))) refl) + (ℤ/2-elim (IsGroupHom.pres1 (invGroupIso (HⁿRP∞≅ℤ/2 (n +' m)) .snd) + ∙ sym (⌣-0ₕ {G'' = ℤ/2Ring} n m (inv (cohomRP∞Iso n) fone)) + ∙ cong₂ (_⌣_ {G'' = ℤ/2Ring}) refl + (sym (IsGroupHom.pres1 (invGroupIso (HⁿRP∞≅ℤ/2 m) .snd)))) + (main n m)) where - help : (g : _) → (λ i → eRP∞ (emloop g i) .fst .fst south) - ≡ emloop g - help g j i = - hcomp (λ k → λ {(i = i0) → embase - ; (i = i1) → emloop g k - ; (j = i0) → eRP∞ (emloop g i) .fst .fst - (merid (λ w → emloop g (i ∧ w)) k) - ; (j = i1) → emloop g (i ∧ k)}) - (eRP∞ (emloop g i) .fst .snd j) - - eRP∞-lem : (x : _) → eRP∞ x .fst .fst south ≡ x - eRP∞-lem = EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) - λ { embase-raw → refl ; (emloop-raw g i) j → help g j i } - -abstract - ⌣RP∞'IsEq : (n : ℕ) → isEquiv (⌣RP∞' n) - ⌣RP∞'IsEq n = transport (λ i → isEquiv (⌣RP∞≡⌣RP∞' n i)) (⌣RP∞IsEq n) - -⌣RP∞Equiv : (n : ℕ) - → (EM ℤ/2 1 → EM ℤ/2 n) - ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (n +' 1)) -fst (⌣RP∞Equiv n) = ⌣RP∞ n -snd (⌣RP∞Equiv n) = ⌣RP∞IsEq n - -⌣RP∞'Equiv : (n : ℕ) - → (EM ℤ/2 1 → EM ℤ/2 n) - ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (1 +' n)) -fst (⌣RP∞'Equiv n) = ⌣RP∞' n -snd (⌣RP∞'Equiv n) = ⌣RP∞'IsEq n - - -+'-suc₁ : (n : ℕ) → 1 +' n ≡ suc n -+'-suc₁ zero = refl -+'-suc₁ (suc n) = refl - -⌣RP∞''Equiv : (n : ℕ) - → (EM ℤ/2 1 → EM ℤ/2 n) - ≃ (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (suc n)) -⌣RP∞''Equiv n = - compEquiv (⌣RP∞'Equiv n) - (isoToEquiv (pre∘∙equiv - ((pathToEquiv (cong (EM ℤ/2) (+'-suc₁ n))) - , (subst-EM-0ₖ (+'-suc₁ n))))) - -RP→Charac₀ : Iso (EM ℤ/2 1 → ℤ/2 .fst) - (ℤ/2 .fst) -Iso.fun RP→Charac₀ f = f embase -Iso.inv RP→Charac₀ a = λ _ → a -Iso.rightInv RP→Charac₀ a = refl -Iso.leftInv RP→Charac₀ f = - funExt (EM→Prop _ 0 (λ _ → is-set (snd ℤ/2Ring) _ _) refl) - -RP→EM-ℤ/2-CharacIso : (n : ℕ) - → Iso (EM ℤ/2 1 → EM ℤ/2 n) - ((EM ℤ/2) ˣ n) -RP→EM-ℤ/2-CharacIso zero = RP→Charac₀ -- RP→Charac₀ -RP→EM-ℤ/2-CharacIso (suc n) = - compIso (EM→-charac {A = EM∙ ℤ/2 1} (suc n)) - (prodIso - (compIso help (RP→EM-ℤ/2-CharacIso n)) - idIso) - where - help : Iso (EM∙ ℤ/2 1 →∙ EM∙ ℤ/2 (suc n)) - (EM ℤ/2 1 → EM ℤ/2 n) - help = equivToIso (invEquiv (⌣RP∞''Equiv n)) - -∥EM-ℤ/2ˣ∥₀-Iso : (n : ℕ) → Iso ∥ EM ℤ/2 ˣ n ∥₂ (fst ℤ/2) -∥EM-ℤ/2ˣ∥₀-Iso zero = setTruncIdempotentIso (is-set (snd ℤ/2Ring)) -∥EM-ℤ/2ˣ∥₀-Iso (suc n) = - compIso - (compIso - setTruncOfProdIso - (compIso (Σ-cong-iso-snd λ _ - → equivToIso - (isContr→≃Unit (∣ 0ₖ (suc n) ∣₂ - , ST.elim (λ _ → isSetPathImplicit) - (EM→Prop _ _ (λ _ → squash₂ _ _) refl)))) - rUnit×Iso)) - (∥EM-ℤ/2ˣ∥₀-Iso n) - -⌣RP∞''Equiv∙ : (n : ℕ) - → ⌣RP∞''Equiv n .fst (λ _ → 0ₖ n) ≡ const∙ _ _ -⌣RP∞''Equiv∙ n = →∙Homogeneous≡ (isHomogeneousEM _) - (funExt λ x → cong (subst (EM ℤ/2) (+'-suc₁ n)) (⌣ₖ-0ₖ 1 n x) - ∙ subst-EM∙ (+'-suc₁ n) .snd) - -cohomRP∞Iso : (n : ℕ) → Iso (coHom n ℤ/2 (EM ℤ/2 1)) (ℤ/2 .fst) -cohomRP∞Iso n = compIso (setTruncIso (RP→EM-ℤ/2-CharacIso n)) (∥EM-ℤ/2ˣ∥₀-Iso n) - -RP→EM-ℤ/2-CharacIso∙ : (n : ℕ) - → Iso.fun (RP→EM-ℤ/2-CharacIso n) (λ _ → 0ₖ n) ≡ EM-ℤ/2ˣ∙ n -RP→EM-ℤ/2-CharacIso∙ zero = refl -RP→EM-ℤ/2-CharacIso∙ (suc n) = - ΣPathP (((cong (fun (RP→EM-ℤ/2-CharacIso n)) - (funExt λ x - → cong (λ f → invEq (⌣RP∞''Equiv n) f x) - (→∙Homogeneous≡ (isHomogeneousEM _) - (funExt (λ x → rCancelₖ (suc n) (0ₖ (suc n))))) - ∙ funExt⁻ (sym (cong (invEq (⌣RP∞''Equiv n)) (⌣RP∞''Equiv∙ n))) x - ∙ λ i → retEq (⌣RP∞''Equiv n) (λ _ → 0ₖ n) i x)) - ∙ RP→EM-ℤ/2-CharacIso∙ n) - , refl) - -HⁿRP∞≅ℤ/2 : (n : ℕ) → AbGroupIso (coHomGr n ℤ/2 (EM ℤ/2 1)) ℤ/2 -fst (HⁿRP∞≅ℤ/2 n) = cohomRP∞Iso n -snd (HⁿRP∞≅ℤ/2 n) = pres0→GroupIso (isoToEquiv (cohomRP∞Iso n)) (cohomRP∞Iso∙ n) - where - ∥EM-ℤ/2ˣ∥₀-Iso∙ : (n : ℕ) → Iso.fun (∥EM-ℤ/2ˣ∥₀-Iso n) ∣ EM-ℤ/2ˣ∙ n ∣₂ ≡ fzero - ∥EM-ℤ/2ˣ∥₀-Iso∙ zero = refl - ∥EM-ℤ/2ˣ∥₀-Iso∙ (suc n) = ∥EM-ℤ/2ˣ∥₀-Iso∙ n - - cohomRP∞Iso∙ : (n : ℕ) → cohomRP∞Iso n .fun (0ₕ n) ≡ fzero - cohomRP∞Iso∙ n = cong (Iso.fun (∥EM-ℤ/2ˣ∥₀-Iso n) ∘ ∣_∣₂) (RP→EM-ℤ/2-CharacIso∙ n) - ∙ ∥EM-ℤ/2ˣ∥₀-Iso∙ n - --- makeSquare : (n : ℕ) --- → (EM ℤ/2 1 --- → EM ℤ/2 n --- → EM ℤ/2 (n +' n)) --- → (EM ℤ/2 n → (EM ℤ/2) ˣ (n +' n)) --- makeSquare n f y = Iso.fun (RP→EM-ℤ/2-CharacIso _) λ x → f x y - --- module _ {ℓ} (X Y : RP∞' ℓ) (A : fst X → fst Y → Pointed ℓ) where --- UnordSmash∙² : Pointed ℓ --- UnordSmash∙² = UnordSmash∙ X (λ x → UnordSmash∙ Y (A x)) - --- UnordSmash∙²-fun : ((x : fst X) (y : fst Y) → A x y .fst) --- → UnordSmash∙² .fst --- UnordSmash∙²-fun f = inr (λ x → inr (f x)) - --- -- open import Cubical.HITs.Join --- -- ×⁴ : ∀ {ℓ} (A : Bool → Bool → Pointed ℓ) → Type ℓ --- -- ×⁴ A = A true true .fst × A true false .fst --- -- × A false true .fst × A false false .fst - --- -- module _ {ℓ} {A : Type ℓ} {B C : A → Type ℓ} (contr : isContr (Σ A B)) where --- -- private --- -- push-c : (a : A) (p : contr .fst .fst ≡ a) --- -- → (c : C a) --- -- → Path (cofib {A = C (contr .fst .fst)} {B = Σ A C} λ x → _ , x) (inl tt) (inr (a , c)) --- -- push-c = J> push - --- -- cofib→join : (ptA : A) (ptB : B ptA) → (cofib {A = C ptA} {B = Σ A C} λ x → _ , x) → Σ[ a ∈ A ] join (B a) (C a) --- -- cofib→join ptA ptB (inl x) = ptA , (inl ptB) -- contr .fst .fst , inl (contr .fst .snd) --- -- cofib→join ptA ptB (inr (a , c)) = a , inr c --- -- cofib→join ptA ptB (push a i) = ptA , push ptB a i - --- -- push-c-coh : (a : A) (p : contr .fst .fst ≡ a) (d : B a) (pp : PathP (λ i → B (p i)) (contr .fst .snd) d) (c : C a) --- -- → PathP (λ i → cofib→join (contr .fst .fst) (contr .fst .snd) (push-c a p c i) ≡ (a , push d c i)) --- -- (ΣPathP (p , (λ j → inl (pp j)))) --- -- refl -- refl --- -- push-c-coh = --- -- J> (J> λ c → flipSquare ((λ j i → cofib→join (contr .fst .fst) (contr .fst .snd) (help c j i)) --- -- ◁ λ i j → (contr .fst .fst) , (push (contr .fst .snd) c j))) --- -- where --- -- help : (c : _) → push-c (contr .fst .fst) refl c ≡ push c --- -- help = funExt⁻ (transportRefl push) - --- -- joinC : Iso (Σ[ a ∈ A ] join (B a) (C a)) --- -- (cofib {A = C (contr .fst .fst)} {B = Σ A C} λ x → _ , x) --- -- Iso.fun joinC (a , inl x) = inl tt --- -- Iso.fun joinC (a , inr x) = inr (a , x) --- -- Iso.fun joinC (a , push b c i) = push-c a (cong fst (contr .snd (a , b))) c i --- -- Iso.inv joinC = cofib→join (contr .fst .fst) (contr .fst .snd) --- -- Iso.rightInv joinC (inl x) = refl --- -- Iso.rightInv joinC (inr x) = refl --- -- Iso.rightInv joinC (push a i) j = --- -- ((λ j → push-c (contr .fst .fst) --- -- (cong fst (isProp→isSet (isContr→isProp contr) _ _ (contr .snd (contr .fst)) refl j)) a) --- -- ∙ lem a) j i --- -- where --- -- lem : (a : _) → Path (Path (cofib (λ x → contr .fst .fst , x)) _ _) (push-c (contr .fst .fst) refl a) (push a) --- -- lem = funExt⁻ (transportRefl push) --- -- Iso.leftInv joinC (a , inl x) = ΣPathP ((cong fst (contr .snd (a , x))) , (λ j → inl (contr .snd (a , x) j .snd))) --- -- Iso.leftInv joinC (a , inr x) = refl --- -- Iso.leftInv joinC (a , push c d i) j = --- -- push-c-coh a (cong fst (contr .snd (a , c))) c (cong snd (contr .snd (a , c))) d i j - --- -- JoinStructureBool : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) --- -- (f : A true true .fst × A true false .fst --- -- × A false true .fst × A false false .fst --- -- → fst B) --- -- → Type --- -- JoinStructureBool A B f = --- -- (g : A true true .fst × A true false .fst --- -- × A false true .fst × A false false .fst) --- -- → join (join (fst g ≡ A true true .snd) --- -- (snd g .fst ≡ A true false .snd)) --- -- (join (snd (snd g) .fst ≡ A false true .snd) --- -- (snd (snd g) .snd ≡ A false false .snd)) --- -- → f g ≡ B .snd - --- -- module DEP (A B : Pointed ℓ-zero) (T : fst A → fst B → Pointed ℓ-zero) --- -- (f : (a : fst A) (b : fst B) → fst (T a b)) where --- -- JS₁ : Type --- -- JS₁ = (a : A .fst) (b : B .fst) --- -- → join (a ≡ A .snd) (b ≡ B .snd) --- -- → f a b ≡ T a b .snd - --- -- JS'l : singl (pt A) → Type --- -- JS'l (a , p) = (b : fst B) → f a b ≡ T a b .snd - --- -- JS'r : singl (pt B) → Type --- -- JS'r (b , p) = (a : fst A) → f a b ≡ T a b .snd - --- -- JS₁' : Type --- -- JS₁' = Σ[ l ∈ ((a : _) → JS'l a) ] --- -- Σ[ r ∈ ((a : _) → JS'r a) ] --- -- ((a : _) (b : _) → l a (fst b) ≡ r b (fst a)) - - --- -- JS₁'' : Type --- -- JS₁'' = Σ[ l ∈ JS'l (pt A , refl) ] --- -- Σ[ r ∈ JS'r (pt B , refl) ] --- -- l (pt B) ≡ r (pt A) - --- -- IS1 : Iso JS₁ JS₁' --- -- fst (Iso.fun IS1 F) (a , p) b = F a b (inl (sym p)) --- -- fst (snd (Iso.fun IS1 F)) (b , p) a = F a b (inr (sym p)) --- -- snd (snd (Iso.fun IS1 F)) (a , p) (b , q) = cong (F a b) (push (sym p) (sym q)) --- -- Iso.inv IS1 (l , r , lr) a b (inl x) = l (a , sym x) b --- -- Iso.inv IS1 (l , r , lr) a b (inr x) = r (b , sym x) a --- -- Iso.inv IS1 (l , r , lr) a b (push p q i) = lr (a , sym p) (b , sym q) i --- -- Iso.rightInv IS1 _ = refl --- -- Iso.leftInv IS1 F = funExt λ a → funExt λ b → funExt --- -- λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} - --- -- IS2 : Iso JS₁' JS₁'' --- -- IS2 = compIso --- -- (Σ-cong-iso --- -- {B = λ l → Σ[ r ∈ ((a : _) → JS'r a) ] --- -- ((a : _) (b : _) → l a (fst b) ≡ r b (fst a))} --- -- {B' = λ l → Σ[ r ∈ ((a : _) → JS'r a) ] --- -- ((b : _) → l (fst b) ≡ r b (pt A))} --- -- (DomainContrIso (isContrSingl (pt A))) --- -- λ x → Σ-cong-iso-snd λ r → DomainContrIso (isContrSingl (pt A))) --- -- (Σ-cong-iso-snd --- -- λ l → Σ-cong-iso {B = λ r → ((b : _) → l (fst b) ≡ r b (pt A))} --- -- {B' = λ r → (l (pt B) ≡ r (pt A))} --- -- (DomainContrIso (isContrSingl (pt B))) --- -- λ _ → DomainContrIso (isContrSingl (pt B))) - --- -- FullIso : Iso JS₁ JS₁'' --- -- FullIso = compIso IS1 IS2 - --- -- move later --- -- RP∞'→SetElim : --- -- ∀ {ℓ} {A : RP∞' ℓ → Type ℓ} (s : (X : _) → isSet (A X)) --- -- → (f : (X : RP∞' ℓ) → fst X → A X) --- -- → ((X : RP∞' ℓ) → (x : fst X) → f X x ≡ f X (RP∞'-fields.notRP∞' X x)) --- -- → ((X : RP∞' ℓ) → A X) --- -- RP∞'→SetElim {A = A} s f f-comm = --- -- uncurry λ X → uncurry λ 2tx --- -- → elim→Set (λ _ → s _) --- -- (λ x → f (X , 2tx , ∣ x ∣₁) x) --- -- λ x → RP∞'-fields.elimRP∞' (X , 2tx , ∣ x ∣₁) x --- -- (λ i → f (X , 2tx , squash₁ ∣ x ∣₁ ∣ x ∣₁ i) x) --- -- (f-comm (X , 2tx , ∣ x ∣₁) x --- -- ◁ λ i → f (X , 2tx , squash₁ ∣ x ∣₁ ∣ fst 2tx x ∣₁ i) (fst 2tx x)) - --- -- RP∞'→SetElimEq : --- -- ∀ {ℓ} {A : RP∞' ℓ → Type ℓ} (s : (X : _) → isSet (A X)) --- -- (f : (X : RP∞' ℓ) → fst X → A X) --- -- (h : (X : RP∞' ℓ) → (x : fst X) → f X x ≡ f X (RP∞'-fields.notRP∞' X x)) --- -- (X : RP∞' ℓ) (x : _) --- -- → RP∞'→SetElim s f h X ≡ f X x --- -- RP∞'→SetElimEq {A = A} s f h = uncurry λ X → uncurry --- -- λ 2x → PT.elim (λ _ → isPropΠ λ _ → s _ _ _) --- -- λ x → RP∞'-fields.elimRP∞' (X , 2x , ∣ x ∣₁) x --- -- (fromPathP λ i → (f (X , 2x , squash₁ ∣ x ∣₁ ∣ x ∣₁ i) x)) --- -- (fromPathP (h (X , 2x , ∣ x ∣₁) x --- -- ◁ λ i → f (X , 2x , squash₁ ∣ x ∣₁ ∣ x ∣₁ i) (fst 2x x))) - --- RP∞'→SetRec : ∀ {ℓ} {A : Type ℓ} (s : isSet A) (X : RP∞' ℓ) --- → (f : fst X → A) --- → ((x : _) → f x ≡ f (RP∞'-fields.notRP∞' X x)) --- → A --- RP∞'→SetRec s = uncurry λ X --- → uncurry λ 2tx --- → elim→Set (λ _ → isSetΠ2 λ _ _ → s) --- (λ x f coh → f x) --- λ x → RP∞'-fields.elimRP∞' (X , 2tx , ∣ x ∣₁) x --- (λ i f coh → f x) --- λ i f coh → coh x i - --- RP∞'→SetRecEq : --- ∀ {ℓ} {A : Type ℓ} (s : isSet A) (X : _) --- (f : fst X → A) --- (h : (x : fst X) → f x ≡ f (RP∞'-fields.notRP∞' X x)) --- (x : fst X) --- → RP∞'→SetRec s X f h ≡ f x --- RP∞'→SetRecEq {A = A} s = uncurry λ X → uncurry --- λ 2x → PT.elim (λ _ → isPropΠ3 λ _ _ _ → s _ _) --- λ x f h → RP∞'-fields.elimRP∞' (X , 2x , ∣ x ∣₁) x --- (λ i → transportRefl (transportRefl (f (transportRefl x i)) i) i) --- ((λ i → transportRefl (transportRefl (f (transportRefl x i)) i) i) ∙ h x) - - --- abstract --- notNotRP∞' : ∀ {ℓ} (X : RP∞' ℓ) (x : fst X) --- → RP∞'-fields.notRP∞' X (RP∞'-fields.notRP∞' X x) ≡ x --- notNotRP∞' = JRP∞' refl - --- RP∞'→GroupoidRec : ∀ {ℓ} {A : Type ℓ} (s : isGroupoid A) (X : RP∞' ℓ) --- → (f : fst X → A) --- → (f-coh : (x : _) → f x ≡ f (RP∞'-fields.notRP∞' X x)) --- → (p : (x : fst X) --- → PathP (λ i → f (notNotRP∞' X x (~ i)) ≡ f (RP∞'-fields.notRP∞' X x)) --- (f-coh x) --- (sym (f-coh (RP∞'-fields.notRP∞' X x)))) --- → A --- RP∞'→GroupoidRec {ℓ = ℓ} {A = A} grA = uncurry λ X --- → uncurry λ 2tx --- → elim→Gpd _ (λ _ → isGroupoidΠ3 λ _ _ _ → grA) --- (F1 X 2tx) --- (F2 X 2tx) --- λ x → F3 (X , 2tx , ∣ x ∣₁) x --- where --- F1 : (X : Type) (2tx : is2Type ℓ X) (x : X) (f : X → A) --- (f-coh : (x' : X) → f x' ≡ f (RP∞'-fields.notRP∞' (X , 2tx , ∣ x ∣₁) x')) → --- ((x' : X) → --- PathP (λ i → f (notNotRP∞' (X , 2tx , ∣ x ∣₁) x' (~ i)) --- ≡ f (RP∞'-fields.notRP∞' (X , 2tx , ∣ x ∣₁) x')) --- (f-coh x') --- (sym (f-coh (RP∞'-fields.notRP∞' (X , 2tx , ∣ x ∣₁) x')))) --- → A --- F1 X 2tx x f coh p = f x - --- F2 : (X : Type) (2tx : is2Type ℓ X) (x y : X) → --- PathP (λ i → --- (f : X → A) --- (f-coh : (x' : X) --- → f x' ≡ f (RP∞'-fields.notRP∞' (X , 2tx , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) x')) --- → ((x' : X) → --- PathP (λ i₁ → --- f (notNotRP∞' (X , 2tx , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) x' (~ i₁)) --- ≡ f (RP∞'-fields.notRP∞' (X , 2tx , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) x')) --- (f-coh x') --- (sym (f-coh (RP∞'-fields.notRP∞' (X , 2tx , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) x')))) → A) --- (F1 X 2tx x) --- (F1 X 2tx y) --- F2 X 2tx x = --- RP∞'-fields.elimRP∞' (X , 2tx , ∣ x ∣₁) x --- (λ i f coh p → f x) --- λ i f coh p → coh x i - --- F3 : (X : RP∞' ℓ) (x y z : fst X) → --- SquareP --- (λ i j → (f : fst X → A) --- (f-coh : (x' : fst X) --- → f x' --- ≡ f (RP∞'-fields.notRP∞' (fst X , fst (snd X) , squash₁ᵗ x y z i j) x')) --- → ((x' : fst X) → --- PathP --- (λ i₁ → --- f (notNotRP∞' (fst X , fst (snd X) , squash₁ᵗ x y z i j) x' (~ i₁)) --- ≡ f (RP∞'-fields.notRP∞' (fst X , fst (snd X) , squash₁ᵗ x y z i j) x')) --- (f-coh x') --- (sym (f-coh (RP∞'-fields.notRP∞' (fst X , fst (snd X) , squash₁ᵗ x y z i j) x')))) → A) --- (F2 (fst X) (fst (snd X)) x y) (F2 (fst X) (fst (snd X)) x z) --- (λ i f f-coh p → f x) (F2 (fst X) (fst (snd X)) y z) --- F3 = JRP∞' (CasesBool true --- (CasesBool true (λ i j f f-coh _ → f true) --- λ i j f f-coh p → f-coh true (i ∧ j)) --- (CasesBool true --- (λ i j f f-coh p --- → hcomp (λ k → λ {(i = i0) → p true (~ k) j --- ; (i = i1) → f true -- f-coh false (k ∧ j) -- f true --- ; (j = i0) → f (notNotRP∞' (RP∞'∙ ℓ) true (k ∨ i)) --- ; (j = i1) → f-coh false i}) --- (hcomp (λ k → λ {(i = i0) → f-coh false (~ j) --- ; (i = i1) → f true --- ; (j = i0) → f (isSetBool _ _ refl (notNotRP∞' (RP∞'∙ ℓ) true) k i) --- ; (j = i1) → f-coh false i}) --- (f-coh false (i ∨ ~ j)))) --- λ i j f f-coh p → f-coh true j)) - --- RP∞'→GroupoidRecId : ∀ {ℓ} {A : Type ℓ} (s : isGroupoid A) (X : RP∞' ℓ) --- → (f : fst X → A) --- → (f-coh : (x : _) → f x ≡ f (RP∞'-fields.notRP∞' X x)) --- → (p : (x : fst X) --- → PathP (λ i → f (notNotRP∞' X x (~ i)) ≡ f (RP∞'-fields.notRP∞' X x)) --- (f-coh x) --- (sym (f-coh (RP∞'-fields.notRP∞' X x)))) --- → Σ[ f-id ∈ ((x : fst X) → RP∞'→GroupoidRec s X f f-coh p ≡ f x) ] --- ((x : fst X) → PathP (λ i → {!!}) {!!} {!!}) -- (RP∞'→GroupoidRec s X f f-coh p) (RP∞'→GroupoidRec s X f f-coh p)) --- RP∞'→GroupoidRecId = {!!} - --- isGroupoidRP∞' : isGroupoid (RP∞' ℓ-zero) --- isGroupoidRP∞' = {!!} - --- EM₁-ℤ/2→RP∞'-emloop : (g : ℤ/2 .fst) → RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero --- EM₁-ℤ/2→RP∞'-emloop (zero , p) = refl --- EM₁-ℤ/2→RP∞'-emloop (suc zero , p) = Σ≡Prop (isPropIsRP∞ ℓ-zero) (ua notEquiv) --- EM₁-ℤ/2→RP∞'-emloop (suc (suc g) , p) = --- ⊥.rec (¬-<-zero (<-k+-cancel {k = 2} p)) - --- EM₁-ℤ/2→RP∞' : EM ℤ/2 1 → RP∞' ℓ-zero --- EM₁-ℤ/2→RP∞' = --- elimGroupoid _ (λ _ → isGroupoidRP∞') (RP∞'∙ ℓ-zero) --- EM₁-ℤ/2→RP∞'-emloop --- λ g h → ΣSquareSet (λ _ → isProp→isSet (isPropIsRP∞ ℓ-zero _)) --- (coh g h) --- where --- coh : (g h : ℤ/2 .fst) --- → Square refl (cong fst (EM₁-ℤ/2→RP∞'-emloop h)) --- (cong fst (EM₁-ℤ/2→RP∞'-emloop g)) --- (cong fst (EM₁-ℤ/2→RP∞'-emloop (g +ₘ h))) --- coh = ℤ/2-elim (ℤ/2-elim refl λ i j → ua notEquiv (i ∧ j)) --- (ℤ/2-elim (λ i j → ua notEquiv i) --- ((λ i j → ua notEquiv (~ j ∧ i)) ▷ sym help)) --- where --- help : ua notEquiv ≡ sym (ua notEquiv) --- help = cong ua (Σ≡Prop isPropIsEquiv (funExt (CasesBool true refl refl))) --- ∙ uaInvEquiv notEquiv - --- RP∞'→EM₁-ℤ/2' : EM ℤ/2 1 ≃ RP∞' ℓ-zero --- RP∞'→EM₁-ℤ/2' = EM₁-ℤ/2→RP∞' --- , record { equiv-proof = RP∞'pt→Prop (λ _ → isPropIsContr) --- ((embase , refl) --- , (isEmbedding→hasPropFibers′ {f = EM₁-ℤ/2→RP∞'} --- (Cubical.HITs.EilenbergMacLane1.elimProp _ (λ _ → isPropΠ λ _ → isPropIsEquiv _) --- (Cubical.HITs.EilenbergMacLane1.elimProp _ (λ _ → isPropIsEquiv _) --- (subst isEquiv (sym main) (isoToEquiv (invIso (compIso p2 p1)) .snd )))) _ _)) } --- where -- --- p1 : Iso Bool (Ω (EM∙ ℤ/2 1) .fst) --- p1 = compIso (fst Bool≅ℤGroup/2) (Iso-EM-ΩEM+1 0) - --- is1 : Iso (RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero) (Bool ≡ Bool) --- fun is1 = cong fst --- inv is1 p = Σ≡Prop (isPropIsRP∞ ℓ-zero) p --- rightInv is1 p = refl --- leftInv is1 p = --- ΣSquareSet (λ _ → isProp→isSet (isPropIsRP∞ ℓ-zero _)) --- λ i j → fst (p j) - --- p2 : Iso (Ω (RP∞' ℓ-zero , RP∞'∙ ℓ-zero) .fst) Bool --- p2 = compIso is1 (compIso (equivToIso univalence) Bool≃Charac) - --- main : cong EM₁-ℤ/2→RP∞' ≡ Iso.inv p2 ∘ Iso.inv p1 --- main = funExt λ x → cong (cong EM₁-ℤ/2→RP∞') (sym (Iso.rightInv p1 x)) --- ∙ sym (Iso.leftInv p2 _) --- ∙ cong (Iso.inv p2) (is3 (Iso.inv p1 x)) --- where --- is3 : (x : Bool) → Iso.fun p2 (cong EM₁-ℤ/2→RP∞' (Iso.fun p1 x)) ≡ x --- is3 false = refl --- is3 true = refl - - --- {- --- help : (X : RP∞' ℓ-zero) → isContr (Σ[ X' ∈ (EM ℤ/2 1) ] EM₁-ℤ/2→RP∞' X' ≡ X) --- help = --- RP∞'pt→Prop (λ _ → isPropIsContr) ((embase , refl) --- , (uncurry (elimSet _ (λ _ → isSetΠ λ _ --- → isOfHLevelPath' 2 --- (isOfHLevelΣ 3 emsquash (λ _ → isOfHLevelPath 3 isGroupoidRP∞' _ _)) _ _) --- (λ y → p1 {!!}) --- {!!}))) --- where --- fibr = Σ[ X' ∈ (EM ℤ/2 1) ] EM₁-ℤ/2→RP∞' X' ≡ RP∞'∙ ℓ-zero - - --- uap' : Bool → RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero --- uap' false = Σ≡Prop (isPropIsRP∞ ℓ-zero) (ua notEquiv) --- uap' true = refl - --- uapEq : Bool ≃ (RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero) --- fst uapEq = uap' --- snd uapEq = subst isEquiv (funExt l) (isoToEquiv (invIso uap) .snd) --- where - - --- l : (x : _) → inv uap x ≡ uap' x --- l false = refl --- l true = ΣSquareSet (λ _ → isProp→isSet (isPropIsRP∞ ℓ-zero _)) --- (uaIdEquiv) - --- p1 : (e : Bool) → Path fibr (embase , refl) (embase , uap' e) --- p1 false = ΣPathP ((emloop fone) , --- (ΣSquareSet (λ _ → isProp→isSet (isPropIsRP∞ ℓ-zero _)) --- ((λ i j → ua (notEquiv) (i ∧ ~ j)) --- ▷ (sym (uaInvEquiv notEquiv) ∙ cong ua (Σ≡Prop isPropIsEquiv refl))))) --- p1 true = refl - --- r : (e : RP∞'∙ ℓ-zero ≡ RP∞'∙ ℓ-zero) --- → Σ[ f ∈ Path fibr (embase , refl) (embase , e) ] --- {!(g : ℤ/2) → PathP !} --- r = {!!} --- -} - - --- RP∞'→EM₁-ℤ/2 : RP∞' ℓ-zero → EM ℤ/2 1 --- RP∞'→EM₁-ℤ/2 X = --- RP∞'→GroupoidRec emsquash X (λ _ → embase) --- (λ x → {!!}) --- {!!} - --- -- ∑RP∞' : (X : RP∞' ℓ-zero) (n : fst X → ℕ) → ℕ --- -- ∑RP∞' X n = --- -- RP∞'→SetRec isSetℕ X --- -- (λ x → n x +' n (RP∞'-fields.notRP∞' X x)) --- -- λ x → +'-comm (n x) _ ∙ cong (n (RP∞'-fields.notRP∞' X x) +'_) --- -- (cong n (sym (notNotRP∞' X x))) - - --- -- ∑RP∞'≡ : (X : RP∞' ℓ-zero) (x : fst X) (n : fst X → ℕ) --- -- → ∑RP∞' X n ≡ n x +' n (RP∞'-fields.notRP∞' X x) --- -- ∑RP∞'≡ = RP∞'pt→Prop (λ _ → isPropΠ2 λ _ _ → isSetℕ _ _) --- -- λ { false n → +'-comm (n true) (n false) --- -- ; true n → refl} - --- -- ∑RP∞'Fubini : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) --- -- → (∑RP∞' Y (λ y → ∑RP∞' X (λ x → n x y))) --- -- ≡ (∑RP∞' X (λ x → ∑RP∞' Y (n x))) --- -- ∑RP∞'Fubini = --- -- RP∞'pt→Prop (λ X → isPropΠ2 λ _ _ → isSetℕ _ _) --- -- (RP∞'pt→Prop ((λ _ → isPropΠ λ _ → isSetℕ _ _)) --- -- λ n → move4 (n true true) (n false true) (n true false) (n false false) _+'_ --- -- +'-assoc --- -- +'-comm) - --- -- module _ {ℓ} (X : RP∞' ℓ) (A : fst X → Pointed ℓ) (B : Pointed ℓ) where --- -- BipointedUnordJoin : (f : ((x : fst X) → A x .fst) → fst B) → Type ℓ --- -- BipointedUnordJoin f = --- -- (g : (x : fst X) → A x .fst) --- -- → UnordJoinR X (λ x → g x ≡ A x .snd) --- -- → f g ≡ B .snd - - --- -- module _ {ℓ} (X Y : RP∞' ℓ) (A : fst X → fst Y → Pointed ℓ) (B : Pointed ℓ) where --- -- QuadpointedUnordJoin : (f : ((x : fst X) (y : fst Y) → A x y .fst) → fst B) → Type ℓ --- -- QuadpointedUnordJoin f = (g : (x : fst X) (y : fst Y) → A x y .fst) --- -- → UnordJoinR X (λ x → UnordJoinR Y λ y → g x y ≡ A x y .snd) --- -- → f g ≡ B .snd - --- -- open import Cubical.HITs.Join --- -- module _ {ℓ} (A B T : Pointed ℓ) --- -- (f : fst A → fst B → fst T) where --- -- BipointedJoinBool : Type ℓ --- -- BipointedJoinBool = (a : A .fst) (b : B .fst) --- -- → join (a ≡ A .snd) (b ≡ B .snd) --- -- → f a b ≡ T .snd - --- -- JS'l : singl (pt A) → Type ℓ --- -- JS'l (a , p) = (b : fst B) → f a b ≡ T .snd - --- -- JS'r : singl (pt B) → Type ℓ --- -- JS'r (b , p) = (a : fst A) → f a b ≡ T .snd - --- -- BipointedJoinBool' : Type ℓ --- -- BipointedJoinBool' = Σ[ l ∈ ((a : _) → JS'l a) ] --- -- Σ[ r ∈ ((a : _) → JS'r a) ] --- -- ((a : _) (b : _) → l a (fst b) ≡ r b (fst a)) - - --- -- BipointedJoinBool'' : Type ℓ --- -- BipointedJoinBool'' = Σ[ l ∈ JS'l (pt A , refl) ] --- -- Σ[ r ∈ JS'r (pt B , refl) ] --- -- l (pt B) ≡ r (pt A) - --- -- IS1 : Iso BipointedJoinBool BipointedJoinBool' --- -- fst (Iso.fun IS1 F) (a , p) b = F a b (inl (sym p)) --- -- fst (snd (Iso.fun IS1 F)) (b , p) a = F a b (inr (sym p)) --- -- snd (snd (Iso.fun IS1 F)) (a , p) (b , q) = cong (F a b) (push (sym p) (sym q)) --- -- Iso.inv IS1 (l , r , lr) a b (inl x) = l (a , sym x) b --- -- Iso.inv IS1 (l , r , lr) a b (inr x) = r (b , sym x) a --- -- Iso.inv IS1 (l , r , lr) a b (push p q i) = lr (a , sym p) (b , sym q) i --- -- Iso.rightInv IS1 _ = refl --- -- Iso.leftInv IS1 F = funExt λ a → funExt λ b → funExt --- -- λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} - --- -- IS2 : Iso BipointedJoinBool' BipointedJoinBool'' --- -- IS2 = compIso --- -- (Σ-cong-iso --- -- {B = λ l → Σ[ r ∈ ((a : _) → JS'r a) ] --- -- ((a : _) (b : _) → l a (fst b) ≡ r b (fst a))} --- -- {B' = λ l → Σ[ r ∈ ((a : _) → JS'r a) ] --- -- ((b : _) → l (fst b) ≡ r b (pt A))} --- -- (DomainContrIso (isContrSingl (pt A))) --- -- λ x → Σ-cong-iso-snd λ r → DomainContrIso (isContrSingl (pt A))) --- -- (Σ-cong-iso-snd --- -- λ l → Σ-cong-iso {B = λ r → ((b : _) → l (fst b) ≡ r b (pt A))} --- -- {B' = λ r → (l (pt B) ≡ r (pt A))} --- -- (DomainContrIso (isContrSingl (pt B))) --- -- λ _ → DomainContrIso (isContrSingl (pt B))) - --- -- FullIso : Iso BipointedJoinBool BipointedJoinBool'' --- -- FullIso = compIso IS1 IS2 - --- -- module _ (A B C D T : Pointed ℓ-zero) --- -- (f : fst A → fst B → fst C → fst D → fst T) where --- -- JS4 : Type --- -- JS4 = (a : fst A) (b : fst B) (c : fst C) (d : fst D) --- -- → join (join (a ≡ snd A) (b ≡ snd B)) (join (c ≡ snd C) (d ≡ snd D)) --- -- → f a b c d ≡ pt T - --- -- open import Cubical.HITs.SmashProduct - --- -- isOfHLevelEM→ : ∀ {ℓ} {G H : AbGroup ℓ} (n : ℕ) --- -- → isOfHLevel 3 (EM∙ G n →∙ EM∙ H (suc n)) --- -- isOfHLevelEM→ {G = G} {H = H} n = --- -- isConnected→∙ (suc n) 2 (isConnectedEM n) --- -- (subst (λ m → isOfHLevel m (EM H (suc n))) (+-comm 2 (suc n)) --- -- (hLevelEM _ (suc n))) - --- -- mainLem' : (A B T : Pointed ℓ-zero) --- -- → Iso (Σ[ f ∈ (fst A → fst B → fst T) ] BipointedJoinBool A B T f) --- -- (A ⋀∙ B →∙ T) --- -- mainLem' A B T = compIso (Σ-cong-iso-snd (λ f → FullIso A B T f)) is2 --- -- where --- -- F→ : (T : Type) (t : T) → Σ[ f ∈ (fst A → fst B → T) ] BipointedJoinBool'' A B (T , t) f → (A ⋀∙ B →∙ (T , t)) --- -- fst (F→ T t (f , fl , fr , flr)) (inl x) = t --- -- fst (F→ T t (f , fl , fr , flr)) (inr x) = f (fst x) (snd x) --- -- fst (F→ T t (f , fl , fr , flr)) (push (inl x) i) = fr x (~ i) --- -- fst (F→ T t (f , fl , fr , flr)) (push (inr x) i) = fl x (~ i) --- -- fst (F→ T t (f , fl , fr , flr)) (push (push a i₁) i) = flr (~ i₁) (~ i) --- -- snd (F→ T t (f , fl , fr , flr)) = refl - --- -- mmB : (T : Type) (f : A ⋀ B → T) --- -- → Σ[ g ∈ (fst A → fst B → T) ] BipointedJoinBool'' A B (T , f (inl tt)) g --- -- fst (mmB T f) a b = f (inr (a , b)) --- -- fst (snd (mmB T f)) b = cong f (sym (push (inr b))) --- -- fst (snd (snd (mmB T f))) c = cong f (sym (push (inl c))) --- -- snd (snd (snd (mmB T f))) j i = f (push (push tt (~ j)) (~ i)) - --- -- mm' : (T : Type) (f : A ⋀ B → T) (t : T) --- -- → (f (inl tt) ≡ t) --- -- → BipointedJoinBool'' A B (T , t) (λ a b → f (inr (a , b))) --- -- mm' T f = J> mmB T f .snd - --- -- mm : (T : Type) (f : A ⋀ B → T) (t : T) --- -- → (f (inl tt) ≡ t) --- -- → Σ[ f ∈ (fst A → fst B → T) ] BipointedJoinBool'' A B (T , t) f --- -- fst (mm T f t x) a b = f (inr (a , b)) --- -- snd (mm T f t x) = mm' T f t x - --- -- c1 : (T : Type) (f : A ⋀ B → T) (t : T) (p : f (inl tt) ≡ t) --- -- → F→ T t (mm T f t p) ≡ (f , p) --- -- c1 T f = --- -- J> cong (F→ T (f (inl tt))) --- -- (ΣPathP (refl , (transportRefl (mmB T f .snd)))) --- -- ∙ ΣPathP ((funExt ( --- -- λ { (inl x) → refl --- -- ; (inr x) → refl --- -- ; (push (inl x) i) → refl --- -- ; (push (inr x) i) → refl --- -- ; (push (push a i₁) i) → refl})) , refl) - --- -- is2 : Iso (Σ[ f ∈ (fst A → fst B → fst T) ] BipointedJoinBool'' A B T f) --- -- (A ⋀∙ B →∙ T) --- -- Iso.fun is2 f = F→ (fst T) (snd T) f --- -- Iso.inv is2 f = mm (fst T) (fst f) (snd T) (snd f) --- -- Iso.rightInv is2 f = c1 (fst T) (fst f) _ (snd f) --- -- Iso.leftInv is2 f = ΣPathP (refl , transportRefl (snd f)) - --- -- mainLem : (A B T : Pointed ℓ-zero) --- -- → Iso (Σ[ f ∈ (fst A → fst B → fst T) ] BipointedJoinBool A B T f) --- -- (A →∙ (B →∙ T ∙)) --- -- mainLem A B T = compIso (mainLem' A B T) SmashAdjIso - --- -- open Iso --- -- open import Cubical.Foundations.Equiv.HalfAdjoint - - --- -- Iso-BipointedUnordJoin-BipointedJoinBool : --- -- ∀ {ℓ} (A : Bool → Pointed ℓ) (B : Pointed ℓ) --- -- → (f : ((x : Bool) → A x .fst) → fst B) --- -- → Iso (BipointedUnordJoin (RP∞'∙ ℓ) A B f) --- -- (BipointedJoinBool (A true) (A false) B --- -- λ a b → f (CasesBool true a b)) --- -- Iso-BipointedUnordJoin-BipointedJoinBool A B f = --- -- compIso (codomainIsoDep (λ g → domIso join-UnordJoinR-iso)) --- -- (compIso (ΠIso ΠBool×Iso λ g --- -- → codomainIso (pathToIso (cong (_≡ B .snd) (cong f (sym (CasesBoolη g)))))) curryIso) - --- -- SteenrodFunType : (X : RP∞' ℓ-zero) (n : fst X → ℕ) → Type --- -- SteenrodFunType X n = --- -- Σ[ f ∈ (((x : fst X) → EM ℤ/2 (n x)) → EM ℤ/2 (∑RP∞' X n)) ] --- -- BipointedUnordJoin X (λ x → EM∙ ℤ/2 (n x)) (EM∙ ℤ/2 (∑RP∞' X n)) f - --- -- isSetSteenrodFunTypeBoolIso : (n : Bool → ℕ) --- -- → Iso (SteenrodFunType (RP∞'∙ ℓ-zero) n) --- -- (EM∙ ℤ/2 (n true) →∙ (EM∙ ℤ/2 (n false) →∙ EM∙ ℤ/2 (n true +' n false) ∙)) --- -- isSetSteenrodFunTypeBoolIso n = --- -- (compIso (Σ-cong-iso-snd (λ f → Iso-BipointedUnordJoin-BipointedJoinBool _ _ _)) --- -- (compIso (invIso (Σ-cong-iso (compIso (invIso curryIso) (invIso (ΠIso ΠBool×Iso λ f → idIso))) --- -- λ g → idIso)) --- -- (mainLem (EM∙ ℤ/2 (n true)) (EM∙ ℤ/2 (n false)) --- -- (EM∙ ℤ/2 (n true +' n false))))) - --- -- isSetSteenrodFunTypeBoolIsoId : (n : Bool → ℕ) (f : _) (x : _) (y : _) --- -- → Iso.fun (isSetSteenrodFunTypeBoolIso n) f .fst x .fst y ≡ fst f (×→ΠBool (x , y)) --- -- isSetSteenrodFunTypeBoolIsoId n f x y = transportRefl _ - --- -- isSetSteenrodFunType : (X : RP∞' ℓ-zero) (n : fst X → ℕ) → isSet (SteenrodFunType X n) --- -- isSetSteenrodFunType = RP∞'pt→Prop (λ _ → isPropΠ λ _ → isPropIsOfHLevel 2) --- -- λ n → isOfHLevelRetractFromIso 2 --- -- (isSetSteenrodFunTypeBoolIso n) --- -- (isConnected→∙ (suc (n true)) 1 --- -- (isConnectedEM (n true)) --- -- (subst (λ m → isOfHLevel m (EM∙ ℤ/2 (n false) →∙ EM∙ ℤ/2 (n true +' n false))) --- -- (cong suc (+-comm 1 (n true)) ) --- -- (isConnected→∙ (suc (n false)) (suc (n true)) --- -- (isConnectedEM (n false)) --- -- (subst (λ m → isOfHLevel (suc m) (EM ℤ/2 (n true +' n false))) --- -- (cong suc (+'≡+ (n true) (n false)) ∙ (+-comm (suc (n true)) (n false))) --- -- (hLevelEM ℤ/2 (n true +' n false)))))) - --- -- SteenrodFunType≡ : (X : RP∞' ℓ-zero) (n : fst X → ℕ) --- -- (f g : SteenrodFunType X n) --- -- → fst f ≡ fst g --- -- → f ≡ g --- -- SteenrodFunType≡ = --- -- RP∞'pt→Prop (λ X → isPropΠ4 λ n _ _ _ → isSetSteenrodFunType X n _ _) --- -- λ n f g p → sym (Iso.leftInv (isSetSteenrodFunTypeBoolIso n) f) --- -- ∙∙ cong (inv (isSetSteenrodFunTypeBoolIso n)) --- -- (→∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) --- -- (funExt (λ x → →∙Homogeneous≡ (isHomogeneousEM _) --- -- (funExt λ y → transportRefl _ --- -- ∙ funExt⁻ p (×→ΠBool (x , y)) --- -- ∙ sym (transportRefl _))))) --- -- ∙∙ Iso.leftInv (isSetSteenrodFunTypeBoolIso n) g - --- -- cp : {n m : ℕ} → EM ℤ/2 n → EM ℤ/2 m → EM ℤ/2 (n +' m) --- -- cp = _⌣ₖ_ {G'' = ℤ/2Ring} - --- -- Sq : (X : RP∞' ℓ-zero) (x : fst X) (n : fst X → ℕ) --- -- (f : ((x₁ : fst X) → EM ℤ/2 (n x₁))) → EM ℤ/2 (∑RP∞' X n) --- -- Sq X x n f = --- -- subst (EM ℤ/2) (sym (∑RP∞'≡ X x n)) --- -- (cp (f x) (f (RP∞'-fields.notRP∞' X x))) - --- -- Sq-bool : (n : Bool → ℕ) (f : ((x₁ : Bool) → EM ℤ/2 (n x₁))) --- -- → EM ℤ/2 (∑RP∞' (RP∞'∙ ℓ-zero) n) --- -- Sq-bool n f = cp (f true) (f false) - --- -- sq≡ : (n : Bool → ℕ) (f : ((x₁ : Bool) → EM ℤ/2 (n x₁))) --- -- → Sq (RP∞'∙ ℓ-zero) true n f ≡ Sq-bool n f --- -- sq≡ n f = (λ j → subst (EM ℤ/2) --- -- (isSetℕ _ _ (sym (∑RP∞'≡ (RP∞'∙ ℓ-zero) true n)) refl j) --- -- (Sq-bool n f)) --- -- ∙ transportRefl _ - - --- -- sq' : (n : Bool → ℕ) --- -- → SmashPt (EM∙ ℤ/2 (n true)) (EM∙ ℤ/2 (n false)) --- -- →∙ EM∙ ℤ/2 (n true +' n false) --- -- fst (sq' n) basel = 0ₖ (n true +' n false) --- -- fst (sq' n) baser = 0ₖ (n true +' n false) --- -- fst (sq' n) (proj x y) = cp x y --- -- fst (sq' n) (gluel a i) = ⌣ₖ-0ₖ {G'' = ℤ/2Ring} (n true) (n false) a i --- -- fst (sq' n) (gluer b i) = 0ₖ-⌣ₖ {G'' = ℤ/2Ring} (n true) (n false) b i --- -- snd (sq' n) = refl - --- -- sq-coh-bool : (n : Bool → ℕ) --- -- → BipointedJoinBool (EM∙ ℤ/2 (n true)) (EM∙ ℤ/2 (n false)) --- -- (EM∙ ℤ/2 (∑RP∞' (RP∞'∙ ℓ-zero) n)) --- -- (λ a b → Sq-bool n (CasesBool true a b)) --- -- sq-coh-bool n = --- -- mainLem' (EM∙ ℤ/2 (n true)) (EM∙ ℤ/2 (n false)) --- -- (EM∙ ℤ/2 (∑RP∞' (RP∞'∙ ℓ-zero) n)) .Iso.inv --- -- (sq' n ∘∙ (⋀→Smash , refl)) .snd - --- -- sq-coh : (X : RP∞' ℓ-zero) (x : fst X) (n : fst X → ℕ) --- -- → BipointedUnordJoin X (λ x₁ → EM∙ ℤ/2 (n x₁)) --- -- (EM∙ ℤ/2 (∑RP∞' X n)) (Sq X x n) --- -- sq-coh = JRP∞' λ n → Iso-BipointedUnordJoin-BipointedJoinBool --- -- (λ x₁ → EM∙ ℤ/2 (n x₁)) --- -- (EM∙ ℤ/2 (∑RP∞' (RP∞'∙ ℓ-zero) n)) --- -- (Sq (RP∞'∙ ℓ-zero) true n) .Iso.inv --- -- λ g x r → sq≡ n _ ∙ sq-coh-bool n g x r - --- -- comm-field : (X : RP∞' ℓ-zero) (x : fst X) (n : fst X → ℕ) --- -- → Path (SteenrodFunType X n) --- -- (Sq X x n --- -- , sq-coh X x n) --- -- (Sq X (RP∞'-fields.notRP∞' X x) n --- -- , sq-coh X (RP∞'-fields.notRP∞' X x) n) --- -- comm-field = JRP∞' λ n → SteenrodFunType≡ (RP∞'∙ ℓ-zero) n _ _ --- -- (funExt λ f → cong (subst (EM ℤ/2) (sym (∑RP∞'≡ (RP∞'∙ ℓ-zero) true n))) --- -- (⌣ₖ-commℤ/2 (n true) (n false) (f true) (f false)) --- -- ∙ help _ (+'-comm (n false) (n true)) _ --- -- (sym (∑RP∞'≡ (RP∞'∙ ℓ-zero) true n)) --- -- (sym (∑RP∞'≡ (RP∞'∙ ℓ-zero) false n)) --- -- (cp (f false) (f true))) --- -- where --- -- help : {n : ℕ} (m : ℕ) (p : n ≡ m) (l : ℕ) (q : m ≡ l) (r : n ≡ l) --- -- → (x : EM ℤ/2 n) --- -- → subst (EM ℤ/2) q (subst (EM ℤ/2) p x) ≡ subst (EM ℤ/2) r x --- -- help = J> (J> λ r x --- -- → transportRefl _ --- -- ∙ λ j → subst (EM ℤ/2) (isSetℕ _ _ refl r j) x) - - --- -- SQ : (X : RP∞' ℓ-zero) (n : fst X → ℕ) → SteenrodFunType X n --- -- SQ X n = RP∞'→SetRec (isSetSteenrodFunType X n) X --- -- (λ x → Sq X x n , sq-coh X x n) --- -- (λ x → comm-field X x n) - --- -- STSQ : (X : RP∞' ℓ-zero) (n : fst X → ℕ) --- -- (f : ((x₁ : fst X) → EM ℤ/2 (n x₁))) → EM ℤ/2 (∑RP∞' X n) --- -- STSQ X n f = SQ X n .fst f - --- -- SQId : (X : RP∞' ℓ-zero) (n : fst X → ℕ) --- -- (x : fst X) (g : (x₁ : fst X) → EM ℤ/2 (n x₁)) --- -- → SQ X n .fst g ≡ Sq X x n g --- -- SQId X n x g i = --- -- RP∞'→SetRecEq (isSetSteenrodFunType X n) X --- -- (λ x → Sq X x n , sq-coh X x n) --- -- (λ x → comm-field X x n) x i .fst g - --- -- SQBool : (n : Bool → ℕ) (g : _) → SQ (RP∞'∙ ℓ-zero) n .fst g ≡ cp (g true) (g false) --- -- SQBool n g = SQId (RP∞'∙ ℓ-zero) n true g ∙ sq≡ n g - --- -- ΣQuadpointTy : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) --- -- → Type --- -- ΣQuadpointTy X Y n = --- -- Σ[ f ∈ (((x : fst X) (y : fst Y) → EM ℤ/2 (n x y)) --- -- → EM ℤ/2 (∑RP∞' X λ x → ∑RP∞' Y λ y → n x y)) ] --- -- QuadpointedUnordJoin X Y --- -- (λ x y → EM∙ ℤ/2 (n x y)) --- -- (EM∙ ℤ/2 (∑RP∞' X λ x → ∑RP∞' Y λ y → n x y)) f - --- -- SQ4 : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) --- -- → ΣQuadpointTy X Y n --- -- fst (SQ4 X Y n) f = SQ X (λ x → ∑RP∞' Y (n x)) .fst λ x → SQ Y (n x) .fst (f x) --- -- snd (SQ4 X Y n) g (inlR p) = SQ X (λ x → ∑RP∞' Y (n x)) .snd _ (inlR (fst p , SQ Y (n (fst p)) .snd _ (snd p))) --- -- snd (SQ4 X Y n) g (inrR f) = SQ X (λ x → ∑RP∞' Y (n x)) .snd _ (inrR λ x → SQ Y (n x) .snd _ (f x)) --- -- snd (SQ4 X Y n) g (pushR p f r i) = --- -- SQ X (λ x → ∑RP∞' Y (n x)) .snd _ --- -- (pushR (fst p , SQ Y (n (fst p)) .snd _ (snd p)) --- -- (λ x → SQ Y (n x) .snd _ (f x)) --- -- (cong (SQ Y (n (fst p)) .snd (λ x₂ → g (fst p) x₂)) r) i) - --- -- SQ4comm' : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) --- -- → Σ[ f ∈ (((x : fst X) (y : fst Y) → EM ℤ/2 (n x y)) --- -- → EM ℤ/2 (∑RP∞' Y λ y → ∑RP∞' X λ x → n x y)) ] --- -- QuadpointedUnordJoin Y X --- -- (λ y x → EM∙ ℤ/2 (n x y)) --- -- (EM∙ ℤ/2 (∑RP∞' Y λ y → ∑RP∞' X λ x → n x y)) --- -- (λ g → f (λ x y → g y x)) --- -- fst (SQ4comm' X Y n) f = --- -- SQ Y (λ z → ∑RP∞' X (λ x → n x z)) .fst --- -- λ y → SQ X (λ x → n x y) .fst λ x → f x y --- -- snd (SQ4comm' X Y n) g (inlR x) = --- -- SQ Y (λ z → ∑RP∞' X (λ x₁ → n x₁ z)) .snd --- -- (λ y → SQ X (λ x₁ → n x₁ y) .fst (λ x₁ → g y x₁)) --- -- (inlR (fst x , SQ X (λ x₁ → n x₁ (fst x)) .snd (g (fst x)) (snd x))) --- -- snd (SQ4comm' X Y n) g (inrR f) = --- -- SQ Y (λ y → ∑RP∞' X (λ x → n x y)) .snd --- -- (λ y → SQ X (λ x → n x y) .fst (g y)) --- -- (inrR λ y → SQ X (λ x → n x y) .snd (g y) (f y)) --- -- snd (SQ4comm' X Y n) g (pushR a b x i) = --- -- SQ Y (λ y → ∑RP∞' X (λ x → n x y)) .snd --- -- (λ y → SQ X (λ x₁ → n x₁ y) .fst (g y)) --- -- (pushR (fst a , SQ X (λ x → n x (fst a)) .snd (g (fst a)) (snd a)) --- -- (λ y → SQ X (λ x → n x y) .snd (g y) (b y)) --- -- (cong (SQ X (λ x₁ → n x₁ (fst a)) .snd (g (fst a))) x) i) - --- -- SQ4comm : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) --- -- → ΣQuadpointTy X Y n --- -- fst (SQ4comm X Y n) f = --- -- subst (EM ℤ/2) (∑RP∞'Fubini X Y n) --- -- (SQ Y (λ z → ∑RP∞' X (λ x → n x z)) .fst --- -- λ y → SQ X (λ x → n x y) .fst λ x → f x y) --- -- snd (SQ4comm X Y n) f p = --- -- cong (subst (EM ℤ/2) (∑RP∞'Fubini X Y n)) --- -- (SQ4comm' X Y n .snd (λ y x → f x y) --- -- (UnordJoinFubiniFun X Y _ p)) --- -- ∙ subst-EM∙ (∑RP∞'Fubini X Y n) .snd - --- -- mainA : (A B C D T : Pointed ℓ-zero) --- -- → Iso (Σ[ f ∈ (fst A → fst B → fst C → fst D → fst T) ] --- -- JS4 A B C D T f) --- -- (A →∙ (B →∙ C →∙ D →∙ T ∙ ∙ ∙)) --- -- mainA A B C D T = --- -- compIso --- -- (compIso IsMain --- -- (Σ-cong-iso --- -- (codomainIso (codomainIso (mainLem C D T))) --- -- λ f → codomainIsoDep λ a --- -- → codomainIsoDep λ b --- -- → codomainIso --- -- (compIso (congIso {x = f a b} (mainLem C D T)) --- -- (pathToIso (cong (fun (mainLem C D T) (f a b) ≡_) --- -- (mainIs .snd)) )))) --- -- (mainLem A B (C →∙ D →∙ T ∙ ∙)) --- -- where --- -- T1 = (Σ[ f ∈ (fst A → fst B → fst C → fst D → fst T) ] JS4 A B C D T f) --- -- T2 = (Σ (fst A → fst B → Σ (fst C → fst D → fst T) (BipointedJoinBool C D T)) --- -- (BipointedJoinBool A B (Σ (fst C → fst D → fst T) (BipointedJoinBool C D T) --- -- , (λ _ _ → snd T) , (λ _ _ _ → refl) ))) - --- -- module _ (a : fst A) (b : fst B) (c : fst C) (d : fst D) --- -- (p : join (a ≡ snd A) (b ≡ snd B)) --- -- (q : join (c ≡ snd C) (d ≡ snd D)) (i j k : I) where --- -- fill₁ : T1 → fst T --- -- fill₁ (f , g) = --- -- hfill (λ k → λ {(i = i0) → g a b c d (push p q k) j --- -- ; (i = i1) → snd T --- -- ; (j = i0) → g a b c d (inl p) i --- -- ; (j = i1) → snd T}) --- -- (inS (g a b c d (inl p) (i ∨ j))) k - --- -- fill₂ : T2 → fst T --- -- fill₂ (f , g) = --- -- hfill (λ k → λ {(i = i0) → g a b p j .snd c d q (~ k) --- -- ; (i = i1) → f a b .snd c d q (~ k ∨ j) --- -- ; (j = i0) → f a b .snd c d q (~ k) --- -- ; (j = i1) → snd T}) --- -- (inS (snd T)) k - --- -- T1→T2 : T1 → T2 --- -- fst (fst (T1→T2 (f , p)) a b) c d = f a b c d --- -- snd (fst (T1→T2 (f , p)) a b) c d t = p a b c d (inr t) --- -- fst (snd (T1→T2 (f , p)) a b t i) c d = p a b c d (inl t) i --- -- snd (snd (T1→T2 (f , p)) a b t i) c d q j = fill₁ a b c d t q i j i1 (f , p) - --- -- T2→T1 : T2 → T1 --- -- fst (T2→T1 (f , p)) a b c d = f a b .fst c d --- -- snd (T2→T1 (f , p)) a b c d (inl x) i = p a b x i .fst c d --- -- snd (T2→T1 (f , p)) a b c d (inr x) = f a b .snd c d x --- -- snd (T2→T1 (f , g)) a b c d (push p q i) j = fill₂ a b c d p q i j i1 (f , g) - --- -- IsMain : Iso T1 T2 --- -- Iso.fun IsMain = T1→T2 --- -- Iso.inv IsMain = T2→T1 --- -- fst (Iso.rightInv IsMain (f , p) i) = fst (T1→T2 (T2→T1 (f , p))) --- -- fst (snd (Iso.rightInv IsMain (f , p) i) a b t j) = p a b t j .fst --- -- snd (snd (Iso.rightInv IsMain (f , g) i) a b p j) c d q k = --- -- hcomp (λ r → λ {(i = i0) → fill₁ a b c d p q j k r ((λ a b c d → f a b .fst c d) , (snd (T2→T1 (f , g)))) --- -- ; (i = i1) → snd (g a b p j) c d q k --- -- ; (j = i0) → sd r i k --- -- ; (j = i1) → snd T --- -- ; (k = i0) → g a b p j .fst c d --- -- ; (k = i1) → snd T}) --- -- (cb k j i) --- -- where --- -- sq : (k i r : I) → fst T --- -- sq k i r = --- -- hfill (λ r → λ {(i = i0) → g a b p k .snd c d q (~ r) --- -- ; (i = i1) → f a b .snd c d q (~ r ∨ k) --- -- ; (k = i0) → f a b .snd c d q (~ r) --- -- ; (k = i1) → snd T}) --- -- (inS (snd T)) --- -- r - --- -- sd : Cube (λ i j → sq j i i1) --- -- (λ i k → f a b .snd c d q k) --- -- (λ r k → fill₂ a b c d p q r k i1 --- -- (f , λ a b p → g a b p)) --- -- (λ r k → snd (f a b) c d q k) --- -- (λ r i → f a b .fst c d) (λ _ _ → pt T) --- -- sd r i k = --- -- hcomp (λ w → λ {(r = i0) → sq k i w --- -- ; (r = i1) → f a b .snd c d q (~ w ∨ k) --- -- ; (i = i0) → fill₂ a b c d p q r k w (f , λ a b p → g a b p) --- -- ; (i = i1) → f a b .snd c d q (~ w ∨ k) --- -- ; (k = i0) → f a b .snd c d q (~ w) --- -- ; (k = i1) → snd T}) --- -- (snd T) - --- -- cb : Cube (λ j i → g a b p j .fst c d) (λ _ _ → pt T) --- -- (λ i j → sq i j i1) (λ _ _ → pt T) --- -- (λ k j → g a b p (j ∨ k) .fst c d) --- -- λ k j → snd (g a b p j) c d q k --- -- cb k j i = --- -- hcomp (λ r → λ {(i = i0) → g a b p (k ∨ j) .snd c d q (~ r) --- -- ; (i = i1) → snd (g a b p j) c d q (~ r ∨ k) --- -- ; (j = i1) → snd T --- -- ; (k = i0) → g a b p j .snd c d q (~ r) --- -- ; (k = i1) → snd T}) --- -- (snd T) --- -- fst (Iso.leftInv IsMain (f , g) i) = f --- -- snd (Iso.leftInv IsMain (f , g) i) a b c d (inl p) = g a b c d (inl p) --- -- snd (Iso.leftInv IsMain (f , g) i) a b c d (inr p) = g a b c d (inr p) --- -- snd (Iso.leftInv IsMain (f , g) i) a b c d (push p q j) k = --- -- hcomp (λ r → λ {(i = i0) → fill₂ a b c d p q j k r --- -- (fst (T1→T2 (f , g)) , snd (T1→T2 (f , g))) --- -- ; (i = i1) → ss r j k --- -- ; (j = i0) → s2 r k i --- -- ; (j = i1) → g a b c d (inr q) (~ r ∨ k) --- -- ; (k = i0) → g a b c d (inr q) (~ r) --- -- ; (k = i1) → pt T}) --- -- (snd T) --- -- where --- -- PP : (i j k : I) → fst T --- -- PP i j k = doubleWhiskFiller {A = λ i → g a b c d (inr q) (~ i) ≡ pt T} refl --- -- (λ i j → g a b c d (inr q) (~ i ∨ j)) --- -- (λ j k → g a b c d (push p q (~ j)) k) --- -- k i j - --- -- s2 : Cube (λ _ _ → pt T) (λ k i → g a b c d (inl p) k) --- -- (λ r i → g a b c d (inr q) (~ r)) (λ _ _ → pt T) --- -- (λ r k → fill₁ a b c d p q k (~ r) i1 (f , g)) --- -- λ r k → PP r k i1 --- -- s2 r k i = --- -- hcomp (λ j → λ {(r = i0) → pt T --- -- ; (r = i1) → g a b c d (push p q (~ j ∧ i)) k --- -- ; (k = i0) → g a b c d (push p q (j ∨ i)) (~ r) --- -- ; (k = i1) → pt T --- -- ; (i = i0) → fill₁ a b c d p q k (~ r) j (f , g) --- -- ; (i = i1) → PP r k j}) --- -- (g a b c d (push p q i) (k ∨ ~ r)) - --- -- ss : Cube (λ _ _ → pt T) (λ j k → g a b c d (push p q j) k) --- -- (λ i j → PP i j i1) --- -- (λ r k → g a b c d (inr q) (~ r ∨ k)) --- -- (λ r j → g a b c d (inr q) (~ r)) --- -- (λ r j → pt T) --- -- ss r j k = --- -- hcomp (λ w → λ {(r = i0) → pt T --- -- ; (r = i1) → g a b c d (push p q (~ w ∨ j)) k --- -- ; (j = i1) → g a b c d (inr q) (~ r ∨ k) --- -- ; (k = i0) → g a b c d (inr q) (~ r) --- -- ; (k = i1) → pt T}) --- -- (g a b c d (inr q) (~ r ∨ k)) - --- -- mainIs : (isoToPath (mainLem C D T) i0 , (λ c d → pt T) , λ _ _ _ → refl) --- -- ≃∙ (C →∙ (D →∙ T ∙) ∙) --- -- fst mainIs = isoToEquiv (mainLem C D T) --- -- snd mainIs = cong (Iso.fun SmashAdjIso) --- -- (ΣPathP ((funExt ( --- -- λ { (inl x) → refl --- -- ; (inr x) → refl --- -- ; (push (inl x) i) → refl --- -- ; (push (inr x) i) → refl --- -- ; (push (push a i₁) i) → refl})) , refl)) --- -- ∙ SmashAdj≃∙ .snd --- -- {- --- -- ((isoToEquiv (mainLem C D T))) --- -- (cong (Iso.fun SmashAdjIso) --- -- (ΣPathP ((funExt ( --- -- λ { (inl x) → refl --- -- ; (inr x) → refl --- -- ; (push (inl x) i) → refl --- -- ; (push (inr x) i) → refl --- -- ; (push (push a i₁) i) → refl})) , refl)) --- -- ∙ SmashAdj≃∙ .snd) --- -- -} - --- -- ΣQuadpointTyBool : (n : Bool → Bool → ℕ) --- -- → Iso (ΣQuadpointTy (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n) --- -- (EM∙ ℤ/2 (n true true) --- -- →∙ (EM∙ ℤ/2 (n true false) --- -- →∙ EM∙ ℤ/2 (n false true) --- -- →∙ EM∙ ℤ/2 (n false false) --- -- →∙ EM∙ ℤ/2 ((n true true +' n true false) --- -- +' (n false true +' n false false)) ∙ ∙ ∙)) --- -- ΣQuadpointTyBool n = --- -- (compIso --- -- (Σ-cong-iso --- -- (invIso (compIso (invIso curryIso) --- -- (compIso (invIso curryIso) --- -- (compIso (invIso curryIso) --- -- (domIso (invIso help)))))) --- -- λ f → invIso (compIso --- -- (compIso (invIso curryIso) --- -- (compIso (invIso curryIso) --- -- (compIso (invIso curryIso) --- -- (invIso --- -- (ΠIso idIso --- -- λ {(((x , y) , z) , w) --- -- → domIso (compIso join-UnordJoinR-iso --- -- (Iso→joinIso --- -- join-UnordJoinR-iso --- -- join-UnordJoinR-iso))}))))) --- -- (ΠIso (invIso help) λ _ → idIso))) --- -- (mainA (EM∙ ℤ/2 (n true true)) --- -- (EM∙ ℤ/2 (n true false)) --- -- (EM∙ ℤ/2 (n false true)) --- -- (EM∙ ℤ/2 (n false false)) --- -- (EM∙ ℤ/2 ((n true true +' n true false) --- -- +' (n false true +' n false false))))) --- -- where --- -- help : Iso ((x y : fst (RP∞'∙ ℓ-zero)) → EM∙ ℤ/2 (n x y) .fst) --- -- (((EM ℤ/2 (n true true) --- -- × EM ℤ/2 (n true false)) --- -- × EM ℤ/2 (n false true)) --- -- × EM ℤ/2 (n false false)) --- -- help = compIso (compIso ΠBool×Iso --- -- (prodIso ΠBool×Iso ΠBool×Iso)) --- -- (invIso Σ-assoc-Iso) - --- -- {- --- -- ΣQuadpointTyPres : (n : Bool → Bool → ℕ) --- -- (f : ΣQuadpointTy (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n) --- -- → Path (EM ℤ/2 (n true true) → --- -- (EM ℤ/2 (n true false) → --- -- EM ℤ/2 (n false true) → --- -- EM ℤ/2 (n false false) → --- -- EM ℤ/2 --- -- ((n true true +' n true false) --- -- +' (n false true +' n false false)))) --- -- (λ x y z w → Iso.fun (ΣQuadpointTyBool n) f .fst x .fst y .fst z .fst w) --- -- λ x y z w → f .fst (CasesBool true --- -- (CasesBool true x y) --- -- (CasesBool true z w)) --- -- ΣQuadpointTyPres n f = refl --- -- -} - --- -- isSetΣQuadPoint : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) --- -- → isSet (ΣQuadpointTy X Y n) --- -- isSetΣQuadPoint = --- -- RP∞'pt→Prop (λ Y → isPropΠ2 (λ _ _ → isPropIsSet)) --- -- (RP∞'pt→Prop (λ Y → isPropΠ (λ _ → isPropIsSet)) --- -- λ n → isOfHLevelRetractFromIso 2 --- -- (ΣQuadpointTyBool n) --- -- (isConnected→∙ (suc (n true true)) 1 --- -- (isConnectedEM (n true true)) --- -- (isConnected→∙ (suc (n true false)) (n true true + 1) --- -- (isConnectedEM (n true false)) --- -- (isConnected→∙ (suc (n false true)) --- -- ((n true false) + (n true true + 1)) --- -- (isConnectedEM (n false true)) --- -- (isConnected→∙ --- -- (suc (n false false)) --- -- (n false true + (n true false + (n true true + 1))) --- -- (isConnectedEM (n false false)) --- -- (subst (λ k → isOfHLevel (suc k) (EM ℤ/2 (N' n))) --- -- (lem n) --- -- (hLevelEM ℤ/2 (N' n)))))))) --- -- where --- -- N' : (n : Bool → Bool → ℕ) → ℕ --- -- N' n = ((n true true +' n true false) +' (n false true +' n false false)) - --- -- lem : (n : _) → suc (N' n) --- -- ≡ (n false false + (n false true + (n true false + (n true true + 1)))) --- -- lem n = cong suc ((cong₂ _+'_ (+'≡+ (n true true) (n true false)) --- -- (+'≡+ (n false true) (n false false)) --- -- ∙ +'≡+ _ _) --- -- ∙ +-assoc (n true true + n true false ) (n false true) (n false false)) --- -- ∙ cong (_+ n false false) --- -- (cong (_+ n false true) --- -- ((λ i → +-comm (+-comm 1 (n true true) i) (n true false) i)) --- -- ∙ +-comm _ (n false true)) --- -- ∙ +-comm _ (n false false) - --- -- ΣQuadPoint≡ : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) --- -- (f g : ΣQuadpointTy X Y n) --- -- → ((t : _) → f .fst t ≡ g .fst t) --- -- → f ≡ g --- -- ΣQuadPoint≡ = --- -- RP∞'pt→Prop (λ X → isPropΠ5 λ Y n _ _ _ → isSetΣQuadPoint X Y n _ _) --- -- (RP∞'pt→Prop (λ Y → isPropΠ4 λ n _ _ _ --- -- → isSetΣQuadPoint (RP∞'∙ ℓ-zero) Y n _ _) --- -- λ n f g p --- -- → sym (Iso.leftInv (ΣQuadpointTyBool n) f) --- -- ∙∙ cong (inv (ΣQuadpointTyBool n)) --- -- (main n f g p) --- -- ∙∙ Iso.leftInv (ΣQuadpointTyBool n) g) --- -- where --- -- module _ (n : Bool → Bool → ℕ) --- -- (f g : ΣQuadpointTy (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n) --- -- (p : (t : (x y : fst (RP∞'∙ ℓ-zero)) → EM ℤ/2 (n x y)) --- -- → f .fst t ≡ g .fst t) where --- -- p' : (x : _) (y : _) (z : _) (w : _) --- -- → fun (ΣQuadpointTyBool n) f .fst x .fst y .fst z .fst w --- -- ≡ fun (ΣQuadpointTyBool n) g .fst x .fst y .fst z .fst w --- -- p' x y z w = p (CasesBool true --- -- (CasesBool true x y) --- -- (CasesBool true z w)) - --- -- module _ {ℓ ℓ' ℓT} {C : Pointed ℓ} {D : Pointed ℓ'} {T : Pointed ℓT} --- -- (hom : isHomogeneous T) where --- -- isHomogeneous→∙₂ : isHomogeneous (C →∙ (D →∙ T ∙) ∙) --- -- isHomogeneous→∙₂ = isHomogeneous→∙ (isHomogeneous→∙ hom) - --- -- module _ {ℓ'' : Level} {B : Pointed ℓ''} where --- -- isHomogeneous→∙₃ : isHomogeneous (B →∙ (C →∙ (D →∙ T ∙) ∙) ∙) --- -- isHomogeneous→∙₃ = isHomogeneous→∙ isHomogeneous→∙₂ - --- -- isHomogeneous→∙₄ : ∀ {ℓ'''} {A : Pointed ℓ'''} --- -- → isHomogeneous (A →∙ (B →∙ (C →∙ (D →∙ T ∙) ∙) ∙) ∙) --- -- isHomogeneous→∙₄ = isHomogeneous→∙ isHomogeneous→∙₃ - - - --- -- T = (n true true +' n true false) +' (n false true +' n false false) - --- -- m4 : (x : EM ℤ/2 (n true true)) (y : EM ℤ/2 (n true false)) --- -- (z : EM ℤ/2 (n false true)) --- -- → fun (ΣQuadpointTyBool n) f .fst x .fst y .fst z --- -- ≡ fun (ΣQuadpointTyBool n) g .fst x .fst y .fst z --- -- m4 x y z = →∙Homogeneous≡ (isHomogeneousEM T) (funExt (p' x y z)) - --- -- m3 : (x : EM ℤ/2 (n true true)) (y : EM ℤ/2 (n true false)) --- -- → fun (ΣQuadpointTyBool n) f .fst x .fst y --- -- ≡ fun (ΣQuadpointTyBool n) g .fst x .fst y --- -- m3 x y = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM T)) --- -- (funExt (m4 x y)) - --- -- m2 : (x : EM ℤ/2 (n true true)) --- -- → fun (ΣQuadpointTyBool n) f .fst x --- -- ≡ fun (ΣQuadpointTyBool n) g .fst x --- -- m2 x = →∙Homogeneous≡ (isHomogeneous→∙₂ (isHomogeneousEM T)) --- -- (funExt (m3 x)) - --- -- main : fun (ΣQuadpointTyBool n) f ≡ fun (ΣQuadpointTyBool n) g --- -- main = →∙Homogeneous≡ (isHomogeneous→∙₃ (isHomogeneousEM T)) --- -- (funExt m2) - --- -- SQ4≡SQ4comm : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) --- -- → SQ4 X Y n ≡ SQ4comm X Y n --- -- SQ4≡SQ4comm = --- -- RP∞'pt→Prop (λ _ → isPropΠ2 λ Y n → isSetΣQuadPoint _ Y n _ _) --- -- (RP∞'pt→Prop (λ Y → isPropΠ λ n → isSetΣQuadPoint _ Y n _ _) --- -- λ n → ΣQuadPoint≡ _ _ _ _ _ --- -- λ f → SQBool (λ x → ∑RP∞' (RP∞'∙ ℓ-zero) (n x)) --- -- (λ x → SQ (RP∞'∙ ℓ-zero) (n x) .fst (f x)) --- -- ∙ cong₂ cp (SQBool (n true) (f true)) --- -- (SQBool (n false) (f false)) --- -- ∙ help (EM ℤ/2) (λ n m x y → cp {n = n} {m = m} x y) --- -- ⌣ₖ-commℤ/2 assoc⌣ₖ --- -- (n true true) (n true false) --- -- (n false true) (n false false) --- -- (f true true) (f true false) --- -- (f false true) (f false false) --- -- (∑RP∞'Fubini (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n) --- -- ∙ cong (subst (EM ℤ/2) (∑RP∞'Fubini (RP∞'∙ ℓ-zero) (RP∞'∙ ℓ-zero) n)) --- -- (sym (SQBool (λ z → ∑RP∞' (RP∞'∙ ℓ-zero) (λ x → n x z)) --- -- (λ y → SQ (RP∞'∙ ℓ-zero) (λ x → n x y) .fst (λ x → f x y)) --- -- ∙ cong₂ cp (SQBool (λ x → n x true) (λ x → f x true)) --- -- (SQBool (λ x → n x false) (λ x → f x false))))) --- -- where --- -- help : ∀ {ℓ} (A : ℕ → Type ℓ) (compA : (n m : ℕ) (x : A n) (y : A m) → A (n +' m)) --- -- → (⌣comm : (n m : ℕ) (x : A n) (y : A m) --- -- → compA n m x y --- -- ≡ subst A (+'-comm m n) (compA m n y x)) --- -- → (⌣assoc : (n m l : ℕ) (x : A n) (y : A m) (z : A l) --- -- → compA (n +' m) l (compA n m x y) z --- -- ≡ subst A (+'-assoc n m l) --- -- (compA n (m +' l) x (compA m l y z))) --- -- → (n m k l : ℕ) (x : A n) (y : A m) (z : A k) (w : A l) --- -- → (p : ((n +' k) +' (m +' l)) ≡ ((n +' m) +' (k +' l))) --- -- → compA (n +' m) (k +' l) (compA n m x y) (compA k l z w) --- -- ≡ subst A p (compA (n +' k) (m +' l) (compA n k x z) (compA m l y w)) --- -- help A compA ⌣comm ⌣assoc n m k l x y z w p = --- -- (sym (transportRefl _) --- -- ∙ (λ i → subst A (isSetℕ _ _ refl (((sym (+'-assoc n m (k +' l))) ∙ p5 ∙ p4) ∙ p) i) --- -- (compA (n +' m) (k +' l) (compA n m x y) (compA k l z w)))) --- -- ∙ substComposite A ((sym (+'-assoc n m (k +' l)) ∙ p5 ∙ p4)) p _ --- -- ∙ cong (subst A p) --- -- ((substComposite A (sym (+'-assoc n m (k +' l))) (p5 ∙ p4) _ --- -- ∙ cong (subst A (p5 ∙ p4)) --- -- (cong (subst A (sym (+'-assoc n m (k +' l)))) --- -- (⌣assoc _ _ _ x y (compA k l z w)) --- -- ∙ subst⁻Subst A (+'-assoc n m (k +' l)) _)) --- -- ∙ substComposite A (cong (_+'_ n) ((p1 ∙ p2) ∙ p3)) p4 --- -- (compA n (m +' (k +' l)) x (compA m (k +' l) y (compA k l z w))) --- -- ∙ cong (subst A (+'-assoc n k (m +' l))) --- -- (sym (substLems _ _ ((p1 ∙ p2) ∙ p3) _ .snd --- -- x (compA m (k +' l) y (compA k l z w))) --- -- ∙ cong (compA n (k +' (m +' l)) x) --- -- (substComposite A (p1 ∙ p2) p3 (compA m (k +' l) y (compA k l z w)) --- -- ∙ cong (subst A p3) --- -- ((substComposite A p1 p2 (compA m (k +' l) y (compA k l z w)) --- -- ∙ cong (subst A (cong (_+' l) (+'-comm m k))) --- -- (sym (⌣assoc m k l y z w))) --- -- ∙ sym (substLems _ _ (+'-comm m k) _ .fst (compA m k y z) w) --- -- ∙ cong (λ z → compA (k +' m) l z w) --- -- (sym (⌣comm k m z y))) --- -- ∙ cong (subst A p3) --- -- (⌣assoc _ _ _ z y w) --- -- ∙ subst⁻Subst A (+'-assoc k m l) _)) --- -- ∙ sym (⌣assoc _ _ _ x z (compA m l y w))) --- -- where --- -- p1 = +'-assoc m k l --- -- p2 = cong (_+' l) (+'-comm m k) --- -- p3 = sym (+'-assoc k m l) --- -- p4 = +'-assoc n k (m +' l) --- -- p5 = cong (n +'_) ((p1 ∙ p2) ∙ p3) - --- -- substLems : (n n' : ℕ) (p : n ≡ n') (m : ℕ) --- -- → ((x : A n) (y : A m) --- -- → compA n' m (subst A p x) y ≡ subst A (cong (_+' m) p) (compA n m x y)) --- -- × ((x : A m) (y : A n) --- -- → compA m n' x (subst A p y) ≡ subst A (cong (m +'_) p) (compA m n x y)) --- -- substLems n = J> λ m --- -- → (λ x y → cong (λ x → compA n m x y) (transportRefl x) --- -- ∙ sym (transportRefl _)) --- -- , ((λ x y → cong (λ y → compA m n x y) (transportRefl y) --- -- ∙ sym (transportRefl _))) - --- -- PreCartan : (X Y : RP∞' ℓ-zero) (n : fst X → fst Y → ℕ) --- -- → (f : (x : fst X) (y : fst Y) → EM ℤ/2 (n x y)) --- -- → STSQ X (λ x → ∑RP∞' Y (n x)) (λ x → STSQ Y (n x) (f x)) --- -- ≡ subst (EM ℤ/2) (∑RP∞'Fubini X Y n) --- -- (STSQ Y (λ y → ∑RP∞' X (λ x → n x y)) --- -- (λ y → STSQ X (λ x → n x y) (λ x → f x y))) --- -- PreCartan X Y n f i = SQ4≡SQ4comm X Y n i .fst f - --- -- -- module _ (A B C D : Pointed ℓ-zero) (T : Pointed ℓ-zero) --- -- -- (f : A .fst × B .fst × C .fst × D .fst --- -- -- → fst T) where --- -- -- BipointedJoinBool : Type --- -- -- BipointedJoinBool = (a : A .fst) (b : B .fst) (c : C .fst) (d : D .fst) --- -- -- → join (a ≡ A .snd) --- -- -- (join (b ≡ B .snd) --- -- -- (join (c ≡ C .snd) --- -- -- (d ≡ D .snd))) --- -- -- → f (a , b , c , d) ≡ T .snd - --- -- -- BipointedJoinBool* : Type --- -- -- BipointedJoinBool* = (b : B .fst) (c : C .fst) (d : D .fst) --- -- -- → Σ[ fr ∈ ((a : A .fst) → join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd)) --- -- -- → f (a , b , c , d) ≡ T .snd) ] --- -- -- ((x : singl (A .snd)) → --- -- -- (Σ[ fl ∈ (f (x .fst , b , c , d) ≡ T .snd) ] --- -- -- ((t : join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd))) --- -- -- → fl ≡ fr (x .fst) t))) - --- -- -- BipointedJoinBool** : Type --- -- -- BipointedJoinBool** = (b : B .fst) (c : C .fst) (d : D .fst) --- -- -- → Σ[ fr ∈ ((a : A .fst) → join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd)) --- -- -- → f (a , b , c , d) ≡ T .snd) ] --- -- -- Σ[ fl ∈ (f (pt A , b , c , d) ≡ T .snd) ] --- -- -- ((t : join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd))) --- -- -- → fl ≡ fr (pt A) t) - --- -- -- JS₂ : Type --- -- -- JS₂ = (c : C .fst) (d : D .fst) --- -- -- → Σ[ fr ∈ ((a : A .fst) (b : fst B) --- -- -- → join (c ≡ C .snd) (d ≡ D .snd) --- -- -- → f (a , b , c , d) ≡ T .snd) ] --- -- -- Σ[ fl ∈ ((b : fst B) → f (pt A , b , c , d) ≡ T .snd) ] --- -- -- Σ[ flp ∈ ((b : fst B) → (t : join (c ≡ C .snd) (d ≡ D .snd)) --- -- -- → fl b ≡ fr (pt A) b t) ] --- -- -- ((x : singl (B .snd)) --- -- -- → Σ[ frl ∈ ((a : fst A) → f (a , fst x , c , d) ≡ T .snd) ] --- -- -- Σ[ frp ∈ ((a : fst A) (t : join (c ≡ C .snd) (d ≡ D .snd)) → frl a ≡ fr a (fst x) t) ] --- -- -- Σ[ r ∈ fl (fst x) ≡ frl (pt A) ] --- -- -- ((t : join (c ≡ C .snd) (d ≡ D .snd)) --- -- -- → Square r (flp (fst x) t) refl (frp (pt A) t))) - --- -- -- JS₂* : Type --- -- -- JS₂* = (c : C .fst) (d : D .fst) --- -- -- → Σ[ fr ∈ ((a : A .fst) (b : fst B) --- -- -- → join (c ≡ C .snd) (d ≡ D .snd) --- -- -- → f (a , b , c , d) ≡ T .snd) ] --- -- -- Σ[ fl ∈ ((b : fst B) → f (pt A , b , c , d) ≡ T .snd) ] --- -- -- Σ[ flp ∈ ((b : fst B) → (t : join (c ≡ C .snd) (d ≡ D .snd)) --- -- -- → fl b ≡ fr (pt A) b t) ] --- -- -- (Σ[ frl ∈ ((a : fst A) → f (a , pt B , c , d) ≡ T .snd) ] --- -- -- Σ[ frp ∈ ((a : fst A) (t : join (c ≡ C .snd) (d ≡ D .snd)) → frl a ≡ fr a (pt B) t) ] --- -- -- Σ[ r ∈ fl (pt B) ≡ frl (pt A) ] --- -- -- ((t : join (c ≡ C .snd) (d ≡ D .snd)) --- -- -- → Square r (flp (pt B) t) refl (frp (pt A) t))) - --- -- -- module _ (fl₁ : ((b : fst B) (c : C .fst) (d : D .fst) → f (pt A , b , c , d) ≡ T .snd)) --- -- -- (fl₂ : ((a : fst A) (c : C .fst) (d : D .fst) → f (a , pt B , c , d) ≡ T .snd)) --- -- -- (fl₁₂ : (c : fst C) (d : fst D) → fl₁ (pt B) c d ≡ fl₂ (pt A) c d) --- -- -- where --- -- -- TL : singl (snd C) → Type --- -- -- TL (c , p) = --- -- -- Σ[ fr ∈ ((a : fst A) (b : fst B) (d : fst D) → f (a , b , c , d) ≡ T .snd) ] --- -- -- Σ[ flp ∈ ((b : fst B) (d : fst D) → fl₁ b c d ≡ fr (pt A) b d) ] --- -- -- Σ[ frp ∈ ((a : fst A) (d : fst D) → fl₂ a c d ≡ fr a (pt B) d) ] --- -- -- ((d : fst D) → Square (fl₁₂ c d) (flp (pt B) d) refl (frp (pt A) d)) --- -- -- TR : singl (snd D) → Type --- -- -- TR (d , p) = --- -- -- Σ[ fr ∈ ((a : fst A) (b : fst B) (c : fst C) → f (a , b , c , d) ≡ T .snd) ] --- -- -- Σ[ flp ∈ ((b : fst B) (c : fst C) → fl₁ b c d ≡ fr (pt A) b c) ] --- -- -- Σ[ frp ∈ ((a : fst A) (c : fst C) → fl₂ a c d ≡ fr a (pt B) c) ] --- -- -- ((c : fst C) → Square (fl₁₂ c d) (flp (pt B) c) refl (frp (pt A) c)) - --- -- -- TLR : (c : singl (snd C)) (d : singl (snd D)) → TL c → TR d → Type --- -- -- TLR (c , p) (d , q) (frₗ , flpₗ , frpₗ , sqₗ) (frᵣ , flpᵣ , frpᵣ , sqᵣ) = --- -- -- Σ[ frₗᵣ ∈ (((a : fst A) (b : fst B) → frₗ a b d ≡ frᵣ a b c)) ] --- -- -- Σ[ flpₗᵣ ∈ ((b : fst B) → PathP (λ i → fl₁ b c d ≡ frₗᵣ (pt A) b i) (flpₗ b d) (flpᵣ b c)) ] --- -- -- Σ[ frpₗᵣ ∈ ((a : fst A) → PathP (λ i → fl₂ a c d ≡ frₗᵣ a (pt B) i) (frpₗ a d) (frpᵣ a c)) ] --- -- -- Cube (sqₗ d) (sqᵣ c) --- -- -- (λ i j → fl₁₂ c d j) (flpₗᵣ (pt B)) --- -- -- (λ j i → fl₁ (pt B) c d) (frpₗᵣ (pt A)) - - --- -- -- JS₃* : Type --- -- -- JS₃* = Σ[ fl₁ ∈ ((b : fst B) (c : C .fst) (d : D .fst) → f (pt A , b , c , d) ≡ T .snd) ] --- -- -- Σ[ fl₂ ∈ ((a : fst A) (c : C .fst) (d : D .fst) → f (a , pt B , c , d) ≡ T .snd) ] --- -- -- Σ[ fl₁₂ ∈ ((c : fst C) (d : fst D) → fl₁ (pt B) c d ≡ fl₂ (pt A) c d) ] --- -- -- Σ[ l ∈ ((c : _) → TL fl₁ fl₂ fl₁₂ c) ] --- -- -- Σ[ r ∈ ((c : _) → TR fl₁ fl₂ fl₁₂ c) ] --- -- -- ((c : singl (snd C)) (d : singl (snd D)) → TLR fl₁ fl₂ fl₁₂ c d (l c) (r d)) - --- -- -- JS₃** : Type --- -- -- JS₃** = Σ[ fl₁ ∈ ((b : fst B) (c : C .fst) (d : D .fst) → f (pt A , b , c , d) ≡ T .snd) ] --- -- -- Σ[ fl₂ ∈ ((a : fst A) (c : C .fst) (d : D .fst) → f (a , pt B , c , d) ≡ T .snd) ] --- -- -- Σ[ fl₁₂ ∈ ((c : fst C) (d : fst D) → fl₁ (pt B) c d ≡ fl₂ (pt A) c d) ] --- -- -- Σ[ l ∈ (TL fl₁ fl₂ fl₁₂ (pt C , refl)) ] --- -- -- Σ[ r ∈ (TR fl₁ fl₂ fl₁₂ (pt D , refl)) ] --- -- -- (TLR fl₁ fl₂ fl₁₂ (pt C , refl) (pt D , refl) l r) - --- -- -- module _ (js : JS₃**) where --- -- -- open import Cubical.HITs.SmashProduct --- -- -- JS₃**→' : (A ⋀∙ (B ⋀∙ (C ⋀∙ D))) →∙ T --- -- -- fst JS₃**→' (inl x) = pt T --- -- -- fst JS₃**→' (inr (a , inl x)) = {!f (a , ?)!} -- pt T --- -- -- fst JS₃**→' (inr (a , inr (b , inl x))) = {!!} -- pt T --- -- -- fst JS₃**→' (inr (a , inr (b , inr (c , d)))) = f (a , b , c , d) --- -- -- fst JS₃**→' (inr (a , inr (b , push (inl x) i))) = snd js .snd .snd .snd .fst .fst a b x (~ i) --- -- -- fst JS₃**→' (inr (a , inr (b , push (inr x) i))) = snd js .snd .snd .fst .fst a b x (~ i) --- -- -- fst JS₃**→' (inr (a , inr (b , push (push tt i₁) i))) = snd js .snd .snd .snd .snd .fst a b (~ i₁) (~ i) --- -- -- fst JS₃**→' (inr (a , push (inl x) i)) = {!f (a , pt B , pt C , ?)!} -- pt T --- -- -- fst JS₃**→' (inr (a , push (inr (inl x)) i)) = {!f (a , pt B , pt C , x)!} -- snd T --- -- -- fst JS₃**→' (inr (a , push (inr (inr (c , d))) i)) = snd js .fst a c d (~ i) --- -- -- fst JS₃**→' (inr (a , push (inr (push (inl x) i₁)) i)) = {!!} --- -- -- fst JS₃**→' (inr (a , push (inr (push (inr x) i₁)) i)) = {!snd js .snd .snd .fst .snd .snd .fst a x (~ i₁) (~ i)!} --- -- -- fst JS₃**→' (inr (a , push (inr (push (push a₁ i₂) i₁)) i)) = {!!} --- -- -- fst JS₃**→' (inr (a , push (push a₁ i₁) i)) = {!!} --- -- -- fst JS₃**→' (push a i) = {!!} --- -- -- snd JS₃**→' = {!!} - --- -- -- JS₃**→ : A →∙ (B →∙ (C →∙ (D →∙ T ∙) ∙) ∙) --- -- -- fst (fst (fst (fst JS₃**→ a) b) c) d = f (a , b , c , d) --- -- -- snd (fst (fst (fst JS₃**→ a) b) c) = snd js .snd .snd .snd .fst .fst a b c --- -- -- fst (snd (fst (fst JS₃**→ a) b) i) d = snd js .snd .snd .fst .fst a b d i --- -- -- snd (snd (fst (fst JS₃**→ a) b) i) = {!snd js .snd .snd .snd .snd .fst a b i!} --- -- -- fst (fst (snd (fst JS₃**→ a) i) c) d = snd js .fst a c d i --- -- -- snd (fst (snd (fst JS₃**→ a) i) c) j = {!!} --- -- -- fst (snd (snd (fst JS₃**→ a) i) j) d = {!!} --- -- -- snd (snd (snd (fst JS₃**→ a) i) j) k = {!!} --- -- -- fst (fst (fst (snd JS₃**→ i) b) c) d = fst js b c d i --- -- -- snd (fst (fst (snd JS₃**→ i) b) c) j = {!!} --- -- -- fst (snd (fst (snd JS₃**→ i) b) j) d = {!!} --- -- -- snd (snd (fst (snd JS₃**→ i) b) j) k = {!!} --- -- -- fst (fst (snd (snd JS₃**→ i) j) c) d = {!!} --- -- -- snd (fst (snd (snd JS₃**→ i) j) c) k = {!!} --- -- -- snd (snd (snd JS₃**→ i) j) = {!!} - --- -- -- Iso-JS₃*-JS₃** : Iso JS₃* JS₃** --- -- -- Iso-JS₃*-JS₃** = --- -- -- Σ-cong-iso-snd λ f' → Σ-cong-iso-snd λ g → Σ-cong-iso-snd λ fg --- -- -- → compIso (Σ-cong-iso-snd (λ r → Σ-cong-iso-snd λ s --- -- -- → compIso (DomainContrIso (isContrSingl (snd C))) --- -- -- (DomainContrIso (isContrSingl (snd D))))) --- -- -- (compIso (ΣDomainContrIso {C = λ c l → Σ ((d : _) → TR f' g fg d) --- -- -- λ r → TLR f' g fg c (pt D , refl) l (r (pt D , refl))} (isContrSingl (snd C))) --- -- -- (Σ-cong-iso-snd λ l → --- -- -- ΣDomainContrIso {C = λ d r → TLR f' g fg (pt C , refl) d l r} (isContrSingl (snd D)))) - --- -- -- Iso₂ : Iso BipointedJoinBool** JS₂ --- -- -- fst (Iso.fun Iso₂ F c d) a b t = F b c d .fst a (inr t) --- -- -- fst (snd (Iso.fun Iso₂ F c d)) b = F b c d .snd .fst --- -- -- fst (snd (snd (Iso.fun Iso₂ F c d))) b t = F b c d .snd .snd (inr t) --- -- -- fst (snd (snd (snd (Iso.fun Iso₂ F c d))) (b , p)) a = --- -- -- F b c d .fst a (inl (sym p)) --- -- -- fst (snd (snd (snd (snd (Iso.fun Iso₂ F c d))) (b , p))) a t = --- -- -- cong (F b c d .fst a) (push (sym p) t) --- -- -- fst (snd (snd (snd (snd (snd (Iso.fun Iso₂ F c d))) (b , p)))) = --- -- -- F b c d .snd .snd (inl (sym p)) --- -- -- snd (snd (snd (snd (snd (snd (Iso.fun Iso₂ F c d))) (b , p)))) t = --- -- -- cong (F b c d .snd .snd) (push (sym p) t) --- -- -- fst (Iso.inv Iso₂ F b c d) a (inl x) = F c d .snd .snd .snd (b , sym x) .fst a --- -- -- fst (Iso.inv Iso₂ F b c d) a (inr t) = F c d .fst a b t --- -- -- fst (Iso.inv Iso₂ F b c d) a (push x t i) = --- -- -- F c d .snd .snd .snd (b , sym x) .snd .fst a t i --- -- -- fst (snd (Iso.inv Iso₂ F b c d)) = F c d .snd .fst b --- -- -- snd (snd (Iso.inv Iso₂ F b c d)) (inl x) = --- -- -- F c d .snd .snd .snd (b , sym x) .snd .snd .fst --- -- -- snd (snd (Iso.inv Iso₂ F b c d)) (inr t) = F c d .snd .snd .fst b t --- -- -- snd (snd (Iso.inv Iso₂ F b c d)) (push a t i) = --- -- -- F c d .snd .snd .snd (b , sym a) .snd .snd .snd t i --- -- -- Iso.rightInv Iso₂ = λ _ → refl --- -- -- Iso.leftInv Iso₂ F = funExt λ b → funExt λ c → funExt λ x → --- -- -- ΣPathP (funExt (λ a → funExt λ { (inl x) → refl --- -- -- ; (inr x) → refl --- -- -- ; (push a x i) → refl}) --- -- -- , ΣPathP (refl , (funExt (λ { (inl x) → refl --- -- -- ; (inr x) → refl --- -- -- ; (push a x i) → refl})))) - --- -- -- Iso₂* : Iso BipointedJoinBool** JS₂* --- -- -- Iso₂* = --- -- -- compIso Iso₂ --- -- -- (codomainIsoDep λ c → codomainIsoDep λ d → --- -- -- Σ-cong-iso-snd λ f → Σ-cong-iso-snd λ g → Σ-cong-iso-snd --- -- -- λ h → DomainContrIso (isContrSingl (B .snd))) - --- -- -- Iso₃ : Iso JS₂* JS₃* --- -- -- fst (Iso.fun Iso₃ F) b c d = F c d .snd .fst b --- -- -- fst (snd (Iso.fun Iso₃ F)) a c d = F c d .snd .snd .snd .fst a --- -- -- fst (snd (snd (Iso.fun Iso₃ F))) c d = F c d .snd .snd .snd .snd .snd .fst --- -- -- fst (fst (snd (snd (snd (Iso.fun Iso₃ F)))) (c , p)) a b d = --- -- -- F c d .fst a b (inl (sym p)) --- -- -- fst (snd (fst (snd (snd (snd (Iso.fun Iso₃ F)))) (c , p))) b d = --- -- -- F c d .snd .snd .fst b (inl (sym p)) --- -- -- fst (snd (snd (fst (snd (snd (snd (Iso.fun Iso₃ F)))) (c , p)))) a d = --- -- -- F c d .snd .snd .snd .snd .fst a (inl (sym p)) --- -- -- snd (snd (snd (fst (snd (snd (snd (Iso.fun Iso₃ F)))) (c , p)))) d = --- -- -- F c d .snd .snd .snd .snd .snd .snd (inl (sym p)) --- -- -- fst (fst (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (d , p)) a b c = --- -- -- F c d .fst a b (inr (sym p)) --- -- -- fst (snd (fst (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (d , p))) b c = --- -- -- F c d .snd .snd .fst b (inr (sym p)) --- -- -- fst (snd (snd (fst (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (d , p)))) a c = --- -- -- F c d .snd .snd .snd .snd .fst a (inr (sym p)) --- -- -- snd (snd (snd (fst (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (d , p)))) c = --- -- -- F c d .snd .snd .snd .snd .snd .snd (inr (sym p)) --- -- -- fst (snd (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (c , p) (d , q)) a b = --- -- -- cong (F c d .fst a b) (push (sym p) (sym q)) --- -- -- fst (snd (snd (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (c , p) (d , q))) b i = --- -- -- F c d .snd .snd .fst b (push (sym p) (sym q) i) --- -- -- fst (snd (snd (snd (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (c , p) (d , q)))) a i = --- -- -- F c d .snd .snd .snd .snd .fst a (push (sym p) (sym q) i) --- -- -- snd (snd (snd (snd (snd (snd (snd (snd (Iso.fun Iso₃ F))))) (c , p) (d , q)))) i = --- -- -- F c d .snd .snd .snd .snd .snd .snd (push (sym p) (sym q) i) --- -- -- fst (Iso.inv Iso₃ F c d) a b (inl x) = --- -- -- F .snd .snd .snd .fst (c , sym x) .fst a b d --- -- -- fst (Iso.inv Iso₃ F c d) a b (inr x) = --- -- -- F .snd .snd .snd .snd .fst (d , sym x) .fst a b c --- -- -- fst (Iso.inv Iso₃ F c d) a b (push p q i) = --- -- -- F .snd .snd .snd .snd .snd (c , sym p) (d , sym q) .fst a b i --- -- -- fst (snd (Iso.inv Iso₃ F c d)) b = F .fst b c d --- -- -- fst (snd (snd (Iso.inv Iso₃ F c d))) b (inl x) = F .snd .snd .snd .fst (c , sym x) .snd .fst b d --- -- -- fst (snd (snd (Iso.inv Iso₃ F c d))) b (inr x) = F .snd .snd .snd .snd .fst (d , sym x) .snd .fst b c --- -- -- fst (snd (snd (Iso.inv Iso₃ F c d))) b (push p q i) = --- -- -- F .snd .snd .snd .snd .snd (c , sym p) (d , sym q) .snd .fst b i --- -- -- fst (snd (snd (snd (Iso.inv Iso₃ F c d)))) a = F .snd .fst a c d --- -- -- fst (snd (snd (snd (snd (Iso.inv Iso₃ F c d))))) a (inl x) = --- -- -- F .snd .snd .snd .fst (c , sym x) .snd .snd .fst a d --- -- -- fst (snd (snd (snd (snd (Iso.inv Iso₃ F c d))))) a (inr x) = --- -- -- F .snd .snd .snd .snd .fst (d , sym x) .snd .snd .fst a c --- -- -- fst (snd (snd (snd (snd (Iso.inv Iso₃ F c d))))) a (push p q i) = --- -- -- F .snd .snd .snd .snd .snd (c , sym p) (d , sym q) .snd .snd .fst a i --- -- -- fst (snd (snd (snd (snd (snd (Iso.inv Iso₃ F c d)))))) = F .snd .snd .fst c d --- -- -- snd (snd (snd (snd (snd (snd (Iso.inv Iso₃ F c d)))))) (inl x) = --- -- -- F .snd .snd .snd .fst (c , sym x) .snd .snd .snd d --- -- -- snd (snd (snd (snd (snd (snd (Iso.inv Iso₃ F c d)))))) (inr x) = --- -- -- F .snd .snd .snd .snd .fst (d , sym x) .snd .snd .snd c --- -- -- snd (snd (snd (snd (snd (snd (Iso.inv Iso₃ F c d)))))) (push p q i) = --- -- -- F .snd .snd .snd .snd .snd (c , sym p) (d , sym q) .snd .snd .snd i --- -- -- Iso.rightInv Iso₃ _ = refl --- -- -- Iso.leftInv Iso₃ F = funExt λ c → funExt λ d → --- -- -- ΣPathP ((funExt (λ a → funExt λ b → funExt --- -- -- λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl})) --- -- -- , ΣPathP (refl , (ΣPathP ((funExt (λ b --- -- -- → funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl})) --- -- -- , (ΣPathP (refl , (ΣPathP ((funExt (λ a → --- -- -- funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl})) --- -- -- , (ΣPathP (refl , (funExt (λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl})))))))))))) - - --- -- -- Σ[ fl ∈ (f (pt A , b , c , d) ≡ T .snd) ] --- -- -- ((t : join (b ≡ B .snd) (join (c ≡ C .snd) (d ≡ D .snd))) --- -- -- → fl ≡ fr (pt A) t) - --- -- -- Iso₁ : Iso BipointedJoinBool BipointedJoinBool* --- -- -- fst (Iso.fun Iso₁ F b c d) a x = F a b c d (inr x) --- -- -- fst (snd (Iso.fun Iso₁ F b c d) (a , p)) = F a b c d (inl (sym p)) --- -- -- snd (snd (Iso.fun Iso₁ F b c d) (a , p)) t = cong (F a b c d) (push (sym p) t) --- -- -- Iso.inv Iso₁ F a b c d (inl x) = F b c d .snd (a , sym x) .fst --- -- -- Iso.inv Iso₁ F a b c d (inr t) = F b c d .fst a t --- -- -- Iso.inv Iso₁ F a b c d (push p t i) = F b c d .snd (a , sym p) .snd t i --- -- -- Iso.rightInv Iso₁ = λ _ → refl --- -- -- Iso.leftInv Iso₁ F = funExt λ a → funExt λ b → funExt λ c → funExt λ d --- -- -- → funExt λ { (inl x) → refl ; (inr x) → refl ; (push a x i) → refl} - --- -- -- Iso₁' : Iso BipointedJoinBool BipointedJoinBool** --- -- -- Iso₁' = compIso Iso₁ (codomainIsoDep λ b → codomainIsoDep λ c → codomainIsoDep λ d --- -- -- → Σ-cong-iso-snd λ f → DomainContrIso (isContrSingl (A .snd))) - --- -- -- JoinStructureBool* : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) --- -- -- (f : A true true .fst × A true false .fst --- -- -- × A false true .fst × A false false .fst --- -- -- → fst B) --- -- -- → Type --- -- -- JoinStructureBool* A B f = --- -- -- (g : A true true .fst × A true false .fst --- -- -- × A false true .fst × A false false .fst) --- -- -- → join (fst g ≡ A true true .snd) --- -- -- (join (snd g .fst ≡ A true false .snd) --- -- -- (join (snd (snd g) .fst ≡ A false true .snd) --- -- -- (snd (snd g) .snd ≡ A false false .snd))) --- -- -- → f g ≡ B .snd - --- -- -- 4→∙ : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) → Type ℓ-zero --- -- -- 4→∙ A B = A true true →∙ (A true false →∙ (A false true →∙ (A false false →∙ B ∙) ∙) ∙) - --- -- -- 4→∙' : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) --- -- -- → Type ℓ-zero --- -- -- 4→∙' A B = --- -- -- Σ[ f ∈ (A true true .fst → A true false .fst --- -- -- → A false true .fst → A false false .fst → fst B) ] --- -- -- Σ[ f-inl-inl ∈ ((a : singl (A true true .snd)) (b : _) (c : _) (d : _) → f (fst a) b c d ≡ pt B) ] --- -- -- Σ[ f-inl-inr ∈ ((b : singl (A true false .snd)) (a : _) (c : _) (d : _) → f a (fst b) c d ≡ pt B) ] --- -- -- Σ[ f-inl-push ∈ (((a : singl (A true true .snd)) (b : singl (A true false .snd)) (c : _) (d : _) --- -- -- → f-inl-inl a (fst b) c d ≡ f-inl-inr b (fst a) c d)) ] --- -- -- Σ[ f-inr-inl ∈ ((c : singl (A false true .snd)) (a : _) (b : _) (d : _) → f a b (fst c) d ≡ pt B) ] --- -- -- Σ[ f-inr-inr ∈ ((d : singl (A false false .snd)) (a : _) (b : _) (c : _) → f a b c (fst d) ≡ pt B) ] --- -- -- Σ[ f-inl-push ∈ ((c : singl (A false true .snd)) (d : singl (A false false .snd)) (a : _) (b : _) --- -- -- → f-inr-inl c a b (fst d) ≡ f-inr-inr d a b (fst c)) ] --- -- -- {!Σ[ f-inl-push ∈ ((c : singl (A false true .snd)) (d : singl (A false false .snd)) (a : _) (b : _) --- -- -- → f-inr-inl c a b (fst d) ≡ f-inr-inr d a b (fst c)) ]!} - --- -- -- pss : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) (f : _) → JoinStructureBool A B f --- -- -- pss A B f (x , y , z , w) (inl (inl p)) = {!p!} --- -- -- pss A B f (x , y , z , w) (inl (inr q)) = {!q!} --- -- -- pss A B f (x , y , z , w) (inl (push p q i)) = {!!} --- -- -- pss A B f (x , y , z , w) (inr (inl p)) = {!!} --- -- -- pss A B f (x , y , z , w) (inr (inr q)) = {!!} --- -- -- pss A B f (x , y , z , w) (inr (push p q i)) = {!!} --- -- -- pss A B f (x , y , z , w) (push (inl p) (inl q) i) = {!!} --- -- -- pss A B f (x , y , z , w) (push (inr p) (inl q) i) = {!!} --- -- -- pss A B f (x , y , z , w) (push (push p q i₁) (inl r) i) = {!!} --- -- -- pss A B f (x , y , z , w) (push p (inr q) i) = {!!} --- -- -- pss A B f (x , y , z , w) (push p (push q r i₁) i) = {!!} - - --- -- -- JoinStructureBoolD : (A : Bool → Bool → Pointed ℓ-zero) (B : Pointed ℓ-zero) --- -- -- → Σ _ (JoinStructureBool A B) --- -- -- → A true true →∙ (A true false →∙ (A false true →∙ (A false false →∙ B ∙) ∙) ∙) --- -- -- fst (fst (fst (fst (JoinStructureBoolD A B (f , p)) x) y) z) w = --- -- -- f (x , y , z , w) --- -- -- snd (fst (fst (fst (JoinStructureBoolD A B (f , p)) x) y) z) = --- -- -- p (x , y , z , snd (A false false)) (inr (inr refl)) --- -- -- fst (snd (fst (fst (JoinStructureBoolD A B (f , p)) x) y) i) w = --- -- -- p (x , y , snd (A false true) , w) (inr (inl refl)) i --- -- -- snd (snd (fst (fst (JoinStructureBoolD A B (f , p)) x) y) i) = {!!} --- -- -- fst (fst (snd (fst (JoinStructureBoolD A B (f , p)) x) i) z) w = --- -- -- p (x , snd (A true false) , z , w) (inl (inr refl)) i --- -- -- snd (fst (snd (fst (JoinStructureBoolD A B (f , p)) x) i) z) = {!!} --- -- -- fst (snd (snd (fst (JoinStructureBoolD A B (f , p)) x) i) j) w = {!!} --- -- -- snd (snd (snd (fst (JoinStructureBoolD A B (f , p)) x) i) j) = {!!} --- -- -- fst (fst (fst (snd (JoinStructureBoolD A B (f , p)) i) y) z) w = --- -- -- p (snd (A true true) , y , z , w) (inl (inl refl)) i --- -- -- snd (fst (fst (snd (JoinStructureBoolD A B (f , p)) i) y) z) j = {!!} --- -- -- snd (fst (snd (JoinStructureBoolD A B (f , p)) i) y) = {!!} --- -- -- snd (snd (JoinStructureBoolD A B (f , p)) i) = {!!} - --- -- -- module _ {ℓ} (X Y : RP∞' ℓ) (A : fst X → fst Y → Pointed ℓ) (B : Type ℓ) where - --- -- -- JoinStructure : ((f : (x : fst X) (y : fst Y) → A x y .fst) → B) → Type ℓ --- -- -- JoinStructure f = --- -- -- (g : (x : fst X) (y : fst Y) → A x y .fst) --- -- -- → UnordJoinR² X Y (λ x y → g x y ≡ pt (A x y)) --- -- -- → f g ≡ f (λ x y → pt (A x y)) - --- -- -- module _ {ℓ} (X Y : RP∞' ℓ) (A : fst X → fst Y → Pointed ℓ) (B : Type ℓ) where - --- -- -- JoinStructure→ : (t : (f : (x : fst X) (y : fst Y) → A x y .fst) → B) --- -- -- → JoinStructure X Y A B t --- -- -- → JoinStructure Y X (λ y x → A x y) B --- -- -- λ f → t λ x y → f y x --- -- -- JoinStructure→ f st g pr = --- -- -- st (λ x y → g y x) --- -- -- (UnordJoinFubiniFun Y X (λ x y → g x y ≡ pt (A y x)) pr) - - --- -- -- inr* : _ → UnordSmash∙² X Y A .fst --- -- -- inr* = inr - --- -- -- inr-case : (g : (x₁ : fst X) (y : fst Y) → A x₁ y .fst) (x : fst X) --- -- -- → UnordJoinR Y (λ y → g x y ≡ pt (A x y)) --- -- -- → Path (UnordSmash Y (A x)) (inr (g x)) (inr (λ y → pt (A x y))) --- -- -- inr-case g x (inlR (y , z)) = {!z!} --- -- -- inr-case g x (inrR z) = {!!} --- -- -- inr-case g x (pushR a b x₁ i) = {!!} - --- -- -- m2 : Σ _ (JoinStructure X Y A B) → UnordSmash∙² X Y A .fst → B --- -- -- m2 (f , p) (inl x) = {!!} --- -- -- m2 (f , p) (inr t) = {!!} --- -- -- m2 (f , p) (push a i) = {!!} - --- -- -- m1 : (f : UnordSmash∙² X Y A .fst → B) --- -- -- → Σ _ (JoinStructure X Y A B) --- -- -- fst (m1 f) g = f (inr λ x → inr (g x)) --- -- -- snd (m1 f) g (inlR (x , inlR r)) = {!cong !} --- -- -- snd (m1 f) g (inlR (x , inrR r)) = {!!} --- -- -- snd (m1 f) g (inlR (x , pushR a b x₁ i)) = {!!} --- -- -- snd (m1 f) g (inrR h) = cong f (cong inr* (funExt λ x → inr-case g x (h x))) --- -- -- snd (m1 f) g (pushR a b x i) = {!!} - --- -- -- -- Smash→ : (Σ[ h ∈ (UnordΠ X (fst ∘ A) → fst B) ] --- -- -- -- ((f : UnordΠ X ((λ r → fst r) ∘ A)) --- -- -- -- → UnordJoinR X (λ x → f x ≡ pt (A x)) --- -- -- -- → h f ≡ pt B)) --- -- -- -- → (UnordSmash X A → fst B) --- -- -- -- Smash→ (h , f) (inl x) = pt B --- -- -- -- Smash→ (h , f) (inr x) = h x --- -- -- -- Smash→ (h , f) (push (x , a) i) = --- -- -- -- f (UnordΣ→UnordΠ X A (x , a)) --- -- -- -- (inlR (RP∞'-fields.notRP∞' X x --- -- -- -- , RP∞'-fields.elimRP∞'β X x a --- -- -- -- (pt (A (RP∞'-fields.notRP∞' X x))) .snd)) (~ i) + main : (n m : ℕ) → (Iso.inv (HⁿRP∞≅ℤ/2 (n +' m) .fst) fone + ≡ _⌣_ {G'' = ℤ/2Ring} (Iso.inv (HⁿRP∞≅ℤ/2 n .fst) fone) + (Iso.inv (HⁿRP∞≅ℤ/2 m .fst) fone)) + main n m = eRP∞^-isGen (n +' m) ∙∙ sym (eRP∞^-comp n m) + ∙∙ sym (cong₂ (_⌣_ {G'' = ℤ/2Ring}) (eRP∞^-isGen n) (eRP∞^-isGen m)) + + +-- open import Cubical.Agebra.Monoids.Instance.NatVec + +module _ {ℓ : Level} (n : ℕ) where + private + RP∞ = EM ℤ/2 1 + + ℤ₂[x] = Poly ℤ/2CommRing 1 + + ℤ₂[X] = CommRing→Ring (PolyCommRing ℤ/2CommRing 1) + + open GradedRing-⊕HIT-index + NatMonoid + (Poly ℤ/2CommRing) + (λ n → CommRing→AbGroup (PolyCommRing ℤ/2CommRing n) .snd) + + open RingStr (snd (H*R ℤ/2Ring RP∞)) + renaming + ( 0r to 0H*RP + ; 1r to 1H*RP + ; _+_ to _+H*RP_ + ; -_ to -H*RP_ + ; _·_ to _cupS_ + ; +Assoc to +H*RPAssoc + ; +IdR to +H*RPIdR + ; +Comm to +H*RPComm + ; ·Assoc to ·H*RPAssoc + ; is-set to isSetH*RP + ; ·DistR+ to ·DistR+H* + ; ·DistL+ to ·DistL+H*) + + open RingStr (snd ℤ₂[X]) + renaming + ( 0r to 0Z₂X + ; 1r to 1Z₂X + ; _+_ to _+Z₂X_ + ; -_ to -Z₂X_ + ; _·_ to _·Z₂X_ + ; +Assoc to +Z₂XAssoc + ; +IdR to +Z₂XIdR + ; +Comm to +Z₂XComm + ; ·Assoc to ·Z₂XAssoc + ; is-set to isSetZ₂X + ; ·DistR+ to ·DistR+Z₂ + ; ·DistL+ to ·DistL+Z₂) + + ℤ/2≅HⁿRP∞ : (n : ℕ) → AbGroupEquiv ℤ/2 (coHomGr n ℤ/2 (EM ℤ/2 1)) + ℤ/2≅HⁿRP∞ n = invGroupEquiv (GroupIso→GroupEquiv (HⁿRP∞≅ℤ/2 n)) + + ℤ/2≅HⁿRP∞pres0 : (n : ℕ) → ℤ/2≅HⁿRP∞ n .fst .fst fzero ≡ 0ₕ n + ℤ/2≅HⁿRP∞pres0 n = IsGroupHom.pres1 (ℤ/2≅HⁿRP∞ n .snd) + + ℤ₂[X]→H*[RP∞,ℤ₂]-main : Vec ℕ 1 → ℤ/2 .fst → fst (H*R ℤ/2Ring RP∞) + ℤ₂[X]→H*[RP∞,ℤ₂]-main (n ∷ []) g = base n (ℤ/2≅HⁿRP∞ n .fst .fst g) + + base* = base {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} + {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} + + base-neutral* = + base-neutral {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} + {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} + + ℤ₂[X]→H*[RP∞,ℤ₂]-main-coh₁ : (r : Vec ℕ 1) → + ℤ₂[X]→H*[RP∞,ℤ₂]-main r fzero ≡ neutral + ℤ₂[X]→H*[RP∞,ℤ₂]-main-coh₁ (n ∷ []) = + (λ i → base* n (ℤ/2≅HⁿRP∞pres0 n i)) + ∙ base-neutral n + + ℤ₂[X]→H*[RP∞,ℤ₂]-fun : ℤ₂[x] → fst (H*R ℤ/2Ring RP∞) + ℤ₂[X]→H*[RP∞,ℤ₂]-fun = + DS-Rec-Set.f _ _ _ _ isSetH*RP + neutral + ℤ₂[X]→H*[RP∞,ℤ₂]-main + _add_ + addAssoc + addRid + addComm + ℤ₂[X]→H*[RP∞,ℤ₂]-main-coh₁ + λ {(n ∷ []) → λ a b → base-add n _ _ + ∙ (λ i → base {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} + {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} n + (IsGroupHom.pres· (ℤ/2≅HⁿRP∞ n .snd) a b (~ i)))} + + open import Cubical.Foundations.Equiv.Properties + open import Cubical.Foundations.Transport + + ℤ/2[X]→H*[RP∞,ℤ/2]-pres· : (x y : ℤ₂[x]) + → ℤ₂[X]→H*[RP∞,ℤ₂]-fun (RingStr._·_ (snd ℤ₂[X]) x y) + ≡ (ℤ₂[X]→H*[RP∞,ℤ₂]-fun x cupS ℤ₂[X]→H*[RP∞,ℤ₂]-fun y) + ℤ/2[X]→H*[RP∞,ℤ/2]-pres· = + DS-Ind-Prop.f _ _ _ _ + (λ _ → isPropΠ λ _ → trunc _ _) + (λ _ → refl) + (λ v a → + DS-Ind-Prop.f _ _ _ _ + (λ _ → trunc _ _) + (cong (ℤ₂[X]→H*[RP∞,ℤ₂]-fun) + (0RightAnnihilates ℤ₂[X] (base v a)) + ∙ sym (0RightAnnihilates (H*R ℤ/2Ring RP∞) + (ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v a)))) + (λ w b → main v w a b) + λ {x} {y} p q → cong ℤ₂[X]→H*[RP∞,ℤ₂]-fun (·DistR+Z₂ (base v a) x y) + ∙ cong₂ _add_ p q + ∙ sym (·DistR+H* (ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v a)) + (ℤ₂[X]→H*[RP∞,ℤ₂]-fun x) + (ℤ₂[X]→H*[RP∞,ℤ₂]-fun y))) + λ {x} {y} p q z + → cong (ℤ₂[X]→H*[RP∞,ℤ₂]-fun) + (·DistL+Z₂ x y z) + ∙ cong₂ _add_ (p z) (q z) + where + main-main : (v w : Vec ℕ 1) + → ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v fone ·Z₂X base w fone) + ≡ ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v fone) + cupS (ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base w fone)) + main-main = {!!} + + + main : (v w : Vec ℕ 1) (a b : ℤ/2 .fst) + → ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v a ·Z₂X base w b) + ≡ ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v a) cupS ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base w b) + main (n ∷ []) (m ∷ []) a b = + cong (base* (n +ℕ m)) {!!} + ∙ {!!} + ∙ ? -- cong (base* (n +' m)) (HⁿRP∞≅ℤ/2⌣ n m a b) + + -- ℤ₂[X]→H*[RP∞,ℤ₂] : RingHom ℤ₂[X] (H*R ℤ/2Ring RP∞) + -- fst ℤ₂[X]→H*[RP∞,ℤ₂] = ℤ₂[X]→H*[RP∞,ℤ₂]-fun + -- snd ℤ₂[X]→H*[RP∞,ℤ₂] = makeIsRingHom refl (λ _ _ → refl) ℤ/2[X]→H*[RP∞,ℤ/2]-pres· + + -- open Quotient-FGideal-CommRing-Ring (PolyCommRing ℤ/2CommRing 1) (H*R ℤ/2Ring RP∞) + -- ℤ₂[X]→H*[RP∞,ℤ₂] ( ℤ/2CommRing 1 0 3) + -- (λ { zero → base-neutral _}) + + -- ℤ₂[X]/→H*[RP∞,ℤ₂] : RingHom (CommRing→Ring ℤ₂[X]/) (H*R ℤ/2Ring RP∞) + -- ℤ₂[X]/→H*[RP∞,ℤ₂] = inducedHom + + -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' : (r : ℕ) → coHom r ℤ/2 RP∞ → ℤ/2 .fst + -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' zero x = H⁰[RP∞,ℤ/2]≅ℤ/2 .fst .fst x + -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' one x = H¹[RP∞,ℤ/2]≅ℤ/2 .fst .fst x + -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' two x = H²[RP∞,ℤ/2]≅ℤ/2 .fst .fst x + -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' (suc (suc (suc r))) x = fzero + + + -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun : (r : ℕ) → coHom r ℤ/2 RP∞ → ℤ₂[X]/ .fst + -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun r x = [ base (r ∷ []) (H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' r x) ] + + -- H*[RP∞,ℤ₂]→ℤ₂[X]/ : H*R ℤ/2Ring RP∞ .fst → ℤ₂[X]/ .fst + -- H*[RP∞,ℤ₂]→ℤ₂[X]/ = DS-Rec-Set.f _ _ _ _ squash/ + -- [ neutral ] + -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun + -- (CommRingStr._+_ (snd ℤ₂[X]/)) + -- (CommRingStr.+Assoc (snd ℤ₂[X]/)) + -- (CommRingStr.+IdR (snd ℤ₂[X]/)) + -- (CommRingStr.+Comm (snd ℤ₂[X]/)) + -- (λ { zero → cong [_] (base-neutral (0 ∷ [])) + -- ; one → cong [_] (cong (base (1 ∷ [])) (IsGroupHom.pres1 (snd (H¹[RP∞,ℤ/2]≅ℤ/2))) + -- ∙ base-neutral (1 ∷ [])) + -- ; two → cong [_] (cong (base (2 ∷ [])) (IsGroupHom.pres1 (snd (H²[RP∞,ℤ/2]≅ℤ/2))) + -- ∙ base-neutral (2 ∷ [])) + -- ; (suc (suc (suc r))) → cong [_] (base-neutral _)}) + -- (λ { zero a b → cong [_] (base-add (0 ∷ []) _ _ + -- ∙ cong (base (0 ∷ [])) + -- (sym (IsGroupHom.pres· (snd (H⁰[RP∞,ℤ/2]≅ℤ/2)) a b))) + -- ; one a b → cong [_] (base-add (1 ∷ []) _ _ + -- ∙ cong (base (1 ∷ [])) + -- (sym (IsGroupHom.pres· (snd (H¹[RP∞,ℤ/2]≅ℤ/2)) a b))) + -- ; two a b → cong [_] (base-add (2 ∷ []) _ _ + -- ∙ cong (base (2 ∷ [])) + -- (sym (IsGroupHom.pres· (snd (H²[RP∞,ℤ/2]≅ℤ/2)) a b))) + -- ; (suc (suc (suc r))) a b + -- → cong [_] (cong₂ _add_ refl (base-neutral (3 +ℕ r ∷ [])) + -- ∙ addRid _)}) + + -- ℤ₂[X]/→H*[RP∞,ℤ₂]→ℤ₂[X]/ : (x : _) + -- → H*[RP∞,ℤ₂]→ℤ₂[X]/ (ℤ₂[X]/→H*[RP∞,ℤ₂] .fst x) ≡ x + -- ℤ₂[X]/→H*[RP∞,ℤ₂]→ℤ₂[X]/ = SQ.elimProp (λ _ → squash/ _ _) + -- (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) + -- refl + -- (λ { (zero ∷ []) a → cong [_] (cong (base (zero ∷ [])) + -- (secEq (H⁰[RP∞,ℤ/2]≅ℤ/2 .fst) a)) + -- ; (one ∷ []) a → cong [_] (cong (base (one ∷ [])) + -- (secEq (H¹[RP∞,ℤ/2]≅ℤ/2 .fst) a)) + -- ; (two ∷ []) a → cong [_] (cong (base (two ∷ [])) + -- (secEq (H²[RP∞,ℤ/2]≅ℤ/2 .fst) a)) + -- ; (suc (suc (suc r)) ∷ []) → ℤ/2-elim refl + -- (sym (eq/ _ _ ∣ (λ {zero → base (r ∷ []) fone}) + -- , ((cong₂ _add_ refl (base-neutral _) + -- ∙ addRid _ + -- ∙ λ i → base (+-comm 3 r i ∷ []) fone)) + -- ∙ sym (addRid _) ∣₁))}) + -- (λ p q → cong₂ (CommRingStr._+_ (snd ℤ₂[X]/)) p q)) + + -- H*[RP∞,ℤ₂]→ℤ₂[X]/→H*[RP∞,ℤ₂] : (x : H*R ℤ/2Ring RP∞ .fst) + -- → ℤ₂[X]/→H*[RP∞,ℤ₂] .fst (H*[RP∞,ℤ₂]→ℤ₂[X]/ x) ≡ x + -- H*[RP∞,ℤ₂]→ℤ₂[X]/→H*[RP∞,ℤ₂] = DS-Ind-Prop.f _ _ _ _ + -- (λ _ → trunc _ _) + -- refl + -- (λ { zero x → cong (base zero) (retEq (H⁰[RP∞,ℤ/2]≅ℤ/2 .fst) x) + -- ; one x → cong (base one) (retEq (H¹[RP∞,ℤ/2]≅ℤ/2 .fst) x) + -- ; two x → cong (base two) (retEq (H²[RP∞,ℤ/2]≅ℤ/2 .fst) x) + -- ; (suc (suc (suc r))) x → cong (base (3 +ℕ r)) + -- (isContr→isProp (isContrH³⁺ⁿ[RP∞,G] r ℤ/2) _ _)}) + -- λ {x} {y} p q + -- → (IsRingHom.pres+ (ℤ₂[X]/→H*[RP∞,ℤ₂] .snd) _ _ ∙ cong₂ _add_ p q) + + -- ℤ₂[X]/≅H*[RP∞,ℤ₂] : RingEquiv (CommRing→Ring ℤ₂[X]/) (H*R ℤ/2Ring RP∞) + -- fst ℤ₂[X]/≅H*[RP∞,ℤ₂] = + -- isoToEquiv (iso (ℤ₂[X]/→H*[RP∞,ℤ₂] .fst) H*[RP∞,ℤ₂]→ℤ₂[X]/ + -- H*[RP∞,ℤ₂]→ℤ₂[X]/→H*[RP∞,ℤ₂] ℤ₂[X]/→H*[RP∞,ℤ₂]→ℤ₂[X]/) + -- snd ℤ₂[X]/≅H*[RP∞,ℤ₂] = ℤ₂[X]/→H*[RP∞,ℤ₂] .snd + + -- H*[RP∞,ℤ₂]≅ℤ₂[X]/ : RingEquiv (H*R ℤ/2Ring RP∞) (CommRing→Ring ℤ₂[X]/) + -- H*[RP∞,ℤ₂]≅ℤ₂[X]/ = RingEquivs.invRingEquiv ℤ₂[X]/≅H*[RP∞,ℤ₂] diff --git a/Cubical/Data/Bool/Properties.agda b/Cubical/Data/Bool/Properties.agda index 56282dcfae..5570fabbd6 100644 --- a/Cubical/Data/Bool/Properties.agda +++ b/Cubical/Data/Bool/Properties.agda @@ -409,3 +409,23 @@ Iso-⊤⊎⊤-Bool .rightInv false = refl separatedBool : Separated Bool separatedBool = Discrete→Separated _≟_ + + +Bool→Bool→∙Bool : Bool → (Bool , true) →∙ (Bool , true) +Bool→Bool→∙Bool false = idfun∙ _ +Bool→Bool→∙Bool true = const∙ _ _ + +Iso-Bool→∙Bool-Bool : Iso ((Bool , true) →∙ (Bool , true)) Bool +Iso.fun Iso-Bool→∙Bool-Bool f = fst f false +Iso.inv Iso-Bool→∙Bool-Bool = Bool→Bool→∙Bool +Iso.rightInv Iso-Bool→∙Bool-Bool false = refl +Iso.rightInv Iso-Bool→∙Bool-Bool true = refl +Iso.leftInv Iso-Bool→∙Bool-Bool f = Σ≡Prop (λ _ → isSetBool _ _) (help _ refl) + where + help : (x : Bool) → fst f false ≡ x + → Bool→Bool→∙Bool (fst f false) .fst ≡ f .fst + help false p = funExt + λ { false → (λ j → Bool→Bool→∙Bool (p j) .fst false) ∙ sym p + ; true → (λ j → Bool→Bool→∙Bool (p j) .fst true) ∙ sym (snd f)} + help true p = (λ j → Bool→Bool→∙Bool (p j) .fst) + ∙ funExt λ { false → sym p ; true → sym (snd f)} diff --git a/Cubical/Data/Nat/Base.agda b/Cubical/Data/Nat/Base.agda index 87ba8e68a1..9f424f947b 100644 --- a/Cubical/Data/Nat/Base.agda +++ b/Cubical/Data/Nat/Base.agda @@ -12,6 +12,7 @@ open import Cubical.Data.Bool.Base open import Cubical.Data.Sum.Base hiding (elim) open import Cubical.Data.Empty.Base hiding (elim) open import Cubical.Data.Unit.Base +open import Cubical.Data.Sigma.Base predℕ : ℕ → ℕ predℕ zero = zero @@ -77,3 +78,13 @@ isZero (suc n) = false _^_ : ℕ → ℕ → ℕ m ^ 0 = 1 m ^ (suc n) = m · m ^ n + + +-- Iterated product +_ˣ_ : ∀ {ℓ} (A : ℕ → Type ℓ) (n : ℕ) → Type ℓ +A ˣ zero = A zero +A ˣ suc n = (A ˣ n) × A (suc n) + +0ˣ : ∀ {ℓ} (A : ℕ → Type ℓ) (0A : (n : ℕ) → A n) → (n : ℕ) → A ˣ n +0ˣ A 0A zero = 0A zero +0ˣ A 0A (suc n) = (0ˣ A 0A n) , (0A (suc n)) diff --git a/Cubical/Foundations/Pointed/Homogeneous.agda b/Cubical/Foundations/Pointed/Homogeneous.agda index 447cac30b0..1e103b664a 100644 --- a/Cubical/Foundations/Pointed/Homogeneous.agda +++ b/Cubical/Foundations/Pointed/Homogeneous.agda @@ -57,6 +57,14 @@ isHomogeneous {ℓ} (A , x) = ∀ y → Path (Pointed ℓ) (A , x) (A , y) }) (sym (h (pt B∙)) ∙ h ((sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) +→∙HomogeneousPathP : + ∀ {ℓ ℓ'} {A∙ A∙' : Pointed ℓ} {B∙ B∙' : Pointed ℓ'} + {f∙ : A∙ →∙ B∙} {g∙ : A∙' →∙ B∙'} (p : A∙ ≡ A∙') (q : B∙ ≡ B∙') + (h : isHomogeneous B∙') + → PathP (λ i → fst (p i) → fst (q i)) (f∙ .fst) (g∙ .fst) + → PathP (λ i → p i →∙ q i) f∙ g∙ +→∙HomogeneousPathP p q h r = toPathP (→∙Homogeneous≡ h (fromPathP r)) + →∙Homogeneous≡Path : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} (h : isHomogeneous B∙) → (p q : f∙ ≡ g∙) → cong fst p ≡ cong fst q → p ≡ q →∙Homogeneous≡Path {A∙ = A∙@(A , a₀)} {B∙@(B , b)} {f∙@(f , f₀)} {g∙@(g , g₀)} h p q r = From 5e8b305aec47a6cac3167a2262d0874de6884643 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 23 Jan 2024 14:17:35 +0100 Subject: [PATCH 14/73] done? --- Cubical/Algebra/AbGroup/Instances/IntMod.agda | 6 +- Cubical/Algebra/AbGroup/Properties.agda | 4 +- Cubical/Algebra/CommRing/Instances/Int.agda | 4 +- .../EilenbergMacLane/Groups/KleinBottle.agda | 2 +- .../EilenbergMacLane/Groups/RP2.agda | 85 +++---- .../EilenbergMacLane/Rings/RP2.agda | 100 ++------ .../EilenbergMacLane/Rings/RP2wedgeS1.agda | 2 +- .../EilenbergMacLane/Rings/RPinf.agda | 225 +++++------------- .../Cohomology/EilenbergMacLane/Rings/Sn.agda | 62 +---- .../ComputationalSyntheticCohomology.agda | 56 +---- .../CohomologyRings/KleinBottle.agda | 2 +- Cubical/ZCohomology/CohomologyRings/RP2.agda | 2 +- 12 files changed, 157 insertions(+), 393 deletions(-) diff --git a/Cubical/Algebra/AbGroup/Instances/IntMod.agda b/Cubical/Algebra/AbGroup/Instances/IntMod.agda index cb6bfcf652..fb3daf8e78 100644 --- a/Cubical/Algebra/AbGroup/Instances/IntMod.agda +++ b/Cubical/Algebra/AbGroup/Instances/IntMod.agda @@ -19,7 +19,7 @@ open import Cubical.Data.Fin open import Cubical.Data.Fin.Arithmetic open import Cubical.Data.Sigma -open import Cubical.HITs.SetQuotients as SQ renaming ([_] to ⟦_⟧) +open import Cubical.HITs.SetQuotients as SQ open import Cubical.HITs.PropositionalTruncation as PT ℤAbGroup/_ : ℕ → AbGroup ℓ-zero @@ -34,7 +34,7 @@ open import Cubical.HITs.PropositionalTruncation as PT ℤ/2 : AbGroup ℓ-zero ℤ/2 = ℤAbGroup/ 2 -ℤ/2[2]≅ℤ/2 : AbGroupIso (ℤ/2 [ 2 ]) ℤ/2 +ℤ/2[2]≅ℤ/2 : AbGroupIso (ℤ/2 [ 2 ]ₜ) ℤ/2 Iso.fun (fst ℤ/2[2]≅ℤ/2) = fst Iso.inv (fst ℤ/2[2]≅ℤ/2) x = x , cong (x +ₘ_) (+ₘ-rUnit x) ∙ x+x x where @@ -59,7 +59,7 @@ Iso.fun (fst ℤ/2/2≅ℤ/2) = (λ p → ⊥.rec (snotz (sym (cong fst p)))) λ p → ⊥.rec (snotz (sym (cong fst p)))))) λ _ → refl) -Iso.inv (fst ℤ/2/2≅ℤ/2) = ⟦_⟧ +Iso.inv (fst ℤ/2/2≅ℤ/2) = [_] Iso.rightInv (fst ℤ/2/2≅ℤ/2) _ = refl Iso.leftInv (fst ℤ/2/2≅ℤ/2) = SQ.elimProp (λ _ → squash/ _ _) λ _ → refl diff --git a/Cubical/Algebra/AbGroup/Properties.agda b/Cubical/Algebra/AbGroup/Properties.agda index 618f5f7025..a80c06397f 100644 --- a/Cubical/Algebra/AbGroup/Properties.agda +++ b/Cubical/Algebra/AbGroup/Properties.agda @@ -112,8 +112,8 @@ G /^ n = λ a b → cong [_] (AbGroupStr.+Comm (snd G) a b)) -- Torsion subgrous -_[_] : (G : AbGroup ℓ) (n : ℕ) → AbGroup ℓ -G [ n ] = +_[_]ₜ : (G : AbGroup ℓ) (n : ℕ) → AbGroup ℓ +G [ n ]ₜ = Group→AbGroup (Subgroup→Group (AbGroup→Group G) (kerSubgroup (multₗHom G (pos n)))) λ {(x , p) (y , q) → Σ≡Prop (λ _ → isPropIsInKer (multₗHom G (pos n)) _) diff --git a/Cubical/Algebra/CommRing/Instances/Int.agda b/Cubical/Algebra/CommRing/Instances/Int.agda index b5a9e710a9..8ff91cd6ae 100644 --- a/Cubical/Algebra/CommRing/Instances/Int.agda +++ b/Cubical/Algebra/CommRing/Instances/Int.agda @@ -35,8 +35,8 @@ isCommRing (snd ℤCommRing) = isCommRingℤ ℤAbGroup : AbGroup ℓ-zero ℤAbGroup = Group→AbGroup ℤGroup +Comm -ℤTorsion : (n : ℕ) → isContr (fst (ℤAbGroup [ (suc n) ])) -fst (ℤTorsion n) = AbGroupStr.0g (snd (ℤAbGroup [ (suc n) ])) +ℤTorsion : (n : ℕ) → isContr (fst (ℤAbGroup [ (suc n) ]ₜ)) +fst (ℤTorsion n) = AbGroupStr.0g (snd (ℤAbGroup [ (suc n) ]ₜ)) snd (ℤTorsion n) (a , p) = Σ≡Prop (λ _ → isSetℤ _ _) (sym (help a (ℤ·≡· (pos (suc n)) a ∙ p))) where diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda index 0e9f39f8ad..ca2190d186 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda @@ -414,7 +414,7 @@ snd (H¹[K²,G]≅G×H¹[RP²,G] G) = lem = isSet× (is-set (snd G)) squash₂ H¹[K²,G]≅G×G[2] : (G : AbGroup ℓ) - → AbGroupEquiv (coHomGr 1 G KleinBottle) (AbDirProd G (G [ 2 ])) + → AbGroupEquiv (coHomGr 1 G KleinBottle) (AbDirProd G (G [ 2 ]ₜ)) H¹[K²,G]≅G×G[2] G = compGroupEquiv (H¹[K²,G]≅G×H¹[RP²,G] G) (GroupEquivDirProd idGroupEquiv (H¹[RP²,G]≅G[2] G)) diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda index 8f79b44abe..274b6cd5f5 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/RP2.agda @@ -43,7 +43,7 @@ open import Cubical.HITs.SetTruncation as ST open import Cubical.HITs.Truncation as TR open import Cubical.HITs.EilenbergMacLane1 as EM1 open import Cubical.HITs.RPn -open import Cubical.HITs.SetQuotients as SQ renaming ([_] to ⟦_⟧) +open import Cubical.HITs.SetQuotients as SQ open AbGroupStr @@ -259,19 +259,19 @@ H⁰[RP²,G]≅G G = ----- H¹(RP², G) ------ module _ {ℓ : Level} (G : AbGroup ℓ) where private - G[2]path : (x : (G [ 2 ]) .fst) → _+_ (snd G) (fst x) (fst x) ≡ 0g (snd G) + G[2]path : (x : (G [ 2 ]ₜ) .fst) → _+_ (snd G) (fst x) (fst x) ≡ 0g (snd G) G[2]path (x , p) = cong (_+_ (snd G) x) (sym (+IdR (snd G) x)) ∙ p - G[2]path' : (x : (G [ 2 ]) .fst) → fst x ≡ -_ (snd G) (fst x) + G[2]path' : (x : (G [ 2 ]ₜ) .fst) → fst x ≡ -_ (snd G) (fst x) G[2]path' x = (sym (+IdR (snd G) _) ∙ cong (_+_ (snd G) (fst x)) (sym (+InvR (snd G) (fst x))) ∙ (+Assoc (snd G) _ _ _)) ∙∙ cong (λ w → _+_ (snd G) w (-_ (snd G) (fst x))) (G[2]path x) ∙∙ +IdL (snd G) _ - H¹[RP²,G]→G[2] : coHom 1 G RP² → (G [ 2 ]) .fst + H¹[RP²,G]→G[2] : coHom 1 G RP² → (G [ 2 ]ₜ) .fst H¹[RP²,G]→G[2] = - ST.rec (is-set (snd (G [ 2 ]))) + ST.rec (is-set (snd (G [ 2 ]ₜ))) λ f → ΩEM+1→EM-gen {G = G} 0 (f point) (cong f line) , cong (_+_ (snd G) (ΩEM+1→EM-gen 0 (f point) (cong f line))) (+IdR (snd G) _) @@ -281,7 +281,7 @@ module _ {ℓ : Level} (G : AbGroup ℓ) where ∙ cong (cong f) (cong (line ∙_) square ∙ rCancel line) ) ∙ ΩEM+1→EM-gen-refl 0 (f point) - G[2]→H¹[RP²,G] : (G [ 2 ]) .fst → coHom 1 G RP² + G[2]→H¹[RP²,G] : (G [ 2 ]ₜ) .fst → coHom 1 G RP² G[2]→H¹[RP²,G] g = ∣ (λ { point → embase ; (line i) → emloop (fst g) i @@ -290,7 +290,7 @@ module _ {ℓ : Level} (G : AbGroup ℓ) where ∙ emloop-inv (AbGroup→Group G) (fst g)) i j }) ∣₂ private - G[2]≅H¹[RP²,G] : AbGroupEquiv (G [ 2 ]) (coHomGr 1 G RP²) + G[2]≅H¹[RP²,G] : AbGroupEquiv (G [ 2 ]ₜ) (coHomGr 1 G RP²) fst G[2]≅H¹[RP²,G] = isoToEquiv (invIso is) where is : Iso _ _ @@ -314,7 +314,7 @@ module _ {ℓ : Level} (G : AbGroup ℓ) where ∙ sym (cong₂+₁ (emloop (fst x)) (emloop (fst y))))))) - H¹[RP²,G]≅G[2] : AbGroupEquiv (coHomGr 1 G RP²) (G [ 2 ]) + H¹[RP²,G]≅G[2] : AbGroupEquiv (coHomGr 1 G RP²) (G [ 2 ]ₜ) H¹[RP²,G]≅G[2] = invGroupEquiv G[2]≅H¹[RP²,G] ----- H²(RP², G) ------ @@ -333,7 +333,6 @@ module _ (G : AbGroup ℓ) where (λ i → ΩEM+1→EM-sym (suc zero) refl ∙ rUnit (sym (ΩEM+1→EM-sym (suc zero) refl)) (~ i)) ∙ rCancel _ - G/2-ord2 : (x : fst (G /^ 2)) → (- snd (G /^ 2)) x ≡ x G/2-ord2 = SQ.elimProp (λ _ → SQ.squash/ _ _) λ a → eq/ _ _ ∣ snd G .-_ a @@ -348,28 +347,30 @@ module _ (G : AbGroup ℓ) where (uncurry (TR.elim (λ _ → isOfHLevelPlus {n = 2 +ℕ n} 2 (isOfHLevelPlus' {n = n} 2 (isSetΠ λ _ → squash₂))) (EM-raw'-trivElim G (suc n) - (λ _ → isOfHLevelΠ (suc (suc n)) λ _ → isOfHLevelPlus' {n = n} 2 squash₂) ∣_∣₂))) + (λ _ → isOfHLevelΠ (suc (suc n)) + λ _ → isOfHLevelPlus' {n = n} 2 squash₂) ∣_∣₂))) Iso.inv (killFirstCompIsoGen n) = ST.map (0ₖ (2 +ℕ n) ,_) Iso.rightInv (killFirstCompIsoGen n) = ST.elim (λ _ → isSetPathImplicit) λ _ → refl Iso.leftInv (killFirstCompIsoGen n) = ST.elim (λ _ → isSetPathImplicit) - (uncurry (TR.elim (λ _ → isProp→isOfHLevelSuc (3 +ℕ n) (isPropΠ (λ _ → squash₂ _ _))) + (uncurry (TR.elim (λ _ → isProp→isOfHLevelSuc (3 +ℕ n) + (isPropΠ (λ _ → squash₂ _ _))) (EM-raw'-trivElim G (suc n) (λ _ → isProp→isOfHLevelSuc (suc n) (isPropΠ λ _ → squash₂ _ _ )) λ _ → refl))) - Iso₁ : Iso (coHom 2 G RP²) ∥ (Σ[ p ∈ Ω (EM∙ G 2) .fst ] p ≡ sym p) ∥₂ - Iso₁ = compIso (setTruncIso RP²FunCharacIso) (killFirstCompIsoGen 0) + H²RP²-Iso₁ : Iso (coHom 2 G RP²) ∥ (Σ[ p ∈ Ω (EM∙ G 2) .fst ] p ≡ sym p) ∥₂ + H²RP²-Iso₁ = compIso (setTruncIso RP²FunCharacIso) (killFirstCompIsoGen 0) - Iso₂ : Iso ((Σ[ p ∈ Ω (EM∙ G 2) .fst ] p ≡ sym p)) (Σ[ p ∈ EM G 1 ] p ≡ -ₖ p) - Iso₂ = (Σ-cong-iso (invIso (Iso-EM-ΩEM+1 1)) + H²RP²-Iso₂ : Iso ((Σ[ p ∈ Ω (EM∙ G 2) .fst ] p ≡ sym p)) (Σ[ p ∈ EM G 1 ] p ≡ -ₖ p) + H²RP²-Iso₂ = (Σ-cong-iso (invIso (Iso-EM-ΩEM+1 1)) λ x → compIso (congIso (invIso (Iso-EM-ΩEM+1 1))) (equivToIso (compPathrEquiv (ΩEM+1→EM-sym' x)))) - RP²Fun+ : (p q : (Ω^ 2) (EM∙ G 2) .fst) (x : _) + RP²Fun-pres+ : (p q : (Ω^ 2) (EM∙ G 2) .fst) (x : _) → RP²Fun _ _ p x +[ 2 ]ₖ RP²Fun _ _ q x ≡ RP²Fun _ _ (p ∙ q) x - RP²Fun+ p q point = refl - RP²Fun+ p q (line i) j = 0ₖ 2 - RP²Fun+ p q (square i j) k = + RP²Fun-pres+ p q point = refl + RP²Fun-pres+ p q (line i) j = 0ₖ 2 + RP²Fun-pres+ p q (square i j) k = hcomp (λ r → λ {(i = i0) → lem3 k j r ; (i = i1) → lem2 k (~ r) j ; (j = i0) → p (i ∨ ~ k) r @@ -403,26 +404,26 @@ module _ (G : AbGroup ℓ) where ; (k = i1) → rUnit (λ _ → 0ₖ 2) (i ∧ ~ r) j}) (((sym≡flipSquare p) ∙ flipSquare≡cong-sym p) j k r)) - Iso₂Iso₂' : (p : (Ω^ 2) (EM∙ G 2) .fst) - → ST.map (Iso.fun Iso₂) (Iso.fun Iso₁ ∣ RP²Fun _ _ p ∣₂) + H²RP²-Iso₁,₂-charac : (p : (Ω^ 2) (EM∙ G 2) .fst) + → ST.map (Iso.fun H²RP²-Iso₂) (Iso.fun H²RP²-Iso₁ ∣ RP²Fun _ _ p ∣₂) ≡ ∣ embase , (cong (ΩEM+1→EM 1) p) ∣₂ - Iso₂Iso₂' p = cong ∣_∣₂ (ΣPathP (refl + H²RP²-Iso₁,₂-charac p = cong ∣_∣₂ (ΣPathP (refl , (λ i → cong (ΩEM+1→EM {G = G} 1) p ∙ ΩEM+1→EM-sym'-refl i) ∙ sym (rUnit (cong (ΩEM+1→EM 1) p)))) - Iso₂Iso₂+ : (p q : (Ω^ 2) (EM∙ G 2) .fst) - → ST.map (Iso.fun Iso₂) (Iso.fun Iso₁ (∣ RP²Fun _ _ p ∣₂ +ₕ ∣ RP²Fun _ _ q ∣₂)) + H²RP²-Iso₁,₂+ : (p q : (Ω^ 2) (EM∙ G 2) .fst) + → ST.map (Iso.fun H²RP²-Iso₂) (Iso.fun H²RP²-Iso₁ (∣ RP²Fun _ _ p ∣₂ +ₕ ∣ RP²Fun _ _ q ∣₂)) ≡ ∣ embase , cong (ΩEM+1→EM 1) p ∙ cong (ΩEM+1→EM 1) q ∣₂ - Iso₂Iso₂+ p q = cong (ST.map (Iso.fun Iso₂) ∘ Iso.fun Iso₁) - (cong ∣_∣₂ (funExt (RP²Fun+ p q))) - ∙ Iso₂Iso₂' (p ∙ q) + H²RP²-Iso₁,₂+ p q = cong (ST.map (Iso.fun H²RP²-Iso₂) ∘ Iso.fun H²RP²-Iso₁) + (cong ∣_∣₂ (funExt (RP²Fun-pres+ p q))) + ∙ H²RP²-Iso₁,₂-charac (p ∙ q) ∙ cong ∣_∣₂ (ΣPathP (refl , cong-∙ (ΩEM+1→EM 1) p q)) - Iso₃→ : (p : EM G 1) → (p ≡ -[ 1 ]ₖ p) → fst (G /^ 2) - Iso₃→ = EM1.elimSet (AbGroup→Group G) (λ _ → isSetΠ λ _ → squash/) - (λ p → ⟦ ΩEM+1→EM 0 p ⟧) + H²RP²-Iso₃→ : (p : EM G 1) → (p ≡ -[ 1 ]ₖ p) → fst (G /^ 2) + H²RP²-Iso₃→ = EM1.elimSet (AbGroup→Group G) (λ _ → isSetΠ λ _ → squash/) + (λ p → [ ΩEM+1→EM 0 p ]) λ g → toPathP (funExt λ q - → cong ⟦_⟧ (transportRefl _ + → cong [_] (transportRefl _ ∙ cong (ΩEM+1→EM 0) ((λ j → transp (λ i → Path (EM G 1) (emloop g (~ i ∧ ~ j)) (emloop g (i ∨ j))) j (doubleCompPath-filler (emloop g) q (emloop g) j)) @@ -439,9 +440,9 @@ module _ (G : AbGroup ℓ) where (sym (+InvR (snd G) (ΩEM+1→EM 0 q))) ∙ +Assoc (snd G) _ _ _ ∣₁) - Iso₃ : Iso ∥ (Σ[ p ∈ EM G 1 ] p ≡ -ₖ p) ∥₂ (fst (G /^ 2)) - Iso.fun Iso₃ = ST.rec squash/ (uncurry Iso₃→) - Iso.inv Iso₃ = + H²RP²-Iso₃ : Iso ∥ (Σ[ p ∈ EM G 1 ] p ≡ -ₖ p) ∥₂ (fst (G /^ 2)) + Iso.fun H²RP²-Iso₃ = ST.rec squash/ (uncurry H²RP²-Iso₃→) + Iso.inv H²RP²-Iso₃ = SQ.elim (λ _ → squash₂) (λ a → ∣ embase , emloop a ∣₂) λ a b → PT.rec (squash₂ _ _) (uncurry λ x y → cong ∣_∣₂ (ΣPathP ((emloop x) @@ -463,9 +464,9 @@ module _ (G : AbGroup ℓ) where ∙ (sym (+Assoc (snd G) _ _ _) ∙ cong (_+_ (snd G) b) (+InvL (snd G) a)) ∙ +IdR (snd G) b))))) - Iso.rightInv Iso₃ = SQ.elimProp (λ _ → squash/ _ _) - λ a → cong ⟦_⟧ (Iso.leftInv (Iso-EM-ΩEM+1 0) a) - Iso.leftInv Iso₃ = ST.elim (λ _ → isSetPathImplicit) + Iso.rightInv H²RP²-Iso₃ = SQ.elimProp (λ _ → squash/ _ _) + λ a → cong [_] (Iso.leftInv (Iso-EM-ΩEM+1 0) a) + Iso.leftInv H²RP²-Iso₃ = ST.elim (λ _ → isSetPathImplicit) (uncurry (EM1.elimProp (AbGroup→Group G) (λ _ → isPropΠ λ _ → squash₂ _ _) λ p → cong ∣_∣₂ (ΣPathP (refl , (Iso.rightInv (Iso-EM-ΩEM+1 0) p))))) @@ -473,17 +474,17 @@ module _ (G : AbGroup ℓ) where fst H²[RP²,G]≅G/2 = isoToEquiv is where is : Iso _ _ - is = compIso (compIso Iso₁ (setTruncIso Iso₂)) Iso₃ + is = compIso (compIso H²RP²-Iso₁ (setTruncIso H²RP²-Iso₂)) H²RP²-Iso₃ snd H²[RP²,G]≅G/2 = makeIsGroupHom (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash/ _ _) (RP²Conn.elim₂ (isConnectedEM 2) (λ _ → isPropΠ λ _ → squash/ _ _) (0ₖ 2) λ p → RP²Conn.elim₂ (isConnectedEM 2) (λ _ → squash/ _ _) - (0ₖ 2) λ q → cong (Iso.fun Iso₃) (Iso₂Iso₂+ p q) - ∙ cong ⟦_⟧ (ΩEM+1→EM-hom 0 (cong (ΩEM+1→EM 1) p) + (0ₖ 2) λ q → cong (Iso.fun H²RP²-Iso₃) (H²RP²-Iso₁,₂+ p q) + ∙ cong [_] (ΩEM+1→EM-hom 0 (cong (ΩEM+1→EM 1) p) (cong (ΩEM+1→EM 1) q)) ∙ cong₂ ((G /^ 2) .snd ._+_) - (cong (Iso.fun Iso₃) (sym (Iso₂Iso₂' p))) - (cong (Iso.fun Iso₃) (sym (Iso₂Iso₂' q))))) + (cong (Iso.fun H²RP²-Iso₃) (sym (H²RP²-Iso₁,₂-charac p))) + (cong (Iso.fun H²RP²-Iso₃) (sym (H²RP²-Iso₁,₂-charac q))))) H³⁺ⁿ[RP²,G]≅0 : (n : ℕ) → AbGroupEquiv (coHomGr (3 +ℕ n) G RP²) (trivialAbGroup {ℓ}) fst (H³⁺ⁿ[RP²,G]≅0 n) = isContr→≃Unit* (0ₕ _ diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda index c594f895f9..db570dcc10 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda @@ -2,92 +2,60 @@ module Cubical.Cohomology.EilenbergMacLane.Rings.RP2 where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Path -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Univalence open import Cubical.Foundations.GroupoidLaws -open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2 - open import Cubical.Cohomology.EilenbergMacLane.Base -open import Cubical.Cohomology.EilenbergMacLane.Groups.Connected open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2 -open import Cubical.Cohomology.EilenbergMacLane.Groups.Sn -open import Cubical.Cohomology.EilenbergMacLane.Groups.Wedge -open import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle open import Cubical.Cohomology.EilenbergMacLane.CupProduct open import Cubical.Cohomology.EilenbergMacLane.Rings.KleinBottle hiding (α↦1) open import Cubical.Cohomology.EilenbergMacLane.Rings.RP2wedgeS1 +open import Cubical.Cohomology.EilenbergMacLane.RingStructure +open import Cubical.Homotopy.EilenbergMacLane.Base open import Cubical.Homotopy.EilenbergMacLane.GroupStructure -open import Cubical.Homotopy.EilenbergMacLane.Order2 open import Cubical.Homotopy.EilenbergMacLane.Properties -open import Cubical.Homotopy.EilenbergMacLane.Base open import Cubical.Homotopy.EilenbergMacLane.CupProduct open import Cubical.Homotopy.EilenbergMacLane.CupProductTensor renaming (_⌣ₖ_ to _⌣ₖ⊗_ ; ⌣ₖ-0ₖ to ⌣ₖ-0ₖ⊗ ; 0ₖ-⌣ₖ to 0ₖ-⌣ₖ⊗) open import Cubical.Homotopy.Loopspace -open import Cubical.Homotopy.Connected -open import Cubical.Cohomology.EilenbergMacLane.RingStructure -open import Cubical.Cohomology.EilenbergMacLane.Rings.Z2-properties open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) -open import Cubical.Data.Nat.Order open import Cubical.Data.Unit open import Cubical.Data.Fin -open import Cubical.Data.Sum open import Cubical.Data.Fin.Arithmetic -open import Cubical.Data.Sigma open import Cubical.Data.Vec -open import Cubical.Data.Empty as ⊥ open import Cubical.Data.FinData renaming (Fin to FinI) -open import Cubical.HITs.KleinBottle renaming (rec to KleinFun) open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.SetTruncation as ST open import Cubical.HITs.Truncation as TR open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _sq/_) open import Cubical.HITs.EilenbergMacLane1 hiding (elimProp ; elimSet) open import Cubical.HITs.Susp -open import Cubical.HITs.Pushout -open import Cubical.HITs.S1 renaming (rec to S¹Fun) -open import Cubical.HITs.Sn open import Cubical.HITs.RPn -open import Cubical.HITs.Wedge -open import Cubical.Algebra.GradedRing.DirectSumHIT -open import Cubical.Algebra.Group +open import Cubical.Algebra.GradedRing.DirectSumHIT open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.IntMod -open import Cubical.Algebra.AbGroup hiding (_[_]) -open import Cubical.Algebra.AbGroup.Instances.Unit +open import Cubical.Algebra.AbGroup open import Cubical.Algebra.DirectSum.DirectSumHIT.Base open import Cubical.Algebra.Ring open import Cubical.Algebra.AbGroup.TensorProduct -open import Cubical.Algebra.Monoid - open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.Quotient open import Cubical.Algebra.CommRing.Instances.IntMod open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly-Quotient -open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly-notationZ2 open import Cubical.Algebra.AbGroup.Instances.IntMod - open import Cubical.Algebra.Monoid.Instances.Nat -open Iso open PlusBis open RingTheory -open AbGroupStr - private α-raw : RP² → EM ℤ/2 1 @@ -98,8 +66,6 @@ private α↦1 : H¹[RP²,ℤ/2]≅ℤ/2 .fst .fst α ≡ 1 α↦1 = refl - w = RP²∨S¹→K₂-fun - γ-raw : RP² → EM ℤ/2 2 γ-raw = RP²Fun (0ₖ 2) refl (Iso.inv Iso-Ω²K₂-ℤ/2 1) @@ -199,10 +165,8 @@ private (cong (EM→ΩEM+1 1) (EM→ΩEM+1 0 1)) (EM→ΩEM+1-0ₖ 1) k i j))) -α²=γ : _⌣_ {G'' = ℤ/2Ring} α α ≡ γ -α²=γ = cong ∣_∣₂ (funExt α²-raw') - - +α²≡γ : _⌣_ {G'' = ℤ/2Ring} α α ≡ γ +α²≡γ = cong ∣_∣₂ (funExt α²-raw') -- open import Cubical.Agebra.Monoids.Instance.NatVec @@ -258,7 +222,7 @@ module _ {ℓ : Level} (n : ℕ) where γ* = base 2 γ α*²≡γ* : α* cupS α* ≡ γ* - α*²≡γ* = cong (base 2) α²=γ + α*²≡γ* = cong (base 2) α²≡γ α-isGen : base 1 (invEq (fst (H¹[RP²,ℤ/2]≅ℤ/2)) fone) ≡ α* α-isGen = cong (base 1) @@ -314,47 +278,11 @@ module _ {ℓ : Level} (n : ℕ) where ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₁ ℤ₂[X]→H*[RP²,ℤ₂]-main-coh₂ - open import Cubical.Foundations.Equiv.Properties - open import Cubical.Foundations.Transport -{- - Hⁿ[Sⁿ,G]≅G-pres⌣ : (n : ℕ) (G : Ring ℓ) (g : fst G) - (x : coHom (suc n) (Ring→AbGroup G) (S₊ (suc n))) - → Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n .fst .fst - (_⌣_ {G'' = G} {n = 0} {suc n} ∣ (λ _ → g) ∣₂ x) - ≡ RingStr._·_ (snd G) g (Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n .fst .fst x) - Hⁿ[Sⁿ,G]≅G-pres⌣ zero G g = - ST.elim (λ _ → isOfHLevelPath 2 (RingStr.is-set (snd G)) _ _) - (S¹-connElim (isConnectedEM 1) - (λ _ → RingStr.is-set (snd G) _ _) - embase - λ p → cong (H¹S¹→G (Ring→AbGroup G)) - (cong (_⌣_ {G'' = G} {n = 0} {1} ∣ (λ _ → g) ∣₂) - (cong (∣_∣₂ ∘ S¹Fun embase) - (sym (Iso.rightInv (Iso-EM-ΩEM+1 0) p)))) - ∙∙ help (ΩEM+1→EM 0 p) - ∙∙ cong (RingStr._·_ (snd G) g) (λ _ → ΩEM+1→EM 0 p)) - where - help : (h : fst G) - → H¹S¹→G (Ring→AbGroup G) - (_⌣_ {G'' = G} {n = 0} {1} ∣ (λ _ → g) ∣₂ - ∣ S¹Fun embase (emloop h) ∣₂) - ≡ RingStr._·_ (snd G) g h - help h = Iso.leftInv (Iso-EM-ΩEM+1 0) (RingStr._·_ (snd G) g h) - Hⁿ[Sⁿ,G]≅G-pres⌣ (suc n) G g = - ST.elim (λ _ → isOfHLevelPath 2 (RingStr.is-set (snd G)) _ _) - (Sⁿ-connElim n (isConnectedSubtr 2 (suc n) - (subst (λ k → isConnected (suc k) (EM (Ring→AbGroup G) (suc (suc n)))) - (+-comm 2 n) (isConnectedEM (2 +ℕ n)))) - (λ _ → RingStr.is-set (snd G) _ _) ∣ north ∣ₕ - λ f → cong (fst (Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n) .fst) - (cong ∣_∣₂ (funExt (λ x → sym (ΩEM+1→EM-distr⌣ₖ0n {G'' = G} (suc n) g _)))) - ∙ Hⁿ[Sⁿ,G]≅G-pres⌣ n G g _) --} - - isContrH³⁺ⁿ[RP²,G] : ∀ {ℓ} (n : ℕ) (G : AbGroup ℓ) → isContr (coHom (3 +ℕ n) G RP²) - isContrH³⁺ⁿ[RP²,G] n G = - isOfHLevelRetractFromIso 0 - (equivToIso (fst (H³⁺ⁿ[RP²,G]≅0 G n))) isContrUnit* + private + isContrH³⁺ⁿ[RP²,G] : ∀ {ℓ} (n : ℕ) (G : AbGroup ℓ) → isContr (coHom (3 +ℕ n) G RP²) + isContrH³⁺ⁿ[RP²,G] n G = + isOfHLevelRetractFromIso 0 + (equivToIso (fst (H³⁺ⁿ[RP²,G]≅0 G n))) isContrUnit* ℤ/2[X]→H*[RP²,ℤ/2]-pres· : (x y : ℤ₂[x]) → ℤ₂[X]→H*[RP²,ℤ₂]-fun (RingStr._·_ (snd ℤ₂[X]) x y) @@ -448,7 +376,6 @@ module _ {ℓ : Level} (n : ℕ) where H*[RP²,ℤ₂]→ℤ₂[X]/-fun' two x = H²[RP²,ℤ/2]≅ℤ/2 .fst .fst x H*[RP²,ℤ₂]→ℤ₂[X]/-fun' (suc (suc (suc r))) x = fzero - H*[RP²,ℤ₂]→ℤ₂[X]/-fun : (r : ℕ) → coHom r ℤ/2 RP² → ℤ₂[X]/ .fst H*[RP²,ℤ₂]→ℤ₂[X]/-fun r x = [ base (r ∷ []) (H*[RP²,ℤ₂]→ℤ₂[X]/-fun' r x) ] @@ -514,7 +441,8 @@ module _ {ℓ : Level} (n : ℕ) where ℤ₂[X]/≅H*[RP²,ℤ₂] : RingEquiv (CommRing→Ring ℤ₂[X]/) (H*R ℤ/2Ring RP²) fst ℤ₂[X]/≅H*[RP²,ℤ₂] = isoToEquiv (iso (ℤ₂[X]/→H*[RP²,ℤ₂] .fst) H*[RP²,ℤ₂]→ℤ₂[X]/ - H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/) + H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] + ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/) snd ℤ₂[X]/≅H*[RP²,ℤ₂] = ℤ₂[X]/→H*[RP²,ℤ₂] .snd H*[RP²,ℤ₂]≅ℤ₂[X]/ : RingEquiv (H*R ℤ/2Ring RP²) (CommRing→Ring ℤ₂[X]/) diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda index 0e632396f1..b4f4150e07 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda @@ -57,7 +57,7 @@ open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.IntMod -open import Cubical.Algebra.AbGroup hiding (_[_]) +open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.Unit open import Cubical.Algebra.AbGroup.Instances.IntMod open import Cubical.Algebra.DirectSum.DirectSumHIT.Base diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda index 12f3287d2c..5be4064ecf 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RPinf.agda @@ -6,92 +6,45 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Path -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Transport -open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2wedgeS1 open import Cubical.Cohomology.EilenbergMacLane.Base -open import Cubical.Cohomology.EilenbergMacLane.Groups.RPinf open import Cubical.Cohomology.EilenbergMacLane.CupProduct +open import Cubical.Cohomology.EilenbergMacLane.Groups.RPinf +open import Cubical.Cohomology.EilenbergMacLane.RingStructure +open import Cubical.Homotopy.EilenbergMacLane.Base open import Cubical.Homotopy.EilenbergMacLane.GroupStructure -open import Cubical.Homotopy.EilenbergMacLane.Order2 open import Cubical.Homotopy.EilenbergMacLane.Properties -open import Cubical.Homotopy.EilenbergMacLane.Base open import Cubical.Homotopy.EilenbergMacLane.CupProduct -open import Cubical.Homotopy.EilenbergMacLane.CupProductTensor - renaming (_⌣ₖ_ to _⌣ₖ⊗_ ; ⌣ₖ-0ₖ to ⌣ₖ-0ₖ⊗ ; 0ₖ-⌣ₖ to 0ₖ-⌣ₖ⊗) -open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected -open import Cubical.Cohomology.EilenbergMacLane.RingStructure -open import Cubical.Cohomology.EilenbergMacLane.Rings.Z2-properties open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) -open import Cubical.Data.Nat.Order -open import Cubical.Data.Unit open import Cubical.Data.Fin -open import Cubical.Data.Sum open import Cubical.Data.Fin.Arithmetic -open import Cubical.Data.Sigma open import Cubical.Data.Vec -open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.FinData renaming (Fin to FinI) -open import Cubical.HITs.KleinBottle renaming (rec to KleinFun) -open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.SetTruncation as ST open import Cubical.HITs.Truncation as TR -open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _sq/_) -open import Cubical.HITs.EilenbergMacLane1 hiding (elimProp ; elimSet) -open import Cubical.HITs.Susp -open import Cubical.HITs.Pushout -open import Cubical.HITs.S1 renaming (rec to S¹Fun) -open import Cubical.HITs.Sn open import Cubical.HITs.RPn -open import Cubical.HITs.Wedge -open import Cubical.Algebra.GradedRing.DirectSumHIT -open import Cubical.Algebra.Group +open import Cubical.Algebra.GradedRing.DirectSumHIT open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.IntMod -open import Cubical.Algebra.AbGroup hiding (_[_]) -open import Cubical.Algebra.AbGroup.Instances.Unit +open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.IntMod - open import Cubical.Algebra.DirectSum.DirectSumHIT.Base open import Cubical.Algebra.Ring -open import Cubical.Algebra.AbGroup.TensorProduct -open import Cubical.Algebra.Monoid - open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.FGIdeal -open import Cubical.Algebra.CommRing.Quotient open import Cubical.Algebra.CommRing.Instances.IntMod open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly -open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly-notationZ2 - open import Cubical.Algebra.Monoid.Instances.Nat -open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties - -open import Cubical.Data.Fin.Arithmetic -open import Cubical.Cohomology.EilenbergMacLane.CupProduct - -open import Cubical.Foundations.Path -open import Cubical.Foundations.Transport -open import Cubical.Homotopy.Connected -open import Cubical.HITs.Truncation as TR - open Iso open PlusBis open RingTheory -open AbGroupStr - eRP∞^ : (n : ℕ) → coHom n ℤ/2 (EM ℤ/2 1) eRP∞^ zero = 1ₕ {G'' = ℤ/2Ring} eRP∞^ (suc n) = @@ -253,18 +206,13 @@ module _ {ℓ : Level} (n : ℕ) where ℤ₂[X]→H*[RP∞,ℤ₂]-main : Vec ℕ 1 → ℤ/2 .fst → fst (H*R ℤ/2Ring RP∞) ℤ₂[X]→H*[RP∞,ℤ₂]-main (n ∷ []) g = base n (ℤ/2≅HⁿRP∞ n .fst .fst g) - base* = base {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} - {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} - - base-neutral* = - base-neutral {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} - {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} - ℤ₂[X]→H*[RP∞,ℤ₂]-main-coh₁ : (r : Vec ℕ 1) → ℤ₂[X]→H*[RP∞,ℤ₂]-main r fzero ≡ neutral ℤ₂[X]→H*[RP∞,ℤ₂]-main-coh₁ (n ∷ []) = - (λ i → base* n (ℤ/2≅HⁿRP∞pres0 n i)) - ∙ base-neutral n + (λ i → base {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} + {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} n (ℤ/2≅HⁿRP∞pres0 n i)) + ∙ base-neutral {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} + {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} n ℤ₂[X]→H*[RP∞,ℤ₂]-fun : ℤ₂[x] → fst (H*R ℤ/2Ring RP∞) ℤ₂[X]→H*[RP∞,ℤ₂]-fun = @@ -309,106 +257,61 @@ module _ {ℓ : Level} (n : ℕ) where (·DistL+Z₂ x y z) ∙ cong₂ _add_ (p z) (q z) where - main-main : (v w : Vec ℕ 1) - → ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v fone ·Z₂X base w fone) - ≡ ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v fone) - cupS (ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base w fone)) - main-main = {!!} - - main : (v w : Vec ℕ 1) (a b : ℤ/2 .fst) → ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v a ·Z₂X base w b) ≡ ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base v a) cupS ℤ₂[X]→H*[RP∞,ℤ₂]-fun (base w b) main (n ∷ []) (m ∷ []) a b = - cong (base* (n +ℕ m)) {!!} - ∙ {!!} - ∙ ? -- cong (base* (n +' m)) (HⁿRP∞≅ℤ/2⌣ n m a b) - - -- ℤ₂[X]→H*[RP∞,ℤ₂] : RingHom ℤ₂[X] (H*R ℤ/2Ring RP∞) - -- fst ℤ₂[X]→H*[RP∞,ℤ₂] = ℤ₂[X]→H*[RP∞,ℤ₂]-fun - -- snd ℤ₂[X]→H*[RP∞,ℤ₂] = makeIsRingHom refl (λ _ _ → refl) ℤ/2[X]→H*[RP∞,ℤ/2]-pres· - - -- open Quotient-FGideal-CommRing-Ring (PolyCommRing ℤ/2CommRing 1) (H*R ℤ/2Ring RP∞) - -- ℤ₂[X]→H*[RP∞,ℤ₂] ( ℤ/2CommRing 1 0 3) - -- (λ { zero → base-neutral _}) - - -- ℤ₂[X]/→H*[RP∞,ℤ₂] : RingHom (CommRing→Ring ℤ₂[X]/) (H*R ℤ/2Ring RP∞) - -- ℤ₂[X]/→H*[RP∞,ℤ₂] = inducedHom - - -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' : (r : ℕ) → coHom r ℤ/2 RP∞ → ℤ/2 .fst - -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' zero x = H⁰[RP∞,ℤ/2]≅ℤ/2 .fst .fst x - -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' one x = H¹[RP∞,ℤ/2]≅ℤ/2 .fst .fst x - -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' two x = H²[RP∞,ℤ/2]≅ℤ/2 .fst .fst x - -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' (suc (suc (suc r))) x = fzero - - - -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun : (r : ℕ) → coHom r ℤ/2 RP∞ → ℤ₂[X]/ .fst - -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun r x = [ base (r ∷ []) (H*[RP∞,ℤ₂]→ℤ₂[X]/-fun' r x) ] - - -- H*[RP∞,ℤ₂]→ℤ₂[X]/ : H*R ℤ/2Ring RP∞ .fst → ℤ₂[X]/ .fst - -- H*[RP∞,ℤ₂]→ℤ₂[X]/ = DS-Rec-Set.f _ _ _ _ squash/ - -- [ neutral ] - -- H*[RP∞,ℤ₂]→ℤ₂[X]/-fun - -- (CommRingStr._+_ (snd ℤ₂[X]/)) - -- (CommRingStr.+Assoc (snd ℤ₂[X]/)) - -- (CommRingStr.+IdR (snd ℤ₂[X]/)) - -- (CommRingStr.+Comm (snd ℤ₂[X]/)) - -- (λ { zero → cong [_] (base-neutral (0 ∷ [])) - -- ; one → cong [_] (cong (base (1 ∷ [])) (IsGroupHom.pres1 (snd (H¹[RP∞,ℤ/2]≅ℤ/2))) - -- ∙ base-neutral (1 ∷ [])) - -- ; two → cong [_] (cong (base (2 ∷ [])) (IsGroupHom.pres1 (snd (H²[RP∞,ℤ/2]≅ℤ/2))) - -- ∙ base-neutral (2 ∷ [])) - -- ; (suc (suc (suc r))) → cong [_] (base-neutral _)}) - -- (λ { zero a b → cong [_] (base-add (0 ∷ []) _ _ - -- ∙ cong (base (0 ∷ [])) - -- (sym (IsGroupHom.pres· (snd (H⁰[RP∞,ℤ/2]≅ℤ/2)) a b))) - -- ; one a b → cong [_] (base-add (1 ∷ []) _ _ - -- ∙ cong (base (1 ∷ [])) - -- (sym (IsGroupHom.pres· (snd (H¹[RP∞,ℤ/2]≅ℤ/2)) a b))) - -- ; two a b → cong [_] (base-add (2 ∷ []) _ _ - -- ∙ cong (base (2 ∷ [])) - -- (sym (IsGroupHom.pres· (snd (H²[RP∞,ℤ/2]≅ℤ/2)) a b))) - -- ; (suc (suc (suc r))) a b - -- → cong [_] (cong₂ _add_ refl (base-neutral (3 +ℕ r ∷ [])) - -- ∙ addRid _)}) - - -- ℤ₂[X]/→H*[RP∞,ℤ₂]→ℤ₂[X]/ : (x : _) - -- → H*[RP∞,ℤ₂]→ℤ₂[X]/ (ℤ₂[X]/→H*[RP∞,ℤ₂] .fst x) ≡ x - -- ℤ₂[X]/→H*[RP∞,ℤ₂]→ℤ₂[X]/ = SQ.elimProp (λ _ → squash/ _ _) - -- (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) - -- refl - -- (λ { (zero ∷ []) a → cong [_] (cong (base (zero ∷ [])) - -- (secEq (H⁰[RP∞,ℤ/2]≅ℤ/2 .fst) a)) - -- ; (one ∷ []) a → cong [_] (cong (base (one ∷ [])) - -- (secEq (H¹[RP∞,ℤ/2]≅ℤ/2 .fst) a)) - -- ; (two ∷ []) a → cong [_] (cong (base (two ∷ [])) - -- (secEq (H²[RP∞,ℤ/2]≅ℤ/2 .fst) a)) - -- ; (suc (suc (suc r)) ∷ []) → ℤ/2-elim refl - -- (sym (eq/ _ _ ∣ (λ {zero → base (r ∷ []) fone}) - -- , ((cong₂ _add_ refl (base-neutral _) - -- ∙ addRid _ - -- ∙ λ i → base (+-comm 3 r i ∷ []) fone)) - -- ∙ sym (addRid _) ∣₁))}) - -- (λ p q → cong₂ (CommRingStr._+_ (snd ℤ₂[X]/)) p q)) - - -- H*[RP∞,ℤ₂]→ℤ₂[X]/→H*[RP∞,ℤ₂] : (x : H*R ℤ/2Ring RP∞ .fst) - -- → ℤ₂[X]/→H*[RP∞,ℤ₂] .fst (H*[RP∞,ℤ₂]→ℤ₂[X]/ x) ≡ x - -- H*[RP∞,ℤ₂]→ℤ₂[X]/→H*[RP∞,ℤ₂] = DS-Ind-Prop.f _ _ _ _ - -- (λ _ → trunc _ _) - -- refl - -- (λ { zero x → cong (base zero) (retEq (H⁰[RP∞,ℤ/2]≅ℤ/2 .fst) x) - -- ; one x → cong (base one) (retEq (H¹[RP∞,ℤ/2]≅ℤ/2 .fst) x) - -- ; two x → cong (base two) (retEq (H²[RP∞,ℤ/2]≅ℤ/2 .fst) x) - -- ; (suc (suc (suc r))) x → cong (base (3 +ℕ r)) - -- (isContr→isProp (isContrH³⁺ⁿ[RP∞,G] r ℤ/2) _ _)}) - -- λ {x} {y} p q - -- → (IsRingHom.pres+ (ℤ₂[X]/→H*[RP∞,ℤ₂] .snd) _ _ ∙ cong₂ _add_ p q) - - -- ℤ₂[X]/≅H*[RP∞,ℤ₂] : RingEquiv (CommRing→Ring ℤ₂[X]/) (H*R ℤ/2Ring RP∞) - -- fst ℤ₂[X]/≅H*[RP∞,ℤ₂] = - -- isoToEquiv (iso (ℤ₂[X]/→H*[RP∞,ℤ₂] .fst) H*[RP∞,ℤ₂]→ℤ₂[X]/ - -- H*[RP∞,ℤ₂]→ℤ₂[X]/→H*[RP∞,ℤ₂] ℤ₂[X]/→H*[RP∞,ℤ₂]→ℤ₂[X]/) - -- snd ℤ₂[X]/≅H*[RP∞,ℤ₂] = ℤ₂[X]/→H*[RP∞,ℤ₂] .snd - - -- H*[RP∞,ℤ₂]≅ℤ₂[X]/ : RingEquiv (H*R ℤ/2Ring RP∞) (CommRing→Ring ℤ₂[X]/) - -- H*[RP∞,ℤ₂]≅ℤ₂[X]/ = RingEquivs.invRingEquiv ℤ₂[X]/≅H*[RP∞,ℤ₂] + (λ i → base {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} + {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} + (+'≡+ n m (~ i)) + (inv (cohomRP∞Iso (+'≡+ n m (~ i))) (a ·ₘ b))) + ∙ λ i → base {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} + {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} (n +' m) + (HⁿRP∞≅ℤ/2⌣ n m a b i) + + ℤ₂[X]→H*[RP∞,ℤ₂] : RingHom ℤ₂[X] (H*R ℤ/2Ring RP∞) + fst ℤ₂[X]→H*[RP∞,ℤ₂] = ℤ₂[X]→H*[RP∞,ℤ₂]-fun + snd ℤ₂[X]→H*[RP∞,ℤ₂] = makeIsRingHom refl (λ _ _ → refl) ℤ/2[X]→H*[RP∞,ℤ/2]-pres· + + H*[RP∞,ℤ₂]→ℤ₂[X] : H*R ℤ/2Ring RP∞ .fst → ℤ₂[X] .fst + H*[RP∞,ℤ₂]→ℤ₂[X] = DS-Rec-Set.f _ _ _ _ trunc + neutral + (λ r x → base (r ∷ []) (invEq (ℤ/2≅HⁿRP∞ r .fst) x)) + (RingStr._+_ (snd ℤ₂[X])) + (RingStr.+Assoc (snd ℤ₂[X])) + (RingStr.+IdR (snd ℤ₂[X])) + (RingStr.+Comm (snd ℤ₂[X])) + (λ r → cong (base (r ∷ [])) + (IsGroupHom.pres1 (invGroupEquiv (ℤ/2≅HⁿRP∞ r) .snd)) + ∙ base-neutral _) + λ r a b → base-add (r ∷ []) _ _ + ∙ cong (base (r ∷ [])) + (sym (IsGroupHom.pres· (invGroupEquiv (ℤ/2≅HⁿRP∞ r) .snd) a b)) + + ℤ₂[X]→H*[RP∞,ℤ₂]→ℤ₂[X] : (x : _) + → H*[RP∞,ℤ₂]→ℤ₂[X] (ℤ₂[X]→H*[RP∞,ℤ₂] .fst x) ≡ x + ℤ₂[X]→H*[RP∞,ℤ₂]→ℤ₂[X] = DS-Ind-Prop.f _ _ _ _ (λ _ → trunc _ _) + refl + (λ{ (x ∷ []) a → cong (base (x ∷ [])) (retEq (ℤ/2≅HⁿRP∞ x .fst) a)}) + (λ p q → cong₂ _add_ p q) + + H*[RP∞,ℤ₂]→ℤ₂[X]→H*[RP∞,ℤ₂] : (x : _) + → ℤ₂[X]→H*[RP∞,ℤ₂] .fst (H*[RP∞,ℤ₂]→ℤ₂[X] x) ≡ x + H*[RP∞,ℤ₂]→ℤ₂[X]→H*[RP∞,ℤ₂] = DS-Ind-Prop.f _ _ _ _ (λ _ → trunc _ _) + refl + (λ{ x a → cong (base {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} + {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd} x) + (secEq (ℤ/2≅HⁿRP∞ x .fst) a)}) + λ p q → cong₂ (_add_ {Idx = ℕ} {P = λ n → coHom n ℤ/2 RP∞} + {AGP = λ n → coHomGr n ℤ/2 RP∞ .snd}) + p q + + + ℤ₂[X]≅H*[RP∞,ℤ₂] : RingEquiv ℤ₂[X] (H*R ℤ/2Ring RP∞) + fst ℤ₂[X]≅H*[RP∞,ℤ₂] = + isoToEquiv (iso (ℤ₂[X]→H*[RP∞,ℤ₂] .fst) H*[RP∞,ℤ₂]→ℤ₂[X] + H*[RP∞,ℤ₂]→ℤ₂[X]→H*[RP∞,ℤ₂] ℤ₂[X]→H*[RP∞,ℤ₂]→ℤ₂[X]) + snd ℤ₂[X]≅H*[RP∞,ℤ₂] = ℤ₂[X]→H*[RP∞,ℤ₂] .snd + + H*[RP∞,ℤ₂]≅ℤ₂[X] : RingEquiv (H*R ℤ/2Ring RP∞) ℤ₂[X] + H*[RP∞,ℤ₂]≅ℤ₂[X] = RingEquivs.invRingEquiv ℤ₂[X]≅H*[RP∞,ℤ₂] diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda index f5d2bc9a2e..5c3facf6a1 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda @@ -2,88 +2,53 @@ module Cubical.Cohomology.EilenbergMacLane.Rings.Sn where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Path -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2wedgeS1 open import Cubical.Cohomology.EilenbergMacLane.Base -open import Cubical.Cohomology.EilenbergMacLane.Groups.Connected -open import Cubical.Cohomology.EilenbergMacLane.Groups.RP2 open import Cubical.Cohomology.EilenbergMacLane.Groups.Sn -open import Cubical.Cohomology.EilenbergMacLane.Groups.Wedge -open import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle open import Cubical.Cohomology.EilenbergMacLane.CupProduct -open import Cubical.Cohomology.EilenbergMacLane.Rings.KleinBottle hiding (α↦1) +open import Cubical.Cohomology.EilenbergMacLane.RingStructure -open import Cubical.Homotopy.EilenbergMacLane.GroupStructure -open import Cubical.Homotopy.EilenbergMacLane.Order2 -open import Cubical.Homotopy.EilenbergMacLane.Properties open import Cubical.Homotopy.EilenbergMacLane.Base +open import Cubical.Homotopy.EilenbergMacLane.Properties open import Cubical.Homotopy.EilenbergMacLane.CupProduct -open import Cubical.Homotopy.EilenbergMacLane.CupProductTensor - renaming (_⌣ₖ_ to _⌣ₖ⊗_ ; ⌣ₖ-0ₖ to ⌣ₖ-0ₖ⊗ ; 0ₖ-⌣ₖ to 0ₖ-⌣ₖ⊗) -open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected -open import Cubical.Cohomology.EilenbergMacLane.RingStructure -open import Cubical.Cohomology.EilenbergMacLane.Rings.Z2-properties open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order open import Cubical.Data.Unit -open import Cubical.Data.Fin -open import Cubical.Data.Sum -open import Cubical.Data.Fin.Arithmetic -open import Cubical.Data.Sigma -open import Cubical.Data.Vec open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.FinData renaming (Fin to FinI) +open import Cubical.Data.Vec +open import Cubical.Data.FinData +open import Cubical.Data.Sum -open import Cubical.HITs.KleinBottle renaming (rec to KleinFun) open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.SetTruncation as ST open import Cubical.HITs.Truncation as TR open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _sq/_) -open import Cubical.HITs.EilenbergMacLane1 hiding (elimProp ; elimSet) +open import Cubical.HITs.EilenbergMacLane1 open import Cubical.HITs.Susp -open import Cubical.HITs.Pushout open import Cubical.HITs.S1 renaming (rec to S¹Fun) open import Cubical.HITs.Sn -open import Cubical.HITs.RPn -open import Cubical.HITs.Wedge -open import Cubical.Algebra.GradedRing.DirectSumHIT -open import Cubical.Algebra.Group +open import Cubical.Algebra.GradedRing.DirectSumHIT open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.IntMod -open import Cubical.Algebra.AbGroup hiding (_[_]) -open import Cubical.Algebra.AbGroup.Instances.Unit +open import Cubical.Algebra.AbGroup open import Cubical.Algebra.DirectSum.DirectSumHIT.Base open import Cubical.Algebra.Ring -open import Cubical.Algebra.AbGroup.TensorProduct -open import Cubical.Algebra.Monoid - open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.Quotient -open import Cubical.Algebra.CommRing.Instances.IntMod open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly-Quotient -open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly-notationZ2 - -open import Cubical.Algebra.Monoid.Instances.Nat +open import Cubical.Algebra.Monoid.Instances.Nat -open Iso open PlusBis open RingTheory -open AbGroupStr -- open import Cubical.Agebra.Monoids.Instance.NatVec @@ -135,7 +100,7 @@ module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where G[X]→H*[Sⁿ,G]-main : Vec ℕ 1 → fst G → fst (H*R GRing (S₊ (suc n))) G[X]→H*[Sⁿ,G]-main (zero ∷ []) g = base 0 (invEquiv (H⁰[Sⁿ,G]≅G GAb n .fst) .fst g) - G[X]→H*[Sⁿ,G]-main (one ∷ []) g = base (suc n) (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g) + G[X]→H*[Sⁿ,G]-main (one ∷ []) g = base (suc n) (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g) G[X]→H*[Sⁿ,G]-main (suc (suc x) ∷ []) g = ⊕HIT.neutral G[X]→H*[Sⁿ,G]-main-coh₁ : (r : Vec ℕ 1) @@ -159,7 +124,7 @@ module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where G[X]→H*[Sⁿ,G]-fun : G[x] → fst (H*R GRing (S₊ (suc n))) G[X]→H*[Sⁿ,G]-fun = DS-Rec-Set.f _ _ _ _ isSetH*S - neutral + neutral G[X]→H*[Sⁿ,G]-main _add_ addAssoc @@ -168,9 +133,6 @@ module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where G[X]→H*[Sⁿ,G]-main-coh₁ G[X]→H*[Sⁿ,G]-main-coh₂ - open import Cubical.Foundations.Equiv.Properties - open import Cubical.Foundations.Transport - Hⁿ[Sⁿ,G]≅G-pres⌣ : (n : ℕ) (G : Ring ℓ) (g : fst G) (x : coHom (suc n) (Ring→AbGroup G) (S₊ (suc n))) → Hⁿ[Sⁿ,G]≅G (Ring→AbGroup G) n .fst .fst diff --git a/Cubical/Papers/ComputationalSyntheticCohomology.agda b/Cubical/Papers/ComputationalSyntheticCohomology.agda index 65ec06c95e..74abe6a870 100644 --- a/Cubical/Papers/ComputationalSyntheticCohomology.agda +++ b/Cubical/Papers/ComputationalSyntheticCohomology.agda @@ -18,9 +18,6 @@ open import Cubical.Algebra.AbGroup open import Cubical.Foundations.GroupoidLaws -- 1: Introduction - -import Cubical.ZCohomology.Groups.S2wedgeS1wedgeS1 as HⁿS²∨S¹∨S¹ - -- 2: Background open import Cubical.Foundations.Prelude as Prelude @@ -30,7 +27,7 @@ import Cubical.HITs.Susp as Suspensions import Cubical.HITs.Pushout as Pushouts import Cubical.Foundations.Path as Paths --- 3: Stuff +-- 3: Eilenberg-MacLane spaces import Cubical.Homotopy.EilenbergMacLane.Base as EMSpace import Cubical.Homotopy.EilenbergMacLane.Properties as EMProps import Cubical.Homotopy.EilenbergMacLane.WedgeConnectivity as WC @@ -43,11 +40,12 @@ import Cubical.Homotopy.EilenbergMacLane.GradedCommTensor as CupComm -- 4: Cohomology import Cubical.Cohomology.EilenbergMacLane.Base as Cohom import Cubical.Cohomology.EilenbergMacLane.CupProduct as CohomCup -import Cubical.Cohomology.EilenbergMacLane.MayerVietoris as MV import Cubical.Axiom.Choice as Choice import Cubical.HITs.Wedge as ⋁ import Cubical.Homotopy.EilenbergSteenrod as Axioms import Cubical.Cohomology.EilenbergMacLane.EilenbergSteenrod as SatAxioms +import Cubical.Cohomology.EilenbergMacLane.MayerVietoris as MV +import Cubical.Cohomology.EilenbergMacLane.Gysin as Gysin -- 5: Computations of cohomology groups and rings import Cubical.Cohomology.EilenbergMacLane.Groups.Unit as HⁿUnit @@ -64,36 +62,7 @@ import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle as CohomK² import Cubical.ZCohomology.CohomologyRings.KleinBottle as ZCohomRingK² import Cubical.Cohomology.EilenbergMacLane.Rings.RP2 as RP²Ring import Cubical.Cohomology.EilenbergMacLane.Rings.KleinBottle as K²Ring --- import Cubical.Cohomology.Groups.KleinBottle as HⁿK² --- import Cubical.ZCohomology.Groups.KleinBottle as HⁿK² - --- import Cubical.Cohomology.EilenbergMacLane.RSnings as CohomRingSn - --- import Cubical.ZCohomology.RingStructure.CohomologyRing as ℤCohomologyRing --- import Cubical.HITs.S1 as S1 - --- import Cubical.HITs.Sn as Sn --- import Cubical.ZCohomology.Groups.Sn as HⁿSᵐ --- import Cubical.ZCohomology.CohomologyRings.S1 as H*S¹ --- import Cubical.ZCohomology.CohomologyRings.Sn as H*Sᵐ --- open import Cubical.Homotopy.Hopf as HopfFibration --- import Cubical.ZCohomology.Groups.CP2 as HⁿℂP² --- import Cubical.ZCohomology.CohomologyRings.CP2 as H*ℂP² --- import Cubical.HITs.Wedge as ⋁ --- import Cubical.ZCohomology.Groups.S2wedgeS4 as HⁿS²∨S⁴ --- import Cubical.ZCohomology.CohomologyRings.S2wedgeS4 as H*S²∨S⁴ --- import Cubical.Cohomology.EilenbergMacLane.RingStructure as GCohomologyRing --- import Cubical.HITs.KleinBottle as 𝕂² - --- import Cubical.ZCohomology.CohomologyRings.KleinBottle as H*𝕂² --- import Cubical.ZCohomology.Groups.RP2wedgeS1 as HⁿℝP²∨S¹ --- import Cubical.ZCohomology.CohomologyRings.RP2wedgeS1 as H*ℝP²∨S¹ --- import Cubical.Cohomology.EilenbergMacLane.Groups.KleinBottle as ℤ/2-Hⁿ𝕂² --- open import Cubical.Cohomology.EilenbergMacLane.Rings.KleinBottle as ℤ/2-H*𝕂² --- import Cubical.Cohomology.EilenbergMacLane.Groups.RP2wedgeS1 as ℤ/2-HⁿℝP²∨S¹ --- import Cubical.Cohomology.EilenbergMacLane.Rings.RP2wedgeS1 as ℤ/2-H*ℝP²∨S¹ --- renaming (RP²∨S¹-CohomologyRing to H*RP²∨S¹≅ℤ/2[X,Y]/) - +import Cubical.Cohomology.EilenbergMacLane.Rings.RPinf as RP∞Ring ----- 1. INNTRODUCTION ----- @@ -255,10 +224,11 @@ open MV.MV using ( Ker-i⊂Im-d ; Im-d⊂Ker-i --- 4.4 The Thom isomorphism and the Gysin sequence -- Theorem 43 (Gysin sequence) --- TODO! Port from Steenrod branch. - --- +open Gysin.Gysin using ( Im-mapᵣ⊂Ker-mapₗ ; Ker-mapₗ⊂Im-mapᵣ + ; Ker-⌣⊂Im-mapₗ ; Im-mapₗ⊂Ker-⌣ + ; Im--⌣⊂Ker-mapᵣ ; Ker-mapᵣ⊂Im--⌣) +--- 5. Computations of cohomology groups and rings -- Proposition 44. Cohom of 1 open HⁿUnit using (H⁰[Unit,G]≅G ; Hⁿ⁺¹[Unit,G]≅0) @@ -330,13 +300,13 @@ open K²Ring using (α²↦1) -- Proposition 67. H*(RP²,Z/2) open RP²Ring using (H*[RP²,ℤ₂]≅ℤ₂[X]/) --- Proposition 68 (Roughly) +-- Proposition 68. (Roughly) open K²Ring using (α²↦1 ; βα↦1 ; β²↦0) --- Proposition 69 H*(K²,ℤ/2) +-- Proposition 69. H*(K²,ℤ/2) open K²Ring using (H*KleinBottle≅ℤ/2[X,Y]/) --- Lemma 70 (implicitly used) +-- Lemma 70. (implicitly used) --- Proposition 71 --- Todo +-- Proposition 71. H*(RP∞, ℤ/2) +open RP∞Ring using (H*[RP∞,ℤ₂]≅ℤ₂[X]) diff --git a/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda b/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda index 4c882195f3..0454d41bbd 100644 --- a/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda +++ b/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda @@ -24,7 +24,7 @@ open import Cubical.Data.FinData open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.AbGroup hiding (_[_]) +open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Group.Instances.Unit open import Cubical.Algebra.Group.Instances.Bool open import Cubical.Algebra.Group.Instances.Int diff --git a/Cubical/ZCohomology/CohomologyRings/RP2.agda b/Cubical/ZCohomology/CohomologyRings/RP2.agda index a56751a623..b17c6f6ab4 100644 --- a/Cubical/ZCohomology/CohomologyRings/RP2.agda +++ b/Cubical/ZCohomology/CohomologyRings/RP2.agda @@ -24,7 +24,7 @@ open import Cubical.Data.FinData open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.AbGroup hiding (_[_]) +open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Group.Instances.Unit open import Cubical.Algebra.Group.Instances.Bool open import Cubical.Algebra.Group.Instances.Int renaming (ℤGroup to ℤG) From 74a85212fdbfedc27cc2f3c9e2c846246065ec3d Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 23 Jan 2024 14:47:14 +0100 Subject: [PATCH 15/73] wups --- Cubical/Algebra/AbGroup/Instances/Int.agda | 11 ++ Cubical/Algebra/AbGroup/Instances/IntMod.agda | 25 ++- Cubical/Algebra/CommRing/Instances/Int.agda | 26 ---- Cubical/Algebra/Group/GroupPath.agda | 42 +---- Cubical/Algebra/Group/Instances/Pi.agda | 1 - .../EilenbergMacLane/Groups/Sn.agda | 2 +- .../EilenbergMacLane/Groups/Torus.agda | 2 +- Cubical/Data/Fin/Properties.agda | 38 ----- Cubical/Data/Nat/Order.agda | 18 --- Cubical/Data/Unit/Properties.agda | 3 - Cubical/Foundations/Path.agda | 6 - Cubical/Foundations/Pointed/Properties.agda | 10 -- Cubical/Foundations/Prelude.agda | 10 -- .../PropositionalTruncation/Properties.agda | 11 -- Cubical/HITs/Pushout/Properties.agda | 23 --- Cubical/HITs/SequentialColimit/Base.agda | 7 - .../HITs/SequentialColimit/Properties.agda | 146 ------------------ Cubical/HITs/Sn/Base.agda | 4 - Cubical/HITs/Sn/Properties.agda | 109 ++++--------- Cubical/HITs/Susp/Properties.agda | 18 --- Cubical/HITs/Truncation/Properties.agda | 8 - Cubical/Homotopy/Connected.agda | 32 ---- .../Homotopy/EilenbergMacLane/CupProduct.agda | 1 - 23 files changed, 65 insertions(+), 488 deletions(-) create mode 100644 Cubical/Algebra/AbGroup/Instances/Int.agda diff --git a/Cubical/Algebra/AbGroup/Instances/Int.agda b/Cubical/Algebra/AbGroup/Instances/Int.agda new file mode 100644 index 0000000000..cfd4dfe51a --- /dev/null +++ b/Cubical/Algebra/AbGroup/Instances/Int.agda @@ -0,0 +1,11 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.AbGroup.Instances.Int where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Int +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Instances.Int + +ℤAbGroup : AbGroup ℓ-zero +ℤAbGroup = Group→AbGroup ℤGroup +Comm diff --git a/Cubical/Algebra/AbGroup/Instances/IntMod.agda b/Cubical/Algebra/AbGroup/Instances/IntMod.agda index fb3daf8e78..a9c8eee522 100644 --- a/Cubical/Algebra/AbGroup/Instances/IntMod.agda +++ b/Cubical/Algebra/AbGroup/Instances/IntMod.agda @@ -7,14 +7,19 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function open import Cubical.Algebra.AbGroup + +open import Cubical.Algebra.Group.Instances.Int +open import Cubical.Algebra.AbGroup.Instances.Int open import Cubical.Algebra.Group.Instances.IntMod open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.ZAction open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) open import Cubical.Data.Nat.Order -open import Cubical.Data.Int renaming (_+_ to _+ℤ_) +open import Cubical.Data.Int + renaming (_+_ to _+ℤ_ ; _·_ to _·ℤ_ ; -_ to -ℤ_) open import Cubical.Data.Fin open import Cubical.Data.Fin.Arithmetic open import Cubical.Data.Sigma @@ -66,3 +71,19 @@ Iso.leftInv (fst ℤ/2/2≅ℤ/2) = snd ℤ/2/2≅ℤ/2 = makeIsGroupHom (SQ.elimProp (λ _ → isPropΠ λ _ → isSetFin _ _) λ a → SQ.elimProp (λ _ → isSetFin _ _) λ b → refl) + +ℤTorsion : (n : ℕ) → isContr (fst (ℤAbGroup [ (suc n) ]ₜ)) +fst (ℤTorsion n) = AbGroupStr.0g (snd (ℤAbGroup [ (suc n) ]ₜ)) +snd (ℤTorsion n) (a , p) = Σ≡Prop (λ _ → isSetℤ _ _) + (sym (help a (ℤ·≡· (pos (suc n)) a ∙ p))) + where + help : (a : ℤ) → a +ℤ pos n ·ℤ a ≡ 0 → a ≡ 0 + help (pos zero) p = refl + help (pos (suc a)) p = + ⊥.rec (snotz (injPos (pos+ (suc a) (n ·ℕ suc a) + ∙ cong (pos (suc a) +ℤ_) (pos·pos n (suc a)) ∙ p))) + help (negsuc a) p = ⊥.rec + (snotz (injPos (cong -ℤ_ (negsuc+ a (n ·ℕ suc a) + ∙ (cong (negsuc a +ℤ_) + (cong (-ℤ_) (pos·pos n (suc a))) + ∙ sym (cong (negsuc a +ℤ_) (pos·negsuc n a)) ∙ p))))) diff --git a/Cubical/Algebra/CommRing/Instances/Int.agda b/Cubical/Algebra/CommRing/Instances/Int.agda index 8ff91cd6ae..87b3a88bc5 100644 --- a/Cubical/Algebra/CommRing/Instances/Int.agda +++ b/Cubical/Algebra/CommRing/Instances/Int.agda @@ -4,16 +4,9 @@ module Cubical.Algebra.CommRing.Instances.Int where open import Cubical.Foundations.Prelude open import Cubical.Algebra.CommRing -open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.Group.ZAction -open import Cubical.Algebra.Group.Instances.Int open import Cubical.Data.Int as Int renaming ( ℤ to ℤ ; _+_ to _+ℤ_; _·_ to _·ℤ_; -_ to -ℤ_) -open import Cubical.Data.Nat - renaming (_+_ to _+ℕ_; _·_ to _·ℕ_) -open import Cubical.Data.Sigma -open import Cubical.Data.Empty as ⊥ open CommRingStr using (0r ; 1r ; _+_ ; _·_ ; -_ ; isCommRing) @@ -31,22 +24,3 @@ isCommRing (snd ℤCommRing) = isCommRingℤ isCommRingℤ = makeIsCommRing isSetℤ Int.+Assoc (λ _ → refl) -Cancel Int.+Comm Int.·Assoc Int.·IdR ·DistR+ Int.·Comm - -ℤAbGroup : AbGroup ℓ-zero -ℤAbGroup = Group→AbGroup ℤGroup +Comm - -ℤTorsion : (n : ℕ) → isContr (fst (ℤAbGroup [ (suc n) ]ₜ)) -fst (ℤTorsion n) = AbGroupStr.0g (snd (ℤAbGroup [ (suc n) ]ₜ)) -snd (ℤTorsion n) (a , p) = Σ≡Prop (λ _ → isSetℤ _ _) - (sym (help a (ℤ·≡· (pos (suc n)) a ∙ p))) - where - help : (a : ℤ) → a +ℤ pos n ·ℤ a ≡ 0 → a ≡ 0 - help (pos zero) p = refl - help (pos (suc a)) p = - ⊥.rec (snotz (injPos (pos+ (suc a) (n ·ℕ suc a) - ∙ cong (pos (suc a) +ℤ_) (pos·pos n (suc a)) ∙ p))) - help (negsuc a) p = ⊥.rec - (snotz (injPos (cong -ℤ_ (negsuc+ a (n ·ℕ suc a) - ∙ (cong (negsuc a +ℤ_) - (cong (-ℤ_) (pos·pos n (suc a))) - ∙ sym (cong (negsuc a +ℤ_) (pos·negsuc n a)) ∙ p))))) diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index 4159140591..ffe86ded07 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -1,5 +1,5 @@ -- The SIP applied to groups -{-# OPTIONS --safe --lossy-unification #-} +{-# OPTIONS --safe #-} module Cubical.Algebra.Group.GroupPath where open import Cubical.Foundations.Prelude @@ -11,6 +11,7 @@ open import Cubical.Foundations.GroupoidLaws hiding (assoc) open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence open import Cubical.Foundations.SIP +open import Cubical.Foundations.Path open import Cubical.Data.Sigma @@ -197,42 +198,3 @@ GroupEquivJ {G = G} P p {H} e = (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof (transportRefl (fst (fst e) (transportRefl x i)) i)))) ∙ retEq (fst e) x)) - - -isGroupoidGroup : ∀ {ℓ} → isGroupoid (Group ℓ) -isGroupoidGroup G H = - isOfHLevelRespectEquiv 2 (GroupPath _ _) - (isOfHLevelΣ 2 (isOfHLevel≃ 2 (GroupStr.is-set (snd G)) (GroupStr.is-set (snd H))) - λ _ → isProp→isSet (isPropIsGroupHom _ _)) - -module _ {ℓ ℓ'} {A : Type ℓ} - (G : A → Group ℓ') - (G-coh : (x y : A) → GroupEquiv (G x) (G y)) - (G-coh-coh : (x y z : A) (g : fst (G x)) - → fst (fst (G-coh y z)) ((fst (fst (G-coh x y)) g)) - ≡ fst (fst (G-coh x z)) g ) where - - PropTrunc→Group-coh : (x y : A) → G x ≡ G y - PropTrunc→Group-coh x y = uaGroup (G-coh x y) - - PropTrunc→Group-coh-coh : (x y z : A) → compGroupEquiv (G-coh x y) (G-coh y z) ≡ G-coh x z - PropTrunc→Group-coh-coh x y z = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (Σ≡Prop (λ _ → isPropIsEquiv _) - (funExt (G-coh-coh x y z))) - - open import Cubical.Foundations.Path - PropTrunc→Group : ∥ A ∥₁ → Group ℓ' - PropTrunc→Group = rec→Gpd isGroupoidGroup - G - (record { link = PropTrunc→Group-coh - ; coh₁ = coh-coh }) - where - coh-coh : (x y z : A) - → Square (PropTrunc→Group-coh x y) (PropTrunc→Group-coh x z) - refl (PropTrunc→Group-coh y z) - coh-coh x y z = - compPathL→PathP - (sym (lUnit _) - ∙∙ sym (uaCompGroupEquiv (G-coh x y) (G-coh y z)) - ∙∙ cong uaGroup (PropTrunc→Group-coh-coh x y z)) diff --git a/Cubical/Algebra/Group/Instances/Pi.agda b/Cubical/Algebra/Group/Instances/Pi.agda index 0d1483652b..bb2bb2e355 100644 --- a/Cubical/Algebra/Group/Instances/Pi.agda +++ b/Cubical/Algebra/Group/Instances/Pi.agda @@ -13,7 +13,6 @@ open IsMonoid open IsSemigroup module _ {ℓ ℓ' : Level} {X : Type ℓ} (G : X → Group ℓ') where - -- TODO: upstream ΠGroup : Group (ℓ-max ℓ ℓ') fst ΠGroup = (x : X) → fst (G x) 1g (snd ΠGroup) x = 1g (G x .snd) diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/Sn.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/Sn.agda index 92dea22524..099a4fdada 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/Sn.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/Sn.agda @@ -298,7 +298,7 @@ module _ (G : AbGroup ℓ) where , (TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) (sphereElim n (λ _ → isProp→isOfHLevelSuc n (isOfHLevelTrunc 2 _ _)) refl))) G - + Hⁿ[Sᵐ,G]Full : (n m : ℕ) → (((n ≡ 0) ⊎ (n ≡ suc m)) → AbGroupEquiv (coHomGr n G (S₊ (suc m))) G) diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/Torus.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/Torus.agda index 1788faa995..241c9c4a34 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/Torus.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/Torus.agda @@ -155,7 +155,7 @@ H¹[T²,G]≅G×G G = Iso.fun typIso = ST.rec (isSet× squash₂ squash₂) λ f → ∣ (λ x → f (x , base)) ∣₂ , ∣ (λ x → f (base , x)) ∣₂ Iso.inv typIso = uncurry (ST.rec2 squash₂ - λ f g → ∣ (λ x → f (fst x) +1 g (snd x)) ∣₂) + λ f g → ∣ (λ x → f (fst x) +1 g (snd x)) ∣₂) Iso.rightInv typIso = uncurry (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSet× squash₂ squash₂) _ _) (S¹-connElim (isConnectedEM {G' = G} 1) diff --git a/Cubical/Data/Fin/Properties.agda b/Cubical/Data/Fin/Properties.agda index d48638d7e9..b57b4ec979 100644 --- a/Cubical/Data/Fin/Properties.agda +++ b/Cubical/Data/Fin/Properties.agda @@ -680,41 +680,3 @@ isProp→Fin≤1 : (n : ℕ) → isProp (Fin n) → n ≤ 1 isProp→Fin≤1 0 _ = ≤-solver 0 1 isProp→Fin≤1 1 _ = ≤-solver 1 1 isProp→Fin≤1 (suc (suc n)) p = Empty.rec (fzero≠fone (p fzero fone)) - --- Characterisation of Π over (Fin (suc n)) -CharacΠFinIso : ∀ {ℓ} (n : ℕ) {B : Fin (suc n) → Type ℓ} - → Iso ((x : _) → B x) (B fzero × ((x : _) → B (fsuc x))) -Iso.fun (CharacΠFinIso n {B = B}) f = f fzero , f ∘ fsuc -Iso.inv (CharacΠFinIso n {B = B}) (x , f) (zero , p) = - subst B (Fin-fst-≡ {i = fzero} {j = zero , p} refl) x -Iso.inv (CharacΠFinIso n {B = B}) (x , f) (suc y , p) = - subst B (Fin-fst-≡ refl) (f (y , pred-≤-pred p)) -Iso.rightInv (CharacΠFinIso n {B = B}) (x , f) = - ΣPathP ((λ j → - transp (λ i → B (isSetFin fzero fzero (Fin-fst-≡ (λ _ → zero)) refl j i)) j x) - , funExt λ x → (λ j → subst B (pathid x j) - (f (fst x , pred-≤-pred (suc-≤-suc (snd x))))) - ∙ (λ i → transp (λ j → B (path₃ x (i ∨ j))) i (f (path₂ x i)))) - where - module _ (x : Fin n) where - path₁ : _ - path₁ = Fin-fst-≡ {i = (fsuc (fst x , pred-≤-pred (snd (fsuc x))))} - {j = fsuc x} refl - - path₂ : _ - path₂ = Fin-fst-≡ refl - - path₃ : _ - path₃ = cong fsuc path₂ - - pathid : path₁ ≡ path₃ - pathid = isSetFin _ _ _ _ - -Iso.leftInv (CharacΠFinIso n {B = B}) f = - funExt λ { (zero , p) j - → transp (λ i → B (Fin-fst-≡ {i = fzero} {j = zero , p} refl (i ∨ j))) - j (f (Fin-fst-≡ {i = fzero} {j = zero , p} refl j)) - ; (suc x , p) j → transp (λ i → B (q x p (i ∨ j))) j (f (q x p j))} - where - q : (x : _) (p : _) → _ - q x p = Fin-fst-≡ {i = (fsuc (x , pred-≤-pred p))} {j = suc x , p} refl diff --git a/Cubical/Data/Nat/Order.agda b/Cubical/Data/Nat/Order.agda index 2698bd7eb8..4d53caf66c 100644 --- a/Cubical/Data/Nat/Order.agda +++ b/Cubical/Data/Nat/Order.agda @@ -191,24 +191,6 @@ predℕ-≤-predℕ {suc m} {suc n} ineq = pred-≤-pred ineq <-k+-trans : (k + m < n) → m < n <-k+-trans {m} {k} {n} p = ≤<-trans (m , refl) p -¬-suc-n PathIdTruncIso - PathIdTrunc : {a b : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) PathIdTrunc n = isoToPath (PathIdTruncIso n) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index b2a5adcb26..5640935b83 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -332,26 +332,6 @@ isConnectedCong n f cf {a₀} {a₁} q = (sym (fiberCong f q)) (isConnectedPath n (cf (f a₁)) (a₀ , q) (a₁ , refl)) -isConnectedCong² : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) - → isConnectedFun (suc (suc n)) f - → ∀ {a₀ a₁ a₂ a₃} {p : a₀ ≡ a₁} {q : a₂ ≡ a₃} - {r : a₀ ≡ a₂} {s : a₁ ≡ a₃} - → isConnectedFun n - {A = Square p q r s} - {B = Square (cong f p) (cong f q) (cong f r) (cong f s)} - (λ p i j → f (p i j)) -isConnectedCong² n {A = A} f cf {a₀} {a₁} {r = r} {s = s} - = isConnectedCong²' _ r _ s - where - isConnectedCong²' : (a₂ : A) (r : a₀ ≡ a₂) (a₃ : A) (s : a₁ ≡ a₃) - {p : a₀ ≡ a₁} {q : a₂ ≡ a₃} - → isConnectedFun n - {A = Square p q r s} - {B = Square (cong f p) (cong f q) (cong f r) (cong f s)} - (λ p i j → f (p i j)) - isConnectedCong²' = - J> (J> isConnectedCong n (cong f) (isConnectedCong (suc n) f cf)) - isConnectedCong' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x : A} {y : B} (n : ℕ) (f : A → B) → isConnectedFun (suc n) f @@ -616,18 +596,6 @@ inrConnected {A = A} {B = B} {C = C} n f g iscon = (equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .snd i a)) -inlConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : ℕ) - → (f : C → A) (g : C → B) - → isConnectedFun n g - → isConnectedFun n {A = A} {B = Pushout f g} inl -inlConnected {A = A} {B = B} {C = C} n f g iscon = - transport (λ i → isConnectedFun n (lem i)) - (inrConnected n g f iscon) - where - lem : PathP (λ i → A → ua (symPushout g f) i) inr inl - lem = toPathP (λ j x → inl (transportRefl (transportRefl x j) j)) - - isConnectedPushout→ : (f₁ : X₀ → X₁) (f₂ : X₀ → X₂) (g₁ : Y₀ → Y₁) (g₂ : Y₀ → Y₂) (h₀ : X₀ → Y₀) (h₁ : X₁ → Y₁) (h₂ : X₂ → Y₂) diff --git a/Cubical/Homotopy/EilenbergMacLane/CupProduct.agda b/Cubical/Homotopy/EilenbergMacLane/CupProduct.agda index 7253f8ec0b..8a0ca8b127 100644 --- a/Cubical/Homotopy/EilenbergMacLane/CupProduct.agda +++ b/Cubical/Homotopy/EilenbergMacLane/CupProduct.agda @@ -271,4 +271,3 @@ syntax ⌣[]ₖ-syntax R x y = x ⌣[ R R]ₖ y syntax ⌣[]Cₖ-syntax R x y = x ⌣[ R R]Cₖ y syntax ⌣[,,]ₖ-syntax n m R x y = x ⌣[ R , n , m ]ₖ y syntax ⌣[,,]Cₖ-syntax n m R x y = x ⌣[ R , n , m ]Cₖ y - From 7875e3ee00816291bf03351f0ea3bfce3fc356fb Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 23 Jan 2024 14:51:05 +0100 Subject: [PATCH 16/73] wups --- Cubical/Algebra/CommRing/Instances/Int.agda | 1 - Cubical/Algebra/Group/GroupPath.agda | 3 --- Cubical/HITs/Pushout/Properties.agda | 1 - 3 files changed, 5 deletions(-) diff --git a/Cubical/Algebra/CommRing/Instances/Int.agda b/Cubical/Algebra/CommRing/Instances/Int.agda index 87b3a88bc5..2e6a6b5360 100644 --- a/Cubical/Algebra/CommRing/Instances/Int.agda +++ b/Cubical/Algebra/CommRing/Instances/Int.agda @@ -4,7 +4,6 @@ module Cubical.Algebra.CommRing.Instances.Int where open import Cubical.Foundations.Prelude open import Cubical.Algebra.CommRing - open import Cubical.Data.Int as Int renaming ( ℤ to ℤ ; _+_ to _+ℤ_; _·_ to _·ℤ_; -_ to -ℤ_) diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index ffe86ded07..84c33a481e 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -11,7 +11,6 @@ open import Cubical.Foundations.GroupoidLaws hiding (assoc) open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence open import Cubical.Foundations.SIP -open import Cubical.Foundations.Path open import Cubical.Data.Sigma @@ -25,8 +24,6 @@ open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.HITs.PropositionalTruncation - private variable ℓ ℓ' ℓ'' : Level diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index 21c7f9f254..c01e0b4b44 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -31,7 +31,6 @@ open import Cubical.Foundations.Function open import Cubical.Data.Sigma open import Cubical.Data.Unit -open import Cubical.Data.Empty as ⊥ open import Cubical.HITs.Pushout.Base From 7313a311f150ab352c767dc716c9f115bce8d46a Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 23 Jan 2024 15:04:18 +0100 Subject: [PATCH 17/73] wups --- Cubical/Axiom/Choice.agda | 44 --------------------------------------- 1 file changed, 44 deletions(-) diff --git a/Cubical/Axiom/Choice.agda b/Cubical/Axiom/Choice.agda index ba2cafa007..a330f67d5b 100644 --- a/Cubical/Axiom/Choice.agda +++ b/Cubical/Axiom/Choice.agda @@ -56,47 +56,3 @@ satAC→satAC∃ F B C = propBiimpl→Equiv squash₁ (isPropΠ (λ _ → squash -- All types satisfy (-2) level axiom of choice satAC₀ : {A : Type ℓ} → satAC ℓ' 0 A satAC₀ B = isoToIsEquiv (iso (λ _ _ → tt*) (λ _ → tt*) (λ _ → refl) λ _ → refl) - - - --- Main theorem Fin m satisfies AC for any level n. -FinSatAC : (n m : ℕ) → ∀ {ℓ} → satAC ℓ n (Fin m) -FinSatAC n zero B = - isoToIsEquiv (iso _ - (λ f → ∣ (λ x → ⊥.rec (¬Fin0 x)) ∣ₕ) - (λ f → funExt λ x → ⊥.rec (¬Fin0 x)) - (TR.elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) - λ a → cong ∣_∣ₕ (funExt λ x → ⊥.rec (¬Fin0 x)))) -FinSatAC n (suc m) B = - subst isEquiv (ac-eq n m {B} (FinSatAC n m)) - (isoToIsEquiv (ac-map' n m B (FinSatAC n m))) - where - ac-map' : ∀ {ℓ} (n m : ℕ) (B : Fin (suc m) → Type ℓ) → (satAC ℓ n (Fin m)) - → Iso (hLevelTrunc n ((x : _) → B x)) ((x : _) → hLevelTrunc n (B x)) - ac-map' n m B ise = - compIso (mapCompIso (CharacΠFinIso m)) - (compIso (truncOfProdIso n) - (compIso (Σ-cong-iso-snd λ _ → equivToIso (_ , ise (λ x → B (fsuc x)))) - (invIso (CharacΠFinIso m)))) - - ac-eq : (n m : ℕ) {B : _} → (eq : satAC ℓ n (Fin m)) - → Iso.fun (ac-map' n m B eq) ≡ choiceMap {B = B} n - ac-eq zero m {B = B} x = refl - ac-eq (suc n) m {B = B} x = - funExt (TR.elim (λ _ → isOfHLevelPath (suc n) - (isOfHLevelΠ (suc n) - (λ _ → isOfHLevelTrunc (suc n))) _ _) - λ F → funExt - λ { (zero , p) j → ∣ transp (λ i → B (p1 p (j ∨ i))) j (F (p1 p j)) ∣ₕ - ; (suc x , p) j → ∣ transp (λ i → B (p2 x p (j ∨ i))) j (F (p2 x p j)) ∣ₕ}) - where - p1 : (p : _ ) → fzero ≡ (zero , p) - p1 p = Fin-fst-≡ refl - - p2 : (x : _) (p : suc x < suc m) - → Path (Fin _) (fsuc (x , pred-≤-pred p)) (suc x , p) - p2 x p = Fin-fst-≡ refl - --- Key result for construction of cw-approx at lvl 0 -satAC∃Fin : (n : ℕ) → satAC∃ ℓ ℓ' (Fin n) -satAC∃Fin n = satAC→satAC∃ (FinSatAC 1 n) From b7716a46e9017d052b7d0a6bd1b2db4db16d40a4 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 23 Jan 2024 19:19:56 +0100 Subject: [PATCH 18/73] fixes --- Cubical/Axiom/Choice.agda | 10 +++- .../EilenbergMacLane/Groups/KleinBottle.agda | 2 +- .../Cohomology/EilenbergMacLane/Rings/Sn.agda | 49 +++++-------------- Cubical/Homotopy/EilenbergSteenrod.agda | 6 ++- Cubical/Papers/CohomologyRings.agda | 2 +- Cubical/ZCohomology/EilenbergSteenrodZ.agda | 21 ++++++++ 6 files changed, 50 insertions(+), 40 deletions(-) diff --git a/Cubical/Axiom/Choice.agda b/Cubical/Axiom/Choice.agda index a330f67d5b..8e3bd511ba 100644 --- a/Cubical/Axiom/Choice.agda +++ b/Cubical/Axiom/Choice.agda @@ -1,5 +1,14 @@ {-# OPTIONS --safe #-} +{- +This file contains a definition of the n-level axiom of coice, +i.e. the statment that the canonical map + +∥ Πₐ Bₐ ∥ₙ → (Πₐ ∥ Bₐ ∥ₙ) + +is an equivalence. +-} + module Cubical.Axiom.Choice where open import Cubical.Foundations.Prelude @@ -21,7 +30,6 @@ private variable ℓ ℓ' ℓ'' : Level --- Stuff about choice: choiceMap : {A : Type ℓ} {B : A → Type ℓ'} (n : ℕ) → hLevelTrunc n ((a : A) → B a) → (a : A) → hLevelTrunc n (B a) diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda index ca2190d186..816f326ec8 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda @@ -342,7 +342,7 @@ snd (isContr-HⁿKleinBottle n G) = ST.elim (λ _ → isSetPathImplicit) ---- With general coefficients --- H⁰(K²,G) : ? +-- H⁰(K²,G) H⁰[K²,G]≅G : (G : AbGroup ℓ) → AbGroupEquiv (coHomGr 0 G KleinBottle) G H⁰[K²,G]≅G G = H⁰conn (∣ point ∣ , (TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda index 5c3facf6a1..7f494a47c9 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda @@ -68,35 +68,12 @@ module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where (Poly G) (λ n → CommRing→AbGroup (PolyCommRing G n) .snd) - open RingStr (snd (H*R GRing (S₊ (suc n)))) - renaming - ( 0r to 0H*S - ; 1r to 1H*S - ; _+_ to _+H*S_ - ; -_ to -H*S_ - ; _·_ to _cupS_ - ; +Assoc to +H*SAssoc - ; +IdR to +H*SIdR - ; +Comm to +H*SComm - ; ·Assoc to ·H*SAssoc - ; is-set to isSetH*S - ; ·DistR+ to ·DistR+H* - ; ·DistL+ to ·DistL+H*) - - open RingStr (snd G[X]) - renaming - ( 0r to 0GX - ; 1r to 1GX - ; _+_ to _+GX_ - ; -_ to -GX_ - ; _·_ to _·GX_ - ; +Assoc to +GXAssoc - ; +IdR to +GXIdR - ; +Comm to +GXComm - ; ·Assoc to ·GXAssoc - ; is-set to isSetGX - ; ·DistR+ to ·DistR+G - ; ·DistL+ to ·DistL+G) + module H*Sⁿ = RingStr (snd (H*R GRing (S₊ (suc n)))) + module G[X]Str = RingStr (snd G[X]) + + private + _cupS_ = H*Sⁿ._·_ + _·GX_ = G[X]Str._·_ G[X]→H*[Sⁿ,G]-main : Vec ℕ 1 → fst G → fst (H*R GRing (S₊ (suc n))) G[X]→H*[Sⁿ,G]-main (zero ∷ []) g = base 0 (invEquiv (H⁰[Sⁿ,G]≅G GAb n .fst) .fst g) @@ -123,7 +100,7 @@ module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where G[X]→H*[Sⁿ,G]-fun : G[x] → fst (H*R GRing (S₊ (suc n))) G[X]→H*[Sⁿ,G]-fun = - DS-Rec-Set.f _ _ _ _ isSetH*S + DS-Rec-Set.f _ _ _ _ H*Sⁿ.is-set neutral G[X]→H*[Sⁿ,G]-main _add_ @@ -170,12 +147,12 @@ module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where → G[X]→H*[Sⁿ,G]-fun (x ·GX y) ≡ (G[X]→H*[Sⁿ,G]-fun x cupS G[X]→H*[Sⁿ,G]-fun y) G[X]→H*[Sⁿ,G]-pres· = DS-Ind-Prop.f _ _ _ _ - (λ _ → isPropΠ λ _ → isSetH*S _ _) + (λ _ → isPropΠ λ _ → H*Sⁿ.is-set _ _) (λ _ → refl) G[X]→H*[Sⁿ,G]-pres·+ - λ {x} {y} p q s → cong (G[X]→H*[Sⁿ,G]-fun) (·DistL+G x y s) + λ {x} {y} p q s → cong (G[X]→H*[Sⁿ,G]-fun) (G[X]Str.·DistL+ x y s) ∙ cong₂ _add_ (p s) (q s) - ∙ sym (·DistL+H* (G[X]→H*[Sⁿ,G]-fun x) + ∙ sym (H*Sⁿ.·DistL+ (G[X]→H*[Sⁿ,G]-fun x) (G[X]→H*[Sⁿ,G]-fun y) (G[X]→H*[Sⁿ,G]-fun s)) where main₀₁ : (g g' : fst G) → _ ≡ _ @@ -217,15 +194,15 @@ module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where G[X]→H*[Sⁿ,G]-pres·+ : (r : Vec ℕ 1) (a : fst G) (y : G[x]) → G[X]→H*[Sⁿ,G]-fun (base r a ·GX y) ≡ G[X]→H*[Sⁿ,G]-fun (base r a) cupS G[X]→H*[Sⁿ,G]-fun y - G[X]→H*[Sⁿ,G]-pres·+ v a = DS-Ind-Prop.f _ _ _ _ (λ _ → isSetH*S _ _) + G[X]→H*[Sⁿ,G]-pres·+ v a = DS-Ind-Prop.f _ _ _ _ (λ _ → H*Sⁿ.is-set _ _) (cong (G[X]→H*[Sⁿ,G]-fun) (0RightAnnihilates G[X] (base v a)) ∙ sym (0RightAnnihilates (H*R GRing (S₊ (suc n))) (G[X]→H*[Sⁿ,G]-fun (base v a) ))) (λ r t → main v r a t) - λ {x} {y} p q → cong (G[X]→H*[Sⁿ,G]-fun) (·DistR+G (base v a) x y) + λ {x} {y} p q → cong (G[X]→H*[Sⁿ,G]-fun) (G[X]Str.·DistR+ (base v a) x y) ∙ cong₂ _add_ p q - ∙ sym (·DistR+H* (G[X]→H*[Sⁿ,G]-fun (base v a)) + ∙ sym (H*Sⁿ.·DistR+ (G[X]→H*[Sⁿ,G]-fun (base v a)) (G[X]→H*[Sⁿ,G]-fun x) (G[X]→H*[Sⁿ,G]-fun y)) G[X]→H*[Sⁿ,G] : RingHom G[X] (H*R GRing (S₊ (suc n))) diff --git a/Cubical/Homotopy/EilenbergSteenrod.agda b/Cubical/Homotopy/EilenbergSteenrod.agda index bf49cd26aa..be22126aad 100644 --- a/Cubical/Homotopy/EilenbergSteenrod.agda +++ b/Cubical/Homotopy/EilenbergSteenrod.agda @@ -41,6 +41,9 @@ record coHomTheory {ℓ ℓ' : Level} (H : (n : ℤ) → Pointed ℓ → AbGroup Boolℓ = Lift Bool , lift true field Hmap : (n : ℤ) → {A B : Pointed ℓ} (f : A →∙ B) → AbGroupHom (H n B) (H n A) + HMapComp : (n : ℤ) → {A B C : Pointed ℓ} (g : B →∙ C) (f : A →∙ B) + → compGroupHom (Hmap n g) (Hmap n f) ≡ Hmap n (g ∘∙ f) + HMapId : (n : ℤ) → {A : Pointed ℓ} → Hmap n (idfun∙ A) ≡ idGroupHom Suspension : Σ[ F ∈ ((n : ℤ) {A : Pointed ℓ} → AbGroupEquiv (H (sucℤ n) (Susp (typ A) , north)) (H n A)) ] ({A B : Pointed ℓ} (f : A →∙ B) (n : ℤ) → fst (Hmap (sucℤ n) (suspFun (fst f) , refl)) ∘ invEq (fst (F n {A = B})) @@ -57,7 +60,8 @@ record coHomTheoryGen {ℓ ℓ' : Level} (H : (n : ℤ) → Pointed ℓ → AbGr Boolℓ = Lift Bool , lift true field Hmap : (n : ℤ) → {A B : Pointed ℓ} (f : A →∙ B) → AbGroupHom (H n B) (H n A) - HMapComp : (n : ℤ) → {A B C : Pointed ℓ} (g : B →∙ C) (f : A →∙ B) → compGroupHom (Hmap n g) (Hmap n f) ≡ Hmap n (g ∘∙ f) + HMapComp : (n : ℤ) → {A B C : Pointed ℓ} (g : B →∙ C) (f : A →∙ B) + → compGroupHom (Hmap n g) (Hmap n f) ≡ Hmap n (g ∘∙ f) HMapId : (n : ℤ) → {A : Pointed ℓ} → Hmap n (idfun∙ A) ≡ idGroupHom Suspension : Σ[ F ∈ ((n : ℤ) {A : Pointed ℓ} → AbGroupEquiv (H (sucℤ n) (Susp (typ A) , north)) (H n A)) ] diff --git a/Cubical/Papers/CohomologyRings.agda b/Cubical/Papers/CohomologyRings.agda index 780f3ba11e..88b2d42db4 100644 --- a/Cubical/Papers/CohomologyRings.agda +++ b/Cubical/Papers/CohomologyRings.agda @@ -275,7 +275,7 @@ open H*𝕂² using (CohomologyRing-𝕂²) open H*ℝP²∨S¹ using (CohomologyRing-RP²⋁S¹) -- ℤ/2ℤ Cohomology groups of the Klein Bottle and ℝP² ⋁ S¹ -open ℤ/2-Hⁿ𝕂² using (H⁰[K²,ℤ/2]≅ℤ/2 ; H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 ; H²[K²,ℤ/2]≅ℤ/2 ; H³⁺ⁿK²≅0) +open ℤ/2-Hⁿ𝕂² using (H⁰[K²,ℤ/2]≅ℤ/2 ; H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 ; H²[K²,ℤ/2]≅ℤ/2 ; H³⁺ⁿ[K²,G]≅0) open ℤ/2-HⁿℝP²∨S¹ using (H⁰[RP²∨S¹,ℤ/2]≅ℤ/2 ; H¹[RP²∨S¹,ℤ/2]≅ℤ/2×ℤ/2 ; H²[RP²∨S¹,ℤ/2]≅ℤ/2 ; H³⁺ⁿ[RP²∨S¹,ℤ/2]≅Unit) -- Proposition 5.7, ℤ/2ℤ cohomology ring of the Klein Bottle diff --git a/Cubical/ZCohomology/EilenbergSteenrodZ.agda b/Cubical/ZCohomology/EilenbergSteenrodZ.agda index 398b2da551..a124865f29 100644 --- a/Cubical/ZCohomology/EilenbergSteenrodZ.agda +++ b/Cubical/ZCohomology/EilenbergSteenrodZ.agda @@ -215,9 +215,30 @@ private fst (theMorph (negsuc n) f) = idfun _ snd (theMorph (negsuc n) f) = makeIsGroupHom λ _ _ → refl + compMorph : ∀ {ℓ} (n : ℤ) {A B C : Pointed ℓ} (g : B →∙ C) (f : A →∙ B) → + compGroupHom (theMorph n g) (theMorph n f) ≡ theMorph n (g ∘∙ f) + compMorph (pos zero) g f = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt (ST.elim (λ _ → isSetPathImplicit) + λ _ → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl))) + compMorph (pos (suc n)) f g = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt (ST.elim (λ _ → isSetPathImplicit) + λ _ → refl)) + compMorph (negsuc n) g f = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl + + idMorph : ∀ {ℓ} (n : ℤ) {A : Pointed ℓ} → theMorph n (idfun∙ A) ≡ idGroupHom + idMorph (pos zero) = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (ST.elim (λ _ → isSetPathImplicit) + λ _ → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl))) + idMorph (pos (suc n)) = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (ST.elim (λ _ → isSetPathImplicit) λ _ → refl)) + idMorph (negsuc n) = refl + open coHomTheory isCohomTheoryZ' : ∀ {ℓ} → coHomTheory {ℓ} coHomFunctor' Hmap isCohomTheoryZ' = theMorph + HMapComp isCohomTheoryZ' = compMorph + HMapId isCohomTheoryZ' = idMorph -------------------------- Suspension -------------------------- -- existence of suspension isomorphism From f7524ebc3fed41d408c158e93748ec2f3bbae462 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 23 Jan 2024 19:44:57 +0100 Subject: [PATCH 19/73] ugh... --- .../EilenbergMacLane/EilenbergSteenrod.agda | 39 ++++++++++--------- 1 file changed, 20 insertions(+), 19 deletions(-) diff --git a/Cubical/Cohomology/EilenbergMacLane/EilenbergSteenrod.agda b/Cubical/Cohomology/EilenbergMacLane/EilenbergSteenrod.agda index b6cdb7aaa0..c669679958 100644 --- a/Cubical/Cohomology/EilenbergMacLane/EilenbergSteenrod.agda +++ b/Cubical/Cohomology/EilenbergMacLane/EilenbergSteenrod.agda @@ -61,6 +61,22 @@ module _ where Hmap∙ (pos n) = coHomHom∙ n Hmap∙ (negsuc n) f = idGroupHom + compAx : (n : ℤ) {A B C : Pointed ℓ'} + (g : B →∙ C) (f : A →∙ B) → + compGroupHom (Hmap∙ n g) (Hmap∙ n f) + ≡ Hmap∙ n (g ∘∙ f) + compAx (pos n) g f = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (ST.elim (λ _ → isSetPathImplicit) + λ a → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl))) + compAx (negsuc n) g f = Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl + + idAx : (n : ℤ) {A : Pointed ℓ'} → + Hmap∙ n (idfun∙ A) ≡ idGroupHom + idAx (pos n) = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (ΣPathP (refl , sym (lUnit (snd f)))))) + idAx (negsuc n) = refl + suspMap : (n : ℤ) {A : Pointed ℓ'} → AbGroupHom (coHomRedℤ G (sucℤ n) (Susp (typ A) , north)) (coHomRedℤ G n A) @@ -175,6 +191,8 @@ module _ where satisfies-ES : ∀ {ℓ ℓ'} (G : AbGroup ℓ) → coHomTheory {ℓ'} (coHomRedℤ G) Hmap (satisfies-ES G) = coHomRedℤ.Hmap∙ + HMapComp (satisfies-ES G) = coHomRedℤ.compAx + HMapId (satisfies-ES G) = coHomRedℤ.idAx fst (Suspension (satisfies-ES G)) n = GroupIso→GroupEquiv (coHomRedℤ.suspMapIso n) snd (Suspension (satisfies-ES G)) f (pos n) = funExt (ST.elim (λ _ → isSetPathImplicit) λ f @@ -286,27 +304,10 @@ module _ where invGroupEquiv (GroupIso→GroupEquiv lUnitGroupIso^) open coHomTheoryGen -private - compAx : ∀ {ℓ ℓ'} (G : AbGroup ℓ) (n : ℤ) {A B C : Pointed ℓ'} - (g : B →∙ C) (f : A →∙ B) → - compGroupHom (coHomRedℤ.Hmap∙ {G = G} n g) (coHomRedℤ.Hmap∙ n f) - ≡ coHomRedℤ.Hmap∙ n (g ∘∙ f) - compAx G (pos n) g f = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (ST.elim (λ _ → isSetPathImplicit) - λ a → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousEM n) refl))) - compAx G (negsuc n) g f = Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl - - idAx : ∀ {ℓ ℓ'} (G : AbGroup ℓ) (n : ℤ) {A : Pointed ℓ'} → - coHomRedℤ.Hmap∙ {G = G} n (idfun∙ A) ≡ idGroupHom - idAx G (pos n) = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (ST.elim (λ _ → isSetPathImplicit) - λ f → cong ∣_∣₂ (ΣPathP (refl , sym (lUnit (snd f)))))) - idAx G (negsuc n) = refl - satisfies-ES-gen : ∀ {ℓ ℓ'} (G : AbGroup ℓ) → coHomTheoryGen {ℓ'} (coHomRedℤ G) Hmap (satisfies-ES-gen G) = coHomTheory.Hmap (satisfies-ES G) -HMapComp (satisfies-ES-gen G) = compAx G -HMapId (satisfies-ES-gen G) = idAx G +HMapComp (satisfies-ES-gen G) = coHomTheory.HMapComp (satisfies-ES G) +HMapId (satisfies-ES-gen G) = coHomTheory.HMapId (satisfies-ES G) Suspension (satisfies-ES-gen G) = coHomTheory.Suspension (satisfies-ES G) Exactness (satisfies-ES-gen G) = coHomTheory.Exactness (satisfies-ES G) Dimension (satisfies-ES-gen G) = coHomTheory.Dimension (satisfies-ES G) From e47bad54dc36571f9fc7f41f551e347131991430 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 24 Jan 2024 19:53:25 +0100 Subject: [PATCH 20/73] wups --- Cubical/Algebra/Group/GroupPath.agda | 42 ++++- Cubical/Data/Nat/Order.agda | 18 +++ Cubical/Data/Unit/Properties.agda | 3 + Cubical/Foundations/Path.agda | 6 + Cubical/Foundations/Pointed/Properties.agda | 11 ++ Cubical/Foundations/Prelude.agda | 10 ++ .../PropositionalTruncation/Properties.agda | 11 ++ Cubical/HITs/Pushout/Properties.agda | 24 +++ Cubical/HITs/SequentialColimit/Base.agda | 7 + .../HITs/SequentialColimit/Properties.agda | 146 ++++++++++++++++++ Cubical/HITs/Sn/Base.agda | 5 + Cubical/HITs/Sn/Properties.agda | 109 +++++++++---- Cubical/HITs/Susp/Properties.agda | 17 ++ Cubical/HITs/Truncation/Properties.agda | 8 + Cubical/Homotopy/Connected.agda | 31 ++++ 15 files changed, 420 insertions(+), 28 deletions(-) diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index 84c33a481e..b4b6b90793 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -1,5 +1,5 @@ -- The SIP applied to groups -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.Group.GroupPath where open import Cubical.Foundations.Prelude @@ -24,6 +24,8 @@ open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.HITs.PropositionalTruncation + private variable ℓ ℓ' ℓ'' : Level @@ -195,3 +197,41 @@ GroupEquivJ {G = G} P p {H} e = (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof (transportRefl (fst (fst e) (transportRefl x i)) i)))) ∙ retEq (fst e) x)) + +isGroupoidGroup : ∀ {ℓ} → isGroupoid (Group ℓ) +isGroupoidGroup G H = + isOfHLevelRespectEquiv 2 (GroupPath _ _) + (isOfHLevelΣ 2 (isOfHLevel≃ 2 (GroupStr.is-set (snd G)) (GroupStr.is-set (snd H))) + λ _ → isProp→isSet (isPropIsGroupHom _ _)) + +module _ {ℓ ℓ'} {A : Type ℓ} + (G : A → Group ℓ') + (G-coh : (x y : A) → GroupEquiv (G x) (G y)) + (G-coh-coh : (x y z : A) (g : fst (G x)) + → fst (fst (G-coh y z)) ((fst (fst (G-coh x y)) g)) + ≡ fst (fst (G-coh x z)) g ) where + + PropTrunc→Group-coh : (x y : A) → G x ≡ G y + PropTrunc→Group-coh x y = uaGroup (G-coh x y) + + PropTrunc→Group-coh-coh : (x y z : A) → compGroupEquiv (G-coh x y) (G-coh y z) ≡ G-coh x z + PropTrunc→Group-coh-coh x y z = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (Σ≡Prop (λ _ → isPropIsEquiv _) + (funExt (G-coh-coh x y z))) + + open import Cubical.Foundations.Path + PropTrunc→Group : ∥ A ∥₁ → Group ℓ' + PropTrunc→Group = rec→Gpd isGroupoidGroup + G + (record { link = PropTrunc→Group-coh + ; coh₁ = coh-coh }) + where + coh-coh : (x y z : A) + → Square (PropTrunc→Group-coh x y) (PropTrunc→Group-coh x z) + refl (PropTrunc→Group-coh y z) + coh-coh x y z = + compPathL→PathP + (sym (lUnit _) + ∙∙ sym (uaCompGroupEquiv (G-coh x y) (G-coh y z)) + ∙∙ cong uaGroup (PropTrunc→Group-coh-coh x y z)) diff --git a/Cubical/Data/Nat/Order.agda b/Cubical/Data/Nat/Order.agda index 4d53caf66c..87ffadf43f 100644 --- a/Cubical/Data/Nat/Order.agda +++ b/Cubical/Data/Nat/Order.agda @@ -194,9 +194,27 @@ predℕ-≤-predℕ {suc m} {suc n} ineq = pred-≤-pred ineq <-+-< : m < n → k < l → m + k < n + l <-+-< m PathIdTruncIso + PathIdTrunc : {a b : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) PathIdTrunc n = isoToPath (PathIdTruncIso n) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 5640935b83..842e58ac3d 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -332,6 +332,26 @@ isConnectedCong n f cf {a₀} {a₁} q = (sym (fiberCong f q)) (isConnectedPath n (cf (f a₁)) (a₀ , q) (a₁ , refl)) +isConnectedCong² : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) + → isConnectedFun (suc (suc n)) f + → ∀ {a₀ a₁ a₂ a₃} {p : a₀ ≡ a₁} {q : a₂ ≡ a₃} + {r : a₀ ≡ a₂} {s : a₁ ≡ a₃} + → isConnectedFun n + {A = Square p q r s} + {B = Square (cong f p) (cong f q) (cong f r) (cong f s)} + (λ p i j → f (p i j)) +isConnectedCong² n {A = A} f cf {a₀} {a₁} {r = r} {s = s} + = isConnectedCong²' _ r _ s + where + isConnectedCong²' : (a₂ : A) (r : a₀ ≡ a₂) (a₃ : A) (s : a₁ ≡ a₃) + {p : a₀ ≡ a₁} {q : a₂ ≡ a₃} + → isConnectedFun n + {A = Square p q r s} + {B = Square (cong f p) (cong f q) (cong f r) (cong f s)} + (λ p i j → f (p i j)) + isConnectedCong²' = + J> (J> isConnectedCong n (cong f) (isConnectedCong (suc n) f cf)) + isConnectedCong' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x : A} {y : B} (n : ℕ) (f : A → B) → isConnectedFun (suc n) f @@ -596,6 +616,17 @@ inrConnected {A = A} {B = B} {C = C} n f g iscon = (equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .snd i a)) +inlConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : ℕ) + → (f : C → A) (g : C → B) + → isConnectedFun n g + → isConnectedFun n {A = A} {B = Pushout f g} inl +inlConnected {A = A} {B = B} {C = C} n f g iscon = + transport (λ i → isConnectedFun n (lem i)) + (inrConnected n g f iscon) + where + lem : PathP (λ i → A → ua (symPushout g f) i) inr inl + lem = toPathP (λ j x → inl (transportRefl (transportRefl x j) j)) + isConnectedPushout→ : (f₁ : X₀ → X₁) (f₂ : X₀ → X₂) (g₁ : Y₀ → Y₁) (g₂ : Y₀ → Y₂) (h₀ : X₀ → Y₀) (h₁ : X₁ → Y₁) (h₂ : X₂ → Y₂) From 90714817e9e9115cdcc0d50fb9f839d17bfef30d Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 31 Jan 2024 10:46:56 +0100 Subject: [PATCH 21/73] stuff --- Cubical/Axiom/Choice.agda | 47 +++++++++++++++++++++++++++- Cubical/Data/Fin/Properties.agda | 38 ++++++++++++++++++++++ Cubical/HITs/Pushout/Properties.agda | 31 ++++++++---------- Cubical/HITs/Sn/Properties.agda | 4 +-- Cubical/HITs/Susp/Properties.agda | 7 +++-- 5 files changed, 103 insertions(+), 24 deletions(-) diff --git a/Cubical/Axiom/Choice.agda b/Cubical/Axiom/Choice.agda index 8e3bd511ba..b780c0f32b 100644 --- a/Cubical/Axiom/Choice.agda +++ b/Cubical/Axiom/Choice.agda @@ -59,8 +59,53 @@ satAC→satAC∃ : {A : Type ℓ} satAC→satAC∃ F B C = propBiimpl→Equiv squash₁ (isPropΠ (λ _ → squash₁)) _ (λ f → invEq propTrunc≃Trunc1 (TR.map (λ f → fst ∘ f , λ a → f a .snd ) - (invEq (_ , F (λ x → Σ (B x) (C x))) λ a → fst propTrunc≃Trunc1 (f a)))) .snd + (invEq (_ , F (λ x → Σ (B x) (C x))) + λ a → fst propTrunc≃Trunc1 (f a)))) .snd -- All types satisfy (-2) level axiom of choice satAC₀ : {A : Type ℓ} → satAC ℓ' 0 A satAC₀ B = isoToIsEquiv (iso (λ _ _ → tt*) (λ _ → tt*) (λ _ → refl) λ _ → refl) + + +-- Fin m satisfies AC for any level n. +FinSatAC : (n m : ℕ) → ∀ {ℓ} → satAC ℓ n (Fin m) +FinSatAC n zero B = + isoToIsEquiv (iso _ + (λ f → ∣ (λ x → ⊥.rec (¬Fin0 x)) ∣ₕ) + (λ f → funExt λ x → ⊥.rec (¬Fin0 x)) + (TR.elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) + λ a → cong ∣_∣ₕ (funExt λ x → ⊥.rec (¬Fin0 x)))) +FinSatAC n (suc m) B = + subst isEquiv (ac-eq n m {B} (FinSatAC n m)) + (isoToIsEquiv (ac-map' n m B (FinSatAC n m))) + where + ac-map' : ∀ {ℓ} (n m : ℕ) (B : Fin (suc m) → Type ℓ) → (satAC ℓ n (Fin m)) + → Iso (hLevelTrunc n ((x : _) → B x)) ((x : _) → hLevelTrunc n (B x)) + ac-map' n m B ise = + compIso (mapCompIso (CharacΠFinIso m)) + (compIso (truncOfProdIso n) + (compIso (Σ-cong-iso-snd λ _ → equivToIso + (_ , ise (λ x → B (fsuc x)))) + (invIso (CharacΠFinIso m)))) + + ac-eq : (n m : ℕ) {B : _} → (eq : satAC ℓ n (Fin m)) + → Iso.fun (ac-map' n m B eq) ≡ choiceMap {B = B} n + ac-eq zero m {B = B} x = refl + ac-eq (suc n) m {B = B} x = + funExt (TR.elim (λ _ → isOfHLevelPath (suc n) + (isOfHLevelΠ (suc n) + (λ _ → isOfHLevelTrunc (suc n))) _ _) + λ F → funExt + λ { (zero , p) j → ∣ transp (λ i → B (lem₁ p (j ∨ i))) j (F (lem₁ p j)) ∣ₕ + ; (suc x , p) j → ∣ transp (λ i → B (lem₂ x p (j ∨ i))) j (F (lem₂ x p j)) ∣ₕ}) + where + lem₁ : (p : _ ) → fzero ≡ (zero , p) + lem₁ p = Fin-fst-≡ refl + + lem₂ : (x : _) (p : suc x < suc m) + → Path (Fin _) (fsuc (x , pred-≤-pred p)) (suc x , p) + lem₂ x p = Fin-fst-≡ refl + +-- Key result for construction of cw-approx at lvl 0 +satAC∃Fin : (n : ℕ) → satAC∃ ℓ ℓ' (Fin n) +satAC∃Fin n = satAC→satAC∃ (FinSatAC 1 n) diff --git a/Cubical/Data/Fin/Properties.agda b/Cubical/Data/Fin/Properties.agda index b57b4ec979..d48638d7e9 100644 --- a/Cubical/Data/Fin/Properties.agda +++ b/Cubical/Data/Fin/Properties.agda @@ -680,3 +680,41 @@ isProp→Fin≤1 : (n : ℕ) → isProp (Fin n) → n ≤ 1 isProp→Fin≤1 0 _ = ≤-solver 0 1 isProp→Fin≤1 1 _ = ≤-solver 1 1 isProp→Fin≤1 (suc (suc n)) p = Empty.rec (fzero≠fone (p fzero fone)) + +-- Characterisation of Π over (Fin (suc n)) +CharacΠFinIso : ∀ {ℓ} (n : ℕ) {B : Fin (suc n) → Type ℓ} + → Iso ((x : _) → B x) (B fzero × ((x : _) → B (fsuc x))) +Iso.fun (CharacΠFinIso n {B = B}) f = f fzero , f ∘ fsuc +Iso.inv (CharacΠFinIso n {B = B}) (x , f) (zero , p) = + subst B (Fin-fst-≡ {i = fzero} {j = zero , p} refl) x +Iso.inv (CharacΠFinIso n {B = B}) (x , f) (suc y , p) = + subst B (Fin-fst-≡ refl) (f (y , pred-≤-pred p)) +Iso.rightInv (CharacΠFinIso n {B = B}) (x , f) = + ΣPathP ((λ j → + transp (λ i → B (isSetFin fzero fzero (Fin-fst-≡ (λ _ → zero)) refl j i)) j x) + , funExt λ x → (λ j → subst B (pathid x j) + (f (fst x , pred-≤-pred (suc-≤-suc (snd x))))) + ∙ (λ i → transp (λ j → B (path₃ x (i ∨ j))) i (f (path₂ x i)))) + where + module _ (x : Fin n) where + path₁ : _ + path₁ = Fin-fst-≡ {i = (fsuc (fst x , pred-≤-pred (snd (fsuc x))))} + {j = fsuc x} refl + + path₂ : _ + path₂ = Fin-fst-≡ refl + + path₃ : _ + path₃ = cong fsuc path₂ + + pathid : path₁ ≡ path₃ + pathid = isSetFin _ _ _ _ + +Iso.leftInv (CharacΠFinIso n {B = B}) f = + funExt λ { (zero , p) j + → transp (λ i → B (Fin-fst-≡ {i = fzero} {j = zero , p} refl (i ∨ j))) + j (f (Fin-fst-≡ {i = fzero} {j = zero , p} refl j)) + ; (suc x , p) j → transp (λ i → B (q x p (i ∨ j))) j (f (q x p j))} + where + q : (x : _) (p : _) → _ + q x p = Fin-fst-≡ {i = (fsuc (x , pred-≤-pred p))} {j = suc x , p} refl diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index 2e832226b5..03916db59d 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -531,23 +531,18 @@ module PushoutDistr {ℓ ℓ' ℓ'' ℓ''' : Level} compPath-filler' (push (g a)) (λ j → inr (push a j)) (~ j) i leftInv PushoutDistrIso (push a i) j = push a (i ∧ j) -PushoutEmptyFam : ∀ {ℓ} {A B C : Type ℓ} - → ¬ A → ¬ C +PushoutEmptyFam : ¬ A → ¬ C → {f : A → B} {g : A → C} → Iso B (Pushout {A = A} {B = B} {C = C} f g) -PushoutEmptyFam {A = A} {B = B} {C = C} ¬A ¬C = - compIso is - (pushoutIso _ _ _ _ - (uninhabEquiv (λ {()}) ¬A) - (idEquiv B) - (uninhabEquiv (λ {()}) ¬C) - (funExt (λ{()})) (funExt λ{()})) - where - lift⊥ : ∀ ℓ → Type ℓ - lift⊥ ℓ = Lift ⊥ - - is : Iso B (Pushout {A = lift⊥ ℓ} {B = B} {C = lift⊥ ℓ''} (λ{()}) λ{()}) - Iso.fun is = inl - Iso.inv is (inl x) = x - Iso.rightInv is (inl x) = refl - Iso.leftInv is x = refl +fun (PushoutEmptyFam ¬A ¬C) = inl +inv (PushoutEmptyFam ¬A ¬C) (inl x) = x +inv (PushoutEmptyFam ¬A ¬C) (inr x) = ⊥.rec (¬C x) +inv (PushoutEmptyFam ¬A ¬C {f = f} {g = g}) (push a i) = + ⊥.rec {A = f a ≡ rec (¬C (g a))} (¬A a) i +rightInv (PushoutEmptyFam {A = A} {B = B} ¬A ¬C) (inl x) = refl +rightInv (PushoutEmptyFam {A = A} {B = B} ¬A ¬C) (inr x) = ⊥.rec (¬C x) +rightInv (PushoutEmptyFam {A = A} {B = B} ¬A ¬C {f = f} {g = g}) (push a i) j = + ⊥.rec {A = Square (λ i → inl (rec {A = f a ≡ rec (¬C (g a))} (¬A a) i)) + (push a) (λ _ → inl (f a)) (rec (¬C (g a)))} + (¬A a) j i +leftInv (PushoutEmptyFam {A = A} {B = B} ¬A ¬C) x = refl diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 50b9689589..21ddcde829 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -300,7 +300,7 @@ sphereConnected : (n : HLevel) → isConnected (suc n) (S₊ n) sphereConnected n = ∣ ptSn n ∣ , elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) (λ a → sym (spoke ∣_∣ (ptSn n)) ∙ spoke ∣_∣ a) -sphereToTrunc : (n : ℕ) {A : S₊ n → Type} +sphereToTrunc : (n : ℕ) {A : S₊ n → Type ℓ} → ((x : S₊ n) → hLevelTrunc (suc n) (A x)) → ∥ ((x : _) → A x) ∥₁ sphereToTrunc zero {A = A} indr = @@ -322,7 +322,7 @@ sphereToTrunc (suc (suc n)) {A = A} indr = lem (sphereToTrunc (suc n)) (indr north) (indr south) λ a → cong indr (merid a) where - lem : ({A : S₊ (suc n) → Type} → + lem : ({A : S₊ (suc n) → Type _} → ((i : S₊ (suc n)) → hLevelTrunc (suc (suc n)) (A i)) → ∥ ((x : S₊ (suc n)) → A x) ∥₁) → (x : hLevelTrunc (3 + n) (A north)) diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index 419116e97b..6c043cf3b9 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -21,21 +21,22 @@ open import Cubical.Homotopy.Loopspace private variable ℓ : Level + A B C : Type ℓ open Iso -suspFunComp : {A B C : Type ℓ} (f : B → C) (g : A → B) +suspFunComp : (f : B → C) (g : A → B) → suspFun (f ∘ g) ≡ (suspFun f) ∘ (suspFun g) suspFunComp f g i north = north suspFunComp f g i south = south suspFunComp f g i (merid a i₁) = merid (f (g a)) i₁ -suspFunConst : {A B : Type ℓ} (b : B) → suspFun (λ (_ : A) → b) ≡ λ _ → north +suspFunConst : (b : B) → suspFun (λ (_ : A) → b) ≡ λ _ → north suspFunConst b i north = north suspFunConst b i south = merid b (~ i) suspFunConst b i (merid a j) = merid b (~ i ∧ j) -suspFunIdFun : {A : Type ℓ} → suspFun (λ (a : A) → a) ≡ λ x → x +suspFunIdFun : suspFun (λ (a : A) → a) ≡ λ x → x suspFunIdFun i north = north suspFunIdFun i south = south suspFunIdFun i (merid a j) = merid a j From ea5fefdda2f7ceb38d0cd1bc799ec4cdff448dac Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 25 Feb 2024 23:02:26 +0100 Subject: [PATCH 22/73] stuff --- .../AbGroup/Instances/FreeAbGroup.agda | 404 ++++++++++++++ Cubical/Algebra/AbGroup/Instances/Pi.agda | 4 + Cubical/Algebra/AbGroup/Properties.agda | 26 +- Cubical/Algebra/ChainComplex/Base.agda | 367 +++++++++++++ Cubical/Algebra/ChainComplex/Natindexed.agda | 80 +++ Cubical/Algebra/Group/Base.agda | 5 + Cubical/Algebra/Group/ChainComplex/Base.agda | 0 Cubical/Algebra/Group/Instances/Int.agda | 11 + Cubical/Algebra/Group/ZAction.agda | 29 +- Cubical/Data/CW.agda | 6 + Cubical/Data/CW/Base.agda | 161 ++++++ Cubical/Data/CW/Properties.agda | 171 ++++++ Cubical/Data/Fin/Base.agda | 8 + Cubical/Data/Fin/Properties.agda | 68 ++- Cubical/Data/Int/Base.agda | 8 + Cubical/Data/Int/Properties.agda | 9 + Cubical/HITs/FreeAbGroup/Base.agda | 42 ++ Cubical/HITs/Sn/Degree.agda | 163 ++++++ Cubical/HITs/Sn/Properties.agda | 15 + Cubical/HITs/SphereBouquet.agda | 5 + Cubical/HITs/SphereBouquet/Base.agda | 16 + Cubical/HITs/SphereBouquet/Degree.agda | 501 ++++++++++++++++++ Cubical/HITs/SphereBouquet/Properties.agda | 375 +++++++++++++ Cubical/HITs/Susp/Base.agda | 6 + Cubical/HITs/Wedge.agda | 1 + Cubical/HITs/Wedge/Properties.agda | 167 ++++++ Cubical/Homotopy/Group/PinSn.agda | 263 ++++++++- Cubical/Structures/Successor.agda | 6 + Cubical/ZCohomology/GroupStructure.agda | 12 + Cubical/ZCohomology/Groups/Sn.agda | 145 ++++- 30 files changed, 3050 insertions(+), 24 deletions(-) create mode 100644 Cubical/Algebra/ChainComplex/Base.agda create mode 100644 Cubical/Algebra/ChainComplex/Natindexed.agda create mode 100644 Cubical/Algebra/Group/ChainComplex/Base.agda create mode 100644 Cubical/Data/CW.agda create mode 100644 Cubical/Data/CW/Base.agda create mode 100644 Cubical/Data/CW/Properties.agda create mode 100644 Cubical/HITs/Sn/Degree.agda create mode 100644 Cubical/HITs/SphereBouquet.agda create mode 100644 Cubical/HITs/SphereBouquet/Base.agda create mode 100644 Cubical/HITs/SphereBouquet/Degree.agda create mode 100644 Cubical/HITs/SphereBouquet/Properties.agda create mode 100644 Cubical/HITs/Wedge/Properties.agda diff --git a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda index 1ddc12a1df..d2757f95fe 100644 --- a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda +++ b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda @@ -2,8 +2,26 @@ module Cubical.Algebra.AbGroup.Instances.FreeAbGroup where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat hiding (_·_) renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) +open import Cubical.Data.Fin +open import Cubical.Data.Empty as ⊥ + open import Cubical.HITs.FreeAbGroup + open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.Pi +open import Cubical.Algebra.AbGroup.Instances.Int +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties + private variable ℓ : Level @@ -12,3 +30,389 @@ module _ {A : Type ℓ} where FAGAbGroup : AbGroup ℓ FAGAbGroup = makeAbGroup {G = FreeAbGroup A} ε _·_ _⁻¹ trunc assoc identityᵣ invᵣ comm + +-- Alternative definition of case when A = Fin n +ℤ[Fin_] : (n : ℕ) → AbGroup ℓ-zero +ℤ[Fin n ] = ΠℤAbGroup (Fin n) + + +-- generator of ℤ[Fin_] +ℤFinGenerator : {n : ℕ} (k : Fin n) → ℤ[Fin n ] .fst +ℤFinGenerator {n = n} k s with (fst k ≟ fst s) +... | lt x = 0 +... | eq x = 1 +... | gt x = 0 + +ℤFinGeneratorComm : {n : ℕ} (x y : Fin n) → ℤFinGenerator x y ≡ ℤFinGenerator y x +ℤFinGeneratorComm x y with (fst x ≟ fst y) | (fst y ≟ fst x) +... | lt x₁ | lt x₂ = refl +... | lt x₁ | eq x₂ = ⊥.rec (¬m 1∈Im→isEquiv G e) diff --git a/Cubical/Data/CW.agda b/Cubical/Data/CW.agda new file mode 100644 index 0000000000..80ce46452a --- /dev/null +++ b/Cubical/Data/CW.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.Data.CW where + +open import Cubical.Data.CW.Base public +open import Cubical.Data.CW.Properties public diff --git a/Cubical/Data/CW/Base.agda b/Cubical/Data/CW/Base.agda new file mode 100644 index 0000000000..b3ca77c06d --- /dev/null +++ b/Cubical/Data/CW/Base.agda @@ -0,0 +1,161 @@ +{-# OPTIONS --cubical --safe --lossy-unification #-} + +{-This file contains: + +-} + +module Cubical.Data.CW.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Fin +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ + +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + +open import Cubical.HITs.SequentialColimit +open Sequence + +open import Cubical.Relation.Nullary + + + +private + variable + ℓ ℓ' : Level + +--- CW complexes --- +-- Definition of (the skeleton) of an arbitrary CW complex +-- New def: X 0 is empty and C (suc n) is pushout +yieldsCWskel : (ℕ → Type ℓ) → Type ℓ +yieldsCWskel X = + Σ[ f ∈ (ℕ → ℕ) ] + Σ[ α ∈ ((n : ℕ) → (Fin (f n) × (S⁻ n)) → X n) ] + ((¬ X zero) × + ((n : ℕ) → X (suc n) ≃ Pushout (α n) fst)) + +CWskel : (ℓ : Level) → Type (ℓ-suc ℓ) +CWskel ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel X) + +module CWskel-fields (C : CWskel ℓ) where + card = C .snd .fst + A = Fin ∘ card + α = C .snd .snd .fst + e = C .snd .snd .snd .snd + + ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero + ℤ[A n ] = ℤ[Fin (snd C .fst n) ] + +-- Technically, a CW complex should be the sequential colimit over the following maps +CW↪ : (T : CWskel ℓ) → (n : ℕ) → fst T n → fst T (suc n) +CW↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inl x) + +-- But if it stabilises, no need for colimits. +yieldsFinCWskel : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ +yieldsFinCWskel n X = + Σ[ CWskel ∈ yieldsCWskel X ] ((k : ℕ) → isEquiv (CW↪ (X , CWskel) (k +ℕ n))) + +-- ... which should give us finite CW complexes. +finCWskel : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) +finCWskel ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel n C) + +finCWskel→CWskel : (n : ℕ) → finCWskel ℓ n → CWskel ℓ +finCWskel→CWskel n C = fst C , fst (snd C) + +realiseSeq : CWskel ℓ → Sequence ℓ +Sequence.obj (realiseSeq (C , p)) = C +Sequence.map (realiseSeq C) = CW↪ C _ + +-- realisation of CW complex from skeleton +realise : CWskel ℓ → Type ℓ +realise C = SeqColim (realiseSeq C) + +-- Finally: definition of CW complexes +isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' + +CW : (ℓ : Level) → Type (ℓ-suc ℓ) +CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ + +CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) + +-- Finite CW complexes +isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +isFinCW {ℓ = ℓ} X = + Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) + +finCW : (ℓ : Level) → Type (ℓ-suc ℓ) +finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ + +finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) + +isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X +isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str + +finCW→CW : finCW ℓ → CW ℓ +finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p + + +-- morphisms +_→ᶜʷ_ : CW ℓ → CW ℓ → Type ℓ +C →ᶜʷ D = fst C → fst D + +--the cofibre of the inclusion of X n into X (suc n) +cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ +cofibCW n C = cofib (CW↪ C n) + +--...is basically a sphere bouquet +cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) + → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) +cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e + where + s = Bouquet≃-gen + α = C .snd .snd .fst + e = C .snd .snd .snd .snd n + +--sending X (suc n) into the cofibCW +to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C +to_cofibCW n C x = inr x + +--the connecting map of the long exact sequence +δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) + → cofib conn → Susp A +δ-pre conn (inl x) = north +δ-pre conn (inr x) = south +δ-pre conn (push a i) = merid a i + +δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) +δ n C = δ-pre (CW↪ C n) + +-- send the stage n to the realization (the same as incl, but with explicit args and type) +CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C +CW↪∞ C n x = incl x + +finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m +fst (finCW↑ m n p C) = fst C +fst (snd (finCW↑ m n p C)) = snd C .fst +snd (snd (finCW↑ m n p C)) k = + subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) + (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) + (snd C .snd (k +ℕ fst p)) diff --git a/Cubical/Data/CW/Properties.agda b/Cubical/Data/CW/Properties.agda new file mode 100644 index 0000000000..02ee60aa47 --- /dev/null +++ b/Cubical/Data/CW/Properties.agda @@ -0,0 +1,171 @@ +{-# OPTIONS --cubical --safe --lossy-unification #-} + +{-This file contains: + +-} + +module Cubical.Data.CW.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Fin +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.CW.Base + +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + +open import Cubical.HITs.SequentialColimit +open Sequence + +open import Cubical.Relation.Nullary + + + +private + variable + ℓ ℓ' : Level + +CW₀-empty : (C : CWskel ℓ) → ¬ fst C 0 +CW₀-empty C = snd (snd (snd C)) .fst + +CW₁-discrete : (C : CWskel ℓ) → fst C 1 ≃ Fin (snd C .fst 0) +CW₁-discrete C = compEquiv (snd C .snd .snd .snd 0) (isoToEquiv main) + where + main : Iso (Pushout (fst (snd C .snd) 0) fst) (Fin (snd C .fst 0)) + Iso.fun main (inl x) = ⊥.rec (CW₀-empty C x) + Iso.fun main (inr x) = x + Iso.inv main = inr + Iso.rightInv main x = refl + Iso.leftInv main (inl x) = ⊥.rec (CW₀-empty C x) + Iso.leftInv main (inr x) = refl + +-- elimination from Cₙ into prop +CWskel→Prop : (C : CWskel ℓ) {A : (n : ℕ) → fst C n → Type ℓ'} + → ((n : ℕ) (x : _) → isProp (A n x)) + → ((a : _) → A (suc zero) a) + → ((n : ℕ) (a : _) → (A (suc n) a → A (suc (suc n)) (CW↪ C (suc n) a))) + → (n : _) (c : fst C n) → A n c +CWskel→Prop C {A = A} pr b eqs zero c = ⊥.rec (CW₀-empty C c) +CWskel→Prop C {A = A} pr b eqs (suc zero) = b +CWskel→Prop C {A = A} pr b eqs (suc (suc n)) c = + subst (A (suc (suc n))) + (retEq (snd C .snd .snd .snd (suc n)) c) + (help (CWskel→Prop C pr b eqs (suc n)) _) + where + help : (inder : (c₁ : fst C (suc n)) → A (suc n) c₁) + → (a : Pushout _ fst) + → A (suc (suc n)) (invEq (snd C .snd .snd .snd (suc n)) a) + help inder = + elimProp _ (λ _ → pr _ _) (λ b → eqs n _ (inder b)) + λ c → subst (A (suc (suc n))) + (cong (invEq (snd C .snd .snd .snd (suc n))) (push (c , ptSn n))) + (eqs n _ (inder _)) + +isSet-CW₀ : (C : CWskel ℓ) → isSet (fst C (suc zero)) +isSet-CW₀ C = + isOfHLevelRetractFromIso 2 (equivToIso (CW₁-discrete C)) + isSetFin + +-- eliminating from CW complex into prop +CW→Prop : (C : CWskel ℓ) {A : realise C → Type ℓ'} + → ((x : _) → isProp (A x)) + → ((a : _) → A (incl {n = suc zero} a)) + → (a : _) → A a +CW→Prop C {A = A} pr ind = + SeqColim→Prop pr (CWskel→Prop C (λ _ _ → pr _) + ind + λ n a → subst A (push a)) + +-- realisation of finite complex +realiseFin : (n : ℕ) (C : finCWskel ℓ n) → Iso (fst C n) (realise (finCWskel→CWskel n C)) +realiseFin n C = converges→ColimIso n (snd C .snd) + +-- elimination principles for CW complexes +module _ {ℓ : Level} (C : CWskel ℓ) where + open CWskel-fields C + module _ (n : ℕ) {B : fst C (suc n) → Type ℓ'} + (inler : (x : fst C n) → B (invEq (e n) (inl x))) + (inrer : (x : A n) → B (invEq (e n) (inr x))) + (pusher : (x : A n) (y : S⁻ n) + → PathP (λ i → B (invEq (e n) (push (x , y) i))) + (inler (α n (x , y))) + (inrer x)) where + private + gen : ∀ {ℓ ℓ'} {A B : Type ℓ} (C : A → Type ℓ') + (e : A ≃ B) + → ((x : B) → C (invEq e x)) + → (x : A) → C x + gen C e h x = subst C (retEq e x) (h (fst e x)) + + gen-coh : ∀ {ℓ ℓ'} {A B : Type ℓ} (C : A → Type ℓ') + (e : A ≃ B) (h : (x : B) → C (invEq e x)) + → (b : B) + → gen C e h (invEq e b) ≡ h b + gen-coh {ℓ' = ℓ'} {A = A} {B = B} C e = + EquivJ (λ A e → (C : A → Type ℓ') (h : (x : B) → C (invEq e x)) + → (b : B) + → gen C e h (invEq e b) ≡ h b) + (λ C h b → transportRefl (h b)) e C + + main : (x : _) → B (invEq (e n) x) + main (inl x) = inler x + main (inr x) = inrer x + main (push (x , y) i) = pusher x y i + + CWskel-elim : (x : _) → B x + CWskel-elim = gen B (e n) main + + CWskel-elim-inl : (x : _) → CWskel-elim (invEq (e n) (inl x)) ≡ inler x + CWskel-elim-inl x = gen-coh B (e n) main (inl x) + + module _ (n : ℕ) {B : fst C (suc (suc n)) → Type ℓ'} + (inler : (x : fst C (suc n)) + → B (invEq (e (suc n)) (inl x))) + (ind : ((x : A (suc n)) (y : S₊ n) + → PathP (λ i → B (invEq (e (suc n)) + ((push (x , y) ∙ sym (push (x , ptSn n))) i))) + (inler (α (suc n) (x , y))) + (inler (α (suc n) (x , ptSn n))))) where + CWskel-elim' : (x : _) → B x + CWskel-elim' = + CWskel-elim (suc n) inler + (λ x → subst (λ t → B (invEq (e (suc n)) t)) + (push (x , ptSn n)) + (inler (α (suc n) (x , ptSn n)))) + λ x y → toPathP (sym (substSubst⁻ (B ∘ invEq (e (suc n))) _ _) + ∙ cong (subst (λ t → B (invEq (e (suc n)) t)) + (push (x , ptSn n))) + (sym (substComposite (B ∘ invEq (e (suc n))) _ _ _) + ∙ fromPathP (ind x y))) + + CWskel-elim'-inl : (x : _) + → CWskel-elim' (invEq (e (suc n)) (inl x)) ≡ inler x + CWskel-elim'-inl = CWskel-elim-inl (suc n) {B = B} inler _ _ + +finCWskel≃ : (n : ℕ) (C : finCWskel ℓ n) (m : ℕ) → n ≤ m → fst C n ≃ fst C m +finCWskel≃ n C m (zero , diff) = substEquiv (λ n → fst C n) diff +finCWskel≃ n C zero (suc x , diff) = ⊥.rec (snotz diff) +finCWskel≃ n C (suc m) (suc x , diff) = + compEquiv (finCWskel≃ n C m (x , cong predℕ diff)) + (compEquiv (substEquiv (λ n → fst C n) (sym (cong predℕ diff))) + (compEquiv (_ , snd C .snd x) + (substEquiv (λ n → fst C n) diff))) diff --git a/Cubical/Data/Fin/Base.agda b/Cubical/Data/Fin/Base.agda index 85286e0f2d..c414d879a7 100644 --- a/Cubical/Data/Fin/Base.agda +++ b/Cubical/Data/Fin/Base.agda @@ -68,6 +68,9 @@ fsplit (suc k , k (J> (J> (J> λ p → sym (rUnit refl) ◁ sym (help m p)))) + where + help : (m : ℕ) (p : _) + → (sym (rUnitₖ m (0ₖ m)) + ∙∙ cong (+ₖ-syntax m (0ₖ m)) + ((sym p ∙ refl) ∙ refl) + ∙∙ rUnitₖ m (0ₖ m)) + ∙ p + ≡ refl + help zero p = isSetℤ _ _ _ _ + help (suc zero) p = + cong (_∙ p) (sym (rUnit _) + ∙ λ i j → lUnitₖ 1 (rUnit (rUnit (sym p) (~ i)) (~ i) j) i) + ∙ lCancel p + help (suc (suc m)) p = + cong (_∙ p) (sym (rUnit _) + ∙ λ i j → lUnitₖ (2 +ℕ m) (rUnit (rUnit (sym p) (~ i)) (~ i) j) i) + ∙ lCancel p + + ⋁gen→∙Kn≡∙ : {A : Type ℓ} {B : A → Pointed ℓ'} (m : ℕ) + (f g : ⋁gen∙ A B →∙ coHomK-ptd m) + → ((x : _) → fst f (inr x) ≡ fst g (inr x)) + → f ≡ g + ⋁gen→∙Kn≡∙ m f g hom = + ΣPathP ((funExt (⋁gen→∙Kn≡ m f g hom)) + , (flipSquare ((λ i + → (λ j → snd f (i ∧ j)) + ∙∙ (λ j → snd f (i ∨ j)) + ∙∙ sym (snd g)) + ◁ λ i → doubleCompPath-filler (snd f) refl (sym (snd g)) (~ i)))) + +-- bouquetDegree preserves id +bouquetDegreeId : {n m : ℕ} + → bouquetDegree (idfun (SphereBouquet n (Fin m))) ≡ idGroupHom +bouquetDegreeId {n = n} {m = m} = + agreeOnℤFinGenerator→≡ λ (x : Fin m) + → funExt λ t + → sym (isGeneratorℤFinGenerator' + (λ w → degree n (λ x₁ → pickPetal t (inr (w , x₁)))) x) + ∙ help x t + where + help : (x t : Fin m) + → degree n (λ x₁ → pickPetal t (inr (x , x₁))) ≡ ℤFinGenerator x t + help x y with (fst x ≟ fst y) + ... | lt p = cong (degree n) (funExt lem) ∙ degree-const n + where + lem : (x₁ : S₊ n) → pickPetal y (inr (x , x₁)) ≡ ptSn n + lem x₁ with (fst y ≟ fst x) + ... | lt x = refl + ... | eq q = ⊥.rec (¬m λ i j → rCancel (λ _ → base) j i +fst (isIso-πₙSⁿ-unpointIso (suc n)) = + ST.rec squash₂ λ f → ST.map (λ p → f , p) (makePted n (f north)) +fst (snd (isIso-πₙSⁿ-unpointIso (suc n))) = + ST.elim (λ _ → isSetPathImplicit) + λ f → ST.rec isSetPathImplicit + (λ p j → πₙSⁿ-unpoint (suc n) + (ST.map (λ p₁ → f , p₁) (makePted-eq n (f north) p j))) + (makePted n (f north)) +snd (snd (isIso-πₙSⁿ-unpointIso (suc n))) = + ST.elim (λ _ → isSetPathImplicit) + λ f i → ST.map (λ p → fst f , p) (makePted-eq n _ (snd f) i) + +πₙSⁿ-unpointIso : (n : ℕ) → Iso ∥ (S₊∙ (suc n) →∙ S₊∙ (suc n)) ∥₂ ∥ (S₊ (suc n) → S₊ (suc n)) ∥₂ +Iso.fun (πₙSⁿ-unpointIso n) = πₙSⁿ-unpoint n +Iso.inv (πₙSⁿ-unpointIso n) = isIso-πₙSⁿ-unpointIso n .fst +Iso.rightInv (πₙSⁿ-unpointIso n) = isIso-πₙSⁿ-unpointIso n .snd .fst +Iso.leftInv (πₙSⁿ-unpointIso n) = isIso-πₙSⁿ-unpointIso n .snd .snd + + + +-- πₙ(Sⁿ) → Hⁿ(Sⁿ,ℤ) +πₙSⁿ→HⁿSⁿ-fun : (n : ℕ) → π'Gr n (S₊∙ (suc n)) .fst → coHom (suc n) (S₊ (suc n)) +πₙSⁿ→HⁿSⁿ-fun n = ST.map λ f x → ∣ fst f x ∣ + +-- to prove it's a homomorphism, we need the following +module suspensionLemmas (n : ℕ) where + Susp⊣Ω-Sn← : ((f : S₊ (suc n) → Ω (S₊∙ (suc (suc n))) .fst) + → S₊ (suc (suc n)) → S₊ (suc (suc n))) + Susp⊣Ω-Sn← f north = north + Susp⊣Ω-Sn← f south = north + Susp⊣Ω-Sn← f (merid a i) = f a i + + Susp⊣Ω-Sn←-σ : (f : S₊∙ (suc n) →∙ Ω (S₊∙ (suc (suc n)))) + → (x : _) + → cong (Susp⊣Ω-Sn← (fst f)) (σSn (suc n) x) ≡ fst f x + Susp⊣Ω-Sn←-σ f x = + cong-∙ (Susp⊣Ω-Sn← (fst f)) (merid x) (sym (merid _)) + ∙ cong (λ z → fst f x ∙ sym z) (snd f) + ∙ sym (rUnit _) + + Susp⊣Ω-Sn←≡ : (f : _) → ∥ Σ[ g ∈ (S₊∙ (suc n) →∙ Ω (S₊∙ (suc (suc n)))) ] f + ≡ Susp⊣Ω-Sn← (fst g) ∥₂ + Susp⊣Ω-Sn←≡ f = + ST.map + (λ p → ((λ x → (sym p ∙ cong f (merid x) + ∙ (cong f (sym (merid (ptSn _))) ∙ p))) + , (cong (sym p ∙_) + (assoc _ _ _ ∙ cong (_∙ p) (rCancel (cong f (merid (ptSn _)))) + ∙ sym (lUnit p)) + ∙ lCancel p)) + , funExt ( + λ { north → p + ; south → cong f (sym (merid (ptSn _))) ∙ p + ; (merid a i) j → compPath-filler' (sym p) + (compPath-filler (cong f (merid a)) + (cong f (sym (merid (ptSn _))) ∙ p) j) j i})) + (makePted n (f north)) + + Susp⊣Ω-Sn←≡∙ : (f : (S₊∙ (2 + n)) →∙ S₊∙ (2 + n)) + → ∃[ g ∈ _ ] f ≡ (Susp⊣Ω-Sn← (fst g) , refl) + Susp⊣Ω-Sn←≡∙ f = + ST.rec (isProp→isSet squash₁) + (uncurry (λ g q → TR.rec (isProp→isOfHLevelSuc n squash₁) + (λ r → ∣ g , ΣPathP (q , (sym r ◁ (λ i j → q (i ∨ j) north))) ∣₁) + (isConnectedPath _ + (isConnectedPathSⁿ (suc n) (fst f north) north) + (funExt⁻ q north) (snd f) .fst ))) + (Susp⊣Ω-Sn←≡ (fst f)) + +-- homomorphism πₙ(Sⁿ) → Hⁿ(Sⁿ,ℤ) +πₙSⁿ→HⁿSⁿ : (n : ℕ) + → GroupHom (π'Gr n (S₊∙ (suc n))) (coHomGr (suc n) (S₊ (suc n))) +fst (πₙSⁿ→HⁿSⁿ n) = πₙSⁿ→HⁿSⁿ-fun n +snd (πₙSⁿ→HⁿSⁿ zero) = + makeIsGroupHom + (ST.elim2 (λ _ _ → isSetPathImplicit) + λ f g → subst2 P (sym (S¹-fun-charac f .snd)) (sym (S¹-fun-charac g .snd)) + (main (S¹-fun-charac f .fst) (S¹-fun-charac g .fst))) + where + mkfun : fst (Ω (S₊∙ 1)) → S¹ → S¹ + mkfun p base = base + mkfun p (loop i) = p i + + main : (a b : Ω (S₊∙ 1) .fst) + → πₙSⁿ→HⁿSⁿ-fun zero (·π' zero ∣ mkfun a , refl ∣₂ ∣ mkfun b , refl ∣₂) + ≡ πₙSⁿ→HⁿSⁿ-fun zero ∣ mkfun a , refl ∣₂ + +ₕ πₙSⁿ→HⁿSⁿ-fun zero ∣ mkfun b , refl ∣₂ + main a b = cong ∣_∣₂ (funExt + λ { base → refl + ; (loop i) j → ((λ i j → ∣ (rUnit a (~ i) ∙ rUnit b (~ i)) j ∣ₕ) + ∙∙ cong-∙ ∣_∣ₕ a b + ∙∙ ∙≡+₁ (cong ∣_∣ₕ a) (cong ∣_∣ₕ b)) j i}) + + S¹-fun-charac : (f : S₊∙ 1 →∙ S₊∙ 1) + → Σ[ g ∈ Ω (S₊∙ 1) .fst ] f ≡ (mkfun g , refl) + S¹-fun-charac (f , p) = (sym p ∙∙ cong f loop ∙∙ p) + , ΣPathP ((funExt ( + λ { base → p + ; (loop i) j → doubleCompPath-filler (sym p) (cong f loop) p j i})) + , λ i j → p (i ∨ j)) + + P : (f g : _) → Type + P f g = πₙSⁿ→HⁿSⁿ-fun zero (·π' zero ∣ f ∣₂ ∣ g ∣₂) + ≡ (πₙSⁿ→HⁿSⁿ-fun zero ∣ f ∣₂) +ₕ (πₙSⁿ→HⁿSⁿ-fun zero ∣ g ∣₂) + +snd (πₙSⁿ→HⁿSⁿ (suc n)) = makeIsGroupHom isGrHom + where + open suspensionLemmas n + isGrHom : (f g : _) + → πₙSⁿ→HⁿSⁿ-fun (suc n) (·π' _ f g) + ≡ πₙSⁿ→HⁿSⁿ-fun (suc n) f +ₕ πₙSⁿ→HⁿSⁿ-fun (suc n) g + isGrHom = + ST.elim2 (λ _ _ → isSetPathImplicit) + λ f g → PT.rec2 (squash₂ _ _) + (uncurry (λ g' gp → uncurry λ h hp + → (λ i → πₙSⁿ→HⁿSⁿ-fun (suc n) (·π' (suc n) ∣ gp i ∣₂ ∣ hp i ∣₂) ) + ∙∙ cong ∣_∣₂ (funExt + (λ { north → refl + ; south → refl + ; (merid a i) j + → hcomp (λ k → + λ {(i = i0) → 0ₖ (2 + n) + ; (i = i1) → 0ₖ (2 + n) + ; (j = i0) → ∣ (rUnit (Susp⊣Ω-Sn←-σ g' a (~ k)) k + ∙ rUnit (Susp⊣Ω-Sn←-σ h a (~ k)) k) i ∣ₕ + ; (j = i1) → ∙≡+₂ _ (cong ∣_∣ₕ (g' .fst a)) (cong ∣_∣ₕ (h .fst a)) k i}) + (cong-∙ ∣_∣ₕ (g' .fst a) (h .fst a) j i)})) + ∙∙ λ i → πₙSⁿ→HⁿSⁿ-fun (suc n) ∣ gp (~ i) ∣₂ + +ₕ πₙSⁿ→HⁿSⁿ-fun (suc n) ∣ hp (~ i) ∣₂)) + (Susp⊣Ω-Sn←≡∙ f) (Susp⊣Ω-Sn←≡∙ g) + +-- isomorphism πₙ(Sⁿ) ≅ Hⁿ(Sⁿ,ℤ) +πₙSⁿ≅HⁿSⁿ : (n : ℕ) + → GroupEquiv (π'Gr n (S₊∙ (suc n))) (coHomGr (suc n) (S₊ (suc n))) +fst (fst (πₙSⁿ≅HⁿSⁿ n)) = fst (πₙSⁿ→HⁿSⁿ n) +snd (fst (πₙSⁿ≅HⁿSⁿ zero)) = + subst isEquiv (sym funid) (isoToIsEquiv (compIso is1 is2)) + where + is1 : Iso (π' 1 (S₊∙ 1)) ∥ (S¹ → S¹) ∥₂ + is1 = πₙSⁿ-unpointIso 0 + + is2 : Iso ∥ (S¹ → S¹) ∥₂ (coHom 1 S¹) + is2 = setTruncIso (codomainIso (invIso (truncIdempotentIso 3 isGroupoidS¹))) + + funid : πₙSⁿ→HⁿSⁿ-fun zero ≡ Iso.fun is2 ∘ Iso.fun is1 + funid = funExt (ST.elim (λ _ → isSetPathImplicit) λ _ → refl) +snd (fst (πₙSⁿ≅HⁿSⁿ (suc n))) = + gen∈Im→isEquiv _ + (GroupIso→GroupEquiv (invGroupIso (πₙ'Sⁿ≅ℤ (suc n)))) _ + (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc n)))) + ∣ ∣_∣ₕ ∣₂ + (sym (HⁿSⁿ-gen (suc n))) + (πₙSⁿ→HⁿSⁿ (suc n)) + ∣ ∣ idfun∙ _ ∣₂ , refl ∣₁ +snd (πₙSⁿ≅HⁿSⁿ n) = snd (πₙSⁿ→HⁿSⁿ n) + +-- the multiplications +module _ (n : ℕ) where + multSⁿ↬ : (f g : ∥ (S₊ (suc n) → S₊ (suc n)) ∥₂) + → ∥ (S₊ (suc n) → S₊ (suc n)) ∥₂ + multSⁿ↬ = ST.rec2 squash₂ λ f g → ∣ f ∘ g ∣₂ + + multπₙ : (f g : π'Gr n (S₊∙ (suc n)) .fst) → π'Gr n (S₊∙ (suc n)) .fst + multπₙ = ST.rec2 squash₂ λ f g → ∣ f ∘∙ g ∣₂ + +-- preservation of multiplication under relevant isos +multπₙ-pres : (n : ℕ) (f g : π'Gr n (S₊∙ (suc n)) .fst) + → Iso.fun (πₙSⁿ-unpointIso n) (multπₙ n f g) + ≡ multSⁿ↬ n (Iso.fun (πₙSⁿ-unpointIso n) f) (Iso.fun (πₙSⁿ-unpointIso n) g) +multπₙ-pres n = ST.elim2 (λ _ _ → isSetPathImplicit) λ f g → refl + +multπₙ-pres' : (n : ℕ) (f g : ∥ (S₊ (suc n) → S₊ (suc n)) ∥₂) + → Iso.inv (πₙSⁿ-unpointIso n) (multSⁿ↬ n f g) + ≡ multπₙ n (Iso.inv (πₙSⁿ-unpointIso n) f) (Iso.inv (πₙSⁿ-unpointIso n) g) +multπₙ-pres' n f g = + (λ i → isIso-πₙSⁿ-unpointIso n .fst + (multSⁿ↬ n (Iso.rightInv (πₙSⁿ-unpointIso n) f (~ i)) + (Iso.rightInv (πₙSⁿ-unpointIso n) g (~ i)))) + ∙∙ sym (cong (Iso.inv (πₙSⁿ-unpointIso n)) + (multπₙ-pres n (Iso.inv (πₙSⁿ-unpointIso n) f) (Iso.inv (πₙSⁿ-unpointIso n) g))) + ∙∙ Iso.leftInv (πₙSⁿ-unpointIso n) _ + +multHⁿSⁿ-pres : (n : ℕ) (f g : π'Gr n (S₊∙ (suc n)) .fst) + → πₙSⁿ→HⁿSⁿ-fun n (multπₙ n f g) + ≡ multHⁿSⁿ n (πₙSⁿ→HⁿSⁿ-fun n f) (πₙSⁿ→HⁿSⁿ-fun n g) +multHⁿSⁿ-pres n = ST.elim2 (λ _ _ → isSetPathImplicit) λ f g → refl diff --git a/Cubical/Structures/Successor.agda b/Cubical/Structures/Successor.agda index ce72551c60..01055702f2 100644 --- a/Cubical/Structures/Successor.agda +++ b/Cubical/Structures/Successor.agda @@ -8,6 +8,7 @@ module Cubical.Structures.Successor where open import Cubical.Foundations.Prelude open import Cubical.Data.Int open import Cubical.Data.Nat +open import Cubical.Data.Fin private variable @@ -27,3 +28,8 @@ open SuccStr ℕ+ : SuccStr ℓ-zero ℕ+ .Index = ℕ ℕ+ .succ = suc + +Fin+ : (n : ℕ) → SuccStr ℓ-zero +Fin+ n .Index = Fin n +Fin+ n .succ (x , zero , p) = x , zero , p +Fin+ n .succ (x , suc diff , p) = suc x , diff , +-suc diff (suc x) ∙ p diff --git a/Cubical/ZCohomology/GroupStructure.agda b/Cubical/ZCohomology/GroupStructure.agda index 7944f1d9b2..f229f79c7d 100644 --- a/Cubical/ZCohomology/GroupStructure.agda +++ b/Cubical/ZCohomology/GroupStructure.agda @@ -15,6 +15,7 @@ open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) open import Cubical.Data.Sigma open import Cubical.Data.Int renaming (_+_ to _+ℤ_ ; -_ to -ℤ_) open import Cubical.Data.Nat renaming (+-assoc to +-assocℕ ; +-comm to +-commℕ) +open import Cubical.Data.Fin open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms @@ -445,6 +446,9 @@ syntax +ₕ-syntax n x y = x +[ n ]ₕ y syntax -ₕ-syntax n x = -[ n ]ₕ x syntax -ₕ'-syntax n x y = x -[ n ]ₕ y +sumFinK : {n m : ℕ} (f : Fin n → coHomK m) → coHomK m +sumFinK {n = n} {m = m} = sumFinGen (λ x y → x +[ m ]ₖ y) (0ₖ m) + 0ₕ : (n : ℕ) → coHom n A 0ₕ n = ∣ (λ _ → (0ₖ n)) ∣₂ @@ -944,3 +948,11 @@ open IsGroupHom -- → (e : GroupIso (coHomGr n A) G) -- → coHomGr n A ≡ inducedCoHom e -- inducedCoHomPath e = InducedGroupPath _ _ _ _ + +sumFinKComm : {n m : ℕ} (f : Fin n → S₊ m → coHomK m) + → sumFinGroup (coHomGr m (S₊ m)) (λ x → ∣ f x ∣₂) + ≡ ∣ (λ x → sumFinK {m = m} λ i → f i x) ∣₂ +sumFinKComm {n = zero} {m = m} f = refl +sumFinKComm {n = suc n} {m = m} f = + cong (λ y → ∣ f flast ∣₂ +[ m ]ₕ y) + (sumFinKComm {n = n} (f ∘ injectSuc)) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index 013c68e189..b37c9ed7bb 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -15,7 +15,8 @@ open import Cubical.Relation.Nullary open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Bool open import Cubical.Data.Nat -open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Int + renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc ; _·_ to _·ℤ_) open import Cubical.Data.Sigma open import Cubical.Data.Sum @@ -25,6 +26,7 @@ open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.Unit open import Cubical.Algebra.Group.Instances.Int +open import Cubical.Algebra.Group.ZAction open import Cubical.HITs.Pushout open import Cubical.HITs.Sn @@ -304,6 +306,147 @@ Hⁿ-Sᵐ≅0 (suc n) (suc m) pf = suspensionAx-Sn n m □ Hⁿ-Sᵐ≅0 n m λ p → pf (cong suc p) +-- generator of Hⁿ(Sⁿ) +HⁿSⁿ-gen : (n : ℕ) → Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n)) ∣ ∣_∣ₕ ∣₂ ≡ 1 +HⁿSⁿ-gen zero = refl +HⁿSⁿ-gen (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) main ∙ HⁿSⁿ-gen n + where + lem : Iso.inv (fst (suspensionAx-Sn n n)) ∣ ∣_∣ₕ ∣₂ ≡ ∣ ∣_∣ₕ ∣₂ + lem = cong ∣_∣₂ + (funExt λ { north → refl + ; south i → ∣ merid (ptSn (suc n)) i ∣ₕ + ; (merid a i) j → ∣ compPath-filler (merid a) + (sym (merid (ptSn (suc n)))) (~ j) i ∣ₕ}) + + main : Iso.fun (fst (suspensionAx-Sn n n)) ∣ ∣_∣ₕ ∣₂ ≡ ∣ ∣_∣ₕ ∣₂ + main = (sym (cong (Iso.fun (fst (suspensionAx-Sn n n))) lem) + ∙ Iso.rightInv (fst (suspensionAx-Sn n n)) ∣ ∣_∣ₕ ∣₂) + +----------------------- multiplication ---------------------------- +-- explicit description of the (ring) multiplication on Hⁿ(Sⁿ) +premultHⁿSⁿ : (n : ℕ) (f g : S₊ (suc n) + → coHomK (suc n)) → (S₊ (suc n) → coHomK (suc n)) +premultHⁿSⁿ n f g x = T.rec (isOfHLevelTrunc (3 + n)) f (g x) + +multHⁿSⁿ : (n : ℕ) (f g : coHom (suc n) (S₊ (suc n))) + → coHom (suc n) (S₊ (suc n)) +multHⁿSⁿ n = ST.rec2 squash₂ (λ f g → ∣ premultHⁿSⁿ n f g ∣₂) + +------------------------- properties ------------------------------ +module multPropsHⁿSⁿ (m : ℕ) where + private + hlevelLem : ∀ {x y : coHomK (suc m)} → isOfHLevel (3 + m) (x ≡ y) + hlevelLem = isOfHLevelPath (3 + m) (isOfHLevelTrunc (3 + m)) _ _ + + cohomImElim : ∀ {ℓ} (n : ℕ) + (P : coHomK (suc n) → Type ℓ) + → ((x : _) → isOfHLevel (3 + n) (P x)) + → (f : S₊ (suc n) → coHomK (suc n)) + → (t : S₊ (suc n)) + → ((r : S₊ (suc n)) → f t ≡ ∣ r ∣ → P ∣ r ∣) + → P (f t) + cohomImElim n P hlev f t ind = l (f t) refl + where + l : (x : _) → f t ≡ x → P x + l = T.elim (λ x → isOfHLevelΠ (3 + n) λ _ → hlev _) ind + + + multHⁿSⁿ-0ₗ : (f : _) → multHⁿSⁿ m (0ₕ (suc m)) f ≡ 0ₕ (suc m) + multHⁿSⁿ-0ₗ = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ + (funExt λ x → cohomImElim m + (λ s → rec₊ (isOfHLevelTrunc (3 + m)) + (λ _ → 0ₖ (suc m)) s ≡ 0ₖ (suc m)) + (λ _ → hlevelLem) + f + x + λ _ _ → refl) + + multHⁿSⁿ-1ₗ : (f : _) → multHⁿSⁿ m (∣ ∣_∣ₕ ∣₂) f ≡ f + multHⁿSⁿ-1ₗ = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ + (funExt λ x → cohomImElim m + (λ s → rec₊ (isOfHLevelTrunc (3 + m)) ∣_∣ₕ s ≡ s) + (λ _ → hlevelLem) + f + x + λ _ _ → refl) + + multHⁿSⁿInvₗ : (f g : _) → multHⁿSⁿ m (-ₕ f) g ≡ -ₕ (multHⁿSⁿ m f g) + multHⁿSⁿInvₗ = ST.elim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ + (funExt λ x → cohomImElim m + (λ gt → rec₊ (isOfHLevelTrunc (3 + m)) + (λ x₁ → -ₖ-syntax (suc m) (f x₁)) gt + ≡ -ₖ (rec₊ (isOfHLevelTrunc (3 + m)) f gt)) + (λ _ → hlevelLem) + g x + λ r s → refl) + + multHⁿSⁿDistrₗ : (f g h : _) + → multHⁿSⁿ m (f +ₕ g) h ≡ (multHⁿSⁿ m f h) +ₕ (multHⁿSⁿ m g h) + multHⁿSⁿDistrₗ = ST.elim3 (λ _ _ _ → isSetPathImplicit) + λ f g h → cong ∣_∣₂ (funExt λ x → cohomImElim m + (λ ht → rec₊ (isOfHLevelTrunc (3 + m)) (λ x → f x +ₖ g x) ht + ≡ rec₊ (isOfHLevelTrunc (3 + m)) f ht + +ₖ rec₊ (isOfHLevelTrunc (3 + m)) g ht) + (λ _ → hlevelLem) h x λ _ _ → refl) + +open multPropsHⁿSⁿ +multHⁿSⁿ-presℤ·pos : (m : ℕ) (a : ℕ) (f g : _) + → multHⁿSⁿ m ((pos a) ℤ[ (coHomGr (suc m) (S₊ (suc m))) ]· f) g + ≡ (pos a) ℤ[ (coHomGr (suc m) (S₊ (suc m))) ]· (multHⁿSⁿ m f g) +multHⁿSⁿ-presℤ·pos m zero f g = multHⁿSⁿ-0ₗ m g +multHⁿSⁿ-presℤ·pos m (suc a) f g = + multHⁿSⁿDistrₗ _ f (((pos a) ℤ[ (coHomGr (suc m) (S₊ (suc m))) ]· f)) g + ∙ cong (multHⁿSⁿ m f g +ₕ_) (multHⁿSⁿ-presℤ·pos m a f g) + +multHⁿSⁿ-presℤ· : (m : ℕ) (a : ℤ) (f g : _) + → multHⁿSⁿ m (a ℤ[ (coHomGr (suc m) (S₊ (suc m))) ]· f) g + ≡ a ℤ[ (coHomGr (suc m) (S₊ (suc m))) ]· (multHⁿSⁿ m f g) +multHⁿSⁿ-presℤ· m (pos a) = multHⁿSⁿ-presℤ·pos m a +multHⁿSⁿ-presℤ· m (negsuc nn) f g = + (λ i → multHⁿSⁿ m (ℤ·-negsuc (coHomGr (suc m) (S₊ (suc m))) nn f i) g) + ∙∙ multHⁿSⁿInvₗ m (pos (suc nn) ℤ[ coHomGr (suc m) (S₊ (suc m)) ]· f) g + ∙ cong -ₕ_ (multHⁿSⁿ-presℤ·pos m (suc nn) f g) + ∙∙ sym (ℤ·-negsuc (coHomGr (suc m) (S₊ (suc m)) ) nn (multHⁿSⁿ m f g)) + +Hⁿ-Sⁿ≅ℤ-pres· : (n : ℕ) (f g : _) + → Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n)) (multHⁿSⁿ n f g) + ≡ Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n)) f ·ℤ Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n)) g +Hⁿ-Sⁿ≅ℤ-pres· n f g = + cong ϕ + (cong₂ (multHⁿSⁿ n) (sym (repl f)) (sym (repl g)) + ∙ multHⁿSⁿ-presℤ· n (ϕ f) (∣ ∣_∣ₕ ∣₂) (ϕ g ℤ[ H ]· ∣ ∣_∣ₕ ∣₂)) + ∙ (homPresℤ· (_ , snd (Hⁿ-Sⁿ≅ℤ n)) + (multHⁿSⁿ n ∣ ∣_∣ₕ ∣₂ (ϕ g ℤ[ H ]· ∣ ∣_∣ₕ ∣₂)) (ϕ f) + ∙ sym (ℤ·≡· (ϕ f) _)) + ∙ cong (ϕ f ·ℤ_) + (cong ϕ (multHⁿSⁿ-1ₗ n (ϕ g ℤ[ H ]· ∣ ∣_∣ₕ ∣₂)) + ∙ homPresℤ· (_ , snd (Hⁿ-Sⁿ≅ℤ n)) ∣ ∣_∣ₕ ∣₂ (ϕ g) + ∙ sym (ℤ·≡· (ϕ g) _) + ∙ cong (ϕ g ·ℤ_) (HⁿSⁿ-gen n) + ∙ ·Comm (ϕ g) 1) + where + ϕ = Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n)) + ϕ⁻ = Iso.inv (fst (Hⁿ-Sⁿ≅ℤ n)) + + H = coHomGr (suc n) (S₊ (suc n)) + + repl : (f : H .fst) → (ϕ f ℤ[ H ]· ∣ ∣_∣ₕ ∣₂) ≡ f + repl f = sym (Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ n)) _) + ∙∙ cong ϕ⁻ lem + ∙∙ Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ n)) f + where + lem : ϕ (ϕ f ℤ[ H ]· ∣ ∣_∣ₕ ∣₂) ≡ ϕ f + lem = homPresℤ· (_ , snd (Hⁿ-Sⁿ≅ℤ n)) ∣ ∣_∣ₕ ∣₂ (ϕ f) + ∙ sym (ℤ·≡· (ϕ f) (fst (Hⁿ-Sⁿ≅ℤ n) .Iso.fun ∣ (λ a → ∣ a ∣) ∣₂)) + ∙ cong (ϕ f ·ℤ_) (HⁿSⁿ-gen n) + ∙ ·Comm (ϕ f) 1 + + -------------- A nice packaging for the Hⁿ-Sⁿ ---------------- code : (n m : ℕ) → Type ℓ-zero From b6941207463c57b23399f97d1709f0d635845e16 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 26 Feb 2024 16:42:26 +0100 Subject: [PATCH 23/73] stuff --- Cubical/Algebra/ChainComplex/Base.agda | 266 ++++++++++--------- Cubical/Algebra/ChainComplex/Natindexed.agda | 64 ++--- Cubical/Structures/Successor.agda | 26 ++ 3 files changed, 196 insertions(+), 160 deletions(-) diff --git a/Cubical/Algebra/ChainComplex/Base.agda b/Cubical/Algebra/ChainComplex/Base.agda index 7eb365182a..e5d6d904f9 100644 --- a/Cubical/Algebra/ChainComplex/Base.agda +++ b/Cubical/Algebra/ChainComplex/Base.agda @@ -25,39 +25,51 @@ open import Cubical.HITs.SetQuotients.Properties as SQ open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.Structures.Successor +open import Cubical.Relation.Nullary + private variable ℓ ℓ' ℓ'' : Level -module CC {ℓ₀ : Level} (I : SuccStr ℓ₀) where - open SuccStr I +module CC {ℓ₀ : Level} (I : InjSuccStr ℓ₀) where + open InjSuccStr I + open SuccStr succStr + + ≠succ : Index → Type ℓ₀ + ≠succ i = ¬ (i ≡ succ i) + + Index⇃ : Type ℓ₀ + Index⇃ = Σ[ i ∈ Index ] (≠succ (succ (succ i))) + record ChainComplex (ℓ : Level) : Type (ℓ-suc (ℓ-max ℓ ℓ₀)) where field - chain : (i : Index) → AbGroup ℓ - bdry : (i : Index) → AbGroupHom (chain (succ i)) (chain i) - bdry²=0 : (i : Index) → compGroupHom (bdry (succ i)) (bdry i) ≡ trivGroupHom + chain : (i : Index) → ≠succ i → AbGroup ℓ + bdry : (i : Index) (p : ≠succ (succ i)) → AbGroupHom (chain (succ i) p) (chain i (semiInj p)) + bdry²=0 : (i : Index) (p : ≠succ (succ (succ i))) + → compGroupHom (bdry (succ i) p) (bdry i (semiInj p)) ≡ trivGroupHom record ChainComplexMap {ℓ ℓ' : Level} (A : ChainComplex ℓ) (B : ChainComplex ℓ') : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ₀) where open ChainComplex field - chainmap : (i : Index) → AbGroupHom (chain A i) (chain B i) - bdrycomm : (i : Index) → compGroupHom (chainmap (succ i)) (bdry B i) - ≡ compGroupHom (bdry A i) (chainmap i) + chainmap : (i : Index) (p : ≠succ i) → AbGroupHom (chain A i p) (chain B i p) + bdrycomm : (i : Index) (p : ≠succ (succ i)) + → compGroupHom (chainmap (succ i) p) (bdry B i p) + ≡ compGroupHom (bdry A i p) (chainmap i (semiInj p)) record ChainHomotopy {ℓ : Level} {A : ChainComplex ℓ} {B : ChainComplex ℓ'} (f g : ChainComplexMap A B) : Type (ℓ-max (ℓ-max ℓ' ℓ) ℓ₀) where open ChainComplex open ChainComplexMap field - htpy : (i : Index) → AbGroupHom (chain A i) (chain B (succ i)) - bdryhtpy : (i : Index) - → subtrGroupHom (chain A (succ i)) (chain B (succ i)) - (chainmap f (succ i)) (chainmap g (succ i)) - ≡ addGroupHom (chain A (succ i)) (chain B (succ i)) - (compGroupHom (htpy (succ i)) (bdry B (succ i))) - (compGroupHom (bdry A i) (htpy i)) + htpy : (i : Index) (p : ≠succ (succ i)) → AbGroupHom (chain A i (semiInj p)) (chain B (succ i) p) + bdryhtpy : (i : Index) (p : ≠succ (succ (succ i))) + → subtrGroupHom (chain A (succ i) (semiInj p)) (chain B (succ i) (semiInj p)) + (chainmap f (succ i) (semiInj p)) (chainmap g (succ i)(semiInj p)) + ≡ addGroupHom (chain A (succ i) (semiInj p)) (chain B (succ i) (semiInj p)) + (compGroupHom (htpy (succ i) p) (bdry B (succ i) p)) + (compGroupHom (bdry A i (semiInj p)) (htpy i (semiInj p))) record CoChainComplex (ℓ : Level) : Type (ℓ-max (ℓ-suc ℓ) ℓ₀) where field @@ -71,14 +83,14 @@ module CC {ℓ₀ : Level} (I : SuccStr ℓ₀) where ChainComplexMap≡ : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {f g : ChainComplexMap A B} - → ((i : Index) → chainmap f i ≡ chainmap g i) + → ((i : Index) (j : _) → chainmap f i j ≡ chainmap g i j) → f ≡ g - chainmap (ChainComplexMap≡ p i) n = p n i - bdrycomm (ChainComplexMap≡ {A = A} {B = B} {f = f} {g = g} p i) n = - isProp→PathP {B = λ i → compGroupHom (p (succ n) i) (ChainComplex.bdry B n) - ≡ compGroupHom (ChainComplex.bdry A n) (p n i)} + chainmap (ChainComplexMap≡ p i) n j = p n j i + bdrycomm (ChainComplexMap≡ {A = A} {B = B} {f = f} {g = g} p i) n r = + isProp→PathP {B = λ i → compGroupHom (p (succ n) r i) (ChainComplex.bdry B n r) + ≡ compGroupHom (ChainComplex.bdry A n r) (p n (semiInj r) i)} (λ i → isSetGroupHom _ _) - (bdrycomm f n) (bdrycomm g n) i + (bdrycomm f n r) (bdrycomm g n r) i compChainMap : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {C : ChainComplex ℓ''} → (f : ChainComplexMap A B) (g : ChainComplexMap B C) @@ -88,42 +100,42 @@ module CC {ℓ₀ : Level} (I : SuccStr ℓ₀) where record { chainmap = ψ ; bdrycomm = commψ } = main where main : ChainComplexMap A C - chainmap main n = compGroupHom (ϕ n) (ψ n) - bdrycomm main n = + chainmap main n r = compGroupHom (ϕ n r) (ψ n r) + bdrycomm main n r = Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x - → (funExt⁻ (cong fst (commψ n)) (ϕ (succ n) .fst x)) - ∙ cong (fst (ψ n)) (funExt⁻ (cong fst (commϕ n)) x)) + → (funExt⁻ (cong fst (commψ n r)) (ϕ (succ n) r .fst x)) + ∙ cong (fst (ψ n (semiInj r))) (funExt⁻ (cong fst (commϕ n r)) x)) isChainEquiv : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} → ChainComplexMap A B → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ₀) - isChainEquiv f = ((n : Index) → isEquiv (chainmap f n .fst)) + isChainEquiv f = ((n : Index) (p : _) → isEquiv (chainmap f n p .fst)) _≃Chain_ : (A : ChainComplex ℓ) (B : ChainComplex ℓ') → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ₀) A ≃Chain B = Σ[ f ∈ ChainComplexMap A B ] (isChainEquiv f) idChainMap : (A : ChainComplex ℓ) → ChainComplexMap A A - chainmap (idChainMap A) _ = idGroupHom - bdrycomm (idChainMap A) _ = + chainmap (idChainMap A) _ _ = idGroupHom + bdrycomm (idChainMap A) _ _ = Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl invChainMap : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} → (A ≃Chain B) → ChainComplexMap B A - chainmap (invChainMap (record { chainmap = ϕ ; bdrycomm = ϕcomm } , eq)) n = - GroupEquiv→GroupHom (invGroupEquiv ((ϕ n .fst , eq n) , snd (ϕ n))) - bdrycomm (invChainMap {B = B} (record { chainmap = ϕ ; bdrycomm = ϕcomm } , eq)) n = + chainmap (invChainMap (record { chainmap = ϕ ; bdrycomm = ϕcomm } , eq)) n p = + GroupEquiv→GroupHom (invGroupEquiv ((ϕ n p .fst , eq n p) , snd (ϕ n p))) + bdrycomm (invChainMap {B = B} (record { chainmap = ϕ ; bdrycomm = ϕcomm } , eq)) n p = Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x - → sym (retEq (_ , eq n) _) - ∙∙ cong (invEq (_ , eq n)) - (sym (funExt⁻ (cong fst (ϕcomm n)) (invEq (_ , eq (succ n)) x))) - ∙∙ cong (invEq (ϕ n .fst , eq n) ∘ fst (ChainComplex.bdry B n)) - (secEq (_ , eq (succ n)) x)) + → sym (retEq (_ , eq n (semiInj p)) _) + ∙∙ cong (invEq (_ , eq n (semiInj p))) + (sym (funExt⁻ (cong fst (ϕcomm n p)) (invEq (_ , eq (succ n) p) x))) + ∙∙ cong (invEq (ϕ n (semiInj p) .fst , eq n (semiInj p)) ∘ fst (ChainComplex.bdry B n p)) + (secEq (_ , eq (succ n) p) x)) invChainEquiv : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} → A ≃Chain B → B ≃Chain A fst (invChainEquiv e) = invChainMap e - snd (invChainEquiv e) n = snd (invEquiv (chainmap (fst e) n .fst , snd e n)) + snd (invChainEquiv e) n p = snd (invEquiv (chainmap (fst e) n p .fst , snd e n p)) -- TODO: upstream these module _ {G H : Group ℓ} (ϕ : GroupHom G H) where @@ -136,12 +148,12 @@ module CC {ℓ₀ : Level} (I : SuccStr ℓ₀) where open ChainComplex open IsGroupHom - homology : (n : Index) → ChainComplex ℓ → Group ℓ - homology n C = ker∂n / img∂+1⊂ker∂n + homology : (n : Index⇃) → ChainComplex ℓ → Group ℓ + homology (n , p) C = ker∂n / img∂+1⊂ker∂n where - Cn+2 = AbGroup→Group (chain C (succ (succ n))) - ∂n = bdry C n - ∂n+1 = bdry C (succ n) + Cn+2 = AbGroup→Group (chain C (succ (succ n)) p) + ∂n = bdry C n (semiInj p) + ∂n+1 = bdry C (succ n) p ker∂n = kerGroup ∂n -- Restrict ∂n+1 to ker∂n @@ -151,7 +163,7 @@ module CC {ℓ₀ : Level} (I : SuccStr ℓ₀) where where opaque t : ⟨ fst (kerSubgroup ∂n) (∂n+1 .fst x) ⟩ - t = funExt⁻ (cong fst (bdry²=0 C n)) x + t = funExt⁻ (cong fst (bdry²=0 C n p)) x ∂' : GroupHom Cn+2 ker∂n fst ∂' = ∂'-fun @@ -167,7 +179,7 @@ module CC {ℓ₀ : Level} (I : SuccStr ℓ₀) where snd img∂+1⊂ker∂n = isNormalImSubGroup where opaque - module C1 = AbGroupStr (chain C (succ n) .snd) + module C1 = AbGroupStr (chain C (succ n) (semiInj p) .snd) isNormalImSubGroup : isNormal (imSubgroup ∂') isNormalImSubGroup = isNormalIm ∂' (λ x y → kerGroup≡ ∂n (C1.+Comm (fst x) (fst y))) @@ -181,37 +193,38 @@ module CC {ℓ₀ : Level} (I : SuccStr ℓ₀) where chainComplexMap→HomologyMap : {C D : ChainComplex ℓ} → (ϕ : ChainComplexMap C D) - → (n : Index) + → (n : Index⇃) → GroupHom (homology n C) (homology n D) - chainComplexMap→HomologyMap {C = C} {D} - record { chainmap = ϕ ; bdrycomm = ϕcomm } n = main + chainComplexMap→HomologyMap {C = C} {D} mp (n , p) = main where - fun : homology n C .fst → homology n D .fst + ϕ = chainmap mp + ϕcomm = bdrycomm mp + fun : homology (n , p) C .fst → homology (n , p) D .fst fun = SQ.elim (λ _ → squash/) f - λ f g → PT.rec (GroupStr.is-set (homology n D .snd) _ _ ) (λ r - → eq/ _ _ ∣ (fst (ϕ (succ (succ n))) (fst r)) - , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) - (funExt⁻ (cong fst (ϕcomm (succ n))) (fst r) - ∙∙ cong (fst (ϕ (succ n))) (cong fst (snd r)) - ∙∙ IsGroupHom.pres· (snd (ϕ (succ n))) _ _ - ∙ cong₂ (AbGroupStr._+_ (snd (chain D (succ n)))) + λ f g → PT.rec (GroupStr.is-set (homology (n , p) D .snd) _ _ ) (λ r + → eq/ _ _ ∣ (fst (ϕ (succ (succ n)) p) (fst r)) + , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) + (funExt⁻ (cong fst (ϕcomm (succ n) p)) (fst r) + ∙∙ cong (fst (ϕ (succ n) (semiInj p))) (cong fst (snd r)) + ∙∙ IsGroupHom.pres· (snd (ϕ (succ n) (semiInj p))) _ _ + ∙ cong₂ (AbGroupStr._+_ (snd (chain D (succ n) (semiInj p)))) refl - (IsGroupHom.presinv (snd (ϕ (succ n))) _)) ∣₁) + (IsGroupHom.presinv (snd (ϕ (succ n) (semiInj p))) _)) ∣₁) where - f : _ → homology n D .fst - f (a , b) = [ (ϕ (succ n) .fst a) - , ((λ i → fst (ϕcomm n i) a) - ∙∙ cong (fst (ϕ n)) b - ∙∙ IsGroupHom.pres1 (snd (ϕ n))) ] + f : _ → homology (n , p) D .fst + f (a , b) = [ (ϕ (succ n) (semiInj p) .fst a) + , ((λ i → fst (ϕcomm n (semiInj p) i) a) + ∙∙ cong (fst (ϕ n (semiInj (semiInj p)))) b + ∙∙ IsGroupHom.pres1 (snd (ϕ n (semiInj (semiInj p))))) ] - main : GroupHom (homology n C) (homology n D) + main : GroupHom (homology (n , p) C) (homology (n , p) D) fst main = fun snd main = makeIsGroupHom - (SQ.elimProp2 (λ _ _ → GroupStr.is-set (snd (homology n D)) _ _) + (SQ.elimProp2 (λ _ _ → GroupStr.is-set (snd (homology (n , p) D)) _ _) λ a b → cong [_] - (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) - (IsGroupHom.pres· (snd (ϕ (succ n))) _ _))) + (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) + (IsGroupHom.pres· (snd (ϕ (succ n) (semiInj p))) _ _))) -- chainComplexMap→HomologyMap : {C D : ChainComplex ℓ} -- → (ϕ : ChainComplexMap C D) @@ -223,145 +236,152 @@ module CC {ℓ₀ : Level} (I : SuccStr ℓ₀) where chainComplexMap→HomologyMapComp : {C D E : ChainComplex ℓ} → (ϕ : ChainComplexMap C D) (ψ : ChainComplexMap D E) - → (n : Index) + → (n : Index⇃) → chainComplexMap→HomologyMap (compChainMap ϕ ψ) n ≡ compGroupHom (chainComplexMap→HomologyMap ϕ n) (chainComplexMap→HomologyMap ψ n) chainComplexMap→HomologyMapComp {E = E} record { chainmap = ϕ ; bdrycomm = commϕ } - record { chainmap = ψ ; bdrycomm = commψ } n = + record { chainmap = ψ ; bdrycomm = commψ } (n , p) = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology n E)) _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (n , p) E)) _ _) λ _ → cong [_] - (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain E n)) _ _) refl))) + (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain E n (semiInj (semiInj p)))) _ _) refl))) - chainComplexMap→HomologyMapId : {C : ChainComplex ℓ} (n : Index) + chainComplexMap→HomologyMapId : {C : ChainComplex ℓ} (n : Index⇃) → chainComplexMap→HomologyMap (idChainMap C) n ≡ idGroupHom - chainComplexMap→HomologyMapId {C = C} n = + chainComplexMap→HomologyMapId {C = C} (n , p) = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology n C)) _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (n , p) C)) _ _) λ _ → cong [_] - (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C n)) _ _) refl))) + (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C n (semiInj (semiInj p)))) _ _) refl))) ChainHomotopy→HomologyPath : {A B : ChainComplex ℓ} (f g : ChainComplexMap A B) → ChainHomotopy f g - → (n : Index) → chainComplexMap→HomologyMap f n - ≡ chainComplexMap→HomologyMap g n - ChainHomotopy→HomologyPath {A = A} {B = B} f g ϕ n = + → (n : Index⇃) → chainComplexMap→HomologyMap f n + ≡ chainComplexMap→HomologyMap g n + ChainHomotopy→HomologyPath {A = A} {B = B} f g ϕ (n , r) = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology n _)) _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (n , r) _)) _ _) λ {(a , p) → eq/ _ _ - ∣ (ChainHomotopy.htpy ϕ (succ n) .fst a) - , (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain B n)) _ _) - (sym ((funExt⁻ (cong fst (ChainHomotopy.bdryhtpy ϕ n)) a) + ∣ (ChainHomotopy.htpy ϕ (succ n) r .fst a) + , (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain B n (semiInj (semiInj r)))) _ _) + (sym ((funExt⁻ (cong fst (ChainHomotopy.bdryhtpy ϕ n r)) a) ∙ cong₂ _+B_ refl - (cong (fst (ChainHomotopy.htpy ϕ n)) p - ∙ IsGroupHom.pres1 (snd (ChainHomotopy.htpy ϕ n))) - ∙ AbGroupStr.+IdR (snd (chain B (succ n))) _))) ∣₁})) + (cong (fst (ChainHomotopy.htpy ϕ n (semiInj r))) p + ∙ IsGroupHom.pres1 (snd (ChainHomotopy.htpy ϕ n (semiInj r)))) + ∙ AbGroupStr.+IdR (snd (chain B (succ n) (semiInj r))) _))) ∣₁})) where - open GroupTheory (AbGroup→Group (chain B (succ n))) - invB = GroupStr.inv (snd (AbGroup→Group (chain B (succ n)))) - _+B_ = AbGroupStr._+_ (snd (chain B (succ n))) + open GroupTheory (AbGroup→Group (chain B (succ n) (semiInj r))) + invB = GroupStr.inv (snd (AbGroup→Group (chain B (succ n) (semiInj r)))) + _+B_ = AbGroupStr._+_ (snd (chain B (succ n) (semiInj r))) chainComplexEquiv→HomoglogyIso : {C D : ChainComplex ℓ} (f : C ≃Chain D) - → (n : Index) → GroupIso (homology n C) (homology n D) + → (n : Index⇃) → GroupIso (homology n C) (homology n D) Iso.fun (fst (chainComplexEquiv→HomoglogyIso (f , eq) n)) = chainComplexMap→HomologyMap f n .fst Iso.inv (fst (chainComplexEquiv→HomoglogyIso (f , eq) n)) = chainComplexMap→HomologyMap (invChainMap (f , eq)) n .fst - Iso.rightInv (fst (chainComplexEquiv→HomoglogyIso (f , eq) n)) = + Iso.rightInv (fst (chainComplexEquiv→HomoglogyIso (f , eq) (n , p))) = funExt⁻ (cong fst (sym (chainComplexMap→HomologyMapComp - (invChainMap (f , eq)) f n)) - ∙∙ cong (λ f → fst (chainComplexMap→HomologyMap f n)) - (ChainComplexMap≡ - (λ n → Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (secEq (_ , eq n))))) - ∙∙ cong fst (chainComplexMap→HomologyMapId n)) + (invChainMap (f , eq)) f (n , p))) + ∙∙ cong (λ f → fst (chainComplexMap→HomologyMap f (n , p))) + (ChainComplexMap≡ λ r i + → Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (secEq (_ , eq r i)))) + ∙∙ cong fst (chainComplexMap→HomologyMapId (n , p))) + Iso.leftInv (fst (chainComplexEquiv→HomoglogyIso (f , eq) n)) = funExt⁻ (cong fst (sym (chainComplexMap→HomologyMapComp f (invChainMap (f , eq)) n)) ∙∙ cong (λ f → fst (chainComplexMap→HomologyMap f n)) (ChainComplexMap≡ - (λ n → Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (retEq (_ , eq n))))) + (λ n i → Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (retEq (_ , eq n i))))) ∙∙ cong fst (chainComplexMap→HomologyMapId n)) snd (chainComplexEquiv→HomoglogyIso (f , eq) n) = chainComplexMap→HomologyMap f n .snd -- More general version - homologyIso : (n : Index) (C D : ChainComplex ℓ) - → (chEq₂ : AbGroupIso (chain C (succ (succ n))) (chain D (succ (succ n)))) - → (chEq₁ : AbGroupIso (chain C (succ n)) (chain D (succ n))) - → (chEq₀ : AbGroupIso (chain C n) (chain D n)) - → Iso.fun (chEq₀ .fst) ∘ bdry C n .fst - ≡ bdry D n .fst ∘ Iso.fun (chEq₁ .fst) - → Iso.fun (chEq₁ .fst) ∘ bdry C (succ n) .fst - ≡ bdry D (succ n) .fst ∘ Iso.fun (chEq₂ .fst) + homologyIso : (n : Index⇃) (C D : ChainComplex ℓ) + → (chEq₂ : AbGroupIso (chain C (succ (succ (fst n))) (snd n)) + (chain D (succ (succ (fst n))) (snd n))) + → (chEq₁ : AbGroupIso (chain C (succ (fst n)) (semiInj (snd n))) + (chain D (succ (fst n)) (semiInj (snd n)))) + → (chEq₀ : AbGroupIso (chain C (fst n) (semiInj (semiInj (snd n)))) + (chain D (fst n) (semiInj (semiInj (snd n))))) + → Iso.fun (chEq₀ .fst) ∘ bdry C (fst n) (semiInj (snd n)) .fst + ≡ bdry D (fst n) (semiInj (snd n)) .fst ∘ Iso.fun (chEq₁ .fst) + → Iso.fun (chEq₁ .fst) ∘ bdry C (succ (fst n)) (snd n) .fst + ≡ bdry D (succ (fst n)) (snd n) .fst ∘ Iso.fun (chEq₂ .fst) → GroupIso (homology n C) (homology n D) - homologyIso n C D chEq₂ chEq₁ chEq₀ eq1 eq2 = main-iso + homologyIso (n , p) C D chEq₂ chEq₁ chEq₀ eq1 eq2 = main-iso where - F : homology n C .fst → homology n D .fst + F : homology (n , p) C .fst → homology (n , p) D .fst F = SQ.elim (λ _ → squash/) f λ a b r → eq/ _ _ (PT.map (λ { (s , t) → (Iso.fun (chEq₂ .fst) s) - , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) + , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) (sym (funExt⁻ eq2 s) ∙ cong (Iso.fun (chEq₁ .fst)) (cong fst t) ∙ IsGroupHom.pres· (chEq₁ .snd) _ _ - ∙ cong₂ (snd (chain D (succ n)) .AbGroupStr._+_) + ∙ cong₂ (snd (chain D (succ n) (semiInj p)) .AbGroupStr._+_) refl (IsGroupHom.presinv (chEq₁ .snd) _))}) r) where - f : _ → homology n D .fst + f : _ → homology (n , p) D .fst f (a , b) = [ Iso.fun (fst chEq₁) a , sym (funExt⁻ eq1 a) ∙ cong (Iso.fun (chEq₀ .fst)) b ∙ IsGroupHom.pres1 (snd chEq₀) ] - G : homology n D .fst → homology n C .fst + G : homology (n , p) D .fst → homology (n , p) C .fst G = SQ.elim (λ _ → squash/) g λ a b r → eq/ _ _ (PT.map (λ {(s , t) → (Iso.inv (chEq₂ .fst) s) - , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C n)) _ _) + , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C n (semiInj (semiInj p)))) _ _) (sym (Iso.leftInv (chEq₁ .fst) _) ∙ cong (Iso.inv (chEq₁ .fst)) (funExt⁻ eq2 (Iso.inv (chEq₂ .fst) s)) - ∙ cong (Iso.inv (chEq₁ .fst) ∘ bdry D (succ n) .fst) + ∙ cong (Iso.inv (chEq₁ .fst) ∘ bdry D (succ n) p .fst) (Iso.rightInv (chEq₂ .fst) s) ∙ cong (Iso.inv (chEq₁ .fst)) (cong fst t) ∙ IsGroupHom.pres· (invGroupIso chEq₁ .snd) _ _ - ∙ cong₂ (snd (chain C (succ n)) .AbGroupStr._+_) + ∙ cong₂ (snd (chain C (succ n) (semiInj p)) .AbGroupStr._+_) refl ((IsGroupHom.presinv (invGroupIso chEq₁ .snd) _)))}) r) where - g : _ → homology n C .fst + g : _ → homology (n , p) C .fst g (a , b) = [ Iso.inv (fst chEq₁) a , sym (Iso.leftInv (chEq₀ .fst) _) ∙ cong (Iso.inv (chEq₀ .fst)) (funExt⁻ eq1 (Iso.inv (chEq₁ .fst) a)) - ∙ cong (Iso.inv (chEq₀ .fst) ∘ bdry D n .fst) + ∙ cong (Iso.inv (chEq₀ .fst) ∘ bdry D n (semiInj p) .fst) (Iso.rightInv (chEq₁ .fst) a) ∙ cong (Iso.inv (chEq₀ .fst)) b ∙ IsGroupHom.pres1 (invGroupIso chEq₀ .snd) ] - F-hom : IsGroupHom (homology n C .snd) F (homology n D .snd) + F-hom : IsGroupHom (homology (n , p) C .snd) F (homology (n , p) D .snd) F-hom = makeIsGroupHom - (SQ.elimProp2 (λ _ _ → GroupStr.is-set (homology n D .snd) _ _) + (SQ.elimProp2 (λ _ _ → GroupStr.is-set (homology (n , p) D .snd) _ _) λ {(a , s) (b , t) - → cong [_] (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) + → cong [_] (Σ≡Prop (λ _ + → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) (IsGroupHom.pres· (snd chEq₁) _ _)) }) - main-iso : GroupIso (homology n C) (homology n D) + main-iso : GroupIso (homology (n , p) C) (homology (n , p) D) Iso.fun (fst main-iso) = F Iso.inv (fst main-iso) = G Iso.rightInv (fst main-iso) = - elimProp (λ _ → GroupStr.is-set (homology n D .snd) _ _) + elimProp (λ _ → GroupStr.is-set (homology (n , p) D .snd) _ _) λ{(a , b) - → cong [_] (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) + → cong [_] (Σ≡Prop (λ _ + → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) (Iso.rightInv (fst chEq₁) a))} Iso.leftInv (fst main-iso) = - elimProp (λ _ → GroupStr.is-set (homology n C .snd) _ _) + elimProp (λ _ → GroupStr.is-set (homology (n , p) C .snd) _ _) λ{(a , b) - → cong [_] (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C n)) _ _) + → cong [_] (Σ≡Prop (λ _ + → AbGroupStr.is-set (snd (chain C n (semiInj (semiInj p)))) _ _) (Iso.leftInv (fst chEq₁) a))} snd main-iso = F-hom diff --git a/Cubical/Algebra/ChainComplex/Natindexed.agda b/Cubical/Algebra/ChainComplex/Natindexed.agda index 28ad012313..592a24d723 100644 --- a/Cubical/Algebra/ChainComplex/Natindexed.agda +++ b/Cubical/Algebra/ChainComplex/Natindexed.agda @@ -6,6 +6,8 @@ open import Cubical.Foundations.Prelude open import Cubical.Data.Nat open import Cubical.Data.Fin open import Cubical.Data.Sigma +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ open import Cubical.Structures.Successor @@ -15,66 +17,54 @@ open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Relation.Nullary private variable ℓ : Level + -- Nat indexed chain complexes -module CCℕ = CC ℕ+ +module CCℕ = CC ℕ+Inj -- Restriction of ℕ indexed chain complexes to Fin n indexed ones module _ (n : ℕ) where open CC - module CCFin = CC (Fin+ n) + -- Fin n indexed chain complexes + module CCFin = CC (Fin+Inj n) open ChainComplex open ChainComplexMap open ChainHomotopy ChainComplex→finChainComplex : CCℕ.ChainComplex ℓ → CCFin.ChainComplex ℓ - chain (ChainComplex→finChainComplex C) (x , zero , p) = trivialAbGroup - chain (ChainComplex→finChainComplex C) (x , suc diff , p) = chain C x - bdry (ChainComplex→finChainComplex C) (x , zero , p) = trivGroupHom - bdry (ChainComplex→finChainComplex C) (x , suc zero , p) = trivGroupHom - bdry (ChainComplex→finChainComplex C) (x , suc (suc diff) , p) = bdry C x - bdry²=0 (ChainComplex→finChainComplex C) (x , zero , p) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl - bdry²=0 (ChainComplex→finChainComplex C) (x , suc zero , p) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl - bdry²=0 (ChainComplex→finChainComplex C) (x , suc (suc zero) , p) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt λ y → IsGroupHom.pres1 (snd (bdry C x))) - bdry²=0 (ChainComplex→finChainComplex C) (x , suc (suc (suc diff)) , p) = bdry²=0 C x + chain (ChainComplex→finChainComplex C) i p = chain C (fst i) ≠suc + bdry (ChainComplex→finChainComplex C) (i , zero , q) p = ⊥.rec (p refl) + bdry (ChainComplex→finChainComplex C) (i , suc diff , q) p = bdry C i ≠suc + bdry²=0 (ChainComplex→finChainComplex C) (i , zero , q) p = ⊥.rec (p refl) + bdry²=0 (ChainComplex→finChainComplex C) (i , suc zero , q) p = ⊥.rec (p refl) + bdry²=0 (ChainComplex→finChainComplex C) (i , suc (suc diff) , q) p = bdry²=0 C i ≠suc ChainComplexMap→finChainComplexMap : {C D : CCℕ.ChainComplex ℓ} → CCℕ.ChainComplexMap C D → CCFin.ChainComplexMap (ChainComplex→finChainComplex C) (ChainComplex→finChainComplex D) - chainmap (ChainComplexMap→finChainComplexMap - record { chainmap = chainmap ; bdrycomm = bdrycomm }) (x , zero , p) = trivGroupHom - chainmap (ChainComplexMap→finChainComplexMap - record { chainmap = chainmap ; bdrycomm = bdrycomm }) (x , suc diff , p) = chainmap x - bdrycomm (ChainComplexMap→finChainComplexMap - record { chainmap = chainmap ; bdrycomm = bdrycomm }) (x , zero , p) = refl - bdrycomm (ChainComplexMap→finChainComplexMap - record { chainmap = chainmap ; bdrycomm = bdrycomm }) (x , suc zero , p) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt λ c → sym (IsGroupHom.pres1 (chainmap x .snd))) - bdrycomm (ChainComplexMap→finChainComplexMap - record { chainmap = chainmap ; bdrycomm = bdrycomm }) (x , suc (suc diff) , p) = bdrycomm x - + chainmap (ChainComplexMap→finChainComplexMap ϕ) (x , zero , q) p = ⊥.rec (p refl) + chainmap (ChainComplexMap→finChainComplexMap ϕ) (x , suc diff , q) _ = chainmap ϕ x _ + bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , zero , q) p = ⊥.rec (p refl) + bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , suc zero , q) p = ⊥.rec (p refl) + bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , suc (suc diff) , p) q = bdrycomm ϕ x _ + chainHom→ : {C D : CCℕ.ChainComplex ℓ} (f g : CCℕ.ChainComplexMap C D) → CCℕ.ChainHomotopy f g → CCFin.ChainHomotopy (ChainComplexMap→finChainComplexMap f) - (ChainComplexMap→finChainComplexMap g) - ChainHomotopy.htpy (chainHom→ f g - record { htpy = htpy ; bdryhtpy = bdryhtpy }) (x , zero , q) = trivGroupHom - ChainHomotopy.htpy (chainHom→ f g p) (x , suc zero , q) = trivGroupHom - ChainHomotopy.htpy (chainHom→ f g p) (x , suc (suc diff) , q) = htpy p x - bdryhtpy (chainHom→ f g p) (x , zero , q) = refl - bdryhtpy (chainHom→ f g p) (x , suc zero , q) = refl - bdryhtpy (chainHom→ f g p) (x , suc (suc zero) , q) = {!bdryhtpy p x!} - bdryhtpy (chainHom→ f g p) (x , suc (suc (suc diff)) , q) = bdryhtpy p x + (ChainComplexMap→finChainComplexMap g) + htpy (chainHom→ f g h) (x , zero , q) p = ⊥.rec (p refl) + htpy (chainHom→ f g h) (x , suc zero , q) p = ⊥.rec (p refl) + htpy (chainHom→ f g h) (x , suc (suc diff) , q) p = htpy h x _ + bdryhtpy (chainHom→ f g h) (x , zero , q) p = ⊥.rec (p refl) + bdryhtpy (chainHom→ f g h) (x , suc zero , q) p = ⊥.rec (p refl) + bdryhtpy (chainHom→ f g h) (x , suc (suc zero) , q) p = ⊥.rec (p refl) + bdryhtpy (chainHom→ f g h) (x , suc (suc (suc diff)) , q) _ = bdryhtpy h x _ diff --git a/Cubical/Structures/Successor.agda b/Cubical/Structures/Successor.agda index 01055702f2..1e4727c90e 100644 --- a/Cubical/Structures/Successor.agda +++ b/Cubical/Structures/Successor.agda @@ -9,6 +9,9 @@ open import Cubical.Foundations.Prelude open import Cubical.Data.Int open import Cubical.Data.Nat open import Cubical.Data.Fin +open import Cubical.Data.Sigma +open import Cubical.Data.Nat.Order +open import Cubical.Relation.Nullary private variable @@ -33,3 +36,26 @@ Fin+ : (n : ℕ) → SuccStr ℓ-zero Fin+ n .Index = Fin n Fin+ n .succ (x , zero , p) = x , zero , p Fin+ n .succ (x , suc diff , p) = suc x , diff , +-suc diff (suc x) ∙ p + +record InjSuccStr (ℓ : Level) : Type (ℓ-suc ℓ) where + field + succStr : SuccStr ℓ + semiInj : {i : Index succStr} + → ¬ (succ succStr i ≡ succ succStr (succ succStr i)) + → ¬ (i ≡ succ succStr i) + +open InjSuccStr + +≠suc : {n : ℕ} → ¬ n ≡ suc n +≠suc {n = n} p = ¬m Date: Tue, 27 Feb 2024 00:10:56 +0100 Subject: [PATCH 24/73] stuf --- Cubical/Data/CW/Base.agda | 6 + Cubical/Data/CW/ChainComplex.agda | 191 ++++++++++++++++++ Cubical/Data/CW/Map.agda | 87 ++++++++ Cubical/Data/CW/Properties.agda | 1 + Cubical/Data/FinSequence.agda | 5 + Cubical/Data/FinSequence/Base.agda | 23 +++ Cubical/Data/FinSequence/Properties.agda | 39 ++++ Cubical/Data/Sequence.agda | 5 + Cubical/Data/Sequence/Base.agda | 28 +++ Cubical/Data/Sequence/Properties.agda | 22 ++ Cubical/HITs/SequentialColimit/Base.agda | 22 +- .../HITs/SequentialColimit/Properties.agda | 54 ++++- Cubical/Sequence.agda | 0 13 files changed, 470 insertions(+), 13 deletions(-) create mode 100644 Cubical/Data/CW/ChainComplex.agda create mode 100644 Cubical/Data/CW/Map.agda create mode 100644 Cubical/Data/FinSequence.agda create mode 100644 Cubical/Data/FinSequence/Base.agda create mode 100644 Cubical/Data/FinSequence/Properties.agda create mode 100644 Cubical/Data/Sequence.agda create mode 100644 Cubical/Data/Sequence/Base.agda create mode 100644 Cubical/Data/Sequence/Properties.agda create mode 100644 Cubical/Sequence.agda diff --git a/Cubical/Data/CW/Base.agda b/Cubical/Data/CW/Base.agda index b3ca77c06d..6a16eb6b40 100644 --- a/Cubical/Data/CW/Base.agda +++ b/Cubical/Data/CW/Base.agda @@ -21,6 +21,9 @@ open import Cubical.Data.Unit open import Cubical.Data.Fin open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence + open import Cubical.HITs.Sn open import Cubical.HITs.Pushout @@ -85,6 +88,9 @@ realiseSeq : CWskel ℓ → Sequence ℓ Sequence.obj (realiseSeq (C , p)) = C Sequence.map (realiseSeq C) = CW↪ C _ +realiseFinSeq : (m : ℕ) → CWskel ℓ → FinSequence m ℓ +realiseFinSeq m C = Sequence→FinSequence m (realiseSeq C) + -- realisation of CW complex from skeleton realise : CWskel ℓ → Type ℓ realise C = SeqColim (realiseSeq C) diff --git a/Cubical/Data/CW/ChainComplex.agda b/Cubical/Data/CW/ChainComplex.agda new file mode 100644 index 0000000000..464a205d7a --- /dev/null +++ b/Cubical/Data/CW/ChainComplex.agda @@ -0,0 +1,191 @@ +{-# OPTIONS --safe --lossy-unification #-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat +open import Cubical.Data.Fin +open import Cubical.Data.Sigma +open import Cubical.Data.CW + +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.SphereBouquet.Degree + +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.ChainComplex.Natindexed +open import Cubical.Structures.Successor + + +-- In this file, we associate to every CW complex its cellular chain complex +-- The group in dimension n is Z[A_n] where A_n is the number of n-cells +-- The boundary map is defined through a bit of homotopical manipulation. +-- The definition loosely follows May's Concise Course in Alg. Top. + +module Cubical.Data.CW.ChainComplex where + +module _ {ℓ} (C : CWskel ℓ) where + + open CWskel-fields C using (ℤ[A_]) + -- in this module we define pre∂, a homotopical version of the boundary map + -- it goes from V_(A_n+2) S^(n+2) to V_(A_n+1) S^(n+2) + module preboundary (n : ℕ) where + Xn = (fst C n) + Xn+1 = (fst C (suc n)) + An = (snd C .fst n) + An+1 = (snd C .fst (suc n)) + αn = (snd C .snd .fst n) + αn+1 = (snd C .snd .fst (suc n)) + + isoSuspBouquet : Susp (SphereBouquet n (Fin An)) + → SphereBouquet (suc n) (Fin An) + isoSuspBouquet = Iso.fun sphereBouquetSuspIso + + isoSuspBouquetInv : SphereBouquet (suc n) (Fin An) + → Susp (SphereBouquet n (Fin An)) + isoSuspBouquetInv = Iso.inv sphereBouquetSuspIso + + isoSuspBouquet↑ : Susp (SphereBouquet (suc n) (Fin An)) + → SphereBouquet (suc (suc n)) (Fin An) + isoSuspBouquet↑ = Iso.fun sphereBouquetSuspIso + + isoSuspBouquetInv↑ : SphereBouquet (suc (suc n)) (Fin An+1) + → Susp (SphereBouquet (suc n) (Fin An+1)) + isoSuspBouquetInv↑ = Iso.inv sphereBouquetSuspIso + + isoCofBouquet : cofibCW n C → SphereBouquet n (Fin An) + isoCofBouquet = Iso.fun (BouquetIso-gen n An αn (snd C .snd .snd .snd n)) + + isoCofBouquetInv↑ : SphereBouquet (suc n) (Fin An+1) → cofibCW (suc n) C + isoCofBouquetInv↑ = Iso.inv (BouquetIso-gen (suc n) An+1 αn+1 (snd C .snd .snd .snd (suc n))) + + -- our homotopical boundary map + pre∂ : SphereBouquet (suc n) (Fin An+1) → SphereBouquet (suc n) (Fin An) + pre∂ = isoSuspBouquet ∘ (suspFun isoCofBouquet) + ∘ (suspFun (to_cofibCW n C)) ∘ (δ (suc n) C) ∘ isoCofBouquetInv↑ + + -- we define a suspended version + -- we cannot prove that pre∂ ∘ pre∂ ≡ 0, because the dimensions do not match + -- instead, we will prove susp-pre∂ ∘ pre∂ ≡ 0 + susp-pre∂ = (suspFun isoSuspBouquet) ∘ (suspFun (suspFun isoCofBouquet)) + ∘ (suspFun (suspFun (to_cofibCW n C))) ∘ (suspFun (δ (suc n) C)) + ∘ (suspFun isoCofBouquetInv↑) + + -- variant of susp-pre∂ where all the suspensions are collected + pre∂↑ : (SphereBouquet (suc (suc n)) (Fin An+1) + → SphereBouquet (suc (suc n)) (Fin An)) + pre∂↑ = isoSuspBouquet↑ ∘ susp-pre∂ ∘ isoSuspBouquetInv↑ + + -- susp is distributive, so susp-pre∂ ≡ pre∂↑ + susp-pre∂-pre∂↑ : bouquetSusp→ pre∂ ≡ pre∂↑ + susp-pre∂-pre∂↑ = congS (λ X → isoSuspBouquet↑ ∘ X ∘ isoSuspBouquetInv↑) susp-distrib + where + susp-distrib : suspFun pre∂ ≡ susp-pre∂ + susp-distrib i north = north + susp-distrib i south = south + susp-distrib i (merid a i₁) = susp-pre∂ (merid a i₁) + + + -- In this module we prove that (susp pre∂) ∘ pre∂ ≡ 0 + -- from that, we will deduce that ∂ ∘ ∂ ≡ 0 + module preboundary-cancellation (n : ℕ) where + open preboundary + + -- the desired equation comes from exactness of the (Baratt-Puppe) long cofibCW sequence + -- we need some choice to prove this lemma -- which is fine, because we use finite sets + -- this is because the spaces we are dealing with are not pointed + cofibCW-exactness : (δ (suc n) C) ∘ (to_cofibCW (suc n) C) ≡ λ x → north + cofibCW-exactness i x = merid (choose-point x) (~ i) + where + choose-aux : (card : ℕ) (α : Fin card × S₊ n → Xn+1 n) + → Pushout α (λ r → fst r) → Xn+1 n + choose-aux zero α (inl x) = x + choose-aux zero α (inr x) with (¬Fin0 x) + choose-aux zero α (inr x) | () + choose-aux zero α (push (x , _) i) with (¬Fin0 x) + choose-aux zero α (push (x , _) i) | () + choose-aux (suc card') α x = α (fzero , ptSn n) + + choose-point : Xn+1 (suc n) → Xn+1 n + choose-point x = choose-aux (An+1 n) (αn (suc n)) (snd C .snd .snd .snd (suc n) .fst x) + + -- combining the previous result with some isomorphism cancellations + cancellation : suspFun (δ (suc n) C) ∘ suspFun (isoCofBouquetInv↑ n) + ∘ (isoSuspBouquetInv↑ n) ∘ (isoSuspBouquet (suc n)) + ∘ (suspFun (isoCofBouquet (suc n))) ∘ suspFun (to_cofibCW (suc n) C) + ≡ λ x → north + cancellation = cong (λ X → suspFun (δ (suc n) C) ∘ suspFun (isoCofBouquetInv↑ n) + ∘ X ∘ (suspFun (isoCofBouquet (suc n))) + ∘ suspFun (to_cofibCW (suc n) C)) + iso-cancel-2 + ∙∙ cong (λ X → suspFun (δ (suc n) C) ∘ X ∘ suspFun (to_cofibCW (suc n) C)) + iso-cancel-1 + ∙∙ cofibCW-exactness↑ + where + cofibCW-exactness↑ : suspFun (δ (suc n) C) ∘ suspFun (to_cofibCW (suc n) C) + ≡ λ x → north + cofibCW-exactness↑ = sym (suspFunComp _ _) + ∙∙ congS suspFun cofibCW-exactness + ∙∙ suspFunConst north + + iso-cancel-1 : suspFun (isoCofBouquetInv↑ n) ∘ suspFun (isoCofBouquet (suc n)) + ≡ λ x → x + iso-cancel-1 = sym (suspFunComp _ _) + ∙∙ cong suspFun (λ i x → Iso.leftInv + (BouquetIso-gen (suc n) (An+1 n) (αn+1 n) + (snd C .snd .snd .snd (suc n))) x i) + ∙∙ suspFunIdFun + iso-cancel-2 : (isoSuspBouquetInv↑ n) ∘ (isoSuspBouquet (suc n)) ≡ λ x → x + iso-cancel-2 i x = Iso.leftInv sphereBouquetSuspIso x i + + left-maps = (isoSuspBouquet↑ n) ∘ (suspFun (isoSuspBouquet n)) + ∘ (suspFun (suspFun (isoCofBouquet n))) ∘ (suspFun (suspFun (to_cofibCW n C))) + + right-maps = (δ (suc (suc n)) C) ∘ (isoCofBouquetInv↑ (suc n)) + + -- the cancellation result but suspension has been distributed on every map + pre∂↑pre∂≡0 : (pre∂↑ n) ∘ (pre∂ (suc n)) ≡ λ _ → inl tt + pre∂↑pre∂≡0 = congS (λ X → left-maps ∘ X ∘ right-maps) cancellation + + -- we factorise the suspensions + -- and use the fact that a pointed map is constant iff its nonpointed part is constant + pre∂pre∂≡0 : (bouquetSusp→ (pre∂ n)) ∘ (pre∂ (suc n)) ≡ (λ _ → inl tt) + pre∂pre∂≡0 = (congS (λ X → X ∘ (pre∂ (suc n))) (susp-pre∂-pre∂↑ n) ∙ pre∂↑pre∂≡0) + + -- the boundary map of the chain complex associated to C + ∂ : (n : ℕ) → AbGroupHom (ℤ[A (suc n) ]) (ℤ[A n ]) + ∂ n = bouquetDegree (preboundary.pre∂ n) + + -- ∂ ∘ ∂ = 0, the fundamental equation of chain complexes + opaque + ∂∂≡0 : (n : ℕ) → compGroupHom (∂ (suc n)) (∂ n) ≡ trivGroupHom + ∂∂≡0 n = congS (compGroupHom (∂ (suc n))) ∂≡∂↑ + ∙∙ sym (bouquetDegreeComp (bouquetSusp→ (pre∂ n)) (pre∂ (suc n))) + ∙∙ (congS bouquetDegree (preboundary-cancellation.pre∂pre∂≡0 n) + ∙ bouquetDegreeConst _ _ _) + where + open preboundary + + ∂↑ : AbGroupHom (ℤ[A (suc n) ]) (ℤ[A n ]) + ∂↑ = bouquetDegree (bouquetSusp→ (pre∂ n)) + + ∂≡∂↑ : ∂ n ≡ ∂↑ + ∂≡∂↑ = bouquetDegreeSusp (pre∂ n) + + + open CCℕ + open ChainComplex + + CW-ChainComplex : ChainComplex ℓ-zero + chain CW-ChainComplex n _ = ℤ[A n ] + bdry CW-ChainComplex n p = ∂ n + bdry²=0 CW-ChainComplex n p = ∂∂≡0 n + + -- Cellular homology + Hᶜʷ : (n : ℕ) → Group₀ + Hᶜʷ n = homology (n , ≠suc) CW-ChainComplex diff --git a/Cubical/Data/CW/Map.agda b/Cubical/Data/CW/Map.agda new file mode 100644 index 0000000000..fc27786d77 --- /dev/null +++ b/Cubical/Data/CW/Map.agda @@ -0,0 +1,87 @@ +{-# OPTIONS --cubical --safe --lossy-unification #-} + +{-This file contains: + +-} + +module Cubical.Data.CW.Map where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Fin +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence + + +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.Data.CW.Base +open import Cubical.Data.CW.Properties + + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + +open import Cubical.HITs.SequentialColimit +open Sequence + +open import Cubical.Relation.Nullary + +private + variable + ℓ ℓ' ℓ'' : Level + C D E : CWskel ℓ + +-- Maps +cellMap : (C : CWskel ℓ) (D : CWskel ℓ') → Type (ℓ-max ℓ ℓ') +cellMap C D = SequenceMap (realiseSeq C) (realiseSeq D) + +-- Extracting a map between the realisations of the CWskel complexes +realiseCellMap : cellMap C D → realise C → realise D +realiseCellMap mp C∞ = realiseSequenceMap mp C∞ + +-- The identity as a cellular map +idCellMap : (C : CWskel ℓ) → cellMap C C +idCellMap C = idSequenceMap _ + +-- Composition of two cellular maps +composeCellMap : (g : cellMap D E) (f : cellMap C D) → cellMap C E +composeCellMap = composeSequenceMap + + +----- finite versions of above ----- +module _ (m : ℕ) where + finCellMap : (C : CWskel ℓ) (D : CWskel ℓ') → Type (ℓ-max ℓ ℓ') + finCellMap C D = FinSequenceMap (realiseFinSeq m C) (realiseFinSeq m D) + + idFinCellMap : (C : CWskel ℓ) → finCellMap C C + idFinCellMap C = idFinSequenceMap m (realiseFinSeq m C) + + composeFinCellMap : (g : finCellMap D E) (f : finCellMap C D) → finCellMap C E + composeFinCellMap = composeFinSequenceMap m + +finCellMap↓ : {m : ℕ} {C : CWskel ℓ} {D : CWskel ℓ'} → finCellMap (suc m) C D → finCellMap m C D +FinSequenceMap.fmap (finCellMap↓ {m = m} ϕ) x = FinSequenceMap.fmap ϕ (injectSuc x) +FinSequenceMap.fcomm (finCellMap↓ {m = m} {C = C} ϕ) x r = FinSequenceMap.fcomm ϕ (injectSuc x) r + ∙ λ t → FinSequenceMap.fmap ϕ (help t) a + where + a : fst C (suc (fst x)) + a = FinSequence.fmap (realiseFinSeq m C) r + help : fsuc (injectSuc x) ≡ injectSuc (fsuc x) + help = Σ≡Prop (λ _ → isProp≤) refl diff --git a/Cubical/Data/CW/Properties.agda b/Cubical/Data/CW/Properties.agda index 02ee60aa47..8bcc684acb 100644 --- a/Cubical/Data/CW/Properties.agda +++ b/Cubical/Data/CW/Properties.agda @@ -22,6 +22,7 @@ open import Cubical.Data.Fin open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ open import Cubical.Data.CW.Base +open import Cubical.Data.Sequence open import Cubical.HITs.Sn open import Cubical.HITs.Pushout diff --git a/Cubical/Data/FinSequence.agda b/Cubical/Data/FinSequence.agda new file mode 100644 index 0000000000..5d3bd1cc60 --- /dev/null +++ b/Cubical/Data/FinSequence.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.FinSequence where + +open import Cubical.Data.FinSequence.Base public +open import Cubical.Data.FinSequence.Properties public diff --git a/Cubical/Data/FinSequence/Base.agda b/Cubical/Data/FinSequence/Base.agda new file mode 100644 index 0000000000..f5cbc17c67 --- /dev/null +++ b/Cubical/Data/FinSequence/Base.agda @@ -0,0 +1,23 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.FinSequence.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Data.Nat +open import Cubical.Data.Fin + + +private + variable + ℓ ℓ' : Level + +record FinSequence (m : ℕ) (ℓ : Level) : Type (ℓ-suc ℓ) where + field + fobj : Fin (suc m) → Type ℓ + fmap : {n : Fin m} → fobj (injectSuc n) → fobj (fsuc n) + +record FinSequenceMap {m : ℕ} (C : FinSequence m ℓ) (D : FinSequence m ℓ') : Type (ℓ-max ℓ ℓ') where + field + fmap : (n : Fin (suc m)) → FinSequence.fobj C n → FinSequence.fobj D n + fcomm : (n : Fin m) (x : FinSequence.fobj C (injectSuc n)) + → FinSequence.fmap D (fmap (injectSuc n) x) ≡ fmap (fsuc n) (FinSequence.fmap C x) diff --git a/Cubical/Data/FinSequence/Properties.agda b/Cubical/Data/FinSequence/Properties.agda new file mode 100644 index 0000000000..ce42c4a359 --- /dev/null +++ b/Cubical/Data/FinSequence/Properties.agda @@ -0,0 +1,39 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.FinSequence.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Data.Nat +open import Cubical.Data.Fin +open import Cubical.Data.FinSequence.Base +open import Cubical.Data.Sequence + + +private + variable + ℓ ℓ' ℓ'' : Level + +open Sequence +open FinSequence +open FinSequenceMap + +idFinSequenceMap : (m : ℕ) (C : FinSequence m ℓ) → FinSequenceMap C C +fmap (idFinSequenceMap m C) p x = x +fcomm (idFinSequenceMap m C) p _ = refl + +composeFinSequenceMap : (m : ℕ) {C : FinSequence m ℓ} {D : FinSequence m ℓ'} {E : FinSequence m ℓ''} + (g : FinSequenceMap D E) (f : FinSequenceMap C D) → FinSequenceMap C E +fmap (composeFinSequenceMap m g f) n x = fmap g n (fmap f n x) +fcomm (composeFinSequenceMap m g f) n x = + fcomm g n _ + ∙ cong (fmap g (fsuc n)) (fcomm f n x) + +Sequence→FinSequence : (m : ℕ) → Sequence ℓ → FinSequence m ℓ +fobj (Sequence→FinSequence m C) x = obj C (fst x) +fmap (Sequence→FinSequence m C) x = map C x + +SequenceMap→FinSequenceMap : (m : ℕ) (C D : Sequence ℓ) + → SequenceMap C D + → FinSequenceMap (Sequence→FinSequence m C) (Sequence→FinSequence m D) +fmap (SequenceMap→FinSequenceMap m C D ϕ) t = SequenceMap.map ϕ (fst t) +fcomm (SequenceMap→FinSequenceMap m C D ϕ) t = SequenceMap.comm ϕ (fst t) diff --git a/Cubical/Data/Sequence.agda b/Cubical/Data/Sequence.agda new file mode 100644 index 0000000000..1e6c435b12 --- /dev/null +++ b/Cubical/Data/Sequence.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Sequence where + +open import Cubical.Data.Sequence.Base public +open import Cubical.Data.Sequence.Properties public diff --git a/Cubical/Data/Sequence/Base.agda b/Cubical/Data/Sequence/Base.agda new file mode 100644 index 0000000000..fbedcf36b5 --- /dev/null +++ b/Cubical/Data/Sequence/Base.agda @@ -0,0 +1,28 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Sequence.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Data.Nat +open import Cubical.Data.Fin + +private + variable + ℓ ℓ' ℓ'' : Level + +record Sequence (ℓ : Level) : Type (ℓ-suc ℓ) where + field + obj : ℕ → Type ℓ + map : {n : ℕ} → obj n → obj (1 + n) + +record SequenceMap (C : Sequence ℓ) (D : Sequence ℓ') : Type (ℓ-max ℓ ℓ') where + field + map : (n : ℕ) → Sequence.obj C n → Sequence.obj D n + comm : (n : ℕ) (x : Sequence.obj C n) + → Sequence.map D (map n x) ≡ map (suc n) (Sequence.map C x) + +converges : ∀ {ℓ} → Sequence ℓ → (n : ℕ) → Type ℓ +converges seq n = (k : ℕ) → isEquiv (Sequence.map seq {n = k + n}) + +finiteSequence : (ℓ : Level) (m : ℕ) → Type (ℓ-suc ℓ) +finiteSequence ℓ m = Σ[ S ∈ Sequence ℓ ] converges S m diff --git a/Cubical/Data/Sequence/Properties.agda b/Cubical/Data/Sequence/Properties.agda new file mode 100644 index 0000000000..b01b6d609b --- /dev/null +++ b/Cubical/Data/Sequence/Properties.agda @@ -0,0 +1,22 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Sequence.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Nat +open import Cubical.Foundations.Equiv +open import Cubical.Data.Sequence.Base + + +private + variable + ℓ ℓ' ℓ'' : Level +idSequenceMap : (C : Sequence ℓ) → SequenceMap C C +SequenceMap.map (idSequenceMap C) n x = x +SequenceMap.comm (idSequenceMap C) n x = refl + +composeSequenceMap : {C : Sequence ℓ} {D : Sequence ℓ'} {E : Sequence ℓ''} + (g : SequenceMap D E) (f : SequenceMap C D) → SequenceMap C E +composeSequenceMap g f .SequenceMap.map n x = g .SequenceMap.map n (f .SequenceMap.map n x) +composeSequenceMap g f .SequenceMap.comm n x = + g .SequenceMap.comm n _ + ∙ cong (g .SequenceMap.map (suc n)) (f .SequenceMap.comm n x) diff --git a/Cubical/HITs/SequentialColimit/Base.agda b/Cubical/HITs/SequentialColimit/Base.agda index 27ca5d7fbc..0b118b6f33 100644 --- a/Cubical/HITs/SequentialColimit/Base.agda +++ b/Cubical/HITs/SequentialColimit/Base.agda @@ -2,17 +2,15 @@ module Cubical.HITs.SequentialColimit.Base where open import Cubical.Foundations.Prelude -open import Cubical.Data.Nat open import Cubical.Foundations.Equiv +open import Cubical.Data.Nat +open import Cubical.Data.Sequence + + private variable - ℓ : Level - -record Sequence (ℓ : Level) : Type (ℓ-suc ℓ) where - field - obj : ℕ → Type ℓ - map : {n : ℕ} → obj n → obj (1 + n) + ℓ ℓ' : Level open Sequence @@ -20,8 +18,8 @@ data SeqColim (X : Sequence ℓ) : Type ℓ where incl : {n : ℕ} → X .obj n → SeqColim X push : {n : ℕ} (x : X .obj n) → incl x ≡ incl (X .map x) -converges : ∀ {ℓ} → Sequence ℓ → (n : ℕ) → Type ℓ -converges seq n = (k : ℕ) → isEquiv (Sequence.map seq {n = k + n}) - -finiteSequence : (ℓ : Level) (m : ℕ) → Type (ℓ-suc ℓ) -finiteSequence ℓ m = Σ[ S ∈ Sequence ℓ ] converges S m +realiseSequenceMap : {C : Sequence ℓ} {D : Sequence ℓ'} + → SequenceMap C D → SeqColim C → SeqColim D +realiseSequenceMap record { map = map ; comm = comm } (incl x) = incl (map _ x) +realiseSequenceMap record { map = map ; comm = comm } (push {n = n} x i) = + (push (map n x) ∙ λ i → incl (comm n x i)) i diff --git a/Cubical/HITs/SequentialColimit/Properties.agda b/Cubical/HITs/SequentialColimit/Properties.agda index 02cfad2f80..656b19ca88 100644 --- a/Cubical/HITs/SequentialColimit/Properties.agda +++ b/Cubical/HITs/SequentialColimit/Properties.agda @@ -15,14 +15,17 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.Univalence +open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.Sequence + open import Cubical.HITs.SequentialColimit.Base open import Cubical.Homotopy.Connected private variable - ℓ ℓ' : Level + ℓ ℓ' ℓ'' : Level module _ (X : Sequence ℓ) where @@ -381,3 +384,52 @@ SeqColim→Prop {C = C} {B = B} pr ind (push x i) = isProp→PathP {B = λ i → B (push x i)} (λ i → pr _) (ind _ x) (ind (suc _) (C .Sequence.map x)) i + +realiseIdSequenceMap : {C : Sequence ℓ} → realiseSequenceMap (idSequenceMap C) ≡ idfun _ +realiseIdSequenceMap {C = C} = + funExt λ { (incl x) → refl + ; (push {n = n} x i) j → rUnit (push {n = n} x) (~ j) i} + +realiseCompSequenceMap : {C : Sequence ℓ} {D : Sequence ℓ'} {E : Sequence ℓ''} + (g : SequenceMap D E) (f : SequenceMap C D) + → realiseSequenceMap (composeSequenceMap g f) + ≡ realiseSequenceMap g ∘ realiseSequenceMap f +realiseCompSequenceMap {C = C} {E = E} g f = + funExt λ { (incl x) → refl + ; (push {n = n} x i) j → main n x j i} + where + module _ (n : ℕ) (x : Sequence.obj C n) where + g₁ = g .SequenceMap.map n + g₊ = g .SequenceMap.map (suc n) + f₁ = f .SequenceMap.map n + + main : Path (Path (SeqColim E) _ _) + (push {n = n} (g₁ (f₁ x)) + ∙ cong incl (SequenceMap.comm g n (f₁ x) + ∙ cong g₊ (SequenceMap.comm f n x))) + (cong (realiseSequenceMap g) (push (f₁ x) ∙ cong incl (SequenceMap.comm f n x))) + main = cong (push (SequenceMap.map g n (f₁ x)) ∙_) + (cong-∙ incl (SequenceMap.comm g n (f₁ x)) + (cong g₊ (SequenceMap.comm f n x))) + ∙ assoc (push (SequenceMap.map g n (f₁ x))) + (cong incl (SequenceMap.comm g n (f₁ x))) + (cong (incl ∘ g₊) (SequenceMap.comm f n x)) + ∙ sym (cong-∙ (realiseSequenceMap g) + (push (f₁ x)) (cong incl (SequenceMap.comm f n x))) + +converges→funId : {seq1 : Sequence ℓ} {seq2 : Sequence ℓ'} + (n m : ℕ) + → converges seq1 n + → converges seq2 m + → (f : SeqColim seq1 → SeqColim seq2) + → (g : SequenceMap seq1 seq2) + → ((n : ℕ) (c : _) → f (incl {n = n} c) ≡ incl (SequenceMap.map g n c)) + → f ≡ realiseSequenceMap g +converges→funId {seq1 = seq1} {seq2} n m t1 t2 f g p = + transport (λ i → help f (~ i) ≡ help (realiseSequenceMap g) (~ i)) + (funExt λ x → p _ x) + where + help : (f : _) → PathP (λ i → isoToPath (converges→ColimIso {seq = seq1} n t1) (~ i) + → SeqColim seq2) + f (f ∘ incl) + help f = toPathP (funExt λ x → (λ i → transportRefl (f (incl (transportRefl x i))) i)) diff --git a/Cubical/Sequence.agda b/Cubical/Sequence.agda new file mode 100644 index 0000000000..e69de29bb2 From df55d5508e8e8029cb3d147636f8a454a008aed9 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 29 Feb 2024 02:49:34 +0100 Subject: [PATCH 25/73] More --- Cubical/Algebra/ChainComplex.agda | 6 + Cubical/Algebra/ChainComplex/Base.agda | 444 ++++--------------- Cubical/Algebra/ChainComplex/Finite.agda | 132 ++++++ Cubical/Algebra/ChainComplex/Homology.agda | 404 +++++++++++++++++ Cubical/Algebra/ChainComplex/Natindexed.agda | 104 ++--- Cubical/Algebra/Group/Subgroup.agda | 7 + Cubical/Data/CW/ChainComplex.agda | 11 +- Cubical/Data/CW/Properties.agda | 17 + Cubical/Data/Fin/Base.agda | 5 + Cubical/Data/Unit/Properties.agda | 18 + Cubical/HITs/SequentialColimit/Base.agda | 336 +++++++++++++- 11 files changed, 1076 insertions(+), 408 deletions(-) create mode 100644 Cubical/Algebra/ChainComplex.agda create mode 100644 Cubical/Algebra/ChainComplex/Finite.agda create mode 100644 Cubical/Algebra/ChainComplex/Homology.agda diff --git a/Cubical/Algebra/ChainComplex.agda b/Cubical/Algebra/ChainComplex.agda new file mode 100644 index 0000000000..e9a1341643 --- /dev/null +++ b/Cubical/Algebra/ChainComplex.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.ChainComplex where + +open import Cubical.Algebra.ChainComplex.Base public +open import Cubical.Algebra.ChainComplex.Homology public +open import Cubical.Algebra.ChainComplex.Finite public diff --git a/Cubical/Algebra/ChainComplex/Base.agda b/Cubical/Algebra/ChainComplex/Base.agda index e5d6d904f9..7e7bbc8868 100644 --- a/Cubical/Algebra/ChainComplex/Base.agda +++ b/Cubical/Algebra/ChainComplex/Base.agda @@ -27,361 +27,107 @@ open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.Structures.Successor open import Cubical.Relation.Nullary +open import Cubical.Data.Nat +open import Cubical.Data.Fin hiding (_/_) + private variable ℓ ℓ' ℓ'' : Level -module CC {ℓ₀ : Level} (I : InjSuccStr ℓ₀) where - open InjSuccStr I - open SuccStr succStr - - ≠succ : Index → Type ℓ₀ - ≠succ i = ¬ (i ≡ succ i) - - Index⇃ : Type ℓ₀ - Index⇃ = Σ[ i ∈ Index ] (≠succ (succ (succ i))) - - record ChainComplex (ℓ : Level) : Type (ℓ-suc (ℓ-max ℓ ℓ₀)) where - field - chain : (i : Index) → ≠succ i → AbGroup ℓ - bdry : (i : Index) (p : ≠succ (succ i)) → AbGroupHom (chain (succ i) p) (chain i (semiInj p)) - bdry²=0 : (i : Index) (p : ≠succ (succ (succ i))) - → compGroupHom (bdry (succ i) p) (bdry i (semiInj p)) ≡ trivGroupHom - - record ChainComplexMap {ℓ ℓ' : Level} - (A : ChainComplex ℓ) (B : ChainComplex ℓ') : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ₀) where - open ChainComplex - field - chainmap : (i : Index) (p : ≠succ i) → AbGroupHom (chain A i p) (chain B i p) - bdrycomm : (i : Index) (p : ≠succ (succ i)) - → compGroupHom (chainmap (succ i) p) (bdry B i p) - ≡ compGroupHom (bdry A i p) (chainmap i (semiInj p)) - - record ChainHomotopy {ℓ : Level} {A : ChainComplex ℓ} {B : ChainComplex ℓ'} - (f g : ChainComplexMap A B) : Type (ℓ-max (ℓ-max ℓ' ℓ) ℓ₀) where - open ChainComplex - open ChainComplexMap - field - htpy : (i : Index) (p : ≠succ (succ i)) → AbGroupHom (chain A i (semiInj p)) (chain B (succ i) p) - bdryhtpy : (i : Index) (p : ≠succ (succ (succ i))) - → subtrGroupHom (chain A (succ i) (semiInj p)) (chain B (succ i) (semiInj p)) - (chainmap f (succ i) (semiInj p)) (chainmap g (succ i)(semiInj p)) - ≡ addGroupHom (chain A (succ i) (semiInj p)) (chain B (succ i) (semiInj p)) - (compGroupHom (htpy (succ i) p) (bdry B (succ i) p)) - (compGroupHom (bdry A i (semiInj p)) (htpy i (semiInj p))) - - record CoChainComplex (ℓ : Level) : Type (ℓ-max (ℓ-suc ℓ) ℓ₀) where - field - cochain : (i : Index) → AbGroup ℓ - cobdry : (i : Index) → AbGroupHom (cochain i) (cochain (succ i)) - cobdry²=0 : (i : Index) → compGroupHom (cobdry i) (cobdry (succ i)) - ≡ trivGroupHom - - open ChainComplexMap - - - ChainComplexMap≡ : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} - {f g : ChainComplexMap A B} - → ((i : Index) (j : _) → chainmap f i j ≡ chainmap g i j) - → f ≡ g - chainmap (ChainComplexMap≡ p i) n j = p n j i - bdrycomm (ChainComplexMap≡ {A = A} {B = B} {f = f} {g = g} p i) n r = - isProp→PathP {B = λ i → compGroupHom (p (succ n) r i) (ChainComplex.bdry B n r) - ≡ compGroupHom (ChainComplex.bdry A n r) (p n (semiInj r) i)} - (λ i → isSetGroupHom _ _) - (bdrycomm f n r) (bdrycomm g n r) i - - compChainMap : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {C : ChainComplex ℓ''} - → (f : ChainComplexMap A B) (g : ChainComplexMap B C) - → ChainComplexMap A C - compChainMap {A = A} {B} {C} - record { chainmap = ϕ ; bdrycomm = commϕ } - record { chainmap = ψ ; bdrycomm = commψ } = main - where - main : ChainComplexMap A C - chainmap main n r = compGroupHom (ϕ n r) (ψ n r) - bdrycomm main n r = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt λ x - → (funExt⁻ (cong fst (commψ n r)) (ϕ (succ n) r .fst x)) - ∙ cong (fst (ψ n (semiInj r))) (funExt⁻ (cong fst (commϕ n r)) x)) - - isChainEquiv : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} - → ChainComplexMap A B → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ₀) - isChainEquiv f = ((n : Index) (p : _) → isEquiv (chainmap f n p .fst)) - - _≃Chain_ : (A : ChainComplex ℓ) (B : ChainComplex ℓ') → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ₀) - A ≃Chain B = Σ[ f ∈ ChainComplexMap A B ] (isChainEquiv f) - - idChainMap : (A : ChainComplex ℓ) → ChainComplexMap A A - chainmap (idChainMap A) _ _ = idGroupHom - bdrycomm (idChainMap A) _ _ = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl - - invChainMap : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} - → (A ≃Chain B) → ChainComplexMap B A - chainmap (invChainMap (record { chainmap = ϕ ; bdrycomm = ϕcomm } , eq)) n p = - GroupEquiv→GroupHom (invGroupEquiv ((ϕ n p .fst , eq n p) , snd (ϕ n p))) - bdrycomm (invChainMap {B = B} (record { chainmap = ϕ ; bdrycomm = ϕcomm } , eq)) n p = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt λ x - → sym (retEq (_ , eq n (semiInj p)) _) - ∙∙ cong (invEq (_ , eq n (semiInj p))) - (sym (funExt⁻ (cong fst (ϕcomm n p)) (invEq (_ , eq (succ n) p) x))) - ∙∙ cong (invEq (ϕ n (semiInj p) .fst , eq n (semiInj p)) ∘ fst (ChainComplex.bdry B n p)) - (secEq (_ , eq (succ n) p) x)) - - invChainEquiv : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} - → A ≃Chain B → B ≃Chain A - fst (invChainEquiv e) = invChainMap e - snd (invChainEquiv e) n p = snd (invEquiv (chainmap (fst e) n p .fst , snd e n p)) - - -- TODO: upstream these - module _ {G H : Group ℓ} (ϕ : GroupHom G H) where - kerGroup : Group ℓ - kerGroup = Subgroup→Group G (kerSubgroup ϕ) - - kerGroup≡ : {x y : ⟨ kerGroup ⟩} → x .fst ≡ y .fst → x ≡ y - kerGroup≡ = Σ≡Prop (isPropIsInKer ϕ) +record ChainComplex (ℓ : Level) : Type (ℓ-suc ℓ) where + field + chain : (i : ℕ) → AbGroup ℓ + bdry : (i : ℕ) → AbGroupHom (chain (suc i)) (chain i) + bdry²=0 : (i : ℕ) → compGroupHom (bdry (suc i)) (bdry i) ≡ trivGroupHom +record ChainComplexMap {ℓ ℓ' : Level} + (A : ChainComplex ℓ) (B : ChainComplex ℓ') : Type (ℓ-max ℓ ℓ') where open ChainComplex - open IsGroupHom + field + chainmap : (i : ℕ) → AbGroupHom (chain A i) (chain B i) + bdrycomm : (i : ℕ) + → compGroupHom (chainmap (suc i)) (bdry B i) ≡ compGroupHom (bdry A i) (chainmap i) - homology : (n : Index⇃) → ChainComplex ℓ → Group ℓ - homology (n , p) C = ker∂n / img∂+1⊂ker∂n - where - Cn+2 = AbGroup→Group (chain C (succ (succ n)) p) - ∂n = bdry C n (semiInj p) - ∂n+1 = bdry C (succ n) p - ker∂n = kerGroup ∂n - - -- Restrict ∂n+1 to ker∂n - ∂'-fun : Cn+2 .fst → ker∂n .fst - fst (∂'-fun x) = ∂n+1 .fst x - snd (∂'-fun x) = t - where - opaque - t : ⟨ fst (kerSubgroup ∂n) (∂n+1 .fst x) ⟩ - t = funExt⁻ (cong fst (bdry²=0 C n p)) x - - ∂' : GroupHom Cn+2 ker∂n - fst ∂' = ∂'-fun - snd ∂' = isHom - where - opaque - isHom : IsGroupHom (Cn+2 .snd) ∂'-fun (ker∂n .snd) - isHom = makeIsGroupHom λ x y - → kerGroup≡ ∂n (∂n+1 .snd .pres· x y) - - img∂+1⊂ker∂n : NormalSubgroup ker∂n - fst img∂+1⊂ker∂n = imSubgroup ∂' - snd img∂+1⊂ker∂n = isNormalImSubGroup - where - opaque - module C1 = AbGroupStr (chain C (succ n) (semiInj p) .snd) - isNormalImSubGroup : isNormal (imSubgroup ∂') - isNormalImSubGroup = isNormalIm ∂' - (λ x y → kerGroup≡ ∂n (C1.+Comm (fst x) (fst y))) - - -- ChainComplexMap→finiteChainComplexMap : {C D : ChainComplex ℓ} - -- → (n : ℕ) → ChainComplexMap C D → finiteChainComplexMap n C D - -- finchainmap (ChainComplexMap→finiteChainComplexMap n - -- record { chainmap = chainmap ; bdrycomm = bdrycomm }) m _ = chainmap m - -- finbdrycomm (ChainComplexMap→finiteChainComplexMap n - -- record { chainmap = chainmap ; bdrycomm = bdrycomm }) m _ = bdrycomm m - - chainComplexMap→HomologyMap : {C D : ChainComplex ℓ} - → (ϕ : ChainComplexMap C D) - → (n : Index⇃) - → GroupHom (homology n C) (homology n D) - chainComplexMap→HomologyMap {C = C} {D} mp (n , p) = main - where - ϕ = chainmap mp - ϕcomm = bdrycomm mp - fun : homology (n , p) C .fst → homology (n , p) D .fst - fun = SQ.elim (λ _ → squash/) f - λ f g → PT.rec (GroupStr.is-set (homology (n , p) D .snd) _ _ ) (λ r - → eq/ _ _ ∣ (fst (ϕ (succ (succ n)) p) (fst r)) - , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) - (funExt⁻ (cong fst (ϕcomm (succ n) p)) (fst r) - ∙∙ cong (fst (ϕ (succ n) (semiInj p))) (cong fst (snd r)) - ∙∙ IsGroupHom.pres· (snd (ϕ (succ n) (semiInj p))) _ _ - ∙ cong₂ (AbGroupStr._+_ (snd (chain D (succ n) (semiInj p)))) - refl - (IsGroupHom.presinv (snd (ϕ (succ n) (semiInj p))) _)) ∣₁) - where - f : _ → homology (n , p) D .fst - f (a , b) = [ (ϕ (succ n) (semiInj p) .fst a) - , ((λ i → fst (ϕcomm n (semiInj p) i) a) - ∙∙ cong (fst (ϕ n (semiInj (semiInj p)))) b - ∙∙ IsGroupHom.pres1 (snd (ϕ n (semiInj (semiInj p))))) ] - - main : GroupHom (homology (n , p) C) (homology (n , p) D) - fst main = fun - snd main = - makeIsGroupHom - (SQ.elimProp2 (λ _ _ → GroupStr.is-set (snd (homology (n , p) D)) _ _) - λ a b → cong [_] - (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) - (IsGroupHom.pres· (snd (ϕ (succ n) (semiInj p))) _ _))) - - -- chainComplexMap→HomologyMap : {C D : ChainComplex ℓ} - -- → (ϕ : ChainComplexMap C D) - -- → (n : ℕ) - -- → GroupHom (homology n C) (homology n D) - -- chainComplexMap→HomologyMap {C = C} {D} ϕ n = - -- finiteChainComplexMap→HomologyMap n (suc (suc n)) (0 , refl) - -- (ChainComplexMap→finiteChainComplexMap (suc (suc n)) ϕ) - - chainComplexMap→HomologyMapComp : {C D E : ChainComplex ℓ} - → (ϕ : ChainComplexMap C D) (ψ : ChainComplexMap D E) - → (n : Index⇃) - → chainComplexMap→HomologyMap (compChainMap ϕ ψ) n - ≡ compGroupHom (chainComplexMap→HomologyMap ϕ n) - (chainComplexMap→HomologyMap ψ n) - chainComplexMap→HomologyMapComp {E = E} - record { chainmap = ϕ ; bdrycomm = commϕ } - record { chainmap = ψ ; bdrycomm = commψ } (n , p) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (n , p) E)) _ _) - λ _ → cong [_] - (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain E n (semiInj (semiInj p)))) _ _) refl))) - - chainComplexMap→HomologyMapId : {C : ChainComplex ℓ} (n : Index⇃) - → chainComplexMap→HomologyMap (idChainMap C) n ≡ idGroupHom - chainComplexMap→HomologyMapId {C = C} (n , p) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (n , p) C)) _ _) - λ _ → cong [_] - (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C n (semiInj (semiInj p)))) _ _) refl))) - - ChainHomotopy→HomologyPath : {A B : ChainComplex ℓ} (f g : ChainComplexMap A B) - → ChainHomotopy f g - → (n : Index⇃) → chainComplexMap→HomologyMap f n - ≡ chainComplexMap→HomologyMap g n - ChainHomotopy→HomologyPath {A = A} {B = B} f g ϕ (n , r) = +record ChainHomotopy {ℓ : Level} {A : ChainComplex ℓ} {B : ChainComplex ℓ'} + (f g : ChainComplexMap A B) : Type (ℓ-max ℓ' ℓ) where + open ChainComplex + open ChainComplexMap + field + htpy : (i : ℕ) → AbGroupHom (chain A i) (chain B (suc i)) + bdryhtpy : (i : ℕ) + → subtrGroupHom (chain A (suc i)) (chain B (suc i)) + (chainmap f (suc i)) (chainmap g (suc i)) + ≡ addGroupHom (chain A (suc i)) (chain B (suc i)) + (compGroupHom (htpy (suc i)) (bdry B (suc i))) + (compGroupHom (bdry A i) (htpy i)) + +record CoChainComplex (ℓ : Level) : Type (ℓ-suc ℓ) where + field + cochain : (i : ℕ) → AbGroup ℓ + cobdry : (i : ℕ) → AbGroupHom (cochain i) (cochain (suc i)) + cobdry²=0 : (i : ℕ) → compGroupHom (cobdry i) (cobdry (suc i)) + ≡ trivGroupHom + +open ChainComplexMap +ChainComplexMap≡ : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} + {f g : ChainComplexMap A B} + → ((i : ℕ) → chainmap f i ≡ chainmap g i) + → f ≡ g +chainmap (ChainComplexMap≡ p i) n = p n i +bdrycomm (ChainComplexMap≡ {A = A} {B = B} {f = f} {g = g} p i) n = + isProp→PathP {B = λ i → compGroupHom (p (suc n) i) (ChainComplex.bdry B n) + ≡ compGroupHom (ChainComplex.bdry A n) (p n i)} + (λ i → isSetGroupHom _ _) + (bdrycomm f n) (bdrycomm g n) i + +compChainMap : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {C : ChainComplex ℓ''} + → (f : ChainComplexMap A B) (g : ChainComplexMap B C) + → ChainComplexMap A C +compChainMap {A = A} {B} {C} ϕ' ψ' = main + where + ϕ = chainmap ϕ' + commϕ = bdrycomm ϕ' + ψ = chainmap ψ' + commψ = bdrycomm ψ' + + main : ChainComplexMap A C + chainmap main n = compGroupHom (ϕ n) (ψ n) + bdrycomm main n = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (n , r) _)) _ _) - λ {(a , p) → eq/ _ _ - ∣ (ChainHomotopy.htpy ϕ (succ n) r .fst a) - , (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain B n (semiInj (semiInj r)))) _ _) - (sym ((funExt⁻ (cong fst (ChainHomotopy.bdryhtpy ϕ n r)) a) - ∙ cong₂ _+B_ refl - (cong (fst (ChainHomotopy.htpy ϕ n (semiInj r))) p - ∙ IsGroupHom.pres1 (snd (ChainHomotopy.htpy ϕ n (semiInj r)))) - ∙ AbGroupStr.+IdR (snd (chain B (succ n) (semiInj r))) _))) ∣₁})) - where - open GroupTheory (AbGroup→Group (chain B (succ n) (semiInj r))) - invB = GroupStr.inv (snd (AbGroup→Group (chain B (succ n) (semiInj r)))) - _+B_ = AbGroupStr._+_ (snd (chain B (succ n) (semiInj r))) - - chainComplexEquiv→HomoglogyIso : {C D : ChainComplex ℓ} (f : C ≃Chain D) - → (n : Index⇃) → GroupIso (homology n C) (homology n D) - Iso.fun (fst (chainComplexEquiv→HomoglogyIso (f , eq) n)) = - chainComplexMap→HomologyMap f n .fst - Iso.inv (fst (chainComplexEquiv→HomoglogyIso (f , eq) n)) = - chainComplexMap→HomologyMap (invChainMap (f , eq)) n .fst - Iso.rightInv (fst (chainComplexEquiv→HomoglogyIso (f , eq) (n , p))) = - funExt⁻ (cong fst (sym (chainComplexMap→HomologyMapComp - (invChainMap (f , eq)) f (n , p))) - ∙∙ cong (λ f → fst (chainComplexMap→HomologyMap f (n , p))) - (ChainComplexMap≡ λ r i - → Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (secEq (_ , eq r i)))) - ∙∙ cong fst (chainComplexMap→HomologyMapId (n , p))) - - Iso.leftInv (fst (chainComplexEquiv→HomoglogyIso (f , eq) n)) = - funExt⁻ (cong fst (sym (chainComplexMap→HomologyMapComp f - (invChainMap (f , eq)) n)) - ∙∙ cong (λ f → fst (chainComplexMap→HomologyMap f n)) - (ChainComplexMap≡ - (λ n i → Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (retEq (_ , eq n i))))) - ∙∙ cong fst (chainComplexMap→HomologyMapId n)) - snd (chainComplexEquiv→HomoglogyIso (f , eq) n) = - chainComplexMap→HomologyMap f n .snd - - -- More general version - homologyIso : (n : Index⇃) (C D : ChainComplex ℓ) - → (chEq₂ : AbGroupIso (chain C (succ (succ (fst n))) (snd n)) - (chain D (succ (succ (fst n))) (snd n))) - → (chEq₁ : AbGroupIso (chain C (succ (fst n)) (semiInj (snd n))) - (chain D (succ (fst n)) (semiInj (snd n)))) - → (chEq₀ : AbGroupIso (chain C (fst n) (semiInj (semiInj (snd n)))) - (chain D (fst n) (semiInj (semiInj (snd n))))) - → Iso.fun (chEq₀ .fst) ∘ bdry C (fst n) (semiInj (snd n)) .fst - ≡ bdry D (fst n) (semiInj (snd n)) .fst ∘ Iso.fun (chEq₁ .fst) - → Iso.fun (chEq₁ .fst) ∘ bdry C (succ (fst n)) (snd n) .fst - ≡ bdry D (succ (fst n)) (snd n) .fst ∘ Iso.fun (chEq₂ .fst) - → GroupIso (homology n C) (homology n D) - homologyIso (n , p) C D chEq₂ chEq₁ chEq₀ eq1 eq2 = main-iso - where - F : homology (n , p) C .fst → homology (n , p) D .fst - F = SQ.elim (λ _ → squash/) f - λ a b r → eq/ _ _ - (PT.map (λ { (s , t) - → (Iso.fun (chEq₂ .fst) s) - , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) - (sym (funExt⁻ eq2 s) - ∙ cong (Iso.fun (chEq₁ .fst)) (cong fst t) - ∙ IsGroupHom.pres· (chEq₁ .snd) _ _ - ∙ cong₂ (snd (chain D (succ n) (semiInj p)) .AbGroupStr._+_) - refl - (IsGroupHom.presinv (chEq₁ .snd) _))}) r) - where - f : _ → homology (n , p) D .fst - f (a , b) = [ Iso.fun (fst chEq₁) a - , sym (funExt⁻ eq1 a) ∙ cong (Iso.fun (chEq₀ .fst)) b - ∙ IsGroupHom.pres1 (snd chEq₀) ] - - G : homology (n , p) D .fst → homology (n , p) C .fst - G = SQ.elim (λ _ → squash/) g - λ a b r → eq/ _ _ - (PT.map (λ {(s , t) - → (Iso.inv (chEq₂ .fst) s) - , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C n (semiInj (semiInj p)))) _ _) - (sym (Iso.leftInv (chEq₁ .fst) _) - ∙ cong (Iso.inv (chEq₁ .fst)) (funExt⁻ eq2 (Iso.inv (chEq₂ .fst) s)) - ∙ cong (Iso.inv (chEq₁ .fst) ∘ bdry D (succ n) p .fst) - (Iso.rightInv (chEq₂ .fst) s) - ∙ cong (Iso.inv (chEq₁ .fst)) (cong fst t) - ∙ IsGroupHom.pres· (invGroupIso chEq₁ .snd) _ _ - ∙ cong₂ (snd (chain C (succ n) (semiInj p)) .AbGroupStr._+_) - refl - ((IsGroupHom.presinv (invGroupIso chEq₁ .snd) _)))}) r) - where - g : _ → homology (n , p) C .fst - g (a , b) = [ Iso.inv (fst chEq₁) a - , sym (Iso.leftInv (chEq₀ .fst) _) - ∙ cong (Iso.inv (chEq₀ .fst)) (funExt⁻ eq1 (Iso.inv (chEq₁ .fst) a)) - ∙ cong (Iso.inv (chEq₀ .fst) ∘ bdry D n (semiInj p) .fst) - (Iso.rightInv (chEq₁ .fst) a) - ∙ cong (Iso.inv (chEq₀ .fst)) b - ∙ IsGroupHom.pres1 (invGroupIso chEq₀ .snd) ] - - F-hom : IsGroupHom (homology (n , p) C .snd) F (homology (n , p) D .snd) - F-hom = - makeIsGroupHom - (SQ.elimProp2 (λ _ _ → GroupStr.is-set (homology (n , p) D .snd) _ _) - λ {(a , s) (b , t) - → cong [_] (Σ≡Prop (λ _ - → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) - (IsGroupHom.pres· (snd chEq₁) _ _)) }) - - main-iso : GroupIso (homology (n , p) C) (homology (n , p) D) - Iso.fun (fst main-iso) = F - Iso.inv (fst main-iso) = G - Iso.rightInv (fst main-iso) = - elimProp (λ _ → GroupStr.is-set (homology (n , p) D .snd) _ _) - λ{(a , b) - → cong [_] (Σ≡Prop (λ _ - → AbGroupStr.is-set (snd (chain D n (semiInj (semiInj p)))) _ _) - (Iso.rightInv (fst chEq₁) a))} - Iso.leftInv (fst main-iso) = - elimProp (λ _ → GroupStr.is-set (homology (n , p) C .snd) _ _) - λ{(a , b) - → cong [_] (Σ≡Prop (λ _ - → AbGroupStr.is-set (snd (chain C n (semiInj (semiInj p)))) _ _) - (Iso.leftInv (fst chEq₁) a))} - snd main-iso = F-hom + (funExt λ x + → (funExt⁻ (cong fst (commψ n)) (ϕ (suc n) .fst x)) + ∙ cong (fst (ψ n)) (funExt⁻ (cong fst (commϕ n)) x)) + +isChainEquiv : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} + → ChainComplexMap A B → Type (ℓ-max ℓ ℓ') +isChainEquiv f = ((n : ℕ) → isEquiv (chainmap f n .fst)) + +_≃Chain_ : (A : ChainComplex ℓ) (B : ChainComplex ℓ') → Type (ℓ-max ℓ ℓ') +A ≃Chain B = Σ[ f ∈ ChainComplexMap A B ] (isChainEquiv f) + +idChainMap : (A : ChainComplex ℓ) → ChainComplexMap A A +chainmap (idChainMap A) _ = idGroupHom +bdrycomm (idChainMap A) _ = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl + +invChainMap : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} + → (A ≃Chain B) → ChainComplexMap B A +chainmap (invChainMap (ϕ , eq)) n = + GroupEquiv→GroupHom (invGroupEquiv ((chainmap ϕ n .fst , eq n) , snd (chainmap ϕ n))) +bdrycomm (invChainMap {B = B} (ϕ' , eq)) n = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x + → sym (retEq (_ , eq n ) _) + ∙∙ cong (invEq (_ , eq n )) + (sym (funExt⁻ (cong fst (ϕcomm n)) (invEq (_ , eq (suc n)) x))) + ∙∙ cong (invEq (ϕ n .fst , eq n ) ∘ fst (ChainComplex.bdry B n)) + (secEq (_ , eq (suc n)) x)) + where + ϕ = chainmap ϕ' + ϕcomm = bdrycomm ϕ' + +invChainEquiv : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} + → A ≃Chain B → B ≃Chain A +fst (invChainEquiv e) = invChainMap e +snd (invChainEquiv e) n = snd (invEquiv (chainmap (fst e) n .fst , snd e n)) diff --git a/Cubical/Algebra/ChainComplex/Finite.agda b/Cubical/Algebra/ChainComplex/Finite.agda new file mode 100644 index 0000000000..1973188b17 --- /dev/null +++ b/Cubical/Algebra/ChainComplex/Finite.agda @@ -0,0 +1,132 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Algebra.ChainComplex.Finite where + +{- When dealing with chain maps and chain homotopies constructively, +it is often the case the case that one only is able to obtain a finite +approximation rather than the full thing. This file contains +definitions of +(1) finite chain maps, +(2) finite chain homotopies +(3) finite chain equivalences +and proof their induced behaviour on homology +-} + +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties -- TODO: why is this there and not exported by the Morphisms file? +open import Cubical.Algebra.Group.Subgroup +open import Cubical.Algebra.Group.QuotientGroup +open import Cubical.Algebra.AbGroup + +open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) +open import Cubical.HITs.SetQuotients.Properties as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.Structures.Successor +open import Cubical.Relation.Nullary + +open import Cubical.Data.Nat +open import Cubical.Data.Fin hiding (_/_) + + +open import Cubical.Algebra.ChainComplex.Base +private + variable + ℓ ℓ' ℓ'' : Level + +module _ where + record finChainComplexMap {ℓ ℓ' : Level} (m : ℕ) + (A : ChainComplex ℓ) (B : ChainComplex ℓ') : Type (ℓ-max ℓ ℓ') where + open ChainComplex + field + fchainmap : (i : Fin (suc m)) → AbGroupHom (chain A (fst i)) (chain B (fst i)) + fbdrycomm : (i : Fin m) + → compGroupHom (fchainmap (fsuc i)) (bdry B (fst i)) + ≡ compGroupHom (bdry A (fst i)) (fchainmap (injectSuc i)) + + record finChainHomotopy {ℓ : Level} (m : ℕ) {A : ChainComplex ℓ} {B : ChainComplex ℓ'} + (f g : finChainComplexMap m A B) : Type (ℓ-max ℓ' ℓ) where + open ChainComplex + open finChainComplexMap + field + fhtpy : (i : Fin (suc m)) → AbGroupHom (chain A (fst i)) (chain B (suc (fst i))) + fbdryhtpy : (i : Fin m) + → subtrGroupHom (chain A (suc (fst i))) (chain B (suc (fst i))) + (fchainmap f (fsuc i)) (fchainmap g (fsuc i)) -- (suc i)) + ≡ addGroupHom (chain A (suc (fst i))) (chain B (suc (fst i))) + (compGroupHom (fhtpy (fsuc i)) (bdry B (suc (fst i)))) + (compGroupHom (bdry A (fst i)) (fhtpy (injectSuc i))) + + open finChainComplexMap + finChainComplexMap≡ : + {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {m : ℕ} + {f g : finChainComplexMap m A B} + → ((i : Fin (suc m)) → fchainmap f i ≡ fchainmap g i) + → f ≡ g + fchainmap (finChainComplexMap≡ p i) n = p n i + fbdrycomm (finChainComplexMap≡ {A = A} {B = B} {f = f} {g = g} p i) n = + isProp→PathP {B = λ i → compGroupHom (p (fsuc n) i) (ChainComplex.bdry B (fst n)) + ≡ compGroupHom (ChainComplex.bdry A (fst n)) (p (injectSuc n) i)} + (λ i → isSetGroupHom _ _) + (fbdrycomm f n) (fbdrycomm g n) i + + compFinChainMap : + {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {C : ChainComplex ℓ''} {m : ℕ} + → (f : finChainComplexMap m A B) (g : finChainComplexMap m B C) + → finChainComplexMap m A C + compFinChainMap {A = A} {B} {C} {m = m} ϕ' ψ' = main + where + ϕ = fchainmap ϕ' + commϕ = fbdrycomm ϕ' + ψ = fchainmap ψ' + commψ = fbdrycomm ψ' + + main : finChainComplexMap m A C + fchainmap main n = compGroupHom (ϕ n) (ψ n) + fbdrycomm main n = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x + → (funExt⁻ (cong fst (commψ n)) (ϕ (fsuc n) .fst x)) + ∙ cong (fst (ψ (injectSuc n))) (funExt⁻ (cong fst (commϕ n)) x)) + + isFinChainEquiv : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {m : ℕ} + → finChainComplexMap m A B → Type (ℓ-max ℓ ℓ') + isFinChainEquiv {m = m} f = ((n : Fin (suc m)) → isEquiv (fchainmap f n .fst)) + + _≃⟨_⟩Chain_ : (A : ChainComplex ℓ) (m : ℕ) (B : ChainComplex ℓ') → Type (ℓ-max ℓ ℓ') + A ≃⟨ m ⟩Chain B = Σ[ f ∈ finChainComplexMap m A B ] (isFinChainEquiv f) + + idFinChainMap : (m : ℕ) (A : ChainComplex ℓ) → finChainComplexMap m A A + fchainmap (idFinChainMap m A) _ = idGroupHom + fbdrycomm (idFinChainMap m A) _ = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl + + invFinChainMap : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {m : ℕ} + → (A ≃⟨ m ⟩Chain B) → finChainComplexMap m B A + fchainmap (invFinChainMap {m = m} (ϕ , eq)) n = + GroupEquiv→GroupHom (invGroupEquiv ((fchainmap ϕ n .fst , eq n) , snd (fchainmap ϕ n))) + fbdrycomm (invFinChainMap {B = B} {m = m} (ϕ' , eq)) n = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x + → sym (retEq (_ , eq (injectSuc n) ) _) + ∙∙ cong (invEq (_ , eq (injectSuc n) )) + (sym (funExt⁻ (cong fst (ϕcomm n)) (invEq (_ , eq (fsuc n)) x))) + ∙∙ cong (invEq (ϕ (injectSuc n) .fst , eq (injectSuc n) ) + ∘ fst (ChainComplex.bdry B (fst n))) + (secEq (_ , eq (fsuc n)) x)) + where + ϕ = fchainmap ϕ' + ϕcomm = fbdrycomm ϕ' + + invFinChainEquiv : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {m : ℕ} + → A ≃⟨ m ⟩Chain B → B ≃⟨ m ⟩Chain A + fst (invFinChainEquiv e) = invFinChainMap e + snd (invFinChainEquiv e) n = snd (invEquiv (fchainmap (fst e) n .fst , snd e n)) diff --git a/Cubical/Algebra/ChainComplex/Homology.agda b/Cubical/Algebra/ChainComplex/Homology.agda new file mode 100644 index 0000000000..33476e9e38 --- /dev/null +++ b/Cubical/Algebra/ChainComplex/Homology.agda @@ -0,0 +1,404 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Algebra.ChainComplex.Homology where +{- + Defines groups and adds the smart constructors [makeGroup-right] and [makeGroup-left] + for constructing groups from less data than the standard [makeGroup] constructor. +-} + +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat.Order + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties -- TODO: why is this there and not exported by the Morphisms file? +open import Cubical.Algebra.Group.Subgroup +open import Cubical.Algebra.Group.QuotientGroup +open import Cubical.Algebra.AbGroup + +open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) +open import Cubical.HITs.SetQuotients.Properties as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.Structures.Successor +open import Cubical.Relation.Nullary + +open import Cubical.Data.Nat +open import Cubical.Data.Fin hiding (_/_) + +open import Cubical.Algebra.ChainComplex.Base +open import Cubical.Algebra.ChainComplex.Finite + + + +private + variable + ℓ ℓ' ℓ'' : Level + +open ChainComplexMap +open ChainComplex +open finChainComplexMap +open IsGroupHom + +-- Definition of homology +homology : (n : ℕ) → ChainComplex ℓ → Group ℓ +homology n C = ker∂n / img∂+1⊂ker∂n + where + Cn+2 = AbGroup→Group (chain C (suc (suc n))) + ∂n = bdry C n + ∂n+1 = bdry C (suc n) + ker∂n = kerGroup ∂n + + -- Restrict ∂n+1 to ker∂n + ∂'-fun : Cn+2 .fst → ker∂n .fst + fst (∂'-fun x) = ∂n+1 .fst x + snd (∂'-fun x) = t + where + opaque + t : ⟨ fst (kerSubgroup ∂n) (∂n+1 .fst x) ⟩ + t = funExt⁻ (cong fst (bdry²=0 C n)) x + + ∂' : GroupHom Cn+2 ker∂n + fst ∂' = ∂'-fun + snd ∂' = isHom + where + opaque + isHom : IsGroupHom (Cn+2 .snd) ∂'-fun (ker∂n .snd) + isHom = makeIsGroupHom λ x y + → kerGroup≡ ∂n (∂n+1 .snd .pres· x y) + + img∂+1⊂ker∂n : NormalSubgroup ker∂n + fst img∂+1⊂ker∂n = imSubgroup ∂' + snd img∂+1⊂ker∂n = isNormalImSubGroup + where + opaque + module C1 = AbGroupStr (chain C (suc n) .snd) + isNormalImSubGroup : isNormal (imSubgroup ∂') + isNormalImSubGroup = isNormalIm ∂' + (λ x y → kerGroup≡ ∂n (C1.+Comm (fst x) (fst y))) + + +-- Induced maps on cohomology by finite chain complex maps/homotopies +module _ where + finChainComplexMap→HomologyMap : {C D : ChainComplex ℓ} (m : ℕ) + → (ϕ : finChainComplexMap (suc m) C D) + → (n : Fin m) + → GroupHom (homology (fst n) C) (homology (fst n) D) + finChainComplexMap→HomologyMap {C = C} {D} m mp (n , p) = main + where + ineq1 : suc n < suc m + ineq1 = suc-≤-suc p + + ineq2 : suc (suc n) < suc (suc m) + ineq2 = suc-≤-suc ineq1 + + ineq3 : suc n < suc (suc m) + ineq3 = <-trans ineq1 (0 , refl) + + ineq4 : n < suc m + ineq4 = <-trans p (0 , refl) + + ϕ = fchainmap mp + ϕcomm = fbdrycomm mp + + lem : (k : ℕ) {p q : _} (f : fst (chain C k)) + → fst (ϕ (k , p)) f ≡ fst (ϕ (k , q)) f + lem k {p} {q} f i = fst (ϕ (k , pq i)) f + where + pq : p ≡ q + pq = isProp≤ _ _ + + fun : homology n C .fst → homology n D .fst + fun = SQ.elim (λ _ → squash/) f + λ f g → PT.rec (GroupStr.is-set (homology n D .snd) _ _ ) + λ r → eq/ _ _ ∣ fst (ϕ (suc (suc n) , ineq2)) (fst r) + , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) + ((funExt⁻ (cong fst (ϕcomm (suc n , _))) (fst r) + ∙∙ cong (fst (ϕ (suc n , _))) (cong fst (snd r)) + ∙∙ (IsGroupHom.pres· (snd (ϕ (suc n , _) )) _ _ + ∙ cong₂ (AbGroupStr._+_ (snd (chain D (suc n)))) + (lem (suc n) (fst f)) + (IsGroupHom.presinv (snd (ϕ (suc n , _) )) _ + ∙ cong (snd (chain D (suc n)) .AbGroupStr.-_) + (lem (suc n) (fst g)))))) ∣₁ + where + f : _ → homology n D .fst + f (a , b) = [ (ϕ (suc n , (suc (fst p)) + , cong suc (+-suc (fst p) (suc n) ∙ cong suc (snd p))) .fst a) + , ((λ i → fst (ϕcomm (n , suc (fst p) , cong suc (snd p)) i) a) + ∙∙ cong (fst (ϕ (n , _))) b + ∙∙ IsGroupHom.pres1 (snd (ϕ (n , _)))) ] + + main : GroupHom (homology n C) (homology n D) + fst main = fun + snd main = + makeIsGroupHom + (SQ.elimProp2 (λ _ _ → GroupStr.is-set (snd (homology n D)) _ _) + λ a b → cong [_] + (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) + (IsGroupHom.pres· (snd (ϕ (suc n , _) )) _ _))) + + finChainComplexMap→HomologyMapComp : {C D E : ChainComplex ℓ} {m : ℕ} + → (ϕ : finChainComplexMap (suc m) C D) (ψ : finChainComplexMap (suc m) D E) + → (n : Fin m) + → finChainComplexMap→HomologyMap m (compFinChainMap ϕ ψ) n + ≡ compGroupHom (finChainComplexMap→HomologyMap m ϕ n) + (finChainComplexMap→HomologyMap m ψ n) + finChainComplexMap→HomologyMapComp {E = E} _ _ n = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (fst n) E)) _ _) + λ _ → cong [_] + (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain E (fst n))) _ _) refl))) + + finChainComplexMap→HomologyMapId : {C : ChainComplex ℓ} {m : ℕ} (n : Fin m) + → finChainComplexMap→HomologyMap m (idFinChainMap (suc m) C) n ≡ idGroupHom + finChainComplexMap→HomologyMapId {C = C} n = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (fst n) C)) _ _) + λ _ → cong [_] + (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C (fst n))) _ _) refl))) + + finChainComplexEquiv→HomoglogyIso : {C D : ChainComplex ℓ} (m : ℕ) (f : C ≃⟨ (suc m) ⟩Chain D) + → (n : Fin m) → GroupIso (homology (fst n) C) (homology (fst n) D) + Iso.fun (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = + finChainComplexMap→HomologyMap m f n .fst + Iso.inv (fst (finChainComplexEquiv→HomoglogyIso m f n)) = + finChainComplexMap→HomologyMap m (invFinChainMap f) n .fst + Iso.rightInv (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = + funExt⁻ (cong fst (sym (finChainComplexMap→HomologyMapComp + (invFinChainMap (f , eqs)) f n)) + ∙∙ cong (λ f → fst (finChainComplexMap→HomologyMap m f n)) + (finChainComplexMap≡ λ r + → Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (secEq (_ , eqs r)))) + ∙∙ cong fst (finChainComplexMap→HomologyMapId n)) + Iso.leftInv (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = + funExt⁻ (cong fst (sym (finChainComplexMap→HomologyMapComp f + (invFinChainMap (f , eqs)) n)) + ∙∙ cong (λ f → fst (finChainComplexMap→HomologyMap m f n)) + (finChainComplexMap≡ + (λ n → Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (retEq (_ , eqs n))))) + ∙∙ cong fst (finChainComplexMap→HomologyMapId n)) + snd (finChainComplexEquiv→HomoglogyIso m (f , eqs) n) = + finChainComplexMap→HomologyMap m f n .snd + + + finChainHomotopy→HomologyPath : {A B : ChainComplex ℓ} {m : ℕ} + (f g : finChainComplexMap (suc m) A B) + → finChainHomotopy (suc m) f g + → (n : Fin m) + → finChainComplexMap→HomologyMap m f n + ≡ finChainComplexMap→HomologyMap m g n + finChainHomotopy→HomologyPath {A = A} {B = B} {m = m} f g ϕ n = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (fst n) _)) _ _) + λ {(a , p) → eq/ _ _ + ∣ (finChainHomotopy.fhtpy ϕ (suc (fst n) , pf) .fst a) + , (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain B (fst n))) _ _) + (sym ((funExt⁻ (cong fst (finChainHomotopy.fbdryhtpy ϕ _)) a) + ∙ cong₂ _+B_ refl + (cong (fst (finChainHomotopy.fhtpy ϕ _)) p + ∙ IsGroupHom.pres1 (snd (finChainHomotopy.fhtpy ϕ _))) + ∙ AbGroupStr.+IdR (snd (chain B (suc (fst n)))) _))) ∣₁})) + where + open GroupTheory (AbGroup→Group (chain B (suc (fst n)))) + pf : suc (fst n) < suc (suc m) + fst pf = suc (fst (snd n)) + snd pf = cong suc (+-suc (fst (snd n)) (suc (fst n)) + ∙ cong suc (snd (snd n))) + + invB = GroupStr.inv (snd (AbGroup→Group (chain B (suc (fst n))))) + _+B_ = AbGroupStr._+_ (snd (chain B (suc (fst n)))) + +-- corresponding lemmas/constructions for full chain complex maps/homotopies +module _ where + chainComplexMap→HomologyMap : {C D : ChainComplex ℓ} + → (ϕ : ChainComplexMap C D) + → (n : ℕ) + → GroupHom (homology n C) (homology n D) + chainComplexMap→HomologyMap {C = C} {D} mp n = main + where + ϕ = chainmap mp + ϕcomm = bdrycomm mp + fun : homology n C .fst → homology n D .fst + fun = SQ.elim (λ _ → squash/) f + λ f g → PT.rec (GroupStr.is-set (homology n D .snd) _ _ ) (λ r + → eq/ _ _ ∣ (fst (ϕ (suc (suc n))) (fst r)) + , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) + (funExt⁻ (cong fst (ϕcomm (suc n))) (fst r) + ∙∙ cong (fst (ϕ (suc n) )) (cong fst (snd r)) + ∙∙ IsGroupHom.pres· (snd (ϕ (suc n) )) _ _ + ∙ cong₂ (AbGroupStr._+_ (snd (chain D (suc n) ))) + refl + (IsGroupHom.presinv (snd (ϕ (suc n) )) _)) ∣₁) + where + f : _ → homology n D .fst + f (a , b) = [ (ϕ (suc n) .fst a) + , ((λ i → fst (ϕcomm n i) a) + ∙∙ cong (fst (ϕ n)) b + ∙∙ IsGroupHom.pres1 (snd (ϕ n))) ] + + main : GroupHom (homology n C) (homology n D) + fst main = fun + snd main = + makeIsGroupHom + (SQ.elimProp2 (λ _ _ → GroupStr.is-set (snd (homology n D)) _ _) + λ a b → cong [_] + (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) + (IsGroupHom.pres· (snd (ϕ (suc n) )) _ _))) + + chainComplexMap→HomologyMapComp : {C D E : ChainComplex ℓ} + → (ϕ : ChainComplexMap C D) (ψ : ChainComplexMap D E) + → (n : ℕ) + → chainComplexMap→HomologyMap (compChainMap ϕ ψ) n + ≡ compGroupHom (chainComplexMap→HomologyMap ϕ n) + (chainComplexMap→HomologyMap ψ n) + chainComplexMap→HomologyMapComp {E = E} _ _ n = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology n E)) _ _) + λ _ → cong [_] + (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain E n)) _ _) refl))) + + chainComplexMap→HomologyMapId : {C : ChainComplex ℓ} (n : ℕ) + → chainComplexMap→HomologyMap (idChainMap C) n ≡ idGroupHom + chainComplexMap→HomologyMapId {C = C} n = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology n C)) _ _) + λ _ → cong [_] + (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C n)) _ _) refl))) + + ChainHomotopy→HomologyPath : {A B : ChainComplex ℓ} (f g : ChainComplexMap A B) + → ChainHomotopy f g + → (n : ℕ) → chainComplexMap→HomologyMap f n + ≡ chainComplexMap→HomologyMap g n + ChainHomotopy→HomologyPath {A = A} {B = B} f g ϕ n = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology n _)) _ _) + λ {(a , p) → eq/ _ _ + ∣ (ChainHomotopy.htpy ϕ (suc n) .fst a) + , (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain B n)) _ _) + (sym ((funExt⁻ (cong fst (ChainHomotopy.bdryhtpy ϕ n)) a) + ∙ cong₂ _+B_ refl + (cong (fst (ChainHomotopy.htpy ϕ n)) p + ∙ IsGroupHom.pres1 (snd (ChainHomotopy.htpy ϕ n))) + ∙ AbGroupStr.+IdR (snd (chain B (suc n))) _))) ∣₁})) + where + open GroupTheory (AbGroup→Group (chain B (suc n))) + invB = GroupStr.inv (snd (AbGroup→Group (chain B (suc n)))) + _+B_ = AbGroupStr._+_ (snd (chain B (suc n))) + + chainComplexEquiv→HomoglogyIso : {C D : ChainComplex ℓ} (f : C ≃Chain D) + → (n : ℕ) → GroupIso (homology n C) (homology n D) + Iso.fun (fst (chainComplexEquiv→HomoglogyIso (f , eqs) n)) = + chainComplexMap→HomologyMap f n .fst + Iso.inv (fst (chainComplexEquiv→HomoglogyIso (f , eqs) n)) = + chainComplexMap→HomologyMap (invChainMap (f , eqs)) n .fst + Iso.rightInv (fst (chainComplexEquiv→HomoglogyIso (f , eqs) n)) = + funExt⁻ (cong fst (sym (chainComplexMap→HomologyMapComp + (invChainMap (f , eqs)) f n)) + ∙∙ cong (λ f → fst (chainComplexMap→HomologyMap f n)) + (ChainComplexMap≡ λ r + → Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (secEq (_ , eqs r)))) + ∙∙ cong fst (chainComplexMap→HomologyMapId n)) + + Iso.leftInv (fst (chainComplexEquiv→HomoglogyIso (f , eqs) n)) = + funExt⁻ (cong fst (sym (chainComplexMap→HomologyMapComp f + (invChainMap (f , eqs)) n)) + ∙∙ cong (λ f → fst (chainComplexMap→HomologyMap f n)) + (ChainComplexMap≡ + (λ n → Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (retEq (_ , eqs n))))) + ∙∙ cong fst (chainComplexMap→HomologyMapId n)) + snd (chainComplexEquiv→HomoglogyIso (f , eqs) n) = + chainComplexMap→HomologyMap f n .snd + +-- More general version +homologyIso : (n : ℕ) (C D : ChainComplex ℓ) + → (chEq₂ : AbGroupIso (chain C (suc (suc n))) + (chain D (suc (suc n)))) + → (chEq₁ : AbGroupIso (chain C (suc n)) + (chain D (suc n))) + → (chEq₀ : AbGroupIso (chain C n) + (chain D n)) + → Iso.fun (chEq₀ .fst) ∘ bdry C n .fst + ≡ bdry D n .fst ∘ Iso.fun (chEq₁ .fst) + → Iso.fun (chEq₁ .fst) ∘ bdry C (suc n) .fst + ≡ bdry D (suc n) .fst ∘ Iso.fun (chEq₂ .fst) + → GroupIso (homology n C) (homology n D) +homologyIso n C D chEq₂ chEq₁ chEq₀ eq1 eq2 = main-iso + where + F : homology n C .fst → homology n D .fst + F = SQ.elim (λ _ → squash/) f + λ a b r → eq/ _ _ + (PT.map (λ { (s , t) + → (Iso.fun (chEq₂ .fst) s) + , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) + (sym (funExt⁻ eq2 s) + ∙ cong (Iso.fun (chEq₁ .fst)) (cong fst t) + ∙ IsGroupHom.pres· (chEq₁ .snd) _ _ + ∙ cong₂ (snd (chain D (suc n) ) .AbGroupStr._+_) + refl + (IsGroupHom.presinv (chEq₁ .snd) _))}) r) + where + f : _ → homology n D .fst + f (a , b) = [ Iso.fun (fst chEq₁) a + , sym (funExt⁻ eq1 a) ∙ cong (Iso.fun (chEq₀ .fst)) b + ∙ IsGroupHom.pres1 (snd chEq₀) ] + + G : homology n D .fst → homology n C .fst + G = SQ.elim (λ _ → squash/) g + λ a b r → eq/ _ _ + (PT.map (λ {(s , t) + → (Iso.inv (chEq₂ .fst) s) + , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C n)) _ _) + (sym (Iso.leftInv (chEq₁ .fst) _) + ∙ cong (Iso.inv (chEq₁ .fst)) (funExt⁻ eq2 (Iso.inv (chEq₂ .fst) s)) + ∙ cong (Iso.inv (chEq₁ .fst) ∘ bdry D (suc n) .fst) + (Iso.rightInv (chEq₂ .fst) s) + ∙ cong (Iso.inv (chEq₁ .fst)) (cong fst t) + ∙ IsGroupHom.pres· (invGroupIso chEq₁ .snd) _ _ + ∙ cong₂ (snd (chain C (suc n) ) .AbGroupStr._+_) + refl + ((IsGroupHom.presinv (invGroupIso chEq₁ .snd) _)))}) r) + where + g : _ → homology n C .fst + g (a , b) = [ Iso.inv (fst chEq₁) a + , sym (Iso.leftInv (chEq₀ .fst) _) + ∙ cong (Iso.inv (chEq₀ .fst)) (funExt⁻ eq1 (Iso.inv (chEq₁ .fst) a)) + ∙ cong (Iso.inv (chEq₀ .fst) ∘ bdry D n .fst) + (Iso.rightInv (chEq₁ .fst) a) + ∙ cong (Iso.inv (chEq₀ .fst)) b + ∙ IsGroupHom.pres1 (invGroupIso chEq₀ .snd) ] + + F-hom : IsGroupHom (homology n C .snd) F (homology n D .snd) + F-hom = + makeIsGroupHom + (SQ.elimProp2 (λ _ _ → GroupStr.is-set (homology n D .snd) _ _) + λ {(a , s) (b , t) + → cong [_] (Σ≡Prop (λ _ + → AbGroupStr.is-set (snd (chain D n)) _ _) + (IsGroupHom.pres· (snd chEq₁) _ _)) }) + + main-iso : GroupIso (homology n C) (homology n D) + Iso.fun (fst main-iso) = F + Iso.inv (fst main-iso) = G + Iso.rightInv (fst main-iso) = + elimProp (λ _ → GroupStr.is-set (homology n D .snd) _ _) + λ{(a , b) + → cong [_] (Σ≡Prop (λ _ + → AbGroupStr.is-set (snd (chain D n)) _ _) + (Iso.rightInv (fst chEq₁) a))} + Iso.leftInv (fst main-iso) = + elimProp (λ _ → GroupStr.is-set (homology n C .snd) _ _) + λ{(a , b) + → cong [_] (Σ≡Prop (λ _ + → AbGroupStr.is-set (snd (chain C n)) _ _) + (Iso.leftInv (fst chEq₁) a))} + snd main-iso = F-hom diff --git a/Cubical/Algebra/ChainComplex/Natindexed.agda b/Cubical/Algebra/ChainComplex/Natindexed.agda index 592a24d723..1a6e918064 100644 --- a/Cubical/Algebra/ChainComplex/Natindexed.agda +++ b/Cubical/Algebra/ChainComplex/Natindexed.agda @@ -1,70 +1,70 @@ {-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.ChainComplex.Natindexed where -open import Cubical.Foundations.Prelude +-- open import Cubical.Foundations.Prelude -open import Cubical.Data.Nat -open import Cubical.Data.Fin -open import Cubical.Data.Sigma -open import Cubical.Data.Nat.Order -open import Cubical.Data.Empty as ⊥ +-- open import Cubical.Data.Nat +-- open import Cubical.Data.Fin +-- open import Cubical.Data.Sigma +-- open import Cubical.Data.Nat.Order +-- open import Cubical.Data.Empty as ⊥ -open import Cubical.Structures.Successor -open import Cubical.Algebra.ChainComplex.Base +-- open import Cubical.Structures.Successor +-- open import Cubical.Algebra.ChainComplex.Base -open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties +-- open import Cubical.Algebra.AbGroup +-- open import Cubical.Algebra.Group.Morphisms +-- open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Relation.Nullary +-- open import Cubical.Relation.Nullary -private variable - ℓ : Level +-- private variable +-- ℓ : Level --- Nat indexed chain complexes -module CCℕ = CC ℕ+Inj +-- -- Nat indexed chain complexes +-- module CCℕ = CC ℕ+Inj --- Restriction of ℕ indexed chain complexes to Fin n indexed ones -module _ (n : ℕ) where - open CC +-- -- Restriction of ℕ indexed chain complexes to Fin n indexed ones +-- module _ (n : ℕ) where +-- open CC - -- Fin n indexed chain complexes - module CCFin = CC (Fin+Inj n) +-- -- Fin n indexed chain complexes +-- module CCFin = CC (Fin+Inj n) - open ChainComplex - open ChainComplexMap - open ChainHomotopy +-- open ChainComplex +-- open ChainComplexMap +-- open ChainHomotopy - ChainComplex→finChainComplex : CCℕ.ChainComplex ℓ → CCFin.ChainComplex ℓ - chain (ChainComplex→finChainComplex C) i p = chain C (fst i) ≠suc - bdry (ChainComplex→finChainComplex C) (i , zero , q) p = ⊥.rec (p refl) - bdry (ChainComplex→finChainComplex C) (i , suc diff , q) p = bdry C i ≠suc - bdry²=0 (ChainComplex→finChainComplex C) (i , zero , q) p = ⊥.rec (p refl) - bdry²=0 (ChainComplex→finChainComplex C) (i , suc zero , q) p = ⊥.rec (p refl) - bdry²=0 (ChainComplex→finChainComplex C) (i , suc (suc diff) , q) p = bdry²=0 C i ≠suc +-- ChainComplex→finChainComplex : CCℕ.ChainComplex ℓ → CCFin.ChainComplex ℓ +-- chain (ChainComplex→finChainComplex C) i p = chain C (fst i) ≠suc +-- bdry (ChainComplex→finChainComplex C) (i , zero , q) p = ⊥.rec (p refl) +-- bdry (ChainComplex→finChainComplex C) (i , suc diff , q) p = bdry C i ≠suc +-- bdry²=0 (ChainComplex→finChainComplex C) (i , zero , q) p = ⊥.rec (p refl) +-- bdry²=0 (ChainComplex→finChainComplex C) (i , suc zero , q) p = ⊥.rec (p refl) +-- bdry²=0 (ChainComplex→finChainComplex C) (i , suc (suc diff) , q) p = bdry²=0 C i ≠suc - ChainComplexMap→finChainComplexMap : {C D : CCℕ.ChainComplex ℓ} - → CCℕ.ChainComplexMap C D - → CCFin.ChainComplexMap (ChainComplex→finChainComplex C) - (ChainComplex→finChainComplex D) - chainmap (ChainComplexMap→finChainComplexMap ϕ) (x , zero , q) p = ⊥.rec (p refl) - chainmap (ChainComplexMap→finChainComplexMap ϕ) (x , suc diff , q) _ = chainmap ϕ x _ - bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , zero , q) p = ⊥.rec (p refl) - bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , suc zero , q) p = ⊥.rec (p refl) - bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , suc (suc diff) , p) q = bdrycomm ϕ x _ +-- ChainComplexMap→finChainComplexMap : {C D : CCℕ.ChainComplex ℓ} +-- → CCℕ.ChainComplexMap C D +-- → CCFin.ChainComplexMap (ChainComplex→finChainComplex C) +-- (ChainComplex→finChainComplex D) +-- chainmap (ChainComplexMap→finChainComplexMap ϕ) (x , zero , q) p = ⊥.rec (p refl) +-- chainmap (ChainComplexMap→finChainComplexMap ϕ) (x , suc diff , q) _ = chainmap ϕ x _ +-- bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , zero , q) p = ⊥.rec (p refl) +-- bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , suc zero , q) p = ⊥.rec (p refl) +-- bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , suc (suc diff) , p) q = bdrycomm ϕ x _ - chainHom→ : {C D : CCℕ.ChainComplex ℓ} (f g : CCℕ.ChainComplexMap C D) - → CCℕ.ChainHomotopy f g - → CCFin.ChainHomotopy - (ChainComplexMap→finChainComplexMap f) - (ChainComplexMap→finChainComplexMap g) - htpy (chainHom→ f g h) (x , zero , q) p = ⊥.rec (p refl) - htpy (chainHom→ f g h) (x , suc zero , q) p = ⊥.rec (p refl) - htpy (chainHom→ f g h) (x , suc (suc diff) , q) p = htpy h x _ - bdryhtpy (chainHom→ f g h) (x , zero , q) p = ⊥.rec (p refl) - bdryhtpy (chainHom→ f g h) (x , suc zero , q) p = ⊥.rec (p refl) - bdryhtpy (chainHom→ f g h) (x , suc (suc zero) , q) p = ⊥.rec (p refl) - bdryhtpy (chainHom→ f g h) (x , suc (suc (suc diff)) , q) _ = bdryhtpy h x _ +-- chainHom→ : {C D : CCℕ.ChainComplex ℓ} (f g : CCℕ.ChainComplexMap C D) +-- → CCℕ.ChainHomotopy f g +-- → CCFin.ChainHomotopy +-- (ChainComplexMap→finChainComplexMap f) +-- (ChainComplexMap→finChainComplexMap g) +-- htpy (chainHom→ f g h) (x , zero , q) p = ⊥.rec (p refl) +-- htpy (chainHom→ f g h) (x , suc zero , q) p = ⊥.rec (p refl) +-- htpy (chainHom→ f g h) (x , suc (suc diff) , q) p = htpy h x _ +-- bdryhtpy (chainHom→ f g h) (x , zero , q) p = ⊥.rec (p refl) +-- bdryhtpy (chainHom→ f g h) (x , suc zero , q) p = ⊥.rec (p refl) +-- bdryhtpy (chainHom→ f g h) (x , suc (suc zero) , q) p = ⊥.rec (p refl) +-- bdryhtpy (chainHom→ f g h) (x , suc (suc (suc diff)) , q) _ = bdryhtpy h x _ diff --git a/Cubical/Algebra/Group/Subgroup.agda b/Cubical/Algebra/Group/Subgroup.agda index 3e060f289b..729095fb73 100644 --- a/Cubical/Algebra/Group/Subgroup.agda +++ b/Cubical/Algebra/Group/Subgroup.agda @@ -207,3 +207,10 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where f x H.· f (G.inv x) ≡⟨ cong (f x H.·_) (ϕ.presinv x) ⟩ f x H.· H.inv (f x) ≡⟨ H.·InvR _ ⟩ H.1g ∎ + +module _ {G H : Group ℓ} (ϕ : GroupHom G H) where + kerGroup : Group ℓ + kerGroup = Subgroup→Group G (kerSubgroup ϕ) + + kerGroup≡ : {x y : ⟨ kerGroup ⟩} → x .fst ≡ y .fst → x ≡ y + kerGroup≡ = Σ≡Prop (isPropIsInKer ϕ) diff --git a/Cubical/Data/CW/ChainComplex.agda b/Cubical/Data/CW/ChainComplex.agda index 464a205d7a..337d5455ab 100644 --- a/Cubical/Data/CW/ChainComplex.agda +++ b/Cubical/Data/CW/ChainComplex.agda @@ -18,7 +18,7 @@ open import Cubical.HITs.SphereBouquet.Degree open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.ChainComplex.Natindexed +open import Cubical.Algebra.ChainComplex open import Cubical.Structures.Successor @@ -178,14 +178,13 @@ module _ {ℓ} (C : CWskel ℓ) where ∂≡∂↑ = bouquetDegreeSusp (pre∂ n) - open CCℕ open ChainComplex CW-ChainComplex : ChainComplex ℓ-zero - chain CW-ChainComplex n _ = ℤ[A n ] - bdry CW-ChainComplex n p = ∂ n - bdry²=0 CW-ChainComplex n p = ∂∂≡0 n + chain CW-ChainComplex n = ℤ[A n ] + bdry CW-ChainComplex n = ∂ n + bdry²=0 CW-ChainComplex n = ∂∂≡0 n -- Cellular homology Hᶜʷ : (n : ℕ) → Group₀ - Hᶜʷ n = homology (n , ≠suc) CW-ChainComplex + Hᶜʷ n = homology n CW-ChainComplex diff --git a/Cubical/Data/CW/Properties.agda b/Cubical/Data/CW/Properties.agda index 8bcc684acb..416d16109c 100644 --- a/Cubical/Data/CW/Properties.agda +++ b/Cubical/Data/CW/Properties.agda @@ -162,6 +162,23 @@ module _ {ℓ : Level} (C : CWskel ℓ) where → CWskel-elim' (invEq (e (suc n)) (inl x)) ≡ inler x CWskel-elim'-inl = CWskel-elim-inl (suc n) {B = B} inler _ _ +-- open import Cubical.Axiom.Choice +-- open import Cubical.HITs.Truncation as TR +-- module _ {ℓ : Level} (C : CWskel ℓ) where +-- open CWskel-fields C +-- CWskel-elim-trunc : (n : ℕ) {B : fst C (suc (suc (suc n))) → Type ℓ'} +-- → (f : ((x : fst C (suc (suc n))) → B (CW↪ C (suc (suc n)) x))) +-- → ∃[ f^ ∈ ((x : _) → B x) ] ((c : _) → f^ (CW↪ C (suc (suc n)) c) ≡ f c) +-- CWskel-elim-trunc n {B = B} f = PT.map (λ F → (CWskel-elim' C (suc n) f F) , CWskel-elim'-inl C (suc n) f F) F +-- where +-- F : ∥ ((x : Fin (card (suc (suc n)))) (y : S₊ (suc n)) +-- → PathP (λ i → B (invEq (e (suc (suc n))) ((push (x , y) ∙ sym (push (x , ptSn (suc n)))) i))) +-- (f (α (suc (suc n)) (x , y))) (f (α (suc (suc n)) (x , ptSn (suc n))))) ∥₁ +-- F = invEq propTrunc≃Trunc1 (invEq (_ , FinSatAC _ _ _) +-- λ x → fst propTrunc≃Trunc1 +-- (sphereToTrunc (suc n) {!!})) +-- q = FinSatAC + finCWskel≃ : (n : ℕ) (C : finCWskel ℓ n) (m : ℕ) → n ≤ m → fst C n ≃ fst C m finCWskel≃ n C m (zero , diff) = substEquiv (λ n → fst C n) diff finCWskel≃ n C zero (suc x , diff) = ⊥.rec (snotz diff) diff --git a/Cubical/Data/Fin/Base.agda b/Cubical/Data/Fin/Base.agda index c414d879a7..180208a6cb 100644 --- a/Cubical/Data/Fin/Base.agda +++ b/Cubical/Data/Fin/Base.agda @@ -44,6 +44,11 @@ fzero≠fone p = znots (cong fst p) fsuc : Fin k → Fin (suc k) fsuc (k , l) = (suc k , suc-≤-suc l) +-- predecessors too +predFin : (m : ℕ) → Fin (suc (suc m)) → Fin (suc m) +predFin m (zero , w) = fzero +predFin m (suc n , w) = n , predℕ-≤-predℕ w + -- Conversion back to ℕ is trivial... toℕ : Fin k → ℕ toℕ = fst diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index 2f8e5cf5dc..85b37dc14b 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -12,6 +12,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Data.Nat open import Cubical.Data.Unit.Base open import Cubical.Data.Prod.Base +open import Cubical.Data.Sigma hiding (_×_) open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv @@ -122,3 +123,20 @@ isContr→≃Unit* contr = compEquiv (isContr→≃Unit contr) Unit≃Unit* isContr→≡Unit* : {A : Type ℓ} → isContr A → A ≡ Unit* isContr→≡Unit* contr = ua (isContr→≃Unit* contr) + +-- J for pointed propositions +JPointedProp : ∀ {ℓ ℓ'} {B : (A : Type ℓ') (a : A) (isPr : isProp A) → Type ℓ} + → B Unit* tt* isPropUnit* + → (A : Type ℓ') (a : A) (isPr : isProp A) → B A a isPr +JPointedProp {ℓ' = ℓ'} {B = B} ind A a isPr = + transport (λ i → B (P (~ i) .fst) (coh i) (P (~ i) .snd)) ind + where + A* : TypeOfHLevel ℓ' 1 + A* = A , isPr + + P : A* ≡ (Unit* , isPropUnit*) + P = Σ≡Prop (λ _ → isPropIsProp) + (ua (propBiimpl→Equiv isPr isPropUnit* (λ _ → tt*) λ _ → a)) + + coh : PathP (λ i → (P (~ i) .fst)) tt* a + coh = toPathP refl diff --git a/Cubical/HITs/SequentialColimit/Base.agda b/Cubical/HITs/SequentialColimit/Base.agda index 0b118b6f33..b8f4add21c 100644 --- a/Cubical/HITs/SequentialColimit/Base.agda +++ b/Cubical/HITs/SequentialColimit/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.HITs.SequentialColimit.Base where open import Cubical.Foundations.Prelude @@ -6,6 +6,8 @@ open import Cubical.Foundations.Equiv open import Cubical.Data.Nat open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence +open import Cubical.Data.Fin private @@ -18,8 +20,340 @@ data SeqColim (X : Sequence ℓ) : Type ℓ where incl : {n : ℕ} → X .obj n → SeqColim X push : {n : ℕ} (x : X .obj n) → incl x ≡ incl (X .map x) +data FinSeqColim (m : ℕ) (X : Sequence ℓ) : Type ℓ where + fincl : {n : Fin (suc m)} → X .obj (fst n) → FinSeqColim m X + fpush : {n : Fin m} (x : X .obj (fst n)) + → fincl {n = injectSuc n} x ≡ fincl {n = fsuc n} (X .map x) + +FinSeqColim→Colim : (m : ℕ) {X : Sequence ℓ} → FinSeqColim m X → SeqColim X +FinSeqColim→Colim m (fincl x) = incl x +FinSeqColim→Colim m (fpush x i) = push x i + realiseSequenceMap : {C : Sequence ℓ} {D : Sequence ℓ'} → SequenceMap C D → SeqColim C → SeqColim D realiseSequenceMap record { map = map ; comm = comm } (incl x) = incl (map _ x) realiseSequenceMap record { map = map ; comm = comm } (push {n = n} x i) = (push (map n x) ∙ λ i → incl (comm n x i)) i + + + + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence + + + + +open import Cubical.Foundations.Path +open import Cubical.Foundations.GroupoidLaws + + + +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Data.Sum hiding (map) + + +-- colim (X₀ → ... → Xₙ) ≃ colim (X₁ → ... → Xₙ) +module _ (X : Sequence ℓ) where + ShiftSeq : Sequence ℓ + obj ShiftSeq m = obj X (suc m) + map ShiftSeq = map X + + -- technical lemmas all concerning the proof irrelevance of the + -- second field of Fin n. + private + ineq-lem₁ : {n m : ℕ} (p : n < m) → suc-≤-suc (snd (injectSuc (n , p))) + ≡ <-trans (snd (fsuc (n , p))) (0 , refl) + ineq-lem₁ p = isProp≤ _ _ + + ineq-lem₂ : {m : ℕ} (w : zero < suc m) + → suc-≤-suc zero-≤ ≡ predℕ-≤-predℕ (suc-≤-suc w) + ineq-lem₂ w = isProp≤ _ _ + + ineq-lem₃ : {n m : ℕ} → (w : suc n < suc m) + → suc-≤-suc (pred-≤-pred w) ≡ pred-≤-pred (suc-≤-suc w) + ineq-lem₃ p = isProp≤ _ _ + + ineq-lem₄ : {n m : ℕ} → (w : suc n < suc m) + → Path (n < suc m) (pred-≤-pred (<-trans w (0 , refl))) + (<-trans (pred-≤-pred w) (0 , refl)) + ineq-lem₄ w = Σ≡Prop (λ _ → isSetℕ _ _) refl + + ineq-lem₅ : {n m : ℕ} → (w : n < m) → predℕ-≤-predℕ (suc-≤-suc w) ≡ w + ineq-lem₅ (w , p) = Σ≡Prop (λ _ → isSetℕ _ _) refl + + ineq-lem₆ : {n : ℕ} (diff : zero < (suc (suc n))) + → <-trans (snd fzero) (0 , refl) ≡ diff + ineq-lem₆ diff = isProp≤ _ _ + + ineq-lem₇ : {n m : ℕ} (x : suc n < suc m) + → suc-≤-suc (predℕ-≤-predℕ x) ≡ x + ineq-lem₇ x = isProp≤ _ _ + + -- two highly technical lemmas to deal with problems arising from + -- the second field of Fin not being a strict prop + lem₁ : ∀ {ℓ} {A : Type ℓ} + (n Date: Thu, 29 Feb 2024 15:34:05 +0100 Subject: [PATCH 26/73] stuff --- Cubical/Data/CW/Properties.agda | 36 +++++++- Cubical/Data/FinData/Properties.agda | 2 +- Cubical/Foundations/Equiv/Properties.agda | 15 ++++ Cubical/HITs/SequentialColimit/Base.agda | 103 ++++++++++++++++++++++ 4 files changed, 152 insertions(+), 4 deletions(-) diff --git a/Cubical/Data/CW/Properties.agda b/Cubical/Data/CW/Properties.agda index 416d16109c..1a3acb7f2f 100644 --- a/Cubical/Data/CW/Properties.agda +++ b/Cubical/Data/CW/Properties.agda @@ -14,6 +14,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.Transport open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.Properties open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order @@ -30,15 +31,16 @@ open import Cubical.HITs.Susp open import Cubical.HITs.SequentialColimit open import Cubical.HITs.SphereBouquet open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as TR open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup -open import Cubical.HITs.SequentialColimit -open Sequence - +open import Cubical.Axiom.Choice open import Cubical.Relation.Nullary +open Sequence + private @@ -187,3 +189,31 @@ finCWskel≃ n C (suc m) (suc x , diff) = (compEquiv (substEquiv (λ n → fst C n) (sym (cong predℕ diff))) (compEquiv (_ , snd C .snd x) (substEquiv (λ n → fst C n) diff))) + +-- C₁ satisfies AC +satAC-CW₁ : ∀ {ℓ ℓ'} (n : ℕ) (C : CWskel ℓ) → satAC ℓ' n (fst C (suc zero)) +satAC-CW₁ {ℓ' = ℓ'} n C A = + subst isEquiv (choicefun≡ n) (isoToIsEquiv (choicefun' n)) + where + fin = Fin (snd C .fst zero) + satAC' : (n : ℕ) → satAC ℓ' n fin + satAC' n = FinSatAC _ _ + + fin→ : fin ≃ fst C 1 + fin→ = invEquiv (CW₁-discrete C) + + + choicefun' : (n : ℕ) → Iso (hLevelTrunc n ((a : fst C 1) → A a)) + ((a : fst C 1) → hLevelTrunc n (A a)) + choicefun' n = compIso (mapCompIso (domIsoDep (equivToIso fin→))) + (compIso (equivToIso (_ , satAC' n (λ a → A (fst fin→ a)))) + (invIso (domIsoDep (equivToIso fin→)))) + + + choicefun≡ : (n : ℕ) → Iso.fun (choicefun' n) ≡ choiceMap n + choicefun≡ zero = refl + choicefun≡ (suc n) = funExt (TR.elim + (λ _ → isOfHLevelPath (suc n) + (isOfHLevelΠ (suc n) (λ _ → isOfHLevelTrunc (suc n))) _ _) + λ f → funExt λ a → cong ∣_∣ + (funExt⁻ ((Iso.leftInv (domIsoDep (equivToIso fin→))) f) a)) diff --git a/Cubical/Data/FinData/Properties.agda b/Cubical/Data/FinData/Properties.agda index f48cbfb58e..f70b09c30b 100644 --- a/Cubical/Data/FinData/Properties.agda +++ b/Cubical/Data/FinData/Properties.agda @@ -13,7 +13,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Powerset open import Cubical.Foundations.Isomorphism -open import Cubical.Data.Sum +open import Cubical.Data.Sum hiding (map) open import Cubical.Data.Sigma open import Cubical.Data.FinData.Base as Fin open import Cubical.Data.Nat renaming (zero to ℕzero ; suc to ℕsuc diff --git a/Cubical/Foundations/Equiv/Properties.agda b/Cubical/Foundations/Equiv/Properties.agda index 05d146f08c..b629b6ecbb 100644 --- a/Cubical/Foundations/Equiv/Properties.agda +++ b/Cubical/Foundations/Equiv/Properties.agda @@ -258,3 +258,18 @@ isPointedTarget→isEquiv→isEquiv : {A B : Type ℓ} (f : A → B) → (B → isEquiv f) → isEquiv f equiv-proof (isPointedTarget→isEquiv→isEquiv f hf) = λ y → equiv-proof (hf y) y + +module _ {ℓ ℓ' ℓ''} {A : Type ℓ} {A' : Type ℓ'} {C : A → Type ℓ''} (is : Iso A' A) where + private + is* = iso→HAEquiv is .snd + + domIsoDep : Iso ((a : A) → C a) ((a : A') → C (Iso.fun is a)) + Iso.fun domIsoDep f x = f (Iso.fun is x) + Iso.inv domIsoDep f x = subst C (isHAEquiv.rinv is* x) (f (Iso.inv is x)) + Iso.rightInv domIsoDep f = + funExt λ x → (λ j → subst C (isHAEquiv.com is* x (~ j)) + (f (Iso.inv is (Iso.fun is x)))) + ∙ λ j → transp (λ i → C (Iso.fun is (isHAEquiv.linv is* x (i ∨ j)))) j + (f (isHAEquiv.linv is* x j)) + Iso.leftInv domIsoDep f j x = + transp (λ i → C (isHAEquiv.rinv is* x (i ∨ j))) j (f (isHAEquiv.rinv is* x j)) diff --git a/Cubical/HITs/SequentialColimit/Base.agda b/Cubical/HITs/SequentialColimit/Base.agda index b8f4add21c..67c8a7a509 100644 --- a/Cubical/HITs/SequentialColimit/Base.agda +++ b/Cubical/HITs/SequentialColimit/Base.agda @@ -357,3 +357,106 @@ private (transport (λ i → (x : isoToPath (invIso (Iso-FinSeqColim-Top X m)) i) → f (ua-unglue (isoToEquiv (invIso (Iso-FinSeqColim-Top X m))) i x) ≡ g (ua-unglue (isoToEquiv (invIso (Iso-FinSeqColim-Top X m))) i x)) h) + + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sigma + +_<*_ : (n m : ℕ) → Type +n <* zero = ⊥ +zero <* suc m = Unit +suc n <* suc m = n <* m + +Fin* : (n : ℕ) → Type +Fin* n = Σ[ m ∈ ℕ ] (m <* n) + +fsuc* : {n : ℕ} → Fin* n → Fin* (suc n) +fst (fsuc* {n = n} (x , p)) = suc x +snd (fsuc* {n = suc n} (x , p)) = p + +<*-trans-suc : {n m : ℕ} → n <* m → n <* suc m +<*-trans-suc {n = zero} {suc m} x = tt +<*-trans-suc {n = suc n} {suc m} x = <*-trans-suc {n = n} x + +injectSuc* : {n : ℕ} → Fin* n → Fin* (suc n) +fst (injectSuc* {n = n} (x , p)) = x +snd (injectSuc* {n = suc n} (x , p)) = <*-trans-suc {n = x} p + +fsuc-injectSuc* : {m : ℕ} (n : Fin* m) → injectSuc* {n = suc m} (fsuc* {n = m} n) ≡ fsuc* (injectSuc* n) +fsuc-injectSuc* {m = suc m} (x , p) = refl + +<*sucm : {m : ℕ} → m <* suc m +<*sucm {m = zero} = tt +<*sucm {m = suc m} = <*sucm {m = m} + +flast* : {m : ℕ} → Fin* (suc m) +fst (flast* {m = m}) = m +snd (flast* {m = m}) = <*sucm {m = m} + +fzero* : {m : ℕ} → Fin* (suc m) +fzero* = 0 , tt + +elimFin* : ∀ {ℓ} {m : ℕ} {A : Fin* (suc m) → Type ℓ} + (max : A flast*) + (f : (x : Fin* m) → A (injectSuc* x)) + → (x : _) → A x +elimFin* {m = zero} {A = A} max f (zero , p) = max +elimFin* {m = suc m} {A = A} max f (zero , p) = f (zero , tt) +elimFin* {m = suc zero} {A = A} max f (suc zero , p) = max +elimFin* {m = suc (suc m)} {A = A} max f (suc x , p) = + elimFin* {m = suc m} {A = λ x → A (fsuc* x)} max (λ t → f (fsuc* t)) (x , p) + +elimFin*-alt : ∀ {ℓ} {m : ℕ} {A : Fin* (suc m) → Type ℓ} + (max : A fzero*) + (f : (x : Fin* m) → A (fsuc* x)) + → (x : _) → A x +elimFin*-alt {m = zero} max f (zero , p) = max +elimFin*-alt {m = suc m} max f (zero , p) = max +elimFin*-alt {m = suc m} max f (suc x , p) = f (x , p) + +elimFin*β : ∀ {ℓ} {m : ℕ} {A : Fin* (suc m) → Type ℓ} + (max : A flast*) + (f : (x : Fin* m) → A (injectSuc* x)) + → ((elimFin* {A = A} max f flast* ≡ max)) + × ((x : Fin* m) → elimFin* {A = A} max f (injectSuc* x) ≡ f x) +elimFin*β {m = zero} {A = A} max f = refl , λ {()} +elimFin*β {m = suc zero} {A = A} max f = refl , λ {(zero , p) → refl} +elimFin*β {m = suc (suc m)} {A = A} max f = + elimFin*β {m = (suc m)} {A = λ x → A (fsuc* x)} max _ .fst + , elimFin*-alt {m = (suc m)} {A = λ x → elimFin* max f (injectSuc* {n = suc (suc m)} x) ≡ f x} + refl + (elimFin*β {m = (suc m)} {A = λ x → A (fsuc* x)} max _ .snd) + +data FinSeqColim* (m : ℕ) (X : Sequence ℓ) : Type ℓ where + f*incl : {n : Fin* (suc m)} → X .obj (fst n) → FinSeqColim* m X + f*push : {n : Fin* m} (x : X .obj (fst n)) + → f*incl {n = injectSuc* n} x ≡ f*incl {n = fsuc* n} (X .map x) + +FinSeqColim*→Colim : (m : ℕ) {X : Sequence ℓ} → FinSeqColim* m X → SeqColim X +FinSeqColim*→Colim m (f*incl x) = incl x +FinSeqColim*→Colim m (f*push x i) = push x i + +module _ (X : Sequence ℓ) where + ShiftSeq* : Sequence ℓ + obj ShiftSeq* m = obj X (suc m) + map ShiftSeq* = map X + + F : (m : ℕ) → FinSeqColim* m ShiftSeq* → FinSeqColim* (suc m) X + F m (f*incl {n = n} x) = f*incl {n = fsuc* n} x + F (suc m) (f*push {n = n , p} x i) = f*push {n = suc n , p} x i + + G : (m : ℕ) → FinSeqColim* (suc m) X → FinSeqColim* m ShiftSeq* + G m (f*incl {n = zero , p} x) = f*incl {n = zero , p} (map X x) + G m (f*incl {n = suc n , p} x) = f*incl {n = n , p} x + G m (f*push {n = zero , p} x i) = f*incl {n = zero , p} (map X x) + G (suc m) (f*push {n = suc n , p} x i) = f*push {n = n , p} x i + + F→G→F : (m : ℕ) → (x : FinSeqColim* m ShiftSeq*) → G m (F m x) ≡ x + F→G→F m (f*incl {n = n} x) = refl + F→G→F (suc m) (f*push {n = n} x i) = refl + + G→F→G : (m : ℕ) → (x : FinSeqColim* (suc m) X) → F m (G m x) ≡ x + G→F→G m (f*incl {n = zero , p} x) = sym (f*push {n = zero , p} x) + G→F→G m (f*incl {n = suc n , p} x) = refl + G→F→G m (f*push {n = zero , p} x i) j = f*push {n = zero , p} x (~ j ∨ i) + G→F→G (suc m) (f*push {n = suc n , p} x i) = refl From 54c7919a18c48e72ccf111c9179fe0f5c3da0a4a Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 1 Mar 2024 19:08:20 +0100 Subject: [PATCH 27/73] stuff --- .../AbGroup/Instances/FreeAbGroup.agda | 193 ++++---- Cubical/Algebra/AbGroup/Properties.agda | 3 +- Cubical/Algebra/ChainComplex/Finite.agda | 5 +- Cubical/Algebra/ChainComplex/Homology.agda | 47 +- Cubical/Algebra/Group/Base.agda | 4 +- Cubical/Algebra/Group/Instances/Int.agda | 12 +- Cubical/Axiom/Choice.agda | 30 +- Cubical/Data/CW/Base.agd | 0 Cubical/Data/CW/Base.agda | 3 +- Cubical/Data/CW/ChainComplex.agda | 3 +- Cubical/Data/CW/Map.agda | 11 +- Cubical/Data/CW/Properties.agda | 5 +- Cubical/Data/Fin/Inductive.agda | 217 +++++++++ Cubical/Data/Fin/Inductive/Base.agda | 75 +++ Cubical/Data/Fin/Inductive/Properties.agda | 164 +++++++ Cubical/Data/FinSequence/Base.agda | 3 +- Cubical/Data/FinSequence/Properties.agda | 3 +- Cubical/Data/Int/Base.agda | 8 +- Cubical/Data/Int/Properties.agda | 8 +- Cubical/HITs/FreeAbGroup/Base.agda | 5 +- Cubical/HITs/SequentialColimit/Base.agda | 445 +++--------------- Cubical/HITs/SphereBouquet/Degree.agda | 41 +- Cubical/HITs/SphereBouquet/Properties.agda | 3 +- Cubical/ZCohomology/GroupStructure.agda | 3 +- 24 files changed, 719 insertions(+), 572 deletions(-) create mode 100644 Cubical/Data/CW/Base.agd create mode 100644 Cubical/Data/Fin/Inductive.agda create mode 100644 Cubical/Data/Fin/Inductive/Base.agda create mode 100644 Cubical/Data/Fin/Inductive/Properties.agda diff --git a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda index d2757f95fe..3a464dfbef 100644 --- a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda +++ b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.AbGroup.Instances.FreeAbGroup where open import Cubical.Foundations.Prelude @@ -10,7 +10,8 @@ open import Cubical.Data.Sigma open import Cubical.Data.Nat hiding (_·_) renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) -open import Cubical.Data.Fin +open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Empty as ⊥ open import Cubical.HITs.FreeAbGroup @@ -38,74 +39,71 @@ module _ {A : Type ℓ} where -- generator of ℤ[Fin_] ℤFinGenerator : {n : ℕ} (k : Fin n) → ℤ[Fin n ] .fst -ℤFinGenerator {n = n} k s with (fst k ≟ fst s) +ℤFinGenerator {n = n} k s with (fst k ≟ᵗ fst s) ... | lt x = 0 ... | eq x = 1 ... | gt x = 0 ℤFinGeneratorComm : {n : ℕ} (x y : Fin n) → ℤFinGenerator x y ≡ ℤFinGenerator y x -ℤFinGeneratorComm x y with (fst x ≟ fst y) | (fst y ≟ fst x) +ℤFinGeneratorComm x y with (fst x ≟ᵗ fst y) | (fst y ≟ᵗ fst x) ... | lt x₁ | lt x₂ = refl -... | lt x₁ | eq x₂ = ⊥.rec (¬m Date: Sun, 3 Mar 2024 03:52:29 +0100 Subject: [PATCH 28/73] stuff --- Cubical/Algebra/ChainComplex/Homology.agda | 131 ++- Cubical/Data/CW/Base.agda | 2 +- Cubical/Data/CW/ChainComplex.agda | 4 +- Cubical/Data/CW/Homotopy.agda | 775 ++++++++++++++++++ Cubical/Data/CW/Map.agda | 376 ++++++++- .../PropositionalTruncation/Properties.agda | 35 +- 6 files changed, 1287 insertions(+), 36 deletions(-) create mode 100644 Cubical/Data/CW/Homotopy.agda diff --git a/Cubical/Algebra/ChainComplex/Homology.agda b/Cubical/Algebra/ChainComplex/Homology.agda index bada452445..f135a86a94 100644 --- a/Cubical/Algebra/ChainComplex/Homology.agda +++ b/Cubical/Algebra/ChainComplex/Homology.agda @@ -83,6 +83,131 @@ homology n C = ker∂n / img∂+1⊂ker∂n isNormalImSubGroup = isNormalIm ∂' (λ x y → kerGroup≡ ∂n (C1.+Comm (fst x) (fst y))) +-- -- Induced maps on cohomology by finite chain complex maps/homotopies +-- module _ where +-- finChainComplexMap→HomologyMap : {C D : ChainComplex ℓ} (m : ℕ) +-- → (ϕ : finChainComplexMap (suc (suc m)) C D) +-- → (n : Fin (suc m)) +-- → GroupHom (homology (fst n) C) (homology (fst n) D) +-- finChainComplexMap→HomologyMap {C = C} {D} m mp (n , p) = main +-- where +-- ineq3 : suc n <ᵗ suc (suc (suc m)) +-- ineq3 = <ᵗ-trans p <ᵗsucm + +-- ineq4 : n <ᵗ suc (suc m) +-- ineq4 = ineq3 + +-- ϕ = fchainmap mp +-- ϕcomm = fbdrycomm mp + +-- lem : (k : ℕ) {p q : _} (f : fst (chain C k)) +-- → fst (ϕ (k , p)) f ≡ fst (ϕ (k , q)) f +-- lem k {p} {q} f i = fst (ϕ (k , pq i)) f +-- where +-- pq : p ≡ q +-- pq = isProp<ᵗ _ _ + +-- fun : homology n C .fst → homology n D .fst +-- fun = SQ.elim (λ _ → squash/) f +-- λ f g → PT.rec (GroupStr.is-set (homology n D .snd) _ _ ) +-- λ r → eq/ _ _ ∣ (ϕ (suc (suc n) , p) .fst (fst r)) +-- , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) +-- ((funExt⁻ (cong fst (ϕcomm (suc n , _))) (fst r) +-- ∙∙ cong (fst (ϕ (suc n , _))) (cong fst (snd r)) +-- ∙∙ (IsGroupHom.pres· (snd (ϕ (suc n , _) )) _ _ +-- ∙ cong₂ (AbGroupStr._+_ (snd (chain D (suc n)))) +-- (lem (suc n) (fst f)) +-- (IsGroupHom.presinv (snd (ϕ (suc n , _) )) _ +-- ∙ cong (snd (chain D (suc n)) .AbGroupStr.-_) +-- (lem (suc n) (fst g)))))) ∣₁ +-- where +-- f : _ → homology n D .fst +-- f (a , b) = [ ϕ (suc n , ineq3) .fst a -- (ϕ {!n!} .fst a) +-- , ((λ i → fst (ϕcomm (n , ineq3) i) a) +-- ∙∙ cong (fst (ϕ (n , _))) b +-- ∙∙ IsGroupHom.pres1 (snd (ϕ (n , _)))) ] + + +-- main : GroupHom (homology n C) (homology n D) +-- fst main = fun +-- snd main = +-- makeIsGroupHom +-- (SQ.elimProp2 (λ _ _ → GroupStr.is-set (snd (homology n D)) _ _) +-- λ a b → cong [_] +-- (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) +-- (IsGroupHom.pres· (snd (ϕ (suc n , _) )) _ _))) + +-- finChainComplexMap→HomologyMapComp : {C D E : ChainComplex ℓ} {m : ℕ} +-- → (ϕ : finChainComplexMap (suc (suc m)) C D) (ψ : finChainComplexMap (suc (suc m)) D E) +-- → (n : Fin (suc m)) +-- → finChainComplexMap→HomologyMap m (compFinChainMap ϕ ψ) n +-- ≡ compGroupHom (finChainComplexMap→HomologyMap m ϕ n) +-- (finChainComplexMap→HomologyMap m ψ n) +-- finChainComplexMap→HomologyMapComp {E = E} _ _ n = +-- Σ≡Prop (λ _ → isPropIsGroupHom _ _) +-- (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (fst n) E)) _ _) +-- λ _ → cong [_] +-- (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain E (fst n))) _ _) refl))) + +-- finChainComplexMap→HomologyMapId : {C : ChainComplex ℓ} {m : ℕ} (n : Fin (suc m)) +-- → finChainComplexMap→HomologyMap m (idFinChainMap (suc (suc m)) C) n ≡ idGroupHom +-- finChainComplexMap→HomologyMapId {C = C} n = +-- Σ≡Prop (λ _ → isPropIsGroupHom _ _) +-- (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (fst n) C)) _ _) +-- λ _ → cong [_] +-- (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C (fst n))) _ _) refl))) + +-- finChainComplexEquiv→HomoglogyIso : {C D : ChainComplex ℓ} (m : ℕ) (f : C ≃⟨ (suc (suc m)) ⟩Chain D) +-- → (n : Fin (suc m)) → GroupIso (homology (fst n) C) (homology (fst n) D) +-- Iso.fun (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = +-- finChainComplexMap→HomologyMap m f n .fst +-- Iso.inv (fst (finChainComplexEquiv→HomoglogyIso m f n)) = +-- finChainComplexMap→HomologyMap m (invFinChainMap f) n .fst +-- Iso.rightInv (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = +-- funExt⁻ (cong fst (sym (finChainComplexMap→HomologyMapComp +-- (invFinChainMap (f , eqs)) f n)) +-- ∙∙ cong (λ f → fst (finChainComplexMap→HomologyMap m f n)) +-- (finChainComplexMap≡ λ r +-- → Σ≡Prop (λ _ → isPropIsGroupHom _ _) +-- (funExt (secEq (_ , eqs r)))) +-- ∙∙ cong fst (finChainComplexMap→HomologyMapId n)) +-- Iso.leftInv (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = +-- funExt⁻ (cong fst (sym (finChainComplexMap→HomologyMapComp f +-- (invFinChainMap (f , eqs)) n)) +-- ∙∙ cong (λ f → fst (finChainComplexMap→HomologyMap m f n)) +-- (finChainComplexMap≡ +-- (λ n → Σ≡Prop (λ _ → isPropIsGroupHom _ _) +-- (funExt (retEq (_ , eqs n))))) +-- ∙∙ cong fst (finChainComplexMap→HomologyMapId n)) +-- snd (finChainComplexEquiv→HomoglogyIso m (f , eqs) n) = +-- finChainComplexMap→HomologyMap m f n .snd + + +-- finChainHomotopy→HomologyPath : {A B : ChainComplex ℓ} {m : ℕ} +-- (f g : finChainComplexMap (suc (suc m)) A B) +-- → finChainHomotopy (suc (suc m)) f g +-- → (n : Fin (suc m)) +-- → finChainComplexMap→HomologyMap m f n +-- ≡ finChainComplexMap→HomologyMap m g n +-- finChainHomotopy→HomologyPath {A = A} {B = B} {m = m} f g ϕ n = +-- Σ≡Prop (λ _ → isPropIsGroupHom _ _) +-- (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (fst n) _)) _ _) +-- λ {(a , p) → eq/ _ _ +-- ∣ (finChainHomotopy.fhtpy ϕ (suc (fst n) , pf) .fst a) +-- , (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain B (fst n))) _ _) +-- (sym ((funExt⁻ (cong fst (finChainHomotopy.fbdryhtpy ϕ _)) a) +-- ∙ cong₂ _+B_ refl +-- (cong (fst (finChainHomotopy.fhtpy ϕ _)) p +-- ∙ IsGroupHom.pres1 (snd (finChainHomotopy.fhtpy ϕ _))) +-- ∙ AbGroupStr.+IdR (snd (chain B (suc (fst n)))) _))) ∣₁})) +-- where +-- open GroupTheory (AbGroup→Group (chain B (suc (fst n)))) +-- pf : suc (fst n) <ᵗ suc (suc (suc m)) +-- pf = <ᵗ-trans (snd n) <ᵗsucm + +-- invB = GroupStr.inv (snd (AbGroup→Group (chain B (suc (fst n))))) +-- _+B_ = AbGroupStr._+_ (snd (chain B (suc (fst n)))) + -- Induced maps on cohomology by finite chain complex maps/homotopies module _ where @@ -92,12 +217,6 @@ module _ where → GroupHom (homology (fst n) C) (homology (fst n) D) finChainComplexMap→HomologyMap {C = C} {D} m mp (n , p) = main where - ineq1 : suc n <ᵗ suc m - ineq1 = p -- suc-≤-suc p - - ineq2 : suc (suc n) <ᵗ suc (suc m) - ineq2 = p - ineq3 : suc n <ᵗ suc (suc m) ineq3 = <ᵗ-trans p <ᵗsucm diff --git a/Cubical/Data/CW/Base.agda b/Cubical/Data/CW/Base.agda index a4d1d70952..b120476c1c 100644 --- a/Cubical/Data/CW/Base.agda +++ b/Cubical/Data/CW/Base.agda @@ -125,7 +125,7 @@ finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p -- morphisms -_→ᶜʷ_ : CW ℓ → CW ℓ → Type ℓ +_→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') C →ᶜʷ D = fst C → fst D --the cofibre of the inclusion of X n into X (suc n) diff --git a/Cubical/Data/CW/ChainComplex.agda b/Cubical/Data/CW/ChainComplex.agda index 720b6c30f6..8d5a49b792 100644 --- a/Cubical/Data/CW/ChainComplex.agda +++ b/Cubical/Data/CW/ChainComplex.agda @@ -187,5 +187,5 @@ module _ {ℓ} (C : CWskel ℓ) where bdry²=0 CW-ChainComplex n = ∂∂≡0 n -- Cellular homology - Hᶜʷ : (n : ℕ) → Group₀ - Hᶜʷ n = homology n CW-ChainComplex + Hˢᵏᵉˡ : (n : ℕ) → Group₀ + Hˢᵏᵉˡ n = homology n CW-ChainComplex diff --git a/Cubical/Data/CW/Homotopy.agda b/Cubical/Data/CW/Homotopy.agda new file mode 100644 index 0000000000..05468ce9a8 --- /dev/null +++ b/Cubical/Data/CW/Homotopy.agda @@ -0,0 +1,775 @@ +{-# OPTIONS --safe --lossy-unification #-} + +{- +This file contains +1. Definitions of cellular homotopies and their finite counterpart +2. A proof that (finite) cellular homotopies induce (finite) chain complex maps +-} + +module Cubical.Data.CW.Homotopy where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) +open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool hiding (_≟_ ; isProp≤) +open import Cubical.Data.CW +open import Cubical.Data.CW.Map +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence +open import Cubical.Data.CW.ChainComplex + + +open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Degree +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.SphereBouquet.Degree + +open import Cubical.Relation.Nullary + +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup +open import Cubical.Algebra.ChainComplex + +private + variable + ℓ ℓ' ℓ'' : Level + +-- A cellular homotopy between two cellular maps +record cellHom {C : CWskel ℓ} {D : CWskel ℓ'} (f g : cellMap C D) : Type (ℓ-max ℓ ℓ') where + open SequenceMap + field + hom : (n : ℕ) → (x : C .fst n) → CW↪ D n (f .map n x) ≡ CW↪ D n (g .map n x) + coh : (n : ℕ) → (c : C .fst n) + → Square (cong (CW↪ D (suc n)) (hom n c)) + (hom (suc n) (CW↪ C n c)) + (cong (CW↪ D (suc n)) (f .comm n c)) + (cong (CW↪ D (suc n)) (g .comm n c)) + +-- Finite cellular homotopies +record finCellHom (m : ℕ) {C : CWskel ℓ} {D : CWskel ℓ'} + (f g : finCellMap m C D) : Type (ℓ-max ℓ ℓ') where + open FinSequenceMap + field + fhom : (n : Fin (suc m)) (x : C .fst (fst n)) + → CW↪ D (fst n) (f .fmap n x) ≡ CW↪ D (fst n) (g .fmap n x) + fcoh : (n : Fin m) (c : C .fst (fst n)) + → Square (cong (CW↪ D (suc (fst n))) (fhom (injectSuc n) c)) + (fhom (fsuc n) (CW↪ C (fst n) c)) + (cong (CW↪ D (suc (fst n))) (f .fcomm n c)) + (cong (CW↪ D (suc (fst n))) (g .fcomm n c)) + +open finCellHom + +finCellHom↓ : {m : ℕ} {C : CWskel ℓ} {D : CWskel ℓ'} + {f g : finCellMap (suc m) C D} + → finCellHom (suc m) f g → finCellHom m (finCellMap↓ f) (finCellMap↓ g) +fhom (finCellHom↓ ϕ) n x = fhom ϕ (injectSuc n) x +fcoh (finCellHom↓ {m = suc m} ϕ) n x = fcoh ϕ (injectSuc n) x + +-- Extracting a map between sphere bouquets from a MMmap +cofibIso : (n : ℕ) (C : CWskel ℓ) + → Iso (Susp (cofibCW n C)) (SphereBouquet (suc n) (CWskel-fields.A C n)) +cofibIso n C = + compIso (congSuspIso (BouquetIso-gen n + (CWskel-fields.card C n) + (CWskel-fields.α C n) + (CWskel-fields.e C n))) + sphereBouquetSuspIso + +-- Building a chain homotopy from a cell homotopy +module preChainHomotopy (m : ℕ) (C : CWskel ℓ) (D : CWskel ℓ') + (f g : finCellMap m C D) (H : finCellHom m f g) (n : Fin m) where + open FinSequenceMap + private + ℤ[AC_] = CWskel-fields.ℤ[A_] C + ℤ[AD_] = CWskel-fields.ℤ[A_] D + + -- the homotopy expressed as a map Susp (cofibCW n C) → cofibCW (suc n) D + Hn+1/Hn : Susp (cofibCW (fst n) C) → cofibCW (suc (fst n)) D + Hn+1/Hn north = inl tt + Hn+1/Hn south = inl tt + Hn+1/Hn (merid (inl tt) i) = inl tt + Hn+1/Hn (merid (inr x) i) = + ((push (f .fmap (fsuc n) x)) + ∙∙ (cong inr (H .fhom (fsuc n) x)) + ∙∙ (sym (push (g .fmap (fsuc n) x)))) i + Hn+1/Hn (merid (push x j) i) = + hcomp (λ k → λ { (i = i0) → push (f .fcomm n x j) (~ k) + ; (i = i1) → push (g .fcomm n x j) (~ k) + ; (j = i0) → push (fhom H (injectSuc n) x i) (~ k) }) + (inr (H .fcoh n x j i)) + + -- the homotopy expressed as a map of sphere bouquets + bouquetHomotopy : SphereBouquet (suc (fst n)) (CWskel-fields.A C (fst n)) + → SphereBouquet (suc (fst n)) (CWskel-fields.A D (suc (fst n))) + bouquetHomotopy = Iso.fun bouquetIso ∘ Hn+1/Hn ∘ Iso.inv (cofibIso (fst n) C) + where + bouquetIso = BouquetIso-gen (suc (fst n)) + (CWskel-fields.card D (suc (fst n))) + (CWskel-fields.α D (suc (fst n))) + (CWskel-fields.e D (suc (fst n))) + + -- the homotopy as a map of abelian groups + chainHomotopy : AbGroupHom (ℤ[AC (fst n) ]) (ℤ[AD (suc (fst n)) ]) + chainHomotopy = bouquetDegree bouquetHomotopy + +-- Now, we would like to prove the chain homotopy equation ∂H + H∂ = f - g +-- MMmaps (Meridian-to-Meridian maps) are a convenient abstraction for the kind of maps +-- that we are going to manipulate +module MMmaps (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) (n : Fin m) where + MMmap : (m1 m2 : (x : C .fst (suc (fst n))) + → cofibCW (fst n) D) → Type (ℓ-max ℓ ℓ') + MMmap m1 m2 = (x : C .fst (fst n)) + → m1 (CW↪ C (fst n) x) ≡ m2 (CW↪ C (fst n) x) + + -- the suspension of a cell map as a MMmap + MMΣcellMap : (f : finCellMap m C D) + → MMmap (λ x → (inr (f .FinSequenceMap.fmap (fsuc n) x))) (λ x → inl tt) + MMΣcellMap f x = + sym (push (f .FinSequenceMap.fmap (injectSuc n) x) + ∙ (cong inr (f .FinSequenceMap.fcomm n x))) + + -- Addition of MMmaps + MMmap-add : (m1 m2 m3 : (x : C .fst (suc (fst n))) + → cofibCW (fst n) D) + → MMmap m1 m2 → MMmap m2 m3 → MMmap m1 m3 + MMmap-add m1 m2 m3 e1 e2 x = (e1 x) ∙ (e2 x) + + -- Extracting a map between suspensions of cofibCWs from a MMmap + realiseMMmap : (m1 m2 : (x : C .fst (suc (fst n))) → cofibCW (fst n) D) + → MMmap m1 m2 → Susp (cofibCW (fst n) C) → Susp (cofibCW (fst n) D) + realiseMMmap m1 m2 e north = north + realiseMMmap m1 m2 e south = north + realiseMMmap m1 m2 e (merid (inl tt) i) = north + realiseMMmap m1 m2 e (merid (inr x) i) = + (merid (m1 x) ∙∙ refl ∙∙ (sym (merid (m2 x)))) i + realiseMMmap m1 m2 e (merid (push x j) i) = + hcomp (λ k → λ { (i = i0) → merid (m1 (CW↪ C (fst n) x)) (~ k) + ; (i = i1) → merid (m2 (CW↪ C (fst n) x)) (~ k) + ; (j = i0) → merid (e x i) (~ k) }) + (south) + + -- Extracting a map between sphere bouquets from a MMmap + bouquetMMmap : (m1 m2 : (x : C .fst (suc (fst n))) → cofibCW (fst n) D) + → MMmap m1 m2 + → SphereBouquet (suc (fst n)) (CWskel-fields.A C (fst n)) + → SphereBouquet (suc (fst n)) (CWskel-fields.A D (fst n)) + bouquetMMmap m1 m2 f = + Iso.fun (cofibIso (fst n) D) + ∘ realiseMMmap m1 m2 f + ∘ Iso.inv (cofibIso (fst n) C) + + +-- Expressing the chain homotopy at the level of MMmaps +-- There, it is easy to prove the chain homotopy equation +module MMchainHomotopy* (m : ℕ) (C : CWskel ℓ) (D : CWskel ℓ') + (f g : finCellMap m C D) (H : finCellHom m f g) (n : Fin m) where + open FinSequenceMap + + open MMmaps C D + + merid-f merid-g merid-tt : (x : C .fst (suc (fst n))) → cofibCW (fst n) D + merid-f = λ x → inr (f .fmap (fsuc n) x) + merid-g = λ x → inr (g .fmap (fsuc n) x) + merid-tt = λ x → inl tt + + MM∂H : MMmap m n merid-f merid-g + MM∂H x = (sym (cong inr (f .fcomm n x))) + ∙∙ (cong inr (fhom H (injectSuc n) x)) + ∙∙ (cong inr (g .fcomm n x)) + + ww = MMΣcellMap + -- the suspension of f as a MMmap + MMΣf : MMmap m n merid-f merid-tt + MMΣf = MMΣcellMap m n f + + -- the suspension of g as a MMmap + MMΣg : MMmap m n merid-g merid-tt + MMΣg = MMΣcellMap m n g + + -- the suspension of H∂ as a MMmap + MMΣH∂ : MMmap m n merid-tt merid-tt + MMΣH∂ x = sym ((push (f .fmap (injectSuc n) x)) + ∙∙ (cong inr (H .fhom (injectSuc n) x)) + ∙∙ (sym (push (g .fmap (injectSuc n) x)))) + + -- the chain homotopy equation at the level of MMmaps + MMchainHomotopy : ∀ x → + MMmap-add m n merid-f merid-tt merid-tt + (MMmap-add m n merid-f merid-g merid-tt MM∂H MMΣg) MMΣH∂ x + ≡ MMΣf x + MMchainHomotopy x = + sym (doubleCompPath-elim (MM∂H x) (MMΣg x) (MMΣH∂ x)) ∙ aux2 + where + aux : Square (MMΣf x) (MMΣg x) (MM∂H x) (sym (MMΣH∂ x)) + aux i j = + hcomp (λ k → + λ {(i = i0) → compPath-filler (push (f .fmap (injectSuc n) x)) + (λ i₁ → inr (f .fcomm n x i₁)) k (~ j) + ; (i = i1) → compPath-filler (push (g .fmap (injectSuc n) x)) + (λ i₁ → inr (g .fcomm n x i₁)) k (~ j) + ; (j = i1) → (push (f .fmap (injectSuc n) x) + ∙∙ (λ i → inr (H .fhom (injectSuc n) x i)) + ∙∙ (λ i₁ → push (g .fmap (injectSuc n) x) (~ i₁))) i}) + (doubleCompPath-filler + (push (f .fmap (injectSuc n) x)) + (λ i → (inr (H .fhom (injectSuc n) x i))) + (λ i₁ → push (g .fmap (injectSuc n) x) (~ i₁)) j i) + + aux2 : (MM∂H x ∙∙ MMΣg x ∙∙ MMΣH∂ x) ≡ MMΣf x + aux2 i j = + hcomp (λ k → λ { (j = i0) → MM∂H x ((~ i) ∧ (~ k)) + ; (j = i1) → MMΣH∂ x (i ∨ k) + ; (i = i1) → MMΣf x j }) + (aux (~ i) j) + +-- Now we want to transform our MMmap equation to the actual equation +-- First, we connect the involved MMmaps to cofibCW maps +module realiseMMmap (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) + (f g : finCellMap m C D) (H : finCellHom m f g) (n : Fin m) where + open FinSequenceMap + open MMmaps C D + open MMchainHomotopy* m C D f g H + open preChainHomotopy m C D f g H + + -- an alternative extraction function, that will be useful in some computations + realiseMMmap2 : (n : Fin m) + (m1 m2 : (x : C .fst (suc (fst n))) → cofibCW (fst n) D) + → MMmap m n m1 m2 → Susp (cofibCW (fst n) C) → Susp (cofibCW (fst n) D) + realiseMMmap2 n m1 m2 e north = north + realiseMMmap2 n m1 m2 e south = north + realiseMMmap2 n m1 m2 e (merid (inl tt) i) = north + realiseMMmap2 n m1 m2 e (merid (inr x) i) = + (merid (m1 x) ∙∙ refl ∙∙ (sym (merid (m2 x)))) i + realiseMMmap2 n m1 m2 e (merid (push x j) i) = + hcomp (λ k → λ { (i = i0) → merid (e x (~ j)) (~ k) + ; (i = i1) → merid (m2 (CW↪ C (fst n) x)) (~ k) + ; (j = i0) → merid (m2 (CW↪ C (fst n) x)) (~ k) }) + (south) + + -- auxiliary lemma which says the two realisation functions are equal + realiseMMmap1≡2 : (n : Fin m) (m1 m2 : (x : C .fst (suc (fst n))) + → cofibCW (fst n) D) (e : MMmap m n m1 m2) (x : Susp (cofibCW (fst n) C)) + → realiseMMmap m n m1 m2 e x ≡ realiseMMmap2 n m1 m2 e x + realiseMMmap1≡2 n m1 m2 e north = refl + realiseMMmap1≡2 n m1 m2 e south = refl + realiseMMmap1≡2 n m1 m2 e (merid (inl tt) i) = refl + realiseMMmap1≡2 n m1 m2 e (merid (inr x) i) = refl + realiseMMmap1≡2 n m1 m2 e (merid (push x j) i) l = + hcomp (λ k → λ { (i = i0) → merid (e x ((~ j) ∧ l)) (~ k) + ; (i = i1) → merid (m2 (CW↪ C (fst n) x)) (~ k) + ; (j = i0) → merid (e x (i ∨ l)) (~ k) }) + south + + -- realisation of MMΣf is equal to Susp f + realiseMMΣcellMap : (f : finCellMap m C D) (x : Susp (cofibCW (fst n) C)) + → realiseMMmap m n (λ x → (inr (f .fmap (fsuc n) x))) + (λ x → inl tt) (MMΣcellMap m n f) x + ≡ suspFun (prefunctoriality.fn+1/fn m f n) x + realiseMMΣcellMap f x = + realiseMMmap1≡2 n (λ x → (inr (f .fmap (fsuc n) x))) (λ x → inl tt) + (MMΣcellMap m n f) x ∙ aux x + where + aux : (x : Susp (cofibCW (fst n) C)) → + realiseMMmap2 n (λ x → (inr (f .fmap (fsuc n) x))) + (λ x → inl tt) (MMΣcellMap m n f) x + ≡ suspFun (prefunctoriality.fn+1/fn m f n) x + aux north = refl + aux south l = merid (inl tt) l + aux (merid (inl tt) i) l = merid (inl tt) (i ∧ l) + aux (merid (inr x) i) l = + hcomp (λ k → + λ { (i = i0) → merid (inr (f .fmap (fsuc n) x)) (~ k) + ; (i = i1) → merid (inl tt) (l ∨ (~ k)) + ; (l = i1) → merid (inr (f .fmap (fsuc n) x)) (~ k ∨ i) }) + south + aux (merid (push x j) i) l = + hcomp (λ k → + λ {(i = i0) → merid ((push (f .fmap (injectSuc n) x) + ∙ (cong inr (f .fcomm n x))) j) (~ k) + ; (i = i1) → merid (inl tt) (l ∨ (~ k)) + ; (j = i0) → merid (inl tt) ((i ∧ l) ∨ (~ k)) + ; (l = i1) → merid ((push (f .fmap (injectSuc n) x) + ∙ (cong inr (f .fcomm n x))) j) (i ∨ (~ k)) }) + south + + -- realisation of MMΣf is equal to Susp f + realiseMMΣf : (x : Susp (cofibCW (fst n) C)) → + realiseMMmap m n (merid-f n) (merid-tt n) (MMΣf n) x + ≡ suspFun (prefunctoriality.fn+1/fn m f n) x + realiseMMΣf = realiseMMΣcellMap f + + -- realisation of MMΣg is equal to Susp g + realiseMMΣg : (x : Susp (cofibCW (fst n) C)) → + realiseMMmap m n (merid-g n) (merid-tt n) (MMΣg n) x + ≡ suspFun (prefunctoriality.fn+1/fn m g n) x + realiseMMΣg = realiseMMΣcellMap g + + -- a compact version of ∂ ∘ H + cof∂H : Susp (cofibCW (fst n) C) → Susp (cofibCW (fst n) D) + cof∂H north = north + cof∂H south = north + cof∂H (merid (inl tt) i) = north + cof∂H (merid (inr x) i) = + ((merid (inr (f .fmap (fsuc n) x))) + ∙∙ refl + ∙∙ (sym (merid (inr (g .fmap (fsuc n) x))))) i + cof∂H (merid (push x j) i) = + hcomp (λ k → λ { (i = i0) → merid (inr (f .fcomm n x j)) (~ k) + ; (i = i1) → merid (inr (g .fcomm n x j)) (~ k) + ; (j = i0) → merid (inr (fhom H (injectSuc n) x i)) (~ k) }) + (south) + + -- realisation of MM∂H is equal to cof∂H + realiseMM∂H : (x : Susp (cofibCW (fst n) C)) → + realiseMMmap m n (merid-f n) (merid-g n) (MM∂H n) x + ≡ suspFun (to_cofibCW (fst n) D) (δ (suc (fst n)) D (Hn+1/Hn n x)) + realiseMM∂H x = aux2 x ∙ aux x + where + aux : (x : Susp (cofibCW (fst n) C)) + → cof∂H x + ≡ suspFun (to_cofibCW (fst n) D) (δ (suc (fst n)) D (Hn+1/Hn n x)) + aux north = refl + aux south = refl + aux (merid (inl tt) i) = refl + aux (merid (inr x) i) j = + hcomp (λ k → + λ { (i = i0) → merid (inr (f .fmap (fsuc n) x)) (~ k) + ; (i = i1) → merid (inr (g .fmap (fsuc n) x)) (~ k) + ; (j = i1) → suspFun (to_cofibCW (fst n) D) (δ (suc (fst n)) D + (doubleCompPath-filler (push (f .fmap (fsuc n) x)) + (cong inr (H .fhom (fsuc n) x)) + (sym (push (g .fmap (fsuc n) x))) k i)) }) + south + aux (merid (push x j) i) k = + hcomp (λ l → + λ { (i = i0) → merid (inr (f .fcomm n x j)) (~ l) + ; (i = i1) → merid (inr (g .fcomm n x j)) (~ l) + ; (j = i0) → merid (inr (fhom H (injectSuc n) x i)) (~ l) + ; (k = i1) → suspFun (to_cofibCW (fst n) D) (δ (suc (fst n)) D + (hfill (λ k → λ { (i = i0) → push (f .fcomm n x j) (~ k) + ; (i = i1) → push (g .fcomm n x j) (~ k) + ; (j = i0) → push (fhom H (injectSuc n) x i) (~ k)}) + (inS (inr (H .fcoh n x j i))) l))}) + south + + aux2 : (x : Susp (cofibCW (fst n) C)) → + realiseMMmap m n (λ x → (inr (f .fmap (fsuc n) x))) + (λ x → (inr (g .fmap (fsuc n) x))) (MM∂H n) x + ≡ cof∂H x + aux2 north = refl + aux2 south = refl + aux2 (merid (inl tt) i) = refl + aux2 (merid (inr x) i) = refl + aux2 (merid (push x j) i) l = + hcomp (λ k → + λ { (i = i0) → merid (inr (f .fcomm n x (j ∨ (~ l)))) (~ k) + ; (i = i1) → merid (inr (g .fcomm n x (j ∨ (~ l)))) (~ k) + ; (j = i0) → merid (doubleCompPath-filler + (sym (cong inr (f .fcomm n x))) + (cong inr (fhom H (injectSuc n) x)) + (cong inr (g .fcomm n x)) (~ l) i) (~ k) }) + south + + -- realisation of MMΣH∂ is equal to Susp H∂ + -- TODO: it is the same code as before. factorise! +realiseMMΣH∂ : (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) + (f g : finCellMap (suc m) C D) (H : finCellHom (suc m) f g) + (n : Fin m) (x : Susp (cofibCW (suc (fst n)) C)) → + MMmaps.realiseMMmap C D (suc m) (fsuc n) (λ x → inl tt) (λ x → inl tt) + (MMchainHomotopy*.MMΣH∂ (suc m) C D f g H (fsuc n) ) x + ≡ suspFun (preChainHomotopy.Hn+1/Hn (suc m) C D f g H (injectSuc n) + ∘ suspFun (to_cofibCW (fst n) C) + ∘ δ (suc (fst n)) C) x +realiseMMΣH∂ C D (suc m) f g H n x = + realiseMMmap1≡2 fzero (fsuc n) (λ x → inl tt) + (λ x → inl tt) (MMΣH∂ (fsuc n)) x ∙ aux x + where + open FinSequenceMap + open MMmaps C D + open MMchainHomotopy* (suc (suc m)) C D f g H + open preChainHomotopy (suc (suc m)) C D f g H + open realiseMMmap C D (suc (suc m)) f g H + + aux : (x : Susp (cofibCW (suc (fst n)) C)) → + realiseMMmap.realiseMMmap2 C D (suc (suc m)) f g H fzero (fsuc n) + (λ x₁ → inl tt) (λ x₁ → inl tt) + (MMchainHomotopy*.MMΣH∂ (suc (suc m)) C D f g H (fsuc n)) x + ≡ suspFun (Hn+1/Hn (injectSuc n) + ∘ (suspFun (to_cofibCW (fst n) C)) + ∘ (δ (suc (fst n)) C)) x + aux north = refl + aux south l = merid (inl tt) l + aux (merid (inl tt) i) l = merid (inl tt) (i ∧ l) + aux (merid (inr x) i) l = + hcomp (λ k → λ { (i = i0) → merid (inl tt) (~ k) + ; (i = i1) → merid (inl tt) (l ∨ (~ k)) + ; (l = i1) → merid (inl tt) (~ k ∨ i) }) + south + aux (merid (push x j) i) l = + hcomp (λ k → + λ { (i = i0) → merid (((push (f .fmap (injectSuc (fsuc n)) x)) + ∙∙ (cong inr (H .fhom (injectSuc (fsuc n)) x)) + ∙∙ (sym (push (g .fmap (injectSuc (fsuc n)) x)))) j) (~ k) + ; (i = i1) → merid (inl tt) (l ∨ (~ k)) + ; (j = i0) → merid (inl tt) ((i ∧ l) ∨ (~ k)) + ; (l = i1) → merid (((push (f .fmap (injectSuc (fsuc n)) x)) + ∙∙ (cong inr (H .fhom (injectSuc (fsuc n)) x)) + ∙∙ (sym (push (g .fmap (injectSuc (fsuc n)) x)))) j) + (i ∨ (~ k))}) + south + +-- Then, we connect the addition of MMmaps to the addition of abelian maps +module bouquetAdd where + -- keeping imports here for now + open import Cubical.ZCohomology.Base + open import Cubical.ZCohomology.Properties + open import Cubical.ZCohomology.GroupStructure + open import Cubical.HITs.Truncation as TR hiding (map) + open import Cubical.HITs.Sn + open import Cubical.HITs.S1 + open import Cubical.Foundations.Path + open import Cubical.ZCohomology.Groups.Sn + open import Cubical.HITs.SetTruncation as ST hiding (map) + + open MMmaps + + module _ (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) (n : Fin m) + (m1 m2 : (x : C .fst (suc (fst n))) → cofibCW (fst n) D) + (f : MMmap C D m n m1 m2) + (a : CWskel-fields.A D (fst n)) where + + bouquetMMmap∈cohom-raw : (t : CWskel-fields.A C (fst n)) + → S₊ (suc (fst n)) → S₊ (suc (fst n)) + bouquetMMmap∈cohom-raw t x = + pickPetal a (bouquetMMmap C D m n m1 m2 f (inr (t , x))) + + bouquetMMmap∈cohom : (t : CWskel-fields.A C (fst n)) + → S₊ (suc (fst n)) → coHomK (suc (fst n)) + bouquetMMmap∈cohom t x = ∣ bouquetMMmap∈cohom-raw t x ∣ₕ + + bouquetMMmap∈cohom' : (x : Susp (cofibCW (fst n) C)) → coHomK (suc (fst n)) + bouquetMMmap∈cohom' x = + ∣ pickPetal a (Iso.fun (cofibIso (fst n) D) + (realiseMMmap C D m n m1 m2 f x)) ∣ₕ + + realiseAdd-merid : (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) (n : Fin m) + (m1 m2 m3 : (x : C .fst (suc (fst n))) → cofibCW (fst n) D) + (f : MMmap C D m n m1 m2) + (g : MMmap C D m n m2 m3) + (b : _) + → Square (λ j → (realiseMMmap C D m n m1 m2 f (merid b j))) + (λ j → (realiseMMmap C D m n m1 m3 + (MMmap-add C D m n m1 m2 m3 f g) (merid b j))) + (λ _ → north) + (λ i → realiseMMmap C D m n m2 m3 g (merid b i)) + realiseAdd-merid C D m n m1 m2 m3 f g (inl tt) i j = north + realiseAdd-merid C D m n m1 m2 m3 f g (inr x) i j = + hcomp (λ k → λ { (i ∨ j = i0) → merid (m1 x) (~ k) + ; (i ∨ (~ j) = i0) → merid (m2 x) (~ k) + ; (i ∧ (~ j) = i1) → merid (m1 x) (~ k) + ; (i ∧ j = i1) → merid (m3 x) (~ k) + ; (j = i0) → merid (m1 x) (~ k) }) + south + realiseAdd-merid C D m n m1 m2 m3 f g (push a l) i j = + hcomp (λ k → + λ { (i ∨ j = i0) → merid (m1 (CW↪ C (fst n) a)) (~ k) + ; (i ∨ (~ j) = i0) → merid (m2 (CW↪ C (fst n) a)) (~ k) + ; (i ∨ l = i0) → merid (f a j) (~ k) + ; (i ∧ (~ j) = i1) → merid (m1 (CW↪ C (fst n) a)) (~ k) + ; (i ∧ j = i1) → merid (m3 (CW↪ C (fst n) a)) (~ k) + ; (i ∧ (~ l) = i1) → merid (MMmap-add C D m n m1 m2 m3 f g a j) (~ k) + ; (j = i0) → merid (m1 (CW↪ C (fst n) a)) (~ k) + ; (j ∧ (~ l) = i1) → merid (g a i) (~ k) + ; (l = i0) → merid (doubleCompPath-filler (refl) (f a) (g a) i j) (~ k)}) + south + + bouquetMMmap∈cohom'+ : + (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) (n : Fin m) + (m1 m2 m3 : (x : C .fst (suc (fst n))) → cofibCW (fst n) D) + (f : MMmap C D m n m1 m2) + (g : MMmap C D m n m2 m3) + (a : CWskel-fields.A D (fst n)) + (x : _) + → bouquetMMmap∈cohom' C D m n m1 m3 (MMmap-add C D m n m1 m2 m3 f g) a x + ≡ bouquetMMmap∈cohom' C D m n m1 m2 f a x + +ₖ bouquetMMmap∈cohom' C D m n m2 m3 g a x + bouquetMMmap∈cohom'+ C D m (zero , p) m1 m2 m3 f g a north = refl + bouquetMMmap∈cohom'+ C D m (zero , p) m1 m2 m3 f g a south = refl + bouquetMMmap∈cohom'+ C D m (zero , p) m1 m2 m3 f g a (merid b i) j = + ((sym (PathP→compPathL (help b)) + ∙ sym (lUnit _)) + ∙ ∙≡+₁ (λ i → bouquetMMmap∈cohom' C D m (zero , p) m1 m2 f a (merid b i)) + (λ i → bouquetMMmap∈cohom' C D m (zero , p) m2 m3 g a (merid b i))) j i + where + help : (b : _) + → PathP (λ i → ∣ base ∣ₕ + ≡ cong (bouquetMMmap∈cohom' C D m (zero , p) m2 m3 g a) + (merid b) i) + (cong (bouquetMMmap∈cohom' C D m (zero , p) m1 m2 f a) + (merid b)) + (cong (bouquetMMmap∈cohom' C D m (zero , p) m1 m3 + (MMmap-add C D m (zero , p) m1 m2 m3 f g) a) + (merid b)) + help b i j = + ∣ pickPetal a (Iso.fun (cofibIso zero D) + (realiseAdd-merid C D m (zero , p) m1 m2 m3 f g b i j)) ∣ₕ + bouquetMMmap∈cohom'+ C D m (suc n , p) m1 m2 m3 f g a north = refl + bouquetMMmap∈cohom'+ C D m (suc n , p) m1 m2 m3 f g a south = refl + bouquetMMmap∈cohom'+ C D m (suc n , p) m1 m2 m3 f g a (merid b i) j = + ((sym (PathP→compPathL (help b)) + ∙ sym (lUnit _)) + ∙ ∙≡+₂ n (λ i → bouquetMMmap∈cohom' C D m (suc n , p) m1 m2 f a (merid b i)) + (λ i → bouquetMMmap∈cohom' C D m (suc n , p) m2 m3 g a (merid b i))) j i + where + help : (b : _) + → PathP (λ i → ∣ north ∣ₕ + ≡ cong (bouquetMMmap∈cohom' C D m (suc n , p) m2 m3 g a) + (merid b) i) + (cong (bouquetMMmap∈cohom' C D m (suc n , p) m1 m2 f a) + (merid b)) + (cong (bouquetMMmap∈cohom' C D m (suc n , p) m1 m3 + (MMmap-add C D m (suc n , p) m1 m2 m3 f g) a) + (merid b)) + help b i j = + ∣ pickPetal a (Iso.fun (cofibIso (suc n) D) + (realiseAdd-merid C D m (suc n , p) m1 m2 m3 f g b i j)) ∣ₕ + + bouquetMMmap∈cohom+ : (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) (n : Fin m) + (m1 m2 m3 : (x : C .fst (suc (fst n))) → cofibCW (fst n) D) + (f : MMmap C D m n m1 m2) + (g : MMmap C D m n m2 m3) + (t : CWskel-fields.A C (fst n)) + (a : CWskel-fields.A D (fst n)) + (x : S₊ (suc (fst n))) + → bouquetMMmap∈cohom C D m n m1 m3 + (MMmap-add C D m n m1 m2 m3 f g) a t x + ≡ bouquetMMmap∈cohom C D m n m1 m2 f a t x + +ₖ bouquetMMmap∈cohom C D m n m2 m3 g a t x + bouquetMMmap∈cohom+ C D m n m1 m2 m3 f g t a x = + bouquetMMmap∈cohom'+ C D m n m1 m2 m3 f g a + (Iso.inv (cofibIso (fst n) C) (inr (t , x))) + + module _ (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) (n : Fin m) + (m1 m2 m3 : (x : C .fst (suc (fst n))) → cofibCW (fst n) D) + (f : MMmap C D m n m1 m2) (g : MMmap C D m n m2 m3) where + realiseMMmap-hom : + bouquetDegree (bouquetMMmap C D m n m1 m3 + (MMmap-add C D m n m1 m2 m3 f g)) + ≡ addGroupHom _ _ (bouquetDegree (bouquetMMmap C D m n m1 m2 f)) + (bouquetDegree (bouquetMMmap C D m n m2 m3 g)) + realiseMMmap-hom = + agreeOnℤFinGenerator→≡ λ t → funExt λ a + → sym (isGeneratorℤFinGenerator' + (λ a₁ → degree (suc (fst n)) + λ x → pickPetal a (bouquetMMmap C D m n m1 m3 + (MMmap-add C D m n m1 m2 m3 f g) + (inr (a₁ , x)))) t) + ∙ cong (fst (Hⁿ-Sⁿ≅ℤ (fst n)) .Iso.fun ∘ ∣_∣₂) + (funExt (bouquetMMmap∈cohom+ C D m n m1 m2 m3 f g t a)) + ∙∙ IsGroupHom.pres· (snd (Hⁿ-Sⁿ≅ℤ (fst n))) + (∣ (λ x → ∣ pickPetal a + (bouquetMMmap C D m n m1 m2 f (inr (t , x))) ∣ₕ) ∣₂) + (∣ (λ x → ∣ pickPetal a + (bouquetMMmap C D m n m2 m3 g (inr (t , x))) ∣ₕ) ∣₂) + ∙∙ cong₂ _+_ (isGeneratorℤFinGenerator' + (λ a₁ → degree (suc (fst n)) + λ x → pickPetal a (bouquetMMmap C D m n m1 m2 f + (inr (a₁ , x)))) t) + (isGeneratorℤFinGenerator' + (λ a₁ → degree (suc (fst n)) + λ x → pickPetal a (bouquetMMmap C D m n m2 m3 g + (inr (a₁ , x)))) t) + +-- Now we have all the ingredients, we can get the chain homotopy equation +module chainHomEquation (m : ℕ) (C : CWskel ℓ) (D : CWskel ℓ') + (f g : finCellMap (suc m) C D) (H : finCellHom (suc m) f g) (n : Fin m) where + open SequenceMap + open MMmaps C D (suc m) (fsuc n) + open MMchainHomotopy* (suc m) C D f g H (fsuc n) + open preChainHomotopy (suc m) C D f g H + -- open realiseMMmap m C D f g H + + private + ℤ[AC_] = CWskel-fields.ℤ[A_] C + ℤ[AD_] = CWskel-fields.ℤ[A_] D + + -- The four abelian group maps that are involved in the equation + ∂H H∂ fn+1 gn+1 : AbGroupHom (ℤ[AC (suc (fst n))]) (ℤ[AD (suc (fst n)) ]) + + ∂H = compGroupHom (chainHomotopy (fsuc n)) (∂ D (suc (fst n))) + H∂ = compGroupHom (∂ C (fst n)) (chainHomotopy (injectSuc n)) + fn+1 = prefunctoriality.chainFunct (suc m) f (fsuc n) + gn+1 = prefunctoriality.chainFunct (suc m) g (fsuc n) + + -- Technical lemma regarding suspensions of Iso's + suspIso-suspFun : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} + {C : Type ℓ''} {D : Type ℓ'''} + (e1 : Iso A B) (e2 : Iso C D) (f : C → A) + → Iso.fun (congSuspIso e1) ∘ (suspFun f) ∘ Iso.inv (congSuspIso e2) + ≡ suspFun (Iso.fun e1 ∘ f ∘ Iso.inv e2) + suspIso-suspFun e1 e2 f i north = north + suspIso-suspFun e1 e2 f i south = south + suspIso-suspFun e1 e2 f i (merid a j) = + merid ((Iso.fun e1 ∘ f ∘ Iso.inv e2) a) j + + BouquetIso : ∀ {ℓ} (C : CWskel ℓ) (n : ℕ) + → Iso (cofibCW n C) (SphereBouquet n (Fin (CWskel-fields.card C n))) + BouquetIso C n = + BouquetIso-gen n + (CWskel-fields.card C n) (CWskel-fields.α C n) (CWskel-fields.e C n) + + -- Technical lemma to pull bouquetSusp out of a suspended cofibCW map + cofibIso-suspFun : (n : ℕ) (C : CWskel ℓ) (D : CWskel ℓ') + (f : cofibCW n C → cofibCW n D) + → Iso.fun (cofibIso n D) ∘ (suspFun f) ∘ Iso.inv (cofibIso n C) + ≡ bouquetSusp→ ((Iso.fun (BouquetIso D n)) ∘ f ∘ Iso.inv (BouquetIso C n)) + cofibIso-suspFun n C D f = + cong (λ X → Iso.fun sphereBouquetSuspIso ∘ X ∘ Iso.inv sphereBouquetSuspIso) + (suspIso-suspFun (BouquetIso D n) (BouquetIso C n) f) + + -- connecting MM∂H to ∂H + bouquet∂H : bouquetDegree (bouquetMMmap merid-f merid-g MM∂H) ≡ ∂H + bouquet∂H = + cong (λ X → bouquetDegree ((Iso.fun (cofibIso (suc (fst n)) D)) + ∘ X ∘ (Iso.inv (cofibIso (suc (fst n)) C)))) + (funExt (realiseMMmap.realiseMM∂H C D (suc m) f g H (fsuc n))) + ∙ cong bouquetDegree ιδH≡pre∂∘H + ∙ bouquetDegreeComp (preboundary.pre∂ D (suc (fst n))) + (bouquetHomotopy (fsuc n)) + where + ιδH : SphereBouquet (suc (suc (fst n))) + (Fin (CWskel-fields.card C (suc (fst n)))) + → SphereBouquet (suc (suc (fst n))) + (Fin (CWskel-fields.card D (suc (fst n)))) + ιδH = Iso.fun (cofibIso (suc (fst n)) D) + ∘ suspFun (to_cofibCW (suc (fst n)) D) + ∘ δ (suc (suc (fst n))) D + ∘ Hn+1/Hn (fsuc n) + ∘ Iso.inv (cofibIso (suc (fst n)) C) + + ιδH≡pre∂∘H : ιδH ≡ (preboundary.pre∂ D (suc (fst n))) + ∘ bouquetHomotopy (fsuc n) + ιδH≡pre∂∘H = + cong (λ X → Iso.fun (cofibIso (suc (fst n)) D) + ∘ suspFun (to_cofibCW (suc (fst n)) D) + ∘ δ (suc (suc (fst n))) D ∘ X ∘ Hn+1/Hn (fsuc n) + ∘ Iso.inv (cofibIso (suc (fst n)) C)) + (sym (funExt (Iso.leftInv (BouquetIso D (suc (suc (fst n))))))) + + -- connecting MMΣH∂ to H∂ + bouquetΣH∂ : bouquetDegree (bouquetMMmap merid-tt merid-tt MMΣH∂) ≡ H∂ + bouquetΣH∂ = + cong (λ X → bouquetDegree ((Iso.fun (cofibIso (suc (fst n)) D)) + ∘ X ∘ (Iso.inv (cofibIso (suc (fst n)) C)))) + (funExt (realiseMMΣH∂ C D m f g H n)) + ∙ cong bouquetDegree (cofibIso-suspFun _ C D (Hn+1/Hn (injectSuc n) + ∘ suspFun (to_cofibCW (fst n) C) ∘ δ (suc (fst n)) C)) + ∙ sym (bouquetDegreeSusp Hιδ) + ∙ cong bouquetDegree Hιδ≡H∘pre∂ + ∙ bouquetDegreeComp (bouquetHomotopy (injectSuc n)) + (preboundary.pre∂ C (fst n)) + where + Hιδ : SphereBouquet (suc (fst n)) (Fin (CWskel-fields.card C (suc (fst n)))) + → SphereBouquet (suc (fst n)) (Fin (CWskel-fields.card D (suc (fst n)))) + Hιδ = Iso.fun (BouquetIso D (suc (fst n))) + ∘ (Hn+1/Hn (injectSuc n)) + ∘ suspFun (to_cofibCW (fst n) C) + ∘ δ (suc (fst n)) C ∘ Iso.inv (BouquetIso C (suc (fst n))) + + Hιδ≡H∘pre∂ : Hιδ ≡ bouquetHomotopy (injectSuc n) ∘ (preboundary.pre∂ C (fst n)) + Hιδ≡H∘pre∂ = cong (λ X → Iso.fun (BouquetIso D (suc (fst n))) + ∘ (Hn+1/Hn (injectSuc n)) ∘ X + ∘ suspFun (to_cofibCW (fst n) C) ∘ δ (suc (fst n)) C + ∘ Iso.inv (BouquetIso C (suc (fst n)))) + (sym (funExt (Iso.leftInv (cofibIso (fst n) C)))) + + -- connecting MMΣf to fn+1 + bouquetΣf : bouquetDegree (bouquetMMmap merid-f merid-tt MMΣf) ≡ fn+1 + bouquetΣf = cong (λ X → bouquetDegree ((Iso.fun (cofibIso (suc (fst n)) D)) + ∘ X ∘ (Iso.inv (cofibIso (suc (fst n)) C)))) + (funExt (realiseMMmap.realiseMMΣf C D (suc m) f g H (fsuc n))) + ∙ (cong bouquetDegree (cofibIso-suspFun (suc (fst n)) C D + (prefunctoriality.fn+1/fn (suc m) f (fsuc n)))) + ∙ sym (bouquetDegreeSusp (prefunctoriality.bouquetFunct (suc m) f (fsuc n))) + + -- connecting MMΣg to gn+1 + bouquetΣg : bouquetDegree (bouquetMMmap merid-g merid-tt MMΣg) ≡ gn+1 + bouquetΣg = cong (λ X → bouquetDegree ((Iso.fun (cofibIso (suc (fst n)) D)) + ∘ X ∘ (Iso.inv (cofibIso (suc (fst n)) C)))) + (funExt (realiseMMmap.realiseMMΣg C D (suc m) f g H (fsuc n))) + ∙ (cong bouquetDegree (cofibIso-suspFun (suc (fst n)) C D + (prefunctoriality.fn+1/fn (suc m) g (fsuc n)))) + ∙ sym (bouquetDegreeSusp (prefunctoriality.bouquetFunct (suc m) g (fsuc n))) + + -- Alternative formulation of the chain homotopy equation + chainHomotopy1 : addGroupHom _ _ (addGroupHom _ _ ∂H gn+1) H∂ ≡ fn+1 + chainHomotopy1 = + cong (λ X → addGroupHom _ _ X H∂) aux + ∙ aux2 + ∙ cong (λ X → bouquetDegree (bouquetMMmap merid-f merid-tt X)) + (funExt MMchainHomotopy) + ∙ bouquetΣf + where + module T = MMchainHomotopy* (suc m) C D f g H (fsuc n) + MM∂H+MMΣg = MMmap-add T.merid-f T.merid-g T.merid-tt T.MM∂H T.MMΣg + MM∂H+MMΣg+MMΣH∂ = MMmap-add merid-f merid-tt merid-tt MM∂H+MMΣg MMΣH∂ + + aux : addGroupHom _ _ ∂H gn+1 + ≡ bouquetDegree (bouquetMMmap merid-f merid-tt MM∂H+MMΣg) + aux = cong₂ (λ X Y → addGroupHom _ _ X Y) (sym bouquet∂H) (sym bouquetΣg) + ∙ sym (bouquetAdd.realiseMMmap-hom C D (suc m) (fsuc n) + T.merid-f T.merid-g T.merid-tt T.MM∂H T.MMΣg) + + aux2 : addGroupHom _ _ + (bouquetDegree (bouquetMMmap merid-f merid-tt MM∂H+MMΣg)) H∂ + ≡ bouquetDegree (bouquetMMmap merid-f merid-tt MM∂H+MMΣg+MMΣH∂) + aux2 = cong (addGroupHom _ _ (bouquetDegree + (bouquetMMmap merid-f merid-tt MM∂H+MMΣg))) + (sym bouquetΣH∂) + ∙ sym (bouquetAdd.realiseMMmap-hom C D (suc m) (fsuc n) + T.merid-f T.merid-tt T.merid-tt MM∂H+MMΣg T.MMΣH∂) + + -- Standard formulation of the chain homotopy equation + chainHomotopy2 : subtrGroupHom _ _ fn+1 gn+1 ≡ addGroupHom _ _ ∂H H∂ + chainHomotopy2 = + GroupHom≡ (funExt λ x → aux (fn+1 .fst x) (∂H .fst x) (gn+1 .fst x) + (H∂ .fst x) (cong (λ X → X .fst x) chainHomotopy1)) + where + open AbGroupStr (snd (ℤ[AD (suc (fst n)) ])) + renaming (_+_ to _+G_ ; -_ to -G_ ; +Assoc to +AssocG ; +Comm to +CommG) + aux : ∀ w x y z → (x +G y) +G z ≡ w → w +G (-G y) ≡ x +G z + aux w x y z H = cong (λ X → X +G (-G y)) (sym H) + ∙ sym (+AssocG (x +G y) z (-G y)) + ∙ cong (λ X → (x +G y) +G X) (+CommG z (-G y)) + ∙ +AssocG (x +G y) (-G y) z + ∙ cong (λ X → X +G z) (sym (+AssocG x y (-G y)) + ∙ cong (λ X → x +G X) (+InvR y) + ∙ +IdR x) + +-- Going from a cell homotopy to a chain homotopy +cellHom-to-ChainHomotopy : {C : CWskel ℓ} {D : CWskel ℓ'} (m : ℕ) + {f g : finCellMap (suc m) C D} (H : finCellHom (suc m) f g) + → finChainHomotopy m (finCellMap→finChainComplexMap m f) + (finCellMap→finChainComplexMap m g) +cellHom-to-ChainHomotopy {C = C} {D} m {f} {g} H .finChainHomotopy.fhtpy n = + preChainHomotopy.chainHomotopy (suc m) C D f g H n +cellHom-to-ChainHomotopy {C = C} {D} m {f} {g} H .finChainHomotopy.fbdryhtpy n = + chainHomEquation.chainHomotopy2 m C D f g H n diff --git a/Cubical/Data/CW/Map.agda b/Cubical/Data/CW/Map.agda index 9f3e7c6bb2..5eeeb833f2 100644 --- a/Cubical/Data/CW/Map.agda +++ b/Cubical/Data/CW/Map.agda @@ -2,6 +2,10 @@ {-This file contains: +1. Definition of cellular maps +2. Definition of finite cellular maps +3. The induced map on chain complexes and homology by (finite) cellular maps + -} module Cubical.Data.CW.Map where @@ -14,8 +18,11 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.Transport open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) +open import Cubical.Data.Bool hiding (_≟_) open import Cubical.Data.Nat.Order open import Cubical.Data.Unit open import Cubical.Data.Fin.Inductive.Base @@ -24,26 +31,30 @@ open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sequence open import Cubical.Data.FinSequence - +open import Cubical.Data.CW.Base +open import Cubical.Data.CW.Properties +open import Cubical.Data.CW.ChainComplex open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Degree open import Cubical.HITs.Pushout open import Cubical.HITs.Susp open import Cubical.HITs.SequentialColimit open import Cubical.HITs.SphereBouquet open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.Data.CW.Base -open import Cubical.Data.CW.Properties - +open import Cubical.HITs.SphereBouquet.Degree +open import Cubical.HITs.SequentialColimit +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.ChainComplex open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup -open import Cubical.HITs.SequentialColimit -open Sequence - open import Cubical.Relation.Nullary +open Sequence + private variable ℓ ℓ' ℓ'' : Level @@ -80,3 +91,354 @@ module _ (m : ℕ) where finCellMap↓ : {m : ℕ} {C : CWskel ℓ} {D : CWskel ℓ'} → finCellMap (suc m) C D → finCellMap m C D FinSequenceMap.fmap (finCellMap↓ {m = m} ϕ) x = FinSequenceMap.fmap ϕ (injectSuc x) FinSequenceMap.fcomm (finCellMap↓ {m = suc m} {C = C} ϕ) x r = FinSequenceMap.fcomm ϕ (injectSuc x) r + +-- A cellular map between two CW complexes + +-- A cellMap from C to D is a family of maps Cₙ → Dₙ that commute with +-- the inclusions Xₙ ↪ Xₙ₊₁ + +-- From a cellMap to a family of maps between free abelian groups +module prefunctoriality (m : ℕ) (f : finCellMap m C D) (n' : Fin m) where + open FinSequenceMap + open CWskel-fields + + n = fst n' + + fn+1/fn : cofibCW (fst n') C → cofibCW (fst n') D + fn+1/fn (inl tt) = inl tt + fn+1/fn (inr x) = inr (f .fmap (fsuc n') x) + fn+1/fn (push x i) = + (push (f .fmap (injectSuc n') x) ∙ (cong inr (f .fcomm n' x))) i + + bouquetFunct : SphereBouquet n (Fin (card C n)) + → SphereBouquet n (Fin (card D n)) + bouquetFunct = Iso.fun (BouquetIso-gen n (card D n) (α D n) (e D n)) + ∘ fn+1/fn + ∘ Iso.inv (BouquetIso-gen n (card C n) (α C n) (e C n)) + + chainFunct : AbGroupHom (ℤ[A C ] n) (ℤ[A D ] n) + chainFunct = bouquetDegree bouquetFunct + +module _ (m : ℕ) (C : CWskel ℓ) (n' : Fin m) where + open prefunctoriality m (idFinCellMap m C) n' + open SequenceMap + open CWskel-fields + + fn+1/fn-id : fn+1/fn ≡ idfun _ + fn+1/fn-id = funExt + λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j → rUnit (push a) (~ j) i} + + bouquetFunct-id : bouquetFunct ≡ idfun _ + bouquetFunct-id = + cong (λ f → Iso.fun (BouquetIso-gen n (card C n) (α C n) (e C n)) + ∘ f + ∘ Iso.inv (BouquetIso-gen n (card C n) (α C n) (e C n))) + fn+1/fn-id + ∙ funExt (Iso.rightInv (BouquetIso-gen n (card C n) (α C n) (e C n))) + + chainFunct-id : chainFunct ≡ idGroupHom + chainFunct-id = cong bouquetDegree bouquetFunct-id ∙ bouquetDegreeId + +module _ (m : ℕ) (g : finCellMap m D E) (f : finCellMap m C D) (n' : Fin m) where + module pf1 = prefunctoriality m f n' + module pf2 = prefunctoriality m g n' + module pf3 = prefunctoriality m (composeFinCellMap m g f) n' + open FinSequenceMap + open CWskel-fields + private + n = fst n' + + fn+1/fn-comp : pf2.fn+1/fn ∘ pf1.fn+1/fn ≡ pf3.fn+1/fn + fn+1/fn-comp = funExt + λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j → help a j i} + where + help : (a : fst C n) + → cong (pf2.fn+1/fn ∘ pf1.fn+1/fn) (push a) ≡ cong pf3.fn+1/fn (push a) + help a = cong-∙ pf2.fn+1/fn (push (f .fmap (injectSuc n') a)) + (λ i₁ → inr (f .fcomm n' a i₁)) + ∙∙ sym (assoc _ _ _) + ∙∙ sym (cong₂ _∙_ refl + (cong-∙ inr (g .fcomm n' (fmap f (injectSuc n') a)) + (cong (g .fmap (fsuc n')) (f .fcomm n' a)))) + + bouquetFunct-comp : pf2.bouquetFunct ∘ pf1.bouquetFunct ≡ pf3.bouquetFunct + bouquetFunct-comp = funExt λ x + → cong (Iso.fun (BouquetIso-gen n (card E n) (α E n) (e E n))) + (cong pf2.fn+1/fn + (Iso.leftInv (BouquetIso-gen n (card D n) (α D n) (e D n)) _) + ∙ funExt⁻ fn+1/fn-comp + (Iso.inv (BouquetIso-gen n (card C n) (α C n) (e C n)) x)) + + chainFunct-comp : compGroupHom pf1.chainFunct pf2.chainFunct ≡ pf3.chainFunct + chainFunct-comp = + sym (bouquetDegreeComp∙ (pf2.bouquetFunct , refl) + (pf1.bouquetFunct , refl)) + ∙ cong bouquetDegree bouquetFunct-comp + +-- Now we prove the commutativity condition to get a fully fledged chain map +module functoriality (m : ℕ) (f : finCellMap (suc m) C D) where + open CWskel-fields + open SequenceMap + module pf* = prefunctoriality m (finCellMap↓ f) + open prefunctoriality (suc m) f + open FinSequenceMap + + -- δ ∘ fn+1/fn ≡ f ∘ δ + commδ : (n : Fin (suc m)) (x : cofibCW (fst n) C) + → δ (fst n) D (fn+1/fn n x) + ≡ suspFun (f .fmap (injectSuc n)) (δ (fst n) C x) + commδ n (inl x) = refl + commδ n (inr x) = refl + commδ n (push a i) j = + hcomp (λ k → λ { (i = i0) → north + ; (i = i1) → south + ; (j = i0) → δ (fst n) D (compPath-filler + (push (f .fmap (injectSuc n) a)) + (cong inr (f .fcomm n a)) k i) + ; (j = i1) → merid (f .fmap (injectSuc n) a) i }) + (merid (f .fmap (injectSuc n) a) i) + + -- Σto_cofibCW ∘ Σf ≡ Σfn+1/fn ∘ Σto_cofibCW + commToCofibCWSusp : (n : Fin (suc m)) (x : Susp (fst C (suc (fst n)))) + → suspFun (to_cofibCW (fst n) D) (suspFun (f .fmap (fsuc n)) x) + ≡ suspFun (fn+1/fn n) (suspFun (to_cofibCW (fst n) C) x) + commToCofibCWSusp n north = refl + commToCofibCWSusp n south = refl + commToCofibCWSusp n (merid a i) = refl + + -- commδ and commToCofibCWSusp give us the chain map equation at the level of cofibCWs + -- now we massage isomorphisms and suspensions to get the proper equation between SphereBouquets + funct∘pre∂ : (n : Fin (suc m)) + → SphereBouquet (suc (fst n)) (Fin (card C (suc (fst n)))) + → SphereBouquet (suc (fst n)) (Fin (card D (fst n))) + funct∘pre∂ n = (bouquetSusp→ (bouquetFunct n)) ∘ (preboundary.pre∂ C (fst n)) + + pre∂∘funct : (n : Fin m) + → (SphereBouquet (suc (fst n)) (Fin (card C (suc (fst n))))) + → SphereBouquet (suc (fst n)) (Fin (card D (fst n))) + pre∂∘funct n = preboundary.pre∂ D (fst n) ∘ bouquetFunct (fsuc n) + + commPre∂Funct : (n : Fin m) → funct∘pre∂ (injectSuc n) ≡ pre∂∘funct n + commPre∂Funct n = funExt λ x → cong (fun (iso1 D (fst n))) (main x) + where + open preboundary + open Iso + + bouquet : (C : CWskel ℓ) (n m : ℕ) → Type + bouquet = λ C n m → SphereBouquet n (Fin (snd C .fst m)) + + iso1 : (C : CWskel ℓ) (n : ℕ) + → Iso (Susp (bouquet C n n)) (bouquet C (suc n) n) + iso1 C n = sphereBouquetSuspIso + + iso2 : (C : CWskel ℓ) (n : ℕ) → Iso (cofibCW n C) (bouquet C n n) + iso2 C n = + BouquetIso-gen n (snd C .fst n) (snd C .snd .fst n) + (snd C .snd .snd .snd n) + + step2aux : ∀ x → suspFun (bouquetFunct (injectSuc n)) x + ≡ suspFun (fun (iso2 D (fst n))) + (suspFun (fn+1/fn (injectSuc n)) + (suspFun (inv (iso2 C (fst n))) x)) + step2aux north = refl + step2aux south = refl + step2aux (merid a i) = refl + + step3aux : ∀ x + → suspFun (inv (iso2 C (fst n))) (suspFun (fun (iso2 C (fst n))) x) ≡ x + step3aux north = refl + step3aux south = refl + step3aux (merid a i) j = merid (leftInv (iso2 C (fst n)) a j) i + + module _ (x : bouquet C (suc (fst n)) (suc (fst n))) where + step1 = cong (suspFun (bouquetFunct (injectSuc n))) + (leftInv (iso1 C (fst n)) + (((suspFun (fun (iso2 C (fst n)))) + ∘ (suspFun (to_cofibCW (fst n) C)) + ∘ (δ (suc (fst n)) C) ∘ (inv (iso2 C (suc (fst n))))) x)) + + step2 = step2aux (((suspFun (fun (iso2 C (fst n)))) + ∘ (suspFun (to_cofibCW (fst n) C)) + ∘ (δ (suc (fst n)) C) ∘ (inv (iso2 C (suc (fst n))))) x) + + step3 = + cong ((suspFun (fun (iso2 D (fst n)))) + ∘ (suspFun (fn+1/fn (injectSuc n)))) + (step3aux (((suspFun (to_cofibCW (fst n) C)) + ∘ (δ (suc (fst n)) C) + ∘ (inv (iso2 C (suc (fst n))))) x)) + + step4 = cong (suspFun (fun (iso2 D (fst n)))) + (sym (commToCofibCWSusp (injectSuc n) + (((δ (suc (fst n)) C) ∘ (inv (iso2 C (suc (fst n))))) x))) + + step5 = λ i → suspFun (fun (iso2 D (fst n))) + (suspFun (to fst (injectSuc n) cofibCW D) + (suspFun (f .fmap (p i)) + (δ (suc (fst n)) C (inv (iso2 C (suc (fst n))) x)))) + where + p : fsuc (injectSuc n) ≡ injectSuc (fsuc n) + p = Σ≡Prop (λ _ → isProp<ᵗ) refl + + step6 = cong ((suspFun (fun (iso2 D (fst n)))) + ∘ (suspFun (to_cofibCW (fst n) D))) + (sym (commδ (fsuc n) (inv (iso2 C (suc (fst n))) x))) + + step7 = cong ((suspFun (fun (iso2 D (fst n)))) + ∘ (suspFun (to_cofibCW (fst n) D)) + ∘ (δ (suc (fst n)) D)) + (sym (leftInv (iso2 D (suc (fst n))) + (((fn+1/fn (fsuc n)) ∘ (inv (iso2 C (suc (fst n))))) x))) + + main = step1 ∙ step2 ∙ step3 ∙ step4 ∙ step5 ∙ step6 ∙ step7 + + -- finally, we take bouquetDegree to get the equation at the level + -- of abelian groups + comm∂Funct : (n : Fin m) + → compGroupHom (chainFunct (fsuc n)) (∂ D (fst n)) + ≡ compGroupHom (∂ C (fst n)) (chainFunct (injectSuc n)) + comm∂Funct n = (sym (degree-pre∂∘funct n)) + ∙∙ cong bouquetDegree (sym (commPre∂Funct n)) + ∙∙ (degree-funct∘pre∂ n) + where + degree-funct∘pre∂ : (n : Fin m) + → bouquetDegree (funct∘pre∂ (injectSuc n)) + ≡ compGroupHom (∂ C (fst n)) (chainFunct (injectSuc n)) + degree-funct∘pre∂ n = + bouquetDegreeComp (bouquetSusp→ (bouquetFunct (injectSuc n))) + (preboundary.pre∂ C (fst n)) + ∙ cong (compGroupHom (∂ C (fst n))) + (sym (bouquetDegreeSusp (bouquetFunct (injectSuc n)))) + + degree-pre∂∘funct : (n : Fin m) + → bouquetDegree (pre∂∘funct n) + ≡ compGroupHom (chainFunct (fsuc n)) (∂ D (fst n)) + degree-pre∂∘funct n = + bouquetDegreeComp (preboundary.pre∂ D (fst n)) (bouquetFunct (fsuc n)) + +open finChainComplexMap +-- Main statement of functoriality +-- From a cellMap, we can get a ChainComplexMap +finCellMap→finChainComplexMap : (m : ℕ) (f : finCellMap (suc m) C D) + → finChainComplexMap m (CW-ChainComplex C) (CW-ChainComplex D) +fchainmap (finCellMap→finChainComplexMap m f) n = + prefunctoriality.chainFunct (suc m) f n +fbdrycomm (finCellMap→finChainComplexMap m f) n = functoriality.comm∂Funct m f n + +finCellMap→finChainComplexMapId : (m : ℕ) + → finCellMap→finChainComplexMap m (idFinCellMap (suc m) C) + ≡ idFinChainMap m (CW-ChainComplex C) +finCellMap→finChainComplexMapId m = finChainComplexMap≡ + λ x → cong bouquetDegree (bouquetFunct-id _ _ x) ∙ bouquetDegreeId + +finCellMap→finChainComplexMapComp : (m : ℕ) + (g : finCellMap (suc m) D E) (f : finCellMap (suc m) C D) + → finCellMap→finChainComplexMap m (composeFinCellMap _ g f) + ≡ compFinChainMap (finCellMap→finChainComplexMap m f) + (finCellMap→finChainComplexMap m g) +finCellMap→finChainComplexMapComp m g f = + finChainComplexMap≡ λ x + → cong bouquetDegree (sym (bouquetFunct-comp _ g f x)) + ∙ bouquetDegreeComp _ _ + +-- And thus a map of homology +finCellMap→HomologyMap : (m : ℕ) (f : finCellMap (suc (suc (suc m))) C D) + → GroupHom (Hˢᵏᵉˡ C m) (Hˢᵏᵉˡ D m) +finCellMap→HomologyMap {C = C} {D = D} m f = + finChainComplexMap→HomologyMap (suc m) + (finCellMap→finChainComplexMap _ f) flast + +finCellMap→HomologyMapId : (m : ℕ) + → finCellMap→HomologyMap m (idFinCellMap (suc (suc (suc m))) C) + ≡ idGroupHom +finCellMap→HomologyMapId m = + cong (λ r → finChainComplexMap→HomologyMap (suc m) r flast) + (finCellMap→finChainComplexMapId _) + ∙ finChainComplexMap→HomologyMapId _ + +finCellMap→HomologyMapComp : (m : ℕ) + (g : finCellMap (suc (suc (suc m))) D E) + (f : finCellMap (suc (suc (suc m))) C D) + → finCellMap→HomologyMap m (composeFinCellMap _ g f) + ≡ compGroupHom (finCellMap→HomologyMap m f) + (finCellMap→HomologyMap m g) +finCellMap→HomologyMapComp m g f = + (cong (λ r → finChainComplexMap→HomologyMap (suc m) r flast) + (finCellMap→finChainComplexMapComp _ _ _)) + ∙ finChainComplexMap→HomologyMapComp _ _ _ + +-- sanity check: chainFunct of a cellular map fₙ : Cₙ → Dₙ +-- is just functoriality of ℤ[-] when n = 1. +module _ (m : ℕ) (f : finCellMap (suc (suc (suc m))) C D) where + open CWskel-fields + open FinSequenceMap + open prefunctoriality _ f + + cellMap↾₁ : Fin (card C 0) → Fin (card D 0) + cellMap↾₁ = fst (CW₁-discrete D) ∘ fmap f (1 , tt) ∘ invEq (CW₁-discrete C) + + chainFunct' : AbGroupHom (ℤ[A C ] 0) (ℤ[A D ] 0) + chainFunct' = ℤFinFunct cellMap↾₁ + + chainFunct₀ : chainFunct' ≡ chainFunct fzero + chainFunct₀ = + agreeOnℤFinGenerator→≡ λ t → funExt λ x + → sumFin-choose _+_ 0 (λ _ → refl) +Comm + (λ a → ℤFinFunctGenerator cellMap↾₁ (ℤFinGenerator t) a x) + (S⁰×S⁰→ℤ true (pickPetal x (bouquetFunct fzero (inr (t , false))))) + t (ℤFinFunctGenerator≡ cellMap↾₁ t x ∙ main₁ t x) + (main₂ cellMap↾₁ x t) + ∙ isGeneratorℤFinGenerator' + (λ a → degree 0 λ s + → pickPetal x (bouquetFunct fzero (inr (a , s)))) t + where + F = Pushout→Bouquet 0 (card D 0) (α D 0) (e D 0) + + + lem₂ : {k : ℕ} (t : Fin k) (x : Fin k) + → ℤFinGenerator x t ≡ S⁰×S⁰→ℤ true (pickPetal x (inr (t , false))) + lem₂ {k = suc k} t x with (fst x ≟ᵗ fst t) + ... | lt x₁ = refl + ... | eq x₁ = refl + ... | gt x₁ = refl + + main₁ : (t : _) (x : _) + → ℤFinGenerator (cellMap↾₁ t) x + ≡ S⁰×S⁰→ℤ true + (pickPetal x + (F (fst (e D 0) (f .fmap (1 , tt) (invEq (CW₁-discrete C) t))))) + main₁ t x = (ℤFinGeneratorComm (cellMap↾₁ t) x + ∙ lem₂ (cellMap↾₁ t) x) + ∙ cong (S⁰×S⁰→ℤ true ∘ pickPetal x ∘ F) + (lem₁ _) + where + lem₀ : (x : Pushout (α D 0) fst) + → inr (CW₁-discrete D .fst (invEq (e D 0) x)) ≡ x + lem₀ (inl x) = ⊥.rec (CW₀-empty D x) + lem₀ (inr x) j = inr (secEq (CW₁-discrete D) x j) + + lem₁ : (x : _) + → inr (CW₁-discrete D .fst x) ≡ fst (e D 0) x + lem₁ x = (λ i → inr (CW₁-discrete D .fst + (retEq (e D 0) x (~ i)))) + ∙ lem₀ (fst (e D 0) x) + + main₂ : (f' : _) (x : _) (t : _) (x' : Fin (card C zero)) + → ¬ x' ≡ t + → ℤFinFunctGenerator {n = card C zero} {m = card D zero} + f' (ℤFinGenerator t) x' x + ≡ pos 0 + main₂ f' x t x' p with (f' x' .fst ≟ᵗ x .fst) | (fst t ≟ᵗ fst x') + ... | lt x₁ | r = refl + ... | eq x₂ | r = lem + where + lem : _ + lem with (fst t ≟ᵗ fst x') + ... | lt x = refl + ... | eq x = ⊥.rec (p (Σ≡Prop (λ _ → isProp<ᵗ) (sym x))) + ... | gt x = refl + ... | gt x₁ | r = refl diff --git a/Cubical/HITs/PropositionalTruncation/Properties.agda b/Cubical/HITs/PropositionalTruncation/Properties.agda index 92f6ea0b08..cb6e21f011 100644 --- a/Cubical/HITs/PropositionalTruncation/Properties.agda +++ b/Cubical/HITs/PropositionalTruncation/Properties.agda @@ -263,20 +263,26 @@ module SetElim (Bset : isSet B) where open SetElim public using (rec→Set; trunc→Set≃) -elim→Set - : {P : ∥ A ∥₁ → Type ℓ} +elim→Set : ∀ {ℓ'} {A : Type ℓ'} {P : ∥ A ∥₁ → Type ℓ} → (∀ t → isSet (P t)) → (f : (x : A) → P ∣ x ∣₁) → (kf : ∀ x y → PathP (λ i → P (squash₁ ∣ x ∣₁ ∣ y ∣₁ i)) (f x) (f y)) → (t : ∥ A ∥₁) → P t -elim→Set {A = A} {P = P} Pset f kf t - = rec→Set (Pset t) g gk t +elim→Set {A = A} {P = P} Pset f kf t = main t .fst .fst where - g : A → P t - g x = transp (λ i → P (squash₁ ∣ x ∣₁ t i)) i0 (f x) - - gk : 2-Constant g - gk x y i = transp (λ j → P (squash₁ (squash₁ ∣ x ∣₁ ∣ y ∣₁ i) t j)) i0 (kf x y i) + main : (t : ∥ A ∥₁) + → isContr (Σ[ x ∈ P t ] + ((a : A) → PathP (λ i → P (squash₁ t ∣ a ∣₁ i)) x (f a))) + main = elim (λ _ → isPropIsContr) + λ a → + (((f a) , kf a) + , λ {(x , p) → Σ≡Prop (λ _ → isPropΠ + λ _ → isOfHLevelPathP' 1 (Pset _) _ _) + (sym (transport (λ j → PathP (λ i → P (sq a j i)) x (f a)) (p a))) + }) + where + sq : (a : A) → squash₁ ∣ a ∣₁ ∣ a ∣₁ ≡ refl + sq a = isProp→isSet squash₁ _ _ _ _ elim2→Set : {P : ∥ A ∥₁ → ∥ B ∥₁ → Type ℓ} @@ -297,17 +303,6 @@ elim2→Set {A = A} {B = B} {P = P} Pset f kf₁ kf₂ sf = → PathP (λ i → (u : ∥ B ∥₁) → P (squash₁ ∣ x ∣₁ ∣ y ∣₁ i) u) (mapHelper x) (mapHelper y) squareHelper x y i = elim→Set (λ _ → Pset _ _) (λ v → kf₁ x y v i) λ v w → sf x y v w i -elim→Setβ : ∀ {ℓ ℓ'} {A : Type ℓ} - {P : ∥ A ∥₁ → Type ℓ'} - → (is-set : ∀ t → isSet (P t)) - → (f : (x : A) → P ∣ x ∣₁) - → (kf : ∀ x y → PathP (λ i → P (squash₁ ∣ x ∣₁ ∣ y ∣₁ i)) (f x) (f y)) - → (t : A) → elim→Set is-set f kf ∣ t ∣₁ ≡ f t -elim→Setβ {A = A} {P = P} Pset f kf t = - cong (λ p → subst P p (f t)) - (isProp→isSet squash₁ _ _ (squash₁ ∣ t ∣₁ ∣ t ∣₁) refl) - ∙ transportRefl (f t) - RecHProp : (P : A → hProp ℓ) (kP : ∀ x y → P x ≡ P y) → ∥ A ∥₁ → hProp ℓ RecHProp P kP = rec→Set isSetHProp P kP From a3bb47e1a71da858c66ee029b4a5dd53a9bfd620 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 3 Mar 2024 20:48:17 +0100 Subject: [PATCH 29/73] done? --- Cubical/Algebra/Group/GroupPath.agda | 67 ++ Cubical/Data/CW/Approximation.agda | 689 +++++++++++++++++++++ Cubical/Data/CW/Base.agda | 6 +- Cubical/Data/CW/Homology.agda | 243 ++++++++ Cubical/Data/CW/Map.agda | 11 +- Cubical/Data/CW/Properties.agda | 75 ++- Cubical/Data/CW/Subcomplex.agda | 195 ++++++ Cubical/Data/Fin/Inductive/Base.agda | 10 +- Cubical/Data/Fin/Inductive/Properties.agda | 5 + Cubical/Data/Subcomplex.agda | 0 10 files changed, 1265 insertions(+), 36 deletions(-) create mode 100644 Cubical/Data/CW/Approximation.agda create mode 100644 Cubical/Data/CW/Homology.agda create mode 100644 Cubical/Data/CW/Subcomplex.agda create mode 100644 Cubical/Data/Subcomplex.agda diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index b4b6b90793..f616401a9d 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -198,6 +198,11 @@ GroupEquivJ {G = G} P p {H} e = (transportRefl (fst (fst e) (transportRefl x i)) i)))) ∙ retEq (fst e) x)) +GroupEquivJ>_ : {ℓ : Level} {ℓ' : Level} {G : Group ℓ} + {P : (H : Group ℓ) → GroupEquiv G H → Type ℓ'} → + P G idGroupEquiv → (H : Group ℓ) (e : GroupEquiv G H) → P H e +GroupEquivJ>_ {G = G} {P} ids H = GroupEquivJ (λ H e → P H e) ids + isGroupoidGroup : ∀ {ℓ} → isGroupoid (Group ℓ) isGroupoidGroup G H = isOfHLevelRespectEquiv 2 (GroupPath _ _) @@ -235,3 +240,65 @@ module _ {ℓ ℓ'} {A : Type ℓ} (sym (lUnit _) ∙∙ sym (uaCompGroupEquiv (G-coh x y) (G-coh y z)) ∙∙ cong uaGroup (PropTrunc→Group-coh-coh x y z)) + +-- action of of uaGroup on GroupHom +module _ {ℓ ℓ' : Level} {G1 : Group ℓ} {H1 : Group ℓ'} where + private + pre-PathPGroupHom : ∀ + (G2 : Group ℓ) + (eG : GroupEquiv G1 G2) + (H2 : Group ℓ') (eH : GroupEquiv H1 H2) + (ϕ : GroupHom G1 H1) (ψ : GroupHom G2 H2) + → compGroupHom (GroupEquiv→GroupHom eG) ψ + ≡ compGroupHom ϕ (GroupEquiv→GroupHom eH) + → PathP (λ i → GroupHom (uaGroup eG i) (uaGroup eH i)) + ϕ ψ + pre-PathPGroupHom = + GroupEquivJ> (GroupEquivJ> + λ ϕ ψ → λ s + → toPathP ((λ s + → transport (λ i → GroupHom (uaGroupId G1 s i) (uaGroupId H1 s i)) ϕ) + ∙ transportRefl ϕ + ∙ Σ≡Prop (λ _ → isPropIsGroupHom _ _) (sym (cong fst s)))) + + PathPGroupHom : {G2 : Group ℓ} (eG : GroupEquiv G1 G2) + {H2 : Group ℓ'} (eH : GroupEquiv H1 H2) + {ϕ : GroupHom G1 H1} {ψ : GroupHom G2 H2} + → compGroupHom (GroupEquiv→GroupHom eG) ψ + ≡ compGroupHom ϕ (GroupEquiv→GroupHom eH) + → PathP (λ i → GroupHom (uaGroup eG i) (uaGroup eH i)) ϕ ψ + PathPGroupHom eG eH p = pre-PathPGroupHom _ eG _ eH _ _ p + + module _ {H2 : Group ℓ'} (eH : GroupEquiv H1 H2) + {ϕ : GroupHom G1 H1} {ψ : GroupHom G1 H2} where + PathPGroupHomₗ : ψ ≡ compGroupHom ϕ (GroupEquiv→GroupHom eH) + → PathP (λ i → GroupHom G1 (uaGroup eH i)) ϕ ψ + PathPGroupHomₗ p = + transport (λ k → PathP (λ i → GroupHom (uaGroupId G1 k i) (uaGroup eH i)) ϕ ψ) + (PathPGroupHom idGroupEquiv eH + (Σ≡Prop (λ _ → isPropIsGroupHom _ _) (cong fst p))) + + PathPGroupHomₗ' : compGroupHom ψ (GroupEquiv→GroupHom (invGroupEquiv eH)) ≡ ϕ + → PathP (λ i → GroupHom G1 (uaGroup eH i)) ϕ ψ + PathPGroupHomₗ' p = + PathPGroupHomₗ + (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (λ s → sym (secEq (fst eH) (fst ψ s)))) + ∙ cong (λ ϕ → compGroupHom ϕ (GroupEquiv→GroupHom eH)) p) + + module _ {G2 : Group ℓ} (eG : GroupEquiv G1 G2) + {ϕ : GroupHom G1 H1} {ψ : GroupHom G2 H1} + where + PathPGroupHomᵣ : compGroupHom (GroupEquiv→GroupHom eG) ψ ≡ ϕ + → PathP (λ i → GroupHom (uaGroup eG i) H1) ϕ ψ + PathPGroupHomᵣ p = + transport (λ k → PathP (λ i → GroupHom (uaGroup eG i) (uaGroupId H1 k i)) ϕ ψ) + (PathPGroupHom eG idGroupEquiv + (Σ≡Prop (λ _ → isPropIsGroupHom _ _) (cong fst p))) + + PathPGroupHomᵣ' : ψ ≡ compGroupHom (GroupEquiv→GroupHom (invGroupEquiv eG)) ϕ + → PathP (λ i → GroupHom (uaGroup eG i) H1) ϕ ψ + PathPGroupHomᵣ' p = PathPGroupHomᵣ + (cong (compGroupHom (GroupEquiv→GroupHom eG)) p + ∙ Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x → cong (fst ϕ) (retEq (fst eG) x))) diff --git a/Cubical/Data/CW/Approximation.agda b/Cubical/Data/CW/Approximation.agda new file mode 100644 index 0000000000..f5d9d36b4a --- /dev/null +++ b/Cubical/Data/CW/Approximation.agda @@ -0,0 +1,689 @@ +{-# OPTIONS --lossy-unification --safe #-} +{- Cellular approximation theorems for +-- cellular maps and homotopies +-} + +module Cubical.Data.CW.Approximation where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.CW +open import Cubical.Data.CW.Map +open import Cubical.Data.CW.Homotopy +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence + +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.PropositionalTruncation as PT hiding (elimFin) +open import Cubical.HITs.Truncation as TR +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout + +open import Cubical.Axiom.Choice + +open import Cubical.Homotopy.Connected + +open Sequence +open FinSequenceMap + +private + variable + ℓ ℓ' ℓ'' : Level + +---- Part 1. Definitions ----- + +-- finite cellular approximations +finCellApprox : (C : CWskel ℓ) (D : CWskel ℓ') + (f : realise C → realise D) (m : ℕ) → Type (ℓ-max ℓ ℓ') +finCellApprox C D f m = + Σ[ ϕ ∈ finCellMap m C D ] + (FinSeqColim→Colim m {X = realiseSeq D} + ∘ finCellMap→FinSeqColim C D ϕ + ≡ f ∘ FinSeqColim→Colim m {X = realiseSeq C}) + +idFinCellApprox : (m : ℕ) + {C : CWskel ℓ} → finCellApprox C C (idfun _) m +fst (idFinCellApprox m {C}) = idFinCellMap m C +snd (idFinCellApprox m {C}) = + →FinSeqColimHomotopy _ _ λ x → refl + +compFinCellApprox : (m : ℕ) + {C : CWskel ℓ} {D : CWskel ℓ'} {E : CWskel ℓ''} + {g : realise D → realise E} + {f : realise C → realise D} + → finCellApprox D E g m → finCellApprox C D f m + → finCellApprox C E (g ∘ f) m +fst (compFinCellApprox m {g = g} {f} (F , p) (G , q)) = composeFinCellMap m F G +snd (compFinCellApprox m {C = C} {g = g} {f} (F , p) (G , q)) = + →FinSeqColimHomotopy _ _ λ x + → funExt⁻ p _ + ∙ cong g (funExt⁻ q (fincl _ x)) + +-- a finite cellular homotopies relative +finCellHomRel : {C : CWskel ℓ} {D : CWskel ℓ'} (m : ℕ) + (f g : finCellMap m C D) + (h∞ : (n : Fin (suc m)) (c : fst C (fst n)) + → Path (realise D) (incl (f .fmap n c)) (incl (g .fmap n c))) + → Type (ℓ-max ℓ ℓ') +finCellHomRel {C = C} {D = D} m f g h∞ = + Σ[ ϕ ∈ finCellHom m f g ] + ((n : Fin (suc m)) (x : fst C (fst n)) + → Square (h∞ n x) + (cong incl (finCellHom.fhom ϕ n x)) + (push (f .fmap n x)) (push (g .fmap n x))) + +---- Part 2. The cellular (n)-approxiation theorem: ----- +-- Given CW skeleta Cₙ and Dₙ with a map f : C∞ → D∞ between their +-- realisations, there exists a finite cellular map fₙ : Cₙ → Dₙ s.t. +-- (n < m) s.t. incl ∘ fₙ = f ∘ incl + +-- Construction +module _ (C : CWskel ℓ) (D : CWskel ℓ') (f : realise C → realise D) where + find-connected-component : (d : realise D) → ∃[ d0 ∈ fst D 1 ] incl d0 ≡ d + find-connected-component = CW→Prop D (λ _ → squash₁) λ a → ∣ a , refl ∣₁ + + find-connected-component-C₀ : (c : fst C 1) + → ∃[ d0 ∈ fst D 1 ] incl d0 ≡ f (incl c) + find-connected-component-C₀ c = find-connected-component (f (incl c)) + + -- existence of f₁ : C₁ → D₁ + approx₁ : ∃[ f₁ ∈ (fst C 1 → fst D 1) ] ((c : _) → incl (f₁ c) ≡ f (incl c)) + approx₁ = + invEq (_ , satAC∃Fin-C0 C (λ _ → fst D 1) (λ c d0 → incl d0 ≡ f (incl c))) + find-connected-component-C₀ + + approx : (m : ℕ) + → ∃[ fₙ ∈ ((n : Fin (suc m)) → Σ[ h ∈ (fst C (fst n) → fst D (fst n)) ] + ((c : _) → incl (h c) ≡ f (incl c))) ] + ((n : Fin m) (c : fst C (fst n)) + → fₙ (fsuc n) .fst (CW↪ C (fst n) c) + ≡ CW↪ D (fst n) (fₙ (injectSuc n) .fst c)) + approx zero = ∣ (λ { (zero , p) + → (λ x → ⊥.rec (CW₀-empty C x)) + , λ x → ⊥.rec (CW₀-empty C x)}) + , (λ {()}) ∣₁ + approx (suc zero) = + PT.map (λ a1 → + (λ { (zero , p) → (λ x → ⊥.rec (CW₀-empty C x)) + , λ x → ⊥.rec (CW₀-empty C x) + ; (suc zero , p) → a1}) + , λ {(zero , p) c → ⊥.rec (CW₀-empty C c)}) + approx₁ + approx (suc (suc m)) = + PT.rec + squash₁ + (λ {(p , f') + → PT.rec squash₁ (λ r + → PT.map (λ ind → mainₗ p f' r ind + , mainᵣ p f' r ind) + (mere-fib-f-coh (p flast .fst) + r (p (suc m , <ᵗsucm) .snd))) + (fst-lem (p flast .fst) + (p flast .snd))}) + (approx (suc m)) + where + open CWskel-fields C + fst-lem : (fm : fst C (suc m) → fst D (suc m)) + → ((c : fst C (suc m)) → incl (fm c) ≡ f (incl c)) + → ∥ ((x : A (suc m)) (y : S₊ m) + → CW↪ D (suc m) (fm (α (suc m) (x , y))) + ≡ CW↪ D (suc m) (fm (α (suc m) (x , ptSn m)))) ∥₁ + fst-lem fm fh = + invEq propTrunc≃Trunc1 + (invEq (_ , InductiveFinSatAC 1 (CWskel-fields.card C (suc m)) _) + λ a → fst propTrunc≃Trunc1 + (sphereToTrunc m λ y → + TR.map fst (isConnectedCong _ _ (isConnected-CW↪∞ (suc (suc m)) D) + (sym (push _) + ∙ (fh (CWskel-fields.α C (suc m) (a , y)) + ∙ cong f (push _ + ∙ cong incl (cong (invEq (CWskel-fields.e C (suc m))) + (push (a , y) ∙ sym (push (a , ptSn m)))) + ∙ sym (push _)) + ∙ sym (fh (CWskel-fields.α C (suc m) (a , ptSn m)))) + ∙ push _) .fst))) + + module _ (fm : fst C (suc m) → fst D (suc m)) + (fm-coh : (x : A (suc m)) (y : S₊ m) → + CW↪ D (suc m) (fm (α (suc m) (x , y))) + ≡ CW↪ D (suc m) (fm (α (suc m) (x , ptSn m)))) where + module _ (ind : ((c : fst C (suc m)) → incl (fm c) ≡ f (incl c))) where + fib-f-incl : (c : fst C (suc (suc m))) → Type _ + fib-f-incl c = Σ[ x ∈ fst D (suc (suc m)) ] incl x ≡ f (incl c) + + fib-f-l : (c : fst C (suc m)) → fib-f-incl (CW↪ C (suc m) c) + fst (fib-f-l c) = CW↪ D (suc m) (fm c) + snd (fib-f-l c) = sym (push _) ∙∙ ind c ∙∙ cong f (push _) + + fib-f-r : (x : A (suc m)) + → fib-f-incl (invEq (e (suc m)) (inr x)) + fib-f-r x = subst fib-f-incl (cong (invEq (e (suc m))) + (push (x , ptSn m))) + (fib-f-l (α (suc m) (x , ptSn m))) + + fib-f-l-coh : (x : A (suc m)) + → PathP (λ i → fib-f-incl (invEq (e (suc m)) (push (x , ptSn m) i))) + (fib-f-l (α (suc m) (x , ptSn m))) + (fib-f-r x) + fib-f-l-coh x i = + transp (λ j → fib-f-incl (invEq (e (suc m)) (push (x , ptSn m) (i ∧ j)))) + (~ i) + (fib-f-l (α (suc m) (x , ptSn m))) + + mere-fib-f-coh : ∥ ((x : A (suc m)) (y : S₊ m) + → PathP (λ i → fib-f-incl (invEq (e (suc m)) (push (x , y) i))) + (fib-f-l (α (suc m) (x , y))) + (fib-f-r x)) ∥₁ + mere-fib-f-coh = invEq propTrunc≃Trunc1 + (invEq (_ , InductiveFinSatAC 1 (card (suc m)) _) + λ a → fst propTrunc≃Trunc1 (sphereToTrunc m + (sphereElim' m + (λ x → isOfHLevelRetractFromIso m + (invIso (PathPIdTruncIso (suc m))) + (isOfHLevelPathP' m (isProp→isOfHLevelSuc m + (isContr→isProp + (isConnected-CW↪∞ (suc (suc m)) D _))) _ _)) + ∣ fib-f-l-coh a ∣ₕ))) + + module _ (ind : ((c : fst C (suc m)) → incl (fm c) ≡ f (incl c))) + (ind2 : ((x : A (suc m)) (y : S₊ m) + → PathP (λ i → fib-f-incl ind (invEq (e (suc m)) (push (x , y) i))) + (fib-f-l ind (α (suc m) (x , y))) + (fib-f-r ind x))) + where + toFiber : (c : fst C (suc (suc m))) + → fiber (incl {n = (suc (suc m))}) (f (incl c)) + toFiber = CWskel-elim C (suc m) (fib-f-l ind) (fib-f-r ind) ind2 + + toFiberβ : (c : fst C (suc m)) → toFiber (CW↪ C (suc m) c) ≡ fib-f-l ind c + toFiberβ = CWskel-elim-inl C (suc m) (fib-f-l ind) (fib-f-r ind) ind2 + + module _ (p : (n : Fin (suc (suc m))) + → Σ[ h ∈ (fst C (fst n) → fst D (fst n)) ] + ((c : fst C (fst n)) → incl (h c) ≡ f (incl c))) + (f' : (n : Fin (suc m)) (c : fst C (fst n)) + → p (fsuc n) .fst (CW↪ C (fst n) c) ≡ + CW↪ D (fst n) (p (injectSuc n) .fst c)) + (r : (n : A (suc m)) (y : S₊ m) → + CW↪ D (suc m) (p flast .fst (α (suc m) (n , y))) + ≡ CW↪ D (suc m) (p flast .fst (α (suc m) (n , ptSn m)))) + (ind : (n : Fin _) (y : S₊ m) → + PathP + (λ i → + fib-f-incl (p flast .fst) r (p flast .snd) + (invEq (e (suc m)) (push (n , y) i))) + (fib-f-l (p flast .fst) r (p flast .snd) + (α (suc m) (n , y))) + (fib-f-r (p flast .fst) r (p flast .snd) n)) where + + FST-base : Σ[ h ∈ (fst C (suc (suc m)) → fst D (suc (suc m))) ] + ((c : fst C (suc (suc m))) → incl (h c) ≡ f (incl c)) + fst FST-base = fst ∘ toFiber _ _ _ ind + snd FST-base = snd ∘ toFiber _ _ _ ind + + Goalₗ : (n : Fin (suc (suc (suc m)))) → Type _ + Goalₗ n = Σ[ h ∈ (fst C (fst n) → fst D (fst n)) ] + ((c : fst C (fst n)) → incl (h c) ≡ f (incl c)) + + mainₗ : (n : Fin (suc (suc (suc m)))) → Goalₗ n + mainₗ = elimFin FST-base p + + mainᵣ : (n : Fin (suc (suc m))) (c : fst C (fst n)) + → mainₗ (fsuc n) .fst (CW↪ C (fst n) c) + ≡ CW↪ D (fst n) (mainₗ (injectSuc n) .fst c) + mainᵣ = elimFin (λ c + → funExt⁻ (cong fst (elimFinβ {A = Goalₗ} FST-base p .fst)) + (CW↪ C (suc m) c) + ∙ cong fst (toFiberβ _ _ _ ind c) + ∙ cong (CW↪ D (suc m)) + (sym (funExt⁻ + (cong fst (elimFinβ {A = Goalₗ} FST-base p .snd flast)) c))) + λ x c → funExt⁻ (cong fst (elimFinβ {A = Goalₗ} + FST-base p .snd (fsuc x))) (CW↪ C (fst x) c) + ∙ f' x c + ∙ cong (CW↪ D (fst x)) + (sym (funExt⁻ (cong fst + ((elimFinβ {A = Goalₗ} FST-base p .snd) (injectSuc x))) c)) + +-- first main result +CWmap→finCellMap : (C : CWskel ℓ) (D : CWskel ℓ') + (f : realise C → realise D) (m : ℕ) + → ∥ finCellApprox C D f m ∥₁ +CWmap→finCellMap C D f m = + PT.map (λ {(g , hom) + → (record { fmap = fst ∘ g ; fcomm = λ r x → sym (hom r x) }) + , →FinSeqColimHomotopy _ _ (g flast .snd)}) + (approx C D f m) + +---- Part 3. The (finite) cellular approximation theorem for cellular homotopies: ----- +-- Given two (m)-finite cellular maps fₙ, gₙ : Cₙ → Dₙ agreeing on +-- Dₘ ↪ D∞, there is a finite cellular homotopy fₙ ∼ₘ gₙ. + +-- some abbreviations +private + module SeqHomotopyTypes {ℓ ℓ'} {C : Sequence ℓ} {D : Sequence ℓ'} (m : ℕ) + (f-c g-c : FinSequenceMap (Sequence→FinSequence m C) + (Sequence→FinSequence m D)) + where + f = fmap f-c + g = fmap g-c + f-hom = fcomm f-c + g-hom = fcomm g-c + + cell-hom : (n : Fin (suc m)) (c : obj C (fst n)) → Type ℓ' + cell-hom n c = Sequence.map D (f n c) ≡ Sequence.map D (g n c) + + cell-hom-coh : (n : Fin m) (c : obj C (fst n)) + → cell-hom (injectSuc n) c + → cell-hom (fsuc n) (Sequence.map C c) → Type ℓ' + cell-hom-coh n c h h' = + Square (cong (Sequence.map D) h) h' + (cong (Sequence.map D) (f-hom n c)) + (cong (Sequence.map D) (g-hom n c)) + +-- construction +module approx {C : CWskel ℓ} {D : CWskel ℓ'} + (m : ℕ) (f-c g-c : finCellMap m C D) + (h∞' : FinSeqColim→Colim m ∘ finCellMap→FinSeqColim C D f-c + ≡ FinSeqColim→Colim m ∘ finCellMap→FinSeqColim C D g-c) where + open SeqHomotopyTypes m f-c g-c + open CWskel-fields C + + h∞ : (n : Fin (suc m)) (c : fst C (fst n)) + → Path (realise D) (incl (fmap f-c n c)) (incl (fmap g-c n c)) + h∞ n c = funExt⁻ h∞' (fincl n c) + + pathToCellularHomotopy₁ : (t : 1 <ᵗ suc m) (c : fst C 1) + → ∃[ h ∈ cell-hom (1 , t) c ] + Square (h∞ (1 , t) c) + (cong incl h) + (push (f (1 , t) c)) + (push (g (1 , t) c)) + pathToCellularHomotopy₁ t c = + TR.rec squash₁ + (λ {(d , p) + → ∣ d + , (λ i j + → hcomp (λ k → + λ {(i = i0) → doubleCompPath-filler + (sym (push (f (1 , t) c))) + (h∞ _ c) + (push (g (1 , t) c)) (~ k) j + ; (i = i1) → incl (d j) + ; (j = i0) → push (f (1 , t) c) (~ k ∨ i) + ; (j = i1) → push (g (1 , t) c) (~ k ∨ i)}) + (p (~ i) j)) ∣₁}) + (isConnectedCong 1 (CW↪∞ D 2) + (isConnected-CW↪∞ 2 D) h∞* .fst) + where + h∞* : CW↪∞ D 2 (CW↪ D 1 (f (1 , t) c)) ≡ CW↪∞ D 2 (CW↪ D 1 (g (1 , t) c)) + h∞* = sym (push (f (1 , t) c)) ∙∙ h∞ _ c ∙∙ push (g (1 , t) c) + + -- induction step + pathToCellularHomotopy-ind : (n : Fin m) + → (hₙ : (c : fst C (fst n)) → Σ[ h ∈ cell-hom (injectSuc n) c ] + (Square (h∞ (injectSuc n) c) + (cong incl h) + (push (f (injectSuc n) c)) + (push (g (injectSuc n) c)))) + → ∃[ hₙ₊₁ ∈ ((c : fst C (suc (fst n))) + → Σ[ h ∈ cell-hom (fsuc n) c ] + (Square (h∞ (fsuc n) c) + (cong incl h) + (push (f (fsuc n) c)) + (push (g (fsuc n) c)))) ] + ((c : _) → cell-hom-coh n c (hₙ c .fst) + (hₙ₊₁ (CW↪ C (fst n) c) .fst)) + + pathToCellularHomotopy-ind (zero , q) ind = + fst (propTrunc≃ (isoToEquiv (compIso (invIso rUnit×Iso) + (Σ-cong-iso-snd + λ _ → isContr→Iso isContrUnit + ((λ x → ⊥.rec (CW₀-empty C x)) + , λ t → funExt λ x → ⊥.rec (CW₀-empty C x)))))) + (invEq propTrunc≃Trunc1 + (invEq (_ , satAC-CW₁ 1 C _) + λ c → fst propTrunc≃Trunc1 + (pathToCellularHomotopy₁ (fsuc (zero , q) .snd) c))) + pathToCellularHomotopy-ind (suc n' , w) ind = + PT.map (λ q → hₙ₊₁ q , hₙ₊₁-coh q) Πfiber-cong²-hₙ₊₁-push∞ + where + n : Fin m + n = (suc n' , w) + Pushout→C = invEq (e (suc n')) + + hₙ'-filler : (x : fst C (suc n')) → _ + hₙ'-filler x = + doubleCompPath-filler + (sym (f-hom n x)) + (ind x .fst) + (g-hom n x) + + hₙ' : (x : fst C (suc n')) + → f (fsuc n) (CW↪ C (suc n') x) + ≡ g (fsuc n) (CW↪ C (suc n') x) + hₙ' x = hₙ'-filler x i1 + + -- homotopy for inl + hₙ₊₁-inl : (x : fst C (suc n')) + → cell-hom (fsuc n) (invEq (e (suc n')) (inl x)) + hₙ₊₁-inl x = cong (CW↪ D (suc (suc n'))) (hₙ' x) + + -- hₙ₊₁-inl coherent with h∞ + h∞-push-coh : (x : fst C (suc n')) + → Square (h∞ (injectSuc n) x) (h∞ (fsuc n) (CW↪ C (fst n) x)) + (push (f (injectSuc n) x) ∙ (λ i₁ → incl (f-hom n x i₁))) + (push (g (injectSuc n) x) ∙ (λ i₁ → incl (g-hom n x i₁))) + h∞-push-coh x i j = + hcomp (λ k + → λ {(i = i0) → h∞' j (fincl (injectSuc n) x) + ; (i = i1) → h∞' j (fincl (fsuc n) (CW↪ C (fst n) x)) + ; (j = i0) → cong-∙ (FinSeqColim→Colim m) + (fpush n (f (injectSuc n) x)) + (λ i₁ → fincl (fsuc n) (f-hom n x i₁)) k i + ; (j = i1) → cong-∙ (FinSeqColim→Colim m) + (fpush n (g (injectSuc n) x)) + (λ i₁ → fincl (fsuc n) (g-hom n x i₁)) k i}) + (h∞' j (fpush n x i)) + + hₙ₊₁-inl-coh : (x : fst C (fst n)) + → Square (h∞ (fsuc n) (invEq (e (fst n)) (inl x))) + (cong incl (hₙ₊₁-inl x)) + (push (f (fsuc n) (invEq (e (fst n)) (inl x)))) + (push (g (fsuc n) (invEq (e (fst n)) (inl x)))) + hₙ₊₁-inl-coh x i j = + hcomp (λ k + → λ {(i = i0) → h∞ (fsuc n) (CW↪ C (fst n) x) j + ; (i = i1) → push (hₙ' x j) k + ; (j = i0) → push (f (fsuc n) (CW↪ C (fst n) x)) (k ∧ i) + ; (j = i1) → push (g (fsuc n) (CW↪ C (fst n) x)) (k ∧ i)}) + (hcomp (λ k + → λ {(i = i0) → h∞-push-coh x k j + ; (i = i1) → incl (hₙ'-filler x k j) + ; (j = i0) → compPath-filler' + (push (f (injectSuc n) x)) + ((λ i₁ → incl (f-hom n x i₁))) (~ i) k + ; (j = i1) → compPath-filler' + (push (g (injectSuc n) x)) + ((λ i₁ → incl (g-hom n x i₁))) (~ i) k}) + (ind x .snd i j)) + + module _ (x : A (suc n')) (y : S₊ n') where + push-path-filler : I → I → Pushout (α (suc n')) fst + push-path-filler i j = + compPath-filler'' (push (x , y)) (sym (push (x , ptSn n'))) i j + + push-path : + Path (Pushout (α (suc n')) fst) + (inl (α (suc n') (x , y))) + (inl (α (suc n') (x , ptSn n'))) + push-path j = push-path-filler i1 j + + D∞PushSquare : Type ℓ' + D∞PushSquare = + Square {A = realise D} + (cong (CW↪∞ D (suc (suc (suc n')))) + (hₙ₊₁-inl (snd C .snd .fst (suc n') (x , y)))) + (cong (CW↪∞ D (suc (suc (suc n')))) + (hₙ₊₁-inl (snd C .snd .fst (suc n') (x , ptSn n')))) + (λ i → incl (CW↪ D (suc (suc n')) + (f (fsuc n) (Pushout→C (push-path i))))) + (λ i → incl (CW↪ D (suc (suc n')) + (g (fsuc n) (Pushout→C (push-path i))))) + + cong² : PathP (λ i → cell-hom (fsuc n) + (Pushout→C (push-path i))) + (hₙ₊₁-inl (α (suc n') (x , y))) + (hₙ₊₁-inl (α (suc n') (x , ptSn n'))) + → D∞PushSquare + cong² p i j = incl (p i j) + + isConnected-cong² : isConnectedFun (suc n') cong² + isConnected-cong² = + isConnectedCong² (suc n') (CW↪∞ D (3 +ℕ n')) + (isConnected-CW↪∞ (3 +ℕ n') D) + + hₙ₊₁-push∞ : D∞PushSquare + hₙ₊₁-push∞ i j = + hcomp (λ k + → λ {(i = i0) → hₙ₊₁-inl-coh (α (suc n') (x , y)) k j + ; (i = i1) → hₙ₊₁-inl-coh (α (suc n') (x , ptSn n')) k j + ; (j = i0) → push (f (fsuc n) (Pushout→C (push-path i))) k + ; (j = i1) → push (g (fsuc n) (Pushout→C (push-path i))) k}) + (hcomp (λ k + → λ {(i = i0) → h∞' j (fincl (fsuc n) + (Pushout→C (push (x , y) (~ k)))) + ; (i = i1) → h∞' j (fincl (fsuc n) + (Pushout→C (push (x , ptSn n') (~ k)))) + ; (j = i0) → incl (f (fsuc n) + (Pushout→C (push-path-filler k i))) + ; (j = i1) → incl (g (fsuc n) + (Pushout→C (push-path-filler k i)))}) + (h∞' j (fincl (fsuc n) (Pushout→C (inr x))))) + + fiber-cong²-hₙ₊₁-push∞ : hLevelTrunc (suc n') (fiber cong² hₙ₊₁-push∞) + fiber-cong²-hₙ₊₁-push∞ = isConnected-cong² hₙ₊₁-push∞ .fst + + hₙ₊₁-coh∞ : (q : fiber cong² hₙ₊₁-push∞) + → Cube (hₙ₊₁-inl-coh (α (suc n') (x , y))) + (hₙ₊₁-inl-coh (α (suc n') (x , ptSn n'))) + (λ j k → h∞' k (fincl (fsuc n) (Pushout→C (push-path j)))) + (λ j k → incl (fst q j k)) + (λ j i → push (f (fsuc n) (Pushout→C (push-path j))) i) + λ j i → push (g (fsuc n) (Pushout→C (push-path j))) i + hₙ₊₁-coh∞ q j i k = + hcomp (λ r + → λ {(i = i0) → h∞' k (fincl (fsuc n) (Pushout→C (push-path j))) + ; (i = i1) → q .snd (~ r) j k + ; (j = i0) → hₙ₊₁-inl-coh (α (suc n') (x , y)) i k + ; (j = i1) → hₙ₊₁-inl-coh (α (suc n') (x , ptSn n')) i k + ; (k = i0) → push (f (fsuc n) (Pushout→C (push-path j))) i + ; (k = i1) → push (g (fsuc n) (Pushout→C (push-path j))) i}) + (hcomp (λ r + → λ {(i = i0) → h∞' k (fincl (fsuc n) (Pushout→C (push-path j))) + ; (j = i0) → hₙ₊₁-inl-coh (α (suc n') (x , y)) (i ∧ r) k + ; (j = i1) → hₙ₊₁-inl-coh (α (suc n') (x , ptSn n')) (i ∧ r) k + ; (k = i0) → push (f (fsuc n) + (Pushout→C (push-path j))) (i ∧ r) + ; (k = i1) → push (g (fsuc n) + (Pushout→C (push-path j))) (i ∧ r)}) + (hcomp (λ r + → λ {(i = i0) → h∞' k (fincl (fsuc n) + (Pushout→C (push-path-filler r j))) + ; (j = i0) → h∞' k (fincl (fsuc n) (invEq (e (suc n')) + (push (x , y) (~ r)))) + ; (j = i1) → h∞' k (fincl (fsuc n) (invEq (e (suc n')) + (push (x , ptSn n') (~ r)))) + ; (k = i0) → incl (f (fsuc n) + (Pushout→C (push-path-filler r j))) + ; (k = i1) → incl (g (fsuc n) + (Pushout→C (push-path-filler r j)))}) + (h∞' k (fincl (fsuc n) (Pushout→C (inr x)))))) + + Πfiber-cong²-hₙ₊₁-push∞ : + ∥ ((x : _) (y : _) → fiber (cong² x y) (hₙ₊₁-push∞ x y)) ∥₁ + Πfiber-cong²-hₙ₊₁-push∞ = + Iso.inv propTruncTrunc1Iso + (invEq (_ , InductiveFinSatAC 1 _ _) + λ x → Iso.fun propTruncTrunc1Iso + (sphereToTrunc n' (fiber-cong²-hₙ₊₁-push∞ x))) + + module _ (q : (x : Fin (fst (snd C) (suc n'))) (y : S₊ n') + → fiber (cong² x y) (hₙ₊₁-push∞ x y)) where + agrees : (x : fst C (suc n')) (ϕ : cell-hom (fsuc n) (CW↪ C (suc n') x)) + → Type _ + agrees x ϕ = Square (h∞ (fsuc n) (CW↪ C (suc n') x)) + (cong incl ϕ) + (push (f (fsuc n) (CW↪ C (suc n') x))) + (push (g (fsuc n) (CW↪ C (suc n') x))) + + main-inl : (x : fst C (suc n')) + → Σ (cell-hom (fsuc n) (CW↪ C (suc n') x)) + (agrees x) + main-inl x = hₙ₊₁-inl x , hₙ₊₁-inl-coh x + + main-push : (x : A (suc n')) (y : S₊ n') + → PathP (λ i → Σ[ ϕ ∈ (cell-hom (fsuc n) (Pushout→C (push-path x y i))) ] + (Square (h∞ (fsuc n) (Pushout→C (push-path x y i))) + (λ j → incl (ϕ j)) + (push (f (fsuc n) (Pushout→C (push-path x y i)))) + (push (g (fsuc n) (Pushout→C (push-path x y i)))))) + (main-inl (α (suc n') (x , y))) + (main-inl (α (suc n') (x , ptSn n'))) + main-push x y = ΣPathP (fst (q x y) , hₙ₊₁-coh∞ x y (q x y)) + + hₙ₊₁ : (c : fst C (fst (fsuc n))) + → Σ[ ϕ ∈ (cell-hom (fsuc n) c) ] + Square (h∞ (fsuc n) c) + (cong incl ϕ) + (push (f (fsuc n) c)) + (push (g (fsuc n) c)) + hₙ₊₁ = CWskel-elim' C n' main-inl main-push + + hₙ₊₁-coh-pre : (c : fst C (suc n')) → + Square (cong (CW↪ D (suc (suc n'))) (ind c .fst)) + (hₙ₊₁-inl c) + (cong (CW↪ D (suc (suc n'))) (f-hom n c)) + (cong (CW↪ D (suc (suc n'))) (g-hom n c)) + hₙ₊₁-coh-pre c i j = CW↪ D (suc (suc n')) (hₙ'-filler c i j) + + hₙ₊₁-coh : (c : fst C (suc n')) → + Square (cong (CW↪ D (suc (suc n'))) (ind c .fst)) + (hₙ₊₁ (CW↪ C (suc n') c) .fst) + (cong (CW↪ D (suc (suc n'))) (f-hom n c)) + (cong (CW↪ D (suc (suc n'))) (g-hom n c)) + hₙ₊₁-coh c = hₙ₊₁-coh-pre c + ▷ λ i → CWskel-elim'-inl C n' + {B = λ c → Σ[ ϕ ∈ cell-hom (fsuc n) c ] + Square (h∞ (fsuc n) c) (cong incl ϕ) + (push (f (fsuc n) c)) (push (g (fsuc n) c))} + main-inl main-push c (~ i) .fst + +-- converting the above to something on the right form +private + pathToCellularHomotopy-main : + {C : CWskel ℓ} {D : CWskel ℓ'} (m : ℕ) (f-c g-c : finCellMap m C D) + (h∞' : FinSeqColim→Colim m ∘ finCellMap→FinSeqColim C D f-c + ≡ FinSeqColim→Colim m ∘ finCellMap→FinSeqColim C D g-c) + → ∥ finCellHomRel m f-c g-c (approx.h∞ m f-c g-c h∞') ∥₁ + pathToCellularHomotopy-main {C = C} zero f-c g-c h∞' = + ∣ (record { fhom = λ {(zero , p) x → ⊥.rec (CW₀-empty C x)} + ; fcoh = λ {()} }) + , (λ { (zero , p) x → ⊥.rec (CW₀-empty C x)}) ∣₁ + pathToCellularHomotopy-main {C = C} {D = D} (suc zero) f-c g-c h∞' = + PT.map (λ {(d , h) + → (record + { fhom = λ {(zero , p) x → ⊥.rec (CW₀-empty C x) + ; (suc zero , p) → d} + ; fcoh = λ {(zero , p) → λ x → ⊥.rec (CW₀-empty C x)} }) + , (λ {(zero , p) x → ⊥.rec (CW₀-empty C x) + ; (suc zero , tt) → h})}) (invEq (_ , satAC∃Fin-C0 C _ _) k) + where + k : (c : _) → _ + k c = (approx.pathToCellularHomotopy₁ (suc zero) f-c g-c + h∞' tt c) + pathToCellularHomotopy-main {C = C} {D = D} (suc (suc m)) f-c g-c h∞' = + PT.rec squash₁ + (λ ind → PT.map + (λ {(f , p) → + (record { fhom = main-hom ind f p + ; fcoh = main-coh ind f p }) + , ∞coh ind f p}) + (pathToCellularHomotopy-ind flast + λ c → (finCellHom.fhom (ind .fst) flast c) + , (ind .snd flast c))) + (pathToCellularHomotopy-main {C = C} {D = D} (suc m) + (finCellMap↓ f-c) + (finCellMap↓ g-c) + h') + where + open approx _ f-c g-c h∞' + finSeqColim↑ : ∀ {ℓ} {X : Sequence ℓ} {m : ℕ} + → FinSeqColim m X → FinSeqColim (suc m) X + finSeqColim↑ (fincl n x) = fincl (injectSuc n) x + finSeqColim↑ {m = suc m} (fpush n x i) = fpush (injectSuc n) x i + + h' : FinSeqColim→Colim (suc m) ∘ + finCellMap→FinSeqColim C D (finCellMap↓ f-c) + ≡ + FinSeqColim→Colim (suc m) ∘ + finCellMap→FinSeqColim C D (finCellMap↓ g-c) + h' = funExt λ { (fincl n x) j → h∞' j (fincl (injectSuc n) x) + ; (fpush n x i) j → h∞' j (fpush (injectSuc n) x i)} + + open SeqHomotopyTypes + + module _ + (ind : finCellHomRel (suc m) + (finCellMap↓ f-c) (finCellMap↓ g-c) + ((approx.h∞ (suc m) (finCellMap↓ f-c) (finCellMap↓ g-c) h'))) + (f : (c : fst C (suc (suc m))) → + Σ[ h ∈ (cell-hom (suc (suc m)) f-c g-c flast c) ] + (Square (h∞ flast c) (cong incl h) + (push (fmap f-c flast c)) (push (fmap g-c flast c)))) + (fp : (c : fst C (suc m)) → + cell-hom-coh (suc (suc m)) f-c g-c flast c + (finCellHom.fhom (ind .fst) flast c) + (f (CW↪ C (suc m) c) .fst)) where + + main-hom-typ : (n : Fin (suc (suc (suc m)))) + → Type _ + main-hom-typ n = + (x : C .fst (fst n)) + → CW↪ D (fst n) (f-c .fmap n x) + ≡ CW↪ D (fst n) (g-c .fmap n x) + + main-hom : (n : Fin (suc (suc (suc m)))) → main-hom-typ n + main-hom = elimFin (fst ∘ f) (finCellHom.fhom (fst ind)) + + main-homβ = elimFinβ {A = main-hom-typ} (fst ∘ f) + (finCellHom.fhom (fst ind)) + + main-coh : (n : Fin (suc (suc m))) (c : C .fst (fst n)) + → Square + (cong (CW↪ D (suc (fst n))) + (main-hom (injectSuc n) c)) + (main-hom (fsuc n) + (CW↪ C (fst n) c)) + (cong (CW↪ D (suc (fst n))) (fcomm f-c n c)) + (cong (CW↪ D (suc (fst n))) (fcomm g-c n c)) + main-coh = + elimFin + (λ c → cong (cong (CW↪ D (suc (suc m)))) + (funExt⁻ (main-homβ .snd flast) c) + ◁ fp c + ▷ sym (funExt⁻ (main-homβ .fst) (CW↪ C (suc m) c))) + λ n c + → cong (cong (CW↪ D (suc (fst n)))) + (funExt⁻ (main-homβ .snd (injectSuc n)) c) + ◁ finCellHom.fcoh (fst ind) n c + ▷ sym (funExt⁻ (main-homβ .snd (fsuc n)) (CW↪ C (fst n) c)) + + ∞coh : (n : Fin (suc (suc (suc m)))) (x : fst C (fst n)) + → Square (h∞ n x) (λ i → incl (main-hom n x i)) + (push (f-c .fmap n x)) (push (g-c .fmap n x)) + ∞coh = elimFin + (λ c → f c .snd ▷ λ i j → incl (main-homβ .fst (~ i) c j)) + λ n c → ind .snd n c ▷ λ i j → incl (main-homβ .snd n (~ i) c j) + +-- second main theorem +pathToCellularHomotopy : + {C : CWskel ℓ} {D : CWskel ℓ'} (m : ℕ) (f-c g-c : finCellMap m C D) + → ((x : fst C m) → Path (realise D) (incl (fmap f-c flast x)) + (incl (fmap g-c flast x))) + → ∥ finCellHom m f-c g-c ∥₁ +pathToCellularHomotopy {C} {D} m f-c g-c h = + PT.map fst + (pathToCellularHomotopy-main m f-c g-c + (→FinSeqColimHomotopy _ _ h)) diff --git a/Cubical/Data/CW/Base.agda b/Cubical/Data/CW/Base.agda index b120476c1c..da4add8e3e 100644 --- a/Cubical/Data/CW/Base.agda +++ b/Cubical/Data/CW/Base.agda @@ -1,8 +1,6 @@ -{-# OPTIONS --cubical --safe --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} -{-This file contains: - --} +-- This file contains definition of CW complexes and skeleta. module Cubical.Data.CW.Base where diff --git a/Cubical/Data/CW/Homology.agda b/Cubical/Data/CW/Homology.agda new file mode 100644 index 0000000000..5ad06b034c --- /dev/null +++ b/Cubical/Data/CW/Homology.agda @@ -0,0 +1,243 @@ +{-# OPTIONS --safe --lossy-unification #-} +{- +This file contains: +1. Functoriality of Hˢᵏᵉˡ +2. Construction of cellular homology, including funtoriality +-} + +module Cubical.Data.CW.Homology where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat +open import Cubical.Data.Fin.Inductive.Base + +open import Cubical.Data.CW +open import Cubical.Data.CW.Map +open import Cubical.Data.CW.Homotopy +open import Cubical.Data.CW.ChainComplex +open import Cubical.Data.CW.Approximation + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.GroupPath +open import Cubical.Algebra.ChainComplex + +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SequentialColimit + +private + variable + ℓ ℓ' ℓ'' : Level + +------ Part 1. Functoriality of Hˢᵏᵉˡ ------ +Hˢᵏᵉˡ→-pre : (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) + (f : finCellMap (suc (suc (suc m))) C D) + → GroupHom (Hˢᵏᵉˡ C m) (Hˢᵏᵉˡ D m) +Hˢᵏᵉˡ→-pre C D m f = finCellMap→HomologyMap m f + +private + Hˢᵏᵉˡ→-pre' : (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) + → (f : realise C → realise D) + → ∥ finCellApprox C D f (suc (suc (suc m))) ∥₁ + → GroupHom (Hˢᵏᵉˡ C m) (Hˢᵏᵉˡ D m) + Hˢᵏᵉˡ→-pre' C D m f = + rec→Set isSetGroupHom + (λ f → Hˢᵏᵉˡ→-pre C D m (f .fst)) + main + where + main : 2-Constant (λ f₁ → Hˢᵏᵉˡ→-pre C D m (f₁ .fst)) + main f g = PT.rec (isSetGroupHom _ _) + (λ q → finChainHomotopy→HomologyPath _ _ + (cellHom-to-ChainHomotopy _ q) flast) + (pathToCellularHomotopy _ (fst f) (fst g) + λ x → funExt⁻ (snd f ∙ sym (snd g)) (fincl flast x)) + +Hˢᵏᵉˡ→ : {C : CWskel ℓ} {D : CWskel ℓ'} (m : ℕ) + (f : realise C → realise D) + → GroupHom (Hˢᵏᵉˡ C m) (Hˢᵏᵉˡ D m) +Hˢᵏᵉˡ→ {C = C} {D} m f = + Hˢᵏᵉˡ→-pre' C D m f + (CWmap→finCellMap C D f (suc (suc (suc m)))) + +Hˢᵏᵉˡ→β : (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) + {f : realise C → realise D} + (ϕ : finCellApprox C D f (suc (suc (suc m)))) + → Hˢᵏᵉˡ→ {C = C} {D} m f ≡ Hˢᵏᵉˡ→-pre C D m (ϕ .fst) +Hˢᵏᵉˡ→β C D m {f = f} ϕ = + cong (Hˢᵏᵉˡ→-pre' C D m f) (squash₁ _ ∣ ϕ ∣₁) + +Hˢᵏᵉˡ→id : (m : ℕ) {C : CWskel ℓ} + → Hˢᵏᵉˡ→ {C = C} m (idfun _) ≡ idGroupHom +Hˢᵏᵉˡ→id m {C = C} = + Hˢᵏᵉˡ→β C C m {idfun _} (idFinCellApprox (suc (suc (suc m)))) + ∙ finCellMap→HomologyMapId _ + +Hˢᵏᵉˡ→comp : (m : ℕ) + {C : CWskel ℓ} {D : CWskel ℓ'} {E : CWskel ℓ''} + (g : realise D → realise E) + (f : realise C → realise D) + → Hˢᵏᵉˡ→ m (g ∘ f) + ≡ compGroupHom (Hˢᵏᵉˡ→ m f) (Hˢᵏᵉˡ→ m g) +Hˢᵏᵉˡ→comp m {C = C} {D = D} {E = E} g f = + PT.rec2 (isSetGroupHom _ _) + main + (CWmap→finCellMap C D f (suc (suc (suc m)))) + (CWmap→finCellMap D E g (suc (suc (suc m)))) + where + module _ (F : finCellApprox C D f (suc (suc (suc m)))) + (G : finCellApprox D E g (suc (suc (suc m)))) + where + comps : finCellApprox C E (g ∘ f) (suc (suc (suc m))) + comps = compFinCellApprox _ {g = g} {f} G F + + eq1 : Hˢᵏᵉˡ→ m (g ∘ f) + ≡ Hˢᵏᵉˡ→-pre C E m (composeFinCellMap _ (fst G) (fst F)) + eq1 = Hˢᵏᵉˡ→β C E m {g ∘ f} comps + + eq2 : compGroupHom (Hˢᵏᵉˡ→ m f) (Hˢᵏᵉˡ→ m g) + ≡ compGroupHom (Hˢᵏᵉˡ→-pre C D m (fst F)) (Hˢᵏᵉˡ→-pre D E m (fst G)) + eq2 = cong₂ compGroupHom (Hˢᵏᵉˡ→β C D m {f = f} F) (Hˢᵏᵉˡ→β D E m {f = g} G) + + main : Hˢᵏᵉˡ→ m (g ∘ f) ≡ compGroupHom (Hˢᵏᵉˡ→ m f) (Hˢᵏᵉˡ→ m g) + main = eq1 ∙∙ finCellMap→HomologyMapComp _ _ _ ∙∙ sym eq2 + +-- preservation of equivalence +Hˢᵏᵉˡ→Iso : {C : CWskel ℓ} {D : CWskel ℓ'} (m : ℕ) + (e : realise C ≃ realise D) → GroupIso (Hˢᵏᵉˡ C m) (Hˢᵏᵉˡ D m) +Iso.fun (fst (Hˢᵏᵉˡ→Iso m e)) = fst (Hˢᵏᵉˡ→ m (fst e)) +Iso.inv (fst (Hˢᵏᵉˡ→Iso m e)) = fst (Hˢᵏᵉˡ→ m (invEq e)) +Iso.rightInv (fst (Hˢᵏᵉˡ→Iso m e)) = + funExt⁻ (cong fst (sym (Hˢᵏᵉˡ→comp m (fst e) (invEq e)) + ∙∙ cong (Hˢᵏᵉˡ→ m) (funExt (secEq e)) + ∙∙ Hˢᵏᵉˡ→id m)) +Iso.leftInv (fst (Hˢᵏᵉˡ→Iso m e)) = + funExt⁻ (cong fst (sym (Hˢᵏᵉˡ→comp m (invEq e) (fst e)) + ∙∙ cong (Hˢᵏᵉˡ→ m) (funExt (retEq e)) + ∙∙ Hˢᵏᵉˡ→id m)) +snd (Hˢᵏᵉˡ→Iso m e) = snd (Hˢᵏᵉˡ→ m (fst e)) + +Hˢᵏᵉˡ→Equiv : {C : CWskel ℓ} {D : CWskel ℓ'} (m : ℕ) + (e : realise C ≃ realise D) → GroupEquiv (Hˢᵏᵉˡ C m) (Hˢᵏᵉˡ D m) +Hˢᵏᵉˡ→Equiv m e = GroupIso→GroupEquiv (Hˢᵏᵉˡ→Iso m e) + + +------ Part 2. Definition of cellular homology ------ +Hᶜʷ : ∀ (C : CW ℓ) (n : ℕ) → Group₀ +Hᶜʷ (C , CWstr) n = + PropTrunc→Group + (λ Cskel → Hˢᵏᵉˡ (Cskel .fst) n) + (λ Cskel Dskel + → Hˢᵏᵉˡ→Equiv n (compEquiv (invEquiv (snd Cskel)) (snd Dskel))) + coh + CWstr + where + coh : (Cskel Dskel Eskel : isCW C) (t : fst (Hˢᵏᵉˡ (Cskel .fst) n)) + → fst (fst (Hˢᵏᵉˡ→Equiv n (compEquiv (invEquiv (snd Dskel)) (snd Eskel)))) + (fst (fst (Hˢᵏᵉˡ→Equiv n (compEquiv (invEquiv (snd Cskel)) (snd Dskel)))) t) + ≡ fst (fst (Hˢᵏᵉˡ→Equiv n (compEquiv (invEquiv (snd Cskel)) (snd Eskel)))) t + coh (C' , e) (D , f) (E , h) = + funExt⁻ (cong fst + (sym (Hˢᵏᵉˡ→comp n (fst (compEquiv (invEquiv f) h)) + (fst (compEquiv (invEquiv e) f))) + ∙ cong (Hˢᵏᵉˡ→ n) (funExt λ x → cong (fst h) (retEq f _)))) + +-- lemmas for functoriality +private + module _ {C : Type ℓ} {D : Type ℓ'} (f : C → D) (n : ℕ) where + right : (cwC : isCW C) (cwD1 : isCW D) (cwD2 : isCW D) + → PathP (λ i → GroupHom (Hᶜʷ (C , ∣ cwC ∣₁) n) + (Hᶜʷ (D , squash₁ ∣ cwD1 ∣₁ ∣ cwD2 ∣₁ i) n)) + (Hˢᵏᵉˡ→ n (λ x → fst (snd cwD1) (f (invEq (snd cwC) x)))) + (Hˢᵏᵉˡ→ n (λ x → fst (snd cwD2) (f (invEq (snd cwC) x)))) + right cwC cwD1 cwD2 = + PathPGroupHomₗ _ + (cong (Hˢᵏᵉˡ→ n) (funExt (λ s + → cong (fst (snd cwD2)) (sym (retEq (snd cwD1) _)))) + ∙ Hˢᵏᵉˡ→comp n _ _) + + left : (cwC1 : isCW C) (cwC2 : isCW C) (cwD : isCW D) + → PathP (λ i → GroupHom (Hᶜʷ (C , squash₁ ∣ cwC1 ∣₁ ∣ cwC2 ∣₁ i) n) + (Hᶜʷ (D , ∣ cwD ∣₁) n)) + (Hˢᵏᵉˡ→ n (λ x → fst (snd cwD) (f (invEq (snd cwC1) x)))) + (Hˢᵏᵉˡ→ n (λ x → fst (snd cwD) (f (invEq (snd cwC2) x)))) + left cwC1 cwC2 cwD = + PathPGroupHomᵣ (Hˢᵏᵉˡ→Equiv n (compEquiv (invEquiv (snd cwC1)) (snd cwC2))) + (sym (Hˢᵏᵉˡ→comp n (λ x → fst (snd cwD) (f (invEq (snd cwC2) x))) + (compEquiv (invEquiv (snd cwC1)) (snd cwC2) .fst)) + ∙ cong (Hˢᵏᵉˡ→ n) (funExt (λ x + → cong (fst (snd cwD) ∘ f) + (retEq (snd cwC2) _)))) + + left-right : (x y : isCW C) (v w : isCW D) → + SquareP + (λ i j → + GroupHom (Hᶜʷ (C , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) n) + (Hᶜʷ (D , squash₁ ∣ v ∣₁ ∣ w ∣₁ j) n)) + (right x v w) (right y v w) (left x y v) (left x y w) + left-right _ _ _ _ = isSet→SquareP + (λ _ _ → isSetGroupHom) _ _ _ _ + + Hᶜʷ→pre : (cwC : ∥ isCW C ∥₁) (cwD : ∥ isCW D ∥₁) + → GroupHom (Hᶜʷ (C , cwC) n) (Hᶜʷ (D , cwD) n) + Hᶜʷ→pre = + elim2→Set (λ _ _ → isSetGroupHom) + (λ cwC cwD → Hˢᵏᵉˡ→ n (fst (snd cwD) ∘ f ∘ invEq (snd cwC))) + left right + (λ _ _ _ _ → isSet→SquareP + (λ _ _ → isSetGroupHom) _ _ _ _) + +-- functoriality +Hᶜʷ→ : {C : CW ℓ} {D : CW ℓ'} (n : ℕ) (f : C →ᶜʷ D) + → GroupHom (Hᶜʷ C n) (Hᶜʷ D n) +Hᶜʷ→ {C = C , cwC} {D = D , cwD} n f = Hᶜʷ→pre f n cwC cwD + +Hᶜʷ→id : {C : CW ℓ} (n : ℕ) + → Hᶜʷ→ {C = C} {C} n (idfun (fst C)) + ≡ idGroupHom +Hᶜʷ→id {C = C , cwC} n = + PT.elim {P = λ cwC + → Hᶜʷ→ {C = C , cwC} {C , cwC} n (idfun C) ≡ idGroupHom} + (λ _ → isSetGroupHom _ _) + (λ cwC* → cong (Hˢᵏᵉˡ→ n) (funExt (secEq (snd cwC*))) + ∙ Hˢᵏᵉˡ→id n) cwC + +Hᶜʷ→comp : {C : CW ℓ} {D : CW ℓ'} {E : CW ℓ''} (n : ℕ) + (g : D →ᶜʷ E) (f : C →ᶜʷ D) + → Hᶜʷ→ {C = C} {E} n (g ∘ f) + ≡ compGroupHom (Hᶜʷ→ {C = C} {D} n f) (Hᶜʷ→ {C = D} {E} n g) +Hᶜʷ→comp {C = C , cwC} {D = D , cwD} {E = E , cwE} n g f = + PT.elim3 {P = λ cwC cwD cwE + → Hᶜʷ→ {C = C , cwC} {E , cwE} n (g ∘ f) + ≡ compGroupHom (Hᶜʷ→ {C = C , cwC} {D , cwD} n f) + (Hᶜʷ→ {C = D , cwD} {E , cwE} n g)} + (λ _ _ _ → isSetGroupHom _ _) + (λ cwC cwD cwE + → cong (Hˢᵏᵉˡ→ n) (funExt (λ x → cong (fst (snd cwE) ∘ g) + (sym (retEq (snd cwD) _)))) + ∙ Hˢᵏᵉˡ→comp n _ _) + cwC cwD cwE + +-- preservation of equivalence +Hᶜʷ→Iso : {C : CW ℓ} {D : CW ℓ'} (m : ℕ) + (e : fst C ≃ fst D) → GroupIso (Hᶜʷ C m) (Hᶜʷ D m) +Iso.fun (fst (Hᶜʷ→Iso m e)) = fst (Hᶜʷ→ m (fst e)) +Iso.inv (fst (Hᶜʷ→Iso m e)) = fst (Hᶜʷ→ m (invEq e)) +Iso.rightInv (fst (Hᶜʷ→Iso m e)) = + funExt⁻ (cong fst (sym (Hᶜʷ→comp m (fst e) (invEq e)) + ∙∙ cong (Hᶜʷ→ m) (funExt (secEq e)) + ∙∙ Hᶜʷ→id m)) +Iso.leftInv (fst (Hᶜʷ→Iso m e)) = + funExt⁻ (cong fst (sym (Hᶜʷ→comp m (invEq e) (fst e)) + ∙∙ cong (Hᶜʷ→ m) (funExt (retEq e)) + ∙∙ Hᶜʷ→id m)) +snd (Hᶜʷ→Iso m e) = snd (Hᶜʷ→ m (fst e)) + +Hᶜʷ→Equiv : {C : CW ℓ} {D : CW ℓ'} (m : ℕ) + (e : fst C ≃ fst D) → GroupEquiv (Hᶜʷ C m) (Hᶜʷ D m) +Hᶜʷ→Equiv m e = GroupIso→GroupEquiv (Hᶜʷ→Iso m e) diff --git a/Cubical/Data/CW/Map.agda b/Cubical/Data/CW/Map.agda index 5eeeb833f2..d0747afbf2 100644 --- a/Cubical/Data/CW/Map.agda +++ b/Cubical/Data/CW/Map.agda @@ -88,9 +88,16 @@ module _ (m : ℕ) where composeFinCellMap : (g : finCellMap D E) (f : finCellMap C D) → finCellMap C E composeFinCellMap = composeFinSequenceMap m +open FinSequenceMap +finCellMap→FinSeqColim : (C : CWskel ℓ) (D : CWskel ℓ') + {m : ℕ} → finCellMap m C D → FinSeqColim m (realiseSeq C) → FinSeqColim m (realiseSeq D) +finCellMap→FinSeqColim C D {m = m} f (fincl n x) = fincl n (fmap f n x) +finCellMap→FinSeqColim C D {m = m} f (fpush n x i) = + (fpush n (fmap f (injectSuc n) x) ∙ cong (fincl (fsuc n)) (fcomm f n x)) i + finCellMap↓ : {m : ℕ} {C : CWskel ℓ} {D : CWskel ℓ'} → finCellMap (suc m) C D → finCellMap m C D -FinSequenceMap.fmap (finCellMap↓ {m = m} ϕ) x = FinSequenceMap.fmap ϕ (injectSuc x) -FinSequenceMap.fcomm (finCellMap↓ {m = suc m} {C = C} ϕ) x r = FinSequenceMap.fcomm ϕ (injectSuc x) r +fmap (finCellMap↓ {m = m} ϕ) x = fmap ϕ (injectSuc x) +fcomm (finCellMap↓ {m = suc m} {C = C} ϕ) x r = fcomm ϕ (injectSuc x) r -- A cellular map between two CW complexes diff --git a/Cubical/Data/CW/Properties.agda b/Cubical/Data/CW/Properties.agda index 78e6412b53..7034cb2b5c 100644 --- a/Cubical/Data/CW/Properties.agda +++ b/Cubical/Data/CW/Properties.agda @@ -1,7 +1,7 @@ -{-# OPTIONS --cubical --safe --lossy-unification #-} - -{-This file contains: +{-# OPTIONS --safe --lossy-unification #-} +{-This file contains elimination principles and basic properties of CW +complexes/skeleta. -} module Cubical.Data.CW.Properties where @@ -40,13 +40,15 @@ open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup open import Cubical.Axiom.Choice open import Cubical.Relation.Nullary +open import Cubical.Homotopy.Connected + open Sequence private variable - ℓ ℓ' : Level + ℓ ℓ' ℓ'' : Level CW₀-empty : (C : CWskel ℓ) → ¬ fst C 0 CW₀-empty C = snd (snd (snd C)) .fst @@ -165,23 +167,6 @@ module _ {ℓ : Level} (C : CWskel ℓ) where → CWskel-elim' (invEq (e (suc n)) (inl x)) ≡ inler x CWskel-elim'-inl = CWskel-elim-inl (suc n) {B = B} inler _ _ --- open import Cubical.Axiom.Choice --- open import Cubical.HITs.Truncation as TR --- module _ {ℓ : Level} (C : CWskel ℓ) where --- open CWskel-fields C --- CWskel-elim-trunc : (n : ℕ) {B : fst C (suc (suc (suc n))) → Type ℓ'} --- → (f : ((x : fst C (suc (suc n))) → B (CW↪ C (suc (suc n)) x))) --- → ∃[ f^ ∈ ((x : _) → B x) ] ((c : _) → f^ (CW↪ C (suc (suc n)) c) ≡ f c) --- CWskel-elim-trunc n {B = B} f = PT.map (λ F → (CWskel-elim' C (suc n) f F) , CWskel-elim'-inl C (suc n) f F) F --- where --- F : ∥ ((x : Fin (card (suc (suc n)))) (y : S₊ (suc n)) --- → PathP (λ i → B (invEq (e (suc (suc n))) ((push (x , y) ∙ sym (push (x , ptSn (suc n)))) i))) --- (f (α (suc (suc n)) (x , y))) (f (α (suc (suc n)) (x , ptSn (suc n))))) ∥₁ --- F = invEq propTrunc≃Trunc1 (invEq (_ , FinSatAC _ _ _) --- λ x → fst propTrunc≃Trunc1 --- (sphereToTrunc (suc n) {!!})) --- q = FinSatAC - finCWskel≃ : (n : ℕ) (C : finCWskel ℓ n) (m : ℕ) → n ≤ m → fst C n ≃ fst C m finCWskel≃ n C m (zero , diff) = substEquiv (λ n → fst C n) diff finCWskel≃ n C zero (suc x , diff) = ⊥.rec (snotz diff) @@ -192,8 +177,8 @@ finCWskel≃ n C (suc m) (suc x , diff) = (substEquiv (λ n → fst C n) diff))) -- C₁ satisfies AC -satAC-CW₁ : ∀ {ℓ ℓ'} (n : ℕ) (C : CWskel ℓ) → satAC ℓ' n (fst C (suc zero)) -satAC-CW₁ {ℓ' = ℓ'} n C A = +satAC-CW₁ : (n : ℕ) (C : CWskel ℓ) → satAC ℓ' n (fst C (suc zero)) +satAC-CW₁ n C A = subst isEquiv (choicefun≡ n) (isoToIsEquiv (choicefun' n)) where fin = Fin (snd C .fst zero) @@ -218,3 +203,47 @@ satAC-CW₁ {ℓ' = ℓ'} n C A = (isOfHLevelΠ (suc n) (λ _ → isOfHLevelTrunc (suc n))) _ _) λ f → funExt λ a → cong ∣_∣ (funExt⁻ ((Iso.leftInv (domIsoDep (equivToIso fin→))) f) a)) + +satAC∃Fin-C0 : (C : CWskel ℓ) → satAC∃ ℓ' ℓ'' (fst C 1) +satAC∃Fin-C0 C = + subst (satAC∃ _ _) + (ua (compEquiv (invEquiv LiftEquiv) (invEquiv (CW₁-discrete C)))) + λ T c → isoToIsEquiv (iso _ + (λ f → PT.map (λ p → (λ { (lift x) → p .fst x}) + , λ { (lift x) → p .snd x}) + (invEq (_ , (InductiveFinSatAC∃ (snd C .fst 0)) + (T ∘ lift) (c ∘ lift)) (f ∘ lift))) + (λ _ → (isPropΠ λ _ → squash₁) _ _) + λ _ → squash₁ _ _) + +--- Connectivity of CW complexes +-- The embedding of stage n into stage n+1 is (n+1)-connected +-- 2 calls to univalence in there +isConnected-CW↪ : (n : ℕ) (C : CWskel ℓ) → isConnectedFun n (CW↪ C n) +isConnected-CW↪ zero C x = isContrUnit* +isConnected-CW↪ (suc n) C = + EquivJ (λ X E → isConnectedFun (suc n) (λ x → invEq E (inl x))) + inPushoutConnected (e₊ (suc n)) + where + A = snd C .fst + α = snd C .snd .fst + e₊ = snd C .snd .snd .snd + + inPushout : fst C (suc n) → Pushout (α (suc n)) fst + inPushout x = inl x + + fstProjPath : (b : Fin (A (suc n))) → S₊ n ≡ fiber fst b + fstProjPath b = ua (fiberProjEquiv (Fin (A (suc n))) (λ _ → S₊ n) b) + + inPushoutConnected : isConnectedFun (suc n) inPushout + inPushoutConnected = inlConnected (suc n) (α (suc n)) fst + λ b → subst (isConnected (suc n)) (fstProjPath b) (sphereConnected n) + +-- The embedding of stage n into the colimit is (n+1)-connected +isConnected-CW↪∞ : (n : ℕ) (C : CWskel ℓ) → isConnectedFun n (CW↪∞ C n) +isConnected-CW↪∞ zero C b = isContrUnit* +isConnected-CW↪∞ (suc n) C = isConnectedIncl∞ (realiseSeq C) (suc n) (suc n) subtr + where + subtr : (k : ℕ) → isConnectedFun (suc n) (CW↪ C (k +ℕ (suc n))) + subtr k = isConnectedFunSubtr (suc n) k (CW↪ C (k +ℕ (suc n))) + (isConnected-CW↪ (k +ℕ (suc n)) C) diff --git a/Cubical/Data/CW/Subcomplex.agda b/Cubical/Data/CW/Subcomplex.agda new file mode 100644 index 0000000000..5e0de1277d --- /dev/null +++ b/Cubical/Data/CW/Subcomplex.agda @@ -0,0 +1,195 @@ +{-# OPTIONS --safe --lossy-unification #-} + +module Cubical.Data.CW.Subcomplex where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat +open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Data.CW +open import Cubical.Data.CW.ChainComplex +open import Cubical.Data.CW.Approximation + +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.PropositionalTruncation as PT + + +open import Cubical.Relation.Nullary + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.ChainComplex + +private + variable + ℓ ℓ' ℓ'' : Level + +-- finite subcomplex of a cw complex +module _ (C : CWskel ℓ) where + subComplexFam : (n : ℕ) (m : ℕ) → Type ℓ + subComplexFam n m with (m ≟ᵗ n) + ... | lt x = fst C m + ... | eq x = fst C m + ... | gt x = fst C n + + subComplexCard : (n : ℕ) → ℕ → ℕ + subComplexCard n m with (m ≟ᵗ n) + ... | lt x = snd C .fst m + ... | eq x = 0 + ... | gt x = 0 + + subComplexα : (n m : ℕ) → Fin (subComplexCard n m) × S⁻ m → subComplexFam n m + subComplexα n m with (m ≟ᵗ n) + ... | lt x = snd C .snd .fst m + ... | eq x = λ x → ⊥.rec (¬Fin0 (fst x)) + ... | gt x = λ x → ⊥.rec (¬Fin0 (fst x)) + + subComplex₀-empty : (n : ℕ) → ¬ subComplexFam n zero + subComplex₀-empty n with (zero ≟ᵗ n) + ... | lt x = CW₀-empty C + ... | eq x = CW₀-empty C + + subComplexFam≃Pushout : (n m : ℕ) + → subComplexFam n (suc m) ≃ Pushout (subComplexα n m) fst + subComplexFam≃Pushout n m with (m ≟ᵗ n) | ((suc m) ≟ᵗ n) + ... | lt x | lt x₁ = snd C .snd .snd .snd m + ... | lt x | eq x₁ = snd C .snd .snd .snd m + ... | lt x | gt x₁ = ⊥.rec (¬squeeze (x , x₁)) + ... | eq x | lt x₁ = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x (<ᵗ-trans <ᵗsucm x₁))) + ... | eq x | eq x₁ = ⊥.rec (¬m<ᵗm (subst (m <ᵗ_) (x₁ ∙ sym x) <ᵗsucm)) + ... | eq x | gt x₁ = + compEquiv (pathToEquiv (λ i → fst C (x (~ i)))) + (isoToEquiv (PushoutEmptyFam (λ x → ¬Fin0 (fst x)) ¬Fin0)) + ... | gt x | lt x₁ = + ⊥.rec (¬squeeze (x₁ , <ᵗ-trans x (<ᵗ-trans <ᵗsucm <ᵗsucm))) + ... | gt x | eq x₁ = + ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) x₁ (<ᵗ-trans x <ᵗsucm))) + ... | gt x | gt x₁ = isoToEquiv (PushoutEmptyFam (λ x → ¬Fin0 (fst x)) ¬Fin0) + + subComplexYieldsCWskel : (n : ℕ) → yieldsCWskel (subComplexFam n) + fst (subComplexYieldsCWskel n) = subComplexCard n + fst (snd (subComplexYieldsCWskel n)) = subComplexα n + fst (snd (snd (subComplexYieldsCWskel n))) = subComplex₀-empty n + snd (snd (snd (subComplexYieldsCWskel n))) m = subComplexFam≃Pushout n m + + subComplex : (n : ℕ) → CWskel ℓ + fst (subComplex n) = subComplexFam n + snd (subComplex n) = subComplexYieldsCWskel n + + -- as a finite CWskel + subComplexFin : (m : ℕ) (n : Fin (suc m)) + → isEquiv (CW↪ (subComplexFam (fst n) , subComplexYieldsCWskel (fst n)) m) + subComplexFin m (n , r) with (m ≟ᵗ n) | (suc m ≟ᵗ n) + ... | lt x | p = ⊥.rec (¬squeeze (x , r)) + ... | eq x | lt x₁ = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x (<ᵗ-trans <ᵗsucm x₁))) + ... | eq x | eq x₁ = ⊥.rec (¬m<ᵗm (subst (m <ᵗ_) (x₁ ∙ sym x) <ᵗsucm)) + ... | eq x | gt x₁ = + subst isEquiv (funExt λ x → cong (help .fst) + (retEq (isoToEquiv (PushoutEmptyFam {A = Fin 0 × fst C m} + (λ x₃ → ¬Fin0 (fst x₃)) ¬Fin0 {f = snd} {g = fst})) x)) + (help .snd) + where + help : fst C m ≃ fst C n + help = invEquiv (pathToEquiv (λ i → fst C (x (~ i)))) + ... | gt x | lt x₁ = + ⊥.rec (¬squeeze (x₁ , <ᵗ-trans x (<ᵗ-trans <ᵗsucm <ᵗsucm))) + ... | gt x | eq x₁ = + ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) x₁ r)) + ... | gt x | gt x₁ = + subst isEquiv (funExt (retEq (isoToEquiv (PushoutEmptyFam {A = Fin 0 × fst C n} + (λ x₃ → ¬Fin0 (fst x₃)) ¬Fin0 {f = snd} {g = fst})))) + (idEquiv _ .snd) + + subComplexYieldsFinCWskel : (n : ℕ) → yieldsFinCWskel n (subComplexFam n) + fst (subComplexYieldsFinCWskel n) = subComplexYieldsCWskel n + snd (subComplexYieldsFinCWskel n) k = subComplexFin (k + n) (n , <ᵗ-+) + + finSubComplex : (n : ℕ) → finCWskel ℓ n + fst (finSubComplex n) = subComplexFam n + snd (finSubComplex n) = subComplexYieldsFinCWskel n + + complex≃subcomplex : (n : ℕ) (m : Fin (suc n)) + → fst C (fst m) ≃ subComplex n .fst (fst m) + complex≃subcomplex n (m , p) with (m ≟ᵗ n) + ... | lt x = idEquiv _ + ... | eq x = idEquiv _ + ... | gt x = ⊥.rec (¬squeeze (x , p)) + +realiseSubComplex : (n : ℕ) (C : CWskel ℓ) → Iso (fst C n) (realise (subComplex C n)) +realiseSubComplex n C = + compIso (equivToIso (complex≃subcomplex C n flast)) + (realiseFin n (finSubComplex C n)) + +-- Goal: Show that a cell complex C and its subcomplex Cₙ share +-- homology in low enough dimensions +module _ (C : CWskel ℓ) where + ChainSubComplex : (n : ℕ) → snd C .fst n ≡ snd (subComplex C (suc n)) .fst n + ChainSubComplex n with (n ≟ᵗ suc n) + ... | lt x = refl + ... | eq x = ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) (sym x) <ᵗsucm)) + ... | gt x = ⊥.rec (¬-suc-n<ᵗn x) + + ≤ChainSubComplex : (n : ℕ) (m : Fin n) + → snd C .fst (fst m) ≡ snd (subComplex C n) .fst (fst m) + ≤ChainSubComplex n (m , p) with (m ≟ᵗ n) + ... | lt x = refl + ... | eq x = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x p)) + ... | gt x = ⊥.rec (¬m<ᵗm (<ᵗ-trans x p)) + +subChainIso : (C : CWskel ℓ) (n : ℕ) (m : Fin n) + → AbGroupIso (CWskel-fields.ℤ[A_] C (fst m)) + (CWskel-fields.ℤ[A_] (subComplex C n) (fst m)) +subChainIso C n (m , p) with (m ≟ᵗ n) +... | lt x = idGroupIso +... | eq x = ⊥.rec (¬m<ᵗm (subst (m <ᵗ_) (sym x) p)) +... | gt x = ⊥.rec (¬m<ᵗm (<ᵗ-trans x p)) + +-- main result +subComplexHomology : (C : CWskel ℓ) (n m : ℕ) (p : suc (suc m) <ᵗ n) + → GroupIso (Hˢᵏᵉˡ C m) (Hˢᵏᵉˡ (subComplex C n) m) +subComplexHomology C n m p = + homologyIso m _ _ + (subChainIso C n (suc (suc m) , p)) + (subChainIso C n (suc m , <ᵗ-trans <ᵗsucm p)) + (subChainIso C n (m , <ᵗ-trans <ᵗsucm (<ᵗ-trans <ᵗsucm p))) + lem₁ + lem₂ + where + lem₁ : {q : _} {r : _} + → Iso.fun (subChainIso C n (m , q) .fst) ∘ ∂ C m .fst + ≡ ∂ (subComplex C n) m .fst + ∘ Iso.fun (subChainIso C n (suc m , r) .fst) + lem₁ {q} {r} with (m ≟ᵗ n) | (suc m ≟ᵗ n) | (suc (suc m) ≟ᵗ n) + ... | lt x | lt x₁ | lt x₂ = refl + ... | lt x | lt x₁ | eq x₂ = refl + ... | lt x | lt x₁ | gt x₂ = ⊥.rec (¬squeeze (x₁ , x₂)) + ... | lt x | eq x₁ | t = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x₁ r)) + ... | lt x | gt x₁ | t = ⊥.rec (¬m<ᵗm (<ᵗ-trans x₁ r)) + ... | eq x | s | t = ⊥.rec (¬-suc-n<ᵗn (subst (suc m <ᵗ_) (sym x) r)) + ... | gt x | s | t = ⊥.rec (¬-suc-n<ᵗn (<ᵗ-trans r x)) + + lem₂ : {q : suc m <ᵗ n} {r : (suc (suc m)) <ᵗ n} + → Iso.fun (subChainIso C n (suc m , q) .fst) + ∘ ∂ C (suc m) .fst + ≡ ∂ (subComplex C n) (suc m) .fst + ∘ Iso.fun (subChainIso C n (suc (suc m) , r) .fst) + lem₂ {q} with (suc m ≟ᵗ n) | (suc (suc m) ≟ᵗ n) | (suc (suc (suc m)) ≟ᵗ n) + ... | lt x | lt x₁ | lt x₂ = refl + ... | lt x | lt x₁ | eq x₂ = refl + ... | lt x | lt x₁ | gt x₂ = ⊥.rec (¬squeeze (x₁ , x₂)) + ... | lt x | eq x₁ | t = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x₁ p)) + ... | lt x | gt x₁ | t = ⊥.rec (¬m<ᵗm (<ᵗ-trans x₁ p)) + ... | eq x | s | t = ⊥.rec (¬m<ᵗm (subst (suc m <ᵗ_) (sym x) q)) + ... | gt x | s | t = ⊥.rec (¬-suc-n<ᵗn (<ᵗ-trans p x)) diff --git a/Cubical/Data/Fin/Inductive/Base.agda b/Cubical/Data/Fin/Inductive/Base.agda index 07a0d8e52b..6b2f918970 100644 --- a/Cubical/Data/Fin/Inductive/Base.agda +++ b/Cubical/Data/Fin/Inductive/Base.agda @@ -3,21 +3,14 @@ module Cubical.Data.Fin.Inductive.Base where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function - open import Cubical.Data.Nat open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Unit -open import Cubical.Data.Sigma -open import Cubical.Data.Nat.Order open import Cubical.Relation.Nullary --- open import Cubical.Algebra.AbGroup.Base using (move4) - -- inductive definition of < _<ᵗ_ : (n m : ℕ) → Type n <ᵗ zero = ⊥ @@ -60,6 +53,9 @@ snd (fsuc {n = suc n} (x , p)) = p <ᵗ-trans-suc {n = zero} {suc m} x = tt <ᵗ-trans-suc {n = suc n} {suc m} x = <ᵗ-trans-suc {n = n} x +¬-suc-n<ᵗn : {n : ℕ} → ¬ (suc n) <ᵗ n +¬-suc-n<ᵗn {suc n} = ¬-suc-n<ᵗn {n} + injectSuc : {n : ℕ} → Fin n → Fin (suc n) fst (injectSuc {n = n} (x , p)) = x snd (injectSuc {n = suc n} (x , p)) = <ᵗ-trans-suc {n = x} p diff --git a/Cubical/Data/Fin/Inductive/Properties.agda b/Cubical/Data/Fin/Inductive/Properties.agda index 8bb895f2ea..4d04631e49 100644 --- a/Cubical/Data/Fin/Inductive/Properties.agda +++ b/Cubical/Data/Fin/Inductive/Properties.agda @@ -70,6 +70,11 @@ elimFinβ {m = suc (suc m)} {A = A} max f = ¬m<ᵗm : {m : ℕ} → ¬ (m <ᵗ m) ¬m<ᵗm {m = suc m} p = ¬m<ᵗm p +<ᵗ-+ : {n k : ℕ} → n <ᵗ suc (k + n) +<ᵗ-+ {n = zero} {k} = tt +<ᵗ-+ {n = suc n} {k} = + subst (n <ᵗ_) (sym (+-suc k n)) (<ᵗ-+ {n = n} {k}) + ¬squeeze : {n m : ℕ} → ¬ ((n <ᵗ m) × (m <ᵗ suc n)) ¬squeeze {n = suc n} {suc m} = ¬squeeze {n = n} {m = m} diff --git a/Cubical/Data/Subcomplex.agda b/Cubical/Data/Subcomplex.agda new file mode 100644 index 0000000000..e69de29bb2 From 4ba3c8eb525864245fa8d2e532827e3bc0631b55 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 3 Mar 2024 20:55:52 +0100 Subject: [PATCH 30/73] stuff --- Cubical/Data/Fin/Inductive/Properties.agda | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/Cubical/Data/Fin/Inductive/Properties.agda b/Cubical/Data/Fin/Inductive/Properties.agda index 4d04631e49..3899afbc86 100644 --- a/Cubical/Data/Fin/Inductive/Properties.agda +++ b/Cubical/Data/Fin/Inductive/Properties.agda @@ -20,7 +20,8 @@ open import Cubical.Algebra.AbGroup.Base using (move4) -- inductive definition of < -fsuc-injectSuc : {m : ℕ} (n : Fin m) → injectSuc {n = suc m} (fsuc {n = m} n) ≡ fsuc (injectSuc n) +fsuc-injectSuc : {m : ℕ} (n : Fin m) + → injectSuc {n = suc m} (fsuc {n = m} n) ≡ fsuc (injectSuc n) fsuc-injectSuc {m = suc m} (x , p) = refl fzero : {m : ℕ} → Fin (suc m) @@ -57,12 +58,9 @@ elimFinβ {m = suc (suc m)} {A = A} max f = refl (elimFinβ {m = (suc m)} {A = λ x → A (fsuc x)} max _ .snd) --- ¬Fin0 : ¬ Fin 0 ¬Fin0 (x , ()) - --- <ᵗ-trans : {n m k : ℕ} → n <ᵗ m → m <ᵗ k → n <ᵗ k <ᵗ-trans {n = zero} {suc m} {suc k} _ _ = tt <ᵗ-trans {n = suc n} {suc m} {suc k} = <ᵗ-trans {n = n} {m} {k} From 6e259aff9109e5c799b9eb15ef45578ffb3965cb Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 Mar 2024 00:38:02 +0100 Subject: [PATCH 31/73] cleanup --- .../AbGroup/Instances/FreeAbGroup.agda | 3 +- Cubical/Algebra/AbGroup/Properties.agda | 3 +- Cubical/Algebra/ChainComplex/Base.agda | 23 +- Cubical/Algebra/ChainComplex/Finite.agda | 48 ++-- Cubical/Algebra/ChainComplex/Homology.agda | 194 +++------------- Cubical/Algebra/ChainComplex/Natindexed.agda | 70 ------ Cubical/Algebra/Group/ChainComplex/Base.agda | 0 Cubical/Algebra/Group/GroupPath.agda | 2 +- Cubical/Axiom/Choice.agda | 3 +- Cubical/{Data => }/CW/Approximation.agda | 10 +- Cubical/{Data => }/CW/Base.agda | 16 +- Cubical/{Data => }/CW/ChainComplex.agda | 9 +- Cubical/{Data => }/CW/Homology.agda | 15 +- Cubical/{Data => }/CW/Homotopy.agda | 22 +- Cubical/{Data => }/CW/Map.agda | 24 +- Cubical/{Data => }/CW/Properties.agda | 6 +- Cubical/{Data => }/CW/Subcomplex.agda | 10 +- Cubical/Data/CW/Base.agd | 0 Cubical/Data/Fin/Inductive.agda | 217 +----------------- Cubical/Data/Fin/Inductive/Base.agda | 10 +- Cubical/Data/Fin/Inductive/Properties.agda | 1 - Cubical/Data/FinSequence/Base.agda | 6 +- Cubical/Data/FinSequence/Properties.agda | 9 +- Cubical/Data/Subcomplex.agda | 0 Cubical/HITs/SequentialColimit/Base.agda | 103 +-------- .../HITs/SequentialColimit/Properties.agda | 100 +++++++- Cubical/HITs/Sn/Degree.agda | 4 +- Cubical/HITs/SphereBouquet/Base.agda | 3 +- Cubical/HITs/SphereBouquet/Degree.agda | 24 +- Cubical/HITs/SphereBouquet/Properties.agda | 27 +-- Cubical/Sequence.agda | 0 Cubical/Structures/Successor.agda | 61 ----- Cubical/ZCohomology/GroupStructure.agda | 3 +- 33 files changed, 215 insertions(+), 811 deletions(-) delete mode 100644 Cubical/Algebra/ChainComplex/Natindexed.agda delete mode 100644 Cubical/Algebra/Group/ChainComplex/Base.agda rename Cubical/{Data => }/CW/Approximation.agda (99%) rename Cubical/{Data => }/CW/Base.agda (92%) rename Cubical/{Data => }/CW/ChainComplex.agda (98%) rename Cubical/{Data => }/CW/Homology.agda (97%) rename Cubical/{Data => }/CW/Homotopy.agda (98%) rename Cubical/{Data => }/CW/Map.agda (96%) rename Cubical/{Data => }/CW/Properties.agda (99%) rename Cubical/{Data => }/CW/Subcomplex.agda (97%) delete mode 100644 Cubical/Data/CW/Base.agd delete mode 100644 Cubical/Data/Subcomplex.agda delete mode 100644 Cubical/Sequence.agda delete mode 100644 Cubical/Structures/Successor.agda diff --git a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda index 3a464dfbef..aab089129a 100644 --- a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda +++ b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda @@ -10,8 +10,7 @@ open import Cubical.Data.Sigma open import Cubical.Data.Nat hiding (_·_) renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) -open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Fin.Inductive open import Cubical.Data.Empty as ⊥ open import Cubical.HITs.FreeAbGroup diff --git a/Cubical/Algebra/AbGroup/Properties.agda b/Cubical/Algebra/AbGroup/Properties.agda index 0fe66202e9..61d4071a74 100644 --- a/Cubical/Algebra/AbGroup/Properties.agda +++ b/Cubical/Algebra/AbGroup/Properties.agda @@ -17,8 +17,7 @@ open import Cubical.HITs.SetQuotients as SQ hiding (_/_) open import Cubical.Data.Int using (ℤ ; pos ; negsuc) open import Cubical.Data.Nat hiding (_+_) open import Cubical.Data.Sigma -open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Fin.Inductive private variable ℓ ℓ' : Level diff --git a/Cubical/Algebra/ChainComplex/Base.agda b/Cubical/Algebra/ChainComplex/Base.agda index 7e7bbc8868..462d9f2f10 100644 --- a/Cubical/Algebra/ChainComplex/Base.agda +++ b/Cubical/Algebra/ChainComplex/Base.agda @@ -1,36 +1,17 @@ {-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.ChainComplex.Base where -{- - Defines groups and adds the smart constructors [makeGroup-right] and [makeGroup-left] - for constructing groups from less data than the standard [makeGroup] constructor. --} -open import Cubical.Foundations.Structure open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function -open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma +open import Cubical.Data.Nat open import Cubical.Algebra.Group -open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties -- TODO: why is this there and not exported by the Morphisms file? -open import Cubical.Algebra.Group.Subgroup -open import Cubical.Algebra.Group.QuotientGroup +open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.AbGroup -open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) -open import Cubical.HITs.SetQuotients.Properties as SQ -open import Cubical.HITs.PropositionalTruncation as PT - -open import Cubical.Structures.Successor -open import Cubical.Relation.Nullary - -open import Cubical.Data.Nat -open import Cubical.Data.Fin hiding (_/_) - - private variable ℓ ℓ' ℓ'' : Level diff --git a/Cubical/Algebra/ChainComplex/Finite.agda b/Cubical/Algebra/ChainComplex/Finite.agda index b172e145b6..011d37d2f0 100644 --- a/Cubical/Algebra/ChainComplex/Finite.agda +++ b/Cubical/Algebra/ChainComplex/Finite.agda @@ -11,34 +11,18 @@ definitions of and proof their induced behaviour on homology -} -open import Cubical.Foundations.Structure open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function -open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma - -open import Cubical.Algebra.Group -open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties -- TODO: why is this there and not exported by the Morphisms file? -open import Cubical.Algebra.Group.Subgroup -open import Cubical.Algebra.Group.QuotientGroup -open import Cubical.Algebra.AbGroup - -open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) -open import Cubical.HITs.SetQuotients.Properties as SQ -open import Cubical.HITs.PropositionalTruncation as PT - -open import Cubical.Structures.Successor -open import Cubical.Relation.Nullary - open import Cubical.Data.Nat open import Cubical.Data.Fin.Inductive.Base hiding (eq) -open import Cubical.Data.Fin.Inductive.Properties - open import Cubical.Algebra.ChainComplex.Base +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup + private variable ℓ ℓ' ℓ'' : Level @@ -48,20 +32,23 @@ module _ where (A : ChainComplex ℓ) (B : ChainComplex ℓ') : Type (ℓ-max ℓ ℓ') where open ChainComplex field - fchainmap : (i : Fin (suc m)) → AbGroupHom (chain A (fst i)) (chain B (fst i)) + fchainmap : (i : Fin (suc m)) + → AbGroupHom (chain A (fst i)) (chain B (fst i)) fbdrycomm : (i : Fin m) → compGroupHom (fchainmap (fsuc i)) (bdry B (fst i)) ≡ compGroupHom (bdry A (fst i)) (fchainmap (injectSuc i)) - record finChainHomotopy {ℓ : Level} (m : ℕ) {A : ChainComplex ℓ} {B : ChainComplex ℓ'} + record finChainHomotopy {ℓ : Level} (m : ℕ) + {A : ChainComplex ℓ} {B : ChainComplex ℓ'} (f g : finChainComplexMap m A B) : Type (ℓ-max ℓ' ℓ) where open ChainComplex open finChainComplexMap field - fhtpy : (i : Fin (suc m)) → AbGroupHom (chain A (fst i)) (chain B (suc (fst i))) + fhtpy : (i : Fin (suc m)) + → AbGroupHom (chain A (fst i)) (chain B (suc (fst i))) fbdryhtpy : (i : Fin m) → subtrGroupHom (chain A (suc (fst i))) (chain B (suc (fst i))) - (fchainmap f (fsuc i)) (fchainmap g (fsuc i)) -- (suc i)) + (fchainmap f (fsuc i)) (fchainmap g (fsuc i)) ≡ addGroupHom (chain A (suc (fst i))) (chain B (suc (fst i))) (compGroupHom (fhtpy (fsuc i)) (bdry B (suc (fst i)))) (compGroupHom (bdry A (fst i)) (fhtpy (injectSuc i))) @@ -74,10 +61,11 @@ module _ where → f ≡ g fchainmap (finChainComplexMap≡ p i) n = p n i fbdrycomm (finChainComplexMap≡ {A = A} {B = B} {f = f} {g = g} p i) n = - isProp→PathP {B = λ i → compGroupHom (p (fsuc n) i) (ChainComplex.bdry B (fst n)) - ≡ compGroupHom (ChainComplex.bdry A (fst n)) (p (injectSuc n) i)} - (λ i → isSetGroupHom _ _) - (fbdrycomm f n) (fbdrycomm g n) i + isProp→PathP {B = λ i + → compGroupHom (p (fsuc n) i) (ChainComplex.bdry B (fst n)) + ≡ compGroupHom (ChainComplex.bdry A (fst n)) (p (injectSuc n) i)} + (λ i → isSetGroupHom _ _) + (fbdrycomm f n) (fbdrycomm g n) i compFinChainMap : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {C : ChainComplex ℓ''} {m : ℕ} @@ -102,7 +90,8 @@ module _ where → finChainComplexMap m A B → Type (ℓ-max ℓ ℓ') isFinChainEquiv {m = m} f = ((n : Fin (suc m)) → isEquiv (fchainmap f n .fst)) - _≃⟨_⟩Chain_ : (A : ChainComplex ℓ) (m : ℕ) (B : ChainComplex ℓ') → Type (ℓ-max ℓ ℓ') + _≃⟨_⟩Chain_ : (A : ChainComplex ℓ) (m : ℕ) (B : ChainComplex ℓ') + → Type (ℓ-max ℓ ℓ') A ≃⟨ m ⟩Chain B = Σ[ f ∈ finChainComplexMap m A B ] (isFinChainEquiv f) idFinChainMap : (m : ℕ) (A : ChainComplex ℓ) → finChainComplexMap m A A @@ -113,7 +102,8 @@ module _ where invFinChainMap : {A : ChainComplex ℓ} {B : ChainComplex ℓ'} {m : ℕ} → (A ≃⟨ m ⟩Chain B) → finChainComplexMap m B A fchainmap (invFinChainMap {m = m} (ϕ , eq)) n = - GroupEquiv→GroupHom (invGroupEquiv ((fchainmap ϕ n .fst , eq n) , snd (fchainmap ϕ n))) + GroupEquiv→GroupHom + (invGroupEquiv ((fchainmap ϕ n .fst , eq n) , snd (fchainmap ϕ n))) fbdrycomm (invFinChainMap {B = B} {m = m} (ϕ' , eq)) n = Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x diff --git a/Cubical/Algebra/ChainComplex/Homology.agda b/Cubical/Algebra/ChainComplex/Homology.agda index f135a86a94..45c6cefd2b 100644 --- a/Cubical/Algebra/ChainComplex/Homology.agda +++ b/Cubical/Algebra/ChainComplex/Homology.agda @@ -1,41 +1,29 @@ {-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.ChainComplex.Homology where -{- - Defines groups and adds the smart constructors [makeGroup-right] and [makeGroup-left] - for constructing groups from less data than the standard [makeGroup] constructor. --} -open import Cubical.Foundations.Structure open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma -open import Cubical.Data.Nat.Order +open import Cubical.Data.Nat +open import Cubical.Data.Fin.Inductive open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties -- TODO: why is this there and not exported by the Morphisms file? +open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Subgroup open import Cubical.Algebra.Group.QuotientGroup open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.ChainComplex.Base +open import Cubical.Algebra.ChainComplex.Finite open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) open import Cubical.HITs.SetQuotients.Properties as SQ open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.Structures.Successor -open import Cubical.Relation.Nullary - -open import Cubical.Data.Nat -open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties - -open import Cubical.Algebra.ChainComplex.Base -open import Cubical.Algebra.ChainComplex.Finite - - private variable @@ -83,132 +71,6 @@ homology n C = ker∂n / img∂+1⊂ker∂n isNormalImSubGroup = isNormalIm ∂' (λ x y → kerGroup≡ ∂n (C1.+Comm (fst x) (fst y))) --- -- Induced maps on cohomology by finite chain complex maps/homotopies --- module _ where --- finChainComplexMap→HomologyMap : {C D : ChainComplex ℓ} (m : ℕ) --- → (ϕ : finChainComplexMap (suc (suc m)) C D) --- → (n : Fin (suc m)) --- → GroupHom (homology (fst n) C) (homology (fst n) D) --- finChainComplexMap→HomologyMap {C = C} {D} m mp (n , p) = main --- where --- ineq3 : suc n <ᵗ suc (suc (suc m)) --- ineq3 = <ᵗ-trans p <ᵗsucm - --- ineq4 : n <ᵗ suc (suc m) --- ineq4 = ineq3 - --- ϕ = fchainmap mp --- ϕcomm = fbdrycomm mp - --- lem : (k : ℕ) {p q : _} (f : fst (chain C k)) --- → fst (ϕ (k , p)) f ≡ fst (ϕ (k , q)) f --- lem k {p} {q} f i = fst (ϕ (k , pq i)) f --- where --- pq : p ≡ q --- pq = isProp<ᵗ _ _ - --- fun : homology n C .fst → homology n D .fst --- fun = SQ.elim (λ _ → squash/) f --- λ f g → PT.rec (GroupStr.is-set (homology n D .snd) _ _ ) --- λ r → eq/ _ _ ∣ (ϕ (suc (suc n) , p) .fst (fst r)) --- , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) --- ((funExt⁻ (cong fst (ϕcomm (suc n , _))) (fst r) --- ∙∙ cong (fst (ϕ (suc n , _))) (cong fst (snd r)) --- ∙∙ (IsGroupHom.pres· (snd (ϕ (suc n , _) )) _ _ --- ∙ cong₂ (AbGroupStr._+_ (snd (chain D (suc n)))) --- (lem (suc n) (fst f)) --- (IsGroupHom.presinv (snd (ϕ (suc n , _) )) _ --- ∙ cong (snd (chain D (suc n)) .AbGroupStr.-_) --- (lem (suc n) (fst g)))))) ∣₁ --- where --- f : _ → homology n D .fst --- f (a , b) = [ ϕ (suc n , ineq3) .fst a -- (ϕ {!n!} .fst a) --- , ((λ i → fst (ϕcomm (n , ineq3) i) a) --- ∙∙ cong (fst (ϕ (n , _))) b --- ∙∙ IsGroupHom.pres1 (snd (ϕ (n , _)))) ] - - --- main : GroupHom (homology n C) (homology n D) --- fst main = fun --- snd main = --- makeIsGroupHom --- (SQ.elimProp2 (λ _ _ → GroupStr.is-set (snd (homology n D)) _ _) --- λ a b → cong [_] --- (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) --- (IsGroupHom.pres· (snd (ϕ (suc n , _) )) _ _))) - --- finChainComplexMap→HomologyMapComp : {C D E : ChainComplex ℓ} {m : ℕ} --- → (ϕ : finChainComplexMap (suc (suc m)) C D) (ψ : finChainComplexMap (suc (suc m)) D E) --- → (n : Fin (suc m)) --- → finChainComplexMap→HomologyMap m (compFinChainMap ϕ ψ) n --- ≡ compGroupHom (finChainComplexMap→HomologyMap m ϕ n) --- (finChainComplexMap→HomologyMap m ψ n) --- finChainComplexMap→HomologyMapComp {E = E} _ _ n = --- Σ≡Prop (λ _ → isPropIsGroupHom _ _) --- (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (fst n) E)) _ _) --- λ _ → cong [_] --- (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain E (fst n))) _ _) refl))) - --- finChainComplexMap→HomologyMapId : {C : ChainComplex ℓ} {m : ℕ} (n : Fin (suc m)) --- → finChainComplexMap→HomologyMap m (idFinChainMap (suc (suc m)) C) n ≡ idGroupHom --- finChainComplexMap→HomologyMapId {C = C} n = --- Σ≡Prop (λ _ → isPropIsGroupHom _ _) --- (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (fst n) C)) _ _) --- λ _ → cong [_] --- (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C (fst n))) _ _) refl))) - --- finChainComplexEquiv→HomoglogyIso : {C D : ChainComplex ℓ} (m : ℕ) (f : C ≃⟨ (suc (suc m)) ⟩Chain D) --- → (n : Fin (suc m)) → GroupIso (homology (fst n) C) (homology (fst n) D) --- Iso.fun (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = --- finChainComplexMap→HomologyMap m f n .fst --- Iso.inv (fst (finChainComplexEquiv→HomoglogyIso m f n)) = --- finChainComplexMap→HomologyMap m (invFinChainMap f) n .fst --- Iso.rightInv (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = --- funExt⁻ (cong fst (sym (finChainComplexMap→HomologyMapComp --- (invFinChainMap (f , eqs)) f n)) --- ∙∙ cong (λ f → fst (finChainComplexMap→HomologyMap m f n)) --- (finChainComplexMap≡ λ r --- → Σ≡Prop (λ _ → isPropIsGroupHom _ _) --- (funExt (secEq (_ , eqs r)))) --- ∙∙ cong fst (finChainComplexMap→HomologyMapId n)) --- Iso.leftInv (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = --- funExt⁻ (cong fst (sym (finChainComplexMap→HomologyMapComp f --- (invFinChainMap (f , eqs)) n)) --- ∙∙ cong (λ f → fst (finChainComplexMap→HomologyMap m f n)) --- (finChainComplexMap≡ --- (λ n → Σ≡Prop (λ _ → isPropIsGroupHom _ _) --- (funExt (retEq (_ , eqs n))))) --- ∙∙ cong fst (finChainComplexMap→HomologyMapId n)) --- snd (finChainComplexEquiv→HomoglogyIso m (f , eqs) n) = --- finChainComplexMap→HomologyMap m f n .snd - - --- finChainHomotopy→HomologyPath : {A B : ChainComplex ℓ} {m : ℕ} --- (f g : finChainComplexMap (suc (suc m)) A B) --- → finChainHomotopy (suc (suc m)) f g --- → (n : Fin (suc m)) --- → finChainComplexMap→HomologyMap m f n --- ≡ finChainComplexMap→HomologyMap m g n --- finChainHomotopy→HomologyPath {A = A} {B = B} {m = m} f g ϕ n = --- Σ≡Prop (λ _ → isPropIsGroupHom _ _) --- (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd (homology (fst n) _)) _ _) --- λ {(a , p) → eq/ _ _ --- ∣ (finChainHomotopy.fhtpy ϕ (suc (fst n) , pf) .fst a) --- , (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain B (fst n))) _ _) --- (sym ((funExt⁻ (cong fst (finChainHomotopy.fbdryhtpy ϕ _)) a) --- ∙ cong₂ _+B_ refl --- (cong (fst (finChainHomotopy.fhtpy ϕ _)) p --- ∙ IsGroupHom.pres1 (snd (finChainHomotopy.fhtpy ϕ _))) --- ∙ AbGroupStr.+IdR (snd (chain B (suc (fst n)))) _))) ∣₁})) --- where --- open GroupTheory (AbGroup→Group (chain B (suc (fst n)))) --- pf : suc (fst n) <ᵗ suc (suc (suc m)) --- pf = <ᵗ-trans (snd n) <ᵗsucm - --- invB = GroupStr.inv (snd (AbGroup→Group (chain B (suc (fst n))))) --- _+B_ = AbGroupStr._+_ (snd (chain B (suc (fst n)))) - - -- Induced maps on cohomology by finite chain complex maps/homotopies module _ where finChainComplexMap→HomologyMap : {C D : ChainComplex ℓ} (m : ℕ) @@ -217,12 +79,6 @@ module _ where → GroupHom (homology (fst n) C) (homology (fst n) D) finChainComplexMap→HomologyMap {C = C} {D} m mp (n , p) = main where - ineq3 : suc n <ᵗ suc (suc m) - ineq3 = <ᵗ-trans p <ᵗsucm - - ineq4 : n <ᵗ suc m - ineq4 = ineq3 - ϕ = fchainmap mp ϕcomm = fbdrycomm mp @@ -236,20 +92,21 @@ module _ where fun : homology n C .fst → homology n D .fst fun = SQ.elim (λ _ → squash/) f λ f g → PT.rec (GroupStr.is-set (homology n D .snd) _ _ ) - λ r → eq/ _ _ ∣ (ϕ (suc (suc n) , p) .fst (fst r)) - , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) - ((funExt⁻ (cong fst (ϕcomm (suc n , _))) (fst r) - ∙∙ cong (fst (ϕ (suc n , _))) (cong fst (snd r)) - ∙∙ (IsGroupHom.pres· (snd (ϕ (suc n , _) )) _ _ - ∙ cong₂ (AbGroupStr._+_ (snd (chain D (suc n)))) - (lem (suc n) (fst f)) - (IsGroupHom.presinv (snd (ϕ (suc n , _) )) _ - ∙ cong (snd (chain D (suc n)) .AbGroupStr.-_) - (lem (suc n) (fst g)))))) ∣₁ + λ r → eq/ _ _ + ∣ (ϕ (suc (suc n) , p) .fst (fst r)) + , Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain D n)) _ _) + ((funExt⁻ (cong fst (ϕcomm (suc n , _))) (fst r) + ∙∙ cong (fst (ϕ (suc n , _))) (cong fst (snd r)) + ∙∙ (IsGroupHom.pres· (snd (ϕ (suc n , _) )) _ _ + ∙ cong₂ (AbGroupStr._+_ (snd (chain D (suc n)))) + (lem (suc n) (fst f)) + (IsGroupHom.presinv (snd (ϕ (suc n , _) )) _ + ∙ cong (snd (chain D (suc n)) .AbGroupStr.-_) + (lem (suc n) (fst g)))))) ∣₁ where f : _ → homology n D .fst - f (a , b) = [ ϕ (suc n , ineq3) .fst a -- (ϕ {!n!} .fst a) - , ((λ i → fst (ϕcomm (n , ineq3) i) a) + f (a , b) = [ ϕ (suc n , <ᵗ-trans p <ᵗsucm) .fst a + , ((λ i → fst (ϕcomm (n , <ᵗ-trans p <ᵗsucm) i) a) ∙∙ cong (fst (ϕ (n , _))) b ∙∙ IsGroupHom.pres1 (snd (ϕ (n , _)))) ] @@ -283,7 +140,8 @@ module _ where λ _ → cong [_] (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain C (fst n))) _ _) refl))) - finChainComplexEquiv→HomoglogyIso : {C D : ChainComplex ℓ} (m : ℕ) (f : C ≃⟨ (suc m) ⟩Chain D) + finChainComplexEquiv→HomoglogyIso : + {C D : ChainComplex ℓ} (m : ℕ) (f : C ≃⟨ (suc m) ⟩Chain D) → (n : Fin m) → GroupIso (homology (fst n) C) (homology (fst n) D) Iso.fun (fst (finChainComplexEquiv→HomoglogyIso m (f , eqs) n)) = finChainComplexMap→HomologyMap m f n .fst @@ -401,11 +259,11 @@ module _ where λ {(a , p) → eq/ _ _ ∣ (ChainHomotopy.htpy ϕ (suc n) .fst a) , (Σ≡Prop (λ _ → AbGroupStr.is-set (snd (chain B n)) _ _) - (sym ((funExt⁻ (cong fst (ChainHomotopy.bdryhtpy ϕ n)) a) - ∙ cong₂ _+B_ refl - (cong (fst (ChainHomotopy.htpy ϕ n)) p - ∙ IsGroupHom.pres1 (snd (ChainHomotopy.htpy ϕ n))) - ∙ AbGroupStr.+IdR (snd (chain B (suc n))) _))) ∣₁})) + (sym ((funExt⁻ (cong fst (ChainHomotopy.bdryhtpy ϕ n)) a) + ∙ cong₂ _+B_ refl + (cong (fst (ChainHomotopy.htpy ϕ n)) p + ∙ IsGroupHom.pres1 (snd (ChainHomotopy.htpy ϕ n))) + ∙ AbGroupStr.+IdR (snd (chain B (suc n))) _))) ∣₁})) where open GroupTheory (AbGroup→Group (chain B (suc n))) invB = GroupStr.inv (snd (AbGroup→Group (chain B (suc n)))) diff --git a/Cubical/Algebra/ChainComplex/Natindexed.agda b/Cubical/Algebra/ChainComplex/Natindexed.agda deleted file mode 100644 index 1a6e918064..0000000000 --- a/Cubical/Algebra/ChainComplex/Natindexed.agda +++ /dev/null @@ -1,70 +0,0 @@ -{-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.ChainComplex.Natindexed where - --- open import Cubical.Foundations.Prelude - --- open import Cubical.Data.Nat --- open import Cubical.Data.Fin --- open import Cubical.Data.Sigma --- open import Cubical.Data.Nat.Order --- open import Cubical.Data.Empty as ⊥ - - --- open import Cubical.Structures.Successor --- open import Cubical.Algebra.ChainComplex.Base - --- open import Cubical.Algebra.AbGroup --- open import Cubical.Algebra.Group.Morphisms --- open import Cubical.Algebra.Group.MorphismProperties - --- open import Cubical.Relation.Nullary - - --- private variable --- ℓ : Level - - --- -- Nat indexed chain complexes --- module CCℕ = CC ℕ+Inj - --- -- Restriction of ℕ indexed chain complexes to Fin n indexed ones --- module _ (n : ℕ) where --- open CC - --- -- Fin n indexed chain complexes --- module CCFin = CC (Fin+Inj n) - --- open ChainComplex --- open ChainComplexMap --- open ChainHomotopy - --- ChainComplex→finChainComplex : CCℕ.ChainComplex ℓ → CCFin.ChainComplex ℓ --- chain (ChainComplex→finChainComplex C) i p = chain C (fst i) ≠suc --- bdry (ChainComplex→finChainComplex C) (i , zero , q) p = ⊥.rec (p refl) --- bdry (ChainComplex→finChainComplex C) (i , suc diff , q) p = bdry C i ≠suc --- bdry²=0 (ChainComplex→finChainComplex C) (i , zero , q) p = ⊥.rec (p refl) --- bdry²=0 (ChainComplex→finChainComplex C) (i , suc zero , q) p = ⊥.rec (p refl) --- bdry²=0 (ChainComplex→finChainComplex C) (i , suc (suc diff) , q) p = bdry²=0 C i ≠suc - --- ChainComplexMap→finChainComplexMap : {C D : CCℕ.ChainComplex ℓ} --- → CCℕ.ChainComplexMap C D --- → CCFin.ChainComplexMap (ChainComplex→finChainComplex C) --- (ChainComplex→finChainComplex D) --- chainmap (ChainComplexMap→finChainComplexMap ϕ) (x , zero , q) p = ⊥.rec (p refl) --- chainmap (ChainComplexMap→finChainComplexMap ϕ) (x , suc diff , q) _ = chainmap ϕ x _ --- bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , zero , q) p = ⊥.rec (p refl) --- bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , suc zero , q) p = ⊥.rec (p refl) --- bdrycomm (ChainComplexMap→finChainComplexMap ϕ) (x , suc (suc diff) , p) q = bdrycomm ϕ x _ - --- chainHom→ : {C D : CCℕ.ChainComplex ℓ} (f g : CCℕ.ChainComplexMap C D) --- → CCℕ.ChainHomotopy f g --- → CCFin.ChainHomotopy --- (ChainComplexMap→finChainComplexMap f) --- (ChainComplexMap→finChainComplexMap g) --- htpy (chainHom→ f g h) (x , zero , q) p = ⊥.rec (p refl) --- htpy (chainHom→ f g h) (x , suc zero , q) p = ⊥.rec (p refl) --- htpy (chainHom→ f g h) (x , suc (suc diff) , q) p = htpy h x _ --- bdryhtpy (chainHom→ f g h) (x , zero , q) p = ⊥.rec (p refl) --- bdryhtpy (chainHom→ f g h) (x , suc zero , q) p = ⊥.rec (p refl) --- bdryhtpy (chainHom→ f g h) (x , suc (suc zero) , q) p = ⊥.rec (p refl) --- bdryhtpy (chainHom→ f g h) (x , suc (suc (suc diff)) , q) _ = bdryhtpy h x _ diff --git a/Cubical/Algebra/Group/ChainComplex/Base.agda b/Cubical/Algebra/Group/ChainComplex/Base.agda deleted file mode 100644 index e69de29bb2..0000000000 diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index f616401a9d..ff7ab15793 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -3,6 +3,7 @@ module Cubical.Algebra.Group.GroupPath where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path open import Cubical.Foundations.Function using (_∘_) open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv @@ -225,7 +226,6 @@ module _ {ℓ ℓ'} {A : Type ℓ} (Σ≡Prop (λ _ → isPropIsEquiv _) (funExt (G-coh-coh x y z))) - open import Cubical.Foundations.Path PropTrunc→Group : ∥ A ∥₁ → Group ℓ' PropTrunc→Group = rec→Gpd isGroupoidGroup G diff --git a/Cubical/Axiom/Choice.agda b/Cubical/Axiom/Choice.agda index f4d8321cdb..e48a844fb5 100644 --- a/Cubical/Axiom/Choice.agda +++ b/Cubical/Axiom/Choice.agda @@ -21,8 +21,7 @@ open import Cubical.Data.Nat open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Fin as FN -open import Cubical.Data.Fin.Inductive.Base as IndF -open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Fin.Inductive as IndF open import Cubical.Data.Nat.Order open import Cubical.HITs.Truncation as TR diff --git a/Cubical/Data/CW/Approximation.agda b/Cubical/CW/Approximation.agda similarity index 99% rename from Cubical/Data/CW/Approximation.agda rename to Cubical/CW/Approximation.agda index f5d9d36b4a..c8f185faeb 100644 --- a/Cubical/Data/CW/Approximation.agda +++ b/Cubical/CW/Approximation.agda @@ -3,7 +3,12 @@ -- cellular maps and homotopies -} -module Cubical.Data.CW.Approximation where +module Cubical.CW.Approximation where + +open import Cubical.CW.Base +open import Cubical.CW.Properties +open import Cubical.CW.Map +open import Cubical.CW.Homotopy open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -18,9 +23,6 @@ open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma open import Cubical.Data.Unit open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.CW -open import Cubical.Data.CW.Map -open import Cubical.Data.CW.Homotopy open import Cubical.Data.Sequence open import Cubical.Data.FinSequence diff --git a/Cubical/Data/CW/Base.agda b/Cubical/CW/Base.agda similarity index 92% rename from Cubical/Data/CW/Base.agda rename to Cubical/CW/Base.agda index da4add8e3e..8f94d5531e 100644 --- a/Cubical/Data/CW/Base.agda +++ b/Cubical/CW/Base.agda @@ -2,28 +2,19 @@ -- This file contains definition of CW complexes and skeleta. -module Cubical.Data.CW.Base where +module Cubical.CW.Base where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed open import Cubical.Foundations.Equiv -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Transport open import Cubical.Foundations.Function open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order -open import Cubical.Data.Unit open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma -open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sequence open import Cubical.Data.FinSequence - open import Cubical.HITs.Sn open import Cubical.HITs.Pushout open import Cubical.HITs.Susp @@ -34,11 +25,10 @@ open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup -open import Cubical.HITs.SequentialColimit -open Sequence - open import Cubical.Relation.Nullary +open Sequence + private diff --git a/Cubical/Data/CW/ChainComplex.agda b/Cubical/CW/ChainComplex.agda similarity index 98% rename from Cubical/Data/CW/ChainComplex.agda rename to Cubical/CW/ChainComplex.agda index 8d5a49b792..54d32db729 100644 --- a/Cubical/Data/CW/ChainComplex.agda +++ b/Cubical/CW/ChainComplex.agda @@ -1,5 +1,10 @@ {-# OPTIONS --safe --lossy-unification #-} +module Cubical.CW.ChainComplex where + +open import Cubical.CW.Base +open import Cubical.CW.Properties + open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function @@ -8,7 +13,6 @@ open import Cubical.Data.Nat open import Cubical.Data.Fin.Inductive.Base open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma -open import Cubical.Data.CW open import Cubical.HITs.Sn open import Cubical.HITs.Pushout @@ -20,7 +24,6 @@ open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.AbGroup open import Cubical.Algebra.ChainComplex -open import Cubical.Structures.Successor -- In this file, we associate to every CW complex its cellular chain complex @@ -28,8 +31,6 @@ open import Cubical.Structures.Successor -- The boundary map is defined through a bit of homotopical manipulation. -- The definition loosely follows May's Concise Course in Alg. Top. -module Cubical.Data.CW.ChainComplex where - module _ {ℓ} (C : CWskel ℓ) where open CWskel-fields C using (ℤ[A_]) diff --git a/Cubical/Data/CW/Homology.agda b/Cubical/CW/Homology.agda similarity index 97% rename from Cubical/Data/CW/Homology.agda rename to Cubical/CW/Homology.agda index 5ad06b034c..c216a36d01 100644 --- a/Cubical/Data/CW/Homology.agda +++ b/Cubical/CW/Homology.agda @@ -5,7 +5,14 @@ This file contains: 2. Construction of cellular homology, including funtoriality -} -module Cubical.Data.CW.Homology where +module Cubical.CW.Homology where + +open import Cubical.CW.Base +open import Cubical.CW.Properties +open import Cubical.CW.Map +open import Cubical.CW.Homotopy +open import Cubical.CW.ChainComplex +open import Cubical.CW.Approximation open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -16,12 +23,6 @@ open import Cubical.Foundations.Function open import Cubical.Data.Nat open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.CW -open import Cubical.Data.CW.Map -open import Cubical.Data.CW.Homotopy -open import Cubical.Data.CW.ChainComplex -open import Cubical.Data.CW.Approximation - open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties diff --git a/Cubical/Data/CW/Homotopy.agda b/Cubical/CW/Homotopy.agda similarity index 98% rename from Cubical/Data/CW/Homotopy.agda rename to Cubical/CW/Homotopy.agda index 05468ce9a8..fdcd2874ef 100644 --- a/Cubical/Data/CW/Homotopy.agda +++ b/Cubical/CW/Homotopy.agda @@ -6,41 +6,31 @@ This file contains 2. A proof that (finite) cellular homotopies induce (finite) chain complex maps -} -module Cubical.Data.CW.Homotopy where +module Cubical.CW.Homotopy where + +open import Cubical.CW.Base +open import Cubical.CW.Properties +open import Cubical.CW.Map +open import Cubical.CW.ChainComplex open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) -open import Cubical.Data.Nat.Order open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) open import Cubical.Data.Fin.Inductive.Base open import Cubical.Data.Fin.Inductive.Properties -open import Cubical.Data.Sigma -open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Bool hiding (_≟_ ; isProp≤) -open import Cubical.Data.CW -open import Cubical.Data.CW.Map open import Cubical.Data.Sequence open import Cubical.Data.FinSequence -open import Cubical.Data.CW.ChainComplex - -open import Cubical.HITs.Sn open import Cubical.HITs.Sn.Degree open import Cubical.HITs.Pushout open import Cubical.HITs.Susp -open import Cubical.HITs.SequentialColimit open import Cubical.HITs.SphereBouquet open import Cubical.HITs.SphereBouquet.Degree -open import Cubical.Relation.Nullary - -open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Group.Morphisms diff --git a/Cubical/Data/CW/Map.agda b/Cubical/CW/Map.agda similarity index 96% rename from Cubical/Data/CW/Map.agda rename to Cubical/CW/Map.agda index d0747afbf2..3f7c35fe8a 100644 --- a/Cubical/Data/CW/Map.agda +++ b/Cubical/CW/Map.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} {-This file contains: @@ -8,40 +8,32 @@ -} -module Cubical.Data.CW.Map where +module Cubical.CW.Map where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed open import Cubical.Foundations.Equiv -open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Transport open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) -open import Cubical.Data.Bool hiding (_≟_) -open import Cubical.Data.Nat.Order -open import Cubical.Data.Unit +open import Cubical.Data.Bool open import Cubical.Data.Fin.Inductive.Base open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sequence open import Cubical.Data.FinSequence -open import Cubical.Data.CW.Base -open import Cubical.Data.CW.Properties -open import Cubical.Data.CW.ChainComplex -open import Cubical.HITs.Sn +open import Cubical.CW.Base +open import Cubical.CW.Properties +open import Cubical.CW.ChainComplex + open import Cubical.HITs.Sn.Degree open import Cubical.HITs.Pushout open import Cubical.HITs.Susp -open import Cubical.HITs.SequentialColimit open import Cubical.HITs.SphereBouquet -open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.SphereBouquet.Degree open import Cubical.HITs.SequentialColimit @@ -90,7 +82,7 @@ module _ (m : ℕ) where open FinSequenceMap finCellMap→FinSeqColim : (C : CWskel ℓ) (D : CWskel ℓ') - {m : ℕ} → finCellMap m C D → FinSeqColim m (realiseSeq C) → FinSeqColim m (realiseSeq D) + {m : ℕ} → finCellMap m C D → FinSeqColim m (realiseSeq C) → FinSeqColim m (realiseSeq D) finCellMap→FinSeqColim C D {m = m} f (fincl n x) = fincl n (fmap f n x) finCellMap→FinSeqColim C D {m = m} f (fpush n x i) = (fpush n (fmap f (injectSuc n) x) ∙ cong (fincl (fsuc n)) (fcomm f n x)) i diff --git a/Cubical/Data/CW/Properties.agda b/Cubical/CW/Properties.agda similarity index 99% rename from Cubical/Data/CW/Properties.agda rename to Cubical/CW/Properties.agda index 7034cb2b5c..4e481c84a8 100644 --- a/Cubical/Data/CW/Properties.agda +++ b/Cubical/CW/Properties.agda @@ -4,7 +4,9 @@ complexes/skeleta. -} -module Cubical.Data.CW.Properties where +module Cubical.CW.Properties where + +open import Cubical.CW.Base open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed @@ -23,7 +25,6 @@ open import Cubical.Data.Fin.Inductive.Base open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.CW.Base open import Cubical.Data.Sequence open import Cubical.HITs.Sn @@ -38,6 +39,7 @@ open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup open import Cubical.Axiom.Choice + open import Cubical.Relation.Nullary open import Cubical.Homotopy.Connected diff --git a/Cubical/Data/CW/Subcomplex.agda b/Cubical/CW/Subcomplex.agda similarity index 97% rename from Cubical/Data/CW/Subcomplex.agda rename to Cubical/CW/Subcomplex.agda index 5e0de1277d..250298b1f2 100644 --- a/Cubical/Data/CW/Subcomplex.agda +++ b/Cubical/CW/Subcomplex.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Data.CW.Subcomplex where +module Cubical.CW.Subcomplex where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed @@ -15,15 +15,15 @@ open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.CW -open import Cubical.Data.CW.ChainComplex -open import Cubical.Data.CW.Approximation +open import Cubical.CW.Base +open import Cubical.CW.Properties +open import Cubical.CW.ChainComplex +open import Cubical.CW.Approximation open import Cubical.HITs.Sn open import Cubical.HITs.Pushout open import Cubical.HITs.PropositionalTruncation as PT - open import Cubical.Relation.Nullary open import Cubical.Algebra.AbGroup diff --git a/Cubical/Data/CW/Base.agd b/Cubical/Data/CW/Base.agd deleted file mode 100644 index e69de29bb2..0000000000 diff --git a/Cubical/Data/Fin/Inductive.agda b/Cubical/Data/Fin/Inductive.agda index 52bcc04da9..47e330a330 100644 --- a/Cubical/Data/Fin/Inductive.agda +++ b/Cubical/Data/Fin/Inductive.agda @@ -1,217 +1,6 @@ -{-# OPTIONS --safe --lossy-unification #-} +{-# OPTIONS --safe #-} module Cubical.Data.Fin.Inductive where --- open import Cubical.Foundations.Prelude --- open import Cubical.Foundations.Isomorphism --- open import Cubical.Foundations.HLevels --- open import Cubical.Foundations.Function - - --- open import Cubical.Data.Nat --- open import Cubical.Data.Empty as ⊥ --- open import Cubical.Data.Unit --- open import Cubical.Data.Sigma --- open import Cubical.Data.Nat.Order - --- open import Cubical.Relation.Nullary - --- -- open import Cubical.Algebra.AbGroup.Base using (move4) - --- -- inductive definition of < --- _<ᵗ_ : (n m : ℕ) → Type --- n <ᵗ zero = ⊥ --- zero <ᵗ suc m = Unit --- suc n <ᵗ suc m = n <ᵗ m - --- data Trichotomyᵗ (m n : ℕ) : Type₀ where --- lt : m <ᵗ n → Trichotomyᵗ m n --- eq : m ≡ n → Trichotomyᵗ m n --- gt : n <ᵗ m → Trichotomyᵗ m n - --- Trichotomyᵗ-suc : {n m : ℕ} → Trichotomyᵗ n m → Trichotomyᵗ (suc n) (suc m) --- Trichotomyᵗ-suc (lt x) = lt x --- Trichotomyᵗ-suc (eq x) = eq (cong suc x) --- Trichotomyᵗ-suc (gt x) = gt x - --- _≟ᵗ_ : ∀ m n → Trichotomyᵗ m n --- zero ≟ᵗ zero = eq refl --- zero ≟ᵗ suc n = lt tt --- suc m ≟ᵗ zero = gt tt --- suc m ≟ᵗ suc n = Trichotomyᵗ-suc (m ≟ᵗ n) - --- isProp<ᵗ : {n m : ℕ} → isProp (n <ᵗ m) --- isProp<ᵗ {n = n} {zero} = isProp⊥ --- isProp<ᵗ {n = zero} {suc m} = isPropUnit --- isProp<ᵗ {n = suc n} {suc m} = isProp<ᵗ {n = n} {m = m} --- -- Fin defined using <ᵗ --- Fin : (n : ℕ) → Type --- Fin n = Σ[ m ∈ ℕ ] (m <ᵗ n) - --- fsuc : {n : ℕ} → Fin n → Fin (suc n) --- fst (fsuc {n = n} (x , p)) = suc x --- snd (fsuc {n = suc n} (x , p)) = p - --- <ᵗ-trans-suc : {n m : ℕ} → n <ᵗ m → n <ᵗ suc m --- <ᵗ-trans-suc {n = zero} {suc m} x = tt --- <ᵗ-trans-suc {n = suc n} {suc m} x = <ᵗ-trans-suc {n = n} x - --- injectSuc : {n : ℕ} → Fin n → Fin (suc n) --- fst (injectSuc {n = n} (x , p)) = x --- snd (injectSuc {n = suc n} (x , p)) = <ᵗ-trans-suc {n = x} p - --- fsuc-injectSuc : {m : ℕ} (n : Fin m) → injectSuc {n = suc m} (fsuc {n = m} n) ≡ fsuc (injectSuc n) --- fsuc-injectSuc {m = suc m} (x , p) = refl - --- <ᵗsucm : {m : ℕ} → m <ᵗ suc m --- <ᵗsucm {m = zero} = tt --- <ᵗsucm {m = suc m} = <ᵗsucm {m = m} - --- flast : {m : ℕ} → Fin (suc m) --- fst (flast {m = m}) = m --- snd (flast {m = m}) = <ᵗsucm {m = m} - --- fzero : {m : ℕ} → Fin (suc m) --- fzero = 0 , tt - --- elimFin : ∀ {ℓ} {m : ℕ} {A : Fin (suc m) → Type ℓ} --- (max : A flast) --- (f : (x : Fin m) → A (injectSuc x)) --- → (x : _) → A x --- elimFin {m = zero} {A = A} max f (zero , p) = max --- elimFin {m = suc m} {A = A} max f (zero , p) = f (zero , tt) --- elimFin {m = suc zero} {A = A} max f (suc zero , p) = max --- elimFin {m = suc (suc m)} {A = A} max f (suc x , p) = --- elimFin {m = suc m} {A = λ x → A (fsuc x)} max (λ t → f (fsuc t)) (x , p) - --- elimFin-alt : ∀ {ℓ} {m : ℕ} {A : Fin (suc m) → Type ℓ} --- (max : A fzero) --- (f : (x : Fin m) → A (fsuc x)) --- → (x : _) → A x --- elimFin-alt {m = zero} max f (zero , p) = max --- elimFin-alt {m = suc m} max f (zero , p) = max --- elimFin-alt {m = suc m} max f (suc x , p) = f (x , p) - --- elimFinβ : ∀ {ℓ} {m : ℕ} {A : Fin (suc m) → Type ℓ} --- (max : A flast) --- (f : (x : Fin m) → A (injectSuc x)) --- → ((elimFin {A = A} max f flast ≡ max)) --- × ((x : Fin m) → elimFin {A = A} max f (injectSuc x) ≡ f x) --- elimFinβ {m = zero} {A = A} max f = refl , λ {()} --- elimFinβ {m = suc zero} {A = A} max f = refl , λ {(zero , p) → refl} --- elimFinβ {m = suc (suc m)} {A = A} max f = --- elimFinβ {m = (suc m)} {A = λ x → A (fsuc x)} max _ .fst --- , elimFin-alt {m = (suc m)} {A = λ x → elimFin max f (injectSuc {n = suc (suc m)} x) ≡ f x} --- refl --- (elimFinβ {m = (suc m)} {A = λ x → A (fsuc x)} max _ .snd) - --- -- --- ¬Fin0 : ¬ Fin 0 --- ¬Fin0 (x , ()) - --- -- Sums --- sumFinGen : ∀ {ℓ} {A : Type ℓ} {n : ℕ} (_+_ : A → A → A) (0A : A) (f : Fin n → A) → A --- sumFinGen {n = zero} _+_ 0A f = 0A --- sumFinGen {n = suc n} _+_ 0A f = f flast + sumFinGen _+_ 0A (f ∘ injectSuc) - - --- -- --- <ᵗ-trans : {n m k : ℕ} → n <ᵗ m → m <ᵗ k → n <ᵗ k --- <ᵗ-trans {n = zero} {suc m} {suc k} _ _ = tt --- <ᵗ-trans {n = suc n} {suc m} {suc k} = <ᵗ-trans {n = n} {m} {k} - --- ¬m<ᵗm : {m : ℕ} → ¬ (m <ᵗ m) --- ¬m<ᵗm {m = suc m} p = ¬m<ᵗm p - --- ¬squeeze : {n m : ℕ} → ¬ ((n <ᵗ m) × (m <ᵗ suc n)) --- ¬squeeze {n = suc n} {suc m} = ¬squeeze {n = n} {m = m} - --- -- properties of finite sums --- module _ {ℓ : Level} {A : Type ℓ} (_+A_ : A → A → A) (0A : A) --- (rUnit : (x : _) → x +A 0A ≡ x) --- where --- sumFinGen0 : (n : ℕ) (f : Fin n → A) --- → ((x : _) → f x ≡ 0A) --- → sumFinGen _+A_ 0A f --- ≡ 0A --- sumFinGen0 zero f ind = refl --- sumFinGen0 (suc n) f ind = --- cong₂ _+A_ --- (ind flast) --- (sumFinGen0 n (f ∘ injectSuc) (λ x → ind (injectSuc x))) ∙ rUnit 0A - --- module _ (comm : (x y : A) → x +A y ≡ y +A x) where --- private --- lUnitA : (x : _) → 0A +A x ≡ x --- lUnitA x = comm _ _ ∙ rUnit x - --- sumFin-choose : --- {n : ℕ} (f : Fin n → A) --- → (a : A) (x : Fin n) --- → (f x ≡ a) --- → ((x' : Fin n) → ¬ (x' ≡ x) → f x' ≡ 0A) --- → sumFinGen {n = n} _+A_ 0A f ≡ a --- sumFin-choose {zero} f a x p t = ⊥.rec (¬Fin0 x) --- sumFin-choose {suc n} f a (x , q) p t with (n ≟ᵗ x) --- ... | lt x₁ = ⊥.rec (¬squeeze (q , x₁)) --- ... | eq x₁ = --- cong (f flast +A_) (sumFinGen0 n _ λ h → t _ --- λ q → ¬m<ᵗm (subst (_<ᵗ n) (cong fst q ∙ sym x₁) (snd h))) --- ∙ rUnit _ ∙ cong f (sym x=flast) ∙ p --- where --- x=flast : (x , q) ≡ flast --- x=flast = Σ≡Prop (λ _ → isProp<ᵗ) (sym x₁) --- ... | gt x₁ = --- cong₂ _+A_ --- (t flast (λ p → ¬m<ᵗm (subst (_<ᵗ n) (sym (cong fst p)) x₁))) --- refl --- ∙ lUnitA _ --- ∙ sumFin-choose {n = n} (f ∘ injectSuc) a (x , x₁) --- (cong f (Σ≡Prop (λ _ → isProp<ᵗ) refl) ∙ p) --- λ a r → t (injectSuc a) λ d → r (Σ≡Prop (λ _ → isProp<ᵗ) --- (cong fst d)) - --- module _ (assocA : (x y z : A) → x +A (y +A z) ≡ (x +A y) +A z) where --- sumFinGenHom : (n : ℕ) (f g : Fin n → A) --- → sumFinGen _+A_ 0A (λ x → f x +A g x) --- ≡ (sumFinGen _+A_ 0A f +A sumFinGen _+A_ 0A g) --- sumFinGenHom zero f g = sym (rUnit 0A) --- sumFinGenHom (suc n) f g = --- cong ((f flast +A g flast) +A_) (sumFinGenHom n (f ∘ injectSuc) (g ∘ injectSuc)) --- ∙ move4 (f flast) (g flast) --- (sumFinGen _+A_ 0A (λ x → f (injectSuc x))) --- (sumFinGen _+A_ 0A (λ x → g (injectSuc x))) --- _+A_ assocA comm - --- sumFinGenId : ∀ {ℓ} {A : Type ℓ} {n : ℕ} (_+_ : A → A → A) (0A : A) --- (f g : Fin n → A) → f ≡ g → sumFinGen _+_ 0A f ≡ sumFinGen _+_ 0A g --- sumFinGenId _+_ 0A f g p i = sumFinGen _+_ 0A (p i) - - --- open import Cubical.Data.Fin.Base renaming (Fin to Fin*) hiding (¬Fin0) --- open import Cubical.Data.Fin.Properties renaming (isSetFin to isSetFin*) - --- <ᵗ→< : {n m : ℕ} → n <ᵗ m → n < m --- <ᵗ→< {n = zero} {suc m} p = m , +-comm m 1 --- <ᵗ→< {n = suc n} {suc m} p = suc-≤-suc (<ᵗ→< {n = n} {m = m} p) - --- <→<ᵗ : {n m : ℕ} → n < m → n <ᵗ m --- <→<ᵗ {n = zero} {m = zero} x = --- snotz (sym (+-suc (fst x) 0) ∙ snd x) --- <→<ᵗ {n = zero} {m = suc m} _ = tt --- <→<ᵗ {n = suc n} {m = zero} x = --- snotz (sym (+-suc (fst x) (suc n)) ∙ snd x) --- <→<ᵗ {n = suc n} {m = suc m} p = <→<ᵗ {n = n} {m = m} (pred-≤-pred p) - --- Iso-Fin-InductiveFin : (m : ℕ) → Iso (Fin* m) (Fin m) --- Iso.fun (Iso-Fin-InductiveFin m) (x , p) = x , <→<ᵗ p --- Iso.inv (Iso-Fin-InductiveFin m) (x , p) = x , <ᵗ→< p --- Iso.rightInv (Iso-Fin-InductiveFin m) (x , p) = --- Σ≡Prop (λ w → isProp<ᵗ {n = w} {m}) refl --- Iso.leftInv (Iso-Fin-InductiveFin m) _ = Σ≡Prop (λ _ → isProp≤) refl - - --- isSetFin : {n : ℕ} → isSet (Fin n) --- isSetFin {n = n} = --- isOfHLevelRetractFromIso 2 (invIso (Iso-Fin-InductiveFin n)) --- isSetFin* +open import Cubical.Data.Fin.Inductive.Base public +open import Cubical.Data.Fin.Inductive.Properties public diff --git a/Cubical/Data/Fin/Inductive/Base.agda b/Cubical/Data/Fin/Inductive/Base.agda index 6b2f918970..5d3e588cd6 100644 --- a/Cubical/Data/Fin/Inductive/Base.agda +++ b/Cubical/Data/Fin/Inductive/Base.agda @@ -1,5 +1,8 @@ {-# OPTIONS --safe --lossy-unification #-} +-- Alternative definition of < on ℕ and thus an alternative definition +-- of Fin n + module Cubical.Data.Fin.Inductive.Base where open import Cubical.Foundations.Prelude @@ -22,7 +25,8 @@ data Trichotomyᵗ (m n : ℕ) : Type₀ where eq : m ≡ n → Trichotomyᵗ m n gt : n <ᵗ m → Trichotomyᵗ m n -Trichotomyᵗ-suc : {n m : ℕ} → Trichotomyᵗ n m → Trichotomyᵗ (suc n) (suc m) +Trichotomyᵗ-suc : {n m : ℕ} → Trichotomyᵗ n m + → Trichotomyᵗ (suc n) (suc m) Trichotomyᵗ-suc (lt x) = lt x Trichotomyᵗ-suc (eq x) = eq (cong suc x) Trichotomyᵗ-suc (gt x) = gt x @@ -64,8 +68,8 @@ flast : {m : ℕ} → Fin (suc m) fst (flast {m = m}) = m snd (flast {m = m}) = <ᵗsucm {m = m} - -- Sums -sumFinGen : ∀ {ℓ} {A : Type ℓ} {n : ℕ} (_+_ : A → A → A) (0A : A) (f : Fin n → A) → A +sumFinGen : ∀ {ℓ} {A : Type ℓ} {n : ℕ} + (_+_ : A → A → A) (0A : A) (f : Fin n → A) → A sumFinGen {n = zero} _+_ 0A f = 0A sumFinGen {n = suc n} _+_ 0A f = f flast + sumFinGen _+_ 0A (f ∘ injectSuc) diff --git a/Cubical/Data/Fin/Inductive/Properties.agda b/Cubical/Data/Fin/Inductive/Properties.agda index 3899afbc86..a448c2fbd9 100644 --- a/Cubical/Data/Fin/Inductive/Properties.agda +++ b/Cubical/Data/Fin/Inductive/Properties.agda @@ -160,7 +160,6 @@ Iso.rightInv (Iso-Fin-InductiveFin m) (x , p) = Σ≡Prop (λ w → isProp<ᵗ {n = w} {m}) refl Iso.leftInv (Iso-Fin-InductiveFin m) _ = Σ≡Prop (λ _ → isProp≤) refl - isSetFin : {n : ℕ} → isSet (Fin n) isSetFin {n = n} = isOfHLevelRetractFromIso 2 (invIso (Iso-Fin-InductiveFin n)) diff --git a/Cubical/Data/FinSequence/Base.agda b/Cubical/Data/FinSequence/Base.agda index 9fd4f69f0f..9d45212744 100644 --- a/Cubical/Data/FinSequence/Base.agda +++ b/Cubical/Data/FinSequence/Base.agda @@ -2,11 +2,8 @@ module Cubical.Data.FinSequence.Base where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv open import Cubical.Data.Nat open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties - private variable @@ -17,7 +14,8 @@ record FinSequence (m : ℕ) (ℓ : Level) : Type (ℓ-suc ℓ) where fobj : Fin (suc m) → Type ℓ fmap : {n : Fin m} → fobj (injectSuc n) → fobj (fsuc n) -record FinSequenceMap {m : ℕ} (C : FinSequence m ℓ) (D : FinSequence m ℓ') : Type (ℓ-max ℓ ℓ') where +record FinSequenceMap {m : ℕ} + (C : FinSequence m ℓ) (D : FinSequence m ℓ') : Type (ℓ-max ℓ ℓ') where field fmap : (n : Fin (suc m)) → FinSequence.fobj C n → FinSequence.fobj D n fcomm : (n : Fin m) (x : FinSequence.fobj C (injectSuc n)) diff --git a/Cubical/Data/FinSequence/Properties.agda b/Cubical/Data/FinSequence/Properties.agda index b361fc1f6d..b47312df77 100644 --- a/Cubical/Data/FinSequence/Properties.agda +++ b/Cubical/Data/FinSequence/Properties.agda @@ -2,10 +2,8 @@ module Cubical.Data.FinSequence.Properties where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv open import Cubical.Data.Nat open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.FinSequence.Base open import Cubical.Data.Sequence @@ -22,19 +20,20 @@ idFinSequenceMap : (m : ℕ) (C : FinSequence m ℓ) → FinSequenceMap C C fmap (idFinSequenceMap m C) p x = x fcomm (idFinSequenceMap m C) p _ = refl -composeFinSequenceMap : (m : ℕ) {C : FinSequence m ℓ} {D : FinSequence m ℓ'} {E : FinSequence m ℓ''} +composeFinSequenceMap : (m : ℕ) {C : FinSequence m ℓ} {D : FinSequence m ℓ'} + {E : FinSequence m ℓ''} (g : FinSequenceMap D E) (f : FinSequenceMap C D) → FinSequenceMap C E fmap (composeFinSequenceMap m g f) n x = fmap g n (fmap f n x) fcomm (composeFinSequenceMap m g f) n x = fcomm g n _ ∙ cong (fmap g (fsuc n)) (fcomm f n x) -Sequence→FinSequence : (m : ℕ) → Sequence ℓ → FinSequence m ℓ +Sequence→FinSequence : (m : ℕ) → Sequence ℓ → FinSequence m ℓ fobj (Sequence→FinSequence m C) x = obj C (fst x) fmap (Sequence→FinSequence m C) x = map C x SequenceMap→FinSequenceMap : (m : ℕ) (C D : Sequence ℓ) → SequenceMap C D - → FinSequenceMap (Sequence→FinSequence m C) (Sequence→FinSequence m D) + → FinSequenceMap (Sequence→FinSequence m C) (Sequence→FinSequence m D) fmap (SequenceMap→FinSequenceMap m C D ϕ) t = SequenceMap.map ϕ (fst t) fcomm (SequenceMap→FinSequenceMap m C D ϕ) t = SequenceMap.comm ϕ (fst t) diff --git a/Cubical/Data/Subcomplex.agda b/Cubical/Data/Subcomplex.agda deleted file mode 100644 index e69de29bb2..0000000000 diff --git a/Cubical/HITs/SequentialColimit/Base.agda b/Cubical/HITs/SequentialColimit/Base.agda index 4b73ee3f8e..0dcfad0daa 100644 --- a/Cubical/HITs/SequentialColimit/Base.agda +++ b/Cubical/HITs/SequentialColimit/Base.agda @@ -7,8 +7,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Data.Nat open import Cubical.Data.Sequence open import Cubical.Data.FinSequence -open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Fin.Inductive private @@ -35,103 +34,3 @@ realiseSequenceMap : {C : Sequence ℓ} {D : Sequence ℓ'} realiseSequenceMap record { map = map ; comm = comm } (incl x) = incl (map _ x) realiseSequenceMap record { map = map ; comm = comm } (push {n = n} x i) = (push (map n x) ∙ λ i → incl (comm n x i)) i - - - -open import Cubical.Data.Unit -open import Cubical.Data.Sigma -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence - - - - -open import Cubical.Foundations.Path -open import Cubical.Foundations.GroupoidLaws - - - -open import Cubical.Foundations.Isomorphism - -open import Cubical.Data.Nat.Order -open import Cubical.Data.Empty as ⊥ - -open import Cubical.Data.Sum hiding (map) - - --- colim (X₀ → ... → Xₙ) ≃ colim (X₁ → ... → Xₙ) - -module _ (X : Sequence ℓ) where - ShiftSeq : Sequence ℓ - obj ShiftSeq m = obj X (suc m) - map ShiftSeq = map X - - Iso-FinSeqColim→FinSeqColim↑ : (m : ℕ) - → Iso (FinSeqColim (suc m) X) (FinSeqColim m ShiftSeq) - Iso-FinSeqColim→FinSeqColim↑ m = iso (G m) (F m) (F→G→F m) (G→F→G m) - where - F : (m : ℕ) → FinSeqColim m ShiftSeq → FinSeqColim (suc m) X - F m (fincl n x) = fincl (fsuc n) x - F (suc m) (fpush (n , p) x i) = fpush (suc n , p) x i - - G : (m : ℕ) → FinSeqColim (suc m) X → FinSeqColim m ShiftSeq - G m (fincl (zero , p) x) = fincl (zero , p) (map X x) - G m (fincl (suc n , p) x) = fincl (n , p) x - G m (fpush (zero , p) x i) = fincl (zero , p) (map X x) - G (suc m) (fpush (suc n , p) x i) = fpush (n , p) x i - - F→G→F : (m : ℕ) → (x : FinSeqColim m ShiftSeq) → G m (F m x) ≡ x - F→G→F m (fincl n x) = refl - F→G→F (suc m) (fpush n x i) = refl - - G→F→G : (m : ℕ) → (x : FinSeqColim (suc m) X) → F m (G m x) ≡ x - G→F→G m (fincl (zero , p) x) = sym (fpush (zero , p) x) - G→F→G m (fincl (suc n , p) x) = refl - G→F→G m (fpush (zero , p) x i) j = fpush (zero , p) x (~ j ∨ i) - G→F→G (suc m) (fpush (suc n , p) x i) = refl - - Iso-FinSeqColim₀-Top : Iso (FinSeqColim 0 X) (obj X zero) - Iso.fun Iso-FinSeqColim₀-Top (fincl (zero , p) x) = x - Iso.inv Iso-FinSeqColim₀-Top a = fincl fzero a - Iso.rightInv Iso-FinSeqColim₀-Top a = refl - Iso.leftInv Iso-FinSeqColim₀-Top (fincl (zero , p) x) = refl - -pre-Iso-FinSeqColim-Top : (X : Sequence ℓ) (m : ℕ) - → Iso (FinSeqColim m X) (obj X m) -pre-Iso-FinSeqColim-Top X zero = Iso-FinSeqColim₀-Top X -pre-Iso-FinSeqColim-Top X (suc m) = - compIso (Iso-FinSeqColim→FinSeqColim↑ X m) - (pre-Iso-FinSeqColim-Top (ShiftSeq X) m) - -characInverse : (X : Sequence ℓ) (m : ℕ) (a : obj X m) - → Iso.inv (pre-Iso-FinSeqColim-Top X m) a ≡ fincl flast a -characInverse X zero a = refl -characInverse X (suc m) a = - cong (Iso.inv (Iso-FinSeqColim→FinSeqColim↑ X m)) (characInverse _ m a) - --- main result -Iso-FinSeqColim-Top : (X : Sequence ℓ) (m : ℕ) - → Iso (FinSeqColim m X) (obj X m) -Iso.fun (Iso-FinSeqColim-Top X m) = Iso.fun (pre-Iso-FinSeqColim-Top X m) -Iso.inv (Iso-FinSeqColim-Top X m) = fincl flast -Iso.rightInv (Iso-FinSeqColim-Top X m) r = - cong (Iso.fun (pre-Iso-FinSeqColim-Top X m)) (sym (characInverse X m r)) - ∙ Iso.rightInv (pre-Iso-FinSeqColim-Top X m) r -Iso.leftInv (Iso-FinSeqColim-Top X m) r = - sym (characInverse X m (Iso.fun (pre-Iso-FinSeqColim-Top X m) r)) - ∙ Iso.leftInv (pre-Iso-FinSeqColim-Top X m) r - - -- main corollary : given two maps (f g : SeqColim Xᵢ → B) and a - -- family of homotopies hᵢ : (x : Xᵢ) → f (incl x) ≡ g (incl x) for - -- i < n, we can improve hᵢ such that they are compatible with push - -→FinSeqColimHomotopy : ∀ {ℓ ℓ'} - {X : Sequence ℓ} {m : ℕ} {B : FinSeqColim m X → Type ℓ'} - (f g : (x : FinSeqColim m X) → B x) - (h : (x : obj X m) → f (fincl flast x) - ≡ g (fincl flast x)) - → f ≡ g -→FinSeqColimHomotopy {X = X} {m} f g h = funExt - (transport (λ i → (x : isoToPath (invIso (Iso-FinSeqColim-Top X m)) i) - → f (ua-unglue (isoToEquiv (invIso (Iso-FinSeqColim-Top X m))) i x) - ≡ g (ua-unglue (isoToEquiv (invIso (Iso-FinSeqColim-Top X m))) i x)) h) diff --git a/Cubical/HITs/SequentialColimit/Properties.agda b/Cubical/HITs/SequentialColimit/Properties.agda index 656b19ca88..7e5679f63e 100644 --- a/Cubical/HITs/SequentialColimit/Properties.agda +++ b/Cubical/HITs/SequentialColimit/Properties.agda @@ -3,6 +3,7 @@ This file contains: - Eliminators of direct limit, especially an index-shifting version; - Connectivity of inclusion maps. + - Characterisation of colimits over finite sequences -} {-# OPTIONS --safe #-} @@ -19,6 +20,7 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Nat hiding (elim) open import Cubical.Data.Sequence +open import Cubical.Data.Fin.Inductive open import Cubical.HITs.SequentialColimit.Base open import Cubical.Homotopy.Connected @@ -267,18 +269,18 @@ module _ {ℓ : Level} (seq : Sequence ℓ) (n : ℕ) (term : converges seq n) (bs : (s : _) → P zero s) (indr : (k : ℕ) → ((y : _) → P k y) → (x : _) → P (suc k) (Sequence.map seq x)) where - terminaates-seq-ind : (k : ℕ) (x : _) → P k x - terminaates-seq-ind zero = bs - terminaates-seq-ind (suc k) x = + terminates-seq-ind : (k : ℕ) (x : _) → P k x + terminates-seq-ind zero = bs + terminates-seq-ind (suc k) x = subst (P (suc k)) (secEq (_ , term k) x) - (indr k (terminaates-seq-ind k) (invEq (_ , term k) x)) + (indr k (terminates-seq-ind k) (invEq (_ , term k) x)) - terminaates-seq-indβ : (k : ℕ) (s : ((y : seq .obj (k + n)) → P k y)) (x : _) - → terminaates-seq-ind (suc k) (Sequence.map seq x) - ≡ indr k (terminaates-seq-ind k) x - terminaates-seq-indβ k s x = + terminates-seq-indβ : (k : ℕ) (s : ((y : seq .obj (k + n)) → P k y)) (x : _) + → terminates-seq-ind (suc k) (Sequence.map seq x) + ≡ indr k (terminates-seq-ind k) x + terminates-seq-indβ k s x = lem (Sequence.map seq , term k) x (P (suc k)) - (indr k (terminaates-seq-ind k)) + (indr k (terminates-seq-ind k)) where lem : ∀ {ℓ ℓ'} {A B : Type ℓ} (e : A ≃ B) (x : A) (P : B → Type ℓ') (πP : (x : A) → P (fst e x)) @@ -302,7 +304,7 @@ module _ {ℓ : Level} (seq : Sequence ℓ) (n : ℕ) (term : converges seq n) shiftEqShifted (incl x))) (incl x) converges→ColimIso-main-lem = - terminaates-seq-ind + terminates-seq-ind 0case (λ k id x → cong incl @@ -318,7 +320,7 @@ module _ {ℓ : Level} (seq : Sequence ℓ) (n : ℕ) (term : converges seq n) ∙∙ converges→ColimIso-main-lem k x ∙∙ push x) converges→ColimIso-main-lemβ k x = - terminaates-seq-indβ + terminates-seq-indβ 0case (λ k id x → cong incl @@ -433,3 +435,79 @@ converges→funId {seq1 = seq1} {seq2} n m t1 t2 f g p = → SeqColim seq2) f (f ∘ incl) help f = toPathP (funExt λ x → (λ i → transportRefl (f (incl (transportRefl x i))) i)) + +-- colim (X₀ → ... → Xₙ) ≃ colim (X₁ → ... → Xₙ) +module _ (X : Sequence ℓ) where + ShiftSeq : Sequence ℓ + obj ShiftSeq m = obj X (suc m) + map ShiftSeq = map X + + Iso-FinSeqColim→FinSeqColim↑ : (m : ℕ) + → Iso (FinSeqColim (suc m) X) (FinSeqColim m ShiftSeq) + Iso-FinSeqColim→FinSeqColim↑ m = iso (G m) (F m) (F→G→F m) (G→F→G m) + where + F : (m : ℕ) → FinSeqColim m ShiftSeq → FinSeqColim (suc m) X + F m (fincl n x) = fincl (fsuc n) x + F (suc m) (fpush (n , p) x i) = fpush (suc n , p) x i + + G : (m : ℕ) → FinSeqColim (suc m) X → FinSeqColim m ShiftSeq + G m (fincl (zero , p) x) = fincl (zero , p) (map X x) + G m (fincl (suc n , p) x) = fincl (n , p) x + G m (fpush (zero , p) x i) = fincl (zero , p) (map X x) + G (suc m) (fpush (suc n , p) x i) = fpush (n , p) x i + + F→G→F : (m : ℕ) → (x : FinSeqColim m ShiftSeq) → G m (F m x) ≡ x + F→G→F m (fincl n x) = refl + F→G→F (suc m) (fpush n x i) = refl + + G→F→G : (m : ℕ) → (x : FinSeqColim (suc m) X) → F m (G m x) ≡ x + G→F→G m (fincl (zero , p) x) = sym (fpush (zero , p) x) + G→F→G m (fincl (suc n , p) x) = refl + G→F→G m (fpush (zero , p) x i) j = fpush (zero , p) x (~ j ∨ i) + G→F→G (suc m) (fpush (suc n , p) x i) = refl + + Iso-FinSeqColim₀-Top : Iso (FinSeqColim 0 X) (obj X zero) + Iso.fun Iso-FinSeqColim₀-Top (fincl (zero , p) x) = x + Iso.inv Iso-FinSeqColim₀-Top a = fincl fzero a + Iso.rightInv Iso-FinSeqColim₀-Top a = refl + Iso.leftInv Iso-FinSeqColim₀-Top (fincl (zero , p) x) = refl + +pre-Iso-FinSeqColim-Top : (X : Sequence ℓ) (m : ℕ) + → Iso (FinSeqColim m X) (obj X m) +pre-Iso-FinSeqColim-Top X zero = Iso-FinSeqColim₀-Top X +pre-Iso-FinSeqColim-Top X (suc m) = + compIso (Iso-FinSeqColim→FinSeqColim↑ X m) + (pre-Iso-FinSeqColim-Top (ShiftSeq X) m) + +characInverse : (X : Sequence ℓ) (m : ℕ) (a : obj X m) + → Iso.inv (pre-Iso-FinSeqColim-Top X m) a ≡ fincl flast a +characInverse X zero a = refl +characInverse X (suc m) a = + cong (Iso.inv (Iso-FinSeqColim→FinSeqColim↑ X m)) (characInverse _ m a) + +-- main result +Iso-FinSeqColim-Top : (X : Sequence ℓ) (m : ℕ) + → Iso (FinSeqColim m X) (obj X m) +Iso.fun (Iso-FinSeqColim-Top X m) = Iso.fun (pre-Iso-FinSeqColim-Top X m) +Iso.inv (Iso-FinSeqColim-Top X m) = fincl flast +Iso.rightInv (Iso-FinSeqColim-Top X m) r = + cong (Iso.fun (pre-Iso-FinSeqColim-Top X m)) (sym (characInverse X m r)) + ∙ Iso.rightInv (pre-Iso-FinSeqColim-Top X m) r +Iso.leftInv (Iso-FinSeqColim-Top X m) r = + sym (characInverse X m (Iso.fun (pre-Iso-FinSeqColim-Top X m) r)) + ∙ Iso.leftInv (pre-Iso-FinSeqColim-Top X m) r + + -- main corollary : given two maps (f g : SeqColim Xᵢ → B) and a + -- family of homotopies hᵢ : (x : Xᵢ) → f (incl x) ≡ g (incl x) for + -- i < n, we can improve hᵢ such that they are compatible with push + +→FinSeqColimHomotopy : ∀ {ℓ ℓ'} + {X : Sequence ℓ} {m : ℕ} {B : FinSeqColim m X → Type ℓ'} + (f g : (x : FinSeqColim m X) → B x) + (h : (x : obj X m) → f (fincl flast x) + ≡ g (fincl flast x)) + → f ≡ g +→FinSeqColimHomotopy {X = X} {m} f g h = funExt + (transport (λ i → (x : isoToPath (invIso (Iso-FinSeqColim-Top X m)) i) + → f (ua-unglue (isoToEquiv (invIso (Iso-FinSeqColim-Top X m))) i x) + ≡ g (ua-unglue (isoToEquiv (invIso (Iso-FinSeqColim-Top X m))) i x)) h) diff --git a/Cubical/HITs/Sn/Degree.agda b/Cubical/HITs/Sn/Degree.agda index 082dc9b1b0..bfb32d0c9a 100644 --- a/Cubical/HITs/Sn/Degree.agda +++ b/Cubical/HITs/Sn/Degree.agda @@ -1,7 +1,7 @@ {-# OPTIONS --safe #-} module Cubical.HITs.Sn.Degree where {- -Contains facts about the degree map Sⁿ → Sⁿ +Contains facts about the degree of maps Sⁿ → Sⁿ -} open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed @@ -136,7 +136,7 @@ degreeComp' n f g = degreeComp n f g ∙ ·Comm (degree n f) (degree n g) (degreeComp (suc n) f g ∙ sym (degreeComp' (suc n) g f)) ∙ Iso.leftInv (degree∥₂Iso n) (∣ g ∘ f ∣₂))) -degreeSusp : (n : ℕ) (f : (S₊ n → S₊ n)) +degreeSusp : (n : ℕ) (f : S₊ n → S₊ n) → degree n f ≡ degree (suc n) (suspFunS∙ f .fst) degreeSusp zero f with (f true) | (f false) ... | false | false = refl diff --git a/Cubical/HITs/SphereBouquet/Base.agda b/Cubical/HITs/SphereBouquet/Base.agda index d223cd5b14..68c8a9e5dc 100644 --- a/Cubical/HITs/SphereBouquet/Base.agda +++ b/Cubical/HITs/SphereBouquet/Base.agda @@ -3,9 +3,10 @@ module Cubical.HITs.SphereBouquet.Base where open import Cubical.Foundations.Prelude -open import Cubical.Data.Nat open import Cubical.Foundations.Pointed +open import Cubical.Data.Nat + open import Cubical.HITs.Wedge open import Cubical.HITs.Sn diff --git a/Cubical/HITs/SphereBouquet/Degree.agda b/Cubical/HITs/SphereBouquet/Degree.agda index 76b9d30b19..0c7526597b 100644 --- a/Cubical/HITs/SphereBouquet/Degree.agda +++ b/Cubical/HITs/SphereBouquet/Degree.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.HITs.SphereBouquet.Degree where @@ -6,21 +6,16 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Path open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Univalence open import Cubical.Foundations.Function -open import Cubical.Data.Bool hiding (_≤_ ; _≟_ ; isProp≤) +open import Cubical.Data.Bool open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) -open import Cubical.Data.Nat.Order -open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Fin.Inductive open import Cubical.Data.Sigma open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Unit open import Cubical.HITs.S1 open import Cubical.HITs.Sn @@ -30,29 +25,22 @@ open import Cubical.HITs.Susp open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.Truncation as TR open import Cubical.HITs.SetTruncation as ST - -open import Cubical.Homotopy.Loopspace +open import Cubical.HITs.Wedge +open import Cubical.HITs.SphereBouquet.Base +open import Cubical.HITs.SphereBouquet.Properties open import Cubical.Relation.Nullary -open import Cubical.Algebra.Group -open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup -open import Cubical.HITs.Wedge - open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.Properties - -open import Cubical.HITs.SphereBouquet.Base -open import Cubical.HITs.SphereBouquet.Properties - private variable ℓ ℓ' : Level diff --git a/Cubical/HITs/SphereBouquet/Properties.agda b/Cubical/HITs/SphereBouquet/Properties.agda index f0e51b287a..8c0284f731 100644 --- a/Cubical/HITs/SphereBouquet/Properties.agda +++ b/Cubical/HITs/SphereBouquet/Properties.agda @@ -7,47 +7,24 @@ open import Cubical.Foundations.Path open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv -open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence open import Cubical.Foundations.Function -open import Cubical.Data.Bool hiding (_≤_ ; _≟_ ; isProp≤) +open import Cubical.Data.Bool open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) -open import Cubical.Data.Nat.Order -open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Fin.Inductive open import Cubical.Data.Sigma -open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Unit open import Cubical.HITs.S1 open import Cubical.HITs.Sn open import Cubical.HITs.Pushout open import Cubical.HITs.Susp open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.HITs.Truncation as TR -open import Cubical.HITs.SetTruncation as ST - -open import Cubical.Homotopy.Loopspace - -open import Cubical.Relation.Nullary - -open import Cubical.Algebra.Group -open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup - open import Cubical.HITs.Wedge open import Cubical.HITs.SphereBouquet.Base -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.GroupStructure -open import Cubical.ZCohomology.Groups.Sn -open import Cubical.ZCohomology.Properties - private variable ℓ ℓ' : Level diff --git a/Cubical/Sequence.agda b/Cubical/Sequence.agda deleted file mode 100644 index e69de29bb2..0000000000 diff --git a/Cubical/Structures/Successor.agda b/Cubical/Structures/Successor.agda deleted file mode 100644 index 1e4727c90e..0000000000 --- a/Cubical/Structures/Successor.agda +++ /dev/null @@ -1,61 +0,0 @@ -{-# OPTIONS --safe #-} -{- - Successor structures for spectra, chain complexes and fiber sequences. - This is an idea from Floris van Doorn's phd thesis. --} -module Cubical.Structures.Successor where - -open import Cubical.Foundations.Prelude -open import Cubical.Data.Int -open import Cubical.Data.Nat -open import Cubical.Data.Fin -open import Cubical.Data.Sigma -open import Cubical.Data.Nat.Order -open import Cubical.Relation.Nullary - -private - variable - ℓ : Level - -record SuccStr (ℓ : Level) : Type (ℓ-suc ℓ) where - field - Index : Type ℓ - succ : Index → Index - -open SuccStr - -ℤ+ : SuccStr ℓ-zero -ℤ+ .Index = ℤ -ℤ+ .succ = sucℤ - -ℕ+ : SuccStr ℓ-zero -ℕ+ .Index = ℕ -ℕ+ .succ = suc - -Fin+ : (n : ℕ) → SuccStr ℓ-zero -Fin+ n .Index = Fin n -Fin+ n .succ (x , zero , p) = x , zero , p -Fin+ n .succ (x , suc diff , p) = suc x , diff , +-suc diff (suc x) ∙ p - -record InjSuccStr (ℓ : Level) : Type (ℓ-suc ℓ) where - field - succStr : SuccStr ℓ - semiInj : {i : Index succStr} - → ¬ (succ succStr i ≡ succ succStr (succ succStr i)) - → ¬ (i ≡ succ succStr i) - -open InjSuccStr - -≠suc : {n : ℕ} → ¬ n ≡ suc n -≠suc {n = n} p = ¬m Date: Mon, 4 Mar 2024 01:30:09 +0100 Subject: [PATCH 32/73] readme --- Cubical/README.agda | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Cubical/README.agda b/Cubical/README.agda index ed23331c73..c8adb8e308 100644 --- a/Cubical/README.agda +++ b/Cubical/README.agda @@ -82,3 +82,6 @@ import Cubical.Axiom.Everything -- Automatic proving, solvers import Cubical.Tactics.Everything + +-- CW complexes, cellular (co)homology +import Cubical.CW.Everything From 2d3fe1c02ac12dcbebc159a47f37a7fb11d08580 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 Mar 2024 01:48:45 +0100 Subject: [PATCH 33/73] wups --- Cubical/Data/CW.agda | 6 ------ 1 file changed, 6 deletions(-) delete mode 100644 Cubical/Data/CW.agda diff --git a/Cubical/Data/CW.agda b/Cubical/Data/CW.agda deleted file mode 100644 index 80ce46452a..0000000000 --- a/Cubical/Data/CW.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Data.CW where - -open import Cubical.Data.CW.Base public -open import Cubical.Data.CW.Properties public From 1b005a64708be0b27eceefd6c5f1628385f00ae7 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 Mar 2024 02:22:00 +0100 Subject: [PATCH 34/73] ugh --- Cubical/HITs/James/Inductive/Base.agda | 1 + 1 file changed, 1 insertion(+) diff --git a/Cubical/HITs/James/Inductive/Base.agda b/Cubical/HITs/James/Inductive/Base.agda index 69b9b2096c..ce540d66ec 100644 --- a/Cubical/HITs/James/Inductive/Base.agda +++ b/Cubical/HITs/James/Inductive/Base.agda @@ -12,6 +12,7 @@ open import Cubical.Foundations.Pointed open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Data.Nat +open import Cubical.Data.Sequence open import Cubical.HITs.SequentialColimit From ac5249e76295d97d20f3f827158e0f7b994e5150 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 Mar 2024 09:39:57 +0100 Subject: [PATCH 35/73] wups --- Cubical/HITs/James/Inductive.agda | 1 + 1 file changed, 1 insertion(+) diff --git a/Cubical/HITs/James/Inductive.agda b/Cubical/HITs/James/Inductive.agda index fb69a15353..8910f4074d 100644 --- a/Cubical/HITs/James/Inductive.agda +++ b/Cubical/HITs/James/Inductive.agda @@ -23,6 +23,7 @@ open import Cubical.Foundations.Pointed open import Cubical.Data.Nat open import Cubical.Data.Unit open import Cubical.Data.Sigma +open import Cubical.Data.Sequence open import Cubical.HITs.Wedge open import Cubical.HITs.Pushout From 3e53282ae92499ae69f28596cd195e55055a76b3 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 Mar 2024 11:31:38 +0100 Subject: [PATCH 36/73] broken code --- .../Localisation/InvertingElements.agda | 40 +++++++------------ 1 file changed, 15 insertions(+), 25 deletions(-) diff --git a/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda b/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda index 1a10d2e56b..6441cd684e 100644 --- a/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda +++ b/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda @@ -720,17 +720,11 @@ module DoubleLoc (R' : CommRing ℓ) (f g : fst R') where , PT.∣ l , ^-respects-/1 R' l ∣₁) , eq/ _ _ ((f ^ l , PT.∣ l , refl ∣₁) , path)) where - path : f ^ l · (g ^ l · transp (λ i → R) i0 r · transp (λ i → R) i0 (g ^ m)) - · (1r · transp (λ i → R) i0 (f ^ m) · transp (λ i → R) i0 1r) - ≡ f ^ l · (g ^ l · transp (λ i → R) i0 r' · transp (λ i → R) i0 (g ^ n)) - · (1r · transp (λ i → R) i0 (f ^ n) · transp (λ i → R) i0 1r) - path = f ^ l · (g ^ l · transp (λ i → R) i0 r · transp (λ i → R) i0 (g ^ m)) - · (1r · transp (λ i → R) i0 (f ^ m) · transp (λ i → R) i0 1r) - - ≡⟨ (λ i → f ^ l · (g ^ l · transportRefl r i · transportRefl (g ^ m) i) - · (1r · transportRefl (f ^ m) i · transportRefl 1r i)) ⟩ - - f ^ l · (g ^ l · r · g ^ m) · (1r · f ^ m · 1r) + path : f ^ l · (g ^ l · r · (g ^ m)) + · (1r · (f ^ m) · 1r) + ≡ f ^ l · (g ^ l · r' · (g ^ n)) + · (1r · (f ^ n) · 1r) + path = f ^ l · (g ^ l · r · g ^ m) · (1r · f ^ m · 1r) ≡⟨ solve! R' ⟩ @@ -766,13 +760,8 @@ module DoubleLoc (R' : CommRing ℓ) (f g : fst R') where ≡⟨ solve! R' ⟩ - f ^ l · (g ^ l · r' · g ^ n) · (1r · f ^ n · 1r) - - ≡⟨ (λ i → f ^ l · (g ^ l · transportRefl r' (~ i) · transportRefl (g ^ n) (~ i)) - · (1r · transportRefl (f ^ n) (~ i) · transportRefl 1r (~ i))) ⟩ - - f ^ l · (g ^ l · transp (λ i → R) i0 r' · transp (λ i → R) i0 (g ^ n)) - · (1r · transp (λ i → R) i0 (f ^ n) · transp (λ i → R) i0 1r) ∎ + f ^ l · (g ^ l · r' · (g ^ n)) + · (1r · (f ^ n) · 1r) ∎ curriedϕcoh : (r s r' s' u : R) → (p : u · r · s' ≡ u · r' · s) @@ -810,13 +799,14 @@ module DoubleLoc (R' : CommRing ℓ) (f g : fst R') where where ℕcase : (r : R) (n : ℕ) → φ [ r , (f · g) ^ n , PT.∣ n , refl ∣₁ ] ≡ χ [ r , (f · g) ^ n , PT.∣ n , refl ∣₁ ] - ℕcase r n = cong [_] (ΣPathP --look into the components of the double-fractions - ( cong [_] (ΣPathP (eq1 , Σ≡Prop (λ x → S'[f] x .snd) (sym (·IdL _)))) - , Σ≡Prop (λ x → S'[f][g] x .snd) --ignore proof that denominator is power of g/1 - ( cong [_] (ΣPathP (sym (·IdL _) , Σ≡Prop (λ x → S'[f] x .snd) (sym (·IdL _))))))) + ℕcase r n = + cong [_] (ΣPathP (cong [_] (ΣPathP (sym (·IdR r) ∙ cong (r ·_) (sym (transportRefl 1r)) + , Σ≡Prop (λ x → S'[f] x .snd) + (sym (·IdL _) ∙ cong (1r ·_) (sym (transportRefl _))))) + , (Σ≡Prop (λ x → S'[f][g] x .snd) -- ignore proof that denominator is power of g/1 + (cong [_] (ΣPathP (sym (·IdL _) ∙ cong (1r ·_) (sym (transportRefl _)) + , Σ≡Prop (λ x → S'[f] x .snd) + (sym (·IdL _) ∙ cong (1r ·_) (sym (transportRefl _))))))))) where S'[f] = ([_ⁿ|n≥0] R' f) S'[f][g] = ([_ⁿ|n≥0] R[1/f]AsCommRing [ g , 1r , powersFormMultClosedSubset R' f .containsOne ]) - - eq1 : transp (λ i → fst R') i0 r ≡ r · transp (λ i → fst R') i0 1r - eq1 = transportRefl r ∙∙ sym (·IdR r) ∙∙ cong (r ·_) (sym (transportRefl 1r)) From beccacd97a903f3140d19c9ac50dfcff40f89b2e Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 Mar 2024 12:03:53 +0100 Subject: [PATCH 37/73] =?UTF-8?q?ojd=C3=A5?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Cubical/Structures/Successor.agda | 29 +++++++++++++++++++++++++++++ 1 file changed, 29 insertions(+) create mode 100644 Cubical/Structures/Successor.agda diff --git a/Cubical/Structures/Successor.agda b/Cubical/Structures/Successor.agda new file mode 100644 index 0000000000..ce72551c60 --- /dev/null +++ b/Cubical/Structures/Successor.agda @@ -0,0 +1,29 @@ +{-# OPTIONS --safe #-} +{- + Successor structures for spectra, chain complexes and fiber sequences. + This is an idea from Floris van Doorn's phd thesis. +-} +module Cubical.Structures.Successor where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Int +open import Cubical.Data.Nat + +private + variable + ℓ : Level + +record SuccStr (ℓ : Level) : Type (ℓ-suc ℓ) where + field + Index : Type ℓ + succ : Index → Index + +open SuccStr + +ℤ+ : SuccStr ℓ-zero +ℤ+ .Index = ℤ +ℤ+ .succ = sucℤ + +ℕ+ : SuccStr ℓ-zero +ℕ+ .Index = ℕ +ℕ+ .succ = suc From 3784dcd118ef25e2123535a01df66d95097b2de6 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 2 May 2024 11:34:11 +0200 Subject: [PATCH 38/73] comments --- .../AbGroup/Instances/FreeAbGroup.agda | 1 + Cubical/Algebra/ChainComplex/Finite.agda | 2 +- Cubical/Algebra/ChainComplex/Homology.agda | 1 + Cubical/CW/Approximation.agda | 31 +++---- Cubical/CW/Homotopy.agda | 23 +++--- Cubical/CW/Map.agda | 1 + Cubical/Data/Fin/Inductive/Base.agda | 51 +++--------- Cubical/Data/Fin/Inductive/Properties.agda | 44 ++-------- Cubical/Data/FinSequence/Base.agda | 2 + Cubical/Data/Nat/Order/Inductive.agda | 81 +++++++++++++++++++ Cubical/Data/Sequence/Base.agda | 2 + Cubical/HITs/SequentialColimit/Base.agda | 4 +- .../HITs/SequentialColimit/Properties.agda | 5 +- Cubical/HITs/SphereBouquet/Degree.agda | 1 + 14 files changed, 139 insertions(+), 110 deletions(-) create mode 100644 Cubical/Data/Nat/Order/Inductive.agda diff --git a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda index aab089129a..53f6b48a79 100644 --- a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda +++ b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda @@ -9,6 +9,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Data.Sigma open import Cubical.Data.Nat hiding (_·_) renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order +open import Cubical.Data.Nat.Order.Inductive open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) open import Cubical.Data.Fin.Inductive open import Cubical.Data.Empty as ⊥ diff --git a/Cubical/Algebra/ChainComplex/Finite.agda b/Cubical/Algebra/ChainComplex/Finite.agda index 011d37d2f0..0a8611d896 100644 --- a/Cubical/Algebra/ChainComplex/Finite.agda +++ b/Cubical/Algebra/ChainComplex/Finite.agda @@ -17,7 +17,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma open import Cubical.Data.Nat -open import Cubical.Data.Fin.Inductive.Base hiding (eq) +open import Cubical.Data.Fin.Inductive.Base open import Cubical.Algebra.ChainComplex.Base open import Cubical.Algebra.Group.MorphismProperties diff --git a/Cubical/Algebra/ChainComplex/Homology.agda b/Cubical/Algebra/ChainComplex/Homology.agda index 45c6cefd2b..99d4dffd2d 100644 --- a/Cubical/Algebra/ChainComplex/Homology.agda +++ b/Cubical/Algebra/ChainComplex/Homology.agda @@ -10,6 +10,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma open import Cubical.Data.Nat open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Nat.Order.Inductive open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms diff --git a/Cubical/CW/Approximation.agda b/Cubical/CW/Approximation.agda index c8f185faeb..89c2dbd594 100644 --- a/Cubical/CW/Approximation.agda +++ b/Cubical/CW/Approximation.agda @@ -25,6 +25,7 @@ open import Cubical.Data.Unit open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sequence open import Cubical.Data.FinSequence +open import Cubical.Data.Nat.Order.Inductive open import Cubical.HITs.SequentialColimit open import Cubical.HITs.PropositionalTruncation as PT hiding (elimFin) @@ -72,7 +73,7 @@ snd (compFinCellApprox m {C = C} {g = g} {f} (F , p) (G , q)) = → funExt⁻ p _ ∙ cong g (funExt⁻ q (fincl _ x)) --- a finite cellular homotopies relative +-- a finite cellular homotopies relative a homotopy in the colimit finCellHomRel : {C : CWskel ℓ} {D : CWskel ℓ'} (m : ℕ) (f g : finCellMap m C D) (h∞ : (n : Fin (suc m)) (c : fst C (fst n)) @@ -264,7 +265,7 @@ CWmap→finCellMap : (C : CWskel ℓ) (D : CWskel ℓ') → ∥ finCellApprox C D f m ∥₁ CWmap→finCellMap C D f m = PT.map (λ {(g , hom) - → (record { fmap = fst ∘ g ; fcomm = λ r x → sym (hom r x) }) + → finsequencemap (fst ∘ g) (λ r x → sym (hom r x)) , →FinSeqColimHomotopy _ _ (g flast .snd)}) (approx C D f m) @@ -580,17 +581,17 @@ private ≡ FinSeqColim→Colim m ∘ finCellMap→FinSeqColim C D g-c) → ∥ finCellHomRel m f-c g-c (approx.h∞ m f-c g-c h∞') ∥₁ pathToCellularHomotopy-main {C = C} zero f-c g-c h∞' = - ∣ (record { fhom = λ {(zero , p) x → ⊥.rec (CW₀-empty C x)} - ; fcoh = λ {()} }) - , (λ { (zero , p) x → ⊥.rec (CW₀-empty C x)}) ∣₁ + ∣ fincellhom (λ {(zero , p) x → ⊥.rec (CW₀-empty C x)}) + (λ { (zero , p) x → ⊥.rec (CW₀-empty C x)}) + , (λ { (zero , p) x → ⊥.rec (CW₀-empty C x)}) ∣₁ pathToCellularHomotopy-main {C = C} {D = D} (suc zero) f-c g-c h∞' = PT.map (λ {(d , h) - → (record - { fhom = λ {(zero , p) x → ⊥.rec (CW₀-empty C x) - ; (suc zero , p) → d} - ; fcoh = λ {(zero , p) → λ x → ⊥.rec (CW₀-empty C x)} }) - , (λ {(zero , p) x → ⊥.rec (CW₀-empty C x) - ; (suc zero , tt) → h})}) (invEq (_ , satAC∃Fin-C0 C _ _) k) + → (fincellhom (λ {(zero , p) x → ⊥.rec (CW₀-empty C x) + ; (suc zero , p) → d}) + λ {(zero , p) → λ x → ⊥.rec (CW₀-empty C x)}) + , (λ {(zero , p) x → ⊥.rec (CW₀-empty C x) + ; (suc zero , tt) → h})}) + (invEq (_ , satAC∃Fin-C0 C _ _) k) where k : (c : _) → _ k c = (approx.pathToCellularHomotopy₁ (suc zero) f-c g-c @@ -598,10 +599,10 @@ private pathToCellularHomotopy-main {C = C} {D = D} (suc (suc m)) f-c g-c h∞' = PT.rec squash₁ (λ ind → PT.map - (λ {(f , p) → - (record { fhom = main-hom ind f p - ; fcoh = main-coh ind f p }) - , ∞coh ind f p}) + (λ {(f , p) + → (fincellhom (main-hom ind f p) + (main-coh ind f p)) + , (∞coh ind f p)}) (pathToCellularHomotopy-ind flast λ c → (finCellHom.fhom (ind .fst) flast c) , (ind .snd flast c))) diff --git a/Cubical/CW/Homotopy.agda b/Cubical/CW/Homotopy.agda index fdcd2874ef..430437223a 100644 --- a/Cubical/CW/Homotopy.agda +++ b/Cubical/CW/Homotopy.agda @@ -14,6 +14,7 @@ open import Cubical.CW.Map open import Cubical.CW.ChainComplex open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws @@ -25,11 +26,15 @@ open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sequence open import Cubical.Data.FinSequence +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn open import Cubical.HITs.Sn.Degree open import Cubical.HITs.Pushout open import Cubical.HITs.Susp open import Cubical.HITs.SphereBouquet open import Cubical.HITs.SphereBouquet.Degree +open import Cubical.HITs.SetTruncation as ST hiding (map) +open import Cubical.HITs.Truncation as TR hiding (map) open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.AbGroup @@ -37,12 +42,18 @@ open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup open import Cubical.Algebra.ChainComplex +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Sn + private variable ℓ ℓ' ℓ'' : Level -- A cellular homotopy between two cellular maps record cellHom {C : CWskel ℓ} {D : CWskel ℓ'} (f g : cellMap C D) : Type (ℓ-max ℓ ℓ') where + constructor cellhom open SequenceMap field hom : (n : ℕ) → (x : C .fst n) → CW↪ D n (f .map n x) ≡ CW↪ D n (g .map n x) @@ -55,6 +66,7 @@ record cellHom {C : CWskel ℓ} {D : CWskel ℓ'} (f g : cellMap C D) : Type ( -- Finite cellular homotopies record finCellHom (m : ℕ) {C : CWskel ℓ} {D : CWskel ℓ'} (f g : finCellMap m C D) : Type (ℓ-max ℓ ℓ') where + constructor fincellhom open FinSequenceMap field fhom : (n : Fin (suc m)) (x : C .fst (fst n)) @@ -428,17 +440,6 @@ realiseMMΣH∂ C D (suc m) f g H n x = -- Then, we connect the addition of MMmaps to the addition of abelian maps module bouquetAdd where - -- keeping imports here for now - open import Cubical.ZCohomology.Base - open import Cubical.ZCohomology.Properties - open import Cubical.ZCohomology.GroupStructure - open import Cubical.HITs.Truncation as TR hiding (map) - open import Cubical.HITs.Sn - open import Cubical.HITs.S1 - open import Cubical.Foundations.Path - open import Cubical.ZCohomology.Groups.Sn - open import Cubical.HITs.SetTruncation as ST hiding (map) - open MMmaps module _ (C : CWskel ℓ) (D : CWskel ℓ') (m : ℕ) (n : Fin m) diff --git a/Cubical/CW/Map.agda b/Cubical/CW/Map.agda index 3f7c35fe8a..b057c67666 100644 --- a/Cubical/CW/Map.agda +++ b/Cubical/CW/Map.agda @@ -17,6 +17,7 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order.Inductive open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) open import Cubical.Data.Bool open import Cubical.Data.Fin.Inductive.Base diff --git a/Cubical/Data/Fin/Inductive/Base.agda b/Cubical/Data/Fin/Inductive/Base.agda index 5d3e588cd6..d2280736a1 100644 --- a/Cubical/Data/Fin/Inductive/Base.agda +++ b/Cubical/Data/Fin/Inductive/Base.agda @@ -1,7 +1,6 @@ -{-# OPTIONS --safe --lossy-unification #-} +{-# OPTIONS --safe #-} --- Alternative definition of < on ℕ and thus an alternative definition --- of Fin n +-- Definition of Fin n in terms of the inductively defined < module Cubical.Data.Fin.Inductive.Base where @@ -9,38 +8,13 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Nat.Order.Inductive open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Unit open import Cubical.Relation.Nullary --- inductive definition of < -_<ᵗ_ : (n m : ℕ) → Type -n <ᵗ zero = ⊥ -zero <ᵗ suc m = Unit -suc n <ᵗ suc m = n <ᵗ m - -data Trichotomyᵗ (m n : ℕ) : Type₀ where - lt : m <ᵗ n → Trichotomyᵗ m n - eq : m ≡ n → Trichotomyᵗ m n - gt : n <ᵗ m → Trichotomyᵗ m n - -Trichotomyᵗ-suc : {n m : ℕ} → Trichotomyᵗ n m - → Trichotomyᵗ (suc n) (suc m) -Trichotomyᵗ-suc (lt x) = lt x -Trichotomyᵗ-suc (eq x) = eq (cong suc x) -Trichotomyᵗ-suc (gt x) = gt x - -_≟ᵗ_ : ∀ m n → Trichotomyᵗ m n -zero ≟ᵗ zero = eq refl -zero ≟ᵗ suc n = lt tt -suc m ≟ᵗ zero = gt tt -suc m ≟ᵗ suc n = Trichotomyᵗ-suc (m ≟ᵗ n) - -isProp<ᵗ : {n m : ℕ} → isProp (n <ᵗ m) -isProp<ᵗ {n = n} {zero} = isProp⊥ -isProp<ᵗ {n = zero} {suc m} = isPropUnit -isProp<ᵗ {n = suc n} {suc m} = isProp<ᵗ {n = n} {m = m} -- Fin defined using <ᵗ Fin : (n : ℕ) → Type Fin n = Σ[ m ∈ ℕ ] (m <ᵗ n) @@ -49,17 +23,6 @@ fsuc : {n : ℕ} → Fin n → Fin (suc n) fst (fsuc {n = n} (x , p)) = suc x snd (fsuc {n = suc n} (x , p)) = p -<ᵗsucm : {m : ℕ} → m <ᵗ suc m -<ᵗsucm {m = zero} = tt -<ᵗsucm {m = suc m} = <ᵗsucm {m = m} - -<ᵗ-trans-suc : {n m : ℕ} → n <ᵗ m → n <ᵗ suc m -<ᵗ-trans-suc {n = zero} {suc m} x = tt -<ᵗ-trans-suc {n = suc n} {suc m} x = <ᵗ-trans-suc {n = n} x - -¬-suc-n<ᵗn : {n : ℕ} → ¬ (suc n) <ᵗ n -¬-suc-n<ᵗn {suc n} = ¬-suc-n<ᵗn {n} - injectSuc : {n : ℕ} → Fin n → Fin (suc n) fst (injectSuc {n = n} (x , p)) = x snd (injectSuc {n = suc n} (x , p)) = <ᵗ-trans-suc {n = x} p @@ -68,8 +31,12 @@ flast : {m : ℕ} → Fin (suc m) fst (flast {m = m}) = m snd (flast {m = m}) = <ᵗsucm {m = m} +fzero : {m : ℕ} → Fin (suc m) +fzero = 0 , tt + -- Sums sumFinGen : ∀ {ℓ} {A : Type ℓ} {n : ℕ} (_+_ : A → A → A) (0A : A) (f : Fin n → A) → A sumFinGen {n = zero} _+_ 0A f = 0A -sumFinGen {n = suc n} _+_ 0A f = f flast + sumFinGen _+_ 0A (f ∘ injectSuc) +sumFinGen {n = suc n} _+_ 0A f = + f flast + (sumFinGen {n = n}) _+_ 0A (f ∘ injectSuc) diff --git a/Cubical/Data/Fin/Inductive/Properties.agda b/Cubical/Data/Fin/Inductive/Properties.agda index a448c2fbd9..38f654fc45 100644 --- a/Cubical/Data/Fin/Inductive/Properties.agda +++ b/Cubical/Data/Fin/Inductive/Properties.agda @@ -13,19 +13,20 @@ open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Unit open import Cubical.Data.Sigma open import Cubical.Data.Nat.Order +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Fin.Base as FinOld + hiding (Fin ; injectSuc ; fsuc ; fzero ; flast ; ¬Fin0 ; sumFinGen) +open import Cubical.Data.Fin.Properties as FinOldProps + hiding (sumFinGen0 ; isSetFin ; sumFin-choose ; sumFinGenHom ; sumFinGenId) open import Cubical.Relation.Nullary open import Cubical.Algebra.AbGroup.Base using (move4) --- inductive definition of < - fsuc-injectSuc : {m : ℕ} (n : Fin m) → injectSuc {n = suc m} (fsuc {n = m} n) ≡ fsuc (injectSuc n) fsuc-injectSuc {m = suc m} (x , p) = refl -fzero : {m : ℕ} → Fin (suc m) -fzero = 0 , tt elimFin : ∀ {ℓ} {m : ℕ} {A : Fin (suc m) → Type ℓ} (max : A flast) @@ -61,21 +62,6 @@ elimFinβ {m = suc (suc m)} {A = A} max f = ¬Fin0 : ¬ Fin 0 ¬Fin0 (x , ()) -<ᵗ-trans : {n m k : ℕ} → n <ᵗ m → m <ᵗ k → n <ᵗ k -<ᵗ-trans {n = zero} {suc m} {suc k} _ _ = tt -<ᵗ-trans {n = suc n} {suc m} {suc k} = <ᵗ-trans {n = n} {m} {k} - -¬m<ᵗm : {m : ℕ} → ¬ (m <ᵗ m) -¬m<ᵗm {m = suc m} p = ¬m<ᵗm p - -<ᵗ-+ : {n k : ℕ} → n <ᵗ suc (k + n) -<ᵗ-+ {n = zero} {k} = tt -<ᵗ-+ {n = suc n} {k} = - subst (n <ᵗ_) (sym (+-suc k n)) (<ᵗ-+ {n = n} {k}) - -¬squeeze : {n m : ℕ} → ¬ ((n <ᵗ m) × (m <ᵗ suc n)) -¬squeeze {n = suc n} {suc m} = ¬squeeze {n = n} {m = m} - -- properties of finite sums module _ {ℓ : Level} {A : Type ℓ} (_+A_ : A → A → A) (0A : A) (rUnit : (x : _) → x +A 0A ≡ x) @@ -137,23 +123,7 @@ sumFinGenId : ∀ {ℓ} {A : Type ℓ} {n : ℕ} (_+_ : A → A → A) (0A : A) (f g : Fin n → A) → f ≡ g → sumFinGen _+_ 0A f ≡ sumFinGen _+_ 0A g sumFinGenId _+_ 0A f g p i = sumFinGen _+_ 0A (p i) - -open import Cubical.Data.Fin.Base renaming (Fin to Fin*) hiding (¬Fin0) -open import Cubical.Data.Fin.Properties renaming (isSetFin to isSetFin*) - -<ᵗ→< : {n m : ℕ} → n <ᵗ m → n < m -<ᵗ→< {n = zero} {suc m} p = m , +-comm m 1 -<ᵗ→< {n = suc n} {suc m} p = suc-≤-suc (<ᵗ→< {n = n} {m = m} p) - -<→<ᵗ : {n m : ℕ} → n < m → n <ᵗ m -<→<ᵗ {n = zero} {m = zero} x = - snotz (sym (+-suc (fst x) 0) ∙ snd x) -<→<ᵗ {n = zero} {m = suc m} _ = tt -<→<ᵗ {n = suc n} {m = zero} x = - snotz (sym (+-suc (fst x) (suc n)) ∙ snd x) -<→<ᵗ {n = suc n} {m = suc m} p = <→<ᵗ {n = n} {m = m} (pred-≤-pred p) - -Iso-Fin-InductiveFin : (m : ℕ) → Iso (Fin* m) (Fin m) +Iso-Fin-InductiveFin : (m : ℕ) → Iso (FinOld.Fin m) (Fin m) Iso.fun (Iso-Fin-InductiveFin m) (x , p) = x , <→<ᵗ p Iso.inv (Iso-Fin-InductiveFin m) (x , p) = x , <ᵗ→< p Iso.rightInv (Iso-Fin-InductiveFin m) (x , p) = @@ -163,4 +133,4 @@ Iso.leftInv (Iso-Fin-InductiveFin m) _ = Σ≡Prop (λ _ → isProp≤) refl isSetFin : {n : ℕ} → isSet (Fin n) isSetFin {n = n} = isOfHLevelRetractFromIso 2 (invIso (Iso-Fin-InductiveFin n)) - isSetFin* + FinOldProps.isSetFin diff --git a/Cubical/Data/FinSequence/Base.agda b/Cubical/Data/FinSequence/Base.agda index 9d45212744..d81e7ee988 100644 --- a/Cubical/Data/FinSequence/Base.agda +++ b/Cubical/Data/FinSequence/Base.agda @@ -10,12 +10,14 @@ private ℓ ℓ' : Level record FinSequence (m : ℕ) (ℓ : Level) : Type (ℓ-suc ℓ) where + constructor finsequence field fobj : Fin (suc m) → Type ℓ fmap : {n : Fin m} → fobj (injectSuc n) → fobj (fsuc n) record FinSequenceMap {m : ℕ} (C : FinSequence m ℓ) (D : FinSequence m ℓ') : Type (ℓ-max ℓ ℓ') where + constructor finsequencemap field fmap : (n : Fin (suc m)) → FinSequence.fobj C n → FinSequence.fobj D n fcomm : (n : Fin m) (x : FinSequence.fobj C (injectSuc n)) diff --git a/Cubical/Data/Nat/Order/Inductive.agda b/Cubical/Data/Nat/Order/Inductive.agda new file mode 100644 index 0000000000..d0d68b39df --- /dev/null +++ b/Cubical/Data/Nat/Order/Inductive.agda @@ -0,0 +1,81 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Nat.Order.Inductive where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.Relation.Nullary + +-- TODO: unify with recursive.agda + +-- inductive definition of < +_<ᵗ_ : (n m : ℕ) → Type +n <ᵗ zero = ⊥ +zero <ᵗ suc m = Unit +suc n <ᵗ suc m = n <ᵗ m + +data Trichotomyᵗ (m n : ℕ) : Type₀ where + lt : m <ᵗ n → Trichotomyᵗ m n + eq : m ≡ n → Trichotomyᵗ m n + gt : n <ᵗ m → Trichotomyᵗ m n + +Trichotomyᵗ-suc : {n m : ℕ} → Trichotomyᵗ n m + → Trichotomyᵗ (suc n) (suc m) +Trichotomyᵗ-suc (lt x) = lt x +Trichotomyᵗ-suc (eq x) = eq (cong suc x) +Trichotomyᵗ-suc (gt x) = gt x + +_≟ᵗ_ : ∀ m n → Trichotomyᵗ m n +zero ≟ᵗ zero = eq refl +zero ≟ᵗ suc n = lt tt +suc m ≟ᵗ zero = gt tt +suc m ≟ᵗ suc n = Trichotomyᵗ-suc (m ≟ᵗ n) + +isProp<ᵗ : {n m : ℕ} → isProp (n <ᵗ m) +isProp<ᵗ {n = n} {zero} = isProp⊥ +isProp<ᵗ {n = zero} {suc m} _ _ = refl +isProp<ᵗ {n = suc n} {suc m} = isProp<ᵗ {n = n} {m = m} + +<ᵗsucm : {m : ℕ} → m <ᵗ suc m +<ᵗsucm {m = zero} = tt +<ᵗsucm {m = suc m} = <ᵗsucm {m = m} + +<ᵗ-trans-suc : {n m : ℕ} → n <ᵗ m → n <ᵗ suc m +<ᵗ-trans-suc {n = zero} {suc m} x = tt +<ᵗ-trans-suc {n = suc n} {suc m} x = <ᵗ-trans-suc {n = n} x + +¬-suc-n<ᵗn : {n : ℕ} → ¬ (suc n) <ᵗ n +¬-suc-n<ᵗn {suc n} = ¬-suc-n<ᵗn {n} + +<ᵗ-trans : {n m k : ℕ} → n <ᵗ m → m <ᵗ k → n <ᵗ k +<ᵗ-trans {n = zero} {suc m} {suc k} _ _ = tt +<ᵗ-trans {n = suc n} {suc m} {suc k} = <ᵗ-trans {n = n} {m} {k} + +¬m<ᵗm : {m : ℕ} → ¬ (m <ᵗ m) +¬m<ᵗm {m = suc m} p = ¬m<ᵗm {m = m} p + +<ᵗ-+ : {n k : ℕ} → n <ᵗ suc (k + n) +<ᵗ-+ {n = zero} {k} = tt +<ᵗ-+ {n = suc n} {k} = + subst (n <ᵗ_) (sym (+-suc k n)) (<ᵗ-+ {n = n} {k}) + +¬squeeze : {n m : ℕ} → ¬ ((n <ᵗ m) × (m <ᵗ suc n)) +¬squeeze {n = suc n} {suc m} = ¬squeeze {n = n} {m = m} + +<ᵗ→< : {n m : ℕ} → n <ᵗ m → n < m +<ᵗ→< {n = zero} {suc m} p = m , +-comm m 1 +<ᵗ→< {n = suc n} {suc m} p = suc-≤-suc (<ᵗ→< {n = n} {m = m} p) + +<→<ᵗ : {n m : ℕ} → n < m → n <ᵗ m +<→<ᵗ {n = zero} {m = zero} x = + snotz (sym (+-suc (fst x) 0) ∙ snd x) +<→<ᵗ {n = zero} {m = suc m} _ = tt +<→<ᵗ {n = suc n} {m = zero} x = + snotz (sym (+-suc (fst x) (suc n)) ∙ snd x) +<→<ᵗ {n = suc n} {m = suc m} p = <→<ᵗ {n = n} {m = m} (pred-≤-pred p) diff --git a/Cubical/Data/Sequence/Base.agda b/Cubical/Data/Sequence/Base.agda index fbedcf36b5..6d2521edde 100644 --- a/Cubical/Data/Sequence/Base.agda +++ b/Cubical/Data/Sequence/Base.agda @@ -11,11 +11,13 @@ private ℓ ℓ' ℓ'' : Level record Sequence (ℓ : Level) : Type (ℓ-suc ℓ) where + constructor sequence field obj : ℕ → Type ℓ map : {n : ℕ} → obj n → obj (1 + n) record SequenceMap (C : Sequence ℓ) (D : Sequence ℓ') : Type (ℓ-max ℓ ℓ') where + constructor sequencemap field map : (n : ℕ) → Sequence.obj C n → Sequence.obj D n comm : (n : ℕ) (x : Sequence.obj C n) diff --git a/Cubical/HITs/SequentialColimit/Base.agda b/Cubical/HITs/SequentialColimit/Base.agda index 0dcfad0daa..de65ed787e 100644 --- a/Cubical/HITs/SequentialColimit/Base.agda +++ b/Cubical/HITs/SequentialColimit/Base.agda @@ -31,6 +31,6 @@ FinSeqColim→Colim m (fpush n x i) = push x i realiseSequenceMap : {C : Sequence ℓ} {D : Sequence ℓ'} → SequenceMap C D → SeqColim C → SeqColim D -realiseSequenceMap record { map = map ; comm = comm } (incl x) = incl (map _ x) -realiseSequenceMap record { map = map ; comm = comm } (push {n = n} x i) = +realiseSequenceMap (sequencemap map comm) (incl x) = incl (map _ x) +realiseSequenceMap (sequencemap map comm) (push {n = n} x i) = (push (map n x) ∙ λ i → incl (comm n x i)) i diff --git a/Cubical/HITs/SequentialColimit/Properties.agda b/Cubical/HITs/SequentialColimit/Properties.agda index 7e5679f63e..6f2fa6a5a5 100644 --- a/Cubical/HITs/SequentialColimit/Properties.agda +++ b/Cubical/HITs/SequentialColimit/Properties.agda @@ -39,11 +39,13 @@ module _ incl∞ n = incl {n = n} record ElimData (P : SeqColim X → Type ℓ') : Type (ℓ-max ℓ ℓ') where + constructor elimdata field inclP : {k : ℕ} (x : X .obj k) → P (incl x) pushP : {k : ℕ} (x : X .obj k) → PathP (λ i → P (push x i)) (inclP x) (inclP (X .map x)) record ElimDataShifted (n : ℕ)(P : SeqColim X → Type ℓ') : Type (ℓ-max ℓ ℓ') where + constructor elimdata-shift field inclP : {k : ℕ} (x : X .obj (k + n)) → P (incl x) pushP : {k : ℕ} (x : X .obj (k + n)) → PathP (λ i → P (push x i)) (inclP x) (inclP (X .map x)) @@ -337,8 +339,7 @@ converges→ColimIso : ∀ {ℓ} {seq : Sequence ℓ} (n : ℕ) Iso.fun (converges→ColimIso {seq = seq} n e) = incl Iso.inv (converges→ColimIso {seq = seq} n e) = elimShifted seq n _ (shiftEqShifted seq n e) Iso.rightInv (converges→ColimIso {seq = seq} n e) = elimShifted seq n _ - (record { inclP = λ {k} → paths k - ; pushP = λ {k} → cohs k}) + (elimdata-shift (λ {k} → paths k) (λ {k} → cohs k)) where zero-case : (x : seq .obj n) → Path (SeqColim seq) diff --git a/Cubical/HITs/SphereBouquet/Degree.agda b/Cubical/HITs/SphereBouquet/Degree.agda index 0c7526597b..4f0a63d79b 100644 --- a/Cubical/HITs/SphereBouquet/Degree.agda +++ b/Cubical/HITs/SphereBouquet/Degree.agda @@ -16,6 +16,7 @@ open import Cubical.Data.Fin.Inductive open import Cubical.Data.Sigma open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat.Order.Inductive open import Cubical.HITs.S1 open import Cubical.HITs.Sn From eb6b12e6b4d9c7b00ffe4dec96d86f17320dd173 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 7 May 2024 10:12:34 +0200 Subject: [PATCH 39/73] Pointed --- Cubical/CW/Pointed.agda | 949 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 949 insertions(+) create mode 100644 Cubical/CW/Pointed.agda diff --git a/Cubical/CW/Pointed.agda b/Cubical/CW/Pointed.agda new file mode 100644 index 0000000000..366f5388c4 --- /dev/null +++ b/Cubical/CW/Pointed.agda @@ -0,0 +1,949 @@ +{-# OPTIONS --safe --lossy-unification #-} + +-- This file contains definition of CW complexes and skeleta. + +module Cubical.CW.Pointed where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Sigma +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence + +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Wedge + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + +open import Cubical.Relation.Nullary + +open import Cubical.CW.Base + +open Sequence + +open import Cubical.Foundations.Equiv.HalfAdjoint + + + +private + variable + ℓ ℓ' ℓ'' : Level + +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + + + + +module _ {C : Type ℓ} {B : Type ℓ'} where + PushoutAlongEquiv→ : {A : Type ℓ} + (e : A ≃ C) (f : A → B) → Pushout (fst e) f → B + PushoutAlongEquiv→ e f (inl x) = f (invEq e x) + PushoutAlongEquiv→ e f (inr x) = x + PushoutAlongEquiv→ e f (push a i) = f (retEq e a i) + + private + PushoutAlongEquiv→Cancel : {A : Type ℓ} (e : A ≃ C) (f : A → B) + → retract (PushoutAlongEquiv→ e f) inr + PushoutAlongEquiv→Cancel = + EquivJ (λ A e → (f : A → B) + → retract (PushoutAlongEquiv→ e f) inr) + λ f → λ { (inl x) → sym (push x) + ; (inr x) → refl + ; (push a i) j → push a (~ j ∨ i)} + + PushoutAlongEquiv : {A : Type ℓ} (e : A ≃ C) (f : A → B) + → Iso (Pushout (fst e) f) B + Iso.fun (PushoutAlongEquiv e f) = PushoutAlongEquiv→ e f + Iso.inv (PushoutAlongEquiv e f) = inr + Iso.rightInv (PushoutAlongEquiv e f) x = refl + Iso.leftInv (PushoutAlongEquiv e f) = PushoutAlongEquiv→Cancel e f + + +--- CW complexes --- +-- Definition of (the skeleton) of an arbitrary CW complex +-- New def: X 0 is empty and C (suc n) is pushout + +open import Cubical.HITs.Pushout +module _ {A : Type} {B : A → Pointed₀} (C : Pointed₀) (f : (⋁gen A B , inl tt) →∙ C) where + + + private + open 3x3-span + inst : 3x3-span + A00 inst = A + A02 inst = Σ A (fst ∘ B) + A04 inst = Σ A (fst ∘ B) + A20 inst = A + A22 inst = A + A24 inst = Σ A (fst ∘ B) + A40 inst = Unit + A42 inst = Unit + A44 inst = fst C + f10 inst = idfun A + f12 inst = λ x → x , snd (B x) + f14 inst = idfun _ + f30 inst = λ _ → tt + f32 inst = λ _ → tt + f34 inst = fst f ∘ inr + f01 inst = fst + f21 inst = idfun A + f41 inst = λ _ → tt + f03 inst = idfun _ + f23 inst = λ x → x , snd (B x) + f43 inst = λ _ → pt C + H11 inst = λ _ → refl + H13 inst = λ _ → refl + H31 inst = λ _ → refl + H33 inst = λ x → sym (snd f) ∙ cong (fst f) (push x) + + A0□Iso : Iso (A0□ inst) A + A0□Iso = compIso (equivToIso (symPushout _ _)) + (PushoutAlongEquiv (idEquiv _) fst) + + A2□Iso : Iso (A2□ inst) (Σ A (fst ∘ B)) + A2□Iso = PushoutAlongEquiv (idEquiv A) _ + + A4□Iso : Iso (A4□ inst) (fst C) + A4□Iso = PushoutAlongEquiv (idEquiv Unit) λ _ → pt C + + A○□Iso : Iso (A○□ inst) (Pushout (fst f ∘ inr) fst) + A○□Iso = compIso (equivToIso (symPushout _ _)) + (invIso (pushoutIso _ _ _ _ + (isoToEquiv (invIso A2□Iso)) + (isoToEquiv (invIso A4□Iso)) + (isoToEquiv (invIso A0□Iso)) + refl + λ i x → push x i)) + + + + A□0Iso : Iso (A□0 inst) Unit + A□0Iso = isContr→Iso + (inr tt , λ { (inl x) → sym (push x) + ; (inr x) → refl + ; (push a i) j → push a (i ∨ ~ j)}) + (tt , (λ _ → refl)) + + A□2Iso : Iso (A□2 inst) (⋁gen A B) + A□2Iso = equivToIso (symPushout _ _) + + + A□4Iso : Iso (A□4 inst) (fst C) + A□4Iso = PushoutAlongEquiv (idEquiv _) _ + + open import Cubical.Foundations.GroupoidLaws + A□○Iso : Iso (A□○ inst) (cofib (fst f)) + A□○Iso = invIso (pushoutIso _ _ _ _ + (isoToEquiv (invIso A□2Iso)) + (isoToEquiv (invIso A□0Iso)) + (isoToEquiv (invIso A□4Iso)) + (funExt (λ { (inl x) → refl + ; (inr x) → sym (push (fst x)) ∙ refl + ; (push a i) j → (sym (push a) ∙ refl) (i ∧ j)})) + (funExt λ { (inl x) i → inr (snd f i) + ; (inr x) → sym (push x) + ; (push a i) j + → hcomp (λ k + → (λ {(i = i0) → inr (compPath-filler' (sym (snd f)) + (cong (fst f) (push a)) j (~ k)) + ; (i = i1) → push (a , snd (B a)) (~ j) + ; (j = i0) → inr (fst f (push a (~ k ∨ i)))})) + (push (a , snd (B a)) (~ i ∨ ~ j))})) + + ⋁-cofib-Iso : Iso (Pushout (fst f ∘ inr) fst) (cofib (fst f)) + ⋁-cofib-Iso = compIso (compIso (invIso A○□Iso) + (invIso (3x3-Iso inst))) + A□○Iso + + + +-- connected comp +open import Cubical.Homotopy.Connected +open import Cubical.CW.Properties +open import Cubical.HITs.Truncation as TR +open import Cubical.Foundations.HLevels + +isConnectedCW→isConnectedSkel : (C : CWskel ℓ) + → isConnected 2 (realise C) + → (n : ℕ) + → isConnected 2 (C .fst (suc (suc n))) +isConnectedCW→isConnectedSkel C c n = + TR.rec (isOfHLevelSuc 1 isPropIsContr) + (λ ⋆ → TR.rec (isProp→isOfHLevelSuc (suc n) isPropIsContr) + (λ {(x , p) → ∣ x ∣ + , (TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (λ a → Iso.inv (PathIdTruncIso 1) + (TR.rec (isOfHLevelTrunc 1) + (λ q → + TR.rec (isOfHLevelPlus' 1 (isOfHLevelTrunc 1)) + (∣_∣ₕ ∘ fst) + (isConnectedCong (suc n) _ (isConnected-CW↪∞ (suc (suc n)) C) + q .fst)) + (Iso.fun (PathIdTruncIso 1) + (sym (c .snd ∣ incl x ∣) ∙ c .snd ∣ incl a ∣)))))}) + (isConnected-CW↪∞ (suc (suc n)) C ⋆ .fst)) + (c .fst) + +open import Cubical.Data.Bool +open import Cubical.Data.Nat.Order.Inductive +-- open import Cubical.Data.Dec + +open import Cubical.Relation.Nullary.Base +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Data.Fin.Inductive.Properties as Ind + +FinSuc : {n : ℕ} → Iso (Fin 1 ⊎ Fin n) (Fin (suc n)) +Iso.fun (FinSuc {n = n}) = ⊎.rec (λ _ → flast) injectSuc +Iso.inv (FinSuc {n = n}) = Ind.elimFin (inl flast) inr +Iso.rightInv (FinSuc {n = n}) = + Ind.elimFin (cong (⊎.rec (λ _ → flast) injectSuc) + (Ind.elimFinβ (inl flast) inr .fst)) + λ x → cong (⊎.rec (λ _ → flast) injectSuc) + (Ind.elimFinβ (inl flast) inr .snd x) +Iso.leftInv (FinSuc {n = n}) (inl (zero , p)) = + Ind.elimFinβ (inl flast) inr .fst +Iso.leftInv (FinSuc {n = n}) (inr x) = Ind.elimFinβ (inl flast) inr .snd x + +Fin+ : {n m : ℕ} → Iso (Fin n ⊎ Fin m) (Fin (n +ℕ m)) +Iso.fun (Fin+ {n = zero} {m = m}) (inr x) = x +Iso.inv (Fin+ {n = zero} {m = m}) x = inr x +Iso.rightInv (Fin+ {n = zero} {m = m}) x = refl +Iso.leftInv (Fin+ {n = zero} {m = m}) (inr x) = refl +Fin+ {n = suc n} {m = m} = + compIso (⊎Iso (invIso FinSuc) idIso) + (compIso ⊎-assoc-Iso + (compIso (⊎Iso idIso (Fin+ {n = n} {m = m})) + FinSuc)) + +Iso-Unit-Fin : Iso Unit (Fin 1) +Iso.fun Iso-Unit-Fin tt = fzero +Iso.inv Iso-Unit-Fin (x , p) = tt +Iso.rightInv Iso-Unit-Fin (zero , p) = Σ≡Prop (λ _ → isProp<ᵗ) refl +Iso.leftInv Iso-Unit-Fin x = refl + +Iso-Bool-Fin : Iso (S₊ 0) (Fin 2) +Iso.fun Iso-Bool-Fin false = flast +Iso.fun Iso-Bool-Fin true = fzero +Iso.inv Iso-Bool-Fin (zero , p) = true +Iso.inv Iso-Bool-Fin (suc x , p) = false +Iso.rightInv Iso-Bool-Fin (zero , p) = refl +Iso.rightInv Iso-Bool-Fin (suc zero , p) = + Σ≡Prop (λ _ → isProp<ᵗ) refl +Iso.leftInv Iso-Bool-Fin false = refl +Iso.leftInv Iso-Bool-Fin true = refl + +Fin× : {n m : ℕ} → Iso (Fin n × Fin m) (Fin (n · m)) +Fin× {n = zero} {m = m} = + iso (λ {()}) (λ{()}) (λ{()}) (λ{()}) +Fin× {n = suc n} {m = m} = + compIso + (compIso + (compIso (Σ-cong-iso-fst (invIso FinSuc)) + (compIso Σ-swap-Iso + (compIso ×DistR⊎Iso + (⊎Iso (compIso + (Σ-cong-iso-snd (λ _ → invIso Iso-Unit-Fin)) rUnit×Iso) + Σ-swap-Iso)))) + (⊎Iso idIso Fin×)) + (Fin+ {n = m} {n · m}) + +DiscreteFin : ∀ {n} → Discrete (Fin n) +DiscreteFin x y with discreteℕ (fst x) (fst y) +... | yes p = yes (Σ≡Prop (λ _ → isProp<ᵗ) p) +... | no ¬p = no λ q → ¬p (cong fst q) + + +decIm : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') +decIm {A = A} {B = B} f = + (y : B) → (Σ[ x ∈ A ] f x ≡ y) + ⊎ ((x : A) → ¬ f x ≡ y) + +decImFin : ∀ {ℓ} {A : Type ℓ} + (da : Discrete A) (n : ℕ) (f : Fin n → A) + → decIm f +decImFin {A = A} da zero f y = inr λ x → ⊥.rec (¬Fin0 x) +decImFin {A = A} da (suc n) f y with da (f fzero) y +... | yes p = inl (fzero , p) +... | no ¬p with (decImFin da n (f ∘ fsuc) y) +... | inl q = inl ((fsuc (fst q)) , snd q) +... | inr q = inr (Ind.elimFin-alt ¬p q) + +Bool×Fin≡Fin : {n : ℕ} → Iso (S₊ 0 × Fin n) (Fin (2 · n)) +Bool×Fin≡Fin = compIso (Σ-cong-iso-fst Iso-Bool-Fin) (Fin× {n = 2}) + +decIm-S⁰×Fin : ∀ {ℓ} {A : Type ℓ} + (da : Discrete A) (n : ℕ) (f : S₊ 0 × Fin n → A) → decIm f +decIm-S⁰×Fin {A = A} da n = + subst (λ C → (f : C → A) → decIm f) + (isoToPath (invIso Bool×Fin≡Fin)) + (decImFin da _) + + +module _ {ℓ : Level} {n : ℕ} {A : Fin n → Type ℓ} (x₀ : Fin n) + (pt : A x₀) (l : (x : Fin n) → ¬ x ≡ x₀ → A x) where + private + x = fst x₀ + p = snd x₀ + elimFinChoose : (x : _) → A x + elimFinChoose (x' , q) with (discreteℕ x x') + ... | yes p₁ = subst A (Σ≡Prop (λ _ → isProp<ᵗ) p₁) pt + ... | no ¬p = l (x' , q) λ r → ¬p (sym (cong fst r)) + + elimFinChooseβ : (elimFinChoose x₀ ≡ pt) + × ((x : _) (q : ¬ x ≡ x₀) → elimFinChoose x ≡ l x q) + fst elimFinChooseβ with (discreteℕ x x) + ... | yes p₁ = (λ j → subst A (isSetFin x₀ x₀ (Σ≡Prop (λ z → isProp<ᵗ) p₁) refl j) pt) + ∙ transportRefl pt + ... | no ¬p = ⊥.rec (¬p refl) + snd elimFinChooseβ (x' , q) with (discreteℕ x x') + ... | yes p₁ = λ q → ⊥.rec (q (Σ≡Prop (λ _ → isProp<ᵗ) (sym p₁))) + ... | no ¬p = λ s → cong (l (x' , q)) (isPropΠ (λ _ → isProp⊥) _ _) + + +pickDifferentFin : {n : ℕ} (x : Fin (suc (suc n))) → Σ[ y ∈ Fin (suc (suc n)) ] ¬ x ≡ y +pickDifferentFin (zero , p) = (1 , p) , λ q → snotz (sym (cong fst q)) +pickDifferentFin (suc x , p) = fzero , λ q → snotz (cong fst q) + +allConst? : ∀ {ℓ} {A : Type ℓ} {n : ℕ} (dis : Discrete A) + → (t : Fin n → S₊ 0 → A) + → ((x : Fin n) → t x ≡ λ _ → t x true) + ⊎ (Σ[ x ∈ Fin n ] ¬ (t x true ≡ t x false)) +allConst? {n = zero} dis t = inl λ {()} +allConst? {n = suc n} dis t + with (dis (t fzero true) (t fzero false)) + | (allConst? {n = n} dis λ x p → t (fsuc x) p) +... | yes p | inl x = + inl (Ind.elimFin-alt (funExt + (λ { false → sym p ; true → refl})) x) +... | yes p | inr x = inr (_ , (snd x)) +... | no ¬p | q = inr (_ , ¬p) + +CW₁Data' : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) + → isConnected 2 (Pushout f snd) + → Σ[ k ∈ ℕ ] Σ[ f' ∈ (S₊ 0 × Fin k → Fin (suc m)) ] + Iso (Pushout f snd) + (Pushout f' snd) +CW₁Data' zero m f c = ⊥.rec {!!} +CW₁Data' (suc zero) m f c = {!!} +CW₁Data' (suc (suc n)) m f c = {!!} + + +isSurj-α₀ : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) + → isConnected 2 (Pushout f snd) + → (y : _) → Σ[ x ∈ _ ] f x ≡ y +isSurj-α₀ n m f c y with (decIm-S⁰×Fin DiscreteFin n f y) +... | inl x = x +isSurj-α₀ n m f c x₀ | inr q = ⊥.rec nope + where + pre-help : Fin (suc (suc m)) → Type + pre-help = elimFinChoose x₀ ⊥ λ _ _ → Unit + + lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit + lem = elimFinChoose x₀ + (λ s t → ⊥.rec (q _ t)) + λ s t b c → elimFinChooseβ x₀ ⊥ (λ _ _ → Unit) .snd s t + + help : Pushout f snd → Type + help (inl x) = pre-help x + help (inr x) = Unit + help (push a i) = lem (f a) a refl i + + nope : ⊥ + nope = TR.rec isProp⊥ + (λ q → transport (elimFinChooseβ x₀ ⊥ (λ _ _ → Unit) .fst) + (subst help (sym q) + (transport (sym (elimFinChooseβ x₀ ⊥ + (λ _ _ → Unit) .snd _ + (pickDifferentFin x₀ .snd ∘ sym))) tt))) + (Iso.fun (PathIdTruncIso 1) + (isContr→isProp c + ∣ inl x₀ ∣ + ∣ inl (pickDifferentFin x₀ .fst) ∣)) + + +notAllLoops-α₀ : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) + → isConnected 2 (Pushout f snd) + → Σ[ x ∈ Fin n ] (¬ f (true , x) ≡ f (false , x)) +notAllLoops-α₀ n m f c with (allConst? DiscreteFin (λ x y → f (y , x))) +... | inr x = x +notAllLoops-α₀ n m f c | inl q = + ⊥.rec (TR.rec isProp⊥ (λ p → subst T p tt) + (Iso.fun(PathIdTruncIso 1) + (isContr→isProp c + ∣ inl flast ∣ ∣ inl fzero ∣))) + where + inrT : Fin n → Type + inrT x with (DiscreteFin (f (true , x)) fzero) + ... | yes p = ⊥ + ... | no ¬p = Unit + + inlT : Fin (suc (suc m)) → Type + inlT (zero , p) = ⊥ + inlT (suc x , p) = Unit + + inlrT-pre : (a : _) → inlT (f (true , a)) ≡ inrT a + inlrT-pre a with ((DiscreteFin (f (true , a)) fzero)) + ... | yes p = cong inlT p + inlrT-pre s | no ¬p with (f (true , s)) + ... | zero , tt = ⊥.rec (¬p refl) + ... | suc q , t = refl + + inlrT : (a : _) → inlT (f a) ≡ inrT (snd a) + inlrT (false , b) = cong inlT (funExt⁻ (q b) false) ∙ inlrT-pre b + inlrT (true , b) = inlrT-pre b + + T : Pushout f snd → Type + T (inl x) = inlT x + T (inr x) = inrT x + T (push a i) = inlrT a i + +module _ {n : ℕ} where + swapFin : (x y : Fin n) → Fin n → Fin n + swapFin (x , xp) (y , yp) (z , zp) with (discreteℕ z x) | (discreteℕ z y) + ... | yes p | yes p₁ = z , zp + ... | yes p | no ¬p = y , yp + ... | no ¬p | yes p = x , xp + ... | no ¬p | no ¬p₁ = (z , zp) + + + + -- swapFinSwap : (x y z : Fin n) → swapFin x y z ≡ swapFin y x z + -- swapFinSwap x y z with (discreteℕ (fst z) (fst x)) | discreteℕ (fst z) (fst y) + -- ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p₁ ∙ p) + -- ... | yes p | no ¬p = {!help!} + -- where + -- help : y ≡ swapFin y x z + -- help = {!!} + -- ... | no ¬p | q = {!!} + + swapFin² : (x y z : Fin n) → swapFin x y (swapFin x y z) ≡ z + swapFin² (x , xp) (y , yp) (z , zp) with discreteℕ z x | discreteℕ z y + ... | yes p | yes p₁ = silly + where + silly : swapFin (x , xp) (y , yp) (z , zp) ≡ (z , zp) + silly with discreteℕ z x | discreteℕ z y + ... | yes p | yes p₁ = refl + ... | yes p | no ¬p = ⊥.rec (¬p p₁) + ... | no ¬p | r = ⊥.rec (¬p p) + ... | yes p | no ¬q = silly + where + silly : swapFin (x , xp) (y , yp) (y , yp) ≡ (z , zp) + silly with discreteℕ y x | discreteℕ y y + ... | yes p' | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (p' ∙ sym p) + ... | no ¬p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p) + ... | p | no ¬p = ⊥.rec (¬p refl) + ... | no ¬p | yes p = silly + where + silly : swapFin (x , xp) (y , yp) (x , xp) ≡ (z , zp) + silly with discreteℕ x y | discreteℕ x x + ... | yes p₁ | yes _ = Σ≡Prop (λ _ → isProp<ᵗ) (p₁ ∙ sym p) + ... | no ¬p | yes _ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p) + ... | s | no ¬p = ⊥.rec (¬p refl) + ... | no ¬p | no ¬q = silly + where + silly : swapFin (x , xp) (y , yp) (z , zp) ≡ (z , zp) + silly with discreteℕ z x | discreteℕ z y + ... | yes p | yes p₁ = refl + ... | yes p | no ¬b = ⊥.rec (¬p p) + ... | no ¬a | yes b = ⊥.rec (¬q b) + ... | no ¬a | no ¬b = refl + + swapFinIso : (x y : Fin n) → Iso (Fin n) (Fin n) + Iso.fun (swapFinIso x y) = swapFin x y + Iso.inv (swapFinIso x y) = swapFin x y + Iso.rightInv (swapFinIso x y) = swapFin² x y + Iso.leftInv (swapFinIso x y) = swapFin² x y + +Pushout∘Equiv : ∀ {ℓ ℓ' ℓ''} {A A' : Type ℓ} {B B' : Type ℓ'} {C : Type ℓ''} + (e1 : A ≃ A') (e2 : B' ≃ B) + (f : A' → B') (g : A → C) + → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) (Pushout f (g ∘ invEq e1)) +Pushout∘Equiv {A = A} {A' = A'} {B} {B'} {C} = + EquivJ (λ A e1 → (e2 : B' ≃ B) (f : A' → B') (g : A → C) + → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) + (Pushout f (g ∘ invEq e1))) + (EquivJ (λ B' e2 → (f : A' → B') (g : A' → C) + → Iso (Pushout (fst e2 ∘ f) g) + (Pushout f g)) + λ f g → idIso) + +module _ {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : Type ℓ'} + (f : Bool × A → Unit ⊎ B) (dB : Discrete B) (dA : Discrete A) (b₀ : B) where + + F : Bool × (Unit ⊎ A) → Unit ⊎ B + F (false , inl tt) = inl tt + F (true , inl tt) = inr b₀ + F (x , inr a) = f (x , a) + + g : Unit ⊎ B → B + g (inl x) = b₀ + g (inr x) = x + + PF→P∘ₗ : (x : Unit ⊎ A) → Pushout (g ∘ f) snd + PF→P∘ₗ (inl x) = inl b₀ + PF→P∘ₗ (inr x) = inr x + + theCoh : (a : _) → inl (g (F a)) ≡ PF→P∘ₗ (snd a) + theCoh (false , inl x) = refl + theCoh (true , inl x) = refl + theCoh (false , inr x) = push (false , x) + theCoh (true , inr x) = push (true , x) + + PF→P∘' : Pushout F snd → Pushout (g ∘ f) snd + PF→P∘' (inl x) = inl (g x) + PF→P∘' (inr x) = PF→P∘ₗ x + PF→P∘' (push a i) = theCoh a i + + theCoh2 : (a : _) (b : _) + → Path (Pushout F snd) (inl (inr (g (f (b , a))))) (inl (f (b , a))) + theCoh2 a b with (f (b , a)) + theCoh2 a b | inl x = (push (true , inl tt)) ∙ sym (push (false , inl tt)) + ... | inr x = refl + + P∘'→PF : Pushout (g ∘ f) snd → Pushout F snd + P∘'→PF (inl x) = inl (inr x) + P∘'→PF (inr x) = inr (inr x) + P∘'→PF (push (false , a) i) = (theCoh2 a false ∙ push (false , inr a)) i + P∘'→PF (push (true , a) i) = (theCoh2 a true ∙ push (true , inr a)) i + + PFpushT : (x : _) → P∘'→PF (PF→P∘' x) ≡ x + PFpushT (inl (inl x)) = push (true , inl tt) ∙ sym (push (false , inl tt)) + PFpushT (inl (inr x)) = refl + PFpushT (inr (inl x)) = push (true , inl tt) + PFpushT (inr (inr x)) = refl + PFpushT (push (false , inl x) i) j = + compPath-filler (push (true , inl tt)) (sym (push (false , inl tt))) (~ i) j + PFpushT (push (false , inr x) i) j = {!!} + PFpushT (push (true , inl x) i) j = push (true , inl x) (i ∧ j) + PFpushT (push (true , inr x) i) j = {!!} + + PF→P∘ : Pushout F snd → Pushout (g ∘ f) snd + PF→P∘ (inl (inl x)) = inl b₀ + PF→P∘ (inl (inr x)) = inl x + PF→P∘ (inr (inl x)) = inl b₀ + PF→P∘ (inr (inr x)) = inr x + PF→P∘ (push (false , inl x) i) = inl b₀ + PF→P∘ (push (true , inl x) i) = inl b₀ + PF→P∘ (push (false , inr x) i) = {!!} + PF→P∘ (push (true , inr x) i) = {!!} + + Is1 : Iso (Pushout F snd) (Pushout (g ∘ f) snd) + Is1 = {!!} + + + + +TheF : {!!} +TheF = {!!} + +FinPred : ∀ {m} → Fin (suc (suc m)) → Fin (suc m) +FinPred {m = m} (zero , p) = zero , p +FinPred {m = m} (suc x , p) = x , p + +fone : ∀ {m} → Fin (suc (suc m)) +fone {m} = suc zero , tt + + + +CW₁DataPre : (n m : ℕ) (f : S₊ 0 × Fin (suc (suc n)) → Fin (suc (suc m))) + → f (true , fzero) ≡ fzero + → f (false , fzero) ≡ fone + → Σ[ k ∈ ℕ ] Σ[ f' ∈ (S₊ 0 × Fin k → Fin (suc m)) ] + Iso (Pushout f snd) + (Pushout f' snd) +CW₁DataPre n m f p q = (suc n) + , F↓ + , (compIso (subst (λ f → Iso (Pushout f snd) + (Pushout f* snd)) (funExt f*≡) idIso) + {!!}) + where + f* : S₊ 0 × Fin (suc (suc n)) → Fin (suc (suc m)) + f* (false , zero , b) = fone + f* (true , zero , b) = fzero + f* (x , suc a , b) = f (x , suc a , b) + + f↓ : S₊ 0 × Fin (suc (suc n)) → Fin (suc (suc m)) + f↓ (x , p) with DiscreteFin (f (x , p)) fzero + ... | yes p₁ = fone + ... | no ¬p = f (x , p) + + F↓ : Bool × Fin (suc n) → Fin (suc m) + F↓ (x , p) = FinPred (f↓ (x , fsuc p)) + + FinPred-lem : (b : _) (x : _) (p : _) + → FinPred (f (b , suc x , p)) ≡ F↓ (b , x , p) + FinPred-lem b x p with (DiscreteFin (f (b , fsuc (x , p))) fzero) + ... | yes p₁ = cong FinPred p₁ + ... | no ¬p = refl + + FinPred-lem2 : (b : _) (x : _) (p : _) → FinPred (f* (b , suc x , p)) ≡ F↓ (b , FinPred (suc x , p)) + FinPred-lem2 false x r = FinPred-lem false x r + FinPred-lem2 true x p = FinPred-lem true x p + + slask : (b : Bool) + → Path (Pushout F↓ (λ r → snd r)) + (inl (FinPred (f* (b , zero , tt)))) + (inl (F↓ (b , zero , tt))) + slask b with DiscreteFin (f (b , fsuc (0 , tt))) (0 , tt) + slask false | yes p = refl + slask true | yes p = refl + slask false | no ¬p = {!!} ∙ {!!} + slask true | no ¬p = {!¬p !} + + f*≡ : (x : _ ) → f* x ≡ f x + f*≡ (false , zero , z) = sym q + f*≡ (true , zero , z) = sym p + f*≡ (false , suc y , z) = refl + f*≡ (true , suc y , z) = refl + + pp : (b : Bool) (x : Fin (suc (suc n))) + → Path (Pushout F↓ (λ r → snd r)) + (inl (FinPred (f* (b , x)))) (inr (FinPred x)) + pp b (zero , p) = ({!!} ∙ push (b , fzero)) + pp b (suc x , p) = (λ i → inl (FinPred-lem2 b x p i)) ∙ push (b , x , p) + -- ((λ i → inl (FinPred-lem2 b x p i)) ∙ push (b , FinPred (x , p))) + + Pf*→ : Pushout f* snd → Pushout F↓ snd + Pf*→ (inl x) = inl (FinPred x) + Pf*→ (inr x) = inr (FinPred x) + Pf*→ (push (b , zero , q) i) = {!!} + Pf*→ (push (b , suc x' , q) i) = {!!} + + {- + main : Iso _ _ + Iso.fun main (inl x) = inl (FinPred x) + Iso.fun main (inr x) = inr (FinPred x) + Iso.fun main (push (false , zero , s) i) = ({!!} ∙∙ push {!!} ∙∙ {!!}) i + Iso.fun main (push (true , zero , s) i) = {!!} + Iso.fun main (push (b , suc a , s) i) = push (b , a , s) i + Iso.inv main x = {!!} + Iso.rightInv main x = {!!} + Iso.leftInv main x = {!!} + -} + +CW₁Data : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) + → isConnected 2 (Pushout f snd) + → Σ[ k ∈ ℕ ] Σ[ f' ∈ (S₊ 0 × Fin k → Fin (suc m)) ] + Iso (Pushout f snd) + (Pushout f' snd) +CW₁Data zero m f c = ⊥.rec (snd (notAllLoops-α₀ zero m f c .fst)) +CW₁Data (suc n) m f c = n , (λ x → {!xpath!}) , {!RetrFin!} + -- n , {!!} , (compIso {!!} {!!}) + where + t = notAllLoops-α₀ (suc n) m f c + + abstract + x₀ : Fin (suc n) + x₀ = fst t + + xpath : ¬ (f (true , x₀) ≡ f (false , x₀)) + xpath = snd t + + Fin0-iso : Iso (S₊ 0 × Fin (suc n)) (S₊ 0 × Fin (suc n)) + Fin0-iso = Σ-cong-iso-snd λ _ → swapFinIso x₀ fzero + + FinIso2 : Iso (Fin (suc (suc m))) (Fin (suc (suc m))) + FinIso2 = {!!} + + Upath = isoToPath (swapFinIso x₀ fzero) + + f' : S₊ 0 × Fin (suc n) → Fin (suc (suc m)) + f' (x , zero , p) = f (x , x₀) + f' (x , suc b , p) with (discreteℕ (suc b) (fst x₀)) + ... | yes q = f (x , fzero) + ... | no ¬q = f (x , suc b , p) + + f'≡ : (x : _) → f (fst x , (swapFin x₀ fzero (snd x))) ≡ f' x + f'≡ (x , zero , p) with (discreteℕ 0 (fst x₀)) + ... | yes p₁ = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) p₁)) + ... | no ¬p = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) refl)) + f'≡ (x , suc b , p) with (discreteℕ (suc b) (fst x₀)) + ... | yes p₁ = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) refl)) + ... | no ¬p = refl + + Pushout-f≡ : (Pushout f snd) ≡ (Pushout f' snd) + Pushout-f≡ = (λ i → Pushout {A = S₊ 0 × Upath i} {B = Fin (suc (suc m))} {C = Upath i} (help i) snd) + ∙ (λ i → Pushout (λ x → f'≡ x i) snd) + where + help : PathP (λ i → Bool × Upath i → Fin (suc (suc m))) + f + (f ∘ (λ a → fst a , (swapFin x₀ fzero (snd a)))) + help = toPathP (funExt λ p → transportRefl _ + ∙ cong f (λ j → fst p , transp (λ j → Upath (~ j)) i0 (snd p)) + ∙ cong f (cong (fst p ,_) + (cong (Iso.fun (swapFinIso x₀ fzero)) (transportRefl _)))) + + +{- + F : Σ[ x ∈ Fin n ] (¬ f (true , x) ≡ f (false , x)) + F with (allConst? DiscreteFin (λ x y → f (y , x))) + ... | inr x = x + ... | inl qs = {!!} + where + pre-help : Fin (suc (suc m)) → Type + pre-help (zero , p) = ⊥ + pre-help (suc x , p) = Unit + + lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit + lem (zero , q) (false , t) p = {!!} + lem (zero , q) (true , t) p = {!!} + lem (suc x , q) a p = refl + + help : Pushout f snd → Type + help (inl x) = {!!} -- pre-help x + help (inr x) = {!x!} -- Unit + help (push a i) = {!!} -- lem (f a) a refl i + -} + +-- pre-help : Fin (suc (suc m)) → Type +-- pre-help (zero , p) = ⊥ +-- pre-help (suc x , p) = Unit + +-- lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit +-- lem (zero , q) a p = +-- ⊥.rec (x a (p ∙ Σ≡Prop (λ _ → isProp<ᵗ) refl)) +-- lem (suc n , q) a p = refl + +-- help : Pushout f snd → Type +-- help (inl x) = pre-help x +-- help (inr x) = Unit +-- help (push a i) = lem (f a) a refl i + +-- Pf≅⊥ : ⊥ +-- Pf≅⊥ = TR.rec isProp⊥ +-- (λ p → subst help (sym p) tt) +-- (Iso.fun (PathIdTruncIso 1) +-- (isContr→isProp c ∣ inl fzero ∣ ∣ inl flast ∣)) + +-- -- where +-- -- fSurj : (x : Fin _) → Σ[ r ∈ _ ] f r ≡ x +-- -- fSurj x = {! -- Fin→Surj? ? ? ? ?!} + +-- -- ptt : Pushout f snd → Type +-- -- ptt (inl x) = Fin (suc (suc m)) +-- -- ptt (inr x) = {!Fin n!} +-- -- ptt (push a i) = {!!} + +-- -- module _ (s : {!!}) where + + + +-- -- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ +-- -- yieldsCWskel∙ X = +-- -- Σ[ f ∈ (ℕ → ℕ) ] +-- -- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] +-- -- ((X zero ≃ Fin (f zero)) × +-- -- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) + +-- -- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) + +-- -- module CWskel∙-fields (C : CWskel∙ ℓ) where +-- -- card = C .snd .fst +-- -- A = Fin ∘ card +-- -- α = C .snd .snd .fst +-- -- e = C .snd .snd .snd .snd + +-- -- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero +-- -- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] + +-- -- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- -- fst (CWpt (C , f) n) = C n +-- -- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) + +-- -- -- Technically, a CW complex should be the sequential colimit over the following maps +-- -- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) +-- -- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) + +-- -- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n +-- -- ptCW T zero = T .snd .snd .fst zero (inl tt) +-- -- ptCW T (suc n) = CW∙↪ T n (ptCW T n) + +-- -- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- -- CW∙ T n = fst T n , ptCW T n + +-- -- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) +-- -- fst (CW∙↪∙ T n) = CW∙↪ T n +-- -- snd (CW∙↪∙ T n) = refl + + +-- -- -- But if it stabilises, no need for colimits. +-- -- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ +-- -- yieldsFinCWskel∙ n X = +-- -- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) + +-- -- -- ... which should give us finite CW complexes. +-- -- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) +-- -- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) + +-- -- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ +-- -- finCWskel→CWskel∙ n C = fst C , fst (snd C) + +-- -- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ +-- -- Sequence.obj (realiseSeq∙ (C , p)) = C +-- -- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ + +-- -- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ +-- -- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) + +-- -- -- realisation of CW complex from skeleton +-- -- realise∙ : CWskel∙ ℓ → Type ℓ +-- -- realise∙ C = SeqColim (realiseSeq∙ C) + +-- -- realise∙∙ : CWskel∙ ℓ → Pointed ℓ +-- -- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) +-- -- open import Cubical.Data.Empty as ⊥ + +-- -- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ +-- -- CWskel∙→CWskel A zero = Lift ⊥ +-- -- CWskel∙→CWskel A (suc n) = A n +-- -- open import Cubical.Foundations.Isomorphism + + +-- -- module _ (A : ℕ → Type ℓ) +-- -- (cwsk : yieldsCWskel∙ A) where + +-- -- private +-- -- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n +-- -- αs (suc n) x = snd cwsk .fst n (inr x) + +-- -- e0 : {!!} +-- -- e0 = {!!} + +-- -- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst +-- -- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) +-- -- es-suc→ n (inr x) = inl x +-- -- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) +-- -- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i +-- -- es-suc→ n (push (push a i₁) i) = {!!} + +-- -- es-suc : (n : ℕ) +-- -- → Iso (cofib (fst (snd cwsk) n)) +-- -- (Pushout (αs (suc n)) fst) +-- -- Iso.fun (es-suc n) = es-suc→ n +-- -- Iso.inv (es-suc n) = {!!} +-- -- Iso.rightInv (es-suc n) = {!!} +-- -- Iso.leftInv (es-suc n) = {!!} + +-- -- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) +-- -- es zero = {!!} +-- -- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) + +-- -- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) +-- -- fst yieldsCWskel∙→' = cwsk .fst +-- -- fst (snd yieldsCWskel∙→') = αs +-- -- fst (snd (snd yieldsCWskel∙→')) () +-- -- snd (snd (snd yieldsCWskel∙→')) = {!!} + +-- -- yieldsCWskel∙→ : (A : ℕ → Type ℓ) +-- -- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) +-- -- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst +-- -- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) +-- -- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () +-- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} +-- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = +-- -- compEquiv (cwsk .snd .snd .snd n) +-- -- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) +-- -- where +-- -- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst +-- -- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) +-- -- theEq→ (inr x) = inl x +-- -- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) +-- -- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) +-- -- theEq→ (push (push a i₁) i) = {!!} + +-- -- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) +-- -- Iso.fun theEq = theEq→ +-- -- Iso.inv theEq = {!!} +-- -- Iso.rightInv theEq x = {!!} +-- -- Iso.leftInv theEq x = {!!} + + +-- -- -- compEquiv {!!} {!!} + + +-- -- -- -- Finally: definition of CW complexes +-- -- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' + +-- -- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ + +-- -- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) + +-- -- -- -- Finite CW complexes +-- -- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- -- isFinCW {ℓ = ℓ} X = +-- -- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) + +-- -- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ + +-- -- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) + +-- -- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X +-- -- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str + +-- -- -- finCW→CW : finCW ℓ → CW ℓ +-- -- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p + + +-- -- -- -- morphisms +-- -- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') +-- -- -- C →ᶜʷ D = fst C → fst D + +-- -- -- --the cofibre of the inclusion of X n into X (suc n) +-- -- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ +-- -- -- cofibCW n C = cofib (CW↪ C n) + +-- -- -- --...is basically a sphere bouquet +-- -- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) +-- -- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) +-- -- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e +-- -- -- where +-- -- -- s = Bouquet≃-gen +-- -- -- α = C .snd .snd .fst +-- -- -- e = C .snd .snd .snd .snd n + +-- -- -- --sending X (suc n) into the cofibCW +-- -- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C +-- -- -- to_cofibCW n C x = inr x + +-- -- -- --the connecting map of the long exact sequence +-- -- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) +-- -- -- → cofib conn → Susp A +-- -- -- δ-pre conn (inl x) = north +-- -- -- δ-pre conn (inr x) = south +-- -- -- δ-pre conn (push a i) = merid a i + +-- -- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) +-- -- -- δ n C = δ-pre (CW↪ C n) + +-- -- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) +-- -- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C +-- -- -- CW↪∞ C n x = incl x + +-- -- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m +-- -- -- fst (finCW↑ m n p C) = fst C +-- -- -- fst (snd (finCW↑ m n p C)) = snd C .fst +-- -- -- snd (snd (finCW↑ m n p C)) k = +-- -- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) +-- -- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) +-- -- -- (snd C .snd (k +ℕ fst p)) From 87fa4c021063e4ec37fa39d4b7bfe78e2c44a881 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 7 May 2024 18:15:55 +0200 Subject: [PATCH 40/73] done? --- Cubical/CW/Pointed.agda | 983 ++++++++++++++++++++++++---------------- 1 file changed, 596 insertions(+), 387 deletions(-) diff --git a/Cubical/CW/Pointed.agda b/Cubical/CW/Pointed.agda index 366f5388c4..4018fbf796 100644 --- a/Cubical/CW/Pointed.agda +++ b/Cubical/CW/Pointed.agda @@ -34,8 +34,6 @@ open Sequence open import Cubical.Foundations.Equiv.HalfAdjoint - - private variable ℓ ℓ' ℓ'' : Level @@ -44,9 +42,6 @@ open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence - - - module _ {C : Type ℓ} {B : Type ℓ'} where PushoutAlongEquiv→ : {A : Type ℓ} (e : A ≃ C) (f : A → B) → Pushout (fst e) f → B @@ -333,15 +328,6 @@ allConst? {n = suc n} dis t ... | yes p | inr x = inr (_ , (snd x)) ... | no ¬p | q = inr (_ , ¬p) -CW₁Data' : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) - → isConnected 2 (Pushout f snd) - → Σ[ k ∈ ℕ ] Σ[ f' ∈ (S₊ 0 × Fin k → Fin (suc m)) ] - Iso (Pushout f snd) - (Pushout f' snd) -CW₁Data' zero m f c = ⊥.rec {!!} -CW₁Data' (suc zero) m f c = {!!} -CW₁Data' (suc (suc n)) m f c = {!!} - isSurj-α₀ : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) → isConnected 2 (Pushout f snd) @@ -420,7 +406,17 @@ module _ {n : ℕ} where ... | no ¬p | yes p = x , xp ... | no ¬p | no ¬p₁ = (z , zp) + swapFinβₗ : (x y : Fin n) → swapFin x y x ≡ y + swapFinβₗ (x , xp) (y , yp) with (discreteℕ x x) | discreteℕ x y + ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) p₁ + ... | yes p | no ¬p = refl + ... | no ¬p | q = ⊥.rec (¬p refl) + swapFinβᵣ : (x y : Fin n) → swapFin x y y ≡ x + swapFinβᵣ (x , xp) (y , yp) with (discreteℕ y y) | discreteℕ y x + ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) p₁ + ... | yes p | no ¬p = refl + ... | no ¬p | q = ⊥.rec (¬p refl) -- swapFinSwap : (x y z : Fin n) → swapFin x y z ≡ swapFin y x z -- swapFinSwap x y z with (discreteℕ (fst z) (fst x)) | discreteℕ (fst z) (fst y) @@ -482,8 +478,8 @@ Pushout∘Equiv {A = A} {A' = A'} {B} {B'} {C} = (Pushout f g)) λ f g → idIso) -module _ {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : Type ℓ'} - (f : Bool × A → Unit ⊎ B) (dB : Discrete B) (dA : Discrete A) (b₀ : B) where +module _ {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} + (f : Bool × A → Unit ⊎ B) (b₀ : B) where F : Bool × (Unit ⊎ A) → Unit ⊎ B F (false , inl tt) = inl tt @@ -521,35 +517,68 @@ module _ {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : Type ℓ'} P∘'→PF (push (false , a) i) = (theCoh2 a false ∙ push (false , inr a)) i P∘'→PF (push (true , a) i) = (theCoh2 a true ∙ push (true , inr a)) i + PFpushTₗ : (x : _) → P∘'→PF (PF→P∘' (inl x)) ≡ (inl x) + PFpushTₗ (inl x) = push (true , inl tt) ∙ sym (push (false , inl tt)) + PFpushTₗ (inr x) = refl + + PFpushTᵣ : (x : _) → P∘'→PF (PF→P∘' (inr x)) ≡ (inr x) + PFpushTᵣ (inl x) = push (true , inl tt) + PFpushTᵣ (inr x) = refl + + pp1 : (x : A) → PFpushTₗ (f (false , x)) ≡ theCoh2 x false + pp1 x with (f (false , x)) + ... | inl x₁ = refl + ... | inr x₁ = refl + + pp2 : (x : A) → PFpushTₗ (f (true , x)) ≡ theCoh2 x true + pp2 x with (f (true , x)) + ... | inl x₁ = refl + ... | inr x₁ = refl + + open import Cubical.Foundations.Path + open import Cubical.Foundations.GroupoidLaws + + PFpushT : (x : _) → P∘'→PF (PF→P∘' x) ≡ x - PFpushT (inl (inl x)) = push (true , inl tt) ∙ sym (push (false , inl tt)) - PFpushT (inl (inr x)) = refl - PFpushT (inr (inl x)) = push (true , inl tt) - PFpushT (inr (inr x)) = refl + PFpushT (inl x) = PFpushTₗ x + PFpushT (inr x) = PFpushTᵣ x PFpushT (push (false , inl x) i) j = compPath-filler (push (true , inl tt)) (sym (push (false , inl tt))) (~ i) j - PFpushT (push (false , inr x) i) j = {!!} + PFpushT (push (false , inr x) i) j = + (pp1 x + ◁ flipSquare + (symP (compPath-filler' + (theCoh2 x false) (push (false , inr x))))) i j PFpushT (push (true , inl x) i) j = push (true , inl x) (i ∧ j) - PFpushT (push (true , inr x) i) j = {!!} - - PF→P∘ : Pushout F snd → Pushout (g ∘ f) snd - PF→P∘ (inl (inl x)) = inl b₀ - PF→P∘ (inl (inr x)) = inl x - PF→P∘ (inr (inl x)) = inl b₀ - PF→P∘ (inr (inr x)) = inr x - PF→P∘ (push (false , inl x) i) = inl b₀ - PF→P∘ (push (true , inl x) i) = inl b₀ - PF→P∘ (push (false , inr x) i) = {!!} - PF→P∘ (push (true , inr x) i) = {!!} - - Is1 : Iso (Pushout F snd) (Pushout (g ∘ f) snd) - Is1 = {!!} - - + PFpushT (push (true , inr x) i) j = + (pp2 x + ◁ flipSquare + (symP (compPath-filler' + (theCoh2 x true) (push (true , inr x))))) i j + + cong-PF→P∘' : (b : _) (a : _) → cong PF→P∘' (theCoh2 b a) ≡ refl + cong-PF→P∘' b a with (f (a , b)) + ... | inl x = cong-∙ PF→P∘' (push (true , inl tt)) (sym (push (false , inl tt))) + ∙ sym (rUnit refl) + ... | inr x = refl + PF→P∘'→PF : (x : _) → PF→P∘' (P∘'→PF x) ≡ x + PF→P∘'→PF (inl x) = refl + PF→P∘'→PF (inr x) = refl + PF→P∘'→PF (push (false , b) i) j = + (cong-∙ PF→P∘' (theCoh2 b false) (push (false , inr b)) + ∙ cong (_∙ push (false , b)) (cong-PF→P∘' b false) + ∙ sym (lUnit _)) j i + PF→P∘'→PF (push (true , b) i) j = + (cong-∙ PF→P∘' (theCoh2 b true) (push (true , inr b)) + ∙ cong (_∙ push (true , b)) (cong-PF→P∘' b true) + ∙ sym (lUnit _)) j i -TheF : {!!} -TheF = {!!} + Is1 : Iso (Pushout F snd) (Pushout (g ∘ f) snd) + Iso.fun Is1 = PF→P∘' + Iso.inv Is1 = P∘'→PF + Iso.rightInv Is1 = PF→P∘'→PF + Iso.leftInv Is1 = PFpushT FinPred : ∀ {m} → Fin (suc (suc m)) → Fin (suc m) FinPred {m = m} (zero , p) = zero , p @@ -558,83 +587,180 @@ FinPred {m = m} (suc x , p) = x , p fone : ∀ {m} → Fin (suc (suc m)) fone {m} = suc zero , tt +module _ {m : ℕ} where + Fin→Unit⊎Fin : (x : Fin (suc m)) → Unit ⊎ Fin m + Fin→Unit⊎Fin = Ind.elimFin (inl tt) inr + + Unit⊎Fin→Fin : Unit ⊎ Fin m → Fin (suc m) + Unit⊎Fin→Fin (inl x) = flast + Unit⊎Fin→Fin (inr x) = injectSuc x + + Iso-Fin-Unit⊎Fin : Iso (Fin (suc m)) (Unit ⊎ Fin m) + Iso.fun Iso-Fin-Unit⊎Fin = Fin→Unit⊎Fin + Iso.inv Iso-Fin-Unit⊎Fin = Unit⊎Fin→Fin + Iso.rightInv Iso-Fin-Unit⊎Fin (inl x) = Ind.elimFinβ (inl tt) inr .fst + Iso.rightInv Iso-Fin-Unit⊎Fin (inr x) = Ind.elimFinβ (inl tt) inr .snd x + Iso.leftInv Iso-Fin-Unit⊎Fin = + Ind.elimFin + (cong Unit⊎Fin→Fin (Ind.elimFinβ (inl tt) inr .fst)) + λ x → (cong Unit⊎Fin→Fin (Ind.elimFinβ (inl tt) inr .snd x)) + +≠flast→<ᵗflast : {n : ℕ} → (x : Fin (suc n)) → ¬ x ≡ flast → fst x <ᵗ n +≠flast→<ᵗflast = Ind.elimFin (λ p → ⊥.rec (p refl)) λ p _ → snd p CW₁DataPre : (n m : ℕ) (f : S₊ 0 × Fin (suc (suc n)) → Fin (suc (suc m))) - → f (true , fzero) ≡ fzero - → f (false , fzero) ≡ fone + → f (false , flast) ≡ flast + → (t : f (true , flast) .fst <ᵗ suc m) → Σ[ k ∈ ℕ ] Σ[ f' ∈ (S₊ 0 × Fin k → Fin (suc m)) ] Iso (Pushout f snd) (Pushout f' snd) CW₁DataPre n m f p q = (suc n) - , F↓ - , (compIso (subst (λ f → Iso (Pushout f snd) - (Pushout f* snd)) (funExt f*≡) idIso) - {!!}) + , _ + , compIso (invIso (pushoutIso _ _ _ _ + (isoToEquiv (Σ-cong-iso-snd λ _ → invIso Iso-Fin-Unit⊎Fin)) + (isoToEquiv (invIso Iso-Fin-Unit⊎Fin)) + (isoToEquiv (invIso Iso-Fin-Unit⊎Fin)) + (funExt (uncurry help)) + (funExt λ x → refl))) + (Is1 {A = Fin (suc n)} {B = Fin (suc m)} + (λ x → Fin→Unit⊎Fin (f (fst x , injectSuc (snd x)))) + (f (true , flast) .fst , q)) where - f* : S₊ 0 × Fin (suc (suc n)) → Fin (suc (suc m)) - f* (false , zero , b) = fone - f* (true , zero , b) = fzero - f* (x , suc a , b) = f (x , suc a , b) - - f↓ : S₊ 0 × Fin (suc (suc n)) → Fin (suc (suc m)) - f↓ (x , p) with DiscreteFin (f (x , p)) fzero - ... | yes p₁ = fone - ... | no ¬p = f (x , p) - - F↓ : Bool × Fin (suc n) → Fin (suc m) - F↓ (x , p) = FinPred (f↓ (x , fsuc p)) - - FinPred-lem : (b : _) (x : _) (p : _) - → FinPred (f (b , suc x , p)) ≡ F↓ (b , x , p) - FinPred-lem b x p with (DiscreteFin (f (b , fsuc (x , p))) fzero) - ... | yes p₁ = cong FinPred p₁ - ... | no ¬p = refl - - FinPred-lem2 : (b : _) (x : _) (p : _) → FinPred (f* (b , suc x , p)) ≡ F↓ (b , FinPred (suc x , p)) - FinPred-lem2 false x r = FinPred-lem false x r - FinPred-lem2 true x p = FinPred-lem true x p - - slask : (b : Bool) - → Path (Pushout F↓ (λ r → snd r)) - (inl (FinPred (f* (b , zero , tt)))) - (inl (F↓ (b , zero , tt))) - slask b with DiscreteFin (f (b , fsuc (0 , tt))) (0 , tt) - slask false | yes p = refl - slask true | yes p = refl - slask false | no ¬p = {!!} ∙ {!!} - slask true | no ¬p = {!¬p !} - - f*≡ : (x : _ ) → f* x ≡ f x - f*≡ (false , zero , z) = sym q - f*≡ (true , zero , z) = sym p - f*≡ (false , suc y , z) = refl - f*≡ (true , suc y , z) = refl - - pp : (b : Bool) (x : Fin (suc (suc n))) - → Path (Pushout F↓ (λ r → snd r)) - (inl (FinPred (f* (b , x)))) (inr (FinPred x)) - pp b (zero , p) = ({!!} ∙ push (b , fzero)) - pp b (suc x , p) = (λ i → inl (FinPred-lem2 b x p i)) ∙ push (b , x , p) - -- ((λ i → inl (FinPred-lem2 b x p i)) ∙ push (b , FinPred (x , p))) - - Pf*→ : Pushout f* snd → Pushout F↓ snd - Pf*→ (inl x) = inl (FinPred x) - Pf*→ (inr x) = inr (FinPred x) - Pf*→ (push (b , zero , q) i) = {!!} - Pf*→ (push (b , suc x' , q) i) = {!!} - - {- - main : Iso _ _ - Iso.fun main (inl x) = inl (FinPred x) - Iso.fun main (inr x) = inr (FinPred x) - Iso.fun main (push (false , zero , s) i) = ({!!} ∙∙ push {!!} ∙∙ {!!}) i - Iso.fun main (push (true , zero , s) i) = {!!} - Iso.fun main (push (b , suc a , s) i) = push (b , a , s) i - Iso.inv main x = {!!} - Iso.rightInv main x = {!!} - Iso.leftInv main x = {!!} - -} + help : (x : Bool) (y : Unit ⊎ Fin (suc n)) + → Unit⊎Fin→Fin + (F (λ x₁ → Ind.elimFin (inl tt) inr (f (fst x₁ , injectSuc (snd x₁)))) + (f (true , flast) .fst , q) (x , y)) + ≡ f (x , Unit⊎Fin→Fin y) + help false (inl a) = sym p + help true (inl b) = Σ≡Prop (λ _ → isProp<ᵗ) refl + help false (inr a) = Iso.leftInv Iso-Fin-Unit⊎Fin _ + help true (inr a) = Iso.leftInv Iso-Fin-Unit⊎Fin _ + +isPropFin1 : isProp (Fin 1) +isPropFin1 (zero , tt) (zero , tt) = refl + + +Iso⊎→Iso : ∀ {ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → (f : Iso A C) + → (e : Iso (A ⊎ B) (C ⊎ D)) + → ((a : A) → Iso.fun e (inl a) ≡ inl (Iso.fun f a)) + → Iso B D +Iso⊎→Iso {A = A} {B} {C} {D} f e p = Iso' + where + ⊥-fib : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ⊎ B → Type + ⊥-fib (inl x) = ⊥ + ⊥-fib (inr x) = Unit + + module _ {ℓ ℓ' ℓ'' ℓ''' : Level} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + (f : Iso A C) + (e : Iso (A ⊎ B) (C ⊎ D)) + (p : (a : A) → Iso.fun e (inl a) ≡ inl (Iso.fun f a)) where + T : (b : B) → Type _ + T b = Σ[ d' ∈ C ⊎ D ] (Iso.fun e (inr b) ≡ d') + + T-elim : ∀ {ℓ} (b : B) {P : (x : T b) → Type ℓ} + → ((d : D) (s : _) → P (inr d , s)) + → (x : _) → P x + T-elim b ind (inl x , q) = + ⊥.rec (subst ⊥-fib (sym (sym (Iso.leftInv e _) + ∙ cong (Iso.inv e) + (p _ ∙ cong inl (Iso.rightInv f x) ∙ sym q) + ∙ Iso.leftInv e _)) tt) + T-elim b ind (inr x , y) = ind x y + + e-pres-inr-help : (b : B) → T f e p b → D + e-pres-inr-help b = T-elim f e p b λ d _ → d + + p' : (a : C) → Iso.inv e (inl a) ≡ inl (Iso.inv f a) + p' c = cong (Iso.inv e ∘ inl) (sym (Iso.rightInv f c)) + ∙∙ cong (Iso.inv e) (sym (p (Iso.inv f c))) + ∙∙ Iso.leftInv e _ + + e⁻-pres-inr-help : (d : D) → T (invIso f) (invIso e) p' d → B + e⁻-pres-inr-help d = T-elim (invIso f) (invIso e) p' d λ b _ → b + + e-pres-inr : B → D + e-pres-inr b = e-pres-inr-help b (_ , refl) + + e⁻-pres-inr : D → B + e⁻-pres-inr d = e⁻-pres-inr-help d (_ , refl) + + lem1 : (b : B) (e : T f e p b) (d : _) + → e⁻-pres-inr-help (e-pres-inr-help b e) d ≡ b + lem1 b = T-elim f e p b λ d s + → T-elim (invIso f) (invIso e) p' _ + λ b' s' → invEq (_ , isEmbedding-inr _ _) + (sym s' ∙ cong (Iso.inv e) (sym s) ∙ Iso.leftInv e _) + + lem2 : (d : D) (e : T (invIso f) (invIso e) p' d ) (t : _) + → e-pres-inr-help (e⁻-pres-inr-help d e) t ≡ d + lem2 d = T-elim (invIso f) (invIso e) p' d + λ b s → T-elim f e p _ λ d' s' + → invEq (_ , isEmbedding-inr _ _) + (sym s' ∙ cong (Iso.fun e) (sym s) ∙ Iso.rightInv e _) + + Iso' : Iso B D + Iso.fun Iso' = e-pres-inr + Iso.inv Iso' = e⁻-pres-inr + Iso.rightInv Iso' x = lem2 x _ _ + Iso.leftInv Iso' x = lem1 x _ _ + +Fin≠Fin : {n m : ℕ} → ¬ (n ≡ m) → ¬ (Iso (Fin n) (Fin m)) +Fin≠Fin {n = zero} {m = zero} p = ⊥.rec (p refl) +Fin≠Fin {n = zero} {m = suc m} p q = Iso.inv q fzero .snd +Fin≠Fin {n = suc n} {m = zero} p q = Iso.fun q fzero .snd +Fin≠Fin {n = suc n} {m = suc m} p q = + Fin≠Fin {n = n} {m = m} (p ∘ cong suc) + (Iso⊎→Iso idIso help λ {(zero , tt) + → cong (Iso.inv FinSuc) (swapFinβₗ (Iso.fun q flast) flast) + ∙ Ind.elimFinβ (inl flast) inr .fst}) + where + q^ : Iso (Fin (suc n)) (Fin (suc m)) + q^ = compIso q (swapFinIso (Iso.fun q flast) flast) + + help : Iso (Fin 1 ⊎ Fin n) (Fin 1 ⊎ Fin m) + help = compIso FinSuc (compIso q^ (invIso FinSuc)) + +CW₁Data₁ : (m : ℕ) (f : S₊ 0 × Fin 1 → Fin (suc (suc m))) + → isConnected 2 (Pushout f snd) + → Iso (S₊ 0 × Fin 1) (Fin (suc (suc m))) +CW₁Data₁ m f c = mainIso + where + f' : Bool → Fin (suc (suc m)) + f' = f ∘ (_, fzero) + + f'-surj : (x : _) → Σ[ t ∈ Bool ] (f' t ≡ x) + f'-surj x = + isSurj-α₀ (suc zero) m f c x .fst .fst + , cong f (Σ≡Prop (λ _ → λ {(zero , p) (zero , q) → refl}) refl) + ∙ isSurj-α₀ (suc zero) m f c x .snd + + f-true≠f-false : (x : _) → f' true ≡ x → ¬ f' true ≡ f' false + f-true≠f-false (zero , q) p r = lem (f'-surj fone) + where + lem : Σ[ t ∈ Bool ] (f' t ≡ fone) → ⊥ + lem (false , s) = snotz (cong fst (sym s ∙ sym r ∙ p)) + lem (true , s) = snotz (cong fst (sym s ∙ p)) + f-true≠f-false (suc x , q) p r = lem (f'-surj fzero) + where + lem : Σ[ t ∈ Bool ] (f' t ≡ fzero) → ⊥ + lem (false , s) = snotz (cong fst (sym p ∙ r ∙ s)) + lem (true , s) = snotz (cong fst (sym p ∙ s)) + + f-inj : (x y : _) → f x ≡ f y → x ≡ y + f-inj (false , zero , tt) (false , zero , tt) p = refl + f-inj (false , zero , tt) (true , zero , tt) p = ⊥.rec (f-true≠f-false _ refl (sym p)) + f-inj (true , zero , tt) (false , zero , tt) p = ⊥.rec (f-true≠f-false _ refl p) + f-inj (true , zero , tt) (true , zero , tt) p = refl + + mainIso : Iso (S₊ 0 × Fin 1) (Fin (suc (suc m))) + Iso.fun mainIso = f + Iso.inv mainIso x = isSurj-α₀ (suc zero) m f c x .fst + Iso.rightInv mainIso x = isSurj-α₀ 1 m f c x .snd + Iso.leftInv mainIso (x , zero , tt) = + (f-inj _ _ (isSurj-α₀ 1 m f c (f (x , fzero)) .snd)) CW₁Data : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) → isConnected 2 (Pushout f snd) @@ -642,308 +768,391 @@ CW₁Data : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) Iso (Pushout f snd) (Pushout f' snd) CW₁Data zero m f c = ⊥.rec (snd (notAllLoops-α₀ zero m f c .fst)) -CW₁Data (suc n) m f c = n , (λ x → {!xpath!}) , {!RetrFin!} - -- n , {!!} , (compIso {!!} {!!}) +CW₁Data (suc zero) zero f c = 0 , ((λ ()) , compIso isoₗ (PushoutEmptyFam (snd ∘ snd) snd)) where - t = notAllLoops-α₀ (suc n) m f c + isoₗ : Iso (Pushout f snd) (Fin 1) + isoₗ = PushoutAlongEquiv (isoToEquiv (CW₁Data₁ zero f c)) _ +CW₁Data (suc zero) (suc m) f c = + ⊥.rec (Fin≠Fin (snotz ∘ sym ∘ cong (predℕ ∘ predℕ)) + mainIso) + where + mainIso : Iso (Fin 2) (Fin (3 +ℕ m)) + mainIso = + compIso + (compIso + (invIso rUnit×Iso) + (Σ-cong-iso + (invIso Iso-Bool-Fin) + (λ _ → isContr→Iso {A = Unit} + (tt , λ _ → refl) (fzero , λ {(zero , tt) → refl})))) + (CW₁Data₁ (suc m) f c) +CW₁Data (suc (suc n)) m f c = + main .fst , main .snd .fst + , compIso correct (main .snd .snd) + where + t = notAllLoops-α₀ (suc (suc n)) m f c abstract - x₀ : Fin (suc n) + x₀ : Fin (suc (suc n)) x₀ = fst t xpath : ¬ (f (true , x₀) ≡ f (false , x₀)) xpath = snd t - Fin0-iso : Iso (S₊ 0 × Fin (suc n)) (S₊ 0 × Fin (suc n)) - Fin0-iso = Σ-cong-iso-snd λ _ → swapFinIso x₀ fzero + Fin0-iso : Iso (S₊ 0 × Fin (suc (suc n))) (S₊ 0 × Fin (suc (suc n))) + Fin0-iso = Σ-cong-iso-snd λ _ → swapFinIso flast x₀ FinIso2 : Iso (Fin (suc (suc m))) (Fin (suc (suc m))) - FinIso2 = {!!} + FinIso2 = swapFinIso (f (false , x₀)) flast - Upath = isoToPath (swapFinIso x₀ fzero) + f' : S₊ 0 × Fin (suc (suc n)) → Fin (suc (suc m)) + f' = Iso.fun FinIso2 ∘ f ∘ Iso.fun Fin0-iso - f' : S₊ 0 × Fin (suc n) → Fin (suc (suc m)) - f' (x , zero , p) = f (x , x₀) - f' (x , suc b , p) with (discreteℕ (suc b) (fst x₀)) - ... | yes q = f (x , fzero) - ... | no ¬q = f (x , suc b , p) - - f'≡ : (x : _) → f (fst x , (swapFin x₀ fzero (snd x))) ≡ f' x - f'≡ (x , zero , p) with (discreteℕ 0 (fst x₀)) - ... | yes p₁ = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) p₁)) - ... | no ¬p = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) refl)) - f'≡ (x , suc b , p) with (discreteℕ (suc b) (fst x₀)) - ... | yes p₁ = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) refl)) - ... | no ¬p = refl - - Pushout-f≡ : (Pushout f snd) ≡ (Pushout f' snd) - Pushout-f≡ = (λ i → Pushout {A = S₊ 0 × Upath i} {B = Fin (suc (suc m))} {C = Upath i} (help i) snd) - ∙ (λ i → Pushout (λ x → f'≡ x i) snd) - where - help : PathP (λ i → Bool × Upath i → Fin (suc (suc m))) - f - (f ∘ (λ a → fst a , (swapFin x₀ fzero (snd a)))) - help = toPathP (funExt λ p → transportRefl _ - ∙ cong f (λ j → fst p , transp (λ j → Upath (~ j)) i0 (snd p)) - ∙ cong f (cong (fst p ,_) - (cong (Iso.fun (swapFinIso x₀ fzero)) (transportRefl _)))) - - -{- - F : Σ[ x ∈ Fin n ] (¬ f (true , x) ≡ f (false , x)) - F with (allConst? DiscreteFin (λ x y → f (y , x))) - ... | inr x = x - ... | inl qs = {!!} - where - pre-help : Fin (suc (suc m)) → Type - pre-help (zero , p) = ⊥ - pre-help (suc x , p) = Unit - - lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit - lem (zero , q) (false , t) p = {!!} - lem (zero , q) (true , t) p = {!!} - lem (suc x , q) a p = refl - - help : Pushout f snd → Type - help (inl x) = {!!} -- pre-help x - help (inr x) = {!x!} -- Unit - help (push a i) = {!!} -- lem (f a) a refl i - -} - --- pre-help : Fin (suc (suc m)) → Type --- pre-help (zero , p) = ⊥ --- pre-help (suc x , p) = Unit - --- lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit --- lem (zero , q) a p = --- ⊥.rec (x a (p ∙ Σ≡Prop (λ _ → isProp<ᵗ) refl)) --- lem (suc n , q) a p = refl + f'≡ : f' (false , flast) ≡ flast + f'≡ = cong (Iso.fun FinIso2 ∘ f) + (cong (false ,_) (swapFinβₗ flast x₀)) + ∙ swapFinβₗ (f (false , x₀)) flast --- help : Pushout f snd → Type --- help (inl x) = pre-help x --- help (inr x) = Unit --- help (push a i) = lem (f a) a refl i + f'¬≡ : ¬ (f' (true , flast) ≡ flast) + f'¬≡ p = xpath (cong f (ΣPathP (refl , sym (swapFinβₗ flast x₀))) + ∙ sym (Iso.rightInv FinIso2 _) + ∙ cong (Iso.inv FinIso2) (p ∙ sym f'≡) + ∙ Iso.rightInv FinIso2 _ + ∙ cong f (ΣPathP (refl , swapFinβₗ flast x₀))) --- Pf≅⊥ : ⊥ --- Pf≅⊥ = TR.rec isProp⊥ --- (λ p → subst help (sym p) tt) --- (Iso.fun (PathIdTruncIso 1) --- (isContr→isProp c ∣ inl fzero ∣ ∣ inl flast ∣)) + f'< : fst (f' (true , flast)) <ᵗ suc m + f'< = ≠flast→<ᵗflast _ f'¬≡ --- -- where --- -- fSurj : (x : Fin _) → Σ[ r ∈ _ ] f r ≡ x --- -- fSurj x = {! -- Fin→Surj? ? ? ? ?!} + main = CW₁DataPre _ _ f' f'≡ f'< --- -- ptt : Pushout f snd → Type --- -- ptt (inl x) = Fin (suc (suc m)) --- -- ptt (inr x) = {!Fin n!} --- -- ptt (push a i) = {!!} - --- -- module _ (s : {!!}) where + Upath = isoToPath (swapFinIso x₀ fzero) + correct : Iso (Pushout f snd) (Pushout f' snd) + correct = pushoutIso _ _ _ _ + (isoToEquiv Fin0-iso) (isoToEquiv FinIso2) (isoToEquiv (swapFinIso flast x₀)) + (funExt (λ x → cong (FinIso2 .Iso.fun ∘ f) (sym (Iso.rightInv Fin0-iso x)))) + refl --- -- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ --- -- yieldsCWskel∙ X = --- -- Σ[ f ∈ (ℕ → ℕ) ] --- -- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] --- -- ((X zero ≃ Fin (f zero)) × --- -- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) - --- -- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) - --- -- module CWskel∙-fields (C : CWskel∙ ℓ) where --- -- card = C .snd .fst --- -- A = Fin ∘ card --- -- α = C .snd .snd .fst --- -- e = C .snd .snd .snd .snd - --- -- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero --- -- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] - --- -- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- -- fst (CWpt (C , f) n) = C n --- -- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) - --- -- -- Technically, a CW complex should be the sequential colimit over the following maps --- -- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) --- -- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) - --- -- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n --- -- ptCW T zero = T .snd .snd .fst zero (inl tt) --- -- ptCW T (suc n) = CW∙↪ T n (ptCW T n) - --- -- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- -- CW∙ T n = fst T n , ptCW T n - --- -- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) --- -- fst (CW∙↪∙ T n) = CW∙↪ T n --- -- snd (CW∙↪∙ T n) = refl - - --- -- -- But if it stabilises, no need for colimits. --- -- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ --- -- yieldsFinCWskel∙ n X = --- -- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) - --- -- -- ... which should give us finite CW complexes. --- -- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) --- -- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) - --- -- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ --- -- finCWskel→CWskel∙ n C = fst C , fst (snd C) - --- -- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ --- -- Sequence.obj (realiseSeq∙ (C , p)) = C --- -- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ - --- -- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ --- -- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) - --- -- -- realisation of CW complex from skeleton --- -- realise∙ : CWskel∙ ℓ → Type ℓ --- -- realise∙ C = SeqColim (realiseSeq∙ C) +CW₁Data' : (n m : ℕ) (f : S₊ 0 × Fin n → Fin m) + → isConnected 2 (Pushout f snd) + → Σ[ k ∈ ℕ ] Σ[ f' ∈ (S₊ 0 × Fin k → Fin 1) ] + Iso (Pushout f snd) + (Pushout f' snd) +CW₁Data' zero zero f c = ⊥.rec (TR.rec (λ()) help (fst c)) + where + help : ¬ Pushout f snd + help = elimProp _ (λ _ → λ ()) snd snd +CW₁Data' (suc n) zero f c = ⊥.rec (f (true , fzero) .snd) +CW₁Data' n (suc zero) f c = n , f , idIso +CW₁Data' zero (suc (suc m)) f c = + ⊥.rec (TR.rec (λ()) (snotz ∘ sym ∘ cong fst) + (Iso.fun (PathIdTruncIso _) + (isContr→isProp (subst (isConnected 2) (isoToPath help) c) + ∣ fzero ∣ ∣ fone ∣))) + where + help : Iso (Pushout f snd) (Fin (suc (suc m))) + help = invIso (PushoutEmptyFam (λ()) λ()) +CW₁Data' (suc n) (suc (suc m)) f c + with (CW₁Data' _ (suc m) (CW₁Data (suc n) m f c .snd .fst) + (subst (isConnected 2) + (isoToPath (CW₁Data (suc n) m f c .snd .snd)) c)) +... | (k , f↓ , e) = k , f↓ , compIso (CW₁Data (suc n) m f c .snd .snd) e + +-- CW₁Data' zero zero f c = ⊥.rec (TR.rec (λ()) help (fst c)) +-- where +-- help : ¬ Pushout f snd +-- help = elimProp _ (λ _ → λ ()) snd snd +-- CW₁Data' (suc n) zero f c = ⊥.rec (f (true , fzero) .snd) +-- CW₁Data' zero (suc m) f c = {!!} +-- CW₁Data' (suc n) (suc zero) f c = suc n , f , idIso +-- CW₁Data' (suc n) (suc (suc m)) f c = {!!} +-- where +-- help : {!!} +-- help = {!!} + +-- -- f' : S₊ 0 × Fin (suc n) → Fin (suc (suc m)) +-- -- f' (x , zero , p) = f (x , x₀) +-- -- f' (x , suc b , p) with (discreteℕ (suc b) (fst x₀)) +-- -- ... | yes q = f (x , fzero) +-- -- ... | no ¬q = f (x , suc b , p) + +-- -- f'≡ : (x : _) → f (fst x , (swapFin x₀ fzero (snd x))) ≡ f' x +-- -- f'≡ (x , zero , p) with (discreteℕ 0 (fst x₀)) +-- -- ... | yes p₁ = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) p₁)) +-- -- ... | no ¬p = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) refl)) +-- -- f'≡ (x , suc b , p) with (discreteℕ (suc b) (fst x₀)) +-- -- ... | yes p₁ = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) refl)) +-- -- ... | no ¬p = refl + +-- -- Pushout-f≡ : (Pushout f snd) ≡ (Pushout f' snd) +-- -- Pushout-f≡ = (λ i → Pushout {A = S₊ 0 × Upath i} {B = Fin (suc (suc m))} {C = Upath i} (help i) snd) +-- -- ∙ (λ i → Pushout (λ x → f'≡ x i) snd) +-- -- where +-- -- help : PathP (λ i → Bool × Upath i → Fin (suc (suc m))) +-- -- f +-- -- (f ∘ (λ a → fst a , (swapFin x₀ fzero (snd a)))) +-- -- help = toPathP (funExt λ p → transportRefl _ +-- -- ∙ cong f (λ j → fst p , transp (λ j → Upath (~ j)) i0 (snd p)) +-- -- ∙ cong f (cong (fst p ,_) +-- -- (cong (Iso.fun (swapFinIso x₀ fzero)) (transportRefl _)))) + + +-- -- {- +-- -- F : Σ[ x ∈ Fin n ] (¬ f (true , x) ≡ f (false , x)) +-- -- F with (allConst? DiscreteFin (λ x y → f (y , x))) +-- -- ... | inr x = x +-- -- ... | inl qs = {!!} +-- -- where +-- -- pre-help : Fin (suc (suc m)) → Type +-- -- pre-help (zero , p) = ⊥ +-- -- pre-help (suc x , p) = Unit + +-- -- lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit +-- -- lem (zero , q) (false , t) p = {!!} +-- -- lem (zero , q) (true , t) p = {!!} +-- -- lem (suc x , q) a p = refl + +-- -- help : Pushout f snd → Type +-- -- help (inl x) = {!!} -- pre-help x +-- -- help (inr x) = {!x!} -- Unit +-- -- help (push a i) = {!!} -- lem (f a) a refl i +-- -- -} + +-- -- -- pre-help : Fin (suc (suc m)) → Type +-- -- -- pre-help (zero , p) = ⊥ +-- -- -- pre-help (suc x , p) = Unit + +-- -- -- lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit +-- -- -- lem (zero , q) a p = +-- -- -- ⊥.rec (x a (p ∙ Σ≡Prop (λ _ → isProp<ᵗ) refl)) +-- -- -- lem (suc n , q) a p = refl + +-- -- -- help : Pushout f snd → Type +-- -- -- help (inl x) = pre-help x +-- -- -- help (inr x) = Unit +-- -- -- help (push a i) = lem (f a) a refl i + +-- -- -- Pf≅⊥ : ⊥ +-- -- -- Pf≅⊥ = TR.rec isProp⊥ +-- -- -- (λ p → subst help (sym p) tt) +-- -- -- (Iso.fun (PathIdTruncIso 1) +-- -- -- (isContr→isProp c ∣ inl fzero ∣ ∣ inl flast ∣)) + +-- -- -- -- where +-- -- -- -- fSurj : (x : Fin _) → Σ[ r ∈ _ ] f r ≡ x +-- -- -- -- fSurj x = {! -- Fin→Surj? ? ? ? ?!} + +-- -- -- -- ptt : Pushout f snd → Type +-- -- -- -- ptt (inl x) = Fin (suc (suc m)) +-- -- -- -- ptt (inr x) = {!Fin n!} +-- -- -- -- ptt (push a i) = {!!} + +-- -- -- -- module _ (s : {!!}) where + + + +-- -- -- -- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ +-- -- -- -- yieldsCWskel∙ X = +-- -- -- -- Σ[ f ∈ (ℕ → ℕ) ] +-- -- -- -- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] +-- -- -- -- ((X zero ≃ Fin (f zero)) × +-- -- -- -- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) + +-- -- -- -- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- -- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) + +-- -- -- -- module CWskel∙-fields (C : CWskel∙ ℓ) where +-- -- -- -- card = C .snd .fst +-- -- -- -- A = Fin ∘ card +-- -- -- -- α = C .snd .snd .fst +-- -- -- -- e = C .snd .snd .snd .snd + +-- -- -- -- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero +-- -- -- -- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] + +-- -- -- -- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- -- -- -- fst (CWpt (C , f) n) = C n +-- -- -- -- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) + +-- -- -- -- -- Technically, a CW complex should be the sequential colimit over the following maps +-- -- -- -- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) +-- -- -- -- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) + +-- -- -- -- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n +-- -- -- -- ptCW T zero = T .snd .snd .fst zero (inl tt) +-- -- -- -- ptCW T (suc n) = CW∙↪ T n (ptCW T n) + +-- -- -- -- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- -- -- -- CW∙ T n = fst T n , ptCW T n + +-- -- -- -- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) +-- -- -- -- fst (CW∙↪∙ T n) = CW∙↪ T n +-- -- -- -- snd (CW∙↪∙ T n) = refl + + +-- -- -- -- -- But if it stabilises, no need for colimits. +-- -- -- -- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ +-- -- -- -- yieldsFinCWskel∙ n X = +-- -- -- -- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) + +-- -- -- -- -- ... which should give us finite CW complexes. +-- -- -- -- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) +-- -- -- -- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) + +-- -- -- -- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ +-- -- -- -- finCWskel→CWskel∙ n C = fst C , fst (snd C) + +-- -- -- -- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ +-- -- -- -- Sequence.obj (realiseSeq∙ (C , p)) = C +-- -- -- -- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ + +-- -- -- -- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ +-- -- -- -- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) + +-- -- -- -- -- realisation of CW complex from skeleton +-- -- -- -- realise∙ : CWskel∙ ℓ → Type ℓ +-- -- -- -- realise∙ C = SeqColim (realiseSeq∙ C) --- -- realise∙∙ : CWskel∙ ℓ → Pointed ℓ --- -- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) --- -- open import Cubical.Data.Empty as ⊥ +-- -- -- -- realise∙∙ : CWskel∙ ℓ → Pointed ℓ +-- -- -- -- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) +-- -- -- -- open import Cubical.Data.Empty as ⊥ --- -- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ --- -- CWskel∙→CWskel A zero = Lift ⊥ --- -- CWskel∙→CWskel A (suc n) = A n --- -- open import Cubical.Foundations.Isomorphism +-- -- -- -- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ +-- -- -- -- CWskel∙→CWskel A zero = Lift ⊥ +-- -- -- -- CWskel∙→CWskel A (suc n) = A n +-- -- -- -- open import Cubical.Foundations.Isomorphism --- -- module _ (A : ℕ → Type ℓ) --- -- (cwsk : yieldsCWskel∙ A) where +-- -- -- -- module _ (A : ℕ → Type ℓ) +-- -- -- -- (cwsk : yieldsCWskel∙ A) where --- -- private --- -- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n --- -- αs (suc n) x = snd cwsk .fst n (inr x) +-- -- -- -- private +-- -- -- -- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n +-- -- -- -- αs (suc n) x = snd cwsk .fst n (inr x) --- -- e0 : {!!} --- -- e0 = {!!} - --- -- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst --- -- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) --- -- es-suc→ n (inr x) = inl x --- -- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) --- -- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i --- -- es-suc→ n (push (push a i₁) i) = {!!} - --- -- es-suc : (n : ℕ) --- -- → Iso (cofib (fst (snd cwsk) n)) --- -- (Pushout (αs (suc n)) fst) --- -- Iso.fun (es-suc n) = es-suc→ n --- -- Iso.inv (es-suc n) = {!!} --- -- Iso.rightInv (es-suc n) = {!!} --- -- Iso.leftInv (es-suc n) = {!!} +-- -- -- -- e0 : {!!} +-- -- -- -- e0 = {!!} + +-- -- -- -- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst +-- -- -- -- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) +-- -- -- -- es-suc→ n (inr x) = inl x +-- -- -- -- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) +-- -- -- -- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i +-- -- -- -- es-suc→ n (push (push a i₁) i) = {!!} + +-- -- -- -- es-suc : (n : ℕ) +-- -- -- -- → Iso (cofib (fst (snd cwsk) n)) +-- -- -- -- (Pushout (αs (suc n)) fst) +-- -- -- -- Iso.fun (es-suc n) = es-suc→ n +-- -- -- -- Iso.inv (es-suc n) = {!!} +-- -- -- -- Iso.rightInv (es-suc n) = {!!} +-- -- -- -- Iso.leftInv (es-suc n) = {!!} --- -- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) --- -- es zero = {!!} --- -- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) - --- -- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) --- -- fst yieldsCWskel∙→' = cwsk .fst --- -- fst (snd yieldsCWskel∙→') = αs --- -- fst (snd (snd yieldsCWskel∙→')) () --- -- snd (snd (snd yieldsCWskel∙→')) = {!!} - --- -- yieldsCWskel∙→ : (A : ℕ → Type ℓ) --- -- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) --- -- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst --- -- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) --- -- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () --- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} --- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = --- -- compEquiv (cwsk .snd .snd .snd n) --- -- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) --- -- where --- -- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst --- -- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) --- -- theEq→ (inr x) = inl x --- -- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) --- -- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) --- -- theEq→ (push (push a i₁) i) = {!!} - --- -- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) --- -- Iso.fun theEq = theEq→ --- -- Iso.inv theEq = {!!} --- -- Iso.rightInv theEq x = {!!} --- -- Iso.leftInv theEq x = {!!} - - --- -- -- compEquiv {!!} {!!} - - --- -- -- -- Finally: definition of CW complexes --- -- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' - --- -- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ - --- -- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) - --- -- -- -- Finite CW complexes --- -- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- -- isFinCW {ℓ = ℓ} X = --- -- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) - --- -- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ - --- -- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) - --- -- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X --- -- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str - --- -- -- finCW→CW : finCW ℓ → CW ℓ --- -- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p - - --- -- -- -- morphisms --- -- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') --- -- -- C →ᶜʷ D = fst C → fst D - --- -- -- --the cofibre of the inclusion of X n into X (suc n) --- -- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ --- -- -- cofibCW n C = cofib (CW↪ C n) - --- -- -- --...is basically a sphere bouquet --- -- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) --- -- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) --- -- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e --- -- -- where --- -- -- s = Bouquet≃-gen --- -- -- α = C .snd .snd .fst --- -- -- e = C .snd .snd .snd .snd n - --- -- -- --sending X (suc n) into the cofibCW --- -- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C --- -- -- to_cofibCW n C x = inr x - --- -- -- --the connecting map of the long exact sequence --- -- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) --- -- -- → cofib conn → Susp A --- -- -- δ-pre conn (inl x) = north --- -- -- δ-pre conn (inr x) = south --- -- -- δ-pre conn (push a i) = merid a i - --- -- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) --- -- -- δ n C = δ-pre (CW↪ C n) - --- -- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) --- -- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C --- -- -- CW↪∞ C n x = incl x - --- -- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m --- -- -- fst (finCW↑ m n p C) = fst C --- -- -- fst (snd (finCW↑ m n p C)) = snd C .fst --- -- -- snd (snd (finCW↑ m n p C)) k = --- -- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) --- -- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) --- -- -- (snd C .snd (k +ℕ fst p)) +-- -- -- -- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) +-- -- -- -- es zero = {!!} +-- -- -- -- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) + +-- -- -- -- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) +-- -- -- -- fst yieldsCWskel∙→' = cwsk .fst +-- -- -- -- fst (snd yieldsCWskel∙→') = αs +-- -- -- -- fst (snd (snd yieldsCWskel∙→')) () +-- -- -- -- snd (snd (snd yieldsCWskel∙→')) = {!!} + +-- -- -- -- yieldsCWskel∙→ : (A : ℕ → Type ℓ) +-- -- -- -- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) +-- -- -- -- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst +-- -- -- -- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) +-- -- -- -- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () +-- -- -- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} +-- -- -- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = +-- -- -- -- compEquiv (cwsk .snd .snd .snd n) +-- -- -- -- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) +-- -- -- -- where +-- -- -- -- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst +-- -- -- -- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) +-- -- -- -- theEq→ (inr x) = inl x +-- -- -- -- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) +-- -- -- -- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) +-- -- -- -- theEq→ (push (push a i₁) i) = {!!} + +-- -- -- -- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) +-- -- -- -- Iso.fun theEq = theEq→ +-- -- -- -- Iso.inv theEq = {!!} +-- -- -- -- Iso.rightInv theEq x = {!!} +-- -- -- -- Iso.leftInv theEq x = {!!} + + +-- -- -- -- -- compEquiv {!!} {!!} + + +-- -- -- -- -- -- Finally: definition of CW complexes +-- -- -- -- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- -- -- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' + +-- -- -- -- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- -- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ + +-- -- -- -- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- -- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) + +-- -- -- -- -- -- Finite CW complexes +-- -- -- -- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- -- -- -- isFinCW {ℓ = ℓ} X = +-- -- -- -- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) + +-- -- -- -- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- -- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ + +-- -- -- -- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- -- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) + +-- -- -- -- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X +-- -- -- -- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str + +-- -- -- -- -- finCW→CW : finCW ℓ → CW ℓ +-- -- -- -- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p + + +-- -- -- -- -- -- morphisms +-- -- -- -- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') +-- -- -- -- -- C →ᶜʷ D = fst C → fst D + +-- -- -- -- -- --the cofibre of the inclusion of X n into X (suc n) +-- -- -- -- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ +-- -- -- -- -- cofibCW n C = cofib (CW↪ C n) + +-- -- -- -- -- --...is basically a sphere bouquet +-- -- -- -- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) +-- -- -- -- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) +-- -- -- -- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e +-- -- -- -- -- where +-- -- -- -- -- s = Bouquet≃-gen +-- -- -- -- -- α = C .snd .snd .fst +-- -- -- -- -- e = C .snd .snd .snd .snd n + +-- -- -- -- -- --sending X (suc n) into the cofibCW +-- -- -- -- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C +-- -- -- -- -- to_cofibCW n C x = inr x + +-- -- -- -- -- --the connecting map of the long exact sequence +-- -- -- -- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) +-- -- -- -- -- → cofib conn → Susp A +-- -- -- -- -- δ-pre conn (inl x) = north +-- -- -- -- -- δ-pre conn (inr x) = south +-- -- -- -- -- δ-pre conn (push a i) = merid a i + +-- -- -- -- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) +-- -- -- -- -- δ n C = δ-pre (CW↪ C n) + +-- -- -- -- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) +-- -- -- -- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C +-- -- -- -- -- CW↪∞ C n x = incl x + +-- -- -- -- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m +-- -- -- -- -- fst (finCW↑ m n p C) = fst C +-- -- -- -- -- fst (snd (finCW↑ m n p C)) = snd C .fst +-- -- -- -- -- snd (snd (finCW↑ m n p C)) k = +-- -- -- -- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) +-- -- -- -- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) +-- -- -- -- -- (snd C .snd (k +ℕ fst p)) From afe51a465bd2d086d2461e1d80b44719d539bb5b Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 8 May 2024 18:20:52 +0200 Subject: [PATCH 41/73] p2 --- Cubical/CW/Pointed.agda | 543 ++++++++--------- Cubical/CW/Pointed2.agda | 1226 ++++++++++++++++++++++++++++++++++++++ 2 files changed, 1480 insertions(+), 289 deletions(-) create mode 100644 Cubical/CW/Pointed2.agda diff --git a/Cubical/CW/Pointed.agda b/Cubical/CW/Pointed.agda index 4018fbf796..b9fcee811a 100644 --- a/Cubical/CW/Pointed.agda +++ b/Cubical/CW/Pointed.agda @@ -859,300 +859,265 @@ CW₁Data' (suc n) (suc (suc m)) f c (isoToPath (CW₁Data (suc n) m f c .snd .snd)) c)) ... | (k , f↓ , e) = k , f↓ , compIso (CW₁Data (suc n) m f c .snd .snd) e --- CW₁Data' zero zero f c = ⊥.rec (TR.rec (λ()) help (fst c)) --- where --- help : ¬ Pushout f snd --- help = elimProp _ (λ _ → λ ()) snd snd --- CW₁Data' (suc n) zero f c = ⊥.rec (f (true , fzero) .snd) --- CW₁Data' zero (suc m) f c = {!!} --- CW₁Data' (suc n) (suc zero) f c = suc n , f , idIso --- CW₁Data' (suc n) (suc (suc m)) f c = {!!} +CW₁Data'' : (n m : ℕ) (f : Fin n × S₊ 0 → Fin m) + → isConnected 2 (Pushout f snd) + → Σ[ k ∈ ℕ ] Σ[ f' ∈ (Fin k × S₊ 0 → Fin 1) ] + Iso (Pushout f snd) + (Pushout f' snd) +CW₁Data'' n m f c = + fst help + , (help .snd .fst ∘ Σ-swap-≃ .fst) + , compIso {!help .snd .snd!} {!!} + where + help : {!!} + help = CW₁Data' n m (f ∘ Σ-swap-≃ .fst) {!!} + + +yieldsConnectedCWskel : (ℕ → Type ℓ) → Type _ +yieldsConnectedCWskel A = Σ[ sk ∈ yieldsCWskel A ] (sk .fst 1 ≡ 1) + +yieldsConnectedCWskel' : (ℕ → Type ℓ) → Type _ +yieldsConnectedCWskel' A = Σ[ sk ∈ yieldsCWskel A ] isConnected 2 (realise (_ , sk)) + +connectedCWskel : (ℓ : Level) → Type (ℓ-suc ℓ) +connectedCWskel ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel X) + +connectedCWskel' : (ℓ : Level) → Type (ℓ-suc ℓ) +connectedCWskel' ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel' X) + +truncCWIso : (A : CWskel ℓ) (n : ℕ) + → Iso (hLevelTrunc n (realise A)) (hLevelTrunc n (fst A n)) +truncCWIso A n = invIso (connectedTruncIso n incl (isConnected-CW↪∞ n A)) + +isConnectedColim→isConnectedSkel : + (A : CWskel ℓ) (n : ℕ) + → isConnected n (realise A) + → isConnected n (fst A n) +isConnectedColim→isConnectedSkel A n c = + isOfHLevelRetractFromIso 0 + (invIso (truncCWIso A n)) c + +module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A) where + private + c = snd sk+c + sk = fst sk+c + + module AC = CWskel-fields (_ , sk) + + liftStr = CW₁Data' {!!} {!!} {!AC.α!} {!isConnectedColim→isConnectedSkel (_ , sk) 2 c!} + + + + + connectedCWskel→ : yieldsConnectedCWskel' A + connectedCWskel→ = ({!!} , {!!}) , {!!} + + + +-- yieldsGoodCWSkel : (ℕ → Type ℓ) → Type ℓ +-- yieldsGoodCWSkel = ? + +-- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ +-- yieldsCWskel∙ X = +-- Σ[ f ∈ (ℕ → ℕ) ] +-- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] +-- ((X zero ≃ Fin (f zero)) × +-- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) + +-- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) +-- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) + +-- module CWskel∙-fields (C : CWskel∙ ℓ) where +-- card = C .snd .fst +-- A = Fin ∘ card +-- α = C .snd .snd .fst +-- e = C .snd .snd .snd .snd + +-- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero +-- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] + +-- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- fst (CWpt (C , f) n) = C n +-- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) + +-- -- Technically, a CW complex should be the sequential colimit over the following maps +-- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) +-- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) + +-- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n +-- ptCW T zero = T .snd .snd .fst zero (inl tt) +-- ptCW T (suc n) = CW∙↪ T n (ptCW T n) + +-- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- CW∙ T n = fst T n , ptCW T n + +-- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) +-- fst (CW∙↪∙ T n) = CW∙↪ T n +-- snd (CW∙↪∙ T n) = refl + + +-- -- But if it stabilises, no need for colimits. +-- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ +-- yieldsFinCWskel∙ n X = +-- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) + +-- -- ... which should give us finite CW complexes. +-- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) +-- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) + +-- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ +-- finCWskel→CWskel∙ n C = fst C , fst (snd C) + +-- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ +-- Sequence.obj (realiseSeq∙ (C , p)) = C +-- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ + +-- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ +-- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) + +-- -- realisation of CW complex from skeleton +-- realise∙ : CWskel∙ ℓ → Type ℓ +-- realise∙ C = SeqColim (realiseSeq∙ C) + +-- realise∙∙ : CWskel∙ ℓ → Pointed ℓ +-- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) +-- open import Cubical.Data.Empty as ⊥ + +-- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ +-- CWskel∙→CWskel A zero = Lift ⊥ +-- CWskel∙→CWskel A (suc n) = A n +-- open import Cubical.Foundations.Isomorphism + + +-- module _ (A : ℕ → Type ℓ) +-- (cwsk : yieldsCWskel∙ A) where + +-- private +-- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n +-- αs (suc n) x = snd cwsk .fst n (inr x) + +-- e0 : {!!} +-- e0 = {!!} + +-- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst +-- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) +-- es-suc→ n (inr x) = inl x +-- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) +-- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i +-- es-suc→ n (push (push a i₁) i) = {!!} + +-- es-suc : (n : ℕ) +-- → Iso (cofib (fst (snd cwsk) n)) +-- (Pushout (αs (suc n)) fst) +-- Iso.fun (es-suc n) = es-suc→ n +-- Iso.inv (es-suc n) = {!!} +-- Iso.rightInv (es-suc n) = {!!} +-- Iso.leftInv (es-suc n) = {!!} + +-- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) +-- es zero = {!!} +-- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) + +-- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) +-- fst yieldsCWskel∙→' = cwsk .fst +-- fst (snd yieldsCWskel∙→') = αs +-- fst (snd (snd yieldsCWskel∙→')) () +-- snd (snd (snd yieldsCWskel∙→')) = {!!} + +-- yieldsCWskel∙→ : (A : ℕ → Type ℓ) +-- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) +-- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst +-- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) +-- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () +-- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} +-- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = +-- compEquiv (cwsk .snd .snd .snd n) +-- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) -- where --- help : {!!} --- help = {!!} - --- -- f' : S₊ 0 × Fin (suc n) → Fin (suc (suc m)) --- -- f' (x , zero , p) = f (x , x₀) --- -- f' (x , suc b , p) with (discreteℕ (suc b) (fst x₀)) --- -- ... | yes q = f (x , fzero) --- -- ... | no ¬q = f (x , suc b , p) - --- -- f'≡ : (x : _) → f (fst x , (swapFin x₀ fzero (snd x))) ≡ f' x --- -- f'≡ (x , zero , p) with (discreteℕ 0 (fst x₀)) --- -- ... | yes p₁ = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) p₁)) --- -- ... | no ¬p = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) refl)) --- -- f'≡ (x , suc b , p) with (discreteℕ (suc b) (fst x₀)) --- -- ... | yes p₁ = cong f (ΣPathP (refl , Σ≡Prop (λ _ → isProp<ᵗ) refl)) --- -- ... | no ¬p = refl - --- -- Pushout-f≡ : (Pushout f snd) ≡ (Pushout f' snd) --- -- Pushout-f≡ = (λ i → Pushout {A = S₊ 0 × Upath i} {B = Fin (suc (suc m))} {C = Upath i} (help i) snd) --- -- ∙ (λ i → Pushout (λ x → f'≡ x i) snd) --- -- where --- -- help : PathP (λ i → Bool × Upath i → Fin (suc (suc m))) --- -- f --- -- (f ∘ (λ a → fst a , (swapFin x₀ fzero (snd a)))) --- -- help = toPathP (funExt λ p → transportRefl _ --- -- ∙ cong f (λ j → fst p , transp (λ j → Upath (~ j)) i0 (snd p)) --- -- ∙ cong f (cong (fst p ,_) --- -- (cong (Iso.fun (swapFinIso x₀ fzero)) (transportRefl _)))) - - --- -- {- --- -- F : Σ[ x ∈ Fin n ] (¬ f (true , x) ≡ f (false , x)) --- -- F with (allConst? DiscreteFin (λ x y → f (y , x))) --- -- ... | inr x = x --- -- ... | inl qs = {!!} --- -- where --- -- pre-help : Fin (suc (suc m)) → Type --- -- pre-help (zero , p) = ⊥ --- -- pre-help (suc x , p) = Unit - --- -- lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit --- -- lem (zero , q) (false , t) p = {!!} --- -- lem (zero , q) (true , t) p = {!!} --- -- lem (suc x , q) a p = refl - --- -- help : Pushout f snd → Type --- -- help (inl x) = {!!} -- pre-help x --- -- help (inr x) = {!x!} -- Unit --- -- help (push a i) = {!!} -- lem (f a) a refl i --- -- -} - --- -- -- pre-help : Fin (suc (suc m)) → Type --- -- -- pre-help (zero , p) = ⊥ --- -- -- pre-help (suc x , p) = Unit - --- -- -- lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit --- -- -- lem (zero , q) a p = --- -- -- ⊥.rec (x a (p ∙ Σ≡Prop (λ _ → isProp<ᵗ) refl)) --- -- -- lem (suc n , q) a p = refl - --- -- -- help : Pushout f snd → Type --- -- -- help (inl x) = pre-help x --- -- -- help (inr x) = Unit --- -- -- help (push a i) = lem (f a) a refl i - --- -- -- Pf≅⊥ : ⊥ --- -- -- Pf≅⊥ = TR.rec isProp⊥ --- -- -- (λ p → subst help (sym p) tt) --- -- -- (Iso.fun (PathIdTruncIso 1) --- -- -- (isContr→isProp c ∣ inl fzero ∣ ∣ inl flast ∣)) - --- -- -- -- where --- -- -- -- fSurj : (x : Fin _) → Σ[ r ∈ _ ] f r ≡ x --- -- -- -- fSurj x = {! -- Fin→Surj? ? ? ? ?!} - --- -- -- -- ptt : Pushout f snd → Type --- -- -- -- ptt (inl x) = Fin (suc (suc m)) --- -- -- -- ptt (inr x) = {!Fin n!} --- -- -- -- ptt (push a i) = {!!} - --- -- -- -- module _ (s : {!!}) where - - - --- -- -- -- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ --- -- -- -- yieldsCWskel∙ X = --- -- -- -- Σ[ f ∈ (ℕ → ℕ) ] --- -- -- -- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] --- -- -- -- ((X zero ≃ Fin (f zero)) × --- -- -- -- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) - --- -- -- -- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- -- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) - --- -- -- -- module CWskel∙-fields (C : CWskel∙ ℓ) where --- -- -- -- card = C .snd .fst --- -- -- -- A = Fin ∘ card --- -- -- -- α = C .snd .snd .fst --- -- -- -- e = C .snd .snd .snd .snd - --- -- -- -- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero --- -- -- -- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] - --- -- -- -- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- -- -- -- fst (CWpt (C , f) n) = C n --- -- -- -- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) - --- -- -- -- -- Technically, a CW complex should be the sequential colimit over the following maps --- -- -- -- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) --- -- -- -- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) - --- -- -- -- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n --- -- -- -- ptCW T zero = T .snd .snd .fst zero (inl tt) --- -- -- -- ptCW T (suc n) = CW∙↪ T n (ptCW T n) - --- -- -- -- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- -- -- -- CW∙ T n = fst T n , ptCW T n - --- -- -- -- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) --- -- -- -- fst (CW∙↪∙ T n) = CW∙↪ T n --- -- -- -- snd (CW∙↪∙ T n) = refl - - --- -- -- -- -- But if it stabilises, no need for colimits. --- -- -- -- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ --- -- -- -- yieldsFinCWskel∙ n X = --- -- -- -- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) - --- -- -- -- -- ... which should give us finite CW complexes. --- -- -- -- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) --- -- -- -- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) - --- -- -- -- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ --- -- -- -- finCWskel→CWskel∙ n C = fst C , fst (snd C) - --- -- -- -- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ --- -- -- -- Sequence.obj (realiseSeq∙ (C , p)) = C --- -- -- -- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ - --- -- -- -- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ --- -- -- -- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) - --- -- -- -- -- realisation of CW complex from skeleton --- -- -- -- realise∙ : CWskel∙ ℓ → Type ℓ --- -- -- -- realise∙ C = SeqColim (realiseSeq∙ C) +-- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst +-- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) +-- theEq→ (inr x) = inl x +-- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) +-- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) +-- theEq→ (push (push a i₁) i) = {!!} + +-- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) +-- Iso.fun theEq = theEq→ +-- Iso.inv theEq = {!!} +-- Iso.rightInv theEq x = {!!} +-- Iso.leftInv theEq x = {!!} + + +-- -- compEquiv {!!} {!!} + + +-- -- -- Finally: definition of CW complexes +-- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' + +-- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ + +-- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) + +-- -- -- Finite CW complexes +-- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- isFinCW {ℓ = ℓ} X = +-- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) + +-- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ + +-- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) + +-- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X +-- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str + +-- -- finCW→CW : finCW ℓ → CW ℓ +-- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p + + +-- -- -- morphisms +-- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') +-- -- C →ᶜʷ D = fst C → fst D --- -- -- -- realise∙∙ : CWskel∙ ℓ → Pointed ℓ --- -- -- -- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) --- -- -- -- open import Cubical.Data.Empty as ⊥ +-- -- --the cofibre of the inclusion of X n into X (suc n) +-- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ +-- -- cofibCW n C = cofib (CW↪ C n) --- -- -- -- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ --- -- -- -- CWskel∙→CWskel A zero = Lift ⊥ --- -- -- -- CWskel∙→CWskel A (suc n) = A n --- -- -- -- open import Cubical.Foundations.Isomorphism +-- -- --...is basically a sphere bouquet +-- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) +-- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) +-- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e +-- -- where +-- -- s = Bouquet≃-gen +-- -- α = C .snd .snd .fst +-- -- e = C .snd .snd .snd .snd n +-- -- --sending X (suc n) into the cofibCW +-- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C +-- -- to_cofibCW n C x = inr x --- -- -- -- module _ (A : ℕ → Type ℓ) --- -- -- -- (cwsk : yieldsCWskel∙ A) where +-- -- --the connecting map of the long exact sequence +-- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) +-- -- → cofib conn → Susp A +-- -- δ-pre conn (inl x) = north +-- -- δ-pre conn (inr x) = south +-- -- δ-pre conn (push a i) = merid a i --- -- -- -- private --- -- -- -- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n --- -- -- -- αs (suc n) x = snd cwsk .fst n (inr x) +-- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) +-- -- δ n C = δ-pre (CW↪ C n) --- -- -- -- e0 : {!!} --- -- -- -- e0 = {!!} - --- -- -- -- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst --- -- -- -- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) --- -- -- -- es-suc→ n (inr x) = inl x --- -- -- -- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) --- -- -- -- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i --- -- -- -- es-suc→ n (push (push a i₁) i) = {!!} - --- -- -- -- es-suc : (n : ℕ) --- -- -- -- → Iso (cofib (fst (snd cwsk) n)) --- -- -- -- (Pushout (αs (suc n)) fst) --- -- -- -- Iso.fun (es-suc n) = es-suc→ n --- -- -- -- Iso.inv (es-suc n) = {!!} --- -- -- -- Iso.rightInv (es-suc n) = {!!} --- -- -- -- Iso.leftInv (es-suc n) = {!!} +-- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) +-- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C +-- -- CW↪∞ C n x = incl x --- -- -- -- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) --- -- -- -- es zero = {!!} --- -- -- -- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) - --- -- -- -- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) --- -- -- -- fst yieldsCWskel∙→' = cwsk .fst --- -- -- -- fst (snd yieldsCWskel∙→') = αs --- -- -- -- fst (snd (snd yieldsCWskel∙→')) () --- -- -- -- snd (snd (snd yieldsCWskel∙→')) = {!!} - --- -- -- -- yieldsCWskel∙→ : (A : ℕ → Type ℓ) --- -- -- -- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) --- -- -- -- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst --- -- -- -- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) --- -- -- -- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () --- -- -- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} --- -- -- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = --- -- -- -- compEquiv (cwsk .snd .snd .snd n) --- -- -- -- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) --- -- -- -- where --- -- -- -- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst --- -- -- -- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) --- -- -- -- theEq→ (inr x) = inl x --- -- -- -- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) --- -- -- -- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) --- -- -- -- theEq→ (push (push a i₁) i) = {!!} - --- -- -- -- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) --- -- -- -- Iso.fun theEq = theEq→ --- -- -- -- Iso.inv theEq = {!!} --- -- -- -- Iso.rightInv theEq x = {!!} --- -- -- -- Iso.leftInv theEq x = {!!} - - --- -- -- -- -- compEquiv {!!} {!!} - - --- -- -- -- -- -- Finally: definition of CW complexes --- -- -- -- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- -- -- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' - --- -- -- -- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- -- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ - --- -- -- -- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- -- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) - --- -- -- -- -- -- Finite CW complexes --- -- -- -- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- -- -- -- isFinCW {ℓ = ℓ} X = --- -- -- -- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) - --- -- -- -- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- -- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ - --- -- -- -- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- -- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) - --- -- -- -- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X --- -- -- -- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str - --- -- -- -- -- finCW→CW : finCW ℓ → CW ℓ --- -- -- -- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p - - --- -- -- -- -- -- morphisms --- -- -- -- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') --- -- -- -- -- C →ᶜʷ D = fst C → fst D - --- -- -- -- -- --the cofibre of the inclusion of X n into X (suc n) --- -- -- -- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ --- -- -- -- -- cofibCW n C = cofib (CW↪ C n) - --- -- -- -- -- --...is basically a sphere bouquet --- -- -- -- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) --- -- -- -- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) --- -- -- -- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e --- -- -- -- -- where --- -- -- -- -- s = Bouquet≃-gen --- -- -- -- -- α = C .snd .snd .fst --- -- -- -- -- e = C .snd .snd .snd .snd n - --- -- -- -- -- --sending X (suc n) into the cofibCW --- -- -- -- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C --- -- -- -- -- to_cofibCW n C x = inr x - --- -- -- -- -- --the connecting map of the long exact sequence --- -- -- -- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) --- -- -- -- -- → cofib conn → Susp A --- -- -- -- -- δ-pre conn (inl x) = north --- -- -- -- -- δ-pre conn (inr x) = south --- -- -- -- -- δ-pre conn (push a i) = merid a i - --- -- -- -- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) --- -- -- -- -- δ n C = δ-pre (CW↪ C n) - --- -- -- -- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) --- -- -- -- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C --- -- -- -- -- CW↪∞ C n x = incl x - --- -- -- -- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m --- -- -- -- -- fst (finCW↑ m n p C) = fst C --- -- -- -- -- fst (snd (finCW↑ m n p C)) = snd C .fst --- -- -- -- -- snd (snd (finCW↑ m n p C)) k = --- -- -- -- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) --- -- -- -- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) --- -- -- -- -- (snd C .snd (k +ℕ fst p)) +-- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m +-- -- fst (finCW↑ m n p C) = fst C +-- -- fst (snd (finCW↑ m n p C)) = snd C .fst +-- -- snd (snd (finCW↑ m n p C)) k = +-- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) +-- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) +-- -- (snd C .snd (k +ℕ fst p)) diff --git a/Cubical/CW/Pointed2.agda b/Cubical/CW/Pointed2.agda new file mode 100644 index 0000000000..418986e3c0 --- /dev/null +++ b/Cubical/CW/Pointed2.agda @@ -0,0 +1,1226 @@ +{-# OPTIONS --safe --lossy-unification #-} + +-- This file contains definition of CW complexes and skeleta. + +module Cubical.CW.Pointed2 where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Sigma +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence + +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Wedge + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + +open import Cubical.Relation.Nullary + +open import Cubical.CW.Base + +open Sequence + +open import Cubical.Foundations.Equiv.HalfAdjoint + +private + variable + ℓ ℓ' ℓ'' : Level + +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +module _ {C : Type ℓ} {B : Type ℓ'} where + PushoutAlongEquiv→ : {A : Type ℓ} + (e : A ≃ C) (f : A → B) → Pushout (fst e) f → B + PushoutAlongEquiv→ e f (inl x) = f (invEq e x) + PushoutAlongEquiv→ e f (inr x) = x + PushoutAlongEquiv→ e f (push a i) = f (retEq e a i) + + private + PushoutAlongEquiv→Cancel : {A : Type ℓ} (e : A ≃ C) (f : A → B) + → retract (PushoutAlongEquiv→ e f) inr + PushoutAlongEquiv→Cancel = + EquivJ (λ A e → (f : A → B) + → retract (PushoutAlongEquiv→ e f) inr) + λ f → λ { (inl x) → sym (push x) + ; (inr x) → refl + ; (push a i) j → push a (~ j ∨ i)} + + PushoutAlongEquiv : {A : Type ℓ} (e : A ≃ C) (f : A → B) + → Iso (Pushout (fst e) f) B + Iso.fun (PushoutAlongEquiv e f) = PushoutAlongEquiv→ e f + Iso.inv (PushoutAlongEquiv e f) = inr + Iso.rightInv (PushoutAlongEquiv e f) x = refl + Iso.leftInv (PushoutAlongEquiv e f) = PushoutAlongEquiv→Cancel e f + + +--- CW complexes --- +-- Definition of (the skeleton) of an arbitrary CW complex +-- New def: X 0 is empty and C (suc n) is pushout + +open import Cubical.HITs.Pushout +module _ {A : Type} {B : A → Pointed₀} (C : Pointed₀) (f : (⋁gen A B , inl tt) →∙ C) where + + + private + open 3x3-span + inst : 3x3-span + A00 inst = A + A02 inst = Σ A (fst ∘ B) + A04 inst = Σ A (fst ∘ B) + A20 inst = A + A22 inst = A + A24 inst = Σ A (fst ∘ B) + A40 inst = Unit + A42 inst = Unit + A44 inst = fst C + f10 inst = idfun A + f12 inst = λ x → x , snd (B x) + f14 inst = idfun _ + f30 inst = λ _ → tt + f32 inst = λ _ → tt + f34 inst = fst f ∘ inr + f01 inst = fst + f21 inst = idfun A + f41 inst = λ _ → tt + f03 inst = idfun _ + f23 inst = λ x → x , snd (B x) + f43 inst = λ _ → pt C + H11 inst = λ _ → refl + H13 inst = λ _ → refl + H31 inst = λ _ → refl + H33 inst = λ x → sym (snd f) ∙ cong (fst f) (push x) + + A0□Iso : Iso (A0□ inst) A + A0□Iso = compIso (equivToIso (symPushout _ _)) + (PushoutAlongEquiv (idEquiv _) fst) + + A2□Iso : Iso (A2□ inst) (Σ A (fst ∘ B)) + A2□Iso = PushoutAlongEquiv (idEquiv A) _ + + A4□Iso : Iso (A4□ inst) (fst C) + A4□Iso = PushoutAlongEquiv (idEquiv Unit) λ _ → pt C + + A○□Iso : Iso (A○□ inst) (Pushout (fst f ∘ inr) fst) + A○□Iso = compIso (equivToIso (symPushout _ _)) + (invIso (pushoutIso _ _ _ _ + (isoToEquiv (invIso A2□Iso)) + (isoToEquiv (invIso A4□Iso)) + (isoToEquiv (invIso A0□Iso)) + refl + λ i x → push x i)) + + + + A□0Iso : Iso (A□0 inst) Unit + A□0Iso = isContr→Iso + (inr tt , λ { (inl x) → sym (push x) + ; (inr x) → refl + ; (push a i) j → push a (i ∨ ~ j)}) + (tt , (λ _ → refl)) + + A□2Iso : Iso (A□2 inst) (⋁gen A B) + A□2Iso = equivToIso (symPushout _ _) + + + A□4Iso : Iso (A□4 inst) (fst C) + A□4Iso = PushoutAlongEquiv (idEquiv _) _ + + open import Cubical.Foundations.GroupoidLaws + A□○Iso : Iso (A□○ inst) (cofib (fst f)) + A□○Iso = invIso (pushoutIso _ _ _ _ + (isoToEquiv (invIso A□2Iso)) + (isoToEquiv (invIso A□0Iso)) + (isoToEquiv (invIso A□4Iso)) + (funExt (λ { (inl x) → refl + ; (inr x) → sym (push (fst x)) ∙ refl + ; (push a i) j → (sym (push a) ∙ refl) (i ∧ j)})) + (funExt λ { (inl x) i → inr (snd f i) + ; (inr x) → sym (push x) + ; (push a i) j + → hcomp (λ k + → (λ {(i = i0) → inr (compPath-filler' (sym (snd f)) + (cong (fst f) (push a)) j (~ k)) + ; (i = i1) → push (a , snd (B a)) (~ j) + ; (j = i0) → inr (fst f (push a (~ k ∨ i)))})) + (push (a , snd (B a)) (~ i ∨ ~ j))})) + + ⋁-cofib-Iso : Iso (Pushout (fst f ∘ inr) fst) (cofib (fst f)) + ⋁-cofib-Iso = compIso (compIso (invIso A○□Iso) + (invIso (3x3-Iso inst))) + A□○Iso + + + +-- connected comp +open import Cubical.Homotopy.Connected +open import Cubical.CW.Properties +open import Cubical.HITs.Truncation as TR +open import Cubical.Foundations.HLevels + +isConnectedCW→isConnectedSkel : (C : CWskel ℓ) + → isConnected 2 (realise C) + → (n : ℕ) + → isConnected 2 (C .fst (suc (suc n))) +isConnectedCW→isConnectedSkel C c n = + TR.rec (isOfHLevelSuc 1 isPropIsContr) + (λ ⋆ → TR.rec (isProp→isOfHLevelSuc (suc n) isPropIsContr) + (λ {(x , p) → ∣ x ∣ + , (TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (λ a → Iso.inv (PathIdTruncIso 1) + (TR.rec (isOfHLevelTrunc 1) + (λ q → + TR.rec (isOfHLevelPlus' 1 (isOfHLevelTrunc 1)) + (∣_∣ₕ ∘ fst) + (isConnectedCong (suc n) _ (isConnected-CW↪∞ (suc (suc n)) C) + q .fst)) + (Iso.fun (PathIdTruncIso 1) + (sym (c .snd ∣ incl x ∣) ∙ c .snd ∣ incl a ∣)))))}) + (isConnected-CW↪∞ (suc (suc n)) C ⋆ .fst)) + (c .fst) + +open import Cubical.Data.Bool +open import Cubical.Data.Nat.Order.Inductive +-- open import Cubical.Data.Dec + +open import Cubical.Relation.Nullary.Base +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Data.Fin.Inductive.Properties as Ind + +FinSuc : {n : ℕ} → Iso (Fin 1 ⊎ Fin n) (Fin (suc n)) +Iso.fun (FinSuc {n = n}) = ⊎.rec (λ _ → flast) injectSuc +Iso.inv (FinSuc {n = n}) = Ind.elimFin (inl flast) inr +Iso.rightInv (FinSuc {n = n}) = + Ind.elimFin (cong (⊎.rec (λ _ → flast) injectSuc) + (Ind.elimFinβ (inl flast) inr .fst)) + λ x → cong (⊎.rec (λ _ → flast) injectSuc) + (Ind.elimFinβ (inl flast) inr .snd x) +Iso.leftInv (FinSuc {n = n}) (inl (zero , p)) = + Ind.elimFinβ (inl flast) inr .fst +Iso.leftInv (FinSuc {n = n}) (inr x) = Ind.elimFinβ (inl flast) inr .snd x + +Fin+ : {n m : ℕ} → Iso (Fin n ⊎ Fin m) (Fin (n +ℕ m)) +Iso.fun (Fin+ {n = zero} {m = m}) (inr x) = x +Iso.inv (Fin+ {n = zero} {m = m}) x = inr x +Iso.rightInv (Fin+ {n = zero} {m = m}) x = refl +Iso.leftInv (Fin+ {n = zero} {m = m}) (inr x) = refl +Fin+ {n = suc n} {m = m} = + compIso (⊎Iso (invIso FinSuc) idIso) + (compIso ⊎-assoc-Iso + (compIso (⊎Iso idIso (Fin+ {n = n} {m = m})) + FinSuc)) + +Iso-Unit-Fin : Iso Unit (Fin 1) +Iso.fun Iso-Unit-Fin tt = fzero +Iso.inv Iso-Unit-Fin (x , p) = tt +Iso.rightInv Iso-Unit-Fin (zero , p) = Σ≡Prop (λ _ → isProp<ᵗ) refl +Iso.leftInv Iso-Unit-Fin x = refl + +Iso-Bool-Fin : Iso (S₊ 0) (Fin 2) +Iso.fun Iso-Bool-Fin false = flast +Iso.fun Iso-Bool-Fin true = fzero +Iso.inv Iso-Bool-Fin (zero , p) = true +Iso.inv Iso-Bool-Fin (suc x , p) = false +Iso.rightInv Iso-Bool-Fin (zero , p) = refl +Iso.rightInv Iso-Bool-Fin (suc zero , p) = + Σ≡Prop (λ _ → isProp<ᵗ) refl +Iso.leftInv Iso-Bool-Fin false = refl +Iso.leftInv Iso-Bool-Fin true = refl + +Fin× : {n m : ℕ} → Iso (Fin n × Fin m) (Fin (n · m)) +Fin× {n = zero} {m = m} = + iso (λ {()}) (λ{()}) (λ{()}) (λ{()}) +Fin× {n = suc n} {m = m} = + compIso + (compIso + (compIso (Σ-cong-iso-fst (invIso FinSuc)) + (compIso Σ-swap-Iso + (compIso ×DistR⊎Iso + (⊎Iso (compIso + (Σ-cong-iso-snd (λ _ → invIso Iso-Unit-Fin)) rUnit×Iso) + Σ-swap-Iso)))) + (⊎Iso idIso Fin×)) + (Fin+ {n = m} {n · m}) + +DiscreteFin : ∀ {n} → Discrete (Fin n) +DiscreteFin x y with discreteℕ (fst x) (fst y) +... | yes p = yes (Σ≡Prop (λ _ → isProp<ᵗ) p) +... | no ¬p = no λ q → ¬p (cong fst q) + + +decIm : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') +decIm {A = A} {B = B} f = + (y : B) → (Σ[ x ∈ A ] f x ≡ y) + ⊎ ((x : A) → ¬ f x ≡ y) + +decImFin : ∀ {ℓ} {A : Type ℓ} + (da : Discrete A) (n : ℕ) (f : Fin n → A) + → decIm f +decImFin {A = A} da zero f y = inr λ x → ⊥.rec (¬Fin0 x) +decImFin {A = A} da (suc n) f y with da (f fzero) y +... | yes p = inl (fzero , p) +... | no ¬p with (decImFin da n (f ∘ fsuc) y) +... | inl q = inl ((fsuc (fst q)) , snd q) +... | inr q = inr (Ind.elimFin-alt ¬p q) + +Bool×Fin≡Fin : {n : ℕ} → Iso (Fin n × S₊ 0) (Fin (2 · n)) +Bool×Fin≡Fin = + compIso (compIso Σ-swap-Iso + (Σ-cong-iso-fst Iso-Bool-Fin)) (Fin× {n = 2}) + +decIm-S⁰×Fin : ∀ {ℓ} {A : Type ℓ} + (da : Discrete A) (n : ℕ) (f : Fin n × S₊ 0 → A) → decIm f +decIm-S⁰×Fin {A = A} da n = + subst (λ C → (f : C → A) → decIm f) + (isoToPath (invIso Bool×Fin≡Fin)) + (decImFin da _) + + +module _ {ℓ : Level} {n : ℕ} {A : Fin n → Type ℓ} (x₀ : Fin n) + (pt : A x₀) (l : (x : Fin n) → ¬ x ≡ x₀ → A x) where + private + x = fst x₀ + p = snd x₀ + elimFinChoose : (x : _) → A x + elimFinChoose (x' , q) with (discreteℕ x x') + ... | yes p₁ = subst A (Σ≡Prop (λ _ → isProp<ᵗ) p₁) pt + ... | no ¬p = l (x' , q) λ r → ¬p (sym (cong fst r)) + + elimFinChooseβ : (elimFinChoose x₀ ≡ pt) + × ((x : _) (q : ¬ x ≡ x₀) → elimFinChoose x ≡ l x q) + fst elimFinChooseβ with (discreteℕ x x) + ... | yes p₁ = (λ j → subst A (isSetFin x₀ x₀ (Σ≡Prop (λ z → isProp<ᵗ) p₁) refl j) pt) + ∙ transportRefl pt + ... | no ¬p = ⊥.rec (¬p refl) + snd elimFinChooseβ (x' , q) with (discreteℕ x x') + ... | yes p₁ = λ q → ⊥.rec (q (Σ≡Prop (λ _ → isProp<ᵗ) (sym p₁))) + ... | no ¬p = λ s → cong (l (x' , q)) (isPropΠ (λ _ → isProp⊥) _ _) + + +pickDifferentFin : {n : ℕ} (x : Fin (suc (suc n))) → Σ[ y ∈ Fin (suc (suc n)) ] ¬ x ≡ y +pickDifferentFin (zero , p) = (1 , p) , λ q → snotz (sym (cong fst q)) +pickDifferentFin (suc x , p) = fzero , λ q → snotz (cong fst q) + +allConst? : ∀ {ℓ} {A : Type ℓ} {n : ℕ} (dis : Discrete A) + → (t : Fin n → S₊ 0 → A) + → ((x : Fin n) → t x ≡ λ _ → t x true) + ⊎ (Σ[ x ∈ Fin n ] ¬ (t x true ≡ t x false)) +allConst? {n = zero} dis t = inl λ {()} +allConst? {n = suc n} dis t + with (dis (t fzero true) (t fzero false)) + | (allConst? {n = n} dis λ x p → t (fsuc x) p) +... | yes p | inl x = + inl (Ind.elimFin-alt (funExt + (λ { false → sym p ; true → refl})) x) +... | yes p | inr x = inr (_ , (snd x)) +... | no ¬p | q = inr (_ , ¬p) + + +isSurj-α₀ : (n m : ℕ) (f : Fin n × S₊ 0 → Fin (suc (suc m))) + → isConnected 2 (Pushout f fst) + → (y : _) → Σ[ x ∈ _ ] f x ≡ y +isSurj-α₀ n m f c y with (decIm-S⁰×Fin DiscreteFin n f y) +... | inl x = x +isSurj-α₀ n m f c x₀ | inr q = ⊥.rec nope + where + pre-help : Fin (suc (suc m)) → Type + pre-help = elimFinChoose x₀ ⊥ λ _ _ → Unit + + lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit + lem = elimFinChoose x₀ + (λ s t → ⊥.rec (q _ t)) + λ s t b c → elimFinChooseβ x₀ ⊥ (λ _ _ → Unit) .snd s t + + help : Pushout f fst → Type + help (inl x) = pre-help x + help (inr x) = Unit + help (push a i) = lem (f a) a refl i + + nope : ⊥ + nope = TR.rec isProp⊥ + (λ q → transport (elimFinChooseβ x₀ ⊥ (λ _ _ → Unit) .fst) + (subst help (sym q) + (transport (sym (elimFinChooseβ x₀ ⊥ + (λ _ _ → Unit) .snd _ + (pickDifferentFin x₀ .snd ∘ sym))) tt))) + (Iso.fun (PathIdTruncIso 1) + (isContr→isProp c + ∣ inl x₀ ∣ + ∣ inl (pickDifferentFin x₀ .fst) ∣)) + +notAllLoops-α₀ : (n m : ℕ) (f : Fin n × S₊ 0 → Fin (suc (suc m))) + → isConnected 2 (Pushout f fst) + → Σ[ x ∈ Fin n ] (¬ f (x , true) ≡ f (x , false)) +notAllLoops-α₀ n m f c with (allConst? DiscreteFin (λ x y → f (x , y))) +... | inr x = x +notAllLoops-α₀ n m f c | inl q = + ⊥.rec (TR.rec isProp⊥ (λ p → subst T p tt) + (Iso.fun(PathIdTruncIso 1) + (isContr→isProp c + ∣ inl flast ∣ ∣ inl fzero ∣))) + where + inrT : Fin n → Type + inrT x with (DiscreteFin (f (x , true)) fzero) + ... | yes p = ⊥ + ... | no ¬p = Unit + + inlT : Fin (suc (suc m)) → Type + inlT (zero , p) = ⊥ + inlT (suc x , p) = Unit + + inlrT-pre : (a : _) → inlT (f (a , true)) ≡ inrT a + inlrT-pre a with ((DiscreteFin (f (a , true)) fzero)) + ... | yes p = cong inlT p + inlrT-pre s | no ¬p with (f (s , true)) + ... | zero , tt = ⊥.rec (¬p refl) + ... | suc q , t = refl + + inlrT : (a : _) → inlT (f a) ≡ inrT (fst a) + inlrT (b , false) = cong inlT (funExt⁻ (q b) false) ∙ inlrT-pre b + inlrT (b , true) = inlrT-pre b + + T : Pushout f fst → Type + T (inl x) = inlT x + T (inr x) = inrT x + T (push a i) = inlrT a i + +module _ {n : ℕ} where + swapFin : (x y : Fin n) → Fin n → Fin n + swapFin (x , xp) (y , yp) (z , zp) with (discreteℕ z x) | (discreteℕ z y) + ... | yes p | yes p₁ = z , zp + ... | yes p | no ¬p = y , yp + ... | no ¬p | yes p = x , xp + ... | no ¬p | no ¬p₁ = (z , zp) + + swapFinβₗ : (x y : Fin n) → swapFin x y x ≡ y + swapFinβₗ (x , xp) (y , yp) with (discreteℕ x x) | discreteℕ x y + ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) p₁ + ... | yes p | no ¬p = refl + ... | no ¬p | q = ⊥.rec (¬p refl) + + swapFinβᵣ : (x y : Fin n) → swapFin x y y ≡ x + swapFinβᵣ (x , xp) (y , yp) with (discreteℕ y y) | discreteℕ y x + ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) p₁ + ... | yes p | no ¬p = refl + ... | no ¬p | q = ⊥.rec (¬p refl) + + -- swapFinSwap : (x y z : Fin n) → swapFin x y z ≡ swapFin y x z + -- swapFinSwap x y z with (discreteℕ (fst z) (fst x)) | discreteℕ (fst z) (fst y) + -- ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p₁ ∙ p) + -- ... | yes p | no ¬p = {!help!} + -- where + -- help : y ≡ swapFin y x z + -- help = {!!} + -- ... | no ¬p | q = {!!} + + swapFin² : (x y z : Fin n) → swapFin x y (swapFin x y z) ≡ z + swapFin² (x , xp) (y , yp) (z , zp) with discreteℕ z x | discreteℕ z y + ... | yes p | yes p₁ = silly + where + silly : swapFin (x , xp) (y , yp) (z , zp) ≡ (z , zp) + silly with discreteℕ z x | discreteℕ z y + ... | yes p | yes p₁ = refl + ... | yes p | no ¬p = ⊥.rec (¬p p₁) + ... | no ¬p | r = ⊥.rec (¬p p) + ... | yes p | no ¬q = silly + where + silly : swapFin (x , xp) (y , yp) (y , yp) ≡ (z , zp) + silly with discreteℕ y x | discreteℕ y y + ... | yes p' | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (p' ∙ sym p) + ... | no ¬p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p) + ... | p | no ¬p = ⊥.rec (¬p refl) + ... | no ¬p | yes p = silly + where + silly : swapFin (x , xp) (y , yp) (x , xp) ≡ (z , zp) + silly with discreteℕ x y | discreteℕ x x + ... | yes p₁ | yes _ = Σ≡Prop (λ _ → isProp<ᵗ) (p₁ ∙ sym p) + ... | no ¬p | yes _ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p) + ... | s | no ¬p = ⊥.rec (¬p refl) + ... | no ¬p | no ¬q = silly + where + silly : swapFin (x , xp) (y , yp) (z , zp) ≡ (z , zp) + silly with discreteℕ z x | discreteℕ z y + ... | yes p | yes p₁ = refl + ... | yes p | no ¬b = ⊥.rec (¬p p) + ... | no ¬a | yes b = ⊥.rec (¬q b) + ... | no ¬a | no ¬b = refl + + swapFinIso : (x y : Fin n) → Iso (Fin n) (Fin n) + Iso.fun (swapFinIso x y) = swapFin x y + Iso.inv (swapFinIso x y) = swapFin x y + Iso.rightInv (swapFinIso x y) = swapFin² x y + Iso.leftInv (swapFinIso x y) = swapFin² x y + +liftFun : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} + (f : A → B) → Lift {j = ℓ''} A → Lift {j = ℓ'''} B +liftFun f (lift a) = lift (f a) + +module _ {ℓ ℓ' ℓ'' : Level} (ℓ''' : Level) + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (f : A → B) (g : A → C) where + PushoutLevel : Level + PushoutLevel = ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')) + + PushoutLift : Type PushoutLevel + PushoutLift = Pushout {A = Lift {j = ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')} A} + {B = Lift {j = ℓ-max ℓ (ℓ-max ℓ'' ℓ''')} B} + {C = Lift {j = ℓ-max ℓ (ℓ-max ℓ' ℓ''')} C} + (liftFun f) + (liftFun g) + + PushoutLiftIso : Iso (Pushout f g) PushoutLift + Iso.fun PushoutLiftIso (inl x) = inl (lift x) + Iso.fun PushoutLiftIso (inr x) = inr (lift x) + Iso.fun PushoutLiftIso (push a i) = push (lift a) i + Iso.inv PushoutLiftIso (inl (lift x)) = inl x + Iso.inv PushoutLiftIso (inr (lift x)) = inr x + Iso.inv PushoutLiftIso (push (lift a) i) = push a i + Iso.rightInv PushoutLiftIso (inl (lift x)) = refl + Iso.rightInv PushoutLiftIso (inr (lift x)) = refl + Iso.rightInv PushoutLiftIso (push (lift a) i) = refl + Iso.leftInv PushoutLiftIso (inl x) = refl + Iso.leftInv PushoutLiftIso (inr x) = refl + Iso.leftInv PushoutLiftIso (push a i) = refl + +open import Cubical.Foundations.Equiv + + +pushoutIso' : ∀ {ℓA₁ ℓB₁ ℓC₁ ℓA₂ ℓB₂ ℓC₂} + {A₁ : Type ℓA₁} {B₁ : Type ℓB₁} {C₁ : Type ℓC₁} + {A₂ : Type ℓA₂} {B₂ : Type ℓB₂} {C₂ : Type ℓC₂} + (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) + (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) + (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) + (id1 : fst B≃ ∘ f₁ ≡ f₂ ∘ fst A≃) + (id2 : fst C≃ ∘ g₁ ≡ g₂ ∘ fst A≃) → + Iso (Pushout f₁ g₁) (Pushout f₂ g₂) +pushoutIso' {ℓA₁ = ℓA₁} {ℓB₁} {ℓC₁} {ℓA₂} {ℓB₂} {ℓC₂} + f₁ g₁ f₂ g₂ A≃ B≃ C≃ id1 id2 = + compIso (PushoutLiftIso ℓ* f₁ g₁) + (compIso (pushoutIso _ _ _ _ + (Lift≃Lift A≃) + (Lift≃Lift B≃) + (Lift≃Lift C≃) + (funExt (λ { (lift x) → cong lift (funExt⁻ id1 x)})) + (funExt (λ { (lift x) → cong lift (funExt⁻ id2 x)}))) + (invIso (PushoutLiftIso ℓ* f₂ g₂))) + where + ℓ* = ℓ-max ℓA₁ (ℓ-max ℓA₂ (ℓ-max ℓB₁ (ℓ-max ℓB₂ (ℓ-max ℓC₁ ℓC₂)))) + +PushoutPostcompEquivIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} + {A : Type ℓ} {B : Type ℓ'} {B' : Type ℓ''} {C : Type ℓ'''} + (e2 : B' ≃ B) + (f : A → B') (g : A → C) + → Iso (Pushout (fst e2 ∘ f) g) (Pushout f g) +PushoutPostcompEquivIso {ℓ = ℓ} {ℓ'} {ℓ''} {ℓ'''} + {A = A} {B} {B'} {C} e2 f g = + compIso (PushoutLiftIso ℓ* _ _) + (compIso (pushoutIso _ _ _ _ (idEquiv _) + (compEquiv (invEquiv LiftEquiv) + (compEquiv (invEquiv e2) LiftEquiv)) + (idEquiv _) + (funExt (λ { (lift l) → cong lift (retEq e2 (f l))})) + refl) + (invIso (PushoutLiftIso ℓ* _ _))) + where + ℓ* = ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')) + +Pushout∘Equiv : ∀ {ℓ ℓ' ℓ''} {A A' : Type ℓ} {B B' : Type ℓ'} {C : Type ℓ''} + (e1 : A ≃ A') (e2 : B' ≃ B) + (f : A' → B') (g : A → C) + → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) (Pushout f (g ∘ invEq e1)) +Pushout∘Equiv {A = A} {A' = A'} {B} {B'} {C} = + EquivJ (λ A e1 → (e2 : B' ≃ B) (f : A' → B') (g : A → C) + → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) + (Pushout f (g ∘ invEq e1))) + (EquivJ (λ B' e2 → (f : A' → B') (g : A' → C) + → Iso (Pushout (fst e2 ∘ f) g) + (Pushout f g)) + λ f g → idIso) + +module _ {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} + (f : A × Bool → Unit ⊎ B) (b₀ : B) where + + F : (Unit ⊎ A) × Bool → Unit ⊎ B + F (inl tt , false) = inl tt + F (inl tt , true) = inr b₀ + F (inr a , x) = f (a , x) + + g : Unit ⊎ B → B + g (inl x) = b₀ + g (inr x) = x + + PF→P∘ₗ : (x : Unit ⊎ A) → Pushout (g ∘ f) fst + PF→P∘ₗ (inl x) = inl b₀ + PF→P∘ₗ (inr x) = inr x + + theCoh : (a : _) → inl (g (F a)) ≡ PF→P∘ₗ (fst a) + theCoh (inl x , false) = refl + theCoh (inl x , true) = refl + theCoh (inr x , false) = push (x , false) + theCoh (inr x , true) = push (x , true) + + PF→P∘' : Pushout F fst → Pushout (g ∘ f) fst + PF→P∘' (inl x) = inl (g x) + PF→P∘' (inr x) = PF→P∘ₗ x + PF→P∘' (push a i) = theCoh a i + + theCoh2 : (a : _) (b : _) + → Path (Pushout F fst) (inl (inr (g (f (a , b))))) (inl (f (a , b))) + theCoh2 a b with (f (a , b)) + theCoh2 a b | inl x = (push (inl tt , true)) ∙ sym (push (inl tt , false)) + ... | inr x = refl + + P∘'→PF : Pushout (g ∘ f) fst → Pushout F fst + P∘'→PF (inl x) = inl (inr x) + P∘'→PF (inr x) = inr (inr x) + P∘'→PF (push (a , false) i) = (theCoh2 a false ∙ push (inr a , false)) i + P∘'→PF (push (a , true) i) = (theCoh2 a true ∙ push (inr a , true)) i + + PFpushTₗ : (x : _) → P∘'→PF (PF→P∘' (inl x)) ≡ (inl x) + PFpushTₗ (inl x) = push (inl tt , true) ∙ sym (push (inl tt , false)) + PFpushTₗ (inr x) = refl + + PFpushTᵣ : (x : _) → P∘'→PF (PF→P∘' (inr x)) ≡ (inr x) + PFpushTᵣ (inl x) = push (inl tt , true) + PFpushTᵣ (inr x) = refl + + pp1 : (x : A) → PFpushTₗ (f (x , false)) ≡ theCoh2 x false + pp1 x with (f (x , false)) + ... | inl x₁ = refl + ... | inr x₁ = refl + + pp2 : (x : A) → PFpushTₗ (f (x , true)) ≡ theCoh2 x true + pp2 x with (f (x , true)) + ... | inl x₁ = refl + ... | inr x₁ = refl + + open import Cubical.Foundations.Path + open import Cubical.Foundations.GroupoidLaws + + + PFpushT : (x : _) → P∘'→PF (PF→P∘' x) ≡ x + PFpushT (inl x) = PFpushTₗ x + PFpushT (inr x) = PFpushTᵣ x + PFpushT (push (inl x , false) i) j = + compPath-filler (push (inl tt , true)) (sym (push (inl tt , false))) (~ i) j + PFpushT (push (inr x , false) i) j = + (pp1 x + ◁ flipSquare + (symP (compPath-filler' + (theCoh2 x false) (push (inr x , false))))) i j + PFpushT (push (inl x , true) i) j = push (inl x , true) (i ∧ j) + PFpushT (push (inr x , true) i) j = + (pp2 x + ◁ flipSquare + (symP (compPath-filler' + (theCoh2 x true) (push (inr x , true))))) i j + + cong-PF→P∘' : (b : _) (a : _) → cong PF→P∘' (theCoh2 b a) ≡ refl + cong-PF→P∘' b a with (f (b , a)) + ... | inl x = cong-∙ PF→P∘' (push (inl tt , true)) (sym (push (inl tt , false))) + ∙ sym (rUnit refl) + ... | inr x = refl + + PF→P∘'→PF : (x : _) → PF→P∘' (P∘'→PF x) ≡ x + PF→P∘'→PF (inl x) = refl + PF→P∘'→PF (inr x) = refl + PF→P∘'→PF (push (b , false) i) j = + (cong-∙ PF→P∘' (theCoh2 b false) (push (inr b , false)) + ∙ cong (_∙ push (b , false)) (cong-PF→P∘' b false) + ∙ sym (lUnit _)) j i + PF→P∘'→PF (push (b , true) i) j = + (cong-∙ PF→P∘' (theCoh2 b true) (push (inr b , true)) + ∙ cong (_∙ push (b , true)) (cong-PF→P∘' b true) + ∙ sym (lUnit _)) j i + + Is1 : Iso (Pushout F fst) (Pushout (g ∘ f) fst) + Iso.fun Is1 = PF→P∘' + Iso.inv Is1 = P∘'→PF + Iso.rightInv Is1 = PF→P∘'→PF + Iso.leftInv Is1 = PFpushT + +FinPred : ∀ {m} → Fin (suc (suc m)) → Fin (suc m) +FinPred {m = m} (zero , p) = zero , p +FinPred {m = m} (suc x , p) = x , p + +fone : ∀ {m} → Fin (suc (suc m)) +fone {m} = suc zero , tt + +module _ {m : ℕ} where + Fin→Unit⊎Fin : (x : Fin (suc m)) → Unit ⊎ Fin m + Fin→Unit⊎Fin = Ind.elimFin (inl tt) inr + + Unit⊎Fin→Fin : Unit ⊎ Fin m → Fin (suc m) + Unit⊎Fin→Fin (inl x) = flast + Unit⊎Fin→Fin (inr x) = injectSuc x + + Iso-Fin-Unit⊎Fin : Iso (Fin (suc m)) (Unit ⊎ Fin m) + Iso.fun Iso-Fin-Unit⊎Fin = Fin→Unit⊎Fin + Iso.inv Iso-Fin-Unit⊎Fin = Unit⊎Fin→Fin + Iso.rightInv Iso-Fin-Unit⊎Fin (inl x) = Ind.elimFinβ (inl tt) inr .fst + Iso.rightInv Iso-Fin-Unit⊎Fin (inr x) = Ind.elimFinβ (inl tt) inr .snd x + Iso.leftInv Iso-Fin-Unit⊎Fin = + Ind.elimFin + (cong Unit⊎Fin→Fin (Ind.elimFinβ (inl tt) inr .fst)) + λ x → (cong Unit⊎Fin→Fin (Ind.elimFinβ (inl tt) inr .snd x)) + + +≠flast→<ᵗflast : {n : ℕ} → (x : Fin (suc n)) → ¬ x ≡ flast → fst x <ᵗ n +≠flast→<ᵗflast = Ind.elimFin (λ p → ⊥.rec (p refl)) λ p _ → snd p + +CW₁DataPre : (n m : ℕ) (f : Fin (suc (suc n)) × S₊ 0 → Fin (suc (suc m))) + → f (flast , false) ≡ flast + → (t : f (flast , true) .fst <ᵗ suc m) + → Σ[ k ∈ ℕ ] Σ[ f' ∈ (Fin k × S₊ 0 → Fin (suc m)) ] + Iso (Pushout f fst) + (Pushout f' fst) +CW₁DataPre n m f p q = (suc n) + , _ + , compIso (invIso (pushoutIso _ _ _ _ + (isoToEquiv (Σ-cong-iso-fst (invIso Iso-Fin-Unit⊎Fin))) + (isoToEquiv (invIso Iso-Fin-Unit⊎Fin)) + (isoToEquiv (invIso Iso-Fin-Unit⊎Fin)) + (funExt (uncurry help)) + (funExt λ x → refl))) + (Is1 {A = Fin (suc n)} {B = Fin (suc m)} + (λ x → Fin→Unit⊎Fin (f ((injectSuc (fst x)) , (snd x)))) + (f (flast , true) .fst , q)) + where + help : (y : Unit ⊎ Fin (suc n)) (x : Bool) + → Unit⊎Fin→Fin + (F (λ x₁ → Ind.elimFin (inl tt) inr (f (injectSuc (fst x₁) , snd x₁))) + (f (flast , true) .fst , q) (y , x)) + ≡ f (Unit⊎Fin→Fin y , x) + help (inl a) false = sym p + help (inl b) true = Σ≡Prop (λ _ → isProp<ᵗ) refl + help (inr a) false = Iso.leftInv Iso-Fin-Unit⊎Fin _ + help (inr a) true = Iso.leftInv Iso-Fin-Unit⊎Fin _ + +isPropFin1 : isProp (Fin 1) +isPropFin1 (zero , tt) (zero , tt) = refl + +Iso⊎→Iso : ∀ {ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → (f : Iso A C) + → (e : Iso (A ⊎ B) (C ⊎ D)) + → ((a : A) → Iso.fun e (inl a) ≡ inl (Iso.fun f a)) + → Iso B D +Iso⊎→Iso {A = A} {B} {C} {D} f e p = Iso' + where + ⊥-fib : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ⊎ B → Type + ⊥-fib (inl x) = ⊥ + ⊥-fib (inr x) = Unit + + module _ {ℓ ℓ' ℓ'' ℓ''' : Level} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + (f : Iso A C) + (e : Iso (A ⊎ B) (C ⊎ D)) + (p : (a : A) → Iso.fun e (inl a) ≡ inl (Iso.fun f a)) where + T : (b : B) → Type _ + T b = Σ[ d' ∈ C ⊎ D ] (Iso.fun e (inr b) ≡ d') + + T-elim : ∀ {ℓ} (b : B) {P : (x : T b) → Type ℓ} + → ((d : D) (s : _) → P (inr d , s)) + → (x : _) → P x + T-elim b ind (inl x , q) = + ⊥.rec (subst ⊥-fib (sym (sym (Iso.leftInv e _) + ∙ cong (Iso.inv e) + (p _ ∙ cong inl (Iso.rightInv f x) ∙ sym q) + ∙ Iso.leftInv e _)) tt) + T-elim b ind (inr x , y) = ind x y + + e-pres-inr-help : (b : B) → T f e p b → D + e-pres-inr-help b = T-elim f e p b λ d _ → d + + p' : (a : C) → Iso.inv e (inl a) ≡ inl (Iso.inv f a) + p' c = cong (Iso.inv e ∘ inl) (sym (Iso.rightInv f c)) + ∙∙ cong (Iso.inv e) (sym (p (Iso.inv f c))) + ∙∙ Iso.leftInv e _ + + e⁻-pres-inr-help : (d : D) → T (invIso f) (invIso e) p' d → B + e⁻-pres-inr-help d = T-elim (invIso f) (invIso e) p' d λ b _ → b + + e-pres-inr : B → D + e-pres-inr b = e-pres-inr-help b (_ , refl) + + e⁻-pres-inr : D → B + e⁻-pres-inr d = e⁻-pres-inr-help d (_ , refl) + + lem1 : (b : B) (e : T f e p b) (d : _) + → e⁻-pres-inr-help (e-pres-inr-help b e) d ≡ b + lem1 b = T-elim f e p b λ d s + → T-elim (invIso f) (invIso e) p' _ + λ b' s' → invEq (_ , isEmbedding-inr _ _) + (sym s' ∙ cong (Iso.inv e) (sym s) ∙ Iso.leftInv e _) + + lem2 : (d : D) (e : T (invIso f) (invIso e) p' d ) (t : _) + → e-pres-inr-help (e⁻-pres-inr-help d e) t ≡ d + lem2 d = T-elim (invIso f) (invIso e) p' d + λ b s → T-elim f e p _ λ d' s' + → invEq (_ , isEmbedding-inr _ _) + (sym s' ∙ cong (Iso.fun e) (sym s) ∙ Iso.rightInv e _) + + Iso' : Iso B D + Iso.fun Iso' = e-pres-inr + Iso.inv Iso' = e⁻-pres-inr + Iso.rightInv Iso' x = lem2 x _ _ + Iso.leftInv Iso' x = lem1 x _ _ + +Fin≠Fin : {n m : ℕ} → ¬ (n ≡ m) → ¬ (Iso (Fin n) (Fin m)) +Fin≠Fin {n = zero} {m = zero} p = ⊥.rec (p refl) +Fin≠Fin {n = zero} {m = suc m} p q = Iso.inv q fzero .snd +Fin≠Fin {n = suc n} {m = zero} p q = Iso.fun q fzero .snd +Fin≠Fin {n = suc n} {m = suc m} p q = + Fin≠Fin {n = n} {m = m} (p ∘ cong suc) + (Iso⊎→Iso idIso help λ {(zero , tt) + → cong (Iso.inv FinSuc) (swapFinβₗ (Iso.fun q flast) flast) + ∙ Ind.elimFinβ (inl flast) inr .fst}) + where + q^ : Iso (Fin (suc n)) (Fin (suc m)) + q^ = compIso q (swapFinIso (Iso.fun q flast) flast) + + help : Iso (Fin 1 ⊎ Fin n) (Fin 1 ⊎ Fin m) + help = compIso FinSuc (compIso q^ (invIso FinSuc)) + +CW₁Data₁ : (m : ℕ) (f : Fin 1 × S₊ 0 → Fin (suc (suc m))) + → isConnected 2 (Pushout f fst) + → Iso (Fin 1 × S₊ 0) (Fin (suc (suc m))) +CW₁Data₁ m f c = mainIso -- mainIso + where + f' : Bool → Fin (suc (suc m)) + f' = f ∘ (fzero ,_) + + f'-surj : (x : _) → Σ[ t ∈ Bool ] (f' t ≡ x) + f'-surj x = + isSurj-α₀ (suc zero) m f c x .fst .snd + , cong f (ΣPathP (isPropFin1 _ _ , refl)) ∙ isSurj-α₀ (suc zero) m f c x .snd + f-true≠f-false : (x : _) → f' true ≡ x → ¬ f' true ≡ f' false + f-true≠f-false (zero , q) p r = lem (f'-surj fone) + where + lem : Σ[ t ∈ Bool ] (f' t ≡ fone) → ⊥ + lem (false , s) = snotz (cong fst (sym s ∙ sym r ∙ p)) + lem (true , s) = snotz (cong fst (sym s ∙ p)) + f-true≠f-false (suc x , q) p r = lem (f'-surj fzero) + where + lem : Σ[ t ∈ Bool ] (f' t ≡ fzero) → ⊥ + lem (false , s) = snotz (cong fst (sym p ∙ r ∙ s)) + lem (true , s) = snotz (cong fst (sym p ∙ s)) + + f-inj : (x y : _) → f x ≡ f y → x ≡ y + f-inj ((zero , tt) , false) ((zero , tt) , false) p = refl + f-inj ((zero , tt) , false) ((zero , tt) , true) p = ⊥.rec (f-true≠f-false _ refl (sym p)) + f-inj ((zero , tt) , true) ((zero , tt) , false) p = ⊥.rec (f-true≠f-false _ refl p) + f-inj ((zero , tt) , true) ((zero , tt) , true) p = refl + + mainIso : Iso (Fin 1 × S₊ 0) (Fin (suc (suc m))) + Iso.fun mainIso = f + Iso.inv mainIso x = isSurj-α₀ (suc zero) m f c x .fst + Iso.rightInv mainIso x = isSurj-α₀ 1 m f c x .snd + Iso.leftInv mainIso ((zero , tt) , x) = + (f-inj _ _ (isSurj-α₀ 1 m f c (f (fzero , x)) .snd)) + +CW₁Data : (n m : ℕ) (f : Fin n × S₊ 0 → Fin (suc (suc m))) + → isConnected 2 (Pushout f fst) + → Σ[ k ∈ ℕ ] Σ[ f' ∈ (Fin k × S₊ 0 → Fin (suc m)) ] + Iso (Pushout f fst) + (Pushout f' fst) +CW₁Data zero m f c = ⊥.rec (snd (notAllLoops-α₀ zero m f c .fst)) +CW₁Data (suc zero) zero f c = 0 , ((λ ()) , compIso isoₗ (PushoutEmptyFam (snd ∘ fst) snd)) + where + isoₗ : Iso (Pushout f fst) (Fin 1) + isoₗ = PushoutAlongEquiv (isoToEquiv (CW₁Data₁ zero f c)) fst +CW₁Data (suc zero) (suc m) f c = + ⊥.rec (Fin≠Fin (snotz ∘ sym ∘ cong (predℕ ∘ predℕ)) + mainIso) + where + mainIso : Iso (Fin 2) (Fin (3 +ℕ m)) + mainIso = + compIso + (compIso + (invIso rUnit×Iso) + (compIso + (Σ-cong-iso + (invIso Iso-Bool-Fin) + (λ _ → isContr→Iso (tt , (λ _ → refl)) + (inhProp→isContr fzero isPropFin1))) + Σ-swap-Iso)) + (CW₁Data₁ (suc m) f c) +CW₁Data (suc (suc n)) m f c = + main .fst , main .snd .fst + , compIso correct (main .snd .snd) + where + t = notAllLoops-α₀ (suc (suc n)) m f c + + abstract + x₀ : Fin (suc (suc n)) + x₀ = fst t + + xpath : ¬ (f (x₀ , true) ≡ f (x₀ , false)) + xpath = snd t + + Fin0-iso : Iso (Fin (suc (suc n)) × S₊ 0) (Fin (suc (suc n)) × S₊ 0) + Fin0-iso = Σ-cong-iso-fst (swapFinIso flast x₀) + + FinIso2 : Iso (Fin (suc (suc m))) (Fin (suc (suc m))) + FinIso2 = swapFinIso (f (x₀ , false)) flast + + f' : Fin (suc (suc n)) × S₊ 0 → Fin (suc (suc m)) + f' = Iso.fun FinIso2 ∘ f ∘ Iso.fun Fin0-iso + + f'≡ : f' (flast , false) ≡ flast + f'≡ = cong (Iso.fun FinIso2 ∘ f) + (cong (_, false) (swapFinβₗ flast x₀)) + ∙ swapFinβₗ (f (x₀ , false)) flast + + f'¬≡ : ¬ (f' (flast , true) ≡ flast) + f'¬≡ p = xpath (cong f (ΣPathP (sym (swapFinβₗ flast x₀) , refl)) + ∙ sym (Iso.rightInv FinIso2 _) + ∙ cong (Iso.inv FinIso2) (p ∙ sym f'≡) + ∙ Iso.rightInv FinIso2 _ + ∙ cong f (ΣPathP (swapFinβₗ flast x₀ , refl))) + + f'< : fst (f' (flast , true)) <ᵗ suc m + f'< = ≠flast→<ᵗflast _ f'¬≡ + + main = CW₁DataPre _ _ f' f'≡ f'< + + Upath = isoToPath (swapFinIso x₀ fzero) + + correct : Iso (Pushout f fst) (Pushout f' fst) + correct = pushoutIso _ _ _ _ + (isoToEquiv Fin0-iso) (isoToEquiv FinIso2) (isoToEquiv (swapFinIso flast x₀)) + (funExt (λ x → cong (FinIso2 .Iso.fun ∘ f) (sym (Iso.rightInv Fin0-iso x)))) + refl + + +CW₁Data' : (n m : ℕ) (f : Fin n × S₊ 0 → Fin m) + → isConnected 2 (Pushout f fst) + → Σ[ k ∈ ℕ ] + Iso (Pushout f fst) + (Pushout {A = Fin k × S₊ 0} {B = Fin 1} (λ _ → fzero) fst) +CW₁Data' zero zero f c = ⊥.rec (TR.rec (λ()) help (fst c)) + where + help : ¬ Pushout f fst + help = elimProp _ (λ _ → λ ()) snd snd +CW₁Data' (suc n) zero f c = ⊥.rec (f (fzero , true) .snd) +CW₁Data' n (suc zero) f c = n + , pushoutIso _ _ _ _ (idEquiv _) (idEquiv _) (idEquiv _) + (funExt (λ x → isPropFin1 _ _)) refl +CW₁Data' zero (suc (suc m)) f c = + ⊥.rec (TR.rec (λ()) (snotz ∘ sym ∘ cong fst) + (Iso.fun (PathIdTruncIso _) + (isContr→isProp (subst (isConnected 2) (isoToPath help) c) + ∣ fzero ∣ ∣ fone ∣))) + where + help : Iso (Pushout f fst) (Fin (suc (suc m))) + help = invIso (PushoutEmptyFam (λ()) λ()) +CW₁Data' (suc n) (suc (suc m)) f c + with (CW₁Data' _ (suc m) (CW₁Data (suc n) m f c .snd .fst) + (subst (isConnected 2) + (isoToPath (CW₁Data (suc n) m f c .snd .snd)) c)) +... | (k , e) = k , compIso (CW₁Data (suc n) m f c .snd .snd) e + +yieldsConnectedCWskel : (ℕ → Type ℓ) → Type _ +yieldsConnectedCWskel A = Σ[ sk ∈ yieldsCWskel A ] (sk .fst 0 ≡ 1) + +yieldsConnectedCWskel' : (ℕ → Type ℓ) → Type _ +yieldsConnectedCWskel' A = Σ[ sk ∈ yieldsCWskel A ] isConnected 2 (realise (_ , sk)) + +connectedCWskel : (ℓ : Level) → Type (ℓ-suc ℓ) +connectedCWskel ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel X) + +connectedCWskel' : (ℓ : Level) → Type (ℓ-suc ℓ) +connectedCWskel' ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel' X) + +truncCWIso : (A : CWskel ℓ) (n : ℕ) + → Iso (hLevelTrunc n (realise A)) (hLevelTrunc n (fst A n)) +truncCWIso A n = invIso (connectedTruncIso n incl (isConnected-CW↪∞ n A)) + +isConnectedColim→isConnectedSkel : + (A : CWskel ℓ) (n : ℕ) + → isConnected n (realise A) + → isConnected n (fst A n) +isConnectedColim→isConnectedSkel A n c = + isOfHLevelRetractFromIso 0 + (invIso (truncCWIso A n)) c + +module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A) where + private + c = snd sk+c + sk = fst sk+c + c' = isConnectedColim→isConnectedSkel (_ , sk) 2 c + + module AC = CWskel-fields (_ , sk) + + e₁ : Iso (Pushout (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) fst) (A 2) + e₁ = compIso (PushoutPostcompEquivIso (CW₁-discrete (A , sk)) (AC.α 1) fst) + (equivToIso (invEquiv (AC.e 1))) + + liftStr = CW₁Data' _ _ (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) + (isConnectedRetractFromIso 2 e₁ c') + + collapse₁card : ℕ → ℕ + collapse₁card zero = 1 + collapse₁card (suc zero) = fst liftStr + collapse₁card (suc (suc x)) = AC.card (suc (suc x)) + + collapse₁CWskel : ℕ → Type _ + collapse₁CWskel zero = Lift ⊥ + collapse₁CWskel (suc zero) = Lift (Fin 1) + collapse₁CWskel (suc (suc n)) = A (suc (suc n)) + + collapse₁α : (n : ℕ) → Fin (collapse₁card n) × S⁻ n → collapse₁CWskel n + collapse₁α (suc zero) (x , p) = lift fzero + collapse₁α (suc (suc n)) = AC.α (2 +ℕ n) + + connectedCWskel→ : yieldsConnectedCWskel collapse₁CWskel + fst (fst connectedCWskel→) = collapse₁card + fst (snd (fst connectedCWskel→)) = collapse₁α + fst (snd (snd (fst connectedCWskel→))) = λ() + snd (snd (snd (fst connectedCWskel→))) zero = + isContr→Equiv (isOfHLevelLift 0 (inhProp→isContr fzero isPropFin1)) + ((inr fzero) + , (λ { (inr x) j → inr (isPropFin1 fzero x j) })) + snd (snd (snd (fst connectedCWskel→))) (suc zero) = + compEquiv (invEquiv (isoToEquiv e₁)) + (compEquiv (isoToEquiv (snd liftStr)) + (isoToEquiv (pushoutIso' _ _ _ _ + (idEquiv _) LiftEquiv (idEquiv _) + (funExt cohₗ) (funExt cohᵣ)))) + where + -- Agda complains if these proofs are inlined... + cohₗ : (x : _) → collapse₁α 1 x ≡ collapse₁α 1 x + cohₗ (x , p) = refl + + cohᵣ : (x : Fin (fst liftStr) × Bool) → fst x ≡ fst x + cohᵣ x = refl + snd (snd (snd (fst connectedCWskel→))) (suc (suc n)) = AC.e (suc (suc n)) + snd connectedCWskel→ = refl + + collapse₁Equiv : (n : ℕ) + → realise (_ , sk) ≃ realise (_ , connectedCWskel→ .fst) + collapse₁Equiv n = {!!} + + + +-- yieldsGoodCWSkel : (ℕ → Type ℓ) → Type ℓ +-- yieldsGoodCWSkel = ? + +-- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ +-- yieldsCWskel∙ X = +-- Σ[ f ∈ (ℕ → ℕ) ] +-- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] +-- ((X zero ≃ Fin (f zero)) × +-- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) + +-- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) +-- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) + +-- module CWskel∙-fields (C : CWskel∙ ℓ) where +-- card = C .snd .fst +-- A = Fin ∘ card +-- α = C .snd .snd .fst +-- e = C .snd .snd .snd .snd + +-- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero +-- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] + +-- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- fst (CWpt (C , f) n) = C n +-- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) + +-- -- Technically, a CW complex should be the sequential colimit over the following maps +-- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) +-- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) + +-- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n +-- ptCW T zero = T .snd .snd .fst zero (inl tt) +-- ptCW T (suc n) = CW∙↪ T n (ptCW T n) + +-- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- CW∙ T n = fst T n , ptCW T n + +-- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) +-- fst (CW∙↪∙ T n) = CW∙↪ T n +-- snd (CW∙↪∙ T n) = refl + + +-- -- But if it stabilises, no need for colimits. +-- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ +-- yieldsFinCWskel∙ n X = +-- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) + +-- -- ... which should give us finite CW complexes. +-- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) +-- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) + +-- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ +-- finCWskel→CWskel∙ n C = fst C , fst (snd C) + +-- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ +-- Sequence.obj (realiseSeq∙ (C , p)) = C +-- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ + +-- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ +-- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) + +-- -- realisation of CW complex from skeleton +-- realise∙ : CWskel∙ ℓ → Type ℓ +-- realise∙ C = SeqColim (realiseSeq∙ C) + +-- realise∙∙ : CWskel∙ ℓ → Pointed ℓ +-- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) +-- open import Cubical.Data.Empty as ⊥ + +-- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ +-- CWskel∙→CWskel A zero = Lift ⊥ +-- CWskel∙→CWskel A (suc n) = A n +-- open import Cubical.Foundations.Isomorphism + + +-- module _ (A : ℕ → Type ℓ) +-- (cwsk : yieldsCWskel∙ A) where + +-- private +-- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n +-- αs (suc n) x = snd cwsk .fst n (inr x) + +-- e0 : {!!} +-- e0 = {!!} + +-- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst +-- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) +-- es-suc→ n (inr x) = inl x +-- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) +-- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i +-- es-suc→ n (push (push a i₁) i) = {!!} + +-- es-suc : (n : ℕ) +-- → Iso (cofib (fst (snd cwsk) n)) +-- (Pushout (αs (suc n)) fst) +-- Iso.fun (es-suc n) = es-suc→ n +-- Iso.inv (es-suc n) = {!!} +-- Iso.rightInv (es-suc n) = {!!} +-- Iso.leftInv (es-suc n) = {!!} + +-- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) +-- es zero = {!!} +-- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) + +-- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) +-- fst yieldsCWskel∙→' = cwsk .fst +-- fst (snd yieldsCWskel∙→') = αs +-- fst (snd (snd yieldsCWskel∙→')) () +-- snd (snd (snd yieldsCWskel∙→')) = {!!} + +-- yieldsCWskel∙→ : (A : ℕ → Type ℓ) +-- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) +-- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst +-- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) +-- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () +-- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} +-- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = +-- compEquiv (cwsk .snd .snd .snd n) +-- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) +-- where +-- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst +-- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) +-- theEq→ (inr x) = inl x +-- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) +-- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) +-- theEq→ (push (push a i₁) i) = {!!} + +-- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) +-- Iso.fun theEq = theEq→ +-- Iso.inv theEq = {!!} +-- Iso.rightInv theEq x = {!!} +-- Iso.leftInv theEq x = {!!} + + +-- -- compEquiv {!!} {!!} + + +-- -- -- Finally: definition of CW complexes +-- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' + +-- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ + +-- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) + +-- -- -- Finite CW complexes +-- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- isFinCW {ℓ = ℓ} X = +-- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) + +-- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ + +-- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) + +-- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X +-- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str + +-- -- finCW→CW : finCW ℓ → CW ℓ +-- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p + + +-- -- -- morphisms +-- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') +-- -- C →ᶜʷ D = fst C → fst D + +-- -- --the cofibre of the inclusion of X n into X (suc n) +-- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ +-- -- cofibCW n C = cofib (CW↪ C n) + +-- -- --...is basically a sphere bouquet +-- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) +-- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) +-- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e +-- -- where +-- -- s = Bouquet≃-gen +-- -- α = C .snd .snd .fst +-- -- e = C .snd .snd .snd .snd n + +-- -- --sending X (suc n) into the cofibCW +-- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C +-- -- to_cofibCW n C x = inr x + +-- -- --the connecting map of the long exact sequence +-- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) +-- -- → cofib conn → Susp A +-- -- δ-pre conn (inl x) = north +-- -- δ-pre conn (inr x) = south +-- -- δ-pre conn (push a i) = merid a i + +-- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) +-- -- δ n C = δ-pre (CW↪ C n) + +-- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) +-- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C +-- -- CW↪∞ C n x = incl x + +-- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m +-- -- fst (finCW↑ m n p C) = fst C +-- -- fst (snd (finCW↑ m n p C)) = snd C .fst +-- -- snd (snd (finCW↑ m n p C)) k = +-- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) +-- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) +-- -- (snd C .snd (k +ℕ fst p)) From dac2fd33e3aaabcb936104afc4e0685f130c0dc6 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 23 May 2024 18:08:06 +0200 Subject: [PATCH 42/73] stuff --- Cubical/CW/Pointed2.agda | 507 +++++++++++------- Cubical/CW/Subcomplex.agda | 34 ++ Cubical/Data/Sequence/Base.agda | 26 + Cubical/Data/Sequence/Properties.agda | 187 ++++++- .../HITs/SequentialColimit/Properties.agda | 105 +++- 5 files changed, 651 insertions(+), 208 deletions(-) diff --git a/Cubical/CW/Pointed2.agda b/Cubical/CW/Pointed2.agda index 418986e3c0..eb82ebabb1 100644 --- a/Cubical/CW/Pointed2.agda +++ b/Cubical/CW/Pointed2.agda @@ -23,6 +23,8 @@ open import Cubical.HITs.SphereBouquet open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.Wedge +open import Cubical.Axiom.Choice + open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup @@ -934,17 +936,18 @@ CW₁Data' (suc n) (suc (suc m)) f c (isoToPath (CW₁Data (suc n) m f c .snd .snd)) c)) ... | (k , e) = k , compIso (CW₁Data (suc n) m f c .snd .snd) e -yieldsConnectedCWskel : (ℕ → Type ℓ) → Type _ -yieldsConnectedCWskel A = Σ[ sk ∈ yieldsCWskel A ] (sk .fst 0 ≡ 1) +yieldsConnectedCWskel : (A : ℕ → Type ℓ) (n : ℕ) → Type _ +yieldsConnectedCWskel A k = + Σ[ sk ∈ yieldsCWskel A ] ((sk .fst 0 ≡ 1) × ((n : ℕ) → n <ᵗ k → sk .fst (suc n) ≡ 0)) -yieldsConnectedCWskel' : (ℕ → Type ℓ) → Type _ -yieldsConnectedCWskel' A = Σ[ sk ∈ yieldsCWskel A ] isConnected 2 (realise (_ , sk)) +yieldsConnectedCWskel' : (A : ℕ → Type ℓ) (n : ℕ) → Type _ +yieldsConnectedCWskel' A n = Σ[ sk ∈ yieldsCWskel A ] isConnected (2 +ℕ n) (realise (_ , sk)) -connectedCWskel : (ℓ : Level) → Type (ℓ-suc ℓ) -connectedCWskel ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel X) +connectedCWskel : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) +connectedCWskel ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel X n) -connectedCWskel' : (ℓ : Level) → Type (ℓ-suc ℓ) -connectedCWskel' ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel' X) +connectedCWskel' : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) +connectedCWskel' ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel' X n) truncCWIso : (A : CWskel ℓ) (n : ℕ) → Iso (hLevelTrunc n (realise A)) (hLevelTrunc n (fst A n)) @@ -958,7 +961,7 @@ isConnectedColim→isConnectedSkel A n c = isOfHLevelRetractFromIso 0 (invIso (truncCWIso A n)) c -module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A) where +module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A 0) where private c = snd sk+c sk = fst sk+c @@ -987,7 +990,7 @@ module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A) where collapse₁α (suc zero) (x , p) = lift fzero collapse₁α (suc (suc n)) = AC.α (2 +ℕ n) - connectedCWskel→ : yieldsConnectedCWskel collapse₁CWskel + connectedCWskel→ : yieldsConnectedCWskel collapse₁CWskel 0 fst (fst connectedCWskel→) = collapse₁card fst (snd (fst connectedCWskel→)) = collapse₁α fst (snd (snd (fst connectedCWskel→))) = λ() @@ -1009,218 +1012,324 @@ module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A) where cohᵣ : (x : Fin (fst liftStr) × Bool) → fst x ≡ fst x cohᵣ x = refl snd (snd (snd (fst connectedCWskel→))) (suc (suc n)) = AC.e (suc (suc n)) - snd connectedCWskel→ = refl + snd connectedCWskel→ = refl , (λ _ → λ ()) collapse₁Equiv : (n : ℕ) → realise (_ , sk) ≃ realise (_ , connectedCWskel→ .fst) - collapse₁Equiv n = {!!} - + collapse₁Equiv n = + compEquiv + (isoToEquiv (Iso-SeqColim→SeqColimShift _ 3)) + (compEquiv (pathToEquiv (cong SeqColim + (cong₂ sequence (λ _ m → A (3 +ℕ m)) + λ _ {z} → CW↪ (A , sk) (suc (suc (suc z)))))) + (invEquiv (isoToEquiv (Iso-SeqColim→SeqColimShift _ 3)))) + + +connectedCWskel'→connectedCWskel : ∀ {ℓ} → connectedCWskel' ℓ 0 → connectedCWskel ℓ 0 +connectedCWskel'→connectedCWskel (A , sk) = + _ , connectedCWskel→ A ((sk .fst) , (sk .snd)) .fst , refl , (λ _ → λ()) + +yieldsGoodCWskel : {ℓ : Level} (A₊₂ : ℕ → Pointed ℓ) → Type _ +yieldsGoodCWskel {ℓ = ℓ} A = + Σ[ f₊₁ ∈ (ℕ → ℕ) ] + Σ[ fin ∈ (A 0) .fst ≃ Fin 1 ] + Σ[ α ∈ ((n : ℕ) → SphereBouquet∙ n (Fin (f₊₁ n)) →∙ A n) ] + ((n : ℕ) → cofib (α n .fst) , inl tt ≃∙ A (suc n)) + +GoodCWskelSeq : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Sequence ℓ +obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) zero = Lift ⊥ +obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) (suc n) = fst (A n) +Sequence.map (GoodCWskelSeq {A = A} (f , fin , α , sq)) {n = suc n} x = fst (fst (sq n)) (inr x) + +realiseGood∙ : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Pointed ℓ +realiseGood∙ {A = A} S = SeqColim (GoodCWskelSeq S) , incl {n = 1} (snd (A 0)) + +module _ {ℓ : Level} (A' : ℕ → Type ℓ) + (con : isConnected 0 (A' 2)) + (C : yieldsCWskel A') + (dim : ℕ) (ind : (n : ℕ) → dim <ᵗ n → fst C n ≡ 0) + (cA : fst C 0 ≡ 1) + where + open CWskel-fields (_ , C) + e₀ : A' 1 ≃ Fin (card 0) + e₀ = CW₁-discrete (_ , C) --- yieldsGoodCWSkel : (ℕ → Type ℓ) → Type ℓ --- yieldsGoodCWSkel = ? + pt0 : A' 1 + pt0 = invEq e₀ (subst Fin (sym cA) flast) + + ptA : (n : ℕ) → A' (suc n) + ptA zero = pt0 + ptA (suc n) = CW↪ (_ , C) (suc n) (ptA n) + + mapMakerNil : ∥ ((n : Fin (suc dim)) + → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) + →∙ (A' (suc (fst n)) , ptA (fst n))) ] ((x : _) → fst h (inr x) ≡ α _ x)) ∥₁ + mapMakerNil = + invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) + λ n → {!!}) + where + help : (n : Fin (suc dim)) → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) → ({!CW↪ ?!} ≡ {!!})) + help = {!!} +{- + mapMaker : (n : ℕ) → SphereBouquet∙ n (Fin (card (suc n))) + →∙ (A' (suc n) , ptA n) + mapMaker n with (n ≟ᵗ dim) + ... | lt x = mapMakerNil (n , <ᵗ-trans-suc x) + ... | eq x = mapMakerNil (n , subst (_<ᵗ suc dim) (sym x) <ᵗsucm) + ... | gt x = const∙ _ _ +-} + getGoodCWskelAux : yieldsGoodCWskel (λ n → A' (suc n) , ptA n) + fst getGoodCWskelAux = card + fst (snd getGoodCWskelAux) = {!!} -- CW₁-discrete (_ , C) + fst (snd (snd getGoodCWskelAux)) = {!!} -- mapMaker + snd (snd (snd getGoodCWskelAux)) n = {!!} + +connectedGoodCWskel' : ∀ {ℓ} (Aₙ₊₂ : ℕ → Pointed ℓ) → {!!} +connectedGoodCWskel' = {!!} + +isGoodCW : ∀ {ℓ} (A : Pointed ℓ) → Pointed (ℓ-suc ℓ) +isGoodCW A = {!!} + +open import Cubical.CW.Subcomplex + +GoodCWSeq→CWSeq : {ℓ : Level} (A₊₂ : ℕ → Pointed ℓ) → {!!} +GoodCWSeq→CWSeq A₊₂ = {!!} + +isGoodCW→isCW : ∀ {ℓ} (A₊₂ : ℕ → Pointed ℓ) → yieldsGoodCWskel A₊₂ → {!!} +isGoodCW→isCW = {!!} + +-- connectedCWskel→CWskel : connectedCWskel ℓ → CWskel ℓ +-- connectedCWskel→CWskel (A , sk , _) = A , sk + +-- module _ (C' : connectedCWskel ℓ) (n : ℕ) (x₀ : fst C' 1) +-- (term : converges (realiseSeq (fst C' , snd C' .fst)) n) where +-- private +-- C = fst C' +-- Csk = connectedCWskel→CWskel C' --- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ --- yieldsCWskel∙ X = --- Σ[ f ∈ (ℕ → ℕ) ] --- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] --- ((X zero ≃ Fin (f zero)) × --- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) +-- open CWskel-fields Csk --- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) --- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) +-- ptC : (n : ℕ) → C (suc n) +-- ptC zero = x₀ +-- ptC (suc n) = CW↪ Csk (suc n) (ptC n) --- module CWskel∙-fields (C : CWskel∙ ℓ) where --- card = C .snd .fst --- A = Fin ∘ card --- α = C .snd .snd .fst --- e = C .snd .snd .snd .snd +-- C∙ : ℕ → Pointed _ +-- C∙ n = C (suc n) , ptC n --- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero --- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] +-- improveFam : ∥ ((m : Fin (suc (suc n))) → ?) ∥₁ +-- improveFam = {!!} --- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- fst (CWpt (C , f) n) = C n --- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) +-- finConnectedCW : {!!} +-- finConnectedCW = {!!} --- -- Technically, a CW complex should be the sequential colimit over the following maps --- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) --- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) --- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n --- ptCW T zero = T .snd .snd .fst zero (inl tt) --- ptCW T (suc n) = CW∙↪ T n (ptCW T n) --- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- CW∙ T n = fst T n , ptCW T n --- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) --- fst (CW∙↪∙ T n) = CW∙↪ T n --- snd (CW∙↪∙ T n) = refl +-- -- yieldsGoodCWSkel : (ℕ → Type ℓ) → Type ℓ +-- -- yieldsGoodCWSkel = ? +-- -- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ +-- -- yieldsCWskel∙ X = +-- -- Σ[ f ∈ (ℕ → ℕ) ] +-- -- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] +-- -- ((X zero ≃ Fin (f zero)) × +-- -- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) --- -- But if it stabilises, no need for colimits. --- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ --- yieldsFinCWskel∙ n X = --- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) +-- -- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) --- -- ... which should give us finite CW complexes. --- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) --- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) +-- -- module CWskel∙-fields (C : CWskel∙ ℓ) where +-- -- card = C .snd .fst +-- -- A = Fin ∘ card +-- -- α = C .snd .snd .fst +-- -- e = C .snd .snd .snd .snd + +-- -- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero +-- -- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] + +-- -- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- -- fst (CWpt (C , f) n) = C n +-- -- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) + +-- -- -- Technically, a CW complex should be the sequential colimit over the following maps +-- -- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) +-- -- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) + +-- -- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n +-- -- ptCW T zero = T .snd .snd .fst zero (inl tt) +-- -- ptCW T (suc n) = CW∙↪ T n (ptCW T n) + +-- -- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ +-- -- CW∙ T n = fst T n , ptCW T n + +-- -- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) +-- -- fst (CW∙↪∙ T n) = CW∙↪ T n +-- -- snd (CW∙↪∙ T n) = refl + + +-- -- -- But if it stabilises, no need for colimits. +-- -- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ +-- -- yieldsFinCWskel∙ n X = +-- -- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) + +-- -- -- ... which should give us finite CW complexes. +-- -- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) +-- -- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) + +-- -- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ +-- -- finCWskel→CWskel∙ n C = fst C , fst (snd C) + +-- -- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ +-- -- Sequence.obj (realiseSeq∙ (C , p)) = C +-- -- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ + +-- -- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ +-- -- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) + +-- -- -- realisation of CW complex from skeleton +-- -- realise∙ : CWskel∙ ℓ → Type ℓ +-- -- realise∙ C = SeqColim (realiseSeq∙ C) + +-- -- realise∙∙ : CWskel∙ ℓ → Pointed ℓ +-- -- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) +-- -- open import Cubical.Data.Empty as ⊥ + +-- -- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ +-- -- CWskel∙→CWskel A zero = Lift ⊥ +-- -- CWskel∙→CWskel A (suc n) = A n +-- -- open import Cubical.Foundations.Isomorphism + + +-- -- module _ (A : ℕ → Type ℓ) +-- -- (cwsk : yieldsCWskel∙ A) where + +-- -- private +-- -- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n +-- -- αs (suc n) x = snd cwsk .fst n (inr x) + +-- -- e0 : {!!} +-- -- e0 = {!!} + +-- -- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst +-- -- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) +-- -- es-suc→ n (inr x) = inl x +-- -- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) +-- -- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i +-- -- es-suc→ n (push (push a i₁) i) = {!!} + +-- -- es-suc : (n : ℕ) +-- -- → Iso (cofib (fst (snd cwsk) n)) +-- -- (Pushout (αs (suc n)) fst) +-- -- Iso.fun (es-suc n) = es-suc→ n +-- -- Iso.inv (es-suc n) = {!!} +-- -- Iso.rightInv (es-suc n) = {!!} +-- -- Iso.leftInv (es-suc n) = {!!} + +-- -- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) +-- -- es zero = {!!} +-- -- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) + +-- -- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) +-- -- fst yieldsCWskel∙→' = cwsk .fst +-- -- fst (snd yieldsCWskel∙→') = αs +-- -- fst (snd (snd yieldsCWskel∙→')) () +-- -- snd (snd (snd yieldsCWskel∙→')) = {!!} + +-- -- yieldsCWskel∙→ : (A : ℕ → Type ℓ) +-- -- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) +-- -- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst +-- -- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) +-- -- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () +-- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} +-- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = +-- -- compEquiv (cwsk .snd .snd .snd n) +-- -- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) +-- -- where +-- -- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst +-- -- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) +-- -- theEq→ (inr x) = inl x +-- -- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) +-- -- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) +-- -- theEq→ (push (push a i₁) i) = {!!} --- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ --- finCWskel→CWskel∙ n C = fst C , fst (snd C) +-- -- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) +-- -- Iso.fun theEq = theEq→ +-- -- Iso.inv theEq = {!!} +-- -- Iso.rightInv theEq x = {!!} +-- -- Iso.leftInv theEq x = {!!} --- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ --- Sequence.obj (realiseSeq∙ (C , p)) = C --- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ --- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ --- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) +-- -- -- compEquiv {!!} {!!} --- -- realisation of CW complex from skeleton --- realise∙ : CWskel∙ ℓ → Type ℓ --- realise∙ C = SeqColim (realiseSeq∙ C) --- realise∙∙ : CWskel∙ ℓ → Pointed ℓ --- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) --- open import Cubical.Data.Empty as ⊥ +-- -- -- -- Finally: definition of CW complexes +-- -- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' --- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ --- CWskel∙→CWskel A zero = Lift ⊥ --- CWskel∙→CWskel A (suc n) = A n --- open import Cubical.Foundations.Isomorphism +-- -- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ +-- -- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) --- module _ (A : ℕ → Type ℓ) --- (cwsk : yieldsCWskel∙ A) where +-- -- -- -- Finite CW complexes +-- -- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +-- -- -- isFinCW {ℓ = ℓ} X = +-- -- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) --- private --- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n --- αs (suc n) x = snd cwsk .fst n (inr x) - --- e0 : {!!} --- e0 = {!!} - --- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst --- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) --- es-suc→ n (inr x) = inl x --- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) --- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i --- es-suc→ n (push (push a i₁) i) = {!!} - --- es-suc : (n : ℕ) --- → Iso (cofib (fst (snd cwsk) n)) --- (Pushout (αs (suc n)) fst) --- Iso.fun (es-suc n) = es-suc→ n --- Iso.inv (es-suc n) = {!!} --- Iso.rightInv (es-suc n) = {!!} --- Iso.leftInv (es-suc n) = {!!} - --- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) --- es zero = {!!} --- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) - --- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) --- fst yieldsCWskel∙→' = cwsk .fst --- fst (snd yieldsCWskel∙→') = αs --- fst (snd (snd yieldsCWskel∙→')) () --- snd (snd (snd yieldsCWskel∙→')) = {!!} - --- yieldsCWskel∙→ : (A : ℕ → Type ℓ) --- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) --- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst --- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) --- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () --- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} --- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = --- compEquiv (cwsk .snd .snd .snd n) --- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) --- where --- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst --- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) --- theEq→ (inr x) = inl x --- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) --- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) --- theEq→ (push (push a i₁) i) = {!!} - --- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) --- Iso.fun theEq = theEq→ --- Iso.inv theEq = {!!} --- Iso.rightInv theEq x = {!!} --- Iso.leftInv theEq x = {!!} - - --- -- compEquiv {!!} {!!} - - --- -- -- Finally: definition of CW complexes --- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' - --- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ - --- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) - --- -- -- Finite CW complexes --- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- isFinCW {ℓ = ℓ} X = --- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) - --- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ - --- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) - --- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X --- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str - --- -- finCW→CW : finCW ℓ → CW ℓ --- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p - - --- -- -- morphisms --- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') --- -- C →ᶜʷ D = fst C → fst D - --- -- --the cofibre of the inclusion of X n into X (suc n) --- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ --- -- cofibCW n C = cofib (CW↪ C n) - --- -- --...is basically a sphere bouquet --- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) --- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) --- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e --- -- where --- -- s = Bouquet≃-gen --- -- α = C .snd .snd .fst --- -- e = C .snd .snd .snd .snd n - --- -- --sending X (suc n) into the cofibCW --- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C --- -- to_cofibCW n C x = inr x - --- -- --the connecting map of the long exact sequence --- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) --- -- → cofib conn → Susp A --- -- δ-pre conn (inl x) = north --- -- δ-pre conn (inr x) = south --- -- δ-pre conn (push a i) = merid a i - --- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) --- -- δ n C = δ-pre (CW↪ C n) - --- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) --- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C --- -- CW↪∞ C n x = incl x - --- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m --- -- fst (finCW↑ m n p C) = fst C --- -- fst (snd (finCW↑ m n p C)) = snd C .fst --- -- snd (snd (finCW↑ m n p C)) k = --- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) --- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) --- -- (snd C .snd (k +ℕ fst p)) +-- -- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ + +-- -- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) +-- -- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) + +-- -- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X +-- -- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str + +-- -- -- finCW→CW : finCW ℓ → CW ℓ +-- -- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p + + +-- -- -- -- morphisms +-- -- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') +-- -- -- C →ᶜʷ D = fst C → fst D + +-- -- -- --the cofibre of the inclusion of X n into X (suc n) +-- -- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ +-- -- -- cofibCW n C = cofib (CW↪ C n) + +-- -- -- --...is basically a sphere bouquet +-- -- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) +-- -- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) +-- -- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e +-- -- -- where +-- -- -- s = Bouquet≃-gen +-- -- -- α = C .snd .snd .fst +-- -- -- e = C .snd .snd .snd .snd n + +-- -- -- --sending X (suc n) into the cofibCW +-- -- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C +-- -- -- to_cofibCW n C x = inr x + +-- -- -- --the connecting map of the long exact sequence +-- -- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) +-- -- -- → cofib conn → Susp A +-- -- -- δ-pre conn (inl x) = north +-- -- -- δ-pre conn (inr x) = south +-- -- -- δ-pre conn (push a i) = merid a i + +-- -- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) +-- -- -- δ n C = δ-pre (CW↪ C n) + +-- -- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) +-- -- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C +-- -- -- CW↪∞ C n x = incl x + +-- -- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m +-- -- -- fst (finCW↑ m n p C) = fst C +-- -- -- fst (snd (finCW↑ m n p C)) = snd C .fst +-- -- -- snd (snd (finCW↑ m n p C)) k = +-- -- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) +-- -- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) +-- -- -- (snd C .snd (k +ℕ fst p)) diff --git a/Cubical/CW/Subcomplex.agda b/Cubical/CW/Subcomplex.agda index 250298b1f2..da297d1da2 100644 --- a/Cubical/CW/Subcomplex.agda +++ b/Cubical/CW/Subcomplex.agda @@ -14,6 +14,7 @@ open import Cubical.Data.Fin.Inductive.Base open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat.Order.Inductive open import Cubical.CW.Base open import Cubical.CW.Properties @@ -132,6 +133,39 @@ realiseSubComplex n C = compIso (equivToIso (complex≃subcomplex C n flast)) (realiseFin n (finSubComplex C n)) + +niceFinCWskel : ∀ {ℓ} (n : ℕ) → finCWskel ℓ n → finCWskel ℓ n +fst (niceFinCWskel n (A , AC , fin)) = finSubComplex (A , AC) n .fst +snd (niceFinCWskel n (A , AC , fin)) = finSubComplex (A , AC) n .snd + +open import Cubical.Foundations.HLevels +makeNiceFinCW : ∀ {ℓ} → finCW ℓ → finCW ℓ +fst (makeNiceFinCW C) = fst C +snd (makeNiceFinCW C) = + PT.map better (snd C) + where + module _ (dim+CW : isFinCW (fst C)) where + dim = fst dim+CW + cwsk = fst (snd dim+CW) + e = snd (snd dim+CW) + improved = finSubComplex (cwsk .fst , cwsk .snd .fst) dim + + better : isFinCW (fst C) + fst better = dim + fst (snd better) = improved + snd (snd better) = + compEquiv + (compEquiv e (invEquiv (isoToEquiv (realiseFin dim cwsk)))) + (isoToEquiv (realiseSubComplex dim (cwsk .fst , cwsk .snd .fst))) + +makeNiceFinCW≡ : ∀ {ℓ} (C : finCW ℓ) → makeNiceFinCW C ≡ C +makeNiceFinCW≡ C = Σ≡Prop (λ _ → squash₁) refl + +makeNiceFinCWElim : ∀ {ℓ ℓ'} {A : finCW ℓ → Type ℓ'} + → ((C : finCW ℓ) → A (makeNiceFinCW C)) + → (C : _) → A C +makeNiceFinCWElim {A = A} ind C = subst A (makeNiceFinCW≡ C) (ind C) + -- Goal: Show that a cell complex C and its subcomplex Cₙ share -- homology in low enough dimensions module _ (C : CWskel ℓ) where diff --git a/Cubical/Data/Sequence/Base.agda b/Cubical/Data/Sequence/Base.agda index 6d2521edde..cab82d8076 100644 --- a/Cubical/Data/Sequence/Base.agda +++ b/Cubical/Data/Sequence/Base.agda @@ -3,6 +3,7 @@ module Cubical.Data.Sequence.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism open import Cubical.Data.Nat open import Cubical.Data.Fin @@ -28,3 +29,28 @@ converges seq n = (k : ℕ) → isEquiv (Sequence.map seq {n = k + n}) finiteSequence : (ℓ : Level) (m : ℕ) → Type (ℓ-suc ℓ) finiteSequence ℓ m = Σ[ S ∈ Sequence ℓ ] converges S m + +shiftSequence : ∀ {ℓ} → Sequence ℓ → (n : ℕ) → Sequence ℓ +Sequence.obj (shiftSequence seq n) m = Sequence.obj seq (m + n) +Sequence.map (shiftSequence seq n) {n = k} x = Sequence.map seq x + +-- Isomorphism of sequences +SequenceIso : (A : Sequence ℓ) (B : Sequence ℓ') → Type (ℓ-max ℓ ℓ') +SequenceIso A B = + Σ[ is ∈ ((n : ℕ) → Iso (Sequence.obj A n) (Sequence.obj B n)) ] + ((n : ℕ) (a : Sequence.obj A n) + → Sequence.map B (Iso.fun (is n) a) ≡ Iso.fun (is (suc n)) + (Sequence.map A a)) + +-- Equivalences of sequences +SequenceEquiv : (A : Sequence ℓ) (B : Sequence ℓ') → Type (ℓ-max ℓ ℓ') +SequenceEquiv A B = + Σ[ e ∈ (SequenceMap A B) ] + ((n : ℕ) → isEquiv (SequenceMap.map e n)) + +-- Iso to equiv +SequenceIso→SequenceEquiv : {A : Sequence ℓ} {B : Sequence ℓ'} + → SequenceIso A B → SequenceEquiv A B +SequenceMap.map (fst (SequenceIso→SequenceEquiv (is , e))) x = Iso.fun (is x) +SequenceMap.comm (fst (SequenceIso→SequenceEquiv (is , e))) = e +snd (SequenceIso→SequenceEquiv (is , e)) n = isoToIsEquiv (is n) diff --git a/Cubical/Data/Sequence/Properties.agda b/Cubical/Data/Sequence/Properties.agda index b01b6d609b..75400dd194 100644 --- a/Cubical/Data/Sequence/Properties.agda +++ b/Cubical/Data/Sequence/Properties.agda @@ -2,21 +2,192 @@ module Cubical.Data.Sequence.Properties where open import Cubical.Foundations.Prelude -open import Cubical.Data.Nat open import Cubical.Foundations.Equiv -open import Cubical.Data.Sequence.Base +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Sequence.Base private variable ℓ ℓ' ℓ'' : Level + +open Sequence +open SequenceMap renaming (map to smap) +open isHAEquiv + + +-- Identity map of sequences idSequenceMap : (C : Sequence ℓ) → SequenceMap C C -SequenceMap.map (idSequenceMap C) n x = x -SequenceMap.comm (idSequenceMap C) n x = refl +smap (idSequenceMap C) n x = x +comm (idSequenceMap C) n x = refl +-- Composition of maps composeSequenceMap : {C : Sequence ℓ} {D : Sequence ℓ'} {E : Sequence ℓ''} (g : SequenceMap D E) (f : SequenceMap C D) → SequenceMap C E -composeSequenceMap g f .SequenceMap.map n x = g .SequenceMap.map n (f .SequenceMap.map n x) -composeSequenceMap g f .SequenceMap.comm n x = - g .SequenceMap.comm n _ - ∙ cong (g .SequenceMap.map (suc n)) (f .SequenceMap.comm n x) +composeSequenceMap g f .smap n x = g .smap n (f .smap n x) +composeSequenceMap g f .comm n x = + g .comm n _ + ∙ cong (g .smap (suc n)) (f .comm n x) + +-- Composition of isos +compSequenceIso : {A : Sequence ℓ} {B : Sequence ℓ'} {C : Sequence ℓ''} + → SequenceIso A B → SequenceIso B C → SequenceIso A C +fst (compSequenceIso e g) n = compIso (fst e n) (fst g n) +snd (compSequenceIso e g) n a = + snd g n _ ∙ cong (Iso.fun (fst g (suc n))) (snd e n a) + +-- Identity equivalence +idSequenceEquiv : (A : Sequence ℓ) → SequenceEquiv A A +fst (idSequenceEquiv A) = idSequenceMap A +snd (idSequenceEquiv A) _ = idIsEquiv _ + +-- inversion of equivalences +invSequenceEquiv : {A : Sequence ℓ} {B : Sequence ℓ'} + → SequenceEquiv A B → SequenceEquiv B A +smap (fst (invSequenceEquiv eqs)) n = invEq (_ , snd eqs n) +comm (fst (invSequenceEquiv {A = A} {B} eqs)) n x = + sym (retEq (_ , snd eqs (suc n)) (map A (invEq (_ , snd eqs n) x))) + ∙∙ cong (invEq (_ , snd eqs (suc n))) (sym (comm (fst eqs) n (invEq (_ , snd eqs n) x))) + ∙∙ cong (invEq (_ , snd eqs (suc n)) ∘ map B) (secEq (_ , snd eqs n) x) +snd (invSequenceEquiv eqs) n = invEquiv (_ , snd eqs n) .snd + +-- inversion of the identity +invIdSequenceEquiv : {A : Sequence ℓ} + → invSequenceEquiv (idSequenceEquiv A) ≡ idSequenceEquiv A +invIdSequenceEquiv {A = A} = + Σ≡Prop (λ _ → isPropΠ (λ _ → isPropIsEquiv _)) main + where + main : invSequenceEquiv (idSequenceEquiv A) .fst ≡ idSequenceMap A + smap (main i) n z = z + comm (main i) n x j = rUnit (λ _ → map A x) (~ i) j + +-- ua for sequences +uaSequence : {A B : Sequence ℓ} → SequenceEquiv A B → A ≡ B +obj (uaSequence {A = A} {B} (e , eqv) i) n = ua (_ , eqv n) i +map (uaSequence {A = A} {B} (e , eqv) i) {n = n} a = + hcomp (λ k → λ {(i = i0) → map A (retEq (_ , eqv n) a k) + ; (i = i1) → (sym (comm e n (invEq (_ , eqv n) a)) + ∙ cong (map B) (secEq (_ , eqv n) a)) k}) + (ua-gluePt (_ , eqv (suc n)) i + (map A (invEq (_ , eqv n) (ua-unglue (_ , eqv n) i a)))) + +uaSequenceIdEquiv : ∀ {ℓ} {A : Sequence ℓ} + → uaSequence (idSequenceEquiv A) ≡ refl +obj (uaSequenceIdEquiv {A = A} i j) n = + uaIdEquiv {A = obj A n} i j +map (uaSequenceIdEquiv {A = A} i j) {n = n} x = + hcomp (λ k → λ {(i = i1) → map A x + ; (j = i0) → map A x + ; (j = i1) → rUnit (refl {x = map A x}) (~ i) k}) + (glue {φ = j ∨ ~ j ∨ i} + (λ { (j = i0) → map A x + ; (j = i1) → map A x + ; (i = i1) → map A x}) + (map A (unglue (j ∨ ~ j ∨ i) x))) + +halfAdjointSequenceEquiv : {A : Sequence ℓ} {B : Sequence ℓ'} + → SequenceEquiv A B → SequenceEquiv A B +fst (halfAdjointSequenceEquiv (e , eqv)) = e +snd (halfAdjointSequenceEquiv (e , eqv)) n = + isHAEquiv→isEquiv (equiv→HAEquiv (_ , eqv n) .snd) + +-- Contractibility of total space of (SequenceEquiv A) +isContrTotalSequence : (A : Sequence ℓ) + → isContr (Σ[ B ∈ Sequence ℓ ] SequenceEquiv A B) +fst (isContrTotalSequence A) = A , idSequenceEquiv A +snd (isContrTotalSequence A) (B , eqv*) = + ΣPathP (uaSequence eqv , main) + where + eqv = halfAdjointSequenceEquiv eqv* + + eqv⁻ : {n : ℕ} → obj B n → obj A n + eqv⁻ {n = n} = invEq (_ , snd eqv n) + + haEqv : (n : ℕ) → HAEquiv (obj A n) (obj B n) + haEqv n = equiv→HAEquiv (_ , eqv* .snd n) + + + help : (n : ℕ) (x : obj A n) + → Square (cong (smap (fst eqv) (suc n) ∘ map A) + (retEq (_ , snd eqv n) x)) + (comm (fst eqv) n x) + (sym (comm (fst eqv) n + (eqv⁻ (smap (fst eqv) n x))) + ∙ cong (map B) (secEq (smap (fst eqv) n , snd eqv n) + (smap (fst eqv) n x))) + refl + help n x = flipSquare ((compPath≡compPath' _ _ + ∙ cong₂ _∙'_ refl λ j i + → map B (com + (equiv→HAEquiv (_ , eqv* .snd n) .snd) x (~ j) i)) + ◁ λ i j → + hcomp (λ k → + λ {(j = i0) → comm (fst eqv) n + (lUnit (cong fst + (eqv* .snd n .equiv-proof (smap (fst eqv*) n x) + .snd (x , refl))) k i) k + ; (j = i1) → comm (fst eqv) n x (i ∧ k) + ; (i = i0) → compPath'-filler + (sym (comm (fst eqv) n (eqv⁻ (smap (fst eqv) n x)))) + (cong (map B) (com (haEqv n .snd) x i0)) k j + ; (i = i1) → comm (fst eqv) n x k + }) + (map B (smap (fst eqv*) n + ((cong fst (eqv* .snd n .equiv-proof (smap (fst eqv*) n x) + .snd (x , (λ _ → smap (fst eqv*) n x)))) (i ∨ j))))) + + main : PathP (λ i → SequenceEquiv A (uaSequence eqv i)) + (idSequenceEquiv A) eqv* + smap (fst (main i)) n x = ua-gluePt (_ , snd eqv n) i x + comm (fst (main i)) n x j = + hcomp (λ k → λ {(i = i0) → map A + (retEq (_ , snd eqv n) x (k ∨ j)) + ; (i = i1) → help n x k j + ; (j = i1) → ua-gluePt (_ , snd eqv (suc n)) i (map A x)}) + (ua-gluePt (_ , snd eqv (suc n)) i + (map A (retEq (_ , snd eqv n) x j))) + snd (main i) n = + isProp→PathP {B = λ i → isEquiv (λ x → ua-gluePt (_ , snd eqv n) i x)} + (λ i → isPropIsEquiv _) + (idIsEquiv (obj A n)) (snd eqv* n) i + + +-- J for sequences +SequenceEquivJ : {A : Sequence ℓ} + (P : (B : Sequence ℓ) → SequenceEquiv A B → Type ℓ') + → P A (idSequenceEquiv A) + → (B : _) (e : _) → P B e +SequenceEquivJ {ℓ = ℓ} {A = A} P ind B e = main (B , e) + where + P' : Σ[ B ∈ Sequence ℓ ] SequenceEquiv A B → Type _ + P' (B , e) = P B e + + main : (x : _) → P' x + main (B , e) = subst P' (isContrTotalSequence A .snd (B , e)) ind + +SequenceEquivJ>_ : {A : Sequence ℓ} + {P : (B : Sequence ℓ) → SequenceEquiv A B → Type ℓ'} + → P A (idSequenceEquiv A) + → (B : _) (e : _) → P B e +SequenceEquivJ>_ {P = P} hyp B e = SequenceEquivJ P hyp B e + +SequenceEquivJRefl : ∀ {ℓ ℓ'} {A : Sequence ℓ} + (P : (B : Sequence ℓ) → SequenceEquiv A B → Type ℓ') + (id : P A (idSequenceEquiv A)) + → SequenceEquivJ P id A (idSequenceEquiv A) ≡ id +SequenceEquivJRefl {ℓ = ℓ} {A = A} P ids = + (λ j → subst P' (isProp→isSet (isContr→isProp (isContrTotalSequence A)) + _ _ (isContrTotalSequence A .snd + (A , idSequenceEquiv A)) refl j) ids) + ∙ transportRefl ids + where + P' : Σ[ B ∈ Sequence ℓ ] SequenceEquiv A B → Type _ + P' (B , e) = P B e diff --git a/Cubical/HITs/SequentialColimit/Properties.agda b/Cubical/HITs/SequentialColimit/Properties.agda index 6f2fa6a5a5..f1d233c139 100644 --- a/Cubical/HITs/SequentialColimit/Properties.agda +++ b/Cubical/HITs/SequentialColimit/Properties.agda @@ -17,10 +17,12 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.Univalence open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Transport open import Cubical.Data.Nat hiding (elim) open import Cubical.Data.Sequence open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Sigma open import Cubical.HITs.SequentialColimit.Base open import Cubical.Homotopy.Connected @@ -247,6 +249,7 @@ module _ open Sequence open ElimDataShifted + -- Proof that colim Aₘ ≃ Aₙ for a converging sequence -- A₀ → A₁ → ... → Aₙ ≃ Aₙ₊₁ ≃ ... module _ {ℓ : Level} (seq : Sequence ℓ) (n : ℕ) (term : converges seq n) @@ -410,7 +413,8 @@ realiseCompSequenceMap {C = C} {E = E} g f = (push {n = n} (g₁ (f₁ x)) ∙ cong incl (SequenceMap.comm g n (f₁ x) ∙ cong g₊ (SequenceMap.comm f n x))) - (cong (realiseSequenceMap g) (push (f₁ x) ∙ cong incl (SequenceMap.comm f n x))) + (cong (realiseSequenceMap g) + (push (f₁ x) ∙ cong incl (SequenceMap.comm f n x))) main = cong (push (SequenceMap.map g n (f₁ x)) ∙_) (cong-∙ incl (SequenceMap.comm g n (f₁ x)) (cong g₊ (SequenceMap.comm f n x))) @@ -420,6 +424,34 @@ realiseCompSequenceMap {C = C} {E = E} g f = ∙ sym (cong-∙ (realiseSequenceMap g) (push (f₁ x)) (cong incl (SequenceMap.comm f n x))) +sequenceEquiv→ColimIso : {A B : Sequence ℓ} + → SequenceEquiv A B → Iso (SeqColim A) (SeqColim B) +sequenceEquiv→ColimIso e = mainIso + where + main : {A : Sequence ℓ} (B : Sequence ℓ) (e : SequenceEquiv A B) + → section (realiseSequenceMap (fst e)) + (realiseSequenceMap (fst (invSequenceEquiv e))) + × retract (realiseSequenceMap (fst e)) + (realiseSequenceMap (fst (invSequenceEquiv e))) + main {A = A} = SequenceEquivJ> + ((λ x → (λ i → realiseIdSequenceMap {C = A} i + (realiseSequenceMap (fst (invIdSequenceEquiv {A = A} i)) x)) + ∙ funExt⁻ realiseIdSequenceMap x) + , λ x → (λ i → realiseSequenceMap (fst (invIdSequenceEquiv {A = A} i)) + (realiseIdSequenceMap {C = A} i x)) + ∙ funExt⁻ realiseIdSequenceMap x) + + mainIso : Iso _ _ + Iso.fun mainIso = realiseSequenceMap (fst e) + Iso.inv mainIso = realiseSequenceMap (fst (invSequenceEquiv e)) + Iso.rightInv mainIso = main _ e .fst + Iso.leftInv mainIso = main _ e .snd + +sequenceIso→ColimIso : {A B : Sequence ℓ} + → SequenceIso A B → Iso (SeqColim A) (SeqColim B) +sequenceIso→ColimIso e = + sequenceEquiv→ColimIso (SequenceIso→SequenceEquiv e) + converges→funId : {seq1 : Sequence ℓ} {seq2 : Sequence ℓ'} (n m : ℕ) → converges seq1 n @@ -512,3 +544,74 @@ Iso.leftInv (Iso-FinSeqColim-Top X m) r = (transport (λ i → (x : isoToPath (invIso (Iso-FinSeqColim-Top X m)) i) → f (ua-unglue (isoToEquiv (invIso (Iso-FinSeqColim-Top X m))) i x) ≡ g (ua-unglue (isoToEquiv (invIso (Iso-FinSeqColim-Top X m))) i x)) h) + + +-- Shifinting colimi + +Iso-SeqColim→SeqColimSuc : (X : Sequence ℓ) + → Iso (SeqColim X) (SeqColim (ShiftSeq X)) +Iso-SeqColim→SeqColimSuc X = iso G F F→G→F G→F→G + where + F : SeqColim (ShiftSeq X) → SeqColim X + F (incl {n = n} x) = incl {n = suc n} x + F (push {n = n} x i) = push {n = suc n} x i + + G : SeqColim X → SeqColim (ShiftSeq X) + G (incl {n = zero} x) = incl {n = zero} (map X x) + G (incl {n = suc n} x) = incl {n = n} x + G (push {n = zero} x i) = incl {n = zero} (map X x) + G (push {n = suc n} x i) = push {n = n} x i + + F→G→F : (x : SeqColim (ShiftSeq X)) → G (F x) ≡ x + F→G→F (incl x) = refl + F→G→F (push x i) = refl + + G→F→G : (x : SeqColim X) → F (G x) ≡ x + G→F→G (incl {n = zero} x) = sym (push {n = zero} x) + G→F→G (incl {n = suc n} x) = refl + G→F→G (push {n = zero} x i) j = push {n = zero} x (i ∨ ~ j) + G→F→G (push {n = suc n} x i) = refl + +ShiftSequence+ : (S : Sequence ℓ) (n : ℕ) → Sequence ℓ +Sequence.obj (ShiftSequence+ S n) m = Sequence.obj S (m + n) +Sequence.map (ShiftSequence+ S n) {n = m} = Sequence.map S + +ShiftSequence+Rec : (S : Sequence ℓ) (n : ℕ) → Sequence ℓ +ShiftSequence+Rec S zero = S +ShiftSequence+Rec S (suc n) = ShiftSeq (ShiftSequence+Rec S n) + +Iso-SeqColim→SeqColimShift : (S : Sequence ℓ) (n : ℕ) + → Iso (SeqColim S) (SeqColim (ShiftSequence+Rec S n)) +Iso-SeqColim→SeqColimShift S zero = idIso +Iso-SeqColim→SeqColimShift S (suc n) = + compIso (Iso-SeqColim→SeqColimShift S n) + (Iso-SeqColim→SeqColimSuc _) + +ShiftSequenceIso : {A : Sequence ℓ} (n : ℕ) + → SequenceIso (ShiftSequence+Rec A n) (ShiftSequence+ A n) +fst (ShiftSequenceIso {A = A} zero) m = + pathToIso λ i → Sequence.obj A (+-comm zero m i) +fst (ShiftSequenceIso {A = A} (suc n)) m = + compIso (fst (ShiftSequenceIso {A = A} n) (suc m)) + (pathToIso λ i → Sequence.obj A (+-suc m n (~ i))) +snd (ShiftSequenceIso {A = A} zero) m a = + sym (substCommSlice (Sequence.obj A) (Sequence.obj A ∘ suc) + (λ _ → Sequence.map A) + (+-comm zero m) a) + ∙ λ t → subst (Sequence.obj A) + (lUnit (cong suc (+-comm zero m)) t) + (Sequence.map A a) +snd (ShiftSequenceIso {A = A} (suc n)) m a = + sym (substCommSlice (Sequence.obj A) (Sequence.obj A ∘ suc) + (λ _ → Sequence.map A) + (λ i → (+-suc m n (~ i))) + (Iso.fun (fst (ShiftSequenceIso n) (suc m)) a)) + ∙ cong (subst (λ x → Sequence.obj A (suc x)) (sym (+-suc m n))) + (snd (ShiftSequenceIso {A = A} n) (suc m) a) + +SeqColimIso : (S : Sequence ℓ) (n : ℕ) + → Iso (SeqColim S) (SeqColim (ShiftSequence+ S n)) +SeqColimIso S n = + compIso (Iso-SeqColim→SeqColimShift S n) + (sequenceEquiv→ColimIso + (SequenceIso→SequenceEquiv (ShiftSequenceIso n))) From 77623bb0a40f5be302202c60e9f3ae5791082c90 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 27 May 2024 03:16:45 +0200 Subject: [PATCH 43/73] wip... --- Cubical/CW/Base.agda | 4 + Cubical/CW/Pointed2.agda | 1119 +++++++++++++++++++------- Cubical/CW/Subcomplex.agda | 43 +- Cubical/Data/List/Base.agda | 4 + Cubical/HITs/Pushout/Properties.agda | 162 ++-- 5 files changed, 953 insertions(+), 379 deletions(-) diff --git a/Cubical/CW/Base.agda b/Cubical/CW/Base.agda index 8f94d5531e..fd10c9309c 100644 --- a/Cubical/CW/Base.agda +++ b/Cubical/CW/Base.agda @@ -7,6 +7,7 @@ module Cubical.CW.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function +open import Cubical.Foundations.Pointed open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order @@ -102,6 +103,9 @@ isFinCW {ℓ = ℓ} X = finCW : (ℓ : Level) → Type (ℓ-suc ℓ) finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ +finCW∙ : (ℓ : Level) → Type (ℓ-suc ℓ) +finCW∙ ℓ = Σ[ A ∈ Pointed ℓ ] ∥ isFinCW (fst A) ∥₁ + finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) diff --git a/Cubical/CW/Pointed2.agda b/Cubical/CW/Pointed2.agda index eb82ebabb1..42b120983d 100644 --- a/Cubical/CW/Pointed2.agda +++ b/Cubical/CW/Pointed2.agda @@ -44,6 +44,62 @@ open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence +liftFun : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} + (f : A → B) → Lift {j = ℓ''} A → Lift {j = ℓ'''} B +liftFun f (lift a) = lift (f a) + +module _ {ℓ ℓ' ℓ'' : Level} (ℓ''' : Level) + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (f : A → B) (g : A → C) where + PushoutLevel : Level + PushoutLevel = ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')) + + PushoutLift : Type PushoutLevel + PushoutLift = Pushout {A = Lift {j = ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')} A} + {B = Lift {j = ℓ-max ℓ (ℓ-max ℓ'' ℓ''')} B} + {C = Lift {j = ℓ-max ℓ (ℓ-max ℓ' ℓ''')} C} + (liftFun f) + (liftFun g) + + PushoutLiftIso : Iso (Pushout f g) PushoutLift + Iso.fun PushoutLiftIso (inl x) = inl (lift x) + Iso.fun PushoutLiftIso (inr x) = inr (lift x) + Iso.fun PushoutLiftIso (push a i) = push (lift a) i + Iso.inv PushoutLiftIso (inl (lift x)) = inl x + Iso.inv PushoutLiftIso (inr (lift x)) = inr x + Iso.inv PushoutLiftIso (push (lift a) i) = push a i + Iso.rightInv PushoutLiftIso (inl (lift x)) = refl + Iso.rightInv PushoutLiftIso (inr (lift x)) = refl + Iso.rightInv PushoutLiftIso (push (lift a) i) = refl + Iso.leftInv PushoutLiftIso (inl x) = refl + Iso.leftInv PushoutLiftIso (inr x) = refl + Iso.leftInv PushoutLiftIso (push a i) = refl + +open import Cubical.Foundations.Equiv + + +pushoutIso' : ∀ {ℓA₁ ℓB₁ ℓC₁ ℓA₂ ℓB₂ ℓC₂} + {A₁ : Type ℓA₁} {B₁ : Type ℓB₁} {C₁ : Type ℓC₁} + {A₂ : Type ℓA₂} {B₂ : Type ℓB₂} {C₂ : Type ℓC₂} + (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) + (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) + (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) + (id1 : fst B≃ ∘ f₁ ≡ f₂ ∘ fst A≃) + (id2 : fst C≃ ∘ g₁ ≡ g₂ ∘ fst A≃) → + Iso (Pushout f₁ g₁) (Pushout f₂ g₂) +pushoutIso' {ℓA₁ = ℓA₁} {ℓB₁} {ℓC₁} {ℓA₂} {ℓB₂} {ℓC₂} + f₁ g₁ f₂ g₂ A≃ B≃ C≃ id1 id2 = + compIso (PushoutLiftIso ℓ* f₁ g₁) + (compIso (pushoutIso _ _ _ _ + (Lift≃Lift A≃) + (Lift≃Lift B≃) + (Lift≃Lift C≃) + (funExt (λ { (lift x) → cong lift (funExt⁻ id1 x)})) + (funExt (λ { (lift x) → cong lift (funExt⁻ id2 x)}))) + (invIso (PushoutLiftIso ℓ* f₂ g₂))) + where + ℓ* = ℓ-max ℓA₁ (ℓ-max ℓA₂ (ℓ-max ℓB₁ (ℓ-max ℓB₂ (ℓ-max ℓC₁ ℓC₂)))) + + module _ {C : Type ℓ} {B : Type ℓ'} where PushoutAlongEquiv→ : {A : Type ℓ} (e : A ≃ C) (f : A → B) → Pushout (fst e) f → B @@ -74,8 +130,8 @@ module _ {C : Type ℓ} {B : Type ℓ'} where -- New def: X 0 is empty and C (suc n) is pushout open import Cubical.HITs.Pushout -module _ {A : Type} {B : A → Pointed₀} (C : Pointed₀) (f : (⋁gen A B , inl tt) →∙ C) where - +module _ {ℓA ℓB ℓC : Level} {A : Type ℓA} {B : A → Pointed ℓB} (C : Pointed ℓC) + (f : (⋁gen A B , inl tt) →∙ C) where private open 3x3-span @@ -118,7 +174,7 @@ module _ {A : Type} {B : A → Pointed₀} (C : Pointed₀) (f : (⋁gen A B , i A○□Iso : Iso (A○□ inst) (Pushout (fst f ∘ inr) fst) A○□Iso = compIso (equivToIso (symPushout _ _)) - (invIso (pushoutIso _ _ _ _ + (invIso (pushoutIso' _ _ _ _ (isoToEquiv (invIso A2□Iso)) (isoToEquiv (invIso A4□Iso)) (isoToEquiv (invIso A0□Iso)) @@ -143,7 +199,7 @@ module _ {A : Type} {B : A → Pointed₀} (C : Pointed₀) (f : (⋁gen A B , i open import Cubical.Foundations.GroupoidLaws A□○Iso : Iso (A□○ inst) (cofib (fst f)) - A□○Iso = invIso (pushoutIso _ _ _ _ + A□○Iso = invIso (pushoutIso' _ _ _ _ (isoToEquiv (invIso A□2Iso)) (isoToEquiv (invIso A□0Iso)) (isoToEquiv (invIso A□4Iso)) @@ -468,60 +524,8 @@ module _ {n : ℕ} where Iso.rightInv (swapFinIso x y) = swapFin² x y Iso.leftInv (swapFinIso x y) = swapFin² x y -liftFun : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} - (f : A → B) → Lift {j = ℓ''} A → Lift {j = ℓ'''} B -liftFun f (lift a) = lift (f a) -module _ {ℓ ℓ' ℓ'' : Level} (ℓ''' : Level) - {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (f : A → B) (g : A → C) where - PushoutLevel : Level - PushoutLevel = ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')) - PushoutLift : Type PushoutLevel - PushoutLift = Pushout {A = Lift {j = ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')} A} - {B = Lift {j = ℓ-max ℓ (ℓ-max ℓ'' ℓ''')} B} - {C = Lift {j = ℓ-max ℓ (ℓ-max ℓ' ℓ''')} C} - (liftFun f) - (liftFun g) - - PushoutLiftIso : Iso (Pushout f g) PushoutLift - Iso.fun PushoutLiftIso (inl x) = inl (lift x) - Iso.fun PushoutLiftIso (inr x) = inr (lift x) - Iso.fun PushoutLiftIso (push a i) = push (lift a) i - Iso.inv PushoutLiftIso (inl (lift x)) = inl x - Iso.inv PushoutLiftIso (inr (lift x)) = inr x - Iso.inv PushoutLiftIso (push (lift a) i) = push a i - Iso.rightInv PushoutLiftIso (inl (lift x)) = refl - Iso.rightInv PushoutLiftIso (inr (lift x)) = refl - Iso.rightInv PushoutLiftIso (push (lift a) i) = refl - Iso.leftInv PushoutLiftIso (inl x) = refl - Iso.leftInv PushoutLiftIso (inr x) = refl - Iso.leftInv PushoutLiftIso (push a i) = refl - -open import Cubical.Foundations.Equiv - - -pushoutIso' : ∀ {ℓA₁ ℓB₁ ℓC₁ ℓA₂ ℓB₂ ℓC₂} - {A₁ : Type ℓA₁} {B₁ : Type ℓB₁} {C₁ : Type ℓC₁} - {A₂ : Type ℓA₂} {B₂ : Type ℓB₂} {C₂ : Type ℓC₂} - (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) - (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) - (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) - (id1 : fst B≃ ∘ f₁ ≡ f₂ ∘ fst A≃) - (id2 : fst C≃ ∘ g₁ ≡ g₂ ∘ fst A≃) → - Iso (Pushout f₁ g₁) (Pushout f₂ g₂) -pushoutIso' {ℓA₁ = ℓA₁} {ℓB₁} {ℓC₁} {ℓA₂} {ℓB₂} {ℓC₂} - f₁ g₁ f₂ g₂ A≃ B≃ C≃ id1 id2 = - compIso (PushoutLiftIso ℓ* f₁ g₁) - (compIso (pushoutIso _ _ _ _ - (Lift≃Lift A≃) - (Lift≃Lift B≃) - (Lift≃Lift C≃) - (funExt (λ { (lift x) → cong lift (funExt⁻ id1 x)})) - (funExt (λ { (lift x) → cong lift (funExt⁻ id2 x)}))) - (invIso (PushoutLiftIso ℓ* f₂ g₂))) - where - ℓ* = ℓ-max ℓA₁ (ℓ-max ℓA₂ (ℓ-max ℓB₁ (ℓ-max ℓB₂ (ℓ-max ℓC₁ ℓC₂)))) PushoutPostcompEquivIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {B' : Type ℓ''} {C : Type ℓ'''} @@ -1025,9 +1029,28 @@ module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A 0) where (invEquiv (isoToEquiv (Iso-SeqColim→SeqColimShift _ 3)))) -connectedCWskel'→connectedCWskel : ∀ {ℓ} → connectedCWskel' ℓ 0 → connectedCWskel ℓ 0 -connectedCWskel'→connectedCWskel (A , sk) = +connectedCWskel'→connectedCWskel : ∀ {ℓ} + → Σ[ t ∈ connectedCWskel' ℓ 0 ] + (Σ[ dim ∈ ℕ ] + ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim)))) + → Σ[ t ∈ connectedCWskel ℓ 0 ] + Σ[ dim ∈ ℕ ] + ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim))) +fst (connectedCWskel'→connectedCWskel ((A , sk) , conv)) = _ , connectedCWskel→ A ((sk .fst) , (sk .snd)) .fst , refl , (λ _ → λ()) +fst (snd (connectedCWskel'→connectedCWskel ((A , sk) , conv))) = suc (fst conv) +snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , zero , T))) k = + ⊥.rec (TR.rec (λ()) + (λ a → TR.rec (λ()) + (λ t → CW₀-empty (_ , fst sk) (invEq (_ , T 0) (t .fst))) + (isConnected-CW↪∞ 1 (_ , fst sk) a .fst)) (sk .snd .fst)) +snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , (suc dim) , T))) k = + transport (λ i → isEquiv (CW↪ (collapse₁CWskel A sk , connectedCWskel→ A sk .fst) + (h i))) + (transport (λ i → isEquiv (CW↪ (A , sk .fst) (suc (+-suc k dim i)))) + (T (suc k))) + where + h = cong suc (sym (+-suc k dim)) ∙ sym (+-suc k (suc dim)) yieldsGoodCWskel : {ℓ : Level} (A₊₂ : ℕ → Pointed ℓ) → Type _ yieldsGoodCWskel {ℓ = ℓ} A = @@ -1044,292 +1067,800 @@ Sequence.map (GoodCWskelSeq {A = A} (f , fin , α , sq)) {n = suc n} x = fst (fs realiseGood∙ : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Pointed ℓ realiseGood∙ {A = A} S = SeqColim (GoodCWskelSeq S) , incl {n = 1} (snd (A 0)) -module _ {ℓ : Level} (A' : ℕ → Type ℓ) - (con : isConnected 0 (A' 2)) - (C : yieldsCWskel A') - (dim : ℕ) (ind : (n : ℕ) → dim <ᵗ n → fst C n ≡ 0) - (cA : fst C 0 ≡ 1) +yieldsFinGoodCWskel : {ℓ : Level} (dim : ℕ) (A₊₂ : ℕ → Pointed ℓ) → Type _ +yieldsFinGoodCWskel {ℓ = ℓ} dim A = + Σ[ A ∈ yieldsGoodCWskel A ] converges (GoodCWskelSeq A) dim + +GoodCWskel : (ℓ : Level) → Type (ℓ-suc ℓ) +GoodCWskel ℓ = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsGoodCWskel A + +FinGoodCWskel : (ℓ : Level) (dim : ℕ) → Type (ℓ-suc ℓ) +FinGoodCWskel ℓ dim = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsFinGoodCWskel dim A + +isGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) +isGoodCWExpl {ℓ} A = Σ[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A + +isFinGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) +isFinGoodCWExpl {ℓ} A = + Σ[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] realiseGood∙ (fst (snd sk)) ≃∙ A + +isGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) +isGoodCW {ℓ} A = ∃[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A + +isFinGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) +isFinGoodCW {ℓ} A = + ∃[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] (realiseGood∙ (fst (snd sk)) ≃∙ A) + +finGoodCW : (ℓ : Level) → Type (ℓ-suc ℓ) +finGoodCW ℓ = Σ[ A ∈ Pointed ℓ ] isFinGoodCW A + +ptCW : {ℓ : Level} {A : ℕ → Type ℓ} → yieldsCWskel A → A 1 + → (n : ℕ) → A (suc n) +ptCW sk a zero = a +ptCW sk a (suc n) = CW↪ (_ , sk) (suc n) (ptCW sk a n) + +-- module TWOO {ℓ : Level} (A' : ℕ → Type ℓ) (pt0 : A' 1) +-- (dim : ℕ) (con : isConnected 2 (A' 2)) +-- (C : yieldsFinCWskel dim A') +-- where + +-- open CWskel-fields (_ , fst C) +-- e₀ : A' 1 ≃ Fin (card 0) +-- e₀ = CW₁-discrete (_ , fst C) + +-- ptA : (n : ℕ) → A' (suc n) +-- ptA = ptCW (fst C) pt0 + +-- conA : (n : ℕ) → isConnected 2 (A' (suc n)) +-- conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) +-- (subst (isConnected 2) (sym (cong Fin cA)) +-- (∣ flast ∣ +-- , TR.elim (λ _ → isOfHLevelPath 2 +-- (isOfHLevelTrunc 2) _ _) +-- λ {(zero , tt) → refl})) +-- conA (suc n) = +-- isConnectedRetractFromIso 2 +-- (equivToIso (e (suc n))) +-- (∣ inl (ptA n) ∣ₕ +-- , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) +-- (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) +-- (conA' (conA n)) +-- λ c → conA' (conA n) _ +-- ∙ cong ∣_∣ₕ (push (c , ptSn n)))) +-- where +-- conA' : (conA : isConnected 2 (A' (suc n))) (c : A' (suc n)) +-- → Path (hLevelTrunc 2 (Pushout (α (suc n)) fst)) +-- ∣ inl (ptA n) ∣ₕ ∣ inl c ∣ₕ +-- conA' conA c = +-- TR.rec (isOfHLevelTrunc 2 _ _) +-- (λ p i → ∣ inl (p i) ∣) +-- (Iso.fun (PathIdTruncIso _) +-- (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) + + +-- private +-- funType = ((n : Fin (suc dim)) +-- → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) +-- →∙ (A' (suc (fst n)) , ptA (fst n))) ] +-- ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) + +-- mapMakerNil : ∥ funType ∥₁ +-- mapMakerNil = +-- invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) +-- λ n → TR.map +-- (λ pted → ((λ { (inl x) → ptA (fst n) +-- ; (inr x) → α _ x +-- ; (push a i) → pted a i}) +-- , refl) +-- , λ _ → refl) (help n)) +-- where +-- help : (n : Fin (suc dim)) +-- → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) +-- → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) +-- help n = invEq (_ , InductiveFinSatAC _ _ _) +-- λ m → Iso.fun (PathIdTruncIso _) +-- (isContr→isProp +-- (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ +-- ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) +-- module _ (F : funType) where +-- funs : (n : ℕ) → SphereBouquet∙ n (Fin (card (suc n))) +-- →∙ (A' (suc n) , ptA n) +-- funs n with (n ≟ᵗ dim) +-- ... | lt x = F (n , <ᵗ-trans-suc x) .fst +-- ... | eq x = F (n , subst (_<ᵗ suc dim) (sym x) <ᵗsucm) .fst +-- ... | gt x = const∙ _ _ + +-- funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) +-- ≃∙ Pushout (funs n .fst ∘ inr) (λ r → fst r) , inl (ptA n) +-- funEqP1 n = invEquiv (isoToEquiv (⋁-cofib-Iso _ (funs n))) , refl + +-- funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) +-- ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) +-- funEq n = isoToEquiv (pushoutIso' _ _ _ _ +-- (idEquiv _) +-- (idEquiv _) +-- (idEquiv _) +-- (funExt (uncurry (main n))) +-- (λ i x → fst x)) +-- , λ _ → inl (ptA n) +-- where +-- main : (n : ℕ) (x : Fin (card (suc n))) (y : S₊ n) +-- → funs n .fst (inr (x , y)) ≡ fst (C .snd) (suc n) (x , y) +-- main n with (n ≟ᵗ dim) +-- ... | lt p = λ x y +-- → F (n , <ᵗ-trans-suc p) .snd (x , y) +-- ... | eq p = λ x y +-- → F (n , subst (_<ᵗ suc dim) (λ i → p (~ i)) <ᵗsucm) .snd (x , y) +-- ... | gt p = λ x +-- → ⊥.rec (¬Fin0 (subst Fin (ind (suc n) (<ᵗ-trans p <ᵗsucm)) x)) + +-- getGoodCWskelAux : ∥ yieldsGoodCWskel (λ n → A' (suc n) , ptA n) ∥₁ +-- getGoodCWskelAux = PT.map help mapMakerNil +-- where +-- help : funType → yieldsGoodCWskel (λ n → A' (suc n) , ptA n) +-- fst (help F) = card ∘ suc +-- fst (snd (help F)) = compEquiv e₀ (pathToEquiv (cong Fin cA)) +-- fst (snd (snd (help F))) n = funs F n +-- snd (snd (snd (help F))) n = +-- compEquiv∙ (compEquiv∙ (funEqP1 F n) (funEq F n)) +-- (invEquiv (e (suc n)) , refl) + + +module BS {ℓ : Level} (A' : ℕ → Type ℓ) + (dim : ℕ) + (C+eq : yieldsFinCWskel dim A') + (cA : fst (fst C+eq) 0 ≡ 1) where + C = fst C+eq + ind = snd C+eq open CWskel-fields (_ , C) e₀ : A' 1 ≃ Fin (card 0) e₀ = CW₁-discrete (_ , C) + + ¬dim≡0 : ¬ (dim ≡ 0) + ¬dim≡0 p = CW₀-empty (_ , C) (subst A' p + (invEq (_ , ind 0) (subst A' (cong suc (sym p)) + (invEq e₀ (subst Fin (sym cA) fzero))))) + pt0 : A' 1 pt0 = invEq e₀ (subst Fin (sym cA) flast) ptA : (n : ℕ) → A' (suc n) - ptA zero = pt0 - ptA (suc n) = CW↪ (_ , C) (suc n) (ptA n) + ptA = ptCW C pt0 + + conA : (n : ℕ) → isConnected 2 (A' (suc n)) + conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) + (subst (isConnected 2) (sym (cong Fin cA)) + (∣ flast ∣ + , TR.elim (λ _ → isOfHLevelPath 2 + (isOfHLevelTrunc 2) _ _) + λ {(zero , tt) → refl})) + conA (suc n) = + isConnectedRetractFromIso 2 + (equivToIso (e (suc n))) + (∣ inl (ptA n) ∣ₕ + , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) + (conA' (conA n)) + λ c → conA' (conA n) _ + ∙ cong ∣_∣ₕ (push (c , ptSn n)))) + where + conA' : (conA : isConnected 2 (A' (suc n))) (c : A' (suc n)) + → Path (hLevelTrunc 2 (Pushout (α (suc n)) fst)) + ∣ inl (ptA n) ∣ₕ ∣ inl c ∣ₕ + conA' conA c = + TR.rec (isOfHLevelTrunc 2 _ _) + (λ p i → ∣ inl (p i) ∣) + (Iso.fun (PathIdTruncIso _) + (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) + + private + funType = ((n : Fin dim) + → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) + →∙ (A' (suc (fst n)) , ptA (fst n))) ] + ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) - mapMakerNil : ∥ ((n : Fin (suc dim)) - → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) - →∙ (A' (suc (fst n)) , ptA (fst n))) ] ((x : _) → fst h (inr x) ≡ α _ x)) ∥₁ + mapMakerNil : ∥ funType ∥₁ mapMakerNil = invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) - λ n → {!!}) + λ n → TR.map + (λ pted → ((λ { (inl x) → ptA (fst n) + ; (inr x) → α _ x + ; (push a i) → pted a i}) + , refl) + , λ _ → refl) (help n)) where - help : (n : Fin (suc dim)) → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) → ({!CW↪ ?!} ≡ {!!})) - help = {!!} -{- - mapMaker : (n : ℕ) → SphereBouquet∙ n (Fin (card (suc n))) - →∙ (A' (suc n) , ptA n) - mapMaker n with (n ≟ᵗ dim) - ... | lt x = mapMakerNil (n , <ᵗ-trans-suc x) - ... | eq x = mapMakerNil (n , subst (_<ᵗ suc dim) (sym x) <ᵗsucm) - ... | gt x = const∙ _ _ + help : (n : Fin dim) + → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) + → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) + help n = invEq (_ , InductiveFinSatAC _ _ _) + λ m → Iso.fun (PathIdTruncIso _) + (isContr→isProp + (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) + + module _ (F : funType) where + card' : ℕ → ℕ + card' n with (n ≟ᵗ dim) + ... | lt x = card (suc n) + ... | eq x = 0 + ... | gt x = 0 + + funs : (n : ℕ) → SphereBouquet∙ n (Fin (card' n)) + →∙ (A' (suc n) , ptA n) + funs n with (n ≟ᵗ dim) + ... | lt x = F (n , x) .fst + ... | eq x = const∙ _ _ + ... | gt x = const∙ _ _ + + funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) + ≃∙ Pushout (funs n .fst ∘ inr) (λ r → fst r) , inl (ptA n) + funEqP1 n = invEquiv (isoToEquiv (⋁-cofib-Iso _ (funs n))) , refl + + funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) + ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) + funEq n with (n ≟ᵗ dim) + ... | lt x = isoToEquiv (pushoutIso' _ _ _ _ + (idEquiv _) + (idEquiv _) + (idEquiv _) + (funExt (uncurry λ x y → F (n , _) .snd (x , y))) + (λ i x → fst x)) + , λ _ → inl (ptA n) + ... | eq x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) + (compEquiv ((CW↪ (_ , C) (suc n)) + , transport (λ i → isEquiv (CW↪ (A' , C) + (suc (x (~ i))))) + (ind 1) + ) + (e (suc n)))) , secEq (e (suc n)) _ + ... | gt x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) + (compEquiv ((CW↪ (_ , C) (suc n)) + , (transport (λ i → isEquiv (CW↪ (A' , C) + (suc ((sym (+-suc (fst (<ᵗ→< x)) dim) + ∙ (<ᵗ→< x .snd)) i)))) + (ind (suc (suc (fst (<ᵗ→< x))))))) + (e (suc n)))) , secEq (e (suc n)) _ + + goodCWmk : yieldsGoodCWskel (λ n → A' (suc n) , ptA n) + fst goodCWmk = card' + fst (snd goodCWmk) = compEquiv e₀ (pathToEquiv (cong Fin cA)) + fst (snd (snd goodCWmk)) = funs + snd (snd (snd goodCWmk)) n = + compEquiv∙ (compEquiv∙ (funEqP1 n) (funEq n)) + (invEquiv (e (suc n)) , refl) + + goodCWmk-converges : converges + (sequence (obj (GoodCWskelSeq goodCWmk)) + (Sequence.map (GoodCWskelSeq goodCWmk))) + dim + goodCWmk-converges zero = help dim refl + where + help : (x : _) (p : dim ≡ x) → isEquiv (Sequence.map (GoodCWskelSeq goodCWmk) {x}) + help zero p = ⊥.rec (¬dim≡0 p) + help (suc x) p with (x ≟ᵗ dim) + ... | lt x₁ = transport (λ i → isEquiv λ z → CW↪ (A' , C) (p i) z) (ind zero) + ... | eq x₁ = ⊥.rec (¬m Date: Wed, 29 May 2024 18:11:46 +0200 Subject: [PATCH 44/73] stuff --- Cubical/CW/Pointed2.agda | 1158 ++++++++++++++------------ Cubical/Foundations/Prelude.agda | 4 + Cubical/HITs/Pushout/Properties.agda | 402 ++++++++- 3 files changed, 1028 insertions(+), 536 deletions(-) diff --git a/Cubical/CW/Pointed2.agda b/Cubical/CW/Pointed2.agda index 42b120983d..027143074d 100644 --- a/Cubical/CW/Pointed2.agda +++ b/Cubical/CW/Pointed2.agda @@ -7,6 +7,8 @@ module Cubical.CW.Pointed2 where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws + open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order @@ -44,86 +46,137 @@ open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence -liftFun : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} - (f : A → B) → Lift {j = ℓ''} A → Lift {j = ℓ'''} B -liftFun f (lift a) = lift (f a) - -module _ {ℓ ℓ' ℓ'' : Level} (ℓ''' : Level) - {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (f : A → B) (g : A → C) where - PushoutLevel : Level - PushoutLevel = ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')) - - PushoutLift : Type PushoutLevel - PushoutLift = Pushout {A = Lift {j = ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')} A} - {B = Lift {j = ℓ-max ℓ (ℓ-max ℓ'' ℓ''')} B} - {C = Lift {j = ℓ-max ℓ (ℓ-max ℓ' ℓ''')} C} - (liftFun f) - (liftFun g) - - PushoutLiftIso : Iso (Pushout f g) PushoutLift - Iso.fun PushoutLiftIso (inl x) = inl (lift x) - Iso.fun PushoutLiftIso (inr x) = inr (lift x) - Iso.fun PushoutLiftIso (push a i) = push (lift a) i - Iso.inv PushoutLiftIso (inl (lift x)) = inl x - Iso.inv PushoutLiftIso (inr (lift x)) = inr x - Iso.inv PushoutLiftIso (push (lift a) i) = push a i - Iso.rightInv PushoutLiftIso (inl (lift x)) = refl - Iso.rightInv PushoutLiftIso (inr (lift x)) = refl - Iso.rightInv PushoutLiftIso (push (lift a) i) = refl - Iso.leftInv PushoutLiftIso (inl x) = refl - Iso.leftInv PushoutLiftIso (inr x) = refl - Iso.leftInv PushoutLiftIso (push a i) = refl - open import Cubical.Foundations.Equiv -pushoutIso' : ∀ {ℓA₁ ℓB₁ ℓC₁ ℓA₂ ℓB₂ ℓC₂} - {A₁ : Type ℓA₁} {B₁ : Type ℓB₁} {C₁ : Type ℓC₁} - {A₂ : Type ℓA₂} {B₂ : Type ℓB₂} {C₂ : Type ℓC₂} - (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) - (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) - (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) - (id1 : fst B≃ ∘ f₁ ≡ f₂ ∘ fst A≃) - (id2 : fst C≃ ∘ g₁ ≡ g₂ ∘ fst A≃) → - Iso (Pushout f₁ g₁) (Pushout f₂ g₂) -pushoutIso' {ℓA₁ = ℓA₁} {ℓB₁} {ℓC₁} {ℓA₂} {ℓB₂} {ℓC₂} - f₁ g₁ f₂ g₂ A≃ B≃ C≃ id1 id2 = - compIso (PushoutLiftIso ℓ* f₁ g₁) - (compIso (pushoutIso _ _ _ _ - (Lift≃Lift A≃) - (Lift≃Lift B≃) - (Lift≃Lift C≃) - (funExt (λ { (lift x) → cong lift (funExt⁻ id1 x)})) - (funExt (λ { (lift x) → cong lift (funExt⁻ id2 x)}))) - (invIso (PushoutLiftIso ℓ* f₂ g₂))) - where - ℓ* = ℓ-max ℓA₁ (ℓ-max ℓA₂ (ℓ-max ℓB₁ (ℓ-max ℓB₂ (ℓ-max ℓC₁ ℓC₂)))) -module _ {C : Type ℓ} {B : Type ℓ'} where - PushoutAlongEquiv→ : {A : Type ℓ} - (e : A ≃ C) (f : A → B) → Pushout (fst e) f → B - PushoutAlongEquiv→ e f (inl x) = f (invEq e x) - PushoutAlongEquiv→ e f (inr x) = x - PushoutAlongEquiv→ e f (push a i) = f (retEq e a i) +module _ {A : Pointed ℓ} {B : Pointed ℓ'} where - private - PushoutAlongEquiv→Cancel : {A : Type ℓ} (e : A ≃ C) (f : A → B) - → retract (PushoutAlongEquiv→ e f) inr - PushoutAlongEquiv→Cancel = - EquivJ (λ A e → (f : A → B) - → retract (PushoutAlongEquiv→ e f) inr) - λ f → λ { (inl x) → sym (push x) - ; (inr x) → refl - ; (push a i) j → push a (~ j ∨ i)} - - PushoutAlongEquiv : {A : Type ℓ} (e : A ≃ C) (f : A → B) - → Iso (Pushout (fst e) f) B - Iso.fun (PushoutAlongEquiv e f) = PushoutAlongEquiv→ e f - Iso.inv (PushoutAlongEquiv e f) = inr - Iso.rightInv (PushoutAlongEquiv e f) x = refl - Iso.leftInv (PushoutAlongEquiv e f) = PushoutAlongEquiv→Cancel e f + cofibConst→⋁ : cofib (const∙ A B .fst) → Susp∙ (fst A) ⋁ B + cofibConst→⋁ (inl x) = inl north + cofibConst→⋁ (inr x) = inr x + cofibConst→⋁ (push a i) = ((λ i → inl (toSusp A a i)) ∙ push tt) i + ⋁→cofibConst : Susp∙ (fst A) ⋁ B → cofib (const∙ A B .fst) + ⋁→cofibConst (inl north) = inl tt + ⋁→cofibConst (inl south) = inr (pt B) + ⋁→cofibConst (inl (merid a i)) = push a i + ⋁→cofibConst (inr x) = inr x + ⋁→cofibConst (push a i) = push (pt A) i + open import Cubical.Foundations.GroupoidLaws + + cofibConst : Iso (cofib (const∙ A B .fst)) + (Susp∙ (fst A) ⋁ B) + Iso.fun cofibConst = cofibConst→⋁ + Iso.inv cofibConst = ⋁→cofibConst + Iso.rightInv cofibConst (inl north) = refl + Iso.rightInv cofibConst (inl south) = sym (push tt) ∙ λ i → inl (merid (pt A) i) + Iso.rightInv cofibConst (inl (merid a i)) j = + hcomp (λ k → + λ {(i = i0) → inl (compPath-filler (merid a) (sym (merid (pt A))) (~ j) (~ k)) + ; (i = i1) → (sym (push tt) ∙ λ i → inl (merid (pt A) i)) j + ; (j = i0) → compPath-filler' (λ i → inl (toSusp A a i)) (push tt) k i + ; (j = i1) → inl (merid a (~ k ∨ i))}) + (compPath-filler' (sym (push tt)) (λ i → inl (merid (pt A) i)) i j) + Iso.rightInv cofibConst (inr x) = refl + Iso.rightInv cofibConst (push tt i) j = + (cong (_∙ push tt) (λ i j → inl (rCancel (merid (pt A)) i j)) + ∙ sym (lUnit _)) j i + Iso.leftInv cofibConst (inl x) = refl + Iso.leftInv cofibConst (inr x) = refl + Iso.leftInv cofibConst (push a i) j = help j i + where + help : cong ⋁→cofibConst (((λ i → inl (toSusp A a i)) ∙ push tt)) ≡ push a + help = cong-∙ ⋁→cofibConst (λ i → inl (toSusp A a i)) (push tt) + ∙ cong₂ _∙_ (cong-∙ (⋁→cofibConst ∘ inl) (merid a) (sym (merid (pt A)))) refl + ∙ sym (assoc _ _ _) + ∙ cong (push a ∙_) (lCancel (push (pt A))) + ∙ sym (rUnit _) + + cofibConst' : (f : A →∙ B) → ((x : fst A) → fst f x ≡ pt B) + → Iso (cofib (fst f)) ((Susp∙ (fst A) ⋁ B)) + cofibConst' f d = + compIso (pushoutIso _ _ _ _ + (idEquiv _) (idEquiv _) (idEquiv _) refl (funExt d)) + cofibConst + +open import Cubical.HITs.Pushout +module _ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where + + private + open 3x3-span + inst : 3x3-span + A00 inst = Unit + A02 inst = Unit + A04 inst = Unit + A20 inst = fst A + A22 inst = Unit + A24 inst = Unit + A40 inst = fst B + A42 inst = fst B + A44 inst = Unit + f10 inst = _ + f12 inst = _ + f14 inst = _ + f30 inst = fst f + f32 inst = λ _ → pt B + f34 inst = _ + f01 inst = _ + f21 inst = λ _ → pt A + f41 inst = idfun (fst B) + f03 inst = _ + f23 inst = _ + f43 inst = _ + H11 inst = (λ _ → refl) + H13 inst = λ _ → refl + H31 inst = λ _ → sym (snd f) + H33 inst = λ _ → refl + + A□0≅cofib-f : Iso (A□0 inst) (cofib (fst f)) + A□0≅cofib-f = idIso + + A□2≅B : Iso (A□2 inst) (fst B) + A□2≅B = PushoutAlongEquiv (idEquiv _) λ _ → pt B + + A□4≅Unit : Iso (A□4 inst) Unit + A□4≅Unit = PushoutAlongEquiv (idEquiv _) λ _ → tt + + A0□≅Unit : Iso (A0□ inst) Unit + A0□≅Unit = PushoutAlongEquiv (idEquiv _) λ _ → tt + + A2□≅A : Iso (A2□ inst) (fst A) + A2□≅A = compIso (equivToIso (invEquiv (symPushout _ _))) + (PushoutAlongEquiv (idEquiv _) λ _ → pt A) + + A4□≅Unit : Iso (A4□ inst) Unit + A4□≅Unit = PushoutAlongEquiv (idEquiv _) λ _ → tt + + A□○≅cofibInr : Iso (A□○ inst) (cofib {B = cofib (fst f)} inr) + A□○≅cofibInr = compIso (invIso (equivToIso (symPushout _ _))) + (pushoutIso _ _ _ _ + (isoToEquiv A□2≅B) + (isoToEquiv A□4≅Unit) + (isoToEquiv A□0≅cofib-f) + refl (funExt λ x + → cong (A□0≅cofib-f .Iso.fun ∘ f□1 inst) + (sym (Iso.leftInv A□2≅B x)))) + + A○□≅ : Iso (A○□ inst) (Susp (typ A)) + A○□≅ = + compIso + (pushoutIso _ _ _ _ + (isoToEquiv A2□≅A) + (isoToEquiv A0□≅Unit) + (isoToEquiv A4□≅Unit) + refl refl) + PushoutSuspIsoSusp + + Iso-cofibInr-Susp : Iso (cofib {B = cofib (fst f)} inr) + (Susp (typ A)) + Iso-cofibInr-Susp = + compIso (compIso (invIso A□○≅cofibInr) + (3x3-Iso inst)) A○□≅ --- CW complexes --- -- Definition of (the skeleton) of an arbitrary CW complex @@ -174,7 +227,7 @@ module _ {ℓA ℓB ℓC : Level} {A : Type ℓA} {B : A → Pointed ℓB} (C : A○□Iso : Iso (A○□ inst) (Pushout (fst f ∘ inr) fst) A○□Iso = compIso (equivToIso (symPushout _ _)) - (invIso (pushoutIso' _ _ _ _ + (invIso (pushoutIso _ _ _ _ (isoToEquiv (invIso A2□Iso)) (isoToEquiv (invIso A4□Iso)) (isoToEquiv (invIso A0□Iso)) @@ -199,7 +252,7 @@ module _ {ℓA ℓB ℓC : Level} {A : Type ℓA} {B : A → Pointed ℓB} (C : open import Cubical.Foundations.GroupoidLaws A□○Iso : Iso (A□○ inst) (cofib (fst f)) - A□○Iso = invIso (pushoutIso' _ _ _ _ + A□○Iso = invIso (pushoutIso _ _ _ _ (isoToEquiv (invIso A□2Iso)) (isoToEquiv (invIso A□0Iso)) (isoToEquiv (invIso A□4Iso)) @@ -228,27 +281,37 @@ open import Cubical.Homotopy.Connected open import Cubical.CW.Properties open import Cubical.HITs.Truncation as TR open import Cubical.Foundations.HLevels +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Fin.Inductive.Properties as Ind isConnectedCW→isConnectedSkel : (C : CWskel ℓ) + → (m : ℕ) (n : Fin (suc m)) + → isConnected (fst n) (realise C) + → isConnected (fst n) (C .fst m) +isConnectedCW→isConnectedSkel C m (n , p) = + isOfHLevelRetractFromIso 0 + (compIso (truncOfTruncIso n t₀) + (compIso + (mapCompIso + (subst (λ k → Iso (hLevelTrunc k (C .fst m)) + (hLevelTrunc k (realise C))) + (sym teq) + (connectedTruncIso m _ + (isConnected-CW↪∞ m C)))) + (invIso (truncOfTruncIso n t₀)))) + where + t = <ᵗ→< p + t₀ = fst t + teq : t₀ +ℕ n ≡ m + teq = cong predℕ (sym (+-suc t₀ n) ∙ snd t) + +is2ConnectedCW→is2ConnectedSkel : (C : CWskel ℓ) → isConnected 2 (realise C) → (n : ℕ) → isConnected 2 (C .fst (suc (suc n))) -isConnectedCW→isConnectedSkel C c n = - TR.rec (isOfHLevelSuc 1 isPropIsContr) - (λ ⋆ → TR.rec (isProp→isOfHLevelSuc (suc n) isPropIsContr) - (λ {(x , p) → ∣ x ∣ - , (TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) - (λ a → Iso.inv (PathIdTruncIso 1) - (TR.rec (isOfHLevelTrunc 1) - (λ q → - TR.rec (isOfHLevelPlus' 1 (isOfHLevelTrunc 1)) - (∣_∣ₕ ∘ fst) - (isConnectedCong (suc n) _ (isConnected-CW↪∞ (suc (suc n)) C) - q .fst)) - (Iso.fun (PathIdTruncIso 1) - (sym (c .snd ∣ incl x ∣) ∙ c .snd ∣ incl a ∣)))))}) - (isConnected-CW↪∞ (suc (suc n)) C ⋆ .fst)) - (c .fst) +is2ConnectedCW→is2ConnectedSkel C c n = + isConnectedCW→isConnectedSkel C (suc (suc n)) + (2 , tt) c open import Cubical.Data.Bool open import Cubical.Data.Nat.Order.Inductive @@ -348,7 +411,6 @@ decIm-S⁰×Fin {A = A} da n = (isoToPath (invIso Bool×Fin≡Fin)) (decImFin da _) - module _ {ℓ : Level} {n : ℕ} {A : Fin n → Type ℓ} (x₀ : Fin n) (pt : A x₀) (l : (x : Fin n) → ¬ x ≡ x₀ → A x) where private @@ -950,6 +1012,9 @@ yieldsConnectedCWskel' A n = Σ[ sk ∈ yieldsCWskel A ] isConnected (2 +ℕ n) connectedCWskel : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) connectedCWskel ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel X n) +isConnectedCW : ∀ {ℓ} (n : ℕ) → Type ℓ → Type (ℓ-suc ℓ) +isConnectedCW {ℓ = ℓ} n A = Σ[ sk ∈ connectedCWskel ℓ n ] realise (_ , (snd sk .fst)) ≃ A + connectedCWskel' : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) connectedCWskel' ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel' X n) @@ -1005,7 +1070,7 @@ module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A 0) where snd (snd (snd (fst connectedCWskel→))) (suc zero) = compEquiv (invEquiv (isoToEquiv e₁)) (compEquiv (isoToEquiv (snd liftStr)) - (isoToEquiv (pushoutIso' _ _ _ _ + (isoToEquiv (pushoutIso _ _ _ _ (idEquiv _) LiftEquiv (idEquiv _) (funExt cohₗ) (funExt cohᵣ)))) where @@ -1176,7 +1241,7 @@ ptCW sk a (suc n) = CW↪ (_ , sk) (suc n) (ptCW sk a n) -- funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) -- ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) --- funEq n = isoToEquiv (pushoutIso' _ _ _ _ +-- funEq n = isoToEquiv (pushoutIso _ _ _ _ -- (idEquiv _) -- (idEquiv _) -- (idEquiv _) @@ -1301,7 +1366,7 @@ module BS {ℓ : Level} (A' : ℕ → Type ℓ) funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) funEq n with (n ≟ᵗ dim) - ... | lt x = isoToEquiv (pushoutIso' _ _ _ _ + ... | lt x = isoToEquiv (pushoutIso _ _ _ _ (idEquiv _) (idEquiv _) (idEquiv _) @@ -1379,158 +1444,36 @@ module BS {ℓ : Level} (A' : ℕ → Type ℓ) getGoodCWskel = PT.map (λ F → (goodCWmk F) , (goodCWmk-converges F)) mapMakerNil - getGoodCWskelAux : ∥ yieldsGoodCWskel (λ n → A' (suc n) , ptA n) ∥₁ - getGoodCWskelAux = PT.map help mapMakerNil - where - help : funType → yieldsGoodCWskel (λ n → A' (suc n) , ptA n) - fst (help F) = card' F - fst (snd (help F)) = compEquiv e₀ (pathToEquiv (cong Fin cA)) - fst (snd (snd (help F))) = funs F - snd (snd (snd (help F))) n = - compEquiv∙ (compEquiv∙ (funEqP1 F n) (funEq F n)) - (invEquiv (e (suc n)) , refl) +isFinConnCW : (n : ℕ) → Type ℓ → Type (ℓ-suc ℓ) +isFinConnCW n A = isFinCW A × isConnected n A +isConnCW : (n : ℕ) → Type ℓ → Type (ℓ-suc ℓ) +isConnCW n A = isCW A × isConnected n A - {- - fst (help F) = card ∘ suc - fst (snd (help F)) = compEquiv e₀ (pathToEquiv (cong Fin cA)) - fst (snd (snd (help F))) n = funs F n - snd (snd (snd (help F))) n = - compEquiv∙ (compEquiv∙ (funEqP1 F n) (funEq F n)) - (invEquiv (e (suc n)) , refl) --} - --- module _ {ℓ : Level} (A' : ℕ → Type ℓ) --- (C : yieldsCWskel A') --- (dim : ℕ) (ind : (n : ℕ) → dim <ᵗ n → fst C n ≡ 0) --- (cA : fst C 0 ≡ 1) --- where - --- open CWskel-fields (_ , C) --- e₀ : A' 1 ≃ Fin (card 0) --- e₀ = CW₁-discrete (_ , C) - - --- pt0 : A' 1 --- pt0 = invEq e₀ (subst Fin (sym cA) flast) - --- ptA : (n : ℕ) → A' (suc n) --- ptA = ptCW C pt0 - --- conA : (n : ℕ) → isConnected 2 (A' (suc n)) --- conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) --- (subst (isConnected 2) (sym (cong Fin cA)) --- (∣ flast ∣ --- , TR.elim (λ _ → isOfHLevelPath 2 --- (isOfHLevelTrunc 2) _ _) --- λ {(zero , tt) → refl})) --- conA (suc n) = --- isConnectedRetractFromIso 2 --- (equivToIso (e (suc n))) --- (∣ inl (ptA n) ∣ₕ --- , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) --- (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) --- (conA' (conA n)) --- λ c → conA' (conA n) _ --- ∙ cong ∣_∣ₕ (push (c , ptSn n)))) --- where --- conA' : (conA : isConnected 2 (A' (suc n))) (c : A' (suc n)) --- → Path (hLevelTrunc 2 (Pushout (α (suc n)) fst)) --- ∣ inl (ptA n) ∣ₕ ∣ inl c ∣ₕ --- conA' conA c = --- TR.rec (isOfHLevelTrunc 2 _ _) --- (λ p i → ∣ inl (p i) ∣) --- (Iso.fun (PathIdTruncIso _) --- (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) - - --- private --- funType = ((n : Fin (suc dim)) --- → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) --- →∙ (A' (suc (fst n)) , ptA (fst n))) ] --- ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) - --- mapMakerNil : ∥ funType ∥₁ --- mapMakerNil = --- invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) --- λ n → TR.map --- (λ pted → ((λ { (inl x) → ptA (fst n) --- ; (inr x) → α _ x --- ; (push a i) → pted a i}) --- , refl) --- , λ _ → refl) (help n)) --- where --- help : (n : Fin (suc dim)) --- → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) --- → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) --- help n = invEq (_ , InductiveFinSatAC _ _ _) --- λ m → Iso.fun (PathIdTruncIso _) --- (isContr→isProp --- (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) --- module _ (F : funType) where --- funs : (n : ℕ) → SphereBouquet∙ n (Fin (card (suc n))) --- →∙ (A' (suc n) , ptA n) --- funs n with (n ≟ᵗ dim) --- ... | lt x = F (n , <ᵗ-trans-suc x) .fst --- ... | eq x = F (n , subst (_<ᵗ suc dim) (sym x) <ᵗsucm) .fst --- ... | gt x = const∙ _ _ - --- funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) --- ≃∙ Pushout (funs n .fst ∘ inr) (λ r → fst r) , inl (ptA n) --- funEqP1 n = invEquiv (isoToEquiv (⋁-cofib-Iso _ (funs n))) , refl - --- funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) --- ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) --- funEq n = isoToEquiv (pushoutIso' _ _ _ _ --- (idEquiv _) --- (idEquiv _) --- (idEquiv _) --- (funExt (uncurry (main n))) --- (λ i x → fst x)) --- , λ _ → inl (ptA n) --- where --- main : (n : ℕ) (x : Fin (card (suc n))) (y : S₊ n) --- → funs n .fst (inr (x , y)) ≡ fst (C .snd) (suc n) (x , y) --- main n with (n ≟ᵗ dim) --- ... | lt p = λ x y --- → F (n , <ᵗ-trans-suc p) .snd (x , y) --- ... | eq p = λ x y --- → F (n , subst (_<ᵗ suc dim) (λ i → p (~ i)) <ᵗsucm) .snd (x , y) --- ... | gt p = λ x --- → ⊥.rec (¬Fin0 (subst Fin (ind (suc n) (<ᵗ-trans p <ᵗsucm)) x)) - --- getGoodCWskelAux : ∥ yieldsGoodCWskel (λ n → A' (suc n) , ptA n) ∥₁ --- getGoodCWskelAux = PT.map help mapMakerNil --- where --- help : funType → yieldsGoodCWskel (λ n → A' (suc n) , ptA n) --- fst (help F) = card ∘ suc --- fst (snd (help F)) = compEquiv e₀ (pathToEquiv (cong Fin cA)) --- fst (snd (snd (help F))) n = funs F n --- snd (snd (snd (help F))) n = --- compEquiv∙ (compEquiv∙ (funEqP1 F n) (funEq F n)) --- (invEquiv (e (suc n)) , refl) - - +finConnCW∙ : (n : ℕ) (ℓ : Level)→ Type (ℓ-suc ℓ) +finConnCW∙ n ℓ = Σ[ A ∈ Pointed ℓ ] ∥ isFinConnCW n (fst A) ∥₁ open import Cubical.CW.Subcomplex -finCW→GoodCW : ∀ {ℓ} (A : finCW ℓ) (a₀ : fst A) - → isConnected 2 (fst A) +finCW→GoodCW : ∀ {ℓ} + → finConnCW∙ 2 ℓ → finGoodCW ℓ -fst (finCW→GoodCW A a₀ cA) = fst A , a₀ -snd (finCW→GoodCW (A , cwA) a₀ cA) = - PT.rec squash₁ (λ cw → PT.rec squash₁ - (λ {(sk , T) → ∣ suc (fst cw) , (_ , sk , T) - , compEquiv - (isoToEquiv - (compIso (invIso (converges→ColimIso _ T)) - (compIso (eqv cw _ refl) - (converges→ColimIso _ (cw .snd .fst .snd .snd))))) - (invEquiv (cw .snd .snd)) , {!!} ∣₁}) - (main cw)) - cwA +fst (finCW→GoodCW A) = fst A +snd (finCW→GoodCW ((A , a₀) , cwA+cA)) = + PT.rec squash₁ (λ{(cw , cA) + → PT.rec squash₁ + (λ {(sk , T) + → TR.rec squash₁ + (λ p → ∣ (suc (fst cw)) + , ((_ , sk , T) , (mainEq (cw , cA) sk T , p)) ∣₁) + (mainEqId (cw , cA) sk T)}) + (main (cw , cA))}) + cwA+cA where - module _ (cw : isFinCW A) where + module _ (cw+cA : isFinConnCW 2 A) where + cA = snd cw+cA + cw = fst cw+cA + makeNice' = makeNiceFinCWskel {A = A} cw inst = connectedCWskel'→connectedCWskel @@ -1540,11 +1483,10 @@ snd (finCW→GoodCW (A , cwA) a₀ cA) = (equivToIso (invEquiv (snd makeNice' .snd))) cA)) , _ , snd makeNice' .fst .snd .snd) - open BS - (inst .fst .fst) - (suc (fst cw)) - ((snd (inst .fst) .fst) , inst .snd .snd) - refl + open BS (inst .fst .fst) + (suc (fst cw)) + ((snd (inst .fst) .fst) , inst .snd .snd) + refl main = BS.getGoodCWskel (inst .fst .fst) @@ -1570,297 +1512,451 @@ snd (finCW→GoodCW (A , cwA) a₀ cA) = ... | lt x₁ = ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (suc (suc x) <ᵗ_) p x₁) <ᵗsucm)) ... | eq x₁ = ⊥.rec (¬m<ᵗm (subst (suc x <ᵗ_) (x₁ ∙ p) <ᵗsucm)) ... | gt x₁ = pathToIso (cong (fst (snd cw) .fst) p) - - {- compIso (equivToIso (invEquiv (complex≃subcomplex (_ , snd cw .fst .snd .fst) {!suc (cw .fst)!} (suc (fst cw) , {!fst cw!})))) {!(fst (finCWskel→CWskel (fst cw) (fst (cw .snd))) (fst cw))!} --} - {- BS.getGoodCWskel _ ? ? -- (snd inst .fst) (fst cw) -- (snd inst .fst) (fst cw) - ? -- FN - ? -- refl -- refl - -} - -{- - inst = connectedCWskel'→connectedCWskel - ((snd cw .fst .fst) - , ((snd cw .fst .snd .fst) - , (isConnectedRetractFromIso 2 - (equivToIso (invEquiv (snd cw .snd))) cA))) - - makeNice = makeNiceFinCWskel {A = A} ((cw .fst) - , ((inst .fst , inst .snd .fst - , {!cw .snd .fst .snd .snd!}) - , compEquiv (snd cw .snd) {!!})) - - check : (n : ℕ) → fst cw <ᵗ n → fst (snd inst .fst) n ≡ 0 - check n p with (suc (fst cw) ≟ᵗ n) - ... | lt x = {!makeNiceFinCWskel!} - ... | eq x = {!x!} - ... | gt x = {!!} - - main = getGoodCWskelAux _ (makeNice .snd .fst .snd .fst) (fst cw) -- (snd inst .fst) (fst cw) - {!makeNice .snd .fst!} - {!!} -- refl - -} - --- ecw = snd (snd cw) - --- FL : (n : ℕ) → fst cw <ᵗ n --- → fst (makeNiceFinCWskel cw .snd .fst .snd .fst) n ≡ 0 --- FL n p = {!inst .snd .fst!} - --- FR : fst (makeNiceFinCWskel cw .snd .fst .snd .fst) 0 ≡ 1 --- FR with (0 ≟ᵗ fst cw) --- ... | lt t = {!!} --- ... | eq x = ⊥.rec {!!} --- main : {!!} --- main = getGoodCWskelAux _ --- (makeNiceFinCWskel cw .snd .fst .snd .fst) (fst cw) --- FL FR --- -- makeNiceFinCWElim - --- connectedGoodCWskel' : ∀ {ℓ} (Aₙ₊₂ : ℕ → Pointed ℓ) → {!!} --- connectedGoodCWskel' = {!!} - - - - --- GoodCWSeq→CWSeq : {ℓ : Level} (A₊₂ : ℕ → Pointed ℓ) → {!!} --- GoodCWSeq→CWSeq A₊₂ = {!!} - --- isGoodCW→isCW : ∀ {ℓ} (A₊₂ : ℕ → Pointed ℓ) → yieldsGoodCWskel A₊₂ → {!!} --- isGoodCW→isCW = {!!} - --- -- connectedCWskel→CWskel : connectedCWskel ℓ → CWskel ℓ --- -- connectedCWskel→CWskel (A , sk , _) = A , sk - --- -- module _ (C' : connectedCWskel ℓ) (n : ℕ) (x₀ : fst C' 1) --- -- (term : converges (realiseSeq (fst C' , snd C' .fst)) n) where --- -- private --- -- C = fst C' --- -- Csk = connectedCWskel→CWskel C' - --- -- open CWskel-fields Csk - --- -- ptC : (n : ℕ) → C (suc n) --- -- ptC zero = x₀ --- -- ptC (suc n) = CW↪ Csk (suc n) (ptC n) - --- -- C∙ : ℕ → Pointed _ --- -- C∙ n = C (suc n) , ptC n - --- -- improveFam : ∥ ((m : Fin (suc (suc n))) → ?) ∥₁ --- -- improveFam = {!!} - --- -- finConnectedCW : {!!} --- -- finConnectedCW = {!!} - - - - --- -- -- yieldsGoodCWSkel : (ℕ → Type ℓ) → Type ℓ --- -- -- yieldsGoodCWSkel = ? - --- -- -- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ --- -- -- yieldsCWskel∙ X = --- -- -- Σ[ f ∈ (ℕ → ℕ) ] --- -- -- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] --- -- -- ((X zero ≃ Fin (f zero)) × --- -- -- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) - --- -- -- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) - --- -- -- module CWskel∙-fields (C : CWskel∙ ℓ) where --- -- -- card = C .snd .fst --- -- -- A = Fin ∘ card --- -- -- α = C .snd .snd .fst --- -- -- e = C .snd .snd .snd .snd - --- -- -- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero --- -- -- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] - --- -- -- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- -- -- fst (CWpt (C , f) n) = C n --- -- -- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) - --- -- -- -- Technically, a CW complex should be the sequential colimit over the following maps --- -- -- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) --- -- -- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) - --- -- -- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n --- -- -- ptCW T zero = T .snd .snd .fst zero (inl tt) --- -- -- ptCW T (suc n) = CW∙↪ T n (ptCW T n) - --- -- -- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- -- -- CW∙ T n = fst T n , ptCW T n - --- -- -- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) --- -- -- fst (CW∙↪∙ T n) = CW∙↪ T n --- -- -- snd (CW∙↪∙ T n) = refl - - --- -- -- -- But if it stabilises, no need for colimits. --- -- -- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ --- -- -- yieldsFinCWskel∙ n X = --- -- -- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) - --- -- -- -- ... which should give us finite CW complexes. --- -- -- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) --- -- -- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) - --- -- -- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ --- -- -- finCWskel→CWskel∙ n C = fst C , fst (snd C) - --- -- -- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ --- -- -- Sequence.obj (realiseSeq∙ (C , p)) = C --- -- -- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ - --- -- -- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ --- -- -- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) - --- -- -- -- realisation of CW complex from skeleton --- -- -- realise∙ : CWskel∙ ℓ → Type ℓ --- -- -- realise∙ C = SeqColim (realiseSeq∙ C) - --- -- -- realise∙∙ : CWskel∙ ℓ → Pointed ℓ --- -- -- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) --- -- -- open import Cubical.Data.Empty as ⊥ + mainEq : (sk : _) (T : converges (GoodCWskelSeq sk) (suc (fst cw))) → _ + mainEq sk T = compEquiv + (isoToEquiv + (compIso (invIso (converges→ColimIso _ T)) + (compIso (eqv _ refl) + (converges→ColimIso _ (cw .snd .fst .snd .snd))))) + (invEquiv (cw .snd .snd)) --- -- -- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ --- -- -- CWskel∙→CWskel A zero = Lift ⊥ --- -- -- CWskel∙→CWskel A (suc n) = A n --- -- -- open import Cubical.Foundations.Isomorphism + mainEqId : (sk : _) (T : _) + → hLevelTrunc 1 (mainEq sk T .fst (pt (realiseGood∙ sk)) ≡ a₀) + mainEqId sk T = Iso.fun (PathIdTruncIso 1) (isContr→isProp cA _ _) --- -- -- module _ (A : ℕ → Type ℓ) --- -- -- (cwsk : yieldsCWskel∙ A) where +module GoodCW→finCWExpl {ℓ : Level} (A : Pointed ℓ) + (dim : ℕ) (sk : FinGoodCWskel ℓ dim) + (eq : realiseGood∙ (fst (snd sk)) ≃∙ A) + where --- -- -- private --- -- -- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n --- -- -- αs (suc n) x = snd cwsk .fst n (inr x) + Fam : ℕ → Type ℓ + Fam zero = Lift ⊥ + Fam (suc n) = fst sk n .fst + + card : ℕ → ℕ + card zero = 1 + card (suc n) = snd sk .fst .fst n + + α : (n : ℕ) → Fin (card n) × S⁻ n → Fam n + α (suc n) (a , b) = snd sk .fst .snd .snd .fst n .fst (inr (a , b)) + + e : (n : ℕ) → Iso (fst sk n .fst) + (Pushout (α n) fst) + e zero = compIso (equivToIso (snd sk .fst .snd .fst)) + (compIso (PushoutEmptyFam (λ()) (λ())) + (equivToIso (symPushout _ _))) + e (suc n) = compIso (equivToIso (invEquiv (snd sk .fst .snd .snd .snd n .fst))) + (invIso (⋁-cofib-Iso (fst sk n) + (fst (snd sk .fst .snd .snd) n))) + + yieldsFinCWskelFam : yieldsFinCWskel dim Fam + fst (fst yieldsFinCWskelFam) = card + fst (snd (fst yieldsFinCWskelFam)) = α + fst (snd (snd (fst yieldsFinCWskelFam))) () + snd (snd (snd (fst yieldsFinCWskelFam))) n = isoToEquiv (e n) + snd yieldsFinCWskelFam zero = help dim refl + where + help : (x : ℕ) (p : x ≡ dim) + → isEquiv (CW↪ (Fam , card , (α , (λ()) , (λ n → isoToEquiv (e n)))) x) + help zero p with + (invEq (_ , transport (λ i → isEquiv (Sequence.map (GoodCWskelSeq (fst (snd sk))) + {n = p (~ i)})) + (sk .snd .snd 0)) (fst sk 0 .snd)) + ... | () + + help (suc x) p = + transport (λ i → isEquiv (Sequence.map (GoodCWskelSeq (fst (snd sk))) {n = p (~ i)})) + (sk .snd .snd 0) + snd yieldsFinCWskelFam (suc k) = snd sk .snd (suc k) + + Skel : finCWskel ℓ dim + fst Skel = Fam + snd Skel = yieldsFinCWskelFam + + SkelEq : fst A ≃ realise (finCWskel→CWskel dim Skel) + SkelEq = compEquiv (invEquiv (eq .fst)) + (isoToEquiv + (compIso (Iso-SeqColim→SeqColimShift _ 2) + (compIso (sequenceIso→ColimIso + ((λ n → idIso {A = (fst (fst sk (suc n)))}) + , λ _ _ → refl)) + (invIso (Iso-SeqColim→SeqColimShift _ 2))))) + + conn : isConnected 2 (fst A) + conn = isConnectedRetractFromIso 2 (equivToIso SkelEq) + (isOfHLevelRetractFromIso 0 + (compIso + (invIso (connectedTruncIso 2 _ + (isConnected-CW↪∞ 2 (finCWskel→CWskel dim Skel)))) + (mapCompIso (equivToIso + (snd (finCWskel→CWskel dim Skel) .snd .snd .snd 1) ))) + (∣ inl (fst sk 0 .snd) ∣ₕ + , (TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) + (λ b i → ∣ inl (isProp0 (fst sk 0 .snd) b i) ∣ₕ) + (uncurry λ x y → (λ i → ∣ inl (isProp0 (fst sk 0 .snd) + (α 1 ((x , y) , true)) i) ∣ₕ) + ∙ cong ∣_∣ₕ (push ((x , y) , true))))))) + where + isProp0 : isProp (fst sk 0 .fst) + isProp0 = isOfHLevelRetractFromIso 1 + (equivToIso(sk .snd .fst .snd .fst)) + isPropFin1 + + +GoodCW→finCW : ∀ {ℓ} + → finGoodCW ℓ → finConnCW∙ 2 ℓ +fst (GoodCW→finCW (A , str)) = A +snd (GoodCW→finCW (A , str)) = + PT.rec squash₁ + (λ {(dim , sk , e) + → ∣ (dim + , (GoodCW→finCWExpl.Skel A dim sk e + , GoodCW→finCWExpl.SkelEq A dim sk e)) + , (GoodCW→finCWExpl.conn A dim sk e) ∣₁}) + str + +finGoodCW≅finConnCW∙ : ∀ {ℓ} → Iso (finGoodCW ℓ) (finConnCW∙ 2 ℓ) +Iso.fun finGoodCW≅finConnCW∙ = GoodCW→finCW +Iso.inv finGoodCW≅finConnCW∙ = finCW→GoodCW +Iso.rightInv finGoodCW≅finConnCW∙ A = Σ≡Prop (λ _ → squash₁) refl +Iso.leftInv finGoodCW≅finConnCW∙ A = Σ≡Prop (λ _ → squash₁) refl + +finGoodCW≡finConnCW∙ : ∀ {ℓ} → finGoodCW ℓ ≡ finConnCW∙ 2 ℓ +finGoodCW≡finConnCW∙ = isoToPath finGoodCW≅finConnCW∙ + +elimFinConnCW∙ : ∀ {ℓ ℓ'} {P : finConnCW∙ 2 ℓ → Type ℓ'} + → ((c : finGoodCW ℓ) → P (GoodCW→finCW c)) + → (x : _) → P x +elimFinConnCW∙ {P = P} ind x = + subst P (Iso.rightInv finGoodCW≅finConnCW∙ x) (ind _) + +isConnectedCW→Contr : ∀ {ℓ} (n : ℕ) + (sk : connectedCWskel ℓ n) (m : Fin (suc n)) + → isContr (fst sk (suc (fst m))) +isConnectedCW→Contr zero sk (zero , p) = + isOfHLevelRetractFromIso 0 (equivToIso (CW₁-discrete (_ , snd sk .fst))) + (subst isContr (cong Fin (sym (snd sk .snd .fst))) + (flast , isPropFin1 _)) +isConnectedCW→Contr (suc n) sk (zero , p) = + subst isContr (ua (symPushout _ _) + ∙ ua (invEquiv (sk .snd .fst .snd .snd .snd 0))) + (isOfHLevelRetractFromIso 0 + (invIso (PushoutEmptyFam (λ()) (snd sk .fst .snd .snd .fst))) + (subst (isContr ∘ Fin) (sym (snd sk .snd .fst)) + (flast , isPropFin1 _))) +isConnectedCW→Contr (suc n) sk (suc x , p) + with (isConnectedCW→Contr n + (fst sk , (snd sk .fst) , ((snd sk .snd .fst) + , (λ p w → snd sk .snd .snd p (<ᵗ-trans w <ᵗsucm)))) + (x , p)) +... | ind = subst isContr + (ua (invEquiv (sk .snd .fst .snd .snd .snd (suc x)))) + (isOfHLevelRetractFromIso 0 + (invIso + (PushoutEmptyFam + (λ p' → ¬Fin0 (subst Fin + (snd sk .snd .snd x p) (fst p'))) + λ p' → ¬Fin0 (subst Fin + (snd sk .snd .snd x p) p'))) + ind) + +SphereBouquet₀Iso : (n : ℕ) + → Iso (SphereBouquet zero (Fin n)) + (Fin (suc n)) +Iso.fun (SphereBouquet₀Iso n) (inl x) = fzero +Iso.fun (SphereBouquet₀Iso n) (inr ((x , p) , false)) = suc x , p +Iso.fun (SphereBouquet₀Iso n) (inr ((x , p) , true)) = fzero +Iso.fun (SphereBouquet₀Iso n) (push a i) = fzero +Iso.inv (SphereBouquet₀Iso n) (zero , p) = inl tt +Iso.inv (SphereBouquet₀Iso n) (suc x , p) = inr ((x , p) , false) +Iso.rightInv (SphereBouquet₀Iso n) (zero , p) = refl +Iso.rightInv (SphereBouquet₀Iso n) (suc x , p) = refl +Iso.leftInv (SphereBouquet₀Iso n) (inl x) = refl +Iso.leftInv (SphereBouquet₀Iso n) (inr (x , false)) = refl +Iso.leftInv (SphereBouquet₀Iso n) (inr (x , true)) = push x +Iso.leftInv (SphereBouquet₀Iso n) (push a i) j = push a (i ∧ j) + +isConnectedCW→ContrIso : ∀ {ℓ} (n : ℕ) + (sk : connectedCWskel ℓ n) + → Iso (Pushout (snd sk .fst .snd .fst n) fst) + (SphereBouquet n (CWskel-fields.A (_ , snd sk .fst) n)) +isConnectedCW→ContrIso zero sk = + compIso (compIso (equivToIso (symPushout _ _)) + (invIso (PushoutEmptyFam + (λ()) (sk .snd .fst .snd .snd .fst)))) + (compIso (isContr→Iso (subst isContr + (cong Fin (sym (snd sk .snd .fst))) (flast , isPropFin1 _)) + (flast , isPropFin1 _)) + (compIso (invIso (SphereBouquet₀Iso 0)) + {!!})) +isConnectedCW→ContrIso (suc n) sk = + compIso ((invIso (PushoutEmptyFam + {!λ !} + λ p → ¬Fin0 (subst Fin (snd sk .snd .snd n <ᵗsucm) p)))) -- (λ()) (sk .snd .fst .snd .snd .fst)))) + {!!} + +connectedCWskelLift : ∀ {ℓ} (n : ℕ) {C : Type ℓ} + → (sk : isConnCW (suc (suc n)) C) + → ∥ isConnectedCW n C ∥₁ +connectedCWskelLift zero {C = C} ((cwsk , e) , cA) = + ∣ (_ , (M .fst , refl , λ p → λ())) + , compEquiv (invEquiv (collapse₁Equiv _ _ 0)) (invEquiv e) ∣₁ + where + M = connectedCWskel→ (cwsk .fst) (snd cwsk , subst (isConnected 2) (ua e) cA) +connectedCWskelLift (suc n) {C = C} ((cwsk , eqv) , cA) with + (connectedCWskelLift n ((cwsk , eqv) , (isConnectedSubtr (suc (suc n)) 1 cA))) +... | s = PT.map {!C*!} s + where + module _ (ind : isConnectedCW n C) where + open CWskel-fields (_ , ind .fst .snd .fst) + C* = ind .fst .fst + + iso1 : Iso (C* (suc n)) (Fin 1) + iso1 = isContr→Iso (isConnectedCW→Contr n (ind .fst) (n , <ᵗsucm)) + (flast , isPropFin1 _) + + ww = Susp→PushoutSusp + + iso2 : Iso (C* (suc (suc n))) (SphereBouquet (suc n) (A (suc n))) + iso2 = compIso (equivToIso + (ind .fst .snd .fst .snd .snd .snd (suc n))) + (compIso + (compIso + (compIso + (pushoutIso _ _ _ _ (idEquiv _) + (isoToEquiv (compIso iso1 + (isContr→Iso (flast , isPropFin1 _) + (tt , λ _ → refl)))) + (idEquiv _) + refl + refl) + (⋁-cofib-Iso _ (const∙ _ _))) + PushoutSuspIsoSusp ) + sphereBouquetSuspIso) --- -- -- e0 : {!!} --- -- -- e0 = {!!} --- -- -- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst --- -- -- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) --- -- -- es-suc→ n (inr x) = inl x --- -- -- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) --- -- -- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i --- -- -- es-suc→ n (push (push a i₁) i) = {!!} + α' : (A (suc (suc n))) × S₊ (suc n) → SphereBouquet (suc n) (A (suc n)) + α' = Iso.fun iso2 ∘ α (suc (suc n)) --- -- -- es-suc : (n : ℕ) --- -- -- → Iso (cofib (fst (snd cwsk) n)) --- -- -- (Pushout (αs (suc n)) fst) --- -- -- Iso.fun (es-suc n) = es-suc→ n --- -- -- Iso.inv (es-suc n) = {!!} --- -- -- Iso.rightInv (es-suc n) = {!!} --- -- -- Iso.leftInv (es-suc n) = {!!} - --- -- -- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) --- -- -- es zero = {!!} --- -- -- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) - --- -- -- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) --- -- -- fst yieldsCWskel∙→' = cwsk .fst --- -- -- fst (snd yieldsCWskel∙→') = αs --- -- -- fst (snd (snd yieldsCWskel∙→')) () --- -- -- snd (snd (snd yieldsCWskel∙→')) = {!!} - --- -- -- yieldsCWskel∙→ : (A : ℕ → Type ℓ) --- -- -- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) --- -- -- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst --- -- -- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) --- -- -- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () --- -- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} --- -- -- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = --- -- -- compEquiv (cwsk .snd .snd .snd n) --- -- -- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) --- -- -- where --- -- -- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst --- -- -- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) --- -- -- theEq→ (inr x) = inl x --- -- -- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) --- -- -- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) --- -- -- theEq→ (push (push a i₁) i) = {!!} - --- -- -- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) --- -- -- Iso.fun theEq = theEq→ --- -- -- Iso.inv theEq = {!!} --- -- -- Iso.rightInv theEq x = {!!} --- -- -- Iso.leftInv theEq x = {!!} - - --- -- -- -- compEquiv {!!} {!!} - - --- -- -- -- -- Finally: definition of CW complexes --- -- -- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- -- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' - --- -- -- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ - --- -- -- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) - --- -- -- -- -- Finite CW complexes --- -- -- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- -- -- isFinCW {ℓ = ℓ} X = --- -- -- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) - --- -- -- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ - --- -- -- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- -- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) - --- -- -- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X --- -- -- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str - --- -- -- -- finCW→CW : finCW ℓ → CW ℓ --- -- -- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p - - --- -- -- -- -- morphisms --- -- -- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') --- -- -- -- C →ᶜʷ D = fst C → fst D - --- -- -- -- --the cofibre of the inclusion of X n into X (suc n) --- -- -- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ --- -- -- -- cofibCW n C = cofib (CW↪ C n) - --- -- -- -- --...is basically a sphere bouquet --- -- -- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) --- -- -- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) --- -- -- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e --- -- -- -- where --- -- -- -- s = Bouquet≃-gen --- -- -- -- α = C .snd .snd .fst --- -- -- -- e = C .snd .snd .snd .snd n - --- -- -- -- --sending X (suc n) into the cofibCW --- -- -- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C --- -- -- -- to_cofibCW n C x = inr x - --- -- -- -- --the connecting map of the long exact sequence --- -- -- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) --- -- -- -- → cofib conn → Susp A --- -- -- -- δ-pre conn (inl x) = north --- -- -- -- δ-pre conn (inr x) = south --- -- -- -- δ-pre conn (push a i) = merid a i - --- -- -- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) --- -- -- -- δ n C = δ-pre (CW↪ C n) - --- -- -- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) --- -- -- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C --- -- -- -- CW↪∞ C n x = incl x - --- -- -- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m --- -- -- -- fst (finCW↑ m n p C) = fst C --- -- -- -- fst (snd (finCW↑ m n p C)) = snd C .fst --- -- -- -- snd (snd (finCW↑ m n p C)) k = --- -- -- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) --- -- -- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) --- -- -- -- (snd C .snd (k +ℕ fst p)) +{- + α↑ : A (suc (suc (suc n))) × S₊ (suc (suc n)) → SphereBouquet (suc n) (A (suc n)) + α↑ = {!α (suc (suc (suc n)))!} +-} + opaque + iso3 : Iso (C* (suc (suc (suc n)))) + (Pushout α' fst) + iso3 = compIso (equivToIso (e (suc (suc n)))) + (pushoutIso _ _ _ _ (idEquiv _) + (isoToEquiv iso2) + (idEquiv _) + refl + refl) + + + module _ (α-pt : (x : _) → α' (x , ptSn (suc n)) ≡ inl tt) + where + α∙ : SphereBouquet∙ (suc n) (A (suc (suc n))) + →∙ SphereBouquet∙ (suc n) (A (suc n)) + fst α∙ (inl x) = inl tt + fst α∙ (inr x) = α' x + fst α∙ (push a i) = α-pt a (~ i) + snd α∙ = refl + + opaque + iso4 : Iso (C* (suc (suc (suc n)))) + (cofib (fst α∙)) + iso4 = compIso iso3 (⋁-cofib-Iso _ α∙) + + module _ (α-pt2 : (x : A (3 +ℕ n)) + → Iso.fun iso4 (α (3 +ℕ n) (x , north)) ≡ inl tt) where + α↑ : SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n)))) + →∙ (cofib (fst α∙) , inl tt) + fst α↑ (inl x) = inl tt + fst α↑ (inr x) = Iso.fun iso4 (α (3 +ℕ n) x) + fst α↑ (push a i) = α-pt2 a (~ i) + snd α↑ = refl + + iso5 : Iso (C* (4 +ℕ n)) (cofib (fst α↑)) + iso5 = compIso (equivToIso (e (3 +ℕ n))) + (compIso + (pushoutIso _ _ _ _ + (idEquiv _) + (isoToEquiv iso4) + (idEquiv _) refl refl) + (⋁-cofib-Iso _ α↑)) + + open commSquare + open 3-span + + PT : commSquare + A0 (sp PT) = cofib (fst α∙) + A2 (sp PT) = SphereBouquet (suc (suc n)) (A (suc (suc (suc n)))) + A4 (sp PT) = Unit + f1 (sp PT) = fst α↑ + f3 (sp PT) = λ _ → tt + P PT = cofib (fst α↑) + inlP PT = inr + inrP PT = inl + comm PT = funExt λ x → sym (push x) + + PTPush : PushoutSquare + fst PTPush = PT + snd PTPush = + subst isEquiv (funExt + (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl})) + (symPushout _ _ .snd) + + opaque + iso6 : Iso (Pushout {B = cofib (fst α∙)} inr (λ _ → tt)) + (SphereBouquet (suc (suc n)) (A (suc (suc n)))) + iso6 = compIso (equivToIso (symPushout _ _)) + (compIso (Iso-cofibInr-Susp α∙) + sphereBouquetSuspIso) + + PL : commSquare + A0 (sp PL) = cofib (fst α∙) + A2 (sp PL) = SphereBouquet (suc n) (A (suc n)) + A4 (sp PL) = Unit + f1 (sp PL) = inr + f3 (sp PL) = _ + P PL = SphereBouquet (suc (suc n)) (A (suc (suc n))) + inlP PL x = Iso.fun iso6 (inl x) + inrP PL _ = Iso.fun iso6 (inr tt) -- + comm PL = funExt λ x → cong (Iso.fun iso6) (push x) + + PLPush : PushoutSquare + fst PLPush = PL + snd PLPush = subst isEquiv (funExt coh) (isoToIsEquiv iso6) + where + coh : (x : _) → Iso.fun iso6 x ≡ Pushout→commSquare PL x + coh (inl x) = refl + coh (inr x) = refl + coh (push x i₁) = refl + + C∨PlaceHolder = Pushout (Iso.inv iso5 ∘ inr) λ x → Iso.fun iso6 (inl x) + + C∨ = (C* (suc (suc (suc (suc n)))) , Iso.inv iso5 (inl tt)) + ⋁ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))) + + T* : {!!} + T* = {!!} + + module L→R = PushoutPasteDown PLPush {B = C∨PlaceHolder} (Iso.inv iso5 ∘ inr) inl inr (funExt push) + + isPushoutTot = L→R.isPushoutBottomSquare→isPushoutTotSquare + (subst isEquiv (funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl})) + (idIsEquiv _)) + + + C4n = C* (suc (suc (suc (suc n)))) + + const* : SphereBouquet (suc n) (A (suc n)) → C* (suc (suc (suc (suc n)))) + const* = λ x → Iso.inv iso5 (inr (inr x)) + + connC* : isConnected (3 +ℕ n) (C* (suc (suc (suc (suc n))))) + connC* = isConnectedCW→isConnectedSkel + (_ , ind .fst .snd .fst) (suc (suc (suc (suc n)))) + (suc (suc (suc n)) , <ᵗ-trans <ᵗsucm <ᵗsucm) + (subst (isConnected (3 +ℕ n)) (ua (invEquiv (ind .snd))) + cA) + + C⋆ = iso5 .Iso.inv (inl tt) + + sphereVanish : (f : S₊ (suc n) → C* (suc (suc (suc (suc n))))) + → ∥ ((x : S₊ (suc n)) → f x ≡ C⋆) ∥₁ + sphereVanish f = sphereToTrunc (suc n) + λ x → isConnectedPath (suc (suc n)) connC* _ _ .fst + + sphereBouquetVanish : ∀ {k : ℕ} (f : SphereBouquet (suc n) (Fin k) → C4n) + → ∥ f ≡ (λ _ → C⋆) ∥₁ + sphereBouquetVanish {k = k} f = + TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) + (λ p → PT.rec squash₁ + (λ q → ∣ funExt + (λ { (inl x) → p + ; (inr (x , y)) → (q x y ∙ sym (q x (ptSn (suc n)))) + ∙ cong f (sym (push x)) ∙ p + ; (push a i) j → (cong (_∙ cong f (sym (push a)) ∙ p) + (rCancel (q a (ptSn (suc n)))) + ∙ sym (lUnit _) + ◁ (symP (compPath-filler' + (cong f (sym (push a))) p))) (~ i) j}) ∣₁) + help) + isPted + where + isPted = Iso.fun (PathIdTruncIso (suc (suc n))) + (isContr→isProp connC* ∣ f (inl tt) ∣ₕ ∣ C⋆ ∣ₕ) + + help : ∥ ((x : Fin k) → (y : _) → f (inr (x , y)) ≡ C⋆) ∥₁ + help = invEq propTrunc≃Trunc1 + (invEq (_ , InductiveFinSatAC _ _ _) + λ x → PT.rec (isOfHLevelTrunc 1) + ∣_∣ₕ + (sphereVanish λ y → f (inr (x , y)))) + + pushoutTotAlt : C∨PlaceHolder ≃ cofib const* + pushoutTotAlt = + compEquiv + (invEquiv (_ , isPushoutTot)) + (symPushout _ _) + + β : {!!} + β = {!!} + + module _ (vanish : const* ≡ λ _ → C⋆) where + S∨C = SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4n , C⋆) + + Iso-C∨PlaceHolder-Wedge : + C∨PlaceHolder ≃ (SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4n , C⋆)) + Iso-C∨PlaceHolder-Wedge = + compEquiv pushoutTotAlt + (isoToEquiv (compIso (cofibConst' {A = _ , inl tt} {B = _ , C⋆} + (const* , funExt⁻ vanish _) + (funExt⁻ vanish)) + (pushoutIso _ _ _ _ + (idEquiv _) + (isoToEquiv sphereBouquetSuspIso) + (idEquiv _) + refl + refl))) + + + + + module T→B = PushoutPasteDown PTPush {B = C∨PlaceHolder} + (λ x → Iso.fun iso6 (inl x)) inr + (λ x → inl (Iso.inv iso5 x)) + (sym (funExt push)) + + helpIso : Iso (spanPushout (sp T→B.bottomSquare)) (P T→B.bottomSquare) + helpIso = compIso {!!} {!!} + + ⋁-as-cofib : cofib {!!} ≃ S∨C + ⋁-as-cofib = + compEquiv (compEquiv {!!} + (_ , T→B.isPushoutBottomSquare→isPushoutTotSquare + {!Pushout→commSquare T→B.bottomSquare!})) + Iso-C∨PlaceHolder-Wedge + + + + {- + snd PL = subst isEquiv {!!} (isoToIsEquiv help) + where + help : Iso _ _ + help = compIso {!!} {!!} + + main : isConnectedCW (suc n) C + main = {!!} +-} +{-Iso.fun help ≡ Pushout→commSquare PL +yieldsConnectedCWskel : ? +yieldsConnectedCWskel = ? +-} diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index a4a20a9f7c..c5e1bcf65b 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -622,3 +622,7 @@ open Lift public liftExt : ∀ {A : Type ℓ} {a b : Lift {ℓ} {ℓ'} A} → (lower a ≡ lower b) → a ≡ b liftExt x i = lift (x i) + +liftFun : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} + (f : A → B) → Lift {j = ℓ''} A → Lift {j = ℓ'''} B +liftFun f (lift a) = lift (f a) diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index 7e5fb52959..fc8a6cf9ba 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -22,12 +22,14 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Path open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence open import Cubical.Foundations.Transport open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Relation.Nullary @@ -502,7 +504,7 @@ private ∙∙ λ j → inr (refl-lem (g₁ a) i j))) ∙∙ sym (rUnit (push a)) - +-- induced isomorphism of pushouts (see later def. pushoutIso/pushoutEquiv for a more universe polymorphic) module _ {ℓ : Level} {A₁ B₁ C₁ A₂ B₂ C₂ : Type ℓ} (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) @@ -514,15 +516,93 @@ module _ {ℓ : Level} {A₁ B₁ C₁ A₂ B₂ C₂ : Type ℓ} module P = PushoutIso A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 l-r-cancel = F-G-cancel A≃ B≃ C≃ f₁ g₁ f₂ g₂ id1 id2 + pushoutIso' : Iso (Pushout f₁ g₁) (Pushout f₂ g₂) + fun pushoutIso' = P.F + inv pushoutIso' = P.G + rightInv pushoutIso' = fst l-r-cancel + leftInv pushoutIso' = snd l-r-cancel + + pushoutEquiv' : Pushout f₁ g₁ ≃ Pushout f₂ g₂ + pushoutEquiv' = isoToEquiv pushoutIso' + +-- lift induces an equivalence on Pushouts +module _ {ℓ ℓ' ℓ'' : Level} (ℓ''' : Level) + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (f : A → B) (g : A → C) where + PushoutLevel : Level + PushoutLevel = ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')) + + PushoutLift : Type PushoutLevel + PushoutLift = Pushout {A = Lift {j = ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')} A} + {B = Lift {j = ℓ-max ℓ (ℓ-max ℓ'' ℓ''')} B} + {C = Lift {j = ℓ-max ℓ (ℓ-max ℓ' ℓ''')} C} + (liftFun f) + (liftFun g) + + PushoutLiftIso : Iso (Pushout f g) PushoutLift + Iso.fun PushoutLiftIso (inl x) = inl (lift x) + Iso.fun PushoutLiftIso (inr x) = inr (lift x) + Iso.fun PushoutLiftIso (push a i) = push (lift a) i + Iso.inv PushoutLiftIso (inl (lift x)) = inl x + Iso.inv PushoutLiftIso (inr (lift x)) = inr x + Iso.inv PushoutLiftIso (push (lift a) i) = push a i + Iso.rightInv PushoutLiftIso (inl (lift x)) = refl + Iso.rightInv PushoutLiftIso (inr (lift x)) = refl + Iso.rightInv PushoutLiftIso (push (lift a) i) = refl + Iso.leftInv PushoutLiftIso (inl x) = refl + Iso.leftInv PushoutLiftIso (inr x) = refl + Iso.leftInv PushoutLiftIso (push a i) = refl + +-- equivalence of pushouts with arbitrary univeses +module _ {ℓA₁ ℓB₁ ℓC₁ ℓA₂ ℓB₂ ℓC₂} + {A₁ : Type ℓA₁} {B₁ : Type ℓB₁} {C₁ : Type ℓC₁} + {A₂ : Type ℓA₂} {B₂ : Type ℓB₂} {C₂ : Type ℓC₂} + (f₁ : A₁ → B₁) (g₁ : A₁ → C₁) + (f₂ : A₂ → B₂) (g₂ : A₂ → C₂) + (A≃ : A₁ ≃ A₂) (B≃ : B₁ ≃ B₂) (C≃ : C₁ ≃ C₂) + (id1 : fst B≃ ∘ f₁ ≡ f₂ ∘ fst A≃) + (id2 : fst C≃ ∘ g₁ ≡ g₂ ∘ fst A≃) where + private + ℓ* = ℓ-max ℓA₁ (ℓ-max ℓA₂ (ℓ-max ℓB₁ (ℓ-max ℓB₂ (ℓ-max ℓC₁ ℓC₂)))) + pushoutIso : Iso (Pushout f₁ g₁) (Pushout f₂ g₂) - fun pushoutIso = P.F - inv pushoutIso = P.G - rightInv pushoutIso = fst l-r-cancel - leftInv pushoutIso = snd l-r-cancel + pushoutIso = + compIso (PushoutLiftIso ℓ* f₁ g₁) + (compIso (pushoutIso' _ _ _ _ + (Lift≃Lift A≃) + (Lift≃Lift B≃) + (Lift≃Lift C≃) + (funExt (λ { (lift x) → cong lift (funExt⁻ id1 x)})) + (funExt (λ { (lift x) → cong lift (funExt⁻ id2 x)}))) + (invIso (PushoutLiftIso ℓ* f₂ g₂))) pushoutEquiv : Pushout f₁ g₁ ≃ Pushout f₂ g₂ pushoutEquiv = isoToEquiv pushoutIso +module _ {C : Type ℓ} {B : Type ℓ'} where + PushoutAlongEquiv→ : {A : Type ℓ} + (e : A ≃ C) (f : A → B) → Pushout (fst e) f → B + PushoutAlongEquiv→ e f (inl x) = f (invEq e x) + PushoutAlongEquiv→ e f (inr x) = x + PushoutAlongEquiv→ e f (push a i) = f (retEq e a i) + + private + PushoutAlongEquiv→Cancel : {A : Type ℓ} (e : A ≃ C) (f : A → B) + → retract (PushoutAlongEquiv→ e f) inr + PushoutAlongEquiv→Cancel = + EquivJ (λ A e → (f : A → B) + → retract (PushoutAlongEquiv→ e f) inr) + λ f → λ { (inl x) → sym (push x) + ; (inr x) → refl + ; (push a i) j → push a (~ j ∨ i)} + + PushoutAlongEquiv : {A : Type ℓ} (e : A ≃ C) (f : A → B) + → Iso (Pushout (fst e) f) B + Iso.fun (PushoutAlongEquiv e f) = PushoutAlongEquiv→ e f + Iso.inv (PushoutAlongEquiv e f) = inr + Iso.rightInv (PushoutAlongEquiv e f) x = refl + Iso.leftInv (PushoutAlongEquiv e f) = PushoutAlongEquiv→Cancel e f + + module PushoutDistr {ℓ ℓ' ℓ'' ℓ''' : Level} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} (f : B → A) (g : C → B) (h : C → D) where @@ -568,3 +648,315 @@ rightInv (PushoutEmptyFam {A = A} {B = B} ¬A ¬C {f = f} {g = g}) (push a i) j (push a) (λ _ → inl (f a)) (rec (¬C (g a)))} (¬A a) j i leftInv (PushoutEmptyFam {A = A} {B = B} ¬A ¬C) x = refl + + +module _ {ℓA ℓB ℓC ℓD : Level} + {A : Type ℓA} {B : Type ℓB} + {C : Type ℓC} {D : Type ℓD} + (f : A → B) (g : A → C) + (inl* : B → D) + (inr* : C → D) + (com : inl* ∘ f ≡ inr* ∘ g) + where + Pushout→AbstractPushout : Pushout f g → D + Pushout→AbstractPushout (inl x) = inl* x + Pushout→AbstractPushout (inr x) = inr* x + Pushout→AbstractPushout (push a i) = com i a + +-- Commutative squares and pushout squares +module _ {ℓ₀ ℓ₂ ℓ₄ ℓP : Level} where + private + ℓ* = ℓ-maxList (ℓ₀ ∷ ℓ₂ ∷ ℓ₄ ∷ ℓP ∷ []) + + record commSquare : Type (ℓ-suc ℓ*) where + open 3-span + field + sp : 3-span {ℓ₀} {ℓ₂} {ℓ₄} + P : Type ℓP + inlP : A0 sp → P + inrP : A4 sp → P + comm : inlP ∘ f1 sp ≡ inrP ∘ f3 sp + + open commSquare + + Pushout→commSquare : (sk : commSquare) → spanPushout (sp sk) → P sk + Pushout→commSquare sk (inl x) = inlP sk x + Pushout→commSquare sk (inr x) = inrP sk x + Pushout→commSquare sk (push a i) = comm sk i a + + isPushoutSquare : commSquare → Type _ + isPushoutSquare sk = isEquiv (Pushout→commSquare sk) + + PushoutSquare : Type (ℓ-suc ℓ*) + PushoutSquare = Σ commSquare isPushoutSquare + +-- Rotations of commutative squares and pushout squares +module _ {ℓ₀ ℓ₂ ℓ₄ ℓP : Level} where + open commSquare + open 3-span + + rotateCommSquare : commSquare {ℓ₀} {ℓ₂} {ℓ₄} {ℓP} → commSquare {ℓ₄} {ℓ₂} {ℓ₀} {ℓP} + A0 (sp (rotateCommSquare sq)) = A4 (sp sq) + A2 (sp (rotateCommSquare sq)) = A2 (sp sq) + A4 (sp (rotateCommSquare sq)) = A0 (sp sq) + f1 (sp (rotateCommSquare sq)) = f3 (sp sq) + f3 (sp (rotateCommSquare sq)) = f1 (sp sq) + P (rotateCommSquare sq) = P sq + inlP (rotateCommSquare sq) = inrP sq + inrP (rotateCommSquare sq) = inlP sq + comm (rotateCommSquare sq) = sym (comm sq) + + rotatePushoutSquare : PushoutSquare {ℓ₀} {ℓ₂} {ℓ₄} {ℓP} + → PushoutSquare {ℓ₄} {ℓ₂} {ℓ₀} {ℓP} + fst (rotatePushoutSquare (sq , eq)) = rotateCommSquare sq + snd (rotatePushoutSquare (sq , eq)) = + subst isEquiv (funExt lem) (compEquiv (symPushout _ _) (_ , eq) .snd) + where + lem : (x : _) → Pushout→commSquare sq (symPushout _ _ .fst x) + ≡ Pushout→commSquare (rotateCommSquare sq) x + lem (inl x) = refl + lem (inr x) = refl + lem (push a i) = refl + + +-- Pushout pasting lemma: +{- Given a diagram consisting of two +commuting squares where the top square is a pushout: + +A2 -—f3-→ A4 + ∣ ∣ +f1 inrP + ↓ ⌜ ↓ +A0 -—inlP→ P + | | + f h + ↓ ↓ + A -—-g-—→ B + +The bottom square is a pushout square iff the outer rectangle is +a pushout square. +-} + +module PushoutPasteDown {ℓ₀ ℓ₂ ℓ₄ ℓP ℓA ℓB : Level} + (SQ : PushoutSquare {ℓ₀} {ℓ₂} {ℓ₄} {ℓP}) + {A : Type ℓA} {B : Type ℓB} + (f : 3-span.A0 (commSquare.sp (fst SQ)) → A) + (g : A → B) (h : commSquare.P (fst SQ) → B) + (com : g ∘ f ≡ h ∘ commSquare.inlP (fst SQ)) + where + private + sq = fst SQ + isP = snd SQ + ℓ* = ℓ-maxList (ℓ₀ ∷ ℓ₂ ∷ ℓ₄ ∷ ℓP ∷ []) + + open commSquare sq + open 3-span sp + + bottomSquare : commSquare + 3-span.A0 (commSquare.sp bottomSquare) = A + 3-span.A2 (commSquare.sp bottomSquare) = A0 + 3-span.A4 (commSquare.sp bottomSquare) = P + 3-span.f1 (commSquare.sp bottomSquare) = f + 3-span.f3 (commSquare.sp bottomSquare) = inlP + commSquare.P bottomSquare = B + commSquare.inlP bottomSquare = g + commSquare.inrP bottomSquare = h + commSquare.comm bottomSquare = com + + totSquare : commSquare + 3-span.A0 (commSquare.sp totSquare) = A + 3-span.A2 (commSquare.sp totSquare) = A2 + 3-span.A4 (commSquare.sp totSquare) = A4 + 3-span.f1 (commSquare.sp totSquare) = f ∘ f1 + 3-span.f3 (commSquare.sp totSquare) = f3 + commSquare.P totSquare = B + commSquare.inlP totSquare = g + commSquare.inrP totSquare = h ∘ inrP + commSquare.comm totSquare = + funExt λ x → funExt⁻ com (f1 x) ∙ cong h (funExt⁻ comm x) + + private + P' : Type _ + P' = Pushout f1 f3 + + Iso-P'-P : P' ≃ P + Iso-P'-P = _ , isP + + P'≃P = equiv→HAEquiv Iso-P'-P + + B'bot = Pushout {C = P'} f inl + + B'bot→BBot : B'bot → Pushout {C = P} f inlP + B'bot→BBot (inl x) = inl x + B'bot→BBot (inr x) = inr (fst Iso-P'-P x) + B'bot→BBot (push a i) = push a i + + Bbot→B'bot : Pushout {C = P} f inlP → B'bot + Bbot→B'bot (inl x) = inl x + Bbot→B'bot (inr x) = inr (invEq Iso-P'-P x) + Bbot→B'bot (push a i) = + (push a ∙ λ i → inr (isHAEquiv.linv (P'≃P .snd) (inl a) (~ i))) i + + Iso-B'bot-Bbot : Iso B'bot (Pushout {C = P} f inlP) + fun Iso-B'bot-Bbot = B'bot→BBot + inv Iso-B'bot-Bbot = Bbot→B'bot + rightInv Iso-B'bot-Bbot (inl x) = refl + rightInv Iso-B'bot-Bbot (inr x) i = inr (isHAEquiv.rinv (P'≃P .snd) x i) + rightInv Iso-B'bot-Bbot (push a i) j = help j i + where + help : Square + (cong B'bot→BBot + (push a ∙ λ i → inr (isHAEquiv.linv (P'≃P .snd) (inl a) (~ i)))) + (push a) + refl + λ i → inr (isHAEquiv.rinv (P'≃P .snd) (inlP a) i) + help = flipSquare ((λ i j → B'bot→BBot (compPath-filler (push a) + (λ i → inr (isHAEquiv.linv (P'≃P .snd) (inl a) (~ i))) (~ j) i)) + ▷ λ j i → inr (isHAEquiv.com (P'≃P .snd) (inl a) j i)) + leftInv Iso-B'bot-Bbot (inl x) = refl + leftInv Iso-B'bot-Bbot (inr x) i = inr (isHAEquiv.linv (P'≃P .snd) x i) + leftInv Iso-B'bot-Bbot (push a i) j = + compPath-filler (push a) + (λ i → inr (isHAEquiv.linv (P'≃P .snd) (inl a) (~ i))) (~ j) i + + B'tot : Type _ + B'tot = Pushout (f ∘ f1) f3 + + B'bot→B'tot : B'bot → B'tot + B'bot→B'tot (inl x) = inl x + B'bot→B'tot (inr (inl x)) = inl (f x) + B'bot→B'tot (inr (inr x)) = inr x + B'bot→B'tot (inr (push a i)) = push a i + B'bot→B'tot (push a i) = inl (f a) + + B'tot→B'bot : B'tot → B'bot + B'tot→B'bot (inl x) = inl x + B'tot→B'bot (inr x) = inr (inr x) + B'tot→B'bot (push a i) = (push (f1 a) ∙ λ i → inr (push a i)) i + + Iso-B'bot→B'tot : Iso B'bot B'tot + Iso.fun Iso-B'bot→B'tot = B'bot→B'tot + Iso.inv Iso-B'bot→B'tot = B'tot→B'bot + Iso.rightInv Iso-B'bot→B'tot (inl x) = refl + Iso.rightInv Iso-B'bot→B'tot (inr x) = refl + Iso.rightInv Iso-B'bot→B'tot (push a i) j = + (cong-∙ B'bot→B'tot (push (f1 a)) (λ i → inr (push a i)) + ∙ sym (lUnit _)) j i + Iso.leftInv Iso-B'bot→B'tot (inl x) = refl + Iso.leftInv Iso-B'bot→B'tot (inr (inl x)) = push x + Iso.leftInv Iso-B'bot→B'tot (inr (inr x)) = refl + Iso.leftInv Iso-B'bot→B'tot (inr (push a i)) j = + compPath-filler' (push (f1 a)) (λ i → inr (push a i)) (~ j) i + Iso.leftInv Iso-B'bot→B'tot (push a i) j = push a (i ∧ j) + + main' : Iso (spanPushout (commSquare.sp bottomSquare)) B'tot + main' = compIso (invIso Iso-B'bot-Bbot) (Iso-B'bot→B'tot) + + mainInv∘ : (x : _) → Pushout→commSquare bottomSquare (main' .inv x) + ≡ Pushout→commSquare totSquare x + mainInv∘ (inl x) = refl + mainInv∘ (inr x) = refl + mainInv∘ (push a i) j = help j i + where + help : cong (Pushout→commSquare bottomSquare) + (cong (Iso.fun Iso-B'bot-Bbot) (push (f1 a) ∙ (λ i → inr (push a i)))) + ≡ funExt⁻ com (f1 a) ∙ cong h (funExt⁻ comm a) + help = cong (cong (Pushout→commSquare bottomSquare)) + (cong-∙ (Iso.fun Iso-B'bot-Bbot) (push (f1 a)) (λ i → inr (push a i))) + ∙ cong-∙ (Pushout→commSquare bottomSquare) + (push (3-span.f1 (commSquare.sp sq) a)) + (λ i → inr (commSquare.comm sq i a)) + + isPushoutBottomSquare→isPushoutTotSquare : + isPushoutSquare bottomSquare → isPushoutSquare totSquare + isPushoutBottomSquare→isPushoutTotSquare eq = + subst isEquiv (funExt mainInv∘) (isoToEquiv main .snd) + where + main : Iso (spanPushout (commSquare.sp totSquare)) B + main = compIso (invIso main') (equivToIso (_ , eq)) + + isPushoutTotSquare→isPushoutBottomSquare : + isPushoutSquare totSquare → isPushoutSquare bottomSquare + isPushoutTotSquare→isPushoutBottomSquare eq = + subst isEquiv (funExt coh) + (snd (isoToEquiv main)) + where + + main : Iso (spanPushout (commSquare.sp bottomSquare)) B + main = compIso + (compIso (invIso Iso-B'bot-Bbot) (Iso-B'bot→B'tot)) + (equivToIso (_ , eq)) + + coh : (x : spanPushout (commSquare.sp bottomSquare)) → + main .fun x ≡ Pushout→commSquare bottomSquare x + coh x = sym (secEq (_ , eq) (fun main x)) + ∙∙ sym (mainInv∘ (invEq (_ , eq) (Iso.fun main x))) + ∙∙ cong (Pushout→commSquare bottomSquare) (Iso.leftInv main x) + +-- Pushout pasting lemma, alternative version: +{- Given a diagram consisting of two +commuting squares where the left square is a pushout: + +A2 -—f3-→ A4 -—-f--→ A + ∣ ∣ ∣ +f1 inrP g + ↓ ⌜ ↓ ↓ +A0 -—inlP→ P -—-h--→ B + +The right square is a pushout square iff the outer rectangle is +a pushout square. +-} +module PushoutPasteLeft {ℓ₀ ℓ₂ ℓ₄ ℓP ℓA ℓB : Level} + (SQ : PushoutSquare {ℓ₀} {ℓ₂} {ℓ₄} {ℓP}) + {A : Type ℓA} {B : Type ℓB} + (f : 3-span.A4 (commSquare.sp (fst SQ)) → A) + (g : A → B) (h : commSquare.P (fst SQ) → B) + (com : h ∘ commSquare.inrP (fst SQ) ≡ g ∘ f) + where + + private + sq = fst SQ + isP = snd SQ + ℓ* = ℓ-maxList (ℓ₀ ∷ ℓ₂ ∷ ℓ₄ ∷ ℓP ∷ []) + + open commSquare sq + open 3-span sp + + rightSquare : commSquare + 3-span.A0 (commSquare.sp rightSquare) = P + 3-span.A2 (commSquare.sp rightSquare) = A4 + 3-span.A4 (commSquare.sp rightSquare) = A + 3-span.f1 (commSquare.sp rightSquare) = inrP + 3-span.f3 (commSquare.sp rightSquare) = f + commSquare.P rightSquare = B + commSquare.inlP rightSquare = h + commSquare.inrP rightSquare = g + commSquare.comm rightSquare = com + + totSquare : commSquare + 3-span.A0 (commSquare.sp totSquare) = A0 + 3-span.A2 (commSquare.sp totSquare) = A2 + 3-span.A4 (commSquare.sp totSquare) = A + 3-span.f1 (commSquare.sp totSquare) = f1 + 3-span.f3 (commSquare.sp totSquare) = f ∘ f3 + commSquare.P totSquare = B + commSquare.inlP totSquare = h ∘ inlP + commSquare.inrP totSquare = g + commSquare.comm totSquare = funExt λ x → + sym (sym (funExt⁻ com (f3 x)) ∙ cong h (sym (funExt⁻ comm x))) + + + module M = PushoutPasteDown (rotatePushoutSquare SQ) f g h (sym com) + + isPushoutRightSquare→isPushoutTotSquare : + isPushoutSquare rightSquare → isPushoutSquare totSquare + isPushoutRightSquare→isPushoutTotSquare e = rotatePushoutSquare (_ , help) .snd + where + help : isPushoutSquare M.totSquare + help = M.isPushoutBottomSquare→isPushoutTotSquare (rotatePushoutSquare (_ , e) .snd) + + isPushoutTotSquare→isPushoutRightSquare : + isPushoutSquare totSquare → isPushoutSquare rightSquare + isPushoutTotSquare→isPushoutRightSquare e = rotatePushoutSquare (_ , help) .snd + where + help = M.isPushoutTotSquare→isPushoutBottomSquare (rotatePushoutSquare (_ , e) .snd) From e4f400342907ec03159b741b3a43900777e1b819 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 31 May 2024 04:09:48 +0200 Subject: [PATCH 45/73] almost --- Cubical/CW/Pointed2.agda | 352 ++++++++++++++++++--- Cubical/HITs/Pushout/Properties.agda | 41 ++- Cubical/HITs/SphereBouquet/Properties.agda | 22 ++ Cubical/HITs/Wedge/Properties.agda | 226 ++++++++++++- Cubical/Homotopy/Connected.agda | 6 + 5 files changed, 593 insertions(+), 54 deletions(-) diff --git a/Cubical/CW/Pointed2.agda b/Cubical/CW/Pointed2.agda index 027143074d..6ef19484fd 100644 --- a/Cubical/CW/Pointed2.agda +++ b/Cubical/CW/Pointed2.agda @@ -607,18 +607,38 @@ PushoutPostcompEquivIso {ℓ = ℓ} {ℓ'} {ℓ''} {ℓ'''} where ℓ* = ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')) -Pushout∘Equiv : ∀ {ℓ ℓ' ℓ''} {A A' : Type ℓ} {B B' : Type ℓ'} {C : Type ℓ''} +private + Pushout∘Equiv' : ∀ {ℓ ℓ' ℓ''} {A A' : Type ℓ} {B B' : Type ℓ'} {C : Type ℓ''} + (e1 : A ≃ A') (e2 : B' ≃ B) + (f : A' → B') (g : A → C) + → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) (Pushout f (g ∘ invEq e1)) + Pushout∘Equiv' {A = A} {A' = A'} {B} {B'} {C} = + EquivJ (λ A e1 → (e2 : B' ≃ B) (f : A' → B') (g : A → C) + → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) + (Pushout f (g ∘ invEq e1))) + (EquivJ (λ B' e2 → (f : A' → B') (g : A' → C) + → Iso (Pushout (fst e2 ∘ f) g) + (Pushout f g)) + λ f g → idIso) +open import Cubical.Data.List +Pushout∘Equiv : ∀ {ℓA ℓA' ℓB ℓB' ℓC} + {A : Type ℓA} {A' : Type ℓA'} {B : Type ℓB} {B' : Type ℓB'} + {C : Type ℓC} (e1 : A ≃ A') (e2 : B' ≃ B) (f : A' → B') (g : A → C) → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) (Pushout f (g ∘ invEq e1)) -Pushout∘Equiv {A = A} {A' = A'} {B} {B'} {C} = - EquivJ (λ A e1 → (e2 : B' ≃ B) (f : A' → B') (g : A → C) - → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) - (Pushout f (g ∘ invEq e1))) - (EquivJ (λ B' e2 → (f : A' → B') (g : A' → C) - → Iso (Pushout (fst e2 ∘ f) g) - (Pushout f g)) - λ f g → idIso) +Pushout∘Equiv {ℓA = ℓA} {ℓA'} {ℓB} {ℓB'} {ℓC} e1 e2 f g = + compIso (pushoutIso _ _ _ _ LiftEquiv LiftEquiv LiftEquiv refl refl) + (compIso (Pushout∘Equiv' {ℓ = ℓ*} {ℓ*} {ℓ*} + (liftEq ℓ* e1) (liftEq ℓ* e2) (liftFun f) (liftFun g)) + (invIso (pushoutIso _ _ _ _ + LiftEquiv LiftEquiv (LiftEquiv {ℓ' = ℓ*}) refl refl))) + where + ℓ* = ℓ-maxList (ℓA ∷ ℓA' ∷ ℓB ∷ ℓB' ∷ ℓC ∷ []) + + liftEq : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ℓ* : Level) + → A ≃ B → Lift {j = ℓ*} A ≃ Lift {j = ℓ*} B + liftEq ℓ* e = compEquiv (invEquiv LiftEquiv) (compEquiv e LiftEquiv) module _ {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A × Bool → Unit ⊎ B) (b₀ : B) where @@ -1679,24 +1699,25 @@ Iso.leftInv (SphereBouquet₀Iso n) (inr (x , false)) = refl Iso.leftInv (SphereBouquet₀Iso n) (inr (x , true)) = push x Iso.leftInv (SphereBouquet₀Iso n) (push a i) j = push a (i ∧ j) -isConnectedCW→ContrIso : ∀ {ℓ} (n : ℕ) - (sk : connectedCWskel ℓ n) - → Iso (Pushout (snd sk .fst .snd .fst n) fst) - (SphereBouquet n (CWskel-fields.A (_ , snd sk .fst) n)) -isConnectedCW→ContrIso zero sk = - compIso (compIso (equivToIso (symPushout _ _)) - (invIso (PushoutEmptyFam - (λ()) (sk .snd .fst .snd .snd .fst)))) - (compIso (isContr→Iso (subst isContr - (cong Fin (sym (snd sk .snd .fst))) (flast , isPropFin1 _)) - (flast , isPropFin1 _)) - (compIso (invIso (SphereBouquet₀Iso 0)) - {!!})) -isConnectedCW→ContrIso (suc n) sk = - compIso ((invIso (PushoutEmptyFam - {!λ !} - λ p → ¬Fin0 (subst Fin (snd sk .snd .snd n <ᵗsucm) p)))) -- (λ()) (sk .snd .fst .snd .snd .fst)))) - {!!} +-- isConnectedCW→ContrIso : ∀ {ℓ} (n : ℕ) +-- (sk : connectedCWskel ℓ n) +-- → Iso (Pushout (snd sk .fst .snd .fst n) fst) +-- (SphereBouquet n (CWskel-fields.A (_ , snd sk .fst) n)) +-- isConnectedCW→ContrIso zero sk = +-- compIso (compIso (equivToIso (symPushout _ _)) +-- (invIso (PushoutEmptyFam +-- (λ()) (sk .snd .fst .snd .snd .fst)))) +-- (compIso (isContr→Iso (subst isContr +-- (cong Fin (sym (snd sk .snd .fst))) (flast , isPropFin1 _)) +-- (flast , isPropFin1 _)) +-- (compIso (invIso (SphereBouquet₀Iso 0)) +-- {!!})) +-- isConnectedCW→ContrIso (suc n) sk = +-- compIso ((invIso (PushoutEmptyFam +-- {!λ !} +-- λ p → ¬Fin0 (subst Fin (snd sk .snd .snd n <ᵗsucm) p)))) -- (λ()) (sk .snd .fst .snd .snd .fst)))) +-- {!!} + connectedCWskelLift : ∀ {ℓ} (n : ℕ) {C : Type ℓ} → (sk : isConnCW (suc (suc n)) C) @@ -1706,7 +1727,7 @@ connectedCWskelLift zero {C = C} ((cwsk , e) , cA) = , compEquiv (invEquiv (collapse₁Equiv _ _ 0)) (invEquiv e) ∣₁ where M = connectedCWskel→ (cwsk .fst) (snd cwsk , subst (isConnected 2) (ua e) cA) -connectedCWskelLift (suc n) {C = C} ((cwsk , eqv) , cA) with +connectedCWskelLift {ℓ = ℓ} (suc n) {C = C} ((cwsk , eqv) , cA) with (connectedCWskelLift n ((cwsk , eqv) , (isConnectedSubtr (suc (suc n)) 1 cA))) ... | s = PT.map {!C*!} s where @@ -1741,10 +1762,6 @@ connectedCWskelLift (suc n) {C = C} ((cwsk , eqv) , cA) with α' : (A (suc (suc n))) × S₊ (suc n) → SphereBouquet (suc n) (A (suc n)) α' = Iso.fun iso2 ∘ α (suc (suc n)) -{- - α↑ : A (suc (suc (suc n))) × S₊ (suc (suc n)) → SphereBouquet (suc n) (A (suc n)) - α↑ = {!α (suc (suc (suc n)))!} --} opaque iso3 : Iso (C* (suc (suc (suc n)))) (Pushout α' fst) @@ -1818,6 +1835,9 @@ connectedCWskelLift (suc n) {C = C} ((cwsk , eqv) , cA) with (compIso (Iso-cofibInr-Susp α∙) sphereBouquetSuspIso) + iso6∙ : Iso.fun iso6 (inl (inl tt)) ≡ inl tt + iso6∙ = refl + PL : commSquare A0 (sp PL) = cofib (fst α∙) A2 (sp PL) = SphereBouquet (suc n) (A (suc n)) @@ -1843,9 +1863,6 @@ connectedCWskelLift (suc n) {C = C} ((cwsk , eqv) , cA) with C∨ = (C* (suc (suc (suc (suc n)))) , Iso.inv iso5 (inl tt)) ⋁ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))) - T* : {!!} - T* = {!!} - module L→R = PushoutPasteDown PLPush {B = C∨PlaceHolder} (Iso.inv iso5 ∘ inr) inl inr (funExt push) isPushoutTot = L→R.isPushoutBottomSquare→isPushoutTotSquare @@ -1907,12 +1924,60 @@ connectedCWskelLift (suc n) {C = C} ((cwsk , eqv) , cA) with (invEquiv (_ , isPushoutTot)) (symPushout _ _) - β : {!!} - β = {!!} + β : SphereBouquet (suc (suc n)) (A (suc (suc (suc n)))) + → SphereBouquet (suc (suc n)) (A (suc (suc n))) + β = (Iso.fun iso6 ∘ inl) ∘ fst α↑ + + β∙ : SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n)))) + →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))) + fst β∙ = β + snd β∙ = iso6∙ module _ (vanish : const* ≡ λ _ → C⋆) where S∨C = SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4n , C⋆) + connS∨C : isConnected (suc (suc (suc n))) S∨C + fst connS∨C = ∣ inl (inl tt) ∣ₕ + snd connS∨C = + TR.elim (λ _ → isOfHLevelPath (suc (suc (suc n))) + (isOfHLevelTrunc (suc (suc (suc n)))) _ _) + λ { (inl x) → inlEq x + ; (inr x) → SP ∙ inrEq x + ; (push tt i) j → + (compPath-filler (inlEq (inl tt)) (cong ∣_∣ₕ (push tt)) + ▷ (rUnit SP ∙ cong (SP ∙_) (sym inrEq∙))) i j} + where + inlEq : (x : _) + → Path (hLevelTrunc (suc (suc (suc n))) S∨C) + ∣ inl (inl tt) ∣ ∣ inl x ∣ + inlEq x = TR.rec (isOfHLevelTrunc (3 +ℕ n) _ _) + (λ p i → ∣ inl (p i) ∣ₕ) + (Iso.fun (PathIdTruncIso _) + (isContr→isProp + (isConnectedSphereBouquet' {n = suc n}) + ∣ inl tt ∣ ∣ x ∣)) + + G : (x : C4n) → ∥ C⋆ ≡ x ∥ suc (suc n) + G x = Iso.fun (PathIdTruncIso _) (isContr→isProp connC* ∣ C⋆ ∣ₕ ∣ x ∣ₕ) + + G∙ : G C⋆ ≡ ∣ refl ∣ₕ + G∙ = cong (Iso.fun (PathIdTruncIso _)) + (isProp→isSet (isContr→isProp connC*) _ _ + (isContr→isProp connC* ∣ C⋆ ∣ₕ ∣ C⋆ ∣ₕ) refl) + ∙ cong ∣_∣ₕ (transportRefl refl) + + inrEq : (x : C4n) + → Path (hLevelTrunc (suc (suc (suc n))) S∨C) + ∣ inr C⋆ ∣ₕ ∣ inr x ∣ₕ + inrEq x = TR.rec (isOfHLevelTrunc (suc (suc (suc n))) _ _) + (λ p i → ∣ inr (p i) ∣ₕ) (G x) + + inrEq∙ : inrEq C⋆ ≡ refl + inrEq∙ = cong (TR.rec (isOfHLevelTrunc (suc (suc (suc n))) _ _) + (λ p i → ∣ inr (p i) ∣ₕ)) G∙ + + SP = inlEq (inl tt) ∙ cong ∣_∣ₕ (push tt) + Iso-C∨PlaceHolder-Wedge : C∨PlaceHolder ≃ (SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4n , C⋆)) Iso-C∨PlaceHolder-Wedge = @@ -1927,24 +1992,223 @@ connectedCWskelLift (suc n) {C = C} ((cwsk , eqv) , cA) with refl refl))) - - - module T→B = PushoutPasteDown PTPush {B = C∨PlaceHolder} (λ x → Iso.fun iso6 (inl x)) inr (λ x → inl (Iso.inv iso5 x)) (sym (funExt push)) helpIso : Iso (spanPushout (sp T→B.bottomSquare)) (P T→B.bottomSquare) - helpIso = compIso {!!} {!!} + helpIso = compIso (equivToIso (symPushout _ _)) + (pushoutIso _ _ _ _ (idEquiv _) + (invEquiv (isoToEquiv iso5)) + (idEquiv _) refl refl) + + helpIsoCoh : (x : _) → Iso.fun helpIso x ≡ Pushout→commSquare T→B.bottomSquare x + helpIsoCoh (inl x) = refl + helpIsoCoh (inr x) = refl + helpIsoCoh (push a i) j = sym (rUnit (sym (push a))) j i - ⋁-as-cofib : cofib {!!} ≃ S∨C + ⋁-as-cofib : cofib β ≃ S∨C ⋁-as-cofib = - compEquiv (compEquiv {!!} + compEquiv (compEquiv (symPushout _ _) (_ , T→B.isPushoutBottomSquare→isPushoutTotSquare - {!Pushout→commSquare T→B.bottomSquare!})) + (subst isEquiv (funExt helpIsoCoh) (isoToIsEquiv helpIso)))) Iso-C∨PlaceHolder-Wedge + -- S∨C-CW-fam : ℕ → Type ℓ + -- S∨C-CW-fam zero = ⊥* + -- S∨C-CW-fam (suc m) with (m ≟ᵗ suc (suc n)) + -- ... | lt x = Unit* + -- ... | eq x = Lift (SphereBouquet (suc (suc n)) (A (suc (suc n)))) + -- ... | gt x = S∨C + + + -- asdd = {!!} + + -- cardS∨C : ℕ → ℕ -- n + 2 + -- cardS∨C zero = 1 + -- cardS∨C (suc m) with (m ≟ᵗ suc n) | (m ≟ᵗ suc (suc n)) + -- ... | lt x | r = 0 + -- ... | eq x | r = card (suc (suc n)) + -- ... | gt x | lt x₁ = 0 + -- ... | gt x | eq x₁ = card (suc (suc (suc n))) + -- ... | gt x | gt x₁ = 0 + + -- αS∨C-non-triv : (m : ℕ) (p : suc (suc n) ≡ m) + -- → A (suc (suc (suc n))) × S₊ m → Lift {j = ℓ} (SphereBouquet (suc (suc n)) (A (suc (suc n)))) + -- αS∨C-non-triv m p (x , y) = lift (β (inr (x , (subst S₊ (sym p) y)))) + + -- αS∨C : (n₁ : ℕ) → Fin (cardS∨C n₁) × S⁻ n₁ → S∨C-CW-fam n₁ + -- αS∨C (suc m) with (m ≟ᵗ suc n) | (m ≟ᵗ suc (suc n)) + -- ... | eq x₁ | lt x = λ _ → tt* + -- ... | eq x₁ | eq x = λ _ → lift (inl tt) + -- ... | eq x₁ | gt x = λ _ → inl (inl tt) + -- ... | gt x₁ | lt x = λ() + -- ... | gt x₁ | eq p = αS∨C-non-triv m (sym p) -- + -- ... | gt x₁ | gt x = λ() + + -- eS∨C : (n₁ : ℕ) → S∨C-CW-fam (suc n₁) ≃ Pushout (αS∨C n₁) fst + -- eS∨C zero = isContr→Equiv (tt* , (λ {tt* → refl})) + -- ((inr fzero) + -- , (λ { (inr (zero , tt)) → refl})) + -- eS∨C (suc m) with (m ≟ᵗ suc n) | (m ≟ᵗ suc (suc n)) + -- ... | lt x | lt x₁ = + -- isContr→Equiv (tt* , (λ {tt* → refl})) ((inl tt*) + -- , λ { (inl tt*) → refl}) + -- ... | lt x | eq x₁ = + -- ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (_<ᵗ suc n) x₁ x) <ᵗsucm)) + -- ... | lt x | gt x₁ = + -- ⊥.rec (¬m<ᵗm (<ᵗ-trans (<ᵗ-trans x₁ x) <ᵗsucm)) + -- ... | eq x | lt x₁ = {!!} + -- ... | eq x | eq x₁ = + -- ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc n) ((cong predℕ (sym x ∙ x₁))) <ᵗsucm)) + -- ... | eq x | gt x₁ = + -- ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (suc (suc n) <ᵗ_) x x₁) <ᵗsucm)) + -- ... | gt x | lt x₁ = ⊥.rec (¬squeeze (x , x₁)) + -- ... | gt x | eq x₁ = {!isoToEquiv (invIso (PushoutEmptyFam ? ?))!} + -- ... | gt x | gt x₁ = + -- isoToEquiv (PushoutEmptyFam (λ()) λ()) + + -- S∨C-CW : CWskel ℓ + -- fst S∨C-CW = S∨C-CW-fam + -- fst (snd S∨C-CW) = cardS∨C + -- fst (snd (snd S∨C-CW)) = αS∨C + -- fst (snd (snd (snd S∨C-CW))) () + -- snd (snd (snd (snd S∨C-CW))) = eS∨C + + -- Sbouq-CW-fam : {!!} + -- Sbouq-CW-fam = {!elim.isConnectedPrecompose!} + + -- SbouqSkel : CWskel ℓ + -- SbouqSkel = {!!} + + open import Cubical.Data.Unit + testβ : isConnectedFun (suc (suc (suc n))) {B = cofib β} inr + testβ = inrConnected (3 +ℕ n) _ _ + λ b → isOfHLevelRetractFromIso 0 (mapCompIso fiberUnitIso) + isConnectedSphereBouquet' + + testβ3 : isConnectedFun (suc (suc (suc n))) {B = S∨C} (⋁-as-cofib .fst ∘ inr) + testβ3 = isConnectedComp (⋁-as-cofib .fst) inr (3 +ℕ n) + (isEquiv→isConnected _ (⋁-as-cofib .snd) (3 +ℕ n)) testβ + + main-inr : ∥ ((x : A (suc n)) (y : S₊ (suc (suc n))) + → Σ[ t ∈ SphereBouquet (suc (suc n)) (A (suc (suc n))) ] + (⋁-as-cofib .fst (inr t) ≡ inl (inr (x , y)))) ∥₁ + main-inr = + invEq propTrunc≃Trunc1 + (invEq (_ , InductiveFinSatAC _ _ _) + λ x → fst propTrunc≃Trunc1 + (sphereToTrunc (suc (suc n)) + λ y → testβ3 (inl (inr (x , y))) .fst)) + + mainF : ∃[ F ∈ (SphereBouquet∙ (suc (suc n)) (A (suc n)) + →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n)))) ] + ((x : _) → Path S∨C (⋁-as-cofib .fst (inr (fst F x))) (inl x)) + mainF = TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) + (λ coh₁ → PT.rec squash₁ (λ F → TR.rec squash₁ + (λ h → TR.rec squash₁ (λ q → ∣ (F* F coh₁ h , refl) , + (coh F coh₁ h q) ∣₁) + (makeCoh F coh₁ h)) + (F∙ F coh₁)) + main-inr) + (isConnectedPath _ connS∨C + (⋁-as-cofib .fst (inr (inl tt))) (inl (inl tt)) .fst) + where + CON = (subst (λ x → isConnected x (SphereBouquet (suc (suc n)) + (A (suc (suc n))))) + (cong suc (+-comm 2 n)) + isConnectedSphereBouquet') + module _ (F : ((x : A (suc n)) (y : S₊ (suc (suc n))) + → Σ[ t ∈ SphereBouquet (suc (suc n)) (A (suc (suc n))) ] + (⋁-as-cofib .fst (inr t) ≡ inl (inr (x , y))))) + (coh₁ : Path S∨C (⋁-as-cofib .fst (inr (inl tt))) + (inl (inl tt))) where + F∙ : hLevelTrunc 1 ((x : Fin _) → F x north .fst ≡ inl tt) + F∙ = + invEq (_ , InductiveFinSatAC _ _ _) + λ a → isConnectedPath 1 (isConnectedSubtr 2 (suc n) CON) _ _ .fst + + module _ (h : (x : Fin _) → F x north .fst ≡ inl tt) where + F* : (SphereBouquet (suc (suc n)) (A (suc n)) + → SphereBouquet (suc (suc n)) (A (suc (suc n)))) + F* (inl x) = inl tt + F* (inr (x , y)) = F x y .fst + F* (push a i) = h a (~ i) + + cohTy : Type _ + cohTy = (a : A (suc n)) + → Square (λ i → ⋁-as-cofib .fst (inr (h a (~ i)))) + (λ i → inl (push a i)) + coh₁ (F a north .snd) + + makeCoh : hLevelTrunc 1 cohTy + makeCoh = invEq (_ , InductiveFinSatAC _ _ _) + λ a → isConnectedPathP 1 + (isConnectedSubtr' n 2 (isConnectedPath _ connS∨C _ _)) _ _ .fst + + module _ (coh₂ : cohTy) where + coh : (x : _) → Path S∨C (⋁-as-cofib .fst (inr (F* x))) (inl x) + coh (inl x) = coh₁ + coh (inr (x , y)) = F x y .snd + coh (push a i) j = coh₂ a j i + module _ (F : (SphereBouquet∙ (suc (suc n)) (A (suc n)) + →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))))) + (h : (x : _) → Path S∨C (⋁-as-cofib .fst (inr (fst F x))) (inl x)) + where + h' : (x : _) → inr (fst F x) ≡ invEq ⋁-as-cofib (inl x) + h' x = sym (retEq ⋁-as-cofib (inr (fst F x))) + ∙ cong (invEq ⋁-as-cofib) (h x) + + open import Cubical.Foundations.Transport + + N = card (suc n) +ℕ card (suc (suc (suc n))) + + iso7 : Iso (SphereBouquet (suc (suc n)) (Fin N)) + (SphereBouquet∙ (suc (suc n)) (A (suc n)) + ⋁ SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n))))) + iso7 = compIso + (pathToIso + ((λ i → ⋁gen (isoToPath + (Fin+ {n = card (suc n)} {m = card (suc (suc (suc n)))}) (~ i)) + (λ _ → S₊∙ (suc (suc n)))) + ∙ cong (⋁gen (A (suc n) ⊎ A (suc (suc (suc n))))) + (funExt (⊎.elim (λ _ → refl) (λ _ → refl))))) + (invIso ⋁gen⊎Iso) + + T : Iso (cofib (F ∨→ β∙)) C4n + T = compIso (cofib∨→compIsoᵣ F β∙) + (compIso + (pathToIso (cong cofib (funExt h'))) + (compIso + (equivToIso (symPushout _ _)) + (compIso + (Pushout∘Equiv + (idEquiv (SphereBouquet (suc (suc n)) (A (suc n)))) + (invEquiv ⋁-as-cofib) inl (λ _ → tt)) + (compIso (equivToIso (symPushout _ _)) cofibInl-⋁)))) + + T∙ : Σ[ c ∈ ℕ ] Σ[ α ∈ (SphereBouquet (suc (suc n)) (Fin c) + → SphereBouquet (suc (suc n)) (A (suc (suc n)))) ] + Iso (Pushout (α ∘ inr) fst) C4n + fst T∙ = N + fst (snd T∙) = F ∨→ β∙ ∘ Iso.fun iso7 + snd (snd T∙) = + compIso + (compIso (⋁-cofib-Iso (_ , fst F (inl tt)) ((F ∨→ β∙ ∘ Iso.fun iso7) , refl)) + (invIso (Pushout∘Equiv + (invEquiv (isoToEquiv iso7)) + (idEquiv Unit) _ _))) T + +-- ⋁gen⊎Iso + +{- + fst S∨C-CW = S∨C-CW-fam + fst (fst (snd S∨C-CW)) = {!S∨C-CW-fam!} + fst (snd (fst (snd S∨C-CW))) = {!!} + snd (snd (fst (snd S∨C-CW))) = {!!} + snd (snd S∨C-CW) = {!!} +-} {- diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index fc8a6cf9ba..671c714434 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -564,16 +564,39 @@ module _ {ℓA₁ ℓB₁ ℓC₁ ℓA₂ ℓB₂ ℓC₂} private ℓ* = ℓ-max ℓA₁ (ℓ-max ℓA₂ (ℓ-max ℓB₁ (ℓ-max ℓB₂ (ℓ-max ℓC₁ ℓC₂)))) + pushoutIso→ : Pushout f₁ g₁ → Pushout f₂ g₂ + pushoutIso→ (inl x) = inl (fst B≃ x) + pushoutIso→ (inr x) = inr (fst C≃ x) + pushoutIso→ (push a i) = + ((λ i → inl (id1 i a)) ∙∙ push (fst A≃ a) ∙∙ λ i → inr (id2 (~ i) a)) i + + pushoutIso* : Iso (Pushout f₁ g₁) (Pushout f₂ g₂) + pushoutIso* = + compIso (PushoutLiftIso ℓ* f₁ g₁) + (compIso (pushoutIso' _ _ _ _ + (Lift≃Lift A≃) + (Lift≃Lift B≃) + (Lift≃Lift C≃) + (funExt (λ { (lift x) → cong lift (funExt⁻ id1 x)})) + (funExt (λ { (lift x) → cong lift (funExt⁻ id2 x)}))) + (invIso (PushoutLiftIso ℓ* f₂ g₂))) + + pushoutIso→≡ : (x : _) → Iso.fun pushoutIso* x ≡ pushoutIso→ x + pushoutIso→≡ (inl x) = refl + pushoutIso→≡ (inr x) = refl + pushoutIso→≡ (push a i) j = + cong-∙∙ (Iso.inv (PushoutLiftIso ℓ* f₂ g₂)) + (λ i → inl (lift (id1 i a))) + (push (lift (fst A≃ a))) + (λ i → inr (lift (id2 (~ i) a))) j i + pushoutIso : Iso (Pushout f₁ g₁) (Pushout f₂ g₂) - pushoutIso = - compIso (PushoutLiftIso ℓ* f₁ g₁) - (compIso (pushoutIso' _ _ _ _ - (Lift≃Lift A≃) - (Lift≃Lift B≃) - (Lift≃Lift C≃) - (funExt (λ { (lift x) → cong lift (funExt⁻ id1 x)})) - (funExt (λ { (lift x) → cong lift (funExt⁻ id2 x)}))) - (invIso (PushoutLiftIso ℓ* f₂ g₂))) + fun pushoutIso = pushoutIso→ + inv pushoutIso = inv pushoutIso* + rightInv pushoutIso x = + sym (pushoutIso→≡ (inv pushoutIso* x)) ∙ rightInv pushoutIso* x + leftInv pushoutIso x = + cong (inv pushoutIso*) (sym (pushoutIso→≡ x)) ∙ leftInv pushoutIso* x pushoutEquiv : Pushout f₁ g₁ ≃ Pushout f₂ g₂ pushoutEquiv = isoToEquiv pushoutIso diff --git a/Cubical/HITs/SphereBouquet/Properties.agda b/Cubical/HITs/SphereBouquet/Properties.agda index 8c0284f731..283b0f0ea7 100644 --- a/Cubical/HITs/SphereBouquet/Properties.agda +++ b/Cubical/HITs/SphereBouquet/Properties.agda @@ -10,6 +10,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels open import Cubical.Data.Bool open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) @@ -22,9 +23,12 @@ open import Cubical.HITs.Sn open import Cubical.HITs.Pushout open import Cubical.HITs.Susp open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as TR open import Cubical.HITs.Wedge open import Cubical.HITs.SphereBouquet.Base +open import Cubical.Homotopy.Connected + private variable ℓ ℓ' : Level @@ -54,6 +58,24 @@ isConnectedSphereBouquet {n = n} {A} = (λ (a , s) → sphereToPropElim n {A = λ x → ∥ inr (a , x) ≡ inl tt ∥₁} (λ x → squash₁) ∣ sym (push a) ∣₁ s) +isConnectedSphereBouquet' : {n : ℕ} {A : Type ℓ} + → isConnected (suc (suc n)) (SphereBouquet (suc n) A) +fst (isConnectedSphereBouquet' {n = n}) = ∣ inl tt ∣ +snd (isConnectedSphereBouquet' {n = n} {A = A}) = + TR.elim (λ _ → isOfHLevelPath (suc (suc n)) + (isOfHLevelTrunc (suc (suc n))) _ _) (lem n) + where + lem : (n : ℕ) → (a : SphereBouquet (suc n) A) + → Path (hLevelTrunc (suc (suc n)) (SphereBouquet (suc n) A)) + ∣ inl tt ∣ ∣ a ∣ + lem n (inl x) = refl + lem n (inr (x , y)) = + sphereElim n {A = λ y → ∣ inl tt ∣ ≡ ∣ inr (x , y) ∣} + (λ _ → isOfHLevelTrunc (suc (suc n)) _ _) + (cong ∣_∣ₕ (push x)) y + lem zero (push a i) j = ∣ push a (i ∧ j) ∣ₕ + lem (suc n) (push a i) j = ∣ push a (i ∧ j) ∣ₕ + sphereBouquetSuspIso₀ : {A : Type ℓ} → Iso (⋁gen A (λ a → Susp∙ (fst (S₊∙ zero)))) (SphereBouquet 1 A) diff --git a/Cubical/HITs/Wedge/Properties.agda b/Cubical/HITs/Wedge/Properties.agda index 5073e6316d..6377feb15a 100644 --- a/Cubical/HITs/Wedge/Properties.agda +++ b/Cubical/HITs/Wedge/Properties.agda @@ -9,17 +9,77 @@ open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma open import Cubical.Data.Unit +open import Cubical.Data.Sum as ⊎ open import Cubical.HITs.Pushout.Base open import Cubical.HITs.Wedge.Base open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout open import Cubical.Homotopy.Loopspace private variable - ℓ ℓ' : Level + ℓ ℓ' ℓ'' : Level + +⋁-commFun : {A : Pointed ℓ} {B : Pointed ℓ'} + → A ⋁ B → B ⋁ A +⋁-commFun (inl x) = inr x +⋁-commFun (inr x) = inl x +⋁-commFun (push a i) = push a (~ i) + +⋁-commFun² : {A : Pointed ℓ} {B : Pointed ℓ'} + → (x : A ⋁ B) → ⋁-commFun (⋁-commFun x) ≡ x +⋁-commFun² (inl x) = refl +⋁-commFun² (inr x) = refl +⋁-commFun² (push a i) = refl + +⋁-commIso : {A : Pointed ℓ} {B : Pointed ℓ'} + → Iso (A ⋁ B) (B ⋁ A) +Iso.fun ⋁-commIso = ⋁-commFun +Iso.inv ⋁-commIso = ⋁-commFun +Iso.rightInv ⋁-commIso = ⋁-commFun² +Iso.leftInv ⋁-commIso = ⋁-commFun² + +cofibInl-⋁ : {A : Pointed ℓ} {B : Pointed ℓ'} + → Iso (cofib {B = A ⋁ B} inl) (fst B) +Iso.fun (cofibInl-⋁ {B = B}) (inl x) = pt B +Iso.fun (cofibInl-⋁ {B = B}) (inr (inl x)) = pt B +Iso.fun cofibInl-⋁ (inr (inr x)) = x +Iso.fun (cofibInl-⋁ {B = B}) (inr (push a i)) = pt B +Iso.fun (cofibInl-⋁ {B = B}) (push a i) = pt B +Iso.inv cofibInl-⋁ x = inr (inr x) +Iso.rightInv cofibInl-⋁ x = refl +Iso.leftInv (cofibInl-⋁ {A = A}) (inl x) = + (λ i → inr (push tt (~ i))) ∙ sym (push (pt A)) +Iso.leftInv (cofibInl-⋁ {A = A}) (inr (inl x)) = + ((λ i → inr (push tt (~ i))) ∙ sym (push (pt A))) ∙ push x +Iso.leftInv cofibInl-⋁ (inr (inr x)) = refl +Iso.leftInv (cofibInl-⋁ {A = A}) (inr (push a i)) j = + help (λ i → inr (push tt (~ i))) (sym (push (pt A))) (~ i) j + where + help : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) + → Square refl ((p ∙ q) ∙ sym q) refl p + help p q = + (λ i j → p (i ∧ j)) + ▷ (rUnit p + ∙∙ cong (p ∙_) (sym (rCancel q)) + ∙∙ assoc p q (sym q)) +Iso.leftInv (cofibInl-⋁ {A = A}) (push a i) j = + compPath-filler + ((λ i → inr (push tt (~ i))) ∙ sym (push (pt A))) + (push a) i j + +cofibInr-⋁ : {A : Pointed ℓ} {B : Pointed ℓ'} + → Iso (cofib {B = A ⋁ B} inr) (fst A) +cofibInr-⋁ {A = A} {B = B} = + compIso (pushoutIso _ _ _ _ + (idEquiv _) + (idEquiv Unit) + (isoToEquiv ⋁-commIso) + refl refl) + (cofibInl-⋁ {A = B} {B = A}) -- Susp (⋁ᵢ Aᵢ) ≃ ⋁ᵢ (Susp Aᵢ) private @@ -165,3 +225,167 @@ Iso.leftInv (Iso-SuspBouquet-Bouquet A B) = SuspBouquet-Bouquet-cancel A B .snd SuspBouquet≃Bouquet : (A : Type ℓ) (B : A → Pointed ℓ') → Susp (⋁gen A B) ≃ ⋁gen A (λ a → Susp∙ (fst (B a))) SuspBouquet≃Bouquet A B = isoToEquiv (Iso-SuspBouquet-Bouquet A B) + +module _ {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (f : A →∙ C) (g : B →∙ C) + where + private + inst : 3x3-span + 3x3-span.A00 inst = Unit + 3x3-span.A02 inst = Unit + 3x3-span.A04 inst = Unit + 3x3-span.A20 inst = fst A + 3x3-span.A22 inst = Unit + 3x3-span.A24 inst = fst B + 3x3-span.A40 inst = fst C + 3x3-span.A42 inst = fst C + 3x3-span.A44 inst = fst C + 3x3-span.f10 inst = _ + 3x3-span.f12 inst = _ + 3x3-span.f14 inst = _ + 3x3-span.f30 inst = fst f + 3x3-span.f32 inst = λ _ → pt C + 3x3-span.f34 inst = fst g + 3x3-span.f01 inst = _ + 3x3-span.f21 inst = λ _ → pt A + 3x3-span.f41 inst = λ c → c + 3x3-span.f03 inst = _ + 3x3-span.f23 inst = λ _ → pt B + 3x3-span.f43 inst = λ c → c + 3x3-span.H11 inst = λ _ → refl + 3x3-span.H13 inst = λ _ → refl + 3x3-span.H31 inst = λ _ → sym (snd f) + 3x3-span.H33 inst = λ _ → sym (snd g) + + open 3x3-span inst + + A□2≅C : A□2 ≃ (fst C) + A□2≅C = isoToEquiv (PushoutAlongEquiv (idEquiv _) λ _ → pt C) + + A□○≅ : Iso (Pushout {B = cofib (fst f)} {C = cofib (fst g)} inr inr) A□○ + A□○≅ = + pushoutIso _ _ _ _ (invEquiv A□2≅C) (idEquiv _) (idEquiv _) + refl + refl + + A0□≅ : A0□ ≃ Unit + A0□≅ = isoToEquiv (PushoutAlongEquiv (idEquiv _) _) + + A4□≅ : A4□ ≃ fst C + A4□≅ = isoToEquiv (PushoutAlongEquiv (idEquiv _) _) + + A○□≅ : Iso A○□ (cofib (f ∨→ g)) + A○□≅ = pushoutIso _ _ _ _ (idEquiv _) A0□≅ A4□≅ refl + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push tt i) j → help j i}) + where + help : cong (fst A4□≅) (cong f3□ (push tt)) ≡ snd f ∙ sym (snd g) + help = cong-∙∙ (fst A4□≅) (λ i → inl (snd f i)) + refl (λ i → inl (snd g (~ i))) + ∙ doubleCompPath≡compPath _ _ _ + ∙ cong (snd f ∙_) (sym (lUnit _)) + + + cofibf-Square : commSquare + 3-span.A0 (commSquare.sp cofibf-Square) = Unit + 3-span.A2 (commSquare.sp cofibf-Square) = fst A + 3-span.A4 (commSquare.sp cofibf-Square) = fst C + 3-span.f1 (commSquare.sp cofibf-Square) = _ + 3-span.f3 (commSquare.sp cofibf-Square) = fst f + commSquare.P cofibf-Square = cofib (fst f) + commSquare.inlP cofibf-Square = λ _ → inl tt + commSquare.inrP cofibf-Square = inr + commSquare.comm cofibf-Square i x = push x i + + cofibf-PSquare : PushoutSquare + fst cofibf-PSquare = cofibf-Square + snd cofibf-PSquare = + subst isEquiv (funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl})) + (idIsEquiv _) + + F : cofib (fst f) + → cofib {B = cofib (fst g)} (λ x → inr (fst f x)) + F (inl x) = inl tt + F (inr x) = inr (inr x) + F (push a i) = push a i + + open PushoutPasteLeft + cofibf-PSquare + {B = cofib {B = cofib (fst g)} (λ x → inr (fst f x))} + inr inr F refl + + isPushoutRight : isPushoutSquare rightSquare + isPushoutRight = isPushoutTotSquare→isPushoutRightSquare + (subst isEquiv + (funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j → lUnit (sym (push a)) j (~ i)})) + (idEquiv _ .snd)) + + cofib∨→distrIso : + Iso (cofib (f ∨→ g)) + (Pushout {B = cofib (fst f)} {C = cofib (fst g)} inr inr) + cofib∨→distrIso = invIso (compIso (compIso A□○≅ 3x3-Iso) A○□≅) + + cofib∨→compIsoᵣ : Iso (cofib (f ∨→ g)) + (cofib {B = cofib (fst g)} (λ x → inr (fst f x))) + cofib∨→compIsoᵣ = compIso cofib∨→distrIso (equivToIso (_ , isPushoutRight)) + +cofib∨→compIsoₗ : + {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (f : A →∙ C) (g : B →∙ C) + → Iso (cofib (f ∨→ g)) + (cofib {B = cofib (fst f)} (λ x → inr (fst g x))) +cofib∨→compIsoₗ f g = + compIso (pushoutIso _ _ _ _ (isoToEquiv ⋁-commIso) + (idEquiv _) (idEquiv _) + refl (funExt + λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j → symDistr (snd g) (sym (snd f)) (~ j) i})) + (cofib∨→compIsoᵣ g f) + + +module _ {A : Type ℓ} {B : Type ℓ'} + {P : A → Pointed ℓ''} {Q : B → Pointed ℓ''} + where + private + ⋁gen⊎→ : ⋁gen∙ A P ⋁ ⋁gen∙ B Q + → ⋁gen (A ⊎ B) (⊎.rec P Q) + ⋁gen⊎→ (inl (inl x)) = inl tt + ⋁gen⊎→ (inl (inr (a , p))) = inr (inl a , p) + ⋁gen⊎→ (inl (push a i)) = push (inl a) i + ⋁gen⊎→ (inr (inl x)) = inl tt + ⋁gen⊎→ (inr (inr (b , q))) = inr (inr b , q) + ⋁gen⊎→ (inr (push b i)) = push (inr b) i + ⋁gen⊎→ (push a i) = inl tt + + ⋁gen⊎← : ⋁gen (A ⊎ B) (⊎.rec P Q) + → ⋁gen∙ A P ⋁ ⋁gen∙ B Q + ⋁gen⊎← (inl x) = inl (inl tt) + ⋁gen⊎← (inr (inl x , p)) = inl (inr (x , p)) + ⋁gen⊎← (inr (inr x , p)) = inr (inr (x , p)) + ⋁gen⊎← (push (inl x) i) = inl (push x i) + ⋁gen⊎← (push (inr x) i) = (push tt ∙ λ i → inr (push x i)) i + + ⋁gen⊎Iso : Iso (⋁gen∙ A P ⋁ ⋁gen∙ B Q) + (⋁gen (A ⊎ B) (⊎.rec P Q)) + Iso.fun ⋁gen⊎Iso = ⋁gen⊎→ + Iso.inv ⋁gen⊎Iso = ⋁gen⊎← + Iso.rightInv ⋁gen⊎Iso (inl tt) = refl + Iso.rightInv ⋁gen⊎Iso (inr (inl x , p)) = refl + Iso.rightInv ⋁gen⊎Iso (inr (inr x , p)) = refl + Iso.rightInv ⋁gen⊎Iso (push (inl x) i) = refl + Iso.rightInv ⋁gen⊎Iso (push (inr x) i) j = + ⋁gen⊎→ (compPath-filler' (push tt) (λ i → inr (push x i)) (~ j) i) + Iso.leftInv ⋁gen⊎Iso (inl (inl x)) = refl + Iso.leftInv ⋁gen⊎Iso (inl (inr x)) = refl + Iso.leftInv ⋁gen⊎Iso (inl (push a i)) = refl + Iso.leftInv ⋁gen⊎Iso (inr (inl x)) = push tt + Iso.leftInv ⋁gen⊎Iso (inr (inr x)) = refl + Iso.leftInv ⋁gen⊎Iso (inr (push a i)) j = + compPath-filler' (push tt) (λ i → inr (push a i)) (~ j) i + Iso.leftInv ⋁gen⊎Iso (push a i) j = push a (i ∧ j) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 842e58ac3d..922a257fe4 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -62,6 +62,12 @@ isConnectedSubtr {A = A} n m iscon = helper zero iscon = isContrUnit* helper (suc n) iscon = ∣ iscon .fst ∣ , (Trunc.elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ a → cong ∣_∣ (iscon .snd a)) +isConnectedSubtr' : ∀ {ℓ} {A : Type ℓ} (n m : HLevel) + → isConnected (m + n) A + → isConnected m A +isConnectedSubtr' {A = A} n m iscon = + isConnectedSubtr m n (subst (λ k → isConnected k A) (+-comm m n) iscon) + isConnectedFunSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : HLevel) (f : A → B) → isConnectedFun (m + n) f → isConnectedFun n f From 6bc36641e78cc27c04811c24236e6eb792ee9656 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 31 May 2024 18:51:51 +0200 Subject: [PATCH 46/73] pretty much done --- Cubical/CW/Pointed2.agda | 1025 ++++++++++++++++++++++---------------- 1 file changed, 586 insertions(+), 439 deletions(-) diff --git a/Cubical/CW/Pointed2.agda b/Cubical/CW/Pointed2.agda index 6ef19484fd..4b266b1c35 100644 --- a/Cubical/CW/Pointed2.agda +++ b/Cubical/CW/Pointed2.agda @@ -49,7 +49,22 @@ open import Cubical.Foundations.Univalence open import Cubical.Foundations.Equiv - +CW↪Iterate : ∀ {ℓ} (T : CWskel ℓ) (n m : ℕ) → fst T n → fst T (m +ℕ n) +CW↪Iterate T n zero = idfun _ +CW↪Iterate T n (suc m) x = CW↪ T (m +ℕ n) (CW↪Iterate T n m x) + +⋁-rUnitIso : ∀ {ℓ ℓ'} {A : Pointed ℓ} → Iso (A ⋁ (Unit* {ℓ'} , tt*)) (fst A) +Iso.fun ⋁-rUnitIso (inl x) = x +Iso.fun (⋁-rUnitIso {A = A}) (inr x) = pt A +Iso.fun (⋁-rUnitIso {A = A}) (push a i) = pt A +Iso.inv ⋁-rUnitIso = inl +Iso.rightInv ⋁-rUnitIso x = refl +Iso.leftInv ⋁-rUnitIso (inl x) = refl +Iso.leftInv ⋁-rUnitIso (inr x) = push tt +Iso.leftInv ⋁-rUnitIso (push tt i) j = push tt (i ∧ j) + +⋁-lUnitIso : ∀ {ℓ ℓ'} {A : Pointed ℓ} → Iso ((Unit* {ℓ'} , tt*) ⋁ A) (fst A) +⋁-lUnitIso = compIso ⋁-commIso ⋁-rUnitIso module _ {A : Pointed ℓ} {B : Pointed ℓ'} where @@ -1729,7 +1744,9 @@ connectedCWskelLift zero {C = C} ((cwsk , e) , cA) = M = connectedCWskel→ (cwsk .fst) (snd cwsk , subst (isConnected 2) (ua e) cA) connectedCWskelLift {ℓ = ℓ} (suc n) {C = C} ((cwsk , eqv) , cA) with (connectedCWskelLift n ((cwsk , eqv) , (isConnectedSubtr (suc (suc n)) 1 cA))) -... | s = PT.map {!C*!} s +... | s = PT.rec squash₁ + (λ s → PT.rec squash₁ + (λ α-pt* → {!!}) (α-pt* s)) s where module _ (ind : isConnectedCW n C) where open CWskel-fields (_ , ind .fst .snd .fst) @@ -1741,6 +1758,14 @@ connectedCWskelLift {ℓ = ℓ} (suc n) {C = C} ((cwsk , eqv) , cA) with ww = Susp→PushoutSusp + connC* : isConnected (3 +ℕ n) (C* (suc (suc (suc (suc n))))) + connC* = isConnectedCW→isConnectedSkel + (_ , ind .fst .snd .fst) (suc (suc (suc (suc n)))) + (suc (suc (suc n)) , <ᵗ-trans <ᵗsucm <ᵗsucm) + (subst (isConnected (3 +ℕ n)) (ua (invEquiv (ind .snd))) + cA) + + iso2 : Iso (C* (suc (suc n))) (SphereBouquet (suc n) (A (suc n))) iso2 = compIso (equivToIso (ind .fst .snd .fst .snd .snd .snd (suc n))) @@ -1772,6 +1797,11 @@ connectedCWskelLift {ℓ = ℓ} (suc n) {C = C} ((cwsk , eqv) , cA) with refl refl) + α-pt* : ∥ ((x : _) → α' (x , ptSn (suc n)) ≡ inl tt) ∥₁ + α-pt* = invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) + λ a → isConnectedPath 1 + (isConnectedSubtr' n 2 + (isConnectedSphereBouquet' {n = n})) _ _ .fst) module _ (α-pt : (x : _) → α' (x , ptSn (suc n)) ≡ inl tt) where @@ -1787,440 +1817,557 @@ connectedCWskelLift {ℓ = ℓ} (suc n) {C = C} ((cwsk , eqv) , cA) with (cofib (fst α∙)) iso4 = compIso iso3 (⋁-cofib-Iso _ α∙) - module _ (α-pt2 : (x : A (3 +ℕ n)) - → Iso.fun iso4 (α (3 +ℕ n) (x , north)) ≡ inl tt) where - α↑ : SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n)))) - →∙ (cofib (fst α∙) , inl tt) - fst α↑ (inl x) = inl tt - fst α↑ (inr x) = Iso.fun iso4 (α (3 +ℕ n) x) - fst α↑ (push a i) = α-pt2 a (~ i) - snd α↑ = refl - - iso5 : Iso (C* (4 +ℕ n)) (cofib (fst α↑)) - iso5 = compIso (equivToIso (e (3 +ℕ n))) - (compIso - (pushoutIso _ _ _ _ - (idEquiv _) - (isoToEquiv iso4) - (idEquiv _) refl refl) - (⋁-cofib-Iso _ α↑)) - - open commSquare - open 3-span - - PT : commSquare - A0 (sp PT) = cofib (fst α∙) - A2 (sp PT) = SphereBouquet (suc (suc n)) (A (suc (suc (suc n)))) - A4 (sp PT) = Unit - f1 (sp PT) = fst α↑ - f3 (sp PT) = λ _ → tt - P PT = cofib (fst α↑) - inlP PT = inr - inrP PT = inl - comm PT = funExt λ x → sym (push x) - - PTPush : PushoutSquare - fst PTPush = PT - snd PTPush = - subst isEquiv (funExt - (λ { (inl x) → refl - ; (inr x) → refl - ; (push a i) → refl})) - (symPushout _ _ .snd) - - opaque - iso6 : Iso (Pushout {B = cofib (fst α∙)} inr (λ _ → tt)) - (SphereBouquet (suc (suc n)) (A (suc (suc n)))) - iso6 = compIso (equivToIso (symPushout _ _)) - (compIso (Iso-cofibInr-Susp α∙) - sphereBouquetSuspIso) - - iso6∙ : Iso.fun iso6 (inl (inl tt)) ≡ inl tt - iso6∙ = refl - - PL : commSquare - A0 (sp PL) = cofib (fst α∙) - A2 (sp PL) = SphereBouquet (suc n) (A (suc n)) - A4 (sp PL) = Unit - f1 (sp PL) = inr - f3 (sp PL) = _ - P PL = SphereBouquet (suc (suc n)) (A (suc (suc n))) - inlP PL x = Iso.fun iso6 (inl x) - inrP PL _ = Iso.fun iso6 (inr tt) -- - comm PL = funExt λ x → cong (Iso.fun iso6) (push x) - - PLPush : PushoutSquare - fst PLPush = PL - snd PLPush = subst isEquiv (funExt coh) (isoToIsEquiv iso6) - where - coh : (x : _) → Iso.fun iso6 x ≡ Pushout→commSquare PL x - coh (inl x) = refl - coh (inr x) = refl - coh (push x i₁) = refl - - C∨PlaceHolder = Pushout (Iso.inv iso5 ∘ inr) λ x → Iso.fun iso6 (inl x) - - C∨ = (C* (suc (suc (suc (suc n)))) , Iso.inv iso5 (inl tt)) - ⋁ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))) - - module L→R = PushoutPasteDown PLPush {B = C∨PlaceHolder} (Iso.inv iso5 ∘ inr) inl inr (funExt push) - - isPushoutTot = L→R.isPushoutBottomSquare→isPushoutTotSquare - (subst isEquiv (funExt (λ { (inl x) → refl - ; (inr x) → refl - ; (push a i) → refl})) - (idIsEquiv _)) - - - C4n = C* (suc (suc (suc (suc n)))) - - const* : SphereBouquet (suc n) (A (suc n)) → C* (suc (suc (suc (suc n)))) - const* = λ x → Iso.inv iso5 (inr (inr x)) - - connC* : isConnected (3 +ℕ n) (C* (suc (suc (suc (suc n))))) - connC* = isConnectedCW→isConnectedSkel - (_ , ind .fst .snd .fst) (suc (suc (suc (suc n)))) - (suc (suc (suc n)) , <ᵗ-trans <ᵗsucm <ᵗsucm) - (subst (isConnected (3 +ℕ n)) (ua (invEquiv (ind .snd))) - cA) - - C⋆ = iso5 .Iso.inv (inl tt) - - sphereVanish : (f : S₊ (suc n) → C* (suc (suc (suc (suc n))))) - → ∥ ((x : S₊ (suc n)) → f x ≡ C⋆) ∥₁ - sphereVanish f = sphereToTrunc (suc n) - λ x → isConnectedPath (suc (suc n)) connC* _ _ .fst - - sphereBouquetVanish : ∀ {k : ℕ} (f : SphereBouquet (suc n) (Fin k) → C4n) - → ∥ f ≡ (λ _ → C⋆) ∥₁ - sphereBouquetVanish {k = k} f = - TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) - (λ p → PT.rec squash₁ - (λ q → ∣ funExt - (λ { (inl x) → p - ; (inr (x , y)) → (q x y ∙ sym (q x (ptSn (suc n)))) - ∙ cong f (sym (push x)) ∙ p - ; (push a i) j → (cong (_∙ cong f (sym (push a)) ∙ p) - (rCancel (q a (ptSn (suc n)))) - ∙ sym (lUnit _) - ◁ (symP (compPath-filler' - (cong f (sym (push a))) p))) (~ i) j}) ∣₁) - help) - isPted - where - isPted = Iso.fun (PathIdTruncIso (suc (suc n))) - (isContr→isProp connC* ∣ f (inl tt) ∣ₕ ∣ C⋆ ∣ₕ) - - help : ∥ ((x : Fin k) → (y : _) → f (inr (x , y)) ≡ C⋆) ∥₁ - help = invEq propTrunc≃Trunc1 - (invEq (_ , InductiveFinSatAC _ _ _) - λ x → PT.rec (isOfHLevelTrunc 1) - ∣_∣ₕ - (sphereVanish λ y → f (inr (x , y)))) - - pushoutTotAlt : C∨PlaceHolder ≃ cofib const* - pushoutTotAlt = - compEquiv - (invEquiv (_ , isPushoutTot)) - (symPushout _ _) - - β : SphereBouquet (suc (suc n)) (A (suc (suc (suc n)))) - → SphereBouquet (suc (suc n)) (A (suc (suc n))) - β = (Iso.fun iso6 ∘ inl) ∘ fst α↑ - - β∙ : SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n)))) - →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))) - fst β∙ = β - snd β∙ = iso6∙ - - module _ (vanish : const* ≡ λ _ → C⋆) where - S∨C = SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4n , C⋆) - - connS∨C : isConnected (suc (suc (suc n))) S∨C - fst connS∨C = ∣ inl (inl tt) ∣ₕ - snd connS∨C = - TR.elim (λ _ → isOfHLevelPath (suc (suc (suc n))) - (isOfHLevelTrunc (suc (suc (suc n)))) _ _) - λ { (inl x) → inlEq x - ; (inr x) → SP ∙ inrEq x - ; (push tt i) j → - (compPath-filler (inlEq (inl tt)) (cong ∣_∣ₕ (push tt)) - ▷ (rUnit SP ∙ cong (SP ∙_) (sym inrEq∙))) i j} - where - inlEq : (x : _) - → Path (hLevelTrunc (suc (suc (suc n))) S∨C) - ∣ inl (inl tt) ∣ ∣ inl x ∣ - inlEq x = TR.rec (isOfHLevelTrunc (3 +ℕ n) _ _) - (λ p i → ∣ inl (p i) ∣ₕ) - (Iso.fun (PathIdTruncIso _) - (isContr→isProp - (isConnectedSphereBouquet' {n = suc n}) - ∣ inl tt ∣ ∣ x ∣)) - - G : (x : C4n) → ∥ C⋆ ≡ x ∥ suc (suc n) - G x = Iso.fun (PathIdTruncIso _) (isContr→isProp connC* ∣ C⋆ ∣ₕ ∣ x ∣ₕ) - - G∙ : G C⋆ ≡ ∣ refl ∣ₕ - G∙ = cong (Iso.fun (PathIdTruncIso _)) - (isProp→isSet (isContr→isProp connC*) _ _ - (isContr→isProp connC* ∣ C⋆ ∣ₕ ∣ C⋆ ∣ₕ) refl) - ∙ cong ∣_∣ₕ (transportRefl refl) - - inrEq : (x : C4n) - → Path (hLevelTrunc (suc (suc (suc n))) S∨C) - ∣ inr C⋆ ∣ₕ ∣ inr x ∣ₕ - inrEq x = TR.rec (isOfHLevelTrunc (suc (suc (suc n))) _ _) - (λ p i → ∣ inr (p i) ∣ₕ) (G x) - - inrEq∙ : inrEq C⋆ ≡ refl - inrEq∙ = cong (TR.rec (isOfHLevelTrunc (suc (suc (suc n))) _ _) - (λ p i → ∣ inr (p i) ∣ₕ)) G∙ - - SP = inlEq (inl tt) ∙ cong ∣_∣ₕ (push tt) - - Iso-C∨PlaceHolder-Wedge : - C∨PlaceHolder ≃ (SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4n , C⋆)) - Iso-C∨PlaceHolder-Wedge = - compEquiv pushoutTotAlt - (isoToEquiv (compIso (cofibConst' {A = _ , inl tt} {B = _ , C⋆} - (const* , funExt⁻ vanish _) - (funExt⁻ vanish)) - (pushoutIso _ _ _ _ - (idEquiv _) - (isoToEquiv sphereBouquetSuspIso) - (idEquiv _) - refl - refl))) - - module T→B = PushoutPasteDown PTPush {B = C∨PlaceHolder} - (λ x → Iso.fun iso6 (inl x)) inr - (λ x → inl (Iso.inv iso5 x)) - (sym (funExt push)) - - helpIso : Iso (spanPushout (sp T→B.bottomSquare)) (P T→B.bottomSquare) - helpIso = compIso (equivToIso (symPushout _ _)) - (pushoutIso _ _ _ _ (idEquiv _) - (invEquiv (isoToEquiv iso5)) - (idEquiv _) refl refl) - - helpIsoCoh : (x : _) → Iso.fun helpIso x ≡ Pushout→commSquare T→B.bottomSquare x - helpIsoCoh (inl x) = refl - helpIsoCoh (inr x) = refl - helpIsoCoh (push a i) j = sym (rUnit (sym (push a))) j i - - ⋁-as-cofib : cofib β ≃ S∨C - ⋁-as-cofib = - compEquiv (compEquiv (symPushout _ _) - (_ , T→B.isPushoutBottomSquare→isPushoutTotSquare - (subst isEquiv (funExt helpIsoCoh) (isoToIsEquiv helpIso)))) - Iso-C∨PlaceHolder-Wedge - - -- S∨C-CW-fam : ℕ → Type ℓ - -- S∨C-CW-fam zero = ⊥* - -- S∨C-CW-fam (suc m) with (m ≟ᵗ suc (suc n)) - -- ... | lt x = Unit* - -- ... | eq x = Lift (SphereBouquet (suc (suc n)) (A (suc (suc n)))) - -- ... | gt x = S∨C - - - -- asdd = {!!} - - -- cardS∨C : ℕ → ℕ -- n + 2 - -- cardS∨C zero = 1 - -- cardS∨C (suc m) with (m ≟ᵗ suc n) | (m ≟ᵗ suc (suc n)) - -- ... | lt x | r = 0 - -- ... | eq x | r = card (suc (suc n)) - -- ... | gt x | lt x₁ = 0 - -- ... | gt x | eq x₁ = card (suc (suc (suc n))) - -- ... | gt x | gt x₁ = 0 - - -- αS∨C-non-triv : (m : ℕ) (p : suc (suc n) ≡ m) - -- → A (suc (suc (suc n))) × S₊ m → Lift {j = ℓ} (SphereBouquet (suc (suc n)) (A (suc (suc n)))) - -- αS∨C-non-triv m p (x , y) = lift (β (inr (x , (subst S₊ (sym p) y)))) - - -- αS∨C : (n₁ : ℕ) → Fin (cardS∨C n₁) × S⁻ n₁ → S∨C-CW-fam n₁ - -- αS∨C (suc m) with (m ≟ᵗ suc n) | (m ≟ᵗ suc (suc n)) - -- ... | eq x₁ | lt x = λ _ → tt* - -- ... | eq x₁ | eq x = λ _ → lift (inl tt) - -- ... | eq x₁ | gt x = λ _ → inl (inl tt) - -- ... | gt x₁ | lt x = λ() - -- ... | gt x₁ | eq p = αS∨C-non-triv m (sym p) -- - -- ... | gt x₁ | gt x = λ() - - -- eS∨C : (n₁ : ℕ) → S∨C-CW-fam (suc n₁) ≃ Pushout (αS∨C n₁) fst - -- eS∨C zero = isContr→Equiv (tt* , (λ {tt* → refl})) - -- ((inr fzero) - -- , (λ { (inr (zero , tt)) → refl})) - -- eS∨C (suc m) with (m ≟ᵗ suc n) | (m ≟ᵗ suc (suc n)) - -- ... | lt x | lt x₁ = - -- isContr→Equiv (tt* , (λ {tt* → refl})) ((inl tt*) - -- , λ { (inl tt*) → refl}) - -- ... | lt x | eq x₁ = - -- ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (_<ᵗ suc n) x₁ x) <ᵗsucm)) - -- ... | lt x | gt x₁ = - -- ⊥.rec (¬m<ᵗm (<ᵗ-trans (<ᵗ-trans x₁ x) <ᵗsucm)) - -- ... | eq x | lt x₁ = {!!} - -- ... | eq x | eq x₁ = - -- ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc n) ((cong predℕ (sym x ∙ x₁))) <ᵗsucm)) - -- ... | eq x | gt x₁ = - -- ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (suc (suc n) <ᵗ_) x x₁) <ᵗsucm)) - -- ... | gt x | lt x₁ = ⊥.rec (¬squeeze (x , x₁)) - -- ... | gt x | eq x₁ = {!isoToEquiv (invIso (PushoutEmptyFam ? ?))!} - -- ... | gt x | gt x₁ = - -- isoToEquiv (PushoutEmptyFam (λ()) λ()) - - -- S∨C-CW : CWskel ℓ - -- fst S∨C-CW = S∨C-CW-fam - -- fst (snd S∨C-CW) = cardS∨C - -- fst (snd (snd S∨C-CW)) = αS∨C - -- fst (snd (snd (snd S∨C-CW))) () - -- snd (snd (snd (snd S∨C-CW))) = eS∨C - - -- Sbouq-CW-fam : {!!} - -- Sbouq-CW-fam = {!elim.isConnectedPrecompose!} - - -- SbouqSkel : CWskel ℓ - -- SbouqSkel = {!!} - - open import Cubical.Data.Unit - testβ : isConnectedFun (suc (suc (suc n))) {B = cofib β} inr - testβ = inrConnected (3 +ℕ n) _ _ - λ b → isOfHLevelRetractFromIso 0 (mapCompIso fiberUnitIso) - isConnectedSphereBouquet' - - testβ3 : isConnectedFun (suc (suc (suc n))) {B = S∨C} (⋁-as-cofib .fst ∘ inr) - testβ3 = isConnectedComp (⋁-as-cofib .fst) inr (3 +ℕ n) - (isEquiv→isConnected _ (⋁-as-cofib .snd) (3 +ℕ n)) testβ - - main-inr : ∥ ((x : A (suc n)) (y : S₊ (suc (suc n))) - → Σ[ t ∈ SphereBouquet (suc (suc n)) (A (suc (suc n))) ] - (⋁-as-cofib .fst (inr t) ≡ inl (inr (x , y)))) ∥₁ - main-inr = - invEq propTrunc≃Trunc1 - (invEq (_ , InductiveFinSatAC _ _ _) - λ x → fst propTrunc≃Trunc1 - (sphereToTrunc (suc (suc n)) - λ y → testβ3 (inl (inr (x , y))) .fst)) - - mainF : ∃[ F ∈ (SphereBouquet∙ (suc (suc n)) (A (suc n)) - →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n)))) ] - ((x : _) → Path S∨C (⋁-as-cofib .fst (inr (fst F x))) (inl x)) - mainF = TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) - (λ coh₁ → PT.rec squash₁ (λ F → TR.rec squash₁ - (λ h → TR.rec squash₁ (λ q → ∣ (F* F coh₁ h , refl) , - (coh F coh₁ h q) ∣₁) - (makeCoh F coh₁ h)) - (F∙ F coh₁)) - main-inr) - (isConnectedPath _ connS∨C - (⋁-as-cofib .fst (inr (inl tt))) (inl (inl tt)) .fst) - where - CON = (subst (λ x → isConnected x (SphereBouquet (suc (suc n)) - (A (suc (suc n))))) - (cong suc (+-comm 2 n)) - isConnectedSphereBouquet') - module _ (F : ((x : A (suc n)) (y : S₊ (suc (suc n))) - → Σ[ t ∈ SphereBouquet (suc (suc n)) (A (suc (suc n))) ] - (⋁-as-cofib .fst (inr t) ≡ inl (inr (x , y))))) - (coh₁ : Path S∨C (⋁-as-cofib .fst (inr (inl tt))) - (inl (inl tt))) where - F∙ : hLevelTrunc 1 ((x : Fin _) → F x north .fst ≡ inl tt) - F∙ = - invEq (_ , InductiveFinSatAC _ _ _) - λ a → isConnectedPath 1 (isConnectedSubtr 2 (suc n) CON) _ _ .fst - - module _ (h : (x : Fin _) → F x north .fst ≡ inl tt) where - F* : (SphereBouquet (suc (suc n)) (A (suc n)) - → SphereBouquet (suc (suc n)) (A (suc (suc n)))) - F* (inl x) = inl tt - F* (inr (x , y)) = F x y .fst - F* (push a i) = h a (~ i) - - cohTy : Type _ - cohTy = (a : A (suc n)) - → Square (λ i → ⋁-as-cofib .fst (inr (h a (~ i)))) - (λ i → inl (push a i)) - coh₁ (F a north .snd) - - makeCoh : hLevelTrunc 1 cohTy - makeCoh = invEq (_ , InductiveFinSatAC _ _ _) - λ a → isConnectedPathP 1 - (isConnectedSubtr' n 2 (isConnectedPath _ connS∨C _ _)) _ _ .fst - - module _ (coh₂ : cohTy) where - coh : (x : _) → Path S∨C (⋁-as-cofib .fst (inr (F* x))) (inl x) - coh (inl x) = coh₁ - coh (inr (x , y)) = F x y .snd - coh (push a i) j = coh₂ a j i - module _ (F : (SphereBouquet∙ (suc (suc n)) (A (suc n)) - →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))))) - (h : (x : _) → Path S∨C (⋁-as-cofib .fst (inr (fst F x))) (inl x)) - where - h' : (x : _) → inr (fst F x) ≡ invEq ⋁-as-cofib (inl x) - h' x = sym (retEq ⋁-as-cofib (inr (fst F x))) - ∙ cong (invEq ⋁-as-cofib) (h x) - - open import Cubical.Foundations.Transport - - N = card (suc n) +ℕ card (suc (suc (suc n))) - - iso7 : Iso (SphereBouquet (suc (suc n)) (Fin N)) - (SphereBouquet∙ (suc (suc n)) (A (suc n)) - ⋁ SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n))))) - iso7 = compIso - (pathToIso - ((λ i → ⋁gen (isoToPath - (Fin+ {n = card (suc n)} {m = card (suc (suc (suc n)))}) (~ i)) - (λ _ → S₊∙ (suc (suc n)))) - ∙ cong (⋁gen (A (suc n) ⊎ A (suc (suc (suc n))))) - (funExt (⊎.elim (λ _ → refl) (λ _ → refl))))) - (invIso ⋁gen⊎Iso) - - T : Iso (cofib (F ∨→ β∙)) C4n - T = compIso (cofib∨→compIsoᵣ F β∙) - (compIso - (pathToIso (cong cofib (funExt h'))) - (compIso - (equivToIso (symPushout _ _)) - (compIso - (Pushout∘Equiv - (idEquiv (SphereBouquet (suc (suc n)) (A (suc n)))) - (invEquiv ⋁-as-cofib) inl (λ _ → tt)) - (compIso (equivToIso (symPushout _ _)) cofibInl-⋁)))) - - T∙ : Σ[ c ∈ ℕ ] Σ[ α ∈ (SphereBouquet (suc (suc n)) (Fin c) - → SphereBouquet (suc (suc n)) (A (suc (suc n)))) ] - Iso (Pushout (α ∘ inr) fst) C4n - fst T∙ = N - fst (snd T∙) = F ∨→ β∙ ∘ Iso.fun iso7 - snd (snd T∙) = - compIso - (compIso (⋁-cofib-Iso (_ , fst F (inl tt)) ((F ∨→ β∙ ∘ Iso.fun iso7) , refl)) - (invIso (Pushout∘Equiv - (invEquiv (isoToEquiv iso7)) - (idEquiv Unit) _ _))) T - --- ⋁gen⊎Iso - -{- - fst S∨C-CW = S∨C-CW-fam - fst (fst (snd S∨C-CW)) = {!S∨C-CW-fam!} - fst (snd (fst (snd S∨C-CW))) = {!!} - snd (snd (fst (snd S∨C-CW))) = {!!} - snd (snd S∨C-CW) = {!!} --} - - - {- - snd PL = subst isEquiv {!!} (isoToIsEquiv help) - where - help : Iso _ _ - help = compIso {!!} {!!} - - main : isConnectedCW (suc n) C - main = {!!} --} -{-Iso.fun help ≡ Pushout→commSquare PL -yieldsConnectedCWskel : ? -yieldsConnectedCWskel = ? --} + α-pt2* : ∥ ((x : A (3 +ℕ n)) + → Iso.fun iso4 (α (3 +ℕ n) (x , north)) ≡ inl tt) ∥₁ + α-pt2* = invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) + λ a → isConnectedPath 1 {!isConnectedCofib!} _ _ .fst) + +-- module _ (α-pt2 : (x : A (3 +ℕ n)) +-- → Iso.fun iso4 (α (3 +ℕ n) (x , north)) ≡ inl tt) where +-- α↑ : SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n)))) +-- →∙ (cofib (fst α∙) , inl tt) +-- fst α↑ (inl x) = inl tt +-- fst α↑ (inr x) = Iso.fun iso4 (α (3 +ℕ n) x) +-- fst α↑ (push a i) = α-pt2 a (~ i) +-- snd α↑ = refl + +-- iso5 : Iso (C* (4 +ℕ n)) (cofib (fst α↑)) +-- iso5 = compIso (equivToIso (e (3 +ℕ n))) +-- (compIso +-- (pushoutIso _ _ _ _ +-- (idEquiv _) +-- (isoToEquiv iso4) +-- (idEquiv _) refl refl) +-- (⋁-cofib-Iso _ α↑)) + +-- open commSquare +-- open 3-span + +-- PT : commSquare +-- A0 (sp PT) = cofib (fst α∙) +-- A2 (sp PT) = SphereBouquet (suc (suc n)) (A (suc (suc (suc n)))) +-- A4 (sp PT) = Unit +-- f1 (sp PT) = fst α↑ +-- f3 (sp PT) = λ _ → tt +-- P PT = cofib (fst α↑) +-- inlP PT = inr +-- inrP PT = inl +-- comm PT = funExt λ x → sym (push x) + +-- PTPush : PushoutSquare +-- fst PTPush = PT +-- snd PTPush = +-- subst isEquiv (funExt +-- (λ { (inl x) → refl +-- ; (inr x) → refl +-- ; (push a i) → refl})) +-- (symPushout _ _ .snd) + +-- opaque +-- iso6 : Iso (Pushout {B = cofib (fst α∙)} inr (λ _ → tt)) +-- (SphereBouquet (suc (suc n)) (A (suc (suc n)))) +-- iso6 = compIso (equivToIso (symPushout _ _)) +-- (compIso (Iso-cofibInr-Susp α∙) +-- sphereBouquetSuspIso) + +-- iso6∙ : Iso.fun iso6 (inl (inl tt)) ≡ inl tt +-- iso6∙ = refl + +-- PL : commSquare +-- A0 (sp PL) = cofib (fst α∙) +-- A2 (sp PL) = SphereBouquet (suc n) (A (suc n)) +-- A4 (sp PL) = Unit +-- f1 (sp PL) = inr +-- f3 (sp PL) = _ +-- P PL = SphereBouquet (suc (suc n)) (A (suc (suc n))) +-- inlP PL x = Iso.fun iso6 (inl x) +-- inrP PL _ = Iso.fun iso6 (inr tt) -- +-- comm PL = funExt λ x → cong (Iso.fun iso6) (push x) + +-- PLPush : PushoutSquare +-- fst PLPush = PL +-- snd PLPush = subst isEquiv (funExt coh) (isoToIsEquiv iso6) +-- where +-- coh : (x : _) → Iso.fun iso6 x ≡ Pushout→commSquare PL x +-- coh (inl x) = refl +-- coh (inr x) = refl +-- coh (push x i₁) = refl + +-- C∨PlaceHolder = Pushout (Iso.inv iso5 ∘ inr) λ x → Iso.fun iso6 (inl x) + +-- C∨ = (C* (suc (suc (suc (suc n)))) , Iso.inv iso5 (inl tt)) +-- ⋁ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))) + +-- module L→R = PushoutPasteDown PLPush {B = C∨PlaceHolder} (Iso.inv iso5 ∘ inr) inl inr (funExt push) + +-- isPushoutTot = L→R.isPushoutBottomSquare→isPushoutTotSquare +-- (subst isEquiv (funExt (λ { (inl x) → refl +-- ; (inr x) → refl +-- ; (push a i) → refl})) +-- (idIsEquiv _)) + + +-- C4n = C* (suc (suc (suc (suc n)))) + +-- const* : SphereBouquet (suc n) (A (suc n)) → C* (suc (suc (suc (suc n)))) +-- const* = λ x → Iso.inv iso5 (inr (inr x)) + + +-- C⋆ = iso5 .Iso.inv (inl tt) + +-- sphereVanish : (f : S₊ (suc n) → C* (suc (suc (suc (suc n))))) +-- → ∥ ((x : S₊ (suc n)) → f x ≡ C⋆) ∥₁ +-- sphereVanish f = sphereToTrunc (suc n) +-- λ x → isConnectedPath (suc (suc n)) connC* _ _ .fst + +-- sphereBouquetVanish : ∀ {k : ℕ} (f : SphereBouquet (suc n) (Fin k) → C4n) +-- → ∥ f ≡ (λ _ → C⋆) ∥₁ +-- sphereBouquetVanish {k = k} f = +-- TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) +-- (λ p → PT.rec squash₁ +-- (λ q → ∣ funExt +-- (λ { (inl x) → p +-- ; (inr (x , y)) → (q x y ∙ sym (q x (ptSn (suc n)))) +-- ∙ cong f (sym (push x)) ∙ p +-- ; (push a i) j → (cong (_∙ cong f (sym (push a)) ∙ p) +-- (rCancel (q a (ptSn (suc n)))) +-- ∙ sym (lUnit _) +-- ◁ (symP (compPath-filler' +-- (cong f (sym (push a))) p))) (~ i) j}) ∣₁) +-- help) +-- isPted +-- where +-- isPted = Iso.fun (PathIdTruncIso (suc (suc n))) +-- (isContr→isProp connC* ∣ f (inl tt) ∣ₕ ∣ C⋆ ∣ₕ) + +-- help : ∥ ((x : Fin k) → (y : _) → f (inr (x , y)) ≡ C⋆) ∥₁ +-- help = invEq propTrunc≃Trunc1 +-- (invEq (_ , InductiveFinSatAC _ _ _) +-- λ x → PT.rec (isOfHLevelTrunc 1) +-- ∣_∣ₕ +-- (sphereVanish λ y → f (inr (x , y)))) + +-- pushoutTotAlt : C∨PlaceHolder ≃ cofib const* +-- pushoutTotAlt = +-- compEquiv +-- (invEquiv (_ , isPushoutTot)) +-- (symPushout _ _) + +-- β : SphereBouquet (suc (suc n)) (A (suc (suc (suc n)))) +-- → SphereBouquet (suc (suc n)) (A (suc (suc n))) +-- β = (Iso.fun iso6 ∘ inl) ∘ fst α↑ + +-- β∙ : SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n)))) +-- →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))) +-- fst β∙ = β +-- snd β∙ = iso6∙ + +-- module _ (vanish : const* ≡ λ _ → C⋆) where +-- S∨C = SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4n , C⋆) + +-- connS∨C : isConnected (suc (suc (suc n))) S∨C +-- fst connS∨C = ∣ inl (inl tt) ∣ₕ +-- snd connS∨C = +-- TR.elim (λ _ → isOfHLevelPath (suc (suc (suc n))) +-- (isOfHLevelTrunc (suc (suc (suc n)))) _ _) +-- λ { (inl x) → inlEq x +-- ; (inr x) → SP ∙ inrEq x +-- ; (push tt i) j → +-- (compPath-filler (inlEq (inl tt)) (cong ∣_∣ₕ (push tt)) +-- ▷ (rUnit SP ∙ cong (SP ∙_) (sym inrEq∙))) i j} +-- where +-- inlEq : (x : _) +-- → Path (hLevelTrunc (suc (suc (suc n))) S∨C) +-- ∣ inl (inl tt) ∣ ∣ inl x ∣ +-- inlEq x = TR.rec (isOfHLevelTrunc (3 +ℕ n) _ _) +-- (λ p i → ∣ inl (p i) ∣ₕ) +-- (Iso.fun (PathIdTruncIso _) +-- (isContr→isProp +-- (isConnectedSphereBouquet' {n = suc n}) +-- ∣ inl tt ∣ ∣ x ∣)) + +-- G : (x : C4n) → ∥ C⋆ ≡ x ∥ suc (suc n) +-- G x = Iso.fun (PathIdTruncIso _) (isContr→isProp connC* ∣ C⋆ ∣ₕ ∣ x ∣ₕ) + +-- G∙ : G C⋆ ≡ ∣ refl ∣ₕ +-- G∙ = cong (Iso.fun (PathIdTruncIso _)) +-- (isProp→isSet (isContr→isProp connC*) _ _ +-- (isContr→isProp connC* ∣ C⋆ ∣ₕ ∣ C⋆ ∣ₕ) refl) +-- ∙ cong ∣_∣ₕ (transportRefl refl) + +-- inrEq : (x : C4n) +-- → Path (hLevelTrunc (suc (suc (suc n))) S∨C) +-- ∣ inr C⋆ ∣ₕ ∣ inr x ∣ₕ +-- inrEq x = TR.rec (isOfHLevelTrunc (suc (suc (suc n))) _ _) +-- (λ p i → ∣ inr (p i) ∣ₕ) (G x) + +-- inrEq∙ : inrEq C⋆ ≡ refl +-- inrEq∙ = cong (TR.rec (isOfHLevelTrunc (suc (suc (suc n))) _ _) +-- (λ p i → ∣ inr (p i) ∣ₕ)) G∙ + +-- SP = inlEq (inl tt) ∙ cong ∣_∣ₕ (push tt) + +-- Iso-C∨PlaceHolder-Wedge : +-- C∨PlaceHolder ≃ (SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4n , C⋆)) +-- Iso-C∨PlaceHolder-Wedge = +-- compEquiv pushoutTotAlt +-- (isoToEquiv (compIso (cofibConst' {A = _ , inl tt} {B = _ , C⋆} +-- (const* , funExt⁻ vanish _) +-- (funExt⁻ vanish)) +-- (pushoutIso _ _ _ _ +-- (idEquiv _) +-- (isoToEquiv sphereBouquetSuspIso) +-- (idEquiv _) +-- refl +-- refl))) + +-- module T→B = PushoutPasteDown PTPush {B = C∨PlaceHolder} +-- (λ x → Iso.fun iso6 (inl x)) inr +-- (λ x → inl (Iso.inv iso5 x)) +-- (sym (funExt push)) + +-- helpIso : Iso (spanPushout (sp T→B.bottomSquare)) (P T→B.bottomSquare) +-- helpIso = compIso (equivToIso (symPushout _ _)) +-- (pushoutIso _ _ _ _ (idEquiv _) +-- (invEquiv (isoToEquiv iso5)) +-- (idEquiv _) refl refl) + +-- helpIsoCoh : (x : _) → Iso.fun helpIso x ≡ Pushout→commSquare T→B.bottomSquare x +-- helpIsoCoh (inl x) = refl +-- helpIsoCoh (inr x) = refl +-- helpIsoCoh (push a i) j = sym (rUnit (sym (push a))) j i + +-- ⋁-as-cofib : cofib β ≃ S∨C +-- ⋁-as-cofib = +-- compEquiv (compEquiv (symPushout _ _) +-- (_ , T→B.isPushoutBottomSquare→isPushoutTotSquare +-- (subst isEquiv (funExt helpIsoCoh) (isoToIsEquiv helpIso)))) +-- Iso-C∨PlaceHolder-Wedge + +-- open import Cubical.Data.Unit +-- testβ : isConnectedFun (suc (suc (suc n))) {B = cofib β} inr +-- testβ = inrConnected (3 +ℕ n) _ _ +-- λ b → isOfHLevelRetractFromIso 0 (mapCompIso fiberUnitIso) +-- isConnectedSphereBouquet' + +-- testβ3 : isConnectedFun (suc (suc (suc n))) {B = S∨C} (⋁-as-cofib .fst ∘ inr) +-- testβ3 = isConnectedComp (⋁-as-cofib .fst) inr (3 +ℕ n) +-- (isEquiv→isConnected _ (⋁-as-cofib .snd) (3 +ℕ n)) testβ + +-- main-inr : ∥ ((x : A (suc n)) (y : S₊ (suc (suc n))) +-- → Σ[ t ∈ SphereBouquet (suc (suc n)) (A (suc (suc n))) ] +-- (⋁-as-cofib .fst (inr t) ≡ inl (inr (x , y)))) ∥₁ +-- main-inr = +-- invEq propTrunc≃Trunc1 +-- (invEq (_ , InductiveFinSatAC _ _ _) +-- λ x → fst propTrunc≃Trunc1 +-- (sphereToTrunc (suc (suc n)) +-- λ y → testβ3 (inl (inr (x , y))) .fst)) + +-- mainF : ∃[ F ∈ (SphereBouquet∙ (suc (suc n)) (A (suc n)) +-- →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n)))) ] +-- ((x : _) → Path S∨C (⋁-as-cofib .fst (inr (fst F x))) (inl x)) +-- mainF = TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) +-- (λ coh₁ → PT.rec squash₁ (λ F → TR.rec squash₁ +-- (λ h → TR.rec squash₁ (λ q → ∣ (F* F coh₁ h , refl) , +-- (coh F coh₁ h q) ∣₁) +-- (makeCoh F coh₁ h)) +-- (F∙ F coh₁)) +-- main-inr) +-- (isConnectedPath _ connS∨C +-- (⋁-as-cofib .fst (inr (inl tt))) (inl (inl tt)) .fst) +-- where +-- CON = (subst (λ x → isConnected x (SphereBouquet (suc (suc n)) +-- (A (suc (suc n))))) +-- (cong suc (+-comm 2 n)) +-- isConnectedSphereBouquet') +-- module _ (F : ((x : A (suc n)) (y : S₊ (suc (suc n))) +-- → Σ[ t ∈ SphereBouquet (suc (suc n)) (A (suc (suc n))) ] +-- (⋁-as-cofib .fst (inr t) ≡ inl (inr (x , y))))) +-- (coh₁ : Path S∨C (⋁-as-cofib .fst (inr (inl tt))) +-- (inl (inl tt))) where +-- F∙ : hLevelTrunc 1 ((x : Fin _) → F x north .fst ≡ inl tt) +-- F∙ = +-- invEq (_ , InductiveFinSatAC _ _ _) +-- λ a → isConnectedPath 1 (isConnectedSubtr 2 (suc n) CON) _ _ .fst + +-- module _ (h : (x : Fin _) → F x north .fst ≡ inl tt) where +-- F* : (SphereBouquet (suc (suc n)) (A (suc n)) +-- → SphereBouquet (suc (suc n)) (A (suc (suc n)))) +-- F* (inl x) = inl tt +-- F* (inr (x , y)) = F x y .fst +-- F* (push a i) = h a (~ i) + +-- cohTy : Type _ +-- cohTy = (a : A (suc n)) +-- → Square (λ i → ⋁-as-cofib .fst (inr (h a (~ i)))) +-- (λ i → inl (push a i)) +-- coh₁ (F a north .snd) + +-- makeCoh : hLevelTrunc 1 cohTy +-- makeCoh = invEq (_ , InductiveFinSatAC _ _ _) +-- λ a → isConnectedPathP 1 +-- (isConnectedSubtr' n 2 (isConnectedPath _ connS∨C _ _)) _ _ .fst + +-- module _ (coh₂ : cohTy) where +-- coh : (x : _) → Path S∨C (⋁-as-cofib .fst (inr (F* x))) (inl x) +-- coh (inl x) = coh₁ +-- coh (inr (x , y)) = F x y .snd +-- coh (push a i) j = coh₂ a j i +-- module _ (F : (SphereBouquet∙ (suc (suc n)) (A (suc n)) +-- →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))))) +-- (h : (x : _) → Path S∨C (⋁-as-cofib .fst (inr (fst F x))) (inl x)) +-- where +-- h' : (x : _) → inr (fst F x) ≡ invEq ⋁-as-cofib (inl x) +-- h' x = sym (retEq ⋁-as-cofib (inr (fst F x))) +-- ∙ cong (invEq ⋁-as-cofib) (h x) + +-- open import Cubical.Foundations.Transport + +-- N = card (suc n) +ℕ card (suc (suc (suc n))) + +-- iso7 : Iso (SphereBouquet (suc (suc n)) (Fin N)) +-- (SphereBouquet∙ (suc (suc n)) (A (suc n)) +-- ⋁ SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n))))) +-- iso7 = compIso +-- (pathToIso +-- ((λ i → ⋁gen (isoToPath +-- (Fin+ {n = card (suc n)} {m = card (suc (suc (suc n)))}) (~ i)) +-- (λ _ → S₊∙ (suc (suc n)))) +-- ∙ cong (⋁gen (A (suc n) ⊎ A (suc (suc (suc n))))) +-- (funExt (⊎.elim (λ _ → refl) (λ _ → refl))))) +-- (invIso ⋁gen⊎Iso) + +-- T : Iso (cofib (F ∨→ β∙)) C4n +-- T = compIso (cofib∨→compIsoᵣ F β∙) +-- (compIso +-- (pathToIso (cong cofib (funExt h'))) +-- (compIso +-- (equivToIso (symPushout _ _)) +-- (compIso +-- (Pushout∘Equiv +-- (idEquiv (SphereBouquet (suc (suc n)) (A (suc n)))) +-- (invEquiv ⋁-as-cofib) inl (λ _ → tt)) +-- (compIso (equivToIso (symPushout _ _)) cofibInl-⋁)))) + + +-- T∙ : Σ[ c ∈ ℕ ] Σ[ α ∈ (SphereBouquet (suc (suc n)) (Fin c) +-- → SphereBouquet (suc (suc n)) (A (suc (suc n)))) ] +-- Iso (Pushout (α ∘ inr) fst) C4n +-- fst T∙ = N +-- fst (snd T∙) = F ∨→ β∙ ∘ Iso.fun iso7 +-- snd (snd T∙) = +-- compIso +-- (compIso (⋁-cofib-Iso (_ , fst F (inl tt)) ((F ∨→ β∙ ∘ Iso.fun iso7) , refl)) +-- (invIso (Pushout∘Equiv +-- (invEquiv (isoToEquiv iso7)) +-- (idEquiv Unit) _ _))) T + +-- tyFamParam : (m : ℕ) → Trichotomyᵗ m (suc (suc (suc n))) → Type ℓ +-- tyFamParam zero (lt x) = ⊥* +-- tyFamParam (suc m) (lt x) = Unit* +-- tyFamParam m (eq x) = Lift (SphereBouquet (suc (suc n)) (A (suc (suc n)))) +-- tyFamParam m (gt x) = C* m + +-- cardParam : (m : ℕ) → Trichotomyᵗ m (suc (suc n)) → Trichotomyᵗ m (suc (suc (suc n))) → ℕ +-- cardParam zero (lt x) q = 1 +-- cardParam (suc m) (lt x) q = 0 +-- cardParam m (eq x) q = card (suc (suc n)) +-- cardParam m (gt x) (lt x₁) = 0 +-- cardParam m (gt x) (eq x₁) = N +-- cardParam m (gt x) (gt x₁) = card m + + +-- tyFamParam∙ : (m : ℕ) (q : _) → tyFamParam (suc m) q +-- tyFamParam∙ m (lt x) = tt* +-- tyFamParam∙ m (eq x) = lift (inl tt) +-- tyFamParam∙ m (gt x) = subst C* (+-comm m 1) +-- (CW↪Iterate (_ , ind .fst .snd .fst) 1 m +-- (invEq (CW₁-discrete (_ , ind .fst .snd .fst)) +-- (subst Fin (sym (ind .fst .snd .snd .fst)) flast))) + +-- carParamConn : (n₁ : ℕ) → n₁ <ᵗ suc n → (p : _) (q : _) → cardParam (suc n₁) p q ≡ 0 +-- carParamConn m ineq (lt x) q = refl +-- carParamConn m ineq (eq x) q = ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc (suc n)) x ineq)) +-- carParamConn m ineq (gt x) q = ⊥.rec (¬m<ᵗm (<ᵗ-trans ineq x)) + +-- newα : (m : ℕ) (p : _) (q : _) → Fin (cardParam m p q) × S⁻ m → tyFamParam m q +-- newα (suc m) (lt x) q () +-- newα (suc m) (eq x) (lt x₁) _ = tt* +-- newα (suc m) (eq x) (eq x₁) _ = tyFamParam∙ m (eq x₁) +-- newα (suc m) (eq x) (gt x₁) _ = tyFamParam∙ m (gt x₁) +-- newα (suc m) (gt x) (lt x₁) _ = tyFamParam∙ m (lt x₁) +-- newα (suc m) (gt x) (eq x₁) (a , p) = lift (snd T∙ .fst (inr (a , subst S₊ (cong predℕ x₁) p))) +-- newα (suc m) (gt x) (gt x₁) = α (suc m) + +-- tyFamParamContr : (m : _) (t : m <ᵗ suc n) (q : _) → isContr (tyFamParam (suc m) q) +-- tyFamParamContr m t (lt x) = tt* , λ {tt* → refl} +-- tyFamParamContr m t (eq x) = +-- ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (_<ᵗ suc n) (cong predℕ x) t) <ᵗsucm)) +-- tyFamParamContr m t (gt x) = +-- ⊥.rec (¬m<ᵗm (<ᵗ-trans (<ᵗ-trans x t) <ᵗsucm)) + +-- newEq : (n₁ : ℕ) (p : _) (q : _) +-- → tyFamParam (suc n₁) (Trichotomyᵗ-suc p) +-- ≃ Pushout (newα n₁ p q) fst +-- newEq zero (lt s) (lt t) = +-- isoToEquiv (isContr→Iso (tt* , (λ {tt* → refl})) +-- ((inr flast) , λ {(inr (zero , tt)) → refl})) +-- newEq zero (lt x) (eq p) = ⊥.rec (snotz (sym p)) +-- newEq zero (eq x) q = ⊥.rec (snotz (sym x)) +-- newEq (suc m) (lt x) q = +-- isoToEquiv (isContr→Iso (tt* , (λ {tt* → refl})) +-- ((inl (tyFamParamContr m x q .fst)) +-- , (λ { (inl t) → λ i +-- → inl (tyFamParamContr m x q .snd t i)}))) +-- newEq (suc m) (eq x) (lt x₁) = +-- invEquiv (isoToEquiv +-- (compIso (⋁-cofib-Iso {B = λ _ → _ , ptSn m} (_ , tt*) (_ , refl)) +-- (compIso (cofibConst {A = SphereBouquet∙ m _}) +-- (compIso ⋁-rUnitIso +-- (compIso sphereBouquetSuspIso +-- (compIso +-- (pathToIso (λ i → SphereBouquet (x i) (A (suc (suc n))))) +-- LiftIso)))))) +-- newEq (suc m) (eq x) (eq x₁) = +-- ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc n) +-- (cong (predℕ ∘ predℕ) (sym x ∙ x₁)) <ᵗsucm)) +-- newEq (suc m) (eq x) (gt y) = +-- ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (suc (suc n) <ᵗ_) +-- (cong predℕ x) y) <ᵗsucm)) +-- newEq (suc m) (gt x) (lt x₁) = ⊥.rec (¬squeeze (x , x₁)) +-- newEq (suc m) (gt x) (eq x₁) = +-- isoToEquiv (compIso (compIso +-- (pathToIso (λ i → C* (suc (x₁ i)))) +-- (invIso (T∙ .snd .snd))) +-- (invIso (pushoutIso _ _ _ _ +-- (Σ-cong-equiv-snd +-- (λ _ → pathToEquiv (cong S₊ (cong predℕ x₁)))) +-- (invEquiv LiftEquiv) (idEquiv _) refl refl))) +-- newEq (suc m) (gt x) (gt x₁) = e (suc m) + +-- buildCW* : connectedCWskel ℓ (suc n) +-- fst buildCW* m = tyFamParam m (m ≟ᵗ (3 +ℕ n)) +-- fst (fst (snd buildCW*)) m = cardParam m (m ≟ᵗ (2 +ℕ n)) (m ≟ᵗ (3 +ℕ n)) +-- fst (snd (fst (snd buildCW*))) m = newα m _ _ +-- fst (snd (snd (fst (snd buildCW*)))) () +-- snd (snd (snd (fst (snd buildCW*)))) m = newEq m _ _ +-- fst (snd (snd buildCW*)) = refl +-- snd (snd (snd buildCW*)) m ineq = carParamConn m ineq _ _ + +-- 3+n = suc (suc (suc n)) +-- 4+n = suc 3+n + +-- buildCW-realise : (n₁ : ℕ) (q : _) +-- → Iso (tyFamParam (n₁ +ℕ 4+n) q) +-- (C* (n₁ +ℕ 4+n)) +-- buildCW-realise m (lt x) = ⊥.rec (¬m Date: Tue, 4 Jun 2024 02:59:42 +0200 Subject: [PATCH 47/73] cleaning --- Cubical/CW/Base.agda | 4 + Cubical/CW/Pointed2.agda | 2180 ++++++-------------- Cubical/CW/Properties.agda | 35 + Cubical/Data/Fin/Inductive/Base.agda | 3 + Cubical/Data/Fin/Inductive/Properties.agda | 205 +- Cubical/Data/Sum/Properties.agda | 64 + Cubical/HITs/Pushout/Properties.agda | 121 +- Cubical/HITs/SphereBouquet/Properties.agda | 16 + Cubical/HITs/Wedge/Properties.agda | 155 ++ Cubical/Relation/Nullary/Properties.agda | 7 + 10 files changed, 1281 insertions(+), 1509 deletions(-) diff --git a/Cubical/CW/Base.agda b/Cubical/CW/Base.agda index fd10c9309c..6b9cf256c5 100644 --- a/Cubical/CW/Base.agda +++ b/Cubical/CW/Base.agda @@ -151,6 +151,10 @@ to_cofibCW n C x = inr x CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C CW↪∞ C n x = incl x +CW↪Iterate : ∀ {ℓ} (T : CWskel ℓ) (n m : ℕ) → fst T n → fst T (m +ℕ n) +CW↪Iterate T n zero = idfun _ +CW↪Iterate T n (suc m) x = CW↪ T (m +ℕ n) (CW↪Iterate T n m x) + finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m fst (finCW↑ m n p C) = fst C fst (snd (finCW↑ m n p C)) = snd C .fst diff --git a/Cubical/CW/Pointed2.agda b/Cubical/CW/Pointed2.agda index 4b266b1c35..870efc17ac 100644 --- a/Cubical/CW/Pointed2.agda +++ b/Cubical/CW/Pointed2.agda @@ -24,6 +24,7 @@ open import Cubical.HITs.SequentialColimit open import Cubical.HITs.SphereBouquet open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.Wedge +open import Cubical.HITs.Pushout open import Cubical.Axiom.Choice @@ -48,249 +49,6 @@ open import Cubical.Foundations.Univalence open import Cubical.Foundations.Equiv - -CW↪Iterate : ∀ {ℓ} (T : CWskel ℓ) (n m : ℕ) → fst T n → fst T (m +ℕ n) -CW↪Iterate T n zero = idfun _ -CW↪Iterate T n (suc m) x = CW↪ T (m +ℕ n) (CW↪Iterate T n m x) - -⋁-rUnitIso : ∀ {ℓ ℓ'} {A : Pointed ℓ} → Iso (A ⋁ (Unit* {ℓ'} , tt*)) (fst A) -Iso.fun ⋁-rUnitIso (inl x) = x -Iso.fun (⋁-rUnitIso {A = A}) (inr x) = pt A -Iso.fun (⋁-rUnitIso {A = A}) (push a i) = pt A -Iso.inv ⋁-rUnitIso = inl -Iso.rightInv ⋁-rUnitIso x = refl -Iso.leftInv ⋁-rUnitIso (inl x) = refl -Iso.leftInv ⋁-rUnitIso (inr x) = push tt -Iso.leftInv ⋁-rUnitIso (push tt i) j = push tt (i ∧ j) - -⋁-lUnitIso : ∀ {ℓ ℓ'} {A : Pointed ℓ} → Iso ((Unit* {ℓ'} , tt*) ⋁ A) (fst A) -⋁-lUnitIso = compIso ⋁-commIso ⋁-rUnitIso - -module _ {A : Pointed ℓ} {B : Pointed ℓ'} where - - cofibConst→⋁ : cofib (const∙ A B .fst) → Susp∙ (fst A) ⋁ B - cofibConst→⋁ (inl x) = inl north - cofibConst→⋁ (inr x) = inr x - cofibConst→⋁ (push a i) = ((λ i → inl (toSusp A a i)) ∙ push tt) i - - ⋁→cofibConst : Susp∙ (fst A) ⋁ B → cofib (const∙ A B .fst) - ⋁→cofibConst (inl north) = inl tt - ⋁→cofibConst (inl south) = inr (pt B) - ⋁→cofibConst (inl (merid a i)) = push a i - ⋁→cofibConst (inr x) = inr x - ⋁→cofibConst (push a i) = push (pt A) i - open import Cubical.Foundations.GroupoidLaws - - cofibConst : Iso (cofib (const∙ A B .fst)) - (Susp∙ (fst A) ⋁ B) - Iso.fun cofibConst = cofibConst→⋁ - Iso.inv cofibConst = ⋁→cofibConst - Iso.rightInv cofibConst (inl north) = refl - Iso.rightInv cofibConst (inl south) = sym (push tt) ∙ λ i → inl (merid (pt A) i) - Iso.rightInv cofibConst (inl (merid a i)) j = - hcomp (λ k → - λ {(i = i0) → inl (compPath-filler (merid a) (sym (merid (pt A))) (~ j) (~ k)) - ; (i = i1) → (sym (push tt) ∙ λ i → inl (merid (pt A) i)) j - ; (j = i0) → compPath-filler' (λ i → inl (toSusp A a i)) (push tt) k i - ; (j = i1) → inl (merid a (~ k ∨ i))}) - (compPath-filler' (sym (push tt)) (λ i → inl (merid (pt A) i)) i j) - Iso.rightInv cofibConst (inr x) = refl - Iso.rightInv cofibConst (push tt i) j = - (cong (_∙ push tt) (λ i j → inl (rCancel (merid (pt A)) i j)) - ∙ sym (lUnit _)) j i - Iso.leftInv cofibConst (inl x) = refl - Iso.leftInv cofibConst (inr x) = refl - Iso.leftInv cofibConst (push a i) j = help j i - where - help : cong ⋁→cofibConst (((λ i → inl (toSusp A a i)) ∙ push tt)) ≡ push a - help = cong-∙ ⋁→cofibConst (λ i → inl (toSusp A a i)) (push tt) - ∙ cong₂ _∙_ (cong-∙ (⋁→cofibConst ∘ inl) (merid a) (sym (merid (pt A)))) refl - ∙ sym (assoc _ _ _) - ∙ cong (push a ∙_) (lCancel (push (pt A))) - ∙ sym (rUnit _) - - cofibConst' : (f : A →∙ B) → ((x : fst A) → fst f x ≡ pt B) - → Iso (cofib (fst f)) ((Susp∙ (fst A) ⋁ B)) - cofibConst' f d = - compIso (pushoutIso _ _ _ _ - (idEquiv _) (idEquiv _) (idEquiv _) refl (funExt d)) - cofibConst - -open import Cubical.HITs.Pushout -module _ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where - - private - open 3x3-span - inst : 3x3-span - A00 inst = Unit - A02 inst = Unit - A04 inst = Unit - A20 inst = fst A - A22 inst = Unit - A24 inst = Unit - A40 inst = fst B - A42 inst = fst B - A44 inst = Unit - f10 inst = _ - f12 inst = _ - f14 inst = _ - f30 inst = fst f - f32 inst = λ _ → pt B - f34 inst = _ - f01 inst = _ - f21 inst = λ _ → pt A - f41 inst = idfun (fst B) - f03 inst = _ - f23 inst = _ - f43 inst = _ - H11 inst = (λ _ → refl) - H13 inst = λ _ → refl - H31 inst = λ _ → sym (snd f) - H33 inst = λ _ → refl - - A□0≅cofib-f : Iso (A□0 inst) (cofib (fst f)) - A□0≅cofib-f = idIso - - A□2≅B : Iso (A□2 inst) (fst B) - A□2≅B = PushoutAlongEquiv (idEquiv _) λ _ → pt B - - A□4≅Unit : Iso (A□4 inst) Unit - A□4≅Unit = PushoutAlongEquiv (idEquiv _) λ _ → tt - - A0□≅Unit : Iso (A0□ inst) Unit - A0□≅Unit = PushoutAlongEquiv (idEquiv _) λ _ → tt - - A2□≅A : Iso (A2□ inst) (fst A) - A2□≅A = compIso (equivToIso (invEquiv (symPushout _ _))) - (PushoutAlongEquiv (idEquiv _) λ _ → pt A) - - A4□≅Unit : Iso (A4□ inst) Unit - A4□≅Unit = PushoutAlongEquiv (idEquiv _) λ _ → tt - - A□○≅cofibInr : Iso (A□○ inst) (cofib {B = cofib (fst f)} inr) - A□○≅cofibInr = compIso (invIso (equivToIso (symPushout _ _))) - (pushoutIso _ _ _ _ - (isoToEquiv A□2≅B) - (isoToEquiv A□4≅Unit) - (isoToEquiv A□0≅cofib-f) - refl (funExt λ x - → cong (A□0≅cofib-f .Iso.fun ∘ f□1 inst) - (sym (Iso.leftInv A□2≅B x)))) - - A○□≅ : Iso (A○□ inst) (Susp (typ A)) - A○□≅ = - compIso - (pushoutIso _ _ _ _ - (isoToEquiv A2□≅A) - (isoToEquiv A0□≅Unit) - (isoToEquiv A4□≅Unit) - refl refl) - PushoutSuspIsoSusp - - Iso-cofibInr-Susp : Iso (cofib {B = cofib (fst f)} inr) - (Susp (typ A)) - Iso-cofibInr-Susp = - compIso (compIso (invIso A□○≅cofibInr) - (3x3-Iso inst)) A○□≅ - ---- CW complexes --- --- Definition of (the skeleton) of an arbitrary CW complex --- New def: X 0 is empty and C (suc n) is pushout - -open import Cubical.HITs.Pushout -module _ {ℓA ℓB ℓC : Level} {A : Type ℓA} {B : A → Pointed ℓB} (C : Pointed ℓC) - (f : (⋁gen A B , inl tt) →∙ C) where - - private - open 3x3-span - inst : 3x3-span - A00 inst = A - A02 inst = Σ A (fst ∘ B) - A04 inst = Σ A (fst ∘ B) - A20 inst = A - A22 inst = A - A24 inst = Σ A (fst ∘ B) - A40 inst = Unit - A42 inst = Unit - A44 inst = fst C - f10 inst = idfun A - f12 inst = λ x → x , snd (B x) - f14 inst = idfun _ - f30 inst = λ _ → tt - f32 inst = λ _ → tt - f34 inst = fst f ∘ inr - f01 inst = fst - f21 inst = idfun A - f41 inst = λ _ → tt - f03 inst = idfun _ - f23 inst = λ x → x , snd (B x) - f43 inst = λ _ → pt C - H11 inst = λ _ → refl - H13 inst = λ _ → refl - H31 inst = λ _ → refl - H33 inst = λ x → sym (snd f) ∙ cong (fst f) (push x) - - A0□Iso : Iso (A0□ inst) A - A0□Iso = compIso (equivToIso (symPushout _ _)) - (PushoutAlongEquiv (idEquiv _) fst) - - A2□Iso : Iso (A2□ inst) (Σ A (fst ∘ B)) - A2□Iso = PushoutAlongEquiv (idEquiv A) _ - - A4□Iso : Iso (A4□ inst) (fst C) - A4□Iso = PushoutAlongEquiv (idEquiv Unit) λ _ → pt C - - A○□Iso : Iso (A○□ inst) (Pushout (fst f ∘ inr) fst) - A○□Iso = compIso (equivToIso (symPushout _ _)) - (invIso (pushoutIso _ _ _ _ - (isoToEquiv (invIso A2□Iso)) - (isoToEquiv (invIso A4□Iso)) - (isoToEquiv (invIso A0□Iso)) - refl - λ i x → push x i)) - - - - A□0Iso : Iso (A□0 inst) Unit - A□0Iso = isContr→Iso - (inr tt , λ { (inl x) → sym (push x) - ; (inr x) → refl - ; (push a i) j → push a (i ∨ ~ j)}) - (tt , (λ _ → refl)) - - A□2Iso : Iso (A□2 inst) (⋁gen A B) - A□2Iso = equivToIso (symPushout _ _) - - - A□4Iso : Iso (A□4 inst) (fst C) - A□4Iso = PushoutAlongEquiv (idEquiv _) _ - - open import Cubical.Foundations.GroupoidLaws - A□○Iso : Iso (A□○ inst) (cofib (fst f)) - A□○Iso = invIso (pushoutIso _ _ _ _ - (isoToEquiv (invIso A□2Iso)) - (isoToEquiv (invIso A□0Iso)) - (isoToEquiv (invIso A□4Iso)) - (funExt (λ { (inl x) → refl - ; (inr x) → sym (push (fst x)) ∙ refl - ; (push a i) j → (sym (push a) ∙ refl) (i ∧ j)})) - (funExt λ { (inl x) i → inr (snd f i) - ; (inr x) → sym (push x) - ; (push a i) j - → hcomp (λ k - → (λ {(i = i0) → inr (compPath-filler' (sym (snd f)) - (cong (fst f) (push a)) j (~ k)) - ; (i = i1) → push (a , snd (B a)) (~ j) - ; (j = i0) → inr (fst f (push a (~ k ∨ i)))})) - (push (a , snd (B a)) (~ i ∨ ~ j))})) - - ⋁-cofib-Iso : Iso (Pushout (fst f ∘ inr) fst) (cofib (fst f)) - ⋁-cofib-Iso = compIso (compIso (invIso A○□Iso) - (invIso (3x3-Iso inst))) - A□○Iso - - - -- connected comp open import Cubical.Homotopy.Connected open import Cubical.CW.Properties @@ -299,34 +57,6 @@ open import Cubical.Foundations.HLevels open import Cubical.Data.Nat.Order.Inductive open import Cubical.Data.Fin.Inductive.Properties as Ind -isConnectedCW→isConnectedSkel : (C : CWskel ℓ) - → (m : ℕ) (n : Fin (suc m)) - → isConnected (fst n) (realise C) - → isConnected (fst n) (C .fst m) -isConnectedCW→isConnectedSkel C m (n , p) = - isOfHLevelRetractFromIso 0 - (compIso (truncOfTruncIso n t₀) - (compIso - (mapCompIso - (subst (λ k → Iso (hLevelTrunc k (C .fst m)) - (hLevelTrunc k (realise C))) - (sym teq) - (connectedTruncIso m _ - (isConnected-CW↪∞ m C)))) - (invIso (truncOfTruncIso n t₀)))) - where - t = <ᵗ→< p - t₀ = fst t - teq : t₀ +ℕ n ≡ m - teq = cong predℕ (sym (+-suc t₀ n) ∙ snd t) - -is2ConnectedCW→is2ConnectedSkel : (C : CWskel ℓ) - → isConnected 2 (realise C) - → (n : ℕ) - → isConnected 2 (C .fst (suc (suc n))) -is2ConnectedCW→is2ConnectedSkel C c n = - isConnectedCW→isConnectedSkel C (suc (suc n)) - (2 , tt) c open import Cubical.Data.Bool open import Cubical.Data.Nat.Order.Inductive @@ -337,324 +67,6 @@ open import Cubical.Data.Sum as ⊎ open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Fin.Inductive.Properties as Ind - -FinSuc : {n : ℕ} → Iso (Fin 1 ⊎ Fin n) (Fin (suc n)) -Iso.fun (FinSuc {n = n}) = ⊎.rec (λ _ → flast) injectSuc -Iso.inv (FinSuc {n = n}) = Ind.elimFin (inl flast) inr -Iso.rightInv (FinSuc {n = n}) = - Ind.elimFin (cong (⊎.rec (λ _ → flast) injectSuc) - (Ind.elimFinβ (inl flast) inr .fst)) - λ x → cong (⊎.rec (λ _ → flast) injectSuc) - (Ind.elimFinβ (inl flast) inr .snd x) -Iso.leftInv (FinSuc {n = n}) (inl (zero , p)) = - Ind.elimFinβ (inl flast) inr .fst -Iso.leftInv (FinSuc {n = n}) (inr x) = Ind.elimFinβ (inl flast) inr .snd x - -Fin+ : {n m : ℕ} → Iso (Fin n ⊎ Fin m) (Fin (n +ℕ m)) -Iso.fun (Fin+ {n = zero} {m = m}) (inr x) = x -Iso.inv (Fin+ {n = zero} {m = m}) x = inr x -Iso.rightInv (Fin+ {n = zero} {m = m}) x = refl -Iso.leftInv (Fin+ {n = zero} {m = m}) (inr x) = refl -Fin+ {n = suc n} {m = m} = - compIso (⊎Iso (invIso FinSuc) idIso) - (compIso ⊎-assoc-Iso - (compIso (⊎Iso idIso (Fin+ {n = n} {m = m})) - FinSuc)) - -Iso-Unit-Fin : Iso Unit (Fin 1) -Iso.fun Iso-Unit-Fin tt = fzero -Iso.inv Iso-Unit-Fin (x , p) = tt -Iso.rightInv Iso-Unit-Fin (zero , p) = Σ≡Prop (λ _ → isProp<ᵗ) refl -Iso.leftInv Iso-Unit-Fin x = refl - -Iso-Bool-Fin : Iso (S₊ 0) (Fin 2) -Iso.fun Iso-Bool-Fin false = flast -Iso.fun Iso-Bool-Fin true = fzero -Iso.inv Iso-Bool-Fin (zero , p) = true -Iso.inv Iso-Bool-Fin (suc x , p) = false -Iso.rightInv Iso-Bool-Fin (zero , p) = refl -Iso.rightInv Iso-Bool-Fin (suc zero , p) = - Σ≡Prop (λ _ → isProp<ᵗ) refl -Iso.leftInv Iso-Bool-Fin false = refl -Iso.leftInv Iso-Bool-Fin true = refl - -Fin× : {n m : ℕ} → Iso (Fin n × Fin m) (Fin (n · m)) -Fin× {n = zero} {m = m} = - iso (λ {()}) (λ{()}) (λ{()}) (λ{()}) -Fin× {n = suc n} {m = m} = - compIso - (compIso - (compIso (Σ-cong-iso-fst (invIso FinSuc)) - (compIso Σ-swap-Iso - (compIso ×DistR⊎Iso - (⊎Iso (compIso - (Σ-cong-iso-snd (λ _ → invIso Iso-Unit-Fin)) rUnit×Iso) - Σ-swap-Iso)))) - (⊎Iso idIso Fin×)) - (Fin+ {n = m} {n · m}) - -DiscreteFin : ∀ {n} → Discrete (Fin n) -DiscreteFin x y with discreteℕ (fst x) (fst y) -... | yes p = yes (Σ≡Prop (λ _ → isProp<ᵗ) p) -... | no ¬p = no λ q → ¬p (cong fst q) - - -decIm : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') -decIm {A = A} {B = B} f = - (y : B) → (Σ[ x ∈ A ] f x ≡ y) - ⊎ ((x : A) → ¬ f x ≡ y) - -decImFin : ∀ {ℓ} {A : Type ℓ} - (da : Discrete A) (n : ℕ) (f : Fin n → A) - → decIm f -decImFin {A = A} da zero f y = inr λ x → ⊥.rec (¬Fin0 x) -decImFin {A = A} da (suc n) f y with da (f fzero) y -... | yes p = inl (fzero , p) -... | no ¬p with (decImFin da n (f ∘ fsuc) y) -... | inl q = inl ((fsuc (fst q)) , snd q) -... | inr q = inr (Ind.elimFin-alt ¬p q) - -Bool×Fin≡Fin : {n : ℕ} → Iso (Fin n × S₊ 0) (Fin (2 · n)) -Bool×Fin≡Fin = - compIso (compIso Σ-swap-Iso - (Σ-cong-iso-fst Iso-Bool-Fin)) (Fin× {n = 2}) - -decIm-S⁰×Fin : ∀ {ℓ} {A : Type ℓ} - (da : Discrete A) (n : ℕ) (f : Fin n × S₊ 0 → A) → decIm f -decIm-S⁰×Fin {A = A} da n = - subst (λ C → (f : C → A) → decIm f) - (isoToPath (invIso Bool×Fin≡Fin)) - (decImFin da _) - -module _ {ℓ : Level} {n : ℕ} {A : Fin n → Type ℓ} (x₀ : Fin n) - (pt : A x₀) (l : (x : Fin n) → ¬ x ≡ x₀ → A x) where - private - x = fst x₀ - p = snd x₀ - elimFinChoose : (x : _) → A x - elimFinChoose (x' , q) with (discreteℕ x x') - ... | yes p₁ = subst A (Σ≡Prop (λ _ → isProp<ᵗ) p₁) pt - ... | no ¬p = l (x' , q) λ r → ¬p (sym (cong fst r)) - - elimFinChooseβ : (elimFinChoose x₀ ≡ pt) - × ((x : _) (q : ¬ x ≡ x₀) → elimFinChoose x ≡ l x q) - fst elimFinChooseβ with (discreteℕ x x) - ... | yes p₁ = (λ j → subst A (isSetFin x₀ x₀ (Σ≡Prop (λ z → isProp<ᵗ) p₁) refl j) pt) - ∙ transportRefl pt - ... | no ¬p = ⊥.rec (¬p refl) - snd elimFinChooseβ (x' , q) with (discreteℕ x x') - ... | yes p₁ = λ q → ⊥.rec (q (Σ≡Prop (λ _ → isProp<ᵗ) (sym p₁))) - ... | no ¬p = λ s → cong (l (x' , q)) (isPropΠ (λ _ → isProp⊥) _ _) - - -pickDifferentFin : {n : ℕ} (x : Fin (suc (suc n))) → Σ[ y ∈ Fin (suc (suc n)) ] ¬ x ≡ y -pickDifferentFin (zero , p) = (1 , p) , λ q → snotz (sym (cong fst q)) -pickDifferentFin (suc x , p) = fzero , λ q → snotz (cong fst q) - -allConst? : ∀ {ℓ} {A : Type ℓ} {n : ℕ} (dis : Discrete A) - → (t : Fin n → S₊ 0 → A) - → ((x : Fin n) → t x ≡ λ _ → t x true) - ⊎ (Σ[ x ∈ Fin n ] ¬ (t x true ≡ t x false)) -allConst? {n = zero} dis t = inl λ {()} -allConst? {n = suc n} dis t - with (dis (t fzero true) (t fzero false)) - | (allConst? {n = n} dis λ x p → t (fsuc x) p) -... | yes p | inl x = - inl (Ind.elimFin-alt (funExt - (λ { false → sym p ; true → refl})) x) -... | yes p | inr x = inr (_ , (snd x)) -... | no ¬p | q = inr (_ , ¬p) - - -isSurj-α₀ : (n m : ℕ) (f : Fin n × S₊ 0 → Fin (suc (suc m))) - → isConnected 2 (Pushout f fst) - → (y : _) → Σ[ x ∈ _ ] f x ≡ y -isSurj-α₀ n m f c y with (decIm-S⁰×Fin DiscreteFin n f y) -... | inl x = x -isSurj-α₀ n m f c x₀ | inr q = ⊥.rec nope - where - pre-help : Fin (suc (suc m)) → Type - pre-help = elimFinChoose x₀ ⊥ λ _ _ → Unit - - lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit - lem = elimFinChoose x₀ - (λ s t → ⊥.rec (q _ t)) - λ s t b c → elimFinChooseβ x₀ ⊥ (λ _ _ → Unit) .snd s t - - help : Pushout f fst → Type - help (inl x) = pre-help x - help (inr x) = Unit - help (push a i) = lem (f a) a refl i - - nope : ⊥ - nope = TR.rec isProp⊥ - (λ q → transport (elimFinChooseβ x₀ ⊥ (λ _ _ → Unit) .fst) - (subst help (sym q) - (transport (sym (elimFinChooseβ x₀ ⊥ - (λ _ _ → Unit) .snd _ - (pickDifferentFin x₀ .snd ∘ sym))) tt))) - (Iso.fun (PathIdTruncIso 1) - (isContr→isProp c - ∣ inl x₀ ∣ - ∣ inl (pickDifferentFin x₀ .fst) ∣)) - -notAllLoops-α₀ : (n m : ℕ) (f : Fin n × S₊ 0 → Fin (suc (suc m))) - → isConnected 2 (Pushout f fst) - → Σ[ x ∈ Fin n ] (¬ f (x , true) ≡ f (x , false)) -notAllLoops-α₀ n m f c with (allConst? DiscreteFin (λ x y → f (x , y))) -... | inr x = x -notAllLoops-α₀ n m f c | inl q = - ⊥.rec (TR.rec isProp⊥ (λ p → subst T p tt) - (Iso.fun(PathIdTruncIso 1) - (isContr→isProp c - ∣ inl flast ∣ ∣ inl fzero ∣))) - where - inrT : Fin n → Type - inrT x with (DiscreteFin (f (x , true)) fzero) - ... | yes p = ⊥ - ... | no ¬p = Unit - - inlT : Fin (suc (suc m)) → Type - inlT (zero , p) = ⊥ - inlT (suc x , p) = Unit - - inlrT-pre : (a : _) → inlT (f (a , true)) ≡ inrT a - inlrT-pre a with ((DiscreteFin (f (a , true)) fzero)) - ... | yes p = cong inlT p - inlrT-pre s | no ¬p with (f (s , true)) - ... | zero , tt = ⊥.rec (¬p refl) - ... | suc q , t = refl - - inlrT : (a : _) → inlT (f a) ≡ inrT (fst a) - inlrT (b , false) = cong inlT (funExt⁻ (q b) false) ∙ inlrT-pre b - inlrT (b , true) = inlrT-pre b - - T : Pushout f fst → Type - T (inl x) = inlT x - T (inr x) = inrT x - T (push a i) = inlrT a i - -module _ {n : ℕ} where - swapFin : (x y : Fin n) → Fin n → Fin n - swapFin (x , xp) (y , yp) (z , zp) with (discreteℕ z x) | (discreteℕ z y) - ... | yes p | yes p₁ = z , zp - ... | yes p | no ¬p = y , yp - ... | no ¬p | yes p = x , xp - ... | no ¬p | no ¬p₁ = (z , zp) - - swapFinβₗ : (x y : Fin n) → swapFin x y x ≡ y - swapFinβₗ (x , xp) (y , yp) with (discreteℕ x x) | discreteℕ x y - ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) p₁ - ... | yes p | no ¬p = refl - ... | no ¬p | q = ⊥.rec (¬p refl) - - swapFinβᵣ : (x y : Fin n) → swapFin x y y ≡ x - swapFinβᵣ (x , xp) (y , yp) with (discreteℕ y y) | discreteℕ y x - ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) p₁ - ... | yes p | no ¬p = refl - ... | no ¬p | q = ⊥.rec (¬p refl) - - -- swapFinSwap : (x y z : Fin n) → swapFin x y z ≡ swapFin y x z - -- swapFinSwap x y z with (discreteℕ (fst z) (fst x)) | discreteℕ (fst z) (fst y) - -- ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p₁ ∙ p) - -- ... | yes p | no ¬p = {!help!} - -- where - -- help : y ≡ swapFin y x z - -- help = {!!} - -- ... | no ¬p | q = {!!} - - swapFin² : (x y z : Fin n) → swapFin x y (swapFin x y z) ≡ z - swapFin² (x , xp) (y , yp) (z , zp) with discreteℕ z x | discreteℕ z y - ... | yes p | yes p₁ = silly - where - silly : swapFin (x , xp) (y , yp) (z , zp) ≡ (z , zp) - silly with discreteℕ z x | discreteℕ z y - ... | yes p | yes p₁ = refl - ... | yes p | no ¬p = ⊥.rec (¬p p₁) - ... | no ¬p | r = ⊥.rec (¬p p) - ... | yes p | no ¬q = silly - where - silly : swapFin (x , xp) (y , yp) (y , yp) ≡ (z , zp) - silly with discreteℕ y x | discreteℕ y y - ... | yes p' | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (p' ∙ sym p) - ... | no ¬p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p) - ... | p | no ¬p = ⊥.rec (¬p refl) - ... | no ¬p | yes p = silly - where - silly : swapFin (x , xp) (y , yp) (x , xp) ≡ (z , zp) - silly with discreteℕ x y | discreteℕ x x - ... | yes p₁ | yes _ = Σ≡Prop (λ _ → isProp<ᵗ) (p₁ ∙ sym p) - ... | no ¬p | yes _ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p) - ... | s | no ¬p = ⊥.rec (¬p refl) - ... | no ¬p | no ¬q = silly - where - silly : swapFin (x , xp) (y , yp) (z , zp) ≡ (z , zp) - silly with discreteℕ z x | discreteℕ z y - ... | yes p | yes p₁ = refl - ... | yes p | no ¬b = ⊥.rec (¬p p) - ... | no ¬a | yes b = ⊥.rec (¬q b) - ... | no ¬a | no ¬b = refl - - swapFinIso : (x y : Fin n) → Iso (Fin n) (Fin n) - Iso.fun (swapFinIso x y) = swapFin x y - Iso.inv (swapFinIso x y) = swapFin x y - Iso.rightInv (swapFinIso x y) = swapFin² x y - Iso.leftInv (swapFinIso x y) = swapFin² x y - - - - -PushoutPostcompEquivIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} - {A : Type ℓ} {B : Type ℓ'} {B' : Type ℓ''} {C : Type ℓ'''} - (e2 : B' ≃ B) - (f : A → B') (g : A → C) - → Iso (Pushout (fst e2 ∘ f) g) (Pushout f g) -PushoutPostcompEquivIso {ℓ = ℓ} {ℓ'} {ℓ''} {ℓ'''} - {A = A} {B} {B'} {C} e2 f g = - compIso (PushoutLiftIso ℓ* _ _) - (compIso (pushoutIso _ _ _ _ (idEquiv _) - (compEquiv (invEquiv LiftEquiv) - (compEquiv (invEquiv e2) LiftEquiv)) - (idEquiv _) - (funExt (λ { (lift l) → cong lift (retEq e2 (f l))})) - refl) - (invIso (PushoutLiftIso ℓ* _ _))) - where - ℓ* = ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')) - -private - Pushout∘Equiv' : ∀ {ℓ ℓ' ℓ''} {A A' : Type ℓ} {B B' : Type ℓ'} {C : Type ℓ''} - (e1 : A ≃ A') (e2 : B' ≃ B) - (f : A' → B') (g : A → C) - → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) (Pushout f (g ∘ invEq e1)) - Pushout∘Equiv' {A = A} {A' = A'} {B} {B'} {C} = - EquivJ (λ A e1 → (e2 : B' ≃ B) (f : A' → B') (g : A → C) - → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) - (Pushout f (g ∘ invEq e1))) - (EquivJ (λ B' e2 → (f : A' → B') (g : A' → C) - → Iso (Pushout (fst e2 ∘ f) g) - (Pushout f g)) - λ f g → idIso) -open import Cubical.Data.List -Pushout∘Equiv : ∀ {ℓA ℓA' ℓB ℓB' ℓC} - {A : Type ℓA} {A' : Type ℓA'} {B : Type ℓB} {B' : Type ℓB'} - {C : Type ℓC} - (e1 : A ≃ A') (e2 : B' ≃ B) - (f : A' → B') (g : A → C) - → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) (Pushout f (g ∘ invEq e1)) -Pushout∘Equiv {ℓA = ℓA} {ℓA'} {ℓB} {ℓB'} {ℓC} e1 e2 f g = - compIso (pushoutIso _ _ _ _ LiftEquiv LiftEquiv LiftEquiv refl refl) - (compIso (Pushout∘Equiv' {ℓ = ℓ*} {ℓ*} {ℓ*} - (liftEq ℓ* e1) (liftEq ℓ* e2) (liftFun f) (liftFun g)) - (invIso (pushoutIso _ _ _ _ - LiftEquiv LiftEquiv (LiftEquiv {ℓ' = ℓ*}) refl refl))) - where - ℓ* = ℓ-maxList (ℓA ∷ ℓA' ∷ ℓB ∷ ℓB' ∷ ℓC ∷ []) - - liftEq : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ℓ* : Level) - → A ≃ B → Lift {j = ℓ*} A ≃ Lift {j = ℓ*} B - liftEq ℓ* e = compEquiv (invEquiv LiftEquiv) (compEquiv e LiftEquiv) - module _ {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A × Bool → Unit ⊎ B) (b₀ : B) where @@ -757,35 +169,6 @@ module _ {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} Iso.rightInv Is1 = PF→P∘'→PF Iso.leftInv Is1 = PFpushT -FinPred : ∀ {m} → Fin (suc (suc m)) → Fin (suc m) -FinPred {m = m} (zero , p) = zero , p -FinPred {m = m} (suc x , p) = x , p - -fone : ∀ {m} → Fin (suc (suc m)) -fone {m} = suc zero , tt - -module _ {m : ℕ} where - Fin→Unit⊎Fin : (x : Fin (suc m)) → Unit ⊎ Fin m - Fin→Unit⊎Fin = Ind.elimFin (inl tt) inr - - Unit⊎Fin→Fin : Unit ⊎ Fin m → Fin (suc m) - Unit⊎Fin→Fin (inl x) = flast - Unit⊎Fin→Fin (inr x) = injectSuc x - - Iso-Fin-Unit⊎Fin : Iso (Fin (suc m)) (Unit ⊎ Fin m) - Iso.fun Iso-Fin-Unit⊎Fin = Fin→Unit⊎Fin - Iso.inv Iso-Fin-Unit⊎Fin = Unit⊎Fin→Fin - Iso.rightInv Iso-Fin-Unit⊎Fin (inl x) = Ind.elimFinβ (inl tt) inr .fst - Iso.rightInv Iso-Fin-Unit⊎Fin (inr x) = Ind.elimFinβ (inl tt) inr .snd x - Iso.leftInv Iso-Fin-Unit⊎Fin = - Ind.elimFin - (cong Unit⊎Fin→Fin (Ind.elimFinβ (inl tt) inr .fst)) - λ x → (cong Unit⊎Fin→Fin (Ind.elimFinβ (inl tt) inr .snd x)) - - -≠flast→<ᵗflast : {n : ℕ} → (x : Fin (suc n)) → ¬ x ≡ flast → fst x <ᵗ n -≠flast→<ᵗflast = Ind.elimFin (λ p → ⊥.rec (p refl)) λ p _ → snd p - CW₁DataPre : (n m : ℕ) (f : Fin (suc (suc n)) × S₊ 0 → Fin (suc (suc m))) → f (flast , false) ≡ flast → (t : f (flast , true) .fst <ᵗ suc m) @@ -814,91 +197,6 @@ CW₁DataPre n m f p q = (suc n) help (inr a) false = Iso.leftInv Iso-Fin-Unit⊎Fin _ help (inr a) true = Iso.leftInv Iso-Fin-Unit⊎Fin _ -isPropFin1 : isProp (Fin 1) -isPropFin1 (zero , tt) (zero , tt) = refl - -Iso⊎→Iso : ∀ {ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} - → (f : Iso A C) - → (e : Iso (A ⊎ B) (C ⊎ D)) - → ((a : A) → Iso.fun e (inl a) ≡ inl (Iso.fun f a)) - → Iso B D -Iso⊎→Iso {A = A} {B} {C} {D} f e p = Iso' - where - ⊥-fib : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ⊎ B → Type - ⊥-fib (inl x) = ⊥ - ⊥-fib (inr x) = Unit - - module _ {ℓ ℓ' ℓ'' ℓ''' : Level} - {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} - (f : Iso A C) - (e : Iso (A ⊎ B) (C ⊎ D)) - (p : (a : A) → Iso.fun e (inl a) ≡ inl (Iso.fun f a)) where - T : (b : B) → Type _ - T b = Σ[ d' ∈ C ⊎ D ] (Iso.fun e (inr b) ≡ d') - - T-elim : ∀ {ℓ} (b : B) {P : (x : T b) → Type ℓ} - → ((d : D) (s : _) → P (inr d , s)) - → (x : _) → P x - T-elim b ind (inl x , q) = - ⊥.rec (subst ⊥-fib (sym (sym (Iso.leftInv e _) - ∙ cong (Iso.inv e) - (p _ ∙ cong inl (Iso.rightInv f x) ∙ sym q) - ∙ Iso.leftInv e _)) tt) - T-elim b ind (inr x , y) = ind x y - - e-pres-inr-help : (b : B) → T f e p b → D - e-pres-inr-help b = T-elim f e p b λ d _ → d - - p' : (a : C) → Iso.inv e (inl a) ≡ inl (Iso.inv f a) - p' c = cong (Iso.inv e ∘ inl) (sym (Iso.rightInv f c)) - ∙∙ cong (Iso.inv e) (sym (p (Iso.inv f c))) - ∙∙ Iso.leftInv e _ - - e⁻-pres-inr-help : (d : D) → T (invIso f) (invIso e) p' d → B - e⁻-pres-inr-help d = T-elim (invIso f) (invIso e) p' d λ b _ → b - - e-pres-inr : B → D - e-pres-inr b = e-pres-inr-help b (_ , refl) - - e⁻-pres-inr : D → B - e⁻-pres-inr d = e⁻-pres-inr-help d (_ , refl) - - lem1 : (b : B) (e : T f e p b) (d : _) - → e⁻-pres-inr-help (e-pres-inr-help b e) d ≡ b - lem1 b = T-elim f e p b λ d s - → T-elim (invIso f) (invIso e) p' _ - λ b' s' → invEq (_ , isEmbedding-inr _ _) - (sym s' ∙ cong (Iso.inv e) (sym s) ∙ Iso.leftInv e _) - - lem2 : (d : D) (e : T (invIso f) (invIso e) p' d ) (t : _) - → e-pres-inr-help (e⁻-pres-inr-help d e) t ≡ d - lem2 d = T-elim (invIso f) (invIso e) p' d - λ b s → T-elim f e p _ λ d' s' - → invEq (_ , isEmbedding-inr _ _) - (sym s' ∙ cong (Iso.fun e) (sym s) ∙ Iso.rightInv e _) - - Iso' : Iso B D - Iso.fun Iso' = e-pres-inr - Iso.inv Iso' = e⁻-pres-inr - Iso.rightInv Iso' x = lem2 x _ _ - Iso.leftInv Iso' x = lem1 x _ _ - -Fin≠Fin : {n m : ℕ} → ¬ (n ≡ m) → ¬ (Iso (Fin n) (Fin m)) -Fin≠Fin {n = zero} {m = zero} p = ⊥.rec (p refl) -Fin≠Fin {n = zero} {m = suc m} p q = Iso.inv q fzero .snd -Fin≠Fin {n = suc n} {m = zero} p q = Iso.fun q fzero .snd -Fin≠Fin {n = suc n} {m = suc m} p q = - Fin≠Fin {n = n} {m = m} (p ∘ cong suc) - (Iso⊎→Iso idIso help λ {(zero , tt) - → cong (Iso.inv FinSuc) (swapFinβₗ (Iso.fun q flast) flast) - ∙ Ind.elimFinβ (inl flast) inr .fst}) - where - q^ : Iso (Fin (suc n)) (Fin (suc m)) - q^ = compIso q (swapFinIso (Iso.fun q flast) flast) - - help : Iso (Fin 1 ⊎ Fin n) (Fin 1 ⊎ Fin m) - help = compIso FinSuc (compIso q^ (invIso FinSuc)) - CW₁Data₁ : (m : ℕ) (f : Fin 1 × S₊ 0 → Fin (suc (suc m))) → isConnected 2 (Pushout f fst) → Iso (Fin 1 × S₊ 0) (Fin (suc (suc m))) @@ -957,7 +255,7 @@ CW₁Data (suc zero) (suc m) f c = (invIso rUnit×Iso) (compIso (Σ-cong-iso - (invIso Iso-Bool-Fin) + (invIso Iso-Bool-Fin2) (λ _ → isContr→Iso (tt , (λ _ → refl)) (inhProp→isContr fzero isPropFin1))) Σ-swap-Iso)) @@ -1009,7 +307,6 @@ CW₁Data (suc (suc n)) m f c = (funExt (λ x → cong (FinIso2 .Iso.fun ∘ f) (sym (Iso.rightInv Fin0-iso x)))) refl - CW₁Data' : (n m : ℕ) (f : Fin n × S₊ 0 → Fin m) → isConnected 2 (Pushout f fst) → Σ[ k ∈ ℕ ] @@ -1037,179 +334,270 @@ CW₁Data' (suc n) (suc (suc m)) f c (isoToPath (CW₁Data (suc n) m f c .snd .snd)) c)) ... | (k , e) = k , compIso (CW₁Data (suc n) m f c .snd .snd) e -yieldsConnectedCWskel : (A : ℕ → Type ℓ) (n : ℕ) → Type _ -yieldsConnectedCWskel A k = - Σ[ sk ∈ yieldsCWskel A ] ((sk .fst 0 ≡ 1) × ((n : ℕ) → n <ᵗ k → sk .fst (suc n) ≡ 0)) - -yieldsConnectedCWskel' : (A : ℕ → Type ℓ) (n : ℕ) → Type _ -yieldsConnectedCWskel' A n = Σ[ sk ∈ yieldsCWskel A ] isConnected (2 +ℕ n) (realise (_ , sk)) - -connectedCWskel : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) -connectedCWskel ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel X n) - -isConnectedCW : ∀ {ℓ} (n : ℕ) → Type ℓ → Type (ℓ-suc ℓ) -isConnectedCW {ℓ = ℓ} n A = Σ[ sk ∈ connectedCWskel ℓ n ] realise (_ , (snd sk .fst)) ≃ A - -connectedCWskel' : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) -connectedCWskel' ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel' X n) - -truncCWIso : (A : CWskel ℓ) (n : ℕ) - → Iso (hLevelTrunc n (realise A)) (hLevelTrunc n (fst A n)) -truncCWIso A n = invIso (connectedTruncIso n incl (isConnected-CW↪∞ n A)) - -isConnectedColim→isConnectedSkel : - (A : CWskel ℓ) (n : ℕ) - → isConnected n (realise A) - → isConnected n (fst A n) -isConnectedColim→isConnectedSkel A n c = - isOfHLevelRetractFromIso 0 - (invIso (truncCWIso A n)) c - -module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A 0) where - private - c = snd sk+c - sk = fst sk+c - c' = isConnectedColim→isConnectedSkel (_ , sk) 2 c - - module AC = CWskel-fields (_ , sk) - - e₁ : Iso (Pushout (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) fst) (A 2) - e₁ = compIso (PushoutPostcompEquivIso (CW₁-discrete (A , sk)) (AC.α 1) fst) - (equivToIso (invEquiv (AC.e 1))) - - liftStr = CW₁Data' _ _ (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) - (isConnectedRetractFromIso 2 e₁ c') - - collapse₁card : ℕ → ℕ - collapse₁card zero = 1 - collapse₁card (suc zero) = fst liftStr - collapse₁card (suc (suc x)) = AC.card (suc (suc x)) - - collapse₁CWskel : ℕ → Type _ - collapse₁CWskel zero = Lift ⊥ - collapse₁CWskel (suc zero) = Lift (Fin 1) - collapse₁CWskel (suc (suc n)) = A (suc (suc n)) - - collapse₁α : (n : ℕ) → Fin (collapse₁card n) × S⁻ n → collapse₁CWskel n - collapse₁α (suc zero) (x , p) = lift fzero - collapse₁α (suc (suc n)) = AC.α (2 +ℕ n) - - connectedCWskel→ : yieldsConnectedCWskel collapse₁CWskel 0 - fst (fst connectedCWskel→) = collapse₁card - fst (snd (fst connectedCWskel→)) = collapse₁α - fst (snd (snd (fst connectedCWskel→))) = λ() - snd (snd (snd (fst connectedCWskel→))) zero = - isContr→Equiv (isOfHLevelLift 0 (inhProp→isContr fzero isPropFin1)) - ((inr fzero) - , (λ { (inr x) j → inr (isPropFin1 fzero x j) })) - snd (snd (snd (fst connectedCWskel→))) (suc zero) = - compEquiv (invEquiv (isoToEquiv e₁)) - (compEquiv (isoToEquiv (snd liftStr)) - (isoToEquiv (pushoutIso _ _ _ _ - (idEquiv _) LiftEquiv (idEquiv _) - (funExt cohₗ) (funExt cohᵣ)))) - where - -- Agda complains if these proofs are inlined... - cohₗ : (x : _) → collapse₁α 1 x ≡ collapse₁α 1 x - cohₗ (x , p) = refl - - cohᵣ : (x : Fin (fst liftStr) × Bool) → fst x ≡ fst x - cohᵣ x = refl - snd (snd (snd (fst connectedCWskel→))) (suc (suc n)) = AC.e (suc (suc n)) - snd connectedCWskel→ = refl , (λ _ → λ ()) - - collapse₁Equiv : (n : ℕ) - → realise (_ , sk) ≃ realise (_ , connectedCWskel→ .fst) - collapse₁Equiv n = - compEquiv - (isoToEquiv (Iso-SeqColim→SeqColimShift _ 3)) - (compEquiv (pathToEquiv (cong SeqColim - (cong₂ sequence (λ _ m → A (3 +ℕ m)) - λ _ {z} → CW↪ (A , sk) (suc (suc (suc z)))))) - (invEquiv (isoToEquiv (Iso-SeqColim→SeqColimShift _ 3)))) - - -connectedCWskel'→connectedCWskel : ∀ {ℓ} - → Σ[ t ∈ connectedCWskel' ℓ 0 ] - (Σ[ dim ∈ ℕ ] - ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim)))) - → Σ[ t ∈ connectedCWskel ℓ 0 ] - Σ[ dim ∈ ℕ ] - ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim))) -fst (connectedCWskel'→connectedCWskel ((A , sk) , conv)) = - _ , connectedCWskel→ A ((sk .fst) , (sk .snd)) .fst , refl , (λ _ → λ()) -fst (snd (connectedCWskel'→connectedCWskel ((A , sk) , conv))) = suc (fst conv) -snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , zero , T))) k = - ⊥.rec (TR.rec (λ()) - (λ a → TR.rec (λ()) - (λ t → CW₀-empty (_ , fst sk) (invEq (_ , T 0) (t .fst))) - (isConnected-CW↪∞ 1 (_ , fst sk) a .fst)) (sk .snd .fst)) -snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , (suc dim) , T))) k = - transport (λ i → isEquiv (CW↪ (collapse₁CWskel A sk , connectedCWskel→ A sk .fst) - (h i))) - (transport (λ i → isEquiv (CW↪ (A , sk .fst) (suc (+-suc k dim i)))) - (T (suc k))) - where - h = cong suc (sym (+-suc k dim)) ∙ sym (+-suc k (suc dim)) - -yieldsGoodCWskel : {ℓ : Level} (A₊₂ : ℕ → Pointed ℓ) → Type _ -yieldsGoodCWskel {ℓ = ℓ} A = - Σ[ f₊₁ ∈ (ℕ → ℕ) ] - Σ[ fin ∈ (A 0) .fst ≃ Fin 1 ] - Σ[ α ∈ ((n : ℕ) → SphereBouquet∙ n (Fin (f₊₁ n)) →∙ A n) ] - ((n : ℕ) → cofib (α n .fst) , inl tt ≃∙ A (suc n)) - -GoodCWskelSeq : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Sequence ℓ -obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) zero = Lift ⊥ -obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) (suc n) = fst (A n) -Sequence.map (GoodCWskelSeq {A = A} (f , fin , α , sq)) {n = suc n} x = fst (fst (sq n)) (inr x) - -realiseGood∙ : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Pointed ℓ -realiseGood∙ {A = A} S = SeqColim (GoodCWskelSeq S) , incl {n = 1} (snd (A 0)) - -yieldsFinGoodCWskel : {ℓ : Level} (dim : ℕ) (A₊₂ : ℕ → Pointed ℓ) → Type _ -yieldsFinGoodCWskel {ℓ = ℓ} dim A = - Σ[ A ∈ yieldsGoodCWskel A ] converges (GoodCWskelSeq A) dim - -GoodCWskel : (ℓ : Level) → Type (ℓ-suc ℓ) -GoodCWskel ℓ = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsGoodCWskel A - -FinGoodCWskel : (ℓ : Level) (dim : ℕ) → Type (ℓ-suc ℓ) -FinGoodCWskel ℓ dim = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsFinGoodCWskel dim A - -isGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) -isGoodCWExpl {ℓ} A = Σ[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A - -isFinGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) -isFinGoodCWExpl {ℓ} A = - Σ[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] realiseGood∙ (fst (snd sk)) ≃∙ A - -isGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) -isGoodCW {ℓ} A = ∃[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A - -isFinGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) -isFinGoodCW {ℓ} A = - ∃[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] (realiseGood∙ (fst (snd sk)) ≃∙ A) +-- module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A 0) where +-- private +-- c = snd sk+c +-- sk = fst sk+c +-- c' = isConnectedColim→isConnectedSkel (_ , sk) 2 c + +-- module AC = CWskel-fields (_ , sk) + +-- e₁ : Iso (Pushout (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) fst) (A 2) +-- e₁ = compIso (PushoutCompEquivIso (idEquiv _) (CW₁-discrete (A , sk)) (AC.α 1) fst) +-- (equivToIso (invEquiv (AC.e 1))) + +-- liftStr = CW₁Data' _ _ (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) +-- (isConnectedRetractFromIso 2 e₁ c') + +-- collapse₁card : ℕ → ℕ +-- collapse₁card zero = 1 +-- collapse₁card (suc zero) = fst liftStr +-- collapse₁card (suc (suc x)) = AC.card (suc (suc x)) + +-- collapse₁CWskel : ℕ → Type _ +-- collapse₁CWskel zero = Lift ⊥ +-- collapse₁CWskel (suc zero) = Lift (Fin 1) +-- collapse₁CWskel (suc (suc n)) = A (suc (suc n)) + +-- collapse₁α : (n : ℕ) → Fin (collapse₁card n) × S⁻ n → collapse₁CWskel n +-- collapse₁α (suc zero) (x , p) = lift fzero +-- collapse₁α (suc (suc n)) = AC.α (2 +ℕ n) + +-- connectedCWskel→ : yieldsConnectedCWskel collapse₁CWskel 0 +-- fst (fst connectedCWskel→) = collapse₁card +-- fst (snd (fst connectedCWskel→)) = collapse₁α +-- fst (snd (snd (fst connectedCWskel→))) = λ() +-- snd (snd (snd (fst connectedCWskel→))) zero = +-- isContr→Equiv (isOfHLevelLift 0 (inhProp→isContr fzero isPropFin1)) +-- ((inr fzero) +-- , (λ { (inr x) j → inr (isPropFin1 fzero x j) })) +-- snd (snd (snd (fst connectedCWskel→))) (suc zero) = +-- compEquiv (invEquiv (isoToEquiv e₁)) +-- (compEquiv (isoToEquiv (snd liftStr)) +-- (isoToEquiv (pushoutIso _ _ _ _ +-- (idEquiv _) LiftEquiv (idEquiv _) +-- (funExt cohₗ) (funExt cohᵣ)))) +-- where +-- -- Agda complains if these proofs are inlined... +-- cohₗ : (x : _) → collapse₁α 1 x ≡ collapse₁α 1 x +-- cohₗ (x , p) = refl + +-- cohᵣ : (x : Fin (fst liftStr) × Bool) → fst x ≡ fst x +-- cohᵣ x = refl +-- snd (snd (snd (fst connectedCWskel→))) (suc (suc n)) = AC.e (suc (suc n)) +-- snd connectedCWskel→ = refl , (λ _ → λ ()) + +-- collapse₁Equiv : (n : ℕ) +-- → realise (_ , sk) ≃ realise (_ , connectedCWskel→ .fst) +-- collapse₁Equiv n = +-- compEquiv +-- (isoToEquiv (Iso-SeqColim→SeqColimShift _ 3)) +-- (compEquiv (pathToEquiv (cong SeqColim +-- (cong₂ sequence (λ _ m → A (3 +ℕ m)) +-- λ _ {z} → CW↪ (A , sk) (suc (suc (suc z)))))) +-- (invEquiv (isoToEquiv (Iso-SeqColim→SeqColimShift _ 3)))) + + +-- connectedCWskel'→connectedCWskel : ∀ {ℓ} +-- → Σ[ t ∈ connectedCWskel' ℓ 0 ] +-- (Σ[ dim ∈ ℕ ] +-- ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim)))) +-- → Σ[ t ∈ connectedCWskel ℓ 0 ] +-- Σ[ dim ∈ ℕ ] +-- ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim))) +-- fst (connectedCWskel'→connectedCWskel ((A , sk) , conv)) = +-- _ , connectedCWskel→ A ((sk .fst) , (sk .snd)) .fst , refl , (λ _ → λ()) +-- fst (snd (connectedCWskel'→connectedCWskel ((A , sk) , conv))) = suc (fst conv) +-- snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , zero , T))) k = +-- ⊥.rec (TR.rec (λ()) +-- (λ a → TR.rec (λ()) +-- (λ t → CW₀-empty (_ , fst sk) (invEq (_ , T 0) (t .fst))) +-- (isConnected-CW↪∞ 1 (_ , fst sk) a .fst)) (sk .snd .fst)) +-- snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , (suc dim) , T))) k = +-- transport (λ i → isEquiv (CW↪ (collapse₁CWskel A sk , connectedCWskel→ A sk .fst) +-- (h i))) +-- (transport (λ i → isEquiv (CW↪ (A , sk .fst) (suc (+-suc k dim i)))) +-- (T (suc k))) +-- where +-- h = cong suc (sym (+-suc k dim)) ∙ sym (+-suc k (suc dim)) + +-- yieldsGoodCWskel : {ℓ : Level} (A₊₂ : ℕ → Pointed ℓ) → Type _ +-- yieldsGoodCWskel {ℓ = ℓ} A = +-- Σ[ f₊₁ ∈ (ℕ → ℕ) ] +-- Σ[ fin ∈ (A 0) .fst ≃ Fin 1 ] +-- Σ[ α ∈ ((n : ℕ) → SphereBouquet∙ n (Fin (f₊₁ n)) →∙ A n) ] +-- ((n : ℕ) → cofib (α n .fst) , inl tt ≃∙ A (suc n)) + +-- GoodCWskelSeq : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Sequence ℓ +-- obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) zero = Lift ⊥ +-- obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) (suc n) = fst (A n) +-- Sequence.map (GoodCWskelSeq {A = A} (f , fin , α , sq)) {n = suc n} x = fst (fst (sq n)) (inr x) + +-- realiseGood∙ : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Pointed ℓ +-- realiseGood∙ {A = A} S = SeqColim (GoodCWskelSeq S) , incl {n = 1} (snd (A 0)) + +-- yieldsFinGoodCWskel : {ℓ : Level} (dim : ℕ) (A₊₂ : ℕ → Pointed ℓ) → Type _ +-- yieldsFinGoodCWskel {ℓ = ℓ} dim A = +-- Σ[ A ∈ yieldsGoodCWskel A ] converges (GoodCWskelSeq A) dim + +-- GoodCWskel : (ℓ : Level) → Type (ℓ-suc ℓ) +-- GoodCWskel ℓ = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsGoodCWskel A + +-- FinGoodCWskel : (ℓ : Level) (dim : ℕ) → Type (ℓ-suc ℓ) +-- FinGoodCWskel ℓ dim = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsFinGoodCWskel dim A + +-- isGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) +-- isGoodCWExpl {ℓ} A = Σ[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A + +-- isFinGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) +-- isFinGoodCWExpl {ℓ} A = +-- Σ[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] realiseGood∙ (fst (snd sk)) ≃∙ A + +-- isGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) +-- isGoodCW {ℓ} A = ∃[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A + +-- isFinGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) +-- isFinGoodCW {ℓ} A = +-- ∃[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] (realiseGood∙ (fst (snd sk)) ≃∙ A) + +-- finGoodCW : (ℓ : Level) → Type (ℓ-suc ℓ) +-- finGoodCW ℓ = Σ[ A ∈ Pointed ℓ ] isFinGoodCW A + +-- ptCW : {ℓ : Level} {A : ℕ → Type ℓ} → yieldsCWskel A → A 1 +-- → (n : ℕ) → A (suc n) +-- ptCW sk a zero = a +-- ptCW sk a (suc n) = CW↪ (_ , sk) (suc n) (ptCW sk a n) + +-- -- module TWOO {ℓ : Level} (A' : ℕ → Type ℓ) (pt0 : A' 1) +-- -- (dim : ℕ) (con : isConnected 2 (A' 2)) +-- -- (C : yieldsFinCWskel dim A') +-- -- where + +-- -- open CWskel-fields (_ , fst C) +-- -- e₀ : A' 1 ≃ Fin (card 0) +-- -- e₀ = CW₁-discrete (_ , fst C) + +-- -- ptA : (n : ℕ) → A' (suc n) +-- -- ptA = ptCW (fst C) pt0 + +-- -- conA : (n : ℕ) → isConnected 2 (A' (suc n)) +-- -- conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) +-- -- (subst (isConnected 2) (sym (cong Fin cA)) +-- -- (∣ flast ∣ +-- -- , TR.elim (λ _ → isOfHLevelPath 2 +-- -- (isOfHLevelTrunc 2) _ _) +-- -- λ {(zero , tt) → refl})) +-- -- conA (suc n) = +-- -- isConnectedRetractFromIso 2 +-- -- (equivToIso (e (suc n))) +-- -- (∣ inl (ptA n) ∣ₕ +-- -- , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) +-- -- (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) +-- -- (conA' (conA n)) +-- -- λ c → conA' (conA n) _ +-- -- ∙ cong ∣_∣ₕ (push (c , ptSn n)))) +-- -- where +-- -- conA' : (conA : isConnected 2 (A' (suc n))) (c : A' (suc n)) +-- -- → Path (hLevelTrunc 2 (Pushout (α (suc n)) fst)) +-- -- ∣ inl (ptA n) ∣ₕ ∣ inl c ∣ₕ +-- -- conA' conA c = +-- -- TR.rec (isOfHLevelTrunc 2 _ _) +-- -- (λ p i → ∣ inl (p i) ∣) +-- -- (Iso.fun (PathIdTruncIso _) +-- -- (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) + + +-- -- private +-- -- funType = ((n : Fin (suc dim)) +-- -- → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) +-- -- →∙ (A' (suc (fst n)) , ptA (fst n))) ] +-- -- ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) + +-- -- mapMakerNil : ∥ funType ∥₁ +-- -- mapMakerNil = +-- -- invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) +-- -- λ n → TR.map +-- -- (λ pted → ((λ { (inl x) → ptA (fst n) +-- -- ; (inr x) → α _ x +-- -- ; (push a i) → pted a i}) +-- -- , refl) +-- -- , λ _ → refl) (help n)) +-- -- where +-- -- help : (n : Fin (suc dim)) +-- -- → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) +-- -- → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) +-- -- help n = invEq (_ , InductiveFinSatAC _ _ _) +-- -- λ m → Iso.fun (PathIdTruncIso _) +-- -- (isContr→isProp +-- -- (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ +-- -- ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) +-- -- module _ (F : funType) where +-- -- funs : (n : ℕ) → SphereBouquet∙ n (Fin (card (suc n))) +-- -- →∙ (A' (suc n) , ptA n) +-- -- funs n with (n ≟ᵗ dim) +-- -- ... | lt x = F (n , <ᵗ-trans-suc x) .fst +-- -- ... | eq x = F (n , subst (_<ᵗ suc dim) (sym x) <ᵗsucm) .fst +-- -- ... | gt x = const∙ _ _ + +-- -- funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) +-- -- ≃∙ Pushout (funs n .fst ∘ inr) (λ r → fst r) , inl (ptA n) +-- -- funEqP1 n = invEquiv (isoToEquiv (⋁-cofib-Iso _ (funs n))) , refl + +-- -- funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) +-- -- ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) +-- -- funEq n = isoToEquiv (pushoutIso _ _ _ _ +-- -- (idEquiv _) +-- -- (idEquiv _) +-- -- (idEquiv _) +-- -- (funExt (uncurry (main n))) +-- -- (λ i x → fst x)) +-- -- , λ _ → inl (ptA n) +-- -- where +-- -- main : (n : ℕ) (x : Fin (card (suc n))) (y : S₊ n) +-- -- → funs n .fst (inr (x , y)) ≡ fst (C .snd) (suc n) (x , y) +-- -- main n with (n ≟ᵗ dim) +-- -- ... | lt p = λ x y +-- -- → F (n , <ᵗ-trans-suc p) .snd (x , y) +-- -- ... | eq p = λ x y +-- -- → F (n , subst (_<ᵗ suc dim) (λ i → p (~ i)) <ᵗsucm) .snd (x , y) +-- -- ... | gt p = λ x +-- -- → ⊥.rec (¬Fin0 (subst Fin (ind (suc n) (<ᵗ-trans p <ᵗsucm)) x)) + +-- -- getGoodCWskelAux : ∥ yieldsGoodCWskel (λ n → A' (suc n) , ptA n) ∥₁ +-- -- getGoodCWskelAux = PT.map help mapMakerNil +-- -- where +-- -- help : funType → yieldsGoodCWskel (λ n → A' (suc n) , ptA n) +-- -- fst (help F) = card ∘ suc +-- -- fst (snd (help F)) = compEquiv e₀ (pathToEquiv (cong Fin cA)) +-- -- fst (snd (snd (help F))) n = funs F n +-- -- snd (snd (snd (help F))) n = +-- -- compEquiv∙ (compEquiv∙ (funEqP1 F n) (funEq F n)) +-- -- (invEquiv (e (suc n)) , refl) + + +-- module BS {ℓ : Level} (A' : ℕ → Type ℓ) +-- (dim : ℕ) +-- (C+eq : yieldsFinCWskel dim A') +-- (cA : fst (fst C+eq) 0 ≡ 1) +-- where +-- C = fst C+eq +-- ind = snd C+eq -finGoodCW : (ℓ : Level) → Type (ℓ-suc ℓ) -finGoodCW ℓ = Σ[ A ∈ Pointed ℓ ] isFinGoodCW A +-- open CWskel-fields (_ , C) +-- e₀ : A' 1 ≃ Fin (card 0) +-- e₀ = CW₁-discrete (_ , C) -ptCW : {ℓ : Level} {A : ℕ → Type ℓ} → yieldsCWskel A → A 1 - → (n : ℕ) → A (suc n) -ptCW sk a zero = a -ptCW sk a (suc n) = CW↪ (_ , sk) (suc n) (ptCW sk a n) --- module TWOO {ℓ : Level} (A' : ℕ → Type ℓ) (pt0 : A' 1) --- (dim : ℕ) (con : isConnected 2 (A' 2)) --- (C : yieldsFinCWskel dim A') --- where +-- ¬dim≡0 : ¬ (dim ≡ 0) +-- ¬dim≡0 p = CW₀-empty (_ , C) (subst A' p +-- (invEq (_ , ind 0) (subst A' (cong suc (sym p)) +-- (invEq e₀ (subst Fin (sym cA) fzero))))) --- open CWskel-fields (_ , fst C) --- e₀ : A' 1 ≃ Fin (card 0) --- e₀ = CW₁-discrete (_ , fst C) +-- pt0 : A' 1 +-- pt0 = invEq e₀ (subst Fin (sym cA) flast) -- ptA : (n : ℕ) → A' (suc n) --- ptA = ptCW (fst C) pt0 +-- ptA = ptCW C pt0 -- conA : (n : ℕ) → isConnected 2 (A' (suc n)) -- conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) @@ -1237,9 +625,8 @@ ptCW sk a (suc n) = CW↪ (_ , sk) (suc n) (ptCW sk a n) -- (Iso.fun (PathIdTruncIso _) -- (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) - -- private --- funType = ((n : Fin (suc dim)) +-- funType = ((n : Fin dim) -- → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) -- →∙ (A' (suc (fst n)) , ptA (fst n))) ] -- ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) @@ -1254,20 +641,26 @@ ptCW sk a (suc n) = CW↪ (_ , sk) (suc n) (ptCW sk a n) -- , refl) -- , λ _ → refl) (help n)) -- where --- help : (n : Fin (suc dim)) +-- help : (n : Fin dim) -- → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) -- → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) -- help n = invEq (_ , InductiveFinSatAC _ _ _) -- λ m → Iso.fun (PathIdTruncIso _) -- (isContr→isProp --- (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ --- ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) +-- (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) + -- module _ (F : funType) where --- funs : (n : ℕ) → SphereBouquet∙ n (Fin (card (suc n))) +-- card' : ℕ → ℕ +-- card' n with (n ≟ᵗ dim) +-- ... | lt x = card (suc n) +-- ... | eq x = 0 +-- ... | gt x = 0 + +-- funs : (n : ℕ) → SphereBouquet∙ n (Fin (card' n)) -- →∙ (A' (suc n) , ptA n) -- funs n with (n ≟ᵗ dim) --- ... | lt x = F (n , <ᵗ-trans-suc x) .fst --- ... | eq x = F (n , subst (_<ᵗ suc dim) (sym x) <ᵗsucm) .fst +-- ... | lt x = F (n , x) .fst +-- ... | eq x = const∙ _ _ -- ... | gt x = const∙ _ _ -- funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) @@ -1276,551 +669,424 @@ ptCW sk a (suc n) = CW↪ (_ , sk) (suc n) (ptCW sk a n) -- funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) -- ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) --- funEq n = isoToEquiv (pushoutIso _ _ _ _ +-- funEq n with (n ≟ᵗ dim) +-- ... | lt x = isoToEquiv (pushoutIso _ _ _ _ -- (idEquiv _) -- (idEquiv _) -- (idEquiv _) --- (funExt (uncurry (main n))) +-- (funExt (uncurry λ x y → F (n , _) .snd (x , y))) -- (λ i x → fst x)) -- , λ _ → inl (ptA n) +-- ... | eq x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) +-- (compEquiv ((CW↪ (_ , C) (suc n)) +-- , transport (λ i → isEquiv (CW↪ (A' , C) +-- (suc (x (~ i))))) +-- (ind 1) +-- ) +-- (e (suc n)))) , secEq (e (suc n)) _ +-- ... | gt x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) +-- (compEquiv ((CW↪ (_ , C) (suc n)) +-- , (transport (λ i → isEquiv (CW↪ (A' , C) +-- (suc ((sym (+-suc (fst (<ᵗ→< x)) dim) +-- ∙ (<ᵗ→< x .snd)) i)))) +-- (ind (suc (suc (fst (<ᵗ→< x))))))) +-- (e (suc n)))) , secEq (e (suc n)) _ + +-- goodCWmk : yieldsGoodCWskel (λ n → A' (suc n) , ptA n) +-- fst goodCWmk = card' +-- fst (snd goodCWmk) = compEquiv e₀ (pathToEquiv (cong Fin cA)) +-- fst (snd (snd goodCWmk)) = funs +-- snd (snd (snd goodCWmk)) n = +-- compEquiv∙ (compEquiv∙ (funEqP1 n) (funEq n)) +-- (invEquiv (e (suc n)) , refl) + +-- goodCWmk-converges : converges +-- (sequence (obj (GoodCWskelSeq goodCWmk)) +-- (Sequence.map (GoodCWskelSeq goodCWmk))) +-- dim +-- goodCWmk-converges zero = help dim refl -- where --- main : (n : ℕ) (x : Fin (card (suc n))) (y : S₊ n) --- → funs n .fst (inr (x , y)) ≡ fst (C .snd) (suc n) (x , y) --- main n with (n ≟ᵗ dim) --- ... | lt p = λ x y --- → F (n , <ᵗ-trans-suc p) .snd (x , y) --- ... | eq p = λ x y --- → F (n , subst (_<ᵗ suc dim) (λ i → p (~ i)) <ᵗsucm) .snd (x , y) --- ... | gt p = λ x --- → ⊥.rec (¬Fin0 (subst Fin (ind (suc n) (<ᵗ-trans p <ᵗsucm)) x)) - --- getGoodCWskelAux : ∥ yieldsGoodCWskel (λ n → A' (suc n) , ptA n) ∥₁ --- getGoodCWskelAux = PT.map help mapMakerNil --- where --- help : funType → yieldsGoodCWskel (λ n → A' (suc n) , ptA n) --- fst (help F) = card ∘ suc --- fst (snd (help F)) = compEquiv e₀ (pathToEquiv (cong Fin cA)) --- fst (snd (snd (help F))) n = funs F n --- snd (snd (snd (help F))) n = --- compEquiv∙ (compEquiv∙ (funEqP1 F n) (funEq F n)) --- (invEquiv (e (suc n)) , refl) - - -module BS {ℓ : Level} (A' : ℕ → Type ℓ) - (dim : ℕ) - (C+eq : yieldsFinCWskel dim A') - (cA : fst (fst C+eq) 0 ≡ 1) - where - C = fst C+eq - ind = snd C+eq - - open CWskel-fields (_ , C) - e₀ : A' 1 ≃ Fin (card 0) - e₀ = CW₁-discrete (_ , C) - - - ¬dim≡0 : ¬ (dim ≡ 0) - ¬dim≡0 p = CW₀-empty (_ , C) (subst A' p - (invEq (_ , ind 0) (subst A' (cong suc (sym p)) - (invEq e₀ (subst Fin (sym cA) fzero))))) - - pt0 : A' 1 - pt0 = invEq e₀ (subst Fin (sym cA) flast) - - ptA : (n : ℕ) → A' (suc n) - ptA = ptCW C pt0 - - conA : (n : ℕ) → isConnected 2 (A' (suc n)) - conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) - (subst (isConnected 2) (sym (cong Fin cA)) - (∣ flast ∣ - , TR.elim (λ _ → isOfHLevelPath 2 - (isOfHLevelTrunc 2) _ _) - λ {(zero , tt) → refl})) - conA (suc n) = - isConnectedRetractFromIso 2 - (equivToIso (e (suc n))) - (∣ inl (ptA n) ∣ₕ - , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) - (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) - (conA' (conA n)) - λ c → conA' (conA n) _ - ∙ cong ∣_∣ₕ (push (c , ptSn n)))) - where - conA' : (conA : isConnected 2 (A' (suc n))) (c : A' (suc n)) - → Path (hLevelTrunc 2 (Pushout (α (suc n)) fst)) - ∣ inl (ptA n) ∣ₕ ∣ inl c ∣ₕ - conA' conA c = - TR.rec (isOfHLevelTrunc 2 _ _) - (λ p i → ∣ inl (p i) ∣) - (Iso.fun (PathIdTruncIso _) - (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) - - private - funType = ((n : Fin dim) - → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) - →∙ (A' (suc (fst n)) , ptA (fst n))) ] - ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) - - mapMakerNil : ∥ funType ∥₁ - mapMakerNil = - invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) - λ n → TR.map - (λ pted → ((λ { (inl x) → ptA (fst n) - ; (inr x) → α _ x - ; (push a i) → pted a i}) - , refl) - , λ _ → refl) (help n)) - where - help : (n : Fin dim) - → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) - → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) - help n = invEq (_ , InductiveFinSatAC _ _ _) - λ m → Iso.fun (PathIdTruncIso _) - (isContr→isProp - (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) - - module _ (F : funType) where - card' : ℕ → ℕ - card' n with (n ≟ᵗ dim) - ... | lt x = card (suc n) - ... | eq x = 0 - ... | gt x = 0 - - funs : (n : ℕ) → SphereBouquet∙ n (Fin (card' n)) - →∙ (A' (suc n) , ptA n) - funs n with (n ≟ᵗ dim) - ... | lt x = F (n , x) .fst - ... | eq x = const∙ _ _ - ... | gt x = const∙ _ _ - - funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) - ≃∙ Pushout (funs n .fst ∘ inr) (λ r → fst r) , inl (ptA n) - funEqP1 n = invEquiv (isoToEquiv (⋁-cofib-Iso _ (funs n))) , refl - - funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) - ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) - funEq n with (n ≟ᵗ dim) - ... | lt x = isoToEquiv (pushoutIso _ _ _ _ - (idEquiv _) - (idEquiv _) - (idEquiv _) - (funExt (uncurry λ x y → F (n , _) .snd (x , y))) - (λ i x → fst x)) - , λ _ → inl (ptA n) - ... | eq x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) - (compEquiv ((CW↪ (_ , C) (suc n)) - , transport (λ i → isEquiv (CW↪ (A' , C) - (suc (x (~ i))))) - (ind 1) - ) - (e (suc n)))) , secEq (e (suc n)) _ - ... | gt x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) - (compEquiv ((CW↪ (_ , C) (suc n)) - , (transport (λ i → isEquiv (CW↪ (A' , C) - (suc ((sym (+-suc (fst (<ᵗ→< x)) dim) - ∙ (<ᵗ→< x .snd)) i)))) - (ind (suc (suc (fst (<ᵗ→< x))))))) - (e (suc n)))) , secEq (e (suc n)) _ - - goodCWmk : yieldsGoodCWskel (λ n → A' (suc n) , ptA n) - fst goodCWmk = card' - fst (snd goodCWmk) = compEquiv e₀ (pathToEquiv (cong Fin cA)) - fst (snd (snd goodCWmk)) = funs - snd (snd (snd goodCWmk)) n = - compEquiv∙ (compEquiv∙ (funEqP1 n) (funEq n)) - (invEquiv (e (suc n)) , refl) - - goodCWmk-converges : converges - (sequence (obj (GoodCWskelSeq goodCWmk)) - (Sequence.map (GoodCWskelSeq goodCWmk))) - dim - goodCWmk-converges zero = help dim refl - where - help : (x : _) (p : dim ≡ x) → isEquiv (Sequence.map (GoodCWskelSeq goodCWmk) {x}) - help zero p = ⊥.rec (¬dim≡0 p) - help (suc x) p with (x ≟ᵗ dim) - ... | lt x₁ = transport (λ i → isEquiv λ z → CW↪ (A' , C) (p i) z) (ind zero) - ... | eq x₁ = ⊥.rec (¬m Date: Tue, 4 Jun 2024 02:59:48 +0200 Subject: [PATCH 48/73] cleaning --- Cubical/CW/Connected.agda | 1782 +++++++++++++++++++++++++++++++++++++ 1 file changed, 1782 insertions(+) create mode 100644 Cubical/CW/Connected.agda diff --git a/Cubical/CW/Connected.agda b/Cubical/CW/Connected.agda new file mode 100644 index 0000000000..fe53f1ed45 --- /dev/null +++ b/Cubical/CW/Connected.agda @@ -0,0 +1,1782 @@ +{-# OPTIONS --safe --lossy-unification #-} + +module Cubical.CW.Connected where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Sigma +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence + +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Wedge +open import Cubical.HITs.Pushout + +open import Cubical.Axiom.Choice + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + +open import Cubical.Relation.Nullary + +open import Cubical.CW.Base + + +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Foundations.Equiv + +-- connected comp +open import Cubical.Homotopy.Connected +open import Cubical.CW.Properties +open import Cubical.HITs.Truncation as TR +open import Cubical.Foundations.HLevels +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Fin.Inductive.Properties as Ind + + +open import Cubical.Data.Bool +open import Cubical.Data.Nat.Order.Inductive +-- open import Cubical.Data.Dec + +open import Cubical.Relation.Nullary.Base +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Data.Fin.Inductive.Properties as Ind + +open Sequence + +open import Cubical.Foundations.Equiv.HalfAdjoint + +private + variable + ℓ ℓ' ℓ'' : Level + +------ Definitions ------ + +-- An n-connected CW complex has one 0-cell and no m Date: Tue, 4 Jun 2024 14:12:57 +0200 Subject: [PATCH 49/73] connected done? --- Cubical/CW/Connected.agda | 763 +++++++++++++++++++++----------------- 1 file changed, 426 insertions(+), 337 deletions(-) diff --git a/Cubical/CW/Connected.agda b/Cubical/CW/Connected.agda index fe53f1ed45..02d60ede6e 100644 --- a/Cubical/CW/Connected.agda +++ b/Cubical/CW/Connected.agda @@ -7,6 +7,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order @@ -16,6 +17,7 @@ open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma open import Cubical.Data.Sequence open import Cubical.Data.FinSequence +open import Cubical.Data.Unit open import Cubical.HITs.Sn open import Cubical.HITs.Pushout @@ -632,7 +634,8 @@ makeConnectedCW zero {C = C} (cwsk , e) cA = ∣ (_ , (M .fst , refl , λ p → λ())) , compEquiv (invEquiv (collapse₁Equiv _ _ 0)) (invEquiv e) ∣₁ where - M = connectedCWskel→ (cwsk .fst) (snd cwsk , subst (isConnected 2) (ua e) cA) + M = connectedCWskel→ (cwsk .fst) + (snd cwsk , subst (isConnected 2) (ua e) cA) makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with (makeConnectedCW n (cwsk , eqv) (isConnectedSubtr (suc (suc n)) 1 cA)) ... | s = PT.rec squash₁ @@ -641,10 +644,10 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with (λ α-pt2* → PT.rec squash₁ (λ vanish* - → PT.map (uncurry (buildCW s α-pt* α-pt2* vanish*)) - (mainF s α-pt* α-pt2* vanish*)) + → PT.map (uncurry (isConnectedCW-C' s α-pt* α-pt2* vanish*)) + (∃Improvedβ s α-pt* α-pt2* vanish*)) (sphereBouquetVanish s α-pt* α-pt2* - (const* s α-pt* α-pt2*))) + (const→C4+n s α-pt* α-pt2*))) (α-pt∙₂ s α-pt*)) (α'2+n∙ s)) s where module _ (ind : isConnectedCW n C) where @@ -655,26 +658,31 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with module _ where C* = ind .fst .fst - C1+n = C* (suc n) - C2+n = C* (suc (suc n)) - C3+n = C* (suc (suc (suc n))) - C4+n = C* (suc (suc (suc (suc n)))) + 1+n = suc n + 2+n = suc 1+n + 3+n = suc 2+n + 4+n = suc 3+n - α2+n = α (suc (suc n)) + C1+n = C* 1+n + C2+n = C* 2+n + C3+n = C* 3+n + C4+n = C* 4+n - isConnectedC4+n : isConnected (3 +ℕ n) C4+n + α2+n = α 2+n + + isConnectedC4+n : isConnected 3+n C4+n isConnectedC4+n = isConnectedCW→isConnectedSkel - (_ , ind .fst .snd .fst) (suc (suc (suc (suc n)))) - (suc (suc (suc n)) , <ᵗ-trans <ᵗsucm <ᵗsucm) - (subst (isConnected (3 +ℕ n)) (ua (invEquiv (ind .snd))) + (_ , ind .fst .snd .fst) 4+n + (suc 2+n , <ᵗ-trans <ᵗsucm <ᵗsucm) + (subst (isConnected 3+n) (ua (invEquiv (ind .snd))) cA) - isConnected3+n : isConnected (2 +ℕ n) C3+n + isConnected3+n : isConnected 2+n C3+n isConnected3+n = isConnectedCW→isConnectedSkel - (_ , ind .fst .snd .fst) (suc (suc (suc n))) - (suc (suc n) , <ᵗ-trans <ᵗsucm <ᵗsucm) - (subst (isConnected (2 +ℕ n)) (ua (invEquiv (ind .snd))) - (isConnectedSubtr (suc (suc n)) 1 cA)) + (_ , ind .fst .snd .fst) 3+n + (2+n , <ᵗ-trans <ᵗsucm <ᵗsucm) + (subst (isConnected 2+n) (ua (invEquiv (ind .snd))) + (isConnectedSubtr 2+n 1 cA)) -- C₁₊ₙ is trivial Iso-C1+n-Fin1 : Iso C1+n (Fin 1) @@ -701,13 +709,13 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with sphereBouquetSuspIso) -- We rewrite α along this iso. - α'2+n : A (suc (suc n)) × S₊ (suc n) → SphereBouquet (suc n) (A (suc n)) - α'2+n = Iso.fun Iso-C2+n-SphereBouquet ∘ α (suc (suc n)) + α'2+n : A 2+n × S₊ (suc n) → SphereBouquet (suc n) (A (suc n)) + α'2+n = Iso.fun Iso-C2+n-SphereBouquet ∘ α 2+n opaque - Iso-C3+n-Pushoutα'2+n : Iso (C* (suc (suc (suc n)))) + Iso-C3+n-Pushoutα'2+n : Iso (C* 3+n) (Pushout α'2+n fst) - Iso-C3+n-Pushoutα'2+n = compIso (equivToIso (e (suc (suc n)))) + Iso-C3+n-Pushoutα'2+n = compIso (equivToIso (e 2+n)) (pushoutIso _ _ _ _ (idEquiv _) (isoToEquiv Iso-C2+n-SphereBouquet) (idEquiv _) @@ -725,7 +733,7 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with module _ (α-pt∙ : (x : _) → α'2+n (x , ptSn (suc n)) ≡ inl tt) where -- Doing so allows us to lift it to a map of sphere bouquets - α∙ : SphereBouquet∙ (suc n) (A (suc (suc n))) + α∙ : SphereBouquet∙ (suc n) (A 2+n) →∙ SphereBouquet∙ (suc n) (A (suc n)) fst α∙ (inl x) = inl tt fst α∙ (inr x) = α'2+n x @@ -739,29 +747,29 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with Iso-C3+n-cofibα = compIso Iso-C3+n-Pushoutα'2+n (⋁-cofib-Iso _ α∙) -- This iso is also merely pointed: - α-pt∙₂ : ∥ ((x : A (3 +ℕ n)) - → Iso.fun Iso-C3+n-cofibα (α (3 +ℕ n) (x , north)) ≡ inl tt) ∥₁ + α-pt∙₂ : ∥ ((x : A 3+n) + → Iso.fun Iso-C3+n-cofibα (α 3+n (x , north)) ≡ inl tt) ∥₁ α-pt∙₂ = invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) λ a → isConnectedPath 1 (isConnectedRetractFromIso 2 (invIso Iso-C3+n-cofibα) (isConnectedSubtr' n 2 isConnected3+n) ) _ _ .fst) -- But again, let us assume it is... - module _ (α-pt∙₂ : (x : A (3 +ℕ n)) - → Iso.fun Iso-C3+n-cofibα (α (3 +ℕ n) (x , north)) ≡ inl tt) where + module _ (α-pt∙₂ : (x : A 3+n) + → Iso.fun Iso-C3+n-cofibα (α 3+n (x , north)) ≡ inl tt) where -- Doing so, we can lift α yet again this time to a map -- α↑ : ⋁S²⁺ⁿ → cofib α - α↑ : SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n)))) + α↑ : SphereBouquet∙ 2+n (A 3+n) →∙ (cofib (fst α∙) , inl tt) fst α↑ (inl x) = inl tt - fst α↑ (inr x) = Iso.fun Iso-C3+n-cofibα (α (3 +ℕ n) x) + fst α↑ (inr x) = Iso.fun Iso-C3+n-cofibα (α 3+n x) fst α↑ (push a i) = α-pt∙₂ a (~ i) snd α↑ = refl -- The cofibre of this map is C₄₊ₙ Iso-C4+n-cofibα↑ : Iso (C* (4 +ℕ n)) (cofib (fst α↑)) - Iso-C4+n-cofibα↑ = compIso (equivToIso (e (3 +ℕ n))) + Iso-C4+n-cofibα↑ = compIso (equivToIso (e 3+n)) (compIso (pushoutIso _ _ _ _ (idEquiv _) @@ -770,10 +778,11 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with (⋁-cofib-Iso _ α↑)) + opaque Iso-cofibαinr-SphereBouquet : Iso (Pushout {B = cofib (fst α∙)} inr (λ _ → tt)) - (SphereBouquet (suc (suc n)) (A (suc (suc n)))) + (SphereBouquet 2+n (A 2+n)) Iso-cofibαinr-SphereBouquet = compIso (equivToIso (symPushout _ _)) (compIso (Iso-cofibInr-Susp α∙) sphereBouquetSuspIso) @@ -782,82 +791,105 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with Iso.fun Iso-cofibαinr-SphereBouquet (inl (inl tt)) ≡ inl tt Iso-cofibαinr-SphereBouquet∙ = refl + -- composing the above iso with α↑ lets us define a map of sphere bouquets: + β : SphereBouquet 2+n (A 3+n) + → SphereBouquet 2+n (A 2+n) + β = (Iso.fun Iso-cofibαinr-SphereBouquet ∘ inl) ∘ fst α↑ + + β∙ : SphereBouquet∙ 2+n (A 3+n) + →∙ SphereBouquet∙ 2+n (A 2+n) + fst β∙ = β + snd β∙ = Iso-cofibαinr-SphereBouquet∙ + + -- The goal now: show that the cofibre of β is C₄₊ₙ ⋁ S²⁺ⁿ + C⋆ = Iso-C4+n-cofibα↑ .Iso.inv (inl tt) + + -- Abbreviation of C₄₊ₙ ⋁ S²⁺ⁿ: + C₄₊ₙ⋁SphereBouquet = (C4+n , C⋆) + ⋁ SphereBouquet∙ 2+n (A 2+n) + + + -- Intermediate alternative definition of C₄₊ₙ ⋁ S²⁺ⁿ: + C₄₊ₙ⋁SphereBouquetAsPushout = + Pushout (Iso.inv Iso-C4+n-cofibα↑ ∘ inr) + (λ x → Iso.fun Iso-cofibαinr-SphereBouquet (inl x)) + + -- We geta a map ⋁Sⁿ⁺¹ → C4+n. It is merely constant for + -- connectedness reasons and its cofibre is + -- C₄₊ₙ⋁SphereBouquetAsPushout by a simple pasting argument. + const→C4+n : SphereBouquet (suc n) (A (suc n)) → C4+n + const→C4+n x = Iso.inv Iso-C4+n-cofibα↑ (inr (inr x)) + + -- Settting up the pushout-pastings: open commSquare open 3-span - PT : commSquare - A0 (sp PT) = cofib (fst α∙) - A2 (sp PT) = SphereBouquet (suc (suc n)) (A (suc (suc (suc n)))) - A4 (sp PT) = Unit - f1 (sp PT) = fst α↑ - f3 (sp PT) = λ _ → tt - P PT = cofib (fst α↑) - inlP PT = inr - inrP PT = inl - comm PT = funExt λ x → sym (push x) - - PTPush : PushoutSquare - fst PTPush = PT - snd PTPush = + commSquare₁ : commSquare + A0 (sp commSquare₁) = cofib (fst α∙) + A2 (sp commSquare₁) = SphereBouquet 2+n (A 3+n) + A4 (sp commSquare₁) = Unit + f1 (sp commSquare₁) = fst α↑ + f3 (sp commSquare₁) = λ _ → tt + P commSquare₁ = cofib (fst α↑) + inlP commSquare₁ = inr + inrP commSquare₁ = inl + comm commSquare₁ = funExt λ x → sym (push x) + + pushoutSquare₁ : PushoutSquare + fst pushoutSquare₁ = commSquare₁ + snd pushoutSquare₁ = subst isEquiv (funExt (λ { (inl x) → refl ; (inr x) → refl ; (push a i) → refl})) (symPushout _ _ .snd) - - PL : commSquare - A0 (sp PL) = cofib (fst α∙) - A2 (sp PL) = SphereBouquet (suc n) (A (suc n)) - A4 (sp PL) = Unit - f1 (sp PL) = inr - f3 (sp PL) = _ - P PL = SphereBouquet (suc (suc n)) (A (suc (suc n))) - inlP PL x = Iso.fun Iso-cofibαinr-SphereBouquet (inl x) - inrP PL _ = Iso.fun Iso-cofibαinr-SphereBouquet (inr tt) - comm PL = + commSquare₂ : commSquare + A0 (sp commSquare₂) = cofib (fst α∙) + A2 (sp commSquare₂) = SphereBouquet (suc n) (A (suc n)) + A4 (sp commSquare₂) = Unit + f1 (sp commSquare₂) = inr + f3 (sp commSquare₂) = _ + P commSquare₂ = SphereBouquet 2+n (A 2+n) + inlP commSquare₂ x = Iso.fun Iso-cofibαinr-SphereBouquet (inl x) + inrP commSquare₂ _ = Iso.fun Iso-cofibαinr-SphereBouquet (inr tt) + comm commSquare₂ = funExt λ x → cong (Iso.fun Iso-cofibαinr-SphereBouquet) (push x) - PLPush : PushoutSquare - fst PLPush = PL - snd PLPush = + pushoutSquare₂ : PushoutSquare + fst pushoutSquare₂ = commSquare₂ + snd pushoutSquare₂ = subst isEquiv (funExt coh) (isoToIsEquiv Iso-cofibαinr-SphereBouquet) where coh : (x : _) → Iso.fun Iso-cofibαinr-SphereBouquet x - ≡ Pushout→commSquare PL x + ≡ Pushout→commSquare commSquare₂ x coh (inl x) = refl coh (inr x) = refl coh (push x i₁) = refl - C∨PlaceHolder = - Pushout (Iso.inv Iso-C4+n-cofibα↑ ∘ inr) - (λ x → Iso.fun Iso-cofibαinr-SphereBouquet (inl x)) - - C∨ = (C4+n , Iso.inv Iso-C4+n-cofibα↑ (inl tt)) - ⋁ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))) - - module L→R = PushoutPasteDown PLPush {B = C∨PlaceHolder} - (Iso.inv Iso-C4+n-cofibα↑ ∘ inr) inl inr (funExt push) - - isPushoutTot = - L→R.isPushoutBottomSquare→isPushoutTotSquare - (subst isEquiv (funExt (λ { (inl x) → refl - ; (inr x) → refl - ; (push a i) → refl})) - (idIsEquiv _)) - - const* : SphereBouquet (suc n) (A (suc n)) → C4+n - const* = λ x → Iso.inv Iso-C4+n-cofibα↑ (inr (inr x)) - - C⋆ = Iso-C4+n-cofibα↑ .Iso.inv (inl tt) + module L→R = + PushoutPasteDown pushoutSquare₂ {B = C₄₊ₙ⋁SphereBouquetAsPushout} + (Iso.inv Iso-C4+n-cofibα↑ ∘ inr) inl inr (funExt push) - sphereVanish : (f : S₊ (suc n) → C4+n) - → ∥ ((x : S₊ (suc n)) → f x ≡ C⋆) ∥₁ - sphereVanish f = sphereToTrunc (suc n) - λ x → isConnectedPath (suc (suc n)) isConnectedC4+n _ _ .fst - - sphereBouquetVanish : ∀ {k : ℕ} (f : SphereBouquet (suc n) (Fin k) → C4+n) + -- First equivalence + C₄₊ₙ⋁SphereBouquetAsPushout≃cofib-const : + C₄₊ₙ⋁SphereBouquetAsPushout ≃ cofib const→C4+n + C₄₊ₙ⋁SphereBouquetAsPushout≃cofib-const = + compEquiv + (invEquiv (_ , isPushoutTot)) + (symPushout _ _) + where + isPushoutTot = + L→R.isPushoutBottomSquare→isPushoutTotSquare + (subst isEquiv (funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl})) + (idIsEquiv _)) + + -- The map we have called constant is actually (merely constant): + sphereBouquetVanish : ∀ {k : ℕ} + (f : SphereBouquet (suc n) (Fin k) → C4+n) → ∥ f ≡ (λ _ → C⋆) ∥₁ sphereBouquetVanish {k = k} f = TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) @@ -871,42 +903,39 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with ∙ sym (lUnit _) ◁ (symP (compPath-filler' (cong f (sym (push a))) p))) (~ i) j}) ∣₁) - help) - isPted + lem) + pted where - isPted = Iso.fun (PathIdTruncIso (suc (suc n))) + sphereVanish : (f : S₊ (suc n) → C4+n) + → ∥ ((x : S₊ (suc n)) → f x ≡ C⋆) ∥₁ + sphereVanish f = + sphereToTrunc (suc n) + λ x → isConnectedPath 2+n isConnectedC4+n _ _ .fst + + pted = Iso.fun (PathIdTruncIso 2+n) (isContr→isProp isConnectedC4+n ∣ f (inl tt) ∣ₕ ∣ C⋆ ∣ₕ) - help : ∥ ((x : Fin k) → (y : _) → f (inr (x , y)) ≡ C⋆) ∥₁ - help = invEq propTrunc≃Trunc1 + lem : ∥ ((x : Fin k) → (y : _) → f (inr (x , y)) ≡ C⋆) ∥₁ + lem = invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) λ x → PT.rec (isOfHLevelTrunc 1) ∣_∣ₕ (sphereVanish λ y → f (inr (x , y)))) - pushoutTotAlt : C∨PlaceHolder ≃ cofib const* - pushoutTotAlt = - compEquiv - (invEquiv (_ , isPushoutTot)) - (symPushout _ _) - - β : SphereBouquet (suc (suc n)) (A (suc (suc (suc n)))) - → SphereBouquet (suc (suc n)) (A (suc (suc n))) - β = (Iso.fun Iso-cofibαinr-SphereBouquet ∘ inl) ∘ fst α↑ - - β∙ : SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n)))) - →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))) - fst β∙ = β - snd β∙ = Iso-cofibαinr-SphereBouquet∙ + -- For now, we asssume it's constant + module _ (vanish : const→C4+n ≡ λ _ → C⋆) where - module _ (vanish : const* ≡ λ _ → C⋆) where - S∨C = SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4+n , C⋆) + -- Abbreviation: (⋁S⁴⁺ⁿ) ∨ C₄₊ₙ + SphereBouquet⋁C₄₊ₙ = + SphereBouquet∙ 2+n (A (suc n)) ⋁ (C4+n , C⋆) - connS∨C : isConnected (suc (suc (suc n))) S∨C - fst connS∨C = ∣ inl (inl tt) ∣ₕ - snd connS∨C = - TR.elim (λ _ → isOfHLevelPath (suc (suc (suc n))) - (isOfHLevelTrunc (suc (suc (suc n)))) _ _) + -- Connectedness of this space + isConnectedSphereBouquet⋁C₄₊ₙ : + isConnected 3+n SphereBouquet⋁C₄₊ₙ + fst isConnectedSphereBouquet⋁C₄₊ₙ = ∣ inl (inl tt) ∣ₕ + snd isConnectedSphereBouquet⋁C₄₊ₙ = + TR.elim (λ _ → isOfHLevelPath 3+n + (isOfHLevelTrunc 3+n) _ _) λ { (inl x) → inlEq x ; (inr x) → SP ∙ inrEq x ; (push tt i) j → @@ -914,16 +943,16 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with ▷ (rUnit SP ∙ cong (SP ∙_) (sym inrEq∙))) i j} where inlEq : (x : _) - → Path (hLevelTrunc (suc (suc (suc n))) S∨C) + → Path (hLevelTrunc 3+n SphereBouquet⋁C₄₊ₙ) ∣ inl (inl tt) ∣ ∣ inl x ∣ - inlEq x = TR.rec (isOfHLevelTrunc (3 +ℕ n) _ _) + inlEq x = TR.rec (isOfHLevelTrunc 3+n _ _) (λ p i → ∣ inl (p i) ∣ₕ) (Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedSphereBouquet' {n = suc n}) ∣ inl tt ∣ ∣ x ∣)) - G : (x : C4+n) → ∥ C⋆ ≡ x ∥ suc (suc n) + G : (x : C4+n) → ∥ C⋆ ≡ x ∥ 2+n G x = Iso.fun (PathIdTruncIso _) (isContr→isProp isConnectedC4+n ∣ C⋆ ∣ₕ ∣ x ∣ₕ) @@ -934,309 +963,369 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with ∙ cong ∣_∣ₕ (transportRefl refl) inrEq : (x : C4+n) - → Path (hLevelTrunc (suc (suc (suc n))) S∨C) + → Path (hLevelTrunc 3+n SphereBouquet⋁C₄₊ₙ) ∣ inr C⋆ ∣ₕ ∣ inr x ∣ₕ - inrEq x = TR.rec (isOfHLevelTrunc (suc (suc (suc n))) _ _) + inrEq x = TR.rec (isOfHLevelTrunc 3+n _ _) (λ p i → ∣ inr (p i) ∣ₕ) (G x) inrEq∙ : inrEq C⋆ ≡ refl - inrEq∙ = cong (TR.rec (isOfHLevelTrunc (suc (suc (suc n))) _ _) + inrEq∙ = cong (TR.rec (isOfHLevelTrunc 3+n _ _) (λ p i → ∣ inr (p i) ∣ₕ)) G∙ SP = inlEq (inl tt) ∙ cong ∣_∣ₕ (push tt) - Iso-C∨PlaceHolder-Wedge : - C∨PlaceHolder ≃ (SphereBouquet∙ (suc (suc n)) (A (suc n)) ⋁ (C4+n , C⋆)) - Iso-C∨PlaceHolder-Wedge = - compEquiv pushoutTotAlt - (isoToEquiv (compIso (cofibConst-⋁-Iso' {A = _ , inl tt} {B = _ , C⋆} - (const* , funExt⁻ vanish _) - (funExt⁻ vanish)) - (pushoutIso _ _ _ _ - (idEquiv _) - (isoToEquiv sphereBouquetSuspIso) - (idEquiv _) - refl - refl))) - - module T→B = PushoutPasteDown PTPush {B = C∨PlaceHolder} + + -- Equivalence of our C₄₊ₙ⋁SphereBouquet and SphereBouquet⋁C₄₊ₙ + C₄₊ₙ⋁SphereBouquetAsPushout≃SphereBouquet⋁C₄₊ₙ : + C₄₊ₙ⋁SphereBouquetAsPushout ≃ SphereBouquet⋁C₄₊ₙ + C₄₊ₙ⋁SphereBouquetAsPushout≃SphereBouquet⋁C₄₊ₙ = + compEquiv C₄₊ₙ⋁SphereBouquetAsPushout≃cofib-const + (isoToEquiv + (compIso (cofibConst-⋁-Iso' {A = _ , inl tt} {B = _ , C⋆} + (const→C4+n , funExt⁻ vanish _) + (funExt⁻ vanish)) + (pushoutIso _ _ _ _ + (idEquiv _) + (isoToEquiv sphereBouquetSuspIso) + (idEquiv _) + refl + refl))) + + module T→B = PushoutPasteDown pushoutSquare₁ + {B = C₄₊ₙ⋁SphereBouquetAsPushout} (λ x → Iso.fun Iso-cofibαinr-SphereBouquet (inl x)) inr (λ x → inl (Iso.inv Iso-C4+n-cofibα↑ x)) (sym (funExt push)) - helpIso : Iso (spanPushout (sp T→B.bottomSquare)) (P T→B.bottomSquare) - helpIso = compIso (equivToIso (symPushout _ _)) - (pushoutIso _ _ _ _ (idEquiv _) - (invEquiv (isoToEquiv Iso-C4+n-cofibα↑)) - (idEquiv _) refl refl) - - helpIsoCoh : (x : _) → Iso.fun helpIso x - ≡ Pushout→commSquare T→B.bottomSquare x - helpIsoCoh (inl x) = refl - helpIsoCoh (inr x) = refl - helpIsoCoh (push a i) j = sym (rUnit (sym (push a))) j i - - ⋁-as-cofib : cofib β ≃ S∨C - ⋁-as-cofib = + -- Finally, the first main lemma: the cofibre of β is S¹⁺ⁿ ∨ C₄₊ₙ + cofibβ≃SphereBouquet⋁C₄₊ₙ : cofib β ≃ SphereBouquet⋁C₄₊ₙ + cofibβ≃SphereBouquet⋁C₄₊ₙ = compEquiv (compEquiv (symPushout _ _) (_ , T→B.isPushoutBottomSquare→isPushoutTotSquare (subst isEquiv (funExt helpIsoCoh) (isoToIsEquiv helpIso)))) - Iso-C∨PlaceHolder-Wedge - - open import Cubical.Data.Unit - testβ : isConnectedFun (suc (suc (suc n))) {B = cofib β} inr - testβ = inrConnected (3 +ℕ n) _ _ - λ b → isOfHLevelRetractFromIso 0 (mapCompIso fiberUnitIso) - isConnectedSphereBouquet' - - testβ3 : isConnectedFun (suc (suc (suc n))) {B = S∨C} (⋁-as-cofib .fst ∘ inr) - testβ3 = isConnectedComp (⋁-as-cofib .fst) inr (3 +ℕ n) - (isEquiv→isConnected _ (⋁-as-cofib .snd) (3 +ℕ n)) testβ - - main-inr : ∥ ((x : A (suc n)) (y : S₊ (suc (suc n))) - → Σ[ t ∈ SphereBouquet (suc (suc n)) (A (suc (suc n))) ] - (⋁-as-cofib .fst (inr t) ≡ inl (inr (x , y)))) ∥₁ - main-inr = - invEq propTrunc≃Trunc1 - (invEq (_ , InductiveFinSatAC _ _ _) - λ x → fst propTrunc≃Trunc1 - (sphereToTrunc (suc (suc n)) - λ y → testβ3 (inl (inr (x , y))) .fst)) - - mainF : ∃[ F ∈ (SphereBouquet∙ (suc (suc n)) (A (suc n)) - →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n)))) ] - ((x : _) → Path S∨C (⋁-as-cofib .fst (inr (fst F x))) (inl x)) - mainF = TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) + C₄₊ₙ⋁SphereBouquetAsPushout≃SphereBouquet⋁C₄₊ₙ + where + helpIso : Iso (spanPushout (sp T→B.bottomSquare)) + (P T→B.bottomSquare) + helpIso = compIso (equivToIso (symPushout _ _)) + (pushoutIso _ _ _ _ (idEquiv _) + (invEquiv (isoToEquiv Iso-C4+n-cofibα↑)) + (idEquiv _) refl refl) + + helpIsoCoh : (x : _) → Iso.fun helpIso x + ≡ Pushout→commSquare T→B.bottomSquare x + helpIsoCoh (inl x) = refl + helpIsoCoh (inr x) = refl + helpIsoCoh (push a i) j = sym (rUnit (sym (push a))) j i + + -- If we could adjust β somewhat, killing off S¹⁺ⁿ in S¹⁺ⁿ ∨ + -- C₄₊ₙ, we would be done. + + + -- We prove the mere existence of such an `improved β': + ∃Improvedβ : ∃[ F ∈ (SphereBouquet∙ 2+n (A (suc n)) + →∙ SphereBouquet∙ 2+n (A 2+n)) ] + ((x : _) → Path SphereBouquet⋁C₄₊ₙ + (cofibβ≃SphereBouquet⋁C₄₊ₙ .fst + (inr (fst F x))) (inl x)) + ∃Improvedβ = TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) (λ coh₁ → PT.rec squash₁ (λ F → TR.rec squash₁ - (λ h → TR.rec squash₁ (λ q → ∣ (F* F coh₁ h , refl) , + (λ h → TR.rec squash₁ (λ q → ∣ (Improvedβ F coh₁ h , refl) , (coh F coh₁ h q) ∣₁) (makeCoh F coh₁ h)) (F∙ F coh₁)) - main-inr) - (isConnectedPath _ connS∨C - (⋁-as-cofib .fst (inr (inl tt))) (inl (inl tt)) .fst) + ∃Improvedβ-inr) + (isConnectedPath _ isConnectedSphereBouquet⋁C₄₊ₙ + (cofibβ≃SphereBouquet⋁C₄₊ₙ .fst + (inr (inl tt))) (inl (inl tt)) .fst) where - CON = (subst (λ x → isConnected x (SphereBouquet (suc (suc n)) - (A (suc (suc n))))) - (cong suc (+-comm 2 n)) - isConnectedSphereBouquet') - module _ (F : ((x : A (suc n)) (y : S₊ (suc (suc n))) - → Σ[ t ∈ SphereBouquet (suc (suc n)) (A (suc (suc n))) ] - (⋁-as-cofib .fst (inr t) ≡ inl (inr (x , y))))) - (coh₁ : Path S∨C (⋁-as-cofib .fst (inr (inl tt))) - (inl (inl tt))) where + ∃Improvedβ-inr : ∥ ((x : A (suc n)) (y : S₊ 2+n) + → Σ[ t ∈ SphereBouquet 2+n (A 2+n) ] + (cofibβ≃SphereBouquet⋁C₄₊ₙ .fst (inr t) + ≡ inl (inr (x , y)))) ∥₁ + ∃Improvedβ-inr = + invEq propTrunc≃Trunc1 + (invEq (_ , InductiveFinSatAC _ _ _) + λ x → fst propTrunc≃Trunc1 + (sphereToTrunc 2+n + λ y → isConnectedInr-cofib∘inr (inl (inr (x , y))) .fst)) + where + isConnectedInr-cofibβ : + isConnectedFun 3+n {B = cofib β} inr + isConnectedInr-cofibβ = inrConnected 3+n _ _ + λ b → isOfHLevelRetractFromIso 0 (mapCompIso fiberUnitIso) + isConnectedSphereBouquet' + + isConnectedInr-cofib∘inr : isConnectedFun 3+n + {B = SphereBouquet⋁C₄₊ₙ} + (cofibβ≃SphereBouquet⋁C₄₊ₙ .fst ∘ inr) + isConnectedInr-cofib∘inr = + isConnectedComp + (cofibβ≃SphereBouquet⋁C₄₊ₙ .fst) inr 3+n + (isEquiv→isConnected _ + (cofibβ≃SphereBouquet⋁C₄₊ₙ .snd) 3+n) + isConnectedInr-cofibβ + + -- We asumme the full existence of such an improved β (F below) + module _ + (F : ((x : A (suc n)) (y : S₊ 2+n) + → Σ[ t ∈ SphereBouquet 2+n (A 2+n) ] + (cofibβ≃SphereBouquet⋁C₄₊ₙ .fst (inr t) + ≡ inl (inr (x , y))))) + (coh₁ : Path SphereBouquet⋁C₄₊ₙ + (cofibβ≃SphereBouquet⋁C₄₊ₙ .fst (inr (inl tt))) + (inl (inl tt))) + where + -- mere pointedness F∙ : hLevelTrunc 1 ((x : Fin _) → F x north .fst ≡ inl tt) F∙ = invEq (_ , InductiveFinSatAC _ _ _) - λ a → isConnectedPath 1 (isConnectedSubtr 2 (suc n) CON) _ _ .fst + λ a → isConnectedPath 1 (isConnectedSubtr 2 (suc n) + (subst (λ x → isConnected x (SphereBouquet 2+n + (A 2+n))) + (cong suc (+-comm 2 n)) + isConnectedSphereBouquet')) _ _ .fst + -- ... which we assume, as usual. module _ (h : (x : Fin _) → F x north .fst ≡ inl tt) where - F* : (SphereBouquet (suc (suc n)) (A (suc n)) - → SphereBouquet (suc (suc n)) (A (suc (suc n)))) - F* (inl x) = inl tt - F* (inr (x , y)) = F x y .fst - F* (push a i) = h a (~ i) + Improvedβ : (SphereBouquet 2+n (A (suc n)) + → SphereBouquet 2+n (A 2+n)) + Improvedβ (inl x) = inl tt + Improvedβ (inr (x , y)) = F x y .fst + Improvedβ (push a i) = h a (~ i) + -- We also want the following coherencet to be satisfied cohTy : Type _ cohTy = (a : A (suc n)) - → Square (λ i → ⋁-as-cofib .fst (inr (h a (~ i)))) + → Square (λ i → cofibβ≃SphereBouquet⋁C₄₊ₙ .fst + (inr (h a (~ i)))) (λ i → inl (push a i)) coh₁ (F a north .snd) + -- It merely is... makeCoh : hLevelTrunc 1 cohTy makeCoh = invEq (_ , InductiveFinSatAC _ _ _) λ a → isConnectedPathP 1 - (isConnectedSubtr' n 2 (isConnectedPath _ connS∨C _ _)) _ _ .fst + (isConnectedSubtr' n 2 + (isConnectedPath _ + isConnectedSphereBouquet⋁C₄₊ₙ _ _)) _ _ .fst + -- so we assume it, giving us the main coherence we need module _ (coh₂ : cohTy) where - coh : (x : _) → Path S∨C (⋁-as-cofib .fst (inr (F* x))) (inl x) + coh : (x : _) → Path SphereBouquet⋁C₄₊ₙ + (cofibβ≃SphereBouquet⋁C₄₊ₙ .fst + (inr (Improvedβ x))) + (inl x) coh (inl x) = coh₁ coh (inr (x , y)) = F x y .snd coh (push a i) j = coh₂ a j i - module _ (F : (SphereBouquet∙ (suc (suc n)) (A (suc n)) - →∙ SphereBouquet∙ (suc (suc n)) (A (suc (suc n))))) - (h : (x : _) → Path S∨C (⋁-as-cofib .fst (inr (fst F x))) (inl x)) - where - h' : (x : _) → inr (fst F x) ≡ invEq ⋁-as-cofib (inl x) - h' x = sym (retEq ⋁-as-cofib (inr (fst F x))) - ∙ cong (invEq ⋁-as-cofib) (h x) - - open import Cubical.Foundations.Transport - - N = card (suc n) +ℕ card (suc (suc (suc n))) - iso7 : Iso (SphereBouquet (suc (suc n)) (Fin N)) - (SphereBouquet∙ (suc (suc n)) (A (suc n)) - ⋁ SphereBouquet∙ (suc (suc n)) (A (suc (suc (suc n))))) - iso7 = compIso + -- Now, assuming the exisetence of Imrpovedβ and the coherence above, + -- we can finally endow C₄₊ₙ with the appropriate cell structure + module _ (F : (SphereBouquet∙ 2+n (A (suc n)) + →∙ SphereBouquet∙ 2+n (A 2+n))) + (h : (x : _) → Path SphereBouquet⋁C₄₊ₙ + (cofibβ≃SphereBouquet⋁C₄₊ₙ .fst + (inr (fst F x))) + (inl x)) + where + -- The cardinality of the ₄₊ₙ-cells + N = card (suc n) +ℕ card 3+n + + -- Techincal ⋁-distributivity lemma + Iso-SphereBouquetN : Iso (SphereBouquet 2+n (Fin N)) + (SphereBouquet∙ 2+n (A (suc n)) + ⋁ SphereBouquet∙ 2+n (A 3+n)) + Iso-SphereBouquetN = compIso (pathToIso ((λ i → ⋁gen (isoToPath - (Iso-Fin⊎Fin-Fin+ {n = card (suc n)} {m = card (suc (suc (suc n)))}) (~ i)) - (λ _ → S₊∙ (suc (suc n)))) - ∙ cong (⋁gen (A (suc n) ⊎ A (suc (suc (suc n))))) + (Iso-Fin⊎Fin-Fin+ + {n = card (suc n)} + {m = card 3+n}) (~ i)) + (λ _ → S₊∙ 2+n)) + ∙ cong (⋁gen (A (suc n) ⊎ A 3+n)) (funExt (⊎.elim (λ _ → refl) (λ _ → refl))))) (invIso ⋁gen⊎Iso) - T : Iso (cofib (F ∨→ β∙)) C4+n - T = compIso (cofib∨→compIsoᵣ F β∙) + -- Finally, we have shown that C₄₊ₙ is the cofibre of F ∨→ + -- β → C4+n + Iso-cofib-C₄₊ₙ : Iso (cofib (F ∨→ β∙)) C4+n + Iso-cofib-C₄₊ₙ = compIso (cofib∨→compIsoᵣ F β∙) (compIso - (pathToIso (cong cofib (funExt h'))) + (pathToIso (cong cofib (funExt lem))) (compIso (equivToIso (symPushout _ _)) (compIso (PushoutCompEquivIso - (idEquiv (SphereBouquet (suc (suc n)) (A (suc n)))) - (invEquiv ⋁-as-cofib) inl (λ _ → tt)) + (idEquiv (SphereBouquet 2+n (A (suc n)))) + (invEquiv cofibβ≃SphereBouquet⋁C₄₊ₙ) inl (λ _ → tt)) (compIso (equivToIso (symPushout _ _)) cofibInl-⋁)))) - - - T∙ : Σ[ c ∈ ℕ ] Σ[ α ∈ (SphereBouquet (suc (suc n)) (Fin c) - → SphereBouquet (suc (suc n)) (A (suc (suc n)))) ] - Iso (Pushout (α ∘ inr) fst) C4+n - fst T∙ = N - fst (snd T∙) = F ∨→ β∙ ∘ Iso.fun iso7 - snd (snd T∙) = + where + lem : (x : _) → inr (fst F x) ≡ invEq cofibβ≃SphereBouquet⋁C₄₊ₙ (inl x) + lem x = sym (retEq cofibβ≃SphereBouquet⋁C₄₊ₙ (inr (fst F x))) + ∙ cong (invEq cofibβ≃SphereBouquet⋁C₄₊ₙ) (h x) + + -- Let us summarise the main data: we have shown that + -- there is some k and a map α : ⋁ₖ Sⁿ⁺² → ⋁ Sⁿ⁺² s.t. α + -- defines an attaching map for C₄₊ₙ. + mainData : Σ[ k ∈ ℕ ] Σ[ α ∈ (SphereBouquet 2+n (Fin k) + → SphereBouquet 2+n (A 2+n)) ] + (Iso (Pushout (α ∘ inr) fst) C4+n) + fst mainData = N + fst (snd mainData) = F ∨→ β∙ ∘ Iso.fun Iso-SphereBouquetN + snd (snd mainData) = compIso - (compIso (⋁-cofib-Iso (_ , fst F (inl tt)) ((F ∨→ β∙ ∘ Iso.fun iso7) , refl)) + (compIso (⋁-cofib-Iso (_ , fst F (inl tt)) + ((F ∨→ β∙ ∘ Iso.fun Iso-SphereBouquetN) , refl)) (invIso (PushoutCompEquivIso - (invEquiv (isoToEquiv iso7)) - (idEquiv Unit) _ _))) T - - tyFamParam : (m : ℕ) → Trichotomyᵗ m (suc (suc (suc n))) → Type ℓ - tyFamParam zero (lt x) = ⊥* - tyFamParam (suc m) (lt x) = Unit* - tyFamParam m (eq x) = Lift (SphereBouquet (suc (suc n)) (A (suc (suc n)))) - tyFamParam m (gt x) = C* m - - cardParam : (m : ℕ) → Trichotomyᵗ m (suc (suc n)) → Trichotomyᵗ m (suc (suc (suc n))) → ℕ - cardParam zero (lt x) q = 1 - cardParam (suc m) (lt x) q = 0 - cardParam m (eq x) q = card (suc (suc n)) - cardParam m (gt x) (lt x₁) = 0 - cardParam m (gt x) (eq x₁) = N - cardParam m (gt x) (gt x₁) = card m - - - tyFamParam∙ : (m : ℕ) (q : _) → tyFamParam (suc m) q - tyFamParam∙ m (lt x) = tt* - tyFamParam∙ m (eq x) = lift (inl tt) - tyFamParam∙ m (gt x) = subst C* (+-comm m 1) + (invEquiv (isoToEquiv Iso-SphereBouquetN)) + (idEquiv Unit) _ _))) Iso-cofib-C₄₊ₙ + + -- Using this, we can finally define a CW structure on C with the + -- desired properties + + -- We parametrise with Trichotomyᵗ as it makes removes some bureaucracy later + + -- The type family + C' : (m : ℕ) + → Trichotomyᵗ m 3+n → Type ℓ + C' zero (lt x) = ⊥* + C' (suc m) (lt x) = Unit* + C' m (eq x) = + Lift (SphereBouquet 2+n (A 2+n)) + C' m (gt x) = C* m + + -- Basepoints (not needed, but useful) + C'∙ : (m : ℕ) (q : _) → C' (suc m) q + C'∙ m (lt x) = tt* + C'∙ m (eq x) = lift (inl tt) + C'∙ m (gt x) = subst C* (+-comm m 1) (CW↪Iterate (_ , ind .fst .snd .fst) 1 m (invEq (CW₁-discrete (_ , ind .fst .snd .fst)) (subst Fin (sym (ind .fst .snd .snd .fst)) flast))) - carParamConn : (n₁ : ℕ) → n₁ <ᵗ suc n → (p : _) (q : _) → cardParam (suc n₁) p q ≡ 0 - carParamConn m ineq (lt x) q = refl - carParamConn m ineq (eq x) q = ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc (suc n)) x ineq)) - carParamConn m ineq (gt x) q = ⊥.rec (¬m<ᵗm (<ᵗ-trans ineq x)) - - newα : (m : ℕ) (p : _) (q : _) → Fin (cardParam m p q) × S⁻ m → tyFamParam m q - newα (suc m) (lt x) q () - newα (suc m) (eq x) (lt x₁) _ = tt* - newα (suc m) (eq x) (eq x₁) _ = tyFamParam∙ m (eq x₁) - newα (suc m) (eq x) (gt x₁) _ = tyFamParam∙ m (gt x₁) - newα (suc m) (gt x) (lt x₁) _ = tyFamParam∙ m (lt x₁) - newα (suc m) (gt x) (eq x₁) (a , p) = lift (snd T∙ .fst (inr (a , subst S₊ (cong predℕ x₁) p))) - newα (suc m) (gt x) (gt x₁) = α (suc m) - - tyFamParamContr : (m : _) (t : m <ᵗ suc n) (q : _) → isContr (tyFamParam (suc m) q) - tyFamParamContr m t (lt x) = tt* , λ {tt* → refl} - tyFamParamContr m t (eq x) = - ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (_<ᵗ suc n) (cong predℕ x) t) <ᵗsucm)) - tyFamParamContr m t (gt x) = - ⊥.rec (¬m<ᵗm (<ᵗ-trans (<ᵗ-trans x t) <ᵗsucm)) - - newEq : (n₁ : ℕ) (p : _) (q : _) - → tyFamParam (suc n₁) (Trichotomyᵗ-suc p) - ≃ Pushout (newα n₁ p q) fst - newEq zero (lt s) (lt t) = + -- Cardinality of cells + card' : (m : ℕ) → Trichotomyᵗ m 2+n + → Trichotomyᵗ m 3+n → ℕ + card' zero (lt x) q = 1 + card' (suc m) (lt x) q = 0 + card' m (eq x) q = card 2+n + card' m (gt x) (lt x₁) = 0 + card' m (gt x) (eq x₁) = N + card' m (gt x) (gt x₁) = card m + + -- Connectedness + C'-connected : (n₁ : ℕ) → n₁ <ᵗ suc n + → (p : _) (q : _) → card' (suc n₁) p q ≡ 0 + C'-connected m ineq (lt x) q = refl + C'-connected m ineq (eq x) q = ⊥.rec (¬m<ᵗm (subst (_<ᵗ 2+n) x ineq)) + C'-connected m ineq (gt x) q = ⊥.rec (¬m<ᵗm (<ᵗ-trans ineq x)) + + -- Attaching maps + α' : (m : ℕ) (p : _) (q : _) → Fin (card' m p q) × S⁻ m → C' m q + α' (suc m) (lt x) q () + α' (suc m) (eq x) (lt x₁) _ = tt* + α' (suc m) (eq x) (eq x₁) _ = C'∙ m (eq x₁) + α' (suc m) (eq x) (gt x₁) _ = C'∙ m (gt x₁) + α' (suc m) (gt x) (lt x₁) _ = C'∙ m (lt x₁) + α' (suc m) (gt x) (eq x₁) (a , p) = + lift (snd mainData .fst (inr (a , subst S₊ (cong predℕ x₁) p))) + α' (suc m) (gt x) (gt x₁) = α (suc m) + + -- Equivalences + e' : (n₁ : ℕ) (p : _) (q : _) + → C' (suc n₁) (Trichotomyᵗ-suc p) + ≃ Pushout (α' n₁ p q) fst + e' zero (lt s) (lt t) = isoToEquiv (isContr→Iso (tt* , (λ {tt* → refl})) ((inr flast) , λ {(inr (zero , tt)) → refl})) - newEq zero (lt x) (eq p) = ⊥.rec (snotz (sym p)) - newEq zero (eq x) q = ⊥.rec (snotz (sym x)) - newEq (suc m) (lt x) q = + e' zero (lt x) (eq p) = ⊥.rec (snotz (sym p)) + e' zero (eq x) q = ⊥.rec (snotz (sym x)) + e' (suc m) (lt x) q = isoToEquiv (isContr→Iso (tt* , (λ {tt* → refl})) - ((inl (tyFamParamContr m x q .fst)) + ((inl (C'Contr m x q .fst)) , (λ { (inl t) → λ i - → inl (tyFamParamContr m x q .snd t i)}))) - newEq (suc m) (eq x) (lt x₁) = + → inl (C'Contr m x q .snd t i)}))) + where + C'Contr : (m : _) (t : m <ᵗ suc n) (q : _) + → isContr (C' (suc m) q) + C'Contr m t (lt x) = tt* , λ {tt* → refl} + C'Contr m t (eq x) = + ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (_<ᵗ suc n) + (cong predℕ x) t) <ᵗsucm)) + C'Contr m t (gt x) = + ⊥.rec (¬m<ᵗm (<ᵗ-trans (<ᵗ-trans x t) <ᵗsucm)) + e' (suc m) (eq x) (lt x₁) = invEquiv (isoToEquiv - (compIso (⋁-cofib-Iso {B = λ _ → _ , ptSn m} (_ , tt*) (_ , refl)) + (compIso (⋁-cofib-Iso + {B = λ _ → _ , ptSn m} (_ , tt*) (_ , refl)) (compIso (cofibConst-⋁-Iso {A = SphereBouquet∙ m _}) (compIso ⋁-rUnitIso (compIso sphereBouquetSuspIso (compIso - (pathToIso (λ i → SphereBouquet (x i) (A (suc (suc n))))) + (pathToIso + (λ i → SphereBouquet (x i) (A 2+n))) LiftIso)))))) - newEq (suc m) (eq x) (eq x₁) = + e' (suc m) (eq x) (eq x₁) = ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc n) (cong (predℕ ∘ predℕ) (sym x ∙ x₁)) <ᵗsucm)) - newEq (suc m) (eq x) (gt y) = - ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (suc (suc n) <ᵗ_) + e' (suc m) (eq x) (gt y) = + ⊥.rec (¬m<ᵗm (<ᵗ-trans (subst (2+n <ᵗ_) (cong predℕ x) y) <ᵗsucm)) - newEq (suc m) (gt x) (lt x₁) = ⊥.rec (¬squeeze (x , x₁)) - newEq (suc m) (gt x) (eq x₁) = + e' (suc m) (gt x) (lt x₁) = ⊥.rec (¬squeeze (x , x₁)) + e' (suc m) (gt x) (eq x₁) = isoToEquiv (compIso (compIso (pathToIso (λ i → C* (suc (x₁ i)))) - (invIso (T∙ .snd .snd))) + (invIso (mainData .snd .snd))) (invIso (pushoutIso _ _ _ _ (Σ-cong-equiv-snd (λ _ → pathToEquiv (cong S₊ (cong predℕ x₁)))) (invEquiv LiftEquiv) (idEquiv _) refl refl))) - newEq (suc m) (gt x) (gt x₁) = e (suc m) - - buildCW* : connectedCWskel ℓ (suc n) - fst buildCW* m = tyFamParam m (m ≟ᵗ (3 +ℕ n)) - fst (fst (snd buildCW*)) m = cardParam m (m ≟ᵗ (2 +ℕ n)) (m ≟ᵗ (3 +ℕ n)) - fst (snd (fst (snd buildCW*))) m = newα m _ _ - fst (snd (snd (fst (snd buildCW*)))) () - snd (snd (snd (fst (snd buildCW*)))) m = newEq m _ _ - fst (snd (snd buildCW*)) = refl - snd (snd (snd buildCW*)) m ineq = carParamConn m ineq _ _ - - 3+n = suc (suc (suc n)) - 4+n = suc 3+n - - buildCW-realise : (n₁ : ℕ) (q : _) - → Iso (tyFamParam (n₁ +ℕ 4+n) q) + e' (suc m) (gt x) (gt x₁) = e (suc m) + + -- packaging it up into a connectedCWskel + C'-connectedCWskel : connectedCWskel ℓ (suc n) + fst C'-connectedCWskel m = C' m (m ≟ᵗ 3+n) + fst (fst (snd C'-connectedCWskel)) m = + card' m (m ≟ᵗ 2+n) (m ≟ᵗ 3+n) + fst (snd (fst (snd C'-connectedCWskel))) m = α' m _ _ + fst (snd (snd (fst (snd C'-connectedCWskel)))) () + snd (snd (snd (fst (snd C'-connectedCWskel)))) m = e' m _ _ + fst (snd (snd C'-connectedCWskel)) = refl + snd (snd (snd C'-connectedCWskel)) m ineq = C'-connected m ineq _ _ + + C'-realise : (n₁ : ℕ) (q : _) + → Iso (C' (n₁ +ℕ 4+n) q) (C* (n₁ +ℕ 4+n)) - buildCW-realise m (lt x) = ⊥.rec (¬m Date: Tue, 4 Jun 2024 14:22:49 +0200 Subject: [PATCH 50/73] connected clean --- Cubical/CW/Connected.agda | 628 +++----------------------------------- 1 file changed, 38 insertions(+), 590 deletions(-) diff --git a/Cubical/CW/Connected.agda b/Cubical/CW/Connected.agda index 02d60ede6e..d5cfbfbe20 100644 --- a/Cubical/CW/Connected.agda +++ b/Cubical/CW/Connected.agda @@ -1,23 +1,42 @@ {-# OPTIONS --safe --lossy-unification #-} +{- +This file contains the definition of an n-connected CW complex. This +is defined by saying that it has non-trivial cells only in dimension +≥n. + +The main result is packaged up in makeConnectedCW. This says that the +usual notion of connectedness in terms of truncations (merely) +coincides with the above definition for CW complexes. +-} + module Cubical.CW.Connected where +open import Cubical.CW.Base +open import Cubical.CW.Properties + open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Path open import Cubical.Foundations.Transport +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order open import Cubical.Data.Nat.Order.Inductive open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Fin.Inductive.Properties as Ind open import Cubical.Data.Sigma open import Cubical.Data.Sequence -open import Cubical.Data.FinSequence open import Cubical.Data.Unit +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool open import Cubical.HITs.Sn open import Cubical.HITs.Pushout @@ -25,48 +44,16 @@ open import Cubical.HITs.Susp open import Cubical.HITs.SequentialColimit open import Cubical.HITs.SphereBouquet open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as TR open import Cubical.HITs.Wedge -open import Cubical.HITs.Pushout open import Cubical.Axiom.Choice - -open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup - open import Cubical.Relation.Nullary -open import Cubical.CW.Base - - -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Univalence - -open import Cubical.Foundations.Equiv - --- connected comp open import Cubical.Homotopy.Connected -open import Cubical.CW.Properties -open import Cubical.HITs.Truncation as TR -open import Cubical.Foundations.HLevels -open import Cubical.Data.Nat.Order.Inductive -open import Cubical.Data.Fin.Inductive.Properties as Ind - - -open import Cubical.Data.Bool -open import Cubical.Data.Nat.Order.Inductive --- open import Cubical.Data.Dec - -open import Cubical.Relation.Nullary.Base -open import Cubical.Data.Sum as ⊎ -open import Cubical.Data.Empty as ⊥ - -open import Cubical.Data.Fin.Inductive.Properties as Ind open Sequence -open import Cubical.Foundations.Equiv.HalfAdjoint - private variable ℓ ℓ' ℓ'' : Level @@ -291,10 +278,10 @@ module shrinkPushoutLemma (A : Type ℓ) (B : Type ℓ') ≡ Unit⊎A→Pushout-g∘f-fst₂ x true lem₂ x with (f (x , true)) ... | inl x₁ = refl - ... | inr x₁ = refl + ... | inr x₁ = refl Pushout-g∘f→PushoutF→Pushout-g∘f : (x : _) - → Pushout-g∘f-fst→Unit⊎A (PushoutF→Pushout-g∘f x) ≡ x + → Pushout-g∘f-fst→Unit⊎A (PushoutF→Pushout-g∘f x) ≡ x Pushout-g∘f→PushoutF→Pushout-g∘f (inl x) = PushoutF→Pushout-g∘f→Unit⊎Aₗ x Pushout-g∘f→PushoutF→Pushout-g∘f (inr x) = PushoutF→Pushout-g∘f→Unit⊎Aᵣ x Pushout-g∘f→PushoutF→Pushout-g∘f (push (inl x , false) i) j = @@ -524,10 +511,13 @@ module CWLemmas-0Connected where help : Iso (Pushout f fst) (Fin (suc (suc m))) help = invIso (PushoutEmptyFam (λ()) λ()) Contract1Skel (suc n) (suc (suc m)) f c - with (Contract1Skel _ (suc m) (shrinkImageAttachingMap (suc n) m f c .snd .fst) + with (Contract1Skel _ (suc m) + (shrinkImageAttachingMap (suc n) m f c .snd .fst) (subst (isConnected 2) - (isoToPath (shrinkImageAttachingMap (suc n) m f c .snd .snd)) c)) - ... | (k , e) = k , compIso (shrinkImageAttachingMap (suc n) m f c .snd .snd) e + (isoToPath + (shrinkImageAttachingMap (suc n) m f c .snd .snd)) c)) + ... | (k , e) = k + , compIso (shrinkImageAttachingMap (suc n) m f c .snd .snd) e -- Uning this, we can show that a 0-connected CW complex can be -- approximated by one with trivial 1-skeleton. @@ -541,7 +531,8 @@ module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A 0) where module AC = CWskel-fields (_ , sk) e₁ : Iso (Pushout (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) fst) (A 2) - e₁ = compIso (PushoutCompEquivIso (idEquiv _) (CW₁-discrete (A , sk)) (AC.α 1) fst) + e₁ = compIso (PushoutCompEquivIso (idEquiv _) + (CW₁-discrete (A , sk)) (AC.α 1) fst) (equivToIso (invEquiv (AC.e 1))) liftStr = Contract1Skel _ _ (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) @@ -658,12 +649,11 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with module _ where C* = ind .fst .fst - 1+n = suc n - 2+n = suc 1+n + 2+n = suc (suc n) 3+n = suc 2+n 4+n = suc 3+n - C1+n = C* 1+n + C1+n = C* (suc n) C2+n = C* 2+n C3+n = C* 3+n C4+n = C* 4+n @@ -673,7 +663,7 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with isConnectedC4+n : isConnected 3+n C4+n isConnectedC4+n = isConnectedCW→isConnectedSkel (_ , ind .fst .snd .fst) 4+n - (suc 2+n , <ᵗ-trans <ᵗsucm <ᵗsucm) + (3+n , <ᵗ-trans <ᵗsucm <ᵗsucm) (subst (isConnected 3+n) (ua (invEquiv (ind .snd))) cA) @@ -690,7 +680,7 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with isContr→Iso (isConnectedCW→Contr n (ind .fst) (n , <ᵗsucm)) (flast , isPropFin1 _) - -- C₂₊ₙ is a bouquet of spheres + -- C₂₊ₙ is a bouquet of spheres Iso-C2+n-SphereBouquet : Iso C2+n (SphereBouquet (suc n) (A (suc n))) Iso-C2+n-SphereBouquet = compIso (equivToIso (ind .fst .snd .fst .snd .snd .snd (suc n))) @@ -778,7 +768,7 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with (⋁-cofib-Iso _ α↑)) - + opaque Iso-cofibαinr-SphereBouquet : Iso (Pushout {B = cofib (fst α∙)} inr (λ _ → tt)) @@ -1288,7 +1278,7 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with fst C'-connectedCWskel m = C' m (m ≟ᵗ 3+n) fst (fst (snd C'-connectedCWskel)) m = card' m (m ≟ᵗ 2+n) (m ≟ᵗ 3+n) - fst (snd (fst (snd C'-connectedCWskel))) m = α' m _ _ + fst (snd (fst (snd C'-connectedCWskel))) m = α' m _ _ fst (snd (snd (fst (snd C'-connectedCWskel)))) () snd (snd (snd (fst (snd C'-connectedCWskel)))) m = e' m _ _ fst (snd (snd C'-connectedCWskel)) = refl @@ -1327,545 +1317,3 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with ((λ m → C'-realise m ((m +ℕ 4+n) ≟ᵗ 3+n)) , (λ n → C'-realise-coh n _ _))) (invIso (SeqColimIso _ (4 +ℕ n)))))) (ind .snd) - - --- isConnCW : (n : ℕ) → Type ℓ → Type (ℓ-suc ℓ) --- isConnCW n A = isCW A × isConnected n A - --- connectedCWskel'→connectedCWskel : ∀ {ℓ} --- → Σ[ t ∈ connectedCWskel' ℓ 0 ] --- (Σ[ dim ∈ ℕ ] --- ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim)))) --- → Σ[ t ∈ connectedCWskel ℓ 0 ] --- Σ[ dim ∈ ℕ ] --- ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim))) --- fst (connectedCWskel'→connectedCWskel ((A , sk) , conv)) = --- _ , connectedCWskel→ A ((sk .fst) , (sk .snd)) .fst , refl , (λ _ → λ()) --- fst (snd (connectedCWskel'→connectedCWskel ((A , sk) , conv))) = suc (fst conv) --- snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , zero , T))) k = --- ⊥.rec (TR.rec (λ()) --- (λ a → TR.rec (λ()) --- (λ t → CW₀-empty (_ , fst sk) (invEq (_ , T 0) (t .fst))) --- (isConnected-CW↪∞ 1 (_ , fst sk) a .fst)) (sk .snd .fst)) --- snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , (suc dim) , T))) k = --- transport (λ i → isEquiv (CW↪ (collapse₁CWskel A sk , connectedCWskel→ A sk .fst) --- (h i))) --- (transport (λ i → isEquiv (CW↪ (A , sk .fst) (suc (+-suc k dim i)))) --- (T (suc k))) --- where --- h = cong suc (sym (+-suc k dim)) ∙ sym (+-suc k (suc dim)) - - --- --- yieldsGoodCWskel : {ℓ : Level} (A₊₂ : ℕ → Pointed ℓ) → Type _ --- yieldsGoodCWskel {ℓ = ℓ} A = --- Σ[ f₊₁ ∈ (ℕ → ℕ) ] --- Σ[ fin ∈ (A 0) .fst ≃ Fin 1 ] --- Σ[ α ∈ ((n : ℕ) → SphereBouquet∙ n (Fin (f₊₁ n)) →∙ A n) ] --- ((n : ℕ) → cofib (α n .fst) , inl tt ≃∙ A (suc n)) - --- GoodCWskelSeq : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Sequence ℓ --- obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) zero = Lift ⊥ --- obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) (suc n) = fst (A n) --- Sequence.map (GoodCWskelSeq {A = A} (f , fin , α , sq)) {n = suc n} x = fst (fst (sq n)) (inr x) - --- realiseGood∙ : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Pointed ℓ --- realiseGood∙ {A = A} S = SeqColim (GoodCWskelSeq S) , incl {n = 1} (snd (A 0)) - --- yieldsFinGoodCWskel : {ℓ : Level} (dim : ℕ) (A₊₂ : ℕ → Pointed ℓ) → Type _ --- yieldsFinGoodCWskel {ℓ = ℓ} dim A = --- Σ[ A ∈ yieldsGoodCWskel A ] converges (GoodCWskelSeq A) dim - --- GoodCWskel : (ℓ : Level) → Type (ℓ-suc ℓ) --- GoodCWskel ℓ = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsGoodCWskel A - --- FinGoodCWskel : (ℓ : Level) (dim : ℕ) → Type (ℓ-suc ℓ) --- FinGoodCWskel ℓ dim = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsFinGoodCWskel dim A - --- isGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) --- isGoodCWExpl {ℓ} A = Σ[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A - --- isFinGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) --- isFinGoodCWExpl {ℓ} A = --- Σ[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] realiseGood∙ (fst (snd sk)) ≃∙ A - --- isGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) --- isGoodCW {ℓ} A = ∃[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A - --- isFinGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) --- isFinGoodCW {ℓ} A = --- ∃[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] (realiseGood∙ (fst (snd sk)) ≃∙ A) - --- finGoodCW : (ℓ : Level) → Type (ℓ-suc ℓ) --- finGoodCW ℓ = Σ[ A ∈ Pointed ℓ ] isFinGoodCW A - --- ptCW : {ℓ : Level} {A : ℕ → Type ℓ} → yieldsCWskel A → A 1 --- → (n : ℕ) → A (suc n) --- ptCW sk a zero = a --- ptCW sk a (suc n) = CW↪ (_ , sk) (suc n) (ptCW sk a n) - --- -- module TWOO {ℓ : Level} (A' : ℕ → Type ℓ) (pt0 : A' 1) --- -- (dim : ℕ) (con : isConnected 2 (A' 2)) --- -- (C : yieldsFinCWskel dim A') --- -- where - --- -- open CWskel-fields (_ , fst C) --- -- e₀ : A' 1 ≃ Fin (card 0) --- -- e₀ = CW₁-discrete (_ , fst C) - --- -- ptA : (n : ℕ) → A' (suc n) --- -- ptA = ptCW (fst C) pt0 - --- -- conA : (n : ℕ) → isConnected 2 (A' (suc n)) --- -- conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) --- -- (subst (isConnected 2) (sym (cong Fin cA)) --- -- (∣ flast ∣ --- -- , TR.elim (λ _ → isOfHLevelPath 2 --- -- (isOfHLevelTrunc 2) _ _) --- -- λ {(zero , tt) → refl})) --- -- conA (suc n) = --- -- isConnectedRetractFromIso 2 --- -- (equivToIso (e (suc n))) --- -- (∣ inl (ptA n) ∣ₕ --- -- , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) --- -- (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) --- -- (conA' (conA n)) --- -- λ c → conA' (conA n) _ --- -- ∙ cong ∣_∣ₕ (push (c , ptSn n)))) --- -- where --- -- conA' : (conA : isConnected 2 (A' (suc n))) (c : A' (suc n)) --- -- → Path (hLevelTrunc 2 (Pushout (α (suc n)) fst)) --- -- ∣ inl (ptA n) ∣ₕ ∣ inl c ∣ₕ --- -- conA' conA c = --- -- TR.rec (isOfHLevelTrunc 2 _ _) --- -- (λ p i → ∣ inl (p i) ∣) --- -- (Iso.fun (PathIdTruncIso _) --- -- (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) - - --- -- private --- -- funType = ((n : Fin (suc dim)) --- -- → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) --- -- →∙ (A' (suc (fst n)) , ptA (fst n))) ] --- -- ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) - --- -- mapMakerNil : ∥ funType ∥₁ --- -- mapMakerNil = --- -- invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) --- -- λ n → TR.map --- -- (λ pted → ((λ { (inl x) → ptA (fst n) --- -- ; (inr x) → α _ x --- -- ; (push a i) → pted a i}) --- -- , refl) --- -- , λ _ → refl) (help n)) --- -- where --- -- help : (n : Fin (suc dim)) --- -- → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) --- -- → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) --- -- help n = invEq (_ , InductiveFinSatAC _ _ _) --- -- λ m → Iso.fun (PathIdTruncIso _) --- -- (isContr→isProp --- -- (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ --- -- ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) --- -- module _ (F : funType) where --- -- funs : (n : ℕ) → SphereBouquet∙ n (Fin (card (suc n))) --- -- →∙ (A' (suc n) , ptA n) --- -- funs n with (n ≟ᵗ dim) --- -- ... | lt x = F (n , <ᵗ-trans-suc x) .fst --- -- ... | eq x = F (n , subst (_<ᵗ suc dim) (sym x) <ᵗsucm) .fst --- -- ... | gt x = const∙ _ _ - --- -- funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) --- -- ≃∙ Pushout (funs n .fst ∘ inr) (λ r → fst r) , inl (ptA n) --- -- funEqP1 n = invEquiv (isoToEquiv (⋁-cofib-Iso _ (funs n))) , refl - --- -- funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) --- -- ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) --- -- funEq n = isoToEquiv (pushoutIso _ _ _ _ --- -- (idEquiv _) --- -- (idEquiv _) --- -- (idEquiv _) --- -- (funExt (uncurry (main n))) --- -- (λ i x → fst x)) --- -- , λ _ → inl (ptA n) --- -- where --- -- main : (n : ℕ) (x : Fin (card (suc n))) (y : S₊ n) --- -- → funs n .fst (inr (x , y)) ≡ fst (C .snd) (suc n) (x , y) --- -- main n with (n ≟ᵗ dim) --- -- ... | lt p = λ x y --- -- → F (n , <ᵗ-trans-suc p) .snd (x , y) --- -- ... | eq p = λ x y --- -- → F (n , subst (_<ᵗ suc dim) (λ i → p (~ i)) <ᵗsucm) .snd (x , y) --- -- ... | gt p = λ x --- -- → ⊥.rec (¬Fin0 (subst Fin (ind (suc n) (<ᵗ-trans p <ᵗsucm)) x)) - --- -- getGoodCWskelAux : ∥ yieldsGoodCWskel (λ n → A' (suc n) , ptA n) ∥₁ --- -- getGoodCWskelAux = commSquare₁.map help mapMakerNil --- -- where --- -- help : funType → yieldsGoodCWskel (λ n → A' (suc n) , ptA n) --- -- fst (help F) = card ∘ suc --- -- fst (snd (help F)) = compEquiv e₀ (pathToEquiv (cong Fin cA)) --- -- fst (snd (snd (help F))) n = funs F n --- -- snd (snd (snd (help F))) n = --- -- compEquiv∙ (compEquiv∙ (funEqP1 F n) (funEq F n)) --- -- (invEquiv (e (suc n)) , refl) - - --- module BS {ℓ : Level} (A' : ℕ → Type ℓ) --- (dim : ℕ) --- (C+eq : yieldsFinCWskel dim A') --- (cA : fst (fst C+eq) 0 ≡ 1) --- where --- C = fst C+eq --- ind = snd C+eq - --- open CWskel-fields (_ , C) --- e₀ : A' 1 ≃ Fin (card 0) --- e₀ = CW₁-discrete (_ , C) - - --- ¬dim≡0 : ¬ (dim ≡ 0) --- ¬dim≡0 p = CW₀-empty (_ , C) (subst A' p --- (invEq (_ , ind 0) (subst A' (cong suc (sym p)) --- (invEq e₀ (subst Fin (sym cA) fzero))))) - --- pt0 : A' 1 --- pt0 = invEq e₀ (subst Fin (sym cA) flast) - --- ptA : (n : ℕ) → A' (suc n) --- ptA = ptCW C pt0 - --- conA : (n : ℕ) → isConnected 2 (A' (suc n)) --- conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) --- (subst (isConnected 2) (sym (cong Fin cA)) --- (∣ flast ∣ --- , TR.elim (λ _ → isOfHLevelPath 2 --- (isOfHLevelTrunc 2) _ _) --- λ {(zero , tt) → refl})) --- conA (suc n) = --- isConnectedRetractFromIso 2 --- (equivToIso (e (suc n))) --- (∣ inl (ptA n) ∣ₕ --- , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) --- (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) --- (conA' (conA n)) --- λ c → conA' (conA n) _ --- ∙ cong ∣_∣ₕ (push (c , ptSn n)))) --- where --- conA' : (conA : isConnected 2 (A' (suc n))) (c : A' (suc n)) --- → Path (hLevelTrunc 2 (Pushout (α (suc n)) fst)) --- ∣ inl (ptA n) ∣ₕ ∣ inl c ∣ₕ --- conA' conA c = --- TR.rec (isOfHLevelTrunc 2 _ _) --- (λ p i → ∣ inl (p i) ∣) --- (Iso.fun (PathIdTruncIso _) --- (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) - --- private --- funType = ((n : Fin dim) --- → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) --- →∙ (A' (suc (fst n)) , ptA (fst n))) ] --- ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) - --- mapMakerNil : ∥ funType ∥₁ --- mapMakerNil = --- invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) --- λ n → TR.map --- (λ pted → ((λ { (inl x) → ptA (fst n) --- ; (inr x) → α _ x --- ; (push a i) → pted a i}) --- , refl) --- , λ _ → refl) (help n)) --- where --- help : (n : Fin dim) --- → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) --- → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) --- help n = invEq (_ , InductiveFinSatAC _ _ _) --- λ m → Iso.fun (PathIdTruncIso _) --- (isContr→isProp --- (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) - --- module _ (F : funType) where --- card' : ℕ → ℕ --- card' n with (n ≟ᵗ dim) --- ... | lt x = card (suc n) --- ... | eq x = 0 --- ... | gt x = 0 - --- funs : (n : ℕ) → SphereBouquet∙ n (Fin (card' n)) --- →∙ (A' (suc n) , ptA n) --- funs n with (n ≟ᵗ dim) --- ... | lt x = F (n , x) .fst --- ... | eq x = const∙ _ _ --- ... | gt x = const∙ _ _ - --- funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) --- ≃∙ Pushout (funs n .fst ∘ inr) (λ r → fst r) , inl (ptA n) --- funEqP1 n = invEquiv (isoToEquiv (⋁-cofib-Iso _ (funs n))) , refl - --- funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) --- ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) --- funEq n with (n ≟ᵗ dim) --- ... | lt x = isoToEquiv (pushoutIso _ _ _ _ --- (idEquiv _) --- (idEquiv _) --- (idEquiv _) --- (funExt (uncurry λ x y → F (n , _) .snd (x , y))) --- (λ i x → fst x)) --- , λ _ → inl (ptA n) --- ... | eq x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) --- (compEquiv ((CW↪ (_ , C) (suc n)) --- , transport (λ i → isEquiv (CW↪ (A' , C) --- (suc (x (~ i))))) --- (ind 1) --- ) --- (e (suc n)))) , secEq (e (suc n)) _ --- ... | gt x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) --- (compEquiv ((CW↪ (_ , C) (suc n)) --- , (transport (λ i → isEquiv (CW↪ (A' , C) --- (suc ((sym (+-suc (fst (<ᵗ→< x)) dim) --- ∙ (<ᵗ→< x .snd)) i)))) --- (ind (suc (suc (fst (<ᵗ→< x))))))) --- (e (suc n)))) , secEq (e (suc n)) _ - --- goodCWmk : yieldsGoodCWskel (λ n → A' (suc n) , ptA n) --- fst goodCWmk = card' --- fst (snd goodCWmk) = compEquiv e₀ (pathToEquiv (cong Fin cA)) --- fst (snd (snd goodCWmk)) = funs --- snd (snd (snd goodCWmk)) n = --- compEquiv∙ (compEquiv∙ (funEqP1 n) (funEq n)) --- (invEquiv (e (suc n)) , refl) - --- goodCWmk-converges : converges --- (sequence (obj (GoodCWskelSeq goodCWmk)) --- (Sequence.map (GoodCWskelSeq goodCWmk))) --- dim --- goodCWmk-converges zero = help dim refl --- where --- help : (x : _) (p : dim ≡ x) → isEquiv (Sequence.map (GoodCWskelSeq goodCWmk) {x}) --- help zero p = ⊥.rec (¬dim≡0 p) --- help (suc x) p with (x ≟ᵗ dim) --- ... | lt x₁ = transport (λ i → isEquiv λ z → CW↪ (A' , C) (p i) z) (ind zero) --- ... | eq x₁ = ⊥.rec (¬m Date: Tue, 4 Jun 2024 14:42:34 +0200 Subject: [PATCH 51/73] minor --- Cubical/CW/Pointed.agda | 1123 -------------------------- Cubical/CW/Pointed2.agda | 1567 ------------------------------------ Cubical/CW/Subcomplex.agda | 5 - 3 files changed, 2695 deletions(-) delete mode 100644 Cubical/CW/Pointed.agda delete mode 100644 Cubical/CW/Pointed2.agda diff --git a/Cubical/CW/Pointed.agda b/Cubical/CW/Pointed.agda deleted file mode 100644 index b9fcee811a..0000000000 --- a/Cubical/CW/Pointed.agda +++ /dev/null @@ -1,1123 +0,0 @@ -{-# OPTIONS --safe --lossy-unification #-} - --- This file contains definition of CW complexes and skeleta. - -module Cubical.CW.Pointed where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Function - -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) -open import Cubical.Data.Nat.Order -open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Sigma -open import Cubical.Data.Sequence -open import Cubical.Data.FinSequence - -open import Cubical.HITs.Sn -open import Cubical.HITs.Pushout -open import Cubical.HITs.Susp -open import Cubical.HITs.SequentialColimit -open import Cubical.HITs.SphereBouquet -open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.HITs.Wedge - -open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup - -open import Cubical.Relation.Nullary - -open import Cubical.CW.Base - -open Sequence - -open import Cubical.Foundations.Equiv.HalfAdjoint - -private - variable - ℓ ℓ' ℓ'' : Level - -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Univalence - -module _ {C : Type ℓ} {B : Type ℓ'} where - PushoutAlongEquiv→ : {A : Type ℓ} - (e : A ≃ C) (f : A → B) → Pushout (fst e) f → B - PushoutAlongEquiv→ e f (inl x) = f (invEq e x) - PushoutAlongEquiv→ e f (inr x) = x - PushoutAlongEquiv→ e f (push a i) = f (retEq e a i) - - private - PushoutAlongEquiv→Cancel : {A : Type ℓ} (e : A ≃ C) (f : A → B) - → retract (PushoutAlongEquiv→ e f) inr - PushoutAlongEquiv→Cancel = - EquivJ (λ A e → (f : A → B) - → retract (PushoutAlongEquiv→ e f) inr) - λ f → λ { (inl x) → sym (push x) - ; (inr x) → refl - ; (push a i) j → push a (~ j ∨ i)} - - PushoutAlongEquiv : {A : Type ℓ} (e : A ≃ C) (f : A → B) - → Iso (Pushout (fst e) f) B - Iso.fun (PushoutAlongEquiv e f) = PushoutAlongEquiv→ e f - Iso.inv (PushoutAlongEquiv e f) = inr - Iso.rightInv (PushoutAlongEquiv e f) x = refl - Iso.leftInv (PushoutAlongEquiv e f) = PushoutAlongEquiv→Cancel e f - - ---- CW complexes --- --- Definition of (the skeleton) of an arbitrary CW complex --- New def: X 0 is empty and C (suc n) is pushout - -open import Cubical.HITs.Pushout -module _ {A : Type} {B : A → Pointed₀} (C : Pointed₀) (f : (⋁gen A B , inl tt) →∙ C) where - - - private - open 3x3-span - inst : 3x3-span - A00 inst = A - A02 inst = Σ A (fst ∘ B) - A04 inst = Σ A (fst ∘ B) - A20 inst = A - A22 inst = A - A24 inst = Σ A (fst ∘ B) - A40 inst = Unit - A42 inst = Unit - A44 inst = fst C - f10 inst = idfun A - f12 inst = λ x → x , snd (B x) - f14 inst = idfun _ - f30 inst = λ _ → tt - f32 inst = λ _ → tt - f34 inst = fst f ∘ inr - f01 inst = fst - f21 inst = idfun A - f41 inst = λ _ → tt - f03 inst = idfun _ - f23 inst = λ x → x , snd (B x) - f43 inst = λ _ → pt C - H11 inst = λ _ → refl - H13 inst = λ _ → refl - H31 inst = λ _ → refl - H33 inst = λ x → sym (snd f) ∙ cong (fst f) (push x) - - A0□Iso : Iso (A0□ inst) A - A0□Iso = compIso (equivToIso (symPushout _ _)) - (PushoutAlongEquiv (idEquiv _) fst) - - A2□Iso : Iso (A2□ inst) (Σ A (fst ∘ B)) - A2□Iso = PushoutAlongEquiv (idEquiv A) _ - - A4□Iso : Iso (A4□ inst) (fst C) - A4□Iso = PushoutAlongEquiv (idEquiv Unit) λ _ → pt C - - A○□Iso : Iso (A○□ inst) (Pushout (fst f ∘ inr) fst) - A○□Iso = compIso (equivToIso (symPushout _ _)) - (invIso (pushoutIso _ _ _ _ - (isoToEquiv (invIso A2□Iso)) - (isoToEquiv (invIso A4□Iso)) - (isoToEquiv (invIso A0□Iso)) - refl - λ i x → push x i)) - - - - A□0Iso : Iso (A□0 inst) Unit - A□0Iso = isContr→Iso - (inr tt , λ { (inl x) → sym (push x) - ; (inr x) → refl - ; (push a i) j → push a (i ∨ ~ j)}) - (tt , (λ _ → refl)) - - A□2Iso : Iso (A□2 inst) (⋁gen A B) - A□2Iso = equivToIso (symPushout _ _) - - - A□4Iso : Iso (A□4 inst) (fst C) - A□4Iso = PushoutAlongEquiv (idEquiv _) _ - - open import Cubical.Foundations.GroupoidLaws - A□○Iso : Iso (A□○ inst) (cofib (fst f)) - A□○Iso = invIso (pushoutIso _ _ _ _ - (isoToEquiv (invIso A□2Iso)) - (isoToEquiv (invIso A□0Iso)) - (isoToEquiv (invIso A□4Iso)) - (funExt (λ { (inl x) → refl - ; (inr x) → sym (push (fst x)) ∙ refl - ; (push a i) j → (sym (push a) ∙ refl) (i ∧ j)})) - (funExt λ { (inl x) i → inr (snd f i) - ; (inr x) → sym (push x) - ; (push a i) j - → hcomp (λ k - → (λ {(i = i0) → inr (compPath-filler' (sym (snd f)) - (cong (fst f) (push a)) j (~ k)) - ; (i = i1) → push (a , snd (B a)) (~ j) - ; (j = i0) → inr (fst f (push a (~ k ∨ i)))})) - (push (a , snd (B a)) (~ i ∨ ~ j))})) - - ⋁-cofib-Iso : Iso (Pushout (fst f ∘ inr) fst) (cofib (fst f)) - ⋁-cofib-Iso = compIso (compIso (invIso A○□Iso) - (invIso (3x3-Iso inst))) - A□○Iso - - - --- connected comp -open import Cubical.Homotopy.Connected -open import Cubical.CW.Properties -open import Cubical.HITs.Truncation as TR -open import Cubical.Foundations.HLevels - -isConnectedCW→isConnectedSkel : (C : CWskel ℓ) - → isConnected 2 (realise C) - → (n : ℕ) - → isConnected 2 (C .fst (suc (suc n))) -isConnectedCW→isConnectedSkel C c n = - TR.rec (isOfHLevelSuc 1 isPropIsContr) - (λ ⋆ → TR.rec (isProp→isOfHLevelSuc (suc n) isPropIsContr) - (λ {(x , p) → ∣ x ∣ - , (TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) - (λ a → Iso.inv (PathIdTruncIso 1) - (TR.rec (isOfHLevelTrunc 1) - (λ q → - TR.rec (isOfHLevelPlus' 1 (isOfHLevelTrunc 1)) - (∣_∣ₕ ∘ fst) - (isConnectedCong (suc n) _ (isConnected-CW↪∞ (suc (suc n)) C) - q .fst)) - (Iso.fun (PathIdTruncIso 1) - (sym (c .snd ∣ incl x ∣) ∙ c .snd ∣ incl a ∣)))))}) - (isConnected-CW↪∞ (suc (suc n)) C ⋆ .fst)) - (c .fst) - -open import Cubical.Data.Bool -open import Cubical.Data.Nat.Order.Inductive --- open import Cubical.Data.Dec - -open import Cubical.Relation.Nullary.Base -open import Cubical.Data.Sum as ⊎ -open import Cubical.Data.Empty as ⊥ - -open import Cubical.Data.Fin.Inductive.Properties as Ind - -FinSuc : {n : ℕ} → Iso (Fin 1 ⊎ Fin n) (Fin (suc n)) -Iso.fun (FinSuc {n = n}) = ⊎.rec (λ _ → flast) injectSuc -Iso.inv (FinSuc {n = n}) = Ind.elimFin (inl flast) inr -Iso.rightInv (FinSuc {n = n}) = - Ind.elimFin (cong (⊎.rec (λ _ → flast) injectSuc) - (Ind.elimFinβ (inl flast) inr .fst)) - λ x → cong (⊎.rec (λ _ → flast) injectSuc) - (Ind.elimFinβ (inl flast) inr .snd x) -Iso.leftInv (FinSuc {n = n}) (inl (zero , p)) = - Ind.elimFinβ (inl flast) inr .fst -Iso.leftInv (FinSuc {n = n}) (inr x) = Ind.elimFinβ (inl flast) inr .snd x - -Fin+ : {n m : ℕ} → Iso (Fin n ⊎ Fin m) (Fin (n +ℕ m)) -Iso.fun (Fin+ {n = zero} {m = m}) (inr x) = x -Iso.inv (Fin+ {n = zero} {m = m}) x = inr x -Iso.rightInv (Fin+ {n = zero} {m = m}) x = refl -Iso.leftInv (Fin+ {n = zero} {m = m}) (inr x) = refl -Fin+ {n = suc n} {m = m} = - compIso (⊎Iso (invIso FinSuc) idIso) - (compIso ⊎-assoc-Iso - (compIso (⊎Iso idIso (Fin+ {n = n} {m = m})) - FinSuc)) - -Iso-Unit-Fin : Iso Unit (Fin 1) -Iso.fun Iso-Unit-Fin tt = fzero -Iso.inv Iso-Unit-Fin (x , p) = tt -Iso.rightInv Iso-Unit-Fin (zero , p) = Σ≡Prop (λ _ → isProp<ᵗ) refl -Iso.leftInv Iso-Unit-Fin x = refl - -Iso-Bool-Fin : Iso (S₊ 0) (Fin 2) -Iso.fun Iso-Bool-Fin false = flast -Iso.fun Iso-Bool-Fin true = fzero -Iso.inv Iso-Bool-Fin (zero , p) = true -Iso.inv Iso-Bool-Fin (suc x , p) = false -Iso.rightInv Iso-Bool-Fin (zero , p) = refl -Iso.rightInv Iso-Bool-Fin (suc zero , p) = - Σ≡Prop (λ _ → isProp<ᵗ) refl -Iso.leftInv Iso-Bool-Fin false = refl -Iso.leftInv Iso-Bool-Fin true = refl - -Fin× : {n m : ℕ} → Iso (Fin n × Fin m) (Fin (n · m)) -Fin× {n = zero} {m = m} = - iso (λ {()}) (λ{()}) (λ{()}) (λ{()}) -Fin× {n = suc n} {m = m} = - compIso - (compIso - (compIso (Σ-cong-iso-fst (invIso FinSuc)) - (compIso Σ-swap-Iso - (compIso ×DistR⊎Iso - (⊎Iso (compIso - (Σ-cong-iso-snd (λ _ → invIso Iso-Unit-Fin)) rUnit×Iso) - Σ-swap-Iso)))) - (⊎Iso idIso Fin×)) - (Fin+ {n = m} {n · m}) - -DiscreteFin : ∀ {n} → Discrete (Fin n) -DiscreteFin x y with discreteℕ (fst x) (fst y) -... | yes p = yes (Σ≡Prop (λ _ → isProp<ᵗ) p) -... | no ¬p = no λ q → ¬p (cong fst q) - - -decIm : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') -decIm {A = A} {B = B} f = - (y : B) → (Σ[ x ∈ A ] f x ≡ y) - ⊎ ((x : A) → ¬ f x ≡ y) - -decImFin : ∀ {ℓ} {A : Type ℓ} - (da : Discrete A) (n : ℕ) (f : Fin n → A) - → decIm f -decImFin {A = A} da zero f y = inr λ x → ⊥.rec (¬Fin0 x) -decImFin {A = A} da (suc n) f y with da (f fzero) y -... | yes p = inl (fzero , p) -... | no ¬p with (decImFin da n (f ∘ fsuc) y) -... | inl q = inl ((fsuc (fst q)) , snd q) -... | inr q = inr (Ind.elimFin-alt ¬p q) - -Bool×Fin≡Fin : {n : ℕ} → Iso (S₊ 0 × Fin n) (Fin (2 · n)) -Bool×Fin≡Fin = compIso (Σ-cong-iso-fst Iso-Bool-Fin) (Fin× {n = 2}) - -decIm-S⁰×Fin : ∀ {ℓ} {A : Type ℓ} - (da : Discrete A) (n : ℕ) (f : S₊ 0 × Fin n → A) → decIm f -decIm-S⁰×Fin {A = A} da n = - subst (λ C → (f : C → A) → decIm f) - (isoToPath (invIso Bool×Fin≡Fin)) - (decImFin da _) - - -module _ {ℓ : Level} {n : ℕ} {A : Fin n → Type ℓ} (x₀ : Fin n) - (pt : A x₀) (l : (x : Fin n) → ¬ x ≡ x₀ → A x) where - private - x = fst x₀ - p = snd x₀ - elimFinChoose : (x : _) → A x - elimFinChoose (x' , q) with (discreteℕ x x') - ... | yes p₁ = subst A (Σ≡Prop (λ _ → isProp<ᵗ) p₁) pt - ... | no ¬p = l (x' , q) λ r → ¬p (sym (cong fst r)) - - elimFinChooseβ : (elimFinChoose x₀ ≡ pt) - × ((x : _) (q : ¬ x ≡ x₀) → elimFinChoose x ≡ l x q) - fst elimFinChooseβ with (discreteℕ x x) - ... | yes p₁ = (λ j → subst A (isSetFin x₀ x₀ (Σ≡Prop (λ z → isProp<ᵗ) p₁) refl j) pt) - ∙ transportRefl pt - ... | no ¬p = ⊥.rec (¬p refl) - snd elimFinChooseβ (x' , q) with (discreteℕ x x') - ... | yes p₁ = λ q → ⊥.rec (q (Σ≡Prop (λ _ → isProp<ᵗ) (sym p₁))) - ... | no ¬p = λ s → cong (l (x' , q)) (isPropΠ (λ _ → isProp⊥) _ _) - - -pickDifferentFin : {n : ℕ} (x : Fin (suc (suc n))) → Σ[ y ∈ Fin (suc (suc n)) ] ¬ x ≡ y -pickDifferentFin (zero , p) = (1 , p) , λ q → snotz (sym (cong fst q)) -pickDifferentFin (suc x , p) = fzero , λ q → snotz (cong fst q) - -allConst? : ∀ {ℓ} {A : Type ℓ} {n : ℕ} (dis : Discrete A) - → (t : Fin n → S₊ 0 → A) - → ((x : Fin n) → t x ≡ λ _ → t x true) - ⊎ (Σ[ x ∈ Fin n ] ¬ (t x true ≡ t x false)) -allConst? {n = zero} dis t = inl λ {()} -allConst? {n = suc n} dis t - with (dis (t fzero true) (t fzero false)) - | (allConst? {n = n} dis λ x p → t (fsuc x) p) -... | yes p | inl x = - inl (Ind.elimFin-alt (funExt - (λ { false → sym p ; true → refl})) x) -... | yes p | inr x = inr (_ , (snd x)) -... | no ¬p | q = inr (_ , ¬p) - - -isSurj-α₀ : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) - → isConnected 2 (Pushout f snd) - → (y : _) → Σ[ x ∈ _ ] f x ≡ y -isSurj-α₀ n m f c y with (decIm-S⁰×Fin DiscreteFin n f y) -... | inl x = x -isSurj-α₀ n m f c x₀ | inr q = ⊥.rec nope - where - pre-help : Fin (suc (suc m)) → Type - pre-help = elimFinChoose x₀ ⊥ λ _ _ → Unit - - lem : (fa : _) (a : _) → f a ≡ fa → pre-help fa ≡ Unit - lem = elimFinChoose x₀ - (λ s t → ⊥.rec (q _ t)) - λ s t b c → elimFinChooseβ x₀ ⊥ (λ _ _ → Unit) .snd s t - - help : Pushout f snd → Type - help (inl x) = pre-help x - help (inr x) = Unit - help (push a i) = lem (f a) a refl i - - nope : ⊥ - nope = TR.rec isProp⊥ - (λ q → transport (elimFinChooseβ x₀ ⊥ (λ _ _ → Unit) .fst) - (subst help (sym q) - (transport (sym (elimFinChooseβ x₀ ⊥ - (λ _ _ → Unit) .snd _ - (pickDifferentFin x₀ .snd ∘ sym))) tt))) - (Iso.fun (PathIdTruncIso 1) - (isContr→isProp c - ∣ inl x₀ ∣ - ∣ inl (pickDifferentFin x₀ .fst) ∣)) - - -notAllLoops-α₀ : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) - → isConnected 2 (Pushout f snd) - → Σ[ x ∈ Fin n ] (¬ f (true , x) ≡ f (false , x)) -notAllLoops-α₀ n m f c with (allConst? DiscreteFin (λ x y → f (y , x))) -... | inr x = x -notAllLoops-α₀ n m f c | inl q = - ⊥.rec (TR.rec isProp⊥ (λ p → subst T p tt) - (Iso.fun(PathIdTruncIso 1) - (isContr→isProp c - ∣ inl flast ∣ ∣ inl fzero ∣))) - where - inrT : Fin n → Type - inrT x with (DiscreteFin (f (true , x)) fzero) - ... | yes p = ⊥ - ... | no ¬p = Unit - - inlT : Fin (suc (suc m)) → Type - inlT (zero , p) = ⊥ - inlT (suc x , p) = Unit - - inlrT-pre : (a : _) → inlT (f (true , a)) ≡ inrT a - inlrT-pre a with ((DiscreteFin (f (true , a)) fzero)) - ... | yes p = cong inlT p - inlrT-pre s | no ¬p with (f (true , s)) - ... | zero , tt = ⊥.rec (¬p refl) - ... | suc q , t = refl - - inlrT : (a : _) → inlT (f a) ≡ inrT (snd a) - inlrT (false , b) = cong inlT (funExt⁻ (q b) false) ∙ inlrT-pre b - inlrT (true , b) = inlrT-pre b - - T : Pushout f snd → Type - T (inl x) = inlT x - T (inr x) = inrT x - T (push a i) = inlrT a i - -module _ {n : ℕ} where - swapFin : (x y : Fin n) → Fin n → Fin n - swapFin (x , xp) (y , yp) (z , zp) with (discreteℕ z x) | (discreteℕ z y) - ... | yes p | yes p₁ = z , zp - ... | yes p | no ¬p = y , yp - ... | no ¬p | yes p = x , xp - ... | no ¬p | no ¬p₁ = (z , zp) - - swapFinβₗ : (x y : Fin n) → swapFin x y x ≡ y - swapFinβₗ (x , xp) (y , yp) with (discreteℕ x x) | discreteℕ x y - ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) p₁ - ... | yes p | no ¬p = refl - ... | no ¬p | q = ⊥.rec (¬p refl) - - swapFinβᵣ : (x y : Fin n) → swapFin x y y ≡ x - swapFinβᵣ (x , xp) (y , yp) with (discreteℕ y y) | discreteℕ y x - ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) p₁ - ... | yes p | no ¬p = refl - ... | no ¬p | q = ⊥.rec (¬p refl) - - -- swapFinSwap : (x y z : Fin n) → swapFin x y z ≡ swapFin y x z - -- swapFinSwap x y z with (discreteℕ (fst z) (fst x)) | discreteℕ (fst z) (fst y) - -- ... | yes p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p₁ ∙ p) - -- ... | yes p | no ¬p = {!help!} - -- where - -- help : y ≡ swapFin y x z - -- help = {!!} - -- ... | no ¬p | q = {!!} - - swapFin² : (x y z : Fin n) → swapFin x y (swapFin x y z) ≡ z - swapFin² (x , xp) (y , yp) (z , zp) with discreteℕ z x | discreteℕ z y - ... | yes p | yes p₁ = silly - where - silly : swapFin (x , xp) (y , yp) (z , zp) ≡ (z , zp) - silly with discreteℕ z x | discreteℕ z y - ... | yes p | yes p₁ = refl - ... | yes p | no ¬p = ⊥.rec (¬p p₁) - ... | no ¬p | r = ⊥.rec (¬p p) - ... | yes p | no ¬q = silly - where - silly : swapFin (x , xp) (y , yp) (y , yp) ≡ (z , zp) - silly with discreteℕ y x | discreteℕ y y - ... | yes p' | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (p' ∙ sym p) - ... | no ¬p | yes p₁ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p) - ... | p | no ¬p = ⊥.rec (¬p refl) - ... | no ¬p | yes p = silly - where - silly : swapFin (x , xp) (y , yp) (x , xp) ≡ (z , zp) - silly with discreteℕ x y | discreteℕ x x - ... | yes p₁ | yes _ = Σ≡Prop (λ _ → isProp<ᵗ) (p₁ ∙ sym p) - ... | no ¬p | yes _ = Σ≡Prop (λ _ → isProp<ᵗ) (sym p) - ... | s | no ¬p = ⊥.rec (¬p refl) - ... | no ¬p | no ¬q = silly - where - silly : swapFin (x , xp) (y , yp) (z , zp) ≡ (z , zp) - silly with discreteℕ z x | discreteℕ z y - ... | yes p | yes p₁ = refl - ... | yes p | no ¬b = ⊥.rec (¬p p) - ... | no ¬a | yes b = ⊥.rec (¬q b) - ... | no ¬a | no ¬b = refl - - swapFinIso : (x y : Fin n) → Iso (Fin n) (Fin n) - Iso.fun (swapFinIso x y) = swapFin x y - Iso.inv (swapFinIso x y) = swapFin x y - Iso.rightInv (swapFinIso x y) = swapFin² x y - Iso.leftInv (swapFinIso x y) = swapFin² x y - -Pushout∘Equiv : ∀ {ℓ ℓ' ℓ''} {A A' : Type ℓ} {B B' : Type ℓ'} {C : Type ℓ''} - (e1 : A ≃ A') (e2 : B' ≃ B) - (f : A' → B') (g : A → C) - → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) (Pushout f (g ∘ invEq e1)) -Pushout∘Equiv {A = A} {A' = A'} {B} {B'} {C} = - EquivJ (λ A e1 → (e2 : B' ≃ B) (f : A' → B') (g : A → C) - → Iso (Pushout (fst e2 ∘ f ∘ fst e1) g) - (Pushout f (g ∘ invEq e1))) - (EquivJ (λ B' e2 → (f : A' → B') (g : A' → C) - → Iso (Pushout (fst e2 ∘ f) g) - (Pushout f g)) - λ f g → idIso) - -module _ {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} - (f : Bool × A → Unit ⊎ B) (b₀ : B) where - - F : Bool × (Unit ⊎ A) → Unit ⊎ B - F (false , inl tt) = inl tt - F (true , inl tt) = inr b₀ - F (x , inr a) = f (x , a) - - g : Unit ⊎ B → B - g (inl x) = b₀ - g (inr x) = x - - PF→P∘ₗ : (x : Unit ⊎ A) → Pushout (g ∘ f) snd - PF→P∘ₗ (inl x) = inl b₀ - PF→P∘ₗ (inr x) = inr x - - theCoh : (a : _) → inl (g (F a)) ≡ PF→P∘ₗ (snd a) - theCoh (false , inl x) = refl - theCoh (true , inl x) = refl - theCoh (false , inr x) = push (false , x) - theCoh (true , inr x) = push (true , x) - - PF→P∘' : Pushout F snd → Pushout (g ∘ f) snd - PF→P∘' (inl x) = inl (g x) - PF→P∘' (inr x) = PF→P∘ₗ x - PF→P∘' (push a i) = theCoh a i - - theCoh2 : (a : _) (b : _) - → Path (Pushout F snd) (inl (inr (g (f (b , a))))) (inl (f (b , a))) - theCoh2 a b with (f (b , a)) - theCoh2 a b | inl x = (push (true , inl tt)) ∙ sym (push (false , inl tt)) - ... | inr x = refl - - P∘'→PF : Pushout (g ∘ f) snd → Pushout F snd - P∘'→PF (inl x) = inl (inr x) - P∘'→PF (inr x) = inr (inr x) - P∘'→PF (push (false , a) i) = (theCoh2 a false ∙ push (false , inr a)) i - P∘'→PF (push (true , a) i) = (theCoh2 a true ∙ push (true , inr a)) i - - PFpushTₗ : (x : _) → P∘'→PF (PF→P∘' (inl x)) ≡ (inl x) - PFpushTₗ (inl x) = push (true , inl tt) ∙ sym (push (false , inl tt)) - PFpushTₗ (inr x) = refl - - PFpushTᵣ : (x : _) → P∘'→PF (PF→P∘' (inr x)) ≡ (inr x) - PFpushTᵣ (inl x) = push (true , inl tt) - PFpushTᵣ (inr x) = refl - - pp1 : (x : A) → PFpushTₗ (f (false , x)) ≡ theCoh2 x false - pp1 x with (f (false , x)) - ... | inl x₁ = refl - ... | inr x₁ = refl - - pp2 : (x : A) → PFpushTₗ (f (true , x)) ≡ theCoh2 x true - pp2 x with (f (true , x)) - ... | inl x₁ = refl - ... | inr x₁ = refl - - open import Cubical.Foundations.Path - open import Cubical.Foundations.GroupoidLaws - - - PFpushT : (x : _) → P∘'→PF (PF→P∘' x) ≡ x - PFpushT (inl x) = PFpushTₗ x - PFpushT (inr x) = PFpushTᵣ x - PFpushT (push (false , inl x) i) j = - compPath-filler (push (true , inl tt)) (sym (push (false , inl tt))) (~ i) j - PFpushT (push (false , inr x) i) j = - (pp1 x - ◁ flipSquare - (symP (compPath-filler' - (theCoh2 x false) (push (false , inr x))))) i j - PFpushT (push (true , inl x) i) j = push (true , inl x) (i ∧ j) - PFpushT (push (true , inr x) i) j = - (pp2 x - ◁ flipSquare - (symP (compPath-filler' - (theCoh2 x true) (push (true , inr x))))) i j - - cong-PF→P∘' : (b : _) (a : _) → cong PF→P∘' (theCoh2 b a) ≡ refl - cong-PF→P∘' b a with (f (a , b)) - ... | inl x = cong-∙ PF→P∘' (push (true , inl tt)) (sym (push (false , inl tt))) - ∙ sym (rUnit refl) - ... | inr x = refl - - PF→P∘'→PF : (x : _) → PF→P∘' (P∘'→PF x) ≡ x - PF→P∘'→PF (inl x) = refl - PF→P∘'→PF (inr x) = refl - PF→P∘'→PF (push (false , b) i) j = - (cong-∙ PF→P∘' (theCoh2 b false) (push (false , inr b)) - ∙ cong (_∙ push (false , b)) (cong-PF→P∘' b false) - ∙ sym (lUnit _)) j i - PF→P∘'→PF (push (true , b) i) j = - (cong-∙ PF→P∘' (theCoh2 b true) (push (true , inr b)) - ∙ cong (_∙ push (true , b)) (cong-PF→P∘' b true) - ∙ sym (lUnit _)) j i - - Is1 : Iso (Pushout F snd) (Pushout (g ∘ f) snd) - Iso.fun Is1 = PF→P∘' - Iso.inv Is1 = P∘'→PF - Iso.rightInv Is1 = PF→P∘'→PF - Iso.leftInv Is1 = PFpushT - -FinPred : ∀ {m} → Fin (suc (suc m)) → Fin (suc m) -FinPred {m = m} (zero , p) = zero , p -FinPred {m = m} (suc x , p) = x , p - -fone : ∀ {m} → Fin (suc (suc m)) -fone {m} = suc zero , tt - -module _ {m : ℕ} where - Fin→Unit⊎Fin : (x : Fin (suc m)) → Unit ⊎ Fin m - Fin→Unit⊎Fin = Ind.elimFin (inl tt) inr - - Unit⊎Fin→Fin : Unit ⊎ Fin m → Fin (suc m) - Unit⊎Fin→Fin (inl x) = flast - Unit⊎Fin→Fin (inr x) = injectSuc x - - Iso-Fin-Unit⊎Fin : Iso (Fin (suc m)) (Unit ⊎ Fin m) - Iso.fun Iso-Fin-Unit⊎Fin = Fin→Unit⊎Fin - Iso.inv Iso-Fin-Unit⊎Fin = Unit⊎Fin→Fin - Iso.rightInv Iso-Fin-Unit⊎Fin (inl x) = Ind.elimFinβ (inl tt) inr .fst - Iso.rightInv Iso-Fin-Unit⊎Fin (inr x) = Ind.elimFinβ (inl tt) inr .snd x - Iso.leftInv Iso-Fin-Unit⊎Fin = - Ind.elimFin - (cong Unit⊎Fin→Fin (Ind.elimFinβ (inl tt) inr .fst)) - λ x → (cong Unit⊎Fin→Fin (Ind.elimFinβ (inl tt) inr .snd x)) - - -≠flast→<ᵗflast : {n : ℕ} → (x : Fin (suc n)) → ¬ x ≡ flast → fst x <ᵗ n -≠flast→<ᵗflast = Ind.elimFin (λ p → ⊥.rec (p refl)) λ p _ → snd p - -CW₁DataPre : (n m : ℕ) (f : S₊ 0 × Fin (suc (suc n)) → Fin (suc (suc m))) - → f (false , flast) ≡ flast - → (t : f (true , flast) .fst <ᵗ suc m) - → Σ[ k ∈ ℕ ] Σ[ f' ∈ (S₊ 0 × Fin k → Fin (suc m)) ] - Iso (Pushout f snd) - (Pushout f' snd) -CW₁DataPre n m f p q = (suc n) - , _ - , compIso (invIso (pushoutIso _ _ _ _ - (isoToEquiv (Σ-cong-iso-snd λ _ → invIso Iso-Fin-Unit⊎Fin)) - (isoToEquiv (invIso Iso-Fin-Unit⊎Fin)) - (isoToEquiv (invIso Iso-Fin-Unit⊎Fin)) - (funExt (uncurry help)) - (funExt λ x → refl))) - (Is1 {A = Fin (suc n)} {B = Fin (suc m)} - (λ x → Fin→Unit⊎Fin (f (fst x , injectSuc (snd x)))) - (f (true , flast) .fst , q)) - where - help : (x : Bool) (y : Unit ⊎ Fin (suc n)) - → Unit⊎Fin→Fin - (F (λ x₁ → Ind.elimFin (inl tt) inr (f (fst x₁ , injectSuc (snd x₁)))) - (f (true , flast) .fst , q) (x , y)) - ≡ f (x , Unit⊎Fin→Fin y) - help false (inl a) = sym p - help true (inl b) = Σ≡Prop (λ _ → isProp<ᵗ) refl - help false (inr a) = Iso.leftInv Iso-Fin-Unit⊎Fin _ - help true (inr a) = Iso.leftInv Iso-Fin-Unit⊎Fin _ - -isPropFin1 : isProp (Fin 1) -isPropFin1 (zero , tt) (zero , tt) = refl - - -Iso⊎→Iso : ∀ {ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} - → (f : Iso A C) - → (e : Iso (A ⊎ B) (C ⊎ D)) - → ((a : A) → Iso.fun e (inl a) ≡ inl (Iso.fun f a)) - → Iso B D -Iso⊎→Iso {A = A} {B} {C} {D} f e p = Iso' - where - ⊥-fib : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ⊎ B → Type - ⊥-fib (inl x) = ⊥ - ⊥-fib (inr x) = Unit - - module _ {ℓ ℓ' ℓ'' ℓ''' : Level} - {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} - (f : Iso A C) - (e : Iso (A ⊎ B) (C ⊎ D)) - (p : (a : A) → Iso.fun e (inl a) ≡ inl (Iso.fun f a)) where - T : (b : B) → Type _ - T b = Σ[ d' ∈ C ⊎ D ] (Iso.fun e (inr b) ≡ d') - - T-elim : ∀ {ℓ} (b : B) {P : (x : T b) → Type ℓ} - → ((d : D) (s : _) → P (inr d , s)) - → (x : _) → P x - T-elim b ind (inl x , q) = - ⊥.rec (subst ⊥-fib (sym (sym (Iso.leftInv e _) - ∙ cong (Iso.inv e) - (p _ ∙ cong inl (Iso.rightInv f x) ∙ sym q) - ∙ Iso.leftInv e _)) tt) - T-elim b ind (inr x , y) = ind x y - - e-pres-inr-help : (b : B) → T f e p b → D - e-pres-inr-help b = T-elim f e p b λ d _ → d - - p' : (a : C) → Iso.inv e (inl a) ≡ inl (Iso.inv f a) - p' c = cong (Iso.inv e ∘ inl) (sym (Iso.rightInv f c)) - ∙∙ cong (Iso.inv e) (sym (p (Iso.inv f c))) - ∙∙ Iso.leftInv e _ - - e⁻-pres-inr-help : (d : D) → T (invIso f) (invIso e) p' d → B - e⁻-pres-inr-help d = T-elim (invIso f) (invIso e) p' d λ b _ → b - - e-pres-inr : B → D - e-pres-inr b = e-pres-inr-help b (_ , refl) - - e⁻-pres-inr : D → B - e⁻-pres-inr d = e⁻-pres-inr-help d (_ , refl) - - lem1 : (b : B) (e : T f e p b) (d : _) - → e⁻-pres-inr-help (e-pres-inr-help b e) d ≡ b - lem1 b = T-elim f e p b λ d s - → T-elim (invIso f) (invIso e) p' _ - λ b' s' → invEq (_ , isEmbedding-inr _ _) - (sym s' ∙ cong (Iso.inv e) (sym s) ∙ Iso.leftInv e _) - - lem2 : (d : D) (e : T (invIso f) (invIso e) p' d ) (t : _) - → e-pres-inr-help (e⁻-pres-inr-help d e) t ≡ d - lem2 d = T-elim (invIso f) (invIso e) p' d - λ b s → T-elim f e p _ λ d' s' - → invEq (_ , isEmbedding-inr _ _) - (sym s' ∙ cong (Iso.fun e) (sym s) ∙ Iso.rightInv e _) - - Iso' : Iso B D - Iso.fun Iso' = e-pres-inr - Iso.inv Iso' = e⁻-pres-inr - Iso.rightInv Iso' x = lem2 x _ _ - Iso.leftInv Iso' x = lem1 x _ _ - -Fin≠Fin : {n m : ℕ} → ¬ (n ≡ m) → ¬ (Iso (Fin n) (Fin m)) -Fin≠Fin {n = zero} {m = zero} p = ⊥.rec (p refl) -Fin≠Fin {n = zero} {m = suc m} p q = Iso.inv q fzero .snd -Fin≠Fin {n = suc n} {m = zero} p q = Iso.fun q fzero .snd -Fin≠Fin {n = suc n} {m = suc m} p q = - Fin≠Fin {n = n} {m = m} (p ∘ cong suc) - (Iso⊎→Iso idIso help λ {(zero , tt) - → cong (Iso.inv FinSuc) (swapFinβₗ (Iso.fun q flast) flast) - ∙ Ind.elimFinβ (inl flast) inr .fst}) - where - q^ : Iso (Fin (suc n)) (Fin (suc m)) - q^ = compIso q (swapFinIso (Iso.fun q flast) flast) - - help : Iso (Fin 1 ⊎ Fin n) (Fin 1 ⊎ Fin m) - help = compIso FinSuc (compIso q^ (invIso FinSuc)) - -CW₁Data₁ : (m : ℕ) (f : S₊ 0 × Fin 1 → Fin (suc (suc m))) - → isConnected 2 (Pushout f snd) - → Iso (S₊ 0 × Fin 1) (Fin (suc (suc m))) -CW₁Data₁ m f c = mainIso - where - f' : Bool → Fin (suc (suc m)) - f' = f ∘ (_, fzero) - - f'-surj : (x : _) → Σ[ t ∈ Bool ] (f' t ≡ x) - f'-surj x = - isSurj-α₀ (suc zero) m f c x .fst .fst - , cong f (Σ≡Prop (λ _ → λ {(zero , p) (zero , q) → refl}) refl) - ∙ isSurj-α₀ (suc zero) m f c x .snd - - f-true≠f-false : (x : _) → f' true ≡ x → ¬ f' true ≡ f' false - f-true≠f-false (zero , q) p r = lem (f'-surj fone) - where - lem : Σ[ t ∈ Bool ] (f' t ≡ fone) → ⊥ - lem (false , s) = snotz (cong fst (sym s ∙ sym r ∙ p)) - lem (true , s) = snotz (cong fst (sym s ∙ p)) - f-true≠f-false (suc x , q) p r = lem (f'-surj fzero) - where - lem : Σ[ t ∈ Bool ] (f' t ≡ fzero) → ⊥ - lem (false , s) = snotz (cong fst (sym p ∙ r ∙ s)) - lem (true , s) = snotz (cong fst (sym p ∙ s)) - - f-inj : (x y : _) → f x ≡ f y → x ≡ y - f-inj (false , zero , tt) (false , zero , tt) p = refl - f-inj (false , zero , tt) (true , zero , tt) p = ⊥.rec (f-true≠f-false _ refl (sym p)) - f-inj (true , zero , tt) (false , zero , tt) p = ⊥.rec (f-true≠f-false _ refl p) - f-inj (true , zero , tt) (true , zero , tt) p = refl - - mainIso : Iso (S₊ 0 × Fin 1) (Fin (suc (suc m))) - Iso.fun mainIso = f - Iso.inv mainIso x = isSurj-α₀ (suc zero) m f c x .fst - Iso.rightInv mainIso x = isSurj-α₀ 1 m f c x .snd - Iso.leftInv mainIso (x , zero , tt) = - (f-inj _ _ (isSurj-α₀ 1 m f c (f (x , fzero)) .snd)) - -CW₁Data : (n m : ℕ) (f : S₊ 0 × Fin n → Fin (suc (suc m))) - → isConnected 2 (Pushout f snd) - → Σ[ k ∈ ℕ ] Σ[ f' ∈ (S₊ 0 × Fin k → Fin (suc m)) ] - Iso (Pushout f snd) - (Pushout f' snd) -CW₁Data zero m f c = ⊥.rec (snd (notAllLoops-α₀ zero m f c .fst)) -CW₁Data (suc zero) zero f c = 0 , ((λ ()) , compIso isoₗ (PushoutEmptyFam (snd ∘ snd) snd)) - where - isoₗ : Iso (Pushout f snd) (Fin 1) - isoₗ = PushoutAlongEquiv (isoToEquiv (CW₁Data₁ zero f c)) _ -CW₁Data (suc zero) (suc m) f c = - ⊥.rec (Fin≠Fin (snotz ∘ sym ∘ cong (predℕ ∘ predℕ)) - mainIso) - where - mainIso : Iso (Fin 2) (Fin (3 +ℕ m)) - mainIso = - compIso - (compIso - (invIso rUnit×Iso) - (Σ-cong-iso - (invIso Iso-Bool-Fin) - (λ _ → isContr→Iso {A = Unit} - (tt , λ _ → refl) (fzero , λ {(zero , tt) → refl})))) - (CW₁Data₁ (suc m) f c) -CW₁Data (suc (suc n)) m f c = - main .fst , main .snd .fst - , compIso correct (main .snd .snd) - where - t = notAllLoops-α₀ (suc (suc n)) m f c - - abstract - x₀ : Fin (suc (suc n)) - x₀ = fst t - - xpath : ¬ (f (true , x₀) ≡ f (false , x₀)) - xpath = snd t - - Fin0-iso : Iso (S₊ 0 × Fin (suc (suc n))) (S₊ 0 × Fin (suc (suc n))) - Fin0-iso = Σ-cong-iso-snd λ _ → swapFinIso flast x₀ - - FinIso2 : Iso (Fin (suc (suc m))) (Fin (suc (suc m))) - FinIso2 = swapFinIso (f (false , x₀)) flast - - f' : S₊ 0 × Fin (suc (suc n)) → Fin (suc (suc m)) - f' = Iso.fun FinIso2 ∘ f ∘ Iso.fun Fin0-iso - - f'≡ : f' (false , flast) ≡ flast - f'≡ = cong (Iso.fun FinIso2 ∘ f) - (cong (false ,_) (swapFinβₗ flast x₀)) - ∙ swapFinβₗ (f (false , x₀)) flast - - f'¬≡ : ¬ (f' (true , flast) ≡ flast) - f'¬≡ p = xpath (cong f (ΣPathP (refl , sym (swapFinβₗ flast x₀))) - ∙ sym (Iso.rightInv FinIso2 _) - ∙ cong (Iso.inv FinIso2) (p ∙ sym f'≡) - ∙ Iso.rightInv FinIso2 _ - ∙ cong f (ΣPathP (refl , swapFinβₗ flast x₀))) - - f'< : fst (f' (true , flast)) <ᵗ suc m - f'< = ≠flast→<ᵗflast _ f'¬≡ - - main = CW₁DataPre _ _ f' f'≡ f'< - - Upath = isoToPath (swapFinIso x₀ fzero) - - correct : Iso (Pushout f snd) (Pushout f' snd) - correct = pushoutIso _ _ _ _ - (isoToEquiv Fin0-iso) (isoToEquiv FinIso2) (isoToEquiv (swapFinIso flast x₀)) - (funExt (λ x → cong (FinIso2 .Iso.fun ∘ f) (sym (Iso.rightInv Fin0-iso x)))) - refl - - -CW₁Data' : (n m : ℕ) (f : S₊ 0 × Fin n → Fin m) - → isConnected 2 (Pushout f snd) - → Σ[ k ∈ ℕ ] Σ[ f' ∈ (S₊ 0 × Fin k → Fin 1) ] - Iso (Pushout f snd) - (Pushout f' snd) -CW₁Data' zero zero f c = ⊥.rec (TR.rec (λ()) help (fst c)) - where - help : ¬ Pushout f snd - help = elimProp _ (λ _ → λ ()) snd snd -CW₁Data' (suc n) zero f c = ⊥.rec (f (true , fzero) .snd) -CW₁Data' n (suc zero) f c = n , f , idIso -CW₁Data' zero (suc (suc m)) f c = - ⊥.rec (TR.rec (λ()) (snotz ∘ sym ∘ cong fst) - (Iso.fun (PathIdTruncIso _) - (isContr→isProp (subst (isConnected 2) (isoToPath help) c) - ∣ fzero ∣ ∣ fone ∣))) - where - help : Iso (Pushout f snd) (Fin (suc (suc m))) - help = invIso (PushoutEmptyFam (λ()) λ()) -CW₁Data' (suc n) (suc (suc m)) f c - with (CW₁Data' _ (suc m) (CW₁Data (suc n) m f c .snd .fst) - (subst (isConnected 2) - (isoToPath (CW₁Data (suc n) m f c .snd .snd)) c)) -... | (k , f↓ , e) = k , f↓ , compIso (CW₁Data (suc n) m f c .snd .snd) e - -CW₁Data'' : (n m : ℕ) (f : Fin n × S₊ 0 → Fin m) - → isConnected 2 (Pushout f snd) - → Σ[ k ∈ ℕ ] Σ[ f' ∈ (Fin k × S₊ 0 → Fin 1) ] - Iso (Pushout f snd) - (Pushout f' snd) -CW₁Data'' n m f c = - fst help - , (help .snd .fst ∘ Σ-swap-≃ .fst) - , compIso {!help .snd .snd!} {!!} - where - help : {!!} - help = CW₁Data' n m (f ∘ Σ-swap-≃ .fst) {!!} - - -yieldsConnectedCWskel : (ℕ → Type ℓ) → Type _ -yieldsConnectedCWskel A = Σ[ sk ∈ yieldsCWskel A ] (sk .fst 1 ≡ 1) - -yieldsConnectedCWskel' : (ℕ → Type ℓ) → Type _ -yieldsConnectedCWskel' A = Σ[ sk ∈ yieldsCWskel A ] isConnected 2 (realise (_ , sk)) - -connectedCWskel : (ℓ : Level) → Type (ℓ-suc ℓ) -connectedCWskel ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel X) - -connectedCWskel' : (ℓ : Level) → Type (ℓ-suc ℓ) -connectedCWskel' ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel' X) - -truncCWIso : (A : CWskel ℓ) (n : ℕ) - → Iso (hLevelTrunc n (realise A)) (hLevelTrunc n (fst A n)) -truncCWIso A n = invIso (connectedTruncIso n incl (isConnected-CW↪∞ n A)) - -isConnectedColim→isConnectedSkel : - (A : CWskel ℓ) (n : ℕ) - → isConnected n (realise A) - → isConnected n (fst A n) -isConnectedColim→isConnectedSkel A n c = - isOfHLevelRetractFromIso 0 - (invIso (truncCWIso A n)) c - -module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A) where - private - c = snd sk+c - sk = fst sk+c - - module AC = CWskel-fields (_ , sk) - - liftStr = CW₁Data' {!!} {!!} {!AC.α!} {!isConnectedColim→isConnectedSkel (_ , sk) 2 c!} - - - - - connectedCWskel→ : yieldsConnectedCWskel' A - connectedCWskel→ = ({!!} , {!!}) , {!!} - - - --- yieldsGoodCWSkel : (ℕ → Type ℓ) → Type ℓ --- yieldsGoodCWSkel = ? - --- yieldsCWskel∙ : (ℕ → Type ℓ) → Type ℓ --- yieldsCWskel∙ X = --- Σ[ f ∈ (ℕ → ℕ) ] --- Σ[ α ∈ ((n : ℕ) → ⋁gen (Fin (f (suc n))) (λ _ → S₊∙ n) → X n) ] --- ((X zero ≃ Fin (f zero)) × --- (((n : ℕ) → X (suc n) ≃ cofib (α n)))) - --- CWskel∙ : (ℓ : Level) → Type (ℓ-suc ℓ) --- CWskel∙ ℓ = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCWskel∙ X) - --- module CWskel∙-fields (C : CWskel∙ ℓ) where --- card = C .snd .fst --- A = Fin ∘ card --- α = C .snd .snd .fst --- e = C .snd .snd .snd .snd - --- ℤ[A_] : (n : ℕ) → AbGroup ℓ-zero --- ℤ[A n ] = ℤ[Fin (snd C .fst n) ] - --- CWpt : ∀ {ℓ} → (C : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- fst (CWpt (C , f) n) = C n --- snd (CWpt (C , f) n) = f .snd .fst n (inl tt) - --- -- Technically, a CW complex should be the sequential colimit over the following maps --- CW∙↪ : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n → fst T (suc n) --- CW∙↪ (X , f , α , e₀ , e₊) n x = invEq (e₊ n) (inr x) - --- ptCW : (T : CWskel∙ ℓ) → (n : ℕ) → fst T n --- ptCW T zero = T .snd .snd .fst zero (inl tt) --- ptCW T (suc n) = CW∙↪ T n (ptCW T n) - --- CW∙ : (T : CWskel∙ ℓ) → (n : ℕ) → Pointed ℓ --- CW∙ T n = fst T n , ptCW T n - --- CW∙↪∙ : (T : CWskel∙ ℓ) → (n : ℕ) → CW∙ T n →∙ CW∙ T (suc n) --- fst (CW∙↪∙ T n) = CW∙↪ T n --- snd (CW∙↪∙ T n) = refl - - --- -- But if it stabilises, no need for colimits. --- yieldsFinCWskel∙ : (n : ℕ) (X : ℕ → Type ℓ) → Type ℓ --- yieldsFinCWskel∙ n X = --- Σ[ CWskel ∈ yieldsCWskel∙ X ] ((k : ℕ) → isEquiv (CW∙↪ (X , CWskel) (k +ℕ n))) - --- -- ... which should give us finite CW complexes. --- finCWskel∙ : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) --- finCWskel∙ ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel∙ n C) - --- finCWskel→CWskel∙ : (n : ℕ) → finCWskel ℓ n → CWskel ℓ --- finCWskel→CWskel∙ n C = fst C , fst (snd C) - --- realiseSeq∙ : CWskel∙ ℓ → Sequence ℓ --- Sequence.obj (realiseSeq∙ (C , p)) = C --- Sequence.map (realiseSeq∙ C) = CW∙↪ C _ - --- realiseFinSeq∙ : (m : ℕ) → CWskel∙ ℓ → FinSequence m ℓ --- realiseFinSeq∙ m C = Sequence→FinSequence m (realiseSeq∙ C) - --- -- realisation of CW complex from skeleton --- realise∙ : CWskel∙ ℓ → Type ℓ --- realise∙ C = SeqColim (realiseSeq∙ C) - --- realise∙∙ : CWskel∙ ℓ → Pointed ℓ --- realise∙∙ C = SeqColim (realiseSeq∙ C) , incl {n = 0} (CW∙ C 0 .snd) --- open import Cubical.Data.Empty as ⊥ - --- CWskel∙→CWskel : (A : ℕ → Type ℓ) → ℕ → Type ℓ --- CWskel∙→CWskel A zero = Lift ⊥ --- CWskel∙→CWskel A (suc n) = A n --- open import Cubical.Foundations.Isomorphism - - --- module _ (A : ℕ → Type ℓ) --- (cwsk : yieldsCWskel∙ A) where - --- private --- αs : (n : ℕ) → Fin (cwsk .fst n) × S⁻ n → CWskel∙→CWskel A n --- αs (suc n) x = snd cwsk .fst n (inr x) - --- e0 : {!!} --- e0 = {!!} - --- es-suc→ : (n : ℕ) → cofib (fst (snd cwsk) n) → Pushout (αs (suc n)) fst --- es-suc→ n (inl x) = inl (snd cwsk .fst n (inl tt)) --- es-suc→ n (inr x) = inl x --- es-suc→ n (push (inl x) i) = inl (fst (snd cwsk) n (inl x)) --- es-suc→ n (push (inr (a , b)) i) = ((({!!} ∙ λ i → inl {!snd cwsk .snd n!}) ∙ push (a , b)) ∙ {!!}) i -- ({!!} ∙ sym (push (a , b))) i --- es-suc→ n (push (push a i₁) i) = {!!} - --- es-suc : (n : ℕ) --- → Iso (cofib (fst (snd cwsk) n)) --- (Pushout (αs (suc n)) fst) --- Iso.fun (es-suc n) = es-suc→ n --- Iso.inv (es-suc n) = {!!} --- Iso.rightInv (es-suc n) = {!!} --- Iso.leftInv (es-suc n) = {!!} - --- es : (n : ℕ) → A n ≃ Pushout (αs n) (λ r → fst r) --- es zero = {!!} --- es (suc n) = compEquiv (snd cwsk .snd .snd n) (isoToEquiv (es-suc n)) - --- yieldsCWskel∙→' : yieldsCWskel (CWskel∙→CWskel A) --- fst yieldsCWskel∙→' = cwsk .fst --- fst (snd yieldsCWskel∙→') = αs --- fst (snd (snd yieldsCWskel∙→')) () --- snd (snd (snd yieldsCWskel∙→')) = {!!} - --- yieldsCWskel∙→ : (A : ℕ → Type ℓ) --- → yieldsCWskel∙ A → yieldsCWskel (CWskel∙→CWskel A) --- fst (yieldsCWskel∙→ A cwsk) = cwsk .fst --- fst (snd (yieldsCWskel∙→ A cwsk)) (suc n) (x , p) = snd cwsk .fst n (inr (x , p)) --- fst (snd (snd (yieldsCWskel∙→ A cwsk))) () --- snd (snd (snd (yieldsCWskel∙→ A cwsk))) zero = {!!} --- snd (snd (snd (yieldsCWskel∙→ A cwsk))) (suc n) = --- compEquiv (cwsk .snd .snd .snd n) --- (isoToEquiv {!(fst (snd (yieldsCWskel∙→ A cwsk)) (suc n))!}) -- theEq) --- where --- theEq→ : cofib (fst (cwsk .snd) n) → Pushout _ fst --- theEq→ (inl x) = inl (cwsk .snd .fst n (inl tt)) --- theEq→ (inr x) = inl x --- theEq→ (push (inl x) i) = inl (cwsk .snd .fst n (inl tt)) --- theEq→ (push (inr (a , b)) i) = ({!push ?!} ∙ push {!!} ∙ {!!}) i -- inl (cwsk .snd .fst n {!!}) --- theEq→ (push (push a i₁) i) = {!!} - --- theEq : Iso (cofib (fst (cwsk .snd) n)) (Pushout _ fst) --- Iso.fun theEq = theEq→ --- Iso.inv theEq = {!!} --- Iso.rightInv theEq x = {!!} --- Iso.leftInv theEq x = {!!} - - --- -- compEquiv {!!} {!!} - - --- -- -- Finally: definition of CW complexes --- -- isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' - --- -- CW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ - --- -- CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) - --- -- -- Finite CW complexes --- -- isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) --- -- isFinCW {ℓ = ℓ} X = --- -- Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) - --- -- finCW : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ - --- -- finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) --- -- finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) - --- -- isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X --- -- isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str - --- -- finCW→CW : finCW ℓ → CW ℓ --- -- finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p - - --- -- -- morphisms --- -- _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') --- -- C →ᶜʷ D = fst C → fst D - --- -- --the cofibre of the inclusion of X n into X (suc n) --- -- cofibCW : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) → Type ℓ --- -- cofibCW n C = cofib (CW↪ C n) - --- -- --...is basically a sphere bouquet --- -- cofibCW≃bouquet : ∀ {ℓ} (n : ℕ) (C : CWskel ℓ) --- -- → cofibCW n C ≃ SphereBouquet n (Fin (snd C .fst n)) --- -- cofibCW≃bouquet n C = Bouquet≃-gen n (snd C .fst n) (α n) e --- -- where --- -- s = Bouquet≃-gen --- -- α = C .snd .snd .fst --- -- e = C .snd .snd .snd .snd n - --- -- --sending X (suc n) into the cofibCW --- -- to_cofibCW : (n : ℕ) (C : CWskel ℓ) → fst C (suc n) → cofibCW n C --- -- to_cofibCW n C x = inr x - --- -- --the connecting map of the long exact sequence --- -- δ-pre : {A : Type ℓ} {B : Type ℓ'} (conn : A → B) --- -- → cofib conn → Susp A --- -- δ-pre conn (inl x) = north --- -- δ-pre conn (inr x) = south --- -- δ-pre conn (push a i) = merid a i - --- -- δ : (n : ℕ) (C : CWskel ℓ) → cofibCW n C → Susp (fst C n) --- -- δ n C = δ-pre (CW↪ C n) - --- -- -- send the stage n to the realization (the same as incl, but with explicit args and type) --- -- CW↪∞ : (C : CWskel ℓ) → (n : ℕ) → fst C n → realise C --- -- CW↪∞ C n x = incl x - --- -- finCW↑ : (n m : ℕ) → (m ≥ n) → finCWskel ℓ n → finCWskel ℓ m --- -- fst (finCW↑ m n p C) = fst C --- -- fst (snd (finCW↑ m n p C)) = snd C .fst --- -- snd (snd (finCW↑ m n p C)) k = --- -- subst (λ r → isEquiv (CW↪ (fst C , snd C .fst) r)) --- -- (sym (+-assoc k (fst p) m) ∙ cong (k +ℕ_) (snd p)) --- -- (snd C .snd (k +ℕ fst p)) diff --git a/Cubical/CW/Pointed2.agda b/Cubical/CW/Pointed2.agda deleted file mode 100644 index 870efc17ac..0000000000 --- a/Cubical/CW/Pointed2.agda +++ /dev/null @@ -1,1567 +0,0 @@ -{-# OPTIONS --safe --lossy-unification #-} - --- This file contains definition of CW complexes and skeleta. - -module Cubical.CW.Pointed2 where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Function -open import Cubical.Foundations.GroupoidLaws - - -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) -open import Cubical.Data.Nat.Order -open import Cubical.Data.Fin.Inductive.Base -open import Cubical.Data.Sigma -open import Cubical.Data.Sequence -open import Cubical.Data.FinSequence - -open import Cubical.HITs.Sn -open import Cubical.HITs.Pushout -open import Cubical.HITs.Susp -open import Cubical.HITs.SequentialColimit -open import Cubical.HITs.SphereBouquet -open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.HITs.Wedge -open import Cubical.HITs.Pushout - -open import Cubical.Axiom.Choice - -open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup - -open import Cubical.Relation.Nullary - -open import Cubical.CW.Base - -open Sequence - -open import Cubical.Foundations.Equiv.HalfAdjoint - -private - variable - ℓ ℓ' ℓ'' : Level - -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Univalence - -open import Cubical.Foundations.Equiv - --- connected comp -open import Cubical.Homotopy.Connected -open import Cubical.CW.Properties -open import Cubical.HITs.Truncation as TR -open import Cubical.Foundations.HLevels -open import Cubical.Data.Nat.Order.Inductive -open import Cubical.Data.Fin.Inductive.Properties as Ind - - -open import Cubical.Data.Bool -open import Cubical.Data.Nat.Order.Inductive --- open import Cubical.Data.Dec - -open import Cubical.Relation.Nullary.Base -open import Cubical.Data.Sum as ⊎ -open import Cubical.Data.Empty as ⊥ - -open import Cubical.Data.Fin.Inductive.Properties as Ind -module _ {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} - (f : A × Bool → Unit ⊎ B) (b₀ : B) where - - F : (Unit ⊎ A) × Bool → Unit ⊎ B - F (inl tt , false) = inl tt - F (inl tt , true) = inr b₀ - F (inr a , x) = f (a , x) - - g : Unit ⊎ B → B - g (inl x) = b₀ - g (inr x) = x - - PF→P∘ₗ : (x : Unit ⊎ A) → Pushout (g ∘ f) fst - PF→P∘ₗ (inl x) = inl b₀ - PF→P∘ₗ (inr x) = inr x - - theCoh : (a : _) → inl (g (F a)) ≡ PF→P∘ₗ (fst a) - theCoh (inl x , false) = refl - theCoh (inl x , true) = refl - theCoh (inr x , false) = push (x , false) - theCoh (inr x , true) = push (x , true) - - PF→P∘' : Pushout F fst → Pushout (g ∘ f) fst - PF→P∘' (inl x) = inl (g x) - PF→P∘' (inr x) = PF→P∘ₗ x - PF→P∘' (push a i) = theCoh a i - - theCoh2 : (a : _) (b : _) - → Path (Pushout F fst) (inl (inr (g (f (a , b))))) (inl (f (a , b))) - theCoh2 a b with (f (a , b)) - theCoh2 a b | inl x = (push (inl tt , true)) ∙ sym (push (inl tt , false)) - ... | inr x = refl - - P∘'→PF : Pushout (g ∘ f) fst → Pushout F fst - P∘'→PF (inl x) = inl (inr x) - P∘'→PF (inr x) = inr (inr x) - P∘'→PF (push (a , false) i) = (theCoh2 a false ∙ push (inr a , false)) i - P∘'→PF (push (a , true) i) = (theCoh2 a true ∙ push (inr a , true)) i - - PFpushTₗ : (x : _) → P∘'→PF (PF→P∘' (inl x)) ≡ (inl x) - PFpushTₗ (inl x) = push (inl tt , true) ∙ sym (push (inl tt , false)) - PFpushTₗ (inr x) = refl - - PFpushTᵣ : (x : _) → P∘'→PF (PF→P∘' (inr x)) ≡ (inr x) - PFpushTᵣ (inl x) = push (inl tt , true) - PFpushTᵣ (inr x) = refl - - pp1 : (x : A) → PFpushTₗ (f (x , false)) ≡ theCoh2 x false - pp1 x with (f (x , false)) - ... | inl x₁ = refl - ... | inr x₁ = refl - - pp2 : (x : A) → PFpushTₗ (f (x , true)) ≡ theCoh2 x true - pp2 x with (f (x , true)) - ... | inl x₁ = refl - ... | inr x₁ = refl - - open import Cubical.Foundations.Path - open import Cubical.Foundations.GroupoidLaws - - - PFpushT : (x : _) → P∘'→PF (PF→P∘' x) ≡ x - PFpushT (inl x) = PFpushTₗ x - PFpushT (inr x) = PFpushTᵣ x - PFpushT (push (inl x , false) i) j = - compPath-filler (push (inl tt , true)) (sym (push (inl tt , false))) (~ i) j - PFpushT (push (inr x , false) i) j = - (pp1 x - ◁ flipSquare - (symP (compPath-filler' - (theCoh2 x false) (push (inr x , false))))) i j - PFpushT (push (inl x , true) i) j = push (inl x , true) (i ∧ j) - PFpushT (push (inr x , true) i) j = - (pp2 x - ◁ flipSquare - (symP (compPath-filler' - (theCoh2 x true) (push (inr x , true))))) i j - - cong-PF→P∘' : (b : _) (a : _) → cong PF→P∘' (theCoh2 b a) ≡ refl - cong-PF→P∘' b a with (f (b , a)) - ... | inl x = cong-∙ PF→P∘' (push (inl tt , true)) (sym (push (inl tt , false))) - ∙ sym (rUnit refl) - ... | inr x = refl - - PF→P∘'→PF : (x : _) → PF→P∘' (P∘'→PF x) ≡ x - PF→P∘'→PF (inl x) = refl - PF→P∘'→PF (inr x) = refl - PF→P∘'→PF (push (b , false) i) j = - (cong-∙ PF→P∘' (theCoh2 b false) (push (inr b , false)) - ∙ cong (_∙ push (b , false)) (cong-PF→P∘' b false) - ∙ sym (lUnit _)) j i - PF→P∘'→PF (push (b , true) i) j = - (cong-∙ PF→P∘' (theCoh2 b true) (push (inr b , true)) - ∙ cong (_∙ push (b , true)) (cong-PF→P∘' b true) - ∙ sym (lUnit _)) j i - - Is1 : Iso (Pushout F fst) (Pushout (g ∘ f) fst) - Iso.fun Is1 = PF→P∘' - Iso.inv Is1 = P∘'→PF - Iso.rightInv Is1 = PF→P∘'→PF - Iso.leftInv Is1 = PFpushT - -CW₁DataPre : (n m : ℕ) (f : Fin (suc (suc n)) × S₊ 0 → Fin (suc (suc m))) - → f (flast , false) ≡ flast - → (t : f (flast , true) .fst <ᵗ suc m) - → Σ[ k ∈ ℕ ] Σ[ f' ∈ (Fin k × S₊ 0 → Fin (suc m)) ] - Iso (Pushout f fst) - (Pushout f' fst) -CW₁DataPre n m f p q = (suc n) - , _ - , compIso (invIso (pushoutIso _ _ _ _ - (isoToEquiv (Σ-cong-iso-fst (invIso Iso-Fin-Unit⊎Fin))) - (isoToEquiv (invIso Iso-Fin-Unit⊎Fin)) - (isoToEquiv (invIso Iso-Fin-Unit⊎Fin)) - (funExt (uncurry help)) - (funExt λ x → refl))) - (Is1 {A = Fin (suc n)} {B = Fin (suc m)} - (λ x → Fin→Unit⊎Fin (f ((injectSuc (fst x)) , (snd x)))) - (f (flast , true) .fst , q)) - where - help : (y : Unit ⊎ Fin (suc n)) (x : Bool) - → Unit⊎Fin→Fin - (F (λ x₁ → Ind.elimFin (inl tt) inr (f (injectSuc (fst x₁) , snd x₁))) - (f (flast , true) .fst , q) (y , x)) - ≡ f (Unit⊎Fin→Fin y , x) - help (inl a) false = sym p - help (inl b) true = Σ≡Prop (λ _ → isProp<ᵗ) refl - help (inr a) false = Iso.leftInv Iso-Fin-Unit⊎Fin _ - help (inr a) true = Iso.leftInv Iso-Fin-Unit⊎Fin _ - -CW₁Data₁ : (m : ℕ) (f : Fin 1 × S₊ 0 → Fin (suc (suc m))) - → isConnected 2 (Pushout f fst) - → Iso (Fin 1 × S₊ 0) (Fin (suc (suc m))) -CW₁Data₁ m f c = mainIso -- mainIso - where - f' : Bool → Fin (suc (suc m)) - f' = f ∘ (fzero ,_) - - f'-surj : (x : _) → Σ[ t ∈ Bool ] (f' t ≡ x) - f'-surj x = - isSurj-α₀ (suc zero) m f c x .fst .snd - , cong f (ΣPathP (isPropFin1 _ _ , refl)) ∙ isSurj-α₀ (suc zero) m f c x .snd - f-true≠f-false : (x : _) → f' true ≡ x → ¬ f' true ≡ f' false - f-true≠f-false (zero , q) p r = lem (f'-surj fone) - where - lem : Σ[ t ∈ Bool ] (f' t ≡ fone) → ⊥ - lem (false , s) = snotz (cong fst (sym s ∙ sym r ∙ p)) - lem (true , s) = snotz (cong fst (sym s ∙ p)) - f-true≠f-false (suc x , q) p r = lem (f'-surj fzero) - where - lem : Σ[ t ∈ Bool ] (f' t ≡ fzero) → ⊥ - lem (false , s) = snotz (cong fst (sym p ∙ r ∙ s)) - lem (true , s) = snotz (cong fst (sym p ∙ s)) - - f-inj : (x y : _) → f x ≡ f y → x ≡ y - f-inj ((zero , tt) , false) ((zero , tt) , false) p = refl - f-inj ((zero , tt) , false) ((zero , tt) , true) p = ⊥.rec (f-true≠f-false _ refl (sym p)) - f-inj ((zero , tt) , true) ((zero , tt) , false) p = ⊥.rec (f-true≠f-false _ refl p) - f-inj ((zero , tt) , true) ((zero , tt) , true) p = refl - - mainIso : Iso (Fin 1 × S₊ 0) (Fin (suc (suc m))) - Iso.fun mainIso = f - Iso.inv mainIso x = isSurj-α₀ (suc zero) m f c x .fst - Iso.rightInv mainIso x = isSurj-α₀ 1 m f c x .snd - Iso.leftInv mainIso ((zero , tt) , x) = - (f-inj _ _ (isSurj-α₀ 1 m f c (f (fzero , x)) .snd)) - -CW₁Data : (n m : ℕ) (f : Fin n × S₊ 0 → Fin (suc (suc m))) - → isConnected 2 (Pushout f fst) - → Σ[ k ∈ ℕ ] Σ[ f' ∈ (Fin k × S₊ 0 → Fin (suc m)) ] - Iso (Pushout f fst) - (Pushout f' fst) -CW₁Data zero m f c = ⊥.rec (snd (notAllLoops-α₀ zero m f c .fst)) -CW₁Data (suc zero) zero f c = 0 , ((λ ()) , compIso isoₗ (PushoutEmptyFam (snd ∘ fst) snd)) - where - isoₗ : Iso (Pushout f fst) (Fin 1) - isoₗ = PushoutAlongEquiv (isoToEquiv (CW₁Data₁ zero f c)) fst -CW₁Data (suc zero) (suc m) f c = - ⊥.rec (Fin≠Fin (snotz ∘ sym ∘ cong (predℕ ∘ predℕ)) - mainIso) - where - mainIso : Iso (Fin 2) (Fin (3 +ℕ m)) - mainIso = - compIso - (compIso - (invIso rUnit×Iso) - (compIso - (Σ-cong-iso - (invIso Iso-Bool-Fin2) - (λ _ → isContr→Iso (tt , (λ _ → refl)) - (inhProp→isContr fzero isPropFin1))) - Σ-swap-Iso)) - (CW₁Data₁ (suc m) f c) -CW₁Data (suc (suc n)) m f c = - main .fst , main .snd .fst - , compIso correct (main .snd .snd) - where - t = notAllLoops-α₀ (suc (suc n)) m f c - - abstract - x₀ : Fin (suc (suc n)) - x₀ = fst t - - xpath : ¬ (f (x₀ , true) ≡ f (x₀ , false)) - xpath = snd t - - Fin0-iso : Iso (Fin (suc (suc n)) × S₊ 0) (Fin (suc (suc n)) × S₊ 0) - Fin0-iso = Σ-cong-iso-fst (swapFinIso flast x₀) - - FinIso2 : Iso (Fin (suc (suc m))) (Fin (suc (suc m))) - FinIso2 = swapFinIso (f (x₀ , false)) flast - - f' : Fin (suc (suc n)) × S₊ 0 → Fin (suc (suc m)) - f' = Iso.fun FinIso2 ∘ f ∘ Iso.fun Fin0-iso - - f'≡ : f' (flast , false) ≡ flast - f'≡ = cong (Iso.fun FinIso2 ∘ f) - (cong (_, false) (swapFinβₗ flast x₀)) - ∙ swapFinβₗ (f (x₀ , false)) flast - - f'¬≡ : ¬ (f' (flast , true) ≡ flast) - f'¬≡ p = xpath (cong f (ΣPathP (sym (swapFinβₗ flast x₀) , refl)) - ∙ sym (Iso.rightInv FinIso2 _) - ∙ cong (Iso.inv FinIso2) (p ∙ sym f'≡) - ∙ Iso.rightInv FinIso2 _ - ∙ cong f (ΣPathP (swapFinβₗ flast x₀ , refl))) - - f'< : fst (f' (flast , true)) <ᵗ suc m - f'< = ≠flast→<ᵗflast _ f'¬≡ - - main = CW₁DataPre _ _ f' f'≡ f'< - - Upath = isoToPath (swapFinIso x₀ fzero) - - correct : Iso (Pushout f fst) (Pushout f' fst) - correct = pushoutIso _ _ _ _ - (isoToEquiv Fin0-iso) (isoToEquiv FinIso2) (isoToEquiv (swapFinIso flast x₀)) - (funExt (λ x → cong (FinIso2 .Iso.fun ∘ f) (sym (Iso.rightInv Fin0-iso x)))) - refl - -CW₁Data' : (n m : ℕ) (f : Fin n × S₊ 0 → Fin m) - → isConnected 2 (Pushout f fst) - → Σ[ k ∈ ℕ ] - Iso (Pushout f fst) - (Pushout {A = Fin k × S₊ 0} {B = Fin 1} (λ _ → fzero) fst) -CW₁Data' zero zero f c = ⊥.rec (TR.rec (λ()) help (fst c)) - where - help : ¬ Pushout f fst - help = elimProp _ (λ _ → λ ()) snd snd -CW₁Data' (suc n) zero f c = ⊥.rec (f (fzero , true) .snd) -CW₁Data' n (suc zero) f c = n - , pushoutIso _ _ _ _ (idEquiv _) (idEquiv _) (idEquiv _) - (funExt (λ x → isPropFin1 _ _)) refl -CW₁Data' zero (suc (suc m)) f c = - ⊥.rec (TR.rec (λ()) (snotz ∘ sym ∘ cong fst) - (Iso.fun (PathIdTruncIso _) - (isContr→isProp (subst (isConnected 2) (isoToPath help) c) - ∣ fzero ∣ ∣ fone ∣))) - where - help : Iso (Pushout f fst) (Fin (suc (suc m))) - help = invIso (PushoutEmptyFam (λ()) λ()) -CW₁Data' (suc n) (suc (suc m)) f c - with (CW₁Data' _ (suc m) (CW₁Data (suc n) m f c .snd .fst) - (subst (isConnected 2) - (isoToPath (CW₁Data (suc n) m f c .snd .snd)) c)) -... | (k , e) = k , compIso (CW₁Data (suc n) m f c .snd .snd) e - --- module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A 0) where --- private --- c = snd sk+c --- sk = fst sk+c --- c' = isConnectedColim→isConnectedSkel (_ , sk) 2 c - --- module AC = CWskel-fields (_ , sk) - --- e₁ : Iso (Pushout (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) fst) (A 2) --- e₁ = compIso (PushoutCompEquivIso (idEquiv _) (CW₁-discrete (A , sk)) (AC.α 1) fst) --- (equivToIso (invEquiv (AC.e 1))) - --- liftStr = CW₁Data' _ _ (fst (CW₁-discrete (_ , sk)) ∘ AC.α 1) --- (isConnectedRetractFromIso 2 e₁ c') - --- collapse₁card : ℕ → ℕ --- collapse₁card zero = 1 --- collapse₁card (suc zero) = fst liftStr --- collapse₁card (suc (suc x)) = AC.card (suc (suc x)) - --- collapse₁CWskel : ℕ → Type _ --- collapse₁CWskel zero = Lift ⊥ --- collapse₁CWskel (suc zero) = Lift (Fin 1) --- collapse₁CWskel (suc (suc n)) = A (suc (suc n)) - --- collapse₁α : (n : ℕ) → Fin (collapse₁card n) × S⁻ n → collapse₁CWskel n --- collapse₁α (suc zero) (x , p) = lift fzero --- collapse₁α (suc (suc n)) = AC.α (2 +ℕ n) - --- connectedCWskel→ : yieldsConnectedCWskel collapse₁CWskel 0 --- fst (fst connectedCWskel→) = collapse₁card --- fst (snd (fst connectedCWskel→)) = collapse₁α --- fst (snd (snd (fst connectedCWskel→))) = λ() --- snd (snd (snd (fst connectedCWskel→))) zero = --- isContr→Equiv (isOfHLevelLift 0 (inhProp→isContr fzero isPropFin1)) --- ((inr fzero) --- , (λ { (inr x) j → inr (isPropFin1 fzero x j) })) --- snd (snd (snd (fst connectedCWskel→))) (suc zero) = --- compEquiv (invEquiv (isoToEquiv e₁)) --- (compEquiv (isoToEquiv (snd liftStr)) --- (isoToEquiv (pushoutIso _ _ _ _ --- (idEquiv _) LiftEquiv (idEquiv _) --- (funExt cohₗ) (funExt cohᵣ)))) --- where --- -- Agda complains if these proofs are inlined... --- cohₗ : (x : _) → collapse₁α 1 x ≡ collapse₁α 1 x --- cohₗ (x , p) = refl - --- cohᵣ : (x : Fin (fst liftStr) × Bool) → fst x ≡ fst x --- cohᵣ x = refl --- snd (snd (snd (fst connectedCWskel→))) (suc (suc n)) = AC.e (suc (suc n)) --- snd connectedCWskel→ = refl , (λ _ → λ ()) - --- collapse₁Equiv : (n : ℕ) --- → realise (_ , sk) ≃ realise (_ , connectedCWskel→ .fst) --- collapse₁Equiv n = --- compEquiv --- (isoToEquiv (Iso-SeqColim→SeqColimShift _ 3)) --- (compEquiv (pathToEquiv (cong SeqColim --- (cong₂ sequence (λ _ m → A (3 +ℕ m)) --- λ _ {z} → CW↪ (A , sk) (suc (suc (suc z)))))) --- (invEquiv (isoToEquiv (Iso-SeqColim→SeqColimShift _ 3)))) - - --- connectedCWskel'→connectedCWskel : ∀ {ℓ} --- → Σ[ t ∈ connectedCWskel' ℓ 0 ] --- (Σ[ dim ∈ ℕ ] --- ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim)))) --- → Σ[ t ∈ connectedCWskel ℓ 0 ] --- Σ[ dim ∈ ℕ ] --- ((k : ℕ) → isEquiv (CW↪ (_ , snd t .fst) (k +ℕ dim))) --- fst (connectedCWskel'→connectedCWskel ((A , sk) , conv)) = --- _ , connectedCWskel→ A ((sk .fst) , (sk .snd)) .fst , refl , (λ _ → λ()) --- fst (snd (connectedCWskel'→connectedCWskel ((A , sk) , conv))) = suc (fst conv) --- snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , zero , T))) k = --- ⊥.rec (TR.rec (λ()) --- (λ a → TR.rec (λ()) --- (λ t → CW₀-empty (_ , fst sk) (invEq (_ , T 0) (t .fst))) --- (isConnected-CW↪∞ 1 (_ , fst sk) a .fst)) (sk .snd .fst)) --- snd (snd (connectedCWskel'→connectedCWskel ((A , sk) , (suc dim) , T))) k = --- transport (λ i → isEquiv (CW↪ (collapse₁CWskel A sk , connectedCWskel→ A sk .fst) --- (h i))) --- (transport (λ i → isEquiv (CW↪ (A , sk .fst) (suc (+-suc k dim i)))) --- (T (suc k))) --- where --- h = cong suc (sym (+-suc k dim)) ∙ sym (+-suc k (suc dim)) - --- yieldsGoodCWskel : {ℓ : Level} (A₊₂ : ℕ → Pointed ℓ) → Type _ --- yieldsGoodCWskel {ℓ = ℓ} A = --- Σ[ f₊₁ ∈ (ℕ → ℕ) ] --- Σ[ fin ∈ (A 0) .fst ≃ Fin 1 ] --- Σ[ α ∈ ((n : ℕ) → SphereBouquet∙ n (Fin (f₊₁ n)) →∙ A n) ] --- ((n : ℕ) → cofib (α n .fst) , inl tt ≃∙ A (suc n)) - --- GoodCWskelSeq : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Sequence ℓ --- obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) zero = Lift ⊥ --- obj (GoodCWskelSeq {A = A} (f , fin , α , sq)) (suc n) = fst (A n) --- Sequence.map (GoodCWskelSeq {A = A} (f , fin , α , sq)) {n = suc n} x = fst (fst (sq n)) (inr x) - --- realiseGood∙ : {ℓ : Level} {A : ℕ → Pointed ℓ} → yieldsGoodCWskel A → Pointed ℓ --- realiseGood∙ {A = A} S = SeqColim (GoodCWskelSeq S) , incl {n = 1} (snd (A 0)) - --- yieldsFinGoodCWskel : {ℓ : Level} (dim : ℕ) (A₊₂ : ℕ → Pointed ℓ) → Type _ --- yieldsFinGoodCWskel {ℓ = ℓ} dim A = --- Σ[ A ∈ yieldsGoodCWskel A ] converges (GoodCWskelSeq A) dim - --- GoodCWskel : (ℓ : Level) → Type (ℓ-suc ℓ) --- GoodCWskel ℓ = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsGoodCWskel A - --- FinGoodCWskel : (ℓ : Level) (dim : ℕ) → Type (ℓ-suc ℓ) --- FinGoodCWskel ℓ dim = Σ[ A ∈ (ℕ → Pointed ℓ) ] yieldsFinGoodCWskel dim A - --- isGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) --- isGoodCWExpl {ℓ} A = Σ[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A - --- isFinGoodCWExpl : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) --- isFinGoodCWExpl {ℓ} A = --- Σ[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] realiseGood∙ (fst (snd sk)) ≃∙ A - --- isGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) --- isGoodCW {ℓ} A = ∃[ sk ∈ GoodCWskel ℓ ] realiseGood∙ (snd sk) ≃∙ A - --- isFinGoodCW : {ℓ : Level} (A : Pointed ℓ) → Type (ℓ-suc ℓ) --- isFinGoodCW {ℓ} A = --- ∃[ dim ∈ ℕ ] Σ[ sk ∈ FinGoodCWskel ℓ dim ] (realiseGood∙ (fst (snd sk)) ≃∙ A) - --- finGoodCW : (ℓ : Level) → Type (ℓ-suc ℓ) --- finGoodCW ℓ = Σ[ A ∈ Pointed ℓ ] isFinGoodCW A - --- ptCW : {ℓ : Level} {A : ℕ → Type ℓ} → yieldsCWskel A → A 1 --- → (n : ℕ) → A (suc n) --- ptCW sk a zero = a --- ptCW sk a (suc n) = CW↪ (_ , sk) (suc n) (ptCW sk a n) - --- -- module TWOO {ℓ : Level} (A' : ℕ → Type ℓ) (pt0 : A' 1) --- -- (dim : ℕ) (con : isConnected 2 (A' 2)) --- -- (C : yieldsFinCWskel dim A') --- -- where - --- -- open CWskel-fields (_ , fst C) --- -- e₀ : A' 1 ≃ Fin (card 0) --- -- e₀ = CW₁-discrete (_ , fst C) - --- -- ptA : (n : ℕ) → A' (suc n) --- -- ptA = ptCW (fst C) pt0 - --- -- conA : (n : ℕ) → isConnected 2 (A' (suc n)) --- -- conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) --- -- (subst (isConnected 2) (sym (cong Fin cA)) --- -- (∣ flast ∣ --- -- , TR.elim (λ _ → isOfHLevelPath 2 --- -- (isOfHLevelTrunc 2) _ _) --- -- λ {(zero , tt) → refl})) --- -- conA (suc n) = --- -- isConnectedRetractFromIso 2 --- -- (equivToIso (e (suc n))) --- -- (∣ inl (ptA n) ∣ₕ --- -- , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) --- -- (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) --- -- (conA' (conA n)) --- -- λ c → conA' (conA n) _ --- -- ∙ cong ∣_∣ₕ (push (c , ptSn n)))) --- -- where --- -- conA' : (conA : isConnected 2 (A' (suc n))) (c : A' (suc n)) --- -- → Path (hLevelTrunc 2 (Pushout (α (suc n)) fst)) --- -- ∣ inl (ptA n) ∣ₕ ∣ inl c ∣ₕ --- -- conA' conA c = --- -- TR.rec (isOfHLevelTrunc 2 _ _) --- -- (λ p i → ∣ inl (p i) ∣) --- -- (Iso.fun (PathIdTruncIso _) --- -- (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) - - --- -- private --- -- funType = ((n : Fin (suc dim)) --- -- → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) --- -- →∙ (A' (suc (fst n)) , ptA (fst n))) ] --- -- ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) - --- -- mapMakerNil : ∥ funType ∥₁ --- -- mapMakerNil = --- -- invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) --- -- λ n → TR.map --- -- (λ pted → ((λ { (inl x) → ptA (fst n) --- -- ; (inr x) → α _ x --- -- ; (push a i) → pted a i}) --- -- , refl) --- -- , λ _ → refl) (help n)) --- -- where --- -- help : (n : Fin (suc dim)) --- -- → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) --- -- → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) --- -- help n = invEq (_ , InductiveFinSatAC _ _ _) --- -- λ m → Iso.fun (PathIdTruncIso _) --- -- (isContr→isProp --- -- (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ --- -- ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) --- -- module _ (F : funType) where --- -- funs : (n : ℕ) → SphereBouquet∙ n (Fin (card (suc n))) --- -- →∙ (A' (suc n) , ptA n) --- -- funs n with (n ≟ᵗ dim) --- -- ... | lt x = F (n , <ᵗ-trans-suc x) .fst --- -- ... | eq x = F (n , subst (_<ᵗ suc dim) (sym x) <ᵗsucm) .fst --- -- ... | gt x = const∙ _ _ - --- -- funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) --- -- ≃∙ Pushout (funs n .fst ∘ inr) (λ r → fst r) , inl (ptA n) --- -- funEqP1 n = invEquiv (isoToEquiv (⋁-cofib-Iso _ (funs n))) , refl - --- -- funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) --- -- ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) --- -- funEq n = isoToEquiv (pushoutIso _ _ _ _ --- -- (idEquiv _) --- -- (idEquiv _) --- -- (idEquiv _) --- -- (funExt (uncurry (main n))) --- -- (λ i x → fst x)) --- -- , λ _ → inl (ptA n) --- -- where --- -- main : (n : ℕ) (x : Fin (card (suc n))) (y : S₊ n) --- -- → funs n .fst (inr (x , y)) ≡ fst (C .snd) (suc n) (x , y) --- -- main n with (n ≟ᵗ dim) --- -- ... | lt p = λ x y --- -- → F (n , <ᵗ-trans-suc p) .snd (x , y) --- -- ... | eq p = λ x y --- -- → F (n , subst (_<ᵗ suc dim) (λ i → p (~ i)) <ᵗsucm) .snd (x , y) --- -- ... | gt p = λ x --- -- → ⊥.rec (¬Fin0 (subst Fin (ind (suc n) (<ᵗ-trans p <ᵗsucm)) x)) - --- -- getGoodCWskelAux : ∥ yieldsGoodCWskel (λ n → A' (suc n) , ptA n) ∥₁ --- -- getGoodCWskelAux = PT.map help mapMakerNil --- -- where --- -- help : funType → yieldsGoodCWskel (λ n → A' (suc n) , ptA n) --- -- fst (help F) = card ∘ suc --- -- fst (snd (help F)) = compEquiv e₀ (pathToEquiv (cong Fin cA)) --- -- fst (snd (snd (help F))) n = funs F n --- -- snd (snd (snd (help F))) n = --- -- compEquiv∙ (compEquiv∙ (funEqP1 F n) (funEq F n)) --- -- (invEquiv (e (suc n)) , refl) - - --- module BS {ℓ : Level} (A' : ℕ → Type ℓ) --- (dim : ℕ) --- (C+eq : yieldsFinCWskel dim A') --- (cA : fst (fst C+eq) 0 ≡ 1) --- where --- C = fst C+eq --- ind = snd C+eq - --- open CWskel-fields (_ , C) --- e₀ : A' 1 ≃ Fin (card 0) --- e₀ = CW₁-discrete (_ , C) - - --- ¬dim≡0 : ¬ (dim ≡ 0) --- ¬dim≡0 p = CW₀-empty (_ , C) (subst A' p --- (invEq (_ , ind 0) (subst A' (cong suc (sym p)) --- (invEq e₀ (subst Fin (sym cA) fzero))))) - --- pt0 : A' 1 --- pt0 = invEq e₀ (subst Fin (sym cA) flast) - --- ptA : (n : ℕ) → A' (suc n) --- ptA = ptCW C pt0 - --- conA : (n : ℕ) → isConnected 2 (A' (suc n)) --- conA zero = isConnectedRetractFromIso 2 (equivToIso e₀) --- (subst (isConnected 2) (sym (cong Fin cA)) --- (∣ flast ∣ --- , TR.elim (λ _ → isOfHLevelPath 2 --- (isOfHLevelTrunc 2) _ _) --- λ {(zero , tt) → refl})) --- conA (suc n) = --- isConnectedRetractFromIso 2 --- (equivToIso (e (suc n))) --- (∣ inl (ptA n) ∣ₕ --- , TR.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) --- (elimProp _ (λ _ → isOfHLevelTrunc 2 _ _) --- (conA' (conA n)) --- λ c → conA' (conA n) _ --- ∙ cong ∣_∣ₕ (push (c , ptSn n)))) --- where --- conA' : (conA : isConnected 2 (A' (suc n))) (c : A' (suc n)) --- → Path (hLevelTrunc 2 (Pushout (α (suc n)) fst)) --- ∣ inl (ptA n) ∣ₕ ∣ inl c ∣ₕ --- conA' conA c = --- TR.rec (isOfHLevelTrunc 2 _ _) --- (λ p i → ∣ inl (p i) ∣) --- (Iso.fun (PathIdTruncIso _) --- (isContr→isProp conA ∣ ptA n ∣ₕ ∣ c ∣ₕ)) - --- private --- funType = ((n : Fin dim) --- → Σ[ h ∈ (SphereBouquet∙ (fst n) (Fin (card (suc (fst n)))) --- →∙ (A' (suc (fst n)) , ptA (fst n))) ] --- ((x : _) → fst h (inr x) ≡ α (suc (fst n)) x)) - --- mapMakerNil : ∥ funType ∥₁ --- mapMakerNil = --- invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) --- λ n → TR.map --- (λ pted → ((λ { (inl x) → ptA (fst n) --- ; (inr x) → α _ x --- ; (push a i) → pted a i}) --- , refl) --- , λ _ → refl) (help n)) --- where --- help : (n : Fin dim) --- → hLevelTrunc 1 ((x : Fin (card (suc (fst n)))) --- → (ptA (fst n) ≡ α (suc (fst n)) (x , ptSn (fst n)))) --- help n = invEq (_ , InductiveFinSatAC _ _ _) --- λ m → Iso.fun (PathIdTruncIso _) --- (isContr→isProp --- (conA (fst n)) ∣ (ptA (fst n)) ∣ₕ ∣ α (suc (fst n)) (m , ptSn (fst n)) ∣ₕ) - --- module _ (F : funType) where --- card' : ℕ → ℕ --- card' n with (n ≟ᵗ dim) --- ... | lt x = card (suc n) --- ... | eq x = 0 --- ... | gt x = 0 - --- funs : (n : ℕ) → SphereBouquet∙ n (Fin (card' n)) --- →∙ (A' (suc n) , ptA n) --- funs n with (n ≟ᵗ dim) --- ... | lt x = F (n , x) .fst --- ... | eq x = const∙ _ _ --- ... | gt x = const∙ _ _ - --- funEqP1 : (n : ℕ) → (cofib (funs n .fst) , inl tt) --- ≃∙ Pushout (funs n .fst ∘ inr) (λ r → fst r) , inl (ptA n) --- funEqP1 n = invEquiv (isoToEquiv (⋁-cofib-Iso _ (funs n))) , refl - --- funEq : (n : ℕ) → Pushout (funs n .fst ∘ inr) fst , inl (ptA n) --- ≃∙ Pushout (fst (C .snd) (suc n)) fst , inl (ptA n) --- funEq n with (n ≟ᵗ dim) --- ... | lt x = isoToEquiv (pushoutIso _ _ _ _ --- (idEquiv _) --- (idEquiv _) --- (idEquiv _) --- (funExt (uncurry λ x y → F (n , _) .snd (x , y))) --- (λ i x → fst x)) --- , λ _ → inl (ptA n) --- ... | eq x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) --- (compEquiv ((CW↪ (_ , C) (suc n)) --- , transport (λ i → isEquiv (CW↪ (A' , C) --- (suc (x (~ i))))) --- (ind 1) --- ) --- (e (suc n)))) , secEq (e (suc n)) _ --- ... | gt x = (compEquiv (isoToEquiv (invIso (PushoutEmptyFam (λ()) λ()))) --- (compEquiv ((CW↪ (_ , C) (suc n)) --- , (transport (λ i → isEquiv (CW↪ (A' , C) --- (suc ((sym (+-suc (fst (<ᵗ→< x)) dim) --- ∙ (<ᵗ→< x .snd)) i)))) --- (ind (suc (suc (fst (<ᵗ→< x))))))) --- (e (suc n)))) , secEq (e (suc n)) _ - --- goodCWmk : yieldsGoodCWskel (λ n → A' (suc n) , ptA n) --- fst goodCWmk = card' --- fst (snd goodCWmk) = compEquiv e₀ (pathToEquiv (cong Fin cA)) --- fst (snd (snd goodCWmk)) = funs --- snd (snd (snd goodCWmk)) n = --- compEquiv∙ (compEquiv∙ (funEqP1 n) (funEq n)) --- (invEquiv (e (suc n)) , refl) - --- goodCWmk-converges : converges --- (sequence (obj (GoodCWskelSeq goodCWmk)) --- (Sequence.map (GoodCWskelSeq goodCWmk))) --- dim --- goodCWmk-converges zero = help dim refl --- where --- help : (x : _) (p : dim ≡ x) → isEquiv (Sequence.map (GoodCWskelSeq goodCWmk) {x}) --- help zero p = ⊥.rec (¬dim≡0 p) --- help (suc x) p with (x ≟ᵗ dim) --- ... | lt x₁ = transport (λ i → isEquiv λ z → CW↪ (A' , C) (p i) z) (ind zero) --- ... | eq x₁ = ⊥.rec (¬m Date: Tue, 17 Sep 2024 01:01:18 +0200 Subject: [PATCH 52/73] fixes --- Cubical/CW/Connected.agda | 32 ++++++++++--------- Cubical/CW/Subcomplex.agda | 36 ++++++++++++++++++---- Cubical/Data/Fin/Inductive/Properties.agda | 10 +++--- Cubical/HITs/Pushout/Properties.agda | 3 +- Cubical/HITs/Wedge/Properties.agda | 14 ++++++--- Cubical/Relation/Nullary/Properties.agda | 11 +++---- 6 files changed, 69 insertions(+), 37 deletions(-) diff --git a/Cubical/CW/Connected.agda b/Cubical/CW/Connected.agda index d5cfbfbe20..07e8376331 100644 --- a/Cubical/CW/Connected.agda +++ b/Cubical/CW/Connected.agda @@ -66,20 +66,24 @@ yieldsConnectedCWskel A k = Σ[ sk ∈ yieldsCWskel A ] ((sk .fst 0 ≡ 1) × ((n : ℕ) → n <ᵗ k → sk .fst (suc n) ≡ 0)) -- Alternatively, we may say that its colimit is n-connected -yieldsConnectedCWskel' : (A : ℕ → Type ℓ) (n : ℕ) → Type _ -yieldsConnectedCWskel' A n = Σ[ sk ∈ yieldsCWskel A ] isConnected (2 +ℕ n) (realise (_ , sk)) +yieldsCombinatorialConnectedCWskel : (A : ℕ → Type ℓ) (n : ℕ) → Type _ +yieldsCombinatorialConnectedCWskel A n = + Σ[ sk ∈ yieldsCWskel A ] isConnected (2 +ℕ n) (realise (_ , sk)) connectedCWskel : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) connectedCWskel ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel X n) -connectedCWskel' : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) -connectedCWskel' ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel' X n) +combinatorialConnectedCWskel : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) +combinatorialConnectedCWskel ℓ n = + Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCombinatorialConnectedCWskel X n) isConnectedCW : ∀ {ℓ} (n : ℕ) → Type ℓ → Type (ℓ-suc ℓ) -isConnectedCW {ℓ = ℓ} n A = Σ[ sk ∈ connectedCWskel ℓ n ] realise (_ , (snd sk .fst)) ≃ A +isConnectedCW {ℓ = ℓ} n A = + Σ[ sk ∈ connectedCWskel ℓ n ] realise (_ , (snd sk .fst)) ≃ A isConnectedCW' : ∀ {ℓ} (n : ℕ) → Type ℓ → Type (ℓ-suc ℓ) -isConnectedCW' {ℓ = ℓ} n A = Σ[ sk ∈ connectedCWskel' ℓ n ] realise (_ , (snd sk .fst)) ≃ A +isConnectedCW' {ℓ = ℓ} n A = + Σ[ sk ∈ combinatorialConnectedCWskel ℓ n ] realise (_ , (snd sk .fst)) ≃ A --- Goal: show that these two definitions coincide (note that indexing is off by 2) --- -- For the base case, we need to analyse α₀ : Fin n × S⁰ → X₁ (recall, @@ -89,13 +93,13 @@ isConnectedCW' {ℓ = ℓ} n A = Σ[ sk ∈ connectedCWskel' ℓ n ] realise (_ -- us to iteratively shrink X₁ by contracting the image of α₀(a,-). -- Decision producedures -hasDecidableImage-Fin×S⁰ : {A : Type ℓ} +inhabitedFibres?-Fin×S⁰ : {A : Type ℓ} (da : Discrete A) (n : ℕ) (f : Fin n × S₊ 0 → A) - → hasDecidableImage f -hasDecidableImage-Fin×S⁰ {A = A} da n = - subst (λ C → (f : C → A) → hasDecidableImage f) + → inhabitedFibres? f +inhabitedFibres?-Fin×S⁰ {A = A} da n = + subst (λ C → (f : C → A) → inhabitedFibres? f) (isoToPath (invIso Iso-Fin×Bool-Fin)) - (hasDecidableImageFin da _) + (inhabitedFibres?Fin da _) private allConst? : {A : Type ℓ} {n : ℕ} (dis : Discrete A) @@ -112,11 +116,11 @@ private ... | yes p | inr x = inr (_ , (snd x)) ... | no ¬p | q = inr (_ , ¬p) --- α₀ must be is surjective +-- α₀ must have a section isSurj-α₀ : (n m : ℕ) (f : Fin n × S₊ 0 → Fin (suc (suc m))) → isConnected 2 (Pushout f fst) → (y : _) → Σ[ x ∈ _ ] f x ≡ y -isSurj-α₀ n m f c y with (hasDecidableImage-Fin×S⁰ DiscreteFin n f y) +isSurj-α₀ n m f c y with (inhabitedFibres?-Fin×S⁰ DiscreteFin n f y) ... | inl x = x isSurj-α₀ n m f c x₀ | inr q = ⊥.rec nope where @@ -521,7 +525,7 @@ module CWLemmas-0Connected where -- Uning this, we can show that a 0-connected CW complex can be -- approximated by one with trivial 1-skeleton. -module _ (A : ℕ → Type ℓ) (sk+c : yieldsConnectedCWskel' A 0) where +module _ (A : ℕ → Type ℓ) (sk+c : yieldsCombinatorialConnectedCWskel A 0) where private open CWLemmas-0Connected c = snd sk+c diff --git a/Cubical/CW/Subcomplex.agda b/Cubical/CW/Subcomplex.agda index 5f865326a6..37fbed2618 100644 --- a/Cubical/CW/Subcomplex.agda +++ b/Cubical/CW/Subcomplex.agda @@ -1,5 +1,19 @@ {-# OPTIONS --safe --lossy-unification #-} +{- +This file contains a definition of a notion of (strict) +subcomplexes of a CW complex. Here, a subomplex of a complex +C = colim ( C₀ → C₁ → ...) +is simply the complex +C⁽ⁿ⁾ := colim (C₀ → ... → Cₙ = Cₙ = ...) + +This file contains. +1. Definition of above +2. A `strictification' of finite CW complexes in terms of above +3. An elmination principle for finite CW complexes +4. A proof that C and C⁽ⁿ⁾ has the same homolog in appropriate dimensions. +-} + module Cubical.CW.Subcomplex where open import Cubical.Foundations.Prelude @@ -8,6 +22,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels open import Cubical.Data.Nat open import Cubical.Data.Fin.Inductive.Base @@ -36,20 +51,24 @@ private variable ℓ ℓ' ℓ'' : Level --- finite subcomplex of a cw complex +-- finite subcomplex C⁽ⁿ⁾ of a cw complex C. module _ (C : CWskel ℓ) where + -- The underlying family of types is that of C but ending at some + -- fixed n. subComplexFam : (n : ℕ) (m : ℕ) → Type ℓ subComplexFam n m with (m ≟ᵗ n) ... | lt x = fst C m ... | eq x = fst C m ... | gt x = fst C n + -- Similarly for the number of cells subComplexCard : (n : ℕ) → ℕ → ℕ subComplexCard n m with (m ≟ᵗ n) ... | lt x = snd C .fst m ... | eq x = 0 ... | gt x = 0 + -- Attaching maps subComplexα : (n m : ℕ) → Fin (subComplexCard n m) × S⁻ m → subComplexFam n m subComplexα n m with (m ≟ᵗ n) ... | lt x = snd C .snd .fst m @@ -61,6 +80,7 @@ module _ (C : CWskel ℓ) where ... | lt x = CW₀-empty C ... | eq x = CW₀-empty C + -- Resulting complex has CW structure subComplexFam≃Pushout : (n m : ℕ) → subComplexFam n (suc m) ≃ Pushout (subComplexα n m) fst subComplexFam≃Pushout n m with (m ≟ᵗ n) | ((suc m) ≟ᵗ n) @@ -78,17 +98,19 @@ module _ (C : CWskel ℓ) where ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) x₁ (<ᵗ-trans x <ᵗsucm))) ... | gt x | gt x₁ = isoToEquiv (PushoutEmptyFam (λ x → ¬Fin0 (fst x)) ¬Fin0) + -- packaging it all together subComplexYieldsCWskel : (n : ℕ) → yieldsCWskel (subComplexFam n) fst (subComplexYieldsCWskel n) = subComplexCard n fst (snd (subComplexYieldsCWskel n)) = subComplexα n fst (snd (snd (subComplexYieldsCWskel n))) = subComplex₀-empty n snd (snd (snd (subComplexYieldsCWskel n))) m = subComplexFam≃Pushout n m + -- main definition subComplex : (n : ℕ) → CWskel ℓ fst (subComplex n) = subComplexFam n snd (subComplex n) = subComplexYieldsCWskel n - -- as a finite CWskel + -- Let us also show that this defines a _finite_ CW complex subComplexFin : (m : ℕ) (n : Fin (suc m)) → isEquiv (CW↪ (subComplexFam (fst n) , subComplexYieldsCWskel (fst n)) m) subComplexFin m (n , r) with (m ≟ᵗ n) | (suc m ≟ᵗ n) @@ -128,17 +150,18 @@ module _ (C : CWskel ℓ) where ... | eq x = idEquiv _ ... | gt x = ⊥.rec (¬squeeze (x , p)) -realiseSubComplex : (n : ℕ) (C : CWskel ℓ) → Iso (fst C n) (realise (subComplex C n)) +-- C⁽ⁿ⁾ₙ ≃ Cₙ +realiseSubComplex : (n : ℕ) (C : CWskel ℓ) + → Iso (fst C n) (realise (subComplex C n)) realiseSubComplex n C = compIso (equivToIso (complex≃subcomplex C n flast)) (realiseFin n (finSubComplex C n)) +-- Strictifying finCWskel niceFinCWskel : ∀ {ℓ} (n : ℕ) → finCWskel ℓ n → finCWskel ℓ n fst (niceFinCWskel n (A , AC , fin)) = finSubComplex (A , AC) n .fst snd (niceFinCWskel n (A , AC , fin)) = finSubComplex (A , AC) n .snd -open import Cubical.Foundations.HLevels - makeNiceFinCWskel : ∀ {ℓ} {A : Type ℓ} → isFinCW A → isFinCW A makeNiceFinCWskel {A = A} (dim , cwsk , e) = better where @@ -161,6 +184,7 @@ snd (makeNiceFinCW C) = makeNiceFinCW≡ : ∀ {ℓ} (C : finCW ℓ) → makeNiceFinCW C ≡ C makeNiceFinCW≡ C = Σ≡Prop (λ _ → squash₁) refl +-- Elimination principles makeNiceFinCWElim : ∀ {ℓ ℓ'} {A : finCW ℓ → Type ℓ'} → ((C : finCW ℓ) → A (makeNiceFinCW C)) → (C : _) → A C @@ -173,7 +197,7 @@ makeNiceFinCWElim' : ∀ {ℓ ℓ'} {C : Type ℓ} {A : ∥ isFinCW C ∥₁ → makeNiceFinCWElim' {A = A} pr ind = PT.elim pr λ cw → subst A (squash₁ _ _) (ind cw) --- Goal: Show that a cell complex C and its subcomplex Cₙ share +-- Goal: Show that a cell complex C and its subcomplex C⁽ⁿ⁾ share -- homology in low enough dimensions module _ (C : CWskel ℓ) where ChainSubComplex : (n : ℕ) → snd C .fst n ≡ snd (subComplex C (suc n)) .fst n diff --git a/Cubical/Data/Fin/Inductive/Properties.agda b/Cubical/Data/Fin/Inductive/Properties.agda index d21f070a8f..9b553fd878 100644 --- a/Cubical/Data/Fin/Inductive/Properties.agda +++ b/Cubical/Data/Fin/Inductive/Properties.agda @@ -145,13 +145,13 @@ DiscreteFin x y with discreteℕ (fst x) (fst y) ... | yes p = yes (Σ≡Prop (λ _ → isProp<ᵗ) p) ... | no ¬p = no λ q → ¬p (cong fst q) -hasDecidableImageFin : ∀ {ℓ} {A : Type ℓ} +inhabitedFibres?Fin : ∀ {ℓ} {A : Type ℓ} (da : Discrete A) (n : ℕ) (f : Fin n → A) - → hasDecidableImage f -hasDecidableImageFin {A = A} da zero f y = inr λ x → ⊥.rec (¬Fin0 x) -hasDecidableImageFin {A = A} da (suc n) f y with da (f fzero) y + → inhabitedFibres? f +inhabitedFibres?Fin {A = A} da zero f y = inr λ x → ⊥.rec (¬Fin0 x) +inhabitedFibres?Fin {A = A} da (suc n) f y with da (f fzero) y ... | yes p = inl (fzero , p) -... | no ¬p with (hasDecidableImageFin da n (f ∘ fsuc) y) +... | no ¬p with (inhabitedFibres?Fin da n (f ∘ fsuc) y) ... | inl q = inl ((fsuc (fst q)) , snd q) ... | inr q = inr (elimFin-alt ¬p q) diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index a47e4c49ae..e8fa8c5cd4 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -705,7 +705,8 @@ PushoutCompEquivIso {ℓA = ℓA} {ℓA'} {ℓB} {ℓB'} {ℓC} e1 e2 f g = (Pushout f g)) λ f g → idIso) --- Computation of cofibre of the quotient map B → B/A +-- Computation of cofibre of the quotient map B → B/A (where B/A +-- denotes the cofibre of some f : B → A) module _ {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where private open 3x3-span diff --git a/Cubical/HITs/Wedge/Properties.agda b/Cubical/HITs/Wedge/Properties.agda index 26e15e7138..c357e94d00 100644 --- a/Cubical/HITs/Wedge/Properties.agda +++ b/Cubical/HITs/Wedge/Properties.agda @@ -86,7 +86,7 @@ cofibInr-⋁ {A = A} {B = B} = (cofibInl-⋁ {A = B} {B = A}) -- A ⋁ Unit ≃ A ≃ Unit ⋁ A -⋁-rUnitIso : {A : Pointed ℓ} → Iso (A ⋁ (Unit* {ℓ'} , tt*)) (fst A) +⋁-rUnitIso : {A : Pointed ℓ} → Iso (A ⋁ Unit*∙ {ℓ'}) (fst A) Iso.fun ⋁-rUnitIso (inl x) = x Iso.fun (⋁-rUnitIso {A = A}) (inr x) = pt A Iso.fun (⋁-rUnitIso {A = A}) (push a i) = pt A @@ -96,7 +96,7 @@ Iso.leftInv ⋁-rUnitIso (inl x) = refl Iso.leftInv ⋁-rUnitIso (inr x) = push tt Iso.leftInv ⋁-rUnitIso (push tt i) j = push tt (i ∧ j) -⋁-lUnitIso : {A : Pointed ℓ} → Iso ((Unit* {ℓ'} , tt*) ⋁ A) (fst A) +⋁-lUnitIso : {A : Pointed ℓ} → Iso (Unit*∙ {ℓ'} ⋁ A) (fst A) ⋁-lUnitIso = compIso ⋁-commIso ⋁-rUnitIso -- cofiber of constant function @@ -294,6 +294,8 @@ SuspBouquet≃Bouquet : (A : Type ℓ) (B : A → Pointed ℓ') → Susp (⋁gen A B) ≃ ⋁gen A (λ a → Susp∙ (fst (B a))) SuspBouquet≃Bouquet A B = isoToEquiv (Iso-SuspBouquet-Bouquet A B) +-- cofibre of f ∨→ g : A ⋁ B → C is cofibre of +-- inr ∘ f : A → cofib g module _ {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} (f : A →∙ C) (g : B →∙ C) where @@ -402,6 +404,8 @@ module _ {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} (cofib {B = cofib (fst g)} (λ x → inr (fst f x))) cofib∨→compIsoᵣ = compIso cofib∨→distrIso (equivToIso (_ , isPushoutRight)) +-- cofibre of f ∨→ g : A ⋁ B → C is cofibre of +-- inr ∘ g : B → cofib f cofib∨→compIsoₗ : {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} (f : A →∙ C) (g : B →∙ C) @@ -416,7 +420,7 @@ cofib∨→compIsoₗ f g = ; (push a i) j → symDistr (snd g) (sym (snd f)) (~ j) i})) (cofib∨→compIsoᵣ g f) - +-- (⋁ᵢ Aᵢ ∨ ⋁ⱼ Bᵢ) ≃ ⋁ᵢ₊ⱼ(Aᵢ + Bⱼ) module _ {A : Type ℓ} {B : Type ℓ'} {P : A → Pointed ℓ''} {Q : B → Pointed ℓ''} where @@ -440,7 +444,7 @@ module _ {A : Type ℓ} {B : Type ℓ'} ⋁gen⊎← (push (inr x) i) = (push tt ∙ λ i → inr (push x i)) i ⋁gen⊎Iso : Iso (⋁gen∙ A P ⋁ ⋁gen∙ B Q) - (⋁gen (A ⊎ B) (⊎.rec P Q)) + (⋁gen (A ⊎ B) (⊎.rec P Q)) Iso.fun ⋁gen⊎Iso = ⋁gen⊎→ Iso.inv ⋁gen⊎Iso = ⋁gen⊎← Iso.rightInv ⋁gen⊎Iso (inl tt) = refl @@ -458,7 +462,7 @@ module _ {A : Type ℓ} {B : Type ℓ'} compPath-filler' (push tt) (λ i → inr (push a i)) (~ j) i Iso.leftInv ⋁gen⊎Iso (push a i) j = push a (i ∧ j) --- computation of the cofibre of a map out of a wedge +-- f : ⋁ₐ Bₐ → C has cofibre the pushout of cofib (f ∘ inr) ← Σₐ → A module _ {ℓA ℓB ℓC : Level} {A : Type ℓA} {B : A → Pointed ℓB} (C : Pointed ℓC) (f : (⋁gen A B , inl tt) →∙ C) where private diff --git a/Cubical/Relation/Nullary/Properties.agda b/Cubical/Relation/Nullary/Properties.agda index 5660958a55..71150a5a8e 100644 --- a/Cubical/Relation/Nullary/Properties.agda +++ b/Cubical/Relation/Nullary/Properties.agda @@ -203,12 +203,11 @@ Discrete→Separated d x y = Dec→Stable (d x y) Discrete→isSet : Discrete A → isSet A Discrete→isSet = Separated→isSet ∘ Discrete→Separated --- Decidable images -hasDecidableImage : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} - (f : A → B) → Type (ℓ-max ℓ ℓ') -hasDecidableImage {A = A} {B = B} f = - (y : B) → (Σ[ x ∈ A ] f x ≡ y) ⊎ ((x : A) → ¬ f x ≡ y) - ≡no : ∀ {A : Type ℓ} x y → Path (Dec A) x (no y) ≡no (yes p) y = ⊥.rec (y p) ≡no (no ¬p) y i = no (isProp¬ _ ¬p y i) + +inhabitedFibres? : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} + (f : A → B) → Type (ℓ-max ℓ ℓ') +inhabitedFibres? {A = A} {B = B} f = + (y : B) → (Σ[ x ∈ A ] f x ≡ y) ⊎ ((x : A) → ¬ f x ≡ y) From b4e74c14767031ec00ebbb43696186a53b3b030c Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Sun, 11 Aug 2024 21:08:23 +0800 Subject: [PATCH 53/73] Susp is Pushout 1 <- X -> 1 --- Cubical/HITs/Susp/Properties.agda | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index 98f7c2a2ab..faa0adaf51 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -13,9 +13,11 @@ open import Cubical.Foundations.Function open import Cubical.Data.Bool open import Cubical.Data.Sigma +open import Cubical.Data.Unit open import Cubical.HITs.Join open import Cubical.HITs.Susp.Base +open import Cubical.HITs.Pushout open import Cubical.Homotopy.Loopspace private @@ -73,6 +75,26 @@ Susp≃joinBool = isoToEquiv Susp-iso-joinBool Susp≡joinBool : ∀ {ℓ} {A : Type ℓ} → Susp A ≡ join A Bool Susp≡joinBool = isoToPath Susp-iso-joinBool +Susp-iso-Pushout : ∀ {ℓ} {A : Type ℓ} → Iso (Susp A) (Pushout (terminal A) (terminal A)) +fun Susp-iso-Pushout north = inl _ +fun Susp-iso-Pushout south = inr _ +fun Susp-iso-Pushout (merid a i) = push a i +inv Susp-iso-Pushout (inl _) = north +inv Susp-iso-Pushout (inr _) = south +inv Susp-iso-Pushout (push a i) = merid a i +rightInv Susp-iso-Pushout (inl _) = refl +rightInv Susp-iso-Pushout (inr _) = refl +rightInv Susp-iso-Pushout (push a i) = refl +leftInv Susp-iso-Pushout north = refl +leftInv Susp-iso-Pushout south = refl +leftInv Susp-iso-Pushout (merid a i) = refl + +Susp≃Pushout : ∀ {ℓ} {A : Type ℓ} → Susp A ≃ Pushout (terminal A) (terminal A) +Susp≃Pushout = isoToEquiv Susp-iso-Pushout + +Susp≡Pushout : ∀ {ℓ} {A : Type ℓ} → Susp A ≡ Pushout (terminal A) (terminal A) +Susp≡Pushout = isoToPath Susp-iso-Pushout + congSuspIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso (Susp A) (Susp B) fun (congSuspIso is) = suspFun (fun is) inv (congSuspIso is) = suspFun (inv is) From d0a06cdb284a25dd095c93c78ca78121ed9e05d9 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Sun, 11 Aug 2024 21:40:16 +0800 Subject: [PATCH 54/73] pushout along identity --- Cubical/HITs/Pushout/Properties.agda | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index e8fa8c5cd4..45b0ae1663 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -75,6 +75,28 @@ pushoutSwitchEquiv = isoToEquiv (iso f inv leftInv rightInv) leftInv = λ {(inl x) → refl; (inr x) → refl; (push a i) → refl} rightInv = λ {(inl x) → refl; (inr x) → refl; (push a i) → refl} + +{- + Pushout along the identity gives an equivalence. +-} +pushoutIdfunEquiv : ∀ {ℓ ℓ'} {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) + → Y ≃ Pushout f (idfun X) +pushoutIdfunEquiv f = isoToEquiv (iso inl inv leftInv λ _ → refl) + where + inv : Pushout f (idfun _) → _ + inv (inl y) = y + inv (inr x) = f x + inv (push x i) = f x + + leftInv : section inl inv + leftInv (inl y) = refl + leftInv (inr x) = push x + leftInv (push a i) j = push a (i ∧ j) + +pushoutIdfunEquiv' : ∀ {ℓ ℓ'} {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) + → Y ≃ Pushout (idfun X) f +pushoutIdfunEquiv' f = compEquiv (pushoutIdfunEquiv _) pushoutSwitchEquiv + {- Definition of pushout diagrams -} From 313b74e0dead200f337708449f2cb1dca68d1cd3 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Mon, 12 Aug 2024 00:22:40 +0800 Subject: [PATCH 55/73] =?UTF-8?q?Use=20pushout=20squares=20for=20cofibInl-?= =?UTF-8?q?=E2=8B=81?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Cubical/HITs/Wedge/Properties.agda | 71 ++++++++++++++++++------------ 1 file changed, 44 insertions(+), 27 deletions(-) diff --git a/Cubical/HITs/Wedge/Properties.agda b/Cubical/HITs/Wedge/Properties.agda index c357e94d00..2161e7c668 100644 --- a/Cubical/HITs/Wedge/Properties.agda +++ b/Cubical/HITs/Wedge/Properties.agda @@ -12,7 +12,6 @@ open import Cubical.Data.Sigma open import Cubical.Data.Unit open import Cubical.Data.Sum as ⊎ -open import Cubical.HITs.Pushout.Base open import Cubical.HITs.Wedge.Base open import Cubical.HITs.Susp open import Cubical.HITs.Pushout @@ -45,34 +44,52 @@ Iso.rightInv ⋁-commIso = ⋁-commFun² Iso.leftInv ⋁-commIso = ⋁-commFun² -- cofibre of A --inl→ A ⋁ B is B +cofibInl-⋁-square : (A : Pointed ℓ) (B : Pointed ℓ') → commSquare +cofibInl-⋁-square A B = record { + sp = record { + A0 = Unit ; + A2 = typ A ; + A4 = A ⋁ B ; + f1 = _ ; + f3 = inl } ; + P = typ B ; + inlP = λ _ → pt B ; + inrP = ⋁proj₂ A B ; + comm = refl } + +cofibInl-⋁-Pushout : (A : Pointed ℓ) (B : Pointed ℓ') + → isPushoutSquare (cofibInl-⋁-square A B) +cofibInl-⋁-Pushout A B = isoToIsEquiv (iso _ inv (λ _ → refl) leftInv) + where + inv : typ B → Pushout _ inl + inv b = inr (inr b) + + leftInv : _ + leftInv (inl x) = + (λ i → inr (push tt (~ i))) ∙ sym (push (pt A)) + leftInv (inr (inl x)) = + ((λ i → inr (push tt (~ i))) ∙ sym (push (pt A))) ∙ push x + leftInv (inr (inr x)) = refl + leftInv (inr (push a i)) j = + help (Pushout.inr ∘ inr) (λ i → inr (push tt (~ i))) (sym (push (pt A))) (~ i) j + where + help : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x : B} {y z : A} + (f : B → A) (p : f x ≡ y) (q : y ≡ z) + → Square refl ((p ∙ q) ∙ sym q) (cong f (refl ∙ refl)) p + help f p q = subst (λ t → Square refl ((p ∙ q) ∙ sym q) t p) + (sym (congFunct f refl refl ∙ sym (lUnit _))) + ((λ i j → p (i ∧ j)) + ▷ (rUnit p + ∙∙ cong (p ∙_) (sym (rCancel q)) + ∙∙ assoc p q (sym q))) + leftInv (push a i) j = + compPath-filler + ((λ i → inr (push tt (~ i))) ∙ sym (push (pt A))) + (push a) i j + cofibInl-⋁ : {A : Pointed ℓ} {B : Pointed ℓ'} → Iso (cofib {B = A ⋁ B} inl) (fst B) -Iso.fun (cofibInl-⋁ {B = B}) (inl x) = pt B -Iso.fun (cofibInl-⋁ {B = B}) (inr (inl x)) = pt B -Iso.fun cofibInl-⋁ (inr (inr x)) = x -Iso.fun (cofibInl-⋁ {B = B}) (inr (push a i)) = pt B -Iso.fun (cofibInl-⋁ {B = B}) (push a i) = pt B -Iso.inv cofibInl-⋁ x = inr (inr x) -Iso.rightInv cofibInl-⋁ x = refl -Iso.leftInv (cofibInl-⋁ {A = A}) (inl x) = - (λ i → inr (push tt (~ i))) ∙ sym (push (pt A)) -Iso.leftInv (cofibInl-⋁ {A = A}) (inr (inl x)) = - ((λ i → inr (push tt (~ i))) ∙ sym (push (pt A))) ∙ push x -Iso.leftInv cofibInl-⋁ (inr (inr x)) = refl -Iso.leftInv (cofibInl-⋁ {A = A}) (inr (push a i)) j = - help (λ i → inr (push tt (~ i))) (sym (push (pt A))) (~ i) j - where - help : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) - → Square refl ((p ∙ q) ∙ sym q) refl p - help p q = - (λ i j → p (i ∧ j)) - ▷ (rUnit p - ∙∙ cong (p ∙_) (sym (rCancel q)) - ∙∙ assoc p q (sym q)) -Iso.leftInv (cofibInl-⋁ {A = A}) (push a i) j = - compPath-filler - ((λ i → inr (push tt (~ i))) ∙ sym (push (pt A))) - (push a) i j +cofibInl-⋁ = equivToIso (_ , cofibInl-⋁-Pushout _ _) -- cofibre of B --inl→ A ⋁ B is A cofibInr-⋁ : {A : Pointed ℓ} {B : Pointed ℓ'} From 54038ae9fb28f9a30810c7518246c9819e5bf18f Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Mon, 12 Aug 2024 02:31:37 +0800 Subject: [PATCH 56/73] join inclusions are null --- Cubical/HITs/Join/Properties.agda | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Cubical/HITs/Join/Properties.agda b/Cubical/HITs/Join/Properties.agda index ae82d7800d..c2a9c1bcb3 100644 --- a/Cubical/HITs/Join/Properties.agda +++ b/Cubical/HITs/Join/Properties.agda @@ -36,6 +36,15 @@ private variable ℓ ℓ' : Level +-- the inclusion maps are null-homotopic +join-inl-null : {X : Pointed ℓ} {Y : Pointed ℓ'} (x : typ X) + → Path (join (typ X) (typ Y)) (inl x) (inl (pt X)) +join-inl-null {X = X} {Y} x = push x (pt Y) ∙ sym (push (pt X) (pt Y)) + +join-inr-null : {X : Pointed ℓ} {Y : Type ℓ'} (y : Y) + → Path (join (typ X) Y) (inr y) (inl (pt X)) +join-inr-null {X = X} y = sym (push (pt X) y) + open Iso -- Characterisation of function type join A B → C From fa2da1a05ecd55bd99203ece9c4cbc1781794cb2 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Mon, 12 Aug 2024 02:50:36 +0800 Subject: [PATCH 57/73] projection functions --- Cubical/HITs/Wedge/Base.agda | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda index 192c7aedcc..cec20c9b41 100644 --- a/Cubical/HITs/Wedge/Base.agda +++ b/Cubical/HITs/Wedge/Base.agda @@ -34,6 +34,14 @@ _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointe (f ∨→ g) (inr x) = fst g x (f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ +⋁proj₁ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → A ⋁ B → typ A +⋁proj₁ A B = idfun∙ A ∨→ ((λ _ → pt A) , refl) + +⋁proj₂ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → A ⋁ B → typ B +⋁proj₂ A B = ((λ _ → pt B) , refl) ∨→ idfun∙ B + -- Pointed version ∨→∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → (f : A →∙ C) (g : B →∙ C) → ((A ⋁∙ₗ B) →∙ C) From 90f1ca9f37e5b2587f34f55e879cfbcc3d1fe78c Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Mon, 12 Aug 2024 02:50:52 +0800 Subject: [PATCH 58/73] Remove redundant definition --- Cubical/HITs/SmashProduct/Base.agda | 7 +------ 1 file changed, 1 insertion(+), 6 deletions(-) diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index 2c591a2797..0a5cfe16e4 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -67,13 +67,8 @@ commK (gluer b x) = refl --- Alternative definition -i∧ : {A : Pointed ℓ} {B : Pointed ℓ'} → A ⋁ B → (typ A) × (typ B) -i∧ {A = A , ptA} {B = B , ptB} (inl x) = x , ptB -i∧ {A = A , ptA} {B = B , ptB} (inr x) = ptA , x -i∧ {A = A , ptA} {B = B , ptB} (push tt i) = ptA , ptB - _⋀_ : Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') -A ⋀ B = Pushout {A = (A ⋁ B)} (λ _ → tt) i∧ +A ⋀ B = Pushout {A = (A ⋁ B)} (λ _ → tt) ⋁↪ _⋀∙_ : Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') A ⋀∙ B = (A ⋀ B) , (inl tt) From 4bc6517beba1c4ad1cd5a79b6364acf1dd748e50 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Mon, 12 Aug 2024 02:51:38 +0800 Subject: [PATCH 59/73] =?UTF-8?q?Pushout=E2=8B=81=E2=89=83Unit?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Cubical/HITs/Pushout/Properties.agda | 16 ++++++- Cubical/HITs/Wedge/Properties.agda | 72 ++++++++++++++++++++++++++++ 2 files changed, 87 insertions(+), 1 deletion(-) diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index 45b0ae1663..9cc1c6cbf4 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -77,7 +77,7 @@ pushoutSwitchEquiv = isoToEquiv (iso f inv leftInv rightInv) {- - Pushout along the identity gives an equivalence. + Direct proof that pushout along the identity gives an equivalence. -} pushoutIdfunEquiv : ∀ {ℓ ℓ'} {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) → Y ≃ Pushout f (idfun X) @@ -831,6 +831,20 @@ module _ {ℓ₀ ℓ₂ ℓ₄ ℓP : Level} where PushoutSquare : Type (ℓ-suc ℓ*) PushoutSquare = Σ commSquare isPushoutSquare +-- Pushout itself fits into a pushout square +pushoutToSquare : 3-span {ℓ} {ℓ'} {ℓ''} → PushoutSquare +pushoutToSquare sp = record { + sp = sp ; + P = spanPushout sp ; + inlP = inl ; + inrP = inr ; + comm = funExt push + } , subst isEquiv (λ { + _ (inl x) → inl x ; + _ (inr x) → inr x ; + _ (push a j) → push a j + }) (idIsEquiv _) + -- Rotations of commutative squares and pushout squares module _ {ℓ₀ ℓ₂ ℓ₄ ℓP : Level} where open commSquare diff --git a/Cubical/HITs/Wedge/Properties.agda b/Cubical/HITs/Wedge/Properties.agda index 2161e7c668..b7c66fbe25 100644 --- a/Cubical/HITs/Wedge/Properties.agda +++ b/Cubical/HITs/Wedge/Properties.agda @@ -565,3 +565,75 @@ module _ {ℓA ℓB ℓC : Level} {A : Type ℓA} {B : A → Pointed ℓB} (C : ⋁-cofib-Iso = compIso (compIso (invIso A○□Iso) (invIso (3x3-Iso inst))) A□○Iso + +{- +We prove the square + X ⋁ Y --> X + ↓ ↓ + Y ----> * +is a pushout. (Note the arrow direction!!) +-} + +module _ (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Pointed ℓ') where + + private + fX : X∙ ⋁ Y∙ → X + fX = ⋁proj₁ X∙ Y∙ + + fY : X∙ ⋁ Y∙ → Y + fY = ⋁proj₂ X∙ Y∙ + + weirdSquare : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {Y : Type ℓ'} {Z : Type ℓ''} + → (f : X → Y) (g : Y → Z) → commSquare + weirdSquare f g = record + { sp = record { f1 = idfun _ ; f3 = f } + ; inlP = f + ; inrP = idfun _ + ; comm = refl ∙ refl ∙ refl } + + weirdPushoutSquare : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {Y : Type ℓ'} {Z : Type ℓ''} + → (f : X → Y) (g : Y → Z) → isPushoutSquare (weirdSquare f g) + weirdPushoutSquare f g = isoToIsEquiv (iso _ inr (λ _ → refl) + λ { (inl x) → sym (push _) + ; (inr x) → refl + ; (push a i) j → subst + (λ t → Square (sym (push _)) refl t (push _)) + (cong (cong Pushout.inr) (lUnit _ ∙ lUnit _)) + (λ i j → push a (i ∨ ~ j)) + i j + }) + + {- + The proof proceeds by applying the pasting lemma twice: + 1 ----> Y + ↓ ↓ + X --> X ⋁ Y --> X + ↓ mid ↓ bot ↓ + 1 ----> Y ----> 1 + -} + + midPushout : isPushoutSquare _ + midPushout = isPushoutTotSquare→isPushoutBottomSquare + (weirdPushoutSquare _ (terminal Y)) + where + open PushoutPasteDown (pushoutToSquare record + { A2 = Unit + ; f1 = λ _ → x₀ + ; f3 = λ _ → y₀ + }) (terminal X) (λ _ → y₀) fY refl + + -- slight help to the unifier here + botPushout : isPushoutSquare record { comm = refl } + botPushout = isPushoutTotSquare→isPushoutBottomSquare $ + rotatePushoutSquare (record { comm = refl } , isoToIsEquiv + (iso _ inl (λ _ → refl) λ { + (inl _) i → inl _ + ; (inr a) i → push a i + ; (push a j) i → push a (i ∧ j) + })) .snd + where + open PushoutPasteDown (rotatePushoutSquare (_ , midPushout)) + fX (terminal X) (terminal Y) refl + + Pushout⋁≃Unit : Pushout fX fY ≃ Unit + Pushout⋁≃Unit = _ , botPushout From 9d50340ea6072c87a4207d38077740545c08585b Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Mon, 12 Aug 2024 03:18:06 +0800 Subject: [PATCH 60/73] =?UTF-8?q?Susp=20(X=E2=88=99=20=E2=8B=80=20Y?= =?UTF-8?q?=E2=88=99)=20=E2=89=A1=20join=20X=20Y?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Cubical/HITs/Join/SuspWedgeEquiv.agda | 106 ++++++++++++++++++++++++++ 1 file changed, 106 insertions(+) create mode 100644 Cubical/HITs/Join/SuspWedgeEquiv.agda diff --git a/Cubical/HITs/Join/SuspWedgeEquiv.agda b/Cubical/HITs/Join/SuspWedgeEquiv.agda new file mode 100644 index 0000000000..88a41efc9d --- /dev/null +++ b/Cubical/HITs/Join/SuspWedgeEquiv.agda @@ -0,0 +1,106 @@ +{- + +For pointed types, join X Y is equivalent to the suspension of their wedge. + +This file is split off to prevent cyclic imports. + +-} + + +{-# OPTIONS --safe --lossy-unification #-} + +module Cubical.HITs.Join.SuspWedgeEquiv where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Join +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SmashProduct + +module _ (X∙ @ (X , x₀) : Pointed₀) (Y∙ @ (Y , y₀) : Pointed₀) where + + {- + We consider the 3×3 lemma applied to + 1 <--- 1 ---> 1 + ↑ ↑ ↑ + X <- X ⋁ Y -> Y + ↓ ↓ ↓ + X <- X × Y -> Y + This admittedly makes it annoying to accomodate for higher universes. + -} + + open 3x3-span + smash-span : 3x3-span + smash-span .A00 = Unit + smash-span .A02 = Unit + smash-span .A04 = Unit + smash-span .A20 = X + smash-span .A22 = X∙ ⋁ Y∙ + smash-span .A24 = Y + smash-span .A40 = X + smash-span .A42 = X × Y + smash-span .A44 = Y + smash-span .f10 = _ -- Unique map + smash-span .f12 = _ + smash-span .f14 = _ + smash-span .f30 = idfun X + smash-span .f32 = ⋁↪ + smash-span .f34 = idfun Y + smash-span .f01 = _ + smash-span .f21 = ⋁proj₁ X∙ Y∙ + smash-span .f41 = fst + smash-span .f03 = _ + smash-span .f23 = ⋁proj₂ X∙ Y∙ + smash-span .f43 = snd + smash-span .H11 _ = refl + smash-span .H13 _ = refl + smash-span .H31 (inl x) = refl + smash-span .H31 (inr y) = refl + smash-span .H31 (push _ i) j = doubleCompPath-filler (refl {x = x₀}) refl refl j i + smash-span .H33 (inl x) = refl + smash-span .H33 (inr x) = refl + smash-span .H33 (push _ i) j = doubleCompPath-filler (refl {x = y₀}) refl refl j i + + A□2≡⋀ : A□2 smash-span ≡ X∙ ⋀ Y∙ + A□2≡⋀ = refl + + A4□≃join : A4□ smash-span ≃ join X Y + A4□≃join = joinPushout≃join X Y + + A2□≃Unit : A2□ smash-span ≃ Unit + A2□≃Unit = Pushout⋁≃Unit _ _ + + A○□≡join : A○□ smash-span ≡ join X Y + A○□≡join = spanEquivToPushoutPath record { + e0 = invEquiv (pushoutIdfunEquiv _) ; + e2 = A2□≃Unit ; + e4 = A4□≃join ; + H1 = λ _ → refl ; + H3 = λ x → subst (λ x → inl x₀ ≡ A4□≃join .fst (f3□ smash-span x)) + (cong fst (A2□≃Unit .snd .equiv-proof _ .snd (x , refl))) + (sym (join-inr-null _)) + } ∙ sym (ua (pushoutIdfunEquiv' _)) + + A□○≡Σ : A□○ smash-span ≡ Susp (X∙ ⋀ Y∙) + A□○≡Σ = spanEquivToPushoutPath record { + e0 = invEquiv (pushoutIdfunEquiv _) ; + e2 = idEquiv _ ; + e4 = invEquiv (pushoutIdfunEquiv _) ; + H1 = λ _ → refl ; + H3 = λ _ → refl + } ∙ sym Susp≡Pushout + + join≡Susp : Susp (X∙ ⋀ Y∙) ≡ join X Y + join≡Susp = sym A□○≡Σ ∙ 3x3-lemma smash-span ∙ A○□≡join + From 72a75579ac9152673b56bda44f67c9352f7483d1 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Mon, 12 Aug 2024 03:58:36 +0800 Subject: [PATCH 61/73] Universes --- Cubical/HITs/Join/SuspWedgeEquiv.agda | 56 ++++++++++++++++++--------- 1 file changed, 38 insertions(+), 18 deletions(-) diff --git a/Cubical/HITs/Join/SuspWedgeEquiv.agda b/Cubical/HITs/Join/SuspWedgeEquiv.agda index 88a41efc9d..18f4542f9d 100644 --- a/Cubical/HITs/Join/SuspWedgeEquiv.agda +++ b/Cubical/HITs/Join/SuspWedgeEquiv.agda @@ -28,7 +28,7 @@ open import Cubical.HITs.Susp open import Cubical.HITs.Wedge open import Cubical.HITs.SmashProduct -module _ (X∙ @ (X , x₀) : Pointed₀) (Y∙ @ (Y , y₀) : Pointed₀) where +module _ {ℓ ℓ'} (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Pointed ℓ') where {- We consider the 3×3 lemma applied to @@ -42,9 +42,9 @@ module _ (X∙ @ (X , x₀) : Pointed₀) (Y∙ @ (Y , y₀) : Pointed₀) where open 3x3-span smash-span : 3x3-span - smash-span .A00 = Unit - smash-span .A02 = Unit - smash-span .A04 = Unit + smash-span .A00 = Unit* {ℓ} + smash-span .A02 = Unit* {ℓ-max ℓ ℓ'} + smash-span .A04 = Unit* {ℓ'} smash-span .A20 = X smash-span .A22 = X∙ ⋁ Y∙ smash-span .A24 = Y @@ -72,34 +72,54 @@ module _ (X∙ @ (X , x₀) : Pointed₀) (Y∙ @ (Y , y₀) : Pointed₀) where smash-span .H33 (inr x) = refl smash-span .H33 (push _ i) j = doubleCompPath-filler (refl {x = y₀}) refl refl j i - A□2≡⋀ : A□2 smash-span ≡ X∙ ⋀ Y∙ - A□2≡⋀ = refl + -- Perhaps some clever leveling will avoid these two trivial proofs + A□2≃Unit* : A0□ smash-span ≃ Unit* {ℓ-max ℓ ℓ'} + A□2≃Unit* = _ , record { equiv-proof = + λ _ → (inl tt* , refl) , + λ { (inl _ , q) → refl + ; (inr _ , q) → ΣPathP (push _ , refl) + ; (push _ i , q) j → push _ (i ∧ j) , refl } } + + open Iso + Iso-A□2-⋀ : Iso (A□2 smash-span) (X∙ ⋀ Y∙) + Iso-A□2-⋀ .fun (inl _) = inl _ + Iso-A□2-⋀ .fun (inr x) = inr x + Iso-A□2-⋀ .fun (push a i) = push a i + Iso-A□2-⋀ .inv (inl _) = inl _ + Iso-A□2-⋀ .inv (inr x) = inr x + Iso-A□2-⋀ .inv (push a i) = push a i + Iso-A□2-⋀ .rightInv (inl _) = refl + Iso-A□2-⋀ .rightInv (inr x) = refl + Iso-A□2-⋀ .rightInv (push a i) = refl + Iso-A□2-⋀ .leftInv (inl _) = refl + Iso-A□2-⋀ .leftInv (inr x) = refl + Iso-A□2-⋀ .leftInv (push a i) = refl A4□≃join : A4□ smash-span ≃ join X Y A4□≃join = joinPushout≃join X Y - A2□≃Unit : A2□ smash-span ≃ Unit - A2□≃Unit = Pushout⋁≃Unit _ _ + A2□≃Unit* : A2□ smash-span ≃ Unit* {ℓ-max ℓ ℓ'} + A2□≃Unit* = compEquiv (Pushout⋁≃Unit _ _) Unit≃Unit* A○□≡join : A○□ smash-span ≡ join X Y - A○□≡join = spanEquivToPushoutPath record { - e0 = invEquiv (pushoutIdfunEquiv _) ; - e2 = A2□≃Unit ; - e4 = A4□≃join ; - H1 = λ _ → refl ; - H3 = λ x → subst (λ x → inl x₀ ≡ A4□≃join .fst (f3□ smash-span x)) - (cong fst (A2□≃Unit .snd .equiv-proof _ .snd (x , refl))) - (sym (join-inr-null _)) + A○□≡join = spanEquivToPushoutPath record + { e0 = A□2≃Unit* + ; e2 = A2□≃Unit* + ; e4 = A4□≃join + ; H1 = λ _ → refl + ; H3 = λ x → subst (λ x → inl x₀ ≡ A4□≃join .fst (f3□ smash-span x)) + (cong fst (A2□≃Unit* .snd .equiv-proof _ .snd (x , refl))) + (sym (join-inr-null _)) } ∙ sym (ua (pushoutIdfunEquiv' _)) A□○≡Σ : A□○ smash-span ≡ Susp (X∙ ⋀ Y∙) A□○≡Σ = spanEquivToPushoutPath record { e0 = invEquiv (pushoutIdfunEquiv _) ; - e2 = idEquiv _ ; + e2 = isoToEquiv Iso-A□2-⋀ ; e4 = invEquiv (pushoutIdfunEquiv _) ; H1 = λ _ → refl ; H3 = λ _ → refl - } ∙ sym Susp≡Pushout + } ∙ sym (Susp≡Pushout {ℓ-max ℓ ℓ'}) join≡Susp : Susp (X∙ ⋀ Y∙) ≡ join X Y join≡Susp = sym A□○≡Σ ∙ 3x3-lemma smash-span ∙ A○□≡join From f9fc7e32263dfe5f5f23b26d2bac87fb9a9c0e86 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Mon, 12 Aug 2024 13:50:15 +0800 Subject: [PATCH 62/73] Use PushoutSquare instead --- Cubical/HITs/Join/SuspWedgeEquiv.agda | 2 +- Cubical/HITs/Susp/Properties.agda | 55 ++++++++++++++++++--------- 2 files changed, 37 insertions(+), 20 deletions(-) diff --git a/Cubical/HITs/Join/SuspWedgeEquiv.agda b/Cubical/HITs/Join/SuspWedgeEquiv.agda index 18f4542f9d..36f514051c 100644 --- a/Cubical/HITs/Join/SuspWedgeEquiv.agda +++ b/Cubical/HITs/Join/SuspWedgeEquiv.agda @@ -80,7 +80,7 @@ module _ {ℓ ℓ'} (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Point ; (inr _ , q) → ΣPathP (push _ , refl) ; (push _ i , q) j → push _ (i ∧ j) , refl } } - open Iso + open Iso -- This simply switches the Unit to Unit* Iso-A□2-⋀ : Iso (A□2 smash-span) (X∙ ⋀ Y∙) Iso-A□2-⋀ .fun (inl _) = inl _ Iso-A□2-⋀ .fun (inr x) = inr x diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index faa0adaf51..5ef2e10b41 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -10,6 +10,7 @@ open import Cubical.Foundations.Pointed open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence open import Cubical.Data.Bool open import Cubical.Data.Sigma @@ -75,25 +76,41 @@ Susp≃joinBool = isoToEquiv Susp-iso-joinBool Susp≡joinBool : ∀ {ℓ} {A : Type ℓ} → Susp A ≡ join A Bool Susp≡joinBool = isoToPath Susp-iso-joinBool -Susp-iso-Pushout : ∀ {ℓ} {A : Type ℓ} → Iso (Susp A) (Pushout (terminal A) (terminal A)) -fun Susp-iso-Pushout north = inl _ -fun Susp-iso-Pushout south = inr _ -fun Susp-iso-Pushout (merid a i) = push a i -inv Susp-iso-Pushout (inl _) = north -inv Susp-iso-Pushout (inr _) = south -inv Susp-iso-Pushout (push a i) = merid a i -rightInv Susp-iso-Pushout (inl _) = refl -rightInv Susp-iso-Pushout (inr _) = refl -rightInv Susp-iso-Pushout (push a i) = refl -leftInv Susp-iso-Pushout north = refl -leftInv Susp-iso-Pushout south = refl -leftInv Susp-iso-Pushout (merid a i) = refl - -Susp≃Pushout : ∀ {ℓ} {A : Type ℓ} → Susp A ≃ Pushout (terminal A) (terminal A) -Susp≃Pushout = isoToEquiv Susp-iso-Pushout - -Susp≡Pushout : ∀ {ℓ} {A : Type ℓ} → Susp A ≡ Pushout (terminal A) (terminal A) -Susp≡Pushout = isoToPath Susp-iso-Pushout +-- Here Unit* types are more convenient for general A +SuspSquare : ∀ {ℓ} ℓ' ℓ'' (A : Type ℓ) → commSquare +SuspSquare ℓ' ℓ'' A = record + { sp = record { A2 = A ; A0 = Unit* {ℓ'} ; A4 = Unit* {ℓ''} } + ; P = Susp A + ; inlP = λ _ → north + ; inrP = λ _ → south + ; comm = funExt merid + } + +SuspPushoutSquare : ∀ {ℓ} ℓ' ℓ'' (A : Type ℓ) + → isPushoutSquare (SuspSquare ℓ' ℓ'' A) +SuspPushoutSquare ℓ' ℓ'' A = isoToIsEquiv (iso _ inverse rInv lInv) + where + inverse : _ + inverse north = inl _ + inverse south = inr _ + inverse (merid a i) = push a i + + rInv : _ + rInv north = refl + rInv south = refl + rInv (merid a i) = refl + + lInv : _ + lInv (inl x) = refl + lInv (inr x) = refl + lInv (push a i) = refl + +Susp≃Pushout : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} → Susp A ≃ Pushout _ _ +Susp≃Pushout {ℓ} {ℓ'} {ℓ''} {A} = invEquiv (_ , SuspPushoutSquare ℓ' ℓ'' A) + +Susp≡Pushout : ∀ {ℓ ℓ' ℓ''} {A : Type _} → Susp A ≡ Pushout _ _ +Susp≡Pushout {ℓ} {ℓ'} {ℓ''} = ua + (Susp≃Pushout {ℓ-max ℓ (ℓ-max ℓ' ℓ'')} {ℓ'} {ℓ''}) congSuspIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso (Susp A) (Susp B) fun (congSuspIso is) = suspFun (fun is) From dac0c0ef7abc0c21de26287154dba5c5f5e0f028 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Tue, 13 Aug 2024 02:07:32 +0800 Subject: [PATCH 63/73] Pushout using Unit* --- Cubical/HITs/Wedge/Properties.agda | 52 +++++++++++++++++++++++------- 1 file changed, 41 insertions(+), 11 deletions(-) diff --git a/Cubical/HITs/Wedge/Properties.agda b/Cubical/HITs/Wedge/Properties.agda index b7c66fbe25..4c5b43eb3f 100644 --- a/Cubical/HITs/Wedge/Properties.agda +++ b/Cubical/HITs/Wedge/Properties.agda @@ -43,19 +43,49 @@ Iso.inv ⋁-commIso = ⋁-commFun Iso.rightInv ⋁-commIso = ⋁-commFun² Iso.leftInv ⋁-commIso = ⋁-commFun² +-- Pushout square using Unit* for convenience +⋁-PushoutSquare : ∀ (A : Pointed ℓ) (B : Pointed ℓ') ℓ'' → PushoutSquare +⋁-PushoutSquare A B ℓ'' = record + { sp = record + { A0 = typ A + ; A2 = Unit* {ℓ''} + ; A4 = typ B + ; f1 = λ _ → pt A + ; f3 = λ _ → pt B } + ; P = A ⋁ B + ; inlP = inl + ; inrP = inr + ; comm = funExt λ _ → push _ } , + isoToIsEquiv (iso _ inv lInv rInv) + where + inv : _ + inv (inl a) = inl a + inv (inr b) = inr b + inv (push _ i) = push _ i + + rInv : _ + rInv (inl a) = refl + rInv (inr b) = refl + rInv (push _ i) = refl + + lInv : _ + lInv (inl a) = refl + lInv (inr b) = refl + lInv (push _ i) = refl + -- cofibre of A --inl→ A ⋁ B is B cofibInl-⋁-square : (A : Pointed ℓ) (B : Pointed ℓ') → commSquare -cofibInl-⋁-square A B = record { - sp = record { - A0 = Unit ; - A2 = typ A ; - A4 = A ⋁ B ; - f1 = _ ; - f3 = inl } ; - P = typ B ; - inlP = λ _ → pt B ; - inrP = ⋁proj₂ A B ; - comm = refl } +cofibInl-⋁-square A B = record + { sp = record + { A0 = Unit + ; A2 = typ A + ; A4 = A ⋁ B + ; f1 = _ + ; f3 = inl } + ; P = typ B + ; inlP = λ _ → pt B + ; inrP = ⋁proj₂ A B + ; comm = refl } cofibInl-⋁-Pushout : (A : Pointed ℓ) (B : Pointed ℓ') → isPushoutSquare (cofibInl-⋁-square A B) From a4b23dd2b9dd08b3d614d447aee9ba435a47977f Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Thu, 15 Aug 2024 17:45:42 +0800 Subject: [PATCH 64/73] Rename --- Cubical/HITs/Join/SuspWedgeEquiv.agda | 2 +- Cubical/HITs/Susp/Properties.agda | 10 +++++----- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Cubical/HITs/Join/SuspWedgeEquiv.agda b/Cubical/HITs/Join/SuspWedgeEquiv.agda index 36f514051c..1751034492 100644 --- a/Cubical/HITs/Join/SuspWedgeEquiv.agda +++ b/Cubical/HITs/Join/SuspWedgeEquiv.agda @@ -119,7 +119,7 @@ module _ {ℓ ℓ'} (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Point e4 = invEquiv (pushoutIdfunEquiv _) ; H1 = λ _ → refl ; H3 = λ _ → refl - } ∙ sym (Susp≡Pushout {ℓ-max ℓ ℓ'}) + } ∙ sym (Susp≡PushoutSusp* {ℓ-max ℓ ℓ'}) join≡Susp : Susp (X∙ ⋀ Y∙) ≡ join X Y join≡Susp = sym A□○≡Σ ∙ 3x3-lemma smash-span ∙ A○□≡join diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index 5ef2e10b41..1cb0000ba1 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -105,12 +105,12 @@ SuspPushoutSquare ℓ' ℓ'' A = isoToIsEquiv (iso _ inverse rInv lInv) lInv (inr x) = refl lInv (push a i) = refl -Susp≃Pushout : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} → Susp A ≃ Pushout _ _ -Susp≃Pushout {ℓ} {ℓ'} {ℓ''} {A} = invEquiv (_ , SuspPushoutSquare ℓ' ℓ'' A) +Susp≃PushoutSusp* : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} → Susp A ≃ Pushout _ _ +Susp≃PushoutSusp* {ℓ} {ℓ'} {ℓ''} {A} = invEquiv (_ , SuspPushoutSquare ℓ' ℓ'' A) -Susp≡Pushout : ∀ {ℓ ℓ' ℓ''} {A : Type _} → Susp A ≡ Pushout _ _ -Susp≡Pushout {ℓ} {ℓ'} {ℓ''} = ua - (Susp≃Pushout {ℓ-max ℓ (ℓ-max ℓ' ℓ'')} {ℓ'} {ℓ''}) +Susp≡PushoutSusp* : ∀ {ℓ ℓ' ℓ''} {A : Type _} → Susp A ≡ Pushout _ _ +Susp≡PushoutSusp* {ℓ} {ℓ'} {ℓ''} = ua + (Susp≃PushoutSusp* {ℓ-max ℓ (ℓ-max ℓ' ℓ'')} {ℓ'} {ℓ''}) congSuspIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso (Susp A) (Susp B) fun (congSuspIso is) = suspFun (fun is) From 98bbb7a28cdb7d3465455615ad5fadf98970dfaf Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Sun, 18 Aug 2024 00:08:43 +0800 Subject: [PATCH 65/73] tweak --- Cubical/HITs/Join/SuspWedgeEquiv.agda | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Cubical/HITs/Join/SuspWedgeEquiv.agda b/Cubical/HITs/Join/SuspWedgeEquiv.agda index 1751034492..05a2092c2f 100644 --- a/Cubical/HITs/Join/SuspWedgeEquiv.agda +++ b/Cubical/HITs/Join/SuspWedgeEquiv.agda @@ -121,6 +121,5 @@ module _ {ℓ ℓ'} (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Point H3 = λ _ → refl } ∙ sym (Susp≡PushoutSusp* {ℓ-max ℓ ℓ'}) - join≡Susp : Susp (X∙ ⋀ Y∙) ≡ join X Y - join≡Susp = sym A□○≡Σ ∙ 3x3-lemma smash-span ∙ A○□≡join - + Susp≡join : Susp (X∙ ⋀ Y∙) ≡ join X Y + Susp≡join = sym A□○≡Σ ∙ 3x3-lemma smash-span ∙ A○□≡join From dee0c5608ea276a7c97744ecbacc9eee05260df1 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Sun, 18 Aug 2024 00:08:50 +0800 Subject: [PATCH 66/73] extend commutative squares --- Cubical/HITs/Pushout/Properties.agda | 24 ++++++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index 9cc1c6cbf4..98cecb20ea 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -831,6 +831,30 @@ module _ {ℓ₀ ℓ₂ ℓ₄ ℓP : Level} where PushoutSquare : Type (ℓ-suc ℓ*) PushoutSquare = Σ commSquare isPushoutSquare +module _ {ℓ₀ ℓ₂ ℓ₄ ℓP ℓP' : Level} + {P' : Type ℓP'} where + open commSquare + extendCommSquare : (sk : commSquare {ℓ₀} {ℓ₂} {ℓ₄} {ℓP}) + → (sk .P → P') → commSquare + extendCommSquare sk f = record + { sp = sk .sp + ; P = P' + ; inlP = f ∘ sk .inlP + ; inrP = f ∘ sk .inrP + ; comm = cong (f ∘_) (sk .comm) + } + + extendPushoutSquare : (sk : PushoutSquare {ℓ₀} {ℓ₂} {ℓ₄} {ℓP}) + → (e : sk .fst .P ≃ P') → PushoutSquare + extendPushoutSquare sk e = (extendCommSquare (sk .fst) (e .fst) , + subst isEquiv H (compEquiv (_ , sk .snd) e .snd)) + where + H : e .fst ∘ _ ≡ _ + H = funExt λ + { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl } + -- Pushout itself fits into a pushout square pushoutToSquare : 3-span {ℓ} {ℓ'} {ℓ''} → PushoutSquare pushoutToSquare sp = record { From 51fa7d02ec85a68ec34a0fb4c405e2298b1cecf9 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Thu, 19 Sep 2024 20:12:35 +0800 Subject: [PATCH 67/73] Remove my version in favor of the other --- Cubical/HITs/Join/SuspWedgeEquiv.agda | 125 -------------------------- 1 file changed, 125 deletions(-) delete mode 100644 Cubical/HITs/Join/SuspWedgeEquiv.agda diff --git a/Cubical/HITs/Join/SuspWedgeEquiv.agda b/Cubical/HITs/Join/SuspWedgeEquiv.agda deleted file mode 100644 index 05a2092c2f..0000000000 --- a/Cubical/HITs/Join/SuspWedgeEquiv.agda +++ /dev/null @@ -1,125 +0,0 @@ -{- - -For pointed types, join X Y is equivalent to the suspension of their wedge. - -This file is split off to prevent cyclic imports. - --} - - -{-# OPTIONS --safe --lossy-unification #-} - -module Cubical.HITs.Join.SuspWedgeEquiv where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Function -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Univalence - -open import Cubical.Data.Unit -open import Cubical.Data.Sigma - -open import Cubical.HITs.Pushout -open import Cubical.HITs.Join -open import Cubical.HITs.Susp -open import Cubical.HITs.Wedge -open import Cubical.HITs.SmashProduct - -module _ {ℓ ℓ'} (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Pointed ℓ') where - - {- - We consider the 3×3 lemma applied to - 1 <--- 1 ---> 1 - ↑ ↑ ↑ - X <- X ⋁ Y -> Y - ↓ ↓ ↓ - X <- X × Y -> Y - This admittedly makes it annoying to accomodate for higher universes. - -} - - open 3x3-span - smash-span : 3x3-span - smash-span .A00 = Unit* {ℓ} - smash-span .A02 = Unit* {ℓ-max ℓ ℓ'} - smash-span .A04 = Unit* {ℓ'} - smash-span .A20 = X - smash-span .A22 = X∙ ⋁ Y∙ - smash-span .A24 = Y - smash-span .A40 = X - smash-span .A42 = X × Y - smash-span .A44 = Y - smash-span .f10 = _ -- Unique map - smash-span .f12 = _ - smash-span .f14 = _ - smash-span .f30 = idfun X - smash-span .f32 = ⋁↪ - smash-span .f34 = idfun Y - smash-span .f01 = _ - smash-span .f21 = ⋁proj₁ X∙ Y∙ - smash-span .f41 = fst - smash-span .f03 = _ - smash-span .f23 = ⋁proj₂ X∙ Y∙ - smash-span .f43 = snd - smash-span .H11 _ = refl - smash-span .H13 _ = refl - smash-span .H31 (inl x) = refl - smash-span .H31 (inr y) = refl - smash-span .H31 (push _ i) j = doubleCompPath-filler (refl {x = x₀}) refl refl j i - smash-span .H33 (inl x) = refl - smash-span .H33 (inr x) = refl - smash-span .H33 (push _ i) j = doubleCompPath-filler (refl {x = y₀}) refl refl j i - - -- Perhaps some clever leveling will avoid these two trivial proofs - A□2≃Unit* : A0□ smash-span ≃ Unit* {ℓ-max ℓ ℓ'} - A□2≃Unit* = _ , record { equiv-proof = - λ _ → (inl tt* , refl) , - λ { (inl _ , q) → refl - ; (inr _ , q) → ΣPathP (push _ , refl) - ; (push _ i , q) j → push _ (i ∧ j) , refl } } - - open Iso -- This simply switches the Unit to Unit* - Iso-A□2-⋀ : Iso (A□2 smash-span) (X∙ ⋀ Y∙) - Iso-A□2-⋀ .fun (inl _) = inl _ - Iso-A□2-⋀ .fun (inr x) = inr x - Iso-A□2-⋀ .fun (push a i) = push a i - Iso-A□2-⋀ .inv (inl _) = inl _ - Iso-A□2-⋀ .inv (inr x) = inr x - Iso-A□2-⋀ .inv (push a i) = push a i - Iso-A□2-⋀ .rightInv (inl _) = refl - Iso-A□2-⋀ .rightInv (inr x) = refl - Iso-A□2-⋀ .rightInv (push a i) = refl - Iso-A□2-⋀ .leftInv (inl _) = refl - Iso-A□2-⋀ .leftInv (inr x) = refl - Iso-A□2-⋀ .leftInv (push a i) = refl - - A4□≃join : A4□ smash-span ≃ join X Y - A4□≃join = joinPushout≃join X Y - - A2□≃Unit* : A2□ smash-span ≃ Unit* {ℓ-max ℓ ℓ'} - A2□≃Unit* = compEquiv (Pushout⋁≃Unit _ _) Unit≃Unit* - - A○□≡join : A○□ smash-span ≡ join X Y - A○□≡join = spanEquivToPushoutPath record - { e0 = A□2≃Unit* - ; e2 = A2□≃Unit* - ; e4 = A4□≃join - ; H1 = λ _ → refl - ; H3 = λ x → subst (λ x → inl x₀ ≡ A4□≃join .fst (f3□ smash-span x)) - (cong fst (A2□≃Unit* .snd .equiv-proof _ .snd (x , refl))) - (sym (join-inr-null _)) - } ∙ sym (ua (pushoutIdfunEquiv' _)) - - A□○≡Σ : A□○ smash-span ≡ Susp (X∙ ⋀ Y∙) - A□○≡Σ = spanEquivToPushoutPath record { - e0 = invEquiv (pushoutIdfunEquiv _) ; - e2 = isoToEquiv Iso-A□2-⋀ ; - e4 = invEquiv (pushoutIdfunEquiv _) ; - H1 = λ _ → refl ; - H3 = λ _ → refl - } ∙ sym (Susp≡PushoutSusp* {ℓ-max ℓ ℓ'}) - - Susp≡join : Susp (X∙ ⋀ Y∙) ≡ join X Y - Susp≡join = sym A□○≡Σ ∙ 3x3-lemma smash-span ∙ A○□≡join From 08a56c498329612d3d6e9d4ae27b6655a91b2873 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Thu, 19 Sep 2024 20:21:45 +0800 Subject: [PATCH 68/73] Oddly specific sigma lemma --- Cubical/Foundations/Pointed/Properties.agda | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Cubical/Foundations/Pointed/Properties.agda b/Cubical/Foundations/Pointed/Properties.agda index ae24e72bf9..65bc99668c 100644 --- a/Cubical/Foundations/Pointed/Properties.agda +++ b/Cubical/Foundations/Pointed/Properties.agda @@ -245,6 +245,14 @@ compPathlEquiv∙ : {A : Type ℓ} {a b c : A} {q : b ≡ c} (p : a ≡ b) fst (compPathlEquiv∙ p) = compPathlEquiv p snd (compPathlEquiv∙ p) = refl +-- Pointed version of Σ-cong-equiv-snd +Σ-cong-equiv-snd∙ : ∀ {ℓ ℓ'} {A : Type ℓ} {B₁ B₂ : A → Type ℓ'} + → {a : A} {b₁ : B₁ a} {b₂ : B₂ a} + → (e : ∀ a → B₁ a ≃ B₂ a) + → fst (e a) b₁ ≡ b₂ + → Path (Pointed _) (Σ A B₁ , a , b₁) (Σ A B₂ , a , b₂) +Σ-cong-equiv-snd∙ e p = ua∙ (Σ-cong-equiv-snd e) (ΣPathP (refl , p)) + -- a pointed map is constant iff its non-pointed part is constant constantPointed≡ : {A B : Pointed ℓ} (f : A →∙ B) → fst f ≡ fst (const∙ A B) → f ≡ const∙ A B From c983de796ea23119cfdc910eb41c3d7a77525962 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Thu, 19 Sep 2024 20:36:36 +0800 Subject: [PATCH 69/73] Splitting of loopspace --- Cubical/Homotopy/Loopspace.agda | 79 +++++++++++++++++++++++++++++++++ 1 file changed, 79 insertions(+) diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 36fef7ea1f..9b74e0a59a 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -13,8 +13,10 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Function open import Cubical.Foundations.Path +open import Cubical.Foundations.Univalence open import Cubical.Functions.Morphism +open import Cubical.Functions.Fibration open import Cubical.Data.Nat open import Cubical.Data.Sigma @@ -460,3 +462,80 @@ EH-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : ∥ typ ((Ω^ (2 + n)) A) → π-comp (1 + n) p q ≡ π-comp (1 + n) q p EH-π n = elim2 (λ x y → isOfHLevelPath 2 isSetSetTrunc _ _) λ p q → cong ∣_∣₂ (EH n p q) + +{- + If A -> B -> C is a fiber sequence, and Ω B -> Ω C has a section, + then Ω B is the product Ω A × Ω C. + + https://gist.github.com/ecavallo/5b75313d36977c51ca3563de82506123 +-} + +module _ {ℓ} {B C : Type ℓ} (b₀ : B) (π : B → C) where + private + c₀ = π b₀ + A = fiber π c₀ + a₀ : A + a₀ = b₀ , refl + ι : A → B + ι = fst + π∘ι≡0 : ∀ a → π (ι a) ≡ c₀ + π∘ι≡0 = snd + + A∙ B∙ C∙ : Pointed _ + A∙ = (A , a₀) + B∙ = (B , b₀) + C∙ = (C , c₀) + + TwistedProduct : Pointed _ + TwistedProduct = ((Σ[ p ∈ typ (Ω C∙) ] (b₀ , p) ≡ a₀) , refl , refl) + + -- The middle term of a fiber sequence always splits + -- as a twisted product without any conditions. + presplit : Ω B∙ ≡ TwistedProduct + presplit = cong Ω (ua∙ (totalEquiv π) refl) + ∙ sym (ua∙ ΣPath≃PathΣ refl) + ∙ Σ-cong-equiv-snd∙ (λ p + → PathP≃Path _ _ _ ∙ₑ + compPathlEquiv (sym (fromPathP + (ΣPathP (refl , λ i j → p (i ∧ j)))))) + (rCancel λ j → transp (λ _ → A) (~ j) a₀) + + module Split (s : typ (Ω C∙) → typ (Ω B∙)) + (section : ∀ p → cong π (s p) ≡ p) where + movePoint : ∀ p → a₀ ≡ (b₀ , sym p) + movePoint p = ΣPathP (s p , + λ i j → slideSquare (section p) i (~ j)) + + splitting : typ (Ω B∙) ≡ typ (Ω C∙) × typ (Ω A∙) + splitting = cong typ presplit ∙ ua + (Σ-cong-equiv-snd λ p → compPathlEquiv (movePoint _)) + + -- If the section furthermore respects the base point, then + -- the splitting does too. + module _ (h : s refl ≡ refl) (coh : cong (cong π) h ≡ section refl) where + movePoint-refl : movePoint refl ≡ refl + movePoint-refl i j = h i j , λ k → + slideSquare (cong sym coh ∙ lemma _ _) (~ i) (~ k) j + where + lemma : ∀ {ℓ} {X : Type ℓ} {x : X} (p : x ≡ x) (q : p ≡ refl) + → sym q ≡ flipSquare (slideSquare q) + lemma {x = x} p q = J (λ p (q : refl ≡ p) + → q ≡ flipSquare (slideSquare (sym q))) + (λ i j k → hfill + (slideSquareFaces + {a₀₀ = x} {a₋ = refl} {a₁₋ = refl} {a₋₀ = refl} + j k) (inS x) i) + (sym q) + + twisted∙ : TwistedProduct ≡ Ω C∙ ×∙ Ω A∙ + twisted∙ = Σ-cong-equiv-snd∙ + (λ p → compPathlEquiv (movePoint _)) + (cong (λ t → compPathlEquiv t .fst refl) movePoint-refl ∙ sym (rUnit _)) + + splitting∙ : Ω B∙ ≡ Ω C∙ ×∙ Ω A∙ + splitting∙ = presplit ∙ twisted∙ + + -- splitting and splitting∙ agrees + splitting-typ : cong typ splitting∙ ≡ splitting + splitting-typ = congFunct typ presplit twisted∙ + \ No newline at end of file From d99570a63326ea5471a694e0f55a00aefef93d45 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Thu, 19 Sep 2024 21:10:57 +0800 Subject: [PATCH 70/73] =?UTF-8?q?Hilton=E2=80=93Milnor?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Cubical/Homotopy/HiltonMilnor.agda | 96 ++++++++++++++++++++++++++++++ 1 file changed, 96 insertions(+) create mode 100644 Cubical/Homotopy/HiltonMilnor.agda diff --git a/Cubical/Homotopy/HiltonMilnor.agda b/Cubical/Homotopy/HiltonMilnor.agda new file mode 100644 index 0000000000..c8f967ac0b --- /dev/null +++ b/Cubical/Homotopy/HiltonMilnor.agda @@ -0,0 +1,96 @@ +{-# OPTIONS --cubical --lossy-unification #-} + +{- + The finitary Hilton–Milnor splitting + Ω(X ⋁ Y) ≃ ΩX × ΩY × ΩΣ(ΩY ⋀ ΩX) + for pointed types X and Y. +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism using (isoToPath; isoToEquiv) +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Pushout.Flattening +open import Cubical.HITs.SmashProduct +open import Cubical.HITs.Susp +open import Cubical.HITs.Join +open import Cubical.HITs.Wedge + +module Cubical.Homotopy.HiltonMilnor {ℓ} (X∙@(X , x₀) Y∙@(Y , y₀) : Pointed ℓ) where + +fun : X∙ ⋁ Y∙ → typ X∙ × typ Y∙ +fun = ⋁↪ {A = X∙} {B = Y∙} + +-- The required section, given by concatenating the paths +sect : typ (Ω (X∙ ×∙ Y∙)) → typ (Ω (X∙ ⋁∙ₗ Y∙)) +sect p = cong (inl ∘ fst) p ∙ push _ ∙ cong (inr ∘ snd) p ∙ sym (push _) + +rInv : ∀ c → cong fun (sect c) ≡ c +rInv c = congFunct fun _ _ + ∙ cong (cong (fun ∘ inl ∘ fst) c ∙_) + (congFunct fun (push _) _ + ∙ sym (lUnit _) + ∙ lemma₂ -- can't be inlined because lossy unification + ∙ sym (rUnit _)) + ∙ lemma₁ (cong fst c) (cong snd c) + where + lemma₁ : (p : x₀ ≡ x₀) (q : y₀ ≡ y₀) + → cong (_, y₀) p ∙ cong (x₀ ,_) q ≡ cong₂ _,_ p q + lemma₁ p q = cong ΣPathP (ΣPathP + (congFunct fst (cong (_, y₀) p) _ ∙ sym (rUnit _) , + congFunct snd (cong (_, y₀) p) _ ∙ sym (lUnit _))) + + lemma₂ : cong fun + (cong (inr ∘ snd) c ∙ sym (push _)) + ≡ cong (⋁↪ ∘ inr ∘ snd) c ∙ refl + lemma₂ = congFunct fun _ _ + +-- Now the splitting lemma gives us something in terms of the fiber of ⋁↪. +-- So we try to characterize the type. +module _ where + open FlatteningLemma + (λ (_ : Unit) → x₀) (λ _ → y₀) + (λ x → (x , y₀) ≡ (x₀ , y₀)) (λ y → (x₀ , y) ≡ (x₀ , y₀)) + (λ _ → idEquiv _) + + private + E≡ : ∀ u → E u ≡ (⋁↪ u ≡ (x₀ , y₀)) + E≡ (inl x) = refl + E≡ (inr x) = refl + E≡ (push a i) j = uaIdEquiv {A = typ (Ω (X∙ ×∙ Y∙))} j i + + Pushout≃ : Pushout Σf Σg ≃ join (typ (Ω Y∙)) (typ (Ω X∙)) + Pushout≃ = pushoutEquiv _ _ _ _ + (ΣUnit _ ∙ₑ invEquiv ΣPathP≃PathPΣ ∙ₑ Σ-swap-≃) + (Σ-cong-equiv-snd (λ _ → invEquiv ΣPathP≃PathPΣ ∙ₑ Σ-swap-≃) + ∙ₑ invEquiv Σ-assoc-≃ + ∙ₑ invEquiv (fiberProjEquiv X _ x₀)) + (Σ-cong-equiv-snd (λ _ → invEquiv ΣPathP≃PathPΣ) + ∙ₑ invEquiv Σ-assoc-≃ + ∙ₑ invEquiv (fiberProjEquiv Y _ y₀)) + (funExt λ _ → transportRefl _) + (funExt λ _ → transportRefl _) + ∙ₑ joinPushout≃join _ _ + + fiber⋁↪≃ : fiber fun (x₀ , y₀) ≃ join (typ (Ω Y∙)) (typ (Ω X∙)) + fiber⋁↪≃ = Σ-cong-equiv-snd (λ a → pathToEquiv (sym (E≡ a))) + ∙ₑ flatten ∙ₑ Pushout≃ + +HMSplit : typ (Ω (X∙ ⋁∙ₗ Y∙)) ≡ typ (Ω X∙) × typ (Ω Y∙) × typ (Ω (Susp∙ (Ω Y∙ ⋀ Ω X∙))) +HMSplit = splitting + ∙ cong₂ (λ A B∙ → A × typ (Ω B∙)) (sym ΣPathP≡PathPΣ) + (ua∙ fiber⋁↪≃ refl ∙ ΣPathP + (sym (isoToPath SmashJoinIso) , toPathP refl)) + ∙ ua Σ-assoc-≃ + where open Split (inl x₀) ⋁↪ sect rInv From f0d49bc456b0798603ef0ed6e9bdc88b9a12ff84 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Thu, 19 Sep 2024 21:27:40 +0800 Subject: [PATCH 71/73] James --- Cubical/HITs/James/Stable.agda | 87 +++++++++++ Cubical/HITs/Susp/SuspProduct.agda | 232 +++++++++++++++++++++++++++++ 2 files changed, 319 insertions(+) create mode 100644 Cubical/HITs/James/Stable.agda create mode 100644 Cubical/HITs/Susp/SuspProduct.agda diff --git a/Cubical/HITs/James/Stable.agda b/Cubical/HITs/James/Stable.agda new file mode 100644 index 0000000000..7c9915f495 --- /dev/null +++ b/Cubical/HITs/James/Stable.agda @@ -0,0 +1,87 @@ +{-# OPTIONS --cubical #-} + +{- + The stable version of the James splitting: + ΣΩΣX ≃ ΣX ⋁ Σ(X ⋀ ΩΣX) +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.HITs.Susp +open import Cubical.HITs.Susp.SuspProduct +open import Cubical.HITs.Pushout +open import Cubical.HITs.Pushout.Flattening +open import Cubical.HITs.Wedge +open import Cubical.HITs.SmashProduct + +open import Cubical.Homotopy.Loopspace + +module Cubical.HITs.James.Stable {ℓ} (X∙@(X , x₀) : Pointed ℓ) where + +module ContrPushout where + Code : Pushout (terminal X) (terminal X) → Type ℓ + Code x = inl _ ≡ x + + ΩΣX = Code (inl _) + ΩΣX∙ : Pointed _ + ΩΣX∙ = ΩΣX , refl + + α : X × ΩΣX → ΩΣX + α (x , p) = (p ∙ push x) ∙ sym (push x₀) + + open FlatteningLemma + (terminal X) (terminal X) + (Code ∘ inl) (Code ∘ inr) + (pathToEquiv ∘ cong (inl tt ≡_) ∘ push) + + pushoutEq : Pushout Σf Σg ≃ Pushout snd α + pushoutEq = pushoutEquiv _ _ _ _ + (idEquiv (X × ΩΣX)) (ΣUnit _) + (ΣUnit _ ∙ₑ compPathrEquiv (sym (push x₀))) + refl (funExt λ (x , p) + → cong (_∙ sym (push x₀)) (substInPathsL (push x) p)) + + Code≡E : ∀ x → Code x ≡ E x + Code≡E (inl _) = refl + Code≡E (inr _) = refl + Code≡E (push a i) j = uaη (cong Code (push a)) (~ j) i + + isContrΣE : isContr (Σ _ E) + isContrΣE = subst isContr (Σ-cong-snd Code≡E) (isContrSingl (inl tt)) + + isContrPushout : isContr (Pushout snd α) + isContrPushout = isOfHLevelRespectEquiv _ (flatten ∙ₑ pushoutEq) isContrΣE + +open ContrPushout + +LoopSuspSquare : commSquare +LoopSuspSquare = record + { sp = record { f1 = snd ; f3 = α } + ; P = Unit* {ℓ} + ; comm = refl } + +LoopSuspPushoutSquare : PushoutSquare +LoopSuspPushoutSquare = (LoopSuspSquare , isContr→≃Unit* isContrPushout .snd) + +open PushoutPasteLeft LoopSuspPushoutSquare + (λ _ → lift {j = ℓ} tt) _ _ (funExt merid) + +cofib-snd-James : cofib (λ (r : X × ΩΣX) → snd r) ≃ Susp ΩΣX +cofib-snd-James = pushoutSwitchEquiv + ∙ₑ pushoutEquiv snd _ snd _ + (idEquiv _) (idEquiv _) Unit≃Unit* refl refl + ∙ₑ (_ , isPushoutRightSquare→isPushoutTotSquare + (SuspPushoutSquare _ _ ΩΣX)) + +StableJames : Susp ΩΣX ≃ Susp∙ X ⋁ Susp∙ (X∙ ⋀ ΩΣX∙) +StableJames = invEquiv cofib-snd-James ∙ₑ cofib-snd X∙ (ΩΣX , refl) diff --git a/Cubical/HITs/Susp/SuspProduct.agda b/Cubical/HITs/Susp/SuspProduct.agda new file mode 100644 index 0000000000..d847ff614e --- /dev/null +++ b/Cubical/HITs/Susp/SuspProduct.agda @@ -0,0 +1,232 @@ +{-# OPTIONS --cubical #-} + +{- + The suspension of a Cartesian product is given by a triple wedge sum + Σ (X × Y) = Σ X ⋁ Σ Y ⋁ Σ (X ⋀ Y) +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.SmashProduct +open import Cubical.HITs.Join +open import Cubical.HITs.Wedge + +module Cubical.HITs.Susp.SuspProduct where + +module PushoutNull {ℓ ℓ'} (X : Type ℓ) (Y : Type ℓ') (y₀ : Y) where + + {- The pushout of a null map is equivalent to a wedge with the suspension + X --> 1 --> Y + ↓ ↓ ↓ + 1 -> Σ X -> ? = B + -} + + module _ where + open PushoutPasteLeft + (rotatePushoutSquare (_ , SuspPushoutSquare ℓ-zero ℓ-zero X)) + {B = (Susp X , north) ⋁ (Y , y₀)} + (λ _ → y₀) inr inl (funExt (λ _ → push _)) + + squareL : commSquare + squareL = totSquare + + pushoutSquareL : PushoutSquare + pushoutSquareL = squareL , isPushoutRightSquare→isPushoutTotSquare + (snd (⋁-PushoutSquare _ _ ℓ-zero)) + + module _ where + open PushoutPasteDown + (_ , SuspPushoutSquare ℓ-zero ℓ-zero X) + {B = (Susp X , north) ⋁ (Y , y₀)} + (λ _ → y₀) inr inl (funExt (λ _ → sym (push _))) + + squareR : commSquare + squareR = totSquare + + pushoutSquareR : PushoutSquare + pushoutSquareR = squareR , isPushoutBottomSquare→isPushoutTotSquare + (snd (rotatePushoutSquare (⋁-PushoutSquare _ _ ℓ-zero))) + +module WedgePushout {ℓ ℓ' ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') where + + open 3x3-span + span : 3x3-span + span .A40 = typ A + span .A42 = typ A + span .A44 = typ A + span .A20 = Unit + span .A22 = Unit + span .A24 = Unit + span .A00 = typ B + span .A02 = Unit + span .A04 = typ C + span .f30 _ = pt A + span .f32 _ = pt A + span .f34 _ = pt A + span .f10 _ = pt B + span .f12 _ = _ + span .f14 _ = pt C + span .f41 = idfun (typ A) + span .f21 = _ + span .f01 _ = pt B + span .f43 = idfun (typ A) + span .f23 = _ + span .f03 _ = pt C + span .H11 _ = refl + span .H13 _ = refl + span .H31 _ = refl + span .H33 _ = refl + + A□2≃ : A□2 span ≃ typ A + A□2≃ = invEquiv (pushoutIdfunEquiv' _) + + f : typ A → B ⋁ A + f = inr + + g : typ A → C ⋁ A + g = inr + + A□○≃ : A□○ span ≃ Pushout f g + A□○≃ = pushoutEquiv _ _ _ _ + A□2≃ (idEquiv _) (idEquiv _) + (funExt λ + { (inl x) → push _ + ; (inr x) → refl + ; (push a i) j → sq (~ i) (~ j) }) + (funExt λ + { (inl x) → push _ + ; (inr x) → refl + ; (push a i) j → sq (~ i) (~ j) }) + where + sq : ∀ {ℓ} {C : Type ℓ} {c : C} + → Square refl (sym (Pushout.push _)) + refl (sym (push {f = λ x → c} {g = λ _ → pt A} + tt ∙ refl)) + sq = slideSquare (sym (rUnit _)) + + A4□≃ : A4□ span ≃ typ A + A4□≃ = invEquiv (pushoutIdfunEquiv _) + + A2□≃ : A2□ span ≃ Unit + A2□≃ = invEquiv (pushoutIdfunEquiv _) + + A2□isContr : isContr (A2□ span) + A2□isContr = isOfHLevelRespectEquiv 0 (pushoutIdfunEquiv _) isContrUnit + + A○□≃ : A○□ span ≃ (B ⋁∙ₗ C) ⋁ A + A○□≃ = pushoutEquiv _ _ _ _ + A2□≃ (idEquiv _) A4□≃ + (lemma A2□isContr refl) (lemma A2□isContr refl) + where + -- functions on contractible domains are determined by the basepoint + lemma : ∀ {ℓ ℓ'} + → {A : Type ℓ} {B : Type ℓ'} + → {f g : A → B} + → (c : isContr A) + → f (c .fst) ≡ g (c .fst) + → f ≡ g + lemma {f = f} {g} c p = funExt λ x + → cong f (sym (c .snd x)) + ∙ p + ∙ cong g (c .snd x) + + wedgePushout : PushoutSquare + wedgePushout = extendPushoutSquare + (pushoutToSquare record { f1 = f ; f3 = g }) + (invEquiv A□○≃ ∙ₑ pathToEquiv (3x3-lemma span) ∙ₑ A○□≃) + + +module _ {ℓ ℓ'} (X : Pointed ℓ) (Y : Pointed ℓ') where + open commSquare + open 3-span + + {- + X × Y ---> Y --------> 1 + ↓ A ↓ C ↓ + X ---> Σ(X∧Y) --> ΣY ∨ Σ(X∧Y) + ↓ B ↓ D ↓ + 1 -> ΣX ∨ Σ(X∧Y) --> ? = ΣX ∨ ΣY ∨ Σ(X∧Y) + -} + + pushoutA₀ : PushoutSquare -- inrP maps to south, but we want north + pushoutA₀ = extendPushoutSquare + (pushoutToSquare record { f1 = fst ; f3 = snd }) + (pathToEquiv (joinPushout≡join _ _) ∙ₑ + invEquiv (isoToEquiv (SmashJoinIso {A = X} {B = Y}))) + + squareA : I → commSquare + squareA i = record + { commSquare (pushoutA₀ .fst) + ; inrP = λ _ → merid (inl _) i + ; comm = transp (λ j → + Path (typ X × typ Y → Susp (X ⋀ Y)) + (λ _ → north) (λ _ → merid (inl _) (i ∨ ~ j))) + i (pushoutA₀ .fst .commSquare.comm) + } + + pushoutA : PushoutSquare + pushoutA = squareA i0 , + subst isPushoutSquare (λ i → squareA (~ i)) (pushoutA₀ .snd) + + module _ where + open PushoutNull (typ X) (Susp (X ⋀ Y)) north + open PushoutPasteDown pushoutA + (squareL .sp .f1) + (squareL .inlP) + (squareL .inrP) + (squareL .commSquare.comm) + + pushoutB : PushoutSquare + pushoutB = pushoutSquareL + + pushoutAB : PushoutSquare + pushoutAB = _ , isPushoutBottomSquare→isPushoutTotSquare (pushoutB .snd) + + cofib-snd : cofib snd ≃ Susp∙ (typ X) ⋁ Susp∙ (X ⋀ Y) + cofib-snd = pushoutEquiv _ _ _ _ + (idEquiv _) Unit≃Unit* (idEquiv _) + refl refl + ∙ₑ (_ , pushoutAB .snd) + + pushoutC : PushoutSquare + pushoutC = PushoutNull.pushoutSquareR (typ Y) (Susp (X ⋀ Y)) north + + pushoutD : PushoutSquare + pushoutD = WedgePushout.wedgePushout + (Susp∙ (X ⋀ Y)) (Susp∙ (typ X)) (Susp∙ (typ Y)) + + pushoutCD : PushoutSquare + pushoutCD = _ , isPushoutBottomSquare→isPushoutTotSquare (pushoutD .snd) + where + open PushoutPasteDown pushoutC + (pushoutD .fst .sp .f1) + (pushoutD .fst .inlP) + (pushoutD .fst .inrP) + (pushoutD .fst .commSquare.comm) + + pushoutTot : PushoutSquare + pushoutTot = _ , isPushoutRightSquare→isPushoutTotSquare (pushoutCD .snd) + where + open PushoutPasteLeft pushoutAB + (pushoutCD .fst .sp .f3) + (pushoutCD .fst .inrP) + (pushoutCD .fst .inlP) + (pushoutCD .fst .commSquare.comm) + + SuspProduct : (Susp∙ (typ X) ⋁∙ₗ Susp∙ (typ Y)) ⋁ Susp∙ (X ⋀ Y) + ≃ Susp (typ X × typ Y) + SuspProduct = invEquiv (_ , pushoutTot .snd) + ∙ₑ (_ , SuspPushoutSquare ℓ-zero ℓ-zero (typ X × typ Y)) From 54a2b1fa49ef4a26a62ff52f0f3ba9fa412ebe36 Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Thu, 19 Sep 2024 21:42:05 +0800 Subject: [PATCH 72/73] Fix whitespace --- Cubical/HITs/Wedge/Properties.agda | 4 ++-- Cubical/Homotopy/Loopspace.agda | 1 - 2 files changed, 2 insertions(+), 3 deletions(-) diff --git a/Cubical/HITs/Wedge/Properties.agda b/Cubical/HITs/Wedge/Properties.agda index ba496a4080..60043a6a86 100644 --- a/Cubical/HITs/Wedge/Properties.agda +++ b/Cubical/HITs/Wedge/Properties.agda @@ -618,7 +618,7 @@ module _ (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Pointed ℓ') wh {- The proof proceeds by applying the pasting lemma twice: 1 ----> Y - ↓ ↓ + ↓ ↓ X --> X ⋁ Y --> X ↓ mid ↓ bot ↓ 1 ----> Y ----> 1 @@ -646,6 +646,6 @@ module _ (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Pointed ℓ') wh where open PushoutPasteDown (rotatePushoutSquare (_ , midPushout)) fX (terminal X) (terminal Y) refl - + Pushout⋁≃Unit : Pushout fX fY ≃ Unit Pushout⋁≃Unit = _ , botPushout diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 9b74e0a59a..2f687d7aba 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -538,4 +538,3 @@ module _ {ℓ} {B C : Type ℓ} (b₀ : B) (π : B → C) where -- splitting and splitting∙ agrees splitting-typ : cong typ splitting∙ ≡ splitting splitting-typ = congFunct typ presplit twisted∙ - \ No newline at end of file From 01fa4208b2110765b2cdf93219756124550eff4d Mon Sep 17 00:00:00 2001 From: Trebor-Huang Date: Thu, 19 Sep 2024 22:16:54 +0800 Subject: [PATCH 73/73] safe flags --- Cubical/HITs/James/Stable.agda | 2 +- Cubical/HITs/Susp/SuspProduct.agda | 2 +- Cubical/Homotopy/HiltonMilnor.agda | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Cubical/HITs/James/Stable.agda b/Cubical/HITs/James/Stable.agda index 7c9915f495..48e6af5b73 100644 --- a/Cubical/HITs/James/Stable.agda +++ b/Cubical/HITs/James/Stable.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical #-} +{-# OPTIONS --safe #-} {- The stable version of the James splitting: diff --git a/Cubical/HITs/Susp/SuspProduct.agda b/Cubical/HITs/Susp/SuspProduct.agda index d847ff614e..6f00c0d62b 100644 --- a/Cubical/HITs/Susp/SuspProduct.agda +++ b/Cubical/HITs/Susp/SuspProduct.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical #-} +{-# OPTIONS --safe #-} {- The suspension of a Cartesian product is given by a triple wedge sum diff --git a/Cubical/Homotopy/HiltonMilnor.agda b/Cubical/Homotopy/HiltonMilnor.agda index c8f967ac0b..425a5e7977 100644 --- a/Cubical/Homotopy/HiltonMilnor.agda +++ b/Cubical/Homotopy/HiltonMilnor.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --lossy-unification #-} +{-# OPTIONS --safe --lossy-unification #-} {- The finitary Hilton–Milnor splitting