From 03244d0d71d48b7bb154ad13713316d4afb388ae Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Tue, 2 Apr 2024 05:55:35 +0200 Subject: [PATCH 01/30] small fix --- Cubical/Categories/Instances/Functors/Currying.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Cubical/Categories/Instances/Functors/Currying.agda b/Cubical/Categories/Instances/Functors/Currying.agda index 4a80b28976..9a34731e10 100644 --- a/Cubical/Categories/Instances/Functors/Currying.agda +++ b/Cubical/Categories/Instances/Functors/Currying.agda @@ -75,7 +75,7 @@ module _ (C : Category ℓC ℓC') (D : Category ℓD ℓD') where F-seq (λF⁻ a) _ (eG , cG) = cong₂ (seq' D) (F-seq (F-ob a _) _ _) (cong (flip N-ob _) (F-seq a _ _)) - ∙ AssocCong₂⋆R {C = D} _ + ∙ AssocCong₂⋆R D (N-hom ((F-hom a _) ●ᵛ (F-hom a _)) _ ∙ (⋆Assoc D _ _ _) ∙ cong (seq' D _) (sym (N-hom (F-hom a eG) cG))) @@ -85,7 +85,7 @@ module _ (C : Category ℓC ℓC') (D : Category ℓD ℓD') where N-ob (F-hom λF⁻Functor x) = uncurry (N-ob ∘→ N-ob x) N-hom ((F-hom λF⁻Functor) {F} {F'} x) {xx} {yy} = uncurry λ hh gg → - AssocCong₂⋆R {C = D} _ (cong N-ob (N-hom x hh) ≡$ _) + AssocCong₂⋆R D (cong N-ob (N-hom x hh) ≡$ _) ∙∙ cong (comp' D _) ((N-ob x (fst xx) .N-hom) gg) ∙∙ D .⋆Assoc _ _ _ From f9c7655951445ae7a6d277a372946c2b4c9b58c4 Mon Sep 17 00:00:00 2001 From: MJG Date: Fri, 12 Apr 2024 14:28:04 +0200 Subject: [PATCH 02/30] Normal form of FreeGroup (#1099) * nomral form of freegroup and Group solver tactic * some cleanup * group solver specialised to paths (aware of double-path-composition) * fixed name conflict * start with boilerplate * functor notation * free cats * write down a candidate def for free wild cat * notes * wip on grupoidSolver * removed solvers * removed solvers * added brief description of module * tweaking comments --------- Co-authored-by: Felix Cherubini --- Cubical/Algebra/Group/Free.agda | 507 +++++++++++++++++++++++ Cubical/Data/List/Base.agda | 28 +- Cubical/Data/List/Properties.agda | 193 ++++++++- Cubical/Data/Nat/Properties.agda | 6 + Cubical/Data/Sigma/Properties.agda | 4 + Cubical/Foundations/Function.agda | 8 +- Cubical/Foundations/Univalence.agda | 18 + Cubical/HITs/Bouquet/Discrete.agda | 129 ++++++ Cubical/HITs/FreeGroup/NormalForm.agda | 187 +++++++++ Cubical/Homotopy/Group/Base.agda | 13 + Cubical/Relation/Nullary/Properties.agda | 4 + 11 files changed, 1081 insertions(+), 16 deletions(-) create mode 100644 Cubical/Algebra/Group/Free.agda create mode 100644 Cubical/HITs/Bouquet/Discrete.agda create mode 100644 Cubical/HITs/FreeGroup/NormalForm.agda diff --git a/Cubical/Algebra/Group/Free.agda b/Cubical/Algebra/Group/Free.agda new file mode 100644 index 0000000000..99c26a9b41 --- /dev/null +++ b/Cubical/Algebra/Group/Free.agda @@ -0,0 +1,507 @@ +{-# OPTIONS --safe #-} + +{- + +This module introduces a way to represent elements of free groups as lists of pairs ([𝟚× A]), where A identifies generators and Bool differentiates a generator from its inverse. + + +The definition of `_⇊1g` encodes the concept of a word being equivalent to the identity element in a free group. It includes three constructors: + - The trivial case `[]` represents the identity itself. + - The `cj` constructor signifies conjugation by a generator, indicating that a word extended by a generator and its inverse is still equivalent to the identity. + - Lastly, the `_·_` constructor asserts that if two words are each equivalent to the identity, their concatenation will also be equivalent to the identity, preserving the identity property under the group operation. + +Then, a `_·_⁻¹≡ε relation`, defined as `λ xs ys → (xs ++ invLi ys) ⇊1g` captures the relationship where the concatenation of a word with the inverse of another equates to the group identity. + +By quotienting by _·_⁻¹≡ε, resulting in the List/↘↙group, a group structure on equivalence classes of lists is established. Here, concatenation acts as the group operation, hence forming the foundation for the free group constructed over A. + +Given that _·_⁻¹≡ε functions is equivalence relation, it permits the retrieval of this relationship from paths in the quotient. This feature facilitates the proof of the uniqueness of the normal form in `Cubical.HITs.FreeGroup.NormalForm`. + +In the Discrete module, the presence of decidable equality on the type of generators (A) enables the definition of groups without requiring truncation. This utility is used in `Cubical.HITs.Bouquet.Discrete` to demonstrate that a bouquet over a discrete type is an hGroupoid without truncation. + +-} + +module Cubical.Algebra.Group.Free where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Nat.Order.Recursive as OR +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Sigma +open import Cubical.Data.List as Li + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Functions.Involution + +open import Cubical.Algebra.Group + +open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _/₂_ ; [_] to [_]/) + + + +private + variable + ℓ : Level + +[𝟚×_] : (A : Type ℓ) → Type ℓ +[𝟚×_] = List ∘ (Bool ×_) + +module NormalForm (A : Type ℓ) where + + -- Below we are defining predicate `HasRedex` over `List (Bool × A)` + -- basing it on presence of sublist - REDucible EXpression of form [(b , x) , (not b , x)] + + not₁ : (Bool × A) → (Bool × A) + not₁ = map-fst not + + not₁not₁ : ∀ x → not₁ (not₁ x) ≡ x + not₁not₁ (x , y) = cong (_, y) (notnot x) + + IsRedex : Bool × A → Bool × A → Type ℓ + IsRedex x x' = x ≡ not₁ x' + + symIsRedex : ∀ x y → IsRedex x y → IsRedex y x + symIsRedex x y p = sym (not₁not₁ _) ∙ cong not₁ (sym p) + + WillReduce : Bool → A → (xs : [𝟚× A ]) → Type ℓ + WillReduce _ _ [] = ⊥* + WillReduce b x (h ∷ _) = IsRedex (b , x) h + + HeadIsRedex : [𝟚× A ] → Type ℓ + HeadIsRedex [] = ⊥* + HeadIsRedex ((b , a) ∷ xs) = WillReduce b a xs + + HasRedex : [𝟚× A ] → Type ℓ + HasRedex [] = ⊥* + HasRedex xs@(_ ∷ xs') = HeadIsRedex xs ⊎ HasRedex xs' + + HasRedex∷ʳ : ∀ {xs} {x} → HasRedex xs → HasRedex (xs ∷ʳ x) + HasRedex∷ʳ {x₂ ∷ xs} (inr x₁) = inr (HasRedex∷ʳ x₁) + HasRedex∷ʳ {x₂ ∷ x₃ ∷ xs} (inl x₁) = inl x₁ + + HasRedex++ : ∀ xs ys → HasRedex xs → HasRedex (xs ++ ys) + HasRedex++ (x₁ ∷ xs) ys (inr x) = inr (HasRedex++ xs ys x) + HasRedex++ (x₁ ∷ x₂ ∷ xs) ys (inl x) = inl x + + ++HasRedex : ∀ xs ys → HasRedex ys → HasRedex (xs ++ ys) + ++HasRedex [] ys x = x + ++HasRedex (x₁ ∷ xs) ys x = + inr (++HasRedex xs ys x) + + -- We use absence of redexes as property of representing normalised term + + IsNormalised : [𝟚× A ] → Type ℓ + IsNormalised xs = (¬ HasRedex xs) + + isPropIsNormalised : ∀ xs → isProp (IsNormalised xs) + isPropIsNormalised xs = isProp¬ _ + + + IsNormalised[] : IsNormalised [] + IsNormalised[] = lower + + IsNormalised[x] : ∀ x → IsNormalised [ x ] + IsNormalised[x] _ = ⊎.rec lower lower + + IsNormalisedTail : ∀ xs → IsNormalised xs → IsNormalised (tail xs) + IsNormalisedTail [] n = n + IsNormalisedTail (x ∷ xs) = _∘ inr + + ¬IsRedex→IsNormalisedPair : + ∀ {x x'} → ¬ IsRedex x x' → IsNormalised (x ∷ [ x' ]) + ¬IsRedex→IsNormalisedPair {x' = x'} ¬ir = ⊎.rec ¬ir (IsNormalised[x] x') + + invLi : [𝟚× A ] → [𝟚× A ] + invLi = rev ∘ Li.map (map-fst not) + + invLi++ : ∀ xs ys → invLi (xs ++ ys) ≡ + invLi ys ++ invLi xs + invLi++ xs ys = + sym (cong rev (map++ _ xs ys)) ∙ + rev-++ (Li.map _ xs) (Li.map _ ys) + + invol-invLi : isInvolution invLi + invol-invLi xs = + sym (rev-map-comm (map-fst not) (invLi xs)) ∙ + cong (Li.map (map-fst not)) + (rev-rev (Li.map (map-fst not) xs)) + ∙ map-∘ (map-fst not) (map-fst not) xs ∙ + (λ i → Li.map (map-fst (λ x → notnot x i) ) xs) ∙ map-id xs + + HasRedexInvLi : ∀ {xs} → HasRedex xs → HasRedex (invLi xs) + HasRedexInvLi {_ ∷ []} x = x + HasRedexInvLi {x₁ ∷ x₂ ∷ xs} = ⊎.rec + (subst HasRedex (sym (++-assoc (invLi xs) _ _)) + ∘S ++HasRedex (invLi xs) (_ ∷ _) + ∘S inl ∘S cong not₁ ∘S symIsRedex _ _ ) + (HasRedex∷ʳ ∘S HasRedexInvLi) + + + IsNormalisedInvLi : ∀ {xs} → IsNormalised xs ≃ IsNormalised (invLi xs) + IsNormalisedInvLi = propBiimpl→Equiv (isPropIsNormalised _) (isPropIsNormalised _) + (_∘S subst HasRedex (invol-invLi _) ∘S HasRedexInvLi) + (_∘S HasRedexInvLi) + + HasRedexSplit++ : ∀ {xs ys} → HasRedex (xs ++ ys) → + HasRedex (take 1 (rev xs) ++ take 1 ys) ⊎ + (HasRedex xs ⊎ HasRedex ys) + HasRedexSplit++ {[]} = inr ∘ inr + HasRedexSplit++ {x ∷ []} {[]} r = ⊥.rec (IsNormalised[x] x r) + HasRedexSplit++ {x ∷ []} {x₁ ∷ ys} = ⊎.rec (inl ∘ inl) (inr ∘ inr) + HasRedexSplit++ {x ∷ x₁ ∷ xs} {ys} = + ⊎.rec (inr ∘ inl ∘ inl) + (⊎.rec (inl ∘ subst (λ zz → HasRedex (zz ++ take 1 ys)) + (w _ (subst (0 <_) (+-comm 1 _ ∙ sym (length++ (rev xs) _)) _))) + (⊎.rec (inr ∘ inl ∘ inr) (inr ∘ inr) ) ∘S HasRedexSplit++ {x₁ ∷ xs} {ys}) + where + w : ∀ xs → 0 < length xs → take 1 xs ≡ take 1 (xs ++ [ x ]) + w (x ∷ xs) _ = refl + + -- This predicate, is encoding fact that given list reduces to [], + -- meaning that if interpreted as term of free group it would be equall to 1g + + infixl 10 _⇊1g + + data _⇊1g : [𝟚× A ] → Type ℓ where + [] : [] ⇊1g + cj : ∀ x → ∀ xs → xs ⇊1g → (x ∷ (xs ∷ʳ not₁ x) ) ⇊1g + _·_ : ∀ xs ys → xs ⇊1g → ys ⇊1g → (xs ++ ys) ⇊1g + + _r·_ : ∀ {xs ys} → xs ⇊1g → ys ⇊1g → (xs ++ ys) ⇊1g + _r·_ = _·_ _ _ + + ¬⇊1g[len≡1] : ∀ xs → length xs ≡ 1 → ¬ xs ⇊1g + ¬⇊1g[len≡1] .[] x [] = znots x + ¬⇊1g[len≡1] .(_ ∷ (_ ∷ʳ _)) x (cj _ xs _) = + snotz ((+-comm 1 _ ∙ sym (length++ xs _)) ∙ injSuc x) + ¬⇊1g[len≡1] .(xs ++ ys) x ((xs · ys) x₁ x₂) = + ⊎.rec (flip (¬⇊1g[len≡1] ys) x₂ ∘ snd) + ((flip (¬⇊1g[len≡1] xs) x₁ ∘ fst)) + (m+n≡1→m≡0×n≡1⊎m≡1n≡0 {length xs} {length ys} (sym (length++ xs ys) ∙ x)) + + ⇊1gWillReduceView : ∀ b a ys → ys ⇊1g → WillReduce b a ys → + Σ ((Σ _ _⇊1g) × (Σ _ _⇊1g)) + λ ((rl , _) , (rr , _)) → + rl ++ [ b , a ] ++ rr ≡ tail ys + ⇊1gWillReduceView b a .(x ∷ (xs ∷ʳ _)) (cj x xs x₃) x₁ = + ((_ , x₃) , (_ , [])) , cong (xs ∷ʳ_) x₁ + ⇊1gWillReduceView b a .([] ++ ys) (([] · ys) x x₂) x₁ = + ⇊1gWillReduceView b a ys x₂ x₁ + ⇊1gWillReduceView b a .((_ ∷ _) ++ ys) ((xs@(_ ∷ _) · ys) x x₂) x₁ = + let (((rl , rlR) , (rr , rrR)) , p) = ⇊1gWillReduceView b a xs x x₁ + in ((_ , rlR) , (_ , (_ · _) rrR x₂)) , + sym (++-assoc rl _ _) ∙ cong (_++ ys) p + + ⇊1g⇒HasRedex : ∀ xs → 0 < length xs → xs ⇊1g → HasRedex xs + ⇊1g⇒HasRedex .(x₁ ∷ ([] ∷ʳ not₁ x₁)) x (cj x₁ [] x₂) = + inl (symIsRedex _ _ refl) + ⇊1g⇒HasRedex .(x₁ ∷ ((x₃ ∷ xs) ∷ʳ not₁ x₁)) x (cj x₁ (x₃ ∷ xs) x₂) = + inr (HasRedex∷ʳ (⇊1g⇒HasRedex _ _ x₂)) + ⇊1g⇒HasRedex .([] ++ ys) x (([] · ys) x₁ x₂) = ⇊1g⇒HasRedex _ x x₂ + ⇊1g⇒HasRedex .((x₃ ∷ xs) ++ ys) x (((x₃ ∷ xs) · ys) x₁ x₂) = + HasRedex++ _ _ (⇊1g⇒HasRedex _ _ x₁) + + isNormalised⇊1g : ∀ xs → IsNormalised xs → xs ⇊1g → xs ≡ [] + isNormalised⇊1g [] _ _ = refl + isNormalised⇊1g (x₂ ∷ xs) x x₁ = ⊥.rec (x (⇊1g⇒HasRedex _ _ x₁)) + + + ⇊1g-invLi : ∀ {xs} → xs ⇊1g → (invLi xs) ⇊1g + ⇊1g-invLi [] = [] + ⇊1g-invLi (cj x xs x₁) = + subst _⇊1g (cong (_∷ invLi (x ∷ xs)) (sym (not₁not₁ _) ) + ∙ cong (_∷ʳ _) (sym (invLi++ xs [ not₁ x ]))) (cj x _ (⇊1g-invLi x₁)) + ⇊1g-invLi ((xs · ys) x x₁) = + subst _⇊1g (sym (invLi++ xs ys)) (⇊1g-invLi x₁ r· ⇊1g-invLi x) + + ⇊1gRot : ∀ xs → xs ⇊1g → _⇊1g (rot xs) + ⇊1gRot [] x = [] + ⇊1gRot xs@(x'@(b , a) ∷ xs') x = + let (((rl , rlR) , (rr , rrR)) , p) = ⇊1gWillReduceView (not b) a xs x refl + in subst _⇊1g ((λ i → (++-assoc rl ([ not b , a ] ++ rr) + [ not₁not₁ x' i ]) (~ i)) ∙ cong (_∷ʳ x') p) + (rlR r· cj (not b , a) _ rrR) + + ⇊1g++comm : ∀ xs ys → (xs ++ ys) ⇊1g → (ys ++ xs) ⇊1g + ⇊1g++comm [] ys = subst _⇊1g (sym (++-unit-r ys)) + ⇊1g++comm (x₁ ∷ xs) ys x = + subst _⇊1g (++-assoc ys _ _) + (⇊1g++comm xs _ (subst _⇊1g (++-assoc xs _ _) (⇊1gRot _ x))) + + pop⇊1gHead : ∀ {xs} → HeadIsRedex xs → xs ⇊1g → _⇊1g (tail (tail xs)) + pop⇊1gHead x (cj x₁ [] r) = [] + pop⇊1gHead x (cj x₁ (x₂ ∷ xs) r) = + subst (_⇊1g ∘ (xs ∷ʳ_)) (symIsRedex _ _ x) (⇊1gRot _ r) + pop⇊1gHead x (([] · ys) r r₁) = pop⇊1gHead x r₁ + pop⇊1gHead x (((x₁ ∷ []) · ys) r r₁) = ⊥.rec (¬⇊1g[len≡1] [ x₁ ] refl r) + pop⇊1gHead x (((_ ∷ _ ∷ _) · ys) r r₁) = pop⇊1gHead x r r· r₁ + + ⇊1gCJ : ∀ xs ys → _⇊1g (ys ++ xs ++ invLi ys) → xs ⇊1g + ⇊1gCJ xs [] = subst _⇊1g (++-unit-r xs) + ⇊1gCJ xs (x₁ ∷ ys) x = + ⇊1gCJ xs ys (pop⇊1gHead refl + (⇊1g++comm (x₁ ∷ ys ++ xs ++ invLi ys) [ not₁ x₁ ] + (subst _⇊1g (cong (x₁ ∷_) (cong (ys ++_) (sym (++-assoc xs _ _)) + ∙ sym (++-assoc ys _ _))) x))) + + nf-uR : ∀ xs ys + → IsNormalised (invLi xs) + → IsNormalised ys + → (invLi xs ++ ys) ⇊1g → xs ≡ ys + nf-uR xs [] nXs x₁ r = sym (invol-invLi xs) ∙ cong invLi + (isNormalised⇊1g _ nXs (subst _⇊1g (++-unit-r _) r)) + nf-uR [] (x₃ ∷ ys) x x₁ x₂ = ⊥.rec (x₁ (⇊1g⇒HasRedex _ _ x₂)) + nf-uR xs@(_ ∷ _) (x₃ ∷ ys) nXs nYs r = + let ww = subst _⇊1g (cong (x₃ ∷_) (sym (++-assoc ys _ _))) + (⇊1g++comm (invLi xs) _ r) + www = subst (0 <_) + (sym (+-suc _ _) + ∙ sym (length++ (invLi xs) _)) _ + in (⊎.rec (⊎.rec (λ p → cong₂ _∷_ + (sym (not₁not₁ _) ∙ sym (symIsRedex _ _ p)) + (nf-uR (tail xs) _ (fst IsNormalisedInvLi + (invEq (IsNormalisedInvLi {xs}) nXs ∘ inr) ) (nYs ∘ inr) + (⇊1g++comm _ (invLi (tail xs)) + (pop⇊1gHead p (⇊1g++comm _ [ _ ] ww))))) + (⊥.rec ∘ IsNormalised[x] x₃) ∘S + subst HasRedex (cong ((_++ _) ∘ take 1) + (rev-rev (Li.map not₁ xs)))) + (⊥.rec ∘ ⊎.rec nXs nYs) + ∘S HasRedexSplit++ {invLi xs} + ∘S ⇊1g⇒HasRedex _ www) r + + infixr 5 _·_⁻¹≡ε + + -- We are defining this relation using record + -- to help later definitions infer implicit arguments. + -- With the help of earlier lemmas, we can prove that it is EquivRel, + -- wich lets us recover witness of this relation from path in SetQuotent + + record _·_⁻¹≡ε (xs ys : _) : Type ℓ where + constructor [_]≡ε + field + ≡ε : (xs ++ invLi ys) ⇊1g + + open _·_⁻¹≡ε public + + open BinaryRelation + open isEquivRel + + ·⁻¹≡ε-refl : isRefl _·_⁻¹≡ε + ·⁻¹≡ε-refl [] = [ [] ]≡ε + ≡ε (·⁻¹≡ε-refl (x ∷ xs)) = + subst _⇊1g (sym (++-assoc [ x ] xs (invLi (x ∷ xs)) ∙ + cong (x ∷_) (sym (++-assoc xs (invLi xs) _)))) + (cj x _ (≡ε (·⁻¹≡ε-refl xs))) + + ·⁻¹≡ε-sym : isSym _·_⁻¹≡ε + ≡ε (·⁻¹≡ε-sym a b [ p ]≡ε) = + subst _⇊1g (invLi++ a (invLi b) ∙ + cong (_++ invLi a) (invol-invLi b)) (⇊1g-invLi p) + + ·⁻¹≡ε-trans : isTrans _·_⁻¹≡ε + ≡ε (·⁻¹≡ε-trans xs ys zs [ p ]≡ε [ q ]≡ε) = + ⇊1g++comm (invLi zs) xs + (⇊1gCJ (invLi zs ++ xs) ys + (subst _⇊1g (++-assoc ys _ _ ∙ + cong (ys ++_) (sym (++-assoc (invLi zs) _ _))) (q r· p))) + + ·⁻¹≡ε-isEquivRel : isEquivRel _·_⁻¹≡ε + reflexive ·⁻¹≡ε-isEquivRel = ·⁻¹≡ε-refl + symmetric ·⁻¹≡ε-isEquivRel = ·⁻¹≡ε-sym + transitive ·⁻¹≡ε-isEquivRel = ·⁻¹≡ε-trans + + ≡→red : ∀ a b → Iso ([ a ]/ ≡ [ b ]/) ∥ a · b ⁻¹≡ε ∥₁ + ≡→red = isEquivRel→TruncIso ·⁻¹≡ε-isEquivRel + + + open Iso + + + _↘↙_ = _·_⁻¹≡ε + + List/↘↙ : Type ℓ + List/↘↙ = _ /₂ _·_⁻¹≡ε + + + -- concatenation of lists, respects `_·_⁻¹≡ε` on both sides, + -- so we can give SetQuotiening by _·_⁻¹≡ε structure of a group + -- by taking concatenation as group operation + + _↘↙++↘↙_ : ∀ {xsL xsR ysL ysR} → + xsL ↘↙ ysL → xsR ↘↙ ysR → + (xsL ++ xsR) ↘↙ (ysL ++ ysR) + ≡ε (_↘↙++↘↙_ {xsL} {xsR} {ysL} {ysR} [ p ]≡ε [ q ]≡ε) = + subst _⇊1g (sym (++-assoc xsL _ _)) + (⇊1g++comm _ xsL + (subst _⇊1g (++-assoc xsR _ _ ∙∙ + (λ i → xsR ++ ++-assoc (invLi ysR) (invLi ysL) xsL (~ i)) ∙∙ + ( λ i → ++-assoc xsR (invLi++ ysL ysR (~ i)) xsL (~ i))) + (q r· ⇊1g++comm xsL _ p))) + + List/↘↙· : List/↘↙ → List/↘↙ → List/↘↙ + List/↘↙· = SQ.rec2 squash/ (λ a b → [ a ++ b ]/) + (λ _ _ c → eq/ _ _ ∘ _↘↙++↘↙ (·⁻¹≡ε-refl c)) + (λ a _ _ → eq/ _ _ ∘ (·⁻¹≡ε-refl a) ↘↙++↘↙_ ) + + List/↘↙GroupStr : GroupStr List/↘↙ + GroupStr.1g List/↘↙GroupStr = [ [] ]/ + GroupStr._·_ List/↘↙GroupStr = List/↘↙· + GroupStr.inv List/↘↙GroupStr = + SQ.rec squash/ ([_]/ ∘ invLi) + λ xs ys → sym ∘S eq/ _ _ ∘S [_]≡ε + ∘S subst (_⇊1g ∘ (invLi ys ++_)) (sym (invol-invLi xs)) + ∘S ⇊1g++comm xs (invLi ys) ∘S ≡ε + + GroupStr.isGroup List/↘↙GroupStr = makeIsGroup squash/ + (SQ.elimProp3 (λ _ _ _ → squash/ _ _) + λ xs _ _ → cong [_]/ (sym (++-assoc xs _ _))) + (SQ.elimProp (λ _ → squash/ _ _) λ xs → cong [_]/ (++-unit-r xs)) + (SQ.elimProp (λ _ → squash/ _ _) λ _ → refl) + (SQ.elimProp (λ _ → squash/ _ _) λ xs → sym (eq/ _ _ + ([ ⇊1g-invLi (≡ε (·⁻¹≡ε-refl xs)) ]≡ε))) + (SQ.elimProp (λ _ → squash/ _ _) λ xs → eq/ _ _ [ + subst _⇊1g (cong (invLi xs ++_) (invol-invLi xs) ∙ + sym (++-unit-r _)) (≡ε (·⁻¹≡ε-refl (invLi xs))) ]≡ε) + + List/↘↙group : Group ℓ + List/↘↙group = _ , List/↘↙GroupStr + + + + module Discrete (_≟_ : Discrete A) where + + -- For discrete set of generators, we can compute normal form, + -- and proove `Iso (Σ _ IsNormalised) List/↘↙` + + isSetA = Discrete→isSet _≟_ + + isSet[𝟚×] = isOfHLevelList 0 (isSet× isSetBool isSetA) + + IsRedex? : ∀ x x' → Dec (IsRedex x x') + IsRedex? _ _ = discreteΣ 𝟚._≟_ (λ _ → _≟_) _ _ + + HeadIsRedex? : ∀ xs → Dec (HeadIsRedex xs) + HeadIsRedex? [] = no lower + HeadIsRedex? (x ∷ []) = no lower + HeadIsRedex? (x ∷ x' ∷ _) = IsRedex? x x' + + preη· : ∀ x xs → Dec (HeadIsRedex (x ∷ xs)) → [𝟚× A ] + preη· _ xs (yes _) = tail xs + preη· x xs (no _) = x ∷ xs + + preη·-N : ∀ {x xs} hir? → IsNormalised xs → IsNormalised (preη· x xs hir?) + preη·-N (yes _) = IsNormalisedTail _ + preη·-N (no ¬p) = ⊎.rec ¬p + + sec-preη· : ∀ x xs p q → IsNormalised xs → preη· (not₁ x) (preη· x xs p) q ≡ xs + sec-preη· x (x₂ ∷ xs) (yes p) (no ¬p) x₁ = + cong (_∷ xs) (sym (symIsRedex _ _ p)) + sec-preη· x (x₂ ∷ x₃ ∷ xs) (yes p) (yes p₁) x₁ = + ⊥.rec (x₁ (inl (symIsRedex _ _ p ∙ p₁))) + sec-preη· x xs (no ¬p) (no ¬p₁) x₁ = ⊥.rec (¬p₁ refl) + sec-preη· x xs (no ¬p) (yes p) _ = refl + + η· : (Bool × A) → [𝟚× A ] → [𝟚× A ] + η· x xs = preη· _ _ (HeadIsRedex? (x ∷ xs)) + + η·∷ : ∀ x xs → (HeadIsRedex (x ∷ xs) → ⊥) → η· x xs ≡ x ∷ xs + η·∷ x xs x₁ = cong (λ u → preη· x xs u) + (≡no (HeadIsRedex? (x ∷ xs)) x₁) + + nη· : (Bool × A) → Σ _ IsNormalised → Σ _ IsNormalised + fst (nη· x x₁) = η· x (fst x₁) + snd (nη· x x₁) = preη·-N (HeadIsRedex? _) (snd x₁) + + + η·iso : (Bool × A) → Iso (Σ [𝟚× A ] IsNormalised) (Σ [𝟚× A ] IsNormalised) + Iso.fun (η·iso x) = nη· x + Iso.inv (η·iso x) = nη· (not₁ x) + Iso.rightInv (η·iso x) b = + Σ≡Prop isPropIsNormalised + (funExt⁻ (cong η· (sym (not₁not₁ x)) ) (η· (not₁ x) (fst b)) + ∙ sec-preη· (not₁ x) _ (HeadIsRedex? _) (HeadIsRedex? _) (snd b)) + Iso.leftInv (η·iso x) a = + Σ≡Prop isPropIsNormalised + (sec-preη· x _ (HeadIsRedex? _) (HeadIsRedex? _) (snd a)) + + η·≃ = isoToEquiv ∘ η·iso + + normalise : ∀ xs → Σ _ λ xs' → + (xs' · xs ⁻¹≡ε) × IsNormalised xs' + normalise = Li.elim ([] , [ [] ]≡ε , lower ) + λ {x} {xs} (xs' , [ u ]≡ε , v) → + let zz : ∀ xs' uu u → (preη· x xs' uu ++ invLi (x ∷ xs)) ⇊1g + zz = + λ { xs' (no ¬p) → subst (_⇊1g ∘S (x ∷_)) (++-assoc xs' _ _) ∘ cj x _ + ; [] (yes ()) + ; (_ ∷ xs') (yes p) → + subst _⇊1g (λ i → ++-assoc xs' (invLi xs) + [ symIsRedex _ _ p i ] i) ∘ ⇊1gRot _ } + h = HeadIsRedex? _ + in _ , [ zz xs' h u ]≡ε , preη·-N h v + + IsoNF : Iso (Σ _ IsNormalised) List/↘↙ + fun IsoNF = [_]/ ∘ fst + Iso.inv IsoNF = + SQ.rec (isSetΣ isSet[𝟚×] (isProp→isSet ∘ isPropIsNormalised)) + ((λ (_ , _ , u) → _ , u) ∘ normalise) + λ _ _ → + let (a' , t , u ) = normalise _ ; (b' , t' , u') = normalise _ + in Σ≡Prop isPropIsNormalised ∘S sym + ∘S nf-uR _ _ (fst (IsNormalisedInvLi {b'}) u') u + ∘S ⇊1g++comm a' (invLi b') ∘S ≡ε + ∘S flip (·⁻¹≡ε-trans _ _ _) (·⁻¹≡ε-sym _ _ t') + ∘S ·⁻¹≡ε-trans _ _ _ t + rightInv IsoNF = SQ.elimProp (λ _ → squash/ _ _) + (eq/ _ _ ∘ fst ∘ snd ∘ normalise) + leftInv IsoNF = Σ≡Prop isPropIsNormalised ∘ uncurry + (Li.elim (λ _ → refl) λ f v → + let lem : ∀ uu → preη· _ _ uu ≡ _ ∷ _ + lem = + λ { (yes p) → ⊥.rec (v (inl (subst (WillReduce _ _) (f (v ∘ inr)) p))) + ; (no ¬p) → refl } + in lem (HeadIsRedex? _) ∙ cong (_ ∷_) (f (v ∘ inr))) + +module NF {ℓ'} {A : Type ℓ} (G : Group ℓ') (η : A → ⟨ G ⟩) where + open NormalForm A + + open GroupStr (snd G) renaming (_·_ to _·fg_) public + + η* : Bool × A → ⟨ G ⟩ + η* (b , a) = (if b then idfun _ else inv) (η a) + + fromList : [𝟚× A ] → ⟨ G ⟩ + fromList = foldr (_·fg_ ∘ η*) 1g + + record NF (g : _) : Type (ℓ-max ℓ ℓ') where + constructor _nf_,_ + field + word : [𝟚× A ] + fromListWord≡ : fromList word ≡ g + isNormalisedWord : IsNormalised word + + + fromList· : ∀ xs ys → fromList (xs ++ ys) ≡ + fromList xs ·fg fromList ys + fromList· [] _ = sym (·IdL _) + fromList· (_ ∷ xs) _ = + cong (_ ·fg_) (fromList· xs _) ∙ + ·Assoc _ _ _ + + redex-ε-η* : ∀ x x' → IsRedex x x' → η* x ·fg η* x' ≡ 1g + redex-ε-η* (false , _) (false , _) = ⊥.rec ∘ false≢true ∘ cong fst + redex-ε-η* (false , x) (true , _) q = + cong (inv (η x) ·fg_) (cong (η) (sym (cong snd q))) ∙ ·InvL (η x) + redex-ε-η* (true , x) (false , _) q = + cong (η x ·fg_) (cong (inv ∘ η) (sym (cong snd q))) ∙ ·InvR (η x) + redex-ε-η* (true , _) (true , _) = ⊥.rec ∘ true≢false ∘ cong fst diff --git a/Cubical/Data/List/Base.agda b/Cubical/Data/List/Base.agda index 0222265954..a78ac00455 100644 --- a/Cubical/Data/List/Base.agda +++ b/Cubical/Data/List/Base.agda @@ -3,8 +3,8 @@ 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 +open import Cubical.Data.Maybe.Base as Maybe hiding (rec ; elim) +open import Cubical.Data.Nat.Base hiding (elim) module _ {ℓ} {A : Type ℓ} where @@ -50,3 +50,27 @@ module _ {ℓ} {A : Type ℓ} where 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 + + drop : ℕ → List A → List A + drop zero xs = xs + drop (suc n) [] = [] + drop (suc n) (x ∷ xs) = drop n xs + + take : ℕ → List A → List A + take zero xs = [] + take (suc n) [] = [] + take (suc n) (x ∷ xs) = x ∷ take n xs + + elim : ∀ {ℓ'} {B : List A → Type ℓ'} + → B [] + → (∀ {a l} → B l → B (a ∷ l)) + → ∀ l → B l + elim b _ [] = b + elim {B = B} b[] b (a ∷ l) = b (elim {B = B} b[] b l ) + + rec : ∀ {ℓ'} {B : Type ℓ'} + → B + → (A → B → B) + → List A → B + rec b _ [] = b + rec b f (x ∷ xs) = f x (rec b f xs) diff --git a/Cubical/Data/List/Properties.agda b/Cubical/Data/List/Properties.agda index c44f9db1df..7b755944f5 100644 --- a/Cubical/Data/List/Properties.agda +++ b/Cubical/Data/List/Properties.agda @@ -6,13 +6,16 @@ open import Cubical.Core.Everything open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.HLevels open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat open import Cubical.Data.Sigma +open import Cubical.Data.Sum as ⊎ hiding (map) open import Cubical.Data.Unit open import Cubical.Relation.Nullary -open import Cubical.Data.List.Base +open import Cubical.Data.List.Base as List module _ {ℓ} {A : Type ℓ} where @@ -117,8 +120,11 @@ isOfHLevelList n ofLevel xs ys = private variable - ℓ : Level + ℓ ℓ' : Level A : Type ℓ + B : Type ℓ' + x y : A + xs ys : List A caseList : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (n c : B) → List A → B caseList n _ [] = n @@ -132,23 +138,32 @@ private safe-tail [] = [] safe-tail (_ ∷ xs) = xs -cons-inj₁ : ∀ {x y : A} {xs ys} → x ∷ xs ≡ y ∷ ys → x ≡ y +cons-inj₁ : 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₂ : 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 []) +snoc-inj₂ : xs ∷ʳ x ≡ ys ∷ʳ y → x ≡ y +snoc-inj₂ {xs = xs} {ys = ys} p = + cons-inj₁ ((sym (rev-++ xs _)) ∙∙ cong rev p ∙∙ (rev-++ ys _)) -¬nil≡cons : ∀ {x : A} {xs} → ¬ ([] ≡ x ∷ xs) -¬nil≡cons {A = A} p = lower (subst (caseList (List A) (Lift ⊥)) p []) +snoc-inj₁ : xs ∷ʳ x ≡ ys ∷ʳ y → xs ≡ ys +snoc-inj₁ {xs = xs} {ys = ys} p = + sym (rev-rev _) ∙∙ cong rev (cons-inj₂ ((sym (rev-++ xs _)) ∙∙ cong rev p ∙∙ (rev-++ ys _))) + ∙∙ rev-rev _ -¬snoc≡nil : ∀ {x : A} {xs} → ¬ (xs ∷ʳ x ≡ []) +¬cons≡nil : ¬ (x ∷ xs ≡ []) +¬cons≡nil {_} {A} p = lower (subst (caseList (Lift ⊥) (List A)) p []) + +¬nil≡cons : ¬ ([] ≡ x ∷ xs) +¬nil≡cons {_} {A} p = lower (subst (caseList (List A) (Lift ⊥)) p []) + +¬snoc≡nil : ¬ (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 : ¬ ([] ≡ xs ∷ʳ x) ¬nil≡snoc contra = ¬snoc≡nil (sym contra) cons≡rev-snoc : (x : A) → (xs : List A) → x ∷ rev xs ≡ rev (xs ∷ʳ x) @@ -158,7 +173,7 @@ 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≡[] : isProp (xs ≡ []) isPropXs≡[] {xs = []} = isOfHLevelSuc 0 isContr[]≡[] isPropXs≡[] {xs = x ∷ xs} = λ p _ → ⊥.rec (¬cons≡nil p) @@ -175,7 +190,161 @@ foldrCons : (xs : List A) → foldr _∷_ [] xs ≡ xs foldrCons [] = refl foldrCons (x ∷ xs) = cong (x ∷_) (foldrCons xs) -length-map : ∀ {ℓA ℓB} {A : Type ℓA} {B : Type ℓB} → (f : A → B) → (as : List A) +length-map : (f : A → B) → (as : List A) → length (map f as) ≡ length as length-map f [] = refl length-map f (a ∷ as) = cong suc (length-map f as) + +map++ : (f : A → B) → (as bs : List A) + → map f as ++ map f bs ≡ map f (as ++ bs) +map++ f [] bs = refl +map++ f (x ∷ as) bs = cong (f x ∷_) (map++ f as bs) + +rev-map-comm : (f : A → B) → (as : List A) + → map f (rev as) ≡ rev (map f as) +rev-map-comm f [] = refl +rev-map-comm f (x ∷ as) = + sym (map++ f (rev as) _) ∙ cong (_++ [ f x ]) (rev-map-comm f as) + +length++ : (xs ys : List A) → length (xs ++ ys) ≡ length xs + length ys +length++ [] ys = refl +length++ (x ∷ xs) ys = cong suc (length++ xs ys) + +drop++ : ∀ (xs ys : List A) k → + drop (length xs + k) (xs ++ ys) ≡ drop k ys +drop++ [] ys k = refl +drop++ (x ∷ xs) ys k = drop++ xs ys k + +dropLength++ : (xs : List A) → drop (length xs) (xs ++ ys) ≡ ys +dropLength++ {ys = ys} xs = + cong (flip drop (xs ++ ys)) (sym (+-zero (length xs))) ∙ drop++ xs ys 0 + + +dropLength : (xs : List A) → drop (length xs) xs ≡ [] +dropLength xs = + cong (drop (length xs)) (sym (++-unit-r xs)) + ∙ dropLength++ xs + + +take++ : ∀ (xs ys : List A) k → take (length xs + k) (xs ++ ys) ≡ xs ++ take k ys +take++ [] ys k = refl +take++ (x ∷ xs) ys k = cong (_ ∷_) (take++ _ _ k) + + +takeLength++ : ∀ ys → take (length xs) (xs ++ ys) ≡ xs +takeLength++ {xs = xs} ys = + cong (flip take (xs ++ ys)) (sym (+-zero (length xs))) + ∙∙ take++ xs ys 0 + ∙∙ ++-unit-r xs + +takeLength : take (length xs) xs ≡ xs +takeLength = cong (take _) (sym (++-unit-r _)) ∙ takeLength++ [] + +map-∘ : ∀ {ℓA ℓB ℓC} {A : Type ℓA} {B : Type ℓB} {C : Type ℓC} + (g : B → C) (f : A → B) (as : List A) + → map g (map f as) ≡ map (λ x → g (f x)) as +map-∘ g f [] = refl +map-∘ g f (x ∷ as) = cong (_ ∷_) (map-∘ g f as) + +map-id : (as : List A) → map (λ x → x) as ≡ as +map-id [] = refl +map-id (x ∷ as) = cong (_ ∷_) (map-id as) + +length≡0→≡[] : ∀ (xs : List A) → length xs ≡ 0 → xs ≡ [] +length≡0→≡[] [] x = refl +length≡0→≡[] (x₁ ∷ xs) x = ⊥.rec (snotz x) + +init : List A → List A +init [] = [] +init (x ∷ []) = [] +init (x ∷ xs@(_ ∷ _)) = x ∷ init xs + +tail : List A → List A +tail [] = [] +tail (x ∷ xs) = xs + +init-red-lem : ∀ (x : A) xs → ¬ (xs ≡ []) → (x ∷ init xs) ≡ (init (x ∷ xs)) +init-red-lem x [] x₁ = ⊥.rec (x₁ refl) +init-red-lem x (x₂ ∷ xs) x₁ = refl + +init∷ʳ : init (xs ∷ʳ x) ≡ xs +init∷ʳ {xs = []} = refl +init∷ʳ {xs = _ ∷ []} = refl +init∷ʳ {xs = _ ∷ _ ∷ _} = cong (_ ∷_) init∷ʳ + +tail∷ʳ : tail (xs ∷ʳ y) ∷ʳ x ≡ tail (xs ∷ʳ y ∷ʳ x) +tail∷ʳ {xs = []} = refl +tail∷ʳ {xs = x ∷ xs} = refl + +init-rev-tail : rev (init xs) ≡ tail (rev xs) +init-rev-tail {xs = []} = refl +init-rev-tail {xs = x ∷ []} = refl +init-rev-tail {xs = x ∷ y ∷ xs} = + cong (_∷ʳ x) (init-rev-tail {xs = y ∷ xs}) + ∙ tail∷ʳ {xs = rev xs} + +init++ : ∀ xs → xs ++ init (x ∷ ys) ≡ init (xs ++ x ∷ ys) +init++ [] = refl +init++ (_ ∷ []) = refl +init++ (_ ∷ _ ∷ _) = cong (_ ∷_) (init++ (_ ∷ _)) + +Split++ : (xs ys xs' ys' zs : List A) → Type _ +Split++ xs ys xs' ys' zs = ((xs ++ zs ≡ xs') × (ys ≡ zs ++ ys')) + +split++ : ∀ (xs' ys' xs ys : List A) → xs' ++ ys' ≡ xs ++ ys → + Σ _ λ zs → + ((Split++ xs' ys' xs ys zs) + ⊎ (Split++ xs ys xs' ys' zs)) +split++ [] ys' xs ys x = xs , inl (refl , x) +split++ xs'@(_ ∷ _) ys' [] ys x = xs' , inr (refl , sym x) +split++ (x₁ ∷ xs') ys' (x₂ ∷ xs) ys x = + let (zs , q) = split++ xs' ys' xs ys (cons-inj₂ x) + p = cons-inj₁ x + in zs , ⊎.map (map-fst (λ q i → p i ∷ q i)) + (map-fst (λ q i → p (~ i) ∷ q i)) q + +rot : List A → List A +rot [] = [] +rot (x ∷ xs) = xs ∷ʳ x + +take[] : ∀ n → take {A = A} n [] ≡ [] +take[] zero = refl +take[] (suc n) = refl + +drop[] : ∀ n → drop {A = A} n [] ≡ [] +drop[] zero = refl +drop[] (suc n) = refl + +lookupAlways : A → List A → ℕ → A +lookupAlways a [] _ = a +lookupAlways _ (x ∷ _) zero = x +lookupAlways a (x ∷ xs) (suc k) = lookupAlways a xs k + +module List₂ where + open import Cubical.HITs.SetTruncation renaming + (rec to rec₂ ; map to map₂ ; elim to elim₂ ) + + ∥List∥₂→List∥∥₂ : ∥ List A ∥₂ → List ∥ A ∥₂ + ∥List∥₂→List∥∥₂ = rec₂ (isOfHLevelList 0 squash₂) (map ∣_∣₂) + + List∥∥₂→∥List∥₂ : List ∥ A ∥₂ → ∥ List A ∥₂ + List∥∥₂→∥List∥₂ [] = ∣ [] ∣₂ + List∥∥₂→∥List∥₂ (x ∷ xs) = + rec2 squash₂ (λ x xs → ∣ x ∷ xs ∣₂) x (List∥∥₂→∥List∥₂ xs) + + Iso∥List∥₂List∥∥₂ : Iso (List ∥ A ∥₂) ∥ List A ∥₂ + Iso.fun Iso∥List∥₂List∥∥₂ = List∥∥₂→∥List∥₂ + Iso.inv Iso∥List∥₂List∥∥₂ = ∥List∥₂→List∥∥₂ + Iso.rightInv Iso∥List∥₂List∥∥₂ = + elim₂ (isProp→isSet ∘ λ _ → squash₂ _ _) + (List.elim refl (cong (rec2 squash₂ (λ x₁ xs → ∣ x₁ ∷ xs ∣₂) ∣ _ ∣₂))) + Iso.leftInv Iso∥List∥₂List∥∥₂ = List.elim refl ((lem _ _ ∙_) ∘S cong (_ ∷_)) + where + lem = elim2 {C = λ a l' → ∥List∥₂→List∥∥₂ + (rec2 squash₂ (λ x₁ xs → ∣ x₁ ∷ xs ∣₂) a l') + ≡ a ∷ ∥List∥₂→List∥∥₂ l'} + (λ _ _ → isProp→isSet (isOfHLevelList 0 squash₂ _ _)) + λ _ _ → refl + + List-comm-∥∥₂ : ∀ {ℓ} → List {ℓ} ∘ ∥_∥₂ ≡ ∥_∥₂ ∘ List + List-comm-∥∥₂ = funExt λ A → isoToPath (Iso∥List∥₂List∥∥₂ {A = A}) diff --git a/Cubical/Data/Nat/Properties.agda b/Cubical/Data/Nat/Properties.agda index e3b3c64889..a0097e49d9 100644 --- a/Cubical/Data/Nat/Properties.agda +++ b/Cubical/Data/Nat/Properties.agda @@ -173,6 +173,12 @@ m+n≡0→m≡0×n≡0 : m + n ≡ 0 → (m ≡ 0) × (n ≡ 0) m+n≡0→m≡0×n≡0 {zero} = refl ,_ m+n≡0→m≡0×n≡0 {suc m} p = ⊥.rec (snotz p) +m+n≡1→m≡0×n≡1⊎m≡1n≡0 : m + n ≡ 1 → ((m ≡ 0) × (n ≡ 1)) ⊎ ((m ≡ 1) × (n ≡ 0)) +m+n≡1→m≡0×n≡1⊎m≡1n≡0 {zero} x = inl (refl , x) +m+n≡1→m≡0×n≡1⊎m≡1n≡0 {suc m} {n} x = + let (m≡0 , n≡0) = m+n≡0→m≡0×n≡0 (injSuc x) + in inr (cong suc m≡0 , n≡0 ) + -- Arithmetic facts about · 0≡m·0 : ∀ m → 0 ≡ m · 0 diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index 7a1a0f53cd..a3db8125dc 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -378,6 +378,10 @@ snd (ΣPathPProp {B = B} {u = u} {v = v} pB p i) = lem i lem : PathP (λ i → B i (p i)) (snd u) (snd v) lem = toPathP (pB _ _ _) +discreteΣProp : Discrete A → ((x : A) → isProp (B x)) → Discrete (Σ A B) +discreteΣProp _≟_ isPropA _ _ = + EquivPresDec (Σ≡PropEquiv isPropA) (_ ≟ _) + ≃-× : ∀ {ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} → A ≃ C → B ≃ D → A × B ≃ C × D ≃-× eq1 eq2 = map-× (fst eq1) (fst eq2) diff --git a/Cubical/Foundations/Function.agda b/Cubical/Foundations/Function.agda index bc37863fd4..d68fbad72a 100644 --- a/Cubical/Foundations/Function.agda +++ b/Cubical/Foundations/Function.agda @@ -9,7 +9,7 @@ open import Cubical.Foundations.Prelude private variable ℓ ℓ' ℓ'' : Level - A : Type ℓ + A A' A'' : Type ℓ B : A → Type ℓ C : (a : A) → B a → Type ℓ D : (a : A) (b : B a) → C a b → Type ℓ @@ -26,12 +26,16 @@ _$_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} → ((a : A) → B a) f $ a = f a {-# INLINE _$_ #-} -infixr 9 _∘_ +infixr 9 _∘_ _∘S_ _∘_ : (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 _∘_ #-} +_∘S_ : (A' → A'') → (A → A') → A → A'' +g ∘S f = λ x → g (f x) +{-# INLINE _∘S_ #-} + ∘-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) diff --git a/Cubical/Foundations/Univalence.agda b/Cubical/Foundations/Univalence.agda index 53c858608c..d6c688e421 100644 --- a/Cubical/Foundations/Univalence.agda +++ b/Cubical/Foundations/Univalence.agda @@ -69,6 +69,13 @@ module _ {A B : Type ℓ} (e : A ≃ B) {x : A} {y : B} where unquoteDef ua-ungluePath-Equiv = defStrictEquiv ua-ungluePath-Equiv ua-ungluePath ua-gluePath +ua-ungluePathExt : {A B : Type ℓ} (e : A ≃ B) → PathP (λ i → ua e i → B) (fst e) (idfun B) +ua-ungluePathExt e i = ua-unglue e i + +ua-gluePathExt : {A B : Type ℓ} (e : A ≃ B) → PathP (λ i → A → ua e i) (idfun _) (fst e) +ua-gluePathExt e i x = + ua-glue e i (λ { (i = i0) → x }) (inS (fst e x)) + -- 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 @@ -137,6 +144,14 @@ unglueEquiv : ∀ (A : Type ℓ) (φ : I) (Glue A f) ≃ A unglueEquiv A φ f = ( unglue φ , unglueIsEquiv A φ f ) +ua-unglueEquiv : ∀ {A B : Type ℓ} (e : A ≃ B) → + PathP (λ i → ua e i ≃ B) + e + (idEquiv _) +fst (ua-unglueEquiv e i) = ua-unglue e i +snd (ua-unglueEquiv e i) = + isProp→PathP (λ i → isPropIsEquiv (ua-unglue e i)) + (snd e) (idIsEquiv _) i -- The following is a formulation of univalence proposed by Martín Escardó: -- https://groups.google.com/forum/#!msg/homotopytypetheory/HfCB_b-PNEU/Ibb48LvUMeUJ @@ -194,6 +209,9 @@ pathToEquivRefl {A = A} = equivEq (λ i x → transp (λ _ → A) i x) 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 ℓ} (e : A ≃ B) (x : B) → transport (sym (ua e)) x ≡ invEq e x +~uaβ e x = cong (invEq e) (transportRefl x) + uaη : ∀ {A B : Type ℓ} → (P : A ≡ B) → ua (pathToEquiv P) ≡ P uaη {A = A} {B = B} P i j = Glue B {φ = φ} sides where -- Adapted from a proof by @dolio, cf. commit e42a6fa1 diff --git a/Cubical/HITs/Bouquet/Discrete.agda b/Cubical/HITs/Bouquet/Discrete.agda new file mode 100644 index 0000000000..5b9c69473e --- /dev/null +++ b/Cubical/HITs/Bouquet/Discrete.agda @@ -0,0 +1,129 @@ +{-# OPTIONS --safe #-} + +{- + +This module, uses normalization and decidable equality from `Cubical.Algebra.Group.Free` to demonstrate that Bouquet over discrete type is hGroupoid by establishing a coding between loops in the bouquet and elements of the FreeGroup represented by normalised words. + +-} + +module Cubical.HITs.Bouquet.Discrete where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Bool +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sigma +open import Cubical.Data.List renaming (rec to recList) +open import Cubical.Data.Sum + +open import Cubical.Relation.Nullary + + +open import Cubical.HITs.Bouquet.Base +open import Cubical.Algebra.Group.Free + + +private + variable + ℓ : Level + +module _ {A : Type ℓ} (_≟_ : Discrete A) where + + open NormalForm A + + open Discrete _≟_ + + + CodeBouquet : Bouquet A → Type ℓ + CodeBouquet base = Σ _ IsNormalised + CodeBouquet (loop a i) = ua (η·≃ (true , a)) i + + co→ : ∀ x → base ≡ x → CodeBouquet x + co→ x p = subst CodeBouquet p ([] , lower) + + co←base-step : Bool × A → Path (Bouquet A) base base + + co←base-step (b , a) = if b then (idfun _) else sym $ loop a + + co←base : [𝟚× A ] → Path (Bouquet A) base base + co←base = recList refl (flip _∙_ ∘ co←base-step) + + co← : ∀ x → CodeBouquet x → base ≡ x + co← base = co←base ∘ fst + co← (loop a i) x j = co←Sq a i j x + where + + co←Sq : (a : A) → SquareP (λ i j → ua (η·≃ (true , a)) i → Bouquet A) + (funExt (co←base ∘ fst)) + (funExt (co←base ∘ fst)) + (λ _ _ → base) + (λ i _ → loop a i) + co←Sq a = congP (λ _ → funExt) (ua→ (uncurry + (λ xs y → co←Sq' a xs y (HeadIsRedex? ((true , a) ∷ xs))))) + where + co←Sq' : (a : A) → (x : [𝟚× A ]) (y : IsNormalised x) → + ∀ u → PathP (λ i → base ≡ loop a i) + (recList (λ _ → base) (flip _∙_ ∘ co←base-step) x) + (recList (λ _ → base) (flip _∙_ ∘ co←base-step) (preη· (true , a) x u )) + co←Sq' a ((false , snd₁) ∷ xs) y (yes p) = + cong (λ x' → co←base ((false , x') ∷ xs)) (cong snd (sym p)) + ◁ symP (compPath-filler (co←base xs) (sym (loop a))) + co←Sq' a xs y (no ¬p) = compPath-filler _ _ + co←Sq' a ((true , snd₁) ∷ xs) y (yes p) = ⊥.rec (true≢false (cong fst p)) + + coSec : ∀ x → section (co← x) (co→ x) + coSec _ = J (λ x b → co← x (co→ x b) ≡ b) refl + + coRet : (x : [𝟚× A ]) (y : IsNormalised x) → + fst (subst CodeBouquet (co← base (x , y)) ([] , lower)) + ≡ x + coRet [] y = refl + coRet (x@(b , a) ∷ xs) y = + cong fst (substComposite CodeBouquet (co← base (xs , y ∘ inr)) + (co←base-step x) _) + ∙∙ + cong (fst ∘ subst CodeBouquet (co←base-step x)) + (Σ≡Prop isPropIsNormalised (coRet xs (y ∘ inr))) ∙∙ + lem b xs (y ∘ inr) ∙ η·∷ x xs (y ∘ inl) + + where + lem : ∀ b xs y → fst + (subst CodeBouquet (co←base-step (b , a)) (xs , y)) + ≡ η· (b , a) xs + lem false xs y = cong fst (~uaβ (η·≃ (true , a)) (xs , y )) + lem true xs y = cong fst (uaβ (η·≃ (true , a)) (xs , y )) + + + codeDecode : Iso (Path (Bouquet A) base base) + (Σ _ IsNormalised) + Iso.fun codeDecode p = subst CodeBouquet p ([] , lower) + Iso.inv codeDecode = co← base + Iso.rightInv codeDecode = Σ≡Prop isPropIsNormalised ∘ uncurry coRet + Iso.leftInv codeDecode = coSec base + + isSetΩBouquet : isSet (Path (Bouquet A) base base) + isSetΩBouquet = isOfHLevelRetractFromIso 2 codeDecode + (isSetΣ isSet[𝟚×] (isProp→isSet ∘ isPropIsNormalised)) + + isGroupoidBouquet : isGroupoid (Bouquet A) + isGroupoidBouquet base base = isSetΩBouquet + isGroupoidBouquet base (loop x i) = isProp→PathP + (λ i → isPropIsSet {A = Path (Bouquet A) base (loop x i)}) isSetΩBouquet isSetΩBouquet i + isGroupoidBouquet (loop x i) base = isProp→PathP + (λ i → isPropIsSet {A = Path (Bouquet A) (loop x i) base}) isSetΩBouquet isSetΩBouquet i + isGroupoidBouquet (loop x i) (loop x₁ i₁) = isSet→SquareP + (λ i i₁ → isProp→isSet (isPropIsSet {A = Path (Bouquet A) (loop x i) (loop x₁ i₁)})) + (isProp→PathP + (λ i → isPropIsSet {A = Path (Bouquet A) base (loop x₁ i)}) isSetΩBouquet isSetΩBouquet) + (isProp→PathP + (λ i → isPropIsSet {A = Path (Bouquet A) base (loop x₁ i)}) isSetΩBouquet isSetΩBouquet) + (isProp→PathP + (λ i → isPropIsSet {A = Path (Bouquet A) (loop x i) base}) isSetΩBouquet isSetΩBouquet) + (isProp→PathP + (λ i → isPropIsSet {A = Path (Bouquet A) (loop x i) base}) isSetΩBouquet isSetΩBouquet) + i i₁ diff --git a/Cubical/HITs/FreeGroup/NormalForm.agda b/Cubical/HITs/FreeGroup/NormalForm.agda new file mode 100644 index 0000000000..a4b3a87e12 --- /dev/null +++ b/Cubical/HITs/FreeGroup/NormalForm.agda @@ -0,0 +1,187 @@ +{-# OPTIONS --safe #-} + +{- + +For an arbitrary type of generators: + - `GroupIso-FG-L/↘↙` demonstrates that the free group defined as a HIT is equivalent to the definition based on quotienting lists by the appropriate relation (from `Cubical.Algebra.Group.Free`). + +The following properties are defined with the assumption that the type of the generators is an hSet. Without this assumption, they can be adapted to be stated "modulo set truncation": + + - `isPropNF` ensures the uniqueness of the normal form. + - `ηInj` establishes that the `η` constructor of FreeGroup is injective. + - `NF⇔DiscreteA` indicates that computing the normal form is feasible if and only if the type of generators is discrete. + - `discreteFreeGroup` asserts that if the generators are discrete, then the FreeGroup has decidable equality. + +-} +module Cubical.HITs.FreeGroup.NormalForm 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.Relation.Nullary +open import Cubical.Relation.Nullary.HLevels + +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.Properties +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Free + +open import Cubical.Data.List renaming (elim to elimList) +open import Cubical.Data.Bool as 𝟚 +open import Cubical.Data.Sigma +open import Cubical.Data.Sum as ⊎ + +open import Cubical.HITs.SetQuotients renaming + (elimProp to elimProp/ ; _/_ to _/₂_ ; [_] to [_]/ ; rec to rec/) +open import Cubical.HITs.PropositionalTruncation renaming + (rec to rec₁ ; map to map₁) + + +open import Cubical.HITs.FreeGroup.Base +open import Cubical.HITs.FreeGroup.Properties renaming (rec to recFreeGroup) + +private + variable + ℓ : Level + +module _ {A : Type ℓ} where + + open NormalForm A + + open NF (freeGroupGroup A) η renaming (inv to invFG) + + FG→L/↘↙ : GroupHom (freeGroupGroup A) (_ , List/↘↙GroupStr) + FG→L/↘↙ = recFreeGroup ([_]/ ∘ [_] ∘ (true ,_)) + + module gh/ = IsGroupHom (snd (FG→L/↘↙)) + open GroupTheory (freeGroupGroup A) + + open IsGroupHom + + ⇊1g→FG≡ : ∀ a → a ⇊1g → fromList a ≡ ε + ⇊1g→FG≡ .[] [] = refl + ⇊1g→FG≡ .(x ∷ (xs ∷ʳ not₁ x)) (cj x xs x₁) = + cong (η* x ·fg_) (fromList· xs [ not₁ x ] ∙ + cong₂ _·fg_ (⇊1g→FG≡ xs x₁) (·IdR _) ∙ ·IdL _) ∙ + redex-ε-η* x (not₁ x) (symIsRedex _ _ refl) + ⇊1g→FG≡ .(xs ++ ys) ((xs · ys) x x₁) = fromList· xs ys + ∙∙ cong₂ _·fg_ (⇊1g→FG≡ _ x) (⇊1g→FG≡ _ x₁) ∙∙ ·IdL _ + + fromListInv : (xs : List (Bool × A)) → + fromList (invLi xs) ≡ invFG (fromList xs) + fromListInv [] = sym (GroupTheory.inv1g (freeGroupGroup A)) + fromListInv (x ∷ xs) = (fromList· (invLi xs) _ ∙ + cong (fromList (invLi xs) ·fg_) (w' x)) + ∙∙ cong (_·fg invFG (η* x)) (fromListInv xs) ∙∙ sym (invDistr _ _) + where + w' : ∀ x → fromList [ not₁ x ] ≡ invFG (η* x) + w' = λ { (false , a) → ·IdR _ ∙ sym (invInv _) ; (true , a) → ·IdR _ } + + fromL/ : List/↘↙ → _ + fromL/ = rec/ trunc fromList + λ a b → + _∙ (sym (fromListInv (invLi b)) + ∙ cong fromList (invol-invLi b)) ∘S invUniqueL + ∘S sym (fromList· a (invLi b)) ∙_ ∘S ⇊1g→FG≡ _ ∘S ≡ε + + section-FG-L/↘↙ : ∀ a → fst (FG→L/↘↙) (fromList a) ≡ [ a ]/ + section-FG-L/↘↙ [] = refl + section-FG-L/↘↙ (x ∷ xs) = gh/.pres· (η* x) (fromList xs) ∙ + cong (List/↘↙· (fst FG→L/↘↙ (η* x))) + (section-FG-L/↘↙ xs) ∙ w x where + w : ∀ x → List/↘↙· (fst FG→L/↘↙ (η* x)) [ xs ]/ ≡ [ x ∷ xs ]/ + w = λ { (false , a) → refl ; (true , a) → refl } + + isGroupHomFromL/ : IsGroupHom List/↘↙GroupStr fromL/ (snd (freeGroupGroup A)) + pres· isGroupHomFromL/ = elimProp2 (λ _ _ → trunc _ _) fromList· + pres1 isGroupHomFromL/ = refl + presinv isGroupHomFromL/ = elimProp/ (λ _ → trunc _ _) fromListInv + + GroupIso-FG-L/↘↙ : GroupIso (freeGroupGroup A) (List/↘↙group) + Iso.fun (fst GroupIso-FG-L/↘↙) = fst FG→L/↘↙ + Iso.inv (fst GroupIso-FG-L/↘↙) = fromL/ + Iso.rightInv (fst GroupIso-FG-L/↘↙) = + elimProp/ (λ _ → squash/ _ _) section-FG-L/↘↙ + Iso.leftInv (fst GroupIso-FG-L/↘↙) = + funExt⁻ (congS fst (freeGroupHom≡ + {f = compGroupHom FG→L/↘↙ (fromL/ , isGroupHomFromL/)} + {g = idGroupHom} (sym ∘ idr ∘ η ))) + snd GroupIso-FG-L/↘↙ = snd FG→L/↘↙ + + module _ (isSetA : isSet A) where + + private + isSet[𝟚×A] = isOfHLevelList 0 (isSet× isSetBool isSetA) + + isPropNF : ∀ g → isProp (NF g) + isPropNF = λ g → + λ (xs nf u , v) (xs' nf u' , v') → + let zz = rec₁ (isSet[𝟚×A] xs xs') + ( sym + ∘S nf-uR _ _ (fst IsNormalisedInvLi v') v + ∘S ⇊1g++comm xs (invLi xs') + ∘S ≡ε ) + (Iso.fun (≡→red xs xs') ( + isoInvInjective (fst (GroupIso-FG-L/↘↙)) + _ _ (u ∙ (sym u')))) + in λ i → zz i + nf isProp→PathP (λ i → trunc (fromList (zz i)) g) u u' i + , isProp→PathP (λ i → (isPropIsNormalised (zz i))) v v' i + + ηInj : ∀ a a' → Path (FreeGroup A) (η a) (η a') → a ≡ a' + ηInj a a' = + rec₁ (isSetA _ _) + ((λ { (inl p) i → snd (p i) + ; (inr (inl ())) ; (inr (inr ()))}) + ∘S ⇊1g⇒HasRedex _ _ ∘S ≡ε ) + ∘S Iso.fun (≡→red _ _) + ∘S isoInvInjective (fst (GroupIso-FG-L/↘↙)) + [ [ true , _ ] ]/ [ [ true , _ ] ]/ + ∘S ·IdR _ ∙∙_∙∙ sym (·IdR _) + + NF-η : ∀ a → (nfa : NF (η a)) → [ true , a ] ≡ NF.word nfa + NF-η a nfa = rec₁ (isSet[𝟚×A] _ _) (λ u → + nf-uR _ _ (IsNormalised[x] (true , a)) + (NF.isNormalisedWord nfa) (⇊1g++comm _ [ false , a ] (≡ε u))) + (Iso.fun (≡→red _ _) (isoInvInjective (fst (GroupIso-FG-L/↘↙)) + [ NF.word nfa ]/ [ [ (true , a) ] ]/ + (NF.fromListWord≡ nfa ∙ (sym (·IdR _))))) + + ΠNF⇒DiscreteA : (∀ g → NF g) → Discrete A + ΠNF⇒DiscreteA nF a a' = + let nff = nF (η a · invFG (η a')) + in rec₁ (isPropDec (isSetA _ _)) + (λ r → ⊎.rec + (yes ∘ sym ∘ cong snd) + (no ∘ ⊎.rec (λ p pp → lower (subst (WillReduce false a) + (isNormalised⇊1g _ (NF.isNormalisedWord nff) + (pop⇊1gHead (cong (true ,_) (sym pp)) r)) p)) + (const ∘ NF.isNormalisedWord nff)) + (⇊1g⇒HasRedex _ _ r)) + (map₁ (⇊1g++comm (NF.word nff) _ ∘ ≡ε) + (Iso.fun (≡→red _ _) (isoInvInjective (fst (GroupIso-FG-L/↘↙)) + [ NF.word nff ]/ [ (true , a) ∷ [ false , a' ] ]/ + (NF.fromListWord≡ nff ∙ cong (η a ·_) (sym (·IdR _)))))) + + NF⇔DiscreteA : (∀ g → NF g) ≃ Discrete A + NF⇔DiscreteA = propBiimpl→Equiv (isPropΠ isPropNF) isPropDiscrete + ΠNF⇒DiscreteA λ _≟_ g → + let e = compIso (fst (GroupIso-FG-L/↘↙)) (invIso (Discrete.IsoNF _≟_)) + (g' , n) = Iso.fun e g + in g' nf Iso.leftInv e g , n + + ≟→normalForm : Discrete A → (∀ g → NF g) + ≟→normalForm _≟_ = invEq (NF⇔DiscreteA (Discrete→isSet _≟_)) _≟_ + + discreteFreeGroup : Discrete A → Discrete (FreeGroup A) + discreteFreeGroup _≟_ = + isoPresDiscrete + (compIso + (Discrete.IsoNF _≟_) + (invIso (fst (GroupIso-FG-L/↘↙)))) + (discreteΣProp + (discreteList (discreteΣ 𝟚._≟_ λ _ → _≟_)) + isPropIsNormalised) diff --git a/Cubical/Homotopy/Group/Base.agda b/Cubical/Homotopy/Group/Base.agda index 5c856ba476..ec0c1f8f4a 100644 --- a/Cubical/Homotopy/Group/Base.agda +++ b/Cubical/Homotopy/Group/Base.agda @@ -14,6 +14,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Univalence open import Cubical.Foundations.Function open import Cubical.Foundations.Transport +open import Cubical.Foundations.Structure open import Cubical.Functions.Morphism @@ -1081,3 +1082,15 @@ snd (fst (πIso e n)) = (setTruncIso (equivToIso (_ , isEquivΩ^→ (suc n) (≃∙map e) (snd (fst e))))) snd (πIso e n) = snd (πHom n (≃∙map e)) + +hGroupoidπ₁ : ∀ {ℓ} (A : hGroupoid ℓ) → ⟨ A ⟩ → Group ℓ +fst (hGroupoidπ₁ A a) = a ≡ a +1g (snd (hGroupoidπ₁ A a)) = refl +GroupStr._·_ (snd (hGroupoidπ₁ A a)) = _∙_ +inv (snd (hGroupoidπ₁ A a)) = sym +is-set (isSemigroup (isMonoid (isGroup (snd (hGroupoidπ₁ A a))))) = snd A a a +·Assoc (isSemigroup (isMonoid (isGroup (snd (hGroupoidπ₁ A a))))) = ∙assoc +·IdR (isMonoid (isGroup (snd (hGroupoidπ₁ A a)))) = sym ∘ rUnit +·IdL (isMonoid (isGroup (snd (hGroupoidπ₁ A a)))) = sym ∘ lUnit +·InvR (isGroup (snd (hGroupoidπ₁ A a))) = rCancel +·InvL (isGroup (snd (hGroupoidπ₁ A a))) = lCancel diff --git a/Cubical/Relation/Nullary/Properties.agda b/Cubical/Relation/Nullary/Properties.agda index a3e99d3c78..77535b9c88 100644 --- a/Cubical/Relation/Nullary/Properties.agda +++ b/Cubical/Relation/Nullary/Properties.agda @@ -201,3 +201,7 @@ Discrete→Separated d x y = Dec→Stable (d x y) Discrete→isSet : Discrete A → isSet A Discrete→isSet = Separated→isSet ∘ Discrete→Separated + +≡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) From 598dfa5c85492f503d04164e79dc4bb0e1e54a24 Mon Sep 17 00:00:00 2001 From: mzeuner <56261690+mzeuner@users.noreply.github.com> Date: Tue, 16 Apr 2024 10:37:31 +0200 Subject: [PATCH 03/30] Algebraic geometry directory, take 2 (#1121) * move Zariski lattice * add to README * restructure * fix paper * typos and fixes * Z functor instance * relative coadjunction --- .../Functorial/ZFunctors/Base.agda | 245 ++++++ .../Functorial/ZFunctors/CompactOpen.agda | 382 +++++++++ .../Functorial/ZFunctors/OpenSubscheme.agda | 201 +++++ .../Functorial/ZFunctors/QcQsScheme.agda | 82 ++ .../ZariskiLattice/Base.agda | 2 +- .../ZariskiLattice/Properties.agda | 6 +- .../ZariskiLattice/StructureSheaf.agda | 6 +- .../StructureSheafPullback.agda | 6 +- .../ZariskiLattice/UniversalProperty.agda | 4 +- Cubical/Categories/Instances/ZFunctors.agda | 801 +----------------- Cubical/Papers/AffineSchemes.agda | 52 +- Cubical/Papers/FunctorialQcQsSchemes.agda | 58 +- Cubical/README.agda | 3 + 13 files changed, 995 insertions(+), 853 deletions(-) create mode 100644 Cubical/AlgebraicGeometry/Functorial/ZFunctors/Base.agda create mode 100644 Cubical/AlgebraicGeometry/Functorial/ZFunctors/CompactOpen.agda create mode 100644 Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda create mode 100644 Cubical/AlgebraicGeometry/Functorial/ZFunctors/QcQsScheme.agda rename Cubical/{Algebra => AlgebraicGeometry}/ZariskiLattice/Base.agda (99%) rename Cubical/{Algebra => AlgebraicGeometry}/ZariskiLattice/Properties.agda (97%) rename Cubical/{Algebra => AlgebraicGeometry}/ZariskiLattice/StructureSheaf.agda (98%) rename Cubical/{Algebra => AlgebraicGeometry}/ZariskiLattice/StructureSheafPullback.agda (98%) rename Cubical/{Algebra => AlgebraicGeometry}/ZariskiLattice/UniversalProperty.agda (99%) diff --git a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/Base.agda b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/Base.agda new file mode 100644 index 0000000000..8524743648 --- /dev/null +++ b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/Base.agda @@ -0,0 +1,245 @@ +{- + + A ℤ-functor is just a functor from rings to sets. + + NOTE: we consider the functor category [ Ring ℓ , Set ℓ ] for some universe level ℓ + and not [ Ring ℓ , Set (ℓ+1) ] as is done in + "Introduction to Algebraic Geometry and Algebraic Groups" + by Demazure & Gabriel! + + The category of ℤ-functors contains the category of (qcqs-) schemes + as a full subcategory and satisfies a "universal property" + similar to the one of schemes: + + There is a coadjunction 𝓞 ⊣ᵢ Sp + (relative to the inclusion i : CommRing ℓ → CommRing (ℓ+1)) + between the "global sections functor" 𝓞 + and the fully-faithful embedding of affines Sp, + whose counit is a natural isomorphism + +-} + +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.AlgebraicGeometry.Functorial.ZFunctors.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.HLevels + +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ) + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing + +open import Cubical.Categories.Category renaming (isIso to isIsoC) +open import Cubical.Categories.Functor +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Instances.CommRings +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Yoneda +open import Cubical.Categories.Site.Sheaf +open import Cubical.Categories.Site.Instances.ZariskiCommRing + +open import Cubical.HITs.PropositionalTruncation as PT + + +open Category hiding (_∘_) + + +module _ {ℓ : Level} where + + open Functor + open NatTrans + open CommRingStr ⦃...⦄ + open IsRingHom + + + -- using the naming conventions of Demazure & Gabriel + ℤFunctor = Functor (CommRingsCategory {ℓ = ℓ}) (SET ℓ) + ℤFUNCTOR = FUNCTOR (CommRingsCategory {ℓ = ℓ}) (SET ℓ) + + -- Yoneda in the notation of Demazure & Gabriel, + -- uses that double op is original category definitionally + Sp : Functor (CommRingsCategory {ℓ = ℓ} ^op) ℤFUNCTOR + Sp = YO {C = (CommRingsCategory {ℓ = ℓ} ^op)} + + isAffine : (X : ℤFunctor) → Type (ℓ-suc ℓ) + isAffine X = ∃[ A ∈ CommRing ℓ ] NatIso (Sp .F-ob A) X + -- TODO: 𝔸¹ ≅ Sp ℤ[x] and 𝔾ₘ ≅ Sp ℤ[x,x⁻¹] ≅ D(x) ↪ 𝔸¹ as first examples of affine schemes + + -- a ℤ-functor that is a sheaf wrt the Zariski coverage is called local + isLocal : ℤFunctor → Type (ℓ-suc ℓ) + isLocal X = isSheaf zariskiCoverage X + + -- the forgetful functor + -- aka the affine line + -- (aka the representable of ℤ[x]) + 𝔸¹ : ℤFunctor + 𝔸¹ = ForgetfulCommRing→Set + + -- the global sections functor + 𝓞 : Functor ℤFUNCTOR (CommRingsCategory {ℓ = ℓ-suc ℓ} ^op) + fst (F-ob 𝓞 X) = X ⇒ 𝔸¹ + + -- ring struncture induced by internal ring object 𝔸¹ + N-ob (CommRingStr.0r (snd (F-ob 𝓞 X))) A _ = 0r + where instance _ = A .snd + N-hom (CommRingStr.0r (snd (F-ob 𝓞 X))) φ = funExt λ _ → sym (φ .snd .pres0) + + N-ob (CommRingStr.1r (snd (F-ob 𝓞 X))) A _ = 1r + where instance _ = A .snd + N-hom (CommRingStr.1r (snd (F-ob 𝓞 X))) φ = funExt λ _ → sym (φ .snd .pres1) + + N-ob ((snd (F-ob 𝓞 X) CommRingStr.+ α) β) A x = α .N-ob A x + β .N-ob A x + where instance _ = A .snd + N-hom ((snd (F-ob 𝓞 X) CommRingStr.+ α) β) {x = A} {y = B} φ = funExt path + where + instance + _ = A .snd + _ = B .snd + path : ∀ x → α .N-ob B (X .F-hom φ x) + β .N-ob B (X .F-hom φ x) + ≡ φ .fst (α .N-ob A x + β .N-ob A x) + path x = α .N-ob B (X .F-hom φ x) + β .N-ob B (X .F-hom φ x) + ≡⟨ cong₂ _+_ (funExt⁻ (α .N-hom φ) x) (funExt⁻ (β .N-hom φ) x) ⟩ + φ .fst (α .N-ob A x) + φ .fst (β .N-ob A x) + ≡⟨ sym (φ .snd .pres+ _ _) ⟩ + φ .fst (α .N-ob A x + β .N-ob A x) ∎ + + N-ob ((snd (F-ob 𝓞 X) CommRingStr.· α) β) A x = α .N-ob A x · β .N-ob A x + where instance _ = A .snd + N-hom ((snd (F-ob 𝓞 X) CommRingStr.· α) β) {x = A} {y = B} φ = funExt path + where + instance + _ = A .snd + _ = B .snd + path : ∀ x → α .N-ob B (X .F-hom φ x) · β .N-ob B (X .F-hom φ x) + ≡ φ .fst (α .N-ob A x · β .N-ob A x) + path x = α .N-ob B (X .F-hom φ x) · β .N-ob B (X .F-hom φ x) + ≡⟨ cong₂ _·_ (funExt⁻ (α .N-hom φ) x) (funExt⁻ (β .N-hom φ) x) ⟩ + φ .fst (α .N-ob A x) · φ .fst (β .N-ob A x) + ≡⟨ sym (φ .snd .pres· _ _) ⟩ + φ .fst (α .N-ob A x · β .N-ob A x) ∎ + + N-ob ((CommRingStr.- snd (F-ob 𝓞 X)) α) A x = - α .N-ob A x + where instance _ = A .snd + N-hom ((CommRingStr.- snd (F-ob 𝓞 X)) α) {x = A} {y = B} φ = funExt path + where + instance + _ = A .snd + _ = B .snd + path : ∀ x → - α .N-ob B (X .F-hom φ x) ≡ φ .fst (- α .N-ob A x) + path x = - α .N-ob B (X .F-hom φ x) ≡⟨ cong -_ (funExt⁻ (α .N-hom φ) x) ⟩ + - φ .fst (α .N-ob A x) ≡⟨ sym (φ .snd .pres- _) ⟩ + φ .fst (- α .N-ob A x) ∎ + + CommRingStr.isCommRing (snd (F-ob 𝓞 X)) = makeIsCommRing + isSetNatTrans + (λ _ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.+Assoc _ _ _)) + (λ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.+IdR _)) + (λ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.+InvR _)) + (λ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.+Comm _ _)) + (λ _ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.·Assoc _ _ _)) + (λ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.·IdR _)) + (λ _ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.·DistR+ _ _ _)) + (λ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.·Comm _ _)) + + -- action on natural transformations + fst (F-hom 𝓞 α) = α ●ᵛ_ + pres0 (snd (F-hom 𝓞 α)) = makeNatTransPath (funExt₂ λ _ _ → refl) + pres1 (snd (F-hom 𝓞 α)) = makeNatTransPath (funExt₂ λ _ _ → refl) + pres+ (snd (F-hom 𝓞 α)) _ _ = makeNatTransPath (funExt₂ λ _ _ → refl) + pres· (snd (F-hom 𝓞 α)) _ _ = makeNatTransPath (funExt₂ λ _ _ → refl) + pres- (snd (F-hom 𝓞 α)) _ = makeNatTransPath (funExt₂ λ _ _ → refl) + + -- functoriality of 𝓞 + F-id 𝓞 = RingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) + F-seq 𝓞 _ _ = RingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) + + + +-- There is a coadjunction 𝓞 ⊣ᵢ Sp +-- (relative to the inclusion i : CommRing ℓ → CommRing (ℓ+1)) +-- between the "global sections functor" 𝓞 +-- and the fully-faithful embedding of affines Sp, +-- whose counit is a natural isomorphism +module AdjBij {ℓ : Level} where + + open Functor + open NatTrans + open Iso + open IsRingHom + + private module _ {A : CommRing ℓ} {X : ℤFunctor {ℓ}} where + _♭ : CommRingHom A (𝓞 .F-ob X) → X ⇒ Sp .F-ob A + fst (N-ob (φ ♭) B x) a = φ .fst a .N-ob B x + + pres0 (snd (N-ob (φ ♭) B x)) = cong (λ y → y .N-ob B x) (φ .snd .pres0) + pres1 (snd (N-ob (φ ♭) B x)) = cong (λ y → y .N-ob B x) (φ .snd .pres1) + pres+ (snd (N-ob (φ ♭) B x)) _ _ = cong (λ y → y .N-ob B x) (φ .snd .pres+ _ _) + pres· (snd (N-ob (φ ♭) B x)) _ _ = cong (λ y → y .N-ob B x) (φ .snd .pres· _ _) + pres- (snd (N-ob (φ ♭) B x)) _ = cong (λ y → y .N-ob B x) (φ .snd .pres- _) + + N-hom (φ ♭) ψ = funExt (λ x → RingHom≡ (funExt λ a → funExt⁻ (φ .fst a .N-hom ψ) x)) + + + -- the other direction is just precomposition modulo Yoneda + _♯ : X ⇒ Sp .F-ob A → CommRingHom A (𝓞 .F-ob X) + fst (α ♯) a = α ●ᵛ yonedaᴾ 𝔸¹ A .inv a + + pres0 (snd (α ♯)) = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres0) + pres1 (snd (α ♯)) = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres1) + pres+ (snd (α ♯)) _ _ = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres+ _ _) + pres· (snd (α ♯)) _ _ = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres· _ _) + pres- (snd (α ♯)) _ = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres- _) + + + -- the two maps are inverse to each other + ♭♯Id : ∀ (α : X ⇒ Sp .F-ob A) → ((α ♯) ♭) ≡ α + ♭♯Id _ = makeNatTransPath (funExt₂ λ _ _ → RingHom≡ (funExt (λ _ → refl))) + + ♯♭Id : ∀ (φ : CommRingHom A (𝓞 .F-ob X)) → ((φ ♭) ♯) ≡ φ + ♯♭Id _ = RingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) + + + -- we get a relative adjunction 𝓞 ⊣ᵢ Sp + -- with respect to the inclusion i : CommRing ℓ → CommRing (ℓ+1) + 𝓞⊣SpIso : {A : CommRing ℓ} {X : ℤFunctor {ℓ}} + → Iso (CommRingHom A (𝓞 .F-ob X)) (X ⇒ Sp .F-ob A) + fun 𝓞⊣SpIso = _♭ + inv 𝓞⊣SpIso = _♯ + rightInv 𝓞⊣SpIso = ♭♯Id + leftInv 𝓞⊣SpIso = ♯♭Id + + 𝓞⊣SpNatℤFunctor : {A : CommRing ℓ} {X Y : ℤFunctor {ℓ}} (α : X ⇒ Sp .F-ob A) (β : Y ⇒ X) + → (β ●ᵛ α) ♯ ≡ (𝓞 .F-hom β) ∘cr (α ♯) + 𝓞⊣SpNatℤFunctor _ _ = RingHom≡ (funExt (λ _ → makeNatTransPath (funExt₂ (λ _ _ → refl)))) + + 𝓞⊣SpNatCommRing : {X : ℤFunctor {ℓ}} {A B : CommRing ℓ} + (φ : CommRingHom A (𝓞 .F-ob X)) (ψ : CommRingHom B A) + → (φ ∘cr ψ) ♭ ≡ (φ ♭) ●ᵛ Sp .F-hom ψ + 𝓞⊣SpNatCommRing _ _ = makeNatTransPath (funExt₂ λ _ _ → RingHom≡ (funExt (λ _ → refl))) + + -- the counit is an equivalence + private + ε : (A : CommRing ℓ) → CommRingHom A ((𝓞 ∘F Sp) .F-ob A) + ε A = (idTrans (Sp .F-ob A)) ♯ + + 𝓞⊣SpCounitEquiv : (A : CommRing ℓ) → CommRingEquiv A ((𝓞 ∘F Sp) .F-ob A) + fst (𝓞⊣SpCounitEquiv A) = isoToEquiv theIso + where + theIso : Iso (A .fst) ((𝓞 ∘F Sp) .F-ob A .fst) + fun theIso = ε A .fst + inv theIso = yonedaᴾ 𝔸¹ A .fun + rightInv theIso α = ℤFUNCTOR .⋆IdL _ ∙ yonedaᴾ 𝔸¹ A .leftInv α + leftInv theIso a = path -- I get yellow otherwise + where + path : yonedaᴾ 𝔸¹ A .fun ((idTrans (Sp .F-ob A)) ●ᵛ yonedaᴾ 𝔸¹ A .inv a) ≡ a + path = cong (yonedaᴾ 𝔸¹ A .fun) (ℤFUNCTOR .⋆IdL _) ∙ yonedaᴾ 𝔸¹ A .rightInv a + snd (𝓞⊣SpCounitEquiv A) = ε A .snd diff --git a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/CompactOpen.agda b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/CompactOpen.agda new file mode 100644 index 0000000000..d24abaf8ba --- /dev/null +++ b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/CompactOpen.agda @@ -0,0 +1,382 @@ +{- + + The definition of compact open subfunctors of a ℤ-functor X: + + U ↪ Sp(A) is compact open if it is given by a f.g. ideal of A, + i.e. if ∃ f₁, ... ,fₙ : A s.t. for all rings B: + U(B) = { φ : Hom(A,B) | ⟨ φf₁ , ... , φfₙ ⟩ = B } + + U ↪ X is compact open, if pulling back along any A-valued point + Sp(A) → X gives a compact open of Sp(A). + + By observing that compact open subfunctors of affine schemes + are in 1-1 correspondence with radicals of f.g. ideals, + we get that compact open subfunctors are classified by the + ℤ-functor that sends a ring to its Zariski lattice. + +-} + + +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.AlgebraicGeometry.Functorial.ZFunctors.CompactOpen where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.HLevels + + +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ) + +open import Cubical.Data.FinData + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Localisation +open import Cubical.Algebra.CommRing.RadicalIdeal +open import Cubical.Algebra.Semilattice +open import Cubical.Algebra.Lattice +open import Cubical.Algebra.DistLattice +open import Cubical.Algebra.DistLattice.BigOps + +open import Cubical.AlgebraicGeometry.ZariskiLattice.Base +open import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty +open import Cubical.AlgebraicGeometry.ZariskiLattice.Properties +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.Base + +open import Cubical.Categories.Category renaming (isIso to isIsoC) +open import Cubical.Categories.Functor +open import Cubical.Categories.Instances.CommRings +open import Cubical.Categories.Instances.DistLattice +open import Cubical.Categories.Instances.DistLattices +open import Cubical.Categories.Site.Coverage +open import Cubical.Categories.Site.Sheaf +open import Cubical.Categories.Site.Instances.ZariskiCommRing +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Yoneda + +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SetQuotients as SQ + +open import Cubical.Relation.Binary.Order.Poset + + +module _ {ℓ : Level} where + + open Iso + open Functor + open NatTrans + open NatIso + open DistLatticeStr ⦃...⦄ + open CommRingStr ⦃...⦄ + open IsRingHom + open IsLatticeHom + open ZarLat + open ZarLatUniversalProp + + -- the Zariski lattice functor classifying compact open subobjects + ZarLatFun : ℤFunctor {ℓ = ℓ} + F-ob ZarLatFun A = ZL A , SQ.squash/ + F-hom ZarLatFun φ = inducedZarLatHom φ .fst + F-id ZarLatFun {A} = cong fst (inducedZarLatHomId A) + F-seq ZarLatFun φ ψ = cong fst (inducedZarLatHomSeq φ ψ) + + -- this is a separated presheaf + -- (TODO: prove this a sheaf) + isSeparatedZarLatFun : isSeparated zariskiCoverage ZarLatFun + isSeparatedZarLatFun A (unimodvec n f 1∈⟨f₁,⋯,fₙ⟩) u w uRest≡wRest = + u ≡⟨ sym (∧lLid _) ⟩ + 1l ∧l u ≡⟨ congL _∧l_ D1≡⋁Dfᵢ ⟩ + (⋁ (D A ∘ f)) ∧l u ≡⟨ ⋁Meetldist _ _ ⟩ + ⋁ (λ i → D A (f i) ∧l u) ≡⟨ ⋁Ext Dfᵢ∧u≡Dfᵢ∧w ⟩ + ⋁ (λ i → D A (f i) ∧l w) ≡⟨ sym (⋁Meetldist _ _) ⟩ + (⋁ (D A ∘ f)) ∧l w ≡⟨ congL _∧l_ (sym D1≡⋁Dfᵢ) ⟩ + 1l ∧l w ≡⟨ ∧lLid _ ⟩ + w ∎ + where + open Join (ZariskiLattice A) + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (ZariskiLattice A))) + using (IndPoset) + open LatticeTheory (DistLattice→Lattice (ZariskiLattice A)) + open PosetStr (IndPoset .snd) + open IsSupport (isSupportD A) + open RadicalIdeal A + instance + _ = A .snd + _ = ZariskiLattice A .snd + + D1≡⋁Dfᵢ : 1l ≡ ⋁ (D A ∘ f) + D1≡⋁Dfᵢ = is-antisym _ _ + (supportRadicalIneq f 1r (∈→∈√ _ _ 1∈⟨f₁,⋯,fₙ⟩)) + (1lRightAnnihilates∨l _) + + Dfᵢ∧u≡Dfᵢ∧w : ∀ i → D A (f i) ∧l u ≡ D A (f i) ∧l w + Dfᵢ∧u≡Dfᵢ∧w i = + D A (f i) ∧l u + ≡⟨ sym (cong fst (funExt⁻ (cong fst toLocToDown≡ToDown) u)) ⟩ + locToDownHom .fst (inducedZarLatHom /1AsCommRingHom .fst u) .fst + ≡⟨ cong (λ x → locToDownHom .fst x .fst) (uRest≡wRest i) ⟩ + locToDownHom .fst (inducedZarLatHom /1AsCommRingHom .fst w) .fst + ≡⟨ cong fst (funExt⁻ (cong fst toLocToDown≡ToDown) w) ⟩ + D A (f i) ∧l w ∎ + where + open InvertingElementsBase.UniversalProp A (f i) + open LocDownSetIso A (f i) + + CompactOpen : ℤFunctor → Type (ℓ-suc ℓ) + CompactOpen X = X ⇒ ZarLatFun + + -- the induced subfunctor + ⟦_⟧ᶜᵒ : {X : ℤFunctor} (U : CompactOpen X) → ℤFunctor + F-ob (⟦_⟧ᶜᵒ {X = X} U) A = (Σ[ x ∈ X .F-ob A .fst ] U .N-ob A x ≡ D A 1r) + , isSetΣSndProp (X .F-ob A .snd) λ _ → squash/ _ _ + where instance _ = snd A + F-hom (⟦_⟧ᶜᵒ {X = X} U) {x = A} {y = B} φ (x , Ux≡D1) = (X .F-hom φ x) , path + where + instance + _ = A .snd + _ = B .snd + open IsLatticeHom + path : U .N-ob B (X .F-hom φ x) ≡ D B 1r + path = U .N-ob B (X .F-hom φ x) ≡⟨ funExt⁻ (U .N-hom φ) x ⟩ + ZarLatFun .F-hom φ (U .N-ob A x) ≡⟨ cong (ZarLatFun .F-hom φ) Ux≡D1 ⟩ + ZarLatFun .F-hom φ (D A 1r) ≡⟨ inducedZarLatHom φ .snd .pres1 ⟩ + D B 1r ∎ + F-id (⟦_⟧ᶜᵒ {X = X} U) = funExt (λ x → Σ≡Prop (λ _ → squash/ _ _) + (funExt⁻ (X .F-id) (x .fst))) + F-seq (⟦_⟧ᶜᵒ {X = X} U) φ ψ = funExt (λ x → Σ≡Prop (λ _ → squash/ _ _) + (funExt⁻ (X .F-seq φ ψ) (x .fst))) + + + isAffineCompactOpen : {X : ℤFunctor} (U : CompactOpen X) → Type (ℓ-suc ℓ) + isAffineCompactOpen U = isAffine ⟦ U ⟧ᶜᵒ + + + -- the (big) dist. lattice of compact opens + CompOpenDistLattice : Functor ℤFUNCTOR (DistLatticesCategory {ℓ = ℓ-suc ℓ} ^op) + fst (F-ob CompOpenDistLattice X) = CompactOpen X + + -- lattice structure induce by internal lattice object ZarLatFun + N-ob (DistLatticeStr.0l (snd (F-ob CompOpenDistLattice X))) A _ = 0l + where instance _ = ZariskiLattice A .snd + N-hom (DistLatticeStr.0l (snd (F-ob CompOpenDistLattice X))) _ = funExt λ _ → refl + + N-ob (DistLatticeStr.1l (snd (F-ob CompOpenDistLattice X))) A _ = 1l + where instance _ = ZariskiLattice A .snd + N-hom (DistLatticeStr.1l (snd (F-ob CompOpenDistLattice X))) {x = A} {y = B} φ = funExt λ _ → path + where + instance + _ = A .snd + _ = B .snd + _ = ZariskiLattice B .snd + path : D B 1r ≡ D B (φ .fst 1r) ∨l 0l + path = cong (D B) (sym (φ .snd .pres1)) ∙ sym (∨lRid _) + + N-ob ((snd (F-ob CompOpenDistLattice X) DistLatticeStr.∨l U) V) A x = U .N-ob A x ∨l V .N-ob A x + where instance _ = ZariskiLattice A .snd + N-hom ((snd (F-ob CompOpenDistLattice X) DistLatticeStr.∨l U) V) {x = A} {y = B} φ = funExt path + where + instance + _ = ZariskiLattice A .snd + _ = ZariskiLattice B .snd + path : ∀ x → U .N-ob B (X .F-hom φ x) ∨l V .N-ob B (X .F-hom φ x) + ≡ ZarLatFun .F-hom φ (U .N-ob A x ∨l V .N-ob A x) + path x = U .N-ob B (X .F-hom φ x) ∨l V .N-ob B (X .F-hom φ x) + ≡⟨ cong₂ _∨l_ (funExt⁻ (U .N-hom φ) x) (funExt⁻ (V .N-hom φ) x) ⟩ + ZarLatFun .F-hom φ (U .N-ob A x) ∨l ZarLatFun .F-hom φ (V .N-ob A x) + ≡⟨ sym (inducedZarLatHom φ .snd .pres∨l _ _) ⟩ + ZarLatFun .F-hom φ (U .N-ob A x ∨l V .N-ob A x) ∎ + + N-ob ((snd (F-ob CompOpenDistLattice X) DistLatticeStr.∧l U) V) A x = U .N-ob A x ∧l V .N-ob A x + where instance _ = ZariskiLattice A .snd + N-hom ((snd (F-ob CompOpenDistLattice X) DistLatticeStr.∧l U) V) {x = A} {y = B} φ = funExt path + where + instance + _ = ZariskiLattice A .snd + _ = ZariskiLattice B .snd + path : ∀ x → U .N-ob B (X .F-hom φ x) ∧l V .N-ob B (X .F-hom φ x) + ≡ ZarLatFun .F-hom φ (U .N-ob A x ∧l V .N-ob A x) + path x = U .N-ob B (X .F-hom φ x) ∧l V .N-ob B (X .F-hom φ x) + ≡⟨ cong₂ _∧l_ (funExt⁻ (U .N-hom φ) x) (funExt⁻ (V .N-hom φ) x) ⟩ + ZarLatFun .F-hom φ (U .N-ob A x) ∧l ZarLatFun .F-hom φ (V .N-ob A x) + ≡⟨ sym (inducedZarLatHom φ .snd .pres∧l _ _) ⟩ + ZarLatFun .F-hom φ (U .N-ob A x ∧l V .N-ob A x) ∎ + + DistLatticeStr.isDistLattice (snd (F-ob CompOpenDistLattice X)) = makeIsDistLattice∧lOver∨l + isSetNatTrans + (λ _ _ _ → makeNatTransPath (funExt₂ + (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∨lAssoc _ _ _))) + (λ _ → makeNatTransPath (funExt₂ (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∨lRid _))) + (λ _ _ → makeNatTransPath (funExt₂ (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∨lComm _ _))) + (λ _ _ _ → makeNatTransPath (funExt₂ + (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∧lAssoc _ _ _))) + (λ _ → makeNatTransPath (funExt₂ (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∧lRid _))) + (λ _ _ → makeNatTransPath (funExt₂ (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∧lComm _ _))) + (λ _ _ → makeNatTransPath (funExt₂ -- don't know why ∧lAbsorb∨l doesn't work + (λ A _ → ZariskiLattice A .snd .DistLatticeStr.absorb _ _ .snd))) + (λ _ _ _ → makeNatTransPath (funExt₂ -- same here + (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∧l-dist-∨l _ _ _ .fst))) + + -- (contravariant) action on morphisms + fst (F-hom CompOpenDistLattice α) = α ●ᵛ_ + pres0 (snd (F-hom CompOpenDistLattice α)) = makeNatTransPath (funExt₂ λ _ _ → refl) + pres1 (snd (F-hom CompOpenDistLattice α)) = makeNatTransPath (funExt₂ λ _ _ → refl) + pres∨l (snd (F-hom CompOpenDistLattice α)) _ _ = makeNatTransPath (funExt₂ λ _ _ → refl) + pres∧l (snd (F-hom CompOpenDistLattice α)) _ _ = makeNatTransPath (funExt₂ λ _ _ → refl) + + -- functoriality + F-id CompOpenDistLattice = LatticeHom≡f _ _ + (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) + F-seq CompOpenDistLattice _ _ = LatticeHom≡f _ _ + (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) + + + -- useful lemmas + module _ (X : ℤFunctor) where + open isIsoC + private instance _ = (CompOpenDistLattice .F-ob X) .snd + + compOpenTopNatIso : NatIso X ⟦ 1l ⟧ᶜᵒ + N-ob (trans compOpenTopNatIso) _ φ = φ , refl + N-hom (trans compOpenTopNatIso) _ = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) refl + inv (nIso compOpenTopNatIso B) = fst + sec (nIso compOpenTopNatIso B) = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) refl + ret (nIso compOpenTopNatIso B) = funExt λ _ → refl + + + module _ (X : ℤFunctor) where + open isIsoC + open Join (CompOpenDistLattice .F-ob X) + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (CompOpenDistLattice .F-ob X))) + open PosetStr (IndPoset .snd) hiding (_≤_) + open LatticeTheory ⦃...⦄ + private instance _ = (CompOpenDistLattice .F-ob X) .snd + + compOpenGlobalIncl : (U : CompactOpen X) → ⟦ U ⟧ᶜᵒ ⇒ X + N-ob (compOpenGlobalIncl U) A = fst + N-hom (compOpenGlobalIncl U) φ = refl + + compOpenIncl : {U V : CompactOpen X} → V ≤ U → ⟦ V ⟧ᶜᵒ ⇒ ⟦ U ⟧ᶜᵒ + N-ob (compOpenIncl {U = U} {V = V} V≤U) A (x , Vx≡D1) = x , path + where + instance + _ = A .snd + _ = ZariskiLattice A .snd + _ = DistLattice→Lattice (ZariskiLattice A) + path : U .N-ob A x ≡ D A 1r + path = U .N-ob A x ≡⟨ funExt⁻ (funExt⁻ (cong N-ob (sym V≤U)) A) x ⟩ + V .N-ob A x ∨l U .N-ob A x ≡⟨ cong (_∨l U .N-ob A x) Vx≡D1 ⟩ + D A 1r ∨l U .N-ob A x ≡⟨ 1lLeftAnnihilates∨l _ ⟩ + D A 1r ∎ + N-hom (compOpenIncl V≤U) φ = funExt λ x → Σ≡Prop (λ _ → squash/ _ _) refl + + -- this is essentially U∧_ + compOpenDownHom : (U : CompactOpen X) + → DistLatticeHom (CompOpenDistLattice .F-ob X) + (CompOpenDistLattice .F-ob ⟦ U ⟧ᶜᵒ) + compOpenDownHom U = CompOpenDistLattice .F-hom (compOpenGlobalIncl U) + + module _ {U V : CompactOpen X} (V≤U : V ≤ U) where + -- We need this separate definition to avoid termination checker issues, + -- but we don't understand why. + private + compOpenDownHomFun : (A : CommRing ℓ) + → ⟦ V ⟧ᶜᵒ .F-ob A .fst + → ⟦ compOpenDownHom U .fst V ⟧ᶜᵒ .F-ob A .fst + compOpenDownHomFun A v = (compOpenIncl V≤U ⟦ A ⟧) v , snd v + + compOpenDownHomNatIso : NatIso ⟦ V ⟧ᶜᵒ ⟦ compOpenDownHom U .fst V ⟧ᶜᵒ + N-ob (trans compOpenDownHomNatIso) = compOpenDownHomFun + N-hom (trans compOpenDownHomNatIso) _ = + funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (Σ≡Prop (λ _ → squash/ _ _) refl) + inv (nIso compOpenDownHomNatIso A) ((x , Ux≡D1) , Vx≡D1) = x , Vx≡D1 + sec (nIso compOpenDownHomNatIso A) = + funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (Σ≡Prop (λ _ → squash/ _ _) refl) + ret (nIso compOpenDownHomNatIso A) = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) refl + + compOpenInclId : ∀ {U : CompactOpen X} → compOpenIncl (is-refl U) ≡ idTrans ⟦ U ⟧ᶜᵒ + compOpenInclId = makeNatTransPath (funExt₂ (λ _ _ → Σ≡Prop (λ _ → squash/ _ _) refl)) + + compOpenInclSeq : ∀ {U V W : CompactOpen X} (U≤V : U ≤ V) (V≤W : V ≤ W) + → compOpenIncl (is-trans _ _ _ U≤V V≤W) + ≡ compOpenIncl U≤V ●ᵛ compOpenIncl V≤W + compOpenInclSeq _ _ = makeNatTransPath + (funExt₂ (λ _ _ → Σ≡Prop (λ _ → squash/ _ _) refl)) + + + -- the structure sheaf + private COᵒᵖ = (DistLatticeCategory (CompOpenDistLattice .F-ob X)) ^op + + strDLSh : Functor COᵒᵖ (CommRingsCategory {ℓ = ℓ-suc ℓ}) + F-ob strDLSh U = 𝓞 .F-ob ⟦ U ⟧ᶜᵒ + F-hom strDLSh U≥V = 𝓞 .F-hom (compOpenIncl U≥V) + F-id strDLSh = cong (𝓞 .F-hom) compOpenInclId ∙ 𝓞 .F-id + F-seq strDLSh _ _ = cong (𝓞 .F-hom) (compOpenInclSeq _ _) ∙ 𝓞 .F-seq _ _ + + + -- important lemma + -- Compact opens of Zariski sheaves are sheaves + presLocalCompactOpen : (X : ℤFunctor) (U : CompactOpen X) → isLocal X → isLocal ⟦ U ⟧ᶜᵒ + presLocalCompactOpen X U isLocalX R um@(unimodvec _ f _) = isoToIsEquiv isoU + where + open Coverage zariskiCoverage + open InvertingElementsBase R + instance _ = R .snd + + fᵢCoverR = covers R .snd um + + isoX : Iso (X .F-ob R .fst) (CompatibleFamily X fᵢCoverR) + isoX = equivToIso (elementToCompatibleFamily _ _ , isLocalX R um) + + compatibleFamIncl : (CompatibleFamily ⟦ U ⟧ᶜᵒ fᵢCoverR) → (CompatibleFamily X fᵢCoverR) + compatibleFamIncl fam = (fst ∘ fst fam) + , λ i j B φ ψ φψComm → cong fst (fam .snd i j B φ ψ φψComm) + + compatibleFamIncl≡ : ∀ (y : Σ[ x ∈ X .F-ob R .fst ] U .N-ob R x ≡ D R 1r) + → compatibleFamIncl (elementToCompatibleFamily ⟦ U ⟧ᶜᵒ fᵢCoverR y) + ≡ elementToCompatibleFamily X fᵢCoverR (y .fst) + compatibleFamIncl≡ y = CompatibleFamily≡ _ _ _ _ λ _ → refl + + isoU : Iso (Σ[ x ∈ X .F-ob R .fst ] U .N-ob R x ≡ D R 1r) + (CompatibleFamily ⟦ U ⟧ᶜᵒ fᵢCoverR) + fun isoU = elementToCompatibleFamily _ _ + fst (inv isoU fam) = isoX .inv (compatibleFamIncl fam) + snd (inv isoU fam) = -- U (x) ≡ D(1) + -- knowing that U(x/1)¸≡ D(1) in R[1/fᵢ] + let x = isoX .inv (compatibleFamIncl fam) in + isSeparatedZarLatFun R um (U .N-ob R x) (D R 1r) + λ i → let open UniversalProp (f i) + instance _ = R[1/ (f i) ]AsCommRing .snd in + + inducedZarLatHom /1AsCommRingHom .fst (U .N-ob R x) + + ≡⟨ funExt⁻ (sym (U .N-hom /1AsCommRingHom)) x ⟩ + + U .N-ob R[1/ (f i) ]AsCommRing (X .F-hom /1AsCommRingHom x) + + ≡⟨ cong (U .N-ob R[1/ f i ]AsCommRing) + (funExt⁻ (cong fst (isoX .rightInv (compatibleFamIncl fam))) i) ⟩ + + U .N-ob R[1/ (f i) ]AsCommRing (fam .fst i .fst) + + ≡⟨ fam .fst i .snd ⟩ + + D R[1/ (f i) ]AsCommRing 1r + + ≡⟨ sym (inducedZarLatHom /1AsCommRingHom .snd .pres1) ⟩ + + inducedZarLatHom /1AsCommRingHom .fst (D R 1r) ∎ + + rightInv isoU fam = + Σ≡Prop (λ _ → isPropIsCompatibleFamily _ _ _) + (funExt λ i → Σ≡Prop (λ _ → squash/ _ _) + (funExt⁻ (cong fst + (isoX .rightInv (compatibleFamIncl fam))) i)) + leftInv isoU y = Σ≡Prop (λ _ → squash/ _ _) + (cong (isoX .inv) (compatibleFamIncl≡ y) + ∙ isoX .leftInv (y .fst)) diff --git a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda new file mode 100644 index 0000000000..83dab7931c --- /dev/null +++ b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda @@ -0,0 +1,201 @@ +{- + + Compact open subfunctors of qcqs-schemes are qcqs-schemes (TODO!!!) + The proof proceeds by + 1. Defining standard/basic compact opens of affines and proving that they are affine + 2. Proving that arbitrary compact opens of affines are qcqs-schemes + +-} + +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.AlgebraicGeometry.Functorial.ZFunctors.OpenSubscheme where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.HLevels + + +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ) + +open import Cubical.Data.FinData + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Localisation +open import Cubical.Algebra.Semilattice +open import Cubical.Algebra.Lattice +open import Cubical.Algebra.DistLattice +open import Cubical.Algebra.DistLattice.BigOps + +open import Cubical.AlgebraicGeometry.ZariskiLattice.Base +open import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty +open import Cubical.AlgebraicGeometry.ZariskiLattice.Properties +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.Base +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.CompactOpen +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.QcQsScheme + +open import Cubical.Categories.Category renaming (isIso to isIsoC) +open import Cubical.Categories.Functor + +open import Cubical.Categories.Site.Instances.ZariskiCommRing +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Yoneda + + +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SetQuotients as SQ + +open import Cubical.Relation.Binary.Order.Poset + + +-- standard affine opens +module StandardOpens {ℓ : Level} (R : CommRing ℓ) (f : R .fst) where + + open Iso + open Functor + open NatTrans + open NatIso + open isIsoC + open DistLatticeStr ⦃...⦄ + open CommRingStr ⦃...⦄ + open IsRingHom + open RingHoms + open IsLatticeHom + open ZarLat + + open InvertingElementsBase R + open UniversalProp f + + private module ZL = ZarLatUniversalProp + + private + instance + _ = R .snd + + D : CompactOpen (Sp ⟅ R ⟆) + D = yonedaᴾ ZarLatFun R .inv (ZL.D R f) + + SpR[1/f]≅⟦Df⟧ : NatIso (Sp .F-ob R[1/ f ]AsCommRing) ⟦ D ⟧ᶜᵒ + N-ob (trans SpR[1/f]≅⟦Df⟧) B φ = (φ ∘r /1AsCommRingHom) , ∨lRid _ ∙ path + where + open CommRingHomTheory φ + open IsSupport (ZL.isSupportD B) + instance + _ = B .snd + _ = ZariskiLattice B .snd + + isUnitφ[f/1] : φ .fst (f /1) ∈ B ˣ + isUnitφ[f/1] = RingHomRespInv (f /1) ⦃ S/1⊆S⁻¹Rˣ f ∣ 1 , sym (·IdR f) ∣₁ ⦄ + + path : ZL.D B (φ .fst (f /1)) ≡ 1l + path = supportUnit _ isUnitφ[f/1] + + N-hom (trans SpR[1/f]≅⟦Df⟧) _ = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (RingHom≡ refl) + + inv (nIso SpR[1/f]≅⟦Df⟧ B) (φ , Dφf≡D1) = invElemUniversalProp B φ isUnitφf .fst .fst + where + instance _ = ZariskiLattice B .snd + isUnitφf : φ .fst f ∈ B ˣ + isUnitφf = unitLemmaZarLat B (φ $r f) (sym (∨lRid _) ∙ Dφf≡D1) + + sec (nIso SpR[1/f]≅⟦Df⟧ B) = + funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (RingHom≡ (invElemUniversalProp _ _ _ .fst .snd)) + ret (nIso SpR[1/f]≅⟦Df⟧ B) = + funExt λ φ → cong fst (invElemUniversalProp B (φ ∘r /1AsCommRingHom) _ .snd (φ , refl)) + + isAffineD : isAffineCompactOpen D + isAffineD = ∣ R[1/ f ]AsCommRing , SpR[1/f]≅⟦Df⟧ ∣₁ + + +-- compact opens of affine schemes are qcqs-schemes +module _ {ℓ : Level} (R : CommRing ℓ) (W : CompactOpen (Sp ⟅ R ⟆)) where + + open StandardOpens + + open Iso + open Functor + open NatTrans + open NatIso + open isIsoC + open DistLatticeStr ⦃...⦄ + open CommRingStr ⦃...⦄ + open PosetStr ⦃...⦄ + open IsRingHom + open RingHoms + open IsLatticeHom + open ZarLat + + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (CompOpenDistLattice .F-ob (Sp .F-ob R)))) using (IndPoset; ind≤bigOp) + open InvertingElementsBase R + open Join + open JoinMap + open AffineCover + private module ZL = ZarLatUniversalProp + + private + instance + _ = R .snd + _ = ZariskiLattice R .snd + _ = CompOpenDistLattice .F-ob (Sp .F-ob R) .snd + _ = CompOpenDistLattice .F-ob ⟦ W ⟧ᶜᵒ .snd + _ = IndPoset .snd + + w : ZL R + w = yonedaᴾ ZarLatFun R .fun W + + -- yoneda is a lattice homomorphsim + isHomYoneda : IsLatticeHom (DistLattice→Lattice (ZariskiLattice R) .snd) + (yonedaᴾ ZarLatFun R .inv) + (DistLattice→Lattice (CompOpenDistLattice ⟅ Sp ⟅ R ⟆ ⟆) .snd) + pres0 isHomYoneda = makeNatTransPath (funExt₂ (λ _ _ → refl)) + pres1 isHomYoneda = + makeNatTransPath (funExt₂ (λ _ φ → inducedZarLatHom φ .snd .pres1)) + pres∨l isHomYoneda u v = + makeNatTransPath (funExt₂ (λ _ φ → inducedZarLatHom φ .snd .pres∨l u v)) + pres∧l isHomYoneda u v = + makeNatTransPath (funExt₂ (λ _ φ → inducedZarLatHom φ .snd .pres∧l u v)) + + module _ {n : ℕ} + (f : FinVec (fst R) n) + (⋁Df≡W : ⋁ (CompOpenDistLattice ⟅ Sp ⟅ R ⟆ ⟆) (D R ∘ f) ≡ W) where + + Df≤W : ∀ i → D R (f i) ≤ W + Df≤W i = subst (D R (f i) ≤_) ⋁Df≡W (ind≤bigOp (D R ∘ f) i) + + toAffineCover : AffineCover ⟦ W ⟧ᶜᵒ + AffineCover.n toAffineCover = n + U toAffineCover i = compOpenDownHom (Sp ⟅ R ⟆) W .fst (D R (f i)) + covers toAffineCover = sym (pres⋁ (compOpenDownHom (Sp ⟅ R ⟆) W) (D R ∘ f)) + ∙ cong (compOpenDownHom (Sp ⟅ R ⟆) W .fst) ⋁Df≡W + ∙ makeNatTransPath (funExt₂ (λ _ → snd)) + isAffineU toAffineCover i = + ∣ _ , seqNatIso (SpR[1/f]≅⟦Df⟧ R (f i)) (compOpenDownHomNatIso _ (Df≤W i)) ∣₁ + + module _ {n : ℕ} + (f : FinVec (fst R) n) + (⋁Df≡w : ⋁ (ZariskiLattice R) (ZL.D R ∘ f) ≡ w) where + + private + ⋁Df≡W : ⋁ (CompOpenDistLattice ⟅ Sp ⟅ R ⟆ ⟆) (D R ∘ f) ≡ W + ⋁Df≡W = sym (pres⋁ (_ , isHomYoneda) (ZL.D R ∘ f)) + ∙ cong (yonedaᴾ ZarLatFun R .inv) ⋁Df≡w + ∙ yonedaᴾ ZarLatFun R .leftInv W + + makeAffineCoverCompOpenOfAffine : AffineCover ⟦ W ⟧ᶜᵒ + makeAffineCoverCompOpenOfAffine = toAffineCover f ⋁Df≡W + + hasAffineCoverCompOpenOfAffine : hasAffineCover ⟦ W ⟧ᶜᵒ + hasAffineCoverCompOpenOfAffine = PT.map truncHelper ([]surjective w) + where + truncHelper : Σ[ n,f ∈ Σ ℕ (FinVec (fst R)) ] [ n,f ] ≡ w → AffineCover ⟦ W ⟧ᶜᵒ + truncHelper ((n , f) , [n,f]≡w) = makeAffineCoverCompOpenOfAffine f (ZL.⋁D≡ R f ∙ [n,f]≡w) + + isQcQsSchemeCompOpenOfAffine : isQcQsScheme ⟦ W ⟧ᶜᵒ + fst isQcQsSchemeCompOpenOfAffine = presLocalCompactOpen _ _ (isSubcanonicalZariskiCoverage R) + snd isQcQsSchemeCompOpenOfAffine = hasAffineCoverCompOpenOfAffine diff --git a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/QcQsScheme.agda b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/QcQsScheme.agda new file mode 100644 index 0000000000..d8f9a77c27 --- /dev/null +++ b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/QcQsScheme.agda @@ -0,0 +1,82 @@ +{- + + A qcqs-scheme is a ℤ-functor that is local (a Zariski-sheaf) + and has an affine cover, where the notion of cover is given + by the lattice structure of compact open subfunctors + +-} + +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.AlgebraicGeometry.Functorial.ZFunctors.QcQsScheme where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.HLevels + + +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ) + +open import Cubical.Data.FinData + +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.BigOps + +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.Base +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.CompactOpen + +open import Cubical.Categories.Category renaming (isIso to isIsoC) +open import Cubical.Categories.Functor +open import Cubical.Categories.Site.Instances.ZariskiCommRing + +open import Cubical.HITs.PropositionalTruncation as PT + + +module _ {ℓ : Level} (X : ℤFunctor {ℓ = ℓ}) where + open Functor + open DistLatticeStr ⦃...⦄ + open Join (CompOpenDistLattice .F-ob X) + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (CompOpenDistLattice .F-ob X))) + private instance _ = (CompOpenDistLattice .F-ob X) .snd + + + record AffineCover : Type (ℓ-suc ℓ) where + field + n : ℕ + U : FinVec (CompactOpen X) n + covers : ⋁ U ≡ 1l -- TODO: equivalent to X ≡ ⟦ ⋁ U ⟧ᶜᵒ + isAffineU : ∀ i → isAffineCompactOpen (U i) + + hasAffineCover : Type (ℓ-suc ℓ) + hasAffineCover = ∥ AffineCover ∥₁ + + +module _ {ℓ : Level} where + -- definition of quasi-compact, quasi-separated schemes + isQcQsScheme : ℤFunctor → Type (ℓ-suc ℓ) + isQcQsScheme X = isLocal X × hasAffineCover X + + -- affine schemes are qcqs-schemes + module _ (A : CommRing ℓ) where + open AffineCover + open DistLatticeStr ⦃...⦄ + private instance _ = (CompOpenDistLattice ⟅ Sp ⟅ A ⟆ ⟆) .snd + + -- the canonical one element affine cover of a representable + singlAffineCover : AffineCover (Sp ⟅ A ⟆) + n singlAffineCover = 1 + U singlAffineCover zero = 1l + covers singlAffineCover = ∨lRid _ + isAffineU singlAffineCover zero = ∣ A , compOpenTopNatIso (Sp ⟅ A ⟆) ∣₁ + + isQcQsSchemeAffine : isQcQsScheme (Sp ⟅ A ⟆) + fst isQcQsSchemeAffine = isSubcanonicalZariskiCoverage A + snd isQcQsSchemeAffine = ∣ singlAffineCover ∣₁ diff --git a/Cubical/Algebra/ZariskiLattice/Base.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/Base.agda similarity index 99% rename from Cubical/Algebra/ZariskiLattice/Base.agda rename to Cubical/AlgebraicGeometry/ZariskiLattice/Base.agda index 06bb4f7909..273a85f069 100644 --- a/Cubical/Algebra/ZariskiLattice/Base.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/Base.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.ZariskiLattice.Base where +module Cubical.AlgebraicGeometry.ZariskiLattice.Base where open import Cubical.Foundations.Prelude diff --git a/Cubical/Algebra/ZariskiLattice/Properties.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/Properties.agda similarity index 97% rename from Cubical/Algebra/ZariskiLattice/Properties.agda rename to Cubical/AlgebraicGeometry/ZariskiLattice/Properties.agda index ba9813b866..3fafe859b3 100644 --- a/Cubical/Algebra/ZariskiLattice/Properties.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/Properties.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.ZariskiLattice.Properties where +module Cubical.AlgebraicGeometry.ZariskiLattice.Properties where open import Cubical.Foundations.Prelude @@ -25,8 +25,8 @@ open import Cubical.Algebra.Lattice open import Cubical.Algebra.DistLattice open import Cubical.Algebra.DistLattice.Downset -open import Cubical.Algebra.ZariskiLattice.Base -open import Cubical.Algebra.ZariskiLattice.UniversalProperty +open import Cubical.AlgebraicGeometry.ZariskiLattice.Base +open import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty open import Cubical.HITs.SetQuotients as SQ import Cubical.HITs.PropositionalTruncation as PT diff --git a/Cubical/Algebra/ZariskiLattice/StructureSheaf.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda similarity index 98% rename from Cubical/Algebra/ZariskiLattice/StructureSheaf.agda rename to Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda index b70970886c..0ca7597380 100644 --- a/Cubical/Algebra/ZariskiLattice/StructureSheaf.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda @@ -8,7 +8,7 @@ -} {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.ZariskiLattice.StructureSheaf where +module Cubical.AlgebraicGeometry.ZariskiLattice.StructureSheaf where open import Cubical.Foundations.Prelude @@ -56,8 +56,8 @@ 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.ZariskiLattice.Base -open import Cubical.Algebra.ZariskiLattice.UniversalProperty +open import Cubical.AlgebraicGeometry.ZariskiLattice.Base +open import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty open import Cubical.Categories.Category.Base hiding (_[_,_]) open import Cubical.Categories.Functor diff --git a/Cubical/Algebra/ZariskiLattice/StructureSheafPullback.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda similarity index 98% rename from Cubical/Algebra/ZariskiLattice/StructureSheafPullback.agda rename to Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda index 61bc82976c..41695c786a 100644 --- a/Cubical/Algebra/ZariskiLattice/StructureSheafPullback.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda @@ -9,7 +9,7 @@ -} {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.ZariskiLattice.StructureSheafPullback where +module Cubical.AlgebraicGeometry.ZariskiLattice.StructureSheafPullback where open import Cubical.Foundations.Prelude @@ -59,8 +59,8 @@ 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.ZariskiLattice.Base -open import Cubical.Algebra.ZariskiLattice.UniversalProperty +open import Cubical.AlgebraicGeometry.ZariskiLattice.Base +open import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty open import Cubical.Categories.Category.Base hiding (_[_,_]) open import Cubical.Categories.Functor diff --git a/Cubical/Algebra/ZariskiLattice/UniversalProperty.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/UniversalProperty.agda similarity index 99% rename from Cubical/Algebra/ZariskiLattice/UniversalProperty.agda rename to Cubical/AlgebraicGeometry/ZariskiLattice/UniversalProperty.agda index 4355fe899c..53b4fc4f53 100644 --- a/Cubical/Algebra/ZariskiLattice/UniversalProperty.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/UniversalProperty.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.ZariskiLattice.UniversalProperty where +module Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty where open import Cubical.Foundations.Prelude @@ -40,7 +40,7 @@ open import Cubical.Algebra.DistLattice.Basis open import Cubical.Algebra.DistLattice.BigOps open import Cubical.Algebra.Matrix -open import Cubical.Algebra.ZariskiLattice.Base +open import Cubical.AlgebraicGeometry.ZariskiLattice.Base open import Cubical.HITs.SetQuotients as SQ open import Cubical.HITs.PropositionalTruncation as PT diff --git a/Cubical/Categories/Instances/ZFunctors.agda b/Cubical/Categories/Instances/ZFunctors.agda index faec14ebef..5aef61581d 100644 --- a/Cubical/Categories/Instances/ZFunctors.agda +++ b/Cubical/Categories/Instances/ZFunctors.agda @@ -1,795 +1,24 @@ -{- - - The impredicative way to define functorial qcqs-schemes (over Spec(ℤ)) - --} -{-# OPTIONS --safe --lossy-unification #-} +{-# OPTIONS --safe #-} module Cubical.Categories.Instances.ZFunctors where -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Function -open import Cubical.Foundations.Powerset -open import Cubical.Foundations.HLevels - - -open import Cubical.Functions.FunExtEquiv - -open import Cubical.Data.Sigma -open import Cubical.Data.Nat using (ℕ) - -open import Cubical.Data.FinData - -open import Cubical.Algebra.Ring -open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Localisation -open import Cubical.Algebra.CommRing.RadicalIdeal -open import Cubical.Algebra.Semilattice -open import Cubical.Algebra.Lattice -open import Cubical.Algebra.DistLattice -open import Cubical.Algebra.DistLattice.BigOps -open import Cubical.Algebra.ZariskiLattice.Base -open import Cubical.Algebra.ZariskiLattice.UniversalProperty -open import Cubical.Algebra.ZariskiLattice.Properties - -open import Cubical.Categories.Category renaming (isIso to isIsoC) -open import Cubical.Categories.Functor -open import Cubical.Categories.Instances.Sets -open import Cubical.Categories.Instances.CommRings -open import Cubical.Categories.Instances.DistLattice -open import Cubical.Categories.Instances.DistLattices -open import Cubical.Categories.Instances.Functors -open import Cubical.Categories.Site.Cover -open import Cubical.Categories.Site.Coverage -open import Cubical.Categories.Site.Sheaf -open import Cubical.Categories.Site.Instances.ZariskiCommRing -open import Cubical.Categories.NaturalTransformation -open import Cubical.Categories.Yoneda - - -open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.HITs.SetQuotients as SQ - -open import Cubical.Relation.Binary.Order.Poset - -open Category hiding (_∘_) renaming (_⋆_ to _⋆c_) - -private - variable - ℓ ℓ' ℓ'' : Level - - -module _ {ℓ : Level} where - - open Functor - open NatTrans - open CommRingStr ⦃...⦄ - open IsRingHom - - -- using the naming conventions of Demazure & Gabriel - ℤFunctor = Functor (CommRingsCategory {ℓ = ℓ}) (SET ℓ) - ℤFUNCTOR = FUNCTOR (CommRingsCategory {ℓ = ℓ}) (SET ℓ) - - -- Yoneda in the notation of Demazure & Gabriel, - -- uses that double op is original category definitionally - Sp : Functor (CommRingsCategory {ℓ = ℓ} ^op) ℤFUNCTOR - Sp = YO {C = (CommRingsCategory {ℓ = ℓ} ^op)} - - -- the forgetful functor - -- aka the affine line - -- (aka the representable of ℤ[x]) - 𝔸¹ : ℤFunctor - 𝔸¹ = ForgetfulCommRing→Set - - -- the global sections functor - 𝓞 : Functor ℤFUNCTOR (CommRingsCategory {ℓ = ℓ-suc ℓ} ^op) - fst (F-ob 𝓞 X) = X ⇒ 𝔸¹ - - -- ring struncture induced by internal ring object 𝔸¹ - N-ob (CommRingStr.0r (snd (F-ob 𝓞 X))) A _ = 0r - where instance _ = A .snd - N-hom (CommRingStr.0r (snd (F-ob 𝓞 X))) φ = funExt λ _ → sym (φ .snd .pres0) - - N-ob (CommRingStr.1r (snd (F-ob 𝓞 X))) A _ = 1r - where instance _ = A .snd - N-hom (CommRingStr.1r (snd (F-ob 𝓞 X))) φ = funExt λ _ → sym (φ .snd .pres1) - - N-ob ((snd (F-ob 𝓞 X) CommRingStr.+ α) β) A x = α .N-ob A x + β .N-ob A x - where instance _ = A .snd - N-hom ((snd (F-ob 𝓞 X) CommRingStr.+ α) β) {x = A} {y = B} φ = funExt path - where - instance - _ = A .snd - _ = B .snd - path : ∀ x → α .N-ob B (X .F-hom φ x) + β .N-ob B (X .F-hom φ x) - ≡ φ .fst (α .N-ob A x + β .N-ob A x) - path x = α .N-ob B (X .F-hom φ x) + β .N-ob B (X .F-hom φ x) - ≡⟨ cong₂ _+_ (funExt⁻ (α .N-hom φ) x) (funExt⁻ (β .N-hom φ) x) ⟩ - φ .fst (α .N-ob A x) + φ .fst (β .N-ob A x) - ≡⟨ sym (φ .snd .pres+ _ _) ⟩ - φ .fst (α .N-ob A x + β .N-ob A x) ∎ - - N-ob ((snd (F-ob 𝓞 X) CommRingStr.· α) β) A x = α .N-ob A x · β .N-ob A x - where instance _ = A .snd - N-hom ((snd (F-ob 𝓞 X) CommRingStr.· α) β) {x = A} {y = B} φ = funExt path - where - instance - _ = A .snd - _ = B .snd - path : ∀ x → α .N-ob B (X .F-hom φ x) · β .N-ob B (X .F-hom φ x) - ≡ φ .fst (α .N-ob A x · β .N-ob A x) - path x = α .N-ob B (X .F-hom φ x) · β .N-ob B (X .F-hom φ x) - ≡⟨ cong₂ _·_ (funExt⁻ (α .N-hom φ) x) (funExt⁻ (β .N-hom φ) x) ⟩ - φ .fst (α .N-ob A x) · φ .fst (β .N-ob A x) - ≡⟨ sym (φ .snd .pres· _ _) ⟩ - φ .fst (α .N-ob A x · β .N-ob A x) ∎ - - N-ob ((CommRingStr.- snd (F-ob 𝓞 X)) α) A x = - α .N-ob A x - where instance _ = A .snd - N-hom ((CommRingStr.- snd (F-ob 𝓞 X)) α) {x = A} {y = B} φ = funExt path - where - instance - _ = A .snd - _ = B .snd - path : ∀ x → - α .N-ob B (X .F-hom φ x) ≡ φ .fst (- α .N-ob A x) - path x = - α .N-ob B (X .F-hom φ x) ≡⟨ cong -_ (funExt⁻ (α .N-hom φ) x) ⟩ - - φ .fst (α .N-ob A x) ≡⟨ sym (φ .snd .pres- _) ⟩ - φ .fst (- α .N-ob A x) ∎ - - CommRingStr.isCommRing (snd (F-ob 𝓞 X)) = makeIsCommRing - isSetNatTrans - (λ _ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.+Assoc _ _ _)) - (λ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.+IdR _)) - (λ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.+InvR _)) - (λ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.+Comm _ _)) - (λ _ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.·Assoc _ _ _)) - (λ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.·IdR _)) - (λ _ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.·DistR+ _ _ _)) - (λ _ _ → makeNatTransPath (funExt₂ λ A _ → A .snd .CommRingStr.·Comm _ _)) - - -- action on natural transformations - fst (F-hom 𝓞 α) = α ●ᵛ_ - pres0 (snd (F-hom 𝓞 α)) = makeNatTransPath (funExt₂ λ _ _ → refl) - pres1 (snd (F-hom 𝓞 α)) = makeNatTransPath (funExt₂ λ _ _ → refl) - pres+ (snd (F-hom 𝓞 α)) _ _ = makeNatTransPath (funExt₂ λ _ _ → refl) - pres· (snd (F-hom 𝓞 α)) _ _ = makeNatTransPath (funExt₂ λ _ _ → refl) - pres- (snd (F-hom 𝓞 α)) _ = makeNatTransPath (funExt₂ λ _ _ → refl) - - -- functoriality of 𝓞 - F-id 𝓞 = RingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) - F-seq 𝓞 _ _ = RingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) - - - --- There is an adjunction 𝓞 ⊣ᵢ Sp --- (relative to the inclusion i : CommRing ℓ → CommRing (ℓ+1)) --- between the "global sections functor" 𝓞 --- and the fully-faithful embedding of affines Sp, --- whose counit is a natural isomorphism -module AdjBij where - - open Functor - open NatTrans - open Iso - open IsRingHom - - private module _ {A : CommRing ℓ} {X : ℤFunctor {ℓ}} where - _♭ : CommRingHom A (𝓞 .F-ob X) → X ⇒ Sp .F-ob A - fst (N-ob (φ ♭) B x) a = φ .fst a .N-ob B x - - pres0 (snd (N-ob (φ ♭) B x)) = cong (λ y → y .N-ob B x) (φ .snd .pres0) - pres1 (snd (N-ob (φ ♭) B x)) = cong (λ y → y .N-ob B x) (φ .snd .pres1) - pres+ (snd (N-ob (φ ♭) B x)) _ _ = cong (λ y → y .N-ob B x) (φ .snd .pres+ _ _) - pres· (snd (N-ob (φ ♭) B x)) _ _ = cong (λ y → y .N-ob B x) (φ .snd .pres· _ _) - pres- (snd (N-ob (φ ♭) B x)) _ = cong (λ y → y .N-ob B x) (φ .snd .pres- _) - - N-hom (φ ♭) ψ = funExt (λ x → RingHom≡ (funExt λ a → funExt⁻ (φ .fst a .N-hom ψ) x)) - - - -- the other direction is just precomposition modulo Yoneda - _♯ : X ⇒ Sp .F-ob A → CommRingHom A (𝓞 .F-ob X) - fst (α ♯) a = α ●ᵛ yonedaᴾ 𝔸¹ A .inv a - - pres0 (snd (α ♯)) = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres0) - pres1 (snd (α ♯)) = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres1) - pres+ (snd (α ♯)) _ _ = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres+ _ _) - pres· (snd (α ♯)) _ _ = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres· _ _) - pres- (snd (α ♯)) _ = makeNatTransPath (funExt₂ λ B x → α .N-ob B x .snd .pres- _) - - - -- the two maps are inverse to each other - ♭♯Id : ∀ (α : X ⇒ Sp .F-ob A) → ((α ♯) ♭) ≡ α - ♭♯Id _ = makeNatTransPath (funExt₂ λ _ _ → RingHom≡ (funExt (λ _ → refl))) - - ♯♭Id : ∀ (φ : CommRingHom A (𝓞 .F-ob X)) → ((φ ♭) ♯) ≡ φ - ♯♭Id _ = RingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) - - - -- we get a relative adjunction 𝓞 ⊣ᵢ Sp - -- with respect to the inclusion i : CommRing ℓ → CommRing (ℓ+1) - 𝓞⊣SpIso : {A : CommRing ℓ} {X : ℤFunctor {ℓ}} - → Iso (CommRingHom A (𝓞 .F-ob X)) (X ⇒ Sp .F-ob A) - fun 𝓞⊣SpIso = _♭ - inv 𝓞⊣SpIso = _♯ - rightInv 𝓞⊣SpIso = ♭♯Id - leftInv 𝓞⊣SpIso = ♯♭Id - - 𝓞⊣SpNatℤFunctor : {A : CommRing ℓ} {X Y : ℤFunctor {ℓ}} (α : X ⇒ Sp .F-ob A) (β : Y ⇒ X) - → (β ●ᵛ α) ♯ ≡ (𝓞 .F-hom β) ∘cr (α ♯) - 𝓞⊣SpNatℤFunctor _ _ = RingHom≡ (funExt (λ _ → makeNatTransPath (funExt₂ (λ _ _ → refl)))) - - 𝓞⊣SpNatCommRing : {X : ℤFunctor {ℓ}} {A B : CommRing ℓ} - (φ : CommRingHom A (𝓞 .F-ob X)) (ψ : CommRingHom B A) - → (φ ∘cr ψ) ♭ ≡ (φ ♭) ●ᵛ Sp .F-hom ψ - 𝓞⊣SpNatCommRing _ _ = makeNatTransPath (funExt₂ λ _ _ → RingHom≡ (funExt (λ _ → refl))) - - -- the counit is an equivalence - private - ε : (A : CommRing ℓ) → CommRingHom A ((𝓞 ∘F Sp) .F-ob A) - ε A = (idTrans (Sp .F-ob A)) ♯ - - 𝓞⊣SpCounitEquiv : (A : CommRing ℓ) → CommRingEquiv A ((𝓞 ∘F Sp) .F-ob A) - fst (𝓞⊣SpCounitEquiv A) = isoToEquiv theIso - where - theIso : Iso (A .fst) ((𝓞 ∘F Sp) .F-ob A .fst) - fun theIso = ε A .fst - inv theIso = yonedaᴾ 𝔸¹ A .fun - rightInv theIso α = ℤFUNCTOR .⋆IdL _ ∙ yonedaᴾ 𝔸¹ A .leftInv α - leftInv theIso a = path -- I get yellow otherwise - where - path : yonedaᴾ 𝔸¹ A .fun ((idTrans (Sp .F-ob A)) ●ᵛ yonedaᴾ 𝔸¹ A .inv a) ≡ a - path = cong (yonedaᴾ 𝔸¹ A .fun) (ℤFUNCTOR .⋆IdL _) ∙ yonedaᴾ 𝔸¹ A .rightInv a - snd (𝓞⊣SpCounitEquiv A) = ε A .snd - - --- Affine schemes -module _ {ℓ : Level} where - open Functor - - isAffine : (X : ℤFunctor) → Type (ℓ-suc ℓ) - isAffine X = ∃[ A ∈ CommRing ℓ ] NatIso (Sp .F-ob A) X - - -- TODO: 𝔸¹ ≅ Sp ℤ[x] and 𝔾ₘ ≅ Sp ℤ[x,x⁻¹] ≅ D(x) ↪ 𝔸¹ as first examples of affine schemes - - --- The unit is an equivalence iff the ℤ-functor is affine. --- Unfortunately, we can't give a natural transformation --- X ⇒ Sp (𝓞 X), because the latter ℤ-functor lives in a higher universe. --- We can however give terms that look just like the unit, --- giving us an alternative def. of affine ℤ-functors -private module AffineDefs {ℓ : Level} where - open Functor - open NatTrans - open Iso - open IsRingHom - η : (X : ℤFunctor) (A : CommRing ℓ) → X .F-ob A .fst → CommRingHom (𝓞 .F-ob X) A - fst (η X A x) α = α .N-ob A x - pres0 (snd (η X A x)) = refl - pres1 (snd (η X A x)) = refl - pres+ (snd (η X A x)) _ _ = refl - pres· (snd (η X A x)) _ _ = refl - pres- (snd (η X A x)) _ = refl - - -- this is basically a natural transformation - ηObHom : (X : ℤFunctor) {A B : CommRing ℓ} (φ : CommRingHom A B) - → η X B ∘ (X .F-hom φ) ≡ (φ ∘cr_) ∘ η X A - ηObHom X φ = funExt (λ x → RingHom≡ (funExt λ α → funExt⁻ (α .N-hom φ) x)) - - -- can only state equality on object part, but that would be enough - ηHom : {X Y : ℤFunctor} (α : X ⇒ Y) (A : CommRing ℓ) (x : X .F-ob A .fst) - → η Y A (α .N-ob A x) ≡ η X A x ∘cr 𝓞 .F-hom α - ηHom _ _ _ = RingHom≡ refl - - isAffine' : (X : ℤFunctor) → Type (ℓ-suc ℓ) - isAffine' X = ∀ (A : CommRing ℓ) → isEquiv (η X A) - -- TODO: isAffine → isAffine' - - --- compact opens and affine covers -module _ {ℓ : Level} where - - open Iso - open Functor - open NatTrans - open NatIso - open DistLatticeStr ⦃...⦄ - open CommRingStr ⦃...⦄ - open IsRingHom - open IsLatticeHom - open ZarLat - open ZarLatUniversalProp - - -- the Zariski lattice functor classifying compact open subobjects - ZarLatFun : ℤFunctor {ℓ = ℓ} - F-ob ZarLatFun A = ZL A , SQ.squash/ - F-hom ZarLatFun φ = inducedZarLatHom φ .fst - F-id ZarLatFun {A} = cong fst (inducedZarLatHomId A) - F-seq ZarLatFun φ ψ = cong fst (inducedZarLatHomSeq φ ψ) - - -- this is a separated presheaf - -- (TODO: prove this a sheaf) - isSeparatedZarLatFun : isSeparated zariskiCoverage ZarLatFun - isSeparatedZarLatFun A (unimodvec n f 1∈⟨f₁,⋯,fₙ⟩) u w uRest≡wRest = - u ≡⟨ sym (∧lLid _) ⟩ - 1l ∧l u ≡⟨ congL _∧l_ D1≡⋁Dfᵢ ⟩ - (⋁ (D A ∘ f)) ∧l u ≡⟨ ⋁Meetldist _ _ ⟩ - ⋁ (λ i → D A (f i) ∧l u) ≡⟨ ⋁Ext Dfᵢ∧u≡Dfᵢ∧w ⟩ - ⋁ (λ i → D A (f i) ∧l w) ≡⟨ sym (⋁Meetldist _ _) ⟩ - (⋁ (D A ∘ f)) ∧l w ≡⟨ congL _∧l_ (sym D1≡⋁Dfᵢ) ⟩ - 1l ∧l w ≡⟨ ∧lLid _ ⟩ - w ∎ - where - open Join (ZariskiLattice A) - open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (ZariskiLattice A))) - using (IndPoset) - open LatticeTheory (DistLattice→Lattice (ZariskiLattice A)) - open PosetStr (IndPoset .snd) - open IsSupport (isSupportD A) - open RadicalIdeal A - instance - _ = A .snd - _ = ZariskiLattice A .snd - - D1≡⋁Dfᵢ : 1l ≡ ⋁ (D A ∘ f) - D1≡⋁Dfᵢ = is-antisym _ _ - (supportRadicalIneq f 1r (∈→∈√ _ _ 1∈⟨f₁,⋯,fₙ⟩)) - (1lRightAnnihilates∨l _) - - Dfᵢ∧u≡Dfᵢ∧w : ∀ i → D A (f i) ∧l u ≡ D A (f i) ∧l w - Dfᵢ∧u≡Dfᵢ∧w i = - D A (f i) ∧l u - ≡⟨ sym (cong fst (funExt⁻ (cong fst toLocToDown≡ToDown) u)) ⟩ - locToDownHom .fst (inducedZarLatHom /1AsCommRingHom .fst u) .fst - ≡⟨ cong (λ x → locToDownHom .fst x .fst) (uRest≡wRest i) ⟩ - locToDownHom .fst (inducedZarLatHom /1AsCommRingHom .fst w) .fst - ≡⟨ cong fst (funExt⁻ (cong fst toLocToDown≡ToDown) w) ⟩ - D A (f i) ∧l w ∎ - where - open InvertingElementsBase.UniversalProp A (f i) - open LocDownSetIso A (f i) - - CompactOpen : ℤFunctor → Type (ℓ-suc ℓ) - CompactOpen X = X ⇒ ZarLatFun - - -- the induced subfunctor - ⟦_⟧ᶜᵒ : {X : ℤFunctor} (U : CompactOpen X) → ℤFunctor - F-ob (⟦_⟧ᶜᵒ {X = X} U) A = (Σ[ x ∈ X .F-ob A .fst ] U .N-ob A x ≡ D A 1r) - , isSetΣSndProp (X .F-ob A .snd) λ _ → squash/ _ _ - where instance _ = snd A - F-hom (⟦_⟧ᶜᵒ {X = X} U) {x = A} {y = B} φ (x , Ux≡D1) = (X .F-hom φ x) , path - where - instance - _ = A .snd - _ = B .snd - open IsLatticeHom - path : U .N-ob B (X .F-hom φ x) ≡ D B 1r - path = U .N-ob B (X .F-hom φ x) ≡⟨ funExt⁻ (U .N-hom φ) x ⟩ - ZarLatFun .F-hom φ (U .N-ob A x) ≡⟨ cong (ZarLatFun .F-hom φ) Ux≡D1 ⟩ - ZarLatFun .F-hom φ (D A 1r) ≡⟨ inducedZarLatHom φ .snd .pres1 ⟩ - D B 1r ∎ - F-id (⟦_⟧ᶜᵒ {X = X} U) = funExt (λ x → Σ≡Prop (λ _ → squash/ _ _) - (funExt⁻ (X .F-id) (x .fst))) - F-seq (⟦_⟧ᶜᵒ {X = X} U) φ ψ = funExt (λ x → Σ≡Prop (λ _ → squash/ _ _) - (funExt⁻ (X .F-seq φ ψ) (x .fst))) - - - isAffineCompactOpen : {X : ℤFunctor} (U : CompactOpen X) → Type (ℓ-suc ℓ) - isAffineCompactOpen U = isAffine ⟦ U ⟧ᶜᵒ - - -- TODO: define basic opens D(f) ↪ Sp A and prove ⟦ D(f) ⟧ᶜᵒ ≅ Sp A[1/f] - - -- the (big) dist. lattice of compact opens - CompOpenDistLattice : Functor ℤFUNCTOR (DistLatticesCategory {ℓ = ℓ-suc ℓ} ^op) - fst (F-ob CompOpenDistLattice X) = CompactOpen X - - -- lattice structure induce by internal lattice object ZarLatFun - N-ob (DistLatticeStr.0l (snd (F-ob CompOpenDistLattice X))) A _ = 0l - where instance _ = ZariskiLattice A .snd - N-hom (DistLatticeStr.0l (snd (F-ob CompOpenDistLattice X))) _ = funExt λ _ → refl - - N-ob (DistLatticeStr.1l (snd (F-ob CompOpenDistLattice X))) A _ = 1l - where instance _ = ZariskiLattice A .snd - N-hom (DistLatticeStr.1l (snd (F-ob CompOpenDistLattice X))) {x = A} {y = B} φ = funExt λ _ → path - where - instance - _ = A .snd - _ = B .snd - _ = ZariskiLattice B .snd - path : D B 1r ≡ D B (φ .fst 1r) ∨l 0l - path = cong (D B) (sym (φ .snd .pres1)) ∙ sym (∨lRid _) - - N-ob ((snd (F-ob CompOpenDistLattice X) DistLatticeStr.∨l U) V) A x = U .N-ob A x ∨l V .N-ob A x - where instance _ = ZariskiLattice A .snd - N-hom ((snd (F-ob CompOpenDistLattice X) DistLatticeStr.∨l U) V) {x = A} {y = B} φ = funExt path - where - instance - _ = ZariskiLattice A .snd - _ = ZariskiLattice B .snd - path : ∀ x → U .N-ob B (X .F-hom φ x) ∨l V .N-ob B (X .F-hom φ x) - ≡ ZarLatFun .F-hom φ (U .N-ob A x ∨l V .N-ob A x) - path x = U .N-ob B (X .F-hom φ x) ∨l V .N-ob B (X .F-hom φ x) - ≡⟨ cong₂ _∨l_ (funExt⁻ (U .N-hom φ) x) (funExt⁻ (V .N-hom φ) x) ⟩ - ZarLatFun .F-hom φ (U .N-ob A x) ∨l ZarLatFun .F-hom φ (V .N-ob A x) - ≡⟨ sym (inducedZarLatHom φ .snd .pres∨l _ _) ⟩ - ZarLatFun .F-hom φ (U .N-ob A x ∨l V .N-ob A x) ∎ - - N-ob ((snd (F-ob CompOpenDistLattice X) DistLatticeStr.∧l U) V) A x = U .N-ob A x ∧l V .N-ob A x - where instance _ = ZariskiLattice A .snd - N-hom ((snd (F-ob CompOpenDistLattice X) DistLatticeStr.∧l U) V) {x = A} {y = B} φ = funExt path - where - instance - _ = ZariskiLattice A .snd - _ = ZariskiLattice B .snd - path : ∀ x → U .N-ob B (X .F-hom φ x) ∧l V .N-ob B (X .F-hom φ x) - ≡ ZarLatFun .F-hom φ (U .N-ob A x ∧l V .N-ob A x) - path x = U .N-ob B (X .F-hom φ x) ∧l V .N-ob B (X .F-hom φ x) - ≡⟨ cong₂ _∧l_ (funExt⁻ (U .N-hom φ) x) (funExt⁻ (V .N-hom φ) x) ⟩ - ZarLatFun .F-hom φ (U .N-ob A x) ∧l ZarLatFun .F-hom φ (V .N-ob A x) - ≡⟨ sym (inducedZarLatHom φ .snd .pres∧l _ _) ⟩ - ZarLatFun .F-hom φ (U .N-ob A x ∧l V .N-ob A x) ∎ - - DistLatticeStr.isDistLattice (snd (F-ob CompOpenDistLattice X)) = makeIsDistLattice∧lOver∨l - isSetNatTrans - (λ _ _ _ → makeNatTransPath (funExt₂ - (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∨lAssoc _ _ _))) - (λ _ → makeNatTransPath (funExt₂ (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∨lRid _))) - (λ _ _ → makeNatTransPath (funExt₂ (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∨lComm _ _))) - (λ _ _ _ → makeNatTransPath (funExt₂ - (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∧lAssoc _ _ _))) - (λ _ → makeNatTransPath (funExt₂ (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∧lRid _))) - (λ _ _ → makeNatTransPath (funExt₂ (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∧lComm _ _))) - (λ _ _ → makeNatTransPath (funExt₂ -- don't know why ∧lAbsorb∨l doesn't work - (λ A _ → ZariskiLattice A .snd .DistLatticeStr.absorb _ _ .snd))) - (λ _ _ _ → makeNatTransPath (funExt₂ -- same here - (λ A _ → ZariskiLattice A .snd .DistLatticeStr.∧l-dist-∨l _ _ _ .fst))) - - -- (contravariant) action on morphisms - fst (F-hom CompOpenDistLattice α) = α ●ᵛ_ - pres0 (snd (F-hom CompOpenDistLattice α)) = makeNatTransPath (funExt₂ λ _ _ → refl) - pres1 (snd (F-hom CompOpenDistLattice α)) = makeNatTransPath (funExt₂ λ _ _ → refl) - pres∨l (snd (F-hom CompOpenDistLattice α)) _ _ = makeNatTransPath (funExt₂ λ _ _ → refl) - pres∧l (snd (F-hom CompOpenDistLattice α)) _ _ = makeNatTransPath (funExt₂ λ _ _ → refl) - - -- functoriality - F-id CompOpenDistLattice = LatticeHom≡f _ _ - (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) - F-seq CompOpenDistLattice _ _ = LatticeHom≡f _ _ - (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) - - - module _ (X : ℤFunctor) where - open isIsoC - private instance _ = (CompOpenDistLattice .F-ob X) .snd - - compOpenTopNatIso : NatIso X ⟦ 1l ⟧ᶜᵒ - N-ob (trans compOpenTopNatIso) _ φ = φ , refl - N-hom (trans compOpenTopNatIso) _ = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) refl - inv (nIso compOpenTopNatIso B) = fst - sec (nIso compOpenTopNatIso B) = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) refl - ret (nIso compOpenTopNatIso B) = funExt λ _ → refl - - - module _ (X : ℤFunctor) where - open isIsoC - open Join (CompOpenDistLattice .F-ob X) - open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (CompOpenDistLattice .F-ob X))) - open PosetStr (IndPoset .snd) hiding (_≤_) - open LatticeTheory ⦃...⦄ - private instance _ = (CompOpenDistLattice .F-ob X) .snd - - compOpenGlobalIncl : (U : CompactOpen X) → ⟦ U ⟧ᶜᵒ ⇒ X - N-ob (compOpenGlobalIncl U) A = fst - N-hom (compOpenGlobalIncl U) φ = refl - - compOpenIncl : {U V : CompactOpen X} → V ≤ U → ⟦ V ⟧ᶜᵒ ⇒ ⟦ U ⟧ᶜᵒ - N-ob (compOpenIncl {U = U} {V = V} V≤U) A (x , Vx≡D1) = x , path - where - instance - _ = A .snd - _ = ZariskiLattice A .snd - _ = DistLattice→Lattice (ZariskiLattice A) - path : U .N-ob A x ≡ D A 1r - path = U .N-ob A x ≡⟨ funExt⁻ (funExt⁻ (cong N-ob (sym V≤U)) A) x ⟩ - V .N-ob A x ∨l U .N-ob A x ≡⟨ cong (_∨l U .N-ob A x) Vx≡D1 ⟩ - D A 1r ∨l U .N-ob A x ≡⟨ 1lLeftAnnihilates∨l _ ⟩ - D A 1r ∎ - N-hom (compOpenIncl V≤U) φ = funExt λ x → Σ≡Prop (λ _ → squash/ _ _) refl - - -- this is essentially U∧_ - compOpenDownHom : (U : CompactOpen X) - → DistLatticeHom (CompOpenDistLattice .F-ob X) - (CompOpenDistLattice .F-ob ⟦ U ⟧ᶜᵒ) - compOpenDownHom U = CompOpenDistLattice .F-hom (compOpenGlobalIncl U) - - module _ {U V : CompactOpen X} (V≤U : V ≤ U) where - -- We need this separate definition to avoid termination checker issues, - -- but we don't understand why. - private - compOpenDownHomFun : (A : CommRing ℓ) - → ⟦ V ⟧ᶜᵒ .F-ob A .fst - → ⟦ compOpenDownHom U .fst V ⟧ᶜᵒ .F-ob A .fst - compOpenDownHomFun A v = (compOpenIncl V≤U ⟦ A ⟧) v , snd v - - compOpenDownHomNatIso : NatIso ⟦ V ⟧ᶜᵒ ⟦ compOpenDownHom U .fst V ⟧ᶜᵒ - N-ob (trans compOpenDownHomNatIso) = compOpenDownHomFun - N-hom (trans compOpenDownHomNatIso) _ = - funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (Σ≡Prop (λ _ → squash/ _ _) refl) - inv (nIso compOpenDownHomNatIso A) ((x , Ux≡D1) , Vx≡D1) = x , Vx≡D1 - sec (nIso compOpenDownHomNatIso A) = - funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (Σ≡Prop (λ _ → squash/ _ _) refl) - ret (nIso compOpenDownHomNatIso A) = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) refl - - compOpenInclId : ∀ {U : CompactOpen X} → compOpenIncl (is-refl U) ≡ idTrans ⟦ U ⟧ᶜᵒ - compOpenInclId = makeNatTransPath (funExt₂ (λ _ _ → Σ≡Prop (λ _ → squash/ _ _) refl)) - - compOpenInclSeq : ∀ {U V W : CompactOpen X} (U≤V : U ≤ V) (V≤W : V ≤ W) - → compOpenIncl (is-trans _ _ _ U≤V V≤W) - ≡ compOpenIncl U≤V ●ᵛ compOpenIncl V≤W - compOpenInclSeq _ _ = makeNatTransPath - (funExt₂ (λ _ _ → Σ≡Prop (λ _ → squash/ _ _) refl)) - - - -- the structure sheaf - private COᵒᵖ = (DistLatticeCategory (CompOpenDistLattice .F-ob X)) ^op - - strDLSh : Functor COᵒᵖ (CommRingsCategory {ℓ = ℓ-suc ℓ}) - F-ob strDLSh U = 𝓞 .F-ob ⟦ U ⟧ᶜᵒ - F-hom strDLSh U≥V = 𝓞 .F-hom (compOpenIncl U≥V) - F-id strDLSh = cong (𝓞 .F-hom) compOpenInclId ∙ 𝓞 .F-id - F-seq strDLSh _ _ = cong (𝓞 .F-hom) (compOpenInclSeq _ _) ∙ 𝓞 .F-seq _ _ - - - -- def. affine cover and locality for definition of qcqs-scheme - module _ (X : ℤFunctor) where - open isIsoC - open Join (CompOpenDistLattice .F-ob X) - open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (CompOpenDistLattice .F-ob X))) - open PosetStr (IndPoset .snd) hiding (_≤_) - open LatticeTheory ⦃...⦄ - private instance _ = (CompOpenDistLattice .F-ob X) .snd - - record AffineCover : Type (ℓ-suc ℓ) where - field - n : ℕ - U : FinVec (CompactOpen X) n - covers : ⋁ U ≡ 1l -- TODO: equivalent to X ≡ ⟦ ⋁ U ⟧ᶜᵒ - isAffineU : ∀ i → isAffineCompactOpen (U i) - - hasAffineCover : Type (ℓ-suc ℓ) - hasAffineCover = ∥ AffineCover ∥₁ - - -- qcqs-schemes as Zariski sheaves (local ℤ-functors) with an affine cover in the sense above - isLocal : ℤFunctor → Type (ℓ-suc ℓ) - isLocal X = isSheaf zariskiCoverage X - - -- Compact opens of Zariski sheaves are sheaves - presLocalCompactOpen : (X : ℤFunctor) (U : CompactOpen X) → isLocal X → isLocal ⟦ U ⟧ᶜᵒ - presLocalCompactOpen X U isLocalX R um@(unimodvec _ f _) = isoToIsEquiv isoU - where - open Coverage zariskiCoverage - open InvertingElementsBase R - instance _ = R .snd - - fᵢCoverR = covers R .snd um - - isoX : Iso (X .F-ob R .fst) (CompatibleFamily X fᵢCoverR) - isoX = equivToIso (elementToCompatibleFamily _ _ , isLocalX R um) - - compatibleFamIncl : (CompatibleFamily ⟦ U ⟧ᶜᵒ fᵢCoverR) → (CompatibleFamily X fᵢCoverR) - compatibleFamIncl fam = (fst ∘ fst fam) - , λ i j B φ ψ φψComm → cong fst (fam .snd i j B φ ψ φψComm) - - compatibleFamIncl≡ : ∀ (y : Σ[ x ∈ X .F-ob R .fst ] U .N-ob R x ≡ D R 1r) - → compatibleFamIncl (elementToCompatibleFamily ⟦ U ⟧ᶜᵒ fᵢCoverR y) - ≡ elementToCompatibleFamily X fᵢCoverR (y .fst) - compatibleFamIncl≡ y = CompatibleFamily≡ _ _ _ _ λ _ → refl - - isoU : Iso (Σ[ x ∈ X .F-ob R .fst ] U .N-ob R x ≡ D R 1r) - (CompatibleFamily ⟦ U ⟧ᶜᵒ fᵢCoverR) - fun isoU = elementToCompatibleFamily _ _ - fst (inv isoU fam) = isoX .inv (compatibleFamIncl fam) - snd (inv isoU fam) = -- U (x) ≡ D(1) - -- knowing that U(x/1)¸≡ D(1) in R[1/fᵢ] - let x = isoX .inv (compatibleFamIncl fam) in - isSeparatedZarLatFun R um (U .N-ob R x) (D R 1r) - λ i → let open UniversalProp (f i) - instance _ = R[1/ (f i) ]AsCommRing .snd in - - inducedZarLatHom /1AsCommRingHom .fst (U .N-ob R x) - - ≡⟨ funExt⁻ (sym (U .N-hom /1AsCommRingHom)) x ⟩ - - U .N-ob R[1/ (f i) ]AsCommRing (X .F-hom /1AsCommRingHom x) - - ≡⟨ cong (U .N-ob R[1/ f i ]AsCommRing) - (funExt⁻ (cong fst (isoX .rightInv (compatibleFamIncl fam))) i) ⟩ - - U .N-ob R[1/ (f i) ]AsCommRing (fam .fst i .fst) - - ≡⟨ fam .fst i .snd ⟩ - - D R[1/ (f i) ]AsCommRing 1r - - ≡⟨ sym (inducedZarLatHom /1AsCommRingHom .snd .pres1) ⟩ - - inducedZarLatHom /1AsCommRingHom .fst (D R 1r) ∎ - - rightInv isoU fam = - Σ≡Prop (λ _ → isPropIsCompatibleFamily _ _ _) - (funExt λ i → Σ≡Prop (λ _ → squash/ _ _) - (funExt⁻ (cong fst - (isoX .rightInv (compatibleFamIncl fam))) i)) - leftInv isoU y = Σ≡Prop (λ _ → squash/ _ _) - (cong (isoX .inv) (compatibleFamIncl≡ y) - ∙ isoX .leftInv (y .fst)) - - - -- definition of quasi-compact, quasi-separated schemes - isQcQsScheme : ℤFunctor → Type (ℓ-suc ℓ) - isQcQsScheme X = isLocal X × hasAffineCover X - - - -- affine schemes are qcqs-schemes - module _ (A : CommRing ℓ) where - open AffineCover - private instance _ = (CompOpenDistLattice ⟅ Sp ⟅ A ⟆ ⟆) .snd - - -- the canonical one element affine cover of a representable - singlAffineCover : AffineCover (Sp .F-ob A) - n singlAffineCover = 1 - U singlAffineCover zero = 1l - covers singlAffineCover = ∨lRid _ - isAffineU singlAffineCover zero = ∣ A , compOpenTopNatIso (Sp ⟅ A ⟆) ∣₁ - - isQcQsSchemeAffine : isQcQsScheme (Sp .F-ob A) - fst isQcQsSchemeAffine = isSubcanonicalZariskiCoverage A - snd isQcQsSchemeAffine = ∣ singlAffineCover ∣₁ - - --- standard affine opens --- TODO: separate file? -module StandardOpens {ℓ : Level} (R : CommRing ℓ) (f : R .fst) where - - open Iso - open Functor - open NatTrans - open NatIso - open isIsoC - open DistLatticeStr ⦃...⦄ - open CommRingStr ⦃...⦄ - open IsRingHom - open RingHoms - open IsLatticeHom - open ZarLat - - open InvertingElementsBase R - open UniversalProp f - - private module ZL = ZarLatUniversalProp - - private - instance - _ = R .snd - - D : CompactOpen (Sp ⟅ R ⟆) - D = yonedaᴾ ZarLatFun R .inv (ZL.D R f) - - SpR[1/f]≅⟦Df⟧ : NatIso (Sp .F-ob R[1/ f ]AsCommRing) ⟦ D ⟧ᶜᵒ - N-ob (trans SpR[1/f]≅⟦Df⟧) B φ = (φ ∘r /1AsCommRingHom) , ∨lRid _ ∙ path - where - open CommRingHomTheory φ - open IsSupport (ZL.isSupportD B) - instance - _ = B .snd - _ = ZariskiLattice B .snd - - isUnitφ[f/1] : φ .fst (f /1) ∈ B ˣ - isUnitφ[f/1] = RingHomRespInv (f /1) ⦃ S/1⊆S⁻¹Rˣ f ∣ 1 , sym (·IdR f) ∣₁ ⦄ - - path : ZL.D B (φ .fst (f /1)) ≡ 1l - path = supportUnit _ isUnitφ[f/1] - - N-hom (trans SpR[1/f]≅⟦Df⟧) _ = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (RingHom≡ refl) - - inv (nIso SpR[1/f]≅⟦Df⟧ B) (φ , Dφf≡D1) = invElemUniversalProp B φ isUnitφf .fst .fst - where - instance _ = ZariskiLattice B .snd - isUnitφf : φ .fst f ∈ B ˣ - isUnitφf = unitLemmaZarLat B (φ $r f) (sym (∨lRid _) ∙ Dφf≡D1) - - sec (nIso SpR[1/f]≅⟦Df⟧ B) = - funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (RingHom≡ (invElemUniversalProp _ _ _ .fst .snd)) - ret (nIso SpR[1/f]≅⟦Df⟧ B) = - funExt λ φ → cong fst (invElemUniversalProp B (φ ∘r /1AsCommRingHom) _ .snd (φ , refl)) - - isAffineD : isAffineCompactOpen D - isAffineD = ∣ R[1/ f ]AsCommRing , SpR[1/f]≅⟦Df⟧ ∣₁ - - --- compact opens of affine schemes are qcqs-schemes -module _ {ℓ : Level} (R : CommRing ℓ) (W : CompactOpen (Sp ⟅ R ⟆)) where - - open StandardOpens - - open Iso - open Functor - open NatTrans - open NatIso - open isIsoC - open DistLatticeStr ⦃...⦄ - open CommRingStr ⦃...⦄ - open PosetStr ⦃...⦄ - open IsRingHom - open RingHoms - open IsLatticeHom - open ZarLat - - open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (CompOpenDistLattice .F-ob (Sp .F-ob R)))) using (IndPoset; ind≤bigOp) - open InvertingElementsBase R - open Join - open JoinMap - open AffineCover - private module ZL = ZarLatUniversalProp - - private - instance - _ = R .snd - _ = ZariskiLattice R .snd - _ = CompOpenDistLattice .F-ob (Sp .F-ob R) .snd - _ = CompOpenDistLattice .F-ob ⟦ W ⟧ᶜᵒ .snd - _ = IndPoset .snd - - w : ZL R - w = yonedaᴾ ZarLatFun R .fun W - - -- yoneda is a lattice homomorphsim - isHomYoneda : IsLatticeHom (DistLattice→Lattice (ZariskiLattice R) .snd) - (yonedaᴾ ZarLatFun R .inv) - (DistLattice→Lattice (CompOpenDistLattice ⟅ Sp ⟅ R ⟆ ⟆) .snd) - pres0 isHomYoneda = makeNatTransPath (funExt₂ (λ _ _ → refl)) - pres1 isHomYoneda = - makeNatTransPath (funExt₂ (λ _ φ → inducedZarLatHom φ .snd .pres1)) - pres∨l isHomYoneda u v = - makeNatTransPath (funExt₂ (λ _ φ → inducedZarLatHom φ .snd .pres∨l u v)) - pres∧l isHomYoneda u v = - makeNatTransPath (funExt₂ (λ _ φ → inducedZarLatHom φ .snd .pres∧l u v)) - - module _ {n : ℕ} - (f : FinVec (fst R) n) - (⋁Df≡W : ⋁ (CompOpenDistLattice ⟅ Sp ⟅ R ⟆ ⟆) (D R ∘ f) ≡ W) where +{- - Df≤W : ∀ i → D R (f i) ≤ W - Df≤W i = subst (D R (f i) ≤_) ⋁Df≡W (ind≤bigOp (D R ∘ f) i) + The category of ℤ-functors is the category of functors from rings to sets. + It contains the category of schemes as a full subcategory and is thus of + special interest to algebraic geometers. The basic notions needed to + develop algebraic geometry using ℤ-functors can be found in: - toAffineCover : AffineCover ⟦ W ⟧ᶜᵒ - AffineCover.n toAffineCover = n - U toAffineCover i = compOpenDownHom (Sp ⟅ R ⟆) W .fst (D R (f i)) - covers toAffineCover = sym (pres⋁ (compOpenDownHom (Sp ⟅ R ⟆) W) (D R ∘ f)) - ∙ cong (compOpenDownHom (Sp ⟅ R ⟆) W .fst) ⋁Df≡W - ∙ makeNatTransPath (funExt₂ (λ _ → snd)) - isAffineU toAffineCover i = - ∣ _ , seqNatIso (SpR[1/f]≅⟦Df⟧ R (f i)) (compOpenDownHomNatIso _ (Df≤W i)) ∣₁ +-} - module _ {n : ℕ} - (f : FinVec (fst R) n) - (⋁Df≡w : ⋁ (ZariskiLattice R) (ZL.D R ∘ f) ≡ w) where +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.Base public - private - ⋁Df≡W : ⋁ (CompOpenDistLattice ⟅ Sp ⟅ R ⟆ ⟆) (D R ∘ f) ≡ W - ⋁Df≡W = sym (pres⋁ (_ , isHomYoneda) (ZL.D R ∘ f)) - ∙ cong (yonedaᴾ ZarLatFun R .inv) ⋁Df≡w - ∙ yonedaᴾ ZarLatFun R .leftInv W +{- - makeAffineCoverCompOpenOfAffine : AffineCover ⟦ W ⟧ᶜᵒ - makeAffineCoverCompOpenOfAffine = toAffineCover f ⋁Df≡W + In Cubical Agda we can define the full subcategory of qcqs-schemes. + This is done in: - hasAffineCoverCompOpenOfAffine : hasAffineCover ⟦ W ⟧ᶜᵒ - hasAffineCoverCompOpenOfAffine = PT.map truncHelper ([]surjective w) - where - truncHelper : Σ[ n,f ∈ Σ ℕ (FinVec (fst R)) ] [ n,f ] ≡ w → AffineCover ⟦ W ⟧ᶜᵒ - truncHelper ((n , f) , [n,f]≡w) = makeAffineCoverCompOpenOfAffine f (ZL.⋁D≡ R f ∙ [n,f]≡w) +-} - isQcQsSchemeCompOpenOfAffine : isQcQsScheme ⟦ W ⟧ᶜᵒ - fst isQcQsSchemeCompOpenOfAffine = presLocalCompactOpen _ _ (isSubcanonicalZariskiCoverage R) - snd isQcQsSchemeCompOpenOfAffine = hasAffineCoverCompOpenOfAffine +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.CompactOpen +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.QcQsScheme +open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.OpenSubscheme diff --git a/Cubical/Papers/AffineSchemes.agda b/Cubical/Papers/AffineSchemes.agda index 2f63ba1f7b..5d679863bd 100644 --- a/Cubical/Papers/AffineSchemes.agda +++ b/Cubical/Papers/AffineSchemes.agda @@ -24,55 +24,55 @@ module Cubical.Papers.AffineSchemes where -- 2: Background -- 2.2: Cubical Agda -import Cubical.Foundations.Prelude as Prelude -import Cubical.Foundations.HLevels as HLevels -import Cubical.Foundations.Univalence as Univalence -import Cubical.Data.Sigma as Sigma -import Cubical.HITs.PropositionalTruncation as PT -import Cubical.Algebra.DistLattice.Basis as DistLatticeBasis -import Cubical.HITs.SetQuotients as SQ +import Cubical.Foundations.Prelude as Prelude +import Cubical.Foundations.HLevels as HLevels +import Cubical.Foundations.Univalence as Univalence +import Cubical.Data.Sigma as Sigma +import Cubical.HITs.PropositionalTruncation as PT +import Cubical.Algebra.DistLattice.Basis as DistLatticeBasis +import Cubical.HITs.SetQuotients as SQ -- 3: Commutative Algebra -- 3.1: Localization -import Cubical.Algebra.CommRing.Localisation.Base as L +import Cubical.Algebra.CommRing.Localisation.Base as L module Localization = L.Loc -import Cubical.Algebra.CommRing.Localisation.UniversalProperty as LocalizationUnivProp -import Cubical.Algebra.CommRing.Localisation.InvertingElements as LocalizationInvEl -import Cubical.Algebra.CommAlgebra as R-Algs -import Cubical.Algebra.CommAlgebra.Localisation as LocalizationR-Alg +import Cubical.Algebra.CommRing.Localisation.UniversalProperty as LocalizationUnivProp +import Cubical.Algebra.CommRing.Localisation.InvertingElements as LocalizationInvEl +import Cubical.Algebra.CommAlgebra as R-Algs +import Cubical.Algebra.CommAlgebra.Localisation as LocalizationR-Alg -- 3.2: The Zariski Lattice -import Cubical.Data.FinData.Base as FiniteTypes -import Cubical.Algebra.Matrix as Matrices -import Cubical.Algebra.CommRing.FGIdeal as FinGenIdeals +import Cubical.Data.FinData.Base as FiniteTypes +import Cubical.Algebra.Matrix as Matrices +import Cubical.Algebra.CommRing.FGIdeal as FinGenIdeals -import Cubical.Algebra.ZariskiLattice.Base as ZLB +import Cubical.AlgebraicGeometry.ZariskiLattice.Base as ZLB module ZariskiLatDef = ZLB.ZarLat -import Cubical.Algebra.ZariskiLattice.UniversalProperty as ZLUP +import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty as ZLUP module ZariskiLatUnivProp = ZLUP.ZarLatUniversalProp -- 4: Category Theory -- background theory not explicitly mentioned -import Cubical.Categories.Category.Base as CatTheory -import Cubical.Categories.Limits as GeneralLimits -import Cubical.Categories.Limits.RightKan as GeneralKanExtension +import Cubical.Categories.Category.Base as CatTheory +import Cubical.Categories.Limits as GeneralLimits +import Cubical.Categories.Limits.RightKan as GeneralKanExtension -import Cubical.Categories.DistLatticeSheaf.Diagram as SheafDiagShapes -import Cubical.Categories.DistLatticeSheaf.Base as Sheaves +import Cubical.Categories.DistLatticeSheaf.Diagram as SheafDiagShapes +import Cubical.Categories.DistLatticeSheaf.Base as Sheaves -import Cubical.Categories.DistLatticeSheaf.Extension as E +import Cubical.Categories.DistLatticeSheaf.Extension as E module SheafExtension = E.PreSheafExtension -- 5: The Structure Sheaf -import Cubical.Categories.Instances.CommAlgebras as R-AlgConstructions -import Cubical.Algebra.CommRing.Localisation.Limit as LocalizationLimit -import Cubical.Algebra.ZariskiLattice.StructureSheaf as StructureSheaf +import Cubical.Categories.Instances.CommAlgebras as R-AlgConstructions +import Cubical.Algebra.CommRing.Localisation.Limit as LocalizationLimit +import Cubical.AlgebraicGeometry.ZariskiLattice.StructureSheaf as StructureSheaf diff --git a/Cubical/Papers/FunctorialQcQsSchemes.agda b/Cubical/Papers/FunctorialQcQsSchemes.agda index 6abba9a036..86c727c96b 100644 --- a/Cubical/Papers/FunctorialQcQsSchemes.agda +++ b/Cubical/Papers/FunctorialQcQsSchemes.agda @@ -21,47 +21,47 @@ module Cubical.Papers.FunctorialQcQsSchemes where -- 2: Background -- 2.1: Univalent type theory in Cubical Agda -import Cubical.Foundations.Prelude as Prelude -import Cubical.Foundations.HLevels as HLevels -import Cubical.Foundations.Univalence as Univalence -import Cubical.Data.Sigma as Sigma -import Cubical.HITs.PropositionalTruncation as PT -import Cubical.HITs.SetQuotients as SQ +import Cubical.Foundations.Prelude as Prelude +import Cubical.Foundations.HLevels as HLevels +import Cubical.Foundations.Univalence as Univalence +import Cubical.Data.Sigma as Sigma +import Cubical.HITs.PropositionalTruncation as PT +import Cubical.HITs.SetQuotients as SQ -- 2.2: Localizations and the Zariski lattice -import Cubical.Algebra.CommRing.Localisation.InvertingElements as LocalizationInvEl +import Cubical.Algebra.CommRing.Localisation.InvertingElements as LocalizationInvEl module LocalizationInvElBase = LocalizationInvEl.InvertingElementsBase module LocalizationInvElUniversalProp = LocalizationInvElBase.UniversalProp -import Cubical.Algebra.ZariskiLattice.Base as ZLB +import Cubical.AlgebraicGeometry.ZariskiLattice.Base as ZLB module ZariskiLatDef = ZLB.ZarLat -import Cubical.Algebra.ZariskiLattice.UniversalProperty as ZLUP +import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty as ZLUP module ZariskiLatUnivProp = ZLUP.ZarLatUniversalProp module Localization&Radicals = LocalizationInvEl.RadicalLemma -import Cubical.Algebra.ZariskiLattice.Properties as ZLP +import Cubical.AlgebraicGeometry.ZariskiLattice.Properties as ZLP -- 3: ℤ-functors -import Cubical.Categories.Instances.ZFunctors as ZFun +import Cubical.AlgebraicGeometry.Functorial.ZFunctors.Base as ZFun module RelativeAdjunction = ZFun.AdjBij -- 4: Local ℤ-functors -import Cubical.Categories.Site.Cover as Cover -import Cubical.Categories.Site.Coverage as Coverage -import Cubical.Categories.Site.Sheaf as Sheaf +import Cubical.Categories.Site.Cover as Cover +import Cubical.Categories.Site.Coverage as Coverage +import Cubical.Categories.Site.Sheaf as Sheaf -import Cubical.Categories.Site.Instances.ZariskiCommRing as ZariskiCoverage +import Cubical.Categories.Site.Instances.ZariskiCommRing as ZariskiCoverage module ZarCovSubcanonical = ZariskiCoverage.SubcanonicalLemmas -import Cubical.Algebra.CommRing.Localisation.Limit as LocalizationLimit +import Cubical.Algebra.CommRing.Localisation.Limit as LocalizationLimit --- !!! note: the ZFunctors file is supposed to be broken up into smaller files !!! -- 5: Compact opens and qcqs-schemes --- import Cubical.Categories.Instances.ZFunctors as ZFun +import Cubical.AlgebraicGeometry.Functorial.ZFunctors.CompactOpen as CO +import Cubical.AlgebraicGeometry.Functorial.ZFunctors.QcQsScheme as QcQsSchemes -- 6: Open subschemes --- import Cubical.Categories.Instances.ZFunctors as ZFun -module StandardOpen = ZFun.StandardOpens +import Cubical.AlgebraicGeometry.Functorial.ZFunctors.OpenSubscheme as OSubscheme +module StandardOpen = OSubscheme.StandardOpens @@ -160,32 +160,32 @@ open ZariskiCoverage using (isSubcanonicalZariskiCoverage) ---------- 4: Compact opens and qcqs-schemes ---------- -- Definition 13 -open ZFun renaming (ZarLatFun to 𝓛) +open CO renaming (ZarLatFun to 𝓛) -- Definition 14 -open ZFun using (CompactOpen ; ⟦_⟧ᶜᵒ) +open CO using (CompactOpen ; ⟦_⟧ᶜᵒ) -- Definition 15 -open ZFun using (CompOpenDistLattice) +open CO using (CompOpenDistLattice) -- Definition 16 -open ZFun using (isQcQsScheme) +open QcQsSchemes using (isQcQsScheme) -- Proposition 17 -open ZFun using (singlAffineCover ; isQcQsSchemeAffine) +open QcQsSchemes using (singlAffineCover ; isQcQsSchemeAffine) -- Remark 18 -open ZFun using (AffineCover ; hasAffineCover) +open QcQsSchemes using (AffineCover ; hasAffineCover) ----------- 5: Open subschemes ---------- -- Lemma 20 -open ZFun using (isSeparatedZarLatFun) +open CO using (isSeparatedZarLatFun) -- Lemma 21 -open ZFun using (presLocalCompactOpen) +open CO using (presLocalCompactOpen) -- Definition 22 open StandardOpen using (D) @@ -194,4 +194,4 @@ open StandardOpen using (D) open StandardOpen using (SpR[1/f]≅⟦Df⟧ ; isAffineD) -- Theorem 24 -open ZFun using (isQcQsSchemeCompOpenOfAffine) +open OSubscheme using (isQcQsSchemeCompOpenOfAffine) diff --git a/Cubical/README.agda b/Cubical/README.agda index ed23331c73..b707a2e03a 100644 --- a/Cubical/README.agda +++ b/Cubical/README.agda @@ -68,6 +68,9 @@ import Cubical.ZCohomology.Everything -- Algebra library (in development) import Cubical.Algebra.Everything +-- Algebraic geometry +import Cubical.AlgebraicGeometry.Everything + -- Various talks import Cubical.Talks.Everything From cb0328e3e7bc4d2dbf375fa8a78f86e78b35bd32 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Mon, 22 Apr 2024 14:45:41 +0530 Subject: [PATCH 04/30] Preorder structure on the category of subobjects (#1115) * Update Powerset module * Make repeated definition private * Remove unnecessary using * Sync with agda/cubical * Add preorder structure on the category of subobjects --- .../Categories/Constructions/Slice/Base.agda | 16 ++++++ .../Categories/Constructions/SubObject.agda | 56 +++++++++++++++++++ 2 files changed, 72 insertions(+) diff --git a/Cubical/Categories/Constructions/Slice/Base.agda b/Cubical/Categories/Constructions/Slice/Base.agda index e2cfc8d81e..c11652b7cf 100644 --- a/Cubical/Categories/Constructions/Slice/Base.agda +++ b/Cubical/Categories/Constructions/Slice/Base.agda @@ -5,6 +5,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path open import Cubical.Foundations.Transport using (transpFill) open import Cubical.Categories.Category renaming (isIso to isIsoC) @@ -69,6 +70,21 @@ module _ {xf yg : SliceOb} where SOPath≡PathΣ = ua (isoToEquiv SOPathIsoPathΣ) +-- If the type of objects of C forms a set then so does the type of objects of the slice cat +module _ (isSetCOb : isSet (C .ob)) where + isSetSliceOb : isSet SliceOb + isSetSliceOb x y = + subst + (λ t → isProp t) + (sym (SOPath≡PathΣ {xf = x} {yg = y})) + (isPropΣ + (isSetCOb (x .S-ob) (y .S-ob)) + λ x≡y → + subst + (λ t → isProp t) + (sym (ua (PathP≃Path (λ i → C [ x≡y i , c ]) (x .S-arr) (y .S-arr)))) + (C .isSetHom (transport (λ i → C [ x≡y i , c ]) (x .S-arr)) (y .S-arr))) + -- 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) diff --git a/Cubical/Categories/Constructions/SubObject.agda b/Cubical/Categories/Constructions/SubObject.agda index 45b420e4cd..2a5040bd49 100644 --- a/Cubical/Categories/Constructions/SubObject.agda +++ b/Cubical/Categories/Constructions/SubObject.agda @@ -2,6 +2,8 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Powerset open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.Univalence open import Cubical.Data.Sigma @@ -9,6 +11,8 @@ open import Cubical.Categories.Category open import Cubical.Categories.Functor open import Cubical.Categories.Morphism +open import Cubical.Relation.Binary.Order.Preorder + open Category module Cubical.Categories.Constructions.SubObject @@ -33,3 +37,55 @@ subObjMorIsMonic : {s t : SubObjCat .ob} (f : SubObjCat [ s , t ]) subObjMorIsMonic {s = s} {t = t} f = postcompCreatesMonic C (S-hom f) (S-arr (t .fst)) comp-is-monic where comp-is-monic = subst (isMonic C) (sym (S-comm f)) (s .snd) + +isPropSubObjMor : (s t : SubObjCat .ob) → isProp (SubObjCat [ s , t ]) +SliceHom.S-hom + (isPropSubObjMor + (sliceob incS , isMonicIncS) + (sliceob incT , isMonicIncT) + (slicehom x xComm) + (slicehom y yComm) i) = + isMonicIncT {a = x} {a' = y} (xComm ∙ sym yComm) i +SliceHom.S-comm + (isPropSubObjMor + (sliceob incS , isMonicIncS) + (sliceob incT , isMonicIncT) + (slicehom x xComm) + (slicehom y yComm) i) = + isProp→PathP (λ i → C .isSetHom (seq' C (isMonicIncT (xComm ∙ sym yComm) i) incT) incS) xComm yComm i + +isReflSubObjMor : (x : SubObjCat .ob) → SubObjCat [ x , x ] +SliceHom.S-hom (isReflSubObjMor (sliceob inc , isMonicInc)) = C .id +SliceHom.S-comm (isReflSubObjMor (sliceob inc , isMonicInc)) = C .⋆IdL inc + +isTransSubObjMor : (x y z : SubObjCat .ob) → SubObjCat [ x , y ] → SubObjCat [ y , z ] → SubObjCat [ x , z ] +SliceHom.S-hom + (isTransSubObjMor + (sliceob incX , isMonicIncX) + (sliceob incY , isMonicIncY) + (sliceob incZ , isMonicIncZ) + (slicehom f fComm) + (slicehom g gComm)) = seq' C f g +SliceHom.S-comm + (isTransSubObjMor + (sliceob incX , isMonicIncX) + (sliceob incY , isMonicIncY) + (sliceob incZ , isMonicIncZ) + (slicehom f fComm) + (slicehom g gComm)) = + seq' C (seq' C f g) incZ + ≡⟨ C .⋆Assoc f g incZ ⟩ + seq' C f (seq' C g incZ) + ≡⟨ cong (λ x → seq' C f x) gComm ⟩ + seq' C f incY + ≡⟨ fComm ⟩ + incX + ∎ + +module _ (isSetCOb : isSet (C .ob)) where + subObjCatPreorderStr : PreorderStr _ (SubObjCat .ob) + PreorderStr._≲_ subObjCatPreorderStr x y = SubObjCat [ x , y ] + IsPreorder.is-set (PreorderStr.isPreorder subObjCatPreorderStr) = isSetΣ (isSetSliceOb isSetCOb) λ x → isProp→isSet (∈-isProp isSubObj x) + IsPreorder.is-prop-valued (PreorderStr.isPreorder subObjCatPreorderStr) = isPropSubObjMor + IsPreorder.is-refl (PreorderStr.isPreorder subObjCatPreorderStr) = isReflSubObjMor + IsPreorder.is-trans (PreorderStr.isPreorder subObjCatPreorderStr) = isTransSubObjMor From 36a14f8d9be3b4a0eac1e42bdc521381599bacf6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Axel=20Ljungstr=C3=B6m?= <49276137+aljungstrom@users.noreply.github.com> Date: Mon, 6 May 2024 16:11:25 +0200 Subject: [PATCH 05/30] CW complexes, cellular homology + a lot more (#1111) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * t * duplicate * wups * rme * ganea stuff * w * w * w * fix * stuff * stuff * more * .. * done? * wups * wups * wups * fixes * ugh... * wups * stuff * stuff * stuff * stuf * More * stuff * stuff * stuff * done? * stuff * cleanup * readme * wups * ugh * wups * broken code * ojdå * comments * minor --- .../AbGroup/Instances/FreeAbGroup.agda | 403 ++++++++- Cubical/Algebra/AbGroup/Instances/Pi.agda | 4 + Cubical/Algebra/AbGroup/Properties.agda | 25 +- Cubical/Algebra/ChainComplex.agda | 6 + Cubical/Algebra/ChainComplex/Base.agda | 114 +++ Cubical/Algebra/ChainComplex/Finite.agda | 123 +++ Cubical/Algebra/ChainComplex/Homology.agda | 381 +++++++++ .../Localisation/InvertingElements.agda | 40 +- Cubical/Algebra/Group/Base.agda | 5 + Cubical/Algebra/Group/GroupPath.agda | 109 ++- Cubical/Algebra/Group/Instances/Int.agda | 13 + Cubical/Algebra/Group/Subgroup.agda | 7 + Cubical/Algebra/Group/ZAction.agda | 29 +- Cubical/Axiom/Choice.agda | 61 +- Cubical/CW/Approximation.agda | 692 ++++++++++++++++ Cubical/CW/Base.agda | 156 ++++ Cubical/CW/ChainComplex.agda | 192 +++++ Cubical/CW/Homology.agda | 244 ++++++ Cubical/CW/Homotopy.agda | 766 ++++++++++++++++++ Cubical/CW/Map.agda | 444 ++++++++++ Cubical/CW/Properties.agda | 251 ++++++ Cubical/CW/Subcomplex.agda | 196 +++++ Cubical/Data/Fin/Base.agda | 13 + Cubical/Data/Fin/Inductive.agda | 6 + Cubical/Data/Fin/Inductive/Base.agda | 42 + Cubical/Data/Fin/Inductive/Properties.agda | 136 ++++ Cubical/Data/Fin/Properties.agda | 106 ++- Cubical/Data/FinData/Properties.agda | 2 +- Cubical/Data/FinSequence.agda | 5 + Cubical/Data/FinSequence/Base.agda | 24 + Cubical/Data/FinSequence/Properties.agda | 39 + Cubical/Data/Int/Base.agda | 8 + Cubical/Data/Int/Properties.agda | 11 + Cubical/Data/Nat/Order.agda | 18 + Cubical/Data/Nat/Order/Inductive.agda | 81 ++ Cubical/Data/Sequence.agda | 5 + Cubical/Data/Sequence/Base.agda | 30 + Cubical/Data/Sequence/Properties.agda | 22 + Cubical/Data/Unit/Properties.agda | 21 + Cubical/Foundations/Equiv/Properties.agda | 15 + Cubical/Foundations/Path.agda | 6 + Cubical/Foundations/Pointed/Properties.agda | 11 + Cubical/Foundations/Prelude.agda | 10 + Cubical/HITs/FreeAbGroup/Base.agda | 43 + Cubical/HITs/James/Inductive.agda | 1 + Cubical/HITs/James/Inductive/Base.agda | 1 + .../PropositionalTruncation/Properties.agda | 24 +- Cubical/HITs/Pushout/Properties.agda | 19 + Cubical/HITs/SequentialColimit/Base.agda | 30 +- .../HITs/SequentialColimit/Properties.agda | 279 ++++++- Cubical/HITs/Sn/Base.agda | 5 + Cubical/HITs/Sn/Degree.agda | 163 ++++ Cubical/HITs/Sn/Properties.agda | 124 ++- Cubical/HITs/SphereBouquet.agda | 5 + Cubical/HITs/SphereBouquet/Base.agda | 17 + Cubical/HITs/SphereBouquet/Degree.agda | 491 +++++++++++ Cubical/HITs/SphereBouquet/Properties.agda | 353 ++++++++ Cubical/HITs/Susp/Base.agda | 6 + Cubical/HITs/Susp/Properties.agda | 18 + Cubical/HITs/Truncation/Properties.agda | 8 + Cubical/HITs/Wedge.agda | 1 + Cubical/HITs/Wedge/Properties.agda | 167 ++++ Cubical/Homotopy/Connected.agda | 31 + Cubical/Homotopy/Group/PinSn.agda | 263 +++++- Cubical/README.agda | 3 + Cubical/ZCohomology/GroupStructure.agda | 12 + Cubical/ZCohomology/Groups/Sn.agda | 145 +++- 67 files changed, 6983 insertions(+), 98 deletions(-) create mode 100644 Cubical/Algebra/ChainComplex.agda create mode 100644 Cubical/Algebra/ChainComplex/Base.agda create mode 100644 Cubical/Algebra/ChainComplex/Finite.agda create mode 100644 Cubical/Algebra/ChainComplex/Homology.agda create mode 100644 Cubical/CW/Approximation.agda create mode 100644 Cubical/CW/Base.agda create mode 100644 Cubical/CW/ChainComplex.agda create mode 100644 Cubical/CW/Homology.agda create mode 100644 Cubical/CW/Homotopy.agda create mode 100644 Cubical/CW/Map.agda create mode 100644 Cubical/CW/Properties.agda create mode 100644 Cubical/CW/Subcomplex.agda 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 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/Nat/Order/Inductive.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/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..53f6b48a79 100644 --- a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda +++ b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda @@ -1,9 +1,28 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} 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.Nat.Order.Inductive +open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) +open import Cubical.Data.Fin.Inductive +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 +31,385 @@ 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<ᵗm (subst (_<ᵗ fst y) (sym x₂) x₁)) +... | lt x₁ | gt x₂ = refl +... | eq x₁ | lt x₂ = ⊥.rec (¬m<ᵗm (subst (fst y <ᵗ_) x₁ x₂)) +... | eq x₁ | eq x₂ = refl +... | eq x₁ | gt x₂ = ⊥.rec (¬m<ᵗm (subst (_<ᵗ fst y) x₁ x₂)) +... | gt x₁ | lt x₂ = refl +... | gt x₁ | eq x₂ = ⊥.rec (¬m<ᵗm (subst (_<ᵗ fst x) x₂ x₁)) +... | gt x₁ | gt x₂ = refl + +isGeneratorℤFinGenerator : {n : ℕ} (f : ℤ[Fin n ] .fst) (a : _) + → f a ≡ sumFinℤ {n = n} λ s → f s ·ℤ (ℤFinGenerator s a) +isGeneratorℤFinGenerator {n = suc n} f = + elimFin basec + λ s → cong f (Σ≡Prop (λ _ → isProp<ᵗ) refl) + ∙∙ isGeneratorℤFinGenerator {n = n} (f ∘ injectSuc) s + ∙∙ (λ i → sumFinℤ (λ x → f (injectSuc x) ·ℤ lem₁ s x i)) + ∙∙ +Comm (F s) 0 + ∙∙ λ i → (sym (·Comm (f flast) 0) + ∙ (cong (f flast ·ℤ_) (sym (lem₂ s flast refl)))) i + + F s + where + basec : f flast ≡ sumFinℤ (λ s → f s ·ℤ ℤFinGenerator s flast) + basec with (n ≟ᵗ n) + ... | lt x = ⊥.rec (¬m<ᵗm x) + ... | eq x = λ i → (·Comm (f flast) (pos 1) (~ i)) + lem₂ (~ i) + where + lem₁ : (s : _) → ℤFinGenerator (injectSuc {n = n} s) flast ≡ 0 + lem₁ s with (fst s ≟ᵗ n) + ... | lt x = refl + ... | eq w = ⊥.rec (¬m<ᵗm (subst (fst s <ᵗ_) (sym w) (snd s))) + ... | gt x = refl + + lem₂ : sumFinℤ (λ s → f (injectSuc s) ·ℤ ℤFinGenerator (injectSuc s) flast) ≡ 0 + lem₂ = (λ i → sumFinℤ λ s → (cong (f (injectSuc s) ·ℤ_) (lem₁ s) + ∙ ·Comm (f (injectSuc s)) 0) i) + ∙ (sumFinℤ0 n) + ... | gt x = ⊥.rec (¬m<ᵗm x) + + module _ (a : Fin n) where + F = sumFinGen _+_ (pos 0) (λ x → f (injectSuc x) + ·ℤ (ℤFinGenerator (injectSuc x) (injectSuc a))) + + lem₁ : (x : _) + → ℤFinGenerator {n = n} x a -- (a , diff , cong predℕ p) + ≡ ℤFinGenerator {n = suc n} (injectSuc x) (injectSuc a) -- (a , suc diff , p) + lem₁ x with (fst x ≟ᵗ fst a) + ... | lt x₁ = refl + ... | eq x₁ = refl + ... | gt x₁ = refl + + lem₂ : (k : Fin (suc n)) → fst k ≡ n + → ℤFinGenerator {n = suc n} k (injectSuc a) ≡ 0 + lem₂ k q with (fst k ≟ᵗ fst a) + ... | lt _ = refl + ... | eq x = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) (sym x ∙ q) (snd a))) + ... | gt _ = refl + + +isGeneratorℤFinGenerator' : {n : ℕ} (f : ℤ[Fin n ] .fst) (a : _) + → f a ≡ sumFinℤ {n = n} λ s → (ℤFinGenerator a s) ·ℤ f s +isGeneratorℤFinGenerator' {n = n} f a = + isGeneratorℤFinGenerator f a + ∙ sumFinℤId n λ x → ·Comm (f x) (ℤFinGenerator x a) + ∙ cong (_·ℤ f x) (ℤFinGeneratorComm x a) + +ℤFinGeneratorVanish : (n : ℕ) (x : _) → ℤFinGenerator {n = suc n} flast (injectSuc x) ≡ 0 +ℤFinGeneratorVanish n x with (n ≟ᵗ (fst x)) +... | lt x₁ = refl +... | eq x₁ = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) (sym x₁) (snd x))) +... | gt x₁ = refl + +-- elimination principle +elimℤFin : ∀ {ℓ} {n : ℕ} + → (P : ℤ[Fin (suc n) ] .fst → Type ℓ) + → ((x : _) → P (ℤFinGenerator x)) + → ((f : _) → P f → P (λ x → -ℤ (f x))) + → ((f g : _) → P f → P g → P (λ x → f x + g x)) + → (x : _) → P x +elimℤFin {n = n} P gen d ind f = + subst P (sym (funExt (isGeneratorℤFinGenerator f))) + (P-resp-sumFinℤ n P gen d ind (suc n) (λ x y → f y ·ℤ ℤFinGenerator y x) + λ t → P-resp· n P gen d ind (f t) (ℤFinGenerator t) (gen t)) + where + module _ {ℓ} (n : ℕ) (P : ℤ[Fin (suc n) ] .fst → Type ℓ) + (gens : (x : _) → P (ℤFinGenerator x)) + (ind- : ((f : _) → P f → P (λ x → -ℤ (f x)))) + (ind+ : (f g : _) → P f → P g → P (λ x → f x + g x)) where + P-resp·pos : (x : ℕ) (f : ℤ[Fin (suc n) ] .fst) → P f → P (λ x₁ → pos x ·ℤ f x₁) + P-resp·pos zero f d = + subst P (funExt (λ x → -Cancel (ℤFinGenerator fzero x))) + (ind+ (ℤFinGenerator fzero) + (λ x → -ℤ (ℤFinGenerator fzero x)) + (gens fzero) (ind- _ (gens fzero))) + P-resp·pos (suc zero) f d = d + P-resp·pos (suc (suc x)) f d = + ind+ f (λ a → pos (suc x) ·ℤ f a) d (P-resp·pos (suc x) f d) + + P-resp· : (x : _) (f : ℤ[Fin (suc n) ] .fst) → P f → P (λ x₁ → x ·ℤ f x₁) + P-resp· (pos n) f d = P-resp·pos n f d + P-resp· (negsuc n) f d = + subst P (funExt (λ r → -DistL· (pos (suc n)) (f r))) + (ind- (λ x → pos (suc n) ·ℤ f x) (P-resp·pos (suc n) _ d)) + + P-resp-sumFinℤ : (m : ℕ) (f : Fin (suc n) → Fin m → ℤ) + → ((t : _) → P (λ x → f x t)) + → P (λ t → sumFinℤ {n = m} (f t)) + P-resp-sumFinℤ zero f r = P-resp·pos zero (ℤFinGenerator fzero) (gens fzero) + P-resp-sumFinℤ (suc m) f r = + ind+ (λ t → f t flast) (λ t → sumFinℤ (λ x → f t (injectSuc x))) + (r flast) + (P-resp-sumFinℤ m (λ x y → f x (injectSuc y)) (r ∘ injectSuc)) + + +-- multiplication +·posFree : {n : ℕ} (a : ℕ) (x : Fin n) → FreeAbGroup (Fin n) +·posFree {n = n} zero x = ε +·posFree {n = n} (suc a) x = ⟦ x ⟧ · (·posFree {n = n} a x) + +·Free : {n : ℕ} (a : ℤ) (x : Fin n) → FreeAbGroup (Fin n) +·Free (pos n) x = ·posFree n x +·Free (negsuc n) x = ·posFree (suc n) x ⁻¹ + +·Free⁻¹ : {n : ℕ} (a : ℤ) (x : Fin n) → ·Free (-ℤ a) x ≡ (·Free a x) ⁻¹ +·Free⁻¹ {n = n} (pos zero) x = sym (GroupTheory.inv1g (AbGroup→Group FAGAbGroup)) +·Free⁻¹ {n = n} (pos (suc n₁)) x = refl +·Free⁻¹ {n = n} (negsuc n₁) x = sym (GroupTheory.invInv (AbGroup→Group FAGAbGroup) _) + +·FreeSuc : {n : ℕ} (a : ℤ) (x : Fin n) + → ·Free (sucℤ a) x ≡ ⟦ x ⟧ · ·Free a x +·FreeSuc (pos n) x = refl +·FreeSuc (negsuc zero) x = + sym (cong (⟦ x ⟧ ·_) (cong (_⁻¹) (identityᵣ ⟦ x ⟧)) ∙ invᵣ ⟦ x ⟧) +·FreeSuc (negsuc (suc n)) x = + sym (cong (⟦ x ⟧ ·_) + (GroupTheory.invDistr (AbGroup→Group FAGAbGroup) ⟦ x ⟧ (⟦ x ⟧ · ·posFree n x) + ∙ comm _ _) + ∙∙ assoc _ _ _ + ∙∙ (cong (_· (⟦ x ⟧ · ·posFree n x) ⁻¹) (invᵣ ⟦ x ⟧) + ∙ comm _ _ + ∙ identityᵣ ((⟦ x ⟧ · ·posFree n x) ⁻¹))) + +·FreeHomPos : {n : ℕ} (a : ℕ) (b : ℤ) (x : Fin n) + → ·Free (pos a) x · ·Free b x ≡ ·Free (pos a + b) x +·FreeHomPos zero b x = comm _ _ ∙ identityᵣ (·Free b x) ∙ cong (λ y → ·Free y x) (+Comm b 0) +·FreeHomPos (suc n) b x = + sym (assoc ⟦ x ⟧ (·Free (pos n) x) (·Free b x)) + ∙ cong (⟦ x ⟧ ·_) (·FreeHomPos n b x) + ∙ l b + where + l : (b : ℤ) → ⟦ x ⟧ · ·Free (pos n + b) x ≡ ·Free (pos (suc n) + b) x + l (pos m) = cong (⟦ x ⟧ ·_) (λ i → ·Free (pos+ n m (~ i)) x) + ∙ λ i → ·Free (pos+ (suc n) m i) x + l (negsuc m) = sym (·FreeSuc (pos n +negsuc m) x) + ∙ λ j → ·Free (sucℤ+negsuc m (pos n) j) x + +·FreeHom : {n : ℕ} (a b : ℤ) (x : Fin n) → ·Free a x · ·Free b x ≡ ·Free (a + b) x +·FreeHom (pos n) b x = ·FreeHomPos n b x +·FreeHom (negsuc n) b x = + cong ((⟦ x ⟧ · ·posFree n x) ⁻¹ ·_) + (sym (cong (_⁻¹) (·Free⁻¹ b x) + ∙ GroupTheory.invInv (AbGroup→Group FAGAbGroup) (·Free b x))) + ∙ comm _ _ + ∙∙ sym (GroupTheory.invDistr (AbGroup→Group FAGAbGroup) (·Free (pos (suc n)) x) (·Free (-ℤ b) x)) + ∙∙ cong (_⁻¹) (·FreeHomPos (suc n) (-ℤ b) x) + ∙∙ sym (·Free⁻¹ (pos (suc n) + -ℤ b) x) + ∙∙ λ i → ·Free (help (~ i)) x + where + help : negsuc n + b ≡ -ℤ (pos (suc n) + -ℤ b) + help = cong (negsuc n +_) (sym (-Involutive b)) + ∙ sym (-Dist+ (pos (suc n)) (-ℤ b)) + +sumFinℤFinGenerator≡1 : (n : ℕ) (f : Fin n) + → sumFinGen _·_ ε (λ x → ·Free (ℤFinGenerator f x) x) ≡ ⟦ f ⟧ +sumFinℤFinGenerator≡1 (suc n) = + elimFin (basec n) + indstep + where + basec : (n : ℕ) → sumFinGen _·_ ε (λ x → ·Free (ℤFinGenerator (flast {m = n}) x) x) ≡ ⟦ flast ⟧ + basec n with (n ≟ᵗ n) + ... | lt x = ⊥.rec (¬m<ᵗm x) + ... | eq x = ((λ i → identityᵣ ⟦ flast ⟧ i + · sumFinGen _·_ ε (λ x → ·Free (ℤFinGenerator flast (injectSuc x)) + (injectSuc x))) + ∙ cong (⟦ flast ⟧ ·_) (sumFinGenId _ _ _ _ + (funExt (λ s → cong₂ ·Free (ℤFinGeneratorVanish n s) refl)) + ∙ sumFinGen0 _·_ ε identityᵣ n + (λ x₁ → ·Free (pos 0) (injectSuc x₁)) (λ _ → refl)) ) + ∙ identityᵣ _ + ... | gt x = ⊥.rec (¬m<ᵗm x) + module _ (x : Fin n) where + FR = Free↑ + indstep : + ·Free (ℤFinGenerator (injectSuc x) flast) flast · sumFinGen _·_ ε + (λ x₁ → ·Free (ℤFinGenerator (injectSuc x) (injectSuc x₁)) (injectSuc x₁)) + ≡ ⟦ injectSuc x ⟧ + indstep with (fst x ≟ᵗ n) + ... | lt a = comm _ _ + ∙ identityᵣ _ + ∙ (λ i → sumFinGen _·_ ε (λ x → lem x i)) + ∙ sym (Free↑sumFinℤ n n (λ x₁ → ·Free (ℤFinGenerator x x₁) x₁)) + ∙ cong (Free↑ n) (sumFinℤFinGenerator≡1 n x) + where + lem : (x₁ : Fin n) + → ·Free (ℤFinGenerator (injectSuc x) (injectSuc x₁)) (injectSuc x₁) + ≡ Free↑ n (·Free (ℤFinGenerator x x₁) x₁) + lem x₁ with (fst x ≟ᵗ fst x₁) + ... | lt x = refl + ... | eq x = refl + ... | gt x = refl + ... | eq a = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) a (snd x))) + ... | gt a = ⊥.rec (¬m<ᵗm (<ᵗ-trans a (snd x))) + + +-- equivalence between two versions of free ab group +ℤFin→Free : (n : ℕ) → ℤ[Fin n ] .fst → FreeAbGroup (Fin n) +ℤFin→Free n f = sumFinGen {A = FreeAbGroup (Fin n)} _·_ ε (λ x → ·Free (f x) x) + +Free→ℤFin : (n : ℕ) → FreeAbGroup (Fin n) → ℤ[Fin n ] .fst +Free→ℤFin n ⟦ x ⟧ = ℤFinGenerator x +Free→ℤFin n ε _ = pos 0 +Free→ℤFin n (a · a₁) x = Free→ℤFin n a x + Free→ℤFin n a₁ x +Free→ℤFin n (a ⁻¹) x = -ℤ Free→ℤFin n a x +Free→ℤFin n (assoc a a₁ a₂ i) x = +Assoc (Free→ℤFin n a x) (Free→ℤFin n a₁ x) (Free→ℤFin n a₂ x) i +Free→ℤFin n (comm a a₁ i) x = +Comm (Free→ℤFin n a x) (Free→ℤFin n a₁ x) i +Free→ℤFin n (identityᵣ a i) = Free→ℤFin n a +Free→ℤFin n (invᵣ a i) x = -Cancel (Free→ℤFin n a x) i +Free→ℤFin n (trunc a a₁ x y i j) k = + isSet→isSet' isSetℤ + (λ _ → Free→ℤFin n a k) (λ _ → Free→ℤFin n a₁ k) + (λ j → Free→ℤFin n (x j) k) (λ j → Free→ℤFin n (y j) k) j i + +ℤFin→Free-ℤFinGenerator : (n : ℕ) (x : _) + → ℤFin→Free n (ℤFinGenerator x) ≡ ⟦ x ⟧ +ℤFin→Free-ℤFinGenerator zero x = isContr→isProp isContr-FreeFin0 _ _ +ℤFin→Free-ℤFinGenerator (suc n) f = sumFinℤFinGenerator≡1 _ f + +ℤFin→FreeHom : (n : ℕ) → (f g : ℤ[Fin n ] .fst) + → ℤFin→Free n (λ x → f x + g x) ≡ ℤFin→Free n f · ℤFin→Free n g +ℤFin→FreeHom n f g = + (λ i → sumFinGen _·_ ε (λ x → ·FreeHom (f x) (g x) x (~ i))) + ∙ sumFinGenHom _·_ ε identityᵣ comm assoc n + (λ x → ·Free (f x) x) λ x → ·Free (g x) x + +ℤFin→FreeInv : (n : ℕ) (f : ℤ[Fin n ] .fst) + → ℤFin→Free n (λ x → -ℤ (f x)) ≡ (ℤFin→Free n f) ⁻¹ +ℤFin→FreeInv n f = + (λ i → sumFinGen _·_ ε (λ x → ·Free⁻¹ (f x) x i)) + ∙ sumFinG-neg n {A = FAGAbGroup} (λ x → ·Free (f x) x) + +ℤFin→Free→ℤFin : (n : ℕ) (x : FreeAbGroup (Fin n)) → ℤFin→Free n (Free→ℤFin n x) ≡ x +ℤFin→Free→ℤFin zero x = isContr→isProp (subst (isContr ∘ FreeAbGroup) (sym lem) isContr-Free⊥) _ _ + where + lem : Fin 0 ≡ ⊥ + lem = isoToPath (iso ¬Fin0 (λ{()}) (λ{()}) λ p → ⊥.rec (¬Fin0 p)) +ℤFin→Free→ℤFin (suc n) = + ElimProp.f (trunc _ _) + (ℤFin→Free-ℤFinGenerator (suc n)) + (comm _ _ + ∙∙ identityᵣ _ + ∙∙ sumFinGen0 _·_ ε identityᵣ n (λ _ → ε) (λ _ → refl)) + (λ {x = x} {y = y} p q + → ℤFin→FreeHom (suc n) (Free→ℤFin (suc n) x) (Free→ℤFin (suc n) y) ∙ cong₂ _·_ p q) + λ {x} p → ℤFin→FreeInv (suc n) (Free→ℤFin (suc n) x) ∙ cong (_⁻¹) p + +Free→ℤFin→Free : (n : ℕ) (x : _) → Free→ℤFin n (ℤFin→Free n x) ≡ x +Free→ℤFin→Free zero f = funExt λ x → ⊥.rec (¬Fin0 x) +Free→ℤFin→Free (suc n) = + elimℤFin _ + (λ x → cong (Free→ℤFin (suc n)) (ℤFin→Free-ℤFinGenerator (suc n) x)) + (λ f p → cong (Free→ℤFin (suc n)) + (ℤFin→FreeInv (suc n) f) ∙ λ i r → -ℤ (p i r)) + λ f g p q → cong (Free→ℤFin (suc n)) + (ℤFin→FreeHom (suc n) f g) ∙ λ i r → p i r + q i r + +Iso-ℤFin-FreeAbGroup : (n : ℕ) → Iso (ℤ[Fin n ] .fst) (FAGAbGroup {A = Fin n} .fst) +Iso.fun (Iso-ℤFin-FreeAbGroup n) = ℤFin→Free n +Iso.inv (Iso-ℤFin-FreeAbGroup n) = Free→ℤFin n +Iso.rightInv (Iso-ℤFin-FreeAbGroup n) = ℤFin→Free→ℤFin n +Iso.leftInv (Iso-ℤFin-FreeAbGroup n) = Free→ℤFin→Free n + +ℤFin≅FreeAbGroup : (n : ℕ) → AbGroupIso (ℤ[Fin n ]) (FAGAbGroup {A = Fin n}) +fst (ℤFin≅FreeAbGroup n) = Iso-ℤFin-FreeAbGroup n +snd (ℤFin≅FreeAbGroup n) = makeIsGroupHom (ℤFin→FreeHom n) + +-- this iso implies that ℤFin inherits the prop elimination principle of FAGAbGroup +elimPropℤFin : ∀ {ℓ} (n : ℕ) + (A : ℤ[Fin n ] .fst → Type ℓ) + → ((x : _) → isProp (A x)) + → (A (λ _ → 0)) + → ((x : _) → A (ℤFinGenerator x)) + → ((f g : _) → A f → A g → A λ x → f x + g x) + → ((f : _) → A f → A (-ℤ_ ∘ f)) + → (x : _) → A x +elimPropℤFin n A pr z t s u w = + subst A (Iso.leftInv (Iso-ℤFin-FreeAbGroup n) w) (help (ℤFin→Free n w)) + where + help : (x : _) → A (Free→ℤFin n x) + help = ElimProp.f (pr _) t z + (λ {x} {y} → s (Free→ℤFin n x) (Free→ℤFin n y)) + λ {x} → u (Free→ℤFin n x) + +-- functoriality of ℤFin in n + +ℤFinFunctGenerator : {n m : ℕ} (f : Fin n → Fin m) (g : ℤ[Fin n ] .fst) + (x : Fin n) → ℤ[Fin m ] .fst +ℤFinFunctGenerator {n = n} {m} f g x y with ((f x .fst) ≟ᵗ y .fst) +... | lt _ = 0 +... | eq _ = g x +... | gt _ = 0 + +ℤFinFunctGenerator≡ : {n m : ℕ} (f : Fin n → Fin m) + (t : Fin n) (y : Fin m) + → ℤFinFunctGenerator f (ℤFinGenerator t) t y + ≡ ℤFinGenerator (f t) y +ℤFinFunctGenerator≡ f t y with (f t .fst ≟ᵗ y .fst) +... | lt _ = refl +... | eq _ = lem + where + lem : _ + lem with (fst t ≟ᵗ fst t) + ... | lt q = ⊥.rec (¬m<ᵗm q) + ... | eq _ = refl + ... | gt q = ⊥.rec (¬m<ᵗm q) +... | gt _ = refl + +ℤFinFunctFun : {n m : ℕ} (f : Fin n → Fin m) + → ℤ[Fin n ] .fst → ℤ[Fin m ] .fst +ℤFinFunctFun {n = n} {m} f g x = + sumFinℤ {n = n} λ y → ℤFinFunctGenerator f g y x + +ℤFinFunct : {n m : ℕ} (f : Fin n → Fin m) + → AbGroupHom (ℤ[Fin n ]) (ℤ[Fin m ]) +fst (ℤFinFunct {n = n} {m} f) = ℤFinFunctFun f +snd (ℤFinFunct {n = n} {m} f) = + makeIsGroupHom λ g h + → funExt λ x → sumFinGenId _+_ 0 + (λ y → ℤFinFunctGenerator f (λ x → g x + h x) y x) + (λ y → ℤFinFunctGenerator f g y x + ℤFinFunctGenerator f h y x) + (funExt (lem g h x)) + ∙ sumFinGenHom _+_ (pos 0) (λ _ → refl) +Comm +Assoc n _ _ + where + lem : (g h : _) (x : _) (y : Fin n) + → ℤFinFunctGenerator f (λ x → g x + h x) y x + ≡ ℤFinFunctGenerator f g y x + ℤFinFunctGenerator f h y x + lem g h x y with (f y . fst ≟ᵗ x .fst) + ... | lt _ = refl + ... | eq _ = refl + ... | gt _ = refl + +-- Homs are equal if they agree on generators +agreeOnℤFinGenerator→≡ : ∀ {n m : ℕ} + → {ϕ ψ : AbGroupHom (ℤ[Fin n ]) (ℤ[Fin m ])} + → ((x : _) → fst ϕ (ℤFinGenerator x) ≡ fst ψ (ℤFinGenerator x)) + → ϕ ≡ ψ +agreeOnℤFinGenerator→≡ {n} {m} {ϕ} {ψ} idr = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt + (elimPropℤFin _ _ (λ _ → isOfHLevelPath' 1 (isSetΠ (λ _ → isSetℤ)) _ _) + (IsGroupHom.pres1 (snd ϕ) ∙ sym (IsGroupHom.pres1 (snd ψ))) + idr + (λ f g p q → IsGroupHom.pres· (snd ϕ) f g + ∙∙ (λ i x → p i x + q i x) + ∙∙ sym (IsGroupHom.pres· (snd ψ) f g )) + λ f p → IsGroupHom.presinv (snd ϕ) f + ∙∙ (λ i x → -ℤ (p i x)) + ∙∙ sym (IsGroupHom.presinv (snd ψ) f))) diff --git a/Cubical/Algebra/AbGroup/Instances/Pi.agda b/Cubical/Algebra/AbGroup/Instances/Pi.agda index 5e61f2928e..d5eb0b1273 100644 --- a/Cubical/Algebra/AbGroup/Instances/Pi.agda +++ b/Cubical/Algebra/AbGroup/Instances/Pi.agda @@ -5,8 +5,12 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Group.Instances.Pi +open import Cubical.Algebra.AbGroup.Instances.Int 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 + +ΠℤAbGroup : ∀ {ℓ} (A : Type ℓ) → AbGroup ℓ +ΠℤAbGroup A = ΠAbGroup {X = A} λ _ → ℤAbGroup diff --git a/Cubical/Algebra/AbGroup/Properties.agda b/Cubical/Algebra/AbGroup/Properties.agda index a80c06397f..61d4071a74 100644 --- a/Cubical/Algebra/AbGroup/Properties.agda +++ b/Cubical/Algebra/AbGroup/Properties.agda @@ -2,6 +2,7 @@ module Cubical.Algebra.AbGroup.Properties where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.Group @@ -16,9 +17,10 @@ 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 private variable - ℓ : Level + ℓ ℓ' : Level module AbGroupTheory (A : AbGroup ℓ) where open GroupTheory (AbGroup→Group A) @@ -36,7 +38,7 @@ 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 +addGroupHom : (A : AbGroup ℓ) (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)) @@ -44,7 +46,7 @@ snd (addGroupHom A B ϕ ψ) = makeIsGroupHom (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 +negGroupHom : (A : AbGroup ℓ) (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 @@ -56,7 +58,7 @@ snd (negGroupHom A B ϕ) = (IsGroupHom.presinv (snd ϕ) x) (IsGroupHom.presinv (snd ϕ) y) -subtrGroupHom : (A B : AbGroup ℓ) (ϕ ψ : AbGroupHom A B) → AbGroupHom A B +subtrGroupHom : (A : AbGroup ℓ) (B : AbGroup ℓ') (ϕ ψ : AbGroupHom A B) → AbGroupHom A B subtrGroupHom A B ϕ ψ = addGroupHom A B ϕ (negGroupHom A B ψ) @@ -118,3 +120,18 @@ G [ n ]ₜ = (kerSubgroup (multₗHom G (pos n)))) λ {(x , p) (y , q) → Σ≡Prop (λ _ → isPropIsInKer (multₗHom G (pos n)) _) (AbGroupStr.+Comm (snd G) _ _)} + +-- finite sums commute with negations +sumFinG-neg : (n : ℕ) {A : AbGroup ℓ} + → (f : Fin n → fst A) + → sumFinGroup (AbGroup→Group A) {n} (λ x → AbGroupStr.-_ (snd A) (f x)) + ≡ AbGroupStr.-_ (snd A) (sumFinGroup (AbGroup→Group A) {n} f) +sumFinG-neg zero {A} f = sym (GroupTheory.inv1g (AbGroup→Group A)) +sumFinG-neg (suc n) {A} f = + cong₂ compA refl (sumFinG-neg n {A = A} (f ∘ injectSuc)) + ∙∙ AbGroupStr.+Comm (snd A) _ _ + ∙∙ sym (GroupTheory.invDistr (AbGroup→Group A) _ _) + where + -A = AbGroupStr.-_ (snd A) + 0A = AbGroupStr.0g (snd A) + compA = AbGroupStr._+_ (snd A) 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 new file mode 100644 index 0000000000..462d9f2f10 --- /dev/null +++ b/Cubical/Algebra/ChainComplex/Base.agda @@ -0,0 +1,114 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Algebra.ChainComplex.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +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.MorphismProperties +open import Cubical.Algebra.AbGroup + +private + variable + ℓ ℓ' ℓ'' : Level + +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 + 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) + +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 λ 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..0a8611d896 --- /dev/null +++ b/Cubical/Algebra/ChainComplex/Finite.agda @@ -0,0 +1,123 @@ +{-# 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.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Fin.Inductive.Base + +open import Cubical.Algebra.ChainComplex.Base +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup + +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)) + ≡ 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..99d4dffd2d --- /dev/null +++ b/Cubical/Algebra/ChainComplex/Homology.agda @@ -0,0 +1,381 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Algebra.ChainComplex.Homology where + +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 +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Nat.Order.Inductive + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +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 + + +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 + ϕ = 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 , <ᵗ-trans p <ᵗsucm) .fst a + , ((λ i → fst (ϕcomm (n , <ᵗ-trans p <ᵗsucm) 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) + pf = <ᵗ-trans (snd n) <ᵗsucm + + 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/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)) diff --git a/Cubical/Algebra/Group/Base.agda b/Cubical/Algebra/Group/Base.agda index 2bde95181d..3ab0291971 100644 --- a/Cubical/Algebra/Group/Base.agda +++ b/Cubical/Algebra/Group/Base.agda @@ -8,6 +8,8 @@ module Cubical.Algebra.Group.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Fin.Inductive.Base open import Cubical.Algebra.Monoid open import Cubical.Algebra.Semigroup @@ -175,3 +177,6 @@ makeGroup-left 1g _·_ inv set ·Assoc ·IdL ·InvL = ≡⟨ ·IdL a ⟩ a ∎ + +sumFinGroup : ∀ {ℓ} (G : Group ℓ) {n : ℕ} (f : Fin n → fst G) → fst G +sumFinGroup G {n = n} f = sumFinGen {n = n} (GroupStr._·_ (snd G)) (GroupStr.1g (snd G)) f diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index 84c33a481e..ff7ab15793 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -1,8 +1,9 @@ -- The SIP applied to groups -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} 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 @@ -24,6 +25,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 +198,107 @@ 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)) + +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 _ _) + (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))) + + 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)) + +-- 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/Algebra/Group/Instances/Int.agda b/Cubical/Algebra/Group/Instances/Int.agda index 131a3ec875..9c257912ff 100644 --- a/Cubical/Algebra/Group/Instances/Int.agda +++ b/Cubical/Algebra/Group/Instances/Int.agda @@ -3,9 +3,12 @@ module Cubical.Algebra.Group.Instances.Int where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function open import Cubical.Data.Int renaming (_+_ to _+ℤ_ ; _-_ to _-ℤ_; -_ to -ℤ_ ; _·_ to _·ℤ_) +open import Cubical.Data.Nat using (ℕ ; zero ; suc) +open import Cubical.Data.Fin.Inductive.Base open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Properties @@ -42,3 +45,13 @@ fst negEquivℤ = (GroupTheory.invInv ℤGroup)) snd negEquivℤ = makeIsGroupHom -Dist+ + +sumFinGroupℤComm : (G : Group₀) (h : GroupIso G ℤGroup) {n : ℕ} + (f : Fin n → fst G) → sumFinℤ {n = n} (λ a → Iso.fun (fst h) (f a)) + ≡ Iso.fun (fst h) (sumFinGroup G {n = n} f) +sumFinGroupℤComm G h {n = zero} f = sym (IsGroupHom.pres1 (snd h)) +sumFinGroupℤComm G h {n = suc n} f = + cong₂ _+ℤ_ (λ _ → Iso.fun (fst h) (f flast)) + (sumFinGroupℤComm G h {n = n} (f ∘ injectSuc {n = n})) + ∙ sym (IsGroupHom.pres· (snd h) (f flast) + (sumFinGroup G {n = n} (λ x → f (injectSuc {n = n} 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/Algebra/Group/ZAction.agda b/Cubical/Algebra/Group/ZAction.agda index 38afec7235..50f164fab9 100644 --- a/Cubical/Algebra/Group/ZAction.agda +++ b/Cubical/Algebra/Group/ZAction.agda @@ -162,6 +162,17 @@ module _ (G : Group ℓ) (g : fst G) where (cong (_ℤ[ G ]· g) (+Comm y (negsuc n)) ∙ distrℤ· (negsuc n) y) ∙ (·Assoc (snd G) _ _ _) +ℤ·-negsuc : ∀ {ℓ} (G : Group ℓ) (a : ℕ) (g : fst G) + → (negsuc a ℤ[ G ]· g) + ≡ GroupStr.inv (snd G) ((pos (suc a)) ℤ[ G ]· g) +ℤ·-negsuc G zero g = + sym (cong (GroupStr.inv (snd G)) + (GroupStr.·IdR (snd G) _)) +ℤ·-negsuc G (suc a) g = + (distrℤ· G g (negsuc a) (negsuc zero)) + ∙ cong₂ (GroupStr._·_ (snd G)) (ℤ·-negsuc G a g) refl + ∙ sym (GroupTheory.invDistr G g ((pos (suc a)) ℤ[ G ]· g)) + GroupHomℤ→ℤpres- : (e : GroupHom ℤGroup ℤGroup) (a : ℤ) → fst e (- a) ≡ - fst e a GroupHomℤ→ℤpres- e a = presinv (snd e) a @@ -709,7 +720,7 @@ GroupEquiv-abstractℤ/abs-gen G H L e r f g ex n p = main (sym (cong predℤ (p ∙ negsuc·negsuc n₁ (suc n) ∙ sym (pos· (suc n₁) (suc (suc n)))))))) -1∈Im→isEquiv : ∀ (G : Group₀) (e : GroupEquiv ℤGroup G) +1∈Im→isEquiv : (G : Group₀) (e : GroupEquiv ℤGroup G) → (h : GroupHom G ℤGroup) → isInIm (_ , snd h) 1 → isEquiv (fst h) @@ -719,3 +730,19 @@ GroupEquiv-abstractℤ/abs-gen G H L e r f g ex n p = main → isInIm (_ , snd h) 1 → isEquiv (fst h)) 1∈Im→isEquivℤ + +gen∈Im→isEquiv : ∀ (G : Group₀) (e : GroupEquiv ℤGroup G) + (H : Group₀) (e' : GroupEquiv ℤGroup H) + → (h₀ : fst H) + → 1 ≡ invEq (fst e') h₀ + → (h : GroupHom G H) + → isInIm (_ , snd h) h₀ + → isEquiv (fst h) +gen∈Im→isEquiv G e H = + GroupEquivJ (λ H e' + → (h₀ : fst H) + → 1 ≡ invEq (fst e') h₀ + → (h : GroupHom G H) + → isInIm (_ , snd h) h₀ + → isEquiv (fst h)) + (J> 1∈Im→isEquiv G e) diff --git a/Cubical/Axiom/Choice.agda b/Cubical/Axiom/Choice.agda index 8e3bd511ba..e48a844fb5 100644 --- a/Cubical/Axiom/Choice.agda +++ b/Cubical/Axiom/Choice.agda @@ -21,6 +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 as IndF open import Cubical.Data.Nat.Order open import Cubical.HITs.Truncation as TR @@ -59,8 +60,64 @@ 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) +satAC₀ B = + isoToIsEquiv (iso (λ _ _ → tt*) (λ _ → tt*) (λ _ → refl) λ _ → refl) + +-- Fin m satisfies AC for any level n. +FinSatAC : (n m : ℕ) → ∀ {ℓ} → satAC ℓ n (FN.Fin m) +FinSatAC n zero B = + isoToIsEquiv (iso _ + (λ f → ∣ (λ x → ⊥.rec (FN.¬Fin0 x)) ∣ₕ) + (λ f → funExt λ x → ⊥.rec (FN.¬Fin0 x)) + (TR.elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) + λ a → cong ∣_∣ₕ (funExt λ x → ⊥.rec (FN.¬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 : FN.Fin (suc m) → Type ℓ) + → (satAC ℓ n (FN.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 (FN.fsuc x)))) + (invIso (CharacΠFinIso m)))) + + ac-eq : (n m : ℕ) {B : _} → (eq : satAC ℓ n (FN.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 : _ ) → FN.fzero ≡ (zero , p) + lem₁ p = Fin-fst-≡ refl + + lem₂ : (x : _) (p : suc x < suc m) + → Path (FN.Fin _) (FN.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∃ ℓ ℓ' (FN.Fin n) +satAC∃Fin n = satAC→satAC∃ (FinSatAC 1 n) + +InductiveFinSatAC : (n m : ℕ) → ∀ {ℓ} → satAC ℓ n (IndF.Fin m) +InductiveFinSatAC n m {ℓ} = + subst (satAC ℓ n) (isoToPath (Iso-Fin-InductiveFin m)) (FinSatAC n m) + +InductiveFinSatAC∃ : (n : ℕ) → satAC∃ ℓ ℓ' (IndF.Fin n) +InductiveFinSatAC∃ {ℓ = ℓ} {ℓ'} n = + subst (satAC∃ ℓ ℓ') (isoToPath (Iso-Fin-InductiveFin n)) (satAC∃Fin n) diff --git a/Cubical/CW/Approximation.agda b/Cubical/CW/Approximation.agda new file mode 100644 index 0000000000..89c2dbd594 --- /dev/null +++ b/Cubical/CW/Approximation.agda @@ -0,0 +1,692 @@ +{-# OPTIONS --lossy-unification --safe #-} +{- Cellular approximation theorems for +-- cellular maps and homotopies +-} + +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 +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.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) +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 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)) + → 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) + → finsequencemap (fst ∘ g) (λ 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∞' = + ∣ 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) + → (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 + 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) + → (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))) + (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/CW/Base.agda b/Cubical/CW/Base.agda new file mode 100644 index 0000000000..8f94d5531e --- /dev/null +++ b/Cubical/CW/Base.agda @@ -0,0 +1,156 @@ +{-# OPTIONS --safe --lossy-unification #-} + +-- This file contains definition of CW complexes and skeleta. + +module Cubical.CW.Base 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.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + +open import Cubical.Relation.Nullary + +open Sequence + + + +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 _ + +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) + +-- 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/ChainComplex.agda b/Cubical/CW/ChainComplex.agda new file mode 100644 index 0000000000..54d32db729 --- /dev/null +++ b/Cubical/CW/ChainComplex.agda @@ -0,0 +1,192 @@ +{-# 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 + +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.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 + + +-- 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 _ {ℓ} (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 ChainComplex + + CW-ChainComplex : ChainComplex ℓ-zero + 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 CW-ChainComplex diff --git a/Cubical/CW/Homology.agda b/Cubical/CW/Homology.agda new file mode 100644 index 0000000000..c216a36d01 --- /dev/null +++ b/Cubical/CW/Homology.agda @@ -0,0 +1,244 @@ +{-# OPTIONS --safe --lossy-unification #-} +{- +This file contains: +1. Functoriality of Hˢᵏᵉˡ +2. Construction of cellular homology, including funtoriality +-} + +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 +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.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/CW/Homotopy.agda b/Cubical/CW/Homotopy.agda new file mode 100644 index 0000000000..430437223a --- /dev/null +++ b/Cubical/CW/Homotopy.agda @@ -0,0 +1,766 @@ +{-# 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.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.Path +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.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) +open import Cubical.Data.Fin.Inductive.Base +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 +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) + 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 + constructor fincellhom + 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 + 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/CW/Map.agda b/Cubical/CW/Map.agda new file mode 100644 index 0000000000..b057c67666 --- /dev/null +++ b/Cubical/CW/Map.agda @@ -0,0 +1,444 @@ +{-# OPTIONS --safe --lossy-unification #-} + +{-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.CW.Map where + +open import Cubical.Foundations.Prelude +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.Inductive +open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; -_ to -ℤ_) +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.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.SphereBouquet +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.Relation.Nullary + +open Sequence + +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 + +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 +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 + +-- 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/CW/Properties.agda b/Cubical/CW/Properties.agda new file mode 100644 index 0000000000..4e481c84a8 --- /dev/null +++ b/Cubical/CW/Properties.agda @@ -0,0 +1,251 @@ +{-# OPTIONS --safe --lossy-unification #-} + +{-This file contains elimination principles and basic properties of CW +complexes/skeleta. +-} + +module Cubical.CW.Properties where + +open import Cubical.CW.Base + +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.Equiv.Properties + +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.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.Truncation as TR + +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 + +open Sequence + + + +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))) + +-- 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 = InductiveFinSatAC _ _ + + 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)) + +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/CW/Subcomplex.agda b/Cubical/CW/Subcomplex.agda new file mode 100644 index 0000000000..8cba619b94 --- /dev/null +++ b/Cubical/CW/Subcomplex.agda @@ -0,0 +1,196 @@ +{-# OPTIONS --safe --lossy-unification #-} + +module Cubical.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.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.Empty as ⊥ + +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 +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/Base.agda b/Cubical/Data/Fin/Base.agda index d632ecfd22..ef6fa50e4e 100644 --- a/Cubical/Data/Fin/Base.agda +++ b/Cubical/Data/Fin/Base.agda @@ -48,6 +48,11 @@ fsuc (k , l) = (suc k , suc-≤-suc l) finj : Fin k → Fin (suc k) finj (k , l) = k , ≤-trans l (1 , refl) +-- 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 @@ -72,6 +77,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<ᵗm (subst (fst x <ᵗ_) q p)) + ... | gt x = refl + ... | eq p = cong (degree n) (funExt lem) ∙ degreeIdfun n + where + lem : (x₁ : S₊ n) → pickPetal y (inr (x , x₁)) ≡ x₁ + lem x₁ with (fst y ≟ᵗ fst x) + ... | lt q = ⊥.rec (¬m<ᵗm (subst (fst y <ᵗ_) p q)) + ... | eq x = refl + ... | gt q = ⊥.rec (¬m<ᵗm (subst (_<ᵗ fst y) p q)) + ... | gt 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<ᵗm (subst (_<ᵗ fst x) q p)) + ... | gt x = refl + +bouquetDegreeConst : (n a b : ℕ) + → bouquetDegree {n} {a} {b} (λ _ → inl tt) ≡ trivGroupHom +bouquetDegreeConst n a b = + GroupHom≡ ((λ i r x → sumFinℤ (λ a → r a ·ℤ (degree-const n i))) + ∙∙ (λ i r x → sumFinℤ (λ a → ·Comm (r a) (pos 0) i)) + ∙∙ λ i r x → sumFinℤ0 a i) + +-- bouquetDegree preserves suspension +bouquetDegreeSusp₀ : {m k : ℕ} + → (f : SphereBouquet zero (Fin m) → SphereBouquet zero (Fin k)) + → bouquetDegree f ≡ bouquetDegree (bouquetSusp→ f) +bouquetDegreeSusp₀ {m = m} {k = zero} f = + subst (λ f → bouquetDegree f ≡ bouquetDegree (bouquetSusp→ f)) + f-const + (bouquetDegreeConst zero _ _ + ∙ sym (bouquetDegreeConst (suc zero) _ _) + ∙ cong bouquetDegree lem₂) + where + f-const : (λ _ → inl tt) ≡ f + f-const = funExt (λ x → isContrSphereBouquetZero 0 .snd (f x)) + lem₂ : (λ _ → inl tt) ≡ bouquetSusp→ {n = zero} (λ _ → inl tt) + lem₂ = funExt (λ x + → isContrSphereBouquetZero 1 .snd (bouquetSusp→ (λ _ → inl tt) x)) +bouquetDegreeSusp₀ {m = m} {k = suc k} f = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ r → funExt λ a → sumFinℤId m λ y → cong (r y ·ℤ_) + (degreeSusp 0 (λ x → pickPetal a (f (inr (y , x)))) + ∙ λ i → degree 1 (λ z → help y a z i))) + where + help : (y : _) (a : _) (z : _) + → suspFunS∙ (λ x → pickPetal a (f (inr (y , x)))) .fst z + ≡ pickPetal a (bouquetSusp→ f (inr (y , z))) + help y a base = refl + help y a (loop i) j = help' j i + where + main : (ft ff : SphereBouquet zero (Fin (suc k))) + → cong SuspBool→S¹ (merid (pickPetal a ff)) + ∙ cong SuspBool→S¹ (sym (merid (pickPetal a ft))) + ≡ cong (pickPetal a) (cong sphereBouquetSuspFun (merid ff) + ∙ sym (cong sphereBouquetSuspFun (merid ft))) + main = + SphereBouquet→Prop 0 fzero (λ _ → isPropΠ λ _ → isGroupoidS¹ _ _ _ _) + λ x y → SphereBouquet→Prop 0 fzero (λ _ → isGroupoidS¹ _ _ _ _) + λ x' y' + → (cong₂ _∙_ (final y' x') (cong sym (final y x)) + ∙ sym (cong-∙ (pickPetal a) + (merid-lem x' y' i1) (sym (merid-lem x y i1)))) + ∙ λ i → cong (pickPetal a) + ((merid-lem x' y' (~ i) ∙ sym (merid-lem x y (~ i)))) + where + merid-lem : (x : Fin (suc k)) (y : Bool) + → cong (sphereBouquetSuspIso .Iso.fun) (merid (inr (x , y))) + ≡ push x ∙∙ (λ i → inr (x , SuspBool→S¹ (merid y i))) ∙∙ sym (push x) + merid-lem x y = cong-∙∙ (Iso.fun sphereBouquetSuspIso₀) + (push x) + (λ i → inr (x , (merid y ∙ sym (merid true)) i)) + (sym (push x)) + ∙ cong (push x ∙∙_∙∙ sym (push x)) + (cong-∙ (λ z → Iso.fun sphereBouquetSuspIso₀ (inr (x , z))) + (merid y) (sym (merid true)) + ∙ sym (rUnit _)) + + pre-final : (y : Bool) (x : Fin (suc k)) + → cong SuspBool→S¹ (merid (pickPetal a (inr (x , y)))) + ≡ cong (pickPetal a) (push x) + ∙∙ cong (pickPetal a) (λ i → inr (x , SuspBool→S¹ (merid y i))) + ∙∙ cong (pickPetal a) (sym (push x)) + pre-final y x with (fst a ≟ᵗ fst x) + ... | lt x₁ = rUnit refl + ... | eq x₁ = rUnit _ + ... | gt x₁ = rUnit refl + + final : (y : Bool) (x : Fin (suc k)) + → cong SuspBool→S¹ (merid (pickPetal a (inr (x , y)))) + ≡ cong (pickPetal a) ((push x) + ∙∙ (λ i → inr (x , SuspBool→S¹ (merid y i))) + ∙∙ sym (push x)) + final y x = + pre-final y x + ∙ cong-∙∙ (pickPetal a) + (push x) + (λ i → inr (x , SuspBool→S¹ (merid y i))) + (sym (push x)) + + help' : cong (suspFunS∙ (λ x → pickPetal a (f (inr (y , x)))) .fst) loop + ≡ cong (pickPetal a) λ i → bouquetSusp→ f (inr (y , loop i)) + help' = + (λ j i → Iso.inv S¹IsoSuspBool + (cong-∙ (suspFun λ x → pickPetal a (f (inr (y , x)))) + (merid false) (sym (merid true)) j i)) + ∙ cong-∙ (Iso.inv S¹IsoSuspBool) + (merid (pickPetal a (f (inr (y , false))))) + (sym (merid (pickPetal a (f (inr (y , true)))))) + ∙ main (f (inr (y , true))) ((f (inr (y , false)))) + ∙ cong (cong (pickPetal a)) (refl + ∙ (sym (cong-∙ sphereBouquetSuspFun + (merid (f (inr (y , false)))) + (sym (merid (f (inr (y , true))))))) + ∙ cong (cong sphereBouquetSuspFun) + (sym (cong-∙ (suspFun f) + (merid (inr (y , false))) (sym (merid (inr (y , true))))) + ∙ cong (cong (suspFun f)) + (sym (cong-∙ (λ x → Iso.inv (Iso-SuspBouquet-Bouquet _ _) + (inr (y , x))) (merid false) (sym (merid true)))))) + ∙ λ j i → pickPetal a (sphereBouquetSuspFun + (suspFun f (Iso.inv (Iso-SuspBouquet-Bouquet _ _) + (inr (y , (merid false ∙ sym (merid true)) i))))) + +bouquetDegreeSusp : {n m k : ℕ} + → (f : SphereBouquet n (Fin m) → SphereBouquet n (Fin k)) + → bouquetDegree f ≡ bouquetDegree (bouquetSusp→ f) +bouquetDegreeSusp {n = zero} = bouquetDegreeSusp₀ +bouquetDegreeSusp {n = suc n} {m = m} {k = k} f = + agreeOnℤFinGenerator→≡ λ (x : Fin m) + → funExt λ (b : Fin k) → sumFinℤId m + λ z → cong (ℤFinGenerator x z ·ℤ_) + (degreeSusp (suc n) (λ x₁ → pickPetal b (f (inr (z , x₁)))) + ∙ cong (Iso.fun (Hⁿ-Sⁿ≅ℤ (suc n) .fst)) + (cong ∣_∣₂ (funExt λ x → help b z x))) + where + f₁ : (b : Fin k) → SphereBouquet∙ (suc n) (Fin k) →∙ coHomK-ptd (suc n) + fst (f₁ b) x = ∣ pickPetal b x ∣ₕ + snd (f₁ b) = refl + + f₂ : (b : Fin k) → SphereBouquet∙ (suc n) (Fin k) →∙ coHomK-ptd (suc n) + fst (f₂ b) x = ΩKn+1→Kn (suc n) + λ i → ∣ pickPetal b (Bouquet→ΩBouquetSusp (Fin k) (λ _ → S₊∙ (suc n)) x i) ∣ + snd (f₂ b) = ΩKn+1→Kn-refl (suc n) + + f₁≡f₂ : (b : Fin k) (x : _) → f₁ b .fst x ≡ f₂ b .fst x + f₁≡f₂ b = ⋁gen→∙Kn≡ (suc n) (f₁ b) (f₂ b) (uncurry main) + where + main : (x : Fin k) (y : S₊ (suc n)) + → f₁ b .fst (inr (x , y)) ≡ f₂ b .fst (inr (x , y)) + main x y = + main' + ∙ cong (ΩKn+1→Kn (suc n)) + (cong (cong ∣_∣ₕ) + (sym (cong-∙∙ (pickPetal {n = 2 +ℕ n} b) + (push x) (λ i → inr (x , σSn (suc n) y i)) (sym (push x))))) + where + main' : f₁ b .fst (inr (x , y)) + ≡ ΩKn+1→Kn (suc n) + (cong ∣_∣ₕ (cong (pickPetal {n = 2 +ℕ n} b) (push x) + ∙∙ (λ i → pickPetal {n = 2 +ℕ n} b + (inr (x , σSn (suc n) y i))) + ∙∙ cong (pickPetal {n = 2 +ℕ n} b) (sym (push x)))) + main' with (fst b ≟ᵗ fst x) + ... | lt x = sym (cong (ΩKn+1→Kn (suc n)) + (sym (rUnit refl)) ∙ ΩKn+1→Kn-refl (suc n)) + ... | eq x = sym (cong (ΩKn+1→Kn (suc n)) + (cong (cong ∣_∣ₕ) (sym (rUnit (σSn (suc n) y)))) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) ∣ y ∣) + ... | gt x = sym (cong (ΩKn+1→Kn (suc n)) (sym (rUnit refl)) + ∙ ΩKn+1→Kn-refl (suc n)) + + help : (b : Fin k) (z : Fin m) (t : _) + → Path (coHomK (2 +ℕ n)) + (∣ suspFun (λ x₁ → pickPetal b (f (inr (z , x₁)))) t ∣ₕ) + ∣ pickPetal b (bouquetSusp→ f (inr (z , t))) ∣ₕ + help b z north = refl + help b z south = cong ∣_∣ₕ (sym (merid (ptSn (suc n)))) + help b z (merid a i) j = + hcomp (λ r + → λ {(i = i0) → ∣ north ∣ + ; (i = i1) → ∣ merid (ptSn (suc n)) (~ j ∧ r) ∣ + ; (j = i0) → ∣ compPath-filler + (merid (pickPetal b (f (inr (z , a))))) + (sym (merid (ptSn (suc n)))) (~ r) i ∣ₕ + ; (j = i1) → (cong (Kn→ΩKn+1 (suc n)) (f₁≡f₂ b (f (inr (z , a)))) + ∙ Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) + (λ i → ∣ pickPetal b (bouquetSusp→ f + (inr (z , merid a i))) ∣)) r i}) + (Kn→ΩKn+1 (suc n) ∣ pickPetal b (f (inr (z , a))) ∣ i) + +-- bouquetDegree preserves composition +bouquetDegreeComp∙Suc : {n m k l : ℕ} + → (f : SphereBouquet∙ (suc n) (Fin m) →∙ SphereBouquet∙ (suc n) (Fin k)) + → (g : SphereBouquet∙ (suc n) (Fin l) →∙ SphereBouquet∙ (suc n) (Fin m)) + → bouquetDegree ((fst f) ∘ (fst g)) + ≡ compGroupHom (bouquetDegree (fst g)) (bouquetDegree (fst f)) +bouquetDegreeComp∙Suc {n} {m} {k} {l} f g = + agreeOnℤFinGenerator→≡ + λ (x : Fin l) + → funExt λ t + → sumFinℤId l + (λ a → + ·Comm (ℤFinGenerator x a) _ + ∙ cong (degree (suc n) + (λ x₁ → pickPetal t (fst f (fst g (inr (a , x₁))))) ·ℤ_) + (ℤFinGeneratorComm x a)) + ∙ sym (isGeneratorℤFinGenerator + (λ a → degree (suc n) + (λ x₁ → pickPetal t (fst f (fst g (inr (a , x₁)))))) x) + ∙ main n m (λ s → fst g (inr (x , s))) + ((λ s → pickPetal t (fst f s)) + , (cong (pickPetal t) (snd f))) + ∙ sumFinℤId m λ a → + degreeComp' (suc n) + (λ x₁ → pickPetal t (fst f (inr (a , x₁)))) + (λ x₁ → pickPetal a (fst g (inr (x , x₁)))) + ∙ λ j → simplify x a (~ j) + ·ℤ degree (suc n) (λ x₁ → pickPetal t (fst f (inr (a , x₁)))) + where + main : (n m : ℕ) (w : S₊ (suc n) → SphereBouquet (suc n) (Fin m)) + (F : ((SphereBouquet (suc n) (Fin m)) , inl tt) →∙ S₊∙ (suc n)) + → degree (suc n) (λ s → fst F (w s)) + ≡ sumFinℤ (λ a → degree (suc n) (λ s → fst F (inr (a , pickPetal a (w s))))) + main n zero w (F , Fp) = + (λ j → degree (suc n) (λ s → F (lem w j s))) + ∙ (λ j → degree (suc n) (λ s → Fp j)) + ∙ degree-const (suc n) + where + lem : (f : S₊ (suc n) → SphereBouquet (suc n) (Fin zero)) + → (f ≡ λ _ → inl tt) + lem f = funExt λ x → sym (isContrSphereBouquetZero (suc n) .snd (f x)) + main n (suc m) w F = + cong (Iso.fun (Hⁿ-Sⁿ≅ℤ n .fst)) + (cong ∣_∣₂ (funExt (λ x → asSum F (w x))) + ∙ sym (sumFinKComm + (λ a s → ∣ fst F (inr (a , pickPetal a (w s))) ∣ₕ))) + ∙ sym (sumFinGroupℤComm _ (Hⁿ-Sⁿ≅ℤ n) + (λ a → ∣ (λ s → ∣ fst F (inr (a , pickPetal a (w s))) ∣ₕ) ∣₂)) + where + asSum : (F : (SphereBouquet (suc n) (Fin (suc m)) , inl tt) →∙ S₊∙ (suc n)) + → (p : SphereBouquet (suc n) (Fin (suc m))) + → ∣ F .fst p ∣ₕ + ≡ sumFinK {m = suc n} (λ i → ∣ fst F (inr (i , pickPetal i p)) ∣ₕ) + asSum F = + ⋁gen→∙Kn≡ (suc n) + ((λ p → ∣ F .fst p ∣ₕ) , cong ∣_∣ₕ (snd F)) + ((λ p → sumFinK {m = suc n} + (λ i → ∣ fst F (inr (i , pickPetal i p)) ∣ₕ)) , sumPt) + (uncurry main-lem) + where + id₁ : (x : Fin (suc m)) (y : _) + → fst F (inr (x , pickPetal x (inr (x , y)))) ≡ fst F (inr (x , y)) + id₁ (x , p) y with (x ≟ᵗ x) + ... | lt x₁ = ⊥.rec (¬m<ᵗm x₁) + ... | eq x₁ = refl + ... | gt x₁ = ⊥.rec (¬m<ᵗm x₁) + + id₂ : (x : _) (x' : Fin (suc m)) (y : _) (q : ¬ x' ≡ x) + → ∣ fst F (inr (x' , pickPetal x' (inr (x , y)))) ∣ₕ ≡ 0ₖ (suc n) + id₂ (x , p) (x' , t) y con with (x' ≟ᵗ x) + ... | lt x₁ = cong ∣_∣ₕ (cong (fst F) (sym (push (x' , t))) ∙ snd F) + ... | eq x₁ = ⊥.rec (con (Σ≡Prop (λ _ → isProp<ᵗ) x₁)) + ... | gt x₁ = cong ∣_∣ₕ (cong (fst F) (sym (push (x' , t))) ∙ snd F) + + sumPt : sumFinK (λ i → ∣ fst F (inr (i , pickPetal i (inl tt))) ∣ₕ) + ≡ 0ₖ (suc n) + sumPt = sumFinGen0 _+ₖ_ (0ₖ (suc n)) (rUnitₖ _) _ + (λ i → ∣ fst F (inr (i , pickPetal i (inl tt))) ∣ₕ) + (λ s → cong ∣_∣ₕ (cong (fst F) (sym (push s)) ∙ snd F)) + + main-lem : (x : Fin (suc m)) (y : S₊ (suc n)) + → ∣ F .fst (inr (x , y)) ∣ₕ + ≡ sumFinK {n = suc m} {m = suc n} + (λ i → ∣ fst F (inr (i , pickPetal i (inr (x , y)))) ∣ₕ) + main-lem x y = sym (sumFin-choose _+ₖ_ (0ₖ (suc n)) (rUnitₖ _) (commₖ _) + (λ i → ∣ fst F (inr (i , pickPetal i (inr (x , y)))) ∣ₕ) + ∣ F .fst (inr (x , y)) ∣ₕ x + (cong ∣_∣ₕ (id₁ x y)) + λ x' → id₂ x x' y) + + simplify : (x : Fin l) (a : Fin m) + → sumFinℤ (λ a₁ → ℤFinGenerator x a₁ + ·ℤ degree (suc n) + (λ x₁ → pickPetal a (fst g (inr (a₁ , x₁))))) + ≡ degree (suc n) (λ x₁ → pickPetal a (fst g (inr (x , x₁)))) + simplify x a = + sumFinℤId l + (λ p → ·Comm (ℤFinGenerator x p) + (degree (suc n) (λ x₁ → pickPetal a (fst g (inr (p , x₁))))) + ∙ λ i → degree (suc n) (λ x₁ → pickPetal a (fst g (inr (p , x₁)))) + ·ℤ ℤFinGeneratorComm x p i) + ∙ sym (isGeneratorℤFinGenerator + (λ a₁ → degree (suc n) + (λ x₁ → pickPetal a (fst g (inr (a₁ , x₁))))) x) + +bouquetDegreeCompPos : {n m k l : ℕ} + → (f : SphereBouquet (suc n) (Fin m) → SphereBouquet (suc n) (Fin k)) + → (g : SphereBouquet (suc n) (Fin l) → SphereBouquet (suc n) (Fin m)) + → bouquetDegree (f ∘ g) ≡ compGroupHom (bouquetDegree g) (bouquetDegree f) +bouquetDegreeCompPos {n = n} {m = m} {k = k} {l = l} f g = + PT.rec2 (isSetGroupHom _ _) + (λ Hf Hg → bouquetDegreeComp∙Suc (f , Hf) (g , Hg)) + (isConnectedSphereBouquet (f (inl tt))) + (isConnectedSphereBouquet (g (inl tt))) + +bouquetDegreeComp : {n m k l : ℕ} + → (f : SphereBouquet n (Fin m) → SphereBouquet n (Fin k)) + → (g : SphereBouquet n (Fin l) → SphereBouquet n (Fin m)) + → bouquetDegree (f ∘ g) ≡ compGroupHom (bouquetDegree g) (bouquetDegree f) +bouquetDegreeComp {n = zero} f g = + bouquetDegreeSusp (f ∘ g) + ∙∙ Σ≡Prop (λ _ → isPropIsGroupHom _ _) + ((λ i → bouquetDegree (bouquetSusp→∘ g f (~ i)) .fst) + ∙ cong fst (bouquetDegreeCompPos (bouquetSusp→ f) (bouquetSusp→ g))) + ∙∙ sym (cong₂ compGroupHom (bouquetDegreeSusp g) (bouquetDegreeSusp f)) +bouquetDegreeComp {n = suc n} f g = bouquetDegreeCompPos f g + +bouquetDegreeComp∙ : {n m k l : ℕ} + (f : SphereBouquet∙ n (Fin m) →∙ SphereBouquet∙ n (Fin k)) + (g : SphereBouquet∙ n (Fin l) →∙ SphereBouquet∙ n (Fin m)) + → bouquetDegree ((fst f) ∘ (fst g)) + ≡ compGroupHom (bouquetDegree (fst g)) (bouquetDegree (fst f)) +bouquetDegreeComp∙ {n = zero} {m} {k = k} {l = l} f g = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (cong fst (bouquetDegreeSusp (fst f ∘ fst g)) + ∙ sym (cong (fst ∘ bouquetDegree) (bouquetSusp→∘ (fst g) (fst f))) + ∙ cong fst (bouquetDegreeComp∙Suc (bouquetSusp→ (fst f) , refl) + (bouquetSusp→ (fst g) , refl)) + ∙ cong fst (cong₂ compGroupHom (sym (bouquetDegreeSusp (fst g))) + (sym (bouquetDegreeSusp (fst f))))) +bouquetDegreeComp∙ {n = suc n} = bouquetDegreeComp∙Suc diff --git a/Cubical/HITs/SphereBouquet/Properties.agda b/Cubical/HITs/SphereBouquet/Properties.agda new file mode 100644 index 0000000000..8c0284f731 --- /dev/null +++ b/Cubical/HITs/SphereBouquet/Properties.agda @@ -0,0 +1,353 @@ +{-# OPTIONS --cubical --safe #-} + +module Cubical.HITs.SphereBouquet.Properties where + +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.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function + +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ + +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.Wedge +open import Cubical.HITs.SphereBouquet.Base + +private + variable + ℓ ℓ' : Level + +SphereBouquet→Prop : (n : ℕ) {A : Type ℓ} (a : A) + → {B : SphereBouquet n A → Type ℓ'} + → ((b : _) → isProp (B b)) + → ((x : _) (y : _) → B (inr (x , y))) + → (s : _) → B s +SphereBouquet→Prop n a {B = B} pr d = + elimProp _ pr (λ t → subst B (sym (push a)) (d a _)) + λ {(x , y) → d x y} + +isContrSphereBouquetZero : (n : ℕ) → isContr (SphereBouquet n (Fin zero)) +fst (isContrSphereBouquetZero n) = inl tt +snd (isContrSphereBouquetZero n) (inl x) = refl +snd (isContrSphereBouquetZero n) (inr x) = ⊥.rec (¬Fin0 (fst x)) +snd (isContrSphereBouquetZero n) (push a i) j = + ⊥.rec {A = Square {A = SphereBouquet n (Fin zero)} + (λ _ → inl tt) (push a) (λ _ → inl tt) + (⊥.rec (¬Fin0 a))} (¬Fin0 a) j i + +isConnectedSphereBouquet : {n : ℕ} {A : Type ℓ} + (x : SphereBouquet (suc n) A) → ∥ x ≡ inl tt ∥₁ +isConnectedSphereBouquet {n = n} {A} = + elimProp (λ x → ∥ x ≡ inl tt ∥₁) (λ x → squash₁) (λ x → ∣ refl ∣₁) + (λ (a , s) → sphereToPropElim n {A = λ x → ∥ inr (a , x) ≡ inl tt ∥₁} + (λ x → squash₁) ∣ sym (push a) ∣₁ s) + +sphereBouquetSuspIso₀ : {A : Type ℓ} + → Iso (⋁gen A (λ a → Susp∙ (fst (S₊∙ zero)))) + (SphereBouquet 1 A) +Iso.fun sphereBouquetSuspIso₀ (inl x) = inl x +Iso.fun sphereBouquetSuspIso₀ (inr (a , b)) = + inr (a , Iso.inv (IsoSucSphereSusp 0) b) +Iso.fun sphereBouquetSuspIso₀ (push a i) = push a i +Iso.inv sphereBouquetSuspIso₀ (inl x) = inl x +Iso.inv sphereBouquetSuspIso₀ (inr (a , b)) = + inr (a , Iso.fun (IsoSucSphereSusp 0) b) +Iso.inv sphereBouquetSuspIso₀ (push a i) = push a i +Iso.rightInv sphereBouquetSuspIso₀ (inl x) = refl +Iso.rightInv sphereBouquetSuspIso₀ (inr (a , y)) i = + inr (a , Iso.leftInv (IsoSucSphereSusp 0) y i) +Iso.rightInv sphereBouquetSuspIso₀ (push a i) = refl +Iso.leftInv sphereBouquetSuspIso₀ (inl x) = refl +Iso.leftInv sphereBouquetSuspIso₀ (inr (a , y)) i = + inr (a , Iso.rightInv (IsoSucSphereSusp 0) y i) +Iso.leftInv sphereBouquetSuspIso₀ (push a i) = refl + +--a sphere bouquet is the wedge sum of A n-dimensional spheres +sphereBouquetSuspIso : {A : Type ℓ} {n : ℕ} + → Iso (Susp (SphereBouquet n A)) (SphereBouquet (suc n) A) +sphereBouquetSuspIso {A = A} {n = zero} = + compIso (Iso-SuspBouquet-Bouquet A λ _ → S₊∙ zero) sphereBouquetSuspIso₀ +sphereBouquetSuspIso {A = A} {n = suc n} = + Iso-SuspBouquet-Bouquet A λ _ → S₊∙ (suc n) + +sphereBouquet≃Susp : {A : Type ℓ} {n : ℕ} + → Susp (SphereBouquet n A) ≃ SphereBouquet (suc n) A +sphereBouquet≃Susp = isoToEquiv (sphereBouquetSuspIso) + +sphereBouquetSuspFun : {A : Type ℓ} {n : ℕ} + → Susp (SphereBouquet n A) → SphereBouquet (suc n) A +sphereBouquetSuspFun {A = A} {n = n} = sphereBouquetSuspIso .Iso.fun + +sphereBouquetSuspInvFun : {A : Type ℓ} {n : ℕ} + → SphereBouquet (suc n) A → Susp (SphereBouquet n A) +sphereBouquetSuspInvFun {A = A} {n = n} = sphereBouquetSuspIso .Iso.inv + +--the suspension of a n-dimensional bouquet is a (n+1)-dimensional bouquet +--here is the action of suspension on morphisms +bouquetSusp→ : {n : ℕ} {A B : Type ℓ} + → (SphereBouquet n A → SphereBouquet n B) + → (SphereBouquet (suc n) A → SphereBouquet (suc n) B) +bouquetSusp→ {n = n} {A} {B} f = + sphereBouquetSuspFun ∘ (suspFun f) ∘ sphereBouquetSuspInvFun + +bouquetSusp→∘ : {n : ℕ} {A B C : Type ℓ} + → (f : SphereBouquet n A → SphereBouquet n B) + → (g : SphereBouquet n B → SphereBouquet n C) + → bouquetSusp→ g ∘ bouquetSusp→ f ≡ bouquetSusp→ (g ∘ f) +bouquetSusp→∘ f g i = + sphereBouquetSuspFun + ∘ ((λ i → suspFun g ∘ (λ x → Iso.leftInv sphereBouquetSuspIso x i) ∘ suspFun f) + ∙ sym (suspFunComp g f)) i + ∘ sphereBouquetSuspInvFun + +------ Sphere bouquets as cofibres ------ +{- +Goal: show that given a pushout square + +Fin m × Sⁿ⁻¹ → Fin m + ∣ ∣ + αₙ ∣ ∣ + ↓ ⌜ ↓ + Cₙ -----→ Cₙ₊₁ + ι + +we get an equivalence cofib ι ≃ ⋁ᵢ Sⁿ where +i ranges over Fin m +-} + +-- Prelims for description of maps +preBTC : {Cₙ Cₙ₊₁ : Type ℓ} (n mₙ : ℕ) + (αₙ : Fin mₙ × S⁻ n → Cₙ) + (e : Cₙ₊₁ ≃ Pushout αₙ fst) + → (x : Fin mₙ) + → S₊∙ n →∙ (cofib (invEq e ∘ inl) , inl tt) +fst (preBTC zero mₙ αₙ e x) false = inr (invEq e (inr x)) +fst (preBTC zero mₙ αₙ e x) true = inl tt +fst (preBTC (suc zero) mₙ αₙ e x) base = inl tt +fst (preBTC (suc zero) mₙ αₙ e x) (loop i) = + (push (αₙ (x , false)) + ∙∙ (λ j → inr (invEq e ((push (x , false) ∙ sym (push (x , true))) j))) + ∙∙ sym (push (αₙ (x , true)))) i +fst (preBTC (suc (suc n)) mₙ αₙ e x) north = inl tt +fst (preBTC (suc (suc n)) mₙ αₙ e x) south = inl tt +fst (preBTC (suc (suc n)) mₙ αₙ e x) (merid a i) = + (push (αₙ (x , a)) + ∙∙ (λ j → inr (invEq e ((push (x , a) ∙ sym (push (x , ptSn (suc n)))) j))) + ∙∙ sym (push (αₙ (x , ptSn (suc n))))) i +snd (preBTC zero mₙ αₙ e x) = refl +snd (preBTC (suc zero) mₙ αₙ e x) = refl +snd (preBTC (suc (suc n)) mₙ αₙ e x) = refl + +module _ where + Pushout→Bouquet : {Cₙ Cₙ₊₁ : Type ℓ} (n mₙ : ℕ) + (αₙ : Fin mₙ × S⁻ n → Cₙ) + (e : Cₙ₊₁ ≃ Pushout αₙ fst) + → Pushout αₙ fst → SphereBouquet n (Fin mₙ) + Pushout→Bouquet n mₙ αₙ e (inl x) = inl tt + Pushout→Bouquet zero mₙ αₙ e (inr x) = inr (x , false) + Pushout→Bouquet (suc n) mₙ αₙ e (inr x) = inr (x , ptSn (suc n)) + Pushout→Bouquet (suc n) mₙ αₙ e (push a i) = + (push (a .fst) ∙ λ i → inr (a .fst , σSn n (a .snd) i)) i + +-- Maps back and forth +module BouquetFuns {Cₙ Cₙ₊₁ : Type ℓ} (n mₙ : ℕ) + (αₙ : Fin mₙ × S⁻ n → Cₙ) + (e : Cₙ₊₁ ≃ Pushout αₙ fst) where + CTB : cofib (invEq e ∘ inl) → SphereBouquet n (Fin mₙ) + CTB (inl x) = inl tt + CTB (inr x) = Pushout→Bouquet n mₙ αₙ e (fst e x) + CTB (push a i) = Pushout→Bouquet n mₙ αₙ e (secEq e (inl a) (~ i)) + + BTC : SphereBouquet n (Fin mₙ) → cofib (invEq e ∘ inl) + BTC (inl x) = inl tt + BTC (inr x) = preBTC n mₙ αₙ e (fst x) .fst (snd x) + BTC (push a i) = preBTC n mₙ αₙ e a .snd (~ i) + +-- Maps cancel out +CTB-BTC-cancel : {Cₙ Cₙ₊₁ : Type ℓ} (n mₙ : ℕ) + (αₙ : Fin mₙ × S⁻ n → Cₙ) + (e : Cₙ₊₁ ≃ Pushout αₙ fst) + → section (BouquetFuns.CTB n mₙ αₙ e) (BouquetFuns.BTC n mₙ αₙ e) + × retract (BouquetFuns.CTB n mₙ αₙ e) (BouquetFuns.BTC n mₙ αₙ e) +CTB-BTC-cancel {Cₙ = Cₙ} n mₙ αₙ = + EquivJ (λ C₊ e → + section (BouquetFuns.CTB n mₙ αₙ e) + (BouquetFuns.BTC n mₙ αₙ e) + × + retract (BouquetFuns.CTB n mₙ αₙ e) + (BouquetFuns.BTC n mₙ αₙ e)) + (retr-main n αₙ , section-main n αₙ) + where + module S (n : ℕ) (αₙ : Fin mₙ × S⁻ n → Cₙ) where + module T = BouquetFuns n mₙ αₙ (idEquiv _) + open T public + + retr-inr : (n : ℕ) (αₙ : Fin mₙ × S⁻ n → Cₙ) (a : _) (b : _) + → S.CTB n αₙ (S.BTC n αₙ (inr (a , b))) ≡ inr (a , b) + retr-inr zero aₙ a false = refl + retr-inr zero aₙ a true = push a + retr-inr (suc zero) αₙ a base = push a + retr-inr (suc zero) αₙ a (loop i) j = + hcomp (λ r → + λ {(i = i0) → push a j + ; (i = i1) → push a j + ; (j = i0) → S.CTB (suc zero) αₙ + (doubleCompPath-filler + (push (αₙ (a , false))) + (λ j → inr ((push (a , false) + ∙ sym (push (a , true))) j)) + (sym (push (αₙ (a , true)))) r i) + ; (j = i1) → inr (a , loop i)}) + (hcomp (λ r → + λ {(i = i0) → push a j + ; (i = i1) → compPath-filler' (push a) refl (~ j) (~ r) + ; (j = i0) → S.CTB (suc zero) αₙ + (inr (compPath-filler + (push (a , false)) + (sym (push (a , true))) r i)) + ; (j = i1) → inr (a , loop i)}) + (hcomp (λ r → + λ {(i = i0) → push a (j ∨ ~ r) + ; (i = i1) → inr (a , base) + ; (j = i0) → compPath-filler' (push a) (λ j → inr (a , loop j)) r i + ; (j = i1) → inr (a , loop i)}) + (inr (a , loop i)))) + retr-inr (suc (suc n)) αₙ a north = push a + retr-inr (suc (suc n)) αₙ a south = + push a ∙ λ j → inr (a , merid (ptSn (suc n)) j) + retr-inr (suc (suc n)) αₙ a (merid a₁ i) j = + hcomp (λ r → + λ {(i = i0) → push a j + ; (i = i1) → compPath-filler + (push a) + (λ j₁ → inr (a , merid (ptSn (suc n)) j₁)) r j + ; (j = i0) → S.CTB (suc (suc n)) αₙ + (doubleCompPath-filler + (push (αₙ (a , a₁))) + (λ i → inr ((push (a , a₁) + ∙ sym (push (a , ptSn (suc n)))) i)) + (sym (push (αₙ (a , ptSn (suc n))))) r i) + ; (j = i1) → inr (a , compPath-filler (merid a₁) + (sym (merid (ptSn (suc n)))) (~ r) i)}) + (hcomp (λ r → + λ {(i = i0) → push a j + ; (i = i1) → compPath-filler' (push a) + (λ i → inr (a , σSn _ (ptSn (suc n)) i)) (~ j) (~ r) + ; (j = i0) → S.CTB (suc (suc n)) αₙ + (inr (compPath-filler (push (a , a₁)) + (sym (push (a , ptSn (suc n)))) r i) ) + ; (j = i1) → inr (a , help r i)}) + (hcomp (λ r → λ {(i = i0) → push a (j ∨ ~ r) + ; (i = i1) → inr (a , north) + ; (j = i0) → compPath-filler' + (push a) (λ i → inr (a , σSn _ a₁ i)) r i + ; (j = i1) → inr (a , σSn _ a₁ i)}) + (inr (a , σSn _ a₁ i)))) + where + help : Square (σSn _ a₁) (σSn _ a₁) refl (sym (σSn _ (ptSn (suc n)))) + help = flipSquare ((λ i _ → σSn _ a₁ i) + ▷ λ i → sym (rCancel (merid (ptSn (suc n))) (~ i))) + + retr-main : (n : _) (αₙ : _) → section (S.CTB n αₙ) (S.BTC n αₙ) + retr-main n αₙ (inl x) = refl + retr-main n αₙ (inr (a , b)) = retr-inr n αₙ a b + retr-main zero αₙ (push a i) j = push a (i ∧ j) + retr-main (suc zero) αₙ (push a i) j = push a (i ∧ j) + retr-main (suc (suc n)) αₙ (push a i) j = push a (i ∧ j) + + section-main : (n : _) (αₙ : _) → retract (S.CTB n αₙ) (S.BTC n αₙ) + section-main n αₙ (inl x) = refl + section-main n αₙ (inr (inl x)) = push x + section-main zero αₙ (inr (inr x)) = refl + section-main (suc zero) αₙ (inr (inr x)) = + push (αₙ (x , true)) ∙ λ i → inr (push (x , true) i) + section-main (suc (suc n)) αₙ (inr (inr x)) = + push (αₙ (x , ptSn (suc n))) ∙ λ i → inr (push (x , ptSn (suc n)) i) + section-main (suc zero) αₙ (inr (push (a , false) i)) j = + hcomp (λ r → λ {(i = i0) → push (αₙ (a , false)) j + ; (i = i1) → compPath-filler (push (αₙ (a , true))) + (λ i₁ → inr (push (a , true) i₁)) i1 j + ; (j = i0) → + S.BTC (suc zero) αₙ + (compPath-filler' + (push a) + (λ i → inr (a , σSn zero false i)) r i) + ; (j = i1) → inr (push (a , false) i)}) + (hcomp (λ r → λ {(i = i0) → push (αₙ (a , false)) (j ∨ ~ r) + ; (i = i1) → + compPath-filler' + (push (αₙ (a , true))) + (λ i → inr (push (a , true) i)) r j + ; (j = i1) → inr (push (a , false) i)}) + (inr (compPath-filler + (push (a , false)) + (sym (push (a , true))) (~ j) i))) + section-main (suc zero) αₙ (inr (push (a , true) i)) j = + hcomp (λ r → λ {(i = i0) → push (αₙ (a , true)) j + ; (i = i1) → compPath-filler (push (αₙ (a , true))) + (λ i₁ → inr (push (a , true) i₁)) r j + ; (j = i0) → S.BTC (suc zero) αₙ + (compPath-filler' + (push a) + (λ i → inr (a , σSn zero true i)) r i) + ; (j = i1) → inr (push (a , true) (i ∧ r))}) + (push (αₙ (a , true)) j) + section-main (suc (suc n)) αₙ (inr (push a i)) j = + hcomp (λ r → + λ {(i = i0) → push (αₙ a) j + ; (i = i1) → (push (αₙ (fst a , ptSn (suc n))) + ∙ (λ i₁ → inr (push (fst a , ptSn (suc n)) i₁))) j + ; (j = i0) → S.BTC (suc (suc n)) αₙ + (compPath-filler' (push (fst a)) + (λ i → inr (fst a , σSn (suc n) (snd a) i)) r i) + ; (j = i1) → inr (push a i)}) + (hcomp (λ r → + λ {(i = i0) → doubleCompPath-filler (push (αₙ a)) + (λ j → inr ((push a ∙ sym (push (fst a , ptSn (suc n)))) j)) + (sym (push (αₙ (fst a , ptSn (suc n))))) (~ j) (~ r) + ; (i = i1) → compPath-filler (push (αₙ (fst a , ptSn (suc n)))) + (λ i → inr (push (fst a , ptSn (suc n)) i)) r j + ; (j = i0) → S.BTC (suc (suc n)) αₙ (inr (fst a + , compPath-filler' (merid (snd a)) + (sym (merid (ptSn (suc n)))) r i)) + ; (j = i1) → inr (compPath-filler' + (push a) + (sym (push (fst a , ptSn (suc n)))) (~ i) (~ r))}) + (hcomp (λ r + → λ {(i = i0) → push (αₙ (fst a , ptSn (suc n))) (j ∨ ~ r) + ; (i = i1) → push (αₙ (fst a , ptSn (suc n))) (j ∨ ~ r) + ; (j = i1) → inr (inl (αₙ (fst a , ptSn (suc n))))}) + (inr (ugh (push (fst a , ptSn (suc n))) j i)))) + where + ugh : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → p ∙' sym p ≡ refl + ugh p = sym (compPath≡compPath' p (sym p)) ∙ rCancel p + section-main n αₙ (push a i) j = push a (i ∧ j) + + +-- main theorem +module _ {Cₙ Cₙ₊₁ : Type ℓ} (n mₙ : ℕ) + (αₙ : Fin mₙ × S⁻ n → Cₙ) + (e : Cₙ₊₁ ≃ Pushout αₙ fst) where + + open BouquetFuns n mₙ αₙ e + + BouquetIso-gen : Iso (cofib (invEq e ∘ inl)) (SphereBouquet n (Fin mₙ)) + Iso.fun BouquetIso-gen = CTB + Iso.inv BouquetIso-gen = BTC + Iso.rightInv BouquetIso-gen = CTB-BTC-cancel n mₙ αₙ e .fst + Iso.leftInv BouquetIso-gen = CTB-BTC-cancel n mₙ αₙ e .snd + + Bouquet≃-gen : cofib (invEq e ∘ inl) ≃ SphereBouquet n (Fin mₙ) + Bouquet≃-gen = isoToEquiv BouquetIso-gen diff --git a/Cubical/HITs/Susp/Base.agda b/Cubical/HITs/Susp/Base.agda index 98a456bee8..184d4a5b03 100644 --- a/Cubical/HITs/Susp/Base.agda +++ b/Cubical/HITs/Susp/Base.agda @@ -36,6 +36,12 @@ suspFun f north = north suspFun f south = south suspFun f (merid a i) = merid (f a) i +-- pointed version +suspFun∙ : {A B : Type ℓ} (f : A → B) + → Susp∙ A →∙ Susp∙ B +fst (suspFun∙ f) = suspFun f +snd (suspFun∙ f) = refl + UnitIsoSuspUnit : Iso Unit (Susp Unit) fun UnitIsoSuspUnit _ = north inv UnitIsoSuspUnit _ = tt diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index 4a9358a86d..6c043cf3b9 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -9,6 +9,7 @@ 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 @@ -20,9 +21,26 @@ open import Cubical.Homotopy.Loopspace private variable ℓ : Level + A B C : Type ℓ open Iso +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 : (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 : 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/HITs/Wedge.agda b/Cubical/HITs/Wedge.agda index 1cf2e97475..0b221bb806 100644 --- a/Cubical/HITs/Wedge.agda +++ b/Cubical/HITs/Wedge.agda @@ -2,3 +2,4 @@ module Cubical.HITs.Wedge where open import Cubical.HITs.Wedge.Base public +open import Cubical.HITs.Wedge.Properties public diff --git a/Cubical/HITs/Wedge/Properties.agda b/Cubical/HITs/Wedge/Properties.agda new file mode 100644 index 0000000000..5073e6316d --- /dev/null +++ b/Cubical/HITs/Wedge/Properties.agda @@ -0,0 +1,167 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Wedge.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.HITs.Pushout.Base +open import Cubical.HITs.Wedge.Base +open import Cubical.HITs.Susp + +open import Cubical.Homotopy.Loopspace + + +private + variable + ℓ ℓ' : Level + +-- Susp (⋁ᵢ Aᵢ) ≃ ⋁ᵢ (Susp Aᵢ) +private + Bouquet→ΩBouquetSusp-filler : (A : Type ℓ) (B : A → Pointed ℓ') + → (a : _) → (i j k : I) → ⋁gen A (λ a → Susp∙ (fst (B a))) + Bouquet→ΩBouquetSusp-filler A B a i j k = + hfill (λ k → λ {(i = i0) → inl tt + ; (i = i1) → doubleCompPath-filler + (push a) + (λ i → inr (a + , rCancel' (merid (snd (B a))) (~ k) i)) + (sym (push a)) k j + ; (j = i0) → push a (~ k ∧ i) + ; (j = i1) → push a (~ k ∧ i)}) + (inS (push a i)) + k + +Bouquet→ΩBouquetSusp : (A : Type ℓ) (B : A → Pointed ℓ') + → ⋁gen A B → Ω (⋁gen∙ A (λ a → Susp∙ (fst (B a)))) .fst +Bouquet→ΩBouquetSusp A B (inl x) = refl +Bouquet→ΩBouquetSusp A B (inr (a , b)) = + (push a ∙∙ (λ i → inr (a , toSusp (B a) b i)) ∙∙ sym (push a)) +Bouquet→ΩBouquetSusp A B (push a i) j = Bouquet→ΩBouquetSusp-filler A B a i j i1 + +SuspBouquet→Bouquet : (A : Type ℓ) (B : A → Pointed ℓ') + → Susp (⋁gen A B) → ⋁gen A (λ a → Susp∙ (fst (B a))) +SuspBouquet→Bouquet A B north = inl tt +SuspBouquet→Bouquet A B south = inl tt +SuspBouquet→Bouquet A B (merid a i) = Bouquet→ΩBouquetSusp A B a i + +Bouquet→SuspBouquet : (A : Type ℓ) (B : A → Pointed ℓ') + → ⋁gen A (λ a → Susp∙ (fst (B a))) → Susp (⋁gen A B) +Bouquet→SuspBouquet A B (inl x) = north +Bouquet→SuspBouquet A B (inr (a , north)) = north +Bouquet→SuspBouquet A B (inr (a , south)) = south +Bouquet→SuspBouquet A B (inr (a , merid b i)) = merid (inr (a , b)) i +Bouquet→SuspBouquet A B (push a i) = north + +SuspBouquet-Bouquet-cancel : (A : Type ℓ) (B : A → Pointed ℓ') + → section (SuspBouquet→Bouquet A B) (Bouquet→SuspBouquet A B) + × retract (SuspBouquet→Bouquet A B) (Bouquet→SuspBouquet A B) +SuspBouquet-Bouquet-cancel A B = sec , ret + where + sec : section (SuspBouquet→Bouquet A B) (Bouquet→SuspBouquet A B) + sec (inl tt) i = inl tt + sec (inr (a , north)) = push a + sec (inr (a , south)) = + (push a) + ∙∙ (λ i → inr (a , merid (pt (B a)) i)) + ∙∙ (λ i → inr (a , south)) + sec (inr (a , merid b j)) i = + hcomp (λ k → λ {(~ i ∧ j = i1) → push a (~ k) + ; (i = i1) → inr (a , merid b j) + ; (j = i0) → push a (i ∨ (~ k)) }) + (inr (a , (hcomp (λ k → λ {(i = i1) → merid b j + ; (j = i0) → north + ; (j = i1) → merid (pt (B a)) (i ∨ (~ k))}) + (merid b j)))) + sec (push a j) i = push a (i ∧ j) + + module _ (a : A) (b : fst (B a)) (i j : I) where + ret-fill₁ : I → Susp (⋁gen A B) + ret-fill₁ k = + hfill (λ k → λ {(j = i0) → north + ; (j = i1) → merid (inr (a , pt (B a))) ((~ k) ∨ i) + ; (i = i0) → Bouquet→SuspBouquet A B + (inr (a , compPath-filler (merid b) + (sym (merid (pt (B a)))) k j)) + ; (i = i1) → merid (inr (a , b)) j}) + (inS (merid (inr (a , b)) j)) k + + ret-fill₂ : I → Susp (⋁gen A B) + ret-fill₂ k = + hfill (λ k → λ {(j = i0) → north + ; (j = i1) → merid (push a (~ k)) i + ; (i = i0) → Bouquet→SuspBouquet A B + (doubleCompPath-filler (push a) + (λ i → inr (a , toSusp (B a) b i)) + (sym (push a)) k j) + ; (i = i1) → merid (inr (a , b)) j}) + (inS (ret-fill₁ i1)) k + + ret : retract (SuspBouquet→Bouquet A B) (Bouquet→SuspBouquet A B) + ret north i = north + ret south = merid (inl tt) + ret (merid (inl tt) j) i = merid (inl tt) (i ∧ j) + ret (merid (inr (a , b)) j) i = ret-fill₂ a b i j i1 + ret (merid (push a k) j) i = + hcomp (λ r → λ {(i = i0) → Bouquet→SuspBouquet A B + (Bouquet→ΩBouquetSusp-filler A B a k j r) + ; (i = i1) → merid (push a (~ r ∨ k)) j + ; (j = i0) → north + ; (j = i1) → merid (push a (~ r)) i + ; (k = i0) → merid (push a (~ r)) (i ∧ j) + ; (k = i1) → side r i j} + ) + (merid (inr (a , pt (B a))) (i ∧ j)) + where + side : Cube {A = Susp (⋁gen A B)} + (λ i j → merid (inr (a , pt (B a))) (i ∧ j)) + (λ i j → ret-fill₂ a (pt (B a)) i j i1) + (λ r j → Bouquet→SuspBouquet A B + (Bouquet→ΩBouquetSusp-filler A B a i1 j r)) + (λ r j → merid (inr (a , (pt (B a)))) j) + (λ r i → north) + λ r i → merid (push a (~ r)) i + side r i j = + hcomp (λ k + → λ {(r = i0) → Bouquet→SuspBouquet A B + (inr (a , rCancel-filler' + (merid (pt (B a))) i k j)) + ; (r = i1) → ret-fill₂ a (pt (B a)) i j k + ; (i = i0) → Bouquet→SuspBouquet A B + (doubleCompPath-filler + (push a) + (λ j → inr (a + , rCancel' (merid (pt (B a))) (~ r ∧ k) j)) + (sym (push a)) (r ∧ k) j) + ; (i = i1) → merid (inr (a , snd (B a))) j + ; (j = i0) → north + ; (j = i1) → merid (push a (~ r ∨ ~ k)) i}) + (hcomp (λ k + → λ {(r = i0) → Bouquet→SuspBouquet A B + (inr (a , rCancel-filler' + (merid (pt (B a))) (~ k ∨ i) i0 j)) + ; (r = i1) → ret-fill₁ a (pt (B a)) i j k + ; (i = i0) → Bouquet→SuspBouquet A B + (inr (a , compPath-filler + (merid (pt (B a))) + (sym (merid (pt (B a)))) k j)) + ; (i = i1) → merid (inr (a , snd (B a))) j + ; (j = i0) → north + ; (j = i1) → merid (inr (a , snd (B a))) (~ k ∨ i)}) + (merid (inr (a , snd (B a))) j)) + +Iso-SuspBouquet-Bouquet : (A : Type ℓ) (B : A → Pointed ℓ') + → Iso (Susp (⋁gen A B)) (⋁gen A (λ a → Susp∙ (fst (B a)))) +Iso.fun (Iso-SuspBouquet-Bouquet A B) = SuspBouquet→Bouquet A B +Iso.inv (Iso-SuspBouquet-Bouquet A B) = Bouquet→SuspBouquet A B +Iso.rightInv (Iso-SuspBouquet-Bouquet A B) = SuspBouquet-Bouquet-cancel A B .fst +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) 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₂) diff --git a/Cubical/Homotopy/Group/PinSn.agda b/Cubical/Homotopy/Group/PinSn.agda index a6064137d2..eeb870c61d 100644 --- a/Cubical/Homotopy/Group/PinSn.agda +++ b/Cubical/Homotopy/Group/PinSn.agda @@ -3,6 +3,8 @@ This file contains: 1. An iso πₙ'(Sⁿ) ≅ ℤ, where πₙ'(Sⁿ) = ∥ Sⁿ →∙ Sⁿ ∥₀ 2. A proof that idfun∙ : Sⁿ →∙ Sⁿ is the generator of πₙ'(Sⁿ) +3. An explicit construction of iso Hⁿ(Sⁿ,ℤ) ≅ πₙ'(Sⁿ) +4. Description of multiplication on πₙ'(Sⁿ) (function composition) -} module Cubical.Homotopy.Group.PinSn where @@ -18,13 +20,13 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Path open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv -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.PropositionalTruncation as PT +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as TR +open import Cubical.HITs.S1 renaming (_·_ to _*_) open import Cubical.HITs.Sn open import Cubical.HITs.Susp @@ -40,6 +42,10 @@ open import Cubical.Algebra.Group.ZAction open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Sn + open Iso -- The goal is to prove that πₙSⁿ ≅ ℤ. This is of course trivial, given @@ -78,7 +84,7 @@ 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) + makeIsGroupHom (ST.elim2 (λ _ _ → isSetPathImplicit) λ f g → IsGroupHom.pres· (suspMapπ'Hom (suc n) .snd) ∣ f ∣₂ ∣ g ∣₂) private @@ -88,19 +94,19 @@ private ∙∙ rCancel (merid base) ∣₂ stLoop₁flip : π 2 (S₊∙ 2) - stLoop₁flip = sMap flipSquare stLoop₁ + stLoop₁flip = ST.map 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 +fun (fst (flipSquareIso n)) = ST.map flipSquare +inv (fst (flipSquareIso n)) = ST.map flipSquare rightInv (fst (flipSquareIso n)) = - sElim (λ _ → isSetPathImplicit) λ _ → refl + ST.elim (λ _ → isSetPathImplicit) λ _ → refl leftInv (fst (flipSquareIso n)) = - sElim (λ _ → isSetPathImplicit) λ _ → refl + ST.elim (λ _ → isSetPathImplicit) λ _ → refl snd (flipSquareIso n) = makeIsGroupHom - (sElim2 (λ _ _ → isSetPathImplicit) + (ST.elim2 (λ _ _ → isSetPathImplicit) λ f g → cong ∣_∣₂ ((sym (sym≡flipSquare (f ∙ g)) ∙∙ symDistr f g @@ -125,7 +131,7 @@ fst π₂S²≅π₁S¹ = (invIso setTruncTrunc2Iso)) snd π₂S²≅π₁S¹ = makeIsGroupHom - (sElim2 (λ _ _ → isSetPathImplicit) + (ST.elim2 (λ _ _ → isSetPathImplicit) λ f g → cong (inv setTruncTrunc2Iso) (cong (fun (PathIdTruncIso 2)) @@ -150,10 +156,10 @@ snd π₂S²≅π₁S¹ = setTruncTrunc2IsoFunct : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : hLevelTrunc 2 (x ≡ x)) → inv setTruncTrunc2Iso - (Cubical.HITs.Truncation.map2 _∙_ p q) + (TR.map2 _∙_ p q) ≡ ·π 0 (inv setTruncTrunc2Iso p) (inv setTruncTrunc2Iso q) setTruncTrunc2IsoFunct = - trElim2 (λ _ _ → isSetPathImplicit) λ _ _ → refl + TR.elim2 (λ _ _ → isSetPathImplicit) λ _ _ → refl π₂'S²≅π₁'S¹ : GroupIso (π'Gr 1 (S₊∙ 2)) (π'Gr 0 (S₊∙ 1)) π₂'S²≅π₁'S¹ = @@ -167,7 +173,7 @@ snd π₂S²≅π₁S¹ = πₙ'Sⁿ≅ℤ zero = compGroupIso (π'Gr≅πGr zero (S₊∙ 1)) ((compIso (setTruncIdempotentIso (isGroupoidS¹ _ _)) ΩS¹Isoℤ) - , makeIsGroupHom (sElim2 (λ _ _ → isProp→isSet (isSetℤ _ _)) + , makeIsGroupHom (ST.elim2 (λ _ _ → isProp→isSet (isSetℤ _ _)) winding-hom)) πₙ'Sⁿ≅ℤ (suc zero) = compGroupIso π₂'S²≅π₁'S¹ (πₙ'Sⁿ≅ℤ zero) πₙ'Sⁿ≅ℤ (suc (suc n)) = @@ -334,3 +340,226 @@ private (pos (suc zero)) (λ h → h , (sym (·Comm h (pos 1)) ∙ ℤ·≡· h (pos 1))) (invGroupIso (πₙ'Sⁿ≅ℤ n))) + +-- Isomorphism πₙ(Sⁿ) ≅ Hⁿ(Sⁿ,ℤ) +private + makePted : (n : ℕ) (fn : S₊ (2 + n)) → ∥ fn ≡ north ∥₂ + makePted n fn = + TR.rec (isOfHLevelPlus' 2 squash₂) ∣_∣₂ + (isConnectedPathSⁿ (suc n) fn north .fst) + makePted-eq : (n : ℕ) (fn : S₊ (2 + n)) (p : fn ≡ north) → makePted n fn ≡ ∣ p ∣₂ + makePted-eq n fn p j = + TR.rec (isOfHLevelPlus' 2 squash₂) ∣_∣₂ (isConnectedPathSⁿ (suc n) fn north .snd ∣ p ∣ j) + +-- Forgetting pointedness gives iso +πₙSⁿ-unpoint : (n : ℕ) → (π'Gr n (S₊∙ (suc n)) .fst) → ∥ (S₊ (suc n) → S₊ (suc n)) ∥₂ +πₙSⁿ-unpoint n = ST.map fst + +isIso-πₙSⁿ-unpointIso : (n : ℕ) → isIso (πₙSⁿ-unpoint n) +fst (isIso-πₙSⁿ-unpointIso zero) = + ST.map λ f → (λ x → f x * (invLooper (f base))) + , sym (rCancelS¹ (f base)) +fst (snd (isIso-πₙSⁿ-unpointIso zero)) = + ST.elim (λ _ → isSetPathImplicit) + λ f → PT.rec (squash₂ _ _) + (λ q → cong ∣_∣₂ (funExt λ x → cong (f x *_) + (cong invLooper (sym q)) ∙ rUnitS¹ (f x))) + (isConnectedS¹ (f base)) +snd (snd (isIso-πₙSⁿ-unpointIso zero)) = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ + (ΣPathP ((funExt + (λ r → cong (fst f r *_) (cong invLooper (snd f)) + ∙ rUnitS¹ (fst f r))) + , help _ (sym (snd f)))) + where + help : (x : _) (p : base ≡ x) + → PathP (λ i → ((λ j → x * invLooper (p (~ j))) ∙ rUnitS¹ x) i ≡ base) + (sym (rCancelS¹ x)) (sym p) + help = J> λ 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/README.agda b/Cubical/README.agda index b707a2e03a..bfe8dee179 100644 --- a/Cubical/README.agda +++ b/Cubical/README.agda @@ -85,3 +85,6 @@ import Cubical.Axiom.Everything -- Automatic proving, solvers import Cubical.Tactics.Everything + +-- CW complexes, cellular (co)homology +import Cubical.CW.Everything diff --git a/Cubical/ZCohomology/GroupStructure.agda b/Cubical/ZCohomology/GroupStructure.agda index 7944f1d9b2..156aa1c47d 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.Inductive 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 5b96b28a42d875ac4f8ed2ff0322152e2cb6807a Mon Sep 17 00:00:00 2001 From: Andrew Swan Date: Mon, 13 May 2024 11:24:22 +0200 Subject: [PATCH 06/30] Topological modalities (#1125) * add basic theory for topological modalities * a couple of useful results about topological modalities * comments and some reformatting * replace implicit variables with declared variables * add topological to list of imports in Cubical.HITs.Nullification * make level declared variable private * suggested corrections --- Cubical/HITs/Nullification.agda | 2 + Cubical/HITs/Nullification/Properties.agda | 92 ++++++++++++---- Cubical/HITs/Nullification/Topological.agda | 110 ++++++++++++++++++++ 3 files changed, 181 insertions(+), 23 deletions(-) create mode 100644 Cubical/HITs/Nullification/Topological.agda diff --git a/Cubical/HITs/Nullification.agda b/Cubical/HITs/Nullification.agda index 5a3c2be8b6..90e090e756 100644 --- a/Cubical/HITs/Nullification.agda +++ b/Cubical/HITs/Nullification.agda @@ -4,3 +4,5 @@ module Cubical.HITs.Nullification where open import Cubical.HITs.Nullification.Base public open import Cubical.HITs.Nullification.Properties public + +open import Cubical.HITs.Nullification.Topological public diff --git a/Cubical/HITs/Nullification/Properties.agda b/Cubical/HITs/Nullification/Properties.agda index 65a8896d72..4032d28d9d 100644 --- a/Cubical/HITs/Nullification/Properties.agda +++ b/Cubical/HITs/Nullification/Properties.agda @@ -9,6 +9,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.Equiv.PathSplit open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence open import Cubical.Functions.FunExtEquiv @@ -18,6 +19,7 @@ open import Cubical.Functions.FunExtEquiv open import Cubical.HITs.Localization renaming (rec to Localize-rec) open import Cubical.Data.Unit open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ; zero; suc) open import Cubical.HITs.Nullification.Base @@ -27,15 +29,17 @@ open isPathSplitEquiv private variable ℓα ℓs ℓ ℓ' : Level + A : Type ℓα + S : A → Type ℓs + X : Type ℓ -isNull≡ : {A : Type ℓα} {S : A → Type ℓs} {X : Type ℓ} (nX : isNull S X) {x y : X} → isNull S (x ≡ y) -isNull≡ {A = A} {S = S} nX {x = x} {y = y} α = +isNull≡ : (nX : isNull S X) {x y : X} → isNull S (x ≡ y) +isNull≡ nX α = fromIsEquiv (λ p _ i → p i) (isEquiv[equivFunA≃B∘f]→isEquiv[f] (λ p _ → p) funExtEquiv (isEquivCong (const , toIsEquiv _ (nX α)))) -isNullΠ : {A : Type ℓα} {S : A → Type ℓs} {X : Type ℓ} {Y : X → Type ℓ'} → ((x : X) → isNull S (Y x)) → - isNull S ((x : X) → Y x) -isNullΠ {S = S} {X = X} {Y = Y} nY α = fromIsEquiv _ (snd e) +isNullΠ : {Y : X → Type ℓ'} → ((x : X) → isNull S (Y x)) → isNull S ((x : X) → Y x) +isNullΠ {X = X} {S = S} {Y = Y} nY α = fromIsEquiv _ (snd e) where flipIso : Iso ((x : X) → S α → Y x) (S α → (x : X) → Y x) Iso.fun flipIso f = flip f @@ -52,9 +56,9 @@ isNullΠ {S = S} {X = X} {Y = Y} nY α = fromIsEquiv _ (snd e) (S α → ((x : X) → Y x)) ■ -isNullΣ : {A : Type ℓα} {S : A → Type ℓs} {X : Type ℓ} {Y : X → Type ℓ'} → (isNull S X) → ((x : X) → isNull S (Y x)) → +isNullΣ : {Y : X → Type ℓ'} → (isNull S X) → ((x : X) → isNull S (Y x)) → isNull S (Σ X Y) -isNullΣ {S = S} {X = X} {Y = Y} nX nY α = fromIsEquiv _ (snd e) +isNullΣ {X = X} {S = S} {Y = Y} nX nY α = fromIsEquiv _ (snd e) where e : Σ X Y ≃ (S α → Σ X Y) e = @@ -67,15 +71,21 @@ isNullΣ {S = S} {X = X} {Y = Y} nX nY α = fromIsEquiv _ (snd e) (S α → Σ X Y) ■ -rec : ∀ {ℓα ℓs ℓ ℓ'} {A : Type ℓα} {S : A → Type ℓs} {X : Type ℓ} {Y : Type ℓ'} - → (nB : isNull S Y) → (X → Y) → Null S X → Y +equivPreservesIsNull : {Y : Type ℓ'} → (e : X ≃ Y) → (isNull S X) → isNull S Y +equivPreservesIsNull e nullX α = + fromIsEquiv _ + (isEquiv[f∘equivFunA≃B]→isEquiv[f] + (λ y _ → y) e + (snd (compEquiv (pathSplitToEquiv ((λ x _ → x) , (nullX α))) (postCompEquiv e)))) + +rec : {Y : Type ℓ'} → (nB : isNull S Y) → (X → Y) → Null S X → Y rec nB g ∣ x ∣ = g x rec nB g (hub α f) = fst (sec (nB α)) (λ s → rec nB g (f s)) rec nB g (spoke α f s i) = snd (sec (nB α)) (λ s → rec nB g (f s)) i s rec nB g (≡hub {x} {y} {α} p i) = fst (secCong (nB α) (rec nB g x) (rec nB g y)) (λ i s → rec nB g (p s i)) i rec nB g (≡spoke {x} {y} {α} p s i j) = snd (secCong (nB α) (rec nB g x) (rec nB g y)) (λ i s → rec nB g (p s i)) i j s -toPathP⁻ : ∀ {ℓ} (A : I → Type ℓ) {x : A i0} {y : A i1} → x ≡ transport⁻ (λ i → A i) y → PathP A x y +toPathP⁻ : (A : I → Type ℓ) {x : A i0} {y : A i1} → x ≡ transport⁻ (λ i → A i) y → PathP A x y toPathP⁻ A p i = toPathP {A = λ i → A (~ i)} (sym p) (~ i) toPathP⁻-sq : ∀ {ℓ} {A : Type ℓ} (x : A) → Square (toPathP⁻ (λ _ → A) (λ _ → transport refl x)) refl @@ -84,7 +94,7 @@ toPathP⁻-sq x j i = hcomp (λ l → λ { (i = i0) → transportRefl x j ; (i = i1) → x ; (j = i1) → x }) (transportRefl x (i ∨ j)) -module _ {ℓα ℓs ℓ ℓ'} {A : Type ℓα} {S : A → Type ℓs} {X : Type ℓ} {Y : Null S X → Type ℓ'} where +module _ {Y : Null S X → Type ℓ'} where private secCongDep' : ∀ (nY : (x : Null S X) → isNull S (Y x)) {x y : Null S X} {α} (p : x ≡ y) @@ -128,29 +138,55 @@ module _ {ℓα ℓs ℓ ℓ'} {A : Type ℓα} {S : A → Type ℓs} {X : Type (λ i → elim nY g (p s i)) q₂ j i = toPathP⁻ (λ j → Y (≡spoke p s j i)) (λ j → q₁ j i) j -NullRecIsPathSplitEquiv : ∀ {ℓα ℓs ℓ ℓ'} {A : Type ℓα} {S : A → Type ℓs} {X : Type ℓ} {Y : Type ℓ'} → (isNull S Y) → - isPathSplitEquiv {A = (Null S X) → Y} (λ f → f ∘ ∣_∣) +NullRecIsPathSplitEquiv : {Y : Type ℓ'} → (isNull S Y) → isPathSplitEquiv {A = (Null S X) → Y} (λ f → f ∘ ∣_∣) sec (NullRecIsPathSplitEquiv nY) = rec nY , λ _ → refl secCong (NullRecIsPathSplitEquiv nY) f f' = (λ p → funExt (elim (λ _ → isNull≡ nY) (funExt⁻ p))) , λ _ → refl -NullRecIsEquiv : ∀ {ℓα ℓs ℓ ℓ'} {A : Type ℓα} {S : A → Type ℓs} {X : Type ℓ} {Y : Type ℓ'} → (isNull S Y) → - isEquiv {A = (Null S X) → Y} (λ f → f ∘ ∣_∣) +NullRecIsEquiv : {Y : Type ℓ'} → (isNull S Y) → isEquiv {A = (Null S X) → Y} (λ f → f ∘ ∣_∣) NullRecIsEquiv nY = toIsEquiv _ (NullRecIsPathSplitEquiv nY) -NullRecEquiv : ∀ {ℓα ℓs ℓ ℓ'} {A : Type ℓα} {S : A → Type ℓs} {X : Type ℓ} {Y : Type ℓ'} → (isNull S Y) → - ((Null S X) → Y) ≃ (X → Y) +NullRecEquiv : {Y : Type ℓ'} → (isNull S Y) → ((Null S X) → Y) ≃ (X → Y) NullRecEquiv nY = (λ f → f ∘ ∣_∣) , (NullRecIsEquiv nY) - -NullPreservesProp : ∀ {ℓα ℓs ℓ} {A : Type ℓα} {S : A → Type ℓs} {X : Type ℓ} → - (isProp X) → isProp (Null S X) +NullPreservesProp : isProp X → isProp (Null S X) NullPreservesProp {S = S} pX u = elim (λ v' → isNull≡ (isNull-Null S)) (λ y → elim (λ u' → isNull≡ (isNull-Null S) {x = u'}) (λ x → cong ∣_∣ (pX x y)) u) +NullPreservesContr : isContr X → isContr (Null S X) +NullPreservesContr l = inhProp→isContr ∣ fst l ∣ (NullPreservesProp (isContr→isProp l)) + +isPropIsNull : isProp (isNull S X) +isPropIsNull = isPropΠ (λ _ → isPropIsPathSplitEquiv _) + +{- + We check that a few common definitions in type theory are null, + assuming they are given null types as input. +-} +isNullIsContr : isNull S X → isNull S (isContr X) +isNullIsContr nullX = isNullΣ nullX λ _ → isNullΠ (λ _ → isNull≡ nullX) + +isNullIsEquiv : + {Y : Type ℓ'} (nullX : isNull S X) + (nullY : isNull S Y) (f : X → Y) → isNull S (isEquiv f) +isNullIsEquiv nullX nullY f = + equivPreservesIsNull (invEquiv (isEquiv≃isEquiv' f)) + (isNullΠ λ _ → isNullIsContr (isNullΣ nullX λ _ → isNull≡ nullY)) + +isNullEquiv : + ∀ {ℓα ℓs ℓ} {A : Type ℓα} {S : A → Type ℓs} + {X Y : Type ℓ} → isNull S X → isNull S Y → isNull S (X ≃ Y) +isNullEquiv nullX nullY = isNullΣ (isNullΠ (λ _ → nullY)) (isNullIsEquiv nullX nullY) + +isNullIsOfHLevel : (n : HLevel) → isNull S X → isNull S (isOfHLevel n X) +isNullIsOfHLevel zero nullX = isNullIsContr nullX +isNullIsOfHLevel (suc zero) nullX = isNullΠ (λ _ → isNullΠ (λ _ → isNull≡ nullX)) +isNullIsOfHLevel (suc (suc n)) nullX = + isNullΠ (λ _ → isNullΠ (λ _ → isNullIsOfHLevel (suc n) (isNull≡ nullX))) + -- nullification is a modality -NullModality : ∀ {ℓα ℓs ℓ} {A : Type ℓα} (S : A → Type ℓs) → Modality (ℓ-max ℓ (ℓ-max ℓα ℓs)) +NullModality : {A : Type ℓα} (S : A → Type ℓs) → Modality (ℓ-max ℓ (ℓ-max ℓα ℓs)) isModal (NullModality S) = isNull S isPropIsModal (NullModality S) = isPropΠ (λ α → isPropIsPathSplitEquiv _) ◯ (NullModality S) = Null S @@ -169,7 +205,7 @@ idemNull {ℓ = ℓ} S A nA = ∣_∣ , isModalToIsEquiv (NullModality {ℓ = -- nullification is localization at a family of maps (S α → 1) -module Null-iso-Localize {ℓα ℓs ℓ} {A : Type ℓα} (S : A → Type ℓs) (X : Type ℓ) where +module Null-iso-Localize (S : A → Type ℓs) (X : Type ℓ) where to : Null S X → Localize {A = A} (λ α → const {B = S α} tt) X to ∣ x ∣ = ∣ x ∣ @@ -202,5 +238,15 @@ module Null-iso-Localize {ℓα ℓs ℓ} {A : Type ℓα} (S : A → Type ℓs) isom : Iso (Null S X) (Localize {A = A} (λ α → const {B = S α} tt) X) isom = iso to from to-from from-to -Null≃Localize : ∀ {ℓα ℓs ℓ} {A : Type ℓα} (S : A → Type ℓs) (X : Type ℓ) → Null S X ≃ Localize (λ α → const {B = S α} tt) X +Null≃Localize : (S : A → Type ℓs) (X : Type ℓ) → Null S X ≃ Localize (λ α → const {B = S α} tt) X Null≃Localize S X = isoToEquiv (Null-iso-Localize.isom S X) + +SeparatedAndInjective→Null : + (X : Type ℓ) (sep : (x y : X) → isNull S (x ≡ y)) + (inj : (α : A) → hasSection (const {A = X} {B = S α})) → + isNull S X +sec (SeparatedAndInjective→Null X sep inj α) = inj α +secCong (SeparatedAndInjective→Null X sep inj α) x y = + fst s ∘ funExt⁻ , λ p i j α → snd s (funExt⁻ p) i α j + where + s = sec (sep x y α) diff --git a/Cubical/HITs/Nullification/Topological.agda b/Cubical/HITs/Nullification/Topological.agda new file mode 100644 index 0000000000..10204d56b7 --- /dev/null +++ b/Cubical/HITs/Nullification/Topological.agda @@ -0,0 +1,110 @@ +{-# OPTIONS --safe #-} +{- + A topological modality is by definition the nullification of a family of + propositions. They are always lex modalities. We show that together + with a couple of useful corollaries. We roughly follow the proofs + from Rijke, Shulman, Spitters, Modalities in homotopy type theory. +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Fiberwise +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.Unit +open import Cubical.Data.Nat using (ℕ; zero; suc) + +open import Cubical.HITs.Nullification.Base +open import Cubical.HITs.Nullification.Properties + +module Cubical.HITs.Nullification.Topological + {ℓα ℓs} + {A : Type ℓα} + (S : A → Type ℓs) + where + +private variable + ℓ ℓ' : Level + +{- + We use the formulation of lexness that the universe of null types is + null itself. + + We choose the level of the universe of null types to be as low as + possible such that it is null when the generators are propositions. +-} +NullType : Type (ℓ-max (ℓ-suc (ℓ-max ℓ ℓs)) ℓα ) +NullType {ℓ} = Σ[ X ∈ Type (ℓ-max ℓ ℓs) ] isNull S X + +{- + We can show that the universe is separated in general. +-} +NullType≡isNull : (X Y : NullType {ℓ = ℓ}) → isNull S (X ≡ Y) +NullType≡isNull {ℓ = ℓ} X Y = + equivPreservesIsNull (invEquiv e) (isNullEquiv (snd X) (snd Y)) + where + e : (X ≡ Y) ≃ ((fst X) ≃ (fst Y)) + e = + X ≡ Y ≃⟨ invEquiv (Σ≡PropEquiv λ _ → isPropIsNull) ⟩ + (fst X) ≡ (fst Y) ≃⟨ univalence ⟩ + (fst X) ≃ (fst Y) ■ + +module _ (isPropS : (α : A) → isProp (S α)) where + {- Recall that a type Z is injective when we can extend any map S α → Z to + an element of Z. We show this is the case for Z = NullType. + -} + NullTypeIsInj : (α : A) → hasSection (const {A = NullType {ℓ = ℓ}} {B = S α}) + fst (NullTypeIsInj α) f = ((z : S α) → fst (f z)) , isNullΠ (λ z → snd (f z)) + snd (NullTypeIsInj α) f = funExt (λ z → Σ≡Prop (λ _ → isPropIsNull) (ua (e z))) + where + e : (z : S α) → ((z : S α) → fst (f z)) ≃ fst (f z) + e z = + ((z : S α) → fst (f z)) + ≃⟨ invEquiv + (equivΠ + (invEquiv + (isContr→≃Unit + (inhProp→isContr z (isPropS α)))) λ _ → idEquiv _) ⟩ + (Unit → fst (f z)) + ≃⟨ ΠUnit (λ _ → fst (f z)) ⟩ + fst (f z) + ■ + + isNullNullTypes : isNull S (NullType {ℓ = ℓ}) + isNullNullTypes {ℓ} = + SeparatedAndInjective→Null (NullType {ℓ = ℓ}) + (NullType≡isNull {ℓ = ℓ}) (NullTypeIsInj {ℓ = ℓ}) + + topUnitWeaklyEmb : {X : Type ℓ} (x y : X) → Path (Null S X) ∣ x ∣ ∣ y ∣ ≃ Null S (x ≡ y) + topUnitWeaklyEmb {ℓ} {X} x y = extensionToE ∣ y ∣ + where + E : (Null S X) → NullType {ℓ = ℓ-max ℓ ℓα} + E = rec (isNullNullTypes {ℓ = ℓ-max ℓ ℓα}) (λ y' → (Null S (x ≡ y')) , isNull-Null S) + + extensionToE : (z : Null S X) → (∣ x ∣ ≡ z) ≃ fst (E z) + extensionToE = recognizeId (fst ∘ E) ∣ refl ∣ + ((∣ x ∣ , ∣ refl ∣) , + (uncurry + (elim (λ _ → isNullΠ isnull) + λ y → elim isnull (J (λ y p → (∣ x ∣ , ∣ refl ∣) ≡ (∣ y ∣ , ∣ p ∣)) refl)))) + + where + isnull : {z : Null S X} → (e : fst (E z)) → + isNull S (Path (Σ _ (fst ∘ E)) (∣ x ∣ , ∣ refl ∣) (z , e)) + isnull e = isNull≡ (isNullΣ (isNull-Null S) (λ z → snd (E z))) + + topUnitWeaklyInj : {X : Type ℓ} (x y : X) → Path (Null S X) ∣ x ∣ ∣ y ∣ → Null S (x ≡ y) + topUnitWeaklyInj x y = fst (topUnitWeaklyEmb x y) + + topPreservesHLevel : {X : Type ℓ} (n : HLevel) → (isOfHLevel n X) → isOfHLevel n (Null S X) + topPreservesHLevel zero = NullPreservesContr + topPreservesHLevel (suc zero) = NullPreservesProp + topPreservesHLevel (suc (suc n)) l = + elim (λ x → isNullΠ (λ y → isNullIsOfHLevel _ (isNull≡ (isNull-Null S)))) + (λ x → elim (λ y → isNullIsOfHLevel _ (isNull≡ (isNull-Null S))) + (λ y → isOfHLevelRespectEquiv (suc n) (invEquiv (topUnitWeaklyEmb x y)) + (topPreservesHLevel (suc n) (l x y)))) From bcb4a68c7cec164bbf71be7012fabd4e20fcc959 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 13 May 2024 14:00:38 +0200 Subject: [PATCH 07/30] Update CONTRIBUTING.md --- CONTRIBUTING.md | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/CONTRIBUTING.md b/CONTRIBUTING.md index 45ed411916..bbd4d0408c 100644 --- a/CONTRIBUTING.md +++ b/CONTRIBUTING.md @@ -99,10 +99,9 @@ When preparing a PR here are some general guidelines: 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`. +- Unless a file is in the `Core`, `Foundations` or `Codata` package you + don't need to add it manually to the `Everything` file as it is + automatically generated when running `make`. - For folders with `Base` and `Properties` submodules, the `Base` file can contain some basic consequences of the main definition, but From 2bd09eb58a09f6473090038f5991ed2ece83e267 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 13 May 2024 14:05:51 +0200 Subject: [PATCH 08/30] Update CONTRIBUTING.md Decided with Anders, that we do not want to discourage bundling of imports. --- CONTRIBUTING.md | 6 ------ 1 file changed, 6 deletions(-) diff --git a/CONTRIBUTING.md b/CONTRIBUTING.md index bbd4d0408c..195698a4e9 100644 --- a/CONTRIBUTING.md +++ b/CONTRIBUTING.md @@ -107,12 +107,6 @@ When preparing a PR here are some general guidelines: can contain some basic consequences of the main definition, but shouldn't include theorems that would require additional imports. -- Avoid importing `Foundations.Everything`; import only the modules in - `Foundations` you are using. Be reasonably specific in general when - importing. - In particular, avoid importing useless files or useless renaming - and try to group them by folder like `Foundations` or `Data` - - Avoid `public` imports, except in modules that are specifically meant to collect and re-export results from several modules. From cc3bf676de0ebbaf56d51ae96e76a31d79c07824 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 13 May 2024 14:09:44 +0200 Subject: [PATCH 09/30] Update CONTRIBUTING.md --- CONTRIBUTING.md | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/CONTRIBUTING.md b/CONTRIBUTING.md index 195698a4e9..c1ae237c13 100644 --- a/CONTRIBUTING.md +++ b/CONTRIBUTING.md @@ -80,6 +80,10 @@ When preparing a PR here are some general guidelines: - Use copattern-matching when instantiating records for efficiency. This seems especially important when constructing Iso's. +- New records should be defined with the ```no-eta-equality``` modifier, + unless you have a good reason to drop it - keep in mind that dropping it + can lead to type checking speed problems. + - 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 From 746071b5e6e0988ef5061c757fae9c8d4680bacb Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 13 May 2024 14:16:50 +0200 Subject: [PATCH 10/30] Update Review table --- README.md | 20 +++++++++----------- 1 file changed, 9 insertions(+), 11 deletions(-) diff --git a/README.md b/README.md index 8d7e2654b6..384c186946 100644 --- a/README.md +++ b/README.md @@ -71,7 +71,7 @@ Huber, Anders Mörtberg. This makes it possible to directly represent higher inductive types. -Reviewing of pull requests +Reviewing of [pull requests](https://github.com/agda/cubical/pulls?q=is%3Apr+is%3Aopen+draft%3Afalse) -------------------------- If you switch your draft pull request (PR) to 'ready to merge', or directly create an open PR, @@ -79,15 +79,13 @@ we should request a review, by one of the reviewers below. If that doesn't happen, you can also request a reviewer yourself (for reviewer expertise see below), to make us aware of the open PR. Feel free to use Discord to get in touch with a reviewer in case reviewing is taking a very long time. -| Reviewer | github handle | Area of expertise | Queue | -|-------------------------------------------------------------------------|---------------|---------------------------------------------|------| -| [Anders Mörtberg](https://staff.math.su.se/anders.mortberg/) | [mortberg](https://github.com/mortberg) | *Most topics* | [link](https://github.com/agda/cubical/pulls?q=is%3Apr+review-requested%3Amortberg+) | -| [Evan Cavallo](https://staff.math.su.se/evan.cavallo/) | [ecavallo](https://github.com/ecavallo) | *Most topics* | [link](https://github.com/agda/cubical/pulls?q=is%3Apr+review-requested%3Aecavallo+) | -| [Felix Cherubini](https://felix-cherubini.de) | [felixwellen](https://github.com/felixwellen) | *Mainly algebra related topics* | [link](https://github.com/agda/cubical/pulls?q=is%3Apr+review-requested%3Afelixwellen+) | -| [Max Zeuner](https://www.su.se/english/profiles/maze1512-1.450461) | [mzeuner](https://github.com/mzeuner) | *Algebra related topics* | [link](https://github.com/agda/cubical/pulls?q=is%3Apr+review-requested%3Amzeuner+) | -| [Axel Ljungström](https://www.su.se/english/profiles/axlj4439-1.450268) | [aljungstrom](https://github.com/aljungstrom) | *Synthetic homotopy theory and cohomology* | [link](https://github.com/agda/cubical/pulls?q=is%3Apr+review-requested%3Aaljungstrom+) | -| [Andrea Vezzosi](http://saizan.github.io/) | [Saizan](https://github.com/Saizan) | *Inactive* | [link](https://github.com/agda/cubical/pulls?q=is%3Apr+review-requested%3ASaizan+) | +| Reviewer | github handle | Area of expertise | +|-------------------------------------------------------------------------|---------------|---------------------------------------------| +| [Anders Mörtberg](https://staff.math.su.se/anders.mortberg/) | [mortberg](https://github.com/mortberg) | *Most topics* | +| [Evan Cavallo](https://ecavallo.net/) | [ecavallo](https://github.com/ecavallo) | *Most topics* | +| [Felix Cherubini](https://felix-cherubini.de) | [felixwellen](https://github.com/felixwellen) | *Mainly algebra related topics* | +| [Max Zeuner](https://www.su.se/english/profiles/maze1512-1.450461) | [mzeuner](https://github.com/mzeuner) | *Algebra related topics* | +| [Axel Ljungström](https://aljungstrom.github.io) | [aljungstrom](https://github.com/aljungstrom) | *Synthetic homotopy theory and cohomology* | +| [Andrea Vezzosi](http://saizan.github.io/) | [Saizan](https://github.com/Saizan) | *Inactive* | [Overview](https://github.com/agda/cubical/pulls?q=is%3Apr+is%3Aopen+sort%3Aupdated-asc+draft%3Afalse) of the current open PRs, descending time since last action. - -[![Build Status](https://travis-ci.org/agda/cubical.svg?branch=master)](https://travis-ci.org/agda/cubical) From b046b02b4277f3326daa2be2ba9faa4791f18392 Mon Sep 17 00:00:00 2001 From: mzeuner <56261690+mzeuner@users.noreply.github.com> Date: Mon, 13 May 2024 15:54:00 +0200 Subject: [PATCH 11/30] fully faithful and pres lims (#1123) --- Cubical/Categories/Instances/Sets.agda | 58 +++++++++++++++++++++++++- 1 file changed, 57 insertions(+), 1 deletion(-) diff --git a/Cubical/Categories/Instances/Sets.agda b/Cubical/Categories/Instances/Sets.agda index fd0589da27..440a8b84ff 100644 --- a/Cubical/Categories/Instances/Sets.agda +++ b/Cubical/Categories/Instances/Sets.agda @@ -54,6 +54,18 @@ LiftF .F-hom f x = lift (f (x .lower)) LiftF .F-id = refl LiftF .F-seq f g = funExt λ x → refl +module _ {ℓ ℓ' : Level} where + isFullyFaithfulLiftF : isFullyFaithful (LiftF {ℓ} {ℓ'}) + isFullyFaithfulLiftF X Y = isoToIsEquiv LiftFIso + where + open Iso + LiftFIso : Iso (X .fst → Y .fst) + (Lift {ℓ}{ℓ'} (X .fst) → Lift {ℓ}{ℓ'} (Y .fst)) + fun LiftFIso = LiftF .F-hom {X} {Y} + inv LiftFIso = λ f x → f (lift x) .lower + rightInv LiftFIso = λ _ → funExt λ _ → refl + leftInv LiftFIso = λ _ → funExt λ _ → refl + module _ {C : Category ℓ ℓ'} {F : Functor C (SET ℓ')} where open NatTrans @@ -121,7 +133,6 @@ isUnivalent.univ isUnivalentSET (A , isSet-A) (B , isSet-B) = (λ x i → transp (λ _ → A) i (ci .snd .cInv (transp (λ _ → B) i x)))) -- SET is complete - open LimCone open Cone @@ -135,3 +146,48 @@ univProp (completeSET J D) c cc = (λ _ → funExt (λ _ → refl)) (λ x → isPropIsConeMor cc (limCone (completeSET J D)) x) (λ x hx → funExt (λ d → cone≡ λ u → funExt (λ _ → sym (funExt⁻ (hx u) d)))) + + +-- LiftF : SET ℓ → SET (ℓ-suc ℓ) preserves "small" limits +-- i.e. limits over diagram shapes J : Category ℓ ℓ +module _ {ℓ : Level} where + preservesLimitsLiftF : preservesLimits {ℓJ = ℓ} {ℓJ' = ℓ} (LiftF {ℓ} {ℓ-suc ℓ}) + preservesLimitsLiftF = preservesLimitsChar _ + completeSET + completeSETSuc + limSetIso + λ _ _ _ → refl + where + -- SET (ℓ-suc ℓ) has limits over shapes J : Category ℓ ℓ + completeSETSuc : Limits {ℓJ = ℓ} {ℓJ' = ℓ} (SET (ℓ-suc ℓ)) + lim (completeSETSuc J D) = Cone D (Unit* , isOfHLevelLift 2 isSetUnit) , isSetCone D _ + coneOut (limCone (completeSETSuc J D)) j e = coneOut e j tt* + coneOutCommutes (limCone (completeSETSuc J D)) j i e = coneOutCommutes e j i tt* + univProp (completeSETSuc 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 (completeSETSuc J D)) x) + (λ x hx → funExt (λ d → cone≡ λ u → funExt (λ _ → sym (funExt⁻ (hx u) d)))) + + lowerCone : ∀ J D + → Cone (LiftF ∘F D) (Unit* , isOfHLevelLift 2 isSetUnit) + → Cone D (Unit* , isOfHLevelLift 2 isSetUnit) + coneOut (lowerCone J D cc) v tt* = cc .coneOut v tt* .lower + coneOutCommutes (lowerCone J D cc) e = + funExt λ { tt* → cong lower (funExt⁻ (cc .coneOutCommutes e) tt*) } + + liftCone : ∀ J D + → Cone D (Unit* , isOfHLevelLift 2 isSetUnit) + → Cone (LiftF ∘F D) (Unit* , isOfHLevelLift 2 isSetUnit) + coneOut (liftCone J D cc) v tt* = lift (cc .coneOut v tt*) + coneOutCommutes (liftCone J D cc) e = + funExt λ { tt* → cong lift (funExt⁻ (cc .coneOutCommutes e) tt*) } + + limSetIso : ∀ J D → CatIso (SET (ℓ-suc ℓ)) + (completeSETSuc J (LiftF ∘F D) .lim) + (LiftF .F-ob (completeSET J D .lim)) + fst (limSetIso J D) cc = lift (lowerCone J D cc) + cInv (snd (limSetIso J D)) cc = liftCone J D (cc .lower) + sec (snd (limSetIso J D)) = funExt (λ _ → liftExt (cone≡ λ _ → refl)) + ret (snd (limSetIso J D)) = funExt (λ _ → cone≡ λ _ → refl) From 9ec6a593910c8ddaefa9a43e9cbf580ef1675ae1 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 13 May 2024 16:01:35 +0200 Subject: [PATCH 12/30] Refactor and extend the path-graph (#1116) * refactor GraphPath, use the same argument order as List in the construction * refactor path-graph: make order of arguments more natural, prove coherences --- Cubical/Categories/Constructions/Free.agda | 10 +- Cubical/Data/Graph/Path.agda | 162 ++++++++++++--------- 2 files changed, 96 insertions(+), 76 deletions(-) diff --git a/Cubical/Categories/Constructions/Free.agda b/Cubical/Categories/Constructions/Free.agda index 5bab3d9140..73a5925bba 100644 --- a/Cubical/Categories/Constructions/Free.agda +++ b/Cubical/Categories/Constructions/Free.agda @@ -19,8 +19,8 @@ module _ {ℓv ℓe : Level} where FreeCategory .ob = Node G FreeCategory .Hom[_,_] = Path G FreeCategory .id = pnil - FreeCategory ._⋆_ = ccat G - FreeCategory .⋆IdL = pnil++ G - FreeCategory .⋆IdR P = refl - FreeCategory .⋆Assoc = ++assoc G - FreeCategory .isSetHom = isSetPath G isSetNode isSetEdge _ _ + FreeCategory ._⋆_ = _++_ + FreeCategory .⋆IdL = pnil++ + FreeCategory .⋆IdR P = ++pnil _ + FreeCategory .⋆Assoc = ++assoc + FreeCategory .isSetHom = isSetPath isSetNode isSetEdge _ _ diff --git a/Cubical/Data/Graph/Path.agda b/Cubical/Data/Graph/Path.agda index d8122924c2..ee19a7cc7a 100644 --- a/Cubical/Data/Graph/Path.agda +++ b/Cubical/Data/Graph/Path.agda @@ -3,79 +3,99 @@ module Cubical.Data.Graph.Path where +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude hiding (Path) + open import Cubical.Data.Graph.Base -open import Cubical.Data.List.Base hiding (_++_) +open import Cubical.Data.List.Base hiding (_++_; map) open import Cubical.Data.Nat.Base open import Cubical.Data.Nat.Properties open import Cubical.Data.Sigma.Base hiding (Path) -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Prelude hiding (Path) - -module _ {ℓv ℓe : Level} where - - module _ (G : Graph ℓv ℓe) where - data Path : (v w : Node G) → Type (ℓ-max ℓv ℓe) where - pnil : ∀ {v} → Path v v - pcons : ∀ {v w x} → Path v w → Edge G w x → Path v x - - -- Path concatenation - ccat : ∀ {v w x} → Path v w → Path w x → Path v x - ccat P pnil = P - ccat P (pcons Q e) = pcons (ccat P Q) e - - private - _++_ = ccat - infixr 20 _++_ - - -- Some properties - pnil++ : ∀ {v w} (P : Path v w) → pnil ++ P ≡ P - pnil++ pnil = refl - pnil++ (pcons P e) = cong (λ P → pcons P e) (pnil++ _) - - ++assoc : ∀ {v w x y} - (P : Path v w) (Q : Path w x) (R : Path x y) - → (P ++ Q) ++ R ≡ P ++ (Q ++ R) - ++assoc P Q pnil = refl - ++assoc P Q (pcons R e) = cong (λ P → pcons P e) (++assoc P Q R) - - -- Paths as lists - pathToList : ∀ {v w} → Path v w - → List (Σ[ x ∈ Node G ] Σ[ y ∈ Node G ] Edge G x y) - pathToList pnil = [] - pathToList (pcons P e) = (_ , _ , e) ∷ (pathToList P) - - -- Path v w is a set - -- Lemma 4.2 of https://arxiv.org/abs/2112.06609 - module _ (isSetNode : isSet (Node G)) - (isSetEdge : ∀ v w → isSet (Edge G v w)) where - - -- This is called ̂W (W-hat) in the paper - PathWithLen : ℕ → Node G → Node G → Type (ℓ-max ℓv ℓe) - PathWithLen 0 v w = Lift {j = ℓe} (v ≡ w) - PathWithLen (suc n) v w = Σ[ k ∈ Node G ] (PathWithLen n v k × Edge G k w) - - isSetPathWithLen : ∀ n v w → isSet (PathWithLen n v w) - isSetPathWithLen 0 _ _ = isOfHLevelLift 2 (isProp→isSet (isSetNode _ _)) - isSetPathWithLen (suc n) _ _ = isSetΣ isSetNode λ _ → - isSet× (isSetPathWithLen _ _ _) (isSetEdge _ _) - - isSet-ΣnPathWithLen : ∀ {v w} → isSet (Σ[ n ∈ ℕ ] PathWithLen n v w) - isSet-ΣnPathWithLen = isSetΣ isSetℕ (λ _ → isSetPathWithLen _ _ _) - - Path→PathWithLen : ∀ {v w} → Path v w → Σ[ n ∈ ℕ ] PathWithLen n v w - Path→PathWithLen pnil = 0 , lift refl - Path→PathWithLen (pcons P e) = suc (Path→PathWithLen P .fst) , - _ , Path→PathWithLen P .snd , e - - PathWithLen→Path : ∀ {v w} → Σ[ n ∈ ℕ ] PathWithLen n v w → Path v w - PathWithLen→Path (0 , q) = subst (Path _) (q .lower) pnil - PathWithLen→Path (suc n , _ , pwl , e) = pcons (PathWithLen→Path (n , pwl)) e - - Path→PWL→Path : ∀ {v w} P → PathWithLen→Path {v} {w} (Path→PathWithLen P) ≡ P - Path→PWL→Path {v} pnil = substRefl {B = Path v} pnil - Path→PWL→Path (pcons P x) = cong₂ pcons (Path→PWL→Path _) refl - - isSetPath : ∀ v w → isSet (Path v w) - isSetPath v w = isSetRetract Path→PathWithLen PathWithLen→Path - Path→PWL→Path isSet-ΣnPathWithLen +private variable + ℓv ℓe ℓv' ℓe' : Level + +module _ (G : Graph ℓv ℓe) where + data Path : (v w : Node G) → Type (ℓ-max ℓv ℓe) where + pnil : ∀ {v} → Path v v + pcons : ∀ {v x w} → Edge G v x → Path x w → Path v w + +module _ {G : Graph ℓv ℓe} where + + -- Path concatenation + ccat : ∀ {v x w} → Path G v x → Path G x w → Path G v w + ccat pnil Q = Q + ccat (pcons e P) Q = pcons e (ccat P Q) + + _++_ = ccat + infixr 20 _++_ + + -- Some properties + pnil++ : ∀ {v w} (P : Path G v w) → pnil ++ P ≡ P + pnil++ pnil = refl + pnil++ (pcons e P) = cong (λ P → pcons e P) (pnil++ _) + + ++pnil : ∀ {v w} (P : Path G v w) → P ++ pnil ≡ P + ++pnil pnil = refl + ++pnil (pcons e P) = cong (λ P → pcons e P) (++pnil P) + + ++assoc : ∀ {v w x y} + (P : Path G v w) (Q : Path G w x) (R : Path G x y) + → (P ++ Q) ++ R ≡ P ++ (Q ++ R) + ++assoc pnil P Q = refl + ++assoc (pcons e P) Q R = cong (λ P → pcons e P) (++assoc P Q R) + + -- Paths as lists + pathToList : ∀ {v w} → Path G v w + → List (Σ[ x ∈ Node G ] Σ[ y ∈ Node G ] Edge G x y) + pathToList pnil = [] + pathToList (pcons e P) = (_ , _ , e) ∷ (pathToList P) + + -- Path v w is a set + -- Lemma 4.2 of https://arxiv.org/abs/2112.06609 + module _ (isSetNode : isSet (Node G)) + (isSetEdge : ∀ v w → isSet (Edge G v w)) where + + -- This is called ̂W (W-hat) in the paper + PathWithLen : ℕ → Node G → Node G → Type (ℓ-max ℓv ℓe) + PathWithLen 0 v w = Lift {j = ℓe} (v ≡ w) + PathWithLen (suc n) v w = Σ[ x ∈ Node G ] (Edge G v x × PathWithLen n x w) + + isSetPathWithLen : ∀ n v w → isSet (PathWithLen n v w) + isSetPathWithLen 0 _ _ = isOfHLevelLift 2 (isProp→isSet (isSetNode _ _)) + isSetPathWithLen (suc n) _ _ = isSetΣ isSetNode λ _ → + isSet× (isSetEdge _ _) (isSetPathWithLen _ _ _) + + isSet-ΣnPathWithLen : ∀ {v w} → isSet (Σ[ n ∈ ℕ ] PathWithLen n v w) + isSet-ΣnPathWithLen = isSetΣ isSetℕ (λ _ → isSetPathWithLen _ _ _) + + Path→PathWithLen : ∀ {v w} → Path G v w → Σ[ n ∈ ℕ ] PathWithLen n v w + Path→PathWithLen pnil = 0 , lift refl + Path→PathWithLen (pcons e P) = suc (Path→PathWithLen P .fst) , _ , e , Path→PathWithLen P .snd + + PathWithLen→Path : ∀ {v w} → Σ[ n ∈ ℕ ] PathWithLen n v w → Path G v w + PathWithLen→Path (0 , q) = subst (Path G _) (q .lower) pnil + PathWithLen→Path (suc n , _ , e , pwl) = pcons e (PathWithLen→Path (n , pwl)) + + Path→PWL→Path : ∀ {v w} P → PathWithLen→Path {v} {w} (Path→PathWithLen P) ≡ P + Path→PWL→Path {v} pnil = substRefl {B = Path G v} pnil + Path→PWL→Path (pcons P x) = cong₂ pcons refl (Path→PWL→Path _) + + isSetPath : ∀ v w → isSet (Path G v w) + isSetPath v w = isSetRetract Path→PathWithLen PathWithLen→Path + Path→PWL→Path isSet-ΣnPathWithLen + +module _ {G : Graph ℓv ℓe} {H : Graph ℓv' ℓe'} where + open GraphHom + map : {x y : Node G} + (F : GraphHom G H) + → Path G x y → Path H (F $g x) (F $g y) + map F pnil = pnil + map F (pcons e p) = pcons (F <$g> e) (map F p) + + map++ : {x y z : Node G} + (F : GraphHom G H) + → (p : Path G x y) → (q : Path G y z) + → map F (p ++ q) ≡ map F p ++ map F q + map++ F pnil q = refl + map++ F (pcons x p) q = cong (λ m → pcons (F <$g> x) m) (map++ F p q) From 542449fa853b640eab5e0f183fed2b32cf0038ea Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 13 May 2024 16:06:00 +0200 Subject: [PATCH 13/30] Construct 'the' free wild category on a graph (#1117) * refactor GraphPath, use the same argument order as List in the construction * refactor path-graph: make order of arguments more natural, prove coherences * start with boilerplate * functor notation * free cats * write down a candidate def for free wild cat * tmp * choose better name, since it appears in goals * define free wild cat and partially define its universal property --- Cubical/WildCat/Base.agda | 8 +-- Cubical/WildCat/Free.agda | 74 ++++++++++++++++++++++++++++ Cubical/WildCat/Functor.agda | 16 ++++++ Cubical/WildCat/UnderlyingGraph.agda | 50 +++++++++++++++++++ 4 files changed, 144 insertions(+), 4 deletions(-) create mode 100644 Cubical/WildCat/Free.agda create mode 100644 Cubical/WildCat/UnderlyingGraph.agda diff --git a/Cubical/WildCat/Base.agda b/Cubical/WildCat/Base.agda index fca6167fe7..23df1fcdb9 100644 --- a/Cubical/WildCat/Base.agda +++ b/Cubical/WildCat/Base.agda @@ -38,11 +38,11 @@ _[_,_] : (C : WildCat ℓ ℓ') → (x y : C .ob) → Type ℓ' _[_,_] = Hom[_,_] -- Needed to define this in order to be able to make the subsequence syntax declaration -seq' : ∀ (C : WildCat ℓ ℓ') {x y z} (f : C [ x , y ]) (g : C [ y , z ]) → C [ x , z ] -seq' = _⋆_ +concatMor : ∀ (C : WildCat ℓ ℓ') {x y z} (f : C [ x , y ]) (g : C [ y , z ]) → C [ x , z ] +concatMor = _⋆_ -infixl 15 seq' -syntax seq' C f g = f ⋆⟨ C ⟩ g +infixl 15 concatMor +syntax concatMor C f g = f ⋆⟨ C ⟩ g -- composition comp' : ∀ (C : WildCat ℓ ℓ') {x y z} (g : C [ y , z ]) (f : C [ x , y ]) → C [ x , z ] diff --git a/Cubical/WildCat/Free.agda b/Cubical/WildCat/Free.agda new file mode 100644 index 0000000000..ab97898150 --- /dev/null +++ b/Cubical/WildCat/Free.agda @@ -0,0 +1,74 @@ +{- + This file defines a wild category, which might be the free wild category over a + directed graph (I do not know). This is intended to be used in a solver for + wild categories. +-} +{-# OPTIONS --safe #-} + +module Cubical.WildCat.Free where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Graph.Base +open import Cubical.Data.Graph.Path renaming (Path to GPath) + +open import Cubical.WildCat.Base +open import Cubical.WildCat.Functor +open import Cubical.WildCat.UnderlyingGraph + +open WildCat +open WildFunctor +open Graph + +private + variable + ℓc ℓc' ℓd ℓd' ℓg ℓg' : Level + +Free : (G : Graph ℓg ℓg') → WildCat ℓg (ℓ-max ℓg ℓg') +ob (Free G) = G .Node +Hom[_,_] (Free G) x y = GPath G x y +id (Free G) = pnil +_⋆_ (Free G) f g = f ++ g +⋆IdL (Free G) = λ _ → refl +⋆IdR (Free G) = ++pnil +⋆Assoc (Free G) f g h = ++assoc f g h + +composeAll : (C : WildCat ℓc ℓc') {x y : ob C} → GPath (Cat→Graph C) x y → C [ x , y ] +composeAll C pnil = id C +composeAll C (pcons m path) = m ⋆⟨ C ⟩ composeAll C path + + +composeAll++ : (C : WildCat ℓc ℓc') {x y z : ob C} + → (p : GPath (Cat→Graph C) x y) → (q : GPath (Cat→Graph C) y z) + → composeAll C (p ++ q) ≡ composeAll C p ⋆⟨ C ⟩ composeAll C q +composeAll++ C pnil q = sym (⋆IdL C (composeAll C q)) +composeAll++ C (pcons m p) q = + composeAll C (pcons m p ++ q) ≡⟨ step1 ⟩ + m ⋆⟨ C ⟩ ((composeAll C p) ⋆⟨ C ⟩ (composeAll C q)) ≡⟨ sym (⋆Assoc C m (composeAll C p) (composeAll C q)) ⟩ + (m ⋆⟨ C ⟩ (composeAll C p)) ⋆⟨ C ⟩ (composeAll C q) ∎ + where step1 = cong (λ z → concatMor C m z) (composeAll++ C p q) + +module UniversalProperty (G : Graph ℓg ℓg') where + + + incFree : GraphHom G (Cat→Graph (Free G)) + incFree $g x = x + incFree <$g> e = pcons e pnil + + {- + G ──→ Free G + \ ∣ + ∀ F \ ∣ ∃ F' + ↘ ↓ + C + -} + inducedMorphism : (C : WildCat ℓc ℓc') → GraphHom G (Cat→Graph C) → WildFunctor (Free G) C + inducedMorphism C F = F' + where F' : WildFunctor (Free G) C + F-ob F' x = F $g x + F-hom F' m = composeAll C (map F m) + F-id F' = refl + F-seq F' f g = + composeAll C (map F (f ⋆⟨ Free G ⟩ g)) ≡⟨ cong (λ u → composeAll C u) (map++ F f g) ⟩ + composeAll C (map F f ++ map F g) ≡⟨ composeAll++ C (map F f) (map F g) ⟩ + (composeAll C (map F f)) ⋆⟨ C ⟩ (composeAll C (map F g)) ∎ diff --git a/Cubical/WildCat/Functor.agda b/Cubical/WildCat/Functor.agda index 352a3a0fff..9dd9879aad 100644 --- a/Cubical/WildCat/Functor.agda +++ b/Cubical/WildCat/Functor.agda @@ -162,3 +162,19 @@ WildFunctor.F-ob commFunctor (x , y) = y , x WildFunctor.F-hom commFunctor (f , g) = g , f WildFunctor.F-id commFunctor = refl WildFunctor.F-seq commFunctor _ _ = refl + + +infixl 20 _$_ +_$_ : {ℓC ℓC' ℓD ℓD' : Level} + {C : WildCat ℓC ℓC'} + {D : WildCat ℓD ℓD'} + → WildFunctor C D → C .ob → D .ob +F $ x = F .F-ob x + +infixl 20 _$→_ +_$→_ : {ℓC ℓC' ℓD ℓD' : Level} + {C : WildCat ℓC ℓC'} + {D : WildCat ℓD ℓD'} + {x y : C .ob} + → (F : WildFunctor C D) → C [ x , y ] → D [ F $ x , F $ y ] +F $→ f = F .F-hom f diff --git a/Cubical/WildCat/UnderlyingGraph.agda b/Cubical/WildCat/UnderlyingGraph.agda new file mode 100644 index 0000000000..1fde0d4927 --- /dev/null +++ b/Cubical/WildCat/UnderlyingGraph.agda @@ -0,0 +1,50 @@ +{-# OPTIONS --safe #-} + +module Cubical.WildCat.UnderlyingGraph where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Graph.Base + +open import Cubical.WildCat.Base +open import Cubical.WildCat.Functor + +private + variable + ℓc ℓc' ℓd ℓd' ℓe ℓe' ℓg ℓg' : Level + +open WildCat + +-- Underlying graph of a category +Cat→Graph : ∀ {ℓc ℓc'} (𝓒 : WildCat ℓc ℓc') → Graph ℓc ℓc' +Cat→Graph 𝓒 .Node = 𝓒 .ob +Cat→Graph 𝓒 .Edge = 𝓒 .Hom[_,_] + +Functor→GraphHom : ∀ {ℓc ℓc' ℓd ℓd'} {𝓒 : WildCat ℓc ℓc'} {𝓓 : WildCat ℓd ℓd'} + (F : WildFunctor 𝓒 𝓓) → GraphHom (Cat→Graph 𝓒) (Cat→Graph 𝓓) +Functor→GraphHom F ._$g_ = WildFunctor.F-ob F +Functor→GraphHom F ._<$g>_ = WildFunctor.F-hom F + + +module _ (G : Graph ℓg ℓg') (𝓒 : WildCat ℓc ℓc') where + -- Interpretation of a graph in a wild category + Interpret : Type _ + Interpret = GraphHom G (Cat→Graph 𝓒) + + +_⋆Interpret_ : ∀ {G : Graph ℓg ℓg'} + {𝓒 : WildCat ℓc ℓc'} + {𝓓 : WildCat ℓd ℓd'} + (ı : Interpret G 𝓒) + (F : WildFunctor 𝓒 𝓓) + → Interpret G 𝓓 +(ı ⋆Interpret F) ._$g_ x = WildFunctor.F-ob F (ı $g x) +(ı ⋆Interpret F) ._<$g>_ e = WildFunctor.F-hom F (ı <$g> e) + +_∘Interpret_ : ∀ {G : Graph ℓg ℓg'} + {𝓒 : WildCat ℓc ℓc'} + {𝓓 : WildCat ℓd ℓd'} + (F : WildFunctor 𝓒 𝓓) + (ı : Interpret G 𝓒) + → Interpret G 𝓓 +F ∘Interpret ı = ı ⋆Interpret F From 75494455f83698be84857c50de202ec761c5f8d0 Mon Sep 17 00:00:00 2001 From: Matthias Hutzler Date: Mon, 13 May 2024 16:08:40 +0200 Subject: [PATCH 14/30] =?UTF-8?q?Simplify=20`Embedding-into-hLevel?= =?UTF-8?q?=E2=86=92hLevel`=20(#1107)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * avoid case distinction * derive special cases from general statement --- Cubical/Functions/Embedding.agda | 34 +++++++++++--------------------- 1 file changed, 11 insertions(+), 23 deletions(-) diff --git a/Cubical/Functions/Embedding.agda b/Cubical/Functions/Embedding.agda index bcc5a7b83e..3453a3fe57 100644 --- a/Cubical/Functions/Embedding.agda +++ b/Cubical/Functions/Embedding.agda @@ -204,32 +204,20 @@ 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-hLevel→hLevel + : ∀ n → A ↪ B → isOfHLevel (suc n) B → isOfHLevel (suc n) A +Embedding-into-hLevel→hLevel n (f , isEmbeddingF) isOfHLevelB = + isOfHLevelPath'⁻ n + (λ a a' → + isOfHLevelRespectEquiv n + (invEquiv (cong f , isEmbeddingF a a')) + (isOfHLevelPath' n isOfHLevelB (f a) (f a'))) + 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-isProp→isProp = Embedding-into-hLevel→hLevel 0 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) +Embedding-into-isSet→isSet = Embedding-into-hLevel→hLevel 1 -- We now show that the powerset is the subtype classifier -- i.e. ℙ X ≃ Σ[A ∈ Type ℓ] (A ↪ X) From 59c59c00107993c0a7e88307cbc097b4bb7980c6 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Tue, 14 May 2024 22:47:03 +0200 Subject: [PATCH 15/30] Clean up: Remove Foundation/Everything and outdated stuff (#1127) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Clean Codata folder and make it --safe * move things around * revert * remove foundation/everything * delete outdated README.md * remove Core/Everything, update and clean up imports * fix imports * fix imports * fix imports * fix imports * fix imports --------- Co-authored-by: Anders Mörtberg --- .gitignore | 2 - Cubical/Codata/Conat.agda | 4 +- Cubical/Codata/Conat/Base.agda | 4 +- Cubical/Codata/Conat/Properties.agda | 2 - Cubical/Codata/Everything.agda | 44 ++--- Cubical/Codata/EverythingSafe.agda | 2 - Cubical/Codata/M.agda | 95 ----------- Cubical/Codata/M/AsLimit/Coalg.agda | 5 - Cubical/Codata/M/AsLimit/M.agda | 6 - Cubical/Codata/M/Bisimilarity.agda | 153 ------------------ Cubical/Codata/M/Coalg.agda | 5 + .../Codata/M/{AsLimit => }/Coalg/Base.agda | 6 +- Cubical/Codata/M/{AsLimit => }/Container.agda | 4 +- Cubical/Codata/M/M.agda | 11 ++ Cubical/Codata/M/{AsLimit => }/M/Base.agda | 9 +- .../Codata/M/{AsLimit => }/M/Properties.agda | 8 +- Cubical/Codata/M/{AsLimit => }/helper.agda | 2 +- Cubical/Codata/M/{AsLimit => }/itree.agda | 8 +- Cubical/Codata/M/{AsLimit => }/stream.agda | 8 +- Cubical/Codata/Stream.agda | 2 +- Cubical/Codata/Stream/Base.agda | 2 +- Cubical/Codata/Stream/Properties.agda | 2 - Cubical/Core/Everything.agda | 8 - Cubical/Core/Glue.agda | 5 +- Cubical/Core/README.md | 16 -- Cubical/Data/Bool/Properties.agda | 2 - Cubical/Data/Bool/SwitchStatement.agda | 2 - Cubical/Data/DescendingList/Strict.agda | 2 +- Cubical/Data/Fin/Properties.agda | 10 +- Cubical/Data/InfNat/Base.agda | 3 +- Cubical/Data/Int/Base.agda | 3 +- Cubical/Data/Int/MoreInts/BiInvInt/Base.agda | 10 +- Cubical/Data/Int/MoreInts/QuoInt/Base.agda | 3 +- .../Data/Int/MoreInts/QuoInt/Properties.agda | 3 +- Cubical/Data/Int/Properties.agda | 3 +- Cubical/Data/List/Base.agda | 4 +- Cubical/Data/List/Properties.agda | 9 +- Cubical/Data/Maybe/Base.agda | 2 +- Cubical/Data/Nat/Algebra.agda | 2 - Cubical/Data/Nat/Base.agda | 2 - Cubical/Data/Prod/Base.agda | 2 - Cubical/Data/Prod/Properties.agda | 10 +- Cubical/Data/Sigma/Properties.agda | 9 +- Cubical/Data/Sum/Base.agda | 2 +- Cubical/Data/Sum/Properties.agda | 11 +- Cubical/Data/Unit/Properties.agda | 7 +- Cubical/Foundations/Equiv/Fiberwise.agda | 3 +- Cubical/Foundations/Equiv/Properties.agda | 6 +- Cubical/Foundations/Everything.agda | 34 ---- Cubical/Foundations/Isomorphism.agda | 2 +- Cubical/Foundations/Structure.agda | 2 +- Cubical/Foundations/Univalence/Universe.agda | 2 +- .../HITs/GroupoidQuotients/Properties.agda | 2 - Cubical/HITs/InfNat/Base.agda | 7 +- Cubical/HITs/InfNat/Properties.agda | 7 +- Cubical/HITs/Interval/Base.agda | 2 - Cubical/HITs/KleinBottle/Base.agda | 2 +- .../PropositionalTruncation/Properties.agda | 2 - Cubical/HITs/SetQuotients/Properties.agda | 2 - Cubical/HITs/SetTruncation/Properties.agda | 4 +- Cubical/Homotopy/Loopspace.agda | 16 +- Cubical/Papers/Pi4S3.agda | 7 +- Cubical/Papers/ZCohomology.agda | 14 +- Cubical/Relation/Binary/Base.agda | 2 - .../ZigZag/Applications/MultiSet.agda | 4 +- Cubical/Relation/ZigZag/Base.agda | 4 +- GNUmakefile | 6 +- 67 files changed, 141 insertions(+), 503 deletions(-) delete mode 100644 Cubical/Codata/EverythingSafe.agda delete mode 100644 Cubical/Codata/M.agda delete mode 100644 Cubical/Codata/M/AsLimit/Coalg.agda delete mode 100644 Cubical/Codata/M/AsLimit/M.agda delete mode 100644 Cubical/Codata/M/Bisimilarity.agda create mode 100644 Cubical/Codata/M/Coalg.agda rename Cubical/Codata/M/{AsLimit => }/Coalg/Base.agda (90%) rename Cubical/Codata/M/{AsLimit => }/Container.agda (96%) create mode 100644 Cubical/Codata/M/M.agda rename Cubical/Codata/M/{AsLimit => }/M/Base.agda (98%) rename Cubical/Codata/M/{AsLimit => }/M/Properties.agda (93%) rename Cubical/Codata/M/{AsLimit => }/helper.agda (98%) rename Cubical/Codata/M/{AsLimit => }/itree.agda (93%) rename Cubical/Codata/M/{AsLimit => }/stream.agda (76%) delete mode 100644 Cubical/Core/Everything.agda delete mode 100644 Cubical/Core/README.md delete mode 100644 Cubical/Foundations/Everything.agda diff --git a/.gitignore b/.gitignore index 7d344f2d84..7448c776f5 100644 --- a/.gitignore +++ b/.gitignore @@ -3,6 +3,4 @@ \.*#* \#* Cubical/*/Everything.agda -!Cubical/Core/Everything.agda -!Cubical/Foundations/Everything.agda !Cubical/Codata/Everything.agda diff --git a/Cubical/Codata/Conat.agda b/Cubical/Codata/Conat.agda index f35a316c92..b695c1d03b 100644 --- a/Cubical/Codata/Conat.agda +++ b/Cubical/Codata/Conat.agda @@ -1,7 +1,5 @@ -{-# OPTIONS --guardedness #-} - +{-# OPTIONS --safe --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 index fdb5ec5fa6..ea21bafe9f 100644 --- a/Cubical/Codata/Conat/Base.agda +++ b/Cubical/Codata/Conat/Base.agda @@ -21,11 +21,11 @@ of Sized Types. {-# OPTIONS --safe --guardedness #-} module Cubical.Codata.Conat.Base where +open import Cubical.Foundations.Prelude + open import Cubical.Data.Unit open import Cubical.Data.Sum -open import Cubical.Core.Everything - record Conat : Type₀ Conat′ = Unit ⊎ Conat record Conat where diff --git a/Cubical/Codata/Conat/Properties.agda b/Cubical/Codata/Conat/Properties.agda index d9cc9c8d7a..b620efb598 100644 --- a/Cubical/Codata/Conat/Properties.agda +++ b/Cubical/Codata/Conat/Properties.agda @@ -33,8 +33,6 @@ open import Cubical.Data.Bool import Cubical.Data.Nat as Nat import Cubical.Data.Nat.Order.Recursive as Nat -open import Cubical.Core.Everything - open import Cubical.Functions.Embedding open import Cubical.Foundations.Prelude diff --git a/Cubical/Codata/Everything.agda b/Cubical/Codata/Everything.agda index 52c952d486..163e335490 100644 --- a/Cubical/Codata/Everything.agda +++ b/Cubical/Codata/Everything.agda @@ -1,33 +1,15 @@ -{-# OPTIONS --guardedness #-} - +{-# OPTIONS --safe --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.Conat.Bounded - -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 +import Cubical.Codata.Conat +import Cubical.Codata.Conat.Bounded +import Cubical.Codata.M.Coalg +import Cubical.Codata.M.Coalg.Base +import Cubical.Codata.M.Container +import Cubical.Codata.M.M +import Cubical.Codata.M.M.Base +import Cubical.Codata.M.M.Properties +import Cubical.Codata.M.helper +import Cubical.Codata.M.itree +import Cubical.Codata.M.stream +import Cubical.Codata.Stream diff --git a/Cubical/Codata/EverythingSafe.agda b/Cubical/Codata/EverythingSafe.agda deleted file mode 100644 index 0f7a364688..0000000000 --- a/Cubical/Codata/EverythingSafe.agda +++ /dev/null @@ -1,2 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Codata.EverythingSafe where diff --git a/Cubical/Codata/M.agda b/Cubical/Codata/M.agda deleted file mode 100644 index 4399a07c62..0000000000 --- a/Cubical/Codata/M.agda +++ /dev/null @@ -1,95 +0,0 @@ -{-# 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 deleted file mode 100644 index 4118df13c7..0000000000 --- a/Cubical/Codata/M/AsLimit/Coalg.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# 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/M.agda b/Cubical/Codata/M/AsLimit/M.agda deleted file mode 100644 index 22c3b9e654..0000000000 --- a/Cubical/Codata/M/AsLimit/M.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# 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/Bisimilarity.agda b/Cubical/Codata/M/Bisimilarity.agda deleted file mode 100644 index 602d2f8989..0000000000 --- a/Cubical/Codata/M/Bisimilarity.agda +++ /dev/null @@ -1,153 +0,0 @@ -{-# OPTIONS --postfix-projections --guardedness #-} -module Cubical.Codata.M.Bisimilarity where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.Fiberwise -open import Cubical.Foundations.Transport -open import Cubical.Foundations.HLevels - - -open import Cubical.Codata.M - - -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/M/Coalg.agda b/Cubical/Codata/M/Coalg.agda new file mode 100644 index 0000000000..184589d90a --- /dev/null +++ b/Cubical/Codata/M/Coalg.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.Coalg where + +open import Cubical.Codata.M.Coalg.Base public diff --git a/Cubical/Codata/M/AsLimit/Coalg/Base.agda b/Cubical/Codata/M/Coalg/Base.agda similarity index 90% rename from Cubical/Codata/M/AsLimit/Coalg/Base.agda rename to Cubical/Codata/M/Coalg/Base.agda index d4c7105c23..324d2d0f0d 100644 --- a/Cubical/Codata/M/AsLimit/Coalg/Base.agda +++ b/Cubical/Codata/M/Coalg/Base.agda @@ -1,6 +1,6 @@ {-# OPTIONS --guardedness --safe #-} -module Cubical.Codata.M.AsLimit.Coalg.Base where +module Cubical.Codata.M.Coalg.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function using ( _∘_ ) @@ -15,8 +15,8 @@ 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 +open import Cubical.Codata.M.Container +open import Cubical.Codata.M.helper ------------------------------- -- Definition of a Coalgebra -- diff --git a/Cubical/Codata/M/AsLimit/Container.agda b/Cubical/Codata/M/Container.agda similarity index 96% rename from Cubical/Codata/M/AsLimit/Container.agda rename to Cubical/Codata/M/Container.agda index f0c9cb6e8f..c96ec2c81d 100644 --- a/Cubical/Codata/M/AsLimit/Container.agda +++ b/Cubical/Codata/M/Container.agda @@ -1,6 +1,6 @@ {-# OPTIONS --guardedness --safe #-} -module Cubical.Codata.M.AsLimit.Container where +module Cubical.Codata.M.Container where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv using (_≃_) @@ -20,7 +20,7 @@ open import Cubical.Data.Sum open import Cubical.Foundations.Structure -open import Cubical.Codata.M.AsLimit.helper +open import Cubical.Codata.M.helper ------------------------------------- -- Container and Container Functor -- diff --git a/Cubical/Codata/M/M.agda b/Cubical/Codata/M/M.agda new file mode 100644 index 0000000000..333260c7c7 --- /dev/null +++ b/Cubical/Codata/M/M.agda @@ -0,0 +1,11 @@ +{- +-- 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 +-} +{-# OPTIONS --guardedness --safe #-} +module Cubical.Codata.M.M where + +open import Cubical.Codata.M.M.Base public +open import Cubical.Codata.M.M.Properties public diff --git a/Cubical/Codata/M/AsLimit/M/Base.agda b/Cubical/Codata/M/M/Base.agda similarity index 98% rename from Cubical/Codata/M/AsLimit/M/Base.agda rename to Cubical/Codata/M/M/Base.agda index a07b2d8116..a1dce924ff 100644 --- a/Cubical/Codata/M/AsLimit/M/Base.agda +++ b/Cubical/Codata/M/M/Base.agda @@ -5,7 +5,7 @@ -- "Non-wellfounded trees in Homotopy Type Theory" -- Benedikt Ahrens, Paolo Capriotti, Régis Spadotti -module Cubical.Codata.M.AsLimit.M.Base where +module Cubical.Codata.M.M.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv using (_≃_) @@ -27,10 +27,9 @@ 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 import Cubical.Codata.M.helper +open import Cubical.Codata.M.Container +open import Cubical.Codata.M.Coalg.Base open Iso diff --git a/Cubical/Codata/M/AsLimit/M/Properties.agda b/Cubical/Codata/M/M/Properties.agda similarity index 93% rename from Cubical/Codata/M/AsLimit/M/Properties.agda rename to Cubical/Codata/M/M/Properties.agda index adabd33a1e..1f5a8112eb 100644 --- a/Cubical/Codata/M/AsLimit/M/Properties.agda +++ b/Cubical/Codata/M/M/Properties.agda @@ -1,6 +1,6 @@ {-# OPTIONS --guardedness --safe #-} -module Cubical.Codata.M.AsLimit.M.Properties where +module Cubical.Codata.M.M.Properties where open import Cubical.Data.Unit open import Cubical.Data.Prod @@ -19,9 +19,9 @@ 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 +open import Cubical.Codata.M.helper +open import Cubical.Codata.M.M.Base +open import Cubical.Codata.M.Container -- in-fun and out-fun are inverse diff --git a/Cubical/Codata/M/AsLimit/helper.agda b/Cubical/Codata/M/helper.agda similarity index 98% rename from Cubical/Codata/M/AsLimit/helper.agda rename to Cubical/Codata/M/helper.agda index 33be276fd7..fb649edd5b 100644 --- a/Cubical/Codata/M/AsLimit/helper.agda +++ b/Cubical/Codata/M/helper.agda @@ -1,6 +1,6 @@ {-# OPTIONS --guardedness --safe #-} -module Cubical.Codata.M.AsLimit.helper where +module Cubical.Codata.M.helper where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv using (_≃_) diff --git a/Cubical/Codata/M/AsLimit/itree.agda b/Cubical/Codata/M/itree.agda similarity index 93% rename from Cubical/Codata/M/AsLimit/itree.agda rename to Cubical/Codata/M/itree.agda index 50d8e10757..6edb0d4ebb 100644 --- a/Cubical/Codata/M/AsLimit/itree.agda +++ b/Cubical/Codata/M/itree.agda @@ -1,6 +1,6 @@ {-# OPTIONS --guardedness --safe #-} -module Cubical.Codata.M.AsLimit.itree where +module Cubical.Codata.M.itree where open import Cubical.Data.Unit open import Cubical.Data.Prod @@ -12,9 +12,9 @@ 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 +open import Cubical.Codata.M.Container +open import Cubical.Codata.M.M +open import Cubical.Codata.M.Coalg -- Delay monad defined as an M-type delay-S : (R : Type₀) -> Container ℓ-zero diff --git a/Cubical/Codata/M/AsLimit/stream.agda b/Cubical/Codata/M/stream.agda similarity index 76% rename from Cubical/Codata/M/AsLimit/stream.agda rename to Cubical/Codata/M/stream.agda index f6fc3c19eb..4e952cac3f 100644 --- a/Cubical/Codata/M/AsLimit/stream.agda +++ b/Cubical/Codata/M/stream.agda @@ -1,14 +1,14 @@ {-# OPTIONS --guardedness --safe #-} -module Cubical.Codata.M.AsLimit.stream where +module Cubical.Codata.M.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 +open import Cubical.Codata.M.M +open import Cubical.Codata.M.helper +open import Cubical.Codata.M.Container -------------------------------------- -- Stream definitions using M.AsLimit -- diff --git a/Cubical/Codata/Stream.agda b/Cubical/Codata/Stream.agda index 45011fd312..07ae80fd54 100644 --- a/Cubical/Codata/Stream.agda +++ b/Cubical/Codata/Stream.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --guardedness #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Codata.Stream where diff --git a/Cubical/Codata/Stream/Base.agda b/Cubical/Codata/Stream/Base.agda index d4a62600a3..dedb0284ec 100644 --- a/Cubical/Codata/Stream/Base.agda +++ b/Cubical/Codata/Stream/Base.agda @@ -1,7 +1,7 @@ {-# OPTIONS --safe --guardedness #-} module Cubical.Codata.Stream.Base where -open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude record Stream (A : Type₀) : Type₀ where coinductive diff --git a/Cubical/Codata/Stream/Properties.agda b/Cubical/Codata/Stream/Properties.agda index 85927262c1..22db042543 100644 --- a/Cubical/Codata/Stream/Properties.agda +++ b/Cubical/Codata/Stream/Properties.agda @@ -1,8 +1,6 @@ {-# OPTIONS --safe --guardedness #-} module Cubical.Codata.Stream.Properties where -open import Cubical.Core.Everything - open import Cubical.Data.Nat open import Cubical.Foundations.Prelude diff --git a/Cubical/Core/Everything.agda b/Cubical/Core/Everything.agda deleted file mode 100644 index c82602edc9..0000000000 --- a/Cubical/Core/Everything.agda +++ /dev/null @@ -1,8 +0,0 @@ -{-# 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 diff --git a/Cubical/Core/Glue.agda b/Cubical/Core/Glue.agda index 89f9dd49ed..68d2b6ae0f 100644 --- a/Cubical/Core/Glue.agda +++ b/Cubical/Core/Glue.agda @@ -125,9 +125,8 @@ private e1 : T1 ≃ A1 e1 = e i1 1=1 - open import Cubical.Foundations.Prelude using (transport) transportA : A0 → A1 - transportA = transport (λ i → A i) + transportA = transp (λ i → A i) i0 -- 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 @@ -136,4 +135,4 @@ private -- 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) + transp-S t0 = inS (transp (λ i → Glue (A i) (Te i)) i0 t0) diff --git a/Cubical/Core/README.md b/Cubical/Core/README.md deleted file mode 100644 index da6b3180d1..0000000000 --- a/Cubical/Core/README.md +++ /dev/null @@ -1,16 +0,0 @@ -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/Bool/Properties.agda b/Cubical/Data/Bool/Properties.agda index 5570fabbd6..be192d0af6 100644 --- a/Cubical/Data/Bool/Properties.agda +++ b/Cubical/Data/Bool/Properties.agda @@ -1,8 +1,6 @@ {-# OPTIONS --safe #-} module Cubical.Data.Bool.Properties where -open import Cubical.Core.Everything - open import Cubical.Functions.Involution open import Cubical.Foundations.Prelude diff --git a/Cubical/Data/Bool/SwitchStatement.agda b/Cubical/Data/Bool/SwitchStatement.agda index 841dbdbda6..9911eca09a 100644 --- a/Cubical/Data/Bool/SwitchStatement.agda +++ b/Cubical/Data/Bool/SwitchStatement.agda @@ -1,8 +1,6 @@ {-# OPTIONS --safe #-} module Cubical.Data.Bool.SwitchStatement where -open import Cubical.Core.Everything - open import Cubical.Foundations.Prelude open import Cubical.Data.Bool.Base diff --git a/Cubical/Data/DescendingList/Strict.agda b/Cubical/Data/DescendingList/Strict.agda index 4141c6d76f..9429c2880f 100644 --- a/Cubical/Data/DescendingList/Strict.agda +++ b/Cubical/Data/DescendingList/Strict.agda @@ -4,7 +4,7 @@ {-# OPTIONS --safe #-} -open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude module Cubical.Data.DescendingList.Strict (A : Type₀) diff --git a/Cubical/Data/Fin/Properties.agda b/Cubical/Data/Fin/Properties.agda index c14c1e236f..dda477bf43 100644 --- a/Cubical/Data/Fin/Properties.agda +++ b/Cubical/Data/Fin/Properties.agda @@ -1,19 +1,17 @@ {-# OPTIONS --safe #-} - module Cubical.Data.Fin.Properties where -open import Cubical.Core.Everything - -open import Cubical.Functions.Embedding -open import Cubical.Functions.Surjection +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.Prelude open import Cubical.Foundations.Univalence open import Cubical.Foundations.Transport +open import Cubical.Functions.Embedding +open import Cubical.Functions.Surjection + open import Cubical.HITs.PropositionalTruncation renaming (rec to ∥∥rec) open import Cubical.Data.Fin.Base as Fin diff --git a/Cubical/Data/InfNat/Base.agda b/Cubical/Data/InfNat/Base.agda index d368954645..020c95c0d7 100644 --- a/Cubical/Data/InfNat/Base.agda +++ b/Cubical/Data/InfNat/Base.agda @@ -2,8 +2,9 @@ module Cubical.Data.InfNat.Base where +open import Cubical.Foundations.Prelude + open import Cubical.Data.Nat as ℕ using (ℕ) -open import Cubical.Core.Primitives data ℕ+∞ : Type₀ where ∞ : ℕ+∞ diff --git a/Cubical/Data/Int/Base.agda b/Cubical/Data/Int/Base.agda index ba7500dfe5..0576b0f241 100644 --- a/Cubical/Data/Int/Base.agda +++ b/Cubical/Data/Int/Base.agda @@ -3,7 +3,8 @@ {-# OPTIONS --safe #-} module Cubical.Data.Int.Base where -open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude + open import Cubical.Data.Bool open import Cubical.Data.Nat hiding (_+_ ; _·_) renaming (isEven to isEvenℕ ; isOdd to isOddℕ) open import Cubical.Data.Fin.Inductive.Base diff --git a/Cubical/Data/Int/MoreInts/BiInvInt/Base.agda b/Cubical/Data/Int/MoreInts/BiInvInt/Base.agda index 83af6718b4..040b45754f 100644 --- a/Cubical/Data/Int/MoreInts/BiInvInt/Base.agda +++ b/Cubical/Data/Int/MoreInts/BiInvInt/Base.agda @@ -16,21 +16,18 @@ This file contains: {-# 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.Data.Nat +open import Cubical.Data.Int + open import Cubical.Relation.Nullary @@ -307,4 +304,3 @@ private predl'-suc : ∀ z → predl' (suc z) ≡ z predl'-suc z = refl - diff --git a/Cubical/Data/Int/MoreInts/QuoInt/Base.agda b/Cubical/Data/Int/MoreInts/QuoInt/Base.agda index 96decf2200..92a3802f7a 100644 --- a/Cubical/Data/Int/MoreInts/QuoInt/Base.agda +++ b/Cubical/Data/Int/MoreInts/QuoInt/Base.agda @@ -2,13 +2,12 @@ {-# 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 () diff --git a/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda b/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda index 8669ba9a0c..6090457ab8 100644 --- a/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda +++ b/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda @@ -1,13 +1,14 @@ {-# 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.Foundations.Isomorphism + open import Cubical.Functions.FunExtEquiv + open import Cubical.Relation.Nullary open import Cubical.Data.Nat as ℕ using (ℕ; zero; suc) diff --git a/Cubical/Data/Int/Properties.agda b/Cubical/Data/Int/Properties.agda index 9a3c1d36fd..1fc854cafa 100644 --- a/Cubical/Data/Int/Properties.agda +++ b/Cubical/Data/Int/Properties.agda @@ -1,12 +1,11 @@ {-# OPTIONS --safe #-} module Cubical.Data.Int.Properties 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.Univalence +open import Cubical.Foundations.Equiv open import Cubical.Relation.Nullary diff --git a/Cubical/Data/List/Base.agda b/Cubical/Data/List/Base.agda index a78ac00455..2695f1de8c 100644 --- a/Cubical/Data/List/Base.agda +++ b/Cubical/Data/List/Base.agda @@ -2,7 +2,9 @@ module Cubical.Data.List.Base where open import Agda.Builtin.List public -open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude + open import Cubical.Data.Maybe.Base as Maybe hiding (rec ; elim) open import Cubical.Data.Nat.Base hiding (elim) diff --git a/Cubical/Data/List/Properties.agda b/Cubical/Data/List/Properties.agda index 7b755944f5..5dc57b71b6 100644 --- a/Cubical/Data/List/Properties.agda +++ b/Cubical/Data/List/Properties.agda @@ -1,22 +1,21 @@ {-# OPTIONS --safe #-} module Cubical.Data.List.Properties where -open import Agda.Builtin.List -open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism + open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat open import Cubical.Data.Sigma open import Cubical.Data.Sum as ⊎ hiding (map) open import Cubical.Data.Unit -open import Cubical.Relation.Nullary - open import Cubical.Data.List.Base as List +open import Cubical.Relation.Nullary + module _ {ℓ} {A : Type ℓ} where ++-unit-r : (xs : List A) → xs ++ [] ≡ xs diff --git a/Cubical/Data/Maybe/Base.agda b/Cubical/Data/Maybe/Base.agda index 1712b822ca..713b9e8d10 100644 --- a/Cubical/Data/Maybe/Base.agda +++ b/Cubical/Data/Maybe/Base.agda @@ -1,7 +1,7 @@ {-# OPTIONS --safe #-} module Cubical.Data.Maybe.Base where -open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude private variable diff --git a/Cubical/Data/Nat/Algebra.agda b/Cubical/Data/Nat/Algebra.agda index 0ee04806b4..d6d9761b89 100644 --- a/Cubical/Data/Nat/Algebra.agda +++ b/Cubical/Data/Nat/Algebra.agda @@ -15,8 +15,6 @@ 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 diff --git a/Cubical/Data/Nat/Base.agda b/Cubical/Data/Nat/Base.agda index 9f424f947b..87e9c4aa86 100644 --- a/Cubical/Data/Nat/Base.agda +++ b/Cubical/Data/Nat/Base.agda @@ -1,8 +1,6 @@ {-# 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 _·_) diff --git a/Cubical/Data/Prod/Base.agda b/Cubical/Data/Prod/Base.agda index 89c8b92bba..b3f3c8795c 100644 --- a/Cubical/Data/Prod/Base.agda +++ b/Cubical/Data/Prod/Base.agda @@ -1,8 +1,6 @@ {-# OPTIONS --safe #-} module Cubical.Data.Prod.Base where -open import Cubical.Core.Everything - open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function diff --git a/Cubical/Data/Prod/Properties.agda b/Cubical/Data/Prod/Properties.agda index 393352c444..58e0f01b6b 100644 --- a/Cubical/Data/Prod/Properties.agda +++ b/Cubical/Data/Prod/Properties.agda @@ -1,11 +1,6 @@ {-# OPTIONS --safe #-} module Cubical.Data.Prod.Properties where -open import Cubical.Core.Everything - -open import Cubical.Data.Prod.Base -open import Cubical.Data.Sigma renaming (_×_ to _×Σ_) hiding (prodIso ; toProdIso ; curryIso) - open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function @@ -13,6 +8,9 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence +open import Cubical.Data.Prod.Base +open import Cubical.Data.Sigma renaming (_×_ to _×Σ_) hiding (prodIso ; toProdIso ; curryIso) + private variable ℓ ℓ' : Level @@ -22,7 +20,7 @@ private -- Swapping is an equivalence ×≡ : {a b : A × B} → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b -×≡ {a = (a1 , b1)} {b = (a2 , b2)} id1 id2 i = (id1 i) , (id2 i) +×≡ {a = (_ , _)} {b = (_ , _)} id1 id2 i = (id1 i) , (id2 i) swap : A × B → B × A swap (x , y) = (y , x) diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index a3db8125dc..fea7559aad 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -15,10 +15,6 @@ Basic properties about Σ-types {-# OPTIONS --safe #-} module Cubical.Data.Sigma.Properties where -open import Cubical.Data.Sigma.Base - -open import Cubical.Core.Everything - open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism @@ -28,10 +24,13 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Path open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence -open import Cubical.Relation.Nullary + +open import Cubical.Data.Sigma.Base open import Cubical.Data.Unit.Base open import Cubical.Data.Empty.Base +open import Cubical.Relation.Nullary + open import Cubical.Reflection.StrictEquiv open Iso diff --git a/Cubical/Data/Sum/Base.agda b/Cubical/Data/Sum/Base.agda index 769fb811a0..b4ddd0f8b5 100644 --- a/Cubical/Data/Sum/Base.agda +++ b/Cubical/Data/Sum/Base.agda @@ -3,7 +3,7 @@ module Cubical.Data.Sum.Base where open import Cubical.Relation.Nullary.Base -open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude private variable diff --git a/Cubical/Data/Sum/Properties.agda b/Cubical/Data/Sum/Properties.agda index 2be7be3c1f..0cc5687221 100644 --- a/Cubical/Data/Sum/Properties.agda +++ b/Cubical/Data/Sum/Properties.agda @@ -1,19 +1,20 @@ {-# OPTIONS --safe #-} module Cubical.Data.Sum.Properties where -open import Cubical.Core.Everything -open import Cubical.Foundations.Function open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels -open import Cubical.Functions.Embedding open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism + +open import Cubical.Functions.Embedding + +open import Cubical.Data.Sum.Base as ⊎ open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat open import Cubical.Data.Sigma -open import Cubical.Relation.Nullary -open import Cubical.Data.Sum.Base as ⊎ +open import Cubical.Relation.Nullary open Iso diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index 85b37dc14b..d219f49696 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -1,23 +1,18 @@ {-# OPTIONS --safe #-} module Cubical.Data.Unit.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.Isomorphism open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence 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 -open import Cubical.Foundations.Univalence - open import Cubical.Reflection.StrictEquiv open Iso diff --git a/Cubical/Foundations/Equiv/Fiberwise.agda b/Cubical/Foundations/Equiv/Fiberwise.agda index 9bcbf5021e..9b2e69a1b1 100644 --- a/Cubical/Foundations/Equiv/Fiberwise.agda +++ b/Cubical/Foundations/Equiv/Fiberwise.agda @@ -1,13 +1,12 @@ {-# 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 diff --git a/Cubical/Foundations/Equiv/Properties.agda b/Cubical/Foundations/Equiv/Properties.agda index b629b6ecbb..0cf3e1eba3 100644 --- a/Cubical/Foundations/Equiv/Properties.agda +++ b/Cubical/Foundations/Equiv/Properties.agda @@ -12,10 +12,6 @@ A couple of general facts about equivalences: {-# 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 @@ -27,6 +23,8 @@ open import Cubical.Foundations.HLevels open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma + private variable ℓ ℓ' ℓ'' : Level diff --git a/Cubical/Foundations/Everything.agda b/Cubical/Foundations/Everything.agda deleted file mode 100644 index def615db0f..0000000000 --- a/Cubical/Foundations/Everything.agda +++ /dev/null @@ -1,34 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Foundations.Everything where - --- Basic cubical prelude -open import Cubical.Foundations.Prelude public - -open import Cubical.Foundations.Function public -open import Cubical.Foundations.Equiv public - hiding (transpEquiv) -- Hide to avoid clash with Transport.transpEquiv -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.Equiv.Dependent -open import Cubical.Foundations.HLevels public -open import Cubical.Foundations.HLevels.Extend -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.Univalence.Dependent -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 -open import Cubical.Foundations.Cubes -open import Cubical.Foundations.Cubes.Subtypes -open import Cubical.Foundations.Cubes.Dependent -open import Cubical.Foundations.Cubes.HLevels diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index a1a128eb11..237a24a251 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -10,7 +10,7 @@ Theory about isomorphisms {-# OPTIONS --safe #-} module Cubical.Foundations.Isomorphism where -open import Cubical.Core.Everything +open import Cubical.Core.Glue open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws diff --git a/Cubical/Foundations/Structure.agda b/Cubical/Foundations/Structure.agda index 099bb8a2eb..eba47e0b78 100644 --- a/Cubical/Foundations/Structure.agda +++ b/Cubical/Foundations/Structure.agda @@ -1,8 +1,8 @@ {-# OPTIONS --safe #-} module Cubical.Foundations.Structure where -open import Cubical.Core.Everything open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv private variable diff --git a/Cubical/Foundations/Univalence/Universe.agda b/Cubical/Foundations/Univalence/Universe.agda index 0857ad7c6d..bf038ddee5 100644 --- a/Cubical/Foundations/Univalence/Universe.agda +++ b/Cubical/Foundations/Univalence/Universe.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe --postfix-projections #-} -open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude open import Cubical.Functions.Embedding open import Cubical.Foundations.Equiv diff --git a/Cubical/HITs/GroupoidQuotients/Properties.agda b/Cubical/HITs/GroupoidQuotients/Properties.agda index 6ca1bf5db7..67882c701c 100644 --- a/Cubical/HITs/GroupoidQuotients/Properties.agda +++ b/Cubical/HITs/GroupoidQuotients/Properties.agda @@ -9,8 +9,6 @@ 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 diff --git a/Cubical/HITs/InfNat/Base.agda b/Cubical/HITs/InfNat/Base.agda index 992ab56fcf..ae590cb3a5 100644 --- a/Cubical/HITs/InfNat/Base.agda +++ b/Cubical/HITs/InfNat/Base.agda @@ -2,13 +2,12 @@ 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 +open import Cubical.Data.Maybe +open import Cubical.Data.Nat + data ℕ+∞ : Type₀ where zero : ℕ+∞ suc : ℕ+∞ → ℕ+∞ diff --git a/Cubical/HITs/InfNat/Properties.agda b/Cubical/HITs/InfNat/Properties.agda index f974ba7a9e..f5fd33b412 100644 --- a/Cubical/HITs/InfNat/Properties.agda +++ b/Cubical/HITs/InfNat/Properties.agda @@ -2,13 +2,12 @@ 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.Data.Maybe +open import Cubical.Data.Nat + open import Cubical.HITs.InfNat.Base import Cubical.Data.InfNat as Coprod diff --git a/Cubical/HITs/Interval/Base.agda b/Cubical/HITs/Interval/Base.agda index 54865c63b3..fc52a4bae2 100644 --- a/Cubical/HITs/Interval/Base.agda +++ b/Cubical/HITs/Interval/Base.agda @@ -1,8 +1,6 @@ {-# OPTIONS --safe #-} module Cubical.HITs.Interval.Base where -open import Cubical.Core.Everything - open import Cubical.Foundations.Prelude data Interval : Type₀ where diff --git a/Cubical/HITs/KleinBottle/Base.agda b/Cubical/HITs/KleinBottle/Base.agda index 3e9645c0b2..2119e43500 100644 --- a/Cubical/HITs/KleinBottle/Base.agda +++ b/Cubical/HITs/KleinBottle/Base.agda @@ -6,7 +6,7 @@ Definition of the Klein bottle as a HIT {-# OPTIONS --safe #-} module Cubical.HITs.KleinBottle.Base where -open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude data KleinBottle : Type where point : KleinBottle diff --git a/Cubical/HITs/PropositionalTruncation/Properties.agda b/Cubical/HITs/PropositionalTruncation/Properties.agda index e0b91cf50f..7f143a450c 100644 --- a/Cubical/HITs/PropositionalTruncation/Properties.agda +++ b/Cubical/HITs/PropositionalTruncation/Properties.agda @@ -8,8 +8,6 @@ This file contains: {-# 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 diff --git a/Cubical/HITs/SetQuotients/Properties.agda b/Cubical/HITs/SetQuotients/Properties.agda index 02ca17a0f3..f2327a424d 100644 --- a/Cubical/HITs/SetQuotients/Properties.agda +++ b/Cubical/HITs/SetQuotients/Properties.agda @@ -9,8 +9,6 @@ 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 diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 33d3deefa0..bc42ebb1d1 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -8,8 +8,6 @@ This file contains: {-# 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 @@ -18,9 +16,11 @@ 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) +open import Cubical.HITs.SetTruncation.Base private variable diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index b5b2d5aecd..36fef7ea1f 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -2,26 +2,26 @@ 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 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.Nat open import Cubical.Data.Sigma + +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Truncation hiding (elim2) renaming (rec to trRec) + open Iso {- loop space of a pointed type -} diff --git a/Cubical/Papers/Pi4S3.agda b/Cubical/Papers/Pi4S3.agda index a1c4ecc21d..6bcc7e8642 100644 --- a/Cubical/Papers/Pi4S3.agda +++ b/Cubical/Papers/Pi4S3.agda @@ -16,10 +16,15 @@ Formalizing π₄(S³) ≅ ℤ/2ℤ and Computing a Brunerie Number in Cubical A module Cubical.Papers.Pi4S3 where -- Misc. +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels + open import Cubical.Data.Int hiding (_+_) open import Cubical.Data.Nat open import Cubical.Data.Nat.Order -open import Cubical.Foundations.Everything open import Cubical.Data.Sum open import Cubical.Data.Sigma diff --git a/Cubical/Papers/ZCohomology.agda b/Cubical/Papers/ZCohomology.agda index 97d9a9dfd9..347427d55e 100644 --- a/Cubical/Papers/ZCohomology.agda +++ b/Cubical/Papers/ZCohomology.agda @@ -20,18 +20,17 @@ 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 +open import Cubical.Foundations.Prelude as Prelude +open import Cubical.Foundations.GroupoidLaws as GroupoidLaws +open import Cubical.Foundations.Isomorphism import Cubical.Foundations.Path as Path -import Cubical.Foundations.Pointed as Pointed - renaming (Pointed to Type∙) +open import Cubical.Foundations.Pointed open import Cubical.HITs.S1 as S1 open import Cubical.HITs.Susp as Suspension open import Cubical.HITs.Sn as Sn @@ -108,9 +107,6 @@ open Prelude using ( transport --- 2.2 Important concepts from HoTT/UF in Cubical Agda --- Pointed Types -open Pointed using (Type∙) - -- The circle, 𝕊¹ open S1 using (S¹) @@ -356,7 +352,7 @@ open Cup using (_⌣_) -- 4.2 -- Lemma 14 -Lem14 : ∀ {ℓ} {A : Type∙ ℓ} (n : ℕ) (f g : A →∙ K∙ n) → fst f ≡ fst g → f ≡ g +Lem14 : ∀ {ℓ} {A : Pointed ℓ} (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 diff --git a/Cubical/Relation/Binary/Base.agda b/Cubical/Relation/Binary/Base.agda index d583115a76..07d31701af 100644 --- a/Cubical/Relation/Binary/Base.agda +++ b/Cubical/Relation/Binary/Base.agda @@ -1,8 +1,6 @@ {-# 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 diff --git a/Cubical/Relation/ZigZag/Applications/MultiSet.agda b/Cubical/Relation/ZigZag/Applications/MultiSet.agda index 5f5b2bfd48..270472e60f 100644 --- a/Cubical/Relation/ZigZag/Applications/MultiSet.agda +++ b/Cubical/Relation/ZigZag/Applications/MultiSet.agda @@ -2,7 +2,6 @@ {-# 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 @@ -10,7 +9,9 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.RelationalStructure open import Cubical.Foundations.Structure open import Cubical.Foundations.SIP +open import Cubical.Foundations.Equiv open import Cubical.Foundations.Univalence + open import Cubical.Data.Unit open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat @@ -20,6 +21,7 @@ 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 diff --git a/Cubical/Relation/ZigZag/Base.agda b/Cubical/Relation/ZigZag/Base.agda index a13ff42989..257be78f87 100644 --- a/Cubical/Relation/ZigZag/Base.agda +++ b/Cubical/Relation/ZigZag/Base.agda @@ -3,12 +3,14 @@ {-# 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.Foundations.Equiv + open import Cubical.Data.Sigma + open import Cubical.HITs.SetQuotients open import Cubical.HITs.PropositionalTruncation as Trunc open import Cubical.Relation.Binary.Base diff --git a/GNUmakefile b/GNUmakefile index 92830cce3f..f9d8a6e9b5 100644 --- a/GNUmakefile +++ b/GNUmakefile @@ -35,12 +35,12 @@ check-everythings: .PHONY : gen-everythings gen-everythings: - $(EVERYTHINGS) gen-except Core Foundations Codata + $(EVERYTHINGS) gen-except Codata .PHONY : gen-and-check-everythings gen-and-check-everythings: - $(EVERYTHINGS) gen-except Core Foundations Codata - $(EVERYTHINGS) check Core Foundations Codata + $(EVERYTHINGS) gen-except Codata + $(EVERYTHINGS) check Codata .PHONY : check-README check-README: From 2ae84643c74750b49865e44c05b508cb2117c740 Mon Sep 17 00:00:00 2001 From: Steven Schaefer <32493857+stschaef@users.noreply.github.com> Date: Fri, 17 May 2024 09:42:48 -0400 Subject: [PATCH 16/30] Displayed Category Constructions (#1108) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Displayed fullsubcategories/preorders/functor introduction principles Co-authored-by: Max New * Partial definiton of \exists F (adjoint to reindexing) * Remove adjoint to reindex * Define left adjoint to reindexing * Rename intro for fullsubcategory * (Left)AdjontToReindex simplification * move intros into separate files, move preorder * fstfaithful * fix imports * Simplify adjoint definition * Preorderᴰ > StructureOverC, Fullsubcategoryᴰ > PropertyOverC, clean up imports * PropertyOverC > PropertyOver, Weaken to own file, total category moved * Fix whitespace + move LeftAdjointToReindex * Add left adjoint to reindexing comments * expanded comments * fix typo * isFaithfulFst --------- Co-authored-by: Max New --- .../Constructions/TotalCategory.agda | 74 ++++++++++ Cubical/Categories/Displayed/Base.agda | 78 +++-------- .../Constructions/LeftAdjointToReindex.agda | 126 ++++++++++++++++++ .../Displayed/Constructions/PropertyOver.agda | 51 +++++++ .../Constructions/StructureOver.agda | 58 ++++++++ .../Constructions/TotalCategory.agda | 82 ++++++++++++ Cubical/Categories/Displayed/Functor.agda | 19 ++- Cubical/Categories/Displayed/HLevels.agda | 78 +++++++++++ .../Displayed/Instances/Terminal.agda | 31 +++++ Cubical/Categories/Displayed/Weaken.agda | 46 +++++++ Cubical/Data/Sigma/Properties.agda | 3 + 11 files changed, 574 insertions(+), 72 deletions(-) create mode 100644 Cubical/Categories/Constructions/TotalCategory.agda create mode 100644 Cubical/Categories/Displayed/Constructions/LeftAdjointToReindex.agda create mode 100644 Cubical/Categories/Displayed/Constructions/PropertyOver.agda create mode 100644 Cubical/Categories/Displayed/Constructions/StructureOver.agda create mode 100644 Cubical/Categories/Displayed/Constructions/TotalCategory.agda create mode 100644 Cubical/Categories/Displayed/HLevels.agda create mode 100644 Cubical/Categories/Displayed/Instances/Terminal.agda create mode 100644 Cubical/Categories/Displayed/Weaken.agda diff --git a/Cubical/Categories/Constructions/TotalCategory.agda b/Cubical/Categories/Constructions/TotalCategory.agda new file mode 100644 index 0000000000..772c440599 --- /dev/null +++ b/Cubical/Categories/Constructions/TotalCategory.agda @@ -0,0 +1,74 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Constructions.TotalCategory where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Functor +open import Cubical.Categories.Displayed.Instances.Terminal + +private + variable + ℓC ℓC' ℓD ℓD' ℓE ℓE' ℓCᴰ ℓCᴰ' ℓDᴰ ℓDᴰ' ℓEᴰ ℓEᴰ' : Level + +-- Total category of a displayed category +module _ {C : Category ℓC ℓC'} (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') where + + open Category + open Categoryᴰ Cᴰ + private + module C = Category C + + ∫C : Category (ℓ-max ℓC ℓCᴰ) (ℓ-max ℓC' ℓCᴰ') + ∫C .ob = Σ _ ob[_] + ∫C .Hom[_,_] (_ , xᴰ) (_ , yᴰ) = Σ _ Hom[_][ xᴰ , yᴰ ] + ∫C .id = _ , idᴰ + ∫C ._⋆_ (_ , fᴰ) (_ , gᴰ) = _ , fᴰ ⋆ᴰ gᴰ + ∫C .⋆IdL _ = ΣPathP (_ , ⋆IdLᴰ _) + ∫C .⋆IdR _ = ΣPathP (_ , ⋆IdRᴰ _) + ∫C .⋆Assoc _ _ _ = ΣPathP (_ , ⋆Assocᴰ _ _ _) + ∫C .isSetHom = isSetΣ C.isSetHom (λ _ → isSetHomᴰ) + +-- Total functor of a displayed functor +module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} + {F : Functor C D} {Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ'} {Dᴰ : Categoryᴰ D ℓDᴰ ℓDᴰ'} + (Fᴰ : Functorᴰ F Cᴰ Dᴰ) + where + + open Functor + private + module Fᴰ = Functorᴰ Fᴰ + + ∫F : Functor (∫C Cᴰ) (∫C Dᴰ) + ∫F .F-ob (x , xᴰ) = _ , Fᴰ.F-obᴰ xᴰ + ∫F .F-hom (_ , fᴰ) = _ , Fᴰ.F-homᴰ fᴰ + ∫F .F-id = ΣPathP (_ , Fᴰ.F-idᴰ) + ∫F .F-seq _ _ = ΣPathP (_ , (Fᴰ.F-seqᴰ _ _)) + +-- Projections out of the total category +module _ {C : Category ℓC ℓC'} {Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ'} where + open Functor + open Functorᴰ + + Fst : Functor (∫C Cᴰ) C + Fst .F-ob = fst + Fst .F-hom = fst + Fst .F-id = refl + Fst .F-seq = + λ f g → cong {x = f ⋆⟨ ∫C Cᴰ ⟩ g} fst refl + + -- Functor into the total category + module _ {D : Category ℓD ℓD'} + (F : Functor D C) + (Fᴰ : Functorᴰ F (Unitᴰ D) Cᴰ) + where + intro : Functor D (∫C Cᴰ) + intro .F-ob d = F ⟅ d ⟆ , Fᴰ .F-obᴰ _ + intro .F-hom f = (F ⟪ f ⟫) , (Fᴰ .F-homᴰ _) + intro .F-id = ΣPathP (F .F-id , Fᴰ .F-idᴰ) + intro .F-seq f g = ΣPathP (F .F-seq f g , Fᴰ .F-seqᴰ _ _) diff --git a/Cubical/Categories/Displayed/Base.agda b/Cubical/Categories/Displayed/Base.agda index a71017d0da..7983698790 100644 --- a/Cubical/Categories/Displayed/Base.agda +++ b/Cubical/Categories/Displayed/Base.agda @@ -6,13 +6,11 @@ module Cubical.Categories.Displayed.Base where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -open import Cubical.Data.Sigma open import Cubical.Categories.Category.Base private variable - ℓC ℓC' ℓCᴰ ℓCᴰ' ℓDᴰ ℓDᴰ' : Level + ℓC ℓC' ℓD ℓD' ℓCᴰ ℓCᴰ' ℓDᴰ ℓDᴰ' : Level -- Displayed categories with hom-sets record Categoryᴰ (C : Category ℓC ℓC') ℓCᴰ ℓCᴰ' : Type (ℓ-suc (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓCᴰ ℓCᴰ'))) where @@ -49,64 +47,6 @@ record Categoryᴰ (C : Category ℓC ℓC') ℓCᴰ ℓCᴰ' : Type (ℓ-suc ( -- Helpful syntax/notation _[_][_,_] = Categoryᴰ.Hom[_][_,_] --- Total category of a displayed category -module _ {C : Category ℓC ℓC'} (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') where - - open Category - open Categoryᴰ Cᴰ - private - module C = Category C - - ∫C : Category (ℓ-max ℓC ℓCᴰ) (ℓ-max ℓC' ℓCᴰ') - ∫C .ob = Σ _ ob[_] - ∫C .Hom[_,_] (_ , xᴰ) (_ , yᴰ) = Σ _ Hom[_][ xᴰ , yᴰ ] - ∫C .id = _ , idᴰ - ∫C ._⋆_ (_ , fᴰ) (_ , gᴰ) = _ , fᴰ ⋆ᴰ gᴰ - ∫C .⋆IdL _ = ΣPathP (_ , ⋆IdLᴰ _) - ∫C .⋆IdR _ = ΣPathP (_ , ⋆IdRᴰ _) - ∫C .⋆Assoc _ _ _ = ΣPathP (_ , ⋆Assocᴰ _ _ _) - ∫C .isSetHom = isSetΣ C.isSetHom (λ _ → isSetHomᴰ) - --- Displayed total category, i.e. Σ for displayed categories -module _ {C : Category ℓC ℓC'} - (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') - (Dᴰ : Categoryᴰ (∫C Cᴰ) ℓDᴰ ℓDᴰ') - where - - open Categoryᴰ - private - module Cᴰ = Categoryᴰ Cᴰ - module Dᴰ = Categoryᴰ Dᴰ - - ∫Cᴰ : Categoryᴰ C (ℓ-max ℓCᴰ ℓDᴰ) (ℓ-max ℓCᴰ' ℓDᴰ') - ∫Cᴰ .ob[_] x = Σ[ xᴰ ∈ Cᴰ.ob[ x ] ] Dᴰ.ob[ x , xᴰ ] - ∫Cᴰ .Hom[_][_,_] f (_ , zᴰ) (_ , wᴰ) = Σ[ fᴰ ∈ Cᴰ.Hom[ f ][ _ , _ ] ] Dᴰ.Hom[ f , fᴰ ][ zᴰ , wᴰ ] - ∫Cᴰ .idᴰ = Cᴰ.idᴰ , Dᴰ.idᴰ - ∫Cᴰ ._⋆ᴰ_ (_ , hᴰ) (_ , kᴰ) = _ , hᴰ Dᴰ.⋆ᴰ kᴰ - ∫Cᴰ .⋆IdLᴰ _ = ΣPathP (_ , Dᴰ.⋆IdLᴰ _) - ∫Cᴰ .⋆IdRᴰ _ = ΣPathP (_ , Dᴰ.⋆IdRᴰ _) - ∫Cᴰ .⋆Assocᴰ _ _ _ = ΣPathP (_ , Dᴰ.⋆Assocᴰ _ _ _) - ∫Cᴰ .isSetHomᴰ = isSetΣ Cᴰ.isSetHomᴰ (λ _ → Dᴰ.isSetHomᴰ) - -module _ {C : Category ℓC ℓC'} - (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') - (Dᴰ : Categoryᴰ C ℓDᴰ ℓDᴰ') - where - - open Categoryᴰ - private - module Dᴰ = Categoryᴰ Dᴰ - - weakenᴰ : Categoryᴰ (∫C Cᴰ) ℓDᴰ ℓDᴰ' - weakenᴰ .ob[_] (x , _) = Dᴰ.ob[ x ] - weakenᴰ .Hom[_][_,_] (f , _) = Dᴰ.Hom[ f ][_,_] - weakenᴰ .idᴰ = Dᴰ.idᴰ - weakenᴰ ._⋆ᴰ_ = Dᴰ._⋆ᴰ_ - weakenᴰ .⋆IdLᴰ = Dᴰ.⋆IdLᴰ - weakenᴰ .⋆IdRᴰ = Dᴰ.⋆IdRᴰ - weakenᴰ .⋆Assocᴰ = Dᴰ.⋆Assocᴰ - weakenᴰ .isSetHomᴰ = Dᴰ.isSetHomᴰ - module _ {C : Category ℓC ℓC'} (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') where open Category C open Categoryᴰ Cᴰ @@ -127,3 +67,19 @@ module _ {C : Category ℓC ℓC'} (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') where idᴰCatIsoᴰ : {x : ob} {xᴰ : ob[ x ]} → CatIsoᴰ idCatIso xᴰ xᴰ idᴰCatIsoᴰ = idᴰ , isisoᴰ idᴰ (⋆IdLᴰ idᴰ) (⋆IdLᴰ idᴰ) + +module _ {C : Category ℓC ℓC'} (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') where + open Category + private + module Cᴰ = Categoryᴰ Cᴰ + + open Categoryᴰ + _^opᴰ : Categoryᴰ (C ^op) ℓCᴰ ℓCᴰ' + _^opᴰ .ob[_] x = Cᴰ.ob[ x ] + _^opᴰ .Hom[_][_,_] f xᴰ yᴰ = Cᴰ.Hom[ f ][ yᴰ , xᴰ ] + _^opᴰ .idᴰ = Cᴰ.idᴰ + _^opᴰ ._⋆ᴰ_ fᴰ gᴰ = gᴰ Cᴰ.⋆ᴰ fᴰ + _^opᴰ .⋆IdLᴰ = Cᴰ .⋆IdRᴰ + _^opᴰ .⋆IdRᴰ = Cᴰ .⋆IdLᴰ + _^opᴰ .⋆Assocᴰ fᴰ gᴰ hᴰ = symP (Cᴰ.⋆Assocᴰ _ _ _) + _^opᴰ .isSetHomᴰ = Cᴰ .isSetHomᴰ diff --git a/Cubical/Categories/Displayed/Constructions/LeftAdjointToReindex.agda b/Cubical/Categories/Displayed/Constructions/LeftAdjointToReindex.agda new file mode 100644 index 0000000000..b65ceef930 --- /dev/null +++ b/Cubical/Categories/Displayed/Constructions/LeftAdjointToReindex.agda @@ -0,0 +1,126 @@ +{-# OPTIONS --safe #-} +-- +module Cubical.Categories.Displayed.Constructions.LeftAdjointToReindex where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.Sigma + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Instances.Terminal +open import Cubical.Categories.Functor + +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Properties +open import Cubical.Categories.Constructions.TotalCategory + +private + variable + ℓC ℓC' ℓCᴰ ℓCᴰ' ℓD ℓD' ℓDᴰ ℓDᴰ' ℓE ℓE' ℓEᴰ ℓEᴰ' : Level + +open Categoryᴰ + +module _ + {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} + (Cᴰ : Categoryᴰ C ℓDᴰ ℓDᴰ') (F : Functor C D) + where + + open Category + open Functor F + private + module Cᴰ = Categoryᴰ Cᴰ + open Iso + + hom-path : ∀ {A B A' B'} (p : A ≡ A') (q : B ≡ B') → + (D [ A , B ]) ≡ (D [ A' , B' ]) + hom-path p q = cong₂ (λ a b → D [ a , b ]) p q + + hom-pathP : ∀ {A B A' B'} (p : A ≡ A') (q : B ≡ B') → + (f : D [ A , B ]) → (f' : D [ A' , B' ]) → + Type ℓD' + hom-pathP p q f f' = PathP (λ i → hom-path p q i) f f' + + isPropHomP : ∀ {A B A' B'} (p : A ≡ A') (q : B ≡ B') → + (f : D [ A , B ]) → (f' : D [ A' , B' ]) → + isProp (hom-pathP p q f f') + isPropHomP pdom pcod f f'' x y = + isoFunInjective (PathPIsoPath _ _ _) x y fromPathPeq + where + fromPathPeq : fromPathP x ≡ fromPathP y + fromPathPeq = D .isSetHom _ _ (fromPathP x) (fromPathP y) + + -- Reindexing a displayed category Dᴰ over D along a functor F : C → D + -- gives a displayed category over C + -- Fiberwise pullback the objects over D along F to display them over C + -- + -- reindex Dᴰ F Dᴰ + -- _ _ + -- | | + -- v F v + -- C -------------------> D + -- + -- Which may be read as a 2-functor from displayed categories over D to + -- displayed categories over C. This operation has a left 2-adjoint, which + -- we call ∃F, that maps displays over C to displays over D + -- + -- Cᴰ ∃F Cᴰ F + -- _ _ + -- | | + -- v F v + -- C -------------------> D + -- + ∃F : Categoryᴰ D (ℓ-max (ℓ-max ℓC ℓD) ℓDᴰ) (ℓ-max (ℓ-max ℓC' ℓD') ℓDᴰ') + ∃F .ob[_] d = Σ[ c ∈ C .ob ] Cᴰ.ob[ c ] × (F-ob c ≡ d) + ∃F .Hom[_][_,_] f (c , x , p) (c' , x' , p') = + Σ[ g ∈ C [ c , c' ] ] Cᴰ.Hom[ g ][ x , x' ] × hom-pathP p p' (F-hom g) f + ∃F .idᴰ {d} {c , x , p} = + C .id , + Cᴰ .idᴰ , + (F-id ◁ (cong (λ v → D .id {v}) p)) + ∃F ._⋆ᴰ_ (g , gᴰ , p) (h , hᴰ , q) = + g ⋆⟨ C ⟩ h , + (Cᴰ._⋆ᴰ_ gᴰ hᴰ) , + (F-seq g h ◁ congP₂ (λ _ a b → a ⋆⟨ D ⟩ b) p q) + ⋆IdLᴰ ∃F (f , x , p) = + ΣPathP (C .⋆IdL f , ΣPathPProp (λ _ → isPropHomP _ _ _ _) (Cᴰ .⋆IdLᴰ x)) + ⋆IdRᴰ ∃F (f , x , p) = + ΣPathP (C .⋆IdR f , ΣPathPProp (λ _ → isPropHomP _ _ _ _) (Cᴰ .⋆IdRᴰ x)) + ⋆Assocᴰ ∃F + (f , pf , fᴰ) + (g , pg , gᴰ) + (h , ph , hᴰ) = + ΣPathP (C .⋆Assoc _ _ _ , + ΣPathPProp (λ _ → isPropHomP _ _ _ _) (Cᴰ .⋆Assocᴰ _ _ _)) + isSetHomᴰ ∃F {d}{d'}{f}{u}{v} = + isSetΣ (C .isSetHom) + (λ g → isSet× (Cᴰ .isSetHomᴰ) (isProp→isSet (isPropHomP _ _ _ _))) + +-- Examples of ∃F instantiated +private + module _ where + -- Can define the displayed total category via ∃F + module _ + {C : Category ℓC ℓC'} + {ℓ : Level} + (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') + (Dᴰ : Categoryᴰ (∫C Cᴰ) ℓDᴰ ℓDᴰ') + where + ∫Cᴰ' : Categoryᴰ C + (ℓ-max (ℓ-max (ℓ-max ℓC ℓCᴰ) ℓC) ℓDᴰ) + (ℓ-max (ℓ-max (ℓ-max ℓC' ℓCᴰ') ℓC') ℓDᴰ') + ∫Cᴰ' = ∃F Dᴰ (Fst {C = C} {Cᴰ = Cᴰ}) + + module _ + {C : Category ℓC ℓC'} + {D : Category ℓD ℓD'} + {ℓ : Level} + (F : Functor (TerminalCategory {ℓ-zero}) C) + (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') + (Dᴰ : Categoryᴰ (∫C Cᴰ) ℓDᴰ ℓDᴰ') + where + + weaken' : Categoryᴰ C (ℓ-max (ℓ-max ℓ-zero ℓC) ℓD) + (ℓ-max (ℓ-max ℓ-zero ℓC') ℓD') + weaken' = ∃F (Category→DispOverTerminal D) F diff --git a/Cubical/Categories/Displayed/Constructions/PropertyOver.agda b/Cubical/Categories/Displayed/Constructions/PropertyOver.agda new file mode 100644 index 0000000000..a1cd569ac8 --- /dev/null +++ b/Cubical/Categories/Displayed/Constructions/PropertyOver.agda @@ -0,0 +1,51 @@ +-- | Property displayed over a category. +{-# OPTIONS --safe #-} +module Cubical.Categories.Displayed.Constructions.PropertyOver where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Unit + +open import Cubical.Categories.Category renaming (isIso to isIsoC) +open import Cubical.Categories.Functor +open import Cubical.Categories.Displayed.Functor +open import Cubical.Categories.Displayed.Constructions.StructureOver +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.HLevels + +private + variable + ℓC ℓC' ℓD ℓD' ℓDᴰ ℓDᴰ' ℓP ℓS ℓR : Level + +open Category +open Functor +open Categoryᴰ +open Functorᴰ + +module _ (C : Category ℓC ℓC') (P : Category.ob C → Type ℓP) where + private + module C = Category C + open Category + open Functor + + PropertyOver : Categoryᴰ C ℓP ℓ-zero + PropertyOver = StructureOver→Catᴰ struct where + open StructureOver + struct : StructureOver C ℓP ℓ-zero + struct .ob[_] = P + struct .Hom[_][_,_] _ _ _ = Unit + struct .idᴰ = tt + struct ._⋆ᴰ_ = λ _ _ → tt + struct .isPropHomᴰ = isPropUnit + + hasContrHomsPropertyOver : hasContrHoms PropertyOver + hasContrHomsPropertyOver _ _ _ = isContrUnit + + module _ {D : Category ℓD ℓD'} {Dᴰ : Categoryᴰ D ℓDᴰ ℓDᴰ'} + (F : Functor D C) + (F-obᴰ : {x : D .ob} → + Dᴰ .ob[_] x → ob[ PropertyOver ] (F .F-ob x)) + where + intro : Functorᴰ F Dᴰ PropertyOver + intro = + mkContrHomsFunctor hasContrHomsPropertyOver F-obᴰ diff --git a/Cubical/Categories/Displayed/Constructions/StructureOver.agda b/Cubical/Categories/Displayed/Constructions/StructureOver.agda new file mode 100644 index 0000000000..7b38c0caaf --- /dev/null +++ b/Cubical/Categories/Displayed/Constructions/StructureOver.agda @@ -0,0 +1,58 @@ +-- | Structure displayed over a category. +{-# OPTIONS --safe #-} +module Cubical.Categories.Displayed.Constructions.StructureOver where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Functor +open import Cubical.Categories.Displayed.HLevels +open import Cubical.Categories.Constructions.TotalCategory + +private + variable + ℓC ℓC' ℓCᴰ ℓCᴰ' ℓD ℓD' ℓDᴰ ℓDᴰ' : Level + +record StructureOver (C : Category ℓC ℓC') ℓCᴰ ℓCᴰ' : + Type (ℓ-suc (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓCᴰ ℓCᴰ'))) where + open Category C + field + ob[_] : ob → Type ℓCᴰ + Hom[_][_,_] : {x y : ob} → Hom[ x , y ] → ob[ x ] → ob[ y ] → Type ℓCᴰ' + idᴰ : ∀ {x} {p : ob[ x ]} → Hom[ id ][ p , p ] + _⋆ᴰ_ : ∀ {x y z} {f : Hom[ x , y ]} {g : Hom[ y , z ]} {xᴰ yᴰ zᴰ} + → Hom[ f ][ xᴰ , yᴰ ] → Hom[ g ][ yᴰ , zᴰ ] → Hom[ f ⋆ g ][ xᴰ , zᴰ ] + isPropHomᴰ : ∀ {x y} {f : Hom[ x , y ]} {xᴰ yᴰ} → isProp Hom[ f ][ xᴰ , yᴰ ] + +module _ {C : Category ℓC ℓC'} (Pᴰ : StructureOver C ℓCᴰ ℓCᴰ') where + open Category + open StructureOver + StructureOver→Catᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ' + StructureOver→Catᴰ = record + { ob[_] = Pᴰ .ob[_] + ; Hom[_][_,_] = Pᴰ .Hom[_][_,_] + ; idᴰ = Pᴰ .idᴰ + ; _⋆ᴰ_ = Pᴰ ._⋆ᴰ_ + ; ⋆IdLᴰ = λ _ → + isProp→PathP ((λ i → Pᴰ .isPropHomᴰ {f = ((C .⋆IdL _) i)})) _ _ + ; ⋆IdRᴰ = λ _ → + isProp→PathP ((λ i → Pᴰ .isPropHomᴰ {f = ((C .⋆IdR _) i)})) _ _ + ; ⋆Assocᴰ = λ _ _ _ → + isProp→PathP ((λ i → Pᴰ .isPropHomᴰ {f = ((C .⋆Assoc _ _ _) i)})) _ _ + ; isSetHomᴰ = isProp→isSet (Pᴰ .isPropHomᴰ) + } + + hasPropHomsStructureOver : hasPropHoms StructureOver→Catᴰ + hasPropHomsStructureOver _ _ _ = Pᴰ .isPropHomᴰ + + open Functor + + isFaithfulFst : isFaithful (Fst {Cᴰ = StructureOver→Catᴰ}) + isFaithfulFst x y f g p = + ΣPathP (p , + isProp→PathP (λ i → Pᴰ .isPropHomᴰ {f = p i}) (f .snd) (g .snd)) diff --git a/Cubical/Categories/Displayed/Constructions/TotalCategory.agda b/Cubical/Categories/Displayed/Constructions/TotalCategory.agda new file mode 100644 index 0000000000..d4c6b9095d --- /dev/null +++ b/Cubical/Categories/Displayed/Constructions/TotalCategory.agda @@ -0,0 +1,82 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Displayed.Constructions.TotalCategory where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Functor +open import Cubical.Categories.Displayed.Instances.Terminal +open import Cubical.Categories.Constructions.TotalCategory hiding (intro) + +private + variable + ℓC ℓC' ℓD ℓD' ℓE ℓE' ℓCᴰ ℓCᴰ' ℓDᴰ ℓDᴰ' ℓEᴰ ℓEᴰ' : Level + +-- Displayed total category, i.e. Σ for displayed categories +-- +-- The ordinary total category (∫C Cᴰ) has as objects +-- pairs (x, xᴰ) where x is an object of C and xᴰ is an object of Cᴰ over x. +-- +-- Whereas if we had a category Dᴰ displayed over (∫C Cᴰ), +-- and a category Cᴰ displayed over C, then the +-- displayed total category (∫Cᴰ Cᴰ Dᴰ) likewise has pairs +-- as displayed objects. +-- +-- In the displayed total category, we have objects (xᴰ, yᴰ) displayed +-- over x, where x is an object of C, xᴰ an object in Cᴰ displayed over x, +-- and yᴰ is an object of Dᴰ over (x, xᴰ). +module _ {C : Category ℓC ℓC'} + (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') + (Dᴰ : Categoryᴰ (∫C Cᴰ) ℓDᴰ ℓDᴰ') + where + + open Categoryᴰ + private + module Cᴰ = Categoryᴰ Cᴰ + module Dᴰ = Categoryᴰ Dᴰ + + ∫Cᴰ : Categoryᴰ C (ℓ-max ℓCᴰ ℓDᴰ) (ℓ-max ℓCᴰ' ℓDᴰ') + ∫Cᴰ .ob[_] x = Σ[ xᴰ ∈ Cᴰ.ob[ x ] ] Dᴰ.ob[ x , xᴰ ] + ∫Cᴰ .Hom[_][_,_] f (_ , zᴰ) (_ , wᴰ) = + Σ[ fᴰ ∈ Cᴰ.Hom[ f ][ _ , _ ] ] Dᴰ.Hom[ f , fᴰ ][ zᴰ , wᴰ ] + ∫Cᴰ .idᴰ = Cᴰ.idᴰ , Dᴰ.idᴰ + ∫Cᴰ ._⋆ᴰ_ (_ , hᴰ) (_ , kᴰ) = _ , hᴰ Dᴰ.⋆ᴰ kᴰ + ∫Cᴰ .⋆IdLᴰ _ = ΣPathP (_ , Dᴰ.⋆IdLᴰ _) + ∫Cᴰ .⋆IdRᴰ _ = ΣPathP (_ , Dᴰ.⋆IdRᴰ _) + ∫Cᴰ .⋆Assocᴰ _ _ _ = ΣPathP (_ , Dᴰ.⋆Assocᴰ _ _ _) + ∫Cᴰ .isSetHomᴰ = isSetΣ Cᴰ.isSetHomᴰ (λ _ → Dᴰ.isSetHomᴰ) + +-- Projection out of the displayed total category +module _ {C : Category ℓC ℓC'} + {Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ'} + (Dᴰ : Categoryᴰ (∫C Cᴰ) ℓDᴰ ℓDᴰ') + where + + open Functorᴰ + private + module Cᴰ = Categoryᴰ Cᴰ + module Dᴰ = Categoryᴰ Dᴰ + + Fstᴰ : Functorᴰ Id (∫Cᴰ Cᴰ Dᴰ) Cᴰ + Fstᴰ .F-obᴰ = fst + Fstᴰ .F-homᴰ = fst + Fstᴰ .F-idᴰ = refl + Fstᴰ .F-seqᴰ _ _ = refl + + -- Functor into the displayed total category + module _ {E : Category ℓE ℓE'} (F : Functor E C) + {Eᴰ : Categoryᴰ E ℓEᴰ ℓEᴰ'} + (Fᴰ : Functorᴰ F Eᴰ Cᴰ) + (Gᴰ : Functorᴰ (∫F Fᴰ) (Unitᴰ (∫C Eᴰ)) Dᴰ) + where + + intro : Functorᴰ F Eᴰ (∫Cᴰ Cᴰ Dᴰ) + intro .F-obᴰ xᴰ = Fᴰ .F-obᴰ xᴰ , Gᴰ .F-obᴰ _ + intro .F-homᴰ fᴰ = (Fᴰ .F-homᴰ fᴰ) , (Gᴰ .F-homᴰ _) + intro .F-idᴰ = ΣPathP (Fᴰ .F-idᴰ , Gᴰ .F-idᴰ) + intro .F-seqᴰ fᴰ gᴰ = ΣPathP (Fᴰ .F-seqᴰ fᴰ gᴰ , Gᴰ .F-seqᴰ _ _) diff --git a/Cubical/Categories/Displayed/Functor.agda b/Cubical/Categories/Displayed/Functor.agda index 263636175c..426edbc610 100644 --- a/Cubical/Categories/Displayed/Functor.agda +++ b/Cubical/Categories/Displayed/Functor.agda @@ -9,9 +9,9 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.HLevels -open import Cubical.Data.Sigma open import Cubical.Categories.Category.Base open import Cubical.Categories.Functor + open import Cubical.Categories.Displayed.Base private @@ -142,18 +142,15 @@ module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} {F : Functor C D} F-rUnitᴰ i .F-idᴰ {x} = rUnitP' (Dᴰ [_][ _ , _ ]) Fᴰ.F-idᴰ (~ i) F-rUnitᴰ i .F-seqᴰ _ _ = rUnitP' (Dᴰ [_][ _ , _ ]) (Fᴰ.F-seqᴰ _ _) (~ i) --- Total functor of a displayed functor +-- Displayed opposite functor module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} {F : Functor C D} {Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ'} {Dᴰ : Categoryᴰ D ℓDᴰ ℓDᴰ'} (Fᴰ : Functorᴰ F Cᴰ Dᴰ) where + open Functorᴰ + _^opFᴰ : Functorᴰ (F ^opF) (Cᴰ ^opᴰ) (Dᴰ ^opᴰ) + _^opFᴰ .F-obᴰ = Fᴰ .F-obᴰ + _^opFᴰ .F-homᴰ = Fᴰ .F-homᴰ + _^opFᴰ .F-idᴰ = Fᴰ .F-idᴰ + _^opFᴰ .F-seqᴰ fᴰ gᴰ = Fᴰ .F-seqᴰ gᴰ fᴰ - open Functor - private - module Fᴰ = Functorᴰ Fᴰ - - ∫F : Functor (∫C Cᴰ) (∫C Dᴰ) - ∫F .F-ob (x , xᴰ) = _ , Fᴰ.F-obᴰ xᴰ - ∫F .F-hom (_ , fᴰ) = _ , Fᴰ.F-homᴰ fᴰ - ∫F .F-id = ΣPathP (_ , Fᴰ.F-idᴰ) - ∫F .F-seq _ _ = ΣPathP (_ , (Fᴰ.F-seqᴰ _ _)) diff --git a/Cubical/Categories/Displayed/HLevels.agda b/Cubical/Categories/Displayed/HLevels.agda new file mode 100644 index 0000000000..839f525a7f --- /dev/null +++ b/Cubical/Categories/Displayed/HLevels.agda @@ -0,0 +1,78 @@ +{-# OPTIONS --safe #-} +{-- This file contains some utilities for reasoning + -- about the HLevels of morphisms in displayed categories. + --} +module Cubical.Categories.Displayed.HLevels where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor + +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Functor + +private + variable + ℓC ℓC' ℓCᴰ ℓCᴰ' ℓD ℓD' ℓDᴰ ℓDᴰ' ℓE ℓE' ℓEᴰ ℓEᴰ' : Level + +open Categoryᴰ + +module _ {C : Category ℓC ℓC'} (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') where + open Category + private + module Cᴰ = Categoryᴰ Cᴰ + + hasContrHoms : Type _ + hasContrHoms = + ∀ {c c' : C .ob}(f : C [ c , c' ])(cᴰ : Cᴰ.ob[ c ])(cᴰ' : Cᴰ.ob[ c' ]) + → isContr Cᴰ.Hom[ f ][ cᴰ , cᴰ' ] + + hasPropHoms : Type _ + hasPropHoms = + ∀ {c c' : C .ob}(f : C [ c , c' ])(cᴰ : Cᴰ.ob[ c ])(cᴰ' : Cᴰ.ob[ c' ]) + → isProp Cᴰ.Hom[ f ][ cᴰ , cᴰ' ] + + hasContrHoms→hasPropHoms : hasContrHoms → hasPropHoms + hasContrHoms→hasPropHoms contrHoms = + λ f cᴰ cᴰ' → isContr→isProp (contrHoms f cᴰ cᴰ') + +module _ + {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} + {F : Functor C D} + {Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ'} + {Dᴰ : Categoryᴰ D ℓDᴰ ℓDᴰ'} + where + open Category + open Functor + private + module Cᴰ = Categoryᴰ Cᴰ + module Dᴰ = Categoryᴰ Dᴰ + + module _ + (propHoms : hasPropHoms Dᴰ) + where + + mkPropHomsFunctor : + (F-obᴰ : {x : C .ob} → Cᴰ.ob[ x ] → Dᴰ.ob[ F .F-ob x ]) + → (F-homᴰ : {x y : C .ob} + {f : C [ x , y ]} {xᴰ : Cᴰ.ob[ x ]} {yᴰ : Cᴰ.ob[ y ]} + → Cᴰ [ f ][ xᴰ , yᴰ ] → Dᴰ [ F .F-hom f ][ F-obᴰ xᴰ , F-obᴰ yᴰ ]) + → Functorᴰ F Cᴰ Dᴰ + mkPropHomsFunctor F-obᴰ F-homᴰ .Functorᴰ.F-obᴰ = F-obᴰ + mkPropHomsFunctor F-obᴰ F-homᴰ .Functorᴰ.F-homᴰ = F-homᴰ + mkPropHomsFunctor F-obᴰ F-homᴰ .Functorᴰ.F-idᴰ = + isProp→PathP (λ i → propHoms _ _ _) _ _ + mkPropHomsFunctor F-obᴰ F-homᴰ .Functorᴰ.F-seqᴰ _ _ = + isProp→PathP (λ i → propHoms _ _ _) _ _ + + module _ + (contrHoms : hasContrHoms Dᴰ) + where + + mkContrHomsFunctor : (F-obᴰ : {x : C .ob} → Cᴰ.ob[ x ] → Dᴰ.ob[ F .F-ob x ]) + → Functorᴰ F Cᴰ Dᴰ + mkContrHomsFunctor F-obᴰ = + mkPropHomsFunctor (hasContrHoms→hasPropHoms Dᴰ contrHoms) F-obᴰ + λ _ → contrHoms _ _ _ .fst diff --git a/Cubical/Categories/Displayed/Instances/Terminal.agda b/Cubical/Categories/Displayed/Instances/Terminal.agda new file mode 100644 index 0000000000..9f21f54a91 --- /dev/null +++ b/Cubical/Categories/Displayed/Instances/Terminal.agda @@ -0,0 +1,31 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Displayed.Instances.Terminal where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Unit + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Instances.Terminal +open import Cubical.Categories.Displayed.Base + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +open Category +open Categoryᴰ +-- Terminal category over a base category +Unitᴰ : (C : Category ℓC ℓC') → Categoryᴰ C ℓ-zero ℓ-zero +Unitᴰ C .ob[_] x = Unit +Unitᴰ C .Hom[_][_,_] f tt tt = Unit +Unitᴰ C .idᴰ = tt +Unitᴰ C ._⋆ᴰ_ = λ _ _ → tt +Unitᴰ C .⋆IdLᴰ fᴰ i = tt +Unitᴰ C .⋆IdRᴰ fᴰ i = tt +Unitᴰ C .⋆Assocᴰ fᴰ gᴰ hᴰ i = tt +Unitᴰ C .isSetHomᴰ x y x₁ y₁ i i₁ = tt + +-- Terminal category over the terminal category +UnitCᴰ : Categoryᴰ (TerminalCategory {ℓ-zero}) ℓ-zero ℓ-zero +UnitCᴰ = Unitᴰ TerminalCategory diff --git a/Cubical/Categories/Displayed/Weaken.agda b/Cubical/Categories/Displayed/Weaken.agda new file mode 100644 index 0000000000..41d645d32c --- /dev/null +++ b/Cubical/Categories/Displayed/Weaken.agda @@ -0,0 +1,46 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Displayed.Weaken where + +open import Cubical.Foundations.Prelude + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Constructions.TotalCategory + +private + variable + ℓC ℓC' ℓD ℓD' ℓCᴰ ℓCᴰ' ℓDᴰ ℓDᴰ' : Level + +-- Display a category D over C +module _ (C : Category ℓC ℓC') (D : Category ℓD ℓD') where + open Category + open Categoryᴰ + weaken : Categoryᴰ C ℓD ℓD' + weaken .ob[_] x = D .ob + weaken .Hom[_][_,_] f d d' = D [ d , d' ] + weaken .idᴰ = D .id + weaken ._⋆ᴰ_ = D ._⋆_ + weaken .⋆IdLᴰ = D .⋆IdL + weaken .⋆IdRᴰ = D .⋆IdR + weaken .⋆Assocᴰ = D .⋆Assoc + weaken .isSetHomᴰ = D .isSetHom + +-- Weaken a displayed category Dᴰ over Cᴰ +module _ {C : Category ℓC ℓC'} + (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') + (Dᴰ : Categoryᴰ C ℓDᴰ ℓDᴰ') + where + + open Categoryᴰ + private + module Dᴰ = Categoryᴰ Dᴰ + + weakenᴰ : Categoryᴰ (∫C Cᴰ) ℓDᴰ ℓDᴰ' + weakenᴰ .ob[_] (x , _) = Dᴰ.ob[ x ] + weakenᴰ .Hom[_][_,_] (f , _) = Dᴰ.Hom[ f ][_,_] + weakenᴰ .idᴰ = Dᴰ.idᴰ + weakenᴰ ._⋆ᴰ_ = Dᴰ._⋆ᴰ_ + weakenᴰ .⋆IdLᴰ = Dᴰ.⋆IdLᴰ + weakenᴰ .⋆IdRᴰ = Dᴰ.⋆IdRᴰ + weakenᴰ .⋆Assocᴰ = Dᴰ.⋆Assocᴰ + weakenᴰ .isSetHomᴰ = Dᴰ.isSetHomᴰ diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index fea7559aad..68514fd5d8 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -90,6 +90,9 @@ module _ {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ'} ×≡Prop : isProp A' → {u v : A × A'} → u .fst ≡ v .fst → u ≡ v ×≡Prop pB {u} {v} p i = (p i) , (pB (u .snd) (v .snd) i) +×≡Prop' : isProp A → {u v : A × A'} → u .snd ≡ v .snd → u ≡ v +×≡Prop' pA {u} {v} p i = (pA (u .fst) (v .fst) i) , p i + -- Useful lemma to prove unique existence uniqueExists : (a : A) (b : B a) (h : (a' : A) → isProp (B a')) (H : (a' : A) → B a' → a ≡ a') → ∃![ a ∈ A ] B a fst (uniqueExists a b h H) = (a , b) From 502b1bb8a47fb8c07d82e1bc05020d5b4f10cede Mon Sep 17 00:00:00 2001 From: pi3js2 <126897438+pi3js2@users.noreply.github.com> Date: Wed, 22 May 2024 12:33:50 +0000 Subject: [PATCH 17/30] Simplify CartesianKanOps/FunExtEquiv using "interpolation" on I (#1001) * Simplify CartesianKanOps/FunExtEquiv using "erp" interpolation on I * Rename "erp" to "interpolateI" --- Cubical/Foundations/CartesianKanOps.agda | 44 ++++++++--------- Cubical/Foundations/Interpolate.agda | 53 ++++++++++++++++++++ Cubical/Functions/FunExtEquiv.agda | 63 +++++++++++++++--------- 3 files changed, 112 insertions(+), 48 deletions(-) create mode 100644 Cubical/Foundations/Interpolate.agda diff --git a/Cubical/Foundations/CartesianKanOps.agda b/Cubical/Foundations/CartesianKanOps.agda index 5d00cec502..63f87e333d 100644 --- a/Cubical/Foundations/CartesianKanOps.agda +++ b/Cubical/Foundations/CartesianKanOps.agda @@ -5,6 +5,7 @@ module Cubical.Foundations.CartesianKanOps where open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Transport +open import Cubical.Foundations.Interpolate coe0→1 : ∀ {ℓ} (A : I → Type ℓ) → A i0 → A i1 coe0→1 A a = transp (\ i → A i) i0 a @@ -45,20 +46,20 @@ 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 +-- "Equality" on the interval, chosen for the next definition: +-- interpolateI k i j is constant in k on eqI i j. Note that eqI i i +-- is not i1 but i ∨ ~ i. +private + eqI : I → I → I + eqI i j = (i ∧ j) ∨ (~ i ∧ ~ j) + -- "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 +coei→j A i j a = transp (λ k → A (interpolateI k i j)) (eqI i j) a -- "squeeze" --- this is just defined as the composite face of the master coe +-- this is just defined as the 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 @@ -83,13 +84,14 @@ coei1→i A i a = refl -- only non-definitional equation, but definitional at the ends coei→i : ∀ {ℓ} (A : I → Type ℓ) (i : I) (a : A i) → coei→j A i i a ≡ a -coei→i A i a j = - comp (λ k → A (i ∧ (j ∨ k))) - (λ k → λ - { (i = i0) → a - ; (i = i1) → coe1→i A (j ∨ k) a - ; (j = i1) → a }) - (transpFill {A = A i0} (~ i) (λ t → inS (A (i ∧ ~ t))) a (~ j)) +coei→i A i a j = transp (λ _ → A i) (interpolateI j (i ∨ ~ i) i1) a + where + -- note: coei→i is almost just transportRefl (but the φ for the + -- transp is i ∨ ~ i, not i0) + _ : Path (PathP (λ i → A i → A i) (λ a → a) (λ a → a)) + (λ i a → coei→j A i i a) + (λ i a → transp (λ _ → A i) (i ∨ ~ i) a) + _ = refl coe0→0 : ∀ {ℓ} (A : I → Type ℓ) (a : A i0) → coei→i A i0 a ≡ refl coe0→0 A a = refl @@ -99,14 +101,8 @@ coe1→1 A a = refl -- coercion when there already exists a path coePath : ∀ {ℓ} (A : I → Type ℓ) (p : (i : I) → A i) → (i j : I) → coei→j A i j (p i) ≡ p j -coePath A p i j = - hcomp (λ k → λ - { (i = i0)(j = i0) → rUnit refl (~ k) - ; (i = i1)(j = i1) → rUnit refl (~ k) }) - (diag ∙ coei→i A j (p j)) - where - diag : coei→j A i j (p i) ≡ coei→j A j j (p j) - diag k = coei→j A _ j (p ((j ∨ (i ∧ ~ k)) ∧ (i ∨ (j ∧ k)))) +coePath A p i j k = + transp (λ l → A (interpolateI l (interpolateI k i j) j)) (interpolateI k (eqI i j) i1) (p (interpolateI k i j)) coePathi0 : ∀ {ℓ} (A : I → Type ℓ) (p : (i : I) → A i) → coePath A p i0 i0 ≡ refl coePathi0 A p = refl diff --git a/Cubical/Foundations/Interpolate.agda b/Cubical/Foundations/Interpolate.agda new file mode 100644 index 0000000000..3d732fc8e6 --- /dev/null +++ b/Cubical/Foundations/Interpolate.agda @@ -0,0 +1,53 @@ +{-# OPTIONS --safe #-} + +module Cubical.Foundations.Interpolate where + +open import Cubical.Foundations.Prelude + +-- An "interpolation" operator on the De Morgan interval. Interpolates +-- along t from i to j (see first two properties below.) +interpolateI : I → I → I → I +interpolateI t i j = (~ t ∧ i) ∨ (t ∧ j) ∨ (i ∧ j) + +-- I believe this is the simplest De Morgan function on three +-- variables which satisfies all (or even most) of the nice properties +-- below. + +module _ + {A : Type} {p : I → A} {i j k l m : I} + where + _ : p (interpolateI i0 i j) ≡ p i + _ = refl + + _ : p (interpolateI i1 i j) ≡ p j + _ = refl + + _ : p (interpolateI (~ i) j k) ≡ p (interpolateI i k j) + _ = refl + + _ : p (interpolateI i i0 i1) ≡ p i + _ = refl + + _ : p (interpolateI i i1 i0) ≡ p (~ i) + _ = refl + + _ : p (interpolateI i j j) ≡ p j + _ = refl + + _ : p (interpolateI i (interpolateI j k l) m) ≡ p (interpolateI j (interpolateI i k m) (interpolateI i l m)) + _ = refl + + _ : p (interpolateI i j (interpolateI k l m)) ≡ p (interpolateI k (interpolateI i j l) (interpolateI i j m)) + _ = refl + + _ : p (interpolateI i i0 j) ≡ p (i ∧ j) + _ = refl + + _ : p (interpolateI i i1 j) ≡ p (~ i ∨ j) + _ = refl + + _ : p (interpolateI i j i0) ≡ p (~ i ∧ j) + _ = refl + + _ : p (interpolateI i j i1) ≡ p (i ∨ j) + _ = refl diff --git a/Cubical/Functions/FunExtEquiv.agda b/Cubical/Functions/FunExtEquiv.agda index 21a5efdab7..3cf3367d0e 100644 --- a/Cubical/Functions/FunExtEquiv.agda +++ b/Cubical/Functions/FunExtEquiv.agda @@ -117,13 +117,7 @@ funExtDep : {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ₁} → ({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) + transp (λ k → B i (coei→i A i x k)) (i ∨ ~ i) (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} @@ -142,32 +136,53 @@ funExtDepEquiv {A = A} {B} {f} {g} = isoToEquiv isom isom .fun = funExtDep isom .inv = funExtDep⁻ isom .rightInv q m i x = - comp + transp (λ 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 - }) + (m ∨ i ∨ ~ i) (q i (coei→i A i x m)) isom .leftInv h m p i = - comp + transp (λ 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 - }) + (m ∨ i ∨ ~ 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→j j k = coePath A (λ i → p i) i j k 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) + lemi→i m k = coei→i A i (p i) (m ∨ k) + +-- Funext for non-dependent functions but where both domain and +-- codomain depend on the interval. In this case we can omit the +-- outer transp in funExtDep. + +funExtNonDep : {A : I → Type ℓ} {B : 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₁)) + → PathP (λ i → A i → B i) f g +funExtNonDep {A = A} h i x = h (λ j → coei→j A i j x) i + +funExtNonDep⁻ : {A : I → Type ℓ} {B : I → Type ℓ₁} + {f : A i0 → B i0} {g : A i1 → B i1} + → PathP (λ i → A i → B i) f g + → ({x₀ : A i0} {x₁ : A i1} → PathP A x₀ x₁ → PathP B (f x₀) (g x₁)) +funExtNonDep⁻ = funExtDep⁻ + +funExtNonDepEquiv : {A : I → Type ℓ} {B : 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₁)) + ≃ PathP (λ i → A i → B i) f g +funExtNonDepEquiv {A = A} = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = funExtNonDep + isom .inv = funExtNonDep⁻ + isom .rightInv q m i x = q i (coei→i A i x m) + isom .leftInv h m 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 k = coePath A (λ i → p i) i j k heteroHomotopy≃Homotopy : {A : I → Type ℓ} {B : (i : I) → Type ℓ₁} {f : A i0 → B i0} {g : A i1 → B i1} From 60f18987bb1819a15fccc325343ef7b469bb2eeb Mon Sep 17 00:00:00 2001 From: mzeuner <56261690+mzeuner@users.noreply.github.com> Date: Wed, 29 May 2024 16:32:58 +0200 Subject: [PATCH 18/30] add lemma to summary file (#1131) --- .../Functorial/ZFunctors/OpenSubscheme.agda | 4 +++- .../ZariskiLattice/Properties.agda | 17 +++++++++++++++++ Cubical/Papers/FunctorialQcQsSchemes.agda | 16 +++++++++++----- 3 files changed, 31 insertions(+), 6 deletions(-) diff --git a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda index 83dab7931c..31be9b51e5 100644 --- a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda +++ b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda @@ -1,10 +1,12 @@ {- - Compact open subfunctors of qcqs-schemes are qcqs-schemes (TODO!!!) + Compact open subfunctors of affine schemes are qcqs-schemes The proof proceeds by 1. Defining standard/basic compact opens of affines and proving that they are affine 2. Proving that arbitrary compact opens of affines are qcqs-schemes + TODO: prove that compact opens of qcqs-schemes are qcqs-schemes + -} {-# OPTIONS --safe --lossy-unification #-} diff --git a/Cubical/AlgebraicGeometry/ZariskiLattice/Properties.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/Properties.agda index 3fafe859b3..f2dc3d1f52 100644 --- a/Cubical/AlgebraicGeometry/ZariskiLattice/Properties.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/Properties.agda @@ -67,6 +67,23 @@ module _ (R : CommRing ℓ) where 1∈√⟨f⟩ = isEquivRel→effectiveIso ∼PropValued ∼EquivRel _ _ .fun D1≤Df .fst zero + module _ {n : ℕ} (f₁,⋯,fₙ : FinVec (fst R) n) where + private + D[f₁,⋯,fₙ] : ZL + D[f₁,⋯,fₙ] = [ n , f₁,⋯,fₙ ] + + oneIdealLemmaZarLat : D[f₁,⋯,fₙ] ≡ D 1r → 1r ∈ ⟨ f₁,⋯,fₙ ⟩ + oneIdealLemmaZarLat D[f₁,⋯,fₙ]≡D1 = 1∈√→1∈ _ 1∈√⟨f₁,⋯,fₙ⟩ + where + D1≤D[f₁,⋯,fₙ] : D 1r ≤ D[f₁,⋯,fₙ] + D1≤D[f₁,⋯,fₙ] = subst (_≤ D[f₁,⋯,fₙ]) D[f₁,⋯,fₙ]≡D1 (is-refl _) + + 1∈√⟨f₁,⋯,fₙ⟩ : 1r ∈ √ ⟨ f₁,⋯,fₙ ⟩ + 1∈√⟨f₁,⋯,fₙ⟩ = isEquivRel→effectiveIso ∼PropValued ∼EquivRel _ _ .fun + D1≤D[f₁,⋯,fₙ] .fst zero + + + module LocDownSetIso (R : CommRing ℓ) (f : R .fst) where open CommRingStr ⦃...⦄ open DistLatticeStr ⦃...⦄ diff --git a/Cubical/Papers/FunctorialQcQsSchemes.agda b/Cubical/Papers/FunctorialQcQsSchemes.agda index 86c727c96b..6c80e0d120 100644 --- a/Cubical/Papers/FunctorialQcQsSchemes.agda +++ b/Cubical/Papers/FunctorialQcQsSchemes.agda @@ -5,15 +5,15 @@ 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: -The Functor of Points approach to Schemes in Cubical Agda +The Functor of Points Approach to Schemes in Cubical Agda Max Zeuner, Matthias Hutzler -Preprint: TODO ArXiv link - +Preprint: https://arxiv.org/abs/2403.13088 -} + -- The "--safe" flag ensures that there are no postulates or unfinished goals {-# OPTIONS --safe #-} module Cubical.Papers.FunctorialQcQsSchemes where @@ -108,10 +108,16 @@ open ZLUP.IsSupport open ZariskiLatUnivProp using (D ; isSupportD) open ZariskiLatUnivProp using (ZLHasUniversalProp ; ⋁D≡) --- facts about Zariski lattice and localization -open Localization&Radicals using (toUnit) +-- facts about Zariski lattice and localization: +-- g ∈ √⟨f⟩ implies that f/1 is a unit in R[1/g] +open Localization&Radicals using (unitHelper; toUnit) + +-- D(f) = D(1) implies that f is a unit open ZLP using (unitLemmaZarLat) +-- D(f₁,...,fₙ) = D(1) implies 1 ∈ ⟨f₁,...,fₙ⟩ +open ZLP using (oneIdealLemmaZarLat) + ----------- 3: ℤ-functors ---------- From a0b4ba3ed18d16e410d14ab335058509e402cdc2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Na=C3=AFm=20Favier?= Date: Mon, 3 Jun 2024 08:58:17 +0200 Subject: [PATCH 19/30] =?UTF-8?q?Add=20`=CE=A0-contractDom`=20(#1132)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Cubical/Foundations/HLevels.agda | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index e08f7b98b2..0c1675e08a 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -25,6 +25,8 @@ open import Cubical.Foundations.Univalence using (ua ; univalenceIso) open import Cubical.Data.Sigma open import Cubical.Data.Nat using (ℕ; zero; suc; _+_; +-zero; +-comm) +open Iso + HLevel : Type₀ HLevel = ℕ @@ -839,3 +841,13 @@ module _ (B : (i j k : I) → Type ℓ) isGroupoid→CubeP : isGroupoid (B i1 i1 i1) → CubeP isGroupoid→CubeP grpd = isOfHLevelPathP' 0 (isOfHLevelPathP' 1 (isOfHLevelPathP' 2 grpd _ _) _ _) _ _ .fst + + +Π-contractDomIso : (c : isContr A) → Iso ((x : A) → B x) (B (c .fst)) +Π-contractDomIso {B = B} c .fun f = f (c .fst) +Π-contractDomIso {B = B} c .inv b x = subst B (c .snd x) b +Π-contractDomIso {B = B} c .rightInv b i = transp (λ j → B (isProp→isSet (isContr→isProp c) _ _ (c .snd (c .fst)) refl i j)) i b +Π-contractDomIso {B = B} c .leftInv f = funExt λ x → fromPathP (cong f (c .snd x)) + +Π-contractDom : (c : isContr A) → ((x : A) → B x) ≃ B (c .fst) +Π-contractDom c = isoToEquiv (Π-contractDomIso c) From c2cf8563034ae0ab6a1b7335ef0780b210137476 Mon Sep 17 00:00:00 2001 From: Evan Cavallo Date: Tue, 18 Jun 2024 16:14:49 +0200 Subject: [PATCH 20/30] =?UTF-8?q?Simplify=20proof=20that=20cong=20?= =?UTF-8?q?=E2=9F=A8=5F=E2=9F=A9=20is=20injective=20on=20groups=20(#1134)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Cubical/Algebra/Group/Base.agda | 2 ++ Cubical/Algebra/Group/GroupPath.agda | 38 +++++---------------------- Cubical/Algebra/Group/Properties.agda | 10 +++++++ Cubical/Foundations/HLevels.agda | 12 +++++++++ 4 files changed, 30 insertions(+), 32 deletions(-) diff --git a/Cubical/Algebra/Group/Base.agda b/Cubical/Algebra/Group/Base.agda index 3ab0291971..30200188d6 100644 --- a/Cubical/Algebra/Group/Base.agda +++ b/Cubical/Algebra/Group/Base.agda @@ -48,6 +48,8 @@ record GroupStr (G : Type ℓ) : Type ℓ where open IsGroup isGroup public +unquoteDecl GroupStrIsoΣ = declareRecordIsoΣ GroupStrIsoΣ (quote GroupStr) + Group : ∀ ℓ → Type (ℓ-suc ℓ) Group ℓ = TypeWithStr ℓ GroupStr diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index ff7ab15793..b9501ac90b 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -4,7 +4,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.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels @@ -20,6 +20,8 @@ open import Cubical.Displayed.Auto open import Cubical.Displayed.Record open import Cubical.Displayed.Universe +open import Cubical.Functions.Embedding + open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms @@ -134,41 +136,13 @@ module _ (G : Group ℓ) {A : Type ℓ} InducedGroupPathFromPres· : G ≡ InducedGroupFromPres· InducedGroupPathFromPres· = GroupPath _ _ .fst InducedGroupEquivFromPres· - - uaGroup : {G H : Group ℓ} → GroupEquiv G H → G ≡ H uaGroup {G = G} {H = H} = equivFun (GroupPath G H) --- Group-ua functoriality -Group≡ : (G H : Group ℓ) → ( - Σ[ p ∈ ⟨ G ⟩ ≡ ⟨ H ⟩ ] - Σ[ q ∈ PathP (λ i → p i) (1g (snd G)) (1g (snd H)) ] - Σ[ r ∈ PathP (λ i → p i → p i → p i) (_·_ (snd G)) (_·_ (snd H)) ] - Σ[ s ∈ PathP (λ i → p i → p i) (inv (snd G)) (inv (snd H)) ] - PathP (λ i → IsGroup (q i) (r i) (s i)) (isGroup (snd G)) (isGroup (snd H))) - ≃ (G ≡ H) -Group≡ G H = isoToEquiv theIso - where - theIso : Iso _ _ - fun theIso (p , q , r , s , t) i = p i , groupstr (q i) (r i) (s i) (t i) - inv theIso x = cong ⟨_⟩ x , cong (1g ∘ snd) x , cong (_·_ ∘ snd) x , cong (inv ∘ snd) x , cong (isGroup ∘ snd) x - rightInv theIso _ = refl - leftInv theIso _ = refl - caracGroup≡ : {G H : Group ℓ} (p q : G ≡ H) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q -caracGroup≡ {G = G} {H = H} p q P = - sym (transportTransport⁻ (ua (Group≡ G H)) p) - ∙∙ cong (transport (ua (Group≡ G H))) helper - ∙∙ transportTransport⁻ (ua (Group≡ G H)) q - where - helper : transport (sym (ua (Group≡ G H))) p ≡ transport (sym (ua (Group≡ G H))) q - helper = Σ≡Prop - (λ _ → isPropΣ - (isOfHLevelPathP' 1 (is-set (snd H)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd H)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd H)) _ _) - λ _ → isOfHLevelPathP 1 (isPropIsGroup _ _ _) _ _) - (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) +caracGroup≡ _ _ α = + isEmbedding→Inj (iso→isEmbedding (invIso ΣPathIsoPathΣ)) _ _ $ + Σ≡Prop (λ _ → isOfHLevelPathP' 1 (isSetGroupStr _) _ _) α uaGroupId : (G : Group ℓ) → uaGroup (idGroupEquiv {G = G}) ≡ refl uaGroupId G = caracGroup≡ _ _ uaIdEquiv diff --git a/Cubical/Algebra/Group/Properties.agda b/Cubical/Algebra/Group/Properties.agda index 844106cd31..cb544588c7 100644 --- a/Cubical/Algebra/Group/Properties.agda +++ b/Cubical/Algebra/Group/Properties.agda @@ -2,6 +2,7 @@ module Cubical.Algebra.Group.Properties 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.GroupoidLaws hiding (assoc) @@ -27,6 +28,15 @@ isPropIsGroup 1g _·_ inv = where open IsMonoid +isSetGroupStr : (G : Type ℓ) → isSet (GroupStr G) +isSetGroupStr G = + isOfHLevelSucIfInhabited→isOfHLevelSuc 1 λ Gstr → + isOfHLevelRetractFromIso 2 GroupStrIsoΣ $ + isSetΣ (Gstr .GroupStr.is-set) λ _ → + isSetΣ (isSet→ (isSet→ (Gstr .GroupStr.is-set))) λ _ → + isSetΣSndProp (isSet→ (Gstr .GroupStr.is-set)) λ _ → + isPropIsGroup _ _ _ + module GroupTheory (G : Group ℓ) where open GroupStr (snd G) abstract diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index 0c1675e08a..18fe6927fb 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -98,6 +98,18 @@ 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₁) +-- When proving a type has h-level n+1, we can assume it is inhabited. +-- To prove a type is a proposition, it suffices to prove it is contractible if inhabited + +isOfHLevelSucIfInhabited→isOfHLevelSuc : ∀ n + → (A → isOfHLevel (suc n) A) → isOfHLevel (suc n) A +isOfHLevelSucIfInhabited→isOfHLevelSuc zero hA a = hA a a +isOfHLevelSucIfInhabited→isOfHLevelSuc (suc n) hA a = hA a a + +isContrIfInhabited→isProp : (A → isContr A) → isProp A +isContrIfInhabited→isProp hA = + isOfHLevelSucIfInhabited→isOfHLevelSuc 0 (isContr→isProp ∘ hA) + -- hlevel of path types isProp→isContrPath : isProp A → (x y : A) → isContr (x ≡ y) From 29a3aa4cefa9d7f6f26c8d7e76d801bd18069654 Mon Sep 17 00:00:00 2001 From: Evan Cavallo Date: Tue, 18 Jun 2024 17:26:58 +0200 Subject: [PATCH 21/30] add credit to @Schippmunk (#1136) --- Cubical/Displayed/Base.agda | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/Cubical/Displayed/Base.agda b/Cubical/Displayed/Base.agda index e91565a559..316833149c 100644 --- a/Cubical/Displayed/Base.agda +++ b/Cubical/Displayed/Base.agda @@ -1,6 +1,13 @@ {- - Definition of univalent and displayed univalent relations + Definition of univalent and displayed univalent relations. + This approach to structures is based on the master's thesis + + Johannes Schipp von Branitz + “Higher Groups via Displayed Univalent Reflexive Graphs in Cubical Type Theory” + https://jsvb.xyz/files/master.pdf + + with slightly modified definitions. -} {-# OPTIONS --safe #-} @@ -84,4 +91,3 @@ module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} compEquiv (Σ-cong-equiv (ua a a') (λ p → uaᴰ b p b')) ΣPath≃PathΣ - From 7d8f0d46fb953d253777e5046ab56bc9f51105eb Mon Sep 17 00:00:00 2001 From: Evan Cavallo Date: Wed, 19 Jun 2024 12:02:17 +0200 Subject: [PATCH 22/30] Add a generic UARel lemma to prove EquivJ for groups (#1135) --- Cubical/Algebra/Group/GroupPath.agda | 20 +++-------- Cubical/Displayed/Properties.agda | 53 +++++++++++++++------------- 2 files changed, 34 insertions(+), 39 deletions(-) diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index b9501ac90b..04747dc4d0 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -17,6 +17,7 @@ open import Cubical.Data.Sigma open import Cubical.Displayed.Base open import Cubical.Displayed.Auto +open import Cubical.Displayed.Properties open import Cubical.Displayed.Record open import Cubical.Displayed.Universe @@ -158,25 +159,14 @@ uaCompGroupEquiv f g = caracGroup≡ _ _ ( -- J-rule for GroupEquivs GroupEquivJ : {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') - → P G idGroupEquiv - → ∀ {H} e → P H e -GroupEquivJ {G = G} P p {H} e = - transport (λ i → P (GroupPath G H .fst e i) - (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) - (subst (P G) (sym lem) p) - where - lem : transport (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) e - ≡ idGroupEquiv - lem = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (Σ≡Prop (λ _ → isPropIsEquiv _) - (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof - (transportRefl (fst (fst e) (transportRefl x i)) i)))) - ∙ retEq (fst e) x)) + → P G idGroupEquiv + → ∀ {H} e → P H e +GroupEquivJ P p e = 𝒮-J-customRefl≅ (∫ 𝒮ᴰ-Group) P p e 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 +GroupEquivJ>_ {P = P} ids H = GroupEquivJ (λ H e → P H e) ids isGroupoidGroup : ∀ {ℓ} → isGroupoid (Group ℓ) isGroupoidGroup G H = diff --git a/Cubical/Displayed/Properties.agda b/Cubical/Displayed/Properties.agda index 8d59bf8dd1..de7a4f81ff 100644 --- a/Cubical/Displayed/Properties.agda +++ b/Cubical/Displayed/Properties.agda @@ -2,9 +2,10 @@ module Cubical.Displayed.Properties 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.Equiv open import Cubical.Foundations.Univalence using (pathToEquiv; univalence; ua-ungluePath-Equiv) open import Cubical.Data.Unit @@ -23,29 +24,33 @@ private 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 + -- Contractibility of ≅-singletons + + 𝒮-isContrSingl : (a : A) → isContr (Σ[ a' ∈ A ] (a ≅ a')) + 𝒮-isContrSingl a = + isOfHLevelRetractFromIso 0 + (Σ-cong-iso-snd λ _ → uaIso _ _) + (isContrSingl a) + + -- Sometimes we have a hard-coded a ≅ a we want to use as the reflexivity + + 𝒮-J-customRefl≅ : {a : A} {myRefl : a ≅ a} + (P : (a' : A) → (p : a ≅ a') → Type ℓ) + (d : P a myRefl) + {a' : A} + (p : a ≅ a') + → P a' p + 𝒮-J-customRefl≅ P d p = + subst (uncurry P) (isContr→isProp (𝒮-isContrSingl _) _ _) d + + 𝒮-J-≅ : {a : A} + (P : (a' : A) → (p : a ≅ a') → Type ℓ) + (d : P a (ρ a)) + {a' : A} + (p : a ≅ a') + → P a' p + 𝒮-J-≅ = 𝒮-J-customRefl≅ -- constructors @@ -61,7 +66,7 @@ module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} 𝒮ᴰ-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 + = 𝒮-J-≅ 𝒮-A (λ y q → (b' : B y) → (b ≅ᴰ⟨ q ⟩ b') ≃ PathP (λ i → B (≅→≡ q i)) b b') (λ b' → uni' b') p From ac1bf363a9d151115ae0cbf4d8de82b54740a487 Mon Sep 17 00:00:00 2001 From: Andrew Swan Date: Mon, 24 Jun 2024 09:25:49 +0200 Subject: [PATCH 23/30] Another minor universe level generalisation (#1138) * minor universe level generalisation * fix problems introduced by universe generalisation * better fix for files broken by universe level generalization --- Cubical/Data/FinSet/FiniteChoice.agda | 5 ++--- Cubical/Data/FinSet/Properties.agda | 2 +- Cubical/Data/Unit/Properties.agda | 2 +- 3 files changed, 4 insertions(+), 5 deletions(-) diff --git a/Cubical/Data/FinSet/FiniteChoice.agda b/Cubical/Data/FinSet/FiniteChoice.agda index a02940c023..8270a45491 100644 --- a/Cubical/Data/FinSet/FiniteChoice.agda +++ b/Cubical/Data/FinSet/FiniteChoice.agda @@ -32,9 +32,8 @@ choice≃Fin : {n : ℕ}(Y : Fin n → Type ℓ) → ((x : Fin n) → ∥ Y x ∥₁) ≃ ∥ ((x : Fin n) → Y x) ∥₁ choice≃Fin {n = 0} Y = isContr→≃Unit (isContrΠ⊥) - ⋆ Unit≃Unit* - ⋆ invEquiv (propTruncIdempotent≃ isPropUnit*) - ⋆ propTrunc≃ (invEquiv (isContr→≃Unit* (isContrΠ⊥ {A = Y}))) + ⋆ invEquiv (propTruncIdempotent≃ isPropUnit) + ⋆ propTrunc≃ (invEquiv (isContr→≃Unit (isContrΠ⊥ {A = Y}))) choice≃Fin {n = suc n} Y = Π⊎≃ ⋆ Σ-cong-equiv-fst (ΠUnit (λ x → ∥ Y (inl x) ∥₁)) diff --git a/Cubical/Data/FinSet/Properties.agda b/Cubical/Data/FinSet/Properties.agda index d547146053..34d927ad99 100644 --- a/Cubical/Data/FinSet/Properties.agda +++ b/Cubical/Data/FinSet/Properties.agda @@ -65,7 +65,7 @@ isFinSet→Discrete : isFinSet A → Discrete A isFinSet→Discrete h = Prop.rec isPropDiscrete (λ p → EquivPresDiscrete (invEquiv p) discreteFin) (h .snd) isContr→isFinSet : isContr A → isFinSet A -isContr→isFinSet h = 1 , ∣ isContr→≃Unit* h ⋆ invEquiv Unit≃Unit* ⋆ invEquiv Fin1≃Unit ∣₁ +isContr→isFinSet h = 1 , ∣ isContr→≃Unit h ⋆ invEquiv Fin1≃Unit ∣₁ isDecProp→isFinSet : isProp A → Dec A → isFinSet A isDecProp→isFinSet h (yes p) = isContr→isFinSet (inhProp→isContr p h) diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index d219f49696..99b240fc4a 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -113,7 +113,7 @@ isOfHLevelUnit* (suc (suc (suc n))) = isOfHLevelPlus 3 (isOfHLevelUnit* n) Unit≃Unit* : ∀ {ℓ} → Unit ≃ Unit* {ℓ} Unit≃Unit* = invEquiv (isContr→≃Unit isContrUnit*) -isContr→≃Unit* : {A : Type ℓ} → isContr A → A ≃ Unit* {ℓ} +isContr→≃Unit* : {A : Type ℓ} → isContr A → A ≃ Unit* {ℓ'} isContr→≃Unit* contr = compEquiv (isContr→≃Unit contr) Unit≃Unit* isContr→≡Unit* : {A : Type ℓ} → isContr A → A ≡ Unit* From 6d6870cbf537294fc3fdf9c5a34909bde30ba914 Mon Sep 17 00:00:00 2001 From: Stefania Damato Date: Tue, 25 Jun 2024 15:43:01 +0200 Subject: [PATCH 24/30] Containers (and W and M) (#1129) * Added W and M * Fixed flags * Added containers proofs * Fixed flags * Refactors to follow conventions * More refactoring * Changes due to --guardedness flag * Moved CoinductiveContainers to Codata * Edited comment * Cleaning with Anders * Added --safe * More --safe flags --------- Co-authored-by: Stefania Damato Co-authored-by: Stefania Damato --- Cubical/Codata/Containers/Coalgebras.agda | 28 +++ .../Containers/CoinductiveContainers.agda | 199 ++++++++++++++++++ Cubical/Codata/Everything.agda | 3 + Cubical/Codata/M/MRecord.agda | 35 +++ Cubical/Data/Containers/Algebras.agda | 47 +++++ Cubical/Data/Containers/Base.agda | 134 ++++++++++++ .../Containers/ContainerExtensionProofs.agda | 106 ++++++++++ .../Data/Containers/InductiveContainers.agda | 72 +++++++ Cubical/Data/W/W.agda | 17 ++ 9 files changed, 641 insertions(+) create mode 100644 Cubical/Codata/Containers/Coalgebras.agda create mode 100644 Cubical/Codata/Containers/CoinductiveContainers.agda create mode 100644 Cubical/Codata/M/MRecord.agda create mode 100644 Cubical/Data/Containers/Algebras.agda create mode 100644 Cubical/Data/Containers/Base.agda create mode 100644 Cubical/Data/Containers/ContainerExtensionProofs.agda create mode 100644 Cubical/Data/Containers/InductiveContainers.agda create mode 100644 Cubical/Data/W/W.agda diff --git a/Cubical/Codata/Containers/Coalgebras.agda b/Cubical/Codata/Containers/Coalgebras.agda new file mode 100644 index 0000000000..ecaf72b582 --- /dev/null +++ b/Cubical/Codata/Containers/Coalgebras.agda @@ -0,0 +1,28 @@ +{-# OPTIONS --safe --guardedness #-} + +module Cubical.Codata.Containers.Coalgebras where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open import Cubical.Codata.M.MRecord +open import Cubical.Data.Containers.Algebras + +private + variable + ℓ ℓ' : Level + +module Coalgs (S : Type ℓ) (Q : S → Type ℓ') where + open Algs S Q + open Iso + open M + + MAlg : ContFuncIso + MAlg = iso (M S Q) isom + where + isom : Iso (Σ[ s ∈ S ] (Q s → M S Q)) (M S Q) + fun isom = uncurry sup-M + inv isom m = shape m , pos m + rightInv isom m = ηEqM m + leftInv isom (s , f) = refl diff --git a/Cubical/Codata/Containers/CoinductiveContainers.agda b/Cubical/Codata/Containers/CoinductiveContainers.agda new file mode 100644 index 0000000000..5e221d43da --- /dev/null +++ b/Cubical/Codata/Containers/CoinductiveContainers.agda @@ -0,0 +1,199 @@ +{- Proof that containers are closed under greatest fixed points + +Adapted from 'Containers: Constructing strictly positive types' +by Abbott, Altenkirch, Ghani + +-} + +{-# OPTIONS --safe --guardedness #-} + +open import Cubical.Codata.M.MRecord +open import Cubical.Data.Sigma +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Codata.Containers.Coalgebras +open import Cubical.Data.Containers.Algebras +open M +open M-R + +module Cubical.Codata.Containers.CoinductiveContainers + (Ind : Type) + (S : Type) + (setS : isSet S) + (P : Ind → S → Type) + (Q : S → Type) + (setM : isSet S → isSet (M S Q)) + (X : Ind → Type) + (Y : Type) + (βs : Y → S) + (βg : (y : Y) → (i : Ind) → P i (βs y) → X i) + (βh : (y : Y) → Q (βs y) → Y) where + + open Algs S Q + open Coalgs S Q + + β̅₁ : Y → M S Q + shape (β̅₁ y) = βs y + pos (β̅₁ y) = β̅₁ ∘ βh y + + β̅₂ : (y : Y) (ind : Ind) → Pos P MAlg ind (β̅₁ y) → X ind + β̅₂ y ind (here p) = βg y ind p + β̅₂ y ind (below q p) = β̅₂ (βh y q) ind p + + β̅ : Y → Σ[ m ∈ M S Q ] ((i : Ind) → Pos P MAlg i m → X i) + β̅ y = β̅₁ y , β̅₂ y + + out : Σ[ m ∈ M S Q ] ((i : Ind) → Pos P MAlg i m → X i) → + Σ[ (s , f) ∈ Σ[ s ∈ S ] (Q s → M S Q) ] + (((i : Ind) → P i s → X i) × + ((i : Ind) (q : Q s) → Pos P MAlg i (f q) → X i)) + out (m , k) = (shape m , pos m) , ((λ i p → k i (here p)) , (λ i q p → k i (below q p))) + + module _ (β̃₁ : Y → M S Q) + (β̃₂ : (y : Y) (ind : Ind) → Pos P MAlg ind (β̃₁ y) → X ind) + (comm : (y : Y) → + out (β̃₁ y , β̃₂ y) ≡ + ((βs y , λ q → (β̃₁ (βh y q))) , (βg y , λ i q → (β̃₂ (βh y q)) i))) where + + -- Diagram commutes + β̅Comm : (y : Y) → out (β̅ y) ≡ ((βs y , β̅₁ ∘ (βh y)) , (βg y , λ i q → β̅₂ (βh y q) i)) + β̅Comm y = refl + + β̃ : Y → Σ (M S Q) (λ m → (i : Ind) → Pos P MAlg i m → X i) + β̃ y = β̃₁ y , β̃₂ y + + ---------- + + comm1 : (y : Y) → shape (β̃₁ y) ≡ shape (β̅₁ y) + comm1 y i = fst (fst (comm y i)) + + comm2 : (y : Y) → + PathP (λ i → Q (comm1 y i) → M S Q) + (pos (β̃₁ y)) (λ q → β̃₁ (βh y q)) + comm2 y i = snd (fst (comm y i)) + + comm3 : (y : Y) → PathP (λ i → (ind : Ind) → P ind (comm1 y i) → X ind) + (λ ind p → β̃₂ y ind (here p)) + (βg y) + comm3 y i = fst (snd (comm y i)) + + comm4 : (y : Y) → PathP (λ i → (ind : Ind) → (q : Q (comm1 y i)) → + Pos P MAlg ind (comm2 y i q) → X ind) + (λ ind q b → β̃₂ y ind (below q b)) + (λ ind q b → β̃₂ (βh y q) ind b) + comm4 y i = snd (snd (comm y i)) + + data R : M S Q → M S Q → Type where + R-intro : (y : Y) → R (β̃₁ y) (β̅₁ y) + + isBisimR : {m₀ : M S Q} {m₁ : M S Q} → R m₀ m₁ → M-R R m₀ m₁ + s-eq (isBisimR (R-intro y)) = comm1 y + p-eq (isBisimR (R-intro y)) q₀ q₁ q-eq = + transport (λ i → R (comm2 y (~ i) (q-eq (~ i))) (β̅₁ (βh y q₁))) (R-intro (βh y q₁)) + + fstEq : (y : Y) → β̃₁ y ≡ β̅₁ y + fstEq y = MCoind {S} {Q} R isBisimR (R-intro y) + where + -- Coinduction principle for M + MCoind : {S : Type} {Q : S → Type} (R : M S Q → M S Q → Type) + (is-bisim : {m₀ m₁ : M S Q} → R m₀ m₁ → M-R R m₀ m₁) + {m₀ m₁ : M S Q} → R m₀ m₁ → m₀ ≡ m₁ + shape (MCoind R is-bisim r i) = s-eq (is-bisim r) i + pos (MCoind {S} {Q} R is-bisim {m₀ = m₀}{m₁ = m₁} r i) q = + MCoind R is-bisim {m₀ = pos m₀ q₀} {m₁ = pos m₁ q₁} (p-eq (is-bisim r) q₀ q₁ q₂) i + where QQ : I → Type + QQ i = Q (s-eq (is-bisim r) i) + + q₀ : QQ i0 + q₀ = transp (λ j → QQ (~ j ∧ i)) (~ i) q + + q₁ : QQ i1 + q₁ = transp (λ j → QQ (j ∨ i)) i q + + q₂ : PathP (λ i → QQ i) q₀ q₁ + q₂ k = transp (λ j → QQ ((~ k ∧ ~ j ∧ i) ∨ (k ∧ (j ∨ i)) ∨ + ((~ j ∧ i) ∧ (j ∨ i)))) ((~ k ∧ ~ i) ∨ (k ∧ i)) q + + sndEqGen : (y : Y) (β̃₁ : Y → M S Q) (p : β̅₁ ≡ β̃₁) + (β̃₂ : (y : Y) (ind : Ind) → Pos P MAlg ind (β̃₁ y) → X ind) + (comm1 : shape ∘ β̃₁ ≡ βs) + (comm2 : (y : Y) → PathP (λ i → Q (comm1 i y) → M S Q) + (pos (β̃₁ y)) (β̃₁ ∘ βh y)) + (comm3 : (y : Y) → PathP (λ i → (ind : Ind) → P ind (comm1 i y) → X ind) + (λ ind p → β̃₂ y ind (here p)) + (βg y)) + (comm4 : (y : Y) → PathP (λ i → (ind : Ind) (q : Q (comm1 i y)) → + Pos P MAlg ind (comm2 y i q) → X ind) + (λ ind q b → β̃₂ y ind (below q b)) + λ ind q b → β̃₂ (βh y q) ind b) → + PathP (λ i → (ind : Ind) → Pos P MAlg ind (p i y) → X ind) + (β̅₂ y) (β̃₂ y) + sndEqGen y = + J>_ -- we're applying J to p : makeβ̅₁ βs βg βh ≡ β̅₁ + {P = λ β̃₁ p → + (β̃₂ : (y : Y) (ind : Ind) → Pos P MAlg ind (β̃₁ y) → X ind) + (comm1 : shape ∘ β̃₁ ≡ βs) + (comm2 : (y : Y) → PathP (λ i → Q (comm1 i y) → M S Q) (pos (β̃₁ y)) (β̃₁ ∘ βh y)) + (comm3 : (y : Y) → PathP (λ i → (ind : Ind) → P ind (comm1 i y) → X ind) + (λ ind p → β̃₂ y ind (here p)) + (βg y)) + (comm4 : (y : Y) → PathP (λ i → (ind : Ind) (q : Q (comm1 i y)) → Pos P MAlg ind + (comm2 y i q) → X ind) + (λ ind q b → β̃₂ y ind (below q b)) + λ ind q b → β̃₂ (βh y q) ind b) → + PathP (λ i → (ind : Ind) → Pos P MAlg ind (p i y) → X ind) + (β̅₂ y) + (β̃₂ y)} + λ β̃₂ comm1 → + propElim -- S is a set so equality on S is a prop + {A = (λ y → βs y) ≡ βs} + (isSetΠ (λ _ → setS) (λ y → βs y) βs) + (λ s-eq → + (comm2 : (y : Y) → PathP (λ i → Q (s-eq i y) → M S Q) + (β̅₁ ∘ βh y) (β̅₁ ∘ βh y)) + (comm3 : (y : Y) → PathP (λ i → (ind : Ind) → P ind (s-eq i y) → X ind) + (λ ind p → β̃₂ y ind (here p)) (βg y)) + (comm4 : (y : Y) → PathP (λ i → (ind : Ind) (q : Q (s-eq i y)) → Pos P MAlg ind + (comm2 y i q) → X ind) + (λ ind q b → β̃₂ y ind (below q b)) + (λ ind q b → β̃₂ (βh y q) ind b)) → + (β̅₂ y) ≡ (β̃₂ y)) + refl + (propElim -- M is a set so equality on M is a prop + {A = (y : Y) → + (λ x → β̅₁ (βh y x)) ≡ (λ x → β̅₁ (βh y x))} + (isPropΠ λ y' → isSetΠ (λ _ → setM setS) (β̅₁ ∘ βh y') (β̅₁ ∘ βh y')) + (λ m-eq → + (comm3 : (y : Y) → (λ ind p → β̃₂ y ind (here p)) ≡ (βg y)) + (comm4 : (y : Y) → PathP (λ i → (ind : Ind) (q : Q (βs y)) → Pos P MAlg ind + (m-eq y i q) → X ind) + (λ ind q b → β̃₂ y ind (below q b)) + (λ ind q b → β̃₂ (βh y q) ind b)) → + (β̅₂ y) ≡ (β̃₂ y)) + (λ _ → refl) + λ comm3 comm4 → funExt (λ ind → funExt (sndEqAux β̃₂ comm3 comm4 y ind))) + comm1 + where + -- Proposition elimination + propElim : ∀ {ℓ ℓ'} {A : Type ℓ} (t : isProp A) → (D : A → Type ℓ') → + (x : A) → D x → (a : A) → D a + propElim t D x pr a = subst D (t x a) pr + + sndEqAux : (β̃₂ : (s : Y) (i : Ind) → Pos P MAlg i (β̅₁ s) → X i) + (c3 : (s : Y) → (λ ind p → β̃₂ s ind (here p)) ≡ βg s) + (c4 : (s : Y) → (λ ind q b → β̃₂ s ind (below q b)) ≡ + (λ ind q → β̃₂ (βh s q) ind)) + (y : Y) (ind : Ind) (pos : Pos P MAlg ind (β̅₁ y)) → β̅₂ y ind pos ≡ β̃₂ y ind pos + sndEqAux β̃₂ c3 c4 y ind (here x) = sym (funExt⁻ (funExt⁻ (c3 y) ind) x) + sndEqAux β̃₂ c3 c4 y ind (below q x) = + sndEqAux β̃₂ c3 c4 (βh y q) ind x ∙ funExt⁻ (funExt⁻ (sym (funExt⁻ (c4 y) ind)) q) x + + sndEq : (y : Y) → PathP (λ i → (ind : Ind) → Pos P MAlg ind (fstEq y i) → X ind) (β̃₂ y) (β̅₂ y) + sndEq y i = sndEqGen y β̃₁ (sym (funExt fstEq)) β̃₂ (funExt comm1) comm2 comm3 comm4 (~ i) + + -- β̅ is unique + β̅Unique : β̃ ≡ β̅ + fst (β̅Unique i y) = fstEq y i + snd (β̅Unique i y) = sndEq y i diff --git a/Cubical/Codata/Everything.agda b/Cubical/Codata/Everything.agda index 163e335490..3ab1d38cc2 100644 --- a/Cubical/Codata/Everything.agda +++ b/Cubical/Codata/Everything.agda @@ -6,6 +6,9 @@ import Cubical.Codata.Conat.Bounded import Cubical.Codata.M.Coalg import Cubical.Codata.M.Coalg.Base import Cubical.Codata.M.Container +import Cubical.Codata.Containers.Coalgebras +import Cubical.Codata.Containers.CoinductiveContainers +import Cubical.Codata.M.MRecord import Cubical.Codata.M.M import Cubical.Codata.M.M.Base import Cubical.Codata.M.M.Properties diff --git a/Cubical/Codata/M/MRecord.agda b/Cubical/Codata/M/MRecord.agda new file mode 100644 index 0000000000..48fb15f210 --- /dev/null +++ b/Cubical/Codata/M/MRecord.agda @@ -0,0 +1,35 @@ +{- Alternative definition of M as a record type, together with some of its properties + +-} + +{-# OPTIONS --safe --guardedness #-} + +module Cubical.Codata.M.MRecord where + +open import Cubical.Foundations.Prelude + +private + variable + ℓ ℓ' ℓ'' : Level + +record M (S : Type ℓ) (P : S → Type ℓ') : Type (ℓ-max ℓ ℓ') where + constructor sup-M + coinductive + field + shape : S + pos : P shape → M S P +open M + +-- Lifting M to relations +record M-R {S : Type ℓ} {Q : S → Type ℓ'} (R : M S Q → M S Q → Type ℓ'') + (m₀ m₁ : M S Q) : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) where + field + s-eq : shape m₀ ≡ shape m₁ + p-eq : (q₀ : Q (shape m₀)) (q₁ : Q (shape m₁)) + (q-eq : PathP (λ i → Q (s-eq i)) q₀ q₁) → R (pos m₀ q₀) (pos m₁ q₁) +open M-R + +-- (Propositional) η-equality for M +ηEqM : {S' : Type ℓ} {Q' : S' → Type ℓ'} (m : M S' Q') → sup-M (shape m) (pos m) ≡ m +shape (ηEqM m i) = shape m +pos (ηEqM m i) = pos m diff --git a/Cubical/Data/Containers/Algebras.agda b/Cubical/Data/Containers/Algebras.agda new file mode 100644 index 0000000000..69ad7060d0 --- /dev/null +++ b/Cubical/Data/Containers/Algebras.agda @@ -0,0 +1,47 @@ +{- Basic definitions required for co/inductive container proofs + +- Definition of Pos + +-} + +{-# OPTIONS --safe #-} + +module Cubical.Data.Containers.Algebras where + +open import Cubical.Data.W.W +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Prelude + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + +module Algs (S : Type ℓ) + (Q : S → Type ℓ') where + + open Iso + + -- Fixed point algebras + record ContFuncIso : Type (ℓ-max (ℓ-suc ℓ'') (ℓ-max ℓ ℓ')) where + constructor iso + field + carrier : Type ℓ'' + χ : Iso (Σ[ s ∈ S ] (Q s → carrier)) carrier + + open ContFuncIso + + WAlg : ContFuncIso + WAlg = iso (W S Q) isom + where + isom : Iso (Σ[ s ∈ S ] (Q s → W S Q)) (W S Q) + fun isom = uncurry sup-W + inv isom (sup-W s f) = s , f + rightInv isom (sup-W s f) = refl + leftInv isom (s , f) = refl + + data Pos {Ind : Type ℓ'''} (P : Ind → S → Type ℓ'') (FP : ContFuncIso {ℓ}) (i : Ind) : + carrier FP → Type (ℓ-max (ℓ-suc ℓ) (ℓ-max ℓ'' ℓ')) where + here : {wm : carrier FP} → P i (fst (FP .χ .inv wm)) → Pos P FP i wm + below : {wm : carrier FP} → (q : Q (fst (FP .χ .inv wm))) → + Pos P FP i (snd (FP .χ .inv wm) q) → Pos P FP i wm diff --git a/Cubical/Data/Containers/Base.agda b/Cubical/Data/Containers/Base.agda new file mode 100644 index 0000000000..0f79e4a38b --- /dev/null +++ b/Cubical/Data/Containers/Base.agda @@ -0,0 +1,134 @@ +{- Basic container definitions: + +- Definition of generalised container parameterised by category C + +- Category Cont whose objects are generalised containers + +- Functor ⟦_⟧ : Cont → Functor C Set + +- Example of List as a generalised container + +-} + +{-# OPTIONS --safe #-} + +module Cubical.Data.Containers.Base where + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.NaturalTransformation hiding (_⟦_⟧) +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude hiding (_◁_) + +private + variable + ℓ ℓ' : Level + +open Category +open Functor +open NatTrans + +-- Definition of generalised container +record GenContainer (C : Category ℓ ℓ') : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + constructor _◁_&_ + field + S : Type ℓ' + P : S → C .ob + isSetS : isSet S + +open GenContainer + +module Conts (C : Category ℓ ℓ') where + + -- Category Cont with objects of type GenContainer C + record _⇒c_ (C₁ C₂ : GenContainer C) : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + constructor _◁_ + field + shape : S C₁ → S C₂ + pos : (s : S C₁) → C [ P C₂ (shape s) , P C₁ s ] + + open _⇒c_ + + id-c : {Con : GenContainer C} → Con ⇒c Con + id-c = (λ s → s) ◁ λ _ → C .id + + _⋆c_ : {C₁ C₂ C₃ : GenContainer C} → C₁ ⇒c C₂ → C₂ ⇒c C₃ → C₁ ⇒c C₃ + _⋆c_ (u ◁ f) (v ◁ g) = (λ a → v (u a)) ◁ λ a → (g (u a)) ⋆⟨ C ⟩ (f a) + + assoc-c : {C₁ C₂ C₃ C₄ : GenContainer C} (f : C₁ ⇒c C₂) (g : C₂ ⇒c C₃) (h : C₃ ⇒c C₄) → + (f ⋆c g) ⋆c h ≡ f ⋆c (g ⋆c h) + assoc-c (u ◁ f) (v ◁ g) (w ◁ h) i = (λ a → w (v (u a))) ◁ λ a → C .⋆Assoc (h (v (u a))) (g (u a)) (f a) (~ i) + + isSet⇒c : {C₁ C₂ : GenContainer C} → isSet (C₁ ⇒c C₂) + shape (isSet⇒c {A ◁ B & set-A} {E ◁ F & set-C} m n p q i j) a = + set-C (shape m a) (shape n a) (λ k → shape (p k) a) (λ k → shape (q k) a) i j + pos (isSet⇒c {A ◁ B & set-A} {E ◁ F & set-C} m n p q i j) a = + isSet→SquareP + {A = λ i j → C [ (F (set-C (shape m a) (shape n a) (λ k → shape (p k) a) (λ k → shape (q k) a) i j)) , B a ]} + (λ i j → C .isSetHom {F (set-C (shape m a) (shape n a) (λ k → shape (p k) a) (λ k → shape (q k) a) i j)} {B a}) + (λ k → pos (p k) a) + (λ k → pos (q k) a) + (λ _ → pos m a) + (λ _ → pos n a) + i j + + Cont : Category (ℓ-suc (ℓ-max ℓ ℓ')) (ℓ-suc (ℓ-max ℓ ℓ')) + ob Cont = GenContainer C + Hom[_,_] Cont = _⇒c_ + id Cont = id-c + _⋆_ Cont = _⋆c_ + ⋆IdL Cont m i = (shape m) ◁ (λ a → C .⋆IdR (pos m a) i) + ⋆IdR Cont m i = (shape m) ◁ (λ a → C .⋆IdL (pos m a) i) + ⋆Assoc Cont = assoc-c + isSetHom Cont = isSet⇒c + + -- Definition of functor ⟦_⟧ : Cont → Functor C Set + + -- Type alias for (S : Type) (P : S → C .ob) (X : C .ob) ↦ Σ S (λ s → C [ P s , X ]) + cont-func : (S : Type ℓ') (P : S → C .ob) (X : C .ob) → Type ℓ' + cont-func S P X = Σ S (λ s → C [ P s , X ]) + + isSetContFunc : (A : Type ℓ') (B : A → C .ob) (X : C .ob) (isSetA : isSet A) → isSet (cont-func A B X) + isSetContFunc A B X setA = isSetΣ setA (λ _ → C .isSetHom) + + cont-mor : {A : Type ℓ'} {B : A → C .ob} {X Y : C .ob} (f : C [ X , Y ]) → + cont-func A B X → cont-func A B Y + cont-mor f (s , g) = s , (g ⋆⟨ C ⟩ f) + + ⟦_⟧-obj : GenContainer C → Functor C (SET ℓ') + F-ob ⟦ A ◁ B & set-A ⟧-obj X = cont-func A B X , isSetContFunc A B X set-A + F-hom ⟦ A ◁ B & set-A ⟧-obj = cont-mor + F-id ⟦ A ◁ B & set-A ⟧-obj = funExt λ {(a , b) i → a , C .⋆IdR b i} + F-seq ⟦ A ◁ B & set-A ⟧-obj f g i (a , b) = a , C .⋆Assoc b f g (~ i) + + ⟦_⟧-mor : {C₁ C₂ : GenContainer C} → C₁ ⇒c C₂ → NatTrans ⟦ C₁ ⟧-obj ⟦ C₂ ⟧-obj + N-ob (⟦_⟧-mor (u ◁ f)) X (s , p) = u s , ((f s) ⋆⟨ C ⟩ p) + N-hom (⟦_⟧-mor (u ◁ f)) xy i (a , b) = u a , C .⋆Assoc (f a) b xy (~ i) + + ⟦_⟧-id : {C₁ : GenContainer C} → ⟦_⟧-mor {C₁} {C₁} (id-c {C₁}) ≡ idTrans ⟦ C₁ ⟧-obj + ⟦_⟧-id = makeNatTransPath λ i X (s , p) → s , C .⋆IdL p i + + ⟦_⟧-comp : {U V W : GenContainer C} (g : U ⇒c V) (h : V ⇒c W) → + ⟦ g ⋆c h ⟧-mor ≡ ⟦ g ⟧-mor ⋆⟨ FUNCTOR C (SET ℓ') ⟩ ⟦ h ⟧-mor + ⟦_⟧-comp (ug ◁ fg) (uh ◁ fh) = + makeNatTransPath λ i A (s , p) → uh (ug s) , C .⋆Assoc (fh (ug s)) (fg s) p i + + ⟦_⟧ : Functor Cont (FUNCTOR C (SET ℓ')) + F-ob ⟦_⟧ = ⟦_⟧-obj + F-hom ⟦_⟧ = ⟦_⟧-mor + F-id ⟦_⟧ = ⟦_⟧-id + F-seq ⟦_⟧ = ⟦_⟧-comp + +module Example where + open Conts (SET ℓ-zero) + + open import Cubical.Data.Fin + open import Cubical.Data.Nat + + ListC : GenContainer (SET ℓ-zero) + ListC = ℕ ◁ (λ n → Fin n , isSetFin) & isSetℕ + + ListF : Functor (SET ℓ-zero) (SET ℓ-zero) + ListF = ⟦ ListC ⟧-obj diff --git a/Cubical/Data/Containers/ContainerExtensionProofs.agda b/Cubical/Data/Containers/ContainerExtensionProofs.agda new file mode 100644 index 0000000000..4445bd85e8 --- /dev/null +++ b/Cubical/Data/Containers/ContainerExtensionProofs.agda @@ -0,0 +1,106 @@ +{- 2 presentations of the proof that the functor ⟦_⟧ : Cont → Func + is full and faithful + +- First one is adapted from 'Containers: Constructing strictly positive types' + by Abbott, Altenkirch, Ghani + +- Second one uses the Yoneda lemma + +-} + +{-# OPTIONS --safe #-} + +module Cubical.Data.Containers.ContainerExtensionProofs where + +open import Cubical.Foundations.Prelude hiding (_◁_) +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation hiding (_⟦_⟧) +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.Yoneda + +open import Cubical.Data.Containers.Base +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ' : Level + +open Category hiding (_∘_) +open Functor +open NatTrans +open Iso + +module _ {C : Category ℓ ℓ'} where + + open Conts {ℓ} {ℓ'} C + + -- Proof 1 that the functor ⟦_⟧ is full and faithful + -- Adapted from 'Containers: Constructing strictly positive types' + + ⟦_⟧FullyFaithful : isFullyFaithful ⟦_⟧ + equiv-proof (⟦_⟧FullyFaithful (S ◁ P & set-S) (T ◁ Q & set-T)) (natTrans mors nat) = + (fib (natTrans mors nat) , fib-pf) , unique + where + fib : NatTrans ⟦ (S ◁ P & set-S) ⟧-obj ⟦ (T ◁ Q & set-T) ⟧-obj → + (S ◁ P & set-S) ⇒c (T ◁ Q & set-T) + fib (natTrans mors _) = (fst ∘ tq) ◁ (snd ∘ tq) + where + tq : (s : S) → cont-func T Q (P s) + tq s = mors (P s) (s , C .id {P s}) + + fib-pf : ⟦ fib (natTrans mors nat) ⟧-mor ≡ (natTrans mors nat) + fib-pf = makeNatTransPath + (funExt₂ λ X (s , h) → sym (funExt⁻ (nat h) (s , C .id)) ∙ + cong (λ Z → mors X (s , Z)) (C .⋆IdL h)) + + ret : ∀ X→Y → fib (⟦ X→Y ⟧-mor) ≡ X→Y + ret (u ◁ f) i = u ◁ λ s → C .⋆IdR (f s) i + + unique : (y : fiber (⟦_⟧-mor) (natTrans mors nat)) → (fib (natTrans mors nat) , fib-pf) ≡ y + unique (m , m-eq) = Σ≡Prop (λ _ → isSetNatTrans _ _) (cong fib (sym m-eq) ∙ ret m) + + -- Proof 2 that the functor ⟦_⟧ is full and faithful + -- Uses the Yoneda lemma + + ⟦_⟧FullyFaithful' : isFullyFaithful ⟦_⟧ + equiv-proof (⟦_⟧FullyFaithful' (S ◁ P & set-S) (T ◁ Q & set-T)) (natTrans mors nat) = + (mor , mor-eq) , unique + where + -- Compose heterogenous with homogenous equality + compHetHomP : {A : I → Type ℓ'} {x : A i0} {y : A i1} {z : A i1} → + PathP (λ i → A i) x y → y ≡ z → PathP (λ i → A i) x z + compHetHomP {A} {x} {y} {z} p eq i = subst (λ c → PathP A x c) eq p i + + nat-trans : (s : S) → FUNCTOR C (SET ℓ') [ C [ P s ,-] , ⟦ T ◁ Q & set-T ⟧-obj ] + N-ob (nat-trans s) X = curry (mors X) s + N-hom (nat-trans s) X→Y i Ps→X = nat X→Y i (s , Ps→X) + + apply-yo : (s : S) → cont-func T Q (P s) + apply-yo s = yoneda ⟦ T ◁ Q & set-T ⟧-obj (P s) .fun (nat-trans s) + + apply-Σ-Π-Iso : Σ (S → T) (λ f → (s : S) → C [ Q (f s) , P s ]) + apply-Σ-Π-Iso = Σ-Π-Iso .fun apply-yo + + mor : (S ◁ P & set-S) ⇒c (T ◁ Q & set-T) + mor = fst apply-Σ-Π-Iso ◁ snd apply-Σ-Π-Iso + + mor-eq : ⟦ mor ⟧-mor ≡ natTrans mors nat + mor-eq = makeNatTransPath + (funExt₂ λ X (s , f) → sym (funExt⁻ (nat f) (s , C .id)) ∙ + cong (λ Z → mors X (s , Z)) (C .⋆IdL f)) + + unique : (y : fiber (⟦_⟧-mor) (natTrans mors nat)) → (mor , mor-eq) ≡ y + unique ((u ◁ f) , m-eq) = Σ≡Prop (λ _ → isSetNatTrans _ _) + λ i → (λ s → fst (N-ob (m-eq (~ i)) (P s) (s , C .id))) ◁ + λ s → compHetHomP + (λ j → snd (N-ob (m-eq (~ j)) (P s) (s , C .id))) + (C .⋆IdR (f s)) i diff --git a/Cubical/Data/Containers/InductiveContainers.agda b/Cubical/Data/Containers/InductiveContainers.agda new file mode 100644 index 0000000000..a744636b02 --- /dev/null +++ b/Cubical/Data/Containers/InductiveContainers.agda @@ -0,0 +1,72 @@ +{- Proof that containers are closed over least fixed points + +Adapted from 'Containers: Constructing strictly positive types' +by Abbott, Altenkirch, Ghani + +-} + +{-# OPTIONS --safe #-} + +open import Cubical.Data.W.W +open import Cubical.Data.Containers.Algebras +open import Cubical.Data.Sigma +open import Cubical.Foundations.Prelude + +module Cubical.Data.Containers.InductiveContainers + (Ind : Type) + (S : Type) + (P : Ind → S → Type) + (Q : S → Type) + (X : Ind → Type) + (Y : Type) + (α : Σ S (λ s → Σ ((i : Ind) → P i s → X i) (λ _ → Q s → Y)) → Y) where + + open Algs S Q + + into : Σ[ (s , f) ∈ Σ[ s ∈ S ] (Q s → W S Q) ] + (((i : Ind) → P i s → X i) × + ((i : Ind) (q : Q s) → Pos P WAlg i (f q) → X i)) → + Σ[ w ∈ W S Q ] ((i : Ind) → Pos P WAlg i w → X i) + into ((s , f) , (g , h)) = sup-W s f , λ i → λ {(here p) → g i p ; (below q b) → h i q b} + + α̅' : (w : W S Q) → ((i : Ind) → Pos P WAlg i w → X i) → Y + α̅' (sup-W s f) k = α (s , g , λ q → α̅' (f q) (λ i → h i q)) + where + g : (i : Ind) → P i s → X i + g i p = k i (here p) + + h : (i : Ind) → (q : Q s) → Pos P WAlg i (f q) → X i + h i q b = k i (below q b) + + α̅ : Σ[ w ∈ W S Q ] ((i : Ind) → Pos P WAlg i w → X i) → Y + α̅ (w , k) = α̅' w k + + -- Diagram commutes + α̅Comm : (s : S) (f : Q s → W S Q) (g : (i : Ind) → P i s → X i) + (h : (i : Ind) (q : Q s) → Pos P WAlg i (f q) → X i) → + α̅ (into ((s , f) , (g , h))) ≡ α (s , g , λ q → α̅ (f q , λ i → h i q)) + α̅Comm s f g h = refl + + -- α̅ is unique + α̅Unique : (α̃ : Σ[ w ∈ W S Q ] ((i : Ind) → Pos P WAlg i w → X i) → Y) → + ((s : S) (f : Q s → W S Q) (g : (i : Ind) → P i s → X i) + (h : (i : Ind) (q : Q s) → Pos P WAlg i (f q) → X i) → + α̃ (into ((s , f) , (g , h))) ≡ α (s , g , λ q → α̃ (f q , λ i → h i q))) → + α̅ ≡ α̃ + α̅Unique α̃ α̃-comm = funExt w-rec + where + lemma : (s : S) (f : Q s → W S Q) (g : (i : Ind) → Pos P WAlg i (sup-W s f) → X i) → + α̃ (into ((s , f) , (λ i p → g i (here p)) , (λ i q b → g i (below q b)))) ≡ + α̃ (sup-W s f , g) + lemma s f g = cong₂ (λ w fun → α̃ (w , fun)) refl (funExt λ i → funExt (λ {(here p) → refl ; (below q b) → refl})) + + w-rec : (x : Σ[ w ∈ W S Q ] ((i : Ind) → Pos P WAlg i w → X i)) → α̅ x ≡ α̃ x + w-rec (w , k) = WInd S Q + (λ w → (k : (i : Ind) → Pos P WAlg i w → X i) → α̅ (w , k) ≡ α̃ (w , k)) + (λ {s'} {f'} ind k → + (λ i → α (s' , (λ i p → k i (here p)) , + λ q → ind q (λ i pos → k i (below q pos)) i)) ∙ + sym (α̃-comm s' f' (λ i p → k i (here p)) + (λ i q b → k i (below q b))) ∙ + lemma s' f' k) + w k diff --git a/Cubical/Data/W/W.agda b/Cubical/Data/W/W.agda new file mode 100644 index 0000000000..ac9c84db7d --- /dev/null +++ b/Cubical/Data/W/W.agda @@ -0,0 +1,17 @@ +{-# OPTIONS --safe #-} + +module Cubical.Data.W.W where + +open import Cubical.Foundations.Prelude + +private + variable + ℓ ℓ' ℓ'' : Level + +data W (S : Type ℓ) (P : S → Type ℓ') : Type (ℓ-max ℓ ℓ') where + sup-W : (s : S) → (P s → W S P) → W S P + +WInd : (S : Type ℓ) (P : S → Type ℓ') (M : W S P → Type ℓ'') → + (e : {s : S} {f : P s → W S P} → ((p : P s) → M (f p)) → M (sup-W s f)) → + (w : W S P) → M w +WInd S P M e (sup-W s f) = e (λ p → WInd S P M e (f p)) From ba0a56264b0501eb02056d7fb860cb185e5c65f2 Mon Sep 17 00:00:00 2001 From: Andreas Abel Date: Wed, 24 Jul 2024 11:21:30 +0200 Subject: [PATCH 25/30] Supply implicit argument for agda/agda#7349 (#1143) This small change makes the code more robust w.r.t. the Agda type checker, in particular enabling Agda PR #7349. --- Cubical/Data/Sigma/Properties.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index 68514fd5d8..bb21e1d364 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -330,7 +330,7 @@ snd (leftInv (Σ-contractFstIso {B = B} c) p j) = -- a special case of the above module _ (A : Unit → Type ℓ) where ΣUnit : Σ Unit A ≃ A tt - unquoteDef ΣUnit = defStrictEquiv ΣUnit snd (λ { x → (tt , x) }) + unquoteDef ΣUnit = defStrictEquiv {B = A tt} ΣUnit snd (tt ,_) Σ-contractSnd : ((a : A) → isContr (B a)) → Σ A B ≃ A Σ-contractSnd c = isoToEquiv isom From 1011201c5d4886d3b80411e25f2ad0d2f508eab1 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 24 Jul 2024 20:18:49 +0200 Subject: [PATCH 26/30] add comment on our notion of field (#1142) --- Cubical/Algebra/Field/Base.agda | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Cubical/Algebra/Field/Base.agda b/Cubical/Algebra/Field/Base.agda index 445238dd6b..c548c1de85 100644 --- a/Cubical/Algebra/Field/Base.agda +++ b/Cubical/Algebra/Field/Base.agda @@ -1,4 +1,12 @@ {-# OPTIONS --safe #-} +{- + Keep in mind, that here are many different notions of "field" in constructive algebra. + In the terminology of "A Course in Constructive Algebra" (by Mines, Richman, Ruitenburg) (p. 45), + the notion of field we use below, would be a nontrivial field (where the apartness relation + used in the definition of field is inequality in our case). + This is a very weak notion of field, but behaves a lot like the classical notion, if the carrier + type has decidable equality. +-} module Cubical.Algebra.Field.Base where open import Cubical.Foundations.Prelude From e2cf0ab7ced7ea57685b138d28761eac245fd327 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Axel=20Ljungstr=C3=B6m?= <49276137+aljungstrom@users.noreply.github.com> Date: Mon, 2 Sep 2024 09:30:55 +0200 Subject: [PATCH 27/30] Pi4S3 Paper (#1151) * t * duplicate * wups * rme * ganea stuff * w * w * w * stuff * stuff * fixes * fixes * fixes * added comment --- Cubical/HITs/Join/Base.agda | 6 + Cubical/HITs/SmashProduct/Base.agda | 373 ++++++++ Cubical/HITs/Sn/Multiplication.agda | 898 ++++++++++++++++++ Cubical/HITs/Sn/Properties.agda | 30 + Cubical/HITs/Susp/Base.agda | 10 +- Cubical/Homotopy/Group/Join.agda | 326 +++++++ .../Group/Pi4S3/BrunerieExperiments.agda | 2 +- Cubical/Homotopy/Whitehead.agda | 103 ++ Cubical/Papers/Pi4S3-JournalVersion.agda | 397 ++++++++ 9 files changed, 2143 insertions(+), 2 deletions(-) create mode 100644 Cubical/HITs/Sn/Multiplication.agda create mode 100644 Cubical/Homotopy/Group/Join.agda create mode 100644 Cubical/Papers/Pi4S3-JournalVersion.agda diff --git a/Cubical/HITs/Join/Base.agda b/Cubical/HITs/Join/Base.agda index d94f9551a6..330bf198db 100644 --- a/Cubical/HITs/Join/Base.agda +++ b/Cubical/HITs/Join/Base.agda @@ -4,6 +4,7 @@ module Cubical.HITs.Join.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed.Base open import Cubical.HITs.S1 open import Cubical.HITs.S3 @@ -15,6 +16,11 @@ data join {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') wh inr : B → join A B push : ∀ a b → inl a ≡ inr b +join∙ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → Pointed _ +fst (join∙ A B) = join (fst A) (fst B) +snd (join∙ A B) = inl (pt A) + 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) diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index 6ce442c4f7..2c591a2797 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -9,6 +9,7 @@ open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.Path open import Cubical.Foundations.Function open import Cubical.Foundations.Transport +open import Cubical.Foundations.Equiv open import Cubical.Data.Unit open import Cubical.Data.Sigma @@ -17,6 +18,10 @@ open import Cubical.Data.Fin open import Cubical.HITs.Pushout.Base open import Cubical.HITs.Wedge +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Join + +open import Cubical.Homotopy.Loopspace data Smash {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') : Type (ℓ-max ℓ ℓ') where basel : Smash A B @@ -279,6 +284,34 @@ _⋀∙→refl_ : ∀ {ℓ ℓ'} {C : Type ℓ} {D : Type ℓ'} fst (f ⋀∙→refl g) = f ⋀→refl g snd (f ⋀∙→refl g) = refl + +⋀≃ : ∀ {ℓ ℓ'} {A B : Pointed ℓ} {C D : Pointed ℓ'} + → (f : A ≃∙ B) (g : C ≃∙ D) + → (A ⋀ C) ≃ (B ⋀ D) +⋀≃ {ℓ = ℓ} {ℓ'} {B = B} {D = D} f g = _ , ⋀≃-isEq f g + where + help : (x : _) → (idfun∙ B ⋀→ idfun∙ D) x ≡ x + help (inl x) = refl + help (inr x) = refl + help (push (inl x) i) j = rUnit (push (inl x)) (~ j) i + help (push (inr x) i) j = rUnit (push (inr x)) (~ j) i + help (push (push a i) j) k = + hcomp (λ r → λ {(i = i0) → rUnit (push (inl (snd B))) (~ k ∧ r) j + ; (i = i1) → rUnit (push (inr (snd D))) (~ k ∧ r) j + ; (j = i0) → inl tt + ; (j = i1) → inr (snd B , snd D) + ; (k = i1) → push (push tt i) j}) + (push (push tt i) j) + + ⋀≃-isEq : {A : Pointed ℓ} {C : Pointed ℓ'} + (f : A ≃∙ B) (g : C ≃∙ D) → isEquiv (≃∙map f ⋀→ ≃∙map g) + ⋀≃-isEq {C = C} = + Equiv∙J (λ A f → (g : C ≃∙ D) + → isEquiv (≃∙map f ⋀→ ≃∙map g)) + (Equiv∙J (λ _ g → isEquiv (idfun∙ _ ⋀→ ≃∙map g)) + (subst isEquiv (sym (funExt help)) (idIsEquiv _))) + + ⋀→Smash : A ⋀ B → Smash A B ⋀→Smash (inl x) = basel ⋀→Smash (inr (x , x₁)) = proj x x₁ @@ -918,3 +951,343 @@ module ⋀-fun≡' {C : Type ℓ} (f g : A ⋀ B → C) main : (x : _) → f x ≡ g x main = ⋀-fun≡ {A = A} {B = B} f g p pr lp lem + +-- Suspension of a smash product is a join +module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} where + private + sm-fillᵣ : ∀ {ℓ} {A : Type ℓ} {x* : A} (y* : A) (p* : x* ≡ y*) + (y : A) → (p : x* ≡ y) + → sym p* ≡ (sym p* ∙∙ p ∙∙ sym p) + sm-fillᵣ y* p* y p j i = + hcomp (λ r → λ {(i = i0) → p* r + ; (i = i1) → p (~ r ∧ j) + ; (j = i0) → p* (~ i ∧ r) + ; (j = i1) → doubleCompPath-filler + (sym p*) p (sym p) r i}) + (p (i ∧ j)) + + sm-fillₗ : ∀ {ℓ} {A : Type ℓ} {x* : A} (y* : A) (p* : x* ≡ y*) + (y : A) (p : x* ≡ y) + → p* ≡ (p ∙∙ sym p ∙∙ p*) + sm-fillₗ y* p* y p j i = + hcomp (λ r → λ {(i = i0) → p (~ r ∧ j) + ; (i = i1) → p* r + ; (j = i0) → p* (r ∧ i) + ; (j = i1) → doubleCompPath-filler + p (sym p) p* r i}) + (p (~ i ∧ j)) + + sm-fillₗᵣ≡ : ∀ {ℓ} {A : Type ℓ} {x* : A} (y* : A) (p* : x* ≡ y*) + → sm-fillₗ _ (sym p*) _ (sym p*) ≡ sm-fillᵣ _ p* _ p* + sm-fillₗᵣ≡ = J> refl + + SuspSmash→Join : Susp (A ⋀ B) → (join (typ A) (typ B)) + SuspSmash→Join north = inr (pt B) + SuspSmash→Join south = inl (pt A) + SuspSmash→Join (merid (inl x) i) = + push (pt A) (pt B) (~ i) + SuspSmash→Join (merid (inr (x , b)) i) = + (sym (push x (pt B)) ∙∙ push x b ∙∙ sym (push (pt A) b)) i + SuspSmash→Join (merid (push (inl x) j) i) = + sm-fillₗ {A = join (typ A) (typ B)} _ + (sym (push (pt A) (pt B))) _ (sym (push x (pt B))) j i + SuspSmash→Join (merid (push (inr x) j) i) = + sm-fillᵣ {A = join (typ A) (typ B)} _ + (push (pt A) (pt B)) _ (push (pt A) x) j i + SuspSmash→Join (merid (push (push a k) j) i) = + sm-fillₗᵣ≡ _ (push (pt A) (pt B)) k j i + + Join→SuspSmash : join (typ A) (typ B) → Susp (A ⋀ B) + Join→SuspSmash (inl x) = north + Join→SuspSmash (inr x) = south + Join→SuspSmash (push a b i) = merid (inr (a , b)) i + + Join→SuspSmash→Join : (x : join (typ A) (typ B)) + → SuspSmash→Join (Join→SuspSmash x) ≡ x + Join→SuspSmash→Join (inl x) = sym (push x (pt B)) + Join→SuspSmash→Join (inr x) = push (pt A) x + Join→SuspSmash→Join (push a b i) j = + doubleCompPath-filler + (sym (push a (pt B))) (push a b) (sym (push (pt A) b)) (~ j) i + + SuspSmash→Join→SuspSmash : (x : Susp (A ⋀ B)) + → Join→SuspSmash (SuspSmash→Join x) ≡ x + SuspSmash→Join→SuspSmash north = sym (merid (inr (pt A , pt B))) + SuspSmash→Join→SuspSmash south = merid (inr (pt A , pt B)) + SuspSmash→Join→SuspSmash (merid a i) j = + hcomp (λ r + → λ {(i = i0) → merid (inr (pt A , pt B)) (~ j ∨ ~ r) + ; (i = i1) → merid (inr (pt A , pt B)) (j ∧ r) + ; (j = i0) → Join→SuspSmash (SuspSmash→Join (merid a i)) + ; (j = i1) → doubleCompPath-filler + (sym (merid (inr (pt A , pt B)))) + (merid a) + (sym (merid (inr (pt A , pt B)))) (~ r) i}) + (f₁₂ j .fst a i) + where + f₁ f₂ : A ⋀∙ B →∙ (Path (Susp (A ⋀ B)) south north + , sym (merid (inr (snd A , snd B)))) + (fst f₁) a i = Join→SuspSmash (SuspSmash→Join (merid a i)) + snd f₁ = refl + (fst f₂) a = + sym (merid (inr (pt A , pt B))) + ∙∙ merid a + ∙∙ sym (merid (inr (pt A , pt B))) + snd f₂ = cong₂ (λ x y → sym x ∙∙ y ∙∙ sym x) + refl (cong merid (push (inl (pt A)))) + ∙ doubleCompPath≡compPath + (sym (merid (inr (pt A , pt B)))) _ _ + ∙ cong₂ _∙_ refl (rCancel (merid (inr (pt A , pt B)))) + ∙ sym (rUnit _) + + f₁₂ : f₁ ≡ f₂ + f₁₂ = ⋀→∙Homogeneous≡ (isHomogeneousPath _ _) + λ x y → cong-∙∙ Join→SuspSmash + (sym (push x (pt B))) + (push x y) + (sym (push (pt A) y)) + ∙ (λ i → sym (merid ((sym (push (inl x)) + ∙ push (inl (pt A))) i)) + ∙∙ merid (inr (x , y)) + ∙∙ sym (merid ((sym (push (inr y)) + ∙ push (inl (pt A))) i))) + + SmashJoinIso : Iso (Susp (A ⋀ B)) (join (typ A) (typ B)) + Iso.fun SmashJoinIso = SuspSmash→Join + Iso.inv SmashJoinIso = Join→SuspSmash + Iso.rightInv SmashJoinIso = Join→SuspSmash→Join + Iso.leftInv SmashJoinIso = SuspSmash→Join→SuspSmash + +-- Suspension commutes with smash products +module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} where + + -- some fillers + merid-fill : typ A → I → I → I → Susp (A ⋀ B) + merid-fill a i j k = + hfill (λ k → λ {(i = i0) → north + ; (i = i1) → merid (push (inl a) k) j + ; (j = i0) → north + ; (j = i1) → merid (inl tt) i}) + (inS (merid (inl tt) (i ∧ j))) k + + inl-fill₁ : typ A → I → I → I → (Susp∙ (typ A)) ⋀ B + inl-fill₁ a i j k = + hfill (λ k → λ {(i = i0) → inl tt + ; (i = i1) → inr (σ A a j , snd B) + ; (j = i0) → push (push tt k) i + ; (j = i1) → push (push tt k) i}) + (inS (push (inl (σ A a j)) i)) k + + inl-fill : typ A → I → I → I → (Susp∙ (typ A)) ⋀ B + inl-fill a i j k = + hfill (λ k → λ {(i = i0) → inl tt + ; (i = i1) → doubleCompPath-filler + (push (inr (snd B))) + (λ i₂ → inr (σ A a i₂ , snd B)) + (sym (push (inr (snd B)))) k j + ; (j = i0) → push (inr (snd B)) (~ k ∧ i) + ; (j = i1) → push (inr (snd B)) (~ k ∧ i)}) + (inS (inl-fill₁ a i j i1)) k + + inr-fill₁ : typ B → I → I → I → (Susp∙ (typ A)) ⋀ B + inr-fill₁ b i j k = + hfill (λ k → λ {(i = i0) → inl tt + ; (i = i1) → inr (rCancel (merid (pt A)) (~ k) j , b) + ; (j = i0) → push (inr b) i + ; (j = i1) → push (inr b) i}) + (inS (push (inr b) i)) k + + inr-fill : typ B → I → I → I → (Susp∙ (typ A)) ⋀ B + inr-fill b i j k = + hfill (λ k → λ {(i = i0) → inl tt + ; (i = i1) → doubleCompPath-filler + (push (inr b)) + (λ i₂ → inr (σ A (pt A) i₂ , b)) + (sym (push (inr b))) k j + ; (j = i0) → push (inr b) (~ k ∧ i) + ; (j = i1) → push (inr b) (~ k ∧ i)}) + (inS (inr-fill₁ b i j i1)) k + + SuspL→Susp⋀ : (Susp∙ (typ A)) ⋀ B → Susp (A ⋀ B) + SuspL→Susp⋀ (inl x) = north + SuspL→Susp⋀ (inr (north , y)) = north + SuspL→Susp⋀ (inr (south , y)) = south + SuspL→Susp⋀ (inr (merid a i , y)) = merid (inr (a , y)) i + SuspL→Susp⋀ (push (inl north) i) = north + SuspL→Susp⋀ (push (inl south) i) = merid (inl tt) i + SuspL→Susp⋀ (push (inl (merid a j)) i) = merid-fill a i j i1 + SuspL→Susp⋀ (push (inr x) i) = north + SuspL→Susp⋀ (push (push a i₁) i) = north + + Susp⋀→SuspL : Susp (A ⋀ B) → (Susp∙ (typ A)) ⋀ B + Susp⋀→SuspL north = inl tt + Susp⋀→SuspL south = inl tt + Susp⋀→SuspL (merid (inl x) i) = inl tt + Susp⋀→SuspL (merid (inr (x , y)) i) = + (push (inr y) ∙∙ (λ i → inr (σ A x i , y)) ∙∙ sym (push (inr y))) i + Susp⋀→SuspL (merid (push (inl x) i₁) i) = inl-fill x i₁ i i1 + Susp⋀→SuspL (merid (push (inr x) i₁) i) = inr-fill x i₁ i i1 + Susp⋀→SuspL (merid (push (push a k) j) i) = + hcomp (λ r → λ {(i = i0) → inl-fill (snd A) j i r + ; (i = i1) → inr-fill (snd B) j i r + ; (j = i0) → inl tt + ; (j = i1) → doubleCompPath-filler + (push (inr (pt B))) + (λ i₂ → inr (σ A (pt A) i₂ , (pt B))) + (sym (push (inr (pt B)))) r i + ; (k = i0) → inl-fill (snd A) j i r + ; (k = i1) → inr-fill (snd B) j i r}) + (hcomp (λ r → λ {(i = i0) → push (push tt (r ∨ k)) j + ; (i = i1) → push (push tt (r ∨ k)) j + ; (j = i0) → inl tt + ; (j = i1) → inr (rCancel (merid (pt A)) + (~ r ∧ k) i , pt B) + ; (k = i0) → inl-fill₁ (snd A) j i r + ; (k = i1) → inr-fill₁ (snd B) j i r}) + (hcomp (λ r → λ {(i = i0) → push (push tt k) j + ; (i = i1) → push (push tt k) j + ; (j = i0) → inl tt + ; (j = i1) → inr (rCancel (merid (pt A)) + (~ r ∨ k) i , pt B) + ; (k = i0) → push (inl (rCancel (merid (pt A)) + (~ r) i)) j + ; (k = i1) → push (inr (snd B)) j}) + (push (push tt k) j))) + + SuspSmashCommIso : Iso (Susp∙ (typ A) ⋀ B) (Susp (A ⋀ B)) + Iso.fun SuspSmashCommIso = SuspL→Susp⋀ + Iso.inv SuspSmashCommIso = Susp⋀→SuspL + Iso.rightInv SuspSmashCommIso north = refl + Iso.rightInv SuspSmashCommIso south = merid (inl tt) + Iso.rightInv SuspSmashCommIso (merid a i) j = + hcomp (λ r → λ {(i = i0) → north + ; (i = i1) → merid (inl tt) (j ∧ r) + ; (j = i0) → f₁≡f₂ (~ r) .fst a i + ; (j = i1) → compPath-filler + (merid a) (sym (merid (inl tt))) (~ r) i }) + (f₂ .fst a i) + where + f₁ f₂ : (A ⋀∙ B) →∙ Ω (Susp∙ (A ⋀ B)) + fst f₁ x = cong SuspL→Susp⋀ (cong Susp⋀→SuspL (merid x)) + snd f₁ = refl + fst f₂ = σ (A ⋀∙ B) + snd f₂ = rCancel (merid (inl tt)) + + inr' : (Susp (typ A)) × (typ B) → (Susp∙ (typ A)) ⋀ B + inr' = inr + + f₁≡f₂ : f₁ ≡ f₂ + f₁≡f₂ = + ⋀→∙Homogeneous≡ (isHomogeneousPath _ _) + λ x y + → cong (cong SuspL→Susp⋀) (cong (push (inr y) ∙∙_∙∙ sym (push (inr y))) + (cong-∙ (λ x → inr' (x , y)) (merid x) (sym (merid (pt A))))) + ∙∙ cong-∙∙ SuspL→Susp⋀ (push (inr y)) + (cong (λ x → inr' (x , y)) (merid x) + ∙ cong (λ x → inr' (x , y)) (sym (merid (pt A)))) (sym (push (inr y))) + ∙∙ (sym (rUnit _) + ∙ cong-∙ SuspL→Susp⋀ + (λ i → inr' (merid x i , y)) (λ i → inr' (merid (pt A) (~ i) , y)) + ∙ cong (merid (inr (x , y)) ∙_) + λ j i → merid (push (inr y) (~ j)) (~ i) ) + Iso.leftInv SuspSmashCommIso = + ⋀-fun≡ _ _ refl + (λ x → main (snd x) (fst x)) + (λ { north i j → sₙ i j i1 + ; south i j → sₛ i j i1 + ; (merid a k) i j → cube a j k i}) + λ x i j → push (inr x) (i ∧ j) + where + inr' : Susp (typ A) × (typ B) → (Susp∙ (typ A)) ⋀ B + inr' = inr + sₙ : I → I → I → (Susp∙ (typ A)) ⋀ B + sₙ i j k = + hfill (λ k → λ {(i = i0) → inl tt + ; (i = i1) → push (inr (pt B)) (j ∨ ~ k) + ; (j = i0) → push (inr (pt B)) (i ∧ ~ k) + ; (j = i1) → push (inl north) i}) + (inS (push (push tt (~ j)) i)) + k + + sₛ : I → I → I → (Susp∙ (typ A)) ⋀ B + sₛ i j k = + hfill (λ k → λ {(i = i0) → inl tt + ; (i = i1) → compPath-filler + (push (inr (pt B))) + (λ i → inr (merid (pt A) i , pt B)) k j + ; (j = i0) → inl tt + ; (j = i1) → push (inl (merid (pt A) k)) i}) + (inS (sₙ i j i1)) + k + + filler : fst A → fst B → I → I → I → (Susp∙ (typ A)) ⋀ B + filler a y i j k = + hfill (λ k → λ {(i = i0) → push (inr y) j + ; (i = i1) → compPath-filler + (push (inr y)) + (λ i₁ → inr (merid (pt A) i₁ , y)) k j + ; (j = i0) → Susp⋀→SuspL (SuspL→Susp⋀ (inr (merid a i , y))) + ; (j = i1) → inr (compPath-filler + (merid a) + (sym (merid (pt A))) (~ k) i , y)}) + (inS (doubleCompPath-filler + (push (inr y)) + (λ i₁ → inr' (σ A a i₁ , y)) + (sym (push (inr y))) (~ j) i)) k + + cube₁ : (a : typ A) + → Cube (λ k i → Iso.inv SuspSmashCommIso (merid-fill a k i i1)) + (λ k i → inl-fill a k i i1) + (λ _ _ → inl tt) + refl + (λ _ _ → inl tt) + λ _ _ → inl tt + cube₁ a j k i = + hcomp (λ r → λ {(i = i0) → inl tt + ; (i = i1) → inl tt + ; (j = i0) → Iso.inv SuspSmashCommIso (merid-fill a k i r) + ; (j = i1) → inl-fill a k i i1 + ; (k = i0) → inl tt + ; (k = i1) → inl-fill a (j ∨ r) i i1}) + (inl-fill a (j ∧ k) i i1) + + cube : (a : typ A) + → Cube (λ k i → Iso.inv SuspSmashCommIso + (Iso.fun SuspSmashCommIso (push (inl (merid a k)) i))) + (λ k i → push (inl (merid a k)) i) + (λ j i → sₙ i j i1) (λ j i → sₛ i j i1) + (λ j k → inl tt) (λ j k → filler a (pt B) k j i1) + cube a = (λ j k i → cube₁ a j i k) ◁ + (λ j k i → + hcomp (λ r + → λ {(i = i0) → inl tt + ; (i = i1) → filler a (pt B) k j r + ; (j = i0) → inl-fill a i k i1 + ; (j = i1) → push (inl (compPath-filler + (merid a) (sym (merid (pt A))) (~ r) k)) i + ; (k = i0) → sₙ i j i1 + ; (k = i1) → sₛ i j r}) + (hcomp (λ r + → λ {(i = i0) → inl tt + ; (i = i1) → doubleCompPath-filler + (push (inr (pt B))) + (λ i₁ → inr' (σ A a i₁ , (pt B))) + (sym (push (inr (pt B)))) (~ j ∧ r) k + ; (j = i0) → inl-fill a i k r + ; (j = i1) → push (inl (σ A a k)) i + ; (k = i0) → sₙ i j r + ; (k = i1) → sₙ i j r}) + (hcomp (λ r + → λ {(i = i0) → inl tt + ; (i = i1) → inl-fill₁ a i k r + ; (j = i0) → inl-fill₁ a i k r + ; (j = i1) → push (inl (σ A a k)) i + ; (k = i0) → push (push tt (r ∧ (~ j))) i + ; (k = i1) → push (push tt (r ∧ (~ j))) i}) + (push (inl (σ A a k)) i)))) + + main : (y : typ B) (x : Susp (typ A)) + → Susp⋀→SuspL (SuspL→Susp⋀ (inr (x , y))) ≡ inr (x , y) + main y north = push (inr y) + main y south = push (inr y) ∙ λ i → inr (merid (pt A) i , y) + main y (merid a i) j = filler a y i j i1 diff --git a/Cubical/HITs/Sn/Multiplication.agda b/Cubical/HITs/Sn/Multiplication.agda new file mode 100644 index 0000000000..00ca912c0c --- /dev/null +++ b/Cubical/HITs/Sn/Multiplication.agda @@ -0,0 +1,898 @@ +{-# OPTIONS --safe #-} + +{- +This file contains: +1. Definition of the multplication Sⁿ × Sᵐ → Sⁿ⁺ᵐ +2. The fact that the multiplication induces an equivalence Sⁿ ∧ Sᵐ ≃ Sⁿ⁺ᵐ +3. The algebraic properties of this map +-} + +module Cubical.HITs.Sn.Multiplication where + +open import Cubical.Foundations.Prelude +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.GroupoidLaws +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Data.Sum +open import Cubical.Data.Bool hiding (elim) +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.Sigma + +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn hiding (IsoSphereJoin) +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Join +open import Cubical.HITs.Pushout +open import Cubical.HITs.SmashProduct + +open import Cubical.Homotopy.Loopspace + +open Iso + +open PlusBis + + +---- Sphere multiplication +-- auxiliary function +sphereFun↑ : {n m k : ℕ} + → (f : S₊ n → S₊ m → S₊ k) + → S₊ (suc n) → S₊ m → S₊ (suc k) +sphereFun↑ {n = zero} {m = m} f base y = ptSn _ +sphereFun↑ {n = zero} {m = m} f (loop i) y = σS (f false y) i +sphereFun↑ {n = suc n} {m = m} f north y = ptSn _ +sphereFun↑ {n = suc n} {m = m} f south y = ptSn _ +sphereFun↑ {n = suc n} {m = m} f (merid a i) y = σS (f a y) i + +-- sphere multiplication +_⌣S_ : {n m : ℕ} → S₊ n → S₊ m → S₊ (n + m) +_⌣S_ {n = zero} {m = m} false y = y +_⌣S_ {n = zero} {m = m} true y = ptSn m +_⌣S_ {n = suc n} {m = m} = sphereFun↑ (_⌣S_ {n = n}) + +-- Left- and right-unit laws +IdR⌣S : {n m : ℕ} (y : S₊ m) + → Path (S₊ (n + m)) (ptSn n ⌣S y) (ptSn (n + m)) +IdR⌣S {n = zero} {m = m} y = refl +IdR⌣S {n = suc zero} {m = m} y = refl +IdR⌣S {n = suc (suc n)} {m = m} y = refl + +IdL⌣S : {n m : ℕ} (x : S₊ n) + → Path (S₊ (n + m)) (x ⌣S (ptSn m)) (ptSn (n + m)) +IdL⌣S {n = zero} false = refl +IdL⌣S {n = zero} true = refl +IdL⌣S {n = suc zero} base = refl +IdL⌣S {n = suc zero} {zero} (loop i) j = base +IdL⌣S {n = suc zero} {suc m} (loop i) j = rCancel (merid (ptSn (suc m))) j i +IdL⌣S {n = suc (suc n)} {m} north j = north +IdL⌣S {n = suc (suc n)} {m} south j = north +IdL⌣S {n = suc (suc n)} {m} (merid a i) j = + (cong σS (IdL⌣S a) + ∙ rCancel (merid (ptSn _))) j i + +IdL⌣S≡IdR⌣S : (n m : ℕ) + → IdL⌣S {n = n} {m = m} (ptSn n) ≡ IdR⌣S (ptSn m) +IdL⌣S≡IdR⌣S zero m = refl +IdL⌣S≡IdR⌣S (suc zero) m = refl +IdL⌣S≡IdR⌣S (suc (suc n)) m = refl + +-- Multiplication induced on smash products of spheres +⋀S∙ : (n m : ℕ) → S₊∙ n ⋀∙ S₊∙ m →∙ S₊∙ (n + m) +fst (⋀S∙ n m) (inl x) = ptSn _ +fst (⋀S∙ n m) (inr x) = (fst x) ⌣S (snd x) +fst (⋀S∙ n m) (push (inl x) i) = IdL⌣S x (~ i) +fst (⋀S∙ n m) (push (inr x) i) = IdR⌣S x (~ i) +fst (⋀S∙ n m) (push (push a i₁) i) = IdL⌣S≡IdR⌣S n m i₁ (~ i) +snd (⋀S∙ n m) = refl + +⋀S : (n m : ℕ) → S₊∙ n ⋀ S₊∙ m → S₊ (n + m) +⋀S n m = fst (⋀S∙ n m) + +-- Proof that it is an equivalence +⋀S-base : (m : ℕ) + → Iso (S₊∙ zero ⋀ S₊∙ m) (S₊ m) +fun (⋀S-base m) = ⋀S zero m +inv (⋀S-base m) x = inr (false , x) +rightInv (⋀S-base m) x = refl +leftInv (⋀S-base m) = + ⋀-fun≡ _ _ + (sym (push (inl false))) + (λ { (false , y) → refl + ; (true , y) → sym (push (inl false)) ∙ push (inr y)}) + (λ { false i j → push (inl false) (i ∨ ~ j) + ; true → compPath-filler (sym (push (inl false))) (push (inl true)) + ▷ cong (sym (push (inl false)) ∙_) + (λ i → push (push tt i) )}) + λ x → compPath-filler (sym (push (inl false))) (push (inr x)) +{- +Proof that ⋀S respects suspension, i.e. that the following diagram commutes + ⋀S +Sⁿ⁺¹ ∧ Sᵐ ---------------> Sⁿ⁺ᵐ⁺¹ +| | +| | +v v +Σ (Sⁿ ∧ Sᵐ) -----------> Σ Sⁿ⁺ᵐ + Σ(⋀S) +-} + +⋀S-ind : (n m : ℕ) (x : _) + → ⋀S (suc n) m x + ≡ Iso.inv (IsoSucSphereSusp (n + m)) + (suspFun (⋀S n m) (Iso.fun SuspSmashCommIso + (((Iso.fun (IsoSucSphereSusp n) , IsoSucSphereSusp∙' n) + ⋀→ idfun∙ (S₊∙ m)) x))) +⋀S-ind zero m = ⋀-fun≡ _ _ + (sym (IsoSucSphereSusp∙ m)) + (λ x → main m (fst x) (snd x)) + (mainₗ m) + mainᵣ + where + F' : (m : ℕ) → Susp ((Bool , true) ⋀ S₊∙ m) → _ + F' m = inv (IsoSucSphereSusp (zero + m)) ∘ suspFun (⋀S zero m) + + F : (m : ℕ) → Susp∙ Bool ⋀ S₊∙ m → _ + F m = F' m ∘ fun SuspSmashCommIso + + G : (m : ℕ) → _ → _ + G m = _⋀→_ {A = S₊∙ 1} {B = S₊∙ m} + (fun (IsoSucSphereSusp zero) , (λ _ → north)) + (idfun∙ (S₊∙ m)) + + main : (m : ℕ) (x : S¹) (y : S₊ m) + → x ⌣S y + ≡ F m (inr (S¹→SuspBool x , y)) + main m base y = sym (IsoSucSphereSusp∙ m) + main zero (loop i) false j = + ((cong-∙ (λ x → F zero (inr (x , false))) + (merid false) (sym (merid true))) + ∙ sym (rUnit loop)) (~ j) i + main zero (loop i) true j = + F zero (inr (rCancel (merid true) (~ j) i , false)) + main (suc m) (loop i) y j = + cong-∙ (λ x → F (suc m) (inr (x , y))) + (merid false) (sym (merid true)) (~ j) i + + mainₗ : (m : ℕ) (x : S¹) + → PathP (λ i → IdL⌣S {m = m} x (~ i) + ≡ F m (G m (push (inl x) i))) + (sym (IsoSucSphereSusp∙ m)) + (main m x (ptSn m)) + mainₗ zero = + toPropElim (λ _ → isOfHLevelPathP' 1 (isGroupoidS¹ _ _) _ _) + (flipSquare (cong (cong (F zero)) (rUnit (push (inl north))))) + mainₗ (suc m) x = flipSquare (help x + ▷ (cong (cong (F (suc m))) (rUnit (push (inl (S¹→SuspBool x)))))) + where + help : (x : S¹) + → PathP (λ i → north ≡ main (suc m) x (ptSn (suc m)) i) + (sym (IdL⌣S {n = 1} x)) + (cong (F (suc m)) (push (inl (S¹→SuspBool x)))) + help base = refl + help (loop i) j k = + hcomp (λ r + → λ {(i = i0) → north + ; (i = i1) → F' (suc m) + (merid-fill + {A = Bool , true} + {B = S₊∙ (suc m)} true (~ r) k j) + ; (j = i0) → rCancel-filler (merid (ptSn (suc m))) r (~ k) i + ; (j = i1) → F (suc m) + (push (inl (compPath-filler + (merid false) (sym (merid true)) r i)) k) + ; (k = i0) → north + ; (k = i1) → cong-∙∙-filler + (λ x₁ → F (suc m) (inr (x₁ , ptSn (suc m)))) + refl (merid false) (sym (merid true)) r (~ j) i}) + (F' (suc m) (merid-fill {A = Bool , true} {B = S₊∙ (suc m)} false k i j)) + mainᵣ : (x : S₊ m) + → PathP (λ i → ptSn (suc m) ≡ F m (G m (push (inr x) i))) + (sym (IsoSucSphereSusp∙ m)) + (sym (IsoSucSphereSusp∙ m)) + mainᵣ x = flipSquare ((λ i j → (IsoSucSphereSusp∙ m) (~ i)) + ▷ cong (cong (F m)) (rUnit (push (inr x)))) +⋀S-ind (suc n) m = ⋀-fun≡ _ _ refl + (λ x → h (fst x) (snd x)) + hₗ + λ x → flipSquare (cong (cong (suspFun (⋀S (suc n) m) + ∘ fun SuspSmashCommIso)) + (rUnit (push (inr x)))) + where + h : (x : S₊ (suc (suc n))) (y : S₊ m) + → (x ⌣S y) + ≡ suspFun (⋀S (suc n) m) + (SuspL→Susp⋀ (inr (idfun (Susp (S₊ (suc n))) x , y))) + h north y = refl + h south y = merid (ptSn _) + h (merid a i) y j = compPath-filler + (merid (a ⌣S y)) (sym (merid (ptSn (suc (n + m))))) (~ j) i + + hₗ-lem : (x : Susp (S₊ (suc n))) + → PathP (λ i → north ≡ h x (ptSn m) i) + (sym (IdL⌣S x)) + (cong (suspFun (⋀S (suc n) m) + ∘ fun SuspSmashCommIso) + (push (inl x))) + hₗ-lem north = refl + hₗ-lem south i j = merid (ptSn (suc (n + m))) (i ∧ j) + hₗ-lem (merid a i) j k = help j k i + where + help : Cube (sym (cong σS + (IdL⌣S {n = suc n} {m = m} a) + ∙ rCancel (merid (ptSn (suc n + m))))) + (λ k i → suspFun (⋀S (suc n) m) + (SuspL→Susp⋀ (push (inl (merid a i)) k))) + (λ j i → north) + (λ j i → compPath-filler + (merid (a ⌣S ptSn m)) + (sym (merid (ptSn (suc (n + m))))) (~ j) i) + (λ j k → north) + λ j k → merid (ptSn (suc (n + m))) (j ∧ k) + help j k i = + hcomp (λ r + → λ {(i = i0) → north + ; (i = i1) → merid (ptSn (suc (n + m))) (j ∧ k) + ; (j = i0) → compPath-filler' + (cong σS + (IdL⌣S {n = suc n} {m = m} a)) + (rCancel (merid (ptSn (suc n + m)))) r (~ k) i + ; (j = i1) → suspFun (⋀S (suc n) m) + (merid-fill a k i r) + ; (k = i0) → north + ; (k = i1) → compPath-filler + (merid (IdL⌣S a (~ r))) + (sym (merid (ptSn (suc (n + m))))) (~ j) i}) + (hcomp (λ r → λ {(i = i0) → north + ; (i = i1) → merid (ptSn (suc (n + m))) ((j ∨ ~ r) ∧ k) + ; (j = i0) → rCancel-filler (merid (ptSn _)) r (~ k) i + ; (j = i1) → merid (ptSn (suc (n + m))) (i ∧ k) + ; (k = i0) → north + ; (k = i1) → compPath-filler + (merid (ptSn _)) + (sym (merid (ptSn (suc (n + m))))) + (~ j ∧ r) i}) + (merid (ptSn (suc (n + m))) (i ∧ k))) + + hₗ : (x : Susp (S₊ (suc n))) + → PathP (λ i → IdL⌣S x (~ i) + ≡ inv (IsoSucSphereSusp (suc n + m)) + (suspFun (⋀S (suc n) m) + (fun SuspSmashCommIso + (((fun (IsoSucSphereSusp (suc n)) , IsoSucSphereSusp∙' (suc n)) ⋀→ + idfun∙ (S₊∙ m)) + (push (inl x) i))))) refl (h x (ptSn m)) + hₗ x = + flipSquare + ((hₗ-lem x + ▷ sym (cong (cong (inv (IsoSucSphereSusp (suc n + m)) + ∘ suspFun (⋀S (suc n) m) + ∘ fun SuspSmashCommIso)) + (sym (rUnit (push (inl x))))))) + + +isEquiv-⋀S : (n m : ℕ) → isEquiv (⋀S n m) +isEquiv-⋀S zero m = isoToIsEquiv (⋀S-base m) +isEquiv-⋀S (suc n) m = + subst isEquiv (sym (funExt (⋀S-ind n m))) + (snd (helpEq (isEquiv-⋀S n m))) + where + r = isoToEquiv (IsoSucSphereSusp n) + + helpEq : isEquiv (⋀S n m) → (S₊∙ (suc n) ⋀ S₊∙ m) ≃ S₊ (suc n + m) + helpEq iseq = + compEquiv + (compEquiv + (compEquiv + (⋀≃ (r , IsoSucSphereSusp∙' n) (idEquiv (S₊ m) , refl)) + (isoToEquiv SuspSmashCommIso)) + (isoToEquiv + (congSuspIso (equivToIso (⋀S n m , iseq))))) + (isoToEquiv (invIso (IsoSucSphereSusp (n + m)))) + +SphereSmashIso : (n m : ℕ) → Iso (S₊∙ n ⋀ S₊∙ m) (S₊ (n + m)) +SphereSmashIso n m = equivToIso (⋀S n m , isEquiv-⋀S n m) + +-- Proof that the pinch map Sⁿ * Sᵐ → Sⁿ⁺ᵐ⁺¹ is an equivalence +join→Sphere : (n m : ℕ) + → join (S₊ n) (S₊ m) → S₊ (suc (n + m)) +join→Sphere n m (inl x) = ptSn _ +join→Sphere n m (inr x) = ptSn _ +join→Sphere n m (push a b i) = σS (a ⌣S b) i + +join→Sphere∙ : (n m : ℕ) + → join∙ (S₊∙ n) (S₊∙ m) →∙ S₊∙ (suc (n + m)) +fst (join→Sphere∙ n m) = join→Sphere n m +snd (join→Sphere∙ n m) = refl + +joinSphereIso' : (n m : ℕ) + → Iso (join (S₊ n) (S₊ m)) (S₊ (suc (n + m))) +joinSphereIso' n m = + compIso (invIso (SmashJoinIso {A = S₊∙ n} {B = S₊∙ m})) + (compIso (congSuspIso (SphereSmashIso n m)) + (invIso (IsoSucSphereSusp (n + m)))) + +join→Sphere≡ : (n m : ℕ) (x : _) + → join→Sphere n m x ≡ joinSphereIso' n m .Iso.fun x +join→Sphere≡ zero zero (inl x) = refl +join→Sphere≡ zero (suc m) (inl x) = refl +join→Sphere≡ (suc n) m (inl x) = refl +join→Sphere≡ zero zero (inr x) = refl +join→Sphere≡ zero (suc m) (inr x) = merid (ptSn (suc m)) +join→Sphere≡ (suc n) zero (inr x) = merid (ptSn (suc n + zero)) +join→Sphere≡ (suc n) (suc m) (inr x) = merid (ptSn (suc n + suc m)) +join→Sphere≡ zero zero (push false false i) j = loop i +join→Sphere≡ zero zero (push false true i) j = base +join→Sphere≡ zero zero (push true b i) j = base +join→Sphere≡ zero (suc m) (push a b i) j = + compPath-filler + (merid (a ⌣S b)) (sym (merid (ptSn (suc m)))) (~ j) i +join→Sphere≡ (suc n) zero (push a b i) j = + compPath-filler + (merid (a ⌣S b)) (sym (merid (ptSn (suc n + zero)))) (~ j) i +join→Sphere≡ (suc n) (suc m) (push a b i) j = + compPath-filler + (merid (a ⌣S b)) (sym (merid (ptSn (suc n + suc m)))) (~ j) i + +-- Todo: integrate with Sn.Properties IsoSphereJoin +IsoSphereJoin : (n m : ℕ) + → Iso (join (S₊ n) (S₊ m)) (S₊ (suc (n + m))) +fun (IsoSphereJoin n m) = join→Sphere n m +inv (IsoSphereJoin n m) = joinSphereIso' n m .Iso.inv +rightInv (IsoSphereJoin n m) x = + join→Sphere≡ n m (joinSphereIso' n m .Iso.inv x) + ∙ joinSphereIso' n m .Iso.rightInv x +leftInv (IsoSphereJoin n m) x = + cong (joinSphereIso' n m .inv) (join→Sphere≡ n m x) + ∙ joinSphereIso' n m .Iso.leftInv x + +joinSphereEquiv∙ : (n m : ℕ) → join∙ (S₊∙ n) (S₊∙ m) ≃∙ S₊∙ (suc (n + m)) +fst (joinSphereEquiv∙ n m) = isoToEquiv (IsoSphereJoin n m) +snd (joinSphereEquiv∙ n m) = refl + + +-- Associativity ⌣S +-- Preliminary lemma +⌣S-false : {n : ℕ} (x : S₊ n) → PathP (λ i → S₊ (+-comm n zero i)) (x ⌣S false) x +⌣S-false {n = zero} false = refl +⌣S-false {n = zero} true = refl +⌣S-false {n = suc zero} base = refl +⌣S-false {n = suc zero} (loop i) = refl +⌣S-false {n = suc (suc n)} north i = north +⌣S-false {n = suc (suc n)} south i = merid (ptSn (suc (+-zero n i))) i +⌣S-false {n = suc (suc n)} (merid a i) j = + hcomp (λ k → λ {(i = i0) → north + ; (i = i1) → merid (ptSn (suc (+-zero n j))) (j ∨ ~ k) + ; (j = i1) → merid a i}) + (merid (⌣S-false a j) i) + +assoc⌣S : {n m l : ℕ} (x : S₊ n) (y : S₊ m) (z : S₊ l) + → PathP (λ i → S₊ (+-assoc n m l i)) + (x ⌣S (y ⌣S z)) ((x ⌣S y) ⌣S z) +assoc⌣S {n = zero} {m = m} false y z = refl +assoc⌣S {n = zero} {m = m} true y z = sym (IdR⌣S z) +assoc⌣S {n = suc zero} {m = m} base y z = sym (IdR⌣S z) +assoc⌣S {n = suc zero} {m = m} (loop i) y z j = help m y j i + where + help : (m : ℕ) (y : S₊ m) + → Square (σS (y ⌣S z)) (cong (λ w → sphereFun↑ _⌣S_ w z) (σS y)) + (sym (IdR⌣S z)) (sym (IdR⌣S z)) + help zero false = refl + help zero true = σS∙ + help (suc m) y = + rUnit (σS (y ⌣S z)) + ∙ cong (σS (y ⌣S z) ∙_) + (cong sym ((sym σS∙) + ∙ congS σS (sym (IdR⌣S z)))) + ∙ sym (cong-∙ (λ k → sphereFun↑ _⌣S_ k z) + (merid y) + (sym (merid (ptSn (suc m))))) +assoc⌣S {n = suc (suc n)} {m = m} north y z k = north +assoc⌣S {n = suc (suc n)} {m = m} south y z k = north +assoc⌣S {n = suc (suc n)} {m = m} {l} (merid a i) y z j = + help m y (assoc⌣S a y z) j i + where + help : (m : ℕ) (y : S₊ m) + → PathP (λ i₁ → S₊ (suc (+-assoc n m l i₁))) + (a ⌣S (y ⌣S z)) + ((a ⌣S y) ⌣S z) + → SquareP (λ j i → S₊ (+-assoc (suc (suc n)) m l j)) + (σS (a ⌣S (y ⌣S z))) + (λ i → (merid a i ⌣S y) ⌣S z) + (λ _ → north) (λ _ → north) + help zero false _ = + (λ i j → σ (S₊∙ (suc (+-assoc n zero l i))) (lem i) j) + ▷ rUnit _ + ∙ cong₂ _∙_ (congS σS (λ _ → (a ⌣S false) ⌣S z)) + (cong sym (sym σS∙ + ∙ congS σS + (sym (IdR⌣S z)))) + ∙ sym (cong-∙ (_⌣S z) + (merid (a ⌣S false)) + (sym (merid (ptSn (suc (n + zero)))))) + where + lem : PathP (λ i → S₊ (suc (+-assoc n zero l i))) + (a ⌣S z) ((a ⌣S false) ⌣S z) + lem = toPathP ((λ i → subst (S₊ ∘ suc) + (isSetℕ _ _ (+-assoc n zero l) + (λ j → +-zero n (~ j) + l) i) + (a ⌣S z)) + ∙ fromPathP (λ i → ⌣S-false a (~ i) ⌣S z)) + help zero true _ = + (congS σS (IdL⌣S a) ∙ σS∙) + ◁ (λ i j → north) + ▷ (cong (cong (_⌣S z)) + (sym σS∙ + ∙ congS σS (sym (IdL⌣S a)))) + help (suc m) y q = + (λ i j → σS (q i) j) + ▷ (rUnit _ + ∙ cong₂ _∙_ refl + (cong sym (sym σS∙ ∙ cong σS (sym (IdR⌣S z)))) + ∙ sym (cong-∙ (_⌣S z) (merid (a ⌣S y)) + (sym (merid (ptSn (suc (n + suc m))))))) + +-- Goal: graded commutativity + +-- To state it: we'll need iterated negations +-S^ : {k : ℕ} (n : ℕ) → S₊ k → S₊ k +-S^ zero x = x +-S^ (suc n) x = invSphere (-S^ n x) + +-- The folowing is an explicit definition of -S^ (n · m) which is +-- often easier to reason about +-S^-expl : {k : ℕ} (n m : ℕ) + → isEvenT n ⊎ isOddT n + → isEvenT m ⊎ isOddT m + → S₊ k → S₊ k +-S^-expl {k = zero} n m (inl x₁) q x = x +-S^-expl {k = zero} n m (inr x₁) (inl x₂) x = x +-S^-expl {k = zero} n m (inr x₁) (inr x₂) false = true +-S^-expl {k = zero} n m (inr x₁) (inr x₂) true = false +-S^-expl {k = suc zero} n m p q base = base +-S^-expl {k = suc zero} n m (inl x) q (loop i) = loop i +-S^-expl {k = suc zero} n m (inr x) (inl x₁) (loop i) = loop i +-S^-expl {k = suc zero} n m (inr x) (inr x₁) (loop i) = loop (~ i) +-S^-expl {k = suc (suc k)} n m p q north = north +-S^-expl {k = suc (suc k)} n m p q south = north +-S^-expl {k = suc (suc k)} n m (inl x) q (merid a i) = σS a i +-S^-expl {k = suc (suc k)} n m (inr x) (inl x₁) (merid a i) = σS a i +-S^-expl {k = suc (suc k)} n m (inr x) (inr x₁) (merid a i) = σS a (~ i) + +-- invSphere commutes with S^ +invSphere-S^ : {k : ℕ} (n : ℕ) (x : S₊ k) + → invSphere (-S^ n x) ≡ -S^ n (invSphere x) +invSphere-S^ zero x = refl +invSphere-S^ (suc n) x = cong invSphere (invSphere-S^ n x) + +-S^² : {k : ℕ} (n : ℕ) (x : S₊ k) → -S^ n (-S^ n x) ≡ x +-S^² zero x = refl +-S^² (suc n) x = + cong invSphere (sym (invSphere-S^ n (-S^ n x))) + ∙ invSphere² _ (-S^ n (-S^ n x)) + ∙ -S^² n x + +-S^Iso : {k : ℕ} (n : ℕ) → Iso (S₊ k) (S₊ k) +fun (-S^Iso n) = -S^ n +inv (-S^Iso n) = -S^ n +rightInv (-S^Iso n) = -S^² n +leftInv (-S^Iso n) = -S^² n + +-S^-comp : {k : ℕ} (n m : ℕ) (x : S₊ k) + → -S^ n (-S^ m x) ≡ -S^ (n + m) x +-S^-comp zero m x = refl +-S^-comp (suc n) m x = cong invSphere (-S^-comp n m x) + +-S^·2 : {k : ℕ} (n : ℕ) (x : S₊ k) → -S^ (n + n) x ≡ x +-S^·2 zero x = refl +-S^·2 (suc n) x = + cong invSphere (λ i → -S^ (+-comm n (suc n) i) x) + ∙ invSphere² _ (-S^ (n + n) x) + ∙ -S^·2 n x + +-- technical transport lemma +private + -S^-transp : {k : ℕ} (m : ℕ) (p : k ≡ m) (n : ℕ) (x : S₊ k) + → subst S₊ p (-S^ n x) ≡ -S^ n (subst S₊ p x) + -S^-transp = + J> λ n x → transportRefl _ + ∙ sym (cong (-S^ n) (transportRefl x)) + +-- Iterated path inversion +sym^ : ∀ {ℓ} {A : Type ℓ} {x : A} (n : ℕ) → x ≡ x → x ≡ x +sym^ zero p = p +sym^ (suc n) p = sym (sym^ n p) + +-- Interaction between -S^ and sym^ +σS-S^ : {k : ℕ} (n : ℕ) (x : S₊ (suc k)) + → σS (-S^ n x) ≡ sym^ n (σS x) +σS-S^ {k = k} zero x = refl +σS-S^ {k = k} (suc n) x = + σ-invSphere k _ ∙ cong sym (σS-S^ n x) + +sphereFun↑-subst : {n m : ℕ} (k' k : ℕ) (p : k' ≡ k) + → (f : S₊ n → S₊ m → S₊ k') (x : S₊ _) (y : S₊ _) + → sphereFun↑ (λ x y → subst S₊ p (f x y)) x y + ≡ subst S₊ (cong suc p) (sphereFun↑ f x y) +sphereFun↑-subst k' = J> λ f x y + → (λ i → sphereFun↑ (λ x₁ y₁ → transportRefl (f x₁ y₁) i) x y) + ∙ sym (transportRefl _) + +sphereFun↑-invSphere : {n m k : ℕ} + → (f : S₊ (suc n) → S₊ (suc m) → S₊ (suc k)) (x : S₊ _) (y : S₊ _) + → sphereFun↑ (λ x y → invSphere' (f x y)) x y + ≡ invSphere' (sphereFun↑ (λ x y → (f x y)) x y) +sphereFun↑-invSphere {n = n} {m = m} {k = k} f north y = refl +sphereFun↑-invSphere {n = n} {m = m} {k = k} f south y = refl +sphereFun↑-invSphere {n = n} {m = m} {k = k} f (merid a i) y j = + lem k (f a y) j i + where + lem : (k : ℕ) (x : S₊ (suc k)) + → (σS (invSphere' x)) ≡ (cong invSphere' (σS x)) + lem k x = + sym (cong-∙ invSphere' (merid x) (sym (merid (ptSn _))) + ∙∙ cong (cong invSphere' (merid x) ∙_) + (rCancel (merid (ptSn _))) + ∙∙ (sym (rUnit _) + ∙ sym (σ-invSphere k x) + ∙ cong (σ (S₊∙ (suc k))) + (sym (invSphere'≡ x)))) + +sphereFun↑^ : {n m k : ℕ} (l : ℕ) + → (f : S₊ (suc n) → S₊ (suc m) → S₊ (suc k)) (x : S₊ _) (y : S₊ _) + → sphereFun↑ (λ x y → -S^ l (f x y)) x y + ≡ -S^ l (sphereFun↑ (λ x y → (f x y)) x y) +sphereFun↑^ zero f x y = refl +sphereFun↑^ (suc l) f x y = + (λ i → sphereFun↑ (λ x₁ y₁ → invSphere'≡ (-S^ l (f x₁ y₁)) (~ i)) x y) + ∙ sphereFun↑-invSphere (λ x₁ y₁ → (-S^ l (f x₁ y₁))) x y + ∙ invSphere'≡ ((sphereFun↑ (λ x₁ y₁ → -S^ l (f x₁ y₁)) x y)) + ∙ cong invSphere (sphereFun↑^ l f x y) + +S^-even : {k : ℕ} (n : ℕ) (x : S₊ k) → isEvenT n → -S^ n x ≡ x +S^-even zero x p = refl +S^-even (suc (suc n)) x p = invSphere² _ (-S^ n x) ∙ S^-even n x p + +private + move-transp-S^ : {k : ℕ} (n : ℕ) (p : k ≡ n) (m : ℕ) + → (x : S₊ k) (y : S₊ n) + → subst S₊ p (-S^ m x) ≡ y + → subst S₊ (sym p) (-S^ m y) ≡ x + move-transp-S^ = + J> λ m x → J> transportRefl _ + ∙ cong (-S^ m) (transportRefl _) + ∙ -S^² m x + + master-lem : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (coh : refl ≡ p) + (q : p ≡ p) + → Cube (λ j k → coh j (~ k)) (λ j k → coh j (~ k)) + (λ i k → q k i) (λ i k → q i (~ k)) + (λ j k → coh (~ k) j) λ j k → coh (~ k) j + master-lem = J> λ q → λ i j k → sym≡flipSquare q j (~ k) i + + comm⌣S₁ : {m : ℕ} → (x : S¹) (y : S₊ (suc m)) + → (x ⌣S y) + ≡ subst S₊ (+-comm (suc m) 1) + (-S^ (suc m) (y ⌣S x)) + comm⌣S₁ {m = zero} x y = + (main x y ∙ invSphere'≡ (y ⌣S x)) + ∙ sym (transportRefl (invSusp (y ⌣S x))) + where + pp-main : (x : S¹) → PathP (λ i → IdL⌣S {m = 1} x i + ≡ IdL⌣S {m = 1} x i) (cong (x ⌣S_) loop) (sym (σS x)) + pp-main base i j = rCancel (merid base) (~ i) (~ j) + pp-main (loop k) i j = + master-lem _ (sym (rCancel (merid base))) + (λ j k → σS (loop j) k) k i j + + pp-help : (x : S¹) → PathP (λ i → IdL⌣S {m = 1} x i + ≡ IdL⌣S {m = 1} x i) + (cong (x ⌣S_) loop) (cong invSphere' (σS x)) + pp-help x = pp-main x + ▷ (rUnit _ + ∙∙ cong (sym (σS x) ∙_) (sym (rCancel (merid base))) + ∙∙ sym (cong-∙ invSphere' (merid x) (sym (merid base)))) + + main : (x y : S¹) → (x ⌣S y) ≡ invSphere' (y ⌣S x) + main x base = IdL⌣S {m = 1} x + main x (loop i) = flipSquare (pp-help x) i + comm⌣S₁ {m = suc m} x y = + (main-lem x y + ∙ sym (transportRefl (invSphere' (sphereFun↑ (λ x₂ x₃ → x₃ ⌣S x₂) y x))) + ∙ sym (compSubstℕ {A = S₊} (cong suc (sym (+-comm (suc m) 1))) + (+-comm (suc (suc m)) 1) refl + {x = invSphere' (sphereFun↑ (λ x₂ x₃ → x₃ ⌣S x₂) y x)})) + ∙ cong (subst S₊ (+-comm (suc (suc m)) 1)) + (cong (subst S₊ (cong suc (sym (+-comm (suc m) 1)))) + (sym (S^-lem (suc m) (sphereFun↑ (λ x₂ x₃ → x₃ ⌣S x₂) y x))) + ∙ -S^-transp _ (cong suc (sym (+-comm (suc m) 1))) + (suc (suc m) + suc m) + (sphereFun↑ (λ x₂ x₃ → x₃ ⌣S x₂) y x) + ∙ sym (-S^-comp (suc (suc m)) (suc m) + (subst S₊ (cong suc (sym (+-comm (suc m) 1))) + (sphereFun↑ (λ x₂ x₃ → x₃ ⌣S x₂) y x)))) + ∙ cong (subst S₊ (+-comm (suc (suc m)) 1) + ∘ -S^ (suc (suc m))) + ((sym (-S^-transp _ (cong suc (sym (+-comm (suc m) 1))) (suc m) + (sphereFun↑ (λ x₂ x₃ → x₃ ⌣S x₂) y x)) + ∙ cong (subst S₊ (cong suc (sym (+-comm (suc m) 1)))) + (sym (sphereFun↑^ (suc m) + (λ x₂ x₃ → (x₃ ⌣S x₂)) y x)) + ∙ sym (sphereFun↑-subst _ _ (sym (+-comm (suc m) 1)) + (λ x₂ x₃ → (-S^ (suc m) (x₃ ⌣S x₂))) y x)) + ∙ cong (λ (s : S₊ (suc m) → S¹ → S₊ (suc m + 1)) + → sphereFun↑ s y x) + (refl + ∙ sym (funExt λ x → funExt λ y + → sym (move-transp-S^ _ (+-comm (suc m) 1) + (suc m) (x ⌣S y) (y ⌣S x) + (sym (comm⌣S₁ y x))) + ∙ refl))) + where + S^-lem : {k : ℕ} (m : ℕ) (x : S₊ k) + → -S^ (suc m + m) x ≡ invSphere' x + S^-lem m x = + sym (invSphere'≡ (-S^ (m + m) x)) + ∙ cong invSphere' (-S^·2 m x) + + ⌣S-south : (x : S¹) → x ⌣S south ≡ north + ⌣S-south base = refl + ⌣S-south (loop i) j = + (cong σS (sym (merid (ptSn (suc m))) ) + ∙ rCancel (merid (ptSn _))) j i + + PathP-main : (x : S¹) (a : S₊ (suc m)) + → PathP (λ i → IdL⌣S x i ≡ ⌣S-south x i) (cong (x ⌣S_) (merid a)) + (sym (σS (x ⌣S a))) + PathP-main base a j i = rCancel (merid north) (~ j) (~ i) + PathP-main (loop k) a j i = + hcomp (λ r → λ {(i = i0) → rCancel (merid north) j k + ; (i = i1) → compPath-filler' + (cong σS (sym (merid (ptSn (suc m))))) + (rCancel (merid north)) r j k + ; (j = i0) → σS (compPath-filler (merid a) + (sym (merid (ptSn (suc m)))) (~ r) i) k + ; (j = i1) → σS (σS a k) (~ i) + ; (k = i0) → rCancel (merid north) (~ j) (~ i) + ; (k = i1) → rCancel (merid north) (~ j) (~ i)}) + (master-lem _ (sym (rCancel (merid north))) + (λ i k → σS (loop i ⌣S a) k) k j i) + + pp : (x : S¹) (a : S₊ (suc m)) + → PathP (λ i → IdL⌣S x i ≡ ⌣S-south x i) (cong (x ⌣S_) (merid a)) + (cong invSphere' (σS (x ⌣S a))) + pp x a = PathP-main x a + ▷ (rUnit _ + ∙∙ cong (sym (σS (x ⌣S a)) ∙_) (sym (rCancel (merid north))) + ∙∙ sym (cong-∙ invSphere' (merid (x ⌣S a)) (sym (merid north)))) + + main-lem : (x : S¹) (y : S₊ (2 + m)) + → (x ⌣S y) + ≡ invSphere' (sphereFun↑ (λ x₂ x₃ → x₃ ⌣S x₂) y x) + main-lem x north = IdL⌣S x + main-lem x south = ⌣S-south x + main-lem x (merid a i) j = pp x a j i + + comm⌣S-lem : {n m : ℕ} + → ((x : S₊ (suc n)) (y : S₊ (suc (suc m))) + → (x ⌣S y) ≡ subst S₊ (+-comm (suc (suc m)) (suc n)) + (-S^ (suc (suc m) · (suc n)) (y ⌣S x))) + → (((x : S₊ (suc m)) (y : S₊ (suc (suc n))) + → (x ⌣S y) ≡ subst S₊ (+-comm (suc (suc n)) (suc m)) + (-S^ ((suc (suc n)) · (suc m)) (y ⌣S x)))) + → (((x : S₊ (suc n)) (y : S₊ (suc m)) + → (y ⌣S x) ≡ subst S₊ (sym (+-comm (suc m) (suc n))) + (-S^ ((suc n) · (suc m)) (x ⌣S y)))) + → (x : S₊ (suc (suc n))) (y : S₊ (suc (suc m))) + → (x ⌣S y) ≡ subst S₊ (+-comm (suc (suc m)) (suc (suc n))) + (-S^ (suc (suc m) · (suc (suc n))) (y ⌣S x)) + comm⌣S-lem {n = n} {m = m} ind1 ind2 ind3 x y = + cong (λ (s : S₊ (suc n) → S₊ (suc (suc m)) + → S₊ ((suc n) + (suc (suc m)))) → sphereFun↑ s x y) + (funExt (λ x → funExt λ y + → ind1 x y)) + ∙ (sphereFun↑-subst _ _ (+-comm (suc (suc m)) (suc n)) + (λ x y → -S^ (suc (suc m) · suc n) (y ⌣S x)) x y + ∙ cong (subst S₊ (cong suc (+-comm (suc (suc m)) (suc n)))) + (sphereFun↑^ (suc (suc m) · suc n) (λ x y → y ⌣S x) x y + ∙ cong (-S^ (suc (suc m) · suc n)) + (cong (λ (s : S₊ (suc n) → S₊ (suc (suc m)) + → S₊ ((suc (suc m)) + (suc n))) → sphereFun↑ s x y) + (funExt (λ x → funExt λ y → + cong (λ (s : S₊ (suc m) → S₊ (suc n) → S₊ ((suc m) + (suc n))) + → sphereFun↑ s y x) + (funExt λ x + → funExt λ y + → ind3 y x) + ∙ sphereFun↑-subst _ _ (sym (+-comm (suc m) (suc n))) + (λ x y → -S^ (suc n · suc m) (y ⌣S x)) y x + ∙ cong (subst S₊ (cong suc (sym (+-comm (suc m) (suc n))))) + (sphereFun↑^ (suc n · suc m) (λ x y → (y ⌣S x)) y x))) + ∙ sphereFun↑-subst _ _ (sym (cong suc (+-comm (suc m) (suc n)))) + ((λ x₁ x₂ → + (-S^ (suc n · suc m) + (sphereFun↑ (λ x₃ y₁ → y₁ ⌣S x₃) x₂ x₁)))) x y + ∙ cong (subst S₊ (sym (cong (suc ∘ suc) (+-comm (suc m) (suc n))))) + ((sphereFun↑^ (suc n · suc m) + ((λ x₁ x₂ → (sphereFun↑ (λ x₃ y₁ → y₁ ⌣S x₃) x₂ x₁))) x y + ∙ cong (-S^ (suc n · suc m)) (lem₃ x y))))) + ∙ big-lem (suc n) (suc m) + _ (λ i → suc (suc (+-comm (suc m) (suc n) (~ i)))) + _ (λ i → suc (+-comm (suc (suc m)) (suc n) i)) _ + (sym (+-comm (suc (suc m)) (suc (suc n)))) + (λ i → suc (+-comm (suc (suc n)) (suc m) i)) + (sphereFun↑ (λ x₁ y₂ → y₂ ⌣S x₁) y x) + ∙ sym (cong (subst S₊ (+-comm (suc (suc m)) (suc (suc n)))) + (cong (-S^ (suc (suc m) · suc (suc n))) + ((λ i → sphereFun↑ (λ x y → ind2 x y i) y x) + ∙ sphereFun↑-subst _ _ + (+-comm (suc (suc n)) (suc m)) + (λ x y → -S^ (suc (suc n) · suc m) (y ⌣S x)) y x + ∙ cong (subst S₊ (cong suc (+-comm (suc (suc n)) (suc m)))) + (sphereFun↑^ (suc (suc n) · suc m) (λ x y → y ⌣S x) y x))))) + where + ℕ-p : (n m : ℕ) + → (suc m · suc n + suc n · m) + ≡ (m + m) + ((n · m + n · m) + (suc n)) + ℕ-p n m = + cong suc (cong (_+ (m + n · m)) (cong (n +_) (·-comm m (suc n))) + ∙ sym (+-assoc n (m + n · m) _) + ∙ +-comm n _ + ∙ cong (_+ n) (+-assoc (m + n · m) m (n · m) + ∙ cong (_+ (n · m)) + (sym (+-assoc m (n · m) m) + ∙ cong (m +_) (+-comm (n · m) m) + ∙ +-assoc m m (n · m)) + ∙ sym (+-assoc (m + m) (n · m) (n · m)))) + ∙ sym (+-suc (m + m + (n · m + n · m)) n) + ∙ sym (+-assoc (m + m) (n · m + n · m) (suc n)) + + ℕ-p2 : (n m : ℕ) → suc m · n + n · m + 1 ≡ (((n · m) + (n · m)) + (suc n)) + ℕ-p2 n m = (λ _ → ((n + m · n) + n · m) + 1) + ∙ cong (_+ 1) (sym (+-assoc n (m · n) (n · m)) + ∙ (λ i → +-comm n ((·-comm m n i) + n · m) i)) + ∙ sym (+-assoc (n · m + n · m) n 1) + ∙ cong (n · m + n · m +_) (+-comm n 1) + + big-lem : (n m : ℕ) {x : ℕ} (y : ℕ) (p : x ≡ y) (z : ℕ) (s : y ≡ z) + (d : ℕ) (r : z ≡ d) (t : x ≡ d) + (a : S₊ x) + → subst S₊ s (-S^ (suc m · n) (subst S₊ p (-S^ (n · m) (invSphere' a)))) + ≡ subst S₊ (sym r) + (-S^ (suc m · suc n) + (subst S₊ t (-S^ (suc n · m) a))) + big-lem n m = + J> (J> (J> λ t a + → transportRefl _ + ∙ cong (-S^ (n + m · n)) (transportRefl _) + ∙ sym (transportRefl _ + ∙ cong (-S^ (suc m · suc n)) + ((λ i → subst S₊ (isSetℕ _ _ t refl i) (-S^ (m + n · m) a)) + ∙ transportRefl (-S^ (m + n · m) a) ) + ∙ -S^-comp (suc m · suc n) (suc n · m) a + ∙ ((funExt⁻ (cong -S^ (ℕ-p n m)) a + ∙ (sym (-S^-comp (m + m) _ a) + ∙ -S^·2 m (-S^ (n · m + n · m + suc n) a)) + ∙ funExt⁻ (cong -S^ (sym (ℕ-p2 n m))) a) + ∙ sym (-S^-comp (suc m · n + n · m) 1 a) + ∙ cong (-S^ (suc m · n + n · m)) + (sym (invSphere'≡ a))) + ∙ sym (-S^-comp (suc m · n) (n · m) (invSphere' a)) ))) + + lem₁ : (x : S₊ (2 + n)) + → sphereFun↑ (λ x₂ x₃ → sphereFun↑ (λ x₄ y₁ → y₁ ⌣S x₄) x₃ x₂) x north ≡ north + lem₁ north = refl + lem₁ south = refl + lem₁ (merid a i) j = rCancel (merid north) j i + + lem₂ : (x : S₊ (2 + n)) + → sphereFun↑ (λ x₂ x₃ + → sphereFun↑ (λ x₄ y₁ → y₁ ⌣S x₄) x₃ x₂) x south + ≡ north + lem₂ north = refl + lem₂ south = refl + lem₂ (merid a i) j = rCancel (merid north) j i + + lem₃ : (x : S₊ (2 + n)) (y : S₊ (2 + m)) + → (sphereFun↑ (λ x₁ x₂ + → sphereFun↑ (λ x₃ y₁ → y₁ ⌣S x₃) x₂ x₁) x y) + ≡ invSphere' (sphereFun↑ (λ x₁ y₂ → y₂ ⌣S x₁) y x) + lem₃ x north = lem₁ x + lem₃ x south = lem₂ x + lem₃ x (merid a i) j = h j i + where + main : (x : _) → PathP (λ i → lem₁ x i ≡ lem₂ x i) + (cong (sphereFun↑ (λ x₂ x₃ + → sphereFun↑ (λ x₄ y₁ → y₁ ⌣S x₄) x₃ x₂) x ) + (merid a)) + (sym (σS (x ⌣S a))) + main north = cong sym (sym (rCancel (merid north))) + main south = cong sym (sym (rCancel (merid north))) + main (merid z i) j k = + master-lem _ + (sym (rCancel (merid north))) + (cong (λ x → σS (x ⌣S a)) (merid z)) i j k + + h : PathP (λ i → lem₁ x i ≡ lem₂ x i) + (cong (sphereFun↑ + (λ x₂ x₃ → sphereFun↑ (λ x₄ y₁ → y₁ ⌣S x₄) x₃ x₂) x) + (merid a)) (cong invSphere' (σS (x ⌣S a))) + h = main x + ▷ ((rUnit _ ∙ cong (sym (σS (x ⌣S a)) ∙_) + (sym (rCancel (merid north)))) + ∙ sym (cong-∙ invSphere' + (merid (x ⌣S a)) (sym (merid (ptSn _))))) + + comm⌣S₀ : (x : Bool) (m : ℕ) (y : S₊ m) + → PathP (λ i → S₊ (+-zero m (~ i))) (x ⌣S y) (y ⌣S x) + comm⌣S₀ false = + elim+2 (λ { false → refl ; true → refl}) + (λ { base → refl ; (loop i) → refl}) + ind + where + ind : (n : ℕ) → + ((y : S₊ (suc n)) → + PathP (λ i → S₊ (suc (+-zero n (~ i)))) y (y ⌣S false)) → + (y : Susp (S₊ (suc n))) → + PathP (λ i → Susp (S₊ (suc (+-zero n (~ i))))) y (y ⌣S false) + ind n p north i = north + ind n p south i = merid (ptSn (suc (+-zero n (~ i)))) (~ i) + ind n p (merid a j) i = + comp (λ k → Susp (S₊ (suc (+-zero n (~ i ∨ ~ k))))) + (λ r → + λ {(i = i0) → merid a j + ; (i = i1) → + σ (S₊∙ (suc (+-zero n (~ r)))) (p a r) j + ; (j = i0) → north + ; (j = i1) → merid (ptSn (suc (+-zero n (~ i ∨ ~ r)))) (~ i)}) + (compPath-filler (merid a) (sym (merid (ptSn _))) i j) + comm⌣S₀ true m y = (λ i → ptSn (+-zero m (~ i))) ▷ (sym (IdL⌣S y)) + +-- Graded commutativity +comm⌣S : {n m : ℕ} (x : S₊ n) (y : S₊ m) + → (x ⌣S y) ≡ subst S₊ (+-comm m n) (-S^ (m · n) (y ⌣S x)) +comm⌣S {n = zero} {m = m} x y = + sym (fromPathP (symP {A = λ i → S₊ (+-zero m i)} ((comm⌣S₀ x m y)))) + ∙ sym (cong (subst S₊ (+-zero m)) + ((λ i → -S^ (0≡m·0 m (~ i)) (y ⌣S x)))) +comm⌣S {n = suc n} {m = zero} x y = + sym (fromPathP (comm⌣S₀ y (suc n) x)) + ∙ (λ i → subst S₊ (isSetℕ _ _ + (sym (+-comm (suc n) zero)) + (+-comm zero (suc n)) i) (y ⌣S x)) +comm⌣S {n = suc zero} {m = suc m} x y = + comm⌣S₁ x y + ∙ cong (subst S₊ (+-comm (suc m) 1)) + λ i → -S^ (·-identityʳ (suc m) (~ i)) (y ⌣S x) +comm⌣S {n = suc (suc n)} {m = suc zero} x y = + sym (substSubst⁻ S₊ (+-comm 1 (suc (suc n))) (x ⌣S y)) + ∙ cong (subst S₊ (+-comm 1 (suc (suc n)))) + ((λ i → subst S₊ (isSetℕ _ _ + (sym (+-comm 1 (suc (suc n)))) (+-comm (suc (suc n)) 1) i) + (x ⌣S y)) + ∙ (sym (sym + (-S^-transp _ (+-comm (suc (suc n)) 1) (1 · suc (suc n)) + (-S^ (suc (suc n)) (x ⌣S y))) + ∙ cong (subst S₊ (+-comm (suc (suc n)) 1)) + (cong (-S^ (1 · suc (suc n))) + (λ i → -S^ (·-identityˡ (suc (suc n)) (~ i)) (x ⌣S y)) + ∙ -S^² (1 · suc (suc n)) (x ⌣S y))))) + ∙ sym (cong (subst S₊ (+-comm 1 (suc (suc n))) + ∘ -S^ (1 · suc (suc n))) (comm⌣S₁ y x)) +comm⌣S {n = suc (suc n)} {m = suc (suc m)} x y = + comm⌣S-lem comm⌣S comm⌣S + (λ x y → (sym (cong (subst S₊ (sym (+-comm (suc m) (suc n)))) + (sym (-S^-transp _ (+-comm (suc m) (suc n)) + (suc n · suc m) (-S^ (suc m · suc n) (y ⌣S x))) + ∙ cong (subst S₊ (+-comm (suc m) (suc n))) + (cong (-S^ (suc n · suc m)) + (λ i → -S^ (·-comm (suc m) (suc n) i) (y ⌣S x)) + ∙ -S^² (suc n · suc m) (y ⌣S x) )) + ∙ subst⁻Subst S₊ (+-comm (suc m) (suc n)) (y ⌣S x) )) + ∙ sym (cong (subst S₊ (sym (+-comm (suc m) (suc n))) + ∘ -S^ (suc n · suc m)) + (comm⌣S x y))) x y diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index a3085a2f82..7eb740a4c0 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -20,6 +20,8 @@ open import Cubical.HITs.Sn.Base open import Cubical.HITs.Susp open import Cubical.HITs.Truncation open import Cubical.HITs.PropositionalTruncation as PT hiding (rec ; elim) +open import Cubical.HITs.SmashProduct.Base +open import Cubical.HITs.Pushout.Base 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) @@ -35,6 +37,13 @@ open Iso σSn zero true = refl σSn (suc n) x = toSusp (S₊∙ (suc n)) x +σS : {n : ℕ} → S₊ n → Path (S₊ (suc n)) (ptSn _) (ptSn _) +σS {n = n} = σSn n + +σS∙ : {n : ℕ} → σS (ptSn n) ≡ refl +σS∙ {n = zero} = refl +σS∙ {n = suc n} = rCancel (merid (ptSn (suc n))) + IsoSucSphereSusp : (n : ℕ) → Iso (S₊ (suc n)) (Susp (S₊ n)) IsoSucSphereSusp zero = S¹IsoSuspBool IsoSucSphereSusp (suc n) = idIso @@ -44,6 +53,11 @@ IsoSucSphereSusp∙ : (n : ℕ) IsoSucSphereSusp∙ zero = refl IsoSucSphereSusp∙ (suc n) = refl +IsoSucSphereSusp∙' : (n : ℕ) + → Iso.fun (IsoSucSphereSusp n) (ptSn (suc n)) ≡ north +IsoSucSphereSusp∙' zero = refl +IsoSucSphereSusp∙' (suc n) = refl + suspFunS∙ : {n : ℕ} → (S₊ n → S₊ n) → S₊∙ (suc n) →∙ S₊∙ (suc n) suspFunS∙ {n = zero} f = (λ x → Iso.inv S¹IsoSuspBool (suspFun f (Iso.fun S¹IsoSuspBool x))) , refl @@ -643,6 +657,22 @@ invSphere {n = zero} = not invSphere {n = (suc zero)} = invLooper invSphere {n = (suc (suc n))} = invSusp +-- sometimes also this version is useful +invSphere' : {n : ℕ} → S₊ n → S₊ n +invSphere' {n = zero} = not +invSphere' {n = (suc zero)} = invLooper +invSphere' {n = suc (suc n)} north = north +invSphere' {n = suc (suc n)} south = north +invSphere' {n = suc (suc n)} (merid a i) = σSn (suc n) a (~ i) + +invSphere'≡ : {n : ℕ} → (x : S₊ n) → invSphere' x ≡ invSphere x +invSphere'≡ {n = zero} x = refl +invSphere'≡ {n = suc zero} x = refl +invSphere'≡ {n = suc (suc n)} north = merid (ptSn _) +invSphere'≡ {n = suc (suc n)} south = refl +invSphere'≡ {n = suc (suc n)} (merid a i) j = + compPath-filler (merid a) (sym (merid (ptSn _))) (~ j) (~ i) + invSphere² : (n : ℕ) (x : S₊ n) → invSphere (invSphere x) ≡ x invSphere² zero = notnot invSphere² (suc zero) base = refl diff --git a/Cubical/HITs/Susp/Base.agda b/Cubical/HITs/Susp/Base.agda index 184d4a5b03..da71aa98ef 100644 --- a/Cubical/HITs/Susp/Base.agda +++ b/Cubical/HITs/Susp/Base.agda @@ -37,11 +37,19 @@ suspFun f south = south suspFun f (merid a i) = merid (f a) i -- pointed version -suspFun∙ : {A B : Type ℓ} (f : A → B) +suspFun∙ : {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Susp∙ A →∙ Susp∙ B fst (suspFun∙ f) = suspFun f snd (suspFun∙ f) = refl +suspFun↑ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (b : B) + → ((a : A) → Path B b b) + → Susp A → B +suspFun↑ b f north = b +suspFun↑ b f south = b +suspFun↑ b f (merid a i) = f a i + UnitIsoSuspUnit : Iso Unit (Susp Unit) fun UnitIsoSuspUnit _ = north inv UnitIsoSuspUnit _ = tt diff --git a/Cubical/Homotopy/Group/Join.agda b/Cubical/Homotopy/Group/Join.agda new file mode 100644 index 0000000000..e3015ab4c7 --- /dev/null +++ b/Cubical/Homotopy/Group/Join.agda @@ -0,0 +1,326 @@ +{-# OPTIONS --safe --lossy-unification #-} +{- +This file contains definition of homotopy groups in terms of joins: +π*ₙₘ(A) := ∥ Sⁿ * Sᵐ →∙ A ∥₀ +and the fact that it agrees with the usual definition of homotopy groups. +-} +module Cubical.Homotopy.Group.Join 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.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Bool + +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.S1 +open import Cubical.HITs.Join +open import Cubical.HITs.Sn.Multiplication + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.GroupPath + +open Iso +open GroupStr + +-- Standard loop in Ω (join A B) +ℓ* : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → fst A → fst B → Ω (join∙ A B) .fst +ℓ* A B a b = push (pt A) (pt B) + ∙ (push a (pt B) ⁻¹ ∙∙ push a b ∙∙ (push (pt A) b ⁻¹)) + +ℓS : {n m : ℕ} → S₊ n → S₊ m → Ω (join∙ (S₊∙ n) (S₊∙ m)) .fst +ℓS {n = n} {m} = ℓ* (S₊∙ n) (S₊∙ m) + +-- Addition of functions join A B → C +_+*_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (f g : join∙ A B →∙ C) → join∙ A B →∙ C +fst (_+*_ {C = C} f g) (inl x) = pt C +fst (_+*_ {C = C} f g) (inr x) = pt C +fst (_+*_ {A = A} {B = B} f g) (push a b i) = + (Ω→ f .fst (ℓ* A B a b) ∙ Ω→ g .fst (ℓ* A B a b)) i +snd (f +* g) = refl + +-- Inversion +-* : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (f : join∙ A B →∙ C) → join∙ A B →∙ C +fst (-* {C = C} f) (inl x) = pt C +fst (-* {C = C} f) (inr x) = pt C +fst (-* {A = A} {B} f) (push a b i) = Ω→ f .fst (ℓ* A B a b) (~ i) +snd (-* f) = refl + +-- -Π is the same as -* +-Π≡-* : ∀ {ℓ} {A : Pointed ℓ} {n m : ℕ} + (f : S₊∙ (suc (n + m)) →∙ A) + → (-Π f) ∘∙ join→Sphere∙ n m + ≡ -* (f ∘∙ join→Sphere∙ n m) +fst (-Π≡-* f i) (inl x) = snd (-Π f) i +fst (-Π≡-* f i) (inr x) = snd (-Π f) i +fst (-Π≡-* {A = A} {n = n} {m = m} f i) (push a b j) = main i j + where + lem : (n : ℕ) (f : S₊∙ (suc n) →∙ A) (a : S₊ n) + → Square (cong (fst (-Π f)) (σS a)) + (sym (snd f) ∙∙ cong (fst f) (sym (σS a)) ∙∙ snd f) + (snd (-Π f)) (snd (-Π f)) + lem zero f false = + doubleCompPath-filler (sym (snd f)) (cong (fst f) (sym loop)) (snd f) + lem zero f true = doubleCompPath-filler (sym (snd f)) refl (snd f) + lem (suc n) f a = + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn (suc n)))) + ∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel _)) + ∙ sym (rUnit _)) + ◁ doubleCompPath-filler + (sym (snd f)) (cong (fst f) (sym (σS a))) (snd f) + + main : Square (cong (fst (-Π f)) (σS (a ⌣S b))) + (sym (Ω→ (f ∘∙ join→Sphere∙ n m) .fst (ℓS a b))) + (snd (-Π f)) (snd (-Π f)) + main = lem _ f (a ⌣S b) + ▷ sym ((λ j → (sym (lUnit (snd f) (~ j)) + ∙∙ sym (cong (fst f) (cong-∙ (join→Sphere n m) + (push (ptSn n) (ptSn m)) + ((push a (ptSn m) ⁻¹) + ∙∙ push a b + ∙∙ (push (ptSn n) b ⁻¹)) j)) + ∙∙ lUnit (snd f) (~ j))) + ∙ cong (sym (snd f) ∙∙_∙∙ snd f) + (cong (cong (fst f)) + (congS sym + ((cong₂ _∙_ + (cong σS (IdL⌣S _) ∙ σS∙) + (cong-∙∙ (join→Sphere n m) + (push a (ptSn m) ⁻¹) (push a b) (push (ptSn n) b ⁻¹) + ∙ (cong₂ (λ p q → p ⁻¹ ∙∙ σS (a ⌣S b) ∙∙ q ⁻¹) + (cong σS (IdL⌣S _) ∙ σS∙) + (cong σS (IdR⌣S _) ∙ σS∙)) + ∙ sym (rUnit (σS (a ⌣S b))))) + ∙ sym (lUnit _))))) +snd (-Π≡-* f i) j = lem i j + where + lem : Square (refl ∙ snd (-Π f)) refl (snd (-Π f)) refl + lem = sym (lUnit (snd (-Π f))) ◁ λ i j → (snd (-Π f)) (i ∨ j) + +-- ·Π is the same as +* +·Π≡+* : ∀ {ℓ} {A : Pointed ℓ} {n m : ℕ} + (f g : S₊∙ (suc (n + m)) →∙ A) + → (∙Π f g ∘∙ join→Sphere∙ n m) + ≡ ((f ∘∙ join→Sphere∙ n m) + +* (g ∘∙ join→Sphere∙ n m)) +fst (·Π≡+* {A = A} f g i) (inl x) = snd (∙Π f g) i +fst (·Π≡+* {A = A} f g i) (inr x) = snd (∙Π f g) i +fst (·Π≡+* {A = A} {n = n} {m} f g i) (push a b j) = main i j + where + help : (n : ℕ) (f g : S₊∙ (suc n) →∙ A) (a : S₊ n) + → Square (cong (fst (∙Π f g)) (σS a)) + (Ω→ f .fst (σS a) ∙ Ω→ g .fst (σS a)) + (snd (∙Π f g)) (snd (∙Π f g)) + help zero f g false = refl + help zero f g true = + rUnit refl + ∙ cong₂ _∙_ (sym (∙∙lCancel (snd f))) (sym (∙∙lCancel (snd g))) + help (suc n) f g a = + cong-∙ (fst (∙Π f g)) (merid a) (sym (merid (ptSn (suc n)))) + ∙ cong (cong (fst (∙Π f g)) (merid a) ∙_) + (congS sym + (cong₂ _∙_ + (cong (sym (snd f) ∙∙_∙∙ snd f) + (cong (cong (fst f)) (rCancel (merid (ptSn (suc n))))) + ∙ Ω→ f .snd) + (cong (sym (snd g) ∙∙_∙∙ snd g) + (cong (cong (fst g)) (rCancel (merid (ptSn (suc n))))) + ∙ Ω→ g .snd) + ∙ sym (rUnit refl))) + ∙ sym (rUnit _) + + Ω→σ : ∀ {ℓ} {A : Pointed ℓ} (f : S₊∙ (suc (n + m)) →∙ A) + → Ω→ f .fst (σS (a ⌣S b)) + ≡ Ω→ (f ∘∙ join→Sphere∙ n m) .fst (ℓS a b) + Ω→σ f = + cong (sym (snd f) ∙∙_∙∙ snd f) + (cong (cong (fst f)) + (sym main)) + ∙ λ i → Ω→ (lem (~ i)) .fst (ℓS a b) + where + main : cong (join→Sphere n m) (ℓS a b) ≡ σS (a ⌣S b) + main = cong-∙ (join→Sphere n m) _ _ + ∙ cong₂ _∙_ + (cong σS (IdL⌣S _) ∙ σS∙) + (cong-∙∙ (join→Sphere n m) _ _ _ + ∙ (λ i → sym ((cong σS (IdL⌣S a) ∙ σS∙) i) + ∙∙ σS (a ⌣S b) + ∙∙ sym ((cong σS (IdR⌣S b) ∙ σS∙) i)) + ∙ sym (rUnit (σS (a ⌣S b))) ) + ∙ sym (lUnit (σS (a ⌣S b))) + lem : f ∘∙ join→Sphere∙ n m ≡ (fst f ∘ join→Sphere n m , snd f) + lem = ΣPathP (refl , (sym (lUnit _))) + + main : Square (cong (fst (∙Π f g)) (σS (a ⌣S b))) + (Ω→ (f ∘∙ join→Sphere∙ n m) .fst (ℓS a b) + ∙ Ω→ (g ∘∙ join→Sphere∙ n m) .fst (ℓS a b)) + (snd (∙Π f g)) (snd (∙Π f g)) + main = help _ f g (a ⌣S b) ▷ cong₂ _∙_ (Ω→σ f) (Ω→σ g) +snd (·Π≡+* {A = A} f g i) j = lem i j + where + lem : Square (refl ∙ snd (∙Π f g)) refl (snd (∙Π f g)) refl + lem = sym (lUnit (snd (∙Π f g))) ◁ λ i j → (snd (∙Π f g)) (i ∨ j) + +-- Homotopy groups in terms of joins +π* : ∀ {ℓ} (n m : ℕ) (A : Pointed ℓ) → Type ℓ +π* n m A = ∥ join∙ (S₊∙ n) (S₊∙ m) →∙ A ∥₂ + +-- multiplication +·π* : ∀ {ℓ} {n m : ℕ} {A : Pointed ℓ} (f g : π* n m A) → π* n m A +·π* = ST.rec2 squash₂ λ f g → ∣ f +* g ∣₂ + +-π* : ∀ {ℓ} {n m : ℕ} {A : Pointed ℓ} (f : π* n m A) → π* n m A +-π* = ST.map -* + +1π* : ∀ {ℓ} {n m : ℕ} {A : Pointed ℓ} → π* n m A +1π* = ∣ const∙ _ _ ∣₂ + +Iso-JoinMap-SphereMap : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → Iso (join∙ (S₊∙ n) (S₊∙ m) →∙ A) + (S₊∙ (suc (n + m)) →∙ A) +Iso-JoinMap-SphereMap n m = post∘∙equiv (joinSphereEquiv∙ n m) + +Iso-π*-π' : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → Iso ∥ (join∙ (S₊∙ n) (S₊∙ m) →∙ A) ∥₂ + ∥ (S₊∙ (suc (n + m)) →∙ A) ∥₂ +Iso-π*-π' n m = setTruncIso (Iso-JoinMap-SphereMap n m) + +private + J≃∙map : ∀ {ℓ ℓ' ℓ''} {A1 A2 : Pointed ℓ} {A : Pointed ℓ'} + → (e : A1 ≃∙ A2) {P : A1 →∙ A → Type ℓ''} + → ((f : A2 →∙ A) → P (f ∘∙ ≃∙map e)) + → (f : _) → P f + J≃∙map {ℓ'' = ℓ''} {A2 = A2} {A = A} = + Equiv∙J (λ A1 e → {P : A1 →∙ A → Type ℓ''} + → ((f : A2 →∙ A) → P (f ∘∙ ≃∙map e)) + → (f : _) → P f) + λ {P} ind f + → subst P (ΣPathP (refl , sym (lUnit (snd f)))) (ind f) + +π*≡π' : ∀ {ℓ} {A : Pointed ℓ} {n m : ℕ} + (f g : π* n m A) + → Iso.fun (Iso-π*-π' n m) (·π* f g) + ≡ ·π' _ (Iso.fun (Iso-π*-π' n m) f) (Iso.fun (Iso-π*-π' n m) g) +π*≡π' {A = A} {n} {m} = ST.elim2 (λ _ _ → isSetPathImplicit) + (J≃∙map (joinSphereEquiv∙ n m) + λ f → J≃∙map (joinSphereEquiv∙ n m) + λ g → cong ∣_∣₂ + (cong (fun (Iso-JoinMap-SphereMap n m)) (sym (·Π≡+* f g)) + ∙ ∘∙-assoc _ _ _ + ∙ cong (∙Π f g ∘∙_) ret + ∙ ∘∙-idˡ (∙Π f g) + ∙ cong₂ ∙Π + ((sym (∘∙-idˡ f) ∙ cong (f ∘∙_) (sym ret)) ∙ sym (∘∙-assoc _ _ _)) + (sym (∘∙-idˡ g) ∙ cong (g ∘∙_) (sym ret) ∙ sym (∘∙-assoc _ _ _)))) + where + ret = ≃∙→ret/sec∙ {B = _ , ptSn (suc (n + m))} + (joinSphereEquiv∙ n m) .snd + +-π*≡-π' : ∀ {ℓ} {A : Pointed ℓ} {n m : ℕ} + (f : π* n m A) + → Iso.fun (Iso-π*-π' n m) (-π* f) + ≡ -π' _ (Iso.fun (Iso-π*-π' n m) f) +-π*≡-π' {n = n} {m} = + ST.elim (λ _ → isSetPathImplicit) + (J≃∙map (joinSphereEquiv∙ n m) + λ f → cong ∣_∣₂ + (cong (_∘∙ (≃∙map (invEquiv∙ (joinSphereEquiv∙ n m)))) + (sym (-Π≡-* f)) + ∙ ∘∙-assoc _ _ _ + ∙ cong (-Π f ∘∙_) ret + ∙ ∘∙-idˡ (-Π f) + ∙ cong -Π (sym (∘∙-assoc _ _ _ ∙ cong (f ∘∙_) ret ∙ ∘∙-idˡ f)))) + where + ret = ≃∙→ret/sec∙ {B = _ , ptSn (suc (n + m))} + (joinSphereEquiv∙ n m) .snd + +-- Homotopy groups in terms of joins +π*Gr : ∀ {ℓ} (n m : ℕ) (A : Pointed ℓ) → Group ℓ +fst (π*Gr n m A) = π* n m A +1g (snd (π*Gr n m A)) = 1π* +GroupStr._·_ (snd (π*Gr n m A)) = ·π* +inv (snd (π*Gr n m A)) = -π* +isGroup (snd (π*Gr n m A)) = + transport (λ i → IsGroup (p1 (~ i)) (p2 (~ i)) (p3 (~ i))) + (isGroup (π'Gr (n + m) A .snd)) + where + p1 : PathP (λ i → isoToPath (Iso-π*-π' {A = A} n m) i) + 1π* (1π' (suc (n + m))) + p1 = toPathP (cong ∣_∣₂ (transportRefl _ ∙ ΣPathP (refl , sym (rUnit refl)))) + + p2 : PathP (λ i → (f g : isoToPath (Iso-π*-π' {A = A} n m) i) + → isoToPath (Iso-π*-π' {A = A} n m) i) + ·π* (·π' _) + p2 = toPathP (funExt λ f + → funExt λ g → transportRefl _ + ∙ π*≡π' _ _ + ∙ cong₂ (·π' (n + m)) + (Iso.rightInv (Iso-π*-π' n m) _ ∙ transportRefl f) + (Iso.rightInv (Iso-π*-π' n m) _ ∙ transportRefl g)) + + p3 : PathP (λ i → isoToPath (Iso-π*-π' {A = A} n m) i + → isoToPath (Iso-π*-π' {A = A} n m) i) + -π* (-π' _) + p3 = toPathP (funExt λ f → transportRefl _ + ∙ -π*≡-π' _ + ∙ cong (-π' (n + m)) + (Iso.rightInv (Iso-π*-π' n m) _ ∙ transportRefl f)) + +-- Homotopy groups in terms of joins agrees with usual definition +π*Gr≅π'Gr : ∀ {ℓ} (n m : ℕ) (A : Pointed ℓ) + → GroupIso (π*Gr n m A) (π'Gr (n + m) A) +fst (π*Gr≅π'Gr n m A) = Iso-π*-π' {A = A} n m +snd (π*Gr≅π'Gr n m A) = makeIsGroupHom π*≡π' + +-- Functoriality +π*∘∙fun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + (n m : ℕ) (f : A →∙ B) + → π* n m A → π* n m B +π*∘∙fun n m f = ST.map (f ∘∙_) + +π*∘∙Hom : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + (n m : ℕ) (f : A →∙ B) + → GroupHom (π*Gr n m A) (π*Gr n m B) +fst (π*∘∙Hom {A = A} {B = B} n m f) = π*∘∙fun n m f +snd (π*∘∙Hom {A = A} {B = B} n m f) = + subst (λ ϕ → IsGroupHom (π*Gr n m A .snd) ϕ (π*Gr n m B .snd)) + π*∘∙Hom'≡ + (snd π*∘∙Hom') + where + GroupHomπ≅π*PathP : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') (n m : ℕ) + → GroupHom (π'Gr (n + m) A) (π'Gr (n + m) B) + ≡ GroupHom (π*Gr n m A) (π*Gr n m B) + GroupHomπ≅π*PathP A B n m i = + GroupHom (fst (GroupPath _ _) (GroupIso→GroupEquiv (π*Gr≅π'Gr n m A)) (~ i)) + (fst (GroupPath _ _) (GroupIso→GroupEquiv (π*Gr≅π'Gr n m B)) (~ i)) + + π*∘∙Hom' : _ + π*∘∙Hom' = transport (λ i → GroupHomπ≅π*PathP A B n m i) + (π'∘∙Hom (n + m) f) + + π*∘∙Hom'≡ : π*∘∙Hom' .fst ≡ π*∘∙fun n m f + π*∘∙Hom'≡ = + funExt (ST.elim (λ _ → isSetPathImplicit) + λ g → cong ∣_∣₂ (cong (inv (Iso-JoinMap-SphereMap n m)) + (transportRefl _ + ∙ cong (f ∘∙_) (transportRefl _)) + ∙ ∘∙-assoc _ _ _ + ∙ cong (f ∘∙_ ) + (∘∙-assoc _ _ _ ∙ cong (g ∘∙_) + (≃∙→ret/sec∙ {B = _ , ptSn (suc (n + m))} + (joinSphereEquiv∙ n m) .fst) + ∙ ∘∙-idˡ g))) diff --git a/Cubical/Homotopy/Group/Pi4S3/BrunerieExperiments.agda b/Cubical/Homotopy/Group/Pi4S3/BrunerieExperiments.agda index c65fba931e..63a49ba510 100644 --- a/Cubical/Homotopy/Group/Pi4S3/BrunerieExperiments.agda +++ b/Cubical/Homotopy/Group/Pi4S3/BrunerieExperiments.agda @@ -21,7 +21,7 @@ open import Cubical.Data.Int open import Cubical.HITs.S1 hiding (encode) open import Cubical.HITs.S2 open import Cubical.HITs.S3 -open import Cubical.HITs.Join +open import Cubical.HITs.Join hiding (join∙) open import Cubical.HITs.SetTruncation as SetTrunc open import Cubical.HITs.GroupoidTruncation as GroupoidTrunc open import Cubical.HITs.2GroupoidTruncation as 2GroupoidTrunc diff --git a/Cubical/Homotopy/Whitehead.agda b/Cubical/Homotopy/Whitehead.agda index 99abb89222..2daec9e530 100644 --- a/Cubical/Homotopy/Whitehead.agda +++ b/Cubical/Homotopy/Whitehead.agda @@ -20,6 +20,7 @@ open import Cubical.HITs.Wedge open import Cubical.HITs.SetTruncation open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Loopspace open Iso open 3x3-span @@ -408,3 +409,105 @@ module _ (A B : Type) (a₀ : A) (b₀ : B) where ∙∙ 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) + +-- Generalised Whitehead products +module _ {ℓ ℓ' ℓ''} {A : Pointed ℓ} + {B : Pointed ℓ'} {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) (g : Susp∙ (typ B) →∙ C) where + + _·w_ : join∙ A B →∙ C + fst _·w_ (inl x) = pt C + fst _·w_ (inr x) = pt C + fst _·w_ (push a b i) = (Ω→ g .fst (σ B b) ∙ Ω→ f .fst (σ A a)) i + snd _·w_ = refl + + -- The generalised Whitehead product vanishes under suspension + isConst-Susp·w : suspFun∙ (_·w_ .fst) ≡ const∙ _ _ + isConst-Susp·w = Susp·w∙ + ∙ cong suspFun∙ (cong fst isConst-const*) + ∙ ΣPathP ((suspFunConst (pt C)) , refl) + where + const* : join∙ A B →∙ C + fst const* (inl x) = pt C + fst const* (inr x) = pt C + fst const* (push a b i) = + (Ω→ f .fst (σ A a) ∙ Ω→ g .fst (σ B b)) i + snd const* = refl + + isConst-const* : const* ≡ const∙ _ _ + fst (isConst-const* i) (inl x) = Ω→ f .fst (σ A x) i + fst (isConst-const* i) (inr x) = Ω→ g .fst (σ B x) (~ i) + fst (isConst-const* i) (push a b j) = + compPath-filler'' (Ω→ f .fst (σ A a)) (Ω→ g .fst (σ B b)) (~ i) j + snd (isConst-const* i) j = + (cong (Ω→ f .fst) (rCancel (merid (pt A))) ∙ Ω→ f .snd) j i + + Susp·w : suspFun (fst _·w_) ≡ suspFun (fst const*) + Susp·w i north = north + Susp·w i south = south + Susp·w i (merid (inl x) j) = merid (pt C) j + Susp·w i (merid (inr x) j) = merid (pt C) j + Susp·w i (merid (push a b k) j) = + hcomp (λ r → + λ {(i = i0) → fill₁ k (~ r) j + ; (i = i1) → fill₂ k (~ r) j + ; (j = i0) → north + ; (j = i1) → merid (pt C) r + ; (k = i0) → compPath-filler (merid (snd C)) (merid (pt C) ⁻¹) (~ r) j + ; (k = i1) → compPath-filler (merid (snd C)) (merid (pt C) ⁻¹) (~ r) j}) + (hcomp (λ r → + λ {(i = i0) → doubleCompPath-filler + (sym (rCancel (merid (pt C)))) + (λ k → fill₁ k i1) + (rCancel (merid (pt C))) (~ r) k j + ; (i = i1) → doubleCompPath-filler + (sym (rCancel (merid (pt C)))) + (λ k → fill₂ k i1) + (rCancel (merid (pt C))) (~ r) k j + ; (j = i0) → north + ; (j = i1) → north + ; (k = i0) → rCancel (merid (pt C)) (~ r) j + ; (k = i1) → rCancel (merid (pt C)) (~ r) j}) + (main i k j)) + where + F : Ω C .fst → (Ω^ 2) (Susp∙ (fst C)) .fst + F p = sym (rCancel (merid (pt C))) + ∙∙ cong (σ C) p + ∙∙ rCancel (merid (pt C)) + + F-hom : (p q : _) → F (p ∙ q) ≡ F p ∙ F q + F-hom p q = + cong (sym (rCancel (merid (pt C))) + ∙∙_∙∙ rCancel (merid (pt C))) + (cong-∙ (σ C) p q) + ∙ doubleCompPath≡compPath (sym (rCancel (merid (pt C)))) _ _ + ∙ cong (sym (rCancel (merid (pt C))) ∙_) + (sym (assoc _ _ _)) + ∙ assoc _ _ _ + ∙ (λ i → (sym (rCancel (merid (pt C))) + ∙ compPath-filler (cong (σ C) p) (rCancel (merid (pt C))) i) + ∙ compPath-filler' (sym (rCancel (merid (pt C)))) + (cong (σ C) q ∙ rCancel (merid (pt C))) i) + ∙ cong₂ _∙_ (sym (doubleCompPath≡compPath _ _ _)) + (sym (doubleCompPath≡compPath _ _ _)) + + main : F ((Ω→ g .fst (σ B b) ∙ Ω→ f .fst (σ A a))) + ≡ F ((Ω→ f .fst (σ A a) ∙ Ω→ g .fst (σ B b))) + main = F-hom (Ω→ g .fst (σ B b)) (Ω→ f .fst (σ A a)) + ∙ EH 0 _ _ + ∙ sym (F-hom (Ω→ f .fst (σ A a)) (Ω→ g .fst (σ B b))) + + fill₁ : (k : I) → _ + fill₁ k = compPath-filler + (merid ((Ω→ g .fst (σ B b) + ∙ Ω→ f .fst (σ A a)) k)) + (merid (pt C) ⁻¹) + + fill₂ : (k : I) → _ + fill₂ k = compPath-filler + (merid ((Ω→ f .fst (σ A a) + ∙ Ω→ g .fst (σ B b)) k)) + (merid (pt C) ⁻¹) + + Susp·w∙ : suspFun∙ (_·w_ .fst) ≡ suspFun∙ (fst const*) + Susp·w∙ = ΣPathP (Susp·w , refl) diff --git a/Cubical/Papers/Pi4S3-JournalVersion.agda b/Cubical/Papers/Pi4S3-JournalVersion.agda new file mode 100644 index 0000000000..212a444a0d --- /dev/null +++ b/Cubical/Papers/Pi4S3-JournalVersion.agda @@ -0,0 +1,397 @@ +{- +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: + + Formalising and computing the fourth homotopy group of the 3-sphere in Cubical Agda +-} + +-- The "--safe" flag ensures that there are no postulates or +-- unfinished goals +{-# OPTIONS --safe --cubical #-} + +module Cubical.Papers.Pi4S3-JournalVersion where + +-- Misc. +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order + +-- 2 +open import Cubical.Data.Bool as Boolean +open import Cubical.Data.Unit as UnitType + +open import Cubical.HITs.S1 as Circle +open import Cubical.Foundations.Prelude as Prelude +open import Cubical.HITs.Susp as Suspensions +open import Cubical.HITs.Sn as Spheres + hiding (S) renaming (S₊ to S) +open import Cubical.HITs.Pushout as Pushouts +open import Cubical.HITs.Wedge as Wedges +open import Cubical.HITs.Join as Joins +open import Cubical.HITs.Susp as Suspension +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as Trunc +open import Cubical.Foundations.Univalence as Univ +open import Cubical.Homotopy.Loopspace as Loopy + +open import Cubical.Homotopy.HSpace as H-Spaces +open import Cubical.Homotopy.Group.Base as HomotopyGroups +open import Cubical.Homotopy.Group.LES as LES +open import Cubical.Homotopy.HopfInvariant.HopfMap as HopfMap +open import Cubical.Homotopy.Hopf as HopfFibration +open import Cubical.Homotopy.Connected as Connectedness +open S¹Hopf +open import Cubical.Homotopy.Freudenthal as Freudenthal +open import Cubical.Homotopy.Group.PinSn as Stable +open import Cubical.Homotopy.Group.Pi3S2 as π₃S² + +-- 3 +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso as James₁ +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 as James₂ +open import Cubical.HITs.S2 as Sphere +open import Cubical.Homotopy.Whitehead as Whitehead +open import Cubical.Homotopy.BlakersMassey +module BM = BlakersMassey□ +open BM +open import Cubical.Homotopy.Group.Pi4S3.BrunerieNumber as BNumber + hiding (W) + +-- 5 +open import Cubical.ZCohomology.Base as cohom +open import Cubical.ZCohomology.GroupStructure as cohomGr +open import Cubical.ZCohomology.Properties as cohomProps +open import Cubical.ZCohomology.RingStructure.CupProduct as cup +open import Cubical.ZCohomology.MayerVietorisUnreduced as MayerVietoris +open import Cubical.Homotopy.HopfInvariant.Base as HI +open import Cubical.Homotopy.HopfInvariant.Homomorphism as HI-hom +open import Cubical.Homotopy.HopfInvariant.Brunerie as HI-β +open import Cubical.ZCohomology.Gysin as GysinSeq +open import Cubical.Homotopy.Group.Pi4S3.Summary as π₄S³ + hiding (π) +open import Cubical.ZCohomology.RingStructure.RingLaws as cupLaws + +-- 6 +open import Cubical.HITs.SmashProduct.Base as SmashProd +open import Cubical.HITs.Sn.Multiplication as SMult +open import Cubical.Homotopy.Group.Join as JoinGroup +open import Cubical.Homotopy.Group.Pi4S3.DirectProof as Direct + + +------ 2. HOMOTOPY TYPE THEORY IN Cubical Agda ------ + +--- 2.1 Elementary HoTT notions and Cubical Agda notations --- + +-- Booleans +open Boolean using (Bool) + +-- Unit +open UnitType renaming (Unit to 𝟙) + +-- S¹ +open Circle using (S¹) + +-- Non-dependent paths and refl +open Prelude using (_≡_ ; refl) + +-- funExt, funExt⁻, cong +open Prelude using (funExt; funExt⁻; cong) + +-- PathP +open Prelude using (PathP) + +-- cirlce-indution +open Circle using (elim) + +--- 2.2 More higher inductive types --- + +-- suspension +open Suspensions using (Susp) + +-- spheres +open Spheres using (S₊) + +-- pushouts +open Pushouts using (Pushout) + +-- wedge sums +open Wedges using (_⋁_) + +-- joins +open Joins using (join) + +-- cofibres +open Pushouts using (cofib) + +-- ∇ and i∨ +open Wedges using (fold⋁ ; ⋁↪) +∇ = fold⋁ +i∨ = ⋁↪ + +--- 2.3 Truncation levels and n-truncations --- + +-- propositional and general truncation +-- note that the indexing is off by 2! +open PT using (∥_∥₁) +open Trunc using (∥_∥_) + +--- 2.4 Univalence, loop spaces, and H-spaces --- + +-- Univalence, ua +open Univ using (univalence ; ua) + +-- Loop spaces +open Loopy using (Ω) + +-- H-spaces +open H-Spaces using (HSpace) + +------ 3. First results on homotopy groups of spheres ------ + +-- homotopy groups (function and loop space definition, respectively) +-- Note that the indexing is off by 1. +open HomotopyGroups using (π'Gr ; πGr) + +-- Proposition 3.2 - Long exact sequence of homotoy groups +module ExactSeq = πLES + +-- σ (definition 3.3) +open Suspensions renaming (toSusp to σ) + +-- Proposition 3.4: Sⁿ * Sᵐ ≃ Sⁿ⁺ᵐ⁺¹ +open Spheres using (IsoSphereJoin) + +-- Definition 3.5 and Proposition 3.6 (Hopf map), +-- Phrased somewhat differently in the paper. +open HopfMap using (HopfMap) +open S¹Hopf using (IsoS³TotalHopf) + +-- Lemma 3.7 (connectedness of spheres) +-- Note that the indexing is off by 2. +open Spheres using (sphereConnected) + +-- Proposition 3.8 (πₙSᵐ vanishishing for n < m) +isContr-πₙSᵐ-low : (n m : ℕ) → n < m → isContr (π n (S₊∙ m)) +isContr-πₙSᵐ-low n m l = + transport (cong isContr (sym (ua h))) + (∣ const∙ (S₊∙ n) _ ∣₂ + , ST.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → refl) + where + open import Cubical.HITs.SetTruncation as ST + con-lem : isConnected (2 + n) (S₊ m) + con-lem = isConnectedSubtr (suc (suc n)) (fst l) + (subst (λ n → isConnected n (S₊ m)) + (sym (+-suc (fst l) (suc n) ∙ cong suc (snd l))) + (sphereConnected m)) + + h : π n (S₊∙ m) ≃ π' n (Unit , tt) + h = compEquiv (isoToEquiv (πTruncIso n)) + (compEquiv (pathToEquiv (cong (π n) + (ua∙ (isoToEquiv (isContr→Iso (con-lem) isContrUnit)) refl))) + (pathToEquiv (cong ∥_∥₂ (isoToPath (IsoΩSphereMap n))))) + +-- Theorem 3.9 (Freudenthal Suspension Theorem) +open Freudenthal using (isConnectedσ) -- formalized by Evan Cavallo + +-- Corollary 3.10 (πₙSⁿ≅ℤ with identity as generator) +open Stable using (πₙ'Sⁿ≅ℤ ; πₙ'Sⁿ≅ℤ-idfun∙) + +-- Proposition 3.11 and Corollary 3.12 (π₃S²≅ℤ with Hopf map as generator) +open π₃S² using (π₃S²≅ℤ ; π₂S³-gen-by-HopfMap) + +------ 4. THE BRUNERIE NUMBER ------ +{- The formalisation of this part does not +follow the paper exactly. For this reason, we only give +the crucial results here -} + +--- 4.1 The James construction --- +-- Expository section only. No formalisation following this part of +-- the paper. + +--- 4.2 The James construction --- + +-- Lemma 3 (the family of automorphisms on ∥J₂S²∥₃ +open James₁ using (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅) + + +---- B. Formalization of the James construction ---- + +-- Definition 4.4: J₂S² +open James₁ using (Pushout⋁↪fold⋁S₊2) + +-- S²-HIT +open Sphere using (S²) + +-- Lemma 4.5 +-- Omitted (used implicitly) + +-- Lemma 4.6 (the family of automorphisms on ∥J₂S²∥₃ +open James₁ using (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅) + +-- Proposition 4.7: Ω ∥S³∥₄ ≃ ∥J₂S²∥₃ +open James₁ using (IsoΩ∥SuspS²∥₅∥Pushout⋁↪fold⋁S²∥₅) + + +--- 4.3. Definition of the Brunerie number --- + +-- Definition 4.8: W + whitehead product +W = joinTo⋁ +open Whitehead using ([_∣_]₂) + +-- Theorem 4.9 is omitted as it is used implicitly in the proof of the main result + +-- Theorem 4.10 Blakers-Massey +open BM using (isConnected-toPullback) -- formalized primarily (in a different form) by Kang Rongji + +-- Definition 4.11: The Brunerie number (note that, in the formalization +-- we have worked defined β as the image of the Hopf Invariant +-- directly) +open BNumber using (Brunerie) + +-- Corollary 4.12: π₄S³ ≅ ℤ/βℤ +open BNumber using (BrunerieIso) + + +------ 5. BRUNERIE'S PROOF THAT |β| ≡ 2 ------ + +---- A. Cohomology Theory / B. Formalization of Chapter 5---- +-- All formalizations marked with (BLM22) are borrowed from Brunerie, +-- Ljungström, and Mörtberg, “Synthetic Integral Cohomology in Cubical +-- Agda" + +--- 5.1 Cohomology and the Hopf invariant --- + +-- Eilenberg-MacLane spaces and cohomology groups (BLM22) +open cohom using (coHomK) +open cohomGr using (coHomGr) + +-- addition (BLM22) +open cohomGr using (_+ₖ_) + +-- the cup product (BLM22) +open cup using (_⌣ₖ_ ; _⌣_) + +-- Kₙ ≃ ΩKₙ₊₁ (BLM22) +open cohomProps using (Kn≃ΩKn+1) + +-- Mayer Vietoris (BLM22) +open MV using ( Ker-i⊂Im-d ; Im-d⊂Ker-i + ; Ker-Δ⊂Im-i ; Im-i⊂Ker-Δ + ; Ker-d⊂Im-Δ ; Im-Δ⊂Ker-d) + +-- Lemma 5.1 (cohomology of cofibers S³ → S²) +open HI using (Hopfβ-Iso) + +-- Definition 5.2 (Hopf Invariant) +open HI using (HopfInvariant-π') + +-- Proposition 5.3 (The Hopf invariant is a homomorphism) +open HI-hom using (GroupHom-HopfInvariant-π') + +-- Proposition 5.4 (The Hopf invariant of the Brunerie element is ±2) +open HI-β using (Brunerie'≡2) + +-- Lemma 5.5 -- only included for presentation, omitted from frmalization + +--- 5.1 The Gysin sequence --- + +-- Proposition 5.6 (Gysin sequence) +open Gysin using (Im-⌣e⊂Ker-p ; Ker-p⊂Im-⌣e + ; Im-Susp∘ϕ⊂Ker-⌣e ; Ker-⌣e⊂Im-Susp∘ϕ + ; Im-ϕ∘j⊂Ker-p ; Ker-p⊂Im-ϕ∘j) + +-- Proposition 5.7 : CP² fibration +-- Indirect, but see in particular +open HopfMap using (fibr) + +-- Proposition 5.8 : Iterated Hopf Construction. +-- Indirect, but see in particular: +open Hopf using (joinIso₂) + +-- Proposition 5.9 : ∣ HI hopf ∣ ≡ 1 +open HopfMap using (HopfInvariant-HopfMap) + +-- Theorem 5.10: π₄S³≅ℤ/2ℤ +open π₄S³ using (π₄S³≃ℤ/2ℤ) + +--- Formalisation of the Gysin sequence. --- +-- Lemma 5.11: (BLM22) +open cupLaws using (assoc-helper) + +-- proof that e₂ : H²(CP²) is a generator by computation +-- (the field with refl is where the computation happens) +open HopfMap using (isGenerator≃ℤ-e) + +------ 6. THE SIMPLIFIED NEW PROOF AND NORMALISATION OF A BRUNERIE NUMBER ------ + +--- 6.1. Interlude: joins and smash products of spheres --- + +-- Smash product (not defined as an implicit HIT) +open SmashProd using (_⋀_) + +-- Lemmas 6.1 and 6.2 +-- Omitted (included for presentation purposes, not used explicitly in +-- formalisation) + +-- Definition of pinch map +open SmashProd renaming (Join→SuspSmash to pinch) + +-- Proposition 6.3 (pinch is an equivalence) +open SmashProd using (SmashJoinIso) + +-- Definition of Fₙₘ +open SMult renaming (join→Sphere to F) + +-- Proposition 6.4 (Fₙ,ₘ is an equivalence) +open SMult using (IsoSphereJoin) + +-- Propositions 6.5 & 6.6 (graded commutativity and associativity) +open SMult using (comm⌣S ; assoc⌣S) + +--- 6.2. Homotopy groups in terms of joins. + +-- Definition 6.7 +open JoinGroup using (π*) + +-- Addition +* +open JoinGroup using (_+*_) + +-- Proposition 6.8 +open JoinGroup using (·Π≡+*) + +-- Proposition 6.9 +open JoinGroup using (π*Gr≅π'Gr) + +-- Proposition 6.10 +open JoinGroup using (π*∘∙Hom) + +--- 6.3. The new synthetic proof that π₄(S³) ≅ ℤ/2ℤ +-- A relatively detailed accound of the proof is given in the formalisation: +open Direct +-- Note that the numbering of the ηs is shifted, with +-- η₁ being ∣ ∇ ∘ W ∣, η₂ being η₁ and η₃ being η₂. +open Direct using (η₁ ; η₂ ; η₃) + +-- computation of η₂: the alternative definition and the computation +open Direct using (η₃' ; computerIsoη₃) + +--- 6.4. A stand-alone proof of Brunerie’s theorem? +-- Theorem 6.18 +-- Not formalised explicitly + +-- Definition of generalised Whitehead products ·w +open Whitehead using (_·w_) + +-- Proposition 6.22 (including Lemmas 19 and 20 and Proposition 6.21) +open Whitehead using (isConst-Susp·w) + +-- Theorem 6.23 +-- Follows directly from above but not formalised explicitly (awaiting +-- refactoring of some code in the library) From 72cf4bacfa60b3fb65cb9d359994cdf4511e1cf8 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 2 Sep 2024 21:35:06 +0200 Subject: [PATCH 28/30] small fixes --- Cubical/Categories/Constructions/Slice/Functor.agda | 6 +++--- Cubical/Categories/Functor/Properties.agda | 2 +- Cubical/Categories/Instances/Sets.agda | 2 +- 3 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Cubical/Categories/Constructions/Slice/Functor.agda b/Cubical/Categories/Constructions/Slice/Functor.agda index 1ab6901ff9..aabe059920 100644 --- a/Cubical/Categories/Constructions/Slice/Functor.agda +++ b/Cubical/Categories/Constructions/Slice/Functor.agda @@ -41,7 +41,7 @@ F-seq (F F/ c) _ _ = SliceHom-≡-intro' _ _ $ F-seq F _ _ ∑_ : ∀ {c d} f → Functor (SliceCat C c) (SliceCat C d) -F-ob (∑_ {C = C} f) (sliceob x) = sliceob (_⋆_ C x f) +F-ob (∑_ {C = C} f) (sliceob x) = sliceob (x ⋆⟨ C ⟩ f) F-hom (∑_ {C = C} f) (slicehom h p) = slicehom _ $ sym (C .⋆Assoc _ _ _) ∙ cong (comp' C f) p F-id (∑ f) = SliceHom-≡-intro' _ _ refl @@ -153,7 +153,7 @@ module _ (Pbs : Pullbacks C) where ∙∙ sym (C .⋆Assoc _ _ _)) .fst .snd .fst inv (adjIso L/b⊣R/b) (slicehom f s) = slicehom _ (D .⋆Assoc _ _ _ - ∙∙ congS (_⋆ᵈ (ε ⟦ _ ⟧ ⋆ᵈ _)) (F-seq L _ _) + ∙∙ congS (_⋆ᵈ (ε ⟦ _ ⟧ ⋆⟨ D ⟩ _)) (F-seq L _ _) ∙∙ D .⋆Assoc _ _ _ ∙ cong (L ⟪ f ⟫ ⋆ᵈ_) (cong (L ⟪ pbPr₂ ⟫ ⋆ᵈ_) (sym (N-hom ε _)) ∙∙ sym (D .⋆Assoc _ _ _) @@ -166,7 +166,7 @@ module _ (Pbs : Pullbacks C) where ∙∙ cong (L ⟪_⟫) s) rightInv (adjIso L/b⊣R/b) h = SliceHom-≡-intro' _ _ $ - let p₂ : ∀ {x} → η ⟦ _ ⟧ ⋆ᶜ R ⟪ L ⟪ x ⟫ ⋆ᵈ ε ⟦ _ ⟧ ⟫ ≡ x + let p₂ : ∀ {x} → η ⟦ _ ⟧ ⋆ᶜ R ⟪ L ⟪ x ⟫ ⋆⟨ D ⟩ ε ⟦ _ ⟧ ⟫ ≡ x p₂ = cong (_ ⋆ᶜ_) (F-seq R _ _) ∙ AssocCong₂⋆L C (sym (N-hom η _)) ∙∙ cong (_ ⋆ᶜ_) (Δ₂ _) ∙∙ C .⋆IdR _ diff --git a/Cubical/Categories/Functor/Properties.agda b/Cubical/Categories/Functor/Properties.agda index def27b2477..33f80a0455 100644 --- a/Cubical/Categories/Functor/Properties.agda +++ b/Cubical/Categories/Functor/Properties.agda @@ -123,7 +123,7 @@ isSetFunctor = isOfHLevelFunctor 0 module _ (F : Functor C D) where - -- functors preserve commutative diagrams (specificallysqures here) + -- functors preserve commutative diagrams (specifically squres 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 ⟫) diff --git a/Cubical/Categories/Instances/Sets.agda b/Cubical/Categories/Instances/Sets.agda index 3084d571c4..a2212ae7ba 100644 --- a/Cubical/Categories/Instances/Sets.agda +++ b/Cubical/Categories/Instances/Sets.agda @@ -213,4 +213,4 @@ module _ {ℓ : Level} where fst (limSetIso J D) cc = lift (lowerCone J D cc) cInv (snd (limSetIso J D)) cc = liftCone J D (cc .lower) sec (snd (limSetIso J D)) = funExt (λ _ → liftExt (cone≡ λ _ → refl)) - ret (snd (limSetIso J D)) = funExt (λ _ → cone≡ λ _ → refl) \ No newline at end of file + ret (snd (limSetIso J D)) = funExt (λ _ → cone≡ λ _ → refl) From cc3264c6d020cc4e25a5c183985da5f6a3b12be3 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 2 Sep 2024 21:50:45 +0200 Subject: [PATCH 29/30] removed unviPropEq, replaced its occurence with existing helper --- .../Constructions/Slice/Functor.agda | 12 +++---- Cubical/Categories/Limits/Pullback.agda | 31 +++++++++---------- 2 files changed, 20 insertions(+), 23 deletions(-) diff --git a/Cubical/Categories/Constructions/Slice/Functor.agda b/Cubical/Categories/Constructions/Slice/Functor.agda index aabe059920..424fa85052 100644 --- a/Cubical/Categories/Constructions/Slice/Functor.agda +++ b/Cubical/Categories/Constructions/Slice/Functor.agda @@ -66,11 +66,11 @@ module _ (Pbs : Pullbacks C) where _* : Functor (SliceCat C d) (SliceCat C c) F-ob _* x = sliceob (pbPr₁ {x = x}) F-hom _* f = slicehom _ (sym (fst (pbU f))) - F-id _* = SliceHom-≡-intro' _ _ $ univPropEq (sym (C .⋆IdL _)) (C .⋆IdR _ ∙ sym (C .⋆IdL _)) + F-id _* = SliceHom-≡-intro' _ _ $ pullbackArrowUnique (sym (C .⋆IdL _)) (C .⋆IdR _ ∙ sym (C .⋆IdL _)) F-seq _* _ _ = let (u₁ , v₁) = pbU _ ; (u₂ , v₂) = pbU _ - in SliceHom-≡-intro' _ _ $ univPropEq + in SliceHom-≡-intro' _ _ $ pullbackArrowUnique (u₂ ∙∙ cong (C ⋆ _) u₁ ∙∙ sym (C .⋆Assoc _ _ _)) (sym (C .⋆Assoc _ _ _) ∙∙ cong (comp' C _) v₂ ∙∙ AssocCong₂⋆R C v₁) @@ -91,14 +91,14 @@ module _ (Pbs : Pullbacks C) where inv (adjIso ∑𝑓⊣𝑓*) (slicehom h o) = slicehom _ $ AssocCong₂⋆R C (sym (pbCommutes)) ∙ cong (_⋆ᶜ 𝑓) o rightInv (adjIso ∑𝑓⊣𝑓*) (slicehom h o) = - SliceHom-≡-intro' _ _ (univPropEq (sym o) refl) + SliceHom-≡-intro' _ _ (pullbackArrowUnique (sym o) refl) leftInv (adjIso ∑𝑓⊣𝑓*) (slicehom h o) = let ((_ , (_ , q)) , _) = univProp _ _ _ in SliceHom-≡-intro' _ _ (sym q) adjNatInD ∑𝑓⊣𝑓* f k = SliceHom-≡-intro' _ _ $ let ((h' , (v' , u')) , _) = univProp _ _ _ ((_ , (v'' , u'')) , _) = univProp _ _ _ - in univPropEq (v' ∙∙ cong (h' ⋆ᶜ_) v'' ∙∙ sym (C .⋆Assoc _ _ _)) + in pullbackArrowUnique (v' ∙∙ cong (h' ⋆ᶜ_) v'' ∙∙ sym (C .⋆Assoc _ _ _)) (cong (_⋆ᶜ _) u' ∙ AssocCong₂⋆R C u'') adjNatInC ∑𝑓⊣𝑓* g h = SliceHom-≡-intro' _ _ $ C .⋆Assoc _ _ _ @@ -172,7 +172,7 @@ module _ (Pbs : Pullbacks C) where ∙∙ cong (_ ⋆ᶜ_) (Δ₂ _) ∙∙ C .⋆IdR _ - in univPropEq (sym (S-comm h)) p₂ + in pullbackArrowUnique (sym (S-comm h)) p₂ leftInv (adjIso L/b⊣R/b) _ = SliceHom-≡-intro' _ _ $ cong ((_⋆ᵈ _) ∘ L ⟪_⟫) (sym (snd (snd (fst (univProp _ _ _))))) @@ -181,7 +181,7 @@ module _ (Pbs : Pullbacks C) where let (h , (u , v)) = univProp _ _ _ .fst (u' , v') = pbU _ - in univPropEq + in pullbackArrowUnique (u ∙∙ cong (h ⋆ᶜ_) u' ∙∙ sym (C .⋆Assoc h _ _)) (cong (_ ⋆ᶜ_) (F-seq R _ _) ∙∙ sym (C .⋆Assoc _ _ _) ∙∙ diff --git a/Cubical/Categories/Limits/Pullback.agda b/Cubical/Categories/Limits/Pullback.agda index 2236cbcad0..bba4e78f1b 100644 --- a/Cubical/Categories/Limits/Pullback.agda +++ b/Cubical/Categories/Limits/Pullback.agda @@ -53,16 +53,23 @@ module _ (C : Category ℓ ℓ') where pbCommutes : pbPr₁ ⋆ cspn .s₁ ≡ pbPr₂ ⋆ cspn .s₂ univProp : isPullback cspn pbPr₁ pbPr₂ pbCommutes - univPropEq : ∀ {c p₁ p₂ H g} p q → fst (fst (univProp {c} p₁ p₂ H)) ≡ g - univPropEq p q = cong fst (snd (univProp _ _ _) (_ , (p , q))) + pullbackArrow : + {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ]) + (H : p₁ ⋆ cspn .s₁ ≡ p₂ ⋆ cspn .s₂) → C [ c , pbOb ] + pullbackArrow p₁ p₂ H = univProp p₁ p₂ H .fst .fst + + pullbackArrowUnique : + {c : ob} {p₁ : C [ c , cspn .l ]} {p₂ : C [ c , cspn .r ]} + {H : p₁ ⋆ cspn .s₁ ≡ p₂ ⋆ cspn .s₂} + {pbArrow' : C [ c , pbOb ]} + (H₁ : p₁ ≡ pbArrow' ⋆ pbPr₁) (H₂ : p₂ ≡ pbArrow' ⋆ pbPr₂) + → pullbackArrow p₁ p₂ H ≡ pbArrow' + pullbackArrowUnique H₁ H₂ = + cong fst (univProp _ _ _ .snd (_ , (H₁ , H₂))) + 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) @@ -78,16 +85,6 @@ module _ (C : Category ℓ ℓ') where 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 From 96ca9667d7e5b8881c114beeecd3b9fb08a9579c Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 2 Sep 2024 21:59:26 +0200 Subject: [PATCH 30/30] added requested comment --- Cubical/Categories/Instances/Sets.agda | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Cubical/Categories/Instances/Sets.agda b/Cubical/Categories/Instances/Sets.agda index a2212ae7ba..7917105ba1 100644 --- a/Cubical/Categories/Instances/Sets.agda +++ b/Cubical/Categories/Instances/Sets.agda @@ -152,6 +152,11 @@ univProp (completeSET J D) c cc = module _ {ℓ} where +-- While pullbacks can be obtained from limits +-- (using `completeSET` & `LimitsOfShapeCospanCat→Pullbacks` from `Cubical.Categories.Limits.Pullback`), +-- this direct construction can be more convenient when only pullbacks are needed. +-- It also has better behavior in terms of inferring implicit arguments + open Pullback PullbacksSET : Pullbacks (SET ℓ)