From b88286d6469979eb8a31a5dd32ae04655cb664af Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Wed, 24 Nov 2021 09:59:35 +0800 Subject: [PATCH 01/24] formalize Rijke's counting result --- Cubical/Data/FinSet/FinType.agda | 146 +++++++++ Cubical/Data/FinSet/FiniteStructure.agda | 67 ++++ Cubical/Data/Sigma/Properties.agda | 27 ++ .../Experiments/CountingFiniteStructure.agda | 73 +++++ Cubical/HITs/SetTruncation/Fibers.agda | 294 ++++++++++++++++++ 5 files changed, 607 insertions(+) create mode 100644 Cubical/Data/FinSet/FinType.agda create mode 100644 Cubical/Data/FinSet/FiniteStructure.agda create mode 100644 Cubical/Experiments/CountingFiniteStructure.agda create mode 100644 Cubical/HITs/SetTruncation/Fibers.agda diff --git a/Cubical/Data/FinSet/FinType.agda b/Cubical/Data/FinSet/FinType.agda new file mode 100644 index 0000000000..a480ca6cab --- /dev/null +++ b/Cubical/Data/FinSet/FinType.agda @@ -0,0 +1,146 @@ +{- + +Definition and properties of finite types + +This file formalize the notion of Rijke finite type, +which is a direct generalization of Bishop finite set. +Basically, a type is (Rijke) n-finite if all its i-th +order homotopy groups πᵢ for i ≤ n are Bishop finite. + +https://github.com/EgbertRijke/OEIS-A000001 + +-} +{-# OPTIONS --safe #-} + +module Cubical.Data.FinSet.FinType where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Path + +open import Cubical.HITs.PropositionalTruncation as Prop +open import Cubical.HITs.SetTruncation as Set +open import Cubical.HITs.SetTruncation.Fibers +open import Cubical.HITs.SetQuotients as SetQuot + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +open import Cubical.Data.FinSet.Base +open import Cubical.Data.FinSet.Properties +open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.Quotients +open import Cubical.Data.FinSet.Cardinality + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +private + variable + ℓ ℓ' : Level + n : ℕ + X : Type ℓ + Y : Type ℓ' + +-- the (Rijke) finite type + +isFinType : (n : ℕ) → Type ℓ → Type ℓ +isFinType 0 X = isFinSet ∥ X ∥₂ +isFinType (suc n) X = (isFinType 0 X) × ((a b : X) → isFinType n (a ≡ b)) + +isPropIsFinType : isProp (isFinType n X) +isPropIsFinType {n = 0} = isPropIsFinSet +isPropIsFinType {n = suc n} = isProp× isPropIsFinSet (isPropΠ2 (λ _ _ → isPropIsFinType {n = n})) + +-- the type of finite types + +FinType : (ℓ : Level)(n : ℕ) → Type (ℓ-suc ℓ) +FinType ℓ n = TypeWithStr ℓ (isFinType n) + +-- basic numerical implications + +isFinTypeSuc→isFinType : isFinType (suc n) X → isFinType n X +isFinTypeSuc→isFinType {n = 0} p = p .fst +isFinTypeSuc→isFinType {n = suc n} p .fst = p .fst +isFinTypeSuc→isFinType {n = suc n} p .snd a b = isFinTypeSuc→isFinType {n = n} (p .snd a b) + +isFinType→isFinType0 : isFinType n X → isFinType 0 X +isFinType→isFinType0 {n = 0} p = p +isFinType→isFinType0 {n = suc n} p = p .fst + +isFinTypeSuc→isFinType1 : isFinType (suc n) X → isFinType 1 X +isFinTypeSuc→isFinType1 {n = 0} p = p +isFinTypeSuc→isFinType1 {n = suc n} p .fst = p .fst +isFinTypeSuc→isFinType1 {n = suc n} p .snd a b = isFinType→isFinType0 (p .snd a b) + +-- useful properties + +EquivPresIsFinType : (n : ℕ) → X ≃ Y → isFinType n X → isFinType n Y +EquivPresIsFinType 0 e = EquivPresIsFinSet (isoToEquiv (setTruncIso (equivToIso e))) +EquivPresIsFinType (suc n) e (p , q) .fst = EquivPresIsFinType 0 e p +EquivPresIsFinType (suc n) e (p , q) .snd a b = + EquivPresIsFinType n (invEquiv (congEquiv (invEquiv e))) (q _ _) + +isFinSet→isFinType : (n : ℕ) → isFinSet X → isFinType n X +isFinSet→isFinType 0 p = EquivPresIsFinSet (invEquiv (setTruncIdempotent≃ (isFinSet→isSet p))) p +isFinSet→isFinType (suc n) p .fst = isFinSet→isFinType 0 p +isFinSet→isFinType (suc n) p .snd a b = isFinSet→isFinType n (isFinSet≡ (_ , p) _ _) + +isPathConnected→isFinType0 : isContr ∥ X ∥₂ → isFinType 0 X +isPathConnected→isFinType0 p = isContr→isFinSet p + +{- finite types are closed under forming Σ-types -} + +module _ + (X : Type ℓ ) + (Y : Type ℓ')(h : isFinType 1 Y) + (f : X → Y) + (q : (y : Y) → isFinType 0 (fiber f y)) where + + ∥f∥₂ : ∥ X ∥₂ → ∥ Y ∥₂ + ∥f∥₂ = Set.map f + + module _ + (y : Y) where + + isDecFiberRel : (x x' : ∥ fiber f y ∥₂) → Dec (fiberRel2 _ _ f y x x') + isDecFiberRel x x' = + isFinSet→Dec∥∥ (isFinSetΣ (_ , h .snd y y) (λ _ → _ , isFinSet≡ (_ , q y) _ _)) + + isFinSetFiber∥∥₂' : isFinSet (fiber ∥f∥₂ ∣ y ∣₂) + isFinSetFiber∥∥₂' = + EquivPresIsFinSet (∥fiber∥₂/Rel≃fiber∥∥₂ _ _ f y) + (isFinSetQuot (_ , q y) (fiberRel2 _ _ _ _) (isEquivRelFiberRel _ _ _ _) isDecFiberRel) + + isFinSetFiber∥∥₂ : (y : ∥ Y ∥₂) → isFinSet (fiber ∥f∥₂ y) + isFinSetFiber∥∥₂ = Set.elim (λ _ → isProp→isSet isPropIsFinSet) isFinSetFiber∥∥₂' + + isFinType0Total : isFinType 0 X + isFinType0Total = isFinSetTotal _ (_ , h .fst) ∥f∥₂ isFinSetFiber∥∥₂ + +module _ + (X : FinType ℓ 1) + (Y : X .fst → FinType ℓ' 0) where + + isFinType0Σ : isFinType 0 (Σ (X .fst) (λ x → Y x .fst)) + isFinType0Σ = + isFinType0Total (Σ (X .fst) (λ x → Y x .fst)) (X .fst) (X .snd) + (λ (x , y) → x) (λ x → EquivPresIsFinType 0 (fiberProjEquiv _ _ x) (Y x .snd)) + +-- the closedness result + +isFinTypeΣ : + {n : ℕ} + (X : FinType ℓ (1 + n)) + (Y : X .fst → FinType ℓ' n) + → isFinType n (Σ (X .fst) (λ x → Y x .fst)) +isFinTypeΣ {n = 0} = isFinType0Σ +isFinTypeΣ {n = suc n} X Y .fst = + isFinType0Σ (_ , isFinTypeSuc→isFinType1 (X .snd)) (λ x → _ , isFinType→isFinType0 (Y x .snd)) +isFinTypeΣ {n = suc n} X Y .snd a b = + EquivPresIsFinType n (ΣPathTransport≃PathΣ a b) + (isFinTypeΣ {n = n} (_ , X .snd .snd _ _) (λ _ → _ , Y _ .snd .snd _ _)) diff --git a/Cubical/Data/FinSet/FiniteStructure.agda b/Cubical/Data/FinSet/FiniteStructure.agda new file mode 100644 index 0000000000..d938dd9bed --- /dev/null +++ b/Cubical/Data/FinSet/FiniteStructure.agda @@ -0,0 +1,67 @@ +{-# OPTIONS --safe #-} + +module Cubical.Data.FinSet.FiniteStructure where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.HITs.PropositionalTruncation as Prop +open import Cubical.HITs.SetTruncation as Set + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +open import Cubical.Data.FinSet.Base +open import Cubical.Data.FinSet.Properties +open import Cubical.Data.FinSet.Induction +open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.Cardinality +open import Cubical.Data.FinSet.FinType + +private + variable + ℓ ℓ' : Level + n : ℕ + S : FinSet ℓ → FinSet ℓ' + +-- type of finite sets with finite structure + +FinSetWithStr : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +FinSetWithStr ℓ S = Σ[ X ∈ FinSet ℓ ] S X .fst + +-- type finite sets of a given cardinal + +FinSetOfCard : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) +FinSetOfCard ℓ n = Σ[ X ∈ FinSet ℓ ] (card X ≡ n) + +FinSetWithStrOfCard : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') (n : ℕ) → Type (ℓ-max (ℓ-suc ℓ) ℓ') +FinSetWithStrOfCard ℓ S n = Σ[ X ∈ FinSetOfCard ℓ n ] S (X .fst) .fst + +FinSetOfCard≡ : (X Y : FinSetOfCard ℓ n) → (X .fst ≡ Y .fst) ≃ (X ≡ Y) +FinSetOfCard≡ _ _ = Σ≡PropEquiv (λ _ → isSetℕ _ _) + +open Iso + +∥FinSetOfCard∥₂≡ : (X Y : FinSetOfCard ℓ n) → ∥ X .fst ≡ Y .fst ∥ → ∣ X ∣₂ ≡ ∣ Y ∣₂ +∥FinSetOfCard∥₂≡ _ _ = + Prop.rec (squash₂ _ _) (λ p → PathIdTrunc₀Iso .inv ∣ FinSetOfCard≡ _ _ .fst p ∣) + +isPathConnectedFinSetOfCard : isContr ∥ FinSetOfCard ℓ n ∥₂ +isPathConnectedFinSetOfCard {n = n} .fst = ∣ 𝔽in n , card𝔽in n ∣₂ +isPathConnectedFinSetOfCard {n = n} .snd = + Set.elim (λ _ → isProp→isSet (squash₂ _ _)) (λ (X , p) → sym (∥FinSetOfCard∥₂≡ _ _ (card≡n p))) + +isFinTypeFinSetOfCard : isFinType 1 (FinSetOfCard ℓ n) +isFinTypeFinSetOfCard .fst = isPathConnected→isFinType0 isPathConnectedFinSetOfCard +isFinTypeFinSetOfCard .snd X Y = + isFinSet→isFinType _ (EquivPresIsFinSet (FinSet≡ _ _ ⋆ FinSetOfCard≡ _ _) (isFinSetType≡ (X .fst) (Y .fst))) + +isFinTypeFinSetWithStrOfCard : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') (n : ℕ) + → isFinType 0 (FinSetWithStrOfCard ℓ S n) +isFinTypeFinSetWithStrOfCard ℓ S n = + isFinTypeΣ (_ , isFinTypeFinSetOfCard) (λ X → _ , isFinSet→isFinType 0 (S (X .fst) .snd)) + +-- we can counting the number of finite sets of a given cardinal with certain structures +cardFinSetWithStrOfCard : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') (n : ℕ) → ℕ +cardFinSetWithStrOfCard ℓ S n = card (_ , isFinTypeFinSetWithStrOfCard ℓ S n) diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index 7352ffabba..c4b75a8060 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -362,6 +362,7 @@ module _ {A : Type ℓ} {B : A → Type ℓ'} {C : ∀ a → B a → Type ℓ''} unquoteDecl curryEquiv = declStrictIsoToEquiv curryEquiv curryIso -- Sigma type with empty base + module _ (A : ⊥ → Type ℓ) where open Iso @@ -371,3 +372,29 @@ module _ (A : ⊥ → Type ℓ) where ΣEmpty : Σ ⊥ A ≃ ⊥ ΣEmpty = isoToEquiv ΣEmptyIso + +-- fiber of projection map + +module _ + (A : Type ℓ) + (B : A → Type ℓ') where + + private + proj : Σ A B → A + proj (a , b) = a + + module _ + (a : A) where + + open Iso + + fiberProjIso : Iso (B a) (fiber proj a) + fiberProjIso .fun b = (a , b) , refl + fiberProjIso .inv ((a' , b') , p) = subst B p b' + fiberProjIso .leftInv b i = substRefl {B = B} b i + fiberProjIso .rightInv (_ , p) i .fst .fst = p (~ i) + fiberProjIso .rightInv ((_ , b') , p) i .fst .snd = subst-filler B p b' (~ i) + fiberProjIso .rightInv (_ , p) i .snd j = p (~ i ∨ j) + + fiberProjEquiv : B a ≃ fiber proj a + fiberProjEquiv = isoToEquiv fiberProjIso diff --git a/Cubical/Experiments/CountingFiniteStructure.agda b/Cubical/Experiments/CountingFiniteStructure.agda new file mode 100644 index 0000000000..bd224686e9 --- /dev/null +++ b/Cubical/Experiments/CountingFiniteStructure.agda @@ -0,0 +1,73 @@ +{- + +This file contains the counting number of finite sets with structure. + +https://github.com/EgbertRijke/OEIS-A000001 + +-} +{-# OPTIONS --safe #-} + +module Cubical.Experiments.CountingFiniteStructure where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +open import Cubical.Data.FinSet +open import Cubical.Data.FinSet.Induction +open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.Cardinality +open import Cubical.Data.FinSet.FinType +open import Cubical.Data.FinSet.FiniteStructure + +private + variable + ℓ : Level + +-- convenient abbreviation + +isFinStrCard : (S : FinSet ℓ-zero → FinSet ℓ) (n : ℕ) → isFinType 0 (FinSetWithStrOfCard ℓ-zero S n) +isFinStrCard S n = isFinTypeFinSetWithStrOfCard ℓ-zero S n + +-- structure that is no structure + +TrivialStr : FinSet ℓ → FinSet ℓ +TrivialStr _ = 𝟙 + +-- identity structure + +IdentityStr : FinSet ℓ → FinSet ℓ +IdentityStr X = X + +-- finite semi-groups + +FinSemiGroupStr : FinSet ℓ → FinSet ℓ +FinSemiGroupStr X .fst = + Σ[ p ∈ (X .fst → X .fst → X .fst) ] ((x y z : X .fst) → p (p x y) z ≡ p x (p y z)) +FinSemiGroupStr X .snd = + isFinSetΣ (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) + (λ p → _ , isFinSetΠ3 X (λ _ → X) (λ _ _ → X) (λ _ _ _ → _ , isFinSet≡ X _ _)) + +FinSemiGroupOfCard : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) +FinSemiGroupOfCard ℓ n = FinSetWithStrOfCard ℓ FinSemiGroupStr n + +isFinSemiGroupOfCard : (ℓ : Level) (n : ℕ) → isFinType 0 (FinSemiGroupOfCard ℓ n) +isFinSemiGroupOfCard ℓ n = isFinTypeFinSetWithStrOfCard _ FinSemiGroupStr _ + +-- two rather trivial numbers +-- but the computation is essentially not that trivial +-- this one can be computed in half-a-minute +g : ℕ +g = card (_ , isFinStrCard TrivialStr 2) + +-- this is already hard to compute +-- don't know if anyone's computer give the final result +f : ℕ +f = card (_ , isFinStrCard IdentityStr 2) + +-- the number of finite semi-groups with cardinal 2 +-- it should be 5 +-- would you like to try? +n2 : ℕ +n2 = card (_ , isFinStrCard FinSemiGroupStr 2) diff --git a/Cubical/HITs/SetTruncation/Fibers.agda b/Cubical/HITs/SetTruncation/Fibers.agda new file mode 100644 index 0000000000..856046c9ad --- /dev/null +++ b/Cubical/HITs/SetTruncation/Fibers.agda @@ -0,0 +1,294 @@ +{- + +This file contains: + +- Fibers of induced map between set truncations is the set truncation of fibers + modulo a certain equivalence relation defined by π₁ of the base. + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.SetTruncation.Fibers where + +open import Cubical.HITs.SetTruncation.Base + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Path +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.HITs.PropositionalTruncation as Prop +open import Cubical.HITs.SetTruncation as Set +open import Cubical.HITs.SetQuotients as SetQuot + +open import Cubical.Relation.Binary + +private + variable + ℓ ℓ' : Level + +module _ + (X : Type ℓ ) + (Y : Type ℓ') + (f : X → Y) where + + private + ∥f∥₂ : ∥ X ∥₂ → ∥ Y ∥₂ + ∥f∥₂ = Set.map f + + module _ + (y : Y) where + + open Iso + + fiber→fiber∥∥₂ : fiber f y → fiber ∥f∥₂ ∣ y ∣₂ + fiber→fiber∥∥₂ (x , p) = ∣ x ∣₂ , cong ∣_∣₂ p + + isSetFiber∥∥₂ : isSet (fiber ∥f∥₂ ∣ y ∣₂) + isSetFiber∥∥₂ = isOfHLevelΣ 2 squash₂ (λ _ → isProp→isSet (squash₂ _ _)) + + ∥fiber∥₂→fiber∥∥₂ : ∥ fiber f y ∥₂ → fiber ∥f∥₂ ∣ y ∣₂ + ∥fiber∥₂→fiber∥∥₂ = Set.rec isSetFiber∥∥₂ fiber→fiber∥∥₂ + + ∣transport∣ : ∥ y ≡ y ∥₂ → ∥ fiber f y ∥₂ → ∥ fiber f y ∥₂ + ∣transport∣ = Set.rec2 squash₂ (λ s (x , q) → ∣ x , q ∙ s ∣₂) + + fiberRel2 : ∥ fiber f y ∥₂ → ∥ fiber f y ∥₂ → Type (ℓ-max ℓ ℓ') + fiberRel2 x x' = + ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s x ≡ x' ∥ + + fiberRel' : fiber f y → fiber f y → Type (ℓ-max ℓ ℓ') + fiberRel' (x , p) (x' , p') = + Σ[ s ∈ y ≡ y ] Σ[ q ∈ x ≡ x' ] PathP (λ i → f (q i) ≡ s i) p p' + + feq' : (a b : fiber f y)(r : fiberRel' a b) → fiber→fiber∥∥₂ a ≡ fiber→fiber∥∥₂ b + feq' (x , p) (x' , p') (s , q , t) i .fst = ∣ q i ∣₂ + feq' (x , p) (x' , p') (s , q , t) i .snd j = + hcomp (λ k → λ { (i = i0) → ∣ p j ∣₂ + ; (i = i1) → ∣ p' j ∣₂ + ; (j = i0) → ∣ f (q i) ∣₂ + ; (j = i1) → squash₂ ∣ y ∣₂ ∣ y ∣₂ (cong ∣_∣₂ s) refl k i }) + (∣ t i j ∣₂) + + fiberRel : ∥ fiber f y ∥₂ → ∥ fiber f y ∥₂ → Type (ℓ-max ℓ ℓ') + fiberRel a b = Set.rec2 isSetHProp (λ a b → ∥ fiberRel' a b ∥ , isPropPropTrunc) a b .fst + + isPropFiberRel : (x x' : ∥ fiber f y ∥₂) → isProp (fiberRel x x') + isPropFiberRel a b = Set.rec2 isSetHProp (λ a b → ∥ fiberRel' a b ∥ , isPropPropTrunc) a b .snd + + fiberRel'→fiberRel2' : ((x , p) x' : fiber f y) → fiberRel' (x , p) x' → Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' + fiberRel'→fiberRel2' _ _ (_ , _ , t) .fst i = t i i1 + fiberRel'→fiberRel2' _ _ (_ , q , _) .snd i .fst = q i + fiberRel'→fiberRel2' (x , p) (x' , p') (s , q , t) .snd i .snd j = + hcomp (λ k → λ { (i = i0) → compPath-filler p s k j + ; (i = i1) → p' j + ; (j = i0) → f (q i) + ; (j = i1) → s (i ∨ k) }) + (t i j) + + truncRel : ((x , p) x' : fiber f y) → Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' + → ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s ∣ x , p ∣₂ ≡ ∣ x' ∣₂ ∥ + truncRel _ _ (s , p) = ∣ ∣ s ∣₂ , cong ∣_∣₂ p ∣ + + truncRel` : ((x , p) x' : fiber f y) → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ + → ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s ∣ x , p ∣₂ ≡ ∣ x' ∣₂ ∥ + truncRel` x x' = Prop.rec isPropPropTrunc (λ r → truncRel x x' r) + + fiberRel→fiberRel2' : (x x' : fiber f y) → fiberRel ∣ x ∣₂ ∣ x' ∣₂ → fiberRel2 ∣ x ∣₂ ∣ x' ∣₂ + fiberRel→fiberRel2' x x' = Prop.rec isPropPropTrunc (λ r → truncRel x x' (fiberRel'→fiberRel2' x x' r)) + + fiberRel→fiberRel2 : (x x' : ∥ fiber f y ∥₂) → fiberRel x x' → fiberRel2 x x' + fiberRel→fiberRel2 = + Set.elim2 (λ _ _ → isProp→isSet (isPropΠ (λ _ → isPropPropTrunc))) + (λ x x' → fiberRel→fiberRel2' x x') + + truncRel'' : ((x , p) x' : fiber f y) → (s : y ≡ y) → ∣ x , p ∙ s ∣₂ ≡ ∣ x' ∣₂ + → ∥ (x , p ∙ s) ≡ x' ∥ + truncRel'' _ _ _ q = Set.PathIdTrunc₀Iso .fun q + + truncRel''' : ((x , p) x' : fiber f y) → (s : ∥ y ≡ y ∥₂) → ∣transport∣ s ∣ x , p ∣₂ ≡ ∣ x' ∣₂ + → ∥ Σ[ s ∈ y ≡ y ] ∥ (x , p ∙ s) ≡ x' ∥ ∥ + truncRel''' x x' = + Set.elim (λ _ → isProp→isSet (isPropΠ (λ _ → isPropPropTrunc))) + (λ s' p' → ∣ s' , truncRel'' x x' s' p' ∣) + + helper-fun : ((x , p) x' : fiber f y) → (s : y ≡ y) + → ∥ (x , p ∙ s) ≡ x' ∥ → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ + helper-fun _ _ s = Prop.rec isPropPropTrunc (λ q → ∣ s , q ∣) + + helper-fun' : ((x , p) x' : fiber f y) → ∥ Σ[ s ∈ y ≡ y ] ∥ (x , p ∙ s) ≡ x' ∥ ∥ + → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ + helper-fun' _ _ = + Prop.rec isPropPropTrunc (λ (s , q) → helper-fun _ _ s q) + + truncRel' : ((x , p) x' : fiber f y) → ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s ∣ x , p ∣₂ ≡ ∣ x' ∣₂ ∥ + → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ + truncRel' _ _ = Prop.rec isPropPropTrunc (λ (s , q) → helper-fun' _ _ (truncRel''' _ _ s q)) + + fiberRel2'→fiberRel' : ((x , p) x' : fiber f y) → (s : y ≡ y)(t : (x , p ∙ s) ≡ x') → fiberRel' (x , p) x' + fiberRel2'→fiberRel' _ _ s _ .fst = s + fiberRel2'→fiberRel' _ _ _ t .snd .fst i = t i .fst + fiberRel2'→fiberRel' (x , p) (x' , p') s t .snd .snd i j = + hcomp (λ k → λ { (i = i0) → compPath-filler p s (~ k) j + ; (i = i1) → t k .snd j + ; (j = i0) → f (t (i ∧ k) .fst) + ; (j = i1) → s (i ∨ (~ k)) }) + ((p ∙ s) j) + + fiberRel2'→fiberRel'-helper : ((x , p) x' : fiber f y) + → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ → fiberRel ∣ x , p ∣₂ ∣ x' ∣₂ + fiberRel2'→fiberRel'-helper _ _ = + Prop.rec isPropPropTrunc (λ (s , p) → ∣ fiberRel2'→fiberRel' _ _ s p ∣) + + fiberRel2→fiberRel : (x x' : ∥ fiber f y ∥₂) → fiberRel2 x x' → fiberRel x x' + fiberRel2→fiberRel = + Set.elim2 (λ x x' → isProp→isSet (isPropΠ (λ _ → isPropFiberRel x x'))) + (λ _ _ r → fiberRel2'→fiberRel'-helper _ _ (truncRel' _ _ r)) + + fiberRel≃fiberRel2 : (x x' : ∥ fiber f y ∥₂) → fiberRel x x' ≃ fiberRel2 x x' + fiberRel≃fiberRel2 x x' = + propBiimpl→Equiv (isPropFiberRel x x') isPropPropTrunc + (fiberRel→fiberRel2 _ _) (fiberRel2→fiberRel _ _) + + feq'' : (a b : fiber f y)(r : fiberRel ∣ a ∣₂ ∣ b ∣₂) → ∥fiber∥₂→fiber∥∥₂ ∣ a ∣₂ ≡ ∥fiber∥₂→fiber∥∥₂ ∣ b ∣₂ + feq'' a b r = Prop.rec (isSetFiber∥∥₂ _ _) (feq' a b) r + + feq : (a b : ∥ fiber f y ∥₂)(r : fiberRel a b) → ∥fiber∥₂→fiber∥∥₂ a ≡ ∥fiber∥₂→fiber∥∥₂ b + feq = Set.elim2 (λ _ _ → isProp→isSet (isPropΠ (λ _ → isSetFiber∥∥₂ _ _))) feq'' + + ∥fiber∥₂/R→fiber∥∥₂ : ∥ fiber f y ∥₂ / fiberRel → fiber ∥f∥₂ ∣ y ∣₂ + ∥fiber∥₂/R→fiber∥∥₂ = SetQuot.rec isSetFiber∥∥₂ ∥fiber∥₂→fiber∥∥₂ feq + + fun1 : (x x' : X)(q : x ≡ x')(p : f x ≡ y)(p' : f x' ≡ y) → fiberRel' (x , p) (x' , p') + fun1 x x' q p p' .fst = (sym p) ∙∙ cong f q ∙∙ p' + fun1 x x' q p p' .snd .fst = q + fun1 x x' q p p' .snd .snd i j = doubleCompPath-filler (sym p) (cong f q) p' j i + + fun2'' : (x : X) → (p p' : f x ≡ y) → fiberRel ∣ x , p ∣₂ ∣ x , p' ∣₂ + fun2'' x p p' = ∣ fun1 x x refl p p' ∣ + + fun2' : (x : X) → (p : ∥ f x ≡ y ∥) → ∥ fiber f y ∥₂ / fiberRel + fun2' x = Prop.rec→Set squash/ (λ p → [ ∣ x , p ∣₂ ]) (λ p p' → eq/ _ _ (fun2'' x p p')) + + fun3' : (x : X) → (p : ∥f∥₂ ∣ x ∣₂ ≡ ∣ y ∣₂) → ∥ fiber f y ∥₂ / fiberRel + fun3' x p = fun2' x (Set.PathIdTrunc₀Iso .fun p) + + fun3 : (x : ∥ X ∥₂) → (p : ∥f∥₂ x ≡ ∣ y ∣₂) → ∥ fiber f y ∥₂ / fiberRel + fun3 = Set.elim (λ _ → isSetΠ (λ _ → squash/)) fun3' + + fiber∥∥₂→∥fiber∥₂/R : fiber ∥f∥₂ ∣ y ∣₂ → ∥ fiber f y ∥₂ / fiberRel + fiber∥∥₂→∥fiber∥₂/R (x , p) = fun3 x p + + leftHtpy : (x : fiber f y) → fiber∥∥₂→∥fiber∥₂/R (fiber→fiber∥∥₂ x) ≡ [ ∣ x ∣₂ ] + leftHtpy (x , p) = eq/ _ _ (fun2'' x _ _) + + ∥fiber∥₂/R→fiber∥∥₂→∥fiber∥₂/R : + (x : ∥ fiber f y ∥₂ / fiberRel) → fiber∥∥₂→∥fiber∥₂/R (∥fiber∥₂/R→fiber∥∥₂ x) ≡ x + ∥fiber∥₂/R→fiber∥∥₂→∥fiber∥₂/R = SetQuot.elimProp (λ _ → squash/ _ _) (Set.elim (λ _ → isProp→isSet (squash/ _ _)) leftHtpy) + + rightHtpy : + (x : X) → (p : f x ≡ y) → ∥fiber∥₂/R→fiber∥∥₂ (fun2' x ∣ p ∣) ≡ (∣ x ∣₂ , (Set.PathIdTrunc₀Iso .inv ∣ p ∣)) + rightHtpy x p i .fst = ∣ x ∣₂ + rightHtpy x p i .snd = + isOfHLevelRespectEquiv 1 (invEquiv (isoToEquiv (PathIdTrunc₀Iso))) isPropPropTrunc + (∥fiber∥₂/R→fiber∥∥₂ (fun2' x ∣ p ∣) .snd) (Set.PathIdTrunc₀Iso .inv ∣ p ∣) i + + rightHtpy' : + (x : X) → (p : ∥ f x ≡ y ∥) → ∥fiber∥₂/R→fiber∥∥₂ (fun2' x p) ≡ (∣ x ∣₂ , (Set.PathIdTrunc₀Iso .inv p)) + rightHtpy' x = Prop.elim (λ _ → isSetFiber∥∥₂ _ _) (rightHtpy x) + + rightHtpy'' : + (x : X) → (p : ∥f∥₂ ∣ x ∣₂ ≡ ∣ y ∣₂) → ∥fiber∥₂/R→fiber∥∥₂ (fun3' x p) ≡ (∣ x ∣₂ , p) + rightHtpy'' x p i .fst = rightHtpy' x (Set.PathIdTrunc₀Iso .fun p) i .fst + rightHtpy'' x p i .snd = + hcomp (λ j → λ { (i = i0) → rightHtpy' x (Set.PathIdTrunc₀Iso .fun p) i0 .snd + ; (i = i1) → Set.PathIdTrunc₀Iso .leftInv p j }) + (rightHtpy' x (Set.PathIdTrunc₀Iso .fun p) i .snd) + + rightHtpy''' : + (x : ∥ X ∥₂) → (p : ∥f∥₂ x ≡ ∣ y ∣₂) → ∥fiber∥₂/R→fiber∥∥₂ (fiber∥∥₂→∥fiber∥₂/R (x , p)) ≡ (x , p) + rightHtpy''' = Set.elim (λ _ → isSetΠ (λ _ → isProp→isSet (isSetFiber∥∥₂ _ _))) rightHtpy'' + + fiber∥∥₂→∥fiber∥₂/R→fiber∥∥₂ : + (x : fiber ∥f∥₂ ∣ y ∣₂) → ∥fiber∥₂/R→fiber∥∥₂ (fiber∥∥₂→∥fiber∥₂/R x) ≡ x + fiber∥∥₂→∥fiber∥₂/R→fiber∥∥₂ (x , p) = rightHtpy''' x p + + Iso-∥fiber∥₂/R-fiber∥∥₂ : Iso (∥ fiber f y ∥₂ / fiberRel) (fiber ∥f∥₂ ∣ y ∣₂) + Iso-∥fiber∥₂/R-fiber∥∥₂ .fun = ∥fiber∥₂/R→fiber∥∥₂ + Iso-∥fiber∥₂/R-fiber∥∥₂ .inv = fiber∥∥₂→∥fiber∥₂/R + Iso-∥fiber∥₂/R-fiber∥∥₂ .leftInv = ∥fiber∥₂/R→fiber∥∥₂→∥fiber∥₂/R + Iso-∥fiber∥₂/R-fiber∥∥₂ .rightInv = fiber∥∥₂→∥fiber∥₂/R→fiber∥∥₂ + + -- main results + + ∥fiber∥₂/R≃fiber∥∥₂ : ∥ fiber f y ∥₂ / fiberRel ≃ fiber ∥f∥₂ ∣ y ∣₂ + ∥fiber∥₂/R≃fiber∥∥₂ = isoToEquiv Iso-∥fiber∥₂/R-fiber∥∥₂ + + ∥fiber∥₂/Rel≃∥fiber∥₂/R : ∥ fiber f y ∥₂ / fiberRel2 ≃ ∥ fiber f y ∥₂ / fiberRel + ∥fiber∥₂/Rel≃∥fiber∥₂/R = + isoToEquiv (relBiimpl→TruncIso + (λ {a} {b} → fiberRel2→fiberRel a b) + (λ {a} {b} → fiberRel→fiberRel2 a b)) + + ∥fiber∥₂/Rel≃fiber∥∥₂ : ∥ fiber f y ∥₂ / fiberRel2 ≃ fiber ∥f∥₂ ∣ y ∣₂ + ∥fiber∥₂/Rel≃fiber∥∥₂ = compEquiv ∥fiber∥₂/Rel≃∥fiber∥₂/R ∥fiber∥₂/R≃fiber∥∥₂ + + -- the relation is equivalence relation + + open BinaryRelation + + isReflFiberRel' : (x : fiber f y) → fiberRel2 ∣ x ∣₂ ∣ x ∣₂ + isReflFiberRel' (x , q) = ∣ ∣ refl ∣₂ , (λ i → ∣ x , rUnit q (~ i) ∣₂) ∣ + + isReflFiberRel : isRefl fiberRel2 + isReflFiberRel = Set.elim (λ _ → isProp→isSet isPropPropTrunc) isReflFiberRel' + + isSymFiberRel'' : ((x , p) (x' , p') : fiber f y) + → Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ (x' , p') → Σ[ s' ∈ y ≡ y ] (x' , p' ∙ s') ≡ (x , p) + isSymFiberRel'' _ _ (s , _) .fst = sym s + isSymFiberRel'' (x , p) (x' , p') (s , q) .snd i .fst = q (~ i) .fst + isSymFiberRel'' (x , p) (x' , p') (s , q) .snd i .snd j = + hcomp (λ k → λ { (i = i0) → compPath-filler p' (sym s) k j + ; (i = i1) → compPath-filler p s (~ k) j + ; (j = i0) → f (q (~ i) .fst) + ; (j = i1) → s (~ k) }) + (q (~ i) .snd j) + + isSymFiberRel' : (x x' : fiber f y) + → fiberRel2 ∣ x ∣₂ ∣ x' ∣₂ → fiberRel2 ∣ x' ∣₂ ∣ x ∣₂ + isSymFiberRel' _ _ = truncRel` _ _ ∘ Prop.rec isPropPropTrunc (λ r → ∣ isSymFiberRel'' _ _ r ∣) ∘ truncRel' _ _ + + isSymFiberRel : isSym fiberRel2 + isSymFiberRel = Set.elim2 (λ _ _ → isProp→isSet (isPropΠ (λ _ → isPropPropTrunc))) isSymFiberRel' + + isTransFiberRel'' : ((x , p) (x' , p') (x'' , p'') : fiber f y) + → Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ (x' , p') → Σ[ s' ∈ y ≡ y ] (x' , p' ∙ s') ≡ (x'' , p'') + → Σ[ s'' ∈ y ≡ y ] (x , p ∙ s'') ≡ (x'' , p'') + isTransFiberRel'' (x , p) (x' , p') (x'' , p'') (s , q) (s' , q') = + let sq : (i j : I) → Y + sq i j = + hcomp (λ k → λ { (i = i0) → p j + ; (i = i1) → fiberRel2'→fiberRel' _ _ s' q' .snd .snd k j + ; (j = i0) → f (compPath-filler (cong fst q) (cong fst q') k i) + ; (j = i1) → compPath-filler s s' k i }) + (fiberRel2'→fiberRel' _ _ s q .snd .snd i j) + in fiberRel'→fiberRel2' _ _ (_ , _ , (λ i j → sq i j)) + + isTransFiberRel' : (x x' x'' : fiber f y) + → fiberRel2 ∣ x ∣₂ ∣ x' ∣₂ → fiberRel2 ∣ x' ∣₂ ∣ x'' ∣₂ → fiberRel2 ∣ x ∣₂ ∣ x'' ∣₂ + isTransFiberRel' _ _ _ a b = + truncRel` _ _ (Prop.rec2 isPropPropTrunc + (λ r r' → ∣ isTransFiberRel'' _ _ _ r r' ∣) (truncRel' _ _ a) (truncRel' _ _ b)) + + isTransFiberRel : isTrans fiberRel2 + isTransFiberRel = + Set.elim3 (λ _ _ _ → isProp→isSet (isPropΠ2 (λ _ _ → isPropPropTrunc))) isTransFiberRel' + + open isEquivRel + + isEquivRelFiberRel : isEquivRel fiberRel2 + isEquivRelFiberRel .reflexive = isReflFiberRel + isEquivRelFiberRel .symmetric = isSymFiberRel + isEquivRelFiberRel .transitive = isTransFiberRel From cf76611afe7a8ca57cf7865639b62f41562f4c72 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Wed, 24 Nov 2021 11:43:47 +0800 Subject: [PATCH 02/24] fix typos --- Cubical/Experiments/CountingFiniteStructure.agda | 16 +++++----------- 1 file changed, 5 insertions(+), 11 deletions(-) diff --git a/Cubical/Experiments/CountingFiniteStructure.agda b/Cubical/Experiments/CountingFiniteStructure.agda index bd224686e9..b473290622 100644 --- a/Cubical/Experiments/CountingFiniteStructure.agda +++ b/Cubical/Experiments/CountingFiniteStructure.agda @@ -49,22 +49,16 @@ FinSemiGroupStr X .snd = isFinSetΣ (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ p → _ , isFinSetΠ3 X (λ _ → X) (λ _ _ → X) (λ _ _ _ → _ , isFinSet≡ X _ _)) -FinSemiGroupOfCard : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) -FinSemiGroupOfCard ℓ n = FinSetWithStrOfCard ℓ FinSemiGroupStr n - -isFinSemiGroupOfCard : (ℓ : Level) (n : ℕ) → isFinType 0 (FinSemiGroupOfCard ℓ n) -isFinSemiGroupOfCard ℓ n = isFinTypeFinSetWithStrOfCard _ FinSemiGroupStr _ - -- two rather trivial numbers -- but the computation is essentially not that trivial -- this one can be computed in half-a-minute -g : ℕ -g = card (_ , isFinStrCard TrivialStr 2) +a2 : ℕ +a2 = card (_ , isFinStrCard TrivialStr 2) -- this is already hard to compute --- don't know if anyone's computer give the final result -f : ℕ -f = card (_ , isFinStrCard IdentityStr 2) +-- it takes less than half-an-hour +b2 : ℕ +b2 = card (_ , isFinStrCard IdentityStr 2) -- the number of finite semi-groups with cardinal 2 -- it should be 5 From c443f9673fbccb5ea5e011355b0f2c5c7caf17d2 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Wed, 24 Nov 2021 17:09:27 +0800 Subject: [PATCH 03/24] add the definition of OEIS-A000001 --- .../Experiments/CountingFiniteStructure.agda | 25 +++++++++++++++++++ Cubical/HITs/SetQuotients/Properties.agda | 1 - 2 files changed, 25 insertions(+), 1 deletion(-) diff --git a/Cubical/Experiments/CountingFiniteStructure.agda b/Cubical/Experiments/CountingFiniteStructure.agda index b473290622..4e650d04ab 100644 --- a/Cubical/Experiments/CountingFiniteStructure.agda +++ b/Cubical/Experiments/CountingFiniteStructure.agda @@ -49,6 +49,26 @@ FinSemiGroupStr X .snd = isFinSetΣ (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ p → _ , isFinSetΠ3 X (λ _ → X) (λ _ _ → X) (λ _ _ _ → _ , isFinSet≡ X _ _)) +-- finite groups + +FinGroupStr : FinSet ℓ → FinSet ℓ +FinGroupStr X .fst = + Σ[ e ∈ X .fst ] + Σ[ inv ∈ (X .fst → X .fst) ] + Σ[ p ∈ (X .fst → X .fst → X .fst) ] + ((x y z : X .fst) → p (p x y) z ≡ p x (p y z)) + × ((x : X .fst) + → (p x e ≡ x) × (p e x ≡ x) × (p (inv x) x ≡ e) × (p x (inv x) ≡ e)) +FinGroupStr X .snd = + isFinSetΣ X (λ _ → _ , + isFinSetΣ (_ , isFinSetΠ X (λ _ → X)) (λ _ → _ , + isFinSetΣ (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ _ → _ , + isFinSet× (_ , isFinSetΠ3 X (λ _ → X) (λ _ _ → X) (λ _ _ _ → _ , isFinSet≡ X _ _)) (_ , + isFinSetΠ X (λ _ → _ , + isFinSet× (_ , isFinSet≡ X _ _) (_ , + isFinSet× (_ , isFinSet≡ X _ _) (_ , + isFinSet× (_ , isFinSet≡ X _ _) (_ , isFinSet≡ X _ _)))))))) + -- two rather trivial numbers -- but the computation is essentially not that trivial -- this one can be computed in half-a-minute @@ -65,3 +85,8 @@ b2 = card (_ , isFinStrCard IdentityStr 2) -- would you like to try? n2 : ℕ n2 = card (_ , isFinStrCard FinSemiGroupStr 2) + +-- OEIS-A000001 +-- I think you'd better not evaluate this function +numberOfFinGroups : ℕ → ℕ +numberOfFinGroups n = card (_ , isFinStrCard FinGroupStr n) diff --git a/Cubical/HITs/SetQuotients/Properties.agda b/Cubical/HITs/SetQuotients/Properties.agda index 4ccaf0238d..7986b3c2b8 100644 --- a/Cubical/HITs/SetQuotients/Properties.agda +++ b/Cubical/HITs/SetQuotients/Properties.agda @@ -18,7 +18,6 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Powerset open import Cubical.Functions.FunExtEquiv From db0b5fcf1aa41c70e298bfb34c95e4b8a7bbcb95 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Wed, 24 Nov 2021 20:35:57 +0800 Subject: [PATCH 04/24] add comments --- Cubical/Data/FinSet/FinType.agda | 7 +++---- Cubical/Data/FinSet/FiniteStructure.agda | 6 ++++++ 2 files changed, 9 insertions(+), 4 deletions(-) diff --git a/Cubical/Data/FinSet/FinType.agda b/Cubical/Data/FinSet/FinType.agda index a480ca6cab..0772961db5 100644 --- a/Cubical/Data/FinSet/FinType.agda +++ b/Cubical/Data/FinSet/FinType.agda @@ -2,10 +2,9 @@ Definition and properties of finite types -This file formalize the notion of Rijke finite type, -which is a direct generalization of Bishop finite set. -Basically, a type is (Rijke) n-finite if all its i-th -order homotopy groups πᵢ for i ≤ n are Bishop finite. +This file formalize the notion of Rijke finite type, which is a direct generalization +of Bishop finite set. Basically, a type is (Rijke) n-finite if all its i-th order +homotopy groups πᵢ for i ≤ n are Bishop finite. https://github.com/EgbertRijke/OEIS-A000001 diff --git a/Cubical/Data/FinSet/FiniteStructure.agda b/Cubical/Data/FinSet/FiniteStructure.agda index d938dd9bed..276337d9a0 100644 --- a/Cubical/Data/FinSet/FiniteStructure.agda +++ b/Cubical/Data/FinSet/FiniteStructure.agda @@ -1,3 +1,9 @@ +{ + +Definition and properties of finite sets equipped with finite structures, +namely the type of structures over a given finte set is a finite set. + +} {-# OPTIONS --safe #-} module Cubical.Data.FinSet.FiniteStructure where From 6f49e5b436b04585d2cbef512ec39932d419b064 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Wed, 24 Nov 2021 20:53:43 +0800 Subject: [PATCH 05/24] fix --- Cubical/Data/FinSet/FiniteStructure.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Cubical/Data/FinSet/FiniteStructure.agda b/Cubical/Data/FinSet/FiniteStructure.agda index 276337d9a0..099c571399 100644 --- a/Cubical/Data/FinSet/FiniteStructure.agda +++ b/Cubical/Data/FinSet/FiniteStructure.agda @@ -1,9 +1,9 @@ -{ +{- Definition and properties of finite sets equipped with finite structures, namely the type of structures over a given finte set is a finite set. -} +-} {-# OPTIONS --safe #-} module Cubical.Data.FinSet.FiniteStructure where From 3e0f71a03baa04be309c20b181fc4897d6ab724b Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Mon, 29 Nov 2021 03:10:56 +0800 Subject: [PATCH 06/24] major refactoring --- Cubical/Data/Bool/Base.agda | 4 + Cubical/Data/Bool/Properties.agda | 125 ++++++++++++- Cubical/Data/Fin/Base.agda | 3 + Cubical/Data/FinSet/Base.agda | 54 +++++- Cubical/Data/FinSet/Cardinality.agda | 138 +++++++------- Cubical/Data/FinSet/Constructors.agda | 119 ++++++------ Cubical/Data/FinSet/Decidability.agda | 165 ++++++++++++++++ Cubical/Data/FinSet/FinType.agda | 13 +- Cubical/Data/FinSet/FiniteChoice.agda | 41 ++-- Cubical/Data/FinSet/FiniteStructure.agda | 8 +- Cubical/Data/FinSet/Induction.agda | 32 ++-- Cubical/Data/FinSet/Properties.agda | 55 +++--- Cubical/Data/FinSet/Quotients.agda | 161 ++++++++-------- Cubical/Data/Sum/Properties.agda | 4 + Cubical/Data/SumFin/Properties.agda | 176 +++++++++++++++++- Cubical/Experiments/Combinatorics.agda | 18 +- .../Experiments/CountingFiniteStructure.agda | 125 ++++++++++++- Cubical/HITs/SetQuotients/EqClass.agda | 23 +-- 18 files changed, 951 insertions(+), 313 deletions(-) create mode 100644 Cubical/Data/FinSet/Decidability.agda diff --git a/Cubical/Data/Bool/Base.agda b/Cubical/Data/Bool/Base.agda index f2676e7354..a9d4cc7de6 100644 --- a/Cubical/Data/Bool/Base.agda +++ b/Cubical/Data/Bool/Base.agda @@ -64,6 +64,10 @@ Bool→Type : Bool → Type₀ Bool→Type true = Unit Bool→Type false = ⊥ +Bool→Type* : Bool → Type ℓ +Bool→Type* true = Unit* +Bool→Type* false = ⊥* + True : Dec A → Type₀ True Q = Bool→Type (Dec→Bool Q) diff --git a/Cubical/Data/Bool/Properties.agda b/Cubical/Data/Bool/Properties.agda index c899f997d7..b521a530e1 100644 --- a/Cubical/Data/Bool/Properties.agda +++ b/Cubical/Data/Bool/Properties.agda @@ -6,6 +6,7 @@ open import Cubical.Core.Everything open import Cubical.Functions.Involution open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Transport @@ -14,11 +15,19 @@ open import Cubical.Foundations.Pointed open import Cubical.Data.Bool.Base open import Cubical.Data.Empty +open import Cubical.Data.Sum hiding (rec) open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation hiding (rec) + open import Cubical.Relation.Nullary open import Cubical.Relation.Nullary.DecidableEq +private + variable + ℓ : Level + A : Type ℓ + notnot : ∀ x → not (not x) ≡ x notnot true = refl notnot false = refl @@ -39,9 +48,6 @@ notEq : Bool ≡ Bool notEq = involPath {f = not} notnot private - variable - ℓ : Level - -- This computes to false as expected nfalse : Bool nfalse = transp (λ i → notEq i) i0 true @@ -219,3 +225,116 @@ Iso.rightInv IsoBool→∙ a = refl Iso.leftInv IsoBool→∙ (f , p) = ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) + +-- relation to hProp + +open import Cubical.Data.Unit + +BoolProp≃BoolProp* : {a : Bool} → Bool→Type a ≃ Bool→Type* {ℓ} a +BoolProp≃BoolProp* {a = true} = Unit≃Unit* +BoolProp≃BoolProp* {a = false} = uninhabEquiv rec rec* + +Bool→PropInj : (a b : Bool) → Bool→Type a ≃ Bool→Type b → a ≡ b +Bool→PropInj true true _ = refl +Bool→PropInj true false p = rec (p .fst tt) +Bool→PropInj false true p = rec (invEq p tt) +Bool→PropInj false false _ = refl + +Bool→PropInj* : (a b : Bool) → Bool→Type* {ℓ} a ≃ Bool→Type* {ℓ} b → a ≡ b +Bool→PropInj* true true _ = refl +Bool→PropInj* true false p = rec* (p .fst tt*) +Bool→PropInj* false true p = rec* (invEq p tt*) +Bool→PropInj* false false _ = refl + +isPropBool→Prop : {a : Bool} → isProp (Bool→Type a) +isPropBool→Prop {a = true} = isPropUnit +isPropBool→Prop {a = false} = isProp⊥ + +isPropBool→Prop* : {a : Bool} → isProp (Bool→Type* {ℓ} a) +isPropBool→Prop* {a = true} = isPropUnit* +isPropBool→Prop* {a = false} = isProp⊥* + +DecBool→Prop : {a : Bool} → Dec (Bool→Type a) +DecBool→Prop {a = true} = yes tt +DecBool→Prop {a = false} = no (λ x → x) + +DecBool→Prop* : {a : Bool} → Dec (Bool→Type* {ℓ} a) +DecBool→Prop* {a = true} = yes tt* +DecBool→Prop* {a = false} = no (λ (lift x) → x) + +Dec→DecBool : {P : Type ℓ} → (dec : Dec P) → P → Bool→Type (Dec→Bool dec) +Dec→DecBool (yes p) _ = tt +Dec→DecBool (no ¬p) q = rec (¬p q) + +DecBool→Dec : {P : Type ℓ} → (dec : Dec P) → Bool→Type (Dec→Bool dec) → P +DecBool→Dec (yes p) _ = p + +Dec≃DecBool : {P : Type ℓ} → (h : isProp P)(dec : Dec P) → P ≃ Bool→Type (Dec→Bool dec) +Dec≃DecBool h dec = propBiimpl→Equiv h isPropBool→Prop (Dec→DecBool dec) (DecBool→Dec dec) + +Bool≡BoolDec : {a : Bool} → a ≡ Dec→Bool (DecBool→Prop {a = a}) +Bool≡BoolDec {a = true} = refl +Bool≡BoolDec {a = false} = refl + +Dec→DecBool* : {P : Type ℓ} → (dec : Dec P) → P → Bool→Type* {ℓ} (Dec→Bool dec) +Dec→DecBool* (yes p) _ = tt* +Dec→DecBool* (no ¬p) q = rec (¬p q) + +DecBool→Dec* : {P : Type ℓ} → (dec : Dec P) → Bool→Type* {ℓ} (Dec→Bool dec) → P +DecBool→Dec* (yes p) _ = p + +Dec≃DecBool* : {P : Type ℓ} → (h : isProp P)(dec : Dec P) → P ≃ Bool→Type* {ℓ} (Dec→Bool dec) +Dec≃DecBool* h dec = propBiimpl→Equiv h isPropBool→Prop* (Dec→DecBool* dec) (DecBool→Dec* dec) + +Bool≡BoolDec* : {a : Bool} → a ≡ Dec→Bool (DecBool→Prop* {ℓ} {a = a}) +Bool≡BoolDec* {a = true} = refl +Bool≡BoolDec* {a = false} = refl + +Bool→Prop× : (a b : Bool) → Bool→Type (a and b) → Bool→Type a × Bool→Type b +Bool→Prop× true true _ = tt , tt +Bool→Prop× true false p = rec p +Bool→Prop× false true p = rec p +Bool→Prop× false false p = rec p + +Bool→Prop×' : (a b : Bool) → Bool→Type a × Bool→Type b → Bool→Type (a and b) +Bool→Prop×' true true _ = tt +Bool→Prop×' true false (_ , p) = rec p +Bool→Prop×' false true (p , _) = rec p +Bool→Prop×' false false (p , _) = rec p + +Bool→Prop≃ : (a b : Bool) → Bool→Type a × Bool→Type b ≃ Bool→Type (a and b) +Bool→Prop≃ a b = + propBiimpl→Equiv (isProp× isPropBool→Prop isPropBool→Prop) isPropBool→Prop + (Bool→Prop×' a b) (Bool→Prop× a b) + +Bool→Prop⊎ : (a b : Bool) → Bool→Type (a or b) → Bool→Type a ⊎ Bool→Type b +Bool→Prop⊎ true true _ = inl tt +Bool→Prop⊎ true false _ = inl tt +Bool→Prop⊎ false true _ = inr tt +Bool→Prop⊎ false false p = rec p + +Bool→Prop⊎' : (a b : Bool) → Bool→Type a ⊎ Bool→Type b → Bool→Type (a or b) +Bool→Prop⊎' true true _ = tt +Bool→Prop⊎' true false _ = tt +Bool→Prop⊎' false true _ = tt +Bool→Prop⊎' false false (inl p) = rec p +Bool→Prop⊎' false false (inr p) = rec p + +PropBoolP→P : (dec : Dec A) → Bool→Type (Dec→Bool dec) → A +PropBoolP→P (yes p) _ = p + +P→PropBoolP : (dec : Dec A) → A → Bool→Type (Dec→Bool dec) +P→PropBoolP (yes p) _ = tt +P→PropBoolP (no ¬p) = ¬p + +Bool≡ : Bool → Bool → Bool +Bool≡ true true = true +Bool≡ true false = false +Bool≡ false true = false +Bool≡ false false = true + +Bool≡≃ : (a b : Bool) → (a ≡ b) ≃ Bool→Type (Bool≡ a b) +Bool≡≃ true true = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) +Bool≡≃ true false = uninhabEquiv true≢false rec +Bool≡≃ false true = uninhabEquiv false≢true rec +Bool≡≃ false false = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) diff --git a/Cubical/Data/Fin/Base.agda b/Cubical/Data/Fin/Base.agda index 306ac502af..2f9b6f5bc8 100644 --- a/Cubical/Data/Fin/Base.agda +++ b/Cubical/Data/Fin/Base.agda @@ -68,6 +68,9 @@ fsplit (suc k , k0→isInhab : card X > 0 → ∥ X .fst ∥ card>0→isInhab p = - rec isPropPropTrunc (λ e → ∣ invEq e (Fin>0 _ p) ∣) (∣≃card∣ X) + Prop.rec isPropPropTrunc (λ e → ∣ invEq e (Fin>0 _ p) ∣) (∣≃card∣ X) card>1→hasNonEqualTerm : card X > 1 → ∥ Σ[ a ∈ X .fst ] Σ[ b ∈ X .fst ] ¬ a ≡ b ∥ card>1→hasNonEqualTerm p = - rec isPropPropTrunc + Prop.rec isPropPropTrunc (λ e → ∣ e .fst (Fin>1 _ p .fst) , e .fst (Fin>1 _ p .snd .fst) , Fin>1 _ p .snd .snd ∘ invEq (congEquiv e) ∣) (∣invEquiv∣ (∣≃card∣ X)) card≡1→isContr : card X ≡ 1 → isContr (X .fst) card≡1→isContr p = - rec isPropIsContr - (λ e → isOfHLevelRespectEquiv 0 (invEquiv (e ⋆ pathToEquiv (cong Fin p))) isContrFin1) (∣≃card∣ X) + Prop.rec isPropIsContr + (λ e → isOfHLevelRespectEquiv 0 (invEquiv (e ⋆ pathToEquiv (cong Fin p))) isContrSumFin1) (∣≃card∣ X) card≤1→isProp : card X ≤ 1 → isProp (X .fst) card≤1→isProp p = - rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 (invEquiv e) (Fin≤1 (card X) p)) (∣≃card∣ X) + Prop.rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 (invEquiv e) (Fin≤1 (card X) p)) (∣≃card∣ X) card≡n : card X ≡ n → ∥ X ≡ 𝔽in n ∥ card≡n {n = n} p = - rec isPropPropTrunc + Prop.rec isPropPropTrunc (λ e → ∣(λ i → ua e i , isProp→PathP {B = λ j → isFinSet (ua e j)} @@ -127,34 +125,34 @@ module _ 1 (FinSet≡ X 𝟙) (isOfHLevel≡ 1 (card≤1→isProp (subst (λ a → a ≤ 1) (sym p) (≤-solver 1 1))) (isPropUnit*))) .fst - (rec isPropPropTrunc (λ q → ∣ q ∙ 𝔽in1≡𝟙 ∣) (card≡n p)) + (Prop.rec isPropPropTrunc (λ q → ∣ q ∙ 𝔽in1≡𝟙 ∣) (card≡n p)) module _ (X : FinSet ℓ) where isEmpty→card≡0 : ¬ X .fst → card X ≡ 0 isEmpty→card≡0 p = - rec (isSetℕ _ _) (λ e → sym (emptyFin _ (p ∘ invEq e))) (∣≃card∣ X) + Prop.rec (isSetℕ _ _) (λ e → sym (emptyFin _ (p ∘ invEq e))) (∣≃card∣ X) isInhab→card>0 : ∥ X .fst ∥ → card X > 0 - isInhab→card>0 = rec2 m≤n-isProp (λ p x → nonEmptyFin _ (p .fst x)) (∣≃card∣ X) + isInhab→card>0 = Prop.rec2 m≤n-isProp (λ p x → nonEmptyFin _ (p .fst x)) (∣≃card∣ X) hasNonEqualTerm→card>1 : {a b : X. fst} → ¬ a ≡ b → card X > 1 hasNonEqualTerm→card>1 {a = a} {b = b} q = - rec m≤n-isProp (λ p → nonEqualTermFin _ (p .fst a) (p .fst b) (q ∘ invEq (congEquiv p))) (∣≃card∣ X) + Prop.rec m≤n-isProp (λ p → nonEqualTermFin _ (p .fst a) (p .fst b) (q ∘ invEq (congEquiv p))) (∣≃card∣ X) isContr→card≡1 : isContr (X .fst) → card X ≡ 1 isContr→card≡1 p = cardEquiv X (_ , isFinSetUnit) ∣ isContr→≃Unit p ∣ isProp→card≤1 : isProp (X .fst) → card X ≤ 1 - isProp→card≤1 p = propFin (card X) (rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 e p) (∣≃card∣ X)) + isProp→card≤1 p = propFin (card X) (Prop.rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 e p) (∣≃card∣ X)) {- formulae about cardinality -} --- results to be used in direct induction on FinSet +-- results to be used in diProp.rect induction on FinSet card𝟘 : card (𝟘 {ℓ}) ≡ 0 -card𝟘 {ℓ = ℓ} = isEmpty→card≡0 (𝟘 {ℓ}) (rec*) +card𝟘 {ℓ = ℓ} = isEmpty→card≡0 (𝟘 {ℓ}) Empty.rec* card𝟙 : card (𝟙 {ℓ}) ≡ 1 card𝟙 {ℓ = ℓ} = isContr→card≡1 (𝟙 {ℓ}) isContrUnit* @@ -169,18 +167,10 @@ module _ (Y : FinSet ℓ') where card+ : card (_ , isFinSet⊎ X Y) ≡ card X + card Y - card+ = - cardEquiv (_ , isFinSet⊎ X Y) (Fin (card X + card Y) , isFinSetFin) - (rec2 isPropPropTrunc - (λ e1 e2 → ∣ ⊎-equiv e1 e2 ⋆ invEquiv (isoToEquiv (Fin+≅Fin⊎Fin _ _)) ∣) - (∣≃card∣ X) (∣≃card∣ Y)) + card+ = refl card× : card (_ , isFinSet× X Y) ≡ card X · card Y - card× = - cardEquiv (_ , isFinSet× X Y) (Fin (card X · card Y) , isFinSetFin) - (rec2 isPropPropTrunc - (λ e1 e2 → ∣ Σ-cong-equiv e1 (λ _ → e2) ⋆ factorEquiv ∣) - (∣≃card∣ X) (∣≃card∣ Y)) + card× = refl -- total summation/product of numerical functions from finite sets @@ -251,10 +241,6 @@ module _ prod𝔽in1+n : prod (𝔽in (1 + n)) f ≡ f (inl tt*) · prod (𝔽in n) (f ∘ inr) prod𝔽in1+n = prod⊎ 𝟙 (𝔽in n) f ∙ (λ i → prod𝟙 (f ∘ inl) i · prod (𝔽in n) (f ∘ inr)) -_^_ : ℕ → ℕ → ℕ -m ^ 0 = 1 -m ^ (suc n) = m · m ^ n - sumConst𝔽in : (n : ℕ)(f : 𝔽in {ℓ} n .fst → ℕ)(c : ℕ)(h : (x : 𝔽in n .fst) → f x ≡ c) → sum (𝔽in n) f ≡ c · n sumConst𝔽in 0 f c _ = sum𝟘 f ∙ 0≡m·0 c sumConst𝔽in (suc n) f c h = @@ -300,7 +286,7 @@ sum≤𝔽in (suc n) f g h = sum<𝔽in : (n : ℕ)(f g : 𝔽in {ℓ} n .fst → ℕ)(t : ∥ 𝔽in {ℓ} n .fst ∥)(h : (x : 𝔽in n .fst) → f x < g x) → sum (𝔽in n) f < sum (𝔽in n) g -sum<𝔽in {ℓ = ℓ} 0 _ _ t _ = EmptyRec (<→≢ (isInhab→card>0 (𝔽in 0) t) (card𝟘 {ℓ = ℓ})) +sum<𝔽in {ℓ = ℓ} 0 _ _ t _ = Empty.rec (<→≢ (isInhab→card>0 (𝔽in 0) t) (card𝟘 {ℓ = ℓ})) sum<𝔽in (suc n) f g t h = ≡< (h (inl tt*)) (sum≤𝔽in n (f ∘ inr) (g ∘ inr) (<-weaken ∘ h ∘ inr)) (sum𝔽in1+n n f) (sum𝔽in1+n n g) @@ -361,13 +347,13 @@ module _ cardΣ : card (_ , isFinSetΣ X Y) ≡ sum X (λ x → card (Y x)) cardΣ = cardEquiv (_ , isFinSetΣ X Y) (_ , isFinSetΣ X (λ x → Fin (card (Y x)) , isFinSetFin)) - (rec isPropPropTrunc (λ e → ∣ Σ-cong-equiv-snd e ∣) + (Prop.rec isPropPropTrunc (λ e → ∣ Σ-cong-equiv-snd e ∣) (choice X (λ x → Y x .fst ≃ Fin (card (Y x))) (λ x → ∣≃card∣ (Y x)))) cardΠ : card (_ , isFinSetΠ X Y) ≡ prod X (λ x → card (Y x)) cardΠ = cardEquiv (_ , isFinSetΠ X Y) (_ , isFinSetΠ X (λ x → Fin (card (Y x)) , isFinSetFin)) - (rec isPropPropTrunc (λ e → ∣ equivΠCod e ∣) + (Prop.rec isPropPropTrunc (λ e → ∣ equivΠCod e ∣) (choice X (λ x → Y x .fst ≃ Fin (card (Y x))) (λ x → ∣≃card∣ (Y x)))) module _ @@ -377,12 +363,11 @@ module _ card→ : card (_ , isFinSet→ X Y) ≡ card Y ^ card X card→ = cardΠ X (λ _ → Y) ∙ prodConst X (λ _ → card Y) (card Y) (λ _ → refl) - card≃ : card (_ , isFinSet≃ X X) ≡ factorial (card X) - card≃ = - cardEquiv (_ , isFinSet≃ X X) (Fin (factorial (card X)) , isFinSetFin) - (rec isPropPropTrunc - (λ e → ∣ equivComp e e ⋆ lehmerEquiv ⋆ lehmerFinEquiv ∣) - (∣≃card∣ X)) +module _ + (X : FinSet ℓ ) where + + cardAut : card (_ , isFinSetAut X) ≡ LehmerCode.factorial (card X) + cardAut = refl module _ (X : FinSet ℓ ) @@ -434,7 +419,7 @@ module _ Σ∥P∥→∥ΣP∥ : Σ X (λ x → ∥ P x ∥) → ∥ Σ X P ∥ Σ∥P∥→∥ΣP∥ (x , p) = - rec isPropPropTrunc (λ q → ∣ x , q ∣) p + Prop.rec isPropPropTrunc (λ q → ∣ x , q ∣) p module _ (f : X .fst → Y .fst) @@ -454,8 +439,8 @@ module _ pigeonHole' : ∥ Σ[ x ∈ X .fst ] Σ[ x' ∈ X .fst ] (¬ x ≡ x') × (f x ≡ f x') ∥ pigeonHole' = - rec isPropPropTrunc (λ p → ∣ nonInj p ∣) - (rec isPropPropTrunc fiberNonEqualTerm + Prop.rec isPropPropTrunc (λ p → ∣ nonInj p ∣) + (Prop.rec isPropPropTrunc fiberNonEqualTerm (pigeonHole {X = X} {Y = Y} f 1 (subst (λ a → _ > a) (sym (·-identityˡ _)) p))) -- cardinality and injection/surjection @@ -485,10 +470,10 @@ module _ (λ y → isInhab→card>0 (_ , isFinSetFiber X Y f y) (p y))) card↪Inequality : ∥ X .fst ↪ Y .fst ∥ → card X ≤ card Y - card↪Inequality = rec m≤n-isProp (λ (f , p) → card↪Inequality' f p) + card↪Inequality = Prop.rec m≤n-isProp (λ (f , p) → card↪Inequality' f p) card↠Inequality : ∥ X .fst ↠ Y .fst ∥ → card X ≥ card Y - card↠Inequality = rec m≤n-isProp (λ (f , p) → card↠Inequality' f p) + card↠Inequality = Prop.rec m≤n-isProp (λ (f , p) → card↠Inequality' f p) -- maximal value of numerical functions @@ -539,7 +524,7 @@ module _ ΣMax⊎-case (x , p) (y , q) (gt r) .snd (inr y') = ≤-trans (q y') (<-weaken r) ∃Max⊎ : ∃Max X (f ∘ inl) → ∃Max Y (f ∘ inr) → ∃Max (X ⊎ Y) f - ∃Max⊎ = rec2 isPropPropTrunc (λ p q → ∣ ΣMax⊎-case p q (_≟_ _ _) ∣) + ∃Max⊎ = Prop.rec2 isPropPropTrunc (λ p q → ∣ ΣMax⊎-case p q (_≟_ _ _) ∣) ΣMax𝟙 : (f : 𝟙 {ℓ} .fst → ℕ) → ΣMax _ f ΣMax𝟙 f .fst = tt* @@ -549,7 +534,7 @@ module _ ∃Max𝟙 f = ∣ ΣMax𝟙 f ∣ ∃Max𝔽in : (n : ℕ)(f : 𝔽in {ℓ} n .fst → ℕ)(x : ∥ 𝔽in {ℓ} n .fst ∥) → ∃Max _ f -∃Max𝔽in {ℓ = ℓ} 0 _ x = EmptyRec (<→≢ (isInhab→card>0 (𝔽in 0) x) (card𝟘 {ℓ = ℓ})) +∃Max𝔽in {ℓ = ℓ} 0 _ x = Empty.rec (<→≢ (isInhab→card>0 (𝔽in 0) x) (card𝟘 {ℓ = ℓ})) ∃Max𝔽in 1 f _ = subst (λ X → (f : X .fst → ℕ) → ∃Max _ f) (sym 𝔽in1≡𝟙) ∃Max𝟙 f ∃Max𝔽in (suc (suc n)) f _ = @@ -574,14 +559,14 @@ module _ open Iso Iso-∥FinSet∥₂-ℕ : Iso ∥ FinSet ℓ ∥₂ ℕ -Iso-∥FinSet∥₂-ℕ .fun = SetRec isSetℕ card +Iso-∥FinSet∥₂-ℕ .fun = Set.rec isSetℕ card Iso-∥FinSet∥₂-ℕ .inv n = ∣ 𝔽in n ∣₂ Iso-∥FinSet∥₂-ℕ .rightInv n = card𝔽in n Iso-∥FinSet∥₂-ℕ {ℓ = ℓ} .leftInv = - SetElim {B = λ X → ∣ 𝔽in (SetRec isSetℕ card X) ∣₂ ≡ X} - (λ X → isSetPathImplicit) - (elimProp (λ X → ∣ 𝔽in (card X) ∣₂ ≡ ∣ X ∣₂) (λ X → squash₂ _ _) - (λ n i → ∣ 𝔽in (card𝔽in {ℓ = ℓ} n i) ∣₂)) + Set.elim {B = λ X → ∣ 𝔽in (Set.rec isSetℕ card X) ∣₂ ≡ X} + (λ X → isSetPathImplicit) + (elimProp (λ X → ∣ 𝔽in (card X) ∣₂ ≡ ∣ X ∣₂) (λ X → squash₂ _ _) + (λ n i → ∣ 𝔽in (card𝔽in {ℓ = ℓ} n i) ∣₂)) -- this is the definition of natural numbers you learned from school ∥FinSet∥₂≃ℕ : ∥ FinSet ℓ ∥₂ ≃ ℕ @@ -596,8 +581,8 @@ Bool→FinProp false = 𝟘 , isProp⊥* injBool→FinProp : (x y : Bool) → Bool→FinProp {ℓ = ℓ} x ≡ Bool→FinProp y → x ≡ y injBool→FinProp true true _ = refl injBool→FinProp false false _ = refl -injBool→FinProp true false p = EmptyRec (snotz (cong (card ∘ fst) p)) -injBool→FinProp false true p = EmptyRec (znots (cong (card ∘ fst) p)) +injBool→FinProp true false p = Empty.rec (snotz (cong (card ∘ fst) p)) +injBool→FinProp false true p = Empty.rec (znots (cong (card ∘ fst) p)) isEmbeddingBool→FinProp : isEmbedding (Bool→FinProp {ℓ = ℓ}) isEmbeddingBool→FinProp = injEmbedding isSetBool isSetFinProp (λ {x} {y} → injBool→FinProp x y) @@ -606,7 +591,7 @@ card-case : (P : FinProp ℓ) → {n : ℕ} → card (P .fst) ≡ n → Σ[ x card-case P {n = 0} p = false , FinProp≡ (𝟘 , isProp⊥*) P .fst (cong fst (sym (card≡0 {X = P .fst} p))) card-case P {n = 1} p = true , FinProp≡ (𝟙 , isPropUnit*) P .fst (cong fst (sym (card≡1 {X = P .fst} p))) card-case P {n = suc (suc n)} p = - EmptyRec (¬-<-zero (pred-≤-pred (subst (λ a → a ≤ 1) p (isProp→card≤1 (P .fst) (P .snd))))) + Empty.rec (¬-<-zero (pred-≤-pred (subst (λ a → a ≤ 1) p (isProp→card≤1 (P .fst) (P .snd))))) isSurjectionBool→FinProp : isSurjection (Bool→FinProp {ℓ = ℓ}) isSurjectionBool→FinProp P = ∣ card-case P refl ∣ @@ -618,3 +603,28 @@ FinProp≃Bool = isFinSetFinProp : isFinSet (FinProp ℓ) isFinSetFinProp = EquivPresIsFinSet (invEquiv FinProp≃Bool) isFinSetBool + +-- a more efficient version of equivalence type + +module _ + (X : FinSet ℓ ) + (Y : FinSet ℓ') where + + isFinSet≃Eff' : Dec (card X ≡ card Y) → isFinSet (X .fst ≃ Y .fst) + isFinSet≃Eff' (yes p) = factorial (card Y) , + Prop.elim2 (λ _ _ → isPropPropTrunc {A = _ ≃ Fin _}) + (λ p1 p2 + → ∣ equivComp (p1 ⋆ pathToEquiv (cong Fin p) ⋆ SumFin≃Fin _) (p2 ⋆ SumFin≃Fin _) + ⋆ lehmerEquiv ⋆ lehmerFinEquiv + ⋆ invEquiv (SumFin≃Fin _) ∣) + (∣≃card∣ X) (∣≃card∣ Y) + isFinSet≃Eff' (no ¬p) = 0 , ∣ uninhabEquiv (¬p ∘ cardEquiv X Y ∘ ∣_∣) (idfun _) ∣ + + isFinSet≃Eff : isFinSet (X .fst ≃ Y .fst) + isFinSet≃Eff = isFinSet≃Eff' (discreteℕ _ _) + +module _ + (X Y : FinSet ℓ) where + + isFinSetType≡Eff : isFinSet (X .fst ≡ Y .fst) + isFinSetType≡Eff = EquivPresIsFinSet (invEquiv univalence) (isFinSet≃Eff X Y) diff --git a/Cubical/Data/FinSet/Constructors.agda b/Cubical/Data/FinSet/Constructors.agda index ad00a39cc9..da926cd1c3 100644 --- a/Cubical/Data/FinSet/Constructors.agda +++ b/Cubical/Data/FinSet/Constructors.agda @@ -12,16 +12,16 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv open import Cubical.Foundations.Univalence -open import Cubical.HITs.PropositionalTruncation renaming (rec to TruncRec) +open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.Data.Nat open import Cubical.Data.Unit -open import Cubical.Data.Empty renaming (rec to EmptyRec) +open import Cubical.Data.Empty as Empty open import Cubical.Data.Sum open import Cubical.Data.Sigma -open import Cubical.Data.Fin -open import Cubical.Data.SumFin renaming (Fin to SumFin) +open import Cubical.Data.Fin.LehmerCode as LehmerCode +open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.FiniteChoice @@ -37,55 +37,46 @@ private ℓ ℓ' ℓ'' ℓ''' : Level module _ - (X : Type ℓ)(p : ≃Fin X) where + (X : Type ℓ)(p : isFinOrd X) where - ≃Fin∥∥ : ≃Fin ∥ X ∥ - ≃Fin∥∥ = ≃SumFin→Fin (_ , compEquiv (propTrunc≃ (≃Fin→SumFin p .snd)) (SumFin∥∥≃ _)) + isFinOrd∥∥ : isFinOrd ∥ X ∥ + isFinOrd∥∥ = _ , propTrunc≃ (p .snd) ⋆ SumFin∥∥≃ _ + + isFinOrd≃ : isFinOrd (X ≃ X) + isFinOrd≃ = _ , equivComp (p .snd) (p .snd) ⋆ SumFin≃≃ _ module _ - (X : Type ℓ )(p : ≃Fin X) - (Y : Type ℓ')(q : ≃Fin Y) where + (X : Type ℓ )(p : isFinOrd X) + (Y : Type ℓ')(q : isFinOrd Y) where - ≃Fin⊎ : ≃Fin (X ⊎ Y) - ≃Fin⊎ = ≃SumFin→Fin (_ , compEquiv (⊎-equiv (≃Fin→SumFin p .snd) (≃Fin→SumFin q .snd)) (SumFin⊎≃ _ _)) + isFinOrd⊎ : isFinOrd (X ⊎ Y) + isFinOrd⊎ = _ , ⊎-equiv (p .snd) (q .snd) ⋆ SumFin⊎≃ _ _ - ≃Fin× : ≃Fin (X × Y) - ≃Fin× = ≃SumFin→Fin (_ , compEquiv (Σ-cong-equiv (≃Fin→SumFin p .snd) (λ _ → ≃Fin→SumFin q .snd)) (SumFin×≃ _ _)) + isFinOrd× : isFinOrd (X × Y) + isFinOrd× = _ , Σ-cong-equiv (p .snd) (λ _ → q .snd) ⋆ SumFin×≃ _ _ module _ - (X : Type ℓ )(p : ≃Fin X) - (Y : X → Type ℓ')(q : (x : X) → ≃Fin (Y x)) where + (X : Type ℓ )(p : isFinOrd X) + (Y : X → Type ℓ')(q : (x : X) → isFinOrd (Y x)) where private - p' = ≃Fin→SumFin p - - m = p' .fst - e = p' .snd - - q' : (x : X) → ≃SumFin (Y x) - q' x = ≃Fin→SumFin (q x) + e = p .snd f : (x : X) → ℕ - f x = q' x .fst + f x = q x .fst - ≃SumFinΣ : ≃SumFin (Σ X Y) - ≃SumFinΣ = _ , - Σ-cong-equiv {B' = λ x → Y (invEq (p' .snd) x)} (p' .snd) (transpFamily p') - ⋆ Σ-cong-equiv-snd (λ x → q' (invEq e x) .snd) + isFinOrdΣ : isFinOrd (Σ X Y) + isFinOrdΣ = _ , + Σ-cong-equiv {B' = λ x → Y (invEq e x)} e (transpFamily p) + ⋆ Σ-cong-equiv-snd (λ x → q (invEq e x) .snd) ⋆ SumFinΣ≃ _ _ - ≃SumFinΠ : ≃SumFin ((x : X) → Y x) - ≃SumFinΠ = _ , - equivΠ {B' = λ x → Y (invEq (p' .snd) x)} (p' .snd) (transpFamily p') - ⋆ equivΠCod (λ x → q' (invEq e x) .snd) + isFinOrdΠ : isFinOrd ((x : X) → Y x) + isFinOrdΠ = _ , + equivΠ {B' = λ x → Y (invEq e x)} e (transpFamily p) + ⋆ equivΠCod (λ x → q (invEq e x) .snd) ⋆ SumFinΠ≃ _ _ - ≃FinΣ : ≃Fin (Σ X Y) - ≃FinΣ = ≃SumFin→Fin ≃SumFinΣ - - ≃FinΠ : ≃Fin ((x : X) → Y x) - ≃FinΠ = ≃SumFin→Fin ≃SumFinΠ - -- closedness under several type constructors module _ @@ -95,14 +86,14 @@ module _ isFinSetΣ : isFinSet (Σ (X .fst) (λ x → Y x .fst)) isFinSetΣ = elim2 (λ _ _ → isPropIsFinSet {A = Σ (X .fst) (λ x → Y x .fst)}) - (λ p q → ∣ ≃FinΣ (X .fst) p (λ x → Y x .fst) q ∣) - (X .snd) (choice X (λ x → ≃Fin (Y x .fst)) (λ x → Y x .snd)) + (λ p q → isFinOrd→isFinSet (isFinOrdΣ (X .fst) (_ , p) (λ x → Y x .fst) q)) + (X .snd .snd) (choice X (λ x → isFinOrd (Y x .fst)) (λ x → isFinSet→isFinSet' (Y x .snd))) isFinSetΠ : isFinSet ((x : X .fst) → Y x .fst) isFinSetΠ = elim2 (λ _ _ → isPropIsFinSet {A = ((x : X .fst) → Y x .fst)}) - (λ p q → ∣ ≃FinΠ (X .fst) p (λ x → Y x .fst) q ∣) - (X .snd) (choice X (λ x → ≃Fin (Y x .fst)) (λ x → Y x .snd)) + (λ p q → isFinOrd→isFinSet (isFinOrdΠ (X .fst) (_ , p) (λ x → Y x .fst) q)) + (X .snd .snd) (choice X (λ x → isFinOrd (Y x .fst)) (λ x → isFinSet→isFinSet' (Y x .snd))) module _ (X : FinSet ℓ) @@ -128,7 +119,7 @@ module _ isFinSet≡ a b = isDecProp→isFinSet (isFinSet→isSet (X .snd) a b) (isFinSet→Discrete (X .snd) a b) isFinSet∥∥ : isFinSet ∥ X .fst ∥ - isFinSet∥∥ = TruncRec isPropIsFinSet (λ p → ∣ ≃Fin∥∥ (X .fst) p ∣) (X .snd) + isFinSet∥∥ = Prop.rec isPropIsFinSet (λ p → isFinOrd→isFinSet (isFinOrd∥∥ (X .fst) (_ , p))) (X .snd .snd) isFinSetIsContr : isFinSet (isContr (X .fst)) isFinSetIsContr = isFinSetΣ X (λ x → _ , (isFinSetΠ X (λ y → _ , isFinSet≡ x y))) @@ -155,10 +146,16 @@ module _ (Y : FinSet ℓ') where isFinSet⊎ : isFinSet (X .fst ⊎ Y .fst) - isFinSet⊎ = elim2 (λ _ _ → isPropIsFinSet) (λ p q → ∣ ≃Fin⊎ (X .fst) p (Y .fst) q ∣) (X .snd) (Y .snd) + isFinSet⊎ = card X + card Y , + elim2 (λ _ _ → isPropPropTrunc {A = _ ≃ Fin (card X + card Y)}) + (λ p q → ∣ isFinOrd⊎ (X .fst) (_ , p) (Y .fst) (_ , q) .snd ∣) + (X .snd .snd) (Y .snd .snd) isFinSet× : isFinSet (X .fst × Y .fst) - isFinSet× = elim2 (λ _ _ → isPropIsFinSet) (λ p q → ∣ ≃Fin× (X .fst) p (Y .fst) q ∣) (X .snd) (Y .snd) + isFinSet× = card X · card Y , + elim2 (λ _ _ → isPropPropTrunc {A = _ ≃ Fin (card X · card Y)}) + (λ p q → ∣ isFinOrd× (X .fst) (_ , p) (Y .fst) (_ , q) .snd ∣) + (X .snd .snd) (Y .snd .snd) isFinSet→ : isFinSet (X .fst → Y .fst) isFinSet→ = isFinSetΠ X (λ _ → Y) @@ -172,11 +169,23 @@ module _ isFinSetType≡ : isFinSet (X .fst ≡ Y .fst) isFinSetType≡ = EquivPresIsFinSet (invEquiv univalence) (isFinSet≃ X Y) +module _ + (X : FinSet ℓ) where + + isFinSetAut : isFinSet (X .fst ≃ X .fst) + isFinSetAut = LehmerCode.factorial (card X) , + Prop.elim (λ _ → isPropPropTrunc {A = _ ≃ Fin _}) + (λ p → ∣ isFinOrd≃ (X .fst) (card X , p) .snd ∣) + (X .snd .snd) + + isFinSetTypeAut : isFinSet (X .fst ≡ X .fst) + isFinSetTypeAut = EquivPresIsFinSet (invEquiv univalence) isFinSetAut + module _ (X : FinSet ℓ) where isFinSet¬ : isFinSet (¬ (X .fst)) - isFinSet¬ = isFinSet→ X (⊥ , ∣ 0 , uninhabEquiv (λ x → x) ¬Fin0 ∣) + isFinSet¬ = isFinSet→ X (Fin 0 , isFinSetFin) module _ (X : FinSet ℓ) where @@ -225,21 +234,5 @@ module _ (Y : Type ℓ')(h : Discrete Y) (f : X. fst → Y) where - isFinSetFiberDec : (y : Y) → isFinSet (fiber f y) - isFinSetFiberDec y = isFinSetΣ X (λ x → _ , isDecProp→isFinSet (Discrete→isSet h _ _) (h (f x) y)) - --- decidable predicate - -module _ - (X : FinSet ℓ) - (P : X .fst → Type ℓ') - (h : (x : X .fst) → isProp (P x)) - (dec : (x : X .fst) → Dec (P x)) where - - DecΣ : Dec ∥ Σ (X .fst) P ∥ - DecΣ = isFinSet→Dec∥∥ (isFinSetΣ X (λ x → _ , isDecProp→isFinSet (h x) (dec x))) - - DecΠ : Dec ((x : X .fst) → P x) - DecΠ = - EquivPresDec (propTruncIdempotent≃ (isPropΠ h)) - (isFinSet→Dec∥∥ (isFinSetΠ X (λ x → _ , isDecProp→isFinSet (h x) (dec x)))) + isFinSetFiberDisc : (y : Y) → isFinSet (fiber f y) + isFinSetFiberDisc y = isFinSetΣ X (λ x → _ , isDecProp→isFinSet (Discrete→isSet h _ _) (h (f x) y)) diff --git a/Cubical/Data/FinSet/Decidability.agda b/Cubical/Data/FinSet/Decidability.agda new file mode 100644 index 0000000000..f8326a81ce --- /dev/null +++ b/Cubical/Data/FinSet/Decidability.agda @@ -0,0 +1,165 @@ +{-# OPTIONS --safe #-} + +module Cubical.Data.FinSet.Decidability where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties + +open import Cubical.HITs.PropositionalTruncation as Prop + +open import Cubical.Data.Bool +open import Cubical.Data.Empty as Empty +open import Cubical.Data.Sigma + +open import Cubical.Data.Fin +open import Cubical.Data.SumFin renaming (Fin to SumFin) +open import Cubical.Data.FinSet.Base +open import Cubical.Data.FinSet.Properties + +open import Cubical.Relation.Nullary + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + +isDecProp : Type ℓ → Type ℓ +isDecProp P = Σ[ t ∈ Bool ] P ≃ Bool→Type t + +isDecProp→isProp : {P : Type ℓ} → isDecProp P → isProp P +isDecProp→isProp h = isOfHLevelRespectEquiv 1 (invEquiv (h .snd)) isPropBool→Prop + +isDecProp→Dec : {P : Type ℓ} → isDecProp P → Dec P +isDecProp→Dec h = EquivPresDec (invEquiv (h .snd)) DecBool→Prop + +helper : (P : Type ℓ) → (t : Bool) → isProp (P ≃ Bool→Type t) +helper _ _ = isOfHLevel⁺≃ᵣ 0 isPropBool→Prop + +helper' : (P : Type ℓ) → (p q : isDecProp P) → (p .fst ≡ q .fst) ≃ (p ≡ q) +helper' _ _ _ = Σ≡PropEquiv (helper _) + +isPropIsDecProp : {P : Type ℓ} → isProp (isDecProp P) +isPropIsDecProp p q = + Σ≡PropEquiv (λ _ → isOfHLevel⁺≃ᵣ 0 isPropBool→Prop) .fst + (Bool→PropInj _ _ (invEquiv (p .snd) ⋆ q .snd)) + +DecProp : (ℓ : Level) → Type (ℓ-suc ℓ) +DecProp ℓ = Σ[ P ∈ Type ℓ ] isDecProp P + +module _ + (X : Type ℓ)(p : isFinOrd X) where + + isDecProp∥∥' : isDecProp ∥ X ∥ + isDecProp∥∥' = _ , propTrunc≃ (p .snd) ⋆ SumFin∥∥DecProp _ + +module _ + (X : Type ℓ )(p : isFinOrd X) + (P : X → Type ℓ') + (dec : (x : X) → isDecProp (P x)) where + + private + e = p .snd + + isFinOrdSub : isFinOrd (Σ X P) + isFinOrdSub = _ , + Σ-cong-equiv {B' = λ x → P (invEq e x)} e (transpFamily p) + ⋆ Σ-cong-equiv-snd (λ x → dec (invEq e x) .snd) + ⋆ SumFinSub≃ _ (fst ∘ dec ∘ invEq e) + + isDecProp∃' : isDecProp ∥ Σ X P ∥ + isDecProp∃' = _ , + Prop.propTrunc≃ ( + Σ-cong-equiv {B' = λ x → P (invEq e x)} e (transpFamily p) + ⋆ Σ-cong-equiv-snd (λ x → dec (invEq e x) .snd)) + ⋆ SumFin∃≃ _ (fst ∘ dec ∘ invEq e) + + isDecProp∀' : isDecProp ((x : X) → P x) + isDecProp∀' = _ , + equivΠ {B' = λ x → P (invEq e x)} e (transpFamily p) + ⋆ equivΠCod (λ x → dec (invEq e x) .snd) + ⋆ SumFin∀≃ _ (fst ∘ dec ∘ invEq e) + +module _ + (X : Type ℓ )(p : isFinOrd X) + (a b : X) where + + private + e = p .snd + + isDecProp≡' : isDecProp (a ≡ b) + isDecProp≡' .fst = SumFin≡ _ (e .fst a) (e .fst b) + isDecProp≡' .snd = congEquiv e ⋆ SumFin≡≃ _ _ _ + +module _ + (X : FinSet ℓ) + (P : X .fst → DecProp ℓ') where + + isFinSetSub : isFinSet (Σ (X .fst) (λ x → P x .fst)) + isFinSetSub = + Prop.elim + (λ _ → isPropIsFinSet) + (λ p → isFinOrd→isFinSet (isFinOrdSub (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd))) + (X .snd .snd) + + isDecProp∃ : isDecProp ∥ Σ (X .fst) (λ x → P x .fst) ∥ + isDecProp∃ = + Prop.elim + (λ _ → isPropIsDecProp) + (λ p → isDecProp∃' (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd)) + (X .snd .snd) + + isDecProp∀ : isDecProp ((x : X .fst) → P x .fst) + isDecProp∀ = + Prop.elim + (λ _ → isPropIsDecProp) + (λ p → isDecProp∀' (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd)) + (X .snd .snd) + +module _ + (X : FinSet ℓ) + (Y : X .fst → FinSet ℓ') + (Z : (x : X .fst) → Y x .fst → DecProp ℓ'') where + + isDecProp∀2 : isDecProp ((x : X .fst) → (y : Y x .fst) → Z x y .fst) + isDecProp∀2 = isDecProp∀ X (λ x → _ , isDecProp∀ (Y x) (Z x)) + +module _ + (X : FinSet ℓ) + (Y : X .fst → FinSet ℓ') + (Z : (x : X .fst) → Y x .fst → FinSet ℓ'') + (W : (x : X .fst) → (y : Y x .fst) → Z x y .fst → DecProp ℓ''') where + + isDecProp∀3 : isDecProp ((x : X .fst) → (y : Y x .fst) → (z : Z x y .fst) → W x y z .fst) + isDecProp∀3 = isDecProp∀ X (λ x → _ , isDecProp∀2 (Y x) (Z x) (W x)) + +module _ + (X : FinSet ℓ) where + + isDecProp≡ : (a b : X .fst) → isDecProp (a ≡ b) + isDecProp≡ a b = + Prop.rec isPropIsDecProp + (λ p → isDecProp≡' (X .fst) (_ , p) a b) + (X .snd .snd) + +module _ + (P : DecProp ℓ ) + (Q : DecProp ℓ') where + + isDecProp× : isDecProp (P .fst × Q .fst) + isDecProp× .fst = P .snd .fst and Q .snd .fst + isDecProp× .snd = Σ-cong-equiv (P .snd .snd) (λ _ → Q .snd .snd) ⋆ Bool→Prop≃ _ _ + +module _ + (X : FinSet ℓ) where +{- + isDecProp¬ : isDecProp (¬ (X .fst)) + isDecProp¬ .fst = {!!} + isDecProp¬ .snd = {!!} -} + + isDecProp∥∥ : isDecProp ∥ X .fst ∥ + isDecProp∥∥ = + Prop.rec isPropIsDecProp + (λ p → isDecProp∥∥' (X .fst) (_ , p)) + (X .snd .snd) diff --git a/Cubical/Data/FinSet/FinType.agda b/Cubical/Data/FinSet/FinType.agda index 0772961db5..81fb59b69f 100644 --- a/Cubical/Data/FinSet/FinType.agda +++ b/Cubical/Data/FinSet/FinType.agda @@ -31,6 +31,7 @@ open import Cubical.Data.Sigma open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties +open import Cubical.Data.FinSet.Decidability open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Quotients open import Cubical.Data.FinSet.Cardinality @@ -74,7 +75,7 @@ isFinType→isFinType0 {n = suc n} p = p .fst isFinTypeSuc→isFinType1 : isFinType (suc n) X → isFinType 1 X isFinTypeSuc→isFinType1 {n = 0} p = p isFinTypeSuc→isFinType1 {n = suc n} p .fst = p .fst -isFinTypeSuc→isFinType1 {n = suc n} p .snd a b = isFinType→isFinType0 (p .snd a b) +isFinTypeSuc→isFinType1 {n = suc n} p .snd a b = isFinType→isFinType0 {n = suc n} (p .snd a b) -- useful properties @@ -106,14 +107,13 @@ module _ module _ (y : Y) where - isDecFiberRel : (x x' : ∥ fiber f y ∥₂) → Dec (fiberRel2 _ _ f y x x') - isDecFiberRel x x' = - isFinSet→Dec∥∥ (isFinSetΣ (_ , h .snd y y) (λ _ → _ , isFinSet≡ (_ , q y) _ _)) + isDecPropFiberRel : (x x' : ∥ fiber f y ∥₂) → isDecProp (fiberRel2 _ _ f y x x') + isDecPropFiberRel x x' = isDecProp∃ (_ , h .snd y y) (λ _ → _ , isDecProp≡ (_ , q y) _ _) isFinSetFiber∥∥₂' : isFinSet (fiber ∥f∥₂ ∣ y ∣₂) isFinSetFiber∥∥₂' = EquivPresIsFinSet (∥fiber∥₂/Rel≃fiber∥∥₂ _ _ f y) - (isFinSetQuot (_ , q y) (fiberRel2 _ _ _ _) (isEquivRelFiberRel _ _ _ _) isDecFiberRel) + (isFinSetQuot (_ , q y) (fiberRel2 _ _ _ _) (isEquivRelFiberRel _ _ _ _) isDecPropFiberRel) isFinSetFiber∥∥₂ : (y : ∥ Y ∥₂) → isFinSet (fiber ∥f∥₂ y) isFinSetFiber∥∥₂ = Set.elim (λ _ → isProp→isSet isPropIsFinSet) isFinSetFiber∥∥₂' @@ -139,7 +139,8 @@ isFinTypeΣ : → isFinType n (Σ (X .fst) (λ x → Y x .fst)) isFinTypeΣ {n = 0} = isFinType0Σ isFinTypeΣ {n = suc n} X Y .fst = - isFinType0Σ (_ , isFinTypeSuc→isFinType1 (X .snd)) (λ x → _ , isFinType→isFinType0 (Y x .snd)) + isFinType0Σ (_ , isFinTypeSuc→isFinType1 {n = suc n} (X .snd)) + (λ x → _ , isFinType→isFinType0 {n = suc n} (Y x .snd)) isFinTypeΣ {n = suc n} X Y .snd a b = EquivPresIsFinType n (ΣPathTransport≃PathΣ a b) (isFinTypeΣ {n = n} (_ , X .snd .snd _ _) (λ _ → _ , Y _ .snd .snd _ _)) diff --git a/Cubical/Data/FinSet/FiniteChoice.agda b/Cubical/Data/FinSet/FiniteChoice.agda index 0fe23135d4..fc933298b0 100644 --- a/Cubical/Data/FinSet/FiniteChoice.agda +++ b/Cubical/Data/FinSet/FiniteChoice.agda @@ -13,7 +13,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv -open import Cubical.HITs.PropositionalTruncation renaming (rec to TruncRec ; elim to TruncElim) +open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.Data.Nat open import Cubical.Data.Unit @@ -21,52 +21,51 @@ open import Cubical.Data.Empty open import Cubical.Data.Sum open import Cubical.Data.Sigma -open import Cubical.Data.SumFin renaming (Fin to SumFin) +open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties private variable - ℓ ℓ' ℓ'' : Level + ℓ ℓ' : Level -choice≃SumFin : - {n : ℕ}(Y : SumFin n → Type ℓ) → ((x : SumFin n) → ∥ Y x ∥) ≃ ∥ ((x : SumFin n) → Y x) ∥ -choice≃SumFin {n = 0} Y = +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}))) -choice≃SumFin {n = suc n} Y = +choice≃Fin {n = suc n} Y = Π⊎≃ ⋆ Σ-cong-equiv-fst (ΠUnit (λ x → ∥ Y (inl x) ∥)) - ⋆ Σ-cong-equiv-snd (λ _ → choice≃SumFin {n = n} (λ x → Y (inr x))) + ⋆ Σ-cong-equiv-snd (λ _ → choice≃Fin {n = n} (λ x → Y (inr x))) ⋆ Σ-cong-equiv-fst (propTrunc≃ (invEquiv (ΠUnit (λ x → Y (inl x))))) ⋆ ∥∥-×-≃ ⋆ propTrunc≃ (invEquiv (Π⊎≃ {E = Y})) module _ - (X : Type ℓ)(p : ≃Fin X) + (X : Type ℓ)(p : isFinOrd X) (Y : X → Type ℓ') where private - q = ≃Fin→SumFin p - e = q .snd + e = p .snd choice≃' : ((x : X) → ∥ Y x ∥) ≃ ∥ ((x : X) → Y x) ∥ choice≃' = - equivΠ {B' = λ x → ∥ Y (invEq e x) ∥} e (transpFamily q) - ⋆ choice≃SumFin _ - ⋆ propTrunc≃ (invEquiv (equivΠ {B' = λ x → Y (invEq e x)} e (transpFamily q))) + equivΠ {B' = λ x → ∥ Y (invEq e x) ∥} e (transpFamily p) + ⋆ choice≃Fin _ + ⋆ propTrunc≃ (invEquiv (equivΠ {B' = λ x → Y (invEq e x)} e (transpFamily p))) module _ - ((X , p) : FinSet ℓ) - (Y : X → Type ℓ') where + (X : FinSet ℓ) + (Y : X .fst → Type ℓ') where - choice≃ : ((x : X) → ∥ Y x ∥) ≃ ∥ ((x : X) → Y x) ∥ + choice≃ : ((x : X .fst) → ∥ Y x ∥) ≃ ∥ ((x : X .fst) → Y x) ∥ choice≃ = - TruncRec - (isOfHLevel≃ 1 (isPropΠ (λ x → isPropPropTrunc)) (isPropPropTrunc)) - (λ p → choice≃' X p Y) p + Prop.rec + (isOfHLevel≃ 1 (isPropΠ (λ x → isPropPropTrunc)) isPropPropTrunc) + (λ p → choice≃' (X .fst) (_ , p) Y) (X .snd .snd) - choice : ((x : X) → ∥ Y x ∥) → ∥ ((x : X) → Y x) ∥ + choice : ((x : X .fst) → ∥ Y x ∥) → ∥ ((x : X .fst) → Y x) ∥ choice = choice≃ .fst diff --git a/Cubical/Data/FinSet/FiniteStructure.agda b/Cubical/Data/FinSet/FiniteStructure.agda index 099c571399..e14f47deed 100644 --- a/Cubical/Data/FinSet/FiniteStructure.agda +++ b/Cubical/Data/FinSet/FiniteStructure.agda @@ -61,13 +61,9 @@ isPathConnectedFinSetOfCard {n = n} .snd = isFinTypeFinSetOfCard : isFinType 1 (FinSetOfCard ℓ n) isFinTypeFinSetOfCard .fst = isPathConnected→isFinType0 isPathConnectedFinSetOfCard isFinTypeFinSetOfCard .snd X Y = - isFinSet→isFinType _ (EquivPresIsFinSet (FinSet≡ _ _ ⋆ FinSetOfCard≡ _ _) (isFinSetType≡ (X .fst) (Y .fst))) + isFinSet→isFinType 0 (EquivPresIsFinSet (FinSet≡ _ _ ⋆ FinSetOfCard≡ _ _) (isFinSetType≡Eff (X .fst) (Y .fst))) isFinTypeFinSetWithStrOfCard : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') (n : ℕ) → isFinType 0 (FinSetWithStrOfCard ℓ S n) isFinTypeFinSetWithStrOfCard ℓ S n = - isFinTypeΣ (_ , isFinTypeFinSetOfCard) (λ X → _ , isFinSet→isFinType 0 (S (X .fst) .snd)) - --- we can counting the number of finite sets of a given cardinal with certain structures -cardFinSetWithStrOfCard : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') (n : ℕ) → ℕ -cardFinSetWithStrOfCard ℓ S n = card (_ , isFinTypeFinSetWithStrOfCard ℓ S n) + isFinTypeΣ {n = 0} (_ , isFinTypeFinSetOfCard) (λ X → _ , isFinSet→isFinType 0 (S (X .fst) .snd)) diff --git a/Cubical/Data/FinSet/Induction.agda b/Cubical/Data/FinSet/Induction.agda index fb4a703630..aae83bdcb0 100644 --- a/Cubical/Data/FinSet/Induction.agda +++ b/Cubical/Data/FinSet/Induction.agda @@ -13,16 +13,16 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence -open import Cubical.HITs.PropositionalTruncation renaming (rec to TruncRec) hiding (elim ; elim') -open import Cubical.HITs.SetTruncation renaming (rec to SetRec ; elim to SetElim) +open import Cubical.HITs.PropositionalTruncation as Prop +open import Cubical.HITs.SetTruncation as Set -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) hiding (elim) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Unit -open import Cubical.Data.Empty hiding (elim) -open import Cubical.Data.Sum hiding (elim) +open import Cubical.Data.Empty as Empty +open import Cubical.Data.Sum as Sum -open import Cubical.Data.Fin hiding (elim) -open import Cubical.Data.SumFin renaming (Fin to SumFin) hiding (elim) +open import Cubical.Data.Fin renaming (Fin to Finℕ) +open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.Constructors @@ -37,7 +37,7 @@ module _ {ℓ : Level} where 𝟘 : FinSet ℓ - 𝟘 = ⊥* , ∣ 0 , uninhabEquiv rec* ¬Fin0 ∣ + 𝟘 = ⊥* , 0 , ∣ uninhabEquiv Empty.rec* Empty.rec ∣ 𝟙 : FinSet ℓ 𝟙 = Unit* , isContr→isFinSet (isContrUnit*) @@ -61,12 +61,12 @@ module _ * : {n : ℕ} → 𝔽in (suc n) .fst * = inl tt* - 𝔽in≃SumFin : (n : ℕ) → 𝔽in n .fst ≃ SumFin n - 𝔽in≃SumFin 0 = 𝟘≃Empty - 𝔽in≃SumFin (suc n) = ⊎-equiv 𝟙≃Unit (𝔽in≃SumFin n) - 𝔽in≃Fin : (n : ℕ) → 𝔽in n .fst ≃ Fin n - 𝔽in≃Fin n = 𝔽in≃SumFin n ⋆ SumFin≃Fin n + 𝔽in≃Fin 0 = 𝟘≃Empty + 𝔽in≃Fin (suc n) = ⊎-equiv 𝟙≃Unit (𝔽in≃Fin n) + + 𝔽in≃Finℕ : (n : ℕ) → 𝔽in n .fst ≃ Finℕ n + 𝔽in≃Finℕ n = 𝔽in≃Fin n ⋆ SumFin≃Fin n -- 𝔽in preserves addition @@ -92,9 +92,9 @@ module _ -- every finite sets are merely equal to some 𝔽in ∣≡𝔽in∣ : (X : FinSet ℓ) → ∥ Σ[ n ∈ ℕ ] X ≡ 𝔽in n ∥ -∣≡𝔽in∣ X = TruncRec isPropPropTrunc (λ (n , p) → ∣ n , path X (n , p) ∣) (X .snd) +∣≡𝔽in∣ X = Prop.rec isPropPropTrunc (λ (n , p) → ∣ n , path X (n , p) ∣) (isFinSet→isFinSet' (X .snd)) where - path : (X : FinSet ℓ) → ((n , _) : ≃Fin (X .fst)) → X ≡ 𝔽in n + path : (X : FinSet ℓ) → ((n , _) : isFinOrd (X .fst)) → X ≡ 𝔽in n path X (n , p) i .fst = ua (p ⋆ invEquiv (𝔽in≃Fin n)) i path X (n , p) i .snd = isProp→PathP {B = λ i → isFinSet (path X (n , p) i .fst)} @@ -110,7 +110,7 @@ module _ (p : (n : ℕ) → P (𝔽in n)) where elimProp : (X : FinSet ℓ) → P X - elimProp X = TruncRec (h X) (λ (n , q) → transport (λ i → P (q (~ i))) (p n)) (∣≡𝔽in∣ X) + elimProp X = Prop.rec (h X) (λ (n , q) → transport (λ i → P (q (~ i))) (p n)) (∣≡𝔽in∣ X) module _ (p0 : P 𝟘) diff --git a/Cubical/Data/FinSet/Properties.agda b/Cubical/Data/FinSet/Properties.agda index eca4963cce..8c42b7f92b 100644 --- a/Cubical/Data/FinSet/Properties.agda +++ b/Cubical/Data/FinSet/Properties.agda @@ -10,16 +10,15 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.Equiv -open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.Data.Nat open import Cubical.Data.Unit open import Cubical.Data.Bool -open import Cubical.Data.Empty renaming (rec to EmptyRec) +open import Cubical.Data.Empty as Empty open import Cubical.Data.Sigma -open import Cubical.Data.Fin -open import Cubical.Data.SumFin renaming (Fin to SumFin) hiding (discreteFin) +open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Relation.Nullary @@ -37,52 +36,48 @@ private module _ {A : Type ℓ}{B : Type ℓ'}{C : Type ℓ''} where - infixr 30 _⋆_ infixr 30 _⋆̂_ - _⋆_ : A ≃ B → B ≃ C → A ≃ C - _⋆_ = compEquiv - _⋆̂_ : ∥ A ≃ B ∥ → ∥ B ≃ C ∥ → ∥ A ≃ C ∥ - _⋆̂_ = rec2 isPropPropTrunc (λ p q → ∣ p ⋆ q ∣) + _⋆̂_ = Prop.rec2 isPropPropTrunc (λ p q → ∣ p ⋆ q ∣) module _ {A : Type ℓ}{B : Type ℓ'} where ∣invEquiv∣ : ∥ A ≃ B ∥ → ∥ B ≃ A ∥ - ∣invEquiv∣ = rec isPropPropTrunc (λ p → ∣ invEquiv p ∣) + ∣invEquiv∣ = Prop.rec isPropPropTrunc (λ p → ∣ invEquiv p ∣) -- useful implications EquivPresIsFinSet : A ≃ B → isFinSet A → isFinSet B -EquivPresIsFinSet e = rec isPropIsFinSet (λ (n , p) → ∣ n , compEquiv (invEquiv e) p ∣) +EquivPresIsFinSet e (_ , p) = _ , ∣ invEquiv e ∣ ⋆̂ p isFinSetFin : {n : ℕ} → isFinSet (Fin n) -isFinSetFin = ∣ _ , pathToEquiv refl ∣ +isFinSetFin = _ , ∣ idEquiv _ ∣ isFinSetUnit : isFinSet Unit -isFinSetUnit = ∣ 1 , Unit≃Fin1 ∣ +isFinSetUnit = 1 , ∣ invEquiv Fin1≃Unit ∣ isFinSetBool : isFinSet Bool -isFinSetBool = ∣ 2 , invEquiv (SumFin2≃Bool) ⋆ SumFin≃Fin 2 ∣ +isFinSetBool = 2 , ∣ invEquiv SumFin2≃Bool ∣ isFinSet→Discrete : isFinSet A → Discrete A -isFinSet→Discrete = rec isPropDiscrete (λ (_ , p) → EquivPresDiscrete (invEquiv p) discreteFin) +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* ) ⋆ Unit≃Fin1 ∣ +isContr→isFinSet h = 1 , ∣ isContr→≃Unit* h ⋆ invEquiv Unit≃Unit* ⋆ invEquiv Fin1≃Unit ∣ isDecProp→isFinSet : isProp A → Dec A → isFinSet A isDecProp→isFinSet h (yes p) = isContr→isFinSet (inhProp→isContr p h) -isDecProp→isFinSet h (no ¬p) = ∣ 0 , uninhabEquiv ¬p ¬Fin0 ∣ +isDecProp→isFinSet h (no ¬p) = 0 , ∣ uninhabEquiv ¬p (idfun _) ∣ isDec→isFinSet∥∥ : Dec A → isFinSet ∥ A ∥ isDec→isFinSet∥∥ dec = isDecProp→isFinSet isPropPropTrunc (Dec∥∥ dec) isFinSet→Dec∥∥ : isFinSet A → Dec ∥ A ∥ -isFinSet→Dec∥∥ = - rec (isPropDec isPropPropTrunc) - (λ (_ , p) → EquivPresDec (propTrunc≃ (invEquiv p)) (∥Fin∥ _)) +isFinSet→Dec∥∥ h = + Prop.rec (isPropDec isPropPropTrunc) + (λ p → EquivPresDec (propTrunc≃ (invEquiv p)) (∥Fin∥ _)) (h .snd) isFinProp→Dec : isFinSet A → isProp A → Dec A isFinProp→Dec p h = subst Dec (propTruncIdempotent h) (isFinSet→Dec∥∥ p) @@ -93,6 +88,14 @@ PeirceLaw∥∥ p = Dec→Stable (isFinSet→Dec∥∥ p) PeirceLaw : isFinSet A → NonEmpty A → ∥ A ∥ PeirceLaw p q = PeirceLaw∥∥ p (λ f → q (λ x → f ∣ x ∣)) +-- transprot family towards Fin + +transpFamily : + {A : Type ℓ}{B : A → Type ℓ'} + → ((n , e) : isFinOrd A) → (x : A) → B x ≃ B (invEq e (e .fst x)) +transpFamily {B = B} (n , e) x = pathToEquiv (λ i → B (retEq e x (~ i))) + +{- {- Alternative definition of finite sets @@ -115,9 +118,9 @@ isPropIsFinSet' {A = A} (n , equivn) (m , equivm) = Σ≡Prop (λ _ → isPropPropTrunc) n≡m where Fin-n≃Fin-m : ∥ Fin n ≃ Fin m ∥ - Fin-n≃Fin-m = rec + Fin-n≃Fin-m = Prop.rec isPropPropTrunc - (rec + (Prop.rec (isPropΠ λ _ → isPropPropTrunc) (λ hm hn → ∣ Fin n ≃⟨ invEquiv hn ⟩ A ≃⟨ hm ⟩ Fin m ■ ∣) equivm @@ -125,13 +128,13 @@ isPropIsFinSet' {A = A} (n , equivn) (m , equivm) = equivn Fin-n≡Fin-m : ∥ Fin n ≡ Fin m ∥ - Fin-n≡Fin-m = rec isPropPropTrunc (∣_∣ ∘ ua) Fin-n≃Fin-m + Fin-n≡Fin-m = Prop.rec isPropPropTrunc (∣_∣ ∘ ua) Fin-n≃Fin-m ∥n≡m∥ : ∥ n ≡ m ∥ - ∥n≡m∥ = rec isPropPropTrunc (∣_∣ ∘ Fin-inj n m) Fin-n≡Fin-m + ∥n≡m∥ = Prop.rec isPropPropTrunc (∣_∣ ∘ Fin-inj n m) Fin-n≡Fin-m n≡m : n ≡ m - n≡m = rec (isSetℕ n m) (λ p → p) ∥n≡m∥ + n≡m = Prop.rec (isSetℕ n m) (λ p → p) ∥n≡m∥ -- logical equivalence of two definitions @@ -179,4 +182,4 @@ FinSet≡FinSet' = ua FinSet≃FinSet' transpFamily : {A : Type ℓ}{B : A → Type ℓ'} → ((n , e) : ≃SumFin A) → (x : A) → B x ≃ B (invEq e (e .fst x)) -transpFamily {B = B} (n , e) x = pathToEquiv (λ i → B (retEq e x (~ i))) +transpFamily {B = B} (n , e) x = pathToEquiv (λ i → B (retEq e x (~ i))) -} diff --git a/Cubical/Data/FinSet/Quotients.agda b/Cubical/Data/FinSet/Quotients.agda index 521aaaa846..6d3db0055f 100644 --- a/Cubical/Data/FinSet/Quotients.agda +++ b/Cubical/Data/FinSet/Quotients.agda @@ -6,15 +6,20 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Univalence open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.HITs.SetQuotients as SetQuot open import Cubical.HITs.SetQuotients.EqClass +open import Cubical.Data.Bool open import Cubical.Data.Sigma +open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties +open import Cubical.Data.FinSet.Decidability open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Cardinality @@ -25,102 +30,112 @@ private variable ℓ ℓ' ℓ'' : Level +LiftDecProp : {ℓ ℓ' : Level}{P : Type ℓ} → (p : isDecProp P) → P ≃ Bool→Type* {ℓ'} (p .fst) +LiftDecProp p = p .snd ⋆ BoolProp≃BoolProp* + open Iso module _ - (ℓ : Level) where - - DecProp : Type (ℓ-suc ℓ) - DecProp = Σ[ P ∈ hProp ℓ ] Dec (P .fst) - - isSetDecProp : isSet DecProp - isSetDecProp = isOfHLevelΣ 2 isSetHProp (λ P → isProp→isSet (isPropDec (P .snd))) - - Iso-DecProp-FinProp : Iso DecProp (FinProp ℓ) - Iso-DecProp-FinProp .fun ((P , p) , dec) = (P , isDecProp→isFinSet p dec) , p - Iso-DecProp-FinProp .inv ((P , h) , p) = (P , p) , isFinProp→Dec h p - Iso-DecProp-FinProp .leftInv ((P , p) , dec) i .fst .fst = P - Iso-DecProp-FinProp .leftInv ((P , p) , dec) i .fst .snd = p - Iso-DecProp-FinProp .leftInv ((P , p) , dec) i .snd = - isProp→PathP {B = λ i → Dec P} - (λ i → isPropDec (Iso-DecProp-FinProp .leftInv ((P , p) , dec) i .fst .snd)) - (Iso-DecProp-FinProp .inv (Iso-DecProp-FinProp .fun ((P , p) , dec)) .snd) dec i - Iso-DecProp-FinProp .rightInv ((P , h) , p) i .fst .fst = P - Iso-DecProp-FinProp .rightInv ((P , h) , p) i .fst .snd = - isProp→PathP {B = λ i → isFinSet P} (λ i → isPropIsFinSet) - (Iso-DecProp-FinProp .fun (Iso-DecProp-FinProp .inv ((P , h) , p)) .fst .snd) h i - Iso-DecProp-FinProp .rightInv ((P , h) , p) i .snd = p + (X : Type ℓ) where + + ℙEff : Type ℓ + ℙEff = X → Bool + + isSetℙEff : isSet ℙEff + isSetℙEff = isSetΠ (λ _ → isSetBool) + + ℙEff→ℙDec : ℙEff → ℙDec {ℓ' = ℓ'} X + ℙEff→ℙDec f .fst x = Bool→Type* (f x) , isPropBool→Prop* + ℙEff→ℙDec f .snd x = DecBool→Prop* + + Iso-ℙEff-ℙDec : Iso ℙEff (ℙDec {ℓ' = ℓ'} X) + Iso-ℙEff-ℙDec .fun = ℙEff→ℙDec + Iso-ℙEff-ℙDec .inv (P , dec) x = Dec→Bool (dec x) + Iso-ℙEff-ℙDec {ℓ' = ℓ'} .leftInv f i x = Bool≡BoolDec* {ℓ = ℓ'} {a = f x} (~ i) + Iso-ℙEff-ℙDec .rightInv (P , dec) i .fst x .fst = ua (Dec≃DecBool* (P x .snd) (dec x)) (~ i) + Iso-ℙEff-ℙDec .rightInv (P , dec) i .fst x .snd = + isProp→PathP {B = λ i → isProp (Iso-ℙEff-ℙDec .rightInv (P , dec) i .fst x .fst)} + (λ i → isPropIsProp) + (Iso-ℙEff-ℙDec .fun (Iso-ℙEff-ℙDec .inv (P , dec)) .fst x .snd) (P x .snd) i + Iso-ℙEff-ℙDec .rightInv (P , dec) i .snd x = + isProp→PathP {B = λ i → Dec (Iso-ℙEff-ℙDec .rightInv (P , dec) i .fst x .fst)} + (λ i → isPropDec (Iso-ℙEff-ℙDec .rightInv (P , dec) i .fst x .snd)) + (Iso-ℙEff-ℙDec .fun (Iso-ℙEff-ℙDec .inv (P , dec)) .snd x) (dec x) i + + ℙEff≃ℙDec : ℙEff ≃ (ℙDec {ℓ' = ℓ'} X) + ℙEff≃ℙDec = isoToEquiv Iso-ℙEff-ℙDec module _ - {ℓ ℓ' : Level} - (X : Type ℓ) where - - ℙFin : Type (ℓ-max ℓ (ℓ-suc ℓ')) - ℙFin = X → FinProp ℓ' - - isSetℙFin : isSet ℙFin - isSetℙFin = isSetΠ (λ _ → isSetFinProp) - - ℙFin→ℙDec : ℙFin → ℙDec {ℓ' = ℓ'} X - ℙFin→ℙDec P .fst x = P x .fst .fst , P x .snd - ℙFin→ℙDec P .snd x = isFinProp→Dec (P x .fst .snd) (P x .snd) - - Iso-ℙFin-ℙDec : Iso ℙFin (ℙDec {ℓ' = ℓ'} X) - Iso-ℙFin-ℙDec .fun = ℙFin→ℙDec - Iso-ℙFin-ℙDec .inv (P , dec) x .fst .fst = P x .fst - Iso-ℙFin-ℙDec .inv (P , dec) x .fst .snd = isDecProp→isFinSet (P x .snd) (dec x) - Iso-ℙFin-ℙDec .inv (P , dec) x .snd = P x .snd - Iso-ℙFin-ℙDec .leftInv P i x .fst .fst = P x .fst .fst - Iso-ℙFin-ℙDec .leftInv P i x .fst .snd = - isProp→PathP {B = λ i → isFinSet (P x .fst .fst)} (λ i → isPropIsFinSet) - (isDecProp→isFinSet (P x .snd) (ℙFin→ℙDec P .snd x)) (P x .fst .snd) i - Iso-ℙFin-ℙDec .leftInv P i x .snd = P x .snd - Iso-ℙFin-ℙDec .rightInv (P , dec) i .fst x = P x .fst , P x .snd - Iso-ℙFin-ℙDec .rightInv (P , dec) i .snd x = - isProp→PathP {B = λ i → Dec (P x .fst)} (λ i → isPropDec (P x .snd)) - (Iso-ℙFin-ℙDec .fun (Iso-ℙFin-ℙDec .inv (P , dec)) .snd x) (dec x) i - - ℙFin≃ℙDec : ℙFin ≃ (ℙDec {ℓ' = ℓ'} X) - ℙFin≃ℙDec = isoToEquiv Iso-ℙFin-ℙDec + (X : Type ℓ)(p : isFinOrd X) where + + private + e = p .snd + + isFinOrdℙEff : isFinOrd (ℙEff X) + isFinOrdℙEff = _ , preCompEquiv (invEquiv e) ⋆ SumFinℙ≃ _ module _ (X : FinSet ℓ) where - isFinSetℙFin : isFinSet (ℙFin {ℓ' = ℓ'} (X .fst)) - isFinSetℙFin = isFinSet→ X (_ , isFinSetFinProp) - -open BinaryRelation + isFinSetℙEff : isFinSet (ℙEff (X .fst)) + isFinSetℙEff = 2 ^ (card X) , + Prop.elim (λ _ → isPropPropTrunc {A = _ ≃ Fin _}) + (λ p → ∣ isFinOrdℙEff (X .fst) (_ , p) .snd ∣) + (X .snd .snd) module _ (X : FinSet ℓ) (R : X .fst → X .fst → Type ℓ') - (dec : (x x' : X .fst) → Dec (R x x')) where + (dec : (x x' : X .fst) → isDecProp (R x x')) where + + isEqClassEff : ℙEff (X .fst) → Type ℓ + isEqClassEff f = ∥ Σ[ x ∈ X .fst ] ((a : X .fst) → f a ≡ dec a x .fst) ∥ - _∥Fin_ : Type (ℓ-max ℓ (ℓ-suc ℓ')) - _∥Fin_ = Σ[ P ∈ ℙFin {ℓ' = ℓ'} (X .fst) ] isEqClass (X .fst) R (ℙFin→ℙDec (X .fst) P .fst) + isDecPropIsEqClassEff : {f : ℙEff (X .fst)} → isDecProp (isEqClassEff f) + isDecPropIsEqClassEff = isDecProp∃ X (λ _ → _ , isDecProp∀ X (λ _ → _ , _ , Bool≡≃ _ _)) - isFinSetIsEqClass : (P : ℙFin {ℓ' = ℓ'} (X .fst)) - → isFinSet (isEqClass (X .fst) R (ℙFin→ℙDec (X .fst) P .fst)) - isFinSetIsEqClass P = - isFinSet∥∥ (_ , isFinSetΣ X (λ x → _ , - isFinSetΠ X (λ a → _ , - isFinSetType≡ (_ , P a .fst .snd) (_ , - isDec→isFinSet∥∥ (dec a x))))) + isEqClassEff→isEqClass' : (f : ℙEff (X .fst))(x : X .fst) + → ((a : X .fst) → f a ≡ dec a x .fst) + → (a : X .fst) → Bool→Type* {ℓ = ℓ'} (f a) ≃ ∥ R a x ∥ + isEqClassEff→isEqClass' f x h a = + pathToEquiv (cong Bool→Type* (h a)) + ⋆ invEquiv (LiftDecProp (dec a x)) + ⋆ invEquiv (propTruncIdempotent≃ (isDecProp→isProp (dec a x))) - ∥Fin≃∥Dec : _∥Fin_ ≃ _∥Dec_ (X .fst) R dec - ∥Fin≃∥Dec = Σ-cong-equiv-fst (ℙFin≃ℙDec (X .fst)) + isEqClass→isEqClassEff' : (f : ℙEff (X .fst))(x : X .fst) + → ((a : X .fst) → Bool→Type* {ℓ = ℓ'} (f a) ≃ ∥ R a x ∥) + → (a : X .fst) → f a ≡ dec a x .fst + isEqClass→isEqClassEff' f x h a = + Bool→PropInj* _ _ + (h a ⋆ propTruncIdempotent≃ (isDecProp→isProp (dec a x)) ⋆ LiftDecProp (dec a x)) - isFinSet∥Fin : isFinSet _∥Fin_ - isFinSet∥Fin = isFinSetΣ (_ , isFinSetℙFin X) (λ P → _ , isFinSetIsEqClass P) + isEqClassRff→isEqClass : (f : ℙEff (X .fst)) → isEqClassEff f ≃ isEqClass {ℓ' = ℓ'} _ R (ℙEff→ℙDec _ f .fst) + isEqClassRff→isEqClass f = + propBiimpl→Equiv isPropPropTrunc isPropPropTrunc + (Prop.map (λ (x , p) → x , isEqClassEff→isEqClass' f x p)) + (Prop.map (λ (x , p) → x , isEqClass→isEqClassEff' f x p)) + + _∥Eff_ : Type ℓ + _∥Eff_ = Σ[ f ∈ ℙEff (X .fst) ] isEqClassEff f + + ∥Eff≃∥Dec : _∥Eff_ ≃ _∥Dec_ (X .fst) R (λ x x' → isDecProp→Dec (dec x x')) + ∥Eff≃∥Dec = Σ-cong-equiv (ℙEff≃ℙDec (X .fst)) isEqClassRff→isEqClass + + isFinSet∥Eff : isFinSet _∥Eff_ + isFinSet∥Eff = isFinSetSub (_ , isFinSetℙEff X) (λ _ → _ , isDecPropIsEqClassEff) + +open BinaryRelation module _ (X : FinSet ℓ) (R : X .fst → X .fst → Type ℓ') (h : isEquivRel R) - (dec : (x x' : X .fst) → Dec (R x x')) where + (dec : (x x' : X .fst) → isDecProp (R x x')) where - -- the quotient of finite set by decidable equivalence relation is still finite set isFinSetQuot : isFinSet (X .fst / R) isFinSetQuot = EquivPresIsFinSet - (∥Fin≃∥Dec X _ dec ⋆ ∥Dec≃∥ _ _ dec ⋆ invEquiv (equivQuot _ _ h)) (isFinSet∥Fin X _ dec) + ( ∥Eff≃∥Dec X _ dec + ⋆ ∥Dec≃∥ _ _ (λ x x' → isDecProp→Dec (dec x x')) + ⋆ invEquiv (equivQuot {ℓ' = ℓ'} _ _ h)) + (isFinSet∥Eff X _ dec) diff --git a/Cubical/Data/Sum/Properties.agda b/Cubical/Data/Sum/Properties.agda index c15ea12fb1..1629874566 100644 --- a/Cubical/Data/Sum/Properties.agda +++ b/Cubical/Data/Sum/Properties.agda @@ -184,3 +184,7 @@ leftInv Σ⊎Iso (inr b , eb) = refl Σ⊎≃ : (Σ (A ⊎ B) E) ≃ ((Σ A (λ a → E (inl a))) ⊎ (Σ B (λ b → E (inr b)))) Σ⊎≃ = isoToEquiv Σ⊎Iso + +map-⊎ : (A → C) → (B → D) → A ⊎ B → C ⊎ D +map-⊎ f _ (inl a) = inl (f a) +map-⊎ _ g (inr b) = inr (g b) diff --git a/Cubical/Data/SumFin/Properties.agda b/Cubical/Data/SumFin/Properties.agda index 943c12a7a3..7bf18cee99 100644 --- a/Cubical/Data/SumFin/Properties.agda +++ b/Cubical/Data/SumFin/Properties.agda @@ -2,22 +2,27 @@ module Cubical.Data.SumFin.Properties where -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Prelude open import Cubical.Foundations.Univalence open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Unit open import Cubical.Data.Bool open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order import Cubical.Data.Fin as Fin +import Cubical.Data.Fin.LehmerCode as LehmerCode open import Cubical.Data.SumFin.Base as SumFin open import Cubical.Data.Sum open import Cubical.Data.Sigma -open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation as Prop + +open import Cubical.Relation.Nullary private variable @@ -76,7 +81,7 @@ SumFin×≃ : (m n : ℕ) → (Fin m × Fin n) ≃ (Fin (m · n)) SumFin×≃ m n = SumFinΣ≃ m (λ _ → n) ⋆ pathToEquiv (λ i → Fin (totalSumConst {m = m} n i)) SumFinΠ≃ : (n : ℕ)(f : Fin n → ℕ) → ((x : Fin n) → Fin (f x)) ≃ (Fin (totalProd f)) -SumFinΠ≃ 0 f = isContr→≃Unit (isContrΠ⊥) ⋆ invEquiv (⊎-⊥-≃) +SumFinΠ≃ 0 _ = isContr→≃Unit (isContrΠ⊥) ⋆ invEquiv (⊎-⊥-≃) SumFinΠ≃ (suc n) f = Π⊎≃ ⋆ Σ-cong-equiv (ΠUnit (λ tt → Fin (f (inl tt)))) (λ _ → SumFinΠ≃ n (λ x → f (inr x))) @@ -92,6 +97,26 @@ SumFin∥∥≃ (suc n) = isContr→≃Unit (inhProp→isContr ∣ inl tt ∣ isPropPropTrunc) ⋆ isContr→≃Unit (isContrUnit) ⋆ invEquiv (⊎-⊥-≃) +SumFin∥∥Dec : (n : ℕ) → Dec ∥ Fin n ∥ +SumFin∥∥Dec 0 = no (Prop.rec isProp⊥ (idfun _)) +SumFin∥∥Dec (suc n) = yes ∣ inl tt ∣ + +isNonZero : ℕ → Bool +isNonZero 0 = false +isNonZero (suc n) = true + +SumFin∥∥DecProp : (n : ℕ) → ∥ Fin n ∥ ≃ Bool→Type (isNonZero n) +SumFin∥∥DecProp 0 = uninhabEquiv (Prop.rec isProp⊥ ⊥.rec) ⊥.rec +SumFin∥∥DecProp (suc n) = isContr→≃Unit (inhProp→isContr ∣ inl tt ∣ isPropPropTrunc) + +-- SumFin 1 is equivalent to unit + +Fin1≃Unit : Fin 1 ≃ Unit +Fin1≃Unit = ⊎-⊥-≃ + +isContrSumFin1 : isContr (Fin 1) +isContrSumFin1 = isOfHLevelRespectEquiv 0 (invEquiv Fin1≃Unit) isContrUnit + -- SumFin 2 is equivalent to Bool open Iso @@ -108,3 +133,146 @@ Iso-⊤⊎⊤-Bool .rightInv false = refl SumFin2≃Bool : Fin 2 ≃ Bool SumFin2≃Bool = ⊎-equiv (idEquiv _) ⊎-⊥-≃ ⋆ isoToEquiv Iso-⊤⊎⊤-Bool + +-- decidable predicate over SumFin + +_^_ : ℕ → ℕ → ℕ +m ^ 0 = 1 +m ^ (suc n) = m · m ^ n + +SumFinℙ≃ : (n : ℕ) → (Fin n → Bool) ≃ Fin (2 ^ n) +SumFinℙ≃ 0 = isContr→≃Unit (isContrΠ⊥) ⋆ invEquiv (⊎-⊥-≃) +SumFinℙ≃ (suc n) = + Π⊎≃ + ⋆ Σ-cong-equiv (UnitToType≃ Bool ⋆ invEquiv SumFin2≃Bool) (λ _ → SumFinℙ≃ n) + ⋆ SumFin×≃ 2 (2 ^ n) + +-- decidable subsets of SumFin + +Bool→ℕ : Bool → ℕ +Bool→ℕ true = 1 +Bool→ℕ false = 0 + +trueCount : {n : ℕ}(f : Fin n → Bool) → ℕ +trueCount {n = 0} _ = 0 +trueCount {n = suc n} f = Bool→ℕ (f (inl tt)) + (trueCount (f ∘ inr)) + +SumFinDec⊎≃ : (n : ℕ)(t : Bool) → (Bool→Type t ⊎ Fin n) ≃ (Fin (Bool→ℕ t + n)) +SumFinDec⊎≃ _ true = idEquiv _ +SumFinDec⊎≃ _ false = ⊎-swap-≃ ⋆ ⊎-⊥-≃ + +SumFinSub≃ : (n : ℕ)(f : Fin n → Bool) → Σ _ (Bool→Type ∘ f) ≃ Fin (trueCount f) +SumFinSub≃ 0 _ = ΣEmpty _ +SumFinSub≃ (suc n) f = + Σ⊎≃ + ⋆ ⊎-equiv (ΣUnit (Bool→Type ∘ f ∘ inl)) (SumFinSub≃ n (f ∘ inr)) + ⋆ SumFinDec⊎≃ _ (f (inl tt)) + +-- decidable quantifier + +trueForSome : (n : ℕ)(f : Fin n → Bool) → Bool +trueForSome 0 _ = false +trueForSome (suc n) f = f (inl tt) or trueForSome n (f ∘ inr) + +trueForAll : (n : ℕ)(f : Fin n → Bool) → Bool +trueForAll 0 _ = true +trueForAll (suc n) f = f (inl tt) and trueForAll n (f ∘ inr) + +SumFin∃→ : (n : ℕ)(f : Fin n → Bool) → Σ _ (Bool→Type ∘ f) → Bool→Type (trueForSome n f) +SumFin∃→ 0 _ = ΣEmpty _ .fst +SumFin∃→ (suc n) f = + Bool→Prop⊎' _ _ + ∘ map-⊎ (ΣUnit (Bool→Type ∘ f ∘ inl) .fst) (SumFin∃→ n (f ∘ inr)) + ∘ Σ⊎≃ .fst + +SumFin∃← : (n : ℕ)(f : Fin n → Bool) → Bool→Type (trueForSome n f) → Σ _ (Bool→Type ∘ f) +SumFin∃← 0 _ = ⊥.rec +SumFin∃← (suc n) f = + invEq Σ⊎≃ + ∘ map-⊎ (invEq (ΣUnit (Bool→Type ∘ f ∘ inl))) (SumFin∃← n (f ∘ inr)) + ∘ Bool→Prop⊎ _ _ + +SumFin∃≃ : (n : ℕ)(f : Fin n → Bool) → ∥ Σ (Fin n) (Bool→Type ∘ f) ∥ ≃ Bool→Type (trueForSome n f) +SumFin∃≃ n f = + propBiimpl→Equiv isPropPropTrunc isPropBool→Prop + (Prop.rec isPropBool→Prop (SumFin∃→ n f)) + (∣_∣ ∘ SumFin∃← n f) + +SumFin∀≃ : (n : ℕ)(f : Fin n → Bool) → ((x : Fin n) → Bool→Type (f x)) ≃ Bool→Type (trueForAll n f) +SumFin∀≃ 0 _ = isContr→≃Unit (isContrΠ⊥) +SumFin∀≃ (suc n) f = + Π⊎≃ + ⋆ Σ-cong-equiv (ΠUnit (Bool→Type ∘ f ∘ inl)) (λ _ → SumFin∀≃ n (f ∘ inr)) + ⋆ Bool→Prop≃ _ _ + +-- internal equality + +SumFin≡ : (n : ℕ) → (a b : Fin n) → Bool +SumFin≡ 0 _ _ = true +SumFin≡ (suc n) (inl tt) (inl tt) = true +SumFin≡ (suc n) (inl tt) (inr y) = false +SumFin≡ (suc n) (inr x) (inl tt) = false +SumFin≡ (suc n) (inr x) (inr y) = SumFin≡ n x y + +isSetSumFin : (n : ℕ)→ isSet (Fin n) +isSetSumFin 0 = isProp→isSet isProp⊥ +isSetSumFin (suc n) = isSet⊎ (isProp→isSet isPropUnit) (isSetSumFin n) + +SumFin≡≃ : (n : ℕ) → (a b : Fin n) → (a ≡ b) ≃ Bool→Type (SumFin≡ n a b) +SumFin≡≃ 0 _ _ = isContr→≃Unit (isProp→isContrPath isProp⊥ _ _) +SumFin≡≃ (suc n) (inl tt) (inl tt) = isContr→≃Unit (inhProp→isContr refl (isSetSumFin _ _ _)) +SumFin≡≃ (suc n) (inl tt) (inr y) = invEquiv (⊎Path.Cover≃Path _ _) ⋆ uninhabEquiv ⊥.rec* ⊥.rec +SumFin≡≃ (suc n) (inr x) (inl tt) = invEquiv (⊎Path.Cover≃Path _ _) ⋆ uninhabEquiv ⊥.rec* ⊥.rec +SumFin≡≃ (suc n) (inr x) (inr y) = invEquiv (_ , isEmbedding-inr x y) ⋆ SumFin≡≃ n x y + +-- propositional truncation of Fin + +∥Fin∥ : (n : ℕ) → Dec ∥ Fin n ∥ +∥Fin∥ 0 = no (Prop.rec isProp⊥ (idfun _)) +∥Fin∥ (suc n) = yes ∣ fzero ∣ + +-- some properties about cardinality + +fzero≠fone : {n : ℕ} → ¬ (fzero ≡ fsuc fzero) +fzero≠fone {n = n} = SumFin≡≃ (suc (suc n)) fzero (fsuc fzero) .fst + +Fin>0 : (n : ℕ) → 0 < n → Fin n +Fin>0 0 p = ⊥.rec (¬-<-zero p) +Fin>0 (suc n) p = fzero + +Fin>1 : (n : ℕ) → 1 < n → Σ[ i ∈ Fin n ] Σ[ j ∈ Fin n ] ¬ i ≡ j +Fin>1 0 p = ⊥.rec (snotz (≤0→≡0 p)) +Fin>1 1 p = ⊥.rec (snotz (≤0→≡0 (pred-≤-pred p))) +Fin>1 (suc (suc n)) _ = fzero , fsuc fzero , fzero≠fone + +emptyFin : (n : ℕ) → ¬ Fin n → 0 ≡ n +emptyFin 0 _ = refl +emptyFin (suc n) p = ⊥.rec (p fzero) + +nonEmptyFin : (n : ℕ) → Fin n → 0 < n +nonEmptyFin 0 i = ⊥.rec i +nonEmptyFin (suc n) _ = suc-≤-suc zero-≤ + +nonEqualTermFin : (n : ℕ) → (i j : Fin n) → ¬ i ≡ j → 1 < n +nonEqualTermFin 0 i _ _ = ⊥.rec i +nonEqualTermFin 1 i j p = ⊥.rec (p (isContr→isProp isContrSumFin1 i j)) +nonEqualTermFin (suc (suc n)) _ _ _ = suc-≤-suc (suc-≤-suc zero-≤) + +Fin≤1 : (n : ℕ) → n ≤ 1 → isProp (Fin n) +Fin≤1 0 _ = isProp⊥ +Fin≤1 1 _ = isContr→isProp isContrSumFin1 +Fin≤1 (suc (suc n)) p = ⊥.rec (¬-<-zero (pred-≤-pred p)) + +propFin : (n : ℕ) → isProp (Fin n) → n ≤ 1 +propFin 0 _ = ≤-solver 0 1 +propFin 1 _ = ≤-solver 1 1 +propFin (suc (suc n)) p = ⊥.rec (fzero≠fone (p fzero (fsuc fzero))) + +-- automorphisms of SumFin + +SumFin≃≃ : (n : ℕ) → (Fin n ≃ Fin n) ≃ Fin (LehmerCode.factorial n) +SumFin≃≃ _ = + equivComp (SumFin≃Fin _) (SumFin≃Fin _) + ⋆ LehmerCode.lehmerEquiv + ⋆ LehmerCode.lehmerFinEquiv + ⋆ invEquiv (SumFin≃Fin _) diff --git a/Cubical/Experiments/Combinatorics.agda b/Cubical/Experiments/Combinatorics.agda index 7f310459fd..0074bff918 100644 --- a/Cubical/Experiments/Combinatorics.agda +++ b/Cubical/Experiments/Combinatorics.agda @@ -12,15 +12,17 @@ open import Cubical.Foundations.Equiv renaming (_≃_ to _≃'_) open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order +open import Cubical.Data.Bool open import Cubical.Data.Sum open import Cubical.Data.Sigma renaming (_×_ to _×'_) open import Cubical.Data.Vec -open import Cubical.Data.Fin renaming (Fin to Fin') -open import Cubical.Data.SumFin renaming (Fin to SumFin) hiding (fzero) -open import Cubical.Data.FinSet +open import Cubical.Data.SumFin renaming (Fin to Fin') +open import Cubical.Data.FinSet.Base +open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Cardinality +open import Cubical.Data.FinSet.Decidability open import Cubical.Data.FinSet.Quotients open import Cubical.HITs.PropositionalTruncation renaming (∥_∥ to ∥_∥') @@ -115,9 +117,9 @@ a3,2 = refl -- construct numerical functions from list getFun : {n : ℕ} → Vec ℕ n → Fin n .fst → ℕ -getFun {n = n} ns x = fun n ns (Fin→SumFin x) +getFun {n = n} ns = fun n ns where - fun : (n : ℕ) → Vec ℕ n → SumFin n → ℕ + fun : (n : ℕ) → Vec ℕ n → Fin' n → ℕ fun 0 _ _ = 0 fun (suc m) (n ∷ ns) (inl tt) = n fun (suc m) (n ∷ ns) (inr x) = fun m ns x @@ -138,15 +140,15 @@ m : maxValue (Fin _) f ∣ fzero ∣ ≡ 9 m = refl -- the number of numeral 1 -n1 : card (_ , isFinSetFiberDec (Fin _) ℕ discreteℕ f 1) ≡ 2 +n1 : card (_ , isFinSetFiberDisc (Fin _) ℕ discreteℕ f 1) ≡ 2 n1 = refl -- a somewhat trivial equivalence relation making everything equivalent R : {n : ℕ} → Fin n .fst → Fin n .fst → Type R _ _ = Unit -isDecR : {n : ℕ} → (x y : Fin n .fst) → Dec (R {n = n} x y) -isDecR _ _ = yes tt +isDecR : {n : ℕ} → (x y : Fin n .fst) → isDecProp (R {n = n} x y) +isDecR _ _ = true , idEquiv _ open BinaryRelation open isEquivRel diff --git a/Cubical/Experiments/CountingFiniteStructure.agda b/Cubical/Experiments/CountingFiniteStructure.agda index 4e650d04ab..80abfcaeb9 100644 --- a/Cubical/Experiments/CountingFiniteStructure.agda +++ b/Cubical/Experiments/CountingFiniteStructure.agda @@ -1,6 +1,6 @@ {- -This file contains the counting number of finite sets with structure. +Counting how many structured finite sets with a given cardinal https://github.com/EgbertRijke/OEIS-A000001 @@ -10,11 +10,15 @@ https://github.com/EgbertRijke/OEIS-A000001 module Cubical.Experiments.CountingFiniteStructure where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv open import Cubical.Data.Nat +open import Cubical.Data.Bool open import Cubical.Data.Sigma -open import Cubical.Data.FinSet +open import Cubical.Data.FinSet.Base +open import Cubical.Data.FinSet.Properties +open import Cubical.Data.FinSet.Decidability open import Cubical.Data.FinSet.Induction open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Cardinality @@ -30,12 +34,12 @@ private isFinStrCard : (S : FinSet ℓ-zero → FinSet ℓ) (n : ℕ) → isFinType 0 (FinSetWithStrOfCard ℓ-zero S n) isFinStrCard S n = isFinTypeFinSetWithStrOfCard ℓ-zero S n --- structure that is no structure +-- structure that there is no structure TrivialStr : FinSet ℓ → FinSet ℓ TrivialStr _ = 𝟙 --- identity structure +-- structure that "picking an element in that set" IdentityStr : FinSet ℓ → FinSet ℓ IdentityStr X = X @@ -80,13 +84,122 @@ a2 = card (_ , isFinStrCard TrivialStr 2) b2 : ℕ b2 = card (_ , isFinStrCard IdentityStr 2) +-- the number of finite semi-groups +numberOfFinSemiGroups : ℕ → ℕ +numberOfFinSemiGroups n = card (_ , isFinStrCard FinSemiGroupStr n) + +-- two trivial cases of semi-groups +-- in a flash +n0 : ℕ +n0 = numberOfFinSemiGroups 0 + +-- nearly one minute +n1 : ℕ +n1 = numberOfFinSemiGroups 1 + -- the number of finite semi-groups with cardinal 2 -- it should be 5 -- would you like to try? n2 : ℕ -n2 = card (_ , isFinStrCard FinSemiGroupStr 2) +n2 = numberOfFinSemiGroups 2 -- OEIS-A000001 --- I think you'd better not evaluate this function +-- I think you'd better not evaluate this function with n > 1 numberOfFinGroups : ℕ → ℕ numberOfFinGroups n = card (_ , isFinStrCard FinGroupStr n) + +-- group with one element +-- it takes about 21 minutes +g1 : ℕ +g1 = numberOfFinGroups 1 + +-- Rijke's challenge +-- seemed to big to do an exhaustive search +g4 : ℕ +g4 = numberOfFinGroups 4 + +--------------- test ----------------- + +open import Cubical.Data.SumFin + +k : ℕ +k = card (FinSemiGroupStr (Fin 3 , isFinSetFin)) + +EndoStr : FinSet ℓ → FinSet ℓ +EndoStr X = (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) + +EndoStr' : FinSet ℓ → FinSet ℓ +EndoStr' {ℓ = ℓ} X = (_ , isFinSetΣ (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ _ → 𝟙 {ℓ})) + +EndoStr'' : FinSet ℓ → FinSet ℓ +EndoStr'' {ℓ = ℓ} X = (_ , isFinSetΣ (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ _ → (Fin 1 , isFinSetFin))) + +open import Cubical.Data.Unit +open import Cubical.Relation.Nullary + +{- +isDecPropUnit : isDecProp Unit +isDecPropUnit .fst = isPropUnit +isDecPropUnit .snd = yes tt +-} +isDecPropUnit : isDecProp Unit +isDecPropUnit .fst = true +isDecPropUnit .snd = idEquiv _ + +EndoStr''' : FinSet ℓ → FinSet ℓ +EndoStr''' {ℓ = ℓ} X = (_ , isFinSetSub (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ _ → Unit , isDecPropUnit)) + +l : ℕ +l = card (EndoStr (Fin 3 , isFinSetFin)) + +l' : ℕ +l' = card (EndoStr' (Fin 3 , isFinSetFin)) + +l'' : ℕ +l'' = card (EndoStr'' (Fin 2 , isFinSetFin)) + +l''' : ℕ +l''' = card (EndoStr''' (Fin 3 , isFinSetFin)) + +FinSemiGroupStr' : FinSet ℓ → FinSet ℓ +FinSemiGroupStr' X .fst = + Σ[ p ∈ (X .fst → X .fst → X .fst) ] ((x y z : X .fst) → p (p x y) z ≡ p x (p y z)) +FinSemiGroupStr' X .snd = + isFinSetSub (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) + (λ p → _ , isDecProp∀3 X (λ _ → X) (λ _ _ → X) (λ _ _ _ → _ , isDecProp≡ X _ _)) + +k' : ℕ +k' = card (FinSemiGroupStr' (Fin 2 , isFinSetFin)) + +k'' : ℕ +k'' = card (FinSemiGroupStr (Fin 2 , isFinSetFin)) + +FinGroupStr' : FinSet ℓ → FinSet ℓ +FinGroupStr' X .fst = + Σ[ e ∈ X .fst ] + Σ[ inv ∈ (X .fst → X .fst) ] + Σ[ p ∈ (X .fst → X .fst → X .fst) ] + ((x y z : X .fst) → p (p x y) z ≡ p x (p y z)) + × ((x : X .fst) + → (p x e ≡ x) × (p e x ≡ x) × (p (inv x) x ≡ e) × (p x (inv x) ≡ e)) +FinGroupStr' X .snd = + isFinSetΣ X (λ _ → _ , + isFinSetΣ (_ , isFinSetΠ X (λ _ → X)) (λ _ → _ , + isFinSetSub (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ _ → _ , + isDecProp× (_ , isDecProp∀3 X (λ _ → X) (λ _ _ → X) (λ _ _ _ → _ , isDecProp≡ X _ _)) (_ , + isDecProp∀ X (λ _ → _ , + isDecProp× (_ , isDecProp≡ X _ _) (_ , + isDecProp× (_ , isDecProp≡ X _ _) (_ , + isDecProp× (_ , isDecProp≡ X _ _) (_ , isDecProp≡ X _ _)))))))) + +r : ℕ +r = card (FinGroupStr' (Fin 2 , isFinSetFin)) + +r' : ℕ +r' = card (FinGroupStr (Fin 2 , isFinSetFin)) + +ff : ℕ → ℕ +ff n = card (_ , isFinStrCard FinSemiGroupStr n) + +ff' : ℕ → ℕ +ff' n = card (_ , isFinStrCard FinSemiGroupStr' n) diff --git a/Cubical/HITs/SetQuotients/EqClass.agda b/Cubical/HITs/SetQuotients/EqClass.agda index 187603201e..618d2f86ef 100644 --- a/Cubical/HITs/SetQuotients/EqClass.agda +++ b/Cubical/HITs/SetQuotients/EqClass.agda @@ -19,7 +19,7 @@ open import Cubical.Functions.Surjection private variable - ℓ ℓ' : Level + ℓ ℓ' ℓ'' : Level -- another definition using equivalence classes @@ -45,15 +45,15 @@ module _ isSetℙDec = isOfHLevelΣ 2 isSetℙ (λ P → isSetΠ (λ x → isProp→isSet (isPropDec (P x .snd)))) module _ - (R : X → X → Type ℓ') where + (R : X → X → Type ℓ'') where - isEqClass : ℙ → Type (ℓ-max ℓ (ℓ-suc ℓ')) - isEqClass P = ∥ Σ[ x ∈ X ] ((a : X) → P a .fst ≡ ∥ R a x ∥) ∥ + isEqClass : ℙ → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isEqClass P = ∥ Σ[ x ∈ X ] ((a : X) → P a .fst ≃ ∥ R a x ∥) ∥ isPropIsEqClass : (P : ℙ) → isProp (isEqClass P) isPropIsEqClass P = isPropPropTrunc - _∥_ : Type (ℓ-max ℓ (ℓ-suc ℓ')) + _∥_ : Type (ℓ-max (ℓ-max ℓ (ℓ-suc ℓ')) ℓ'') _∥_ = Σ[ P ∈ ℙ ] isEqClass P isSet∥ : isSet _∥_ @@ -62,12 +62,12 @@ module _ module _ (dec : (x x' : X) → Dec (R x x')) where - _∥Dec_ : Type (ℓ-max ℓ (ℓ-suc ℓ')) + _∥Dec_ : Type (ℓ-max (ℓ-max ℓ (ℓ-suc ℓ')) ℓ'') _∥Dec_ = Σ[ P ∈ ℙDec ] isEqClass (P .fst) isDecEqClass : (P : _∥_) → (x : X) → Dec (P .fst x .fst) isDecEqClass (P , h) a = - Prop.rec (isPropDec (P a .snd)) (λ (x , p) → subst Dec (sym (p a)) (Dec∥∥ (dec a x))) h + Prop.rec (isPropDec (P a .snd)) (λ (x , p) → EquivPresDec (invEquiv (p a)) (Dec∥∥ (dec a x))) h Iso-∥Dec-∥ : Iso _∥Dec_ _∥_ Iso-∥Dec-∥ .fun P = P .fst .fst , P .snd @@ -84,8 +84,9 @@ module _ ∥Dec≃∥ = isoToEquiv Iso-∥Dec-∥ module _ + {ℓ ℓ' ℓ'' : Level} (X : Type ℓ) - (R : X → X → Type ℓ') + (R : X → X → Type ℓ'') (h : isEquivRel R) where ∥Rx∥Iso : (x x' : X)(r : R x x') → (a : X) → Iso ∥ R a x ∥ ∥ R a x' ∥ @@ -95,7 +96,7 @@ module _ ∥Rx∥Iso x x' r a .rightInv _ = isPropPropTrunc _ _ isEqClass∥Rx∥ : (x : X) → isEqClass X R (λ a → ∥ R a x ∥ , isPropPropTrunc) - isEqClass∥Rx∥ x = ∣ x , (λ _ → refl) ∣ + isEqClass∥Rx∥ x = ∣ x , (λ _ → idEquiv _) ∣ ∥R∥ : (x : X) → X ∥ R ∥R∥ x = (λ a → ∥ R a x ∥ , isPropPropTrunc) , isEqClass∥Rx∥ x @@ -125,8 +126,8 @@ module _ isEmbedding/→∥ : isEmbedding /→∥ isEmbedding/→∥ = injEmbedding squash/ (isSet∥ X R) (λ {x} {y} → inj/→∥ x y) - surj/→∥ : (P : X ∥ R) → ((x , _) : Σ[ x ∈ X ] ((a : X) → P .fst a .fst ≡ ∥ R a x ∥)) → ∥R∥ x ≡ P - surj/→∥ P (x , p) i .fst a .fst = p a (~ i) + surj/→∥ : (P : X ∥ R) → ((x , _) : Σ[ x ∈ X ] ((a : X) → P .fst a .fst ≃ ∥ R a x ∥)) → ∥R∥ x ≡ P + surj/→∥ P (x , p) i .fst a .fst = ua (p a) (~ i) surj/→∥ P (x , p) i .fst a .snd = isProp→PathP {B = λ i → isProp (surj/→∥ P (x , p) i .fst a .fst)} (λ i → isPropIsProp) isPropPropTrunc (P .fst a .snd) i From e6d398ad94175fe733fc94b93cfcad3e05156ab7 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Mon, 29 Nov 2021 03:24:25 +0800 Subject: [PATCH 07/24] resolve contradiction --- Cubical/Data/FinInd.agda | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Cubical/Data/FinInd.agda b/Cubical/Data/FinInd.agda index 20286fff7f..cebab10420 100644 --- a/Cubical/Data/FinInd.agda +++ b/Cubical/Data/FinInd.agda @@ -16,7 +16,7 @@ This definition is weaker than `isFinSet`. module Cubical.Data.FinInd where open import Cubical.Data.Nat -open import Cubical.Data.Fin +open import Cubical.Data.SumFin open import Cubical.Data.Sigma open import Cubical.Data.FinSet open import Cubical.Foundations.Prelude @@ -34,10 +34,10 @@ isFinInd : Type ℓ → Type ℓ isFinInd A = ∃[ n ∈ ℕ ] Fin n ↠ A isFinSet→isFinInd : isFinSet A → isFinInd A -isFinSet→isFinInd = PT.rec - squash - λ (n , equiv) → - ∣ n , invEq equiv , section→isSurjection (retEq equiv) ∣ +isFinSet→isFinInd h = PT.elim + (λ _ → squash) + (λ equiv → + ∣ _ , invEq equiv , section→isSurjection (retEq equiv) ∣) (h .snd) isFinInd-S¹ : isFinInd S¹ isFinInd-S¹ = ∣ 1 , f , isSurjection-f ∣ From 9992c70a15a809736dd849dadcb74dbcdf02027c Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Mon, 29 Nov 2021 03:36:19 +0800 Subject: [PATCH 08/24] resolving --- Cubical/Data/SubFinSet.agda | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Cubical/Data/SubFinSet.agda b/Cubical/Data/SubFinSet.agda index eb975216b1..443ec4aaf8 100644 --- a/Cubical/Data/SubFinSet.agda +++ b/Cubical/Data/SubFinSet.agda @@ -15,10 +15,11 @@ Every subfinite set is guaranteed to be a set and discrete. module Cubical.Data.SubFinSet where open import Cubical.Data.Nat -open import Cubical.Data.Fin +open import Cubical.Data.SumFin open import Cubical.Data.Sigma open import Cubical.Data.FinSet open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels open import Cubical.Functions.Embedding open import Cubical.HITs.PropositionalTruncation as PT @@ -34,8 +35,8 @@ isSubFinSet A = ∃[ n ∈ ℕ ] A ↪ Fin n isFinSet→isSubFinSet : isFinSet A → isSubFinSet A isFinSet→isSubFinSet = PT.map - λ (n , f , isEquiv-f) → - n , f , isEquiv→isEmbedding isEquiv-f + (λ (n , f , isEquiv-f) → + n , f , isEquiv→isEmbedding isEquiv-f) ∘ isFinSet→isFinSet' isSubFinSet→isSet : isSubFinSet A → isSet A isSubFinSet→isSet = PT.rec From 43632fa5107d3e153fd55287820ede8f5244c731 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Mon, 29 Nov 2021 12:35:00 +0800 Subject: [PATCH 09/24] clean --- Cubical/Data/Fin/Properties.agda | 40 ------------ Cubical/Data/FinSet/Base.agda | 32 ++++++---- Cubical/Data/FinSet/Cardinality.agda | 4 +- Cubical/Data/FinSet/Decidability.agda | 17 +++-- Cubical/Data/FinSet/Properties.agda | 89 --------------------------- 5 files changed, 31 insertions(+), 151 deletions(-) diff --git a/Cubical/Data/Fin/Properties.agda b/Cubical/Data/Fin/Properties.agda index 59aee3c80e..dcd188bf1e 100644 --- a/Cubical/Data/Fin/Properties.agda +++ b/Cubical/Data/Fin/Properties.agda @@ -636,43 +636,3 @@ FinData≃Fin N = isoToEquiv (FinDataIsoFin N) FinData≡Fin : (N : ℕ) → FinData N ≡ Fin N FinData≡Fin N = ua (FinData≃Fin N) - --- propositional truncation of Fin - -∥Fin∥ : (n : ℕ) → Dec ∥ Fin n ∥ -∥Fin∥ 0 = no (∥∥rec isProp⊥ ¬Fin0) -∥Fin∥ (suc n) = yes ∣ fzero ∣ - --- some properties about cardinality - -Fin>0 : (n : ℕ) → 0 < n → Fin n -Fin>0 0 p = Empty.rec (¬-<-zero p) -Fin>0 (suc n) p = fzero - -Fin>1 : (n : ℕ) → 1 < n → Σ[ i ∈ Fin n ] Σ[ j ∈ Fin n ] ¬ i ≡ j -Fin>1 0 p = Empty.rec (snotz (≤0→≡0 p)) -Fin>1 1 p = Empty.rec (snotz (≤0→≡0 (pred-≤-pred p))) -Fin>1 (suc (suc n)) _ = fzero , fone , fzero≠fone - -emptyFin : (n : ℕ) → ¬ Fin n → 0 ≡ n -emptyFin 0 _ = refl -emptyFin (suc n) p = Empty.rec (p fzero) - -nonEmptyFin : (n : ℕ) → Fin n → 0 < n -nonEmptyFin 0 i = Empty.rec (¬Fin0 i) -nonEmptyFin (suc n) _ = suc-≤-suc zero-≤ - -nonEqualTermFin : (n : ℕ) → (i j : Fin n) → ¬ i ≡ j → 1 < n -nonEqualTermFin 0 i _ _ = Empty.rec (¬Fin0 i) -nonEqualTermFin 1 i j p = Empty.rec (p (isContr→isProp isContrFin1 i j)) -nonEqualTermFin (suc (suc n)) _ _ _ = suc-≤-suc (suc-≤-suc zero-≤) - -Fin≤1 : (n : ℕ) → n ≤ 1 → isProp (Fin n) -Fin≤1 0 _ = isPropFin0 -Fin≤1 1 _ = isContr→isProp isContrFin1 -Fin≤1 (suc (suc n)) p = Empty.rec (¬-<-zero (pred-≤-pred p)) - -propFin : (n : ℕ) → isProp (Fin n) → n ≤ 1 -propFin 0 _ = ≤-solver 0 1 -propFin 1 _ = ≤-solver 1 1 -propFin (suc (suc n)) p = Empty.rec (fzero≠fone (p fzero fone)) diff --git a/Cubical/Data/FinSet/Base.agda b/Cubical/Data/FinSet/Base.agda index 2e9c311fe8..43f2ae0e78 100644 --- a/Cubical/Data/FinSet/Base.agda +++ b/Cubical/Data/FinSet/Base.agda @@ -29,6 +29,7 @@ private A : Type ℓ -- operators to more conveniently compose equivalences + module _ {A : Type ℓ}{B : Type ℓ'}{C : Type ℓ''} where @@ -48,33 +49,40 @@ isFinOrd A = Σ[ n ∈ ℕ ] A ≃ Fin n isFinOrd→isFinSet : isFinOrd A → isFinSet A isFinOrd→isFinSet (_ , p) = _ , ∣ p ∣ --- another definition of isFinSet -isFinSet' : Type ℓ → Type ℓ -isFinSet' A = ∥ Σ[ n ∈ ℕ ] A ≃ Fin n ∥ - -isFinSet→isFinSet' : isFinSet A → isFinSet' A -isFinSet→isFinSet' (_ , p) = Prop.rec isPropPropTrunc (λ p → ∣ _ , p ∣) p - -- finite sets are sets + isFinSet→isSet : isFinSet A → isSet A isFinSet→isSet p = rec isPropIsSet (λ e → isOfHLevelRespectEquiv 2 (invEquiv e) isSetFin) (p .snd) -- isFinSet is proposition -private - cardUniq : (m n : ℕ) → (p : A ≃ Fin m)(q : A ≃ Fin n) → m ≡ n - cardUniq _ _ p q = Fin-inj _ _ (ua (invEquiv (SumFin≃Fin _) ⋆ (invEquiv p) ⋆ q ⋆ SumFin≃Fin _)) isPropIsFinSet : isProp (isFinSet A) isPropIsFinSet p q = Σ≡PropEquiv (λ _ → isPropPropTrunc) .fst ( Prop.elim2 (λ _ _ → isSetℕ _ _) - (λ p q → cardUniq _ _ p q) + (λ p q → Fin-inj _ _ (ua (invEquiv (SumFin≃Fin _) ⋆ (invEquiv p) ⋆ q ⋆ SumFin≃Fin _))) (p .snd) (q .snd)) --- isFinSet is Set +-- isFinOrd is Set +-- ordering can be seen as extra structures over finite sets + isSetIsFinOrd : isSet (isFinOrd A) isSetIsFinOrd = isOfHLevelΣ 2 isSetℕ (λ _ → isOfHLevel⁺≃ᵣ 1 isSetFin) +-- alternative definition of isFinSet + +isFinSet' : Type ℓ → Type ℓ +isFinSet' A = ∥ Σ[ n ∈ ℕ ] A ≃ Fin n ∥ + +isFinSet→isFinSet' : isFinSet A → isFinSet' A +isFinSet→isFinSet' (_ , p) = Prop.rec isPropPropTrunc (λ p → ∣ _ , p ∣) p + +isFinSet'→isFinSet : isFinSet' A → isFinSet A +isFinSet'→isFinSet = Prop.rec isPropIsFinSet (λ (n , p) → _ , ∣ p ∣ ) + +isFinSet≡isFinSet' : isFinSet A ≡ isFinSet' A +isFinSet≡isFinSet' = hPropExt isPropIsFinSet isPropPropTrunc isFinSet→isFinSet' isFinSet'→isFinSet + -- the type of finite sets/propositions FinSet : (ℓ : Level) → Type (ℓ-suc ℓ) diff --git a/Cubical/Data/FinSet/Cardinality.agda b/Cubical/Data/FinSet/Cardinality.agda index 64849721a0..2d61dfc671 100644 --- a/Cubical/Data/FinSet/Cardinality.agda +++ b/Cubical/Data/FinSet/Cardinality.agda @@ -554,6 +554,8 @@ module _ maxValue : ℕ maxValue = ∃Max→maxValue _ _ ∃MaxFinSet +{- some formal consequences of card -} + -- card induces equivalence from set truncation of FinSet to natural numbers open Iso @@ -604,7 +606,7 @@ FinProp≃Bool = isFinSetFinProp : isFinSet (FinProp ℓ) isFinSetFinProp = EquivPresIsFinSet (invEquiv FinProp≃Bool) isFinSetBool --- a more efficient version of equivalence type +-- a more computationally efficient version of equivalence type module _ (X : FinSet ℓ ) diff --git a/Cubical/Data/FinSet/Decidability.agda b/Cubical/Data/FinSet/Decidability.agda index f8326a81ce..e720cfd6d0 100644 --- a/Cubical/Data/FinSet/Decidability.agda +++ b/Cubical/Data/FinSet/Decidability.agda @@ -34,12 +34,6 @@ isDecProp→isProp h = isOfHLevelRespectEquiv 1 (invEquiv (h .snd)) isPropBool isDecProp→Dec : {P : Type ℓ} → isDecProp P → Dec P isDecProp→Dec h = EquivPresDec (invEquiv (h .snd)) DecBool→Prop -helper : (P : Type ℓ) → (t : Bool) → isProp (P ≃ Bool→Type t) -helper _ _ = isOfHLevel⁺≃ᵣ 0 isPropBool→Prop - -helper' : (P : Type ℓ) → (p q : isDecProp P) → (p .fst ≡ q .fst) ≃ (p ≡ q) -helper' _ _ _ = Σ≡PropEquiv (helper _) - isPropIsDecProp : {P : Type ℓ} → isProp (isDecProp P) isPropIsDecProp p q = Σ≡PropEquiv (λ _ → isOfHLevel⁺≃ᵣ 0 isPropBool→Prop) .fst @@ -51,6 +45,9 @@ DecProp ℓ = Σ[ P ∈ Type ℓ ] isDecProp P module _ (X : Type ℓ)(p : isFinOrd X) where + isDecProp¬' : isDecProp (¬ X) + isDecProp¬' = _ , invEquiv (preCompEquiv (p .snd)) ⋆ SumFin¬ _ + isDecProp∥∥' : isDecProp ∥ X ∥ isDecProp∥∥' = _ , propTrunc≃ (p .snd) ⋆ SumFin∥∥DecProp _ @@ -153,10 +150,12 @@ module _ module _ (X : FinSet ℓ) where -{- + isDecProp¬ : isDecProp (¬ (X .fst)) - isDecProp¬ .fst = {!!} - isDecProp¬ .snd = {!!} -} + isDecProp¬ = + Prop.rec isPropIsDecProp + (λ p → isDecProp¬' (X .fst) (_ , p)) + (X .snd .snd) isDecProp∥∥ : isDecProp ∥ X .fst ∥ isDecProp∥∥ = diff --git a/Cubical/Data/FinSet/Properties.agda b/Cubical/Data/FinSet/Properties.agda index 8c42b7f92b..3024f0e6fd 100644 --- a/Cubical/Data/FinSet/Properties.agda +++ b/Cubical/Data/FinSet/Properties.agda @@ -94,92 +94,3 @@ transpFamily : {A : Type ℓ}{B : A → Type ℓ'} → ((n , e) : isFinOrd A) → (x : A) → B x ≃ B (invEq e (e .fst x)) transpFamily {B = B} (n , e) x = pathToEquiv (λ i → B (retEq e x (~ i))) - -{- -{- - -Alternative definition of finite sets - -A set is finite if it is merely equivalent to `Fin n` for some `n`. We -can translate this to code in two ways: a truncated sigma of a nat and -an equivalence, or a sigma of a nat and a truncated equivalence. We -prove that both formulations are equivalent. - --} - -isFinSet' : Type ℓ → Type ℓ -isFinSet' A = Σ[ n ∈ ℕ ] ∥ A ≃ Fin n ∥ - -FinSet' : (ℓ : Level) → Type (ℓ-suc ℓ) -FinSet' ℓ = TypeWithStr _ isFinSet' - -isPropIsFinSet' : isProp (isFinSet' A) -isPropIsFinSet' {A = A} (n , equivn) (m , equivm) = - Σ≡Prop (λ _ → isPropPropTrunc) n≡m - where - Fin-n≃Fin-m : ∥ Fin n ≃ Fin m ∥ - Fin-n≃Fin-m = Prop.rec - isPropPropTrunc - (Prop.rec - (isPropΠ λ _ → isPropPropTrunc) - (λ hm hn → ∣ Fin n ≃⟨ invEquiv hn ⟩ A ≃⟨ hm ⟩ Fin m ■ ∣) - equivm - ) - equivn - - Fin-n≡Fin-m : ∥ Fin n ≡ Fin m ∥ - Fin-n≡Fin-m = Prop.rec isPropPropTrunc (∣_∣ ∘ ua) Fin-n≃Fin-m - - ∥n≡m∥ : ∥ n ≡ m ∥ - ∥n≡m∥ = Prop.rec isPropPropTrunc (∣_∣ ∘ Fin-inj n m) Fin-n≡Fin-m - - n≡m : n ≡ m - n≡m = Prop.rec (isSetℕ n m) (λ p → p) ∥n≡m∥ - --- logical equivalence of two definitions - -isFinSet→isFinSet' : isFinSet A → isFinSet' A -isFinSet→isFinSet' ∣ n , equiv ∣ = n , ∣ equiv ∣ -isFinSet→isFinSet' (squash p q i) = isPropIsFinSet' (isFinSet→isFinSet' p) (isFinSet→isFinSet' q) i - -isFinSet'→isFinSet : isFinSet' A → isFinSet A -isFinSet'→isFinSet (n , ∣ isFinSet-A ∣) = ∣ n , isFinSet-A ∣ -isFinSet'→isFinSet (n , squash p q i) = isPropIsFinSet (isFinSet'→isFinSet (n , p)) (isFinSet'→isFinSet (n , q)) i - -isFinSet≡isFinSet' : isFinSet A ≡ isFinSet' A -isFinSet≡isFinSet' {A = A} = hPropExt isPropIsFinSet isPropIsFinSet' isFinSet→isFinSet' isFinSet'→isFinSet - -FinSet→FinSet' : FinSet ℓ → FinSet' ℓ -FinSet→FinSet' (A , isFinSetA) = A , isFinSet→isFinSet' isFinSetA - -FinSet'→FinSet : FinSet' ℓ → FinSet ℓ -FinSet'→FinSet (A , isFinSet'A) = A , isFinSet'→isFinSet isFinSet'A - -FinSet≃FinSet' : FinSet ℓ ≃ FinSet' ℓ -FinSet≃FinSet' = - isoToEquiv - (iso FinSet→FinSet' FinSet'→FinSet - (λ _ → Σ≡Prop (λ _ → isPropIsFinSet') refl) - (λ _ → Σ≡Prop (λ _ → isPropIsFinSet) refl)) - -FinSet≡FinSet' : FinSet ℓ ≡ FinSet' ℓ -FinSet≡FinSet' = ua FinSet≃FinSet' - --- definitions to reduce problems about FinSet to SumFin - -≃Fin : Type ℓ → Type ℓ -≃Fin A = Σ[ n ∈ ℕ ] A ≃ Fin n - -≃SumFin : Type ℓ → Type ℓ -≃SumFin A = Σ[ n ∈ ℕ ] A ≃ SumFin n - -≃Fin→SumFin : ≃Fin A → ≃SumFin A -≃Fin→SumFin (n , e) = n , compEquiv e (invEquiv (SumFin≃Fin _)) - -≃SumFin→Fin : ≃SumFin A → ≃Fin A -≃SumFin→Fin (n , e) = n , compEquiv e (SumFin≃Fin _) - -transpFamily : - {A : Type ℓ}{B : A → Type ℓ'} - → ((n , e) : ≃SumFin A) → (x : A) → B x ≃ B (invEq e (e .fst x)) -transpFamily {B = B} (n , e) x = pathToEquiv (λ i → B (retEq e x (~ i))) -} From 5d658a66ee7a3c7559bbb8e88149d6be412f6e9e Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Mon, 29 Nov 2021 13:36:44 +0800 Subject: [PATCH 10/24] clean --- Cubical/Data/FinSet/Cardinality.agda | 2 +- .../{Constructors.agda => Constructor.agda} | 2 +- ...idability.agda => DecidablePredicate.agda} | 7 +- Cubical/Data/FinSet/FiniteStructure.agda | 4 +- .../FinSet/{FinType.agda => FiniteType.agda} | 8 +- Cubical/Data/FinSet/Induction.agda | 2 +- .../FinSet/{Quotients.agda => Quotient.agda} | 6 +- Cubical/Data/SumFin/Properties.agda | 10 ++ Cubical/Experiments/Combinatorics.agda | 91 ++++--------------- .../Experiments/CountingFiniteStructure.agda | 6 +- 10 files changed, 47 insertions(+), 91 deletions(-) rename Cubical/Data/FinSet/{Constructors.agda => Constructor.agda} (99%) rename Cubical/Data/FinSet/{Decidability.agda => DecidablePredicate.agda} (97%) rename Cubical/Data/FinSet/{FinType.agda => FiniteType.agda} (96%) rename Cubical/Data/FinSet/{Quotients.agda => Quotient.agda} (97%) diff --git a/Cubical/Data/FinSet/Cardinality.agda b/Cubical/Data/FinSet/Cardinality.agda index 2d61dfc671..f368568ca7 100644 --- a/Cubical/Data/FinSet/Cardinality.agda +++ b/Cubical/Data/FinSet/Cardinality.agda @@ -32,7 +32,7 @@ open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.FiniteChoice -open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.Constructor open import Cubical.Data.FinSet.Induction hiding (_+_) open import Cubical.Relation.Nullary diff --git a/Cubical/Data/FinSet/Constructors.agda b/Cubical/Data/FinSet/Constructor.agda similarity index 99% rename from Cubical/Data/FinSet/Constructors.agda rename to Cubical/Data/FinSet/Constructor.agda index da926cd1c3..d049250ac9 100644 --- a/Cubical/Data/FinSet/Constructors.agda +++ b/Cubical/Data/FinSet/Constructor.agda @@ -5,7 +5,7 @@ This files contains lots of useful properties about constructions on finite sets -} {-# OPTIONS --safe #-} -module Cubical.Data.FinSet.Constructors where +module Cubical.Data.FinSet.Constructor where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels diff --git a/Cubical/Data/FinSet/Decidability.agda b/Cubical/Data/FinSet/DecidablePredicate.agda similarity index 97% rename from Cubical/Data/FinSet/Decidability.agda rename to Cubical/Data/FinSet/DecidablePredicate.agda index e720cfd6d0..57656c7806 100644 --- a/Cubical/Data/FinSet/Decidability.agda +++ b/Cubical/Data/FinSet/DecidablePredicate.agda @@ -1,6 +1,11 @@ +{- + +This files contains lots of useful properties about decidable predicates on finite sets + +-} {-# OPTIONS --safe #-} -module Cubical.Data.FinSet.Decidability where +module Cubical.Data.FinSet.DecidablePredicate where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function diff --git a/Cubical/Data/FinSet/FiniteStructure.agda b/Cubical/Data/FinSet/FiniteStructure.agda index e14f47deed..fce9dd664b 100644 --- a/Cubical/Data/FinSet/FiniteStructure.agda +++ b/Cubical/Data/FinSet/FiniteStructure.agda @@ -21,9 +21,9 @@ open import Cubical.Data.Sigma open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.Induction -open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.Constructor open import Cubical.Data.FinSet.Cardinality -open import Cubical.Data.FinSet.FinType +open import Cubical.Data.FinSet.FiniteType private variable diff --git a/Cubical/Data/FinSet/FinType.agda b/Cubical/Data/FinSet/FiniteType.agda similarity index 96% rename from Cubical/Data/FinSet/FinType.agda rename to Cubical/Data/FinSet/FiniteType.agda index 81fb59b69f..eba37ee21d 100644 --- a/Cubical/Data/FinSet/FinType.agda +++ b/Cubical/Data/FinSet/FiniteType.agda @@ -11,7 +11,7 @@ https://github.com/EgbertRijke/OEIS-A000001 -} {-# OPTIONS --safe #-} -module Cubical.Data.FinSet.FinType where +module Cubical.Data.FinSet.FiniteType where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels @@ -31,9 +31,9 @@ open import Cubical.Data.Sigma open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties -open import Cubical.Data.FinSet.Decidability -open import Cubical.Data.FinSet.Constructors -open import Cubical.Data.FinSet.Quotients +open import Cubical.Data.FinSet.DecidablePredicate +open import Cubical.Data.FinSet.Constructor +open import Cubical.Data.FinSet.Quotient open import Cubical.Data.FinSet.Cardinality open import Cubical.Relation.Nullary diff --git a/Cubical/Data/FinSet/Induction.agda b/Cubical/Data/FinSet/Induction.agda index aae83bdcb0..9826da7761 100644 --- a/Cubical/Data/FinSet/Induction.agda +++ b/Cubical/Data/FinSet/Induction.agda @@ -25,7 +25,7 @@ open import Cubical.Data.Fin renaming (Fin to Finℕ) open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties -open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.Constructor private variable diff --git a/Cubical/Data/FinSet/Quotients.agda b/Cubical/Data/FinSet/Quotient.agda similarity index 97% rename from Cubical/Data/FinSet/Quotients.agda rename to Cubical/Data/FinSet/Quotient.agda index 6d3db0055f..c9dc2d95f4 100644 --- a/Cubical/Data/FinSet/Quotients.agda +++ b/Cubical/Data/FinSet/Quotient.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Data.FinSet.Quotients where +module Cubical.Data.FinSet.Quotient where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels @@ -19,8 +19,8 @@ open import Cubical.Data.Sigma open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties -open import Cubical.Data.FinSet.Decidability -open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.DecidablePredicate +open import Cubical.Data.FinSet.Constructor open import Cubical.Data.FinSet.Cardinality open import Cubical.Relation.Nullary diff --git a/Cubical/Data/SumFin/Properties.agda b/Cubical/Data/SumFin/Properties.agda index 7bf18cee99..dcc5784fc9 100644 --- a/Cubical/Data/SumFin/Properties.agda +++ b/Cubical/Data/SumFin/Properties.agda @@ -109,6 +109,16 @@ SumFin∥∥DecProp : (n : ℕ) → ∥ Fin n ∥ ≃ Bool→Type (isNonZero n) SumFin∥∥DecProp 0 = uninhabEquiv (Prop.rec isProp⊥ ⊥.rec) ⊥.rec SumFin∥∥DecProp (suc n) = isContr→≃Unit (inhProp→isContr ∣ inl tt ∣ isPropPropTrunc) +-- negation of SumFin + +isZero : ℕ → Bool +isZero 0 = true +isZero (suc n) = false + +SumFin¬ : (n : ℕ) → (¬ Fin n) ≃ Bool→Type (isZero n) +SumFin¬ 0 = isContr→≃Unit isContr⊥→A +SumFin¬ (suc n) = uninhabEquiv (λ f → f fzero) ⊥.rec + -- SumFin 1 is equivalent to unit Fin1≃Unit : Fin 1 ≃ Unit diff --git a/Cubical/Experiments/Combinatorics.agda b/Cubical/Experiments/Combinatorics.agda index 0074bff918..0bce7bd8ce 100644 --- a/Cubical/Experiments/Combinatorics.agda +++ b/Cubical/Experiments/Combinatorics.agda @@ -7,112 +7,53 @@ Computable stuff constructed from the Combinatorics of Finite Sets module Cubical.Experiments.Combinatorics where -open import Cubical.Foundations.Prelude renaming (_≡_ to _≡'_) -open import Cubical.Foundations.Equiv renaming (_≃_ to _≃'_) +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat open import Cubical.Data.Nat.Order open import Cubical.Data.Bool open import Cubical.Data.Sum -open import Cubical.Data.Sigma renaming (_×_ to _×'_) +open import Cubical.Data.Sigma open import Cubical.Data.Vec open import Cubical.Data.SumFin renaming (Fin to Fin') open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties -open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.Constructor open import Cubical.Data.FinSet.Cardinality -open import Cubical.Data.FinSet.Decidability -open import Cubical.Data.FinSet.Quotients +open import Cubical.Data.FinSet.DecidablePredicate +open import Cubical.Data.FinSet.Quotient -open import Cubical.HITs.PropositionalTruncation renaming (∥_∥ to ∥_∥') +open import Cubical.HITs.PropositionalTruncation -open import Cubical.Relation.Nullary renaming (¬_ to ¬'_ ; ∥_∥ to ∥_∥') +open import Cubical.Relation.Nullary open import Cubical.Relation.Binary -open import Cubical.Functions.Embedding renaming (_↪_ to _↪'_) -open import Cubical.Functions.Surjection renaming (_↠_ to _↠'_) +open import Cubical.Functions.Embedding +open import Cubical.Functions.Surjection private variable ℓ ℓ' ℓ'' : Level --- renaming of type constructors - -_+_ : FinSet ℓ → FinSet ℓ' → FinSet (ℓ-max ℓ ℓ') -X + Y = _ , isFinSet⊎ X Y - -_×_ : FinSet ℓ → FinSet ℓ' → FinSet (ℓ-max ℓ ℓ') -X × Y = _ , isFinSet× X Y - -_⇒_ : FinSet ℓ → FinSet ℓ' → FinSet (ℓ-max ℓ ℓ') -X ⇒ Y = _ , isFinSet→ X Y - -_≃_ : FinSet ℓ → FinSet ℓ' → FinSet (ℓ-max ℓ ℓ') -X ≃ Y = _ , isFinSet≃ X Y - -_↪_ : FinSet ℓ → FinSet ℓ' → FinSet (ℓ-max ℓ ℓ') -X ↪ Y = _ , isFinSet↪ X Y - -_↠_ : FinSet ℓ → FinSet ℓ' → FinSet (ℓ-max ℓ ℓ') -X ↠ Y = _ , isFinSet↠ X Y - -𝝨 : (X : FinSet ℓ) → (Y : X .fst → FinSet ℓ') → FinSet (ℓ-max ℓ ℓ') -𝝨 X Y = _ , isFinSetΣ X Y - -𝝥 : (X : FinSet ℓ) → (Y : X .fst → FinSet ℓ') → FinSet (ℓ-max ℓ ℓ') -𝝥 X Y = _ , isFinSetΠ X Y - -≋ : (X : FinSet ℓ) → (a b : X .fst) → FinSet ℓ -≋ X a b = _ , isFinSet≡ X a b - -¬_ : FinSet ℓ → FinSet ℓ -¬ X = _ , isFinSet¬ X - -∥_∥ : FinSet ℓ → FinSet ℓ -∥ X ∥ = _ , isFinSet∥∥ X +-- convenient renaming Fin : ℕ → FinSet ℓ-zero Fin n = _ , isFinSetFin {n = n} --- some computable functions - --- this should be addition -Fin+ : ℕ → ℕ → ℕ -Fin+ m n = card (Fin m + Fin n) - --- this should be multiplication -Fin× : ℕ → ℕ → ℕ -Fin× m n = card (Fin m × Fin n) - --- this should be exponential -Fin⇒ : ℕ → ℕ → ℕ -Fin⇒ m n = card (Fin m ⇒ Fin n) - --- this should be factorial or zero -Fin≃ : ℕ → ℕ → ℕ -Fin≃ m n = card (Fin m ≃ Fin n) - --- this should be permutation number -Fin↪ : ℕ → ℕ → ℕ -Fin↪ m n = card (Fin m ↪ Fin n) - --- this should be something that I don't know the name -Fin↠ : ℕ → ℕ → ℕ -Fin↠ m n = card (Fin m ↠ Fin n) - -- explicit numbers -s2 : Fin≃ 2 2 ≡ 2 +s2 : card (_ , isFinSet≃ (Fin 2) (Fin 2)) ≡ 2 s2 = refl -s3 : Fin≃ 3 3 ≡ 6 +s3 : card (_ , isFinSet≃ (Fin 3) (Fin 3)) ≡ 6 s3 = refl -a3,2 : Fin↪ 2 3 ≡ 6 +a3,2 : card (_ , isFinSet↪ (Fin 2) (Fin 3)) ≡ 6 a3,2 = refl -2^4 : Fin⇒ 4 2 ≡ 16 +2^4 : card (_ , isFinSet→ (Fin 4) (Fin 2)) ≡ 16 2^4 = refl -- construct numerical functions from list diff --git a/Cubical/Experiments/CountingFiniteStructure.agda b/Cubical/Experiments/CountingFiniteStructure.agda index 80abfcaeb9..60dca86983 100644 --- a/Cubical/Experiments/CountingFiniteStructure.agda +++ b/Cubical/Experiments/CountingFiniteStructure.agda @@ -18,11 +18,11 @@ open import Cubical.Data.Sigma open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties -open import Cubical.Data.FinSet.Decidability +open import Cubical.Data.FinSet.DecidablePredicate open import Cubical.Data.FinSet.Induction -open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.Constructor open import Cubical.Data.FinSet.Cardinality -open import Cubical.Data.FinSet.FinType +open import Cubical.Data.FinSet.FiniteType open import Cubical.Data.FinSet.FiniteStructure private From 1d96d214f64e2692055a20b0727c68a067158f26 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Mon, 29 Nov 2021 17:01:19 +0800 Subject: [PATCH 11/24] updata experiments --- .../Experiments/CountingFiniteStructure.agda | 96 +------------------ 1 file changed, 5 insertions(+), 91 deletions(-) diff --git a/Cubical/Experiments/CountingFiniteStructure.agda b/Cubical/Experiments/CountingFiniteStructure.agda index 60dca86983..3913530c12 100644 --- a/Cubical/Experiments/CountingFiniteStructure.agda +++ b/Cubical/Experiments/CountingFiniteStructure.agda @@ -75,12 +75,12 @@ FinGroupStr X .snd = -- two rather trivial numbers -- but the computation is essentially not that trivial --- this one can be computed in half-a-minute +-- Time: 5 ms a2 : ℕ a2 = card (_ , isFinStrCard TrivialStr 2) -- this is already hard to compute --- it takes less than half-an-hour +-- Time: 443 ms b2 : ℕ b2 = card (_ , isFinStrCard IdentityStr 2) @@ -89,11 +89,11 @@ numberOfFinSemiGroups : ℕ → ℕ numberOfFinSemiGroups n = card (_ , isFinStrCard FinSemiGroupStr n) -- two trivial cases of semi-groups --- in a flash +-- Time: 29 ms n0 : ℕ n0 = numberOfFinSemiGroups 0 --- nearly one minute +-- Time: 2,787ms n1 : ℕ n1 = numberOfFinSemiGroups 1 @@ -109,7 +109,7 @@ numberOfFinGroups : ℕ → ℕ numberOfFinGroups n = card (_ , isFinStrCard FinGroupStr n) -- group with one element --- it takes about 21 minutes +-- Time: 26,925ms g1 : ℕ g1 = numberOfFinGroups 1 @@ -117,89 +117,3 @@ g1 = numberOfFinGroups 1 -- seemed to big to do an exhaustive search g4 : ℕ g4 = numberOfFinGroups 4 - ---------------- test ----------------- - -open import Cubical.Data.SumFin - -k : ℕ -k = card (FinSemiGroupStr (Fin 3 , isFinSetFin)) - -EndoStr : FinSet ℓ → FinSet ℓ -EndoStr X = (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) - -EndoStr' : FinSet ℓ → FinSet ℓ -EndoStr' {ℓ = ℓ} X = (_ , isFinSetΣ (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ _ → 𝟙 {ℓ})) - -EndoStr'' : FinSet ℓ → FinSet ℓ -EndoStr'' {ℓ = ℓ} X = (_ , isFinSetΣ (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ _ → (Fin 1 , isFinSetFin))) - -open import Cubical.Data.Unit -open import Cubical.Relation.Nullary - -{- -isDecPropUnit : isDecProp Unit -isDecPropUnit .fst = isPropUnit -isDecPropUnit .snd = yes tt --} -isDecPropUnit : isDecProp Unit -isDecPropUnit .fst = true -isDecPropUnit .snd = idEquiv _ - -EndoStr''' : FinSet ℓ → FinSet ℓ -EndoStr''' {ℓ = ℓ} X = (_ , isFinSetSub (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ _ → Unit , isDecPropUnit)) - -l : ℕ -l = card (EndoStr (Fin 3 , isFinSetFin)) - -l' : ℕ -l' = card (EndoStr' (Fin 3 , isFinSetFin)) - -l'' : ℕ -l'' = card (EndoStr'' (Fin 2 , isFinSetFin)) - -l''' : ℕ -l''' = card (EndoStr''' (Fin 3 , isFinSetFin)) - -FinSemiGroupStr' : FinSet ℓ → FinSet ℓ -FinSemiGroupStr' X .fst = - Σ[ p ∈ (X .fst → X .fst → X .fst) ] ((x y z : X .fst) → p (p x y) z ≡ p x (p y z)) -FinSemiGroupStr' X .snd = - isFinSetSub (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) - (λ p → _ , isDecProp∀3 X (λ _ → X) (λ _ _ → X) (λ _ _ _ → _ , isDecProp≡ X _ _)) - -k' : ℕ -k' = card (FinSemiGroupStr' (Fin 2 , isFinSetFin)) - -k'' : ℕ -k'' = card (FinSemiGroupStr (Fin 2 , isFinSetFin)) - -FinGroupStr' : FinSet ℓ → FinSet ℓ -FinGroupStr' X .fst = - Σ[ e ∈ X .fst ] - Σ[ inv ∈ (X .fst → X .fst) ] - Σ[ p ∈ (X .fst → X .fst → X .fst) ] - ((x y z : X .fst) → p (p x y) z ≡ p x (p y z)) - × ((x : X .fst) - → (p x e ≡ x) × (p e x ≡ x) × (p (inv x) x ≡ e) × (p x (inv x) ≡ e)) -FinGroupStr' X .snd = - isFinSetΣ X (λ _ → _ , - isFinSetΣ (_ , isFinSetΠ X (λ _ → X)) (λ _ → _ , - isFinSetSub (_ , isFinSetΠ2 X (λ _ → X) (λ _ _ → X)) (λ _ → _ , - isDecProp× (_ , isDecProp∀3 X (λ _ → X) (λ _ _ → X) (λ _ _ _ → _ , isDecProp≡ X _ _)) (_ , - isDecProp∀ X (λ _ → _ , - isDecProp× (_ , isDecProp≡ X _ _) (_ , - isDecProp× (_ , isDecProp≡ X _ _) (_ , - isDecProp× (_ , isDecProp≡ X _ _) (_ , isDecProp≡ X _ _)))))))) - -r : ℕ -r = card (FinGroupStr' (Fin 2 , isFinSetFin)) - -r' : ℕ -r' = card (FinGroupStr (Fin 2 , isFinSetFin)) - -ff : ℕ → ℕ -ff n = card (_ , isFinStrCard FinSemiGroupStr n) - -ff' : ℕ → ℕ -ff' n = card (_ , isFinStrCard FinSemiGroupStr' n) From 3d01616801fd4c1b4dd874d302fb3989ea2f8b3f Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Tue, 28 Dec 2021 17:04:44 +0800 Subject: [PATCH 12/24] fix name --- Cubical/HITs/SetQuotients/EqClass.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/HITs/SetQuotients/EqClass.agda b/Cubical/HITs/SetQuotients/EqClass.agda index 618d2f86ef..917cadff1c 100644 --- a/Cubical/HITs/SetQuotients/EqClass.agda +++ b/Cubical/HITs/SetQuotients/EqClass.agda @@ -119,7 +119,7 @@ module _ inj/→∥ : (x y : X / R) → /→∥ x ≡ /→∥ y → x ≡ y inj/→∥ = - SetQuot.elimProp2 {C = λ x y → /→∥ x ≡ /→∥ y → x ≡ y} + SetQuot.elimProp2 {P = λ x y → /→∥ x ≡ /→∥ y → x ≡ y} (λ _ _ → isPropΠ (λ _ → squash/ _ _)) (λ x y q → Prop.rec (squash/ _ _) (λ r → eq/ _ _ r) (inj/→∥' x y q)) From f4acb241d0bb78fe7165be72d7049bff098ba7c0 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Tue, 15 Mar 2022 22:08:55 +0800 Subject: [PATCH 13/24] merge --- .github/workflows/ci-ubuntu.yml | 14 + Cubical/Algebra/Algebra.agda | 1 + Cubical/Algebra/Algebra/Base.agda | 60 +- Cubical/Algebra/Algebra/Properties.agda | 170 ++++ Cubical/Algebra/CommAlgebra.agda | 1 + Cubical/Algebra/CommAlgebra/Base.agda | 35 +- .../Instances/FreeCommAlgebra.agda | 1 + .../Algebra/CommAlgebra/Instances/Unit.agda | 25 + Cubical/Algebra/CommAlgebra/Localisation.agda | 104 ++- Cubical/Algebra/CommAlgebra/Properties.agda | 198 ++++- Cubical/Algebra/CommRing/Base.agda | 9 + Cubical/Algebra/CommRing/BinomialThm.agda | 2 +- Cubical/Algebra/CommRing/FGIdeal.agda | 100 ++- Cubical/Algebra/CommRing/FiberedProduct.agda | 1 - Cubical/Algebra/CommRing/Ideal.agda | 3 + .../CommRing/Instances/Polynomials.agda | 32 + Cubical/Algebra/CommRing/Instances/Unit.agda | 18 +- .../Algebra/CommRing/Localisation/Base.agda | 2 +- .../Localisation/InvertingElements.agda | 467 +++++++---- .../CommRing/Localisation/PullbackSquare.agda | 596 ++++++++++++++ .../Localisation/UniversalProperty.agda | 17 +- Cubical/Algebra/CommRing/Properties.agda | 111 ++- Cubical/Algebra/CommRing/RadicalIdeal.agda | 2 +- Cubical/Algebra/DistLattice/Base.agda | 3 + Cubical/Algebra/DistLattice/Basis.agda | 16 +- Cubical/Algebra/Group/Exact.agda | 88 +- Cubical/Algebra/Group/Instances/Bool.agda | 3 +- Cubical/Algebra/Group/Instances/Int.agda | 25 +- Cubical/Algebra/Group/Instances/IntMod.agda | 139 +++- Cubical/Algebra/Group/Instances/Unit.agda | 4 + .../Algebra/Group/IsomorphismTheorems.agda | 41 +- Cubical/Algebra/Group/MorphismProperties.agda | 13 +- Cubical/Algebra/Group/Morphisms.agda | 3 + Cubical/Algebra/Group/QuotientGroup.agda | 32 + Cubical/Algebra/Group/Subgroup.agda | 13 + Cubical/Algebra/Group/ZAction.agda | 442 +++++++++- Cubical/Algebra/Lattice/Base.agda | 7 + Cubical/Algebra/NatSolver/HornerForms.agda | 2 +- Cubical/Algebra/Polynomials.agda | 766 ++++++++++++++++++ Cubical/Algebra/Ring/Base.agda | 20 +- Cubical/Algebra/Ring/Properties.agda | 112 ++- Cubical/Algebra/RingSolver/RawAlgebra.agda | 1 - Cubical/Algebra/RingSolver/RawRing.agda | 2 - Cubical/Algebra/ZariskiLattice/Base.agda | 265 +----- .../ZariskiLattice/StructureSheaf.agda | 455 +++++++++++ .../ZariskiLattice/UniversalProperty.agda | 310 +++++++ Cubical/Categories/Category/Base.agda | 12 + Cubical/Categories/DistLatticeSheaf.agda | 100 ++- Cubical/Categories/Functor/Base.agda | 7 + Cubical/Categories/Instances/Algebras.agda | 28 + .../Categories/Instances/CommAlgebras.agda | 365 +++++++++ Cubical/Categories/Instances/CommRings.agda | 9 +- Cubical/Categories/Instances/Lattice.agda | 3 +- Cubical/Categories/Instances/Semilattice.agda | 14 +- Cubical/Categories/Instances/Sets.agda | 72 +- Cubical/Categories/Limits/Limits.agda | 203 +++-- Cubical/Categories/Limits/Pullback.agda | 40 +- .../NaturalTransformation/Properties.agda | 66 +- Cubical/Data/Bool/Base.agda | 3 - Cubical/Data/Bool/Properties.agda | 66 +- Cubical/Data/Empty/Base.agda | 1 - Cubical/Data/Fin/Base.agda | 7 +- Cubical/Data/Fin/Properties.agda | 45 + Cubical/Data/FinSet/Base.agda | 14 +- Cubical/Data/FinSet/Cardinality.agda | 95 ++- Cubical/Data/FinSet/Constructor.agda | 5 +- Cubical/Data/FinSet/DecidablePredicate.agda | 2 +- Cubical/Data/FinSet/FiniteChoice.agda | 5 +- Cubical/Data/FinSet/FiniteStructure.agda | 2 +- Cubical/Data/FinSet/FiniteType.agda | 2 +- Cubical/Data/FinSet/Induction.agda | 17 +- Cubical/Data/FinSet/Properties.agda | 10 +- .../FinSet/{Quotient.agda => Quotients.agda} | 11 +- Cubical/Data/HomotopyGroup.agda | 4 - Cubical/Data/HomotopyGroup/Base.agda | 52 -- Cubical/Data/Int/Properties.agda | 16 + Cubical/Data/Nat/Base.agda | 10 + Cubical/Data/Nat/Mod.agda | 53 +- Cubical/Data/Nat/Order.agda | 10 +- Cubical/Data/Nat/Properties.agda | 10 + Cubical/Data/Sigma/Properties.agda | 24 + Cubical/Data/SumFin/Properties.agda | 104 +-- Cubical/Data/Unit/Properties.agda | 20 +- Cubical/Displayed/Base.agda | 6 +- Cubical/Displayed/Function.agda | 28 - Cubical/Displayed/Properties.agda | 10 +- Cubical/Displayed/Subst.agda | 2 +- Cubical/Experiments/Brunerie.agda | 4 +- .../ZCohomologyOld/Properties.agda | 1 - .../ZariskiLatticeBasicOpens.agda} | 23 +- Cubical/Foundations/Equiv.agda | 44 +- Cubical/Foundations/Equiv/HalfAdjoint.agda | 41 +- Cubical/Foundations/GroupoidLaws.agda | 16 +- Cubical/Foundations/HLevels.agda | 24 +- Cubical/Foundations/Isomorphism.agda | 2 - Cubical/Foundations/Path.agda | 13 +- Cubical/Foundations/Pointed/Base.agda | 42 + Cubical/Foundations/Pointed/Homotopy.agda | 2 +- Cubical/Foundations/Pointed/Properties.agda | 7 + Cubical/Foundations/Prelude.agda | 25 +- Cubical/Foundations/Univalence.agda | 8 + Cubical/Functions/Morphism.agda | 1 - Cubical/HITs/AssocList/Properties.agda | 36 + Cubical/HITs/FreeComMonoids.agda | 6 + Cubical/HITs/FreeComMonoids/Base.agda | 73 ++ Cubical/HITs/FreeComMonoids/Properties.agda | 93 +++ .../PropositionalTruncation/Properties.agda | 29 + Cubical/HITs/Pushout/Base.agda | 54 ++ Cubical/HITs/Pushout/Properties.agda | 31 + Cubical/HITs/S2/Base.agda | 46 +- Cubical/HITs/SetQuotients/Properties.agda | 12 +- Cubical/HITs/SetTruncation/Base.agda | 6 + Cubical/HITs/SetTruncation/Properties.agda | 6 + Cubical/HITs/Sn/Properties.agda | 1 - Cubical/HITs/Truncation/Base.agda | 7 + Cubical/HITs/Wedge/Base.agda | 13 + Cubical/Homotopy/BlakersMassey.agda | 189 ++++- Cubical/Homotopy/Connected.agda | 128 ++- Cubical/Homotopy/Group/Base.agda | 424 +++++++++- Cubical/Homotopy/Group/LES.agda | 665 +++++++++++++++ Cubical/Homotopy/Group/Pi3S2.agda | 167 ++++ .../Homotopy/Group/Pi4S3/BrunerieNumber.agda | 444 ++++++++++ .../Homotopy/Group/Pi4S3/S3PushoutIso.agda | 682 ++++++++++++++++ .../Homotopy/Group/Pi4S3/S3PushoutIso2.agda | 64 ++ Cubical/Homotopy/Group/Pi4S3/Summary.agda | 98 +++ Cubical/Homotopy/Group/PinSn.agda | 334 ++++++++ Cubical/Homotopy/Group/S3.agda | 182 ----- Cubical/Homotopy/HopfInvariant/Brunerie.agda | 399 +++++++++ Cubical/Homotopy/Loopspace.agda | 137 +++- Cubical/Homotopy/WedgeConnectivity.agda | 2 - Cubical/Relation/Binary/Base.agda | 2 +- .../Nullary/DecidablePropositions.agda | 18 + Cubical/ZCohomology/GroupStructure.agda | 10 + Cubical/ZCohomology/Groups/SphereProduct.agda | 351 ++++++++ Cubical/ZCohomology/Groups/Unit.agda | 40 +- .../ZCohomology/RingStructure/RingLaws.agda | 16 + INSTALL.md | 11 +- README.md | 14 +- 138 files changed, 9906 insertions(+), 1421 deletions(-) create mode 100644 Cubical/Algebra/Algebra/Properties.agda create mode 100644 Cubical/Algebra/CommAlgebra/Instances/Unit.agda create mode 100644 Cubical/Algebra/CommRing/Instances/Polynomials.agda create mode 100644 Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda create mode 100644 Cubical/Algebra/Polynomials.agda create mode 100644 Cubical/Algebra/ZariskiLattice/StructureSheaf.agda create mode 100644 Cubical/Algebra/ZariskiLattice/UniversalProperty.agda create mode 100644 Cubical/Categories/Instances/Algebras.agda create mode 100644 Cubical/Categories/Instances/CommAlgebras.agda rename Cubical/Data/FinSet/{Quotient.agda => Quotients.agda} (94%) delete mode 100644 Cubical/Data/HomotopyGroup.agda delete mode 100644 Cubical/Data/HomotopyGroup/Base.agda rename Cubical/{Algebra/ZariskiLattice/BasicOpens.agda => Experiments/ZariskiLatticeBasicOpens.agda} (93%) create mode 100644 Cubical/HITs/FreeComMonoids.agda create mode 100644 Cubical/HITs/FreeComMonoids/Base.agda create mode 100644 Cubical/HITs/FreeComMonoids/Properties.agda create mode 100644 Cubical/Homotopy/Group/LES.agda create mode 100644 Cubical/Homotopy/Group/Pi3S2.agda create mode 100644 Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda create mode 100644 Cubical/Homotopy/Group/Pi4S3/S3PushoutIso.agda create mode 100644 Cubical/Homotopy/Group/Pi4S3/S3PushoutIso2.agda create mode 100644 Cubical/Homotopy/Group/Pi4S3/Summary.agda create mode 100644 Cubical/Homotopy/Group/PinSn.agda delete mode 100644 Cubical/Homotopy/Group/S3.agda create mode 100644 Cubical/Homotopy/HopfInvariant/Brunerie.agda create mode 100644 Cubical/Relation/Nullary/DecidablePropositions.agda create mode 100644 Cubical/ZCohomology/Groups/SphereProduct.agda diff --git a/.github/workflows/ci-ubuntu.yml b/.github/workflows/ci-ubuntu.yml index 5a8dbd34c5..5c4230c14c 100644 --- a/.github/workflows/ci-ubuntu.yml +++ b/.github/workflows/ci-ubuntu.yml @@ -123,3 +123,17 @@ jobs: AGDA_EXEC='~/.cabal/bin/agda -W error' \ FIX_WHITESPACE='~/.cabal/bin/fix-whitespace' \ RUNHASKELL='~/.ghcup/bin/runhaskell' + + - name: Htmlize cubical + run: | + make listings \ + AGDA_EXEC='~/.cabal/bin/agda -W error' \ + FIX_WHITESPACE='~/.cabal/bin/fix-whitespace' \ + RUNHASKELL='~/.ghcup/bin/runhaskell' + + - name: Deploy to GitHub Pages + if: github.event_name == 'push' && github.ref_name == 'master' + uses: JamesIves/github-pages-deploy-action@4.1.8 + with: + branch: gh-pages + folder: html diff --git a/Cubical/Algebra/Algebra.agda b/Cubical/Algebra/Algebra.agda index 2f9c7d9081..acdb869da5 100644 --- a/Cubical/Algebra/Algebra.agda +++ b/Cubical/Algebra/Algebra.agda @@ -2,3 +2,4 @@ module Cubical.Algebra.Algebra where open import Cubical.Algebra.Algebra.Base public +open import Cubical.Algebra.Algebra.Properties public diff --git a/Cubical/Algebra/Algebra/Base.agda b/Cubical/Algebra/Algebra/Base.agda index 6a108996b4..8db230f92e 100644 --- a/Cubical/Algebra/Algebra/Base.agda +++ b/Cubical/Algebra/Algebra/Base.agda @@ -156,27 +156,18 @@ open IsAlgebraHom AlgebraHom : {R : Ring ℓ} (M : Algebra R ℓ') (N : Algebra R ℓ'') → Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) AlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ → ⟨ N ⟩) ] IsAlgebraHom (M .snd) f (N .snd) -idAlgHom : {R : Ring ℓ} {A : Algebra R ℓ'} → AlgebraHom A A -fst idAlgHom x = x -pres0 (snd idAlgHom) = refl -pres1 (snd idAlgHom) = refl -pres+ (snd idAlgHom) x y = refl -pres· (snd idAlgHom) x y = refl -pres- (snd idAlgHom) x = refl -pres⋆ (snd idAlgHom) r x = refl - -IsAlgebraEquiv : {R : Ring ℓ} {A B : Type ℓ'} +IsAlgebraEquiv : {R : Ring ℓ} {A : Type ℓ'} {B : Type ℓ''} (M : AlgebraStr R A) (e : A ≃ B) (N : AlgebraStr R B) - → Type (ℓ-max ℓ ℓ') + → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') IsAlgebraEquiv M e N = IsAlgebraHom M (e .fst) N -AlgebraEquiv : {R : Ring ℓ} (M N : Algebra R ℓ') → Type (ℓ-max ℓ ℓ') +AlgebraEquiv : {R : Ring ℓ} (M : Algebra R ℓ') (N : Algebra R ℓ'') → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') AlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsAlgebraEquiv (M .snd) e (N .snd) _$a_ : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} → AlgebraHom A B → ⟨ A ⟩ → ⟨ B ⟩ f $a x = fst f x -AlgebraEquiv→AlgebraHom : {R : Ring ℓ} {A B : Algebra R ℓ'} +AlgebraEquiv→AlgebraHom : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} → AlgebraEquiv A B → AlgebraHom A B AlgebraEquiv→AlgebraHom (e , eIsHom) = e .fst , eIsHom @@ -218,6 +209,9 @@ isSetAlgebraEquiv : {R : Ring ℓ} (M N : Algebra R ℓ') isSetAlgebraEquiv M N = isSetΣ (isOfHLevel≃ 2 (isSetAlgebra M) (isSetAlgebra N)) λ _ → isProp→isSet (isPropIsAlgebraHom _ _ _ _) +AlgebraHom≡ : {R : Ring ℓ} {A B : Algebra R ℓ'} {φ ψ : AlgebraHom A B} → fst φ ≡ fst ψ → φ ≡ ψ +AlgebraHom≡ = Σ≡Prop λ f → isPropIsAlgebraHom _ _ f _ + 𝒮ᴰ-Algebra : (R : Ring ℓ) → DUARel (𝒮-Univ ℓ') (AlgebraStr R) (ℓ-max ℓ ℓ') 𝒮ᴰ-Algebra R = 𝒮ᴰ-Record (𝒮-Univ _) (IsAlgebraEquiv {R = R}) @@ -239,46 +233,16 @@ isSetAlgebraEquiv M N = isSetΣ (isOfHLevel≃ 2 (isSetAlgebra M) (isSetAlgebra AlgebraPath : {R : Ring ℓ} (A B : Algebra R ℓ') → (AlgebraEquiv A B) ≃ (A ≡ B) AlgebraPath {R = R} = ∫ (𝒮ᴰ-Algebra R) .UARel.ua -compIsAlgebraHom : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} - {g : ⟨ B ⟩ → ⟨ C ⟩} {f : ⟨ A ⟩ → ⟨ B ⟩} - → IsAlgebraHom (B .snd) g (C .snd) - → IsAlgebraHom (A .snd) f (B .snd) - → IsAlgebraHom (A .snd) (g ∘ f) (C .snd) -compIsAlgebraHom {g = g} {f} gh fh .pres0 = cong g (fh .pres0) ∙ gh .pres0 -compIsAlgebraHom {g = g} {f} gh fh .pres1 = cong g (fh .pres1) ∙ gh .pres1 -compIsAlgebraHom {g = g} {f} gh fh .pres+ x y = cong g (fh .pres+ x y) ∙ gh .pres+ (f x) (f y) -compIsAlgebraHom {g = g} {f} gh fh .pres· x y = cong g (fh .pres· x y) ∙ gh .pres· (f x) (f y) -compIsAlgebraHom {g = g} {f} gh fh .pres- x = cong g (fh .pres- x) ∙ gh .pres- (f x) -compIsAlgebraHom {g = g} {f} gh fh .pres⋆ r x = cong g (fh .pres⋆ r x) ∙ gh .pres⋆ r (f x) - -_∘a_ : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} - → AlgebraHom B C → AlgebraHom A B → AlgebraHom A C -_∘a_ g f .fst = g .fst ∘ f .fst -_∘a_ g f .snd = compIsAlgebraHom (g .snd) (f .snd) - -module AlgebraTheory (R : Ring ℓ) (A : Algebra R ℓ') where - open RingStr (snd R) renaming (_+_ to _+r_ ; _·_ to _·r_) - open AlgebraStr (A .snd) - - 0-actsNullifying : (x : ⟨ A ⟩) → 0r ⋆ x ≡ 0a - 0-actsNullifying x = - let idempotent-+ = 0r ⋆ x ≡⟨ cong (λ u → u ⋆ x) (sym (RingTheory.0Idempotent R)) ⟩ - (0r +r 0r) ⋆ x ≡⟨ ⋆-ldist 0r 0r x ⟩ - (0r ⋆ x) + (0r ⋆ x) ∎ - in RingTheory.+Idempotency→0 (Algebra→Ring A) (0r ⋆ x) idempotent-+ - - ⋆Dist· : (x y : ⟨ R ⟩) (a b : ⟨ A ⟩) → (x ·r y) ⋆ (a · b) ≡ (x ⋆ a) · (y ⋆ b) - ⋆Dist· x y a b = (x ·r y) ⋆ (a · b) ≡⟨ ⋆-rassoc _ _ _ ⟩ - a · ((x ·r y) ⋆ b) ≡⟨ cong (a ·_) (⋆-assoc _ _ _) ⟩ - a · (x ⋆ (y ⋆ b)) ≡⟨ sym (⋆-rassoc _ _ _) ⟩ - x ⋆ (a · (y ⋆ b)) ≡⟨ sym (⋆-lassoc _ _ _) ⟩ - (x ⋆ a) · (y ⋆ b) ∎ +uaAlgebra : {R : Ring ℓ} {A B : Algebra R ℓ'} → AlgebraEquiv A B → A ≡ B +uaAlgebra {A = A} {B = B} = equivFun (AlgebraPath A B) +isGroupoidAlgebra : {R : Ring ℓ} → isGroupoid (Algebra R ℓ') +isGroupoidAlgebra _ _ = isOfHLevelRespectEquiv 2 (AlgebraPath _ _) (isSetAlgebraEquiv _ _) -- Smart constructor for ring homomorphisms -- that infers the other equations from pres1, pres+, and pres· -module _ {R : Ring ℓ} {A : Algebra R ℓ} {B : Algebra R ℓ'} {f : ⟨ A ⟩ → ⟨ B ⟩} where +module _ {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} {f : ⟨ A ⟩ → ⟨ B ⟩} where private module A = AlgebraStr (A .snd) diff --git a/Cubical/Algebra/Algebra/Properties.agda b/Cubical/Algebra/Algebra/Properties.agda new file mode 100644 index 0000000000..7cac6d54bf --- /dev/null +++ b/Cubical/Algebra/Algebra/Properties.agda @@ -0,0 +1,170 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Algebra.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Sigma + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Auto +open import Cubical.Structures.Macro +open import Cubical.Algebra.Module +open import Cubical.Algebra.Ring +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Algebra.Base + +open Iso + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + + + + +module AlgebraTheory (R : Ring ℓ) (A : Algebra R ℓ') where + open RingStr (snd R) renaming (_+_ to _+r_ ; _·_ to _·r_) + open AlgebraStr (A .snd) + + 0-actsNullifying : (x : ⟨ A ⟩) → 0r ⋆ x ≡ 0a + 0-actsNullifying x = + let idempotent-+ = 0r ⋆ x ≡⟨ cong (λ u → u ⋆ x) (sym (RingTheory.0Idempotent R)) ⟩ + (0r +r 0r) ⋆ x ≡⟨ ⋆-ldist 0r 0r x ⟩ + (0r ⋆ x) + (0r ⋆ x) ∎ + in RingTheory.+Idempotency→0 (Algebra→Ring A) (0r ⋆ x) idempotent-+ + + ⋆Dist· : (x y : ⟨ R ⟩) (a b : ⟨ A ⟩) → (x ·r y) ⋆ (a · b) ≡ (x ⋆ a) · (y ⋆ b) + ⋆Dist· x y a b = (x ·r y) ⋆ (a · b) ≡⟨ ⋆-rassoc _ _ _ ⟩ + a · ((x ·r y) ⋆ b) ≡⟨ cong (a ·_) (⋆-assoc _ _ _) ⟩ + a · (x ⋆ (y ⋆ b)) ≡⟨ sym (⋆-rassoc _ _ _) ⟩ + x ⋆ (a · (y ⋆ b)) ≡⟨ sym (⋆-lassoc _ _ _) ⟩ + (x ⋆ a) · (y ⋆ b) ∎ + + +module AlgebraHoms {R : Ring ℓ} where + open IsAlgebraHom + + idAlgebraHom : (A : Algebra R ℓ') → AlgebraHom A A + fst (idAlgebraHom A) x = x + pres0 (snd (idAlgebraHom A)) = refl + pres1 (snd (idAlgebraHom A)) = refl + pres+ (snd (idAlgebraHom A)) x y = refl + pres· (snd (idAlgebraHom A)) x y = refl + pres- (snd (idAlgebraHom A)) x = refl + pres⋆ (snd (idAlgebraHom A)) r x = refl + + compIsAlgebraHom : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + {g : ⟨ B ⟩ → ⟨ C ⟩} {f : ⟨ A ⟩ → ⟨ B ⟩} + → IsAlgebraHom (B .snd) g (C .snd) + → IsAlgebraHom (A .snd) f (B .snd) + → IsAlgebraHom (A .snd) (g ∘ f) (C .snd) + compIsAlgebraHom {g = g} {f} gh fh .pres0 = cong g (fh .pres0) ∙ gh .pres0 + compIsAlgebraHom {g = g} {f} gh fh .pres1 = cong g (fh .pres1) ∙ gh .pres1 + compIsAlgebraHom {g = g} {f} gh fh .pres+ x y = cong g (fh .pres+ x y) ∙ gh .pres+ (f x) (f y) + compIsAlgebraHom {g = g} {f} gh fh .pres· x y = cong g (fh .pres· x y) ∙ gh .pres· (f x) (f y) + compIsAlgebraHom {g = g} {f} gh fh .pres- x = cong g (fh .pres- x) ∙ gh .pres- (f x) + compIsAlgebraHom {g = g} {f} gh fh .pres⋆ r x = cong g (fh .pres⋆ r x) ∙ gh .pres⋆ r (f x) + + compAlgebraHom : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + → AlgebraHom A B → AlgebraHom B C → AlgebraHom A C + compAlgebraHom f g .fst = g .fst ∘ f .fst + compAlgebraHom f g .snd = compIsAlgebraHom (g .snd) (f .snd) + + syntax compAlgebraHom f g = g ∘a f + + compIdAlgebraHom : {A B : Algebra R ℓ'} (φ : AlgebraHom A B) → compAlgebraHom (idAlgebraHom A) φ ≡ φ + compIdAlgebraHom φ = AlgebraHom≡ refl + + idCompAlgebraHom : {A B : Algebra R ℓ'} (φ : AlgebraHom A B) → compAlgebraHom φ (idAlgebraHom B) ≡ φ + idCompAlgebraHom φ = AlgebraHom≡ refl + + compAssocAlgebraHom : {A B C D : Algebra R ℓ'} + (φ : AlgebraHom A B) (ψ : AlgebraHom B C) (χ : AlgebraHom C D) + → compAlgebraHom (compAlgebraHom φ ψ) χ ≡ compAlgebraHom φ (compAlgebraHom ψ χ) + compAssocAlgebraHom _ _ _ = AlgebraHom≡ refl + + +module AlgebraEquivs {R : Ring ℓ} where + open IsAlgebraHom + open AlgebraHoms + + compIsAlgebraEquiv : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + {g : ⟨ B ⟩ ≃ ⟨ C ⟩} {f : ⟨ A ⟩ ≃ ⟨ B ⟩} + → IsAlgebraEquiv (B .snd) g (C .snd) + → IsAlgebraEquiv (A .snd) f (B .snd) + → IsAlgebraEquiv (A .snd) (compEquiv f g) (C .snd) + compIsAlgebraEquiv {g = g} {f} gh fh = compIsAlgebraHom {g = g .fst} {f .fst} gh fh + + compAlgebraEquiv : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + → AlgebraEquiv A B → AlgebraEquiv B C → AlgebraEquiv A C + fst (compAlgebraEquiv f g) = compEquiv (f .fst) (g .fst) + snd (compAlgebraEquiv f g) = compIsAlgebraEquiv {g = g .fst} {f = f .fst} (g .snd) (f .snd) + + + +-- the Algebra version of uaCompEquiv +module AlgebraUAFunctoriality {R : Ring ℓ} where + open AlgebraStr + open AlgebraEquivs + + Algebra≡ : (A B : Algebra R ℓ') → ( + Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] + Σ[ q0 ∈ PathP (λ i → p i) (0a (snd A)) (0a (snd B)) ] + Σ[ q1 ∈ PathP (λ i → p i) (1a (snd A)) (1a (snd B)) ] + Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] + Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] + Σ[ s- ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] + Σ[ s⋆ ∈ PathP (λ i → ⟨ R ⟩ → p i → p i) (_⋆_ (snd A)) (_⋆_ (snd B)) ] + PathP (λ i → IsAlgebra R (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i)) (isAlgebra (snd A)) + (isAlgebra (snd B))) + ≃ (A ≡ B) + Algebra≡ A B = isoToEquiv theIso + where + open Iso + theIso : Iso _ _ + fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i = p i + , algebrastr (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i) (t i) + inv theIso x = cong ⟨_⟩ x , cong (0a ∘ snd) x , cong (1a ∘ snd) x + , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (_⋆_ ∘ snd) x + , cong (isAlgebra ∘ snd) x + rightInv theIso _ = refl + leftInv theIso _ = refl + + caracAlgebra≡ : {A B : Algebra R ℓ'} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q + caracAlgebra≡ {A = A} {B = B} p q P = + sym (transportTransport⁻ (ua (Algebra≡ A B)) p) + ∙∙ cong (transport (ua (Algebra≡ A B))) helper + ∙∙ transportTransport⁻ (ua (Algebra≡ A B)) q + where + helper : transport (sym (ua (Algebra≡ A B))) p ≡ transport (sym (ua (Algebra≡ A B))) q + helper = Σ≡Prop + (λ _ → isPropΣ + (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isOfHLevelPathP 1 (isPropIsAlgebra _ _ _ _ _ _ _) _ _) + (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + + uaCompAlgebraEquiv : {A B C : Algebra R ℓ'} (f : AlgebraEquiv A B) (g : AlgebraEquiv B C) + → uaAlgebra (compAlgebraEquiv f g) ≡ uaAlgebra f ∙ uaAlgebra g + uaCompAlgebraEquiv f g = caracAlgebra≡ _ _ ( + cong ⟨_⟩ (uaAlgebra (compAlgebraEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaAlgebra f) ∙ cong ⟨_⟩ (uaAlgebra g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaAlgebra f) (uaAlgebra g)) ⟩ + cong ⟨_⟩ (uaAlgebra f ∙ uaAlgebra g) ∎) diff --git a/Cubical/Algebra/CommAlgebra.agda b/Cubical/Algebra/CommAlgebra.agda index ba2d1ee872..d150261beb 100644 --- a/Cubical/Algebra/CommAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra.agda @@ -2,3 +2,4 @@ module Cubical.Algebra.CommAlgebra where open import Cubical.Algebra.CommAlgebra.Base public +open import Cubical.Algebra.CommAlgebra.Properties public diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index 8e80df43a3..54def12f04 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -23,7 +23,7 @@ open import Cubical.Reflection.RecordEquiv private variable - ℓ ℓ' : Level + ℓ ℓ' ℓ'' ℓ''' : Level record IsCommAlgebra (R : CommRing ℓ) {A : Type ℓ'} (0a : A) (1a : A) @@ -125,7 +125,7 @@ module _ {R : CommRing ℓ} where x · (r ⋆ y) ∎) ·-comm - module _ (S : CommRing ℓ) where + module _ (S : CommRing ℓ') where open CommRingStr (snd S) renaming (1r to 1S) open CommRingStr (snd R) using () renaming (_·_ to _·R_; _+_ to _+R_; 1r to 1R) commAlgebraFromCommRing : @@ -135,31 +135,45 @@ module _ {R : CommRing ℓ} where → ((r : fst R) (x y : fst S) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) → ((x : fst S) → 1R ⋆ x ≡ x) → ((r : fst R) (x y : fst S) → (r ⋆ x) · y ≡ r ⋆ (x · y)) - → CommAlgebra R ℓ + → CommAlgebra R ℓ' commAlgebraFromCommRing _⋆_ ·Assoc⋆ ⋆DistR ⋆DistL ⋆Lid ⋆Assoc· = fst S , commalgebrastr 0r 1S _+_ _·_ -_ _⋆_ (makeIsCommAlgebra is-set +Assoc +Rid +Rinv +Comm ·Assoc ·Lid ·Ldist+ ·Comm ·Assoc⋆ ⋆DistR ⋆DistL ⋆Lid ⋆Assoc·) - IsCommAlgebraEquiv : {A B : Type ℓ'} + IsCommAlgebraEquiv : {A : Type ℓ'} {B : Type ℓ''} (M : CommAlgebraStr R A) (e : A ≃ B) (N : CommAlgebraStr R B) - → Type (ℓ-max ℓ ℓ') + → Type _ IsCommAlgebraEquiv M e N = IsAlgebraHom (CommAlgebraStr→AlgebraStr M) (e .fst) (CommAlgebraStr→AlgebraStr N) - CommAlgebraEquiv : (M N : CommAlgebra R ℓ') → Type (ℓ-max ℓ ℓ') + CommAlgebraEquiv : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') → Type _ CommAlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsCommAlgebraEquiv (M .snd) e (N .snd) - IsCommAlgebraHom : {A B : Type ℓ'} + IsCommAlgebraHom : {A : Type ℓ'} {B : Type ℓ''} (M : CommAlgebraStr R A) (f : A → B) (N : CommAlgebraStr R B) - → Type (ℓ-max ℓ ℓ') + → Type _ IsCommAlgebraHom M f N = IsAlgebraHom (CommAlgebraStr→AlgebraStr M) f (CommAlgebraStr→AlgebraStr N) - CommAlgebraHom : (M N : CommAlgebra R ℓ') → Type (ℓ-max ℓ ℓ') + CommAlgebraHom : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') → Type _ CommAlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ → ⟨ N ⟩) ] IsCommAlgebraHom (M .snd) f (N .snd) + CommAlgebraEquiv→CommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraEquiv A B → CommAlgebraHom A B + CommAlgebraEquiv→CommAlgebraHom (e , eIsHom) = e .fst , eIsHom + + CommAlgebraHom→CommRingHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') + → CommAlgebraHom A B + → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) + fst (CommAlgebraHom→CommRingHom A B f) = fst f + IsRingHom.pres0 (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres0 (snd f) + IsRingHom.pres1 (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres1 (snd f) + IsRingHom.pres+ (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres+ (snd f) + IsRingHom.pres· (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres· (snd f) + IsRingHom.pres- (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres- (snd f) + module _ {M N : CommAlgebra R ℓ'} where open CommAlgebraStr {{...}} open IsAlgebraHom @@ -239,5 +253,8 @@ isPropIsCommAlgebra R _ _ _ _ _ _ = CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) CommAlgebraPath R = ∫ (𝒮ᴰ-CommAlgebra R) .UARel.ua +uaCommAlgebra : {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} → CommAlgebraEquiv A B → A ≡ B +uaCommAlgebra {R = R} {A = A} {B = B} = equivFun (CommAlgebraPath R A B) + isGroupoidCommAlgebra : {R : CommRing ℓ} → isGroupoid (CommAlgebra R ℓ') isGroupoidCommAlgebra A B = isOfHLevelRespectEquiv 2 (CommAlgebraPath _ _ _) (isSetAlgebraEquiv _ _) diff --git a/Cubical/Algebra/CommAlgebra/Instances/FreeCommAlgebra.agda b/Cubical/Algebra/CommAlgebra/Instances/FreeCommAlgebra.agda index 2638d4cec4..c6f30f99ad 100644 --- a/Cubical/Algebra/CommAlgebra/Instances/FreeCommAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/Instances/FreeCommAlgebra.agda @@ -427,6 +427,7 @@ inducedHom : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') inducedHom A φ = Theory.inducedHom A φ module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ''} where + open AlgebraHoms A′ = CommAlgebra→Algebra A B′ = CommAlgebra→Algebra B R′ = (CommRing→Ring R) diff --git a/Cubical/Algebra/CommAlgebra/Instances/Unit.agda b/Cubical/Algebra/CommAlgebra/Instances/Unit.agda new file mode 100644 index 0000000000..cbccbeaf91 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Instances/Unit.agda @@ -0,0 +1,25 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.Instances.Unit where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Unit + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Instances.Unit +open import Cubical.Algebra.CommAlgebra + +private + variable + ℓ ℓ' : Level + +open CommAlgebraStr + +module _ (R : CommRing ℓ) where + + UnitCommAlgebra : CommAlgebra R ℓ' + UnitCommAlgebra = commAlgebraFromCommRing + UnitCommRing + (λ _ _ → tt*) (λ _ _ _ → refl) (λ _ _ _ → refl) + (λ _ _ _ → refl) (λ _ → refl) (λ _ _ _ → refl) diff --git a/Cubical/Algebra/CommAlgebra/Localisation.agda b/Cubical/Algebra/CommAlgebra/Localisation.agda index a0d9ecaab4..eaa98ef1c7 100644 --- a/Cubical/Algebra/CommAlgebra/Localisation.agda +++ b/Cubical/Algebra/CommAlgebra/Localisation.agda @@ -7,10 +7,13 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Foundations.SIP open import Cubical.Foundations.Powerset open import Cubical.Data.Sigma +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) +open import Cubical.Data.FinData open import Cubical.Reflection.StrictEquiv @@ -19,13 +22,19 @@ open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Monoid open import Cubical.Algebra.CommRing.Base open import Cubical.Algebra.CommRing.Properties +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.RadicalIdeal open import Cubical.Algebra.CommRing.Localisation.Base open import Cubical.Algebra.CommRing.Localisation.UniversalProperty +open import Cubical.Algebra.CommRing.Localisation.InvertingElements open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommAlgebra.Base open import Cubical.Algebra.CommAlgebra.Properties +open import Cubical.Algebra.RingSolver.Reflection + open import Cubical.HITs.SetQuotients as SQ open import Cubical.HITs.PropositionalTruncation as PT @@ -124,7 +133,21 @@ module AlgLoc (R' : CommRing ℓ) isContrHomS⁻¹RS⁻¹R : isContr (CommAlgebraHom S⁻¹RAsCommAlg S⁻¹RAsCommAlg) isContrHomS⁻¹RS⁻¹R = S⁻¹RHasAlgUniversalProp S⁻¹RAsCommAlg S⋆1⊆S⁻¹Rˣ + S⁻¹RAlgCharEquiv : (A' : CommRing ℓ) (φ : CommRingHom R' A') + → PathToS⁻¹R R' S' SMultClosedSubset A' φ + → CommAlgebraEquiv S⁻¹RAsCommAlg (toCommAlg (A' , φ)) + S⁻¹RAlgCharEquiv A' φ cond = toCommAlgebraEquiv _ _ + (S⁻¹RCharEquiv R' S' SMultClosedSubset A' φ cond) + (RingHom≡ (S⁻¹RHasUniversalProp A' φ (cond .φS⊆Aˣ) .fst .snd)) + where open PathToS⁻¹R + +-- the special case of localizing at a single element +R[1/_]AsCommAlgebra : {R : CommRing ℓ} → fst R → CommAlgebra R ℓ +R[1/_]AsCommAlgebra {R = R} f = S⁻¹RAsCommAlg [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + where + open AlgLoc R + open InvertingElementsBase R module AlgLocTwoSubsets (R' : CommRing ℓ) @@ -147,6 +170,8 @@ module AlgLocTwoSubsets (R' : CommRing ℓ) ; isContrHomS⁻¹RS⁻¹R to isContrHomS₂⁻¹RS₂⁻¹R) open IsAlgebraHom + open AlgebraHoms + open CommAlgebraHoms open CommAlgebraStr ⦃...⦄ private @@ -171,10 +196,10 @@ module AlgLocTwoSubsets (R' : CommRing ℓ) χ₂ : CommAlgebraHom S₂⁻¹RAsCommAlg S₁⁻¹RAsCommAlg χ₂ = S₂⁻¹RHasAlgUniversalProp S₁⁻¹RAsCommAlg S₂⊆S₁⁻¹Rˣ .fst - χ₁∘χ₂≡id : χ₁ ∘a χ₂ ≡ idAlgHom + χ₁∘χ₂≡id : χ₁ ∘a χ₂ ≡ idCommAlgebraHom _ χ₁∘χ₂≡id = isContr→isProp isContrHomS₂⁻¹RS₂⁻¹R _ _ - χ₂∘χ₁≡id : χ₂ ∘a χ₁ ≡ idAlgHom + χ₂∘χ₁≡id : χ₂ ∘a χ₁ ≡ idCommAlgebraHom _ χ₂∘χ₁≡id = isContr→isProp isContrHomS₁⁻¹RS₁⁻¹R _ _ IsoS₁⁻¹RS₂⁻¹R : Iso S₁⁻¹R S₂⁻¹R @@ -215,3 +240,76 @@ module AlgLocTwoSubsets (R' : CommRing ℓ) S₂⊆S₁⁻¹Rˣ s₂ s₂∈S₂ = transport (λ i → _⋆_ ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄ s₂ (1a ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄) ∈ (CommAlgebra→CommRing ((sym S₁⁻¹R≡S₂⁻¹R) i)) ˣ) (S₂⋆1⊆S₂⁻¹Rˣ s₂ s₂∈S₂) + + + +-- A crucial result for the construction of the structure sheaf +module DoubleAlgLoc (R : CommRing ℓ) (f g : (fst R)) where + open Exponentiation R + open InvertingElementsBase + open CommRingStr (snd R) hiding (·Rid) + open isMultClosedSubset + open DoubleLoc R f g hiding (R[1/fg]≡R[1/f][1/g]) + open CommAlgChar R + open AlgLoc R ([_ⁿ|n≥0] R (f · g)) (powersFormMultClosedSubset R (f · g)) + renaming (S⁻¹RAlgCharEquiv to R[1/fg]AlgCharEquiv) + open CommIdeal R hiding (subst-∈) renaming (_∈_ to _∈ᵢ_) + open isCommIdeal + open RadicalIdeal R + + private + ⟨_⟩ : (fst R) → CommIdeal + ⟨ f ⟩ = ⟨ replicateFinVec 1 f ⟩[ R ] + + R[1/fg]AsCommAlgebra = R[1/_]AsCommAlgebra {R = R} (f · g) + R[1/fg]ˣ = R[1/_]AsCommRing R (f · g) ˣ + + R[1/g]AsCommAlgebra = R[1/_]AsCommAlgebra {R = R} g + R[1/g]ˣ = R[1/_]AsCommRing R g ˣ + + R[1/f][1/g]AsCommRing = R[1/_]AsCommRing (R[1/_]AsCommRing R f) + [ g , 1r , powersFormMultClosedSubset R f .containsOne ] + R[1/f][1/g]AsCommAlgebra = toCommAlg (R[1/f][1/g]AsCommRing , /1/1AsCommRingHom) + + R[1/fg]≡R[1/f][1/g] : R[1/fg]AsCommAlgebra ≡ R[1/f][1/g]AsCommAlgebra + R[1/fg]≡R[1/f][1/g] = uaCommAlgebra (R[1/fg]AlgCharEquiv _ _ pathtoR[1/fg]) + + doubleLocCancel : g ∈ᵢ √ ⟨ f ⟩ → R[1/f][1/g]AsCommAlgebra ≡ R[1/g]AsCommAlgebra + doubleLocCancel g∈√⟨f⟩ = sym R[1/fg]≡R[1/f][1/g] ∙ isContrR[1/fg]≡R[1/g] toUnit1 toUnit2 .fst + where + open S⁻¹RUniversalProp R ([_ⁿ|n≥0] R g) (powersFormMultClosedSubset R g) + renaming (_/1 to _/1ᵍ) + open S⁻¹RUniversalProp R ([_ⁿ|n≥0] R (f · g)) (powersFormMultClosedSubset R (f · g)) + renaming (_/1 to _/1ᶠᵍ) + open AlgLocTwoSubsets R ([_ⁿ|n≥0] R (f · g)) (powersFormMultClosedSubset R (f · g)) + ([_ⁿ|n≥0] R g) (powersFormMultClosedSubset R g) + renaming (isContrS₁⁻¹R≡S₂⁻¹R to isContrR[1/fg]≡R[1/g]) + open CommAlgebraStr ⦃...⦄ hiding (_·_ ; _+_) + instance + _ = snd R[1/fg]AsCommAlgebra + _ = snd R[1/g]AsCommAlgebra + + toUnit1 : ∀ s → s ∈ [_ⁿ|n≥0] R (f · g) → s ⋆ 1a ∈ R[1/g]ˣ + toUnit1 s s∈[fgⁿ|n≥0] = subst-∈ R[1/g]ˣ (sym (·Rid (s /1ᵍ))) + (RadicalLemma.toUnit R g (f · g) (radHelper _ _ g∈√⟨f⟩) s s∈[fgⁿ|n≥0]) + where + radHelper : ∀ x y → x ∈ᵢ √ ⟨ y ⟩ → x ∈ᵢ √ ⟨ y · x ⟩ + radHelper x y = PT.rec ((√ ⟨ y · x ⟩) .fst x .snd) (uncurry helper1) + where + helper1 : (n : ℕ) → x ^ n ∈ᵢ ⟨ y ⟩ → x ∈ᵢ √ ⟨ y · x ⟩ + helper1 n = PT.rec ((√ ⟨ y · x ⟩) .fst x .snd) (uncurry helper2) + where + helper2 : (α : FinVec (fst R) 1) + → x ^ n ≡ linearCombination R α (replicateFinVec 1 y) + → x ∈ᵢ √ ⟨ y · x ⟩ + helper2 α p = ∣ (suc n) , ∣ α , cong (x ·_) p ∙ useSolver x y (α zero) ∣ ∣ + where + useSolver : ∀ x y a → x · (a · y + 0r) ≡ a · (y · x) + 0r + useSolver = solve R + + toUnit2 : ∀ s → s ∈ [_ⁿ|n≥0] R g → s ⋆ 1a ∈ R[1/fg]ˣ + toUnit2 s s∈[gⁿ|n≥0] = subst-∈ R[1/fg]ˣ (sym (·Rid (s /1ᶠᵍ))) + (RadicalLemma.toUnit R (f · g) g radHelper s s∈[gⁿ|n≥0]) + where + radHelper : (f · g) ∈ᵢ √ ⟨ g ⟩ + radHelper = ·Closed (snd (√ ⟨ g ⟩)) f (∈→∈√ _ _ (indInIdeal R _ zero)) diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index cfa46fbc15..8ffe8d2946 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -2,11 +2,16 @@ module Cubical.Algebra.CommAlgebra.Properties where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport open import Cubical.Foundations.SIP +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path open import Cubical.Data.Sigma @@ -20,9 +25,11 @@ open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.HITs.PropositionalTruncation + private variable - ℓ ℓ′ : Level + ℓ ℓ' ℓ'' ℓ''' : Level -- An R-algebra is the same as a CommRing A with a CommRingHom φ : R → A @@ -111,3 +118,190 @@ module CommAlgChar (R : CommRing ℓ) where inv CommAlgIso = toCommAlg rightInv CommAlgIso = CommRingWithHomRoundTrip leftInv CommAlgIso = CommAlgRoundTrip + + open RingHoms + isCommRingWithHomHom : (A B : CommRingWithHom) → CommRingHom (fst A) (fst B) → Type ℓ + isCommRingWithHomHom (_ , f) (_ , g) h = h ∘r f ≡ g + + CommRingWithHomHom : CommRingWithHom → CommRingWithHom → Type ℓ + CommRingWithHomHom (A , f) (B , g) = Σ[ h ∈ CommRingHom A B ] h ∘r f ≡ g + + toCommAlgebraHom : (A B : CommRingWithHom) (h : CommRingHom (fst A) (fst B)) + → isCommRingWithHomHom A B h + → CommAlgebraHom (toCommAlg A) (toCommAlg B) + toCommAlgebraHom (A , f) (B , g) h commDiag = + makeCommAlgebraHom (fst h) (pres1 (snd h)) (pres+ (snd h)) (pres· (snd h)) pres⋆h + where + open CommRingStr ⦃...⦄ + instance + _ = snd A + _ = snd B + pres⋆h : ∀ r x → fst h (fst f r · x) ≡ fst g r · fst h x + pres⋆h r x = fst h (fst f r · x) ≡⟨ pres· (snd h) _ _ ⟩ + fst h (fst f r) · fst h x ≡⟨ cong (λ φ → fst φ r · fst h x) commDiag ⟩ + fst g r · fst h x ∎ + + fromCommAlgebraHom : (A B : CommAlgebra R ℓ) → CommAlgebraHom A B + → CommRingWithHomHom (fromCommAlg A) (fromCommAlg B) + fst (fst (fromCommAlgebraHom A B f)) = fst f + pres0 (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres0 (snd f) + pres1 (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres1 (snd f) + pres+ (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres+ (snd f) + pres· (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres· (snd f) + pres- (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres- (snd f) + snd (fromCommAlgebraHom A B f) = + RingHom≡ (funExt (λ x → IsAlgebraHom.pres⋆ (snd f) x 1a ∙ cong (x ⋆_) (IsAlgebraHom.pres1 (snd f)))) + where + open CommAlgebraStr (snd A) using (1a) + open CommAlgebraStr (snd B) using (_⋆_) + + isCommRingWithHomEquiv : (A B : CommRingWithHom) → CommRingEquiv (fst A) (fst B) → Type ℓ + isCommRingWithHomEquiv A B e = isCommRingWithHomHom A B (RingEquiv→RingHom e) + + CommRingWithHomEquiv : CommRingWithHom → CommRingWithHom → Type ℓ + CommRingWithHomEquiv A B = Σ[ e ∈ CommRingEquiv (fst A) (fst B) ] isCommRingWithHomEquiv A B e + + toCommAlgebraEquiv : (A B : CommRingWithHom) (e : CommRingEquiv (fst A) (fst B)) + → isCommRingWithHomEquiv A B e + → CommAlgebraEquiv (toCommAlg A) (toCommAlg B) + fst (toCommAlgebraEquiv A B e eCommDiag) = e .fst + snd (toCommAlgebraEquiv A B e eCommDiag) = toCommAlgebraHom A B _ eCommDiag .snd + + + +module CommAlgebraHoms {R : CommRing ℓ} where + open AlgebraHoms + + idCommAlgebraHom : (A : CommAlgebra R ℓ') → CommAlgebraHom A A + idCommAlgebraHom A = idAlgebraHom (CommAlgebra→Algebra A) + + compCommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') (C : CommAlgebra R ℓ''') + → CommAlgebraHom A B → CommAlgebraHom B C → CommAlgebraHom A C + compCommAlgebraHom A B C = compAlgebraHom {A = CommAlgebra→Algebra A} + {CommAlgebra→Algebra B} {CommAlgebra→Algebra C} + + _∘ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} + → CommAlgebraHom B C → CommAlgebraHom A B → CommAlgebraHom A C + g ∘ca f = compCommAlgebraHom _ _ _ f g + + compIdCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) + → compCommAlgebraHom _ _ _ (idCommAlgebraHom A) f ≡ f + compIdCommAlgebraHom = compIdAlgebraHom + + idCompCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) + → compCommAlgebraHom _ _ _ f (idCommAlgebraHom B) ≡ f + idCompCommAlgebraHom = idCompAlgebraHom + + compAssocCommAlgebraHom : {A B C D : CommAlgebra R ℓ'} + (f : CommAlgebraHom A B) (g : CommAlgebraHom B C) (h : CommAlgebraHom C D) + → compCommAlgebraHom _ _ _ (compCommAlgebraHom _ _ _ f g) h + ≡ compCommAlgebraHom _ _ _ f (compCommAlgebraHom _ _ _ g h) + compAssocCommAlgebraHom = compAssocAlgebraHom + +module CommAlgebraEquivs {R : CommRing ℓ} where + open AlgebraEquivs + + compCommAlgebraEquiv : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} + → CommAlgebraEquiv A B → CommAlgebraEquiv B C → CommAlgebraEquiv A C + compCommAlgebraEquiv {A = A} {B = B} {C = C} = compAlgebraEquiv {A = CommAlgebra→Algebra A} + {B = CommAlgebra→Algebra B} + {C = CommAlgebra→Algebra C} + + +-- the CommAlgebra version of uaCompEquiv +module CommAlgebraUAFunctoriality {R : CommRing ℓ} where + open CommAlgebraStr + open CommAlgebraEquivs + + CommAlgebra≡ : (A B : CommAlgebra R ℓ') → ( + Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] + Σ[ q0 ∈ PathP (λ i → p i) (0a (snd A)) (0a (snd B)) ] + Σ[ q1 ∈ PathP (λ i → p i) (1a (snd A)) (1a (snd B)) ] + Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] + Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] + Σ[ s- ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] + Σ[ s⋆ ∈ PathP (λ i → ⟨ R ⟩ → p i → p i) (_⋆_ (snd A)) (_⋆_ (snd B)) ] + PathP (λ i → IsCommAlgebra R (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i)) (isCommAlgebra (snd A)) + (isCommAlgebra (snd B))) + ≃ (A ≡ B) + CommAlgebra≡ A B = isoToEquiv theIso + where + open Iso + theIso : Iso _ _ + fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i = p i + , commalgebrastr (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i) (t i) + inv theIso x = cong ⟨_⟩ x , cong (0a ∘ snd) x , cong (1a ∘ snd) x + , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (_⋆_ ∘ snd) x + , cong (isCommAlgebra ∘ snd) x + rightInv theIso _ = refl + leftInv theIso _ = refl + + caracCommAlgebra≡ : {A B : CommAlgebra R ℓ'} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q + caracCommAlgebra≡ {A = A} {B = B} p q P = + sym (transportTransport⁻ (ua (CommAlgebra≡ A B)) p) + ∙∙ cong (transport (ua (CommAlgebra≡ A B))) helper + ∙∙ transportTransport⁻ (ua (CommAlgebra≡ A B)) q + where + helper : transport (sym (ua (CommAlgebra≡ A B))) p ≡ transport (sym (ua (CommAlgebra≡ A B))) q + helper = Σ≡Prop + (λ _ → isPropΣ + (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isOfHLevelPathP 1 (isPropIsCommAlgebra _ _ _ _ _ _ _) _ _) + (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + + uaCompCommAlgebraEquiv : {A B C : CommAlgebra R ℓ'} (f : CommAlgebraEquiv A B) (g : CommAlgebraEquiv B C) + → uaCommAlgebra (compCommAlgebraEquiv f g) ≡ uaCommAlgebra f ∙ uaCommAlgebra g + uaCompCommAlgebraEquiv f g = caracCommAlgebra≡ _ _ ( + cong ⟨_⟩ (uaCommAlgebra (compCommAlgebraEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaCommAlgebra f) ∙ cong ⟨_⟩ (uaCommAlgebra g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaCommAlgebra f) (uaCommAlgebra g)) ⟩ + cong ⟨_⟩ (uaCommAlgebra f ∙ uaCommAlgebra g) ∎) + + +open CommAlgebraHoms +open CommAlgebraEquivs +open CommAlgebraUAFunctoriality +recPT→CommAlgebra : {R : CommRing ℓ} {A : Type ℓ'} (𝓕 : A → CommAlgebra R ℓ'') + → (σ : ∀ x y → CommAlgebraEquiv (𝓕 x) (𝓕 y)) + → (∀ x y z → σ x z ≡ compCommAlgebraEquiv (σ x y) (σ y z)) + ------------------------------------------------------ + → ∥ A ∥ → CommAlgebra R ℓ'' +recPT→CommAlgebra 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommAlgebra 𝓕 + (3-ConstantCompChar 𝓕 (λ x y → uaCommAlgebra (σ x y)) + λ x y z → sym ( cong uaCommAlgebra (compCoh x y z) + ∙ uaCompCommAlgebraEquiv (σ x y) (σ y z))) + + +contrCommAlgebraHom→contrCommAlgebraEquiv : {R : CommRing ℓ} {A : Type ℓ'} + (σ : A → CommAlgebra R ℓ'') + → (∀ x y → isContr (CommAlgebraHom (σ x) (σ y))) + ---------------------------------------------------------------------------- + → ∀ x y → isContr (CommAlgebraEquiv (σ x) (σ y)) +contrCommAlgebraHom→contrCommAlgebraEquiv σ contrHom x y = σEquiv , + λ e → Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) + (Σ≡Prop isPropIsEquiv (cong fst (contrHom _ _ .snd (CommAlgebraEquiv→CommAlgebraHom e)))) + where + open Iso + χ₁ = contrHom x y .fst + χ₂ = contrHom y x .fst + χ₁∘χ₂≡id : χ₁ ∘ca χ₂ ≡ idCommAlgebraHom _ + χ₁∘χ₂≡id = isContr→isProp (contrHom _ _) _ _ + + χ₂∘χ₁≡id : χ₂ ∘ca χ₁ ≡ idCommAlgebraHom _ + χ₂∘χ₁≡id = isContr→isProp (contrHom _ _) _ _ + + σIso : Iso ⟨ σ x ⟩ ⟨ σ y ⟩ + fun σIso = fst χ₁ + inv σIso = fst χ₂ + rightInv σIso = funExt⁻ (cong fst χ₁∘χ₂≡id) + leftInv σIso = funExt⁻ (cong fst χ₂∘χ₁≡id) + + σEquiv : CommAlgebraEquiv (σ x) (σ y) + fst σEquiv = isoToEquiv σIso + snd σEquiv = snd χ₁ diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda index 802edc9155..8b973f9c44 100644 --- a/Cubical/Algebra/CommRing/Base.agda +++ b/Cubical/Algebra/CommRing/Base.agda @@ -107,6 +107,9 @@ IsCommRingEquiv R e S = IsRingHom (CommRingStr→RingStr R) (e .fst) (CommRingSt CommRingEquiv : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') CommRingEquiv R S = Σ[ e ∈ (R .fst ≃ S .fst) ] IsCommRingEquiv (R .snd) e (S .snd) +CommRingEquiv→CommRingHom : {A : CommRing ℓ} {B : CommRing ℓ'} → CommRingEquiv A B → CommRingHom A B +CommRingEquiv→CommRingHom (e , eIsHom) = e .fst , eIsHom + isPropIsCommRing : {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) → isProp (IsCommRing 0r 1r _+_ _·_ -_) isPropIsCommRing 0r 1r _+_ _·_ -_ (iscommring RR RC) (iscommring SR SC) = @@ -140,5 +143,11 @@ isPropIsCommRing 0r 1r _+_ _·_ -_ (iscommring RR RC) (iscommring SR SC) = CommRingPath : (R S : CommRing ℓ) → CommRingEquiv R S ≃ (R ≡ S) CommRingPath = ∫ 𝒮ᴰ-CommRing .UARel.ua +uaCommRing : {A B : CommRing ℓ} → CommRingEquiv A B → A ≡ B +uaCommRing {A = A} {B = B} = equivFun (CommRingPath A B) + isSetCommRing : ((R , str) : CommRing ℓ) → isSet R isSetCommRing (R , str) = str .CommRingStr.is-set + +isGroupoidCommRing : isGroupoid (CommRing ℓ) +isGroupoidCommRing _ _ = isOfHLevelRespectEquiv 2 (CommRingPath _ _) (isSetRingEquiv _ _) diff --git a/Cubical/Algebra/CommRing/BinomialThm.agda b/Cubical/Algebra/CommRing/BinomialThm.agda index f8e49041d9..3ed7696a91 100644 --- a/Cubical/Algebra/CommRing/BinomialThm.agda +++ b/Cubical/Algebra/CommRing/BinomialThm.agda @@ -4,7 +4,7 @@ module Cubical.Algebra.CommRing.BinomialThm where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function -open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ ; +-comm to +ℕ-comm ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm ; _choose_ to _ℕchoose_ ; snotz to ℕsnotz) diff --git a/Cubical/Algebra/CommRing/FGIdeal.agda b/Cubical/Algebra/CommRing/FGIdeal.agda index d874c53cb5..0ac9b8d8e9 100644 --- a/Cubical/Algebra/CommRing/FGIdeal.agda +++ b/Cubical/Algebra/CommRing/FGIdeal.agda @@ -13,16 +13,19 @@ open import Cubical.Foundations.Transport open import Cubical.Foundations.HLevels open import Cubical.Data.Sigma +open import Cubical.Data.Sum hiding (map ; elim ; rec) open import Cubical.Data.FinData hiding (elim ; rec) open import Cubical.Data.Nat renaming ( zero to ℕzero ; suc to ℕsuc - ; _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ ; +-assoc to +ℕ-assoc ; +-comm to +ℕ-comm ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) - hiding (elim) + hiding (elim ; _choose_) +open import Cubical.Data.Nat.Order open import Cubical.HITs.PropositionalTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.BinomialThm open import Cubical.Algebra.Ring.QuotientRing open import Cubical.Algebra.Ring.Properties open import Cubical.Algebra.Ring.BigOps @@ -247,3 +250,96 @@ module _ (R' : CommRing ℓ) where FGIdealMultLemma : {n m : ℕ} (U : FinVec R n) (V : FinVec R m) → ⟨ U ··Fin V ⟩ ≡ ⟨ U ⟩ ·i ⟨ V ⟩ FGIdealMultLemma U V = CommIdeal≡Char (FGIdealMultLemmaLIncl U V) (FGIdealMultLemmaRIncl U V) + + + + +-- A useful lemma for constructing the structure sheaf +module GeneratingExponents (R' : CommRing ℓ) (f g : fst R') (n : ℕ) where + open CommRingStr (snd R') + open RingTheory (CommRing→Ring R') + open Sum (CommRing→Ring R') + open Exponentiation R' + open BinomialThm R' + open CommIdeal R' + + private + R = fst R' + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + fgVec : FinVec R 2 + fgVec zero = f + fgVec (suc zero) = g + ⟨f,g⟩ = ⟨ fgVec ⟩ + + fⁿgⁿVec : FinVec R 2 + fⁿgⁿVec zero = f ^ n + fⁿgⁿVec (suc zero) = g ^ n + + ⟨fⁿ,gⁿ⟩ = ⟨ fⁿgⁿVec ⟩ + + lemma : 1r ∈ ⟨f,g⟩ → 1r ∈ ⟨fⁿ,gⁿ⟩ + lemma = elim (λ _ → isPropPropTrunc) Σlemma + where + Σlemma : Σ[ α ∈ FinVec R 2 ] (1r ≡ linearCombination R' α fgVec) → 1r ∈ ⟨fⁿ,gⁿ⟩ + Σlemma (α , p) = subst-∈ ⟨fⁿ,gⁿ⟩ path ∑Binomial∈⟨fⁿ,gⁿ⟩ + where + α₀f = α zero · f + α₁g = α (suc zero) · g + + binomialSummand∈⟨fⁿ,gⁿ⟩ : (i : Fin (ℕsuc (n +ℕ n))) → BinomialVec (n +ℕ n) α₀f α₁g i ∈ ⟨fⁿ,gⁿ⟩ + binomialSummand∈⟨fⁿ,gⁿ⟩ i with ≤-+-split n n (toℕ i) (pred-≤-pred (toℕ R[1/f] + ⌟ + _/1 | | χ₁ + v v + + R[1/g] -> R[1/fg] + χ₂ + + where the morphisms χ are induced by the universal property + of localization. + + -} + + +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommRing.Localisation.PullbackSquare where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset renaming (_∈_ to _∈ₚ_) +open import Cubical.Foundations.Transport +open import Cubical.Functions.FunExtEquiv + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Vec +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Algebra.Group +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Localisation.Base +open import Cubical.Algebra.CommRing.Localisation.UniversalProperty +open import Cubical.Algebra.CommRing.Localisation.InvertingElements +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.RadicalIdeal + +open import Cubical.Algebra.RingSolver.Reflection + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Instances.CommRings +open import Cubical.Categories.Limits.Pullback + +open Iso + +private + variable + ℓ ℓ' : Level + A : Type ℓ + + +module _ (R' : CommRing ℓ) (f g : (fst R')) where + + open IsRingHom + open isMultClosedSubset + open CommRingTheory R' + open RingTheory (CommRing→Ring R') + open CommIdeal R' + open RadicalIdeal R' + open Exponentiation R' + open InvertingElementsBase R' + + open CommRingStr ⦃...⦄ + private + R = R' .fst + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + fgVec : FinVec R 2 + fgVec zero = f + fgVec (suc zero) = g + ⟨f,g⟩ = ⟨ fgVec ⟩ + fⁿgⁿVec : (n : ℕ) → FinVec R 2 + fⁿgⁿVec n zero = f ^ n + fⁿgⁿVec n (suc zero) = g ^ n + ⟨fⁿ,gⁿ⟩ : (n : ℕ) → CommIdeal + ⟨fⁿ,gⁿ⟩ n = ⟨ (fⁿgⁿVec n) ⟩ + + instance + _ = R' .snd + _ = R[1/ f ]AsCommRing .snd + _ = R[1/ g ]AsCommRing .snd + _ = R[1/ (R' .snd .CommRingStr._·_ f g) ]AsCommRing .snd + + open Loc R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + renaming (_≈_ to _≈ᶠ_ ; locIsEquivRel to locIsEquivRelᶠ ; S to Sᶠ) + open S⁻¹RUniversalProp R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + renaming ( _/1 to _/1ᶠ ; /1AsCommRingHom to /1ᶠAsCommRingHom + ; S⁻¹RHasUniversalProp to R[1/f]HasUniversalProp) + open Loc R' [ g ⁿ|n≥0] (powersFormMultClosedSubset g) + renaming (_≈_ to _≈ᵍ_ ; locIsEquivRel to locIsEquivRelᵍ ; S to Sᵍ) + open S⁻¹RUniversalProp R' [ g ⁿ|n≥0] (powersFormMultClosedSubset g) + renaming ( _/1 to _/1ᵍ ; /1AsCommRingHom to /1ᵍAsCommRingHom + ; S⁻¹RHasUniversalProp to R[1/g]HasUniversalProp) + open Loc R' [ (f · g) ⁿ|n≥0] (powersFormMultClosedSubset (f · g)) + renaming (_≈_ to _≈ᶠᵍ_ ; locIsEquivRel to locIsEquivRelᶠᵍ ; S to Sᶠᵍ) + open S⁻¹RUniversalProp R' [ (f · g) ⁿ|n≥0] (powersFormMultClosedSubset (f · g)) + renaming ( _/1 to _/1ᶠᵍ ; /1AsCommRingHom to /1ᶠᵍAsCommRingHom) + + + -- the pullback legs + private + -- using RadicalLemma doesn't compute... + χ₁ : CommRingHom R[1/ f ]AsCommRing R[1/ (f · g) ]AsCommRing + χ₁ = R[1/f]HasUniversalProp _ /1ᶠᵍAsCommRingHom unitHelper .fst .fst + where + unitHelper : ∀ s → s ∈ₚ [ f ⁿ|n≥0] → s /1ᶠᵍ ∈ₚ (R[1/ (f · g) ]AsCommRing) ˣ + unitHelper = powersPropElim (λ s → Units.inverseUniqueness _ (s /1ᶠᵍ)) + λ n → [ g ^ n , (f · g) ^ n , ∣ n , refl ∣ ] + , eq/ _ _ ((1r , powersFormMultClosedSubset (f · g) .containsOne) + , path n) + where + useSolver1 : ∀ a b → 1r · (a · b) · 1r ≡ a · b + useSolver1 = solve R' + useSolver2 : ∀ a → a ≡ (1r · 1r) · (1r · a) + useSolver2 = solve R' + + path : (n : ℕ) → 1r · (f ^ n · g ^ n) · 1r ≡ (1r · 1r) · (1r · ((f · g) ^ n)) + path n = useSolver1 _ _ ∙ sym (^-ldist-· f g n) ∙ useSolver2 _ + + χ₂ : CommRingHom R[1/ g ]AsCommRing R[1/ (f · g) ]AsCommRing + χ₂ = R[1/g]HasUniversalProp _ /1ᶠᵍAsCommRingHom unitHelper .fst .fst + where + unitHelper : ∀ s → s ∈ₚ [ g ⁿ|n≥0] → s /1ᶠᵍ ∈ₚ (R[1/ (f · g) ]AsCommRing) ˣ + unitHelper = powersPropElim (λ s → Units.inverseUniqueness _ (s /1ᶠᵍ)) + λ n → [ f ^ n , (f · g) ^ n , ∣ n , refl ∣ ] + , eq/ _ _ ((1r , powersFormMultClosedSubset (f · g) .containsOne) + , path n) + where + useSolver1 : ∀ a b → 1r · (a · b) · 1r ≡ b · a + useSolver1 = solve R' + useSolver2 : ∀ a → a ≡ (1r · 1r) · (1r · a) + useSolver2 = solve R' + + path : (n : ℕ) → 1r · (g ^ n · f ^ n) · 1r ≡ (1r · 1r) · (1r · ((f · g) ^ n)) + path n = useSolver1 _ _ ∙ sym (^-ldist-· f g n) ∙ useSolver2 _ + + + + injectivityLemma : 1r ∈ ⟨f,g⟩ → ∀ (x : R) → x /1ᶠ ≡ 0r → x /1ᵍ ≡ 0r → x ≡ 0r + injectivityLemma 1∈⟨f,g⟩ x x≡0overF x≡0overG = + PT.rec2 (is-set _ _) annihilatorHelper exAnnihilatorF exAnnihilatorG + where + exAnnihilatorF : ∃[ s ∈ Sᶠ ] (fst s · x · 1r ≡ fst s · 0r · 1r) + exAnnihilatorF = isEquivRel→TruncIso locIsEquivRelᶠ _ _ .fun x≡0overF + + exAnnihilatorG : ∃[ s ∈ Sᵍ ] (fst s · x · 1r ≡ fst s · 0r · 1r) + exAnnihilatorG = isEquivRel→TruncIso locIsEquivRelᵍ _ _ .fun x≡0overG + + annihilatorHelper : Σ[ s ∈ Sᶠ ] (fst s · x · 1r ≡ fst s · 0r · 1r) + → Σ[ s ∈ Sᵍ ] (fst s · x · 1r ≡ fst s · 0r · 1r) + → x ≡ 0r + annihilatorHelper ((u , u∈[fⁿ]) , p) ((v , v∈[gⁿ]) , q) = + PT.rec2 (is-set _ _) exponentHelper u∈[fⁿ] v∈[gⁿ] + where + ux≡0 : u · x ≡ 0r + ux≡0 = sym (·Rid _) ∙ p ∙ cong (_· 1r) (0RightAnnihilates _) ∙ (·Rid _) + + vx≡0 : v · x ≡ 0r + vx≡0 = sym (·Rid _) ∙ q ∙ cong (_· 1r) (0RightAnnihilates _) ∙ (·Rid _) + + exponentHelper : Σ[ n ∈ ℕ ] u ≡ f ^ n + → Σ[ n ∈ ℕ ] v ≡ g ^ n + → x ≡ 0r + exponentHelper (n , u≡fⁿ) (m , v≡gᵐ) = + PT.rec (is-set _ _) Σhelper (GeneratingExponents.lemma R' f g l 1∈⟨f,g⟩) + where + l = max n m + + fˡx≡0 : f ^ l · x ≡ 0r + fˡx≡0 = f ^ l · x ≡⟨ cong (λ k → f ^ k · x) (sym (≤-∸-+-cancel left-≤-max)) ⟩ + f ^ ((l ∸ n) +ℕ n) · x ≡⟨ cong (_· x) (sym (·-of-^-is-^-of-+ _ _ _)) ⟩ + f ^ (l ∸ n) · f ^ n · x ≡⟨ cong (λ y → f ^ (l ∸ n) · y · x) (sym u≡fⁿ) ⟩ + f ^ (l ∸ n) · u · x ≡⟨ sym (·Assoc _ _ _) ⟩ + f ^ (l ∸ n) · (u · x) ≡⟨ cong (f ^ (l ∸ n) ·_) ux≡0 ⟩ + f ^ (l ∸ n) · 0r ≡⟨ 0RightAnnihilates _ ⟩ + 0r ∎ + + gˡx≡0 : g ^ l · x ≡ 0r + gˡx≡0 = g ^ l · x ≡⟨ cong (λ k → g ^ k · x) (sym (≤-∸-+-cancel right-≤-max)) ⟩ + g ^ ((l ∸ m) +ℕ m) · x ≡⟨ cong (_· x) (sym (·-of-^-is-^-of-+ _ _ _)) ⟩ + g ^ (l ∸ m) · g ^ m · x ≡⟨ cong (λ y → g ^ (l ∸ m) · y · x) (sym v≡gᵐ) ⟩ + g ^ (l ∸ m) · v · x ≡⟨ sym (·Assoc _ _ _) ⟩ + g ^ (l ∸ m) · (v · x) ≡⟨ cong (g ^ (l ∸ m) ·_) vx≡0 ⟩ + g ^ (l ∸ m) · 0r ≡⟨ 0RightAnnihilates _ ⟩ + 0r ∎ + + Σhelper : Σ[ α ∈ FinVec R 2 ] 1r ≡ α zero · f ^ l + (α (suc zero) · g ^ l + 0r) + → x ≡ 0r + Σhelper (α , 1≡α₀fˡ+α₁gˡ+0) = path + where + α₀ = α zero + α₁ = α (suc zero) + + 1≡α₀fˡ+α₁gˡ : 1r ≡ α₀ · f ^ l + α₁ · g ^ l + 1≡α₀fˡ+α₁gˡ = 1≡α₀fˡ+α₁gˡ+0 ∙ cong (α₀ · f ^ l +_) (+Rid _) + + path : x ≡ 0r + path = x ≡⟨ sym (·Lid _) ⟩ + 1r · x ≡⟨ cong (_· x) 1≡α₀fˡ+α₁gˡ ⟩ + (α₀ · f ^ l + α₁ · g ^ l) · x ≡⟨ ·Ldist+ _ _ _ ⟩ + α₀ · f ^ l · x + α₁ · g ^ l · x ≡⟨ cong₂ _+_ (sym (·Assoc _ _ _)) + (sym (·Assoc _ _ _)) ⟩ + α₀ · (f ^ l · x) + α₁ · (g ^ l · x) ≡⟨ cong₂ (λ y z → α₀ · y + α₁ · z) + fˡx≡0 gˡx≡0 ⟩ + α₀ · 0r + α₁ · 0r ≡⟨ cong₂ _+_ (0RightAnnihilates _) + (0RightAnnihilates _) ∙ +Rid _ ⟩ + 0r ∎ + + + equalizerLemma : 1r ∈ ⟨f,g⟩ + → ∀ (x : R[1/ f ]) (y : R[1/ g ]) + → χ₁ .fst x ≡ χ₂ .fst y + → ∃![ z ∈ R ] (z /1ᶠ ≡ x) × (z /1ᵍ ≡ y) + equalizerLemma 1∈⟨f,g⟩ = invElPropElim2 (λ _ _ → isPropΠ (λ _ → isProp∃!)) baseCase + where + baseCase : ∀ (x y : R) (n : ℕ) + → fst χ₁ ([ x , f ^ n , ∣ n , refl ∣ ]) ≡ fst χ₂ ([ y , g ^ n , ∣ n , refl ∣ ]) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + baseCase x y n χ₁[x/fⁿ]≡χ₂[y/gⁿ] = PT.rec isProp∃! annihilatorHelper exAnnihilator + where + -- doesn't compute that well but at least it computes... + exAnnihilator : ∃[ s ∈ Sᶠᵍ ] -- s.t. + (fst s · (x · transport refl (g ^ n)) · (1r · transport refl ((f · g) ^ n)) + ≡ fst s · (y · transport refl (f ^ n)) · (1r · transport refl ((f · g) ^ n))) + exAnnihilator = isEquivRel→TruncIso locIsEquivRelᶠᵍ _ _ .fun χ₁[x/fⁿ]≡χ₂[y/gⁿ] + + annihilatorHelper : Σ[ s ∈ Sᶠᵍ ] + (fst s · (x · transport refl (g ^ n)) · (1r · transport refl ((f · g) ^ n)) + ≡ fst s · (y · transport refl (f ^ n)) · (1r · transport refl ((f · g) ^ n))) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + annihilatorHelper ((s , s∈[fgⁿ]) , p) = PT.rec isProp∃! exponentHelper s∈[fgⁿ] + where + sxgⁿ[fg]ⁿ≡syfⁿ[fg]ⁿ : s · x · g ^ n · (f · g) ^ n ≡ s · y · f ^ n · (f · g) ^ n + sxgⁿ[fg]ⁿ≡syfⁿ[fg]ⁿ = + s · x · g ^ n · (f · g) ^ n + + ≡⟨ transpHelper _ _ _ _ ⟩ + + s · x · transport refl (g ^ n) · transport refl ((f · g) ^ n) + + ≡⟨ useSolver _ _ _ _ ⟩ + + s · (x · transport refl (g ^ n)) · (1r · transport refl ((f · g) ^ n)) + + ≡⟨ p ⟩ + + s · (y · transport refl (f ^ n)) · (1r · transport refl ((f · g) ^ n)) + + ≡⟨ sym (useSolver _ _ _ _) ⟩ + + s · y · transport refl (f ^ n) · transport refl ((f · g) ^ n) + + ≡⟨ sym (transpHelper _ _ _ _) ⟩ + + s · y · f ^ n · (f · g) ^ n ∎ + + where + transpHelper : ∀ a b c d → a · b · c · d + ≡ a · b · transport refl c · transport refl d + transpHelper a b c d i = a · b · transportRefl c (~ i) · transportRefl d (~ i) + useSolver : ∀ a b c d → a · b · c · d ≡ a · (b · c) · (1r · d) + useSolver = solve R' + + + exponentHelper : Σ[ m ∈ ℕ ] s ≡ (f · g) ^ m + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + exponentHelper (m , s≡[fg]ᵐ) = + PT.rec isProp∃! Σhelper (GeneratingExponents.lemma R' f g 2n+m 1∈⟨f,g⟩) + where + -- the path we'll actually work with + xgⁿ[fg]ⁿ⁺ᵐ≡yfⁿ[fg]ⁿ⁺ᵐ : x · g ^ n · (f · g) ^ (n +ℕ m) ≡ y · f ^ n · (f · g) ^ (n +ℕ m) + xgⁿ[fg]ⁿ⁺ᵐ≡yfⁿ[fg]ⁿ⁺ᵐ = + x · g ^ n · (f · g) ^ (n +ℕ m) + + ≡⟨ cong (x · (g ^ n) ·_) (sym (·-of-^-is-^-of-+ _ _ _)) ⟩ + + x · g ^ n · ((f · g) ^ n · (f · g) ^ m) + + ≡⟨ useSolver _ _ _ _ ⟩ + + (f · g) ^ m · x · g ^ n · (f · g) ^ n + + ≡⟨ cong (λ a → a · x · g ^ n · (f · g) ^ n) (sym s≡[fg]ᵐ) ⟩ + + s · x · g ^ n · (f · g) ^ n + + ≡⟨ sxgⁿ[fg]ⁿ≡syfⁿ[fg]ⁿ ⟩ + + s · y · f ^ n · (f · g) ^ n + + ≡⟨ cong (λ a → a · y · f ^ n · (f · g) ^ n) s≡[fg]ᵐ ⟩ + + (f · g) ^ m · y · f ^ n · (f · g) ^ n + + ≡⟨ sym (useSolver _ _ _ _) ⟩ + + y · f ^ n · ((f · g) ^ n · (f · g) ^ m) + + ≡⟨ cong (y · (f ^ n) ·_) (·-of-^-is-^-of-+ _ _ _) ⟩ + + y · f ^ n · (f · g) ^ (n +ℕ m) ∎ + + where + useSolver : ∀ a b c d → a · b · (c · d) ≡ d · a · b · c + useSolver = solve R' + + -- critical exponent + 2n+m = n +ℕ (n +ℕ m) + -- extracting information from the fact that R=⟨f,g⟩ + Σhelper : Σ[ α ∈ FinVec R 2 ] 1r ≡ linearCombination R' α (fⁿgⁿVec 2n+m) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + Σhelper (α , linCombi) = uniqueExists z (z/1≡x/fⁿ , z/1≡y/gⁿ) + (λ _ → isProp× (is-set _ _) (is-set _ _)) + uniqueness + where + α₀ = α zero + α₁ = α (suc zero) + + 1≡α₀f²ⁿ⁺ᵐ+α₁g²ⁿ⁺ᵐ : 1r ≡ α₀ · f ^ 2n+m + α₁ · g ^ 2n+m + 1≡α₀f²ⁿ⁺ᵐ+α₁g²ⁿ⁺ᵐ = linCombi ∙ cong (α₀ · f ^ 2n+m +_) (+Rid _) + + -- definition of the element + z = α₀ · x · f ^ (n +ℕ m) + α₁ · y · g ^ (n +ℕ m) + + z/1≡x/fⁿ : (z /1ᶠ) ≡ [ x , f ^ n , ∣ n , refl ∣ ] + z/1≡x/fⁿ = eq/ _ _ ((f ^ (n +ℕ m) , ∣ n +ℕ m , refl ∣) , path) + where + useSolver1 : ∀ x y α₀ α₁ fⁿ⁺ᵐ gⁿ⁺ᵐ fⁿ + → fⁿ⁺ᵐ · (α₀ · x · fⁿ⁺ᵐ + α₁ · y · gⁿ⁺ᵐ) · fⁿ + ≡ fⁿ⁺ᵐ · (α₀ · x · (fⁿ · fⁿ⁺ᵐ)) + α₁ · (y · fⁿ · (fⁿ⁺ᵐ · gⁿ⁺ᵐ)) + useSolver1 = solve R' + + useSolver2 : ∀ x α₀ α₁ fⁿ⁺ᵐ g²ⁿ⁺ᵐ f²ⁿ⁺ᵐ + → fⁿ⁺ᵐ · (α₀ · x · f²ⁿ⁺ᵐ) + α₁ · (x · fⁿ⁺ᵐ · g²ⁿ⁺ᵐ) + ≡ fⁿ⁺ᵐ · x · (α₀ · f²ⁿ⁺ᵐ + α₁ · g²ⁿ⁺ᵐ) + useSolver2 = solve R' + + path : f ^ (n +ℕ m) · z · f ^ n ≡ f ^ (n +ℕ m) · x · 1r + path = + f ^ (n +ℕ m) · z · f ^ n + + ≡⟨ useSolver1 _ _ _ _ _ _ _ ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ n · f ^ (n +ℕ m))) + α₁ · (y · f ^ n · (f ^ (n +ℕ m) · g ^ (n +ℕ m))) + + ≡⟨ cong₂ (λ a b → f ^ (n +ℕ m) · (α₀ · x · a) + α₁ · ((y · f ^ n) · b)) + (·-of-^-is-^-of-+ _ _ _) (sym (^-ldist-· _ _ _)) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (y · f ^ n · (f · g) ^ (n +ℕ m)) + + ≡⟨ cong (λ a → f ^ (n +ℕ m) · (α₀ · x · f ^ 2n+m) + α₁ · a) + (sym xgⁿ[fg]ⁿ⁺ᵐ≡yfⁿ[fg]ⁿ⁺ᵐ) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · g ^ n · (f · g) ^ (n +ℕ m)) + + ≡⟨ cong (λ a → f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · g ^ n · a)) (^-ldist-· _ _ _) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · g ^ n · (f ^ (n +ℕ m) · g ^ (n +ℕ m))) + + ≡⟨ cong (λ a → f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · a) (·CommAssocSwap _ _ _ _) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · f ^ (n +ℕ m) · (g ^ n · g ^ (n +ℕ m))) + + ≡⟨ cong (λ a → f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · f ^ (n +ℕ m) · a)) + (·-of-^-is-^-of-+ _ _ _) ⟩ + + f ^ (n +ℕ m) · (α₀ · x · (f ^ 2n+m)) + α₁ · (x · f ^ (n +ℕ m) · g ^ 2n+m) + + ≡⟨ useSolver2 _ _ _ _ _ _ ⟩ + + f ^ (n +ℕ m) · x · (α₀ · f ^ 2n+m + α₁ · g ^ 2n+m) + + ≡⟨ cong (f ^ (n +ℕ m) · x ·_) (sym 1≡α₀f²ⁿ⁺ᵐ+α₁g²ⁿ⁺ᵐ) ⟩ + + f ^ (n +ℕ m) · x · 1r ∎ + + z/1≡y/gⁿ : (z /1ᵍ) ≡ [ y , g ^ n , ∣ n , refl ∣ ] + z/1≡y/gⁿ = eq/ _ _ ((g ^ (n +ℕ m) , ∣ n +ℕ m , refl ∣) , path) + where + useSolver1 : ∀ x y α₀ α₁ fⁿ⁺ᵐ gⁿ⁺ᵐ gⁿ + → gⁿ⁺ᵐ · (α₀ · x · fⁿ⁺ᵐ + α₁ · y · gⁿ⁺ᵐ) · gⁿ + ≡ α₀ · (x · gⁿ · (fⁿ⁺ᵐ · gⁿ⁺ᵐ)) + gⁿ⁺ᵐ · (α₁ · y · (gⁿ · gⁿ⁺ᵐ)) + useSolver1 = solve R' + + useSolver2 : ∀ y α₀ α₁ gⁿ⁺ᵐ g²ⁿ⁺ᵐ f²ⁿ⁺ᵐ + → α₀ · (y · f²ⁿ⁺ᵐ · gⁿ⁺ᵐ) + gⁿ⁺ᵐ · (α₁ · y · g²ⁿ⁺ᵐ) + ≡ gⁿ⁺ᵐ · y · (α₀ · f²ⁿ⁺ᵐ + α₁ · g²ⁿ⁺ᵐ) + useSolver2 = solve R' + + path : g ^ (n +ℕ m) · z · g ^ n ≡ g ^ (n +ℕ m) · y · 1r + path = + g ^ (n +ℕ m) · z · g ^ n + + ≡⟨ useSolver1 _ _ _ _ _ _ _ ⟩ + + α₀ · (x · g ^ n · (f ^ (n +ℕ m) · g ^ (n +ℕ m))) + g ^ (n +ℕ m) · (α₁ · y · (g ^ n · g ^ (n +ℕ m))) + + ≡⟨ cong₂ (λ a b → α₀ · (x · g ^ n · a) + g ^ (n +ℕ m) · (α₁ · y · b)) + (sym (^-ldist-· _ _ _)) (·-of-^-is-^-of-+ _ _ _) ⟩ + + α₀ · (x · g ^ n · (f · g) ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ cong (λ a → α₀ · a + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m)) + xgⁿ[fg]ⁿ⁺ᵐ≡yfⁿ[fg]ⁿ⁺ᵐ ⟩ + + α₀ · (y · f ^ n · (f · g) ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ cong (λ a → α₀ · (y · f ^ n · a) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m)) (^-ldist-· _ _ _) ⟩ + + α₀ · (y · f ^ n · (f ^ (n +ℕ m) · g ^ (n +ℕ m))) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ cong (λ a → α₀ · a + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m)) (·-assoc2 _ _ _ _) ⟩ + + α₀ · (y · (f ^ n · f ^ (n +ℕ m)) · g ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ cong (λ a → α₀ · (y · a · g ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m)) + (·-of-^-is-^-of-+ _ _ _) ⟩ + + α₀ · (y · f ^ 2n+m · g ^ (n +ℕ m)) + g ^ (n +ℕ m) · (α₁ · y · g ^ 2n+m) + + ≡⟨ useSolver2 _ _ _ _ _ _ ⟩ + + g ^ (n +ℕ m) · y · (α₀ · f ^ 2n+m + α₁ · g ^ 2n+m) + + ≡⟨ cong (g ^ (n +ℕ m) · y ·_) (sym 1≡α₀f²ⁿ⁺ᵐ+α₁g²ⁿ⁺ᵐ) ⟩ + + g ^ (n +ℕ m) · y · 1r ∎ + + + uniqueness : ∀ a → ((a /1ᶠ) ≡ [ x , f ^ n , ∣ n , refl ∣ ]) + × ((a /1ᵍ) ≡ [ y , g ^ n , ∣ n , refl ∣ ]) + → z ≡ a + uniqueness a (a/1≡x/fⁿ , a/1≡y/gⁿ) = equalByDifference _ _ + (injectivityLemma 1∈⟨f,g⟩ (z - a) [z-a]/1≡0overF [z-a]/1≡0overG) + where + [z-a]/1≡0overF : (z - a) /1ᶠ ≡ 0r + [z-a]/1≡0overF = (z - a) /1ᶠ + + ≡⟨ /1ᶠAsCommRingHom .snd .pres+ _ _ ⟩ -- (-a)/1=-(a/1) by refl! + + (z /1ᶠ) - (a /1ᶠ) + + ≡⟨ cong₂ _-_ z/1≡x/fⁿ a/1≡x/fⁿ ⟩ + + [ x , f ^ n , ∣ n , refl ∣ ] - [ x , f ^ n , ∣ n , refl ∣ ] + + ≡⟨ +Rinv ([ x , f ^ n , ∣ n , refl ∣ ]) ⟩ + + 0r ∎ + + [z-a]/1≡0overG : (z - a) /1ᵍ ≡ 0r + [z-a]/1≡0overG = (z - a) /1ᵍ + + ≡⟨ /1ᵍAsCommRingHom .snd .pres+ _ _ ⟩ -- (-a)/1=-(a/1) by refl! + + (z /1ᵍ) - (a /1ᵍ) + + ≡⟨ cong₂ _-_ z/1≡y/gⁿ a/1≡y/gⁿ ⟩ + + [ y , g ^ n , ∣ n , refl ∣ ] - [ y , g ^ n , ∣ n , refl ∣ ] + + ≡⟨ +Rinv ([ y , g ^ n , ∣ n , refl ∣ ]) ⟩ + + 0r ∎ + + + {- + putting everything together with the pullback machinery: + If ⟨f,g⟩ = R then we get a square + + _/1ᶠ + R ----> R[1/f] + ⌟ + _/1ᵍ | | χ₁ + v v + + R[1/g] -> R[1/fg] + χ₂ + -} + + open Category (CommRingsCategory {ℓ}) + open Cospan + + fgCospan : Cospan CommRingsCategory + l fgCospan = R[1/ g ]AsCommRing + m fgCospan = R[1/ (f · g) ]AsCommRing + r fgCospan = R[1/ f ]AsCommRing + s₁ fgCospan = χ₂ + s₂ fgCospan = χ₁ + + -- the commutative square + private + /1χComm : ∀ (x : R) → χ₂ .fst (x /1ᵍ) ≡ χ₁ .fst (x /1ᶠ) + /1χComm x = eq/ _ _ ((1r , powersFormMultClosedSubset (f · g) .containsOne) , refl) + + /1χHomComm : /1ᵍAsCommRingHom ⋆ χ₂ ≡ /1ᶠAsCommRingHom ⋆ χ₁ + /1χHomComm = RingHom≡ (funExt /1χComm) + + fgSquare : 1r ∈ ⟨f,g⟩ + → isPullback _ fgCospan /1ᵍAsCommRingHom /1ᶠAsCommRingHom /1χHomComm + fgSquare 1∈⟨f,g⟩ {d = A} ψ φ ψχ₂≡φχ₁ = (χ , χCoh) , χUniqueness + where + instance + _ = snd A + + applyEqualizerLemma : ∀ a → ∃![ χa ∈ R ] (χa /1ᶠ ≡ fst φ a) × (χa /1ᵍ ≡ fst ψ a) + applyEqualizerLemma a = + equalizerLemma 1∈⟨f,g⟩ (fst φ a) (fst ψ a) (cong (_$ a) (sym ψχ₂≡φχ₁)) + + χ : CommRingHom A R' + fst χ a = applyEqualizerLemma a .fst .fst + snd χ = makeIsRingHom + (cong fst (applyEqualizerLemma 1r .snd (1r , 1Coh))) + (λ x y → cong fst (applyEqualizerLemma (x + y) .snd (_ , +Coh x y))) + (λ x y → cong fst (applyEqualizerLemma (x · y) .snd (_ , ·Coh x y))) + where + open IsRingHom + 1Coh : (1r /1ᶠ ≡ fst φ 1r) × (1r /1ᵍ ≡ fst ψ 1r) + 1Coh = (sym (φ .snd .pres1)) , sym (ψ .snd .pres1) + + +Coh : ∀ x y → ((fst χ x + fst χ y) /1ᶠ ≡ fst φ (x + y)) + × ((fst χ x + fst χ y) /1ᵍ ≡ fst ψ (x + y)) + fst (+Coh x y) = /1ᶠAsCommRingHom .snd .pres+ _ _ + ∙∙ cong₂ _+_ (applyEqualizerLemma x .fst .snd .fst) + (applyEqualizerLemma y .fst .snd .fst) + ∙∙ sym (φ .snd .pres+ x y) + snd (+Coh x y) = /1ᵍAsCommRingHom .snd .pres+ _ _ + ∙∙ cong₂ _+_ (applyEqualizerLemma x .fst .snd .snd) + (applyEqualizerLemma y .fst .snd .snd) + ∙∙ sym (ψ .snd .pres+ x y) + + ·Coh : ∀ x y → ((fst χ x · fst χ y) /1ᶠ ≡ fst φ (x · y)) + × ((fst χ x · fst χ y) /1ᵍ ≡ fst ψ (x · y)) + fst (·Coh x y) = /1ᶠAsCommRingHom .snd .pres· _ _ + ∙∙ cong₂ _·_ (applyEqualizerLemma x .fst .snd .fst) + (applyEqualizerLemma y .fst .snd .fst) + ∙∙ sym (φ .snd .pres· x y) + snd (·Coh x y) = /1ᵍAsCommRingHom .snd .pres· _ _ + ∙∙ cong₂ _·_ (applyEqualizerLemma x .fst .snd .snd) + (applyEqualizerLemma y .fst .snd .snd) + ∙∙ sym (ψ .snd .pres· x y) + + + χCoh : (ψ ≡ χ ⋆ /1ᵍAsCommRingHom) × (φ ≡ χ ⋆ /1ᶠAsCommRingHom) + fst χCoh = RingHom≡ (funExt (λ a → sym (applyEqualizerLemma a .fst .snd .snd))) + snd χCoh = RingHom≡ (funExt (λ a → sym (applyEqualizerLemma a .fst .snd .fst))) + + χUniqueness : (y : Σ[ θ ∈ CommRingHom A R' ] + (ψ ≡ θ ⋆ /1ᵍAsCommRingHom) × (φ ≡ θ ⋆ /1ᶠAsCommRingHom)) + → (χ , χCoh) ≡ y + χUniqueness (θ , θCoh) = Σ≡Prop (λ _ → isProp× (isSetRingHom _ _ _ _) + (isSetRingHom _ _ _ _)) + (RingHom≡ (funExt (λ a → cong fst (applyEqualizerLemma a .snd (θtriple a))))) + where + θtriple : ∀ a → Σ[ x ∈ R ] (x /1ᶠ ≡ fst φ a) × (x /1ᵍ ≡ fst ψ a) + θtriple a = fst θ a , sym (cong (_$ a) (θCoh .snd)) + , sym (cong (_$ a) (θCoh .fst)) + + + -- packaging it all up + open Pullback + fgPullback : 1r ∈ ⟨f,g⟩ → Pullback _ fgCospan + pbOb (fgPullback 1r∈⟨f,g⟩) = _ + pbPr₁ (fgPullback 1r∈⟨f,g⟩) = _ + pbPr₂ (fgPullback 1r∈⟨f,g⟩) = _ + pbCommutes (fgPullback 1r∈⟨f,g⟩) = /1χHomComm + univProp (fgPullback 1r∈⟨f,g⟩) = fgSquare 1r∈⟨f,g⟩ diff --git a/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda b/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda index 2d2e00cf81..b0424fafbd 100644 --- a/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda +++ b/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda @@ -21,7 +21,7 @@ open import Cubical.Functions.Embedding import Cubical.Data.Empty as ⊥ open import Cubical.Data.Bool -open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) open import Cubical.Data.Vec @@ -370,11 +370,10 @@ module _ (R' : CommRing ℓ) (S' : ℙ (fst R')) (SMultClosedSubset : isMultClos kerφ⊆annS : ∀ r → fst φ r ≡ 0a → ∃[ s ∈ S ] (s .fst) · r ≡ 0r surχ : ∀ a → ∃[ x ∈ R × S ] fst φ (x .fst) ·A (fst φ (x .snd .fst) ⁻¹) ⦃ φS⊆Aˣ _ (x .snd .snd) ⦄ ≡ a - S⁻¹RChar : (A' : CommRing ℓ) (φ : CommRingHom R' A') - → PathToS⁻¹R A' φ - → S⁻¹RAsCommRing ≡ A' - S⁻¹RChar A' φ cond = CommRingPath S⁻¹RAsCommRing A' .fst - (S⁻¹R≃A , χ .snd) + S⁻¹RCharEquiv : (A' : CommRing ℓ) (φ : CommRingHom R' A') + → PathToS⁻¹R A' φ + → CommRingEquiv S⁻¹RAsCommRing A' + S⁻¹RCharEquiv A' φ cond = S⁻¹R≃A , χ .snd where open CommRingStr (snd A') renaming ( is-set to Aset ; 0r to 0a ; _·_ to _·A_ ; 1r to 1a ; ·Rid to ·A-rid) @@ -415,3 +414,9 @@ module _ (R' : CommRing ℓ) (S' : ℙ (fst R')) (SMultClosedSubset : isMultClos Surχ : isSurjection (fst χ) Surχ a = PT.rec isPropPropTrunc (λ x → PT.∣ [ x .fst ] , x .snd ∣) (surχ a) + + + S⁻¹RChar : (A' : CommRing ℓ) (φ : CommRingHom R' A') + → PathToS⁻¹R A' φ + → S⁻¹RAsCommRing ≡ A' + S⁻¹RChar A' φ cond = uaCommRing (S⁻¹RCharEquiv A' φ cond) diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index db4cc2df9f..847cb45f36 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -1,18 +1,21 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Algebra.CommRing.Properties where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence +open import Cubical.Foundations.GroupoidLaws hiding (_⁻¹) open import Cubical.Foundations.Transport open import Cubical.Foundations.SIP open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path open import Cubical.Data.Sigma -open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) open import Cubical.Structures.Axioms @@ -24,9 +27,11 @@ open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing.Base +open import Cubical.HITs.PropositionalTruncation + private variable - ℓ : Level + ℓ ℓ' ℓ'' : Level module Units (R' : CommRing ℓ) where open CommRingStr (snd R') @@ -141,21 +146,36 @@ module CommRingHoms where idCommRingHom : (R : CommRing ℓ) → CommRingHom R R idCommRingHom R = idRingHom (CommRing→Ring R) - compCommRingHom : (R S T : CommRing ℓ) + compCommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') (T : CommRing ℓ'') → CommRingHom R S → CommRingHom S T → CommRingHom R T compCommRingHom R S T = compRingHom {R = CommRing→Ring R} {CommRing→Ring S} {CommRing→Ring T} - compIdCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) → compCommRingHom R R S (idCommRingHom R) f ≡ f + _∘cr_ : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} + → CommRingHom S T → CommRingHom R S → CommRingHom R T + g ∘cr f = compCommRingHom _ _ _ f g + + compIdCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) + → compCommRingHom _ _ _ (idCommRingHom R) f ≡ f compIdCommRingHom = compIdRingHom - idCompCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) → compCommRingHom R S S f (idCommRingHom S) ≡ f + idCompCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) + → compCommRingHom _ _ _ f (idCommRingHom S) ≡ f idCompCommRingHom = idCompRingHom - compAssocCommRingHom : {R S T U : CommRing ℓ} (f : CommRingHom R S) (g : CommRingHom S T) (h : CommRingHom T U) → - compCommRingHom R T U (compCommRingHom R S T f g) h ≡ - compCommRingHom R S U f (compCommRingHom S T U g h) + compAssocCommRingHom : {R S T U : CommRing ℓ} + (f : CommRingHom R S) (g : CommRingHom S T) (h : CommRingHom T U) + → compCommRingHom _ _ _ (compCommRingHom _ _ _ f g) h + ≡ compCommRingHom _ _ _ f (compCommRingHom _ _ _ g h) compAssocCommRingHom = compAssocRingHom +module CommRingEquivs where + open RingEquivs + + compCommRingEquiv : {A : CommRing ℓ} {B : CommRing ℓ'} {C : CommRing ℓ''} + → CommRingEquiv A B → CommRingEquiv B C → CommRingEquiv A C + compCommRingEquiv {A = A} {B = B} {C = C} = compRingEquiv {A = CommRing→Ring A} + {B = CommRing→Ring B} + {C = CommRing→Ring C} module CommRingHomTheory {A' B' : CommRing ℓ} (φ : CommRingHom A' B') where open Units A' renaming (Rˣ to Aˣ ; _⁻¹ to _⁻¹ᵃ ; ·-rinv to ·A-rinv ; ·-linv to ·A-linv) @@ -225,6 +245,13 @@ module Exponentiation (R' : CommRing ℓ) where ∙∙ cong (f ^ m ·_) (^-rdist-·ℕ f n m) ∙∙ sym (^-ldist-· f (f ^ n) m) + -- interaction of exponentiation and units + open Units R' + + ^-presUnit : ∀ (f : R) (n : ℕ) → f ∈ Rˣ → f ^ n ∈ Rˣ + ^-presUnit f zero f∈Rˣ = RˣContainsOne + ^-presUnit f (suc n) f∈Rˣ = RˣMultClosed f (f ^ n) ⦃ f∈Rˣ ⦄ ⦃ ^-presUnit f n f∈Rˣ ⦄ + -- like in Ring.Properties we provide helpful lemmas here module CommRingTheory (R' : CommRing ℓ) where @@ -244,3 +271,69 @@ module CommRingTheory (R' : CommRing ℓ) where ·CommAssocSwap x y z w = ·Assoc (x · y) z w ∙∙ cong (_· w) (·CommAssocr x y z) ∙∙ sym (·Assoc (x · z) y w) + + +-- the CommRing version of uaCompEquiv +module CommRingUAFunctoriality where + open CommRingStr + open CommRingEquivs + + CommRing≡ : (A B : CommRing ℓ) → ( + Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] + Σ[ q0 ∈ PathP (λ i → p i) (0r (snd A)) (0r (snd B)) ] + Σ[ q1 ∈ PathP (λ i → p i) (1r (snd A)) (1r (snd B)) ] + Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] + Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] + Σ[ s ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] + PathP (λ i → IsCommRing (q0 i) (q1 i) (r+ i) (r· i) (s i)) (isCommRing (snd A)) (isCommRing (snd B))) + ≃ (A ≡ B) + CommRing≡ A B = isoToEquiv theIso + where + open Iso + theIso : Iso _ _ + fun theIso (p , q0 , q1 , r+ , r· , s , t) i = p i + , commringstr (q0 i) (q1 i) (r+ i) (r· i) (s i) (t i) + inv theIso x = cong ⟨_⟩ x , cong (0r ∘ snd) x , cong (1r ∘ snd) x + , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (isCommRing ∘ snd) x + rightInv theIso _ = refl + leftInv theIso _ = refl + + caracCommRing≡ : {A B : CommRing ℓ} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q + caracCommRing≡ {A = A} {B = B} p q P = + sym (transportTransport⁻ (ua (CommRing≡ A B)) p) + ∙∙ cong (transport (ua (CommRing≡ A B))) helper + ∙∙ transportTransport⁻ (ua (CommRing≡ A B)) q + where + helper : transport (sym (ua (CommRing≡ A B))) p ≡ transport (sym (ua (CommRing≡ A B))) q + helper = Σ≡Prop + (λ _ → isPropΣ + (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) + λ _ → isOfHLevelPathP 1 (isPropIsCommRing _ _ _ _ _) _ _) + (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + + uaCompCommRingEquiv : {A B C : CommRing ℓ} (f : CommRingEquiv A B) (g : CommRingEquiv B C) + → uaCommRing (compCommRingEquiv f g) ≡ uaCommRing f ∙ uaCommRing g + uaCompCommRingEquiv f g = caracCommRing≡ _ _ ( + cong ⟨_⟩ (uaCommRing (compCommRingEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaCommRing f) ∙ cong ⟨_⟩ (uaCommRing g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaCommRing f) (uaCommRing g)) ⟩ + cong ⟨_⟩ (uaCommRing f ∙ uaCommRing g) ∎) + + + +open CommRingEquivs +open CommRingUAFunctoriality +recPT→CommRing : {A : Type ℓ'} (𝓕 : A → CommRing ℓ) + → (σ : ∀ x y → CommRingEquiv (𝓕 x) (𝓕 y)) + → (∀ x y z → σ x z ≡ compCommRingEquiv (σ x y) (σ y z)) + ------------------------------------------------------ + → ∥ A ∥ → CommRing ℓ +recPT→CommRing 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommRing 𝓕 + (3-ConstantCompChar 𝓕 (λ x y → uaCommRing (σ x y)) + λ x y z → sym ( cong uaCommRing (compCoh x y z) + ∙ uaCompCommRingEquiv (σ x y) (σ y z))) diff --git a/Cubical/Algebra/CommRing/RadicalIdeal.agda b/Cubical/Algebra/CommRing/RadicalIdeal.agda index 6d4740cd3b..fc44baba07 100644 --- a/Cubical/Algebra/CommRing/RadicalIdeal.agda +++ b/Cubical/Algebra/CommRing/RadicalIdeal.agda @@ -10,7 +10,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Data.Sigma open import Cubical.Data.Sum hiding (map) open import Cubical.Data.FinData hiding (elim) -open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ ; +-comm to +ℕ-comm ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm ; _choose_ to _ℕchoose_ ; snotz to ℕsnotz) diff --git a/Cubical/Algebra/DistLattice/Base.agda b/Cubical/Algebra/DistLattice/Base.agda index 2115bbbafd..09b45863e4 100644 --- a/Cubical/Algebra/DistLattice/Base.agda +++ b/Cubical/Algebra/DistLattice/Base.agda @@ -226,6 +226,9 @@ DistLattice→Lattice (_ , distlatticestr _ _ _ _ H) = DistLatticeHom : (L : DistLattice ℓ) (M : DistLattice ℓ') → Type (ℓ-max ℓ ℓ') DistLatticeHom L M = LatticeHom (DistLattice→Lattice L) (DistLattice→Lattice M) +idDistLatticeHom : (L : DistLattice ℓ) → DistLatticeHom L L +idDistLatticeHom L = idLatticeHom (DistLattice→Lattice L) + IsDistLatticeEquiv : {A : Type ℓ} {B : Type ℓ'} (L : DistLatticeStr A) (e : A ≃ B) (M : DistLatticeStr B) → Type (ℓ-max ℓ ℓ') IsDistLatticeEquiv L e M = diff --git a/Cubical/Algebra/DistLattice/Basis.agda b/Cubical/Algebra/DistLattice/Basis.agda index 2ef7979572..da44dab48c 100644 --- a/Cubical/Algebra/DistLattice/Basis.agda +++ b/Cubical/Algebra/DistLattice/Basis.agda @@ -59,7 +59,21 @@ module _ (L' : DistLattice ℓ) where constructor isbasis field - contains0 : 0l ∈ S + contains1 : 1l ∈ S ∧lClosed : ∀ (x y : L) → x ∈ S → y ∈ S → x ∧l y ∈ S ⋁Basis : ∀ (x : L) → ∃[ n ∈ ℕ ] Σ[ α ∈ FinVec L n ] (∀ i → α i ∈ S) × (⋁ α ≡ x) + open IsBasis + open SemilatticeStr + Basis→MeetSemilattice : (S : ℙ L) → IsBasis S → Semilattice ℓ + fst (Basis→MeetSemilattice S isBasisS) = Σ[ l ∈ L ] (l ∈ S) + ε (snd (Basis→MeetSemilattice S isBasisS)) = 1l , isBasisS .contains1 + _·_ (snd (Basis→MeetSemilattice S isBasisS)) x y = fst x ∧l fst y + , isBasisS .∧lClosed _ _ (snd x) (snd y) + isSemilattice (snd (Basis→MeetSemilattice S isBasisS)) = makeIsSemilattice + (isSetΣ (isSetDistLattice L') λ _ → isProp→isSet (S _ .snd)) + (λ _ _ _ → Σ≡Prop (λ _ → S _ .snd) (∧lAssoc _ _ _)) + (λ _ → Σ≡Prop (λ _ → S _ .snd) (∧lRid _)) + (λ _ → Σ≡Prop (λ _ → S _ .snd) (∧lLid _)) + (λ _ _ → Σ≡Prop (λ _ → S _ .snd) (∧lComm _ _)) + λ _ → Σ≡Prop (λ _ → S _ .snd) (∧lIdem _) diff --git a/Cubical/Algebra/Group/Exact.agda b/Cubical/Algebra/Group/Exact.agda index c2fa70aff0..245cae92bf 100644 --- a/Cubical/Algebra/Group/Exact.agda +++ b/Cubical/Algebra/Group/Exact.agda @@ -3,13 +3,16 @@ module Cubical.Algebra.Group.Exact where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; map to pMap) open import Cubical.Algebra.Group.GroupPath - open import Cubical.Algebra.Group.Instances.Unit +open import Cubical.Data.Unit + renaming (Unit to UnitType) -- TODO : Define exact sequences -- (perhaps short, finite, ℕ-indexed and ℤ-indexed) @@ -52,3 +55,84 @@ SES→isEquiv {R = R} {G = G} {H = H} = (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) (lexact _ inker) BijectionIso.surj bijIso' x = rexact x refl + +-- exact sequence of 4 groups. Useful for the proof of π₄S³ +record Exact4 {ℓ ℓ' ℓ'' ℓ''' : Level} (G : Group ℓ) + (H : Group ℓ') (L : Group ℓ'') (R : Group ℓ''') + (G→H : GroupHom G H) (H→L : GroupHom H L) (L→R : GroupHom L R) + : Type (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ'''))) where + field + ImG→H⊂KerH→L : (x : fst H) → isInIm G→H x → isInKer H→L x + KerH→L⊂ImG→H : (x : fst H) → isInKer H→L x → isInIm G→H x + + ImH→L⊂KerL→R : (x : fst L) → isInIm H→L x → isInKer L→R x + KerL→R⊂ImH→L : (x : fst L) → isInKer L→R x → isInIm H→L x + +open Exact4 + +extendExact4Surjective : {ℓ ℓ' ℓ'' ℓ''' ℓ'''' : Level} + (G : Group ℓ) (H : Group ℓ') (L : Group ℓ'') (R : Group ℓ''') (S : Group ℓ'''') + (G→H : GroupHom G H) (H→L : GroupHom H L) (L→R : GroupHom L R) (R→S : GroupHom R S) + → isSurjective G→H + → Exact4 H L R S H→L L→R R→S + → Exact4 G L R S (compGroupHom G→H H→L) L→R R→S +ImG→H⊂KerH→L (extendExact4Surjective G H L R S G→H H→L L→R R→S surj ex) x = + pRec (GroupStr.is-set (snd R) _ _) + (uncurry λ g → J (λ x _ → isInKer L→R x) + (ImG→H⊂KerH→L ex (fst H→L (fst G→H g)) + ∣ (fst G→H g) , refl ∣)) +KerH→L⊂ImG→H (extendExact4Surjective G H L R S G→H H→L L→R R→S surj ex) x ker = + pRec squash + (uncurry λ y → J (λ x _ → isInIm (compGroupHom G→H H→L) x) + (pMap (uncurry + (λ y → J (λ y _ → Σ[ g ∈ fst G ] fst H→L (fst G→H g) ≡ H→L .fst y) + (y , refl))) (surj y))) + (KerH→L⊂ImG→H ex x ker) +ImH→L⊂KerL→R (extendExact4Surjective G H L R S G→H H→L L→R R→S surj ex) = + ImH→L⊂KerL→R ex +KerL→R⊂ImH→L (extendExact4Surjective G H L R S G→H H→L L→R R→S surj ex) = + KerL→R⊂ImH→L ex + +-- Useful lemma in the proof of π₄S³≅ℤ +transportExact4 : {ℓ ℓ' ℓ'' : Level} + {G G₂ : Group ℓ} {H H₂ : Group ℓ'} {L L₂ : Group ℓ''} {R : Group₀} + (G≡G₂ : G ≡ G₂) (H≡H₂ : H ≡ H₂) (L≡L₂ : L ≡ L₂) + → Unit ≡ R + → (G→H : GroupHom G H) (G₂→H₂ : GroupHom G₂ H₂) + (H→L : GroupHom H L) (H₂→L₂ : GroupHom H₂ L₂) + (L→R : GroupHom L R) + → Exact4 G H L R G→H H→L L→R + → PathP (λ i → GroupHom (G≡G₂ i) (H≡H₂ i)) G→H G₂→H₂ + → PathP (λ i → GroupHom (H≡H₂ i) (L≡L₂ i)) H→L H₂→L₂ + → Exact4 G₂ H₂ L₂ Unit G₂→H₂ H₂→L₂ (→UnitHom L₂) +transportExact4 {G = G} {G₂ = G₂} {H = H} {H₂ = H₂} {L = L} {L₂ = L₂} {R = R} = + J4 (λ G₂ H₂ L₂ R G≡G₂ H≡H₂ L≡L₂ Unit≡R + → (G→H : GroupHom G H) (G₂→H₂ : GroupHom G₂ H₂) + (H→L : GroupHom H L) (H₂→L₂ : GroupHom H₂ L₂) + (L→R : GroupHom L R) + → Exact4 G H L R G→H H→L L→R + → PathP (λ i → GroupHom (G≡G₂ i) (H≡H₂ i)) G→H G₂→H₂ + → PathP (λ i → GroupHom (H≡H₂ i) (L≡L₂ i)) H→L H₂→L₂ + → Exact4 G₂ H₂ L₂ Unit G₂→H₂ H₂→L₂ (→UnitHom L₂)) + (λ G→H G₂→H₂ H→L H₂→L₂ L→R ex pp1 pp2 + → J4 (λ G₂→H₂ H₂→L₂ (x : UnitType) (y : UnitType) + pp1 pp2 (_ : tt ≡ x) (_ : tt ≡ x) + → Exact4 G H L Unit G₂→H₂ H₂→L₂ (→UnitHom L)) + ex G₂→H₂ H₂→L₂ tt tt pp1 pp2 refl refl ) + G₂ H₂ L₂ R + where + J4 : ∀ {ℓ ℓ₂ ℓ₃ ℓ₄ ℓ'} {A : Type ℓ} + {A₂ : Type ℓ₂} {A₃ : Type ℓ₃} {A₄ : Type ℓ₄} + {x : A} {x₂ : A₂} {x₃ : A₃} {x₄ : A₄} + (B : (y : A) (z : A₂) (w : A₃) (u : A₄) + → x ≡ y → x₂ ≡ z → x₃ ≡ w → x₄ ≡ u → Type ℓ') + → B x x₂ x₃ x₄ refl refl refl refl + → (y : A) (z : A₂) (w : A₃) (u : A₄) + (p : x ≡ y) (q : x₂ ≡ z) (r : x₃ ≡ w) (s : x₄ ≡ u) + → B y z w u p q r s + J4 {x = x} {x₂ = x₂} {x₃ = x₃} {x₄ = x₄} B b y z w u = + J (λ y p → (q : x₂ ≡ z) (r : x₃ ≡ w) (s : x₄ ≡ u) → + B y z w u p q r s) + (J (λ z q → (r : x₃ ≡ w) (s : x₄ ≡ u) → B x z w u refl q r s) + (J (λ w r → (s : x₄ ≡ u) → B x x₂ w u refl refl r s) + (J (λ u s → B x x₂ x₃ u refl refl refl s) b))) diff --git a/Cubical/Algebra/Group/Instances/Bool.agda b/Cubical/Algebra/Group/Instances/Bool.agda index f9adc941df..1e64a8c9fb 100644 --- a/Cubical/Algebra/Group/Instances/Bool.agda +++ b/Cubical/Algebra/Group/Instances/Bool.agda @@ -25,8 +25,7 @@ fst Bool = BoolType (snd Bool GroupStr.· false) false = true (snd Bool GroupStr.· false) true = false (snd Bool GroupStr.· true) y = y -(inv (snd Bool)) false = false -(inv (snd Bool)) true = true +(inv (snd Bool)) x = x is-set (isSemigroup (isMonoid (isGroup (snd Bool)))) = isSetBool assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) false false false = refl assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) false false true = refl diff --git a/Cubical/Algebra/Group/Instances/Int.agda b/Cubical/Algebra/Group/Instances/Int.agda index 03db1b2f07..2878504172 100644 --- a/Cubical/Algebra/Group/Instances/Int.agda +++ b/Cubical/Algebra/Group/Instances/Int.agda @@ -2,8 +2,13 @@ module Cubical.Algebra.Group.Instances.Int where open import Cubical.Foundations.Prelude -open import Cubical.Data.Int renaming (ℤ to ℤType ; _+_ to _+ℤ_ ; _-_ to _-ℤ_; -_ to -ℤ_ ; _·_ to _·ℤ_) +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.Int + renaming (ℤ to ℤType ; _+_ to _+ℤ_ ; _-_ to _-ℤ_; -_ to -ℤ_ ; _·_ to _·ℤ_) open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Properties open GroupStr @@ -19,3 +24,21 @@ isGroup (snd ℤ) = isGroupℤ isGroupℤ = makeIsGroup isSetℤ +Assoc (λ _ → refl) (+Comm 0) (λ x → +Comm x (pos 0 -ℤ x) ∙ minusPlus x 0) (λ x → minusPlus x 0) + +ℤHom : (n : ℤType) → GroupHom ℤ ℤ +fst (ℤHom n) x = n ·ℤ x +snd (ℤHom n) = + makeIsGroupHom λ x y → ·DistR+ n x y + +negEquivℤ : GroupEquiv ℤ ℤ +fst negEquivℤ = + isoToEquiv + (iso (GroupStr.inv (snd ℤ)) + (GroupStr.inv (snd ℤ)) + (GroupTheory.invInv ℤ) + (GroupTheory.invInv ℤ)) +snd negEquivℤ = + makeIsGroupHom λ x y + → +Comm (pos 0) (-ℤ (x +ℤ y)) + ∙ -Dist+ x y + ∙ cong₂ _+ℤ_ (+Comm (-ℤ x) (pos 0)) (+Comm (-ℤ y) (pos 0)) diff --git a/Cubical/Algebra/Group/Instances/IntMod.agda b/Cubical/Algebra/Group/Instances/IntMod.agda index 6639441bcc..1f1f71ea5c 100644 --- a/Cubical/Algebra/Group/Instances/IntMod.agda +++ b/Cubical/Algebra/Group/Instances/IntMod.agda @@ -3,6 +3,7 @@ module Cubical.Algebra.Group.Instances.IntMod where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Monoid.Base @@ -11,14 +12,20 @@ open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Bool open import Cubical.Data.Fin open import Cubical.Data.Fin.Arithmetic -open import Cubical.Data.Nat +open import Cubical.Data.Int renaming (_+_ to _+ℤ_ ; ℤ to ℤType) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Unit +open import Cubical.Data.Nat.Mod open import Cubical.Data.Nat.Order +open import Cubical.Algebra.Group.Instances.Int + renaming (ℤ to ℤGroup) open import Cubical.Algebra.Group.Instances.Unit renaming (Unit to UnitGroup) open import Cubical.Algebra.Group.Instances.Bool renaming (Bool to BoolGroup) open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.Properties open import Cubical.Foundations.HLevels open import Cubical.Data.Sigma @@ -27,7 +34,7 @@ open IsGroup open IsMonoid ℤ/_ : ℕ → Group₀ -ℤ/ zero = UnitGroup +ℤ/ zero = ℤGroup fst (ℤ/ suc n) = Fin (suc n) 1g (snd (ℤ/ suc n)) = 0 GroupStr._·_ (snd (ℤ/ suc n)) = _+ₘ_ @@ -67,3 +74,131 @@ snd Bool≅ℤ/2 = ℤ/2≅Bool : GroupIso (ℤ/ 2) BoolGroup ℤ/2≅Bool = invGroupIso Bool≅ℤ/2 + +-- Definition of the quotient map homomorphism ℤ → ℤ/ (suc n) +-- as a group homomorphism. +ℤ→Fin : (n : ℕ) → ℤType → Fin (suc n) +ℤ→Fin n (pos x) = x mod (suc n) , mod< n x +ℤ→Fin n (negsuc x) = -ₘ (suc x mod suc n , mod< n (suc x)) + +ℤ→Fin-presinv : (n : ℕ) (x : ℤType) → ℤ→Fin n (- x) ≡ -ₘ ℤ→Fin n x +ℤ→Fin-presinv n (pos zero) = + Σ≡Prop (λ _ → m≤n-isProp) ((λ _ → zero) ∙ sym (cong fst help)) + where + help : (-ₘ_ {n = n} 0) ≡ 0 + help = GroupTheory.inv1g (ℤ/ (suc n)) +ℤ→Fin-presinv n (pos (suc x)) = Σ≡Prop (λ _ → m≤n-isProp) refl +ℤ→Fin-presinv n (negsuc x) = + sym (GroupTheory.invInv (ℤ/ (suc n)) _) + + +-ₘ1-id : (n : ℕ) + → Path (Fin (suc n)) (-ₘ (1 mod (suc n) , mod< n 1)) + (n mod (suc n) , mod< n n) +-ₘ1-id zero = refl +-ₘ1-id (suc n) = + cong -ₘ_ (FinPathℕ ((1 mod suc (suc n)) , mod< (suc n) 1) 1 + (modIndBase (suc n) 1 (n , +-comm n 2)) .snd) + ∙ Σ≡Prop (λ _ → m≤n-isProp) + ((+inductionBase (suc n) _ + (λ x _ → ((suc (suc n)) ∸ x) mod (suc (suc n))) λ _ x → x) 1 + (n , (+-comm n 2))) + +suc-ₘ1 : (n y : ℕ) + → ((suc y mod suc n) , mod< n (suc y)) -ₘ (1 mod (suc n) , mod< n 1) + ≡ (y mod suc n , mod< n y) +suc-ₘ1 zero y = + isContr→isProp + (isOfHLevelRetractFromIso 0 (fst ℤ/1≅Unit) isContrUnit) _ _ +suc-ₘ1 (suc n) y = + (λ i → ((suc y mod suc (suc n)) , mod< (suc n) (suc y)) + +ₘ (-ₘ1-id (suc n) i)) + ∙ Σ≡Prop (λ _ → m≤n-isProp) + (cong (_mod (2 +ℕ n)) + (cong (_+ℕ (suc n) mod (2 +ℕ n)) + (mod+mod≡mod (suc (suc n)) 1 y)) + ∙∙ sym (mod+mod≡mod (suc (suc n)) + ((1 mod suc (suc n)) + +ℕ (y mod suc (suc n))) (suc n)) + ∙∙ (mod-rCancel (suc (suc n)) ((1 mod suc (suc n)) + +ℕ (y mod suc (suc n))) (suc n) + ∙ cong (_mod (suc (suc n))) + (cong (_+ℕ (suc n mod suc (suc n))) + (+-comm (1 mod suc (suc n)) (y mod suc (suc n))) + ∙ sym (+-assoc (y mod suc (suc n)) + (1 mod suc (suc n)) (suc n mod suc (suc n)))) + ∙∙ mod-rCancel (suc (suc n)) (y mod suc (suc n)) + ((1 mod suc (suc n)) +ℕ (suc n mod suc (suc n))) + ∙∙ (cong (_mod (2 +ℕ n)) + (cong ((y mod suc (suc n)) +ℕ_) + (sym (mod+mod≡mod (suc (suc n)) 1 (suc n)) + ∙ zero-charac (suc (suc n))) + ∙ +-comm _ 0) + ∙ mod-idempotent y))) + +1-ₘsuc : (n y : ℕ) + → ((1 mod (suc n) , mod< n 1) + +ₘ (-ₘ (((suc y mod suc n) , mod< n (suc y))))) + ≡ -ₘ ((y mod suc n) , mod< n y) +1-ₘsuc n y = + sym (GroupTheory.invInv (ℤ/ (suc n)) _) + ∙ cong -ₘ_ + (GroupTheory.invDistr (ℤ/ (suc n)) + (modInd n 1 , mod< n 1) (-ₘ (modInd n (suc y) , mod< n (suc y))) + ∙ cong (_-ₘ (modInd n 1 , mod< n 1)) + (GroupTheory.invInv (ℤ/ (suc n)) (modInd n (suc y) , mod< n (suc y))) + ∙ suc-ₘ1 n y) + +isHomℤ→Fin : (n : ℕ) → IsGroupHom (snd ℤGroup) (ℤ→Fin n) (snd (ℤ/ (suc n))) +isHomℤ→Fin n = + makeIsGroupHom + λ { (pos x) y → pos+case x y + ; (negsuc x) (pos y) → + cong (ℤ→Fin n) (+Comm (negsuc x) (pos y)) + ∙∙ pos+case y (negsuc x) + ∙∙ +ₘ-comm (ℤ→Fin n (pos y)) (ℤ→Fin n (negsuc x)) + ; (negsuc x) (negsuc y) → + sym (cong (ℤ→Fin n) (-Dist+ (pos (suc x)) (pos (suc y)))) + ∙∙ ℤ→Fin-presinv n (pos (suc x) +ℤ (pos (suc y))) + ∙∙ cong -ₘ_ (pos+case (suc x) (pos (suc y))) + ∙∙ GroupTheory.invDistr (ℤ/ (suc n)) + (modInd n (suc x) + , mod< n (suc x)) (modInd n (suc y) , mod< n (suc y)) + ∙∙ +ₘ-comm (ℤ→Fin n (negsuc y)) (ℤ→Fin n (negsuc x))} + where + +1case : (y : ℤType) → ℤ→Fin n (1 +ℤ y) ≡ ℤ→Fin n 1 +ₘ ℤ→Fin n y + +1case (pos zero) = sym (GroupStr.rid (snd (ℤ/ (suc n))) _) + +1case (pos (suc y)) = + cong (ℤ→Fin n) (+Comm 1 (pos (suc y))) + ∙ Σ≡Prop (λ _ → m≤n-isProp) (mod+mod≡mod (suc n) 1 (suc y)) + +1case (negsuc zero) = + Σ≡Prop (λ _ → m≤n-isProp) refl + ∙ sym (GroupStr.invr (snd (ℤ/ (suc n))) (modInd n 1 , mod< n 1)) + +1case (negsuc (suc y)) = + Σ≡Prop (λ _ → m≤n-isProp) + (cong fst (cong (ℤ→Fin n) (+Comm 1 (negsuc (suc y)))) + ∙∙ cong fst (cong -ₘ_ (refl {x = suc y mod suc n , mod< n (suc y)})) + ∙∙ cong fst (sym (1-ₘsuc n (suc y))) + ∙ λ i → fst ((1 mod (suc n) , mod< n 1) + +ₘ (-ₘ (((suc (suc y) mod suc n) , mod< n (suc (suc y))))))) + + pos+case : (x : ℕ) (y : ℤType) + → ℤ→Fin n (pos x +ℤ y) ≡ ℤ→Fin n (pos x) +ₘ ℤ→Fin n y + pos+case zero y = + cong (ℤ→Fin n) (+Comm 0 y) + ∙ sym (GroupStr.lid (snd (ℤ/ (suc n))) (ℤ→Fin n y)) + pos+case (suc zero) y = +1case y + pos+case (suc (suc x)) y = + cong (ℤ→Fin n) (cong (_+ℤ y) (+Comm (pos (suc x)) 1) + ∙ sym (+Assoc 1 (pos (suc x)) y)) + ∙∙ +1case (pos (suc x) +ℤ y) + ∙∙ (cong ((modInd n 1 , mod< n 1) +ₘ_) (pos+case (suc x) y) + ∙∙ sym (+ₘ-assoc (modInd n 1 , mod< n 1) + (modInd n (suc x) , mod< n (suc x)) (ℤ→Fin n y)) + ∙∙ cong (_+ₘ ℤ→Fin n y) (lem x)) + where + lem : (x : ℕ) + → (modInd n 1 , mod< n 1) +ₘ (modInd n (suc x) , mod< n (suc x)) + ≡ ℤ→Fin n (pos (suc (suc x))) + lem x = + Σ≡Prop (λ _ → m≤n-isProp) (sym (mod+mod≡mod (suc n) 1 (suc x))) diff --git a/Cubical/Algebra/Group/Instances/Unit.agda b/Cubical/Algebra/Group/Instances/Unit.agda index daf64d9308..802b7599d2 100644 --- a/Cubical/Algebra/Group/Instances/Unit.agda +++ b/Cubical/Algebra/Group/Instances/Unit.agda @@ -68,3 +68,7 @@ GroupIsoUnitGroup→isContr : {G : Group ℓ-zero} → GroupIso Unit G → isContr (fst G) GroupIsoUnitGroup→isContr is = isOfHLevelRetractFromIso 0 (invIso (fst is)) isContrUnit + +→UnitHom : ∀ {ℓ} (G : Group ℓ) → GroupHom G Unit +fst (→UnitHom G) _ = tt +snd (→UnitHom G) = makeIsGroupHom λ _ _ → refl diff --git a/Cubical/Algebra/Group/IsomorphismTheorems.agda b/Cubical/Algebra/Group/IsomorphismTheorems.agda index 7b2784a216..02c342db0b 100644 --- a/Cubical/Algebra/Group/IsomorphismTheorems.agda +++ b/Cubical/Algebra/Group/IsomorphismTheorems.agda @@ -12,6 +12,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Structure open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Function open import Cubical.Data.Sigma open import Cubical.HITs.SetQuotients renaming (_/_ to _/s_ ; rec to recS ; elim to elimS) open import Cubical.HITs.PropositionalTruncation @@ -35,20 +36,28 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where open Iso open GroupTheory - private - kerϕ : NormalSubgroup G - kerϕ = kerSubgroup ϕ , isNormalKer ϕ + kerNormalSubgroup : NormalSubgroup G + kerNormalSubgroup = kerSubgroup ϕ , isNormalKer ϕ + + private imϕ : Group ℓ imϕ = imGroup ϕ + -- for completeness: + imNormalSubgroup : ((x y : ⟨ H ⟩) + → GroupStr._·_ (snd H) x y ≡ GroupStr._·_ (snd H) y x) + → NormalSubgroup H + imNormalSubgroup comm = imSubgroup ϕ , isNormalIm ϕ comm + + private module G = GroupStr (snd G) module H = GroupStr (snd H) module imG = GroupStr (snd imϕ) - module kerG = GroupStr (snd (G / kerϕ)) + module kerG = GroupStr (snd (G / kerNormalSubgroup)) module ϕ = IsGroupHom (ϕ .snd) - f1 : ⟨ imϕ ⟩ → ⟨ G / kerϕ ⟩ + f1 : ⟨ imϕ ⟩ → ⟨ G / kerNormalSubgroup ⟩ f1 (x , Hx) = rec→Set ( squash/) (λ { (y , hy) → [ y ]}) (λ { (y , hy) (z , hz) → eq/ y z (rem y z hy hz) }) @@ -62,7 +71,7 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where x H.· H.inv x ≡⟨ H.invr x ⟩ H.1g ∎ - f2 : ⟨ G / kerϕ ⟩ → ⟨ imϕ ⟩ + f2 : ⟨ G / kerNormalSubgroup ⟩ → ⟨ imϕ ⟩ f2 = recS imG.is-set (λ y → ϕ .fst y , ∣ y , refl ∣) (λ x y r → Σ≡Prop (λ _ → squash) (rem x y r)) @@ -78,7 +87,7 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where H.1g H.· ϕ .fst y ≡⟨ H.lid _ ⟩ ϕ .fst y ∎ - f12 : (x : ⟨ G / kerϕ ⟩) → f1 (f2 x) ≡ x + f12 : (x : ⟨ G / kerNormalSubgroup ⟩) → f1 (f2 x) ≡ x f12 = elimProp (λ _ → squash/ _ _) (λ _ → refl) f21 : (x : ⟨ imϕ ⟩) → f2 (f1 x) ≡ x @@ -93,7 +102,7 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where (λ _ _ → kerG.is-set _ _) (λ _ _ → refl) hx hy -- The first isomorphism theorem for groups - isoThm1 : GroupIso imϕ (G / kerϕ) + isoThm1 : GroupIso imϕ (G / kerNormalSubgroup) fun (fst isoThm1) = f1 inv (fst isoThm1) = f2 rightInv (fst isoThm1) = f12 @@ -101,5 +110,19 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where snd isoThm1 = makeIsGroupHom f1-isHom -- The SIP lets us turn the isomorphism theorem into a path - pathThm1 : imϕ ≡ G / kerϕ + pathThm1 : imϕ ≡ G / kerNormalSubgroup pathThm1 = uaGroup (GroupIso→GroupEquiv isoThm1) + + surjImIso : isSurjective ϕ → GroupIso imϕ H + surjImIso surj = + BijectionIso→GroupIso + (bijIso indHom + (uncurry + (λ h → elim (λ _ → isPropΠ (λ _ → imG.is-set _ _)) + (uncurry λ g y + → λ inker → Σ≡Prop (λ _ → squash) inker))) + λ b → map (λ x → (b , ∣ x ∣) , refl) (surj b)) + where + indHom : GroupHom imϕ H + fst indHom = fst + snd indHom = makeIsGroupHom λ _ _ → refl diff --git a/Cubical/Algebra/Group/MorphismProperties.agda b/Cubical/Algebra/Group/MorphismProperties.agda index a20cd38ee6..319fab7307 100644 --- a/Cubical/Algebra/Group/MorphismProperties.agda +++ b/Cubical/Algebra/Group/MorphismProperties.agda @@ -25,7 +25,7 @@ open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms -open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation renaming (map to pMap) private variable @@ -174,6 +174,17 @@ snd (GroupHom≡ {G = G} {H = H} {f = f} {g = g} p i) = p-hom i p-hom : PathP (λ i → IsGroupHom (G .snd) (p i) (H .snd)) (f .snd) (g .snd) p-hom = toPathP (isPropIsGroupHom G H _ _) +-- The composition of surjective maps is surjective +compSurjective : ∀ {ℓ ℓ' ℓ''} {G : Group ℓ} {H : Group ℓ'} {L : Group ℓ''} + → (G→H : GroupHom G H) (H→L : GroupHom H L) + → isSurjective G→H → isSurjective H→L + → isSurjective (compGroupHom G→H H→L) +compSurjective G→H H→L surj1 surj2 l = + rec squash + (λ {(h , p) + → pMap (λ {(g , q) → g , (cong (fst H→L) q ∙ p)}) + (surj1 h)}) + (surj2 l) -- GroupEquiv identity, composition and inversion idGroupEquiv : GroupEquiv G G diff --git a/Cubical/Algebra/Group/Morphisms.agda b/Cubical/Algebra/Group/Morphisms.agda index d22f50034b..38ff0c7d44 100644 --- a/Cubical/Algebra/Group/Morphisms.agda +++ b/Cubical/Algebra/Group/Morphisms.agda @@ -62,6 +62,9 @@ IsGroupEquiv M e N = IsGroupHom M (e .fst) N GroupEquiv : (G : Group ℓ) (H : Group ℓ') → Type (ℓ-max ℓ ℓ') GroupEquiv G H = Σ[ e ∈ (G .fst ≃ H .fst) ] IsGroupEquiv (G .snd) e (H .snd) +groupEquivFun : {G : Group ℓ} {H : Group ℓ'} → GroupEquiv G H → G .fst → H .fst +groupEquivFun e = e .fst .fst + -- Image, kernel, surjective, injective, and bijections open IsGroupHom diff --git a/Cubical/Algebra/Group/QuotientGroup.agda b/Cubical/Algebra/Group/QuotientGroup.agda index 7f2a90efc0..64ffafbf26 100644 --- a/Cubical/Algebra/Group/QuotientGroup.agda +++ b/Cubical/Algebra/Group/QuotientGroup.agda @@ -103,3 +103,35 @@ G / H = asGroup G (H .fst) (H .snd) [_]/G : {G : Group ℓ} {H : NormalSubgroup G} → ⟨ G ⟩ → ⟨ G / H ⟩ [ x ]/G = [ x ] + +-- Quotienting by a trivial subgroup +module _ {G' : Group ℓ} (H' : NormalSubgroup G') + (contrH : (x y : fst G') → _~_ G' (fst H') (snd H') x y → x ≡ y) where + private + -- open isSubgroup (snd H') + open GroupStr (snd G') + open GroupTheory G' + G = fst G' + G/H' = fst (G' / H') + + Code : (g : G) → G/H' → hProp ℓ + Code g = + elim (λ _ → isSetHProp) + (λ a → (g ≡ a) , is-set _ _) + λ a b r → Σ≡Prop (λ _ → isPropIsProp) (cong (g ≡_) (contrH a b r)) + + decode : (g : G) (x : G/H') → [ g ] ≡ x → Code g x .fst + decode g x = J (λ x _ → Code g x .fst) refl + + trivialRel→elimPath : {g h : G} → Path G/H' [ g ] [ h ] → g ≡ h + trivialRel→elimPath {g = g} {h = h} = decode g [ h ] + + trivialRelIso : GroupIso G' (G' / H') + Iso.fun (fst trivialRelIso) g = [ g ] + Iso.inv (fst trivialRelIso) = + rec is-set (λ g → g) contrH + Iso.rightInv (fst trivialRelIso) = + elimProp (λ _ → squash/ _ _) λ _ → refl + Iso.leftInv (fst trivialRelIso) _ = refl + snd trivialRelIso = + makeIsGroupHom λ _ _ → refl diff --git a/Cubical/Algebra/Group/Subgroup.agda b/Cubical/Algebra/Group/Subgroup.agda index d190962ae0..da17604f26 100644 --- a/Cubical/Algebra/Group/Subgroup.agda +++ b/Cubical/Algebra/Group/Subgroup.agda @@ -173,6 +173,19 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where imGroup : Group ℓ imGroup = Subgroup→Group _ imSubgroup + isNormalIm : ((x y : ⟨ H ⟩) → x H.· y ≡ y H.· x) + → isNormal imSubgroup + isNormalIm comm x y = + map λ {(g , p) + → g , + (ϕ .fst g ≡⟨ p ⟩ + y ≡⟨ sym (H.rid y) ⟩ + (y H.· H.1g) ≡⟨ cong (y H.·_) (sym (H.invr x)) ⟩ + (y H.· (x H.· H.inv x)) ≡⟨ H.assoc y x (H.inv x) ⟩ + ((y H.· x) H.· H.inv x) ≡⟨ cong (H._· H.inv x) (comm y x) ⟩ + ((x H.· y) H.· H.inv x) ≡⟨ sym (H.assoc x y (H.inv x)) ⟩ + x H.· y H.· H.inv x ∎ )} + kerSubset : ℙ ⟨ G ⟩ kerSubset x = isInKer ϕ x , isPropIsInKer ϕ x diff --git a/Cubical/Algebra/Group/ZAction.agda b/Cubical/Algebra/Group/ZAction.agda index 0202489277..9967c5082a 100644 --- a/Cubical/Algebra/Group/ZAction.agda +++ b/Cubical/Algebra/Group/ZAction.agda @@ -1,25 +1,45 @@ -- Left ℤ-multiplication on groups and some of its properties + +-- TODO: lots of the content here should be moved elsewhere {-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Algebra.Group.ZAction where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset open import Cubical.Data.Sigma open import Cubical.Data.Int - renaming (_·_ to _*_ ; _+_ to _+ℤ_ ; _-_ to _-ℤ_) -open import Cubical.Data.Nat - hiding (_·_) renaming (_+_ to _+ℕ_) + renaming + (_·_ to _*_ ; _+_ to _+ℤ_ ; _-_ to _-ℤ_ ; pos·pos to pos·) +open import Cubical.Data.Nat renaming (_·_ to _·ℕ_ ; _+_ to _+ℕ_) +open import Cubical.Data.Nat.Mod +open import Cubical.Data.Nat.Order open import Cubical.Data.Empty renaming (rec to ⊥-rec) -open import Cubical.Data.Sum +open import Cubical.Data.Sum renaming (rec to ⊎-rec) +open import Cubical.Data.Unit +open import Cubical.Data.Fin hiding (_/_) +open import Cubical.Data.Fin.Arithmetic + +open import Cubical.HITs.SetQuotients renaming (_/_ to _/s_ ; rec to sRec ; elim to sElim) +open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Unit renaming (Unit to UnitGroup) +open import Cubical.Algebra.Group.Instances.IntMod open import Cubical.Algebra.Group.DirProd +open import Cubical.Algebra.Group.Subgroup +open import Cubical.Algebra.Group.GroupPath +open import Cubical.Algebra.Group.IsomorphismTheorems +open import Cubical.Algebra.Group.QuotientGroup +open import Cubical.Algebra.Group.Exact open import Cubical.Relation.Nullary @@ -30,6 +50,7 @@ private open Iso open GroupStr open IsGroupHom +open Exact4 _ℤ[_]·_ : ℤ → (G : Group ℓ) → fst G → fst G pos zero ℤ[ G' ]· g = 1g (snd G') @@ -158,10 +179,11 @@ GroupHomℤ→ℤPres· e a b = cong (fst e) (ℤ·≡· a b) ∙∙ homPresℤ· e b a ∙∙ sym (ℤ·≡· a (fst e b)) -- Generators in terms of ℤ-multiplication --- Todo : generalise + +-- TODO : generalise and develop theory of cyclic groups gen₁-by : (G : Group ℓ) → (g : fst G) → Type _ gen₁-by G g = (h : fst G) - → Σ[ a ∈ ℤ ] h ≡ (a ℤ[ G ]· g) + → Σ[ a ∈ ℤ ] h ≡ (a ℤ[ G ]· g) gen₂-by : ∀ {ℓ} (G : Group ℓ) → (g₁ g₂ : fst G) → Type _ gen₂-by G g₁ g₂ = @@ -169,27 +191,43 @@ gen₂-by G g₁ g₂ = Iso-pres-gen₁ : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (g : fst G) → gen₁-by G g → (e : GroupIso G H) - → gen₁-by H (Iso.fun (fst e) g) + → gen₁-by H (fun (fst e) g) Iso-pres-gen₁ G H g genG is h = - (fst (genG (Iso.inv (fst is) h))) - , (sym (Iso.rightInv (fst is) h) - ∙∙ cong (Iso.fun (fst is)) (snd (genG (Iso.inv (fst is) h))) - ∙∙ (homPresℤ· (_ , snd is) g (fst (genG (Iso.inv (fst is) h))))) + (fst (genG (inv (fst is) h))) + , (sym (rightInv (fst is) h) + ∙∙ cong (fun (fst is)) (snd (genG (inv (fst is) h))) + ∙∙ (homPresℤ· (_ , snd is) g (fst (genG (inv (fst is) h))))) Iso-pres-gen₂ : (G : Group ℓ) (H : Group ℓ') (g₁ g₂ : fst G) → gen₂-by G g₁ g₂ → (e : GroupIso G H) - → gen₂-by H (Iso.fun (fst e) g₁) (Iso.fun (fst e) g₂) -fst (Iso-pres-gen₂ G H g₁ g₂ genG is h) = genG (Iso.inv (fst is) h) .fst + → gen₂-by H (fun (fst e) g₁) (fun (fst e) g₂) +fst (Iso-pres-gen₂ G H g₁ g₂ genG is h) = genG (inv (fst is) h) .fst snd (Iso-pres-gen₂ G H g₁ g₂ genG is h) = - sym (Iso.rightInv (fst is) h) - ∙∙ cong (fun (fst is)) (snd (genG (Iso.inv (fst is) h))) + sym (rightInv (fst is) h) + ∙∙ cong (fun (fst is)) (snd (genG (inv (fst is) h))) ∙∙ (pres· (snd is) _ _ ∙ cong₂ (_·_ (snd H)) (homPresℤ· (_ , snd is) g₁ (fst (fst (genG (inv (fst is) h))))) (homPresℤ· (_ , snd is) g₂ (snd (fst (genG (inv (fst is) h)))))) - -private +-- TODO: upstream and express using divisibility? +¬1=x·suc-suc : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x * pos (suc (suc n)) +¬1=x·suc-suc n (pos zero) p = snotz (injPos p) +¬1=x·suc-suc n (pos (suc n₁)) p = + snotz (injPos (sym (cong predℤ (snd (intLem₂ n n₁))) ∙ sym (cong predℤ p))) + where + intLem₂ : (n x : ℕ) + → Σ[ a ∈ ℕ ] ((pos (suc x)) * pos (suc (suc n)) ≡ pos (suc (suc a))) + intLem₂ n zero = n , refl + intLem₂ n (suc x) = lem _ _ (intLem₂ n x) + where + lem : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) + → Σ[ a ∈ ℕ ] pos n +ℤ x ≡ pos (suc (suc a)) + lem x n (a , p) = n +ℕ a + , cong (pos n +ℤ_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) + ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) +¬1=x·suc-suc n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ intLem₁ _ _ .snd) + where intLem₁ : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n * pos (suc m)) ≡ negsuc a intLem₁ n zero = n , ·Comm (negsuc n) (pos 1) intLem₁ n (suc m) = lem _ _ .fst , @@ -205,29 +243,12 @@ private (lem (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((lem (suc (suc x))) y)) - intLem₂ : (n x : ℕ) - → Σ[ a ∈ ℕ ] ((pos (suc x)) * pos (suc (suc n)) ≡ pos (suc (suc a))) - intLem₂ n zero = n , refl - intLem₂ n (suc x) = lem _ _ (intLem₂ n x) - where - lem : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) - → Σ[ a ∈ ℕ ] pos n +ℤ x ≡ pos (suc (suc a)) - lem x n (a , p) = n +ℕ a - , cong (pos n +ℤ_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) - ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) - - ¬1=x·suc-suc : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x * pos (suc (suc n)) - ¬1=x·suc-suc n (pos zero) p = snotz (injPos p) - ¬1=x·suc-suc n (pos (suc n₁)) p = - snotz (injPos (sym (cong predℤ (snd (intLem₂ n n₁))) ∙ sym (cong predℤ p))) - ¬1=x·suc-suc n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ intLem₁ _ _ .snd) - GroupEquivℤ-pres1 : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 GroupEquivℤ-pres1 e (pos zero) p = ⊥-rec (snotz (injPos (sym (retEq (fst e) 1) ∙∙ (cong (fst (fst (invGroupEquiv e))) p) - ∙∙ IsGroupHom.pres1 (snd (invGroupEquiv e))))) + ∙∙ pres1 (snd (invGroupEquiv e))))) GroupEquivℤ-pres1 e (pos (suc zero)) p = cong abs p GroupEquivℤ-pres1 e (pos (suc (suc n))) p = ⊥-rec (¬1=x·suc-suc _ _ (h3 ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) @@ -270,7 +291,7 @@ groupEquivPresGen G (ϕeq , ϕhom) x (inl r) (ψeq , ψhom) = groupEquivPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) - ∙ cong abs (IsGroupHom.presinv + ∙ cong abs (presinv (snd (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom))) 1) ∙ absLem _ (GroupEquivℤ-pres1 @@ -313,3 +334,352 @@ snd (gen₂ℤ×ℤ (x , y)) = distrLem x y (negsuc (suc n)) = cong₂ _+''_ (ΣPathP (sym (-lem x) , sym (-lem y))) (distrLem x y (negsuc n)) + +gen₁ℤGroup-⊎ : (g : ℤ) → gen₁-by ℤGroup g → (g ≡ 1) ⊎ (g ≡ -1) +gen₁ℤGroup-⊎ (pos zero) h = ⊥-rec (negsucNotpos 0 0 (h (negsuc 0) .snd ∙ rUnitℤ· ℤGroup _)) +gen₁ℤGroup-⊎ (pos (suc zero)) h = inl refl +gen₁ℤGroup-⊎ (pos (suc (suc n))) h = ⊥-rec (¬1=x·suc-suc n _ (rem (pos 1))) + where + rem : (x : ℤ) → x ≡ fst (h x) * pos (suc (suc n)) + rem x = h x .snd ∙ sym (ℤ·≡· (fst (h x)) (pos (suc (suc n)))) +gen₁ℤGroup-⊎ (negsuc zero) h = inr refl +gen₁ℤGroup-⊎ (negsuc (suc n)) h = ⊥-rec (¬1=x·suc-suc n _ (rem (pos 1) ∙ ℤ·negsuc (fst (h (pos 1))) (suc n) ∙ -DistL· _ _)) + where + rem : (x : ℤ) → x ≡ fst (h x) * negsuc (suc n) + rem x = h x .snd ∙ sym (ℤ·≡· (fst (h x)) (negsuc (suc n))) + +-- Properties of homomorphisms of ℤ wrt generators (should be moved) +module _ (ϕ : GroupHom ℤGroup ℤGroup) where + ℤHomId : fst ϕ 1 ≡ 1 → fst ϕ ≡ idfun _ + ℤHomId p = funExt rem + where + remPos : (x : ℕ) → fst ϕ (pos x) ≡ idfun ℤ (pos x) + remPos zero = pres1 (snd ϕ) + remPos (suc zero) = p + remPos (suc (suc n)) = + pres· (snd ϕ) (pos (suc n)) 1 ∙ cong₂ _+ℤ_ (remPos (suc n)) p + + rem : (x : ℤ) → fst ϕ x ≡ idfun ℤ x + rem (pos n) = remPos n + rem (negsuc zero) = + presinv (snd ϕ) 1 ∙ cong (λ x → 0 -ℤ x) p + rem (negsuc (suc n)) = + (cong (fst ϕ) (sym (+Comm 0 (negsuc (suc n)))) + ∙ presinv (snd ϕ) (pos (suc (suc n)))) + ∙ +Comm 0 _ + ∙ cong (-_) (remPos (suc (suc n))) + + ℤHomId- : fst ϕ -1 ≡ -1 → fst ϕ ≡ idfun _ + ℤHomId- p = ℤHomId (presinv (snd ϕ) (negsuc 0) ∙ sym (minus≡0- _) ∙ cong -_ p) + + ℤHom1- : fst ϕ 1 ≡ -1 → fst ϕ ≡ -_ + ℤHom1- p = funExt rem + where + rem-1 : fst ϕ (negsuc zero) ≡ pos 1 + rem-1 = presinv (snd ϕ) (pos 1) ∙ sym (minus≡0- (fst ϕ 1)) ∙ cong -_ p + + rem : (n : ℤ) → fst ϕ n ≡ - n + rem (pos zero) = pres1 (snd ϕ) + rem (pos (suc zero)) = p + rem (pos (suc (suc n))) = pres· (snd ϕ) (pos (suc n)) (pos 1) ∙ cong₂ _+ℤ_ (rem (pos (suc n))) p + rem (negsuc zero) = rem-1 + rem (negsuc (suc n)) = pres· (snd ϕ) (negsuc n) -1 ∙ cong₂ _+ℤ_ (rem (negsuc n)) rem-1 + + ℤHom-1 : fst ϕ -1 ≡ 1 → fst ϕ ≡ -_ + ℤHom-1 p = ℤHom1- (presinv (snd ϕ) -1 ∙∙ +Comm 0 _ ∙∙ cong -_ p) + + +-- Properties of equivalences of ℤ wrt generators (should be moved) +module _ (ϕ : GroupEquiv ℤGroup ℤGroup) where + ℤEquiv1 : (groupEquivFun ϕ 1 ≡ 1) ⊎ (groupEquivFun ϕ 1 ≡ -1) + ℤEquiv1 = + groupEquivPresGen ℤGroup (compGroupEquiv ϕ (invGroupEquiv ϕ)) + (pos 1) (inl (funExt⁻ (cong fst (invEquiv-is-rinv (ϕ .fst))) (pos 1))) ϕ + + ℤEquivIsIdOr- : (g : ℤ) → (groupEquivFun ϕ g ≡ g) ⊎ (groupEquivFun ϕ g ≡ - g) + ℤEquivIsIdOr- g = + ⊎-rec (λ h → inl (funExt⁻ (ℤHomId (_ , ϕ .snd) h) g)) + (λ h → inr (funExt⁻ (ℤHom1- (_ , ϕ .snd) h) g)) + ℤEquiv1 + + absℤEquiv : (g : ℤ) → abs g ≡ abs (fst (fst ϕ) g) + absℤEquiv g = + ⊎-rec (λ h → sym (cong abs h)) + (λ h → sym (abs- _) ∙ sym (cong abs h)) + (ℤEquivIsIdOr- g) + +-- A few consequences of the above lemmas +characℤ≅ℤ : (e : GroupEquiv ℤGroup ℤGroup) + → (e ≡ idGroupEquiv) ⊎ (e ≡ negEquivℤ) +characℤ≅ℤ e = + ⊎-rec + (λ p → inl (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (Σ≡Prop (λ _ → isPropIsEquiv _) + (funExt λ x → + cong (e .fst .fst) (·Comm 1 x) + ∙ GroupHomℤ→ℤPres· (fst (fst e) , snd e) x 1 + ∙ cong (x *_) p + ∙ ·Comm x 1)))) + (λ p → inr (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (Σ≡Prop (λ _ → isPropIsEquiv _) + (funExt λ x → + cong (e .fst .fst) (·Comm 1 x) + ∙∙ GroupHomℤ→ℤPres· (fst (fst e) , snd e) x 1 + ∙ cong (x *_) p + ∙ ·Comm x -1 + ∙∙ +Comm (- x) 0)))) + (ℤEquiv1 e) + +absGroupEquivℤGroup : {G : Group₀} (ϕ ψ : GroupEquiv ℤGroup G) (g : fst G) + → abs (invEq (fst ϕ) g) ≡ abs (invEq (fst ψ) g) +absGroupEquivℤGroup = + GroupEquivJ (λ G ϕ → (ψ : GroupEquiv ℤGroup G) (g : fst G) + → abs (invEq (fst ϕ) g) ≡ abs (invEq (fst ψ) g)) + (λ ψ → absℤEquiv (invGroupEquiv ψ)) + +GroupEquivℤ-isEquiv : {G : Group₀} + → GroupEquiv ℤGroup G + → (g : fst G) + → gen₁-by G g + → (ϕ : GroupHom G ℤGroup) + → (fst ϕ g ≡ 1) ⊎ (fst ϕ g ≡ -1) + → isEquiv (fst ϕ) +GroupEquivℤ-isEquiv {G = G} = + GroupEquivJ + (λ G _ → (g : fst G) + → gen₁-by G g + → (ϕ : GroupHom G ℤGroup) + → (fst ϕ g ≡ 1) ⊎ (fst ϕ g ≡ -1) + → isEquiv (fst ϕ)) + rem + where + rem : (g : ℤ) + → gen₁-by ℤGroup g + → (ϕ : GroupHom ℤGroup ℤGroup) + → (fst ϕ g ≡ 1) ⊎ (fst ϕ g ≡ -1) + → isEquiv (fst ϕ) + rem g gen ϕ (inl h₁) = + ⊎-rec (λ h₂ → subst isEquiv (sym (ℤHomId ϕ (sym (cong (fst ϕ) h₂) ∙ h₁))) (idIsEquiv _)) + (λ h₂ → subst isEquiv (sym (ℤHom-1 ϕ (sym (cong (fst ϕ) h₂) ∙ h₁))) isEquiv-) + (gen₁ℤGroup-⊎ g gen) + rem g gen ϕ (inr h₁) = + ⊎-rec (λ h₂ → subst isEquiv (sym (ℤHom1- ϕ (sym (cong (fst ϕ) h₂) ∙ h₁))) isEquiv-) + (λ h₂ → subst isEquiv (sym (ℤHomId- ϕ (sym (cong (fst ϕ) h₂) ∙ h₁))) (idIsEquiv _)) + (gen₁ℤGroup-⊎ g gen) + +-- Characterisation of ℤ→ℤ +characGroupHomℤ : (e : GroupHom ℤGroup ℤGroup) → e ≡ ℤHom (fst e (pos 1)) +characGroupHomℤ e = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ { (pos n) → lem n + ; (negsuc n) + → cong (fst e) (+Comm (- (pos (suc n))) 0) + ∙ presinv (snd e) (pos (suc n)) + ∙ (+Comm 0 (- fst e (pos (suc n))) + ∙ cong -_ (lem (suc n))) + ∙ sym (ℤ·negsuc (fst e 1) n)}) + where + lem : (n : ℕ) → fst e (pos n) ≡ fst e 1 * (pos n) + lem zero = pres1 (snd e) ∙ ·Comm 0 (fst e 1) + lem (suc zero) = ·Comm 1 (fst e 1) + lem (suc (suc n)) = + pres· (snd e) (pos (suc n)) 1 + ∙ cong (_+ℤ fst e 1) (lem (suc n)) + ∙ cong (fst e (pos 1) * pos (suc n) +ℤ_) (·Comm 1 (fst e 1)) + ∙ sym (·DistR+ (fst e 1) (pos (suc n)) 1) + +imℤHomSubGroup : (f : GroupHom ℤGroup ℤGroup) → NormalSubgroup ℤGroup +imℤHomSubGroup f = imNormalSubgroup f +Comm + +-- Equivalence between ℤ/ (abs (f 1)) and ℤ/ (im f) using the two different +-- definitions of ℤ quotients. We start with the case f 1 ≥ 0. +module _ (f : GroupHom ℤGroup ℤGroup) where + trivHom→ℤ≅ℤ/im : fst f 1 ≡ 0 + → GroupIso ℤGroup (ℤGroup / imℤHomSubGroup f) + trivHom→ℤ≅ℤ/im q = + trivialRelIso + (imℤHomSubGroup f) + λ x y → Prop.rec (isSetℤ _ _) + λ {(c , p) → + GroupTheory.invUniqueL ℤGroup + {g = x} {h = (GroupStr.inv (snd ℤGroup) y)} + (sym p ∙ (funExt⁻ (cong fst (characGroupHomℤ f ∙ cong ℤHom q)) c)) + ∙ GroupTheory.invInv ℤGroup y} + + ℤHom→ℤ/im≅ℤ/im1 : (n : ℕ) → fst f 1 ≡ pos (suc n) + → BijectionIso (ℤGroup / imℤHomSubGroup f) (ℤ/ (suc n)) + fst (BijectionIso.fun (ℤHom→ℤ/im≅ℤ/im1 n p)) = + sRec isSetFin (ℤ→Fin n) λ a b + → rec (isSetFin _ _) (uncurry λ x q + → cong (ℤ→Fin n) + (cancel-lem a b _ + (sym q + ∙ (funExt⁻ (cong fst (characGroupHomℤ f ∙ cong ℤHom p))) x)) + ∙ pres· (isHomℤ→Fin n) (pos (suc n) * x) b + ∙ cong (_+ₘ ℤ→Fin n b) (lem x) + ∙ GroupStr.lid (snd (ℤ/ (suc n))) (ℤ→Fin n b)) + where + lem : (x : ℤ) → ℤ→Fin n (pos (suc n) * x) ≡ 0 + lem (pos x) = cong (ℤ→Fin n) (sym (pos· (suc n) x)) + ∙ Σ≡Prop (λ _ → m≤n-isProp) + (cong (_mod (suc n)) (·-comm (suc n) x) + ∙ zero-charac-gen (suc n) x) + lem (negsuc x) = + cong (ℤ→Fin n) (pos·negsuc (suc n) x + ∙ cong -_ (sym (pos· (suc n) (suc x)))) + ∙∙ cong -ₘ_ (Σ≡Prop (λ _ → m≤n-isProp) + (cong (_mod (suc n)) (·-comm (suc n) (suc x)) + ∙ zero-charac-gen (suc n) (suc x))) + ∙∙ GroupTheory.inv1g (ℤ/ (suc n)) + + cancel-lem : (a b x : ℤ) → a +ℤ (0 -ℤ b) ≡ x → a ≡ x +ℤ b + cancel-lem a b x p = + sym (minusPlus b a) + ∙ (λ i → (a +ℤ (+Comm (- b) 0 i)) +ℤ b) + ∙ cong (_+ℤ b) p + + snd (BijectionIso.fun (ℤHom→ℤ/im≅ℤ/im1 n p)) = + makeIsGroupHom (elimProp2 (λ _ _ → isSetFin _ _) + (pres· (isHomℤ→Fin n))) + BijectionIso.inj (ℤHom→ℤ/im≅ℤ/im1 n p) = + elimProp (λ _ → isPropΠ λ _ → squash/ _ _) + λ { (pos x) q + → eq/ (pos x) 0 ∣ (pos (quotient x / (suc n))) + , funExt⁻ (cong fst (characGroupHomℤ f ∙ cong ℤHom p)) + (pos (quotient x / (suc n))) + ∙ sym (pos· (suc n) (quotient x / (suc n))) + ∙ cong pos (λ i → q (~ i) .fst +ℕ suc n ·ℕ + (quotient x / suc n)) + ∙ cong pos (≡remainder+quotient (suc n) x) ∣ + ; (negsuc x) q → eq/ (negsuc x) 0 + ∣ (0 +ℤ (- pos (quotient suc x / (suc n)))) + , (presinv (snd f) (pos (quotient suc x / (suc n))) + ∙ +Comm 0 (- _)) + ∙ cong -_ (funExt⁻ (cong fst (characGroupHomℤ f ∙ cong ℤHom p)) + (pos (quotient (suc x) / (suc n)))) + ∙∙ cong -_ (sym (pos· (suc n) (quotient suc x / (suc n))) + ∙ (λ i → pos (fst ((sym (GroupTheory.invInv + (ℤ/ (suc n)) + ((suc x mod suc n) , mod< n (suc x))) + ∙ cong -ₘ_ q + ∙ GroupTheory.inv1g (ℤ/ (suc n))) (~ i)) + +ℕ suc n ·ℕ quotient (suc x) / suc n))) + ∙∙ cong -_ (cong pos (≡remainder+quotient (suc n) (suc x))) ∣} + BijectionIso.surj (ℤHom→ℤ/im≅ℤ/im1 n p) x = + ∣ [ pos (fst x) ] + , (Σ≡Prop (λ _ → m≤n-isProp) (modIndBase n (fst x) (snd x))) ∣ + +-- main result +ℤ/imIso : (f : GroupHom ℤGroup ℤGroup) + → GroupIso (ℤGroup / imℤHomSubGroup f) (ℤ/ abs (fst f 1)) +ℤ/imIso f = helpIso _ refl + where + helpIso : (a : ℤ) + → fst f 1 ≡ a → GroupIso (ℤGroup / imℤHomSubGroup f) (ℤ/ abs a) + helpIso (pos zero) p = invGroupIso (trivHom→ℤ≅ℤ/im f p) + helpIso (pos (suc n)) p = BijectionIso→GroupIso (ℤHom→ℤ/im≅ℤ/im1 f n p) + helpIso (negsuc n) p = + subst (λ x → GroupIso (ℤGroup / x) (ℤ/ abs (negsuc n))) + (sym lem1) + (BijectionIso→GroupIso + (ℤHom→ℤ/im≅ℤ/im1 extendHom n (+Comm (pos 0) (- (fst f (pos 1))) ∙ cong -_ p))) + where + extendHom : GroupHom ℤGroup ℤGroup + extendHom = compGroupHom f (fst (fst negEquivℤ) , snd negEquivℤ) + + lem1 : imℤHomSubGroup f ≡ imℤHomSubGroup extendHom + lem1 = Σ≡Prop (λ _ → isPropIsNormal _) + (Σ≡Prop (λ _ → isPropIsSubgroup _ _) + (funExt λ x → Σ≡Prop (λ _ → isPropIsProp) + (isoToPath + (iso (Prop.map (λ {(x , q) → (pos 0 -ℤ x) + , cong (pos 0 -ℤ_) + (presinv (snd f) x) + ∙ GroupTheory.invInv ℤGroup (fst f x) + ∙ q})) + (Prop.map (λ {(x , q) → (pos 0 -ℤ x) + , (presinv (snd f) x + ∙ q)})) + (λ _ → squash _ _) (λ _ → squash _ _))))) + +-- Goal: given G -ᶠ→ H → L → Unit exact, with G ≅ H ≅ ℤ, we get +-- an iso ℤ/abs (f 1) ≅ H, where f 1 and 1 are viewed as integers +-- via the isomorphisms. We start with the case when G = H = ℤ +module _ (f : GroupHom ℤGroup ℤGroup) (G : Group₀) + (g : GroupHom ℤGroup G) + (ex : Exact4 ℤGroup ℤGroup G UnitGroup f g (→UnitHom G)) where + + private + imf≡kerg : imℤHomSubGroup f ≡ kerNormalSubgroup g + imf≡kerg = + Σ≡Prop (λ _ → isPropIsNormal _) + (Σ≡Prop (λ _ → isPropIsSubgroup _ _) + (funExt λ x → Σ≡Prop (λ _ → isPropIsProp) + (isoToPath + (isProp→Iso + (isPropIsInIm _ _) + (isPropIsInKer _ _) + (ImG→H⊂KerH→L ex x) + (KerH→L⊂ImG→H ex x))))) + + ℤ/im≅ℤ/ker : GroupIso (ℤGroup / imℤHomSubGroup f) (ℤGroup / kerNormalSubgroup g) + ℤ/im≅ℤ/ker = + GroupEquiv→GroupIso (invEq (GroupPath _ _) (cong (ℤGroup /_) imf≡kerg)) + + GroupIsoℤ/abs : GroupIso (ℤ/ abs (fst f (pos 1))) G + GroupIsoℤ/abs = + compGroupIso + (invGroupIso (ℤ/imIso f)) + (compGroupIso + ℤ/im≅ℤ/ker + (compGroupIso + (invGroupIso (isoThm1 g)) + (surjImIso g λ a → KerL→R⊂ImH→L ex a refl))) + +-- The general case +GroupEquivℤ/abs-gen : (G H L : Group₀) + → (e : GroupEquiv ℤGroup G) + → (r : GroupEquiv ℤGroup H) + → (f : GroupHom G H) (g : GroupHom H L) + → Exact4 G H L UnitGroup f g (→UnitHom L) + → GroupEquiv (ℤ/ abs (invEq (fst r) (fst f (fst (fst e) 1)))) L +GroupEquivℤ/abs-gen G H L = + GroupEquivJ (λ G e + → (r : GroupEquiv ℤGroup H) + → (f : GroupHom G H) (g : GroupHom H L) + → Exact4 G H L UnitGroup f g (→UnitHom L) + → GroupEquiv (ℤ/ abs (invEq (fst r) (fst f (fst (fst e) 1)))) L) + (GroupEquivJ (λ H r + → (f : GroupHom ℤGroup H) (g : GroupHom H L) → + Exact4 ℤGroup H L UnitGroup f g (→UnitHom L) → + GroupEquiv + (ℤ/ abs (invEq (fst r) (fst f 1))) L) + λ f g ex → GroupIso→GroupEquiv (GroupIsoℤ/abs f L g ex)) + +-- for type checking reasons, let's also do it with an abstract type +abstract + abstractℤ/_ : ℕ → Group₀ + abstractℤ/_ n = ℤ/ n + + abstractℤ/≡ℤ : abstractℤ/_ ≡ ℤ/_ + abstractℤ/≡ℤ = refl + + abstractℤ/≅ℤ : (n : ℕ) → GroupEquiv (abstractℤ/ n) (ℤ/ n) + abstractℤ/≅ℤ n = idGroupEquiv + +GroupEquiv-abstractℤ/abs-gen : (G H L : Group₀) + → (e : GroupEquiv ℤGroup G) + → (r : GroupEquiv ℤGroup H) + → (f : GroupHom G H) (g : GroupHom H L) + → Exact4 G H L UnitGroup f g (→UnitHom L) + → (n : ℕ) + → abs (invEq (fst r) (fst f (fst (fst e) 1))) ≡ n + → GroupEquiv (abstractℤ/ n) L +GroupEquiv-abstractℤ/abs-gen G H L e r f g ex n p = main + where + abstract + main : GroupEquiv (abstractℤ/ n) L + main = + transport (λ i + → GroupEquiv (abstractℤ/≡ℤ (~ i) (p i)) L) + (GroupEquivℤ/abs-gen G H L e r f g ex) diff --git a/Cubical/Algebra/Lattice/Base.agda b/Cubical/Algebra/Lattice/Base.agda index 2904e8515a..8e3f7424cb 100644 --- a/Cubical/Algebra/Lattice/Base.agda +++ b/Cubical/Algebra/Lattice/Base.agda @@ -166,6 +166,13 @@ unquoteDecl IsLatticeHomIsoΣ = declareRecordIsoΣ IsLatticeHomIsoΣ (quote IsLa LatticeHom : (L : Lattice ℓ) (M : Lattice ℓ') → Type (ℓ-max ℓ ℓ') LatticeHom L M = Σ[ f ∈ (⟨ L ⟩ → ⟨ M ⟩) ] IsLatticeHom (L .snd) f (M .snd) +idLatticeHom : (L : Lattice ℓ) → LatticeHom L L +fst (idLatticeHom L) x = x +IsLatticeHom.pres0 (snd (idLatticeHom L)) = refl +IsLatticeHom.pres1 (snd (idLatticeHom L)) = refl +IsLatticeHom.pres∨l (snd (idLatticeHom L)) x y = refl +IsLatticeHom.pres∧l (snd (idLatticeHom L)) x y = refl + IsLatticeEquiv : {A : Type ℓ} {B : Type ℓ'} (M : LatticeStr A) (e : A ≃ B) (N : LatticeStr B) → Type (ℓ-max ℓ ℓ') IsLatticeEquiv M e N = IsLatticeHom M (e .fst) N diff --git a/Cubical/Algebra/NatSolver/HornerForms.agda b/Cubical/Algebra/NatSolver/HornerForms.agda index 50c4ae94e3..1c63ca1bc8 100644 --- a/Cubical/Algebra/NatSolver/HornerForms.agda +++ b/Cubical/Algebra/NatSolver/HornerForms.agda @@ -3,7 +3,7 @@ module Cubical.Algebra.NatSolver.HornerForms where open import Cubical.Foundations.Prelude -open import Cubical.Data.Nat +open import Cubical.Data.Nat hiding (isZero) open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.Data.Bool using (Bool; true; false; if_then_else_) diff --git a/Cubical/Algebra/Polynomials.agda b/Cubical/Algebra/Polynomials.agda new file mode 100644 index 0000000000..6f5700c844 --- /dev/null +++ b/Cubical/Algebra/Polynomials.agda @@ -0,0 +1,766 @@ +{-A +Polynomials over commutative rings +================================== +-} +{-# OPTIONS --safe #-} + +---------------------------------- + +module Cubical.Algebra.Polynomials where + +open import Cubical.HITs.PropositionalTruncation +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat renaming (_+_ to _Nat+_; _·_ to _Nat·_) hiding (·-comm) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty.Base renaming (rec to ⊥rec ) +open import Cubical.Data.Bool + +open import Cubical.Algebra.Group hiding (Bool) +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing + +------------------------------------------------------------------------------------ + +private + variable + ℓ ℓ' : Level + A : Type ℓ + +module PolyMod (R' : CommRing ℓ) where + private + R = fst R' + open CommRingStr (snd R') public + +------------------------------------------------------------------------------------------- +-- First definition of a polynomial. +-- A polynomial a₁ + a₂x + ... + aⱼxʲ of degree j is represented as a list [a₁, a₂, ...,aⱼ] +-- modulo trailing zeros. +------------------------------------------------------------------------------------------- + + data Poly : Type ℓ where + [] : Poly + _∷_ : (a : R) → (p : Poly) → Poly + drop0 : 0r ∷ [] ≡ [] + + infixr 5 _∷_ + + + module Elim (B : Poly → Type ℓ') + ([]* : B []) + (cons* : (r : R) (p : Poly) (b : B p) → B (r ∷ p)) + (drop0* : PathP (λ i → B (drop0 i)) (cons* 0r [] []*) []*) where + + f : (p : Poly) → B p + f [] = []* + f (x ∷ p) = cons* x p (f p) + f (drop0 i) = drop0* i + + + -- Given a proposition (as type) ϕ ranging over polynomials, we prove it by: + -- ElimProp.f ϕ ⌜proof for base case []⌝ ⌜proof for induction case a ∷ p⌝ + -- ⌜proof that ϕ actually is a proposition over the domain of polynomials⌝ + module ElimProp (B : Poly → Type ℓ') + ([]* : B []) + (cons* : (r : R) (p : Poly) (b : B p) → B (r ∷ p)) + (BProp : {p : Poly} → isProp (B p)) where + f : (p : Poly) → B p + f = Elim.f B []* cons* (toPathP (BProp (transport (λ i → B (drop0 i)) (cons* 0r [] []*)) []*)) + + + module Rec (B : Type ℓ') + ([]* : B) + (cons* : R → B → B) + (drop0* : cons* 0r []* ≡ []*) + (Bset : isSet B) where + f : Poly → B + f = Elim.f (λ _ → B) []* (λ r p b → cons* r b) drop0* + + + module RecPoly ([]* : Poly) (cons* : R → Poly → Poly) (drop0* : cons* 0r []* ≡ []*) where + f : Poly → Poly + f [] = []* + f (a ∷ p) = cons* a (f p) + f (drop0 i) = drop0* i + + + +-------------------------------------------------------------------------------------------------- +-- Second definition of a polynomial. The purpose of this second definition is to +-- facilitate the proof that the first definition is a set. The two definitions are +-- then shown to be equivalent. +-- A polynomial a₀ + a₁x + ... + aⱼxʲ of degree j is represented as a function f +-- such that for i ∈ ℕ we have f(i) = aᵢ if i ≤ j, and 0 for i > j +-------------------------------------------------------------------------------------------------- + + PolyFun : Type ℓ + PolyFun = Σ[ p ∈ (ℕ → R) ] (∃[ n ∈ ℕ ] ((m : ℕ) → n ≤ m → p m ≡ 0r)) + + + isSetPolyFun : isSet PolyFun + isSetPolyFun = isSetΣSndProp (isSetΠ (λ x → isSetCommRing R')) λ f x y → squash x y + + + --construction of the function that represents the polynomial + Poly→Fun : Poly → (ℕ → R) + Poly→Fun [] = (λ _ → 0r) + Poly→Fun (a ∷ p) = (λ n → if isZero n then a else Poly→Fun p (predℕ n)) + Poly→Fun (drop0 i) = lemma i + where + lemma : (λ n → if isZero n then 0r else 0r) ≡ (λ _ → 0r) + lemma i zero = 0r + lemma i (suc n) = 0r + + + Poly→Prf : (p : Poly) → ∃[ n ∈ ℕ ] ((m : ℕ) → n ≤ m → (Poly→Fun p m ≡ 0r)) + Poly→Prf = ElimProp.f (λ p → ∃[ n ∈ ℕ ] ((m : ℕ) → n ≤ m → (Poly→Fun p m ≡ 0r))) + ∣ 0 , (λ m ineq → refl) ∣ + (λ r p → map ( λ (n , ineq) → (suc n) , + λ { zero h → ⊥rec (znots (sym (≤0→≡0 h))) ; + (suc m) h → ineq m (pred-≤-pred h) + } + ) + ) + squash + + Poly→PolyFun : Poly → PolyFun + Poly→PolyFun p = (Poly→Fun p) , (Poly→Prf p) + + +---------------------------------------------------- +-- Start of code by Anders Mörtberg and Evan Cavallo + + + at0 : (ℕ → R) → R + at0 f = f 0 + + atS : (ℕ → R) → (ℕ → R) + atS f n = f (suc n) + + polyEq : (p p' : Poly) → Poly→Fun p ≡ Poly→Fun p' → p ≡ p' + polyEq [] [] _ = refl + polyEq [] (a ∷ p') α = + sym drop0 ∙∙ cong₂ _∷_ (cong at0 α) (polyEq [] p' (cong atS α)) ∙∙ refl + polyEq [] (drop0 j) α k = + hcomp + (λ l → λ + { (j = i1) → drop0 l + ; (k = i0) → drop0 l + ; (k = i1) → drop0 (j ∧ l) + }) + (is-set 0r 0r (cong at0 α) refl j k ∷ []) + polyEq (a ∷ p) [] α = + refl ∙∙ cong₂ _∷_ (cong at0 α) (polyEq p [] (cong atS α)) ∙∙ drop0 + polyEq (a ∷ p) (a₁ ∷ p') α = + cong₂ _∷_ (cong at0 α) (polyEq p p' (cong atS α)) + polyEq (a ∷ p) (drop0 j) α k = + hcomp -- filler + (λ l → λ + { (k = i0) → a ∷ p + ; (k = i1) → drop0 (j ∧ l) + ; (j = i0) → at0 (α k) ∷ polyEq p [] (cong atS α) k + }) + (at0 (α k) ∷ polyEq p [] (cong atS α) k) + polyEq (drop0 i) [] α k = + hcomp + (λ l → λ + { (i = i1) → drop0 l + ; (k = i0) → drop0 (i ∧ l) + ; (k = i1) → drop0 l + }) + (is-set 0r 0r (cong at0 α) refl i k ∷ []) + polyEq (drop0 i) (a ∷ p') α k = + hcomp -- filler + (λ l → λ + { (k = i0) → drop0 (i ∧ l) + ; (k = i1) → a ∷ p' + ; (i = i0) → at0 (α k) ∷ polyEq [] p' (cong atS α) k + }) + (at0 (α k) ∷ polyEq [] p' (cong atS α) k) + polyEq (drop0 i) (drop0 j) α k = + hcomp + (λ l → λ + { (k = i0) → drop0 (i ∧ l) + ; (k = i1) → drop0 (j ∧ l) + ; (i = i0) (j = i0) → at0 (α k) ∷ [] + ; (i = i1) (j = i1) → drop0 l + }) + (is-set 0r 0r (cong at0 α) refl (i ∧ j) k ∷ []) + + + PolyFun→Poly+ : (q : PolyFun) → Σ[ p ∈ Poly ] Poly→Fun p ≡ q .fst + PolyFun→Poly+ (f , pf) = rec lem (λ x → rem1 f (x .fst) (x .snd) , + funExt (rem2 f (fst x) (snd x)) + ) pf + where + lem : isProp (Σ[ p ∈ Poly ] Poly→Fun p ≡ f) + lem (p , α) (p' , α') = + ΣPathP (polyEq p p' (α ∙ sym α'), isProp→PathP (λ i → (isSetΠ λ _ → is-set) _ _) _ _) + + rem1 : (p : ℕ → R) (n : ℕ) → ((m : ℕ) → n ≤ m → p m ≡ 0r) → Poly + rem1 p zero h = [] + rem1 p (suc n) h = p 0 ∷ rem1 (λ x → p (suc x)) n (λ m x → h (suc m) (suc-≤-suc x)) + + rem2 : (f : ℕ → R) (n : ℕ) → (h : (m : ℕ) → n ≤ m → f m ≡ 0r) (m : ℕ) → + Poly→Fun (rem1 f n h) m ≡ f m + rem2 f zero h m = sym (h m zero-≤) + rem2 f (suc n) h zero = refl + rem2 f (suc n) h (suc m) = rem2 (λ x → f (suc x)) n (λ k p → h (suc k) (suc-≤-suc p)) m + + PolyFun→Poly : PolyFun → Poly + PolyFun→Poly q = PolyFun→Poly+ q .fst + + PolyFun→Poly→PolyFun : (p : Poly) → PolyFun→Poly (Poly→PolyFun p) ≡ p + PolyFun→Poly→PolyFun p = polyEq _ _ (PolyFun→Poly+ (Poly→PolyFun p) .snd) + + + +--End of code by Mörtberg and Cavallo +------------------------------------- + + isSetPoly : isSet Poly + isSetPoly = isSetRetract Poly→PolyFun + PolyFun→Poly + PolyFun→Poly→PolyFun + isSetPolyFun + + +------------------------------------------------- +-- The rest of the file uses the first definition +------------------------------------------------- + + open CommRingTheory R' + open RingTheory (CommRing→Ring R') + open GroupTheory (Ring→Group (CommRing→Ring R')) + + + pattern [_] x = x ∷ [] + + +--------------------------------------- +-- Definition +-- Identity for addition of polynomials +--------------------------------------- + 0P : Poly + 0P = [] + + + --ReplicatePoly(n,p) returns 0 ∷ 0 ∷ ... ∷ [] (n zeros) + ReplicatePoly0 : (n : ℕ) → Poly + ReplicatePoly0 zero = 0P + ReplicatePoly0 (suc n) = 0r ∷ ReplicatePoly0 n + + + --The empty polynomial has multiple equal representations on the form 0 + 0x + 0 x² + ... + replicatePoly0Is0P : ∀ (n : ℕ) → ReplicatePoly0 n ≡ 0P + replicatePoly0Is0P zero = refl + replicatePoly0Is0P (suc n) = (cong (0r ∷_) (replicatePoly0Is0P n)) ∙ drop0 + + +----------------------------- +-- Definition +-- subtraction of polynomials +----------------------------- + Poly- : Poly → Poly + Poly- [] = [] + Poly- (a ∷ p) = (- a) ∷ (Poly- p) + Poly- (drop0 i) = (cong (_∷ []) (inv1g) ∙ drop0) i + + -- Double negation (of subtraction of polynomials) is the identity mapping + Poly-Poly- : (p : Poly) → Poly- (Poly- p) ≡ p + Poly-Poly- = ElimProp.f (λ x → Poly- (Poly- x) ≡ x) + refl + (λ a p e → cong (_∷ (Poly- (Poly- p))) + (-Idempotent a) ∙ cong (a ∷_ ) (e)) + (isSetPoly _ _) + + +--------------------------- +-- Definition +-- addition for polynomials +--------------------------- + _Poly+_ : Poly → Poly → Poly + p Poly+ [] = p + [] Poly+ (drop0 i) = drop0 i + [] Poly+ (b ∷ q) = b ∷ q + (a ∷ p) Poly+ (b ∷ q) = (a + b) ∷ (p Poly+ q) + (a ∷ p) Poly+ (drop0 i) = +Rid a i ∷ p + (drop0 i) Poly+ (a ∷ q) = lem q i where + lem : ∀ q → (0r + a) ∷ ([] Poly+ q) ≡ a ∷ q + lem = ElimProp.f (λ q → (0r + a) ∷ ([] Poly+ q) ≡ a ∷ q) + (λ i → (+Lid a i ∷ [])) + (λ r p _ → λ i → +Lid a i ∷ r ∷ p ) + (isSetPoly _ _) + (drop0 i) Poly+ (drop0 j) = isSet→isSet' isSetPoly (cong ([_] ) (+Rid 0r)) drop0 + (cong ([_] ) (+Lid 0r)) drop0 i j + + + -- [] is the left identity for Poly+ + Poly+Lid : ∀ p → ([] Poly+ p ≡ p) + Poly+Lid = ElimProp.f (λ p → ([] Poly+ p ≡ p) ) + refl + (λ r p prf → refl) + (λ x y → isSetPoly _ _ x y) + + + + -- [] is the right identity for Poly+ + Poly+Rid : ∀ p → (p Poly+ [] ≡ p) + Poly+Rid p = refl + + + + --Poly+ is Associative + Poly+Assoc : ∀ p q r → p Poly+ (q Poly+ r) ≡ (p Poly+ q) Poly+ r + Poly+Assoc = + ElimProp.f (λ p → (∀ q r → p Poly+ (q Poly+ r) ≡ (p Poly+ q) Poly+ r)) + (λ q r → Poly+Lid (q Poly+ r) ∙ cong (_Poly+ r) (sym (Poly+Lid q))) + (λ a p prf → ElimProp.f ((λ q → ∀ r → ((a ∷ p) Poly+ (q Poly+ r)) ≡ + (((a ∷ p) Poly+ q) Poly+ r))) + (λ r → cong ((a ∷ p) Poly+_) (Poly+Lid r)) + (λ b q prf2 → + ElimProp.f + (λ r → ((a ∷ p) Poly+ ((b ∷ q) Poly+ r)) ≡ + ((a + b ∷ (p Poly+ q)) Poly+ r)) + refl + (λ c r prfp → cong ((a + (b + c))∷_) + (prf q r) ∙ + (cong (_∷ ((p Poly+ q) Poly+ r)) + (+Assoc a b c))) + (isSetPoly _ _)) + λ x y i r → isSetPoly (x r i0) (y r i1) (x r) (y r) i) + λ x y i q r → isSetPoly _ _ (x q r) (y q r) i + + + + -- for any polynomial, p, the additive inverse is given by Poly- p + Poly+Inverses : ∀ p → p Poly+ (Poly- p) ≡ [] + Poly+Inverses = ElimProp.f ( λ p → p Poly+ (Poly- p) ≡ []) + refl --(Poly+Lid (Poly- [])) + (λ r p prf → cong (r + - r ∷_) prf ∙ + (cong (_∷ []) (+Rinv r) ∙ drop0)) + (isSetPoly _ _) + + + + --Poly+ is commutative + Poly+Comm : ∀ p q → p Poly+ q ≡ q Poly+ p + Poly+Comm = ElimProp.f (λ p → (∀ q → p Poly+ q ≡ q Poly+ p)) + (λ q → Poly+Lid q) + (λ a p prf → ElimProp.f (λ q → ((a ∷ p) Poly+ q) ≡ (q Poly+ (a ∷ p))) + refl + (λ b q prf2 → cong (_∷ (p Poly+ q)) (+Comm a b) ∙ + cong ((b + a) ∷_) (prf q)) + (isSetPoly _ _) + ) + (λ {p} → isPropΠ (λ q → isSetPoly (p Poly+ q) (q Poly+ p))) + +-------------------------------------------------------------- +-- Definition +-- multiplication of a polynomial by a (constant) ring element +-------------------------------------------------------------- + _PolyConst*_ : (R) → Poly → Poly + r PolyConst* [] = [] + r PolyConst* (a ∷ p) = (r · a) ∷ (r PolyConst* p) + r PolyConst* (drop0 i) = lem r i where + lem : ∀ r → [ r · 0r ] ≡ [] + lem = λ r → [ r · 0r ] ≡⟨ cong (_∷ []) (0RightAnnihilates r) ⟩ + [ 0r ] ≡⟨ drop0 ⟩ + [] ∎ + + + -- For any polynomial p we have: 0 _PolyConst*_ p = [] + 0rLeftAnnihilatesPoly : ∀ q → 0r PolyConst* q ≡ [ 0r ] + 0rLeftAnnihilatesPoly = ElimProp.f (λ q → 0r PolyConst* q ≡ [ 0r ]) + (sym drop0) + (λ r p prf → cong ((0r · r) ∷_) prf ∙ + cong (_∷ [ 0r ]) (0LeftAnnihilates r) ∙ + cong (0r ∷_) drop0 ) + λ x y → isSetPoly _ _ x y + + + -- For any polynomial p we have: 1 _PolyConst*_ p = p + PolyConst*Lid : ∀ q → 1r PolyConst* q ≡ q + PolyConst*Lid = ElimProp.f (λ q → 1r PolyConst* q ≡ q ) refl + (λ a p prf → cong (_∷ (1r PolyConst* p)) (·Lid a) ∙ + cong (a ∷_) (prf) ) + λ x y → isSetPoly _ _ x y + + +-------------------------------- +-- Definition +-- Multiplication of polynomials +-------------------------------- + _Poly*_ : Poly → Poly → Poly + [] Poly* q = [] + (a ∷ p) Poly* q = (a PolyConst* q) Poly+ (0r ∷ (p Poly* q)) + (drop0 i) Poly* q = lem q i where + lem : ∀ q → (0r PolyConst* q) Poly+ [ 0r ] ≡ [] + lem = λ q → ((0r PolyConst* q) Poly+ [ 0r ]) ≡⟨ cong ( _Poly+ [ 0r ] ) (0rLeftAnnihilatesPoly q)⟩ + ([ 0r ] Poly+ [ 0r ]) ≡⟨ cong (_∷ []) 0Idempotent ∙ drop0 ⟩ + [] ∎ + + +-------------------- +--Definition +--Identity for Poly* +-------------------- + 1P : Poly + 1P = [ 1r ] + + + -- For any polynomial p we have: p Poly* [] = [] + 0PRightAnnihilates : ∀ q → 0P Poly* q ≡ 0P + 0PRightAnnihilates = ElimProp.f (λ q → 0P Poly* q ≡ 0P) + refl + (λ r p prf → prf) + λ x y → isSetPoly _ _ x y + + + -- For any polynomial p we have: [] Poly* p = [] + 0PLeftAnnihilates : ∀ p → p Poly* 0P ≡ 0P + 0PLeftAnnihilates = ElimProp.f (λ p → p Poly* 0P ≡ 0P ) + refl + (λ r p prf → cong (0r ∷_) prf ∙ drop0) + λ x y → isSetPoly _ _ x y + + + -- For any polynomial p we have: p Poly* [ 1r ] = p + Poly*Lid : ∀ q → 1P Poly* q ≡ q + Poly*Lid = + ElimProp.f (λ q → 1P Poly* q ≡ q) + drop0 + (λ r p prf → lemma r p) + (λ x y → isSetPoly _ _ x y) + where + lemma : ∀ r p → 1r · r + 0r ∷ (1r PolyConst* p) ≡ r ∷ p + lemma = + λ r p → 1r · r + 0r ∷ (1r PolyConst* p) ≡⟨ cong (_∷ (1r PolyConst* p) ) + (+Rid (1r · r)) ⟩ + 1r · r ∷ (1r PolyConst* p) ≡⟨ cong (_∷ 1r PolyConst* p) (·Lid r) ⟩ + r ∷ (1r PolyConst* p) ≡⟨ cong (r ∷_) (PolyConst*Lid p) ⟩ + r ∷ p ∎ + + + -- Distribution of indeterminate: (p + q)x = px + qx + XLDistrPoly+ : ∀ p q → (0r ∷ (p Poly+ q)) ≡ ((0r ∷ p) Poly+ (0r ∷ q)) + XLDistrPoly+ = + ElimProp.f (λ p → ∀ q → (0r ∷ (p Poly+ q)) ≡ ((0r ∷ p) Poly+ (0r ∷ q)) ) + (λ q → (cong (0r ∷_) (Poly+Lid q)) ∙ + cong (0r ∷_) (sym (Poly+Lid q)) ∙ + sym (cong (_∷ [] Poly+ q) (+Lid 0r))) + (λ a p prf → ElimProp.f (λ q → 0r ∷ ((a ∷ p) Poly+ q) ≡ + ((0r ∷ a ∷ p) Poly+ (0r ∷ q))) + (cong (_∷ a ∷ p ) (sym (+Lid 0r))) + (λ b q prf2 → cong (_∷ a + b ∷ (p Poly+ q)) (sym (+Lid 0r))) + (λ x y i → isSetPoly (x i0) (x i1) x y i)) + (λ x y i q → isSetPoly (x q i0) (x q i1) (x q) (y q) i) + + + -- Distribution of a constant ring element over added polynomials p, q: a (p + q) = ap + aq + PolyConst*LDistrPoly+ : ∀ a p q → a PolyConst* (p Poly+ q) ≡ + (a PolyConst* p) Poly+ (a PolyConst* q) + PolyConst*LDistrPoly+ = + λ a → ElimProp.f (λ p → ∀ q → a PolyConst* (p Poly+ q) ≡ + (a PolyConst* p) Poly+ (a PolyConst* q)) + (λ q → cong (a PolyConst*_) (Poly+Lid q) ∙ + (sym (Poly+Lid (a PolyConst* q)))) + (λ b p prf → ElimProp.f (λ q → (a PolyConst* ((b ∷ p) Poly+ q)) ≡ + (a PolyConst* (b ∷ p)) Poly+ (a PolyConst* q)) + refl + (λ c q prf2 → cong (_∷ (a PolyConst* (p Poly+ q))) + (·Rdist+ a b c) ∙ + cong (a · b + a · c ∷_) (prf q)) + (isSetPoly _ _)) + (λ x y i q → isSetPoly (x q i0) (x q i1) (x q) (y q) i) + + + + --Poly* left distributes over Poly+ + Poly*LDistrPoly+ : ∀ p q r → p Poly* (q Poly+ r) ≡ (p Poly* q) Poly+ (p Poly* r) + Poly*LDistrPoly+ = + ElimProp.f + (λ p → ∀ q r → p Poly* (q Poly+ r) ≡ (p Poly* q) Poly+ (p Poly* r)) + (λ _ _ → refl) + (λ a p prf q r → ((a PolyConst* (q Poly+ r)) Poly+ + (0r ∷(p Poly*(q Poly+ r)))) ≡⟨ + cong (_Poly+ (0r ∷ (p Poly* (q Poly+ r)))) + (PolyConst*LDistrPoly+ a q r) + ⟩ + (((a PolyConst* q) Poly+ (a PolyConst* r)) Poly+ + (0r ∷ (p Poly* (q Poly+ r)))) ≡⟨ + cong (((a PolyConst* q) Poly+ (a PolyConst* r)) Poly+_) + (cong (0r ∷_) (prf q r)) + ⟩ + (((a PolyConst* q) Poly+ (a PolyConst* r)) Poly+ + (0r ∷ ((p Poly* q) Poly+ (p Poly* r)))) ≡⟨ + cong (((a PolyConst* q) Poly+ + (a PolyConst* r)) Poly+_) + (XLDistrPoly+ (p Poly* q) (p Poly* r)) + ⟩ + (((a PolyConst* q) Poly+ (a PolyConst* r)) Poly+ + ((0r ∷ (p Poly* q)) Poly+ (0r ∷ (p Poly* r)))) ≡⟨ + Poly+Assoc ((a PolyConst* q) Poly+ + (a PolyConst* r)) + (0r ∷ (p Poly* q)) + (0r ∷ (p Poly* r)) + ⟩ + (((a PolyConst* q) Poly+ (a PolyConst* r)) Poly+ + (0r ∷ (p Poly* q))) Poly+ (0r ∷ (p Poly* r)) ≡⟨ cong (_Poly+ (0r ∷ (p Poly* r))) + (sym (Poly+Assoc (a PolyConst* q) + (a PolyConst* r) + (0r ∷ (p Poly* q)))) + ⟩ + (((a PolyConst* q) Poly+ ((a PolyConst* r) Poly+ + (0r ∷ (p Poly* q)))) Poly+ (0r ∷ (p Poly* r))) ≡⟨ + cong (_Poly+ (0r ∷ (p Poly* r))) + (cong ((a PolyConst* q) Poly+_) + (Poly+Comm (a PolyConst* r) + (0r ∷ (p Poly* q)))) + ⟩ + (((a PolyConst* q) Poly+ ((0r ∷ (p Poly* q)) Poly+ + (a PolyConst* r))) Poly+ (0r ∷ (p Poly* r))) ≡⟨ + cong (_Poly+ (0r ∷ (p Poly* r))) + (Poly+Assoc (a PolyConst* q) + (0r ∷ (p Poly* q)) + (a PolyConst* r)) + ⟩ + ((((a PolyConst* q) Poly+ (0r ∷ (p Poly* q))) Poly+ + (a PolyConst* r)) Poly+ (0r ∷ (p Poly* r))) ≡⟨ + sym (Poly+Assoc ((a PolyConst* q) Poly+ + (0r ∷ (p Poly* q))) + ((a PolyConst* r)) + ((0r ∷ (p Poly* r)))) + ⟩ + ((a PolyConst* q) Poly+ (0r ∷ (p Poly* q))) Poly+ + ((a PolyConst* r) Poly+ (0r ∷ (p Poly* r))) ∎) + (λ x y i q r → isSetPoly _ _ (x q r) (y q r) i) + + + -- The constant multiplication of a ring element, a, with a polynomial, p, can be + -- expressed by polynomial multiplication with the singleton polynomial [ a ] + PolyConst*r=Poly*[r] : ∀ a p → a PolyConst* p ≡ p Poly* [ a ] + PolyConst*r=Poly*[r] = + λ a → ElimProp.f (λ p → a PolyConst* p ≡ p Poly* [ a ]) + refl + ( λ r p prf → a · r ∷ (a PolyConst* p) ≡⟨ + cong (a · r ∷_) prf + ⟩ + a · r ∷ (p Poly* [ a ]) ≡⟨ + cong (a · r ∷_) + (sym (Poly+Lid (p Poly* [ a ]))) + ⟩ + a · r ∷ ([] Poly+ (p Poly* [ a ])) ≡⟨ + cong (_∷ ([] Poly+ (p Poly* [ a ]))) + (·Comm a r ) + ⟩ + r · a ∷ ([] Poly+ (p Poly* [ a ])) ≡⟨ + cong (_∷ ([] Poly+ (p Poly* [ a ]))) + (sym (+Rid (r · a))) + ⟩ + r · a + 0r ∷ ([] Poly+ (p Poly* [ a ])) ∎) + ( λ x y i → isSetPoly (x i0) (x i1) x y i) + + + -- Connection between the constant multiplication and the multiplication in the ring + PolyConst*Nested· : ∀ a b p → a PolyConst* (b PolyConst* p) ≡ (a · b) PolyConst* p + PolyConst*Nested· = + λ a b → ElimProp.f (λ p → a PolyConst* (b PolyConst* p) ≡ (a · b) PolyConst* p) + refl + (λ c p prf → cong ((a · (b · c)) ∷_) prf ∙ + cong (_∷ ((a · b) PolyConst* p)) (·Assoc a b c)) + (isSetPoly _ _) + + + -- We can move the indeterminate from left to outside: px * q = (p * q)x + 0r∷LeftAssoc : ∀ p q → (0r ∷ p) Poly* q ≡ 0r ∷ (p Poly* q) + 0r∷LeftAssoc = + ElimProp.f (λ p → ∀ q → (0r ∷ p) Poly* q ≡ 0r ∷ (p Poly* q)) + (λ q → cong (_Poly+ [ 0r ])((cong (_Poly+ []) (0rLeftAnnihilatesPoly q))) ∙ + cong (_∷ []) (+Lid 0r)) + (λ r p b q → cong (_Poly+ (0r ∷ ((r PolyConst* q) Poly+ (0r ∷ (p Poly* q))))) + ((0rLeftAnnihilatesPoly q) ∙ drop0)) + (λ x y i q → isSetPoly _ _ (x q) (y q) i) + + + --Associativity of constant multiplication in relation to polynomial multiplication + PolyConst*AssocPoly* : ∀ a p q → a PolyConst* (p Poly* q) ≡ (a PolyConst* p) Poly* q + PolyConst*AssocPoly* = + λ a → ElimProp.f (λ p → ∀ q → a PolyConst* (p Poly* q) ≡ (a PolyConst* p) Poly* q) + (λ q → refl) + (λ b p prf q → a PolyConst* ((b PolyConst* q) Poly+ + (0r ∷ (p Poly* q))) ≡⟨ + PolyConst*LDistrPoly+ a + (b PolyConst* q) + (0r ∷ (p Poly* q)) + ⟩ + (a PolyConst* (b PolyConst* q)) Poly+ + (a PolyConst* (0r ∷ (p Poly* q))) ≡⟨ + cong (_Poly+ (a · 0r ∷ (a PolyConst* (p Poly* q)))) + (PolyConst*Nested· a b q) + ⟩ + ((a · b) PolyConst* q) Poly+ + (a PolyConst* (0r ∷ (p Poly* q))) ≡⟨ + cong (((a · b) PolyConst* q) Poly+_) + (cong (a · 0r ∷_) + (PolyConst*r=Poly*[r] a + (p Poly* q))) + ⟩ + ((a · b) PolyConst* q) Poly+ + (a · 0r ∷ ((p Poly* q) Poly* [ a ])) ≡⟨ + cong (((a · b) PolyConst* q) Poly+_) + (cong (_∷ ((p Poly* q) Poly* [ a ])) + (0RightAnnihilates a)) ⟩ + ((a · b) PolyConst* q) Poly+ + (0r ∷ ((p Poly* q) Poly* [ a ])) ≡⟨ + cong (((a · b) PolyConst* q) Poly+_) + (cong (0r ∷_) + (sym (PolyConst*r=Poly*[r] a (p Poly* q)))) + ⟩ + ((a · b) PolyConst* q) Poly+ + (0r ∷ (a PolyConst* (p Poly* q))) ≡⟨ + cong (((a · b) PolyConst* q) Poly+_) + (cong (0r ∷_) (prf q)) + ⟩ + ((a · b) PolyConst* q) Poly+ + (0r ∷ ((a PolyConst* p) Poly* q)) ∎) + (λ x y i q → isSetPoly (x q i0) (x q i1) (x q) (y q) i) + + + -- We can move the indeterminate from left to outside: p * qx = (p * q)x + 0r∷RightAssoc : ∀ p q → p Poly* (0r ∷ q) ≡ 0r ∷ (p Poly* q) + 0r∷RightAssoc = + ElimProp.f (λ p → ∀ q → p Poly* (0r ∷ q) ≡ 0r ∷ (p Poly* q)) + (λ q → sym drop0) + (λ a p prf q → ((a ∷ p) Poly* (0r ∷ q)) ≡⟨ + cong ( a · 0r + 0r ∷_) + (cong ((a PolyConst* q) Poly+_ ) + (prf q)) + ⟩ + a · 0r + 0r ∷ ((a PolyConst* q) Poly+ + (0r ∷ (p Poly* q))) ≡⟨ + cong (_∷ ((a PolyConst* q) Poly+ (0r ∷ (p Poly* q)))) + ((+Rid (a · 0r))) + ⟩ + a · 0r ∷ ((a PolyConst* q) Poly+ + (0r ∷ (p Poly* q))) ≡⟨ + cong (_∷ ((a PolyConst* q) Poly+ (0r ∷ (p Poly* q)))) + (0RightAnnihilates a) ⟩ + 0r ∷ ((a PolyConst* q) Poly+ (0r ∷ (p Poly* q))) ∎) + (λ x y i q → isSetPoly (x q i0) (x q i1) (x q) (y q) i) + + + -- We can move the indeterminate around: px * q = p * qx + 0r∷Comm : ∀ p q → (0r ∷ p) Poly* q ≡ p Poly* (0r ∷ q) + 0r∷Comm = ElimProp.f (λ p → ∀ q → (0r ∷ p) Poly* q ≡ p Poly* (0r ∷ q)) + (λ q → (cong ((0r PolyConst* q) Poly+_) drop0) ∙ + 0rLeftAnnihilatesPoly q ∙ + drop0 ) + (λ a p prf q → ((0r ∷ a ∷ p) Poly* q) ≡⟨ 0r∷LeftAssoc (a ∷ p) q ⟩ + 0r ∷ ((a ∷ p) Poly* q) ≡⟨ sym (0r∷RightAssoc (a ∷ p) q) ⟩ + ((a ∷ p) Poly* (0r ∷ q)) ∎ + ) + λ x y i q → isSetPoly (x q i0) (x q i1) (x q) (y q) i + + + + --Poly* is commutative + Poly*Commutative : ∀ p q → p Poly* q ≡ q Poly* p + Poly*Commutative = + ElimProp.f (λ p → ∀ q → p Poly* q ≡ q Poly* p) + (λ q → sym (0PLeftAnnihilates q)) + (λ a p prf q → (a PolyConst* q) Poly+ (0r ∷ (p Poly* q)) ≡⟨ + cong ((a PolyConst* q) Poly+_) + (cong (0r ∷_) (prf q)) ⟩ + ((a PolyConst* q) Poly+ (0r ∷ (q Poly* p))) ≡⟨ + cong ((a PolyConst* q) Poly+_) + (sym (0r∷LeftAssoc q p)) + ⟩ + ((a PolyConst* q) Poly+ ((0r ∷ q) Poly* p)) ≡⟨ + cong (_Poly+ ((0r PolyConst* p) Poly+ (0r ∷ (q Poly* p)))) + (PolyConst*r=Poly*[r] a q) ⟩ + ((q Poly* [ a ]) Poly+ ((0r ∷ q) Poly* p)) ≡⟨ + cong ((q Poly* [ a ]) Poly+_) + (0r∷Comm q p) + ⟩ + ((q Poly* [ a ]) Poly+ (q Poly* (0r ∷ p))) ≡⟨ + sym (Poly*LDistrPoly+ q [ a ] (0r ∷ p)) + ⟩ + (((q Poly* ([ a ] Poly+ (0r ∷ p))))) ≡⟨ + cong (q Poly*_) + (Poly+Comm [ a ] (0r ∷ p)) + ⟩ + ((q Poly* ((0r ∷ p) Poly+ [ a ]))) ≡⟨ + refl + ⟩ + (q Poly* ((0r + a) ∷ p)) ≡⟨ cong (q Poly*_) + (cong (_∷ p) (+Lid a)) + ⟩ + (q Poly* (a ∷ p)) ∎) + (λ x y i q → isSetPoly _ _ (x q ) (y q) i) + + + + --1P is the right identity of Poly*. + Poly*Rid : ∀ p → p Poly* 1P ≡ p + Poly*Rid = λ p → (Poly*Commutative p 1P ∙ Poly*Lid p) + + + --Polynomial multiplication right distributes over polynomial addition. + Poly*RDistrPoly+ : ∀ p q r → (p Poly+ q) Poly* r ≡ (p Poly* r) Poly+ (q Poly* r) + Poly*RDistrPoly+ = λ p q r → sym (Poly*Commutative r (p Poly+ q)) ∙ + Poly*LDistrPoly+ r p q ∙ + cong (_Poly+ (r Poly* q)) (Poly*Commutative r p) ∙ + cong ((p Poly* r) Poly+_) (Poly*Commutative r q) + + + --Polynomial multiplication is associative + Poly*Associative : ∀ p q r → p Poly* (q Poly* r) ≡ (p Poly* q) Poly* r + Poly*Associative = + ElimProp.f (λ p → ∀ q r → p Poly* (q Poly* r) ≡ (p Poly* q) Poly* r ) + (λ _ _ → refl) + (λ a p prf q r → + ((a ∷ p) Poly* (q Poly* r)) ≡⟨ + cong (_Poly+ (0r ∷ (p Poly* (q Poly* r)))) + (PolyConst*AssocPoly* a q r) + ⟩ + (((a PolyConst* q) Poly* r) Poly+ + (0r ∷ (p Poly* (q Poly* r)))) ≡⟨ + sym (cong (((a PolyConst* q) Poly* r) Poly+_) + (cong (_∷ (p Poly* (q Poly* r))) + (+Lid 0r))) + ⟩ + (((a PolyConst* q) Poly* r) Poly+ + (0r + 0r ∷ (p Poly* (q Poly* r)))) ≡⟨ + cong (((a PolyConst* q) Poly* r) Poly+_) + (cong (0r + 0r ∷_) + (sym (Poly+Lid (p Poly* (q Poly* r))))) + ⟩ + (((a PolyConst* q) Poly* r) Poly+ + (0r + 0r ∷ ([] Poly+ (p Poly* (q Poly* r))))) ≡⟨ + cong (((a PolyConst* q) Poly* r) Poly+_) + (cong (0r + 0r ∷_) + (cong ([] Poly+_) + (prf q r))) + ⟩ + (((a PolyConst* q) Poly* r) Poly+ + (0r + 0r ∷ ([] Poly+ ((p Poly* q) Poly* r)))) ≡⟨ + cong (((a PolyConst* q) Poly* r) Poly+_) + (cong (_Poly+ (0r ∷ ((p Poly* q) Poly* r))) + (sym (0rLeftAnnihilatesPoly r))) + ⟩ + (((a PolyConst* q) Poly* r) Poly+ + ((0r PolyConst* r) Poly+ (0r ∷ ((p Poly* q) Poly* r)))) ≡⟨ + sym (Poly*RDistrPoly+ (a PolyConst* q) + (0r ∷ (p Poly* q)) r) + ⟩ + ((((a ∷ p) Poly* q) Poly* r)) ∎) + (λ x y i q r → isSetPoly _ _ (x q r) (y q r) i) + + + +---------------------------------------------------------------------------------------------- +-- An instantiation of Polynomials as a commutative ring can be found in CommRing/Instances -- +---------------------------------------------------------------------------------------------- diff --git a/Cubical/Algebra/Ring/Base.agda b/Cubical/Algebra/Ring/Base.agda index d12f9dddeb..1c7bba2f78 100644 --- a/Cubical/Algebra/Ring/Base.agda +++ b/Cubical/Algebra/Ring/Base.agda @@ -29,7 +29,7 @@ open Iso private variable - ℓ ℓ' : Level + ℓ ℓ' ℓ'' : Level record IsRing {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) : Type ℓ where @@ -166,6 +166,12 @@ _$_ : {R S : Ring ℓ} → (φ : RingHom R S) → (x : ⟨ R ⟩) → ⟨ S ⟩ RingEquiv→RingHom : {A B : Ring ℓ} → RingEquiv A B → RingHom A B RingEquiv→RingHom (e , eIsHom) = e .fst , eIsHom +RingHomIsEquiv→RingEquiv : {A B : Ring ℓ} (f : RingHom A B) + → isEquiv (fst f) → RingEquiv A B +fst (fst (RingHomIsEquiv→RingEquiv f fIsEquiv)) = fst f +snd (fst (RingHomIsEquiv→RingEquiv f fIsEquiv)) = fIsEquiv +snd (RingHomIsEquiv→RingEquiv f fIsEquiv) = snd f + isPropIsRing : {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) → isProp (IsRing 0r 1r _+_ _·_ -_) isPropIsRing 0r 1r _+_ _·_ -_ (isring RG RM RD) (isring SG SM SD) = @@ -192,6 +198,10 @@ isPropIsRingHom R f S = isOfHLevelRetractFromIso 1 IsRingHomIsoΣ isSetRingHom : (R : Ring ℓ) (S : Ring ℓ') → isSet (RingHom R S) isSetRingHom R S = isSetΣSndProp (isSetΠ (λ _ → isSetRing S)) (λ f → isPropIsRingHom (snd R) f (snd S)) +isSetRingEquiv : (A : Ring ℓ) (B : Ring ℓ') → isSet (RingEquiv A B) +isSetRingEquiv A B = isSetΣSndProp (isOfHLevel≃ 2 (isSetRing A) (isSetRing B)) + (λ e → isPropIsRingHom (snd A) (fst e) (snd B)) + RingHomPathP : (R S T : Ring ℓ) (p : S ≡ T) (φ : RingHom R S) (ψ : RingHom R T) → PathP (λ i → R .fst → p i .fst) (φ .fst) (ψ .fst) → PathP (λ i → RingHom R (p i)) φ ψ @@ -222,6 +232,13 @@ RingHom≡ = Σ≡Prop λ f → isPropIsRingHom _ f _ RingPath : (R S : Ring ℓ) → RingEquiv R S ≃ (R ≡ S) RingPath = ∫ 𝒮ᴰ-Ring .UARel.ua +uaRing : {A B : Ring ℓ} → RingEquiv A B → A ≡ B +uaRing {A = A} {B = B} = equivFun (RingPath A B) + +isGroupoidRing : isGroupoid (Ring ℓ) +isGroupoidRing _ _ = isOfHLevelRespectEquiv 2 (RingPath _ _) (isSetRingEquiv _ _) + + -- Rings have an abelian group and a monoid Ring→AbGroup : Ring ℓ → AbGroup ℓ @@ -236,6 +253,7 @@ Ring→AddMonoid = Group→Monoid ∘ Ring→Group Ring→MultMonoid : Ring ℓ → Monoid ℓ Ring→MultMonoid (A , ringstr _ _ _ _ _ R) = monoid _ _ _ (IsRing.·IsMonoid R) + -- Smart constructor for ring homomorphisms -- that infers the other equations from pres1, pres+, and pres· diff --git a/Cubical/Algebra/Ring/Properties.agda b/Cubical/Algebra/Ring/Properties.agda index 5fffe6744a..954d0ea080 100644 --- a/Cubical/Algebra/Ring/Properties.agda +++ b/Cubical/Algebra/Ring/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Algebra.Ring.Properties where open import Cubical.Foundations.Prelude @@ -10,8 +10,11 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.Transport open import Cubical.Foundations.SIP +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path open import Cubical.Data.Sigma +open import Cubical.Relation.Binary.Poset open import Cubical.Structures.Axioms open import Cubical.Structures.Auto @@ -21,9 +24,11 @@ open import Cubical.Algebra.Monoid open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Ring.Base +open import Cubical.HITs.PropositionalTruncation + private variable - ℓ : Level + ℓ ℓ' ℓ'' ℓ''' ℓ'''' : Level {- some basic calculations (used for example in QuotientRing.agda), @@ -171,11 +176,23 @@ module RingHoms where fst (idRingHom R) = idfun (fst R) snd (idRingHom R) = makeIsRingHom refl (λ _ _ → refl) (λ _ _ → refl) - compRingHom : {R S T : Ring ℓ} → RingHom R S → RingHom S T → RingHom R T + compIsRingHom : {A : Ring ℓ} {B : Ring ℓ'} {C : Ring ℓ''} + {g : ⟨ B ⟩ → ⟨ C ⟩} {f : ⟨ A ⟩ → ⟨ B ⟩} + → IsRingHom (B .snd) g (C .snd) + → IsRingHom (A .snd) f (B .snd) + → IsRingHom (A .snd) (g ∘ f) (C .snd) + compIsRingHom {g = g} {f} gh fh .pres0 = cong g (fh .pres0) ∙ gh .pres0 + compIsRingHom {g = g} {f} gh fh .pres1 = cong g (fh .pres1) ∙ gh .pres1 + compIsRingHom {g = g} {f} gh fh .pres+ x y = cong g (fh .pres+ x y) ∙ gh .pres+ (f x) (f y) + compIsRingHom {g = g} {f} gh fh .pres· x y = cong g (fh .pres· x y) ∙ gh .pres· (f x) (f y) + compIsRingHom {g = g} {f} gh fh .pres- x = cong g (fh .pres- x) ∙ gh .pres- (f x) + + compRingHom : {R : Ring ℓ} {S : Ring ℓ'} {T : Ring ℓ''} + → RingHom R S → RingHom S T → RingHom R T fst (compRingHom f g) x = g .fst (f .fst x) - snd (compRingHom f g) = makeIsRingHom (cong (g .fst) (pres1 (snd f)) ∙ pres1 (snd g)) - (λ x y → cong (g .fst) (pres+ (snd f) _ _) ∙ pres+ (snd g) _ _) - (λ x y → cong (g .fst) (pres· (snd f) _ _) ∙ pres· (snd g) _ _) + snd (compRingHom f g) = compIsRingHom (g .snd) (f .snd) + + syntax compRingHom f g = g ∘r f compIdRingHom : {R S : Ring ℓ} (φ : RingHom R S) → compRingHom (idRingHom R) φ ≡ φ compIdRingHom φ = RingHom≡ refl @@ -188,6 +205,23 @@ module RingHoms where compAssocRingHom _ _ _ = RingHom≡ refl +module RingEquivs where + open IsRingHom + open RingHoms + + compIsRingEquiv : {A : Ring ℓ} {B : Ring ℓ'} {C : Ring ℓ''} + {g : ⟨ B ⟩ ≃ ⟨ C ⟩} {f : ⟨ A ⟩ ≃ ⟨ B ⟩} + → IsRingEquiv (B .snd) g (C .snd) + → IsRingEquiv (A .snd) f (B .snd) + → IsRingEquiv (A .snd) (compEquiv f g) (C .snd) + compIsRingEquiv {g = g} {f} gh fh = compIsRingHom {g = g .fst} {f .fst} gh fh + + compRingEquiv : {A : Ring ℓ} {B : Ring ℓ'} {C : Ring ℓ''} + → RingEquiv A B → RingEquiv B C → RingEquiv A C + fst (compRingEquiv f g) = compEquiv (f .fst) (g .fst) + snd (compRingEquiv f g) = compIsRingEquiv {g = g .fst} {f = f .fst} (g .snd) (f .snd) + + module RingHomTheory {R S : Ring ℓ} (φ : RingHom R S) where open RingTheory ⦃...⦄ open RingStr ⦃...⦄ @@ -210,3 +244,69 @@ module RingHomTheory {R S : Ring ℓ} (φ : RingHom R S) where f x - f y ≡⟨ cong (_- f y) p ⟩ f y - f y ≡⟨ +Rinv _ ⟩ 0r ∎ + + +-- the Ring version of uaCompEquiv +module RingUAFunctoriality where + open RingStr + open RingEquivs + + Ring≡ : (A B : Ring ℓ) → ( + Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] + Σ[ q0 ∈ PathP (λ i → p i) (0r (snd A)) (0r (snd B)) ] + Σ[ q1 ∈ PathP (λ i → p i) (1r (snd A)) (1r (snd B)) ] + Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] + Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] + Σ[ s ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] + PathP (λ i → IsRing (q0 i) (q1 i) (r+ i) (r· i) (s i)) (isRing (snd A)) (isRing (snd B))) + ≃ (A ≡ B) + Ring≡ A B = isoToEquiv theIso + where + open Iso + theIso : Iso _ _ + fun theIso (p , q0 , q1 , r+ , r· , s , t) i = p i + , ringstr (q0 i) (q1 i) (r+ i) (r· i) (s i) (t i) + inv theIso x = cong ⟨_⟩ x , cong (0r ∘ snd) x , cong (1r ∘ snd) x + , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (isRing ∘ snd) x + rightInv theIso _ = refl + leftInv theIso _ = refl + + caracRing≡ : {A B : Ring ℓ} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q + caracRing≡ {A = A} {B = B} p q P = + sym (transportTransport⁻ (ua (Ring≡ A B)) p) + ∙∙ cong (transport (ua (Ring≡ A B))) helper + ∙∙ transportTransport⁻ (ua (Ring≡ A B)) q + where + helper : transport (sym (ua (Ring≡ A B))) p ≡ transport (sym (ua (Ring≡ A B))) q + helper = Σ≡Prop + (λ _ → isPropΣ + (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) + λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) + λ _ → isOfHLevelPathP 1 (isPropIsRing _ _ _ _ _) _ _) + (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + + uaCompRingEquiv : {A B C : Ring ℓ} (f : RingEquiv A B) (g : RingEquiv B C) + → uaRing (compRingEquiv f g) ≡ uaRing f ∙ uaRing g + uaCompRingEquiv f g = caracRing≡ _ _ ( + cong ⟨_⟩ (uaRing (compRingEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaRing f) ∙ cong ⟨_⟩ (uaRing g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaRing f) (uaRing g)) ⟩ + cong ⟨_⟩ (uaRing f ∙ uaRing g) ∎) + + + +open RingEquivs +open RingUAFunctoriality +recPT→Ring : {A : Type ℓ'} (𝓕 : A → Ring ℓ) + → (σ : ∀ x y → RingEquiv (𝓕 x) (𝓕 y)) + → (∀ x y z → σ x z ≡ compRingEquiv (σ x y) (σ y z)) + ------------------------------------------------------ + → ∥ A ∥ → Ring ℓ +recPT→Ring 𝓕 σ compCoh = rec→Gpd isGroupoidRing 𝓕 + (3-ConstantCompChar 𝓕 (λ x y → uaRing (σ x y)) + λ x y z → sym ( cong uaRing (compCoh x y z) + ∙ uaCompRingEquiv (σ x y) (σ y z))) diff --git a/Cubical/Algebra/RingSolver/RawAlgebra.agda b/Cubical/Algebra/RingSolver/RawAlgebra.agda index f7f1b9476b..a0ba97aa6f 100644 --- a/Cubical/Algebra/RingSolver/RawAlgebra.agda +++ b/Cubical/Algebra/RingSolver/RawAlgebra.agda @@ -2,7 +2,6 @@ module Cubical.Algebra.RingSolver.RawAlgebra where open import Cubical.Foundations.Prelude -open import Cubical.Data.Sigma open import Cubical.Data.Nat using (ℕ) open import Cubical.Data.Int renaming (_+_ to _+ℤ_ ; _·_ to _·ℤ_ ; -_ to -ℤ_ ; _-_ to _-ℤ_ ; +Assoc to +ℤAssoc ; +Comm to +ℤComm ; -DistL· to -ℤDistL·ℤ) diff --git a/Cubical/Algebra/RingSolver/RawRing.agda b/Cubical/Algebra/RingSolver/RawRing.agda index 605f42d54c..46c5e1e7a3 100644 --- a/Cubical/Algebra/RingSolver/RawRing.agda +++ b/Cubical/Algebra/RingSolver/RawRing.agda @@ -2,8 +2,6 @@ module Cubical.Algebra.RingSolver.RawRing where open import Cubical.Foundations.Prelude -open import Cubical.Data.Sigma -open import Cubical.Data.Nat using (ℕ) private variable diff --git a/Cubical/Algebra/ZariskiLattice/Base.agda b/Cubical/Algebra/ZariskiLattice/Base.agda index a8571ca743..fd78ec0894 100644 --- a/Cubical/Algebra/ZariskiLattice/Base.agda +++ b/Cubical/Algebra/ZariskiLattice/Base.agda @@ -9,17 +9,17 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.HLevels open import Cubical.Foundations.Transport -open import Cubical.Foundations.Powerset using (⊆-refl-consequence) import Cubical.Data.Empty as ⊥ open import Cubical.Data.Bool -open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm ; ·-identityʳ to ·ℕ-rid) open import Cubical.Data.Sigma.Base open import Cubical.Data.Sigma.Properties open import Cubical.Data.FinData +open import Cubical.Data.Unit open import Cubical.Relation.Nullary open import Cubical.Relation.Binary open import Cubical.Relation.Binary.Poset @@ -27,27 +27,18 @@ open import Cubical.Relation.Binary.Poset open import Cubical.Algebra.Ring open import Cubical.Algebra.Ring.Properties open import Cubical.Algebra.Ring.BigOps -open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.BinomialThm open import Cubical.Algebra.CommRing.Ideal open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.RadicalIdeal -open import Cubical.Algebra.CommRing.Localisation.Base -open import Cubical.Algebra.CommRing.Localisation.UniversalProperty -open import Cubical.Algebra.CommRing.Localisation.InvertingElements -open import Cubical.Algebra.CommAlgebra.Base -open import Cubical.Algebra.CommAlgebra.Properties -open import Cubical.Algebra.CommAlgebra.Localisation open import Cubical.Algebra.RingSolver.Reflection open import Cubical.Algebra.Semilattice open import Cubical.Algebra.Lattice open import Cubical.Algebra.DistLattice -open import Cubical.Algebra.DistLattice.BigOps open import Cubical.Algebra.Matrix open import Cubical.HITs.SetQuotients as SQ -open import Cubical.HITs.PropositionalTruncation as PT open Iso open BinaryRelation @@ -194,258 +185,6 @@ module ZarLat (R' : CommRing ℓ) where ∧zAssoc ∧zRid ∧zComm ∧zAbsorb∨z ∧zLDist∨z - -module _ (R' : CommRing ℓ) (L' : DistLattice ℓ') where - - open CommRingStr (R' .snd) - open RingTheory (CommRing→Ring R') - open Sum (CommRing→Ring R') - open CommRingTheory R' - open Exponentiation R' - - open DistLatticeStr (L' .snd) renaming (is-set to isSetL) - open Join L' - open LatticeTheory (DistLattice→Lattice L') - open Order (DistLattice→Lattice L') - open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice L')) - open PosetReasoning IndPoset - open PosetStr (IndPoset .snd) hiding (_≤_) - private - R = fst R' - L = fst L' - - record IsZarMap (d : R → L) : Type (ℓ-max ℓ ℓ') where - constructor iszarmap - - field - pres0 : d 0r ≡ 0l - pres1 : d 1r ≡ 1l - ·≡∧ : ∀ x y → d (x · y) ≡ d x ∧l d y - +≤∨ : ∀ x y → d (x + y) ≤ d x ∨l d y - - ∑≤⋁ : {n : ℕ} (U : FinVec R n) → d (∑ U) ≤ ⋁ λ i → d (U i) - ∑≤⋁ {n = zero} U = ∨lRid _ ∙ pres0 - ∑≤⋁ {n = suc n} U = d (∑ U) ≤⟨ ∨lIdem _ ⟩ - d (U zero + ∑ (U ∘ suc)) ≤⟨ +≤∨ _ _ ⟩ - d (U zero) ∨l d (∑ (U ∘ suc)) ≤⟨ ≤-∨LPres _ _ _ (∑≤⋁ (U ∘ suc)) ⟩ - d (U zero) ∨l ⋁ (d ∘ U ∘ suc) ≤⟨ ∨lIdem _ ⟩ - ⋁ (d ∘ U) ◾ - - d·LCancel : ∀ x y → d (x · y) ≤ d y - d·LCancel x y = subst (λ a → a ≤ d y) (sym (·≡∧ x y)) (∧≤LCancelJoin _ _) - - linearCombination≤LCancel : {n : ℕ} (α β : FinVec R n) - → d (linearCombination R' α β) ≤ ⋁ (d ∘ β) - linearCombination≤LCancel α β = is-trans _ _ _ (∑≤⋁ (λ i → α i · β i)) - (≤-⋁Ext _ _ λ i → d·LCancel (α i) (β i)) - - ZarMapIdem : ∀ (n : ℕ) (x : R) → d (x ^ (suc n)) ≡ d x - ZarMapIdem zero x = ·≡∧ _ _ ∙∙ cong (d x ∧l_) pres1 ∙∙ ∧lRid _ - ZarMapIdem (suc n) x = ·≡∧ _ _ ∙∙ cong (d x ∧l_) (ZarMapIdem n x) ∙∙ ∧lIdem _ - - ZarMapExpIneq : ∀ (n : ℕ) (x : R) → d x ≤ d (x ^ n) - ZarMapExpIneq zero x = cong (d x ∨l_) pres1 ∙∙ 1lRightAnnihilates∨l _ ∙∙ sym pres1 - ZarMapExpIneq (suc n) x = subst (λ y → d x ≤ y) (sym (ZarMapIdem _ x)) (∨lIdem _) - - -- the crucial lemma about "Zariski maps" - open CommIdeal R' - open RadicalIdeal R' - open isCommIdeal - private - ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal - ⟨ V ⟩ = ⟨ V ⟩[ R' ] - - ZarMapRadicalIneq : ∀ {n : ℕ} (α : FinVec R n) (x : R) - → x ∈ √ ⟨ α ⟩ → d x ≤ ⋁ (d ∘ α) - ZarMapRadicalIneq α x = PT.elim (λ _ → isSetL _ _) - (uncurry (λ n → (PT.elim (λ _ → isSetL _ _) (uncurry (curriedHelper n))))) - where - curriedHelper : (n : ℕ) (β : FinVec R _) - → x ^ n ≡ linearCombination R' β α → d x ≤ ⋁ (d ∘ α) - curriedHelper n β xⁿ≡∑βα = d x ≤⟨ ZarMapExpIneq n x ⟩ - d (x ^ n) - ≤⟨ subst (λ y → y ≤ ⋁ (d ∘ α)) (sym (cong d xⁿ≡∑βα)) (linearCombination≤LCancel β α) ⟩ - ⋁ (d ∘ α) ◾ - -module ZarLatUniversalProp (R' : CommRing ℓ) where - open CommRingStr (snd R') - open RingTheory (CommRing→Ring R') - open Sum (CommRing→Ring R') - open CommRingTheory R' - open Exponentiation R' - open BinomialThm R' - open CommIdeal R' - open RadicalIdeal R' - open isCommIdeal - open ProdFin R' - - open ZarLat R' - open IsZarMap - - private - R = fst R' - ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal - ⟨ V ⟩ = ⟨ V ⟩[ R' ] - - - D : R → ZL - D x = [ 1 , replicateFinVec 1 x ] -- λ x → √⟨x⟩ - - isZarMapD : IsZarMap R' ZariskiLattice D - pres0 isZarMapD = eq/ _ _ (cong √ (0FGIdeal _ ∙ sym (emptyFGIdeal _ _))) - pres1 isZarMapD = refl - ·≡∧ isZarMapD x y = cong {B = λ _ → ZL} (λ U → [ 1 , U ]) (Length1··Fin x y) - +≤∨ isZarMapD x y = eq/ _ _ (cong √ (CommIdeal≡Char (inclOfFGIdeal _ 3Vec ⟨ 2Vec ⟩ 3Vec⊆2Vec) - (inclOfFGIdeal _ 2Vec ⟨ 3Vec ⟩ 2Vec⊆3Vec))) - where - 2Vec = replicateFinVec 1 x ++Fin replicateFinVec 1 y - 3Vec = replicateFinVec 1 (x + y) ++Fin (replicateFinVec 1 x ++Fin replicateFinVec 1 y) - - 3Vec⊆2Vec : ∀ (i : Fin 3) → 3Vec i ∈ ⟨ 2Vec ⟩ - 3Vec⊆2Vec zero = ⟨ 2Vec ⟩ .snd .+Closed (indInIdeal _ _ zero) (indInIdeal _ _ (suc zero)) - 3Vec⊆2Vec (suc zero) = indInIdeal _ _ zero - 3Vec⊆2Vec (suc (suc zero)) = indInIdeal _ _ (suc zero) - - 2Vec⊆3Vec : ∀ (i : Fin 2) → 2Vec i ∈ ⟨ 3Vec ⟩ - 2Vec⊆3Vec zero = indInIdeal _ _ (suc zero) - 2Vec⊆3Vec (suc zero) = indInIdeal _ _ (suc (suc zero)) - - - -- defintion of the universal property - hasZarLatUniversalProp : (L : DistLattice ℓ') (D : R → fst L) - → IsZarMap R' L D - → Type _ - hasZarLatUniversalProp {ℓ' = ℓ'} L D _ = ∀ (L' : DistLattice ℓ') (d : R → fst L') - → IsZarMap R' L' d - → ∃![ χ ∈ DistLatticeHom L L' ] (fst χ) ∘ D ≡ d - - isPropZarLatUniversalProp : (L : DistLattice ℓ') (D : R → fst L) (isZarMapD : IsZarMap R' L D) - → isProp (hasZarLatUniversalProp L D isZarMapD) - isPropZarLatUniversalProp L D isZarMapD = isPropΠ3 (λ _ _ _ → isPropIsContr) - - ZLHasUniversalProp : hasZarLatUniversalProp ZariskiLattice D isZarMapD - ZLHasUniversalProp L' d isZarMapd = (χ , funExt χcomp) , χunique - where - open DistLatticeStr (snd L') renaming (is-set to isSetL) - open LatticeTheory (DistLattice→Lattice L') - open Join L' - open IsLatticeHom - L = fst L' - - χ : DistLatticeHom ZariskiLattice L' - fst χ = SQ.rec isSetL (λ (_ , α) → ⋁ (d ∘ α)) - λ (_ , α) (_ , β) → curriedHelper α β - where - curriedHelper : {n m : ℕ} (α : FinVec R n) (β : FinVec R m) - → √ ⟨ α ⟩ ≡ √ ⟨ β ⟩ → ⋁ (d ∘ α) ≡ ⋁ (d ∘ β) - curriedHelper α β √⟨α⟩≡√⟨β⟩ = is-antisym _ _ ineq1 ineq2 - where - open Order (DistLattice→Lattice L') - open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice L')) - open PosetReasoning IndPoset - open PosetStr (IndPoset .snd) hiding (_≤_) - - incl1 : √ ⟨ α ⟩ ⊆ √ ⟨ β ⟩ - incl1 = ⊆-refl-consequence _ _ (cong fst √⟨α⟩≡√⟨β⟩) .fst - - ineq1 : ⋁ (d ∘ α) ≤ ⋁ (d ∘ β) - ineq1 = ⋁IsMax (d ∘ α) (⋁ (d ∘ β)) - λ i → ZarMapRadicalIneq isZarMapd β (α i) (√FGIdealCharLImpl α ⟨ β ⟩ incl1 i) - - incl2 : √ ⟨ β ⟩ ⊆ √ ⟨ α ⟩ - incl2 = ⊆-refl-consequence _ _ (cong fst √⟨α⟩≡√⟨β⟩) .snd - - ineq2 : ⋁ (d ∘ β) ≤ ⋁ (d ∘ α) - ineq2 = ⋁IsMax (d ∘ β) (⋁ (d ∘ α)) - λ i → ZarMapRadicalIneq isZarMapd α (β i) (√FGIdealCharLImpl β ⟨ α ⟩ incl2 i) - - - pres0 (snd χ) = refl - pres1 (snd χ) = ∨lRid _ ∙ isZarMapd .pres1 - pres∨l (snd χ) = elimProp2 (λ _ _ → isSetL _ _) (uncurry (λ n α → uncurry (curriedHelper n α))) - where - curriedHelper : (n : ℕ) (α : FinVec R n) (m : ℕ) (β : FinVec R m) - → ⋁ (d ∘ (α ++Fin β)) ≡ ⋁ (d ∘ α) ∨l ⋁ (d ∘ β) - curriedHelper zero α _ β = sym (∨lLid _) - curriedHelper (suc n) α _ β = - ⋁ (d ∘ (α ++Fin β)) ≡⟨ refl ⟩ - d (α zero) ∨l ⋁ (d ∘ ((α ∘ suc) ++Fin β)) - - ≡⟨ cong (d (α zero) ∨l_) (curriedHelper _ (α ∘ suc) _ β) ⟩ - - d (α zero) ∨l (⋁ (d ∘ α ∘ suc) ∨l ⋁ (d ∘ β)) - ≡⟨ ∨lAssoc _ _ _ ⟩ - - ⋁ (d ∘ α) ∨l ⋁ (d ∘ β) ∎ - - pres∧l (snd χ) = elimProp2 (λ _ _ → isSetL _ _) (uncurry (λ n α → uncurry (curriedHelper n α))) - where - -- have to repeat this one here so the termination checker won't complain - oldHelper : (n : ℕ) (α : FinVec R n) (m : ℕ) (β : FinVec R m) - → ⋁ (d ∘ (α ++Fin β)) ≡ ⋁ (d ∘ α) ∨l ⋁ (d ∘ β) - oldHelper zero α _ β = sym (∨lLid _) - oldHelper (suc n) α _ β = cong (d (α zero) ∨l_) (oldHelper _ (α ∘ suc) _ β) ∙ ∨lAssoc _ _ _ - - curriedHelper : (n : ℕ) (α : FinVec R n) (m : ℕ) (β : FinVec R m) - → ⋁ (d ∘ (α ··Fin β)) ≡ ⋁ (d ∘ α) ∧l ⋁ (d ∘ β) - curriedHelper zero α _ β = sym (0lLeftAnnihilates∧l _) - curriedHelper (suc n) α _ β = - ⋁ (d ∘ (α ··Fin β)) ≡⟨ refl ⟩ - ⋁ (d ∘ ((λ j → α zero · β j) ++Fin ((α ∘ suc) ··Fin β))) - - ≡⟨ oldHelper _ (λ j → α zero · β j) _ ((α ∘ suc) ··Fin β) ⟩ - - ⋁ (d ∘ (λ j → α zero · β j)) ∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β)) - - ≡⟨ cong (_∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β))) (⋁Ext (λ j → isZarMapd .·≡∧ (α zero) (β j))) ⟩ - - ⋁ (λ j → d (α zero) ∧l d (β j)) ∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β)) - - ≡⟨ cong (_∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β))) (sym (⋁Meetrdist _ _)) ⟩ - - (d (α zero) ∧l ⋁ (d ∘ β)) ∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β)) - - ≡⟨ cong ((d (α zero) ∧l ⋁ (d ∘ β)) ∨l_) (curriedHelper _ (α ∘ suc) _ β) ⟩ - - (d (α zero) ∧l ⋁ (d ∘ β)) ∨l (⋁ (d ∘ α ∘ suc) ∧l ⋁ (d ∘ β)) - - ≡⟨ sym (∧lRdist∨l _ _ _) ⟩ - - ⋁ (d ∘ α) ∧l ⋁ (d ∘ β) ∎ - - - χcomp : ∀ (f : R) → χ .fst (D f) ≡ d f - χcomp f = ∨lRid (d f) - - χunique : (y : Σ[ χ' ∈ DistLatticeHom ZariskiLattice L' ] fst χ' ∘ D ≡ d) - → (χ , funExt χcomp) ≡ y - χunique (χ' , χ'∘D≡d) = Σ≡Prop (λ _ → isSetΠ (λ _ → isSetL) _ _) (LatticeHom≡f _ _ - (funExt (elimProp (λ _ → isSetL _ _) (uncurry uniqHelper)))) - where - uniqHelper : (n : ℕ) (α : FinVec R n) → fst χ [ n , α ] ≡ fst χ' [ n , α ] - uniqHelper zero _ = sym (cong (λ α → fst χ' [ 0 , α ]) (funExt (λ ())) ∙ χ' .snd .pres0) - uniqHelper (suc n) α = - ⋁ (d ∘ α) ≡⟨ refl ⟩ - d (α zero) ∨l ⋁ (d ∘ α ∘ suc) - - ≡⟨ cong (d (α zero) ∨l_) (uniqHelper n (α ∘ suc)) ⟩ -- the inductive step - - d (α zero) ∨l fst χ' [ n , α ∘ suc ] - - ≡⟨ cong (_∨l fst χ' [ n , α ∘ suc ]) (sym (funExt⁻ χ'∘D≡d (α zero))) ⟩ - - fst χ' (D (α zero)) ∨l fst χ' [ n , α ∘ suc ] - - ≡⟨ sym (χ' .snd .pres∨l _ _) ⟩ - - fst χ' (D (α zero) ∨z [ n , α ∘ suc ]) - - ≡⟨ cong (λ β → fst χ' [ suc n , β ]) (funExt (λ { zero → refl ; (suc i) → refl })) ⟩ - - fst χ' [ suc n , α ] ∎ - - - -- An equivalent definition that doesn't bump up the unviverse level module SmallZarLat (R' : CommRing ℓ) where open CommRingStr (snd R') diff --git a/Cubical/Algebra/ZariskiLattice/StructureSheaf.agda b/Cubical/Algebra/ZariskiLattice/StructureSheaf.agda new file mode 100644 index 0000000000..93df73cade --- /dev/null +++ b/Cubical/Algebra/ZariskiLattice/StructureSheaf.agda @@ -0,0 +1,455 @@ +{- + + This module defines the basic opens of the Zariski lattice and proves that they're a basis of the lattice. + It also contains the construction of the structure presheaf and a proof of the sheaf property on basic opens, + using the theory developed in the module PreSheafFromUniversalProp and its toSheaf.lemma. + Note that the structure sheaf is a functor into R-algebras and not just commutative rings. + +-} + +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.ZariskiLattice.StructureSheaf where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Powerset using (ℙ ; ⊆-refl-consequence) + renaming (_∈_ to _∈ₚ_ ; subst-∈ to subst-∈ₚ) + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm + ; ·-identityʳ to ·ℕ-rid) +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary +open import Cubical.Relation.Binary.Poset + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Ring.Properties +open import Cubical.Algebra.Ring.BigOps +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.BinomialThm +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.RadicalIdeal +open import Cubical.Algebra.CommRing.Localisation.Base +open import Cubical.Algebra.CommRing.Localisation.UniversalProperty +open import Cubical.Algebra.CommRing.Localisation.InvertingElements +open import Cubical.Algebra.CommRing.Localisation.PullbackSquare +open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Properties +open import Cubical.Algebra.CommAlgebra.Localisation +open import Cubical.Algebra.CommAlgebra.Instances.Unit +open import Cubical.Algebra.RingSolver.Reflection +open import Cubical.Algebra.Semilattice +open import Cubical.Algebra.Lattice +open import Cubical.Algebra.DistLattice +open import Cubical.Algebra.DistLattice.Basis +open import Cubical.Algebra.DistLattice.BigOps +open import Cubical.Algebra.ZariskiLattice.Base +open import Cubical.Algebra.ZariskiLattice.UniversalProperty + +open import Cubical.Categories.Category.Base hiding (_[_,_]) +open import Cubical.Categories.Functor +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Instances.CommAlgebras +open import Cubical.Categories.Instances.DistLattice +open import Cubical.Categories.Instances.Semilattice +open import Cubical.Categories.DistLatticeSheaf + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open Iso +open BinaryRelation +open isEquivRel + +private + variable + ℓ ℓ' : Level + + + +module _ (R' : CommRing ℓ) where + open CommRingStr ⦃...⦄ + open RingTheory (CommRing→Ring R') + open CommIdeal R' + open isCommIdeal + + open ZarLat R' + open ZarLatUniversalProp R' + open IsZarMap + + open Join ZariskiLattice + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice ZariskiLattice)) + open IsBasis + + private + R = fst R' + instance + _ = snd R' + ⟨_⟩ : R → CommIdeal + ⟨ f ⟩ = ⟨ replicateFinVec 1 f ⟩[ R' ] + ⟨_⟩ₚ : R × R → CommIdeal -- p is for pair + ⟨ f , g ⟩ₚ = ⟨ replicateFinVec 1 f ++Fin replicateFinVec 1 g ⟩[ R' ] + + + BasicOpens : ℙ ZL + BasicOpens 𝔞 = (∃[ f ∈ R ] (D f ≡ 𝔞)) , isPropPropTrunc + + BO : Type (ℓ-suc ℓ) + BO = Σ[ 𝔞 ∈ ZL ] (𝔞 ∈ₚ BasicOpens) + + basicOpensAreBasis : IsBasis ZariskiLattice BasicOpens + contains1 basicOpensAreBasis = ∣ 1r , isZarMapD .pres1 ∣ + ∧lClosed basicOpensAreBasis 𝔞 𝔟 = map2 + λ (f , Df≡𝔞) (g , Dg≡𝔟) → (f · g) , isZarMapD .·≡∧ f g ∙ cong₂ (_∧z_) Df≡𝔞 Dg≡𝔟 + ⋁Basis basicOpensAreBasis = elimProp (λ _ → isPropPropTrunc) Σhelper + where + Σhelper : (a : Σ[ n ∈ ℕ ] FinVec R n) + → ∃[ n ∈ ℕ ] Σ[ α ∈ FinVec ZL n ] (∀ i → α i ∈ₚ BasicOpens) × (⋁ α ≡ [ a ]) + Σhelper (n , α) = ∣ n , (D ∘ α) , (λ i → ∣ α i , refl ∣) , path ∣ + where + path : ⋁ (D ∘ α) ≡ [ n , α ] + path = funExt⁻ (cong fst ZLUniversalPropCorollary) _ + + + -- The structure presheaf on BO + ZariskiCat = DistLatticeCategory ZariskiLattice + + BOCat : Category (ℓ-suc ℓ) (ℓ-suc ℓ) + BOCat = ΣPropCat ZariskiCat BasicOpens + + private + P : ZL → Type _ + P 𝔞 = Σ[ f ∈ R ] (D f ≡ 𝔞) -- the untruncated defining property + + 𝓕 : Σ ZL P → CommAlgebra R' _ + 𝓕 (_ , f , _) = R[1/ f ]AsCommAlgebra -- D(f) ↦ R[1/f] + + uniqueHom : ∀ (x y : Σ ZL P) → (fst x) ≤ (fst y) → isContr (CommAlgebraHom (𝓕 y) (𝓕 x)) + uniqueHom (𝔞 , f , p) (𝔟 , g , q) = contrHoms 𝔞 𝔟 f g p q + where + open InvertingElementsBase R' + + contrHoms : (𝔞 𝔟 : ZL) (f g : R) (p : D f ≡ 𝔞) (q : D g ≡ 𝔟) + → 𝔞 ≤ 𝔟 → isContr (CommAlgebraHom R[1/ g ]AsCommAlgebra R[1/ f ]AsCommAlgebra) + contrHoms 𝔞 𝔟 f g p q 𝔞≤𝔟 = R[1/g]HasAlgUniversalProp R[1/ f ]AsCommAlgebra + λ s s∈[gⁿ|n≥0] → subst-∈ₚ (R[1/ f ]AsCommRing ˣ) + (sym (·Rid (s /1))) --can't apply the lemma directly as we get mult with 1 somewhere + (RadicalLemma.toUnit R' f g f∈√⟨g⟩ s s∈[gⁿ|n≥0]) + where + open AlgLoc R' [ g ⁿ|n≥0] (powersFormMultClosedSubset g) + renaming (S⁻¹RHasAlgUniversalProp to R[1/g]HasAlgUniversalProp) + open S⁻¹RUniversalProp R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) using (_/1) + open RadicalIdeal R' + + private + instance + _ = snd R[1/ f ]AsCommRing + + Df≤Dg : D f ≤ D g + Df≤Dg = subst2 _≤_ (sym p) (sym q) 𝔞≤𝔟 + + radicalHelper : √ ⟨ f , g ⟩ₚ ≡ √ ⟨ g ⟩ + radicalHelper = + isEquivRel→effectiveIso (λ _ _ → isSetCommIdeal _ _) ∼EquivRel _ _ .fun Df≤Dg + + f∈√⟨g⟩ : f ∈ √ ⟨ g ⟩ + f∈√⟨g⟩ = subst (f ∈_) radicalHelper (∈→∈√ _ _ (indInIdeal _ _ zero)) + + + open PreSheafFromUniversalProp ZariskiCat P 𝓕 uniqueHom + BasisStructurePShf : Functor (BOCat ^op) (CommAlgebrasCategory R') + BasisStructurePShf = universalPShf + + + -- now prove the sheaf properties + open SheafOnBasis ZariskiLattice (CommAlgebrasCategory R' {ℓ' = ℓ}) + (TerminalCommAlgebra R') BasicOpens basicOpensAreBasis + + isSheafBasisStructurePShf : isDLBasisSheaf BasisStructurePShf + fst isSheafBasisStructurePShf 0∈BO = + transport (λ i → F-ob (0z , canonical0∈BO≡0∈BO i) ≡ UnitCommAlgebra R') R[1/0]≡0 + where + open Functor ⦃...⦄ + instance + _ = BasisStructurePShf + + canonical0∈BO : 0z ∈ₚ BasicOpens + canonical0∈BO = ∣ 0r , isZarMapD .pres0 ∣ + + canonical0∈BO≡0∈BO : canonical0∈BO ≡ 0∈BO + canonical0∈BO≡0∈BO = BasicOpens 0z .snd _ _ + + R[1/0]≡0 : R[1/ 0r ]AsCommAlgebra ≡ UnitCommAlgebra R' + R[1/0]≡0 = uaCommAlgebra (e , eIsRHom) + where + open InvertingElementsBase R' using (isContrR[1/0]) + open IsAlgebraHom + + e : R[1/ 0r ]AsCommAlgebra .fst ≃ UnitCommAlgebra R' .fst + e = isContr→Equiv isContrR[1/0] isContrUnit* + + eIsRHom : IsCommAlgebraEquiv (R[1/ 0r ]AsCommAlgebra .snd) e (UnitCommAlgebra R' .snd) + pres0 eIsRHom = refl + pres1 eIsRHom = refl + pres+ eIsRHom _ _ = refl + pres· eIsRHom _ _ = refl + pres- eIsRHom _ = refl + pres⋆ eIsRHom _ _ = refl + + snd isSheafBasisStructurePShf (𝔞 , 𝔞∈BO) (𝔟 , 𝔟∈BO) 𝔞∨𝔟∈BO = curriedHelper 𝔞 𝔟 𝔞∈BO 𝔟∈BO 𝔞∨𝔟∈BO + where + open condSquare + {- + here: + BFsq (𝔞 , 𝔞∈BO) (𝔟 , 𝔟∈BO) 𝔞∨𝔟∈BO BasisStructurePShf = + + 𝓞 (𝔞∨𝔟) → 𝓞 (𝔞) + + ↓ ↓ + + 𝓞 (𝔟) → 𝓞 (𝔞∧𝔟) + + -} + curriedHelper : (𝔞 𝔟 : ZL) (𝔞∈BO : 𝔞 ∈ₚ BasicOpens) (𝔟∈BO : 𝔟 ∈ₚ BasicOpens) + (𝔞∨𝔟∈BO : 𝔞 ∨z 𝔟 ∈ₚ BasicOpens) + → isPullback (CommAlgebrasCategory R') _ _ _ + (BFsq (𝔞 , 𝔞∈BO) (𝔟 , 𝔟∈BO) 𝔞∨𝔟∈BO BasisStructurePShf) + curriedHelper 𝔞 𝔟 = elim3 (λ 𝔞∈BO 𝔟∈BO 𝔞∨𝔟∈BO → isPropIsPullback _ _ _ _ + (BFsq (𝔞 , 𝔞∈BO) (𝔟 , 𝔟∈BO) 𝔞∨𝔟∈BO BasisStructurePShf)) + Σhelper + where + -- write everything explicitly so things can type-check + thePShfCospan : (a : Σ[ f ∈ R ] D f ≡ 𝔞) (b : Σ[ g ∈ R ] D g ≡ 𝔟) + → Cospan (CommAlgebrasCategory R') + Cospan.l (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-ob (𝔟 , ∣ g , Dg≡𝔟 ∣) + Cospan.m (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-ob + (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣ ∣ g , Dg≡𝔟 ∣) + Cospan.r (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-ob (𝔞 , ∣ f , Df≡𝔞 ∣) + Cospan.s₁ (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-hom + {x = (𝔟 , ∣ g , Dg≡𝔟 ∣)} + {y = (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣ ∣ g , Dg≡𝔟 ∣)} + (hom-∧₂ ZariskiLattice (CommAlgebrasCategory R' {ℓ' = ℓ}) (TerminalCommAlgebra R') 𝔞 𝔟) + Cospan.s₂ (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-hom + {x = (𝔞 , ∣ f , Df≡𝔞 ∣)} + {y = (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣ ∣ g , Dg≡𝔟 ∣)} + (hom-∧₁ ZariskiLattice (CommAlgebrasCategory R' {ℓ' = ℓ}) (TerminalCommAlgebra R') 𝔞 𝔟) + + + Σhelper : (a : Σ[ f ∈ R ] D f ≡ 𝔞) (b : Σ[ g ∈ R ] D g ≡ 𝔟) (c : Σ[ h ∈ R ] D h ≡ 𝔞 ∨z 𝔟) + → isPullback (CommAlgebrasCategory R') (thePShfCospan a b) _ _ + (BFsq (𝔞 , ∣ a ∣) (𝔟 , ∣ b ∣) ∣ c ∣ BasisStructurePShf) + Σhelper (f , Df≡𝔞) (g , Dg≡𝔟) (h , Dh≡𝔞∨𝔟) = toSheaf.lemma + (𝔞 ∨z 𝔟 , ∣ h , Dh≡𝔞∨𝔟 ∣) + (𝔞 , ∣ f , Df≡𝔞 ∣) + (𝔟 , ∣ g , Dg≡𝔟 ∣) + (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣ ∣ g , Dg≡𝔟 ∣) + (Bsq (𝔞 , ∣ f , Df≡𝔞 ∣) (𝔟 , ∣ g , Dg≡𝔟 ∣) ∣ h , Dh≡𝔞∨𝔟 ∣) + theAlgebraCospan theAlgebraPullback refl gPath fPath fgPath + where + open Exponentiation R' + open RadicalIdeal R' + open InvertingElementsBase R' + open DoubleLoc R' h + open S⁻¹RUniversalProp R' [ h ⁿ|n≥0] (powersFormMultClosedSubset h) + open CommIdeal R[1/ h ]AsCommRing using () renaming (CommIdeal to CommIdealₕ ; _∈_ to _∈ₕ_) + + instance + _ = snd R[1/ h ]AsCommRing + + ⟨_⟩ₕ : R[1/ h ] × R[1/ h ] → CommIdealₕ + ⟨ x , y ⟩ₕ = ⟨ replicateFinVec 1 x ++Fin replicateFinVec 1 y ⟩[ R[1/ h ]AsCommRing ] + + -- the crucial algebraic fact: + radicalPath : √ ⟨ h ⟩ ≡ √ ⟨ f , g ⟩ₚ + radicalPath = isEquivRel→effectiveIso (λ _ _ → isSetCommIdeal _ _) ∼EquivRel _ _ .fun DHelper + where + DHelper : D h ≡ D f ∨z D g + DHelper = Dh≡𝔞∨𝔟 ∙ cong₂ (_∨z_) (sym Df≡𝔞) (sym Dg≡𝔟) + + f∈√⟨h⟩ : f ∈ √ ⟨ h ⟩ + f∈√⟨h⟩ = subst (f ∈_) (sym radicalPath) (∈→∈√ _ _ (indInIdeal _ _ zero)) + + g∈√⟨h⟩ : g ∈ √ ⟨ h ⟩ + g∈√⟨h⟩ = subst (g ∈_) (sym radicalPath) (∈→∈√ _ _ (indInIdeal _ _ (suc zero))) + + fg∈√⟨h⟩ : (f · g) ∈ √ ⟨ h ⟩ + fg∈√⟨h⟩ = √ ⟨ h ⟩ .snd .·Closed f g∈√⟨h⟩ + + 1∈fgIdeal : 1r ∈ₕ ⟨ (f /1) , (g /1) ⟩ₕ + 1∈fgIdeal = helper1 (subst (h ∈_) radicalPath (∈→∈√ _ _ (indInIdeal _ _ zero))) + where + helper1 : h ∈ √ ⟨ f , g ⟩ₚ + → 1r ∈ₕ ⟨ (f /1) , (g /1) ⟩ₕ + helper1 = PT.rec isPropPropTrunc (uncurry helper2) + where + helper2 : (n : ℕ) + → h ^ n ∈ ⟨ f , g ⟩ₚ + → 1r ∈ₕ ⟨ (f /1) , (g /1) ⟩ₕ + helper2 n = map helper3 + where + helper3 : Σ[ α ∈ FinVec R 2 ] + h ^ n ≡ linearCombination R' α (λ { zero → f ; (suc zero) → g }) + → Σ[ β ∈ FinVec R[1/ h ] 2 ] + 1r ≡ linearCombination R[1/ h ]AsCommRing β + λ { zero → f /1 ; (suc zero) → g /1 } + helper3 (α , p) = β , path + where + β : FinVec R[1/ h ] 2 + β zero = [ α zero , h ^ n , ∣ n , refl ∣ ] + β (suc zero) = [ α (suc zero) , h ^ n , ∣ n , refl ∣ ] + + path : 1r ≡ linearCombination R[1/ h ]AsCommRing β + λ { zero → f /1 ; (suc zero) → g /1 } + path = eq/ _ _ ((1r , ∣ 0 , refl ∣) , bigPath) + ∙ cong (β zero · (f /1) +_) (sym (+Rid (β (suc zero) · (g /1)))) + where + useSolver1 : ∀ hn → 1r · 1r · ((hn · 1r) · (hn · 1r)) ≡ hn · hn + useSolver1 = solve R' + + useSolver2 : ∀ az f hn as g → hn · (az · f + (as · g + 0r)) + ≡ 1r · (az · f · (hn · 1r) + as · g · (hn · 1r)) · 1r + useSolver2 = solve R' + + bigPath : 1r · 1r · ((h ^ n · 1r) · (h ^ n · 1r)) + ≡ 1r · (α zero · f · (h ^ n · 1r) + α (suc zero) · g · (h ^ n · 1r)) · 1r + bigPath = useSolver1 (h ^ n) ∙ cong (h ^ n ·_) p ∙ useSolver2 _ _ _ _ _ + + {- + + We get the following pullback square in CommRings + + R[1/h] → R[1/h][1/f] + ⌟ + ↓ ↓ + + R[1/h][1/g] → R[1/h][1/fg] + + this lifts to a pullback in R-Algebras using PullbackFromCommRing + as all for rings have canonical morphisms coming from R + and all the triangles commute. + + Then using toSheaf.lemma we get the desired square + + R[1/h] → R[1/f] + ⌟ + ↓ ↓ + + R[1/g] → R[1/fg] + + by only providing paths between the corresponding vertices of the square. + These paths are constructed using S⁻¹RAlgCharEquiv, which gives + an equivalent characterization of the universal property of localization + that can be found in e.g. Cor 3.2 of Atiyah-MacDonald + + -} + + theRingCospan = fgCospan R[1/ h ]AsCommRing (f /1) (g /1) + theRingPullback = fgPullback R[1/ h ]AsCommRing (f /1) (g /1) 1∈fgIdeal + + R[1/h][1/f] = InvertingElementsBase.R[1/_] R[1/ h ]AsCommRing (f /1) + R[1/h][1/f]AsCommRing = InvertingElementsBase.R[1/_]AsCommRing R[1/ h ]AsCommRing (f /1) + R[1/h][1/g] = InvertingElementsBase.R[1/_] R[1/ h ]AsCommRing (g /1) + R[1/h][1/g]AsCommRing = InvertingElementsBase.R[1/_]AsCommRing R[1/ h ]AsCommRing (g /1) + R[1/h][1/fg] = InvertingElementsBase.R[1/_] R[1/ h ]AsCommRing ((f /1) · (g /1)) + R[1/h][1/fg]AsCommRing = InvertingElementsBase.R[1/_]AsCommRing + R[1/ h ]AsCommRing ((f /1) · (g /1)) + + open IsRingHom + /1/1AsCommRingHomFG : CommRingHom R' R[1/h][1/fg]AsCommRing + fst /1/1AsCommRingHomFG r = [ [ r , 1r , ∣ 0 , refl ∣ ] , 1r , ∣ 0 , refl ∣ ] + pres0 (snd /1/1AsCommRingHomFG) = refl + pres1 (snd /1/1AsCommRingHomFG) = refl + pres+ (snd /1/1AsCommRingHomFG) x y = cong [_] (≡-× (cong [_] (≡-× + (cong₂ _+_ (useSolver x) (useSolver y)) + (Σ≡Prop (λ _ → isPropPropTrunc) (useSolver 1r)))) + (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·Rid 1r)))) + where + useSolver : ∀ a → a ≡ a · 1r · (1r · 1r) + useSolver = solve R' + pres· (snd /1/1AsCommRingHomFG) x y = cong [_] (≡-× (cong [_] (≡-× refl + (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·Rid 1r))))) + (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·Rid 1r)))) + pres- (snd /1/1AsCommRingHomFG) x = refl + + open Cospan + open Pullback + open RingHoms + isRHomR[1/h]→R[1/h][1/f] : theRingPullback .pbPr₂ ∘r /1AsCommRingHom ≡ /1/1AsCommRingHom f + isRHomR[1/h]→R[1/h][1/f] = RingHom≡ (funExt (λ x → refl)) + + isRHomR[1/h]→R[1/h][1/g] : theRingPullback .pbPr₁ ∘r /1AsCommRingHom ≡ /1/1AsCommRingHom g + isRHomR[1/h]→R[1/h][1/g] = RingHom≡ (funExt (λ x → refl)) + + isRHomR[1/h][1/f]→R[1/h][1/fg] : theRingCospan .s₂ ∘r /1/1AsCommRingHom f ≡ /1/1AsCommRingHomFG + isRHomR[1/h][1/f]→R[1/h][1/fg] = RingHom≡ (funExt + (λ x → cong [_] (≡-× (cong [_] (≡-× (cong (x ·_) (transportRefl 1r) ∙ ·Rid x) + (Σ≡Prop (λ _ → isPropPropTrunc) (cong (1r ·_) (transportRefl 1r) ∙ ·Rid 1r)))) + (Σ≡Prop (λ _ → isPropPropTrunc) (cong (1r ·_) (transportRefl 1r) ∙ ·Rid 1r))))) + + isRHomR[1/h][1/g]→R[1/h][1/fg] : theRingCospan .s₁ ∘r /1/1AsCommRingHom g ≡ /1/1AsCommRingHomFG + isRHomR[1/h][1/g]→R[1/h][1/fg] = RingHom≡ (funExt + (λ x → cong [_] (≡-× (cong [_] (≡-× (cong (x ·_) (transportRefl 1r) ∙ ·Rid x) + (Σ≡Prop (λ _ → isPropPropTrunc) (cong (1r ·_) (transportRefl 1r) ∙ ·Rid 1r)))) + (Σ≡Prop (λ _ → isPropPropTrunc) (cong (1r ·_) (transportRefl 1r) ∙ ·Rid 1r))))) + + + open PullbackFromCommRing R' theRingCospan theRingPullback + /1AsCommRingHom (/1/1AsCommRingHom f) (/1/1AsCommRingHom g) /1/1AsCommRingHomFG + theAlgebraCospan = algCospan isRHomR[1/h]→R[1/h][1/f] + isRHomR[1/h]→R[1/h][1/g] + isRHomR[1/h][1/f]→R[1/h][1/fg] + isRHomR[1/h][1/g]→R[1/h][1/fg] + theAlgebraPullback = algPullback isRHomR[1/h]→R[1/h][1/f] + isRHomR[1/h]→R[1/h][1/g] + isRHomR[1/h][1/f]→R[1/h][1/fg] + isRHomR[1/h][1/g]→R[1/h][1/fg] + + --and the three remaining paths + fPath : theAlgebraCospan .r ≡ R[1/ f ]AsCommAlgebra + fPath = doubleLocCancel f∈√⟨h⟩ + where + open DoubleAlgLoc R' h f + + gPath : theAlgebraCospan .l ≡ R[1/ g ]AsCommAlgebra + gPath = doubleLocCancel g∈√⟨h⟩ + where + open DoubleAlgLoc R' h g + + fgPath : theAlgebraCospan .m ≡ R[1/ (f · g) ]AsCommAlgebra + fgPath = path ∙ doubleLocCancel fg∈√⟨h⟩ + where + open DoubleAlgLoc R' h (f · g) + open CommAlgChar R' + + R[1/h][1/fg]AsCommRing' = InvertingElementsBase.R[1/_]AsCommRing R[1/ h ]AsCommRing ((f · g) /1) + + path : toCommAlg (R[1/h][1/fg]AsCommRing , /1/1AsCommRingHomFG) + ≡ toCommAlg (R[1/h][1/fg]AsCommRing' , /1/1AsCommRingHom (f · g)) + path = cong toCommAlg (ΣPathP (p , q)) + where + eqInR[1/h] : (f /1) · (g /1) ≡ (f · g) /1 + eqInR[1/h] = sym (/1AsCommRingHom .snd .pres· f g) + + p : R[1/h][1/fg]AsCommRing ≡ R[1/h][1/fg]AsCommRing' + p i = InvertingElementsBase.R[1/_]AsCommRing R[1/ h ]AsCommRing (eqInR[1/h] i) + + q : PathP (λ i → CommRingHom R' (p i)) /1/1AsCommRingHomFG (/1/1AsCommRingHom (f · g)) + q = toPathP (RingHom≡ (funExt ( + λ x → cong [_] (≡-× (cong [_] (≡-× (transportRefl _ ∙ transportRefl x) + (Σ≡Prop (λ _ → isPropPropTrunc) (transportRefl 1r)))) + (Σ≡Prop (λ _ → isPropPropTrunc) (transportRefl 1r)))))) diff --git a/Cubical/Algebra/ZariskiLattice/UniversalProperty.agda b/Cubical/Algebra/ZariskiLattice/UniversalProperty.agda new file mode 100644 index 0000000000..3aabf29e0f --- /dev/null +++ b/Cubical/Algebra/ZariskiLattice/UniversalProperty.agda @@ -0,0 +1,310 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.ZariskiLattice.UniversalProperty where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Powerset using (ℙ ; ⊆-refl-consequence) + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm + ; ·-identityʳ to ·ℕ-rid) +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary +open import Cubical.Relation.Binary.Poset + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Ring.Properties +open import Cubical.Algebra.Ring.BigOps +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.BinomialThm +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.RadicalIdeal +open import Cubical.Algebra.RingSolver.Reflection +open import Cubical.Algebra.Semilattice +open import Cubical.Algebra.Lattice +open import Cubical.Algebra.DistLattice +open import Cubical.Algebra.DistLattice.Basis +open import Cubical.Algebra.DistLattice.BigOps +open import Cubical.Algebra.Matrix + +open import Cubical.Algebra.ZariskiLattice.Base + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +private + variable + ℓ ℓ' : Level + + +module _ (R' : CommRing ℓ) (L' : DistLattice ℓ') where + + open CommRingStr (R' .snd) + open RingTheory (CommRing→Ring R') + open Sum (CommRing→Ring R') + open CommRingTheory R' + open Exponentiation R' + + open DistLatticeStr (L' .snd) renaming (is-set to isSetL) + open Join L' + open LatticeTheory (DistLattice→Lattice L') + open Order (DistLattice→Lattice L') + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice L')) + open PosetReasoning IndPoset + open PosetStr (IndPoset .snd) hiding (_≤_) + private + R = fst R' + L = fst L' + + record IsZarMap (d : R → L) : Type (ℓ-max ℓ ℓ') where + constructor iszarmap + + field + pres0 : d 0r ≡ 0l + pres1 : d 1r ≡ 1l + ·≡∧ : ∀ x y → d (x · y) ≡ d x ∧l d y + +≤∨ : ∀ x y → d (x + y) ≤ d x ∨l d y + + ∑≤⋁ : {n : ℕ} (U : FinVec R n) → d (∑ U) ≤ ⋁ λ i → d (U i) + ∑≤⋁ {n = zero} U = ∨lRid _ ∙ pres0 + ∑≤⋁ {n = suc n} U = d (∑ U) ≤⟨ ∨lIdem _ ⟩ + d (U zero + ∑ (U ∘ suc)) ≤⟨ +≤∨ _ _ ⟩ + d (U zero) ∨l d (∑ (U ∘ suc)) ≤⟨ ≤-∨LPres _ _ _ (∑≤⋁ (U ∘ suc)) ⟩ + d (U zero) ∨l ⋁ (d ∘ U ∘ suc) ≤⟨ ∨lIdem _ ⟩ + ⋁ (d ∘ U) ◾ + + d·LCancel : ∀ x y → d (x · y) ≤ d y + d·LCancel x y = subst (λ a → a ≤ d y) (sym (·≡∧ x y)) (∧≤LCancelJoin _ _) + + linearCombination≤LCancel : {n : ℕ} (α β : FinVec R n) + → d (linearCombination R' α β) ≤ ⋁ (d ∘ β) + linearCombination≤LCancel α β = is-trans _ _ _ (∑≤⋁ (λ i → α i · β i)) + (≤-⋁Ext _ _ λ i → d·LCancel (α i) (β i)) + + ZarMapIdem : ∀ (n : ℕ) (x : R) → d (x ^ (suc n)) ≡ d x + ZarMapIdem zero x = ·≡∧ _ _ ∙∙ cong (d x ∧l_) pres1 ∙∙ ∧lRid _ + ZarMapIdem (suc n) x = ·≡∧ _ _ ∙∙ cong (d x ∧l_) (ZarMapIdem n x) ∙∙ ∧lIdem _ + + ZarMapExpIneq : ∀ (n : ℕ) (x : R) → d x ≤ d (x ^ n) + ZarMapExpIneq zero x = cong (d x ∨l_) pres1 ∙∙ 1lRightAnnihilates∨l _ ∙∙ sym pres1 + ZarMapExpIneq (suc n) x = subst (λ y → d x ≤ y) (sym (ZarMapIdem _ x)) (∨lIdem _) + + -- the crucial lemma about "Zariski maps" + open CommIdeal R' + open RadicalIdeal R' + open isCommIdeal + private + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + + ZarMapRadicalIneq : ∀ {n : ℕ} (α : FinVec R n) (x : R) + → x ∈ √ ⟨ α ⟩ → d x ≤ ⋁ (d ∘ α) + ZarMapRadicalIneq α x = PT.elim (λ _ → isSetL _ _) + (uncurry (λ n → (PT.elim (λ _ → isSetL _ _) (uncurry (curriedHelper n))))) + where + curriedHelper : (n : ℕ) (β : FinVec R _) + → x ^ n ≡ linearCombination R' β α → d x ≤ ⋁ (d ∘ α) + curriedHelper n β xⁿ≡∑βα = d x ≤⟨ ZarMapExpIneq n x ⟩ + d (x ^ n) + ≤⟨ subst (λ y → y ≤ ⋁ (d ∘ α)) (sym (cong d xⁿ≡∑βα)) (linearCombination≤LCancel β α) ⟩ + ⋁ (d ∘ α) ◾ + +module ZarLatUniversalProp (R' : CommRing ℓ) where + open CommRingStr (snd R') + open RingTheory (CommRing→Ring R') + open Sum (CommRing→Ring R') + open CommRingTheory R' + open Exponentiation R' + open BinomialThm R' + open CommIdeal R' + open RadicalIdeal R' + open isCommIdeal + open ProdFin R' + + open ZarLat R' + open IsZarMap + + private + R = fst R' + ⟨_⟩ : {n : ℕ} → FinVec R n → CommIdeal + ⟨ V ⟩ = ⟨ V ⟩[ R' ] + + + D : R → ZL + D x = [ 1 , replicateFinVec 1 x ] -- λ x → √⟨x⟩ + + isZarMapD : IsZarMap R' ZariskiLattice D + pres0 isZarMapD = eq/ _ _ (cong √ (0FGIdeal _ ∙ sym (emptyFGIdeal _ _))) + pres1 isZarMapD = refl + ·≡∧ isZarMapD x y = cong {B = λ _ → ZL} (λ U → [ 1 , U ]) (Length1··Fin x y) + +≤∨ isZarMapD x y = eq/ _ _ (cong √ (CommIdeal≡Char (inclOfFGIdeal _ 3Vec ⟨ 2Vec ⟩ 3Vec⊆2Vec) + (inclOfFGIdeal _ 2Vec ⟨ 3Vec ⟩ 2Vec⊆3Vec))) + where + 2Vec = replicateFinVec 1 x ++Fin replicateFinVec 1 y + 3Vec = replicateFinVec 1 (x + y) ++Fin (replicateFinVec 1 x ++Fin replicateFinVec 1 y) + + 3Vec⊆2Vec : ∀ (i : Fin 3) → 3Vec i ∈ ⟨ 2Vec ⟩ + 3Vec⊆2Vec zero = ⟨ 2Vec ⟩ .snd .+Closed (indInIdeal _ _ zero) (indInIdeal _ _ (suc zero)) + 3Vec⊆2Vec (suc zero) = indInIdeal _ _ zero + 3Vec⊆2Vec (suc (suc zero)) = indInIdeal _ _ (suc zero) + + 2Vec⊆3Vec : ∀ (i : Fin 2) → 2Vec i ∈ ⟨ 3Vec ⟩ + 2Vec⊆3Vec zero = indInIdeal _ _ (suc zero) + 2Vec⊆3Vec (suc zero) = indInIdeal _ _ (suc (suc zero)) + + + -- defintion of the universal property + hasZarLatUniversalProp : (L : DistLattice ℓ') (D : R → fst L) + → IsZarMap R' L D + → Type _ + hasZarLatUniversalProp {ℓ' = ℓ'} L D _ = ∀ (L' : DistLattice ℓ') (d : R → fst L') + → IsZarMap R' L' d + → ∃![ χ ∈ DistLatticeHom L L' ] (fst χ) ∘ D ≡ d + + isPropZarLatUniversalProp : (L : DistLattice ℓ') (D : R → fst L) (isZarMapD : IsZarMap R' L D) + → isProp (hasZarLatUniversalProp L D isZarMapD) + isPropZarLatUniversalProp L D isZarMapD = isPropΠ3 (λ _ _ _ → isPropIsContr) + + ZLHasUniversalProp : hasZarLatUniversalProp ZariskiLattice D isZarMapD + ZLHasUniversalProp L' d isZarMapd = (χ , funExt χcomp) , χunique + where + open DistLatticeStr (snd L') renaming (is-set to isSetL) + open LatticeTheory (DistLattice→Lattice L') + open Join L' + open IsLatticeHom + L = fst L' + + χ : DistLatticeHom ZariskiLattice L' + fst χ = SQ.rec isSetL (λ (_ , α) → ⋁ (d ∘ α)) + λ (_ , α) (_ , β) → curriedHelper α β + where + curriedHelper : {n m : ℕ} (α : FinVec R n) (β : FinVec R m) + → √ ⟨ α ⟩ ≡ √ ⟨ β ⟩ → ⋁ (d ∘ α) ≡ ⋁ (d ∘ β) + curriedHelper α β √⟨α⟩≡√⟨β⟩ = is-antisym _ _ ineq1 ineq2 + where + open Order (DistLattice→Lattice L') + open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice L')) + open PosetReasoning IndPoset + open PosetStr (IndPoset .snd) hiding (_≤_) + + incl1 : √ ⟨ α ⟩ ⊆ √ ⟨ β ⟩ + incl1 = ⊆-refl-consequence _ _ (cong fst √⟨α⟩≡√⟨β⟩) .fst + + ineq1 : ⋁ (d ∘ α) ≤ ⋁ (d ∘ β) + ineq1 = ⋁IsMax (d ∘ α) (⋁ (d ∘ β)) + λ i → ZarMapRadicalIneq isZarMapd β (α i) (√FGIdealCharLImpl α ⟨ β ⟩ incl1 i) + + incl2 : √ ⟨ β ⟩ ⊆ √ ⟨ α ⟩ + incl2 = ⊆-refl-consequence _ _ (cong fst √⟨α⟩≡√⟨β⟩) .snd + + ineq2 : ⋁ (d ∘ β) ≤ ⋁ (d ∘ α) + ineq2 = ⋁IsMax (d ∘ β) (⋁ (d ∘ α)) + λ i → ZarMapRadicalIneq isZarMapd α (β i) (√FGIdealCharLImpl β ⟨ α ⟩ incl2 i) + + + pres0 (snd χ) = refl + pres1 (snd χ) = ∨lRid _ ∙ isZarMapd .pres1 + pres∨l (snd χ) = elimProp2 (λ _ _ → isSetL _ _) (uncurry (λ n α → uncurry (curriedHelper n α))) + where + curriedHelper : (n : ℕ) (α : FinVec R n) (m : ℕ) (β : FinVec R m) + → ⋁ (d ∘ (α ++Fin β)) ≡ ⋁ (d ∘ α) ∨l ⋁ (d ∘ β) + curriedHelper zero α _ β = sym (∨lLid _) + curriedHelper (suc n) α _ β = + ⋁ (d ∘ (α ++Fin β)) ≡⟨ refl ⟩ + d (α zero) ∨l ⋁ (d ∘ ((α ∘ suc) ++Fin β)) + + ≡⟨ cong (d (α zero) ∨l_) (curriedHelper _ (α ∘ suc) _ β) ⟩ + + d (α zero) ∨l (⋁ (d ∘ α ∘ suc) ∨l ⋁ (d ∘ β)) + ≡⟨ ∨lAssoc _ _ _ ⟩ + + ⋁ (d ∘ α) ∨l ⋁ (d ∘ β) ∎ + + pres∧l (snd χ) = elimProp2 (λ _ _ → isSetL _ _) (uncurry (λ n α → uncurry (curriedHelper n α))) + where + -- have to repeat this one here so the termination checker won't complain + oldHelper : (n : ℕ) (α : FinVec R n) (m : ℕ) (β : FinVec R m) + → ⋁ (d ∘ (α ++Fin β)) ≡ ⋁ (d ∘ α) ∨l ⋁ (d ∘ β) + oldHelper zero α _ β = sym (∨lLid _) + oldHelper (suc n) α _ β = cong (d (α zero) ∨l_) (oldHelper _ (α ∘ suc) _ β) ∙ ∨lAssoc _ _ _ + + curriedHelper : (n : ℕ) (α : FinVec R n) (m : ℕ) (β : FinVec R m) + → ⋁ (d ∘ (α ··Fin β)) ≡ ⋁ (d ∘ α) ∧l ⋁ (d ∘ β) + curriedHelper zero α _ β = sym (0lLeftAnnihilates∧l _) + curriedHelper (suc n) α _ β = + ⋁ (d ∘ (α ··Fin β)) ≡⟨ refl ⟩ + ⋁ (d ∘ ((λ j → α zero · β j) ++Fin ((α ∘ suc) ··Fin β))) + + ≡⟨ oldHelper _ (λ j → α zero · β j) _ ((α ∘ suc) ··Fin β) ⟩ + + ⋁ (d ∘ (λ j → α zero · β j)) ∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β)) + + ≡⟨ cong (_∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β))) (⋁Ext (λ j → isZarMapd .·≡∧ (α zero) (β j))) ⟩ + + ⋁ (λ j → d (α zero) ∧l d (β j)) ∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β)) + + ≡⟨ cong (_∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β))) (sym (⋁Meetrdist _ _)) ⟩ + + (d (α zero) ∧l ⋁ (d ∘ β)) ∨l ⋁ (d ∘ ((α ∘ suc) ··Fin β)) + + ≡⟨ cong ((d (α zero) ∧l ⋁ (d ∘ β)) ∨l_) (curriedHelper _ (α ∘ suc) _ β) ⟩ + + (d (α zero) ∧l ⋁ (d ∘ β)) ∨l (⋁ (d ∘ α ∘ suc) ∧l ⋁ (d ∘ β)) + + ≡⟨ sym (∧lRdist∨l _ _ _) ⟩ + + ⋁ (d ∘ α) ∧l ⋁ (d ∘ β) ∎ + + + χcomp : ∀ (f : R) → χ .fst (D f) ≡ d f + χcomp f = ∨lRid (d f) + + χunique : (y : Σ[ χ' ∈ DistLatticeHom ZariskiLattice L' ] fst χ' ∘ D ≡ d) + → (χ , funExt χcomp) ≡ y + χunique (χ' , χ'∘D≡d) = Σ≡Prop (λ _ → isSetΠ (λ _ → isSetL) _ _) (LatticeHom≡f _ _ + (funExt (elimProp (λ _ → isSetL _ _) (uncurry uniqHelper)))) + where + uniqHelper : (n : ℕ) (α : FinVec R n) → fst χ [ n , α ] ≡ fst χ' [ n , α ] + uniqHelper zero _ = sym (cong (λ α → fst χ' [ 0 , α ]) (funExt (λ ())) ∙ χ' .snd .pres0) + uniqHelper (suc n) α = + ⋁ (d ∘ α) ≡⟨ refl ⟩ + d (α zero) ∨l ⋁ (d ∘ α ∘ suc) + + ≡⟨ cong (d (α zero) ∨l_) (uniqHelper n (α ∘ suc)) ⟩ -- the inductive step + + d (α zero) ∨l fst χ' [ n , α ∘ suc ] + + ≡⟨ cong (_∨l fst χ' [ n , α ∘ suc ]) (sym (funExt⁻ χ'∘D≡d (α zero))) ⟩ + + fst χ' (D (α zero)) ∨l fst χ' [ n , α ∘ suc ] + + ≡⟨ sym (χ' .snd .pres∨l _ _) ⟩ + + fst χ' (D (α zero) ∨z [ n , α ∘ suc ]) + + ≡⟨ cong (λ β → fst χ' [ suc n , β ]) (funExt (λ { zero → refl ; (suc i) → refl })) ⟩ + + fst χ' [ suc n , α ] ∎ + + + -- the map induced by applying the universal property to the Zariski lattice + -- itself is the identity hom + ZLUniversalPropCorollary : ZLHasUniversalProp ZariskiLattice D isZarMapD .fst .fst + ≡ idDistLatticeHom ZariskiLattice + ZLUniversalPropCorollary = cong fst + (ZLHasUniversalProp ZariskiLattice D isZarMapD .snd + (idDistLatticeHom ZariskiLattice , refl)) diff --git a/Cubical/Categories/Category/Base.agda b/Cubical/Categories/Category/Base.agda index 36606c14dd..6f31d4ad82 100644 --- a/Cubical/Categories/Category/Base.agda +++ b/Cubical/Categories/Category/Base.agda @@ -4,6 +4,7 @@ module Cubical.Categories.Category.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset private variable @@ -88,3 +89,14 @@ _⋆_ (C ^op) f g = g ⋆⟨ C ⟩ f ⋆IdR (C ^op) = C .⋆IdL ⋆Assoc (C ^op) f g h = sym (C .⋆Assoc _ _ _) isSetHom (C ^op) = C .isSetHom + + +ΣPropCat : (C : Category ℓ ℓ') (P : ℙ (ob C)) → Category ℓ ℓ' +ob (ΣPropCat C P) = Σ[ x ∈ ob C ] x ∈ P +Hom[_,_] (ΣPropCat C P) x y = C [ fst x , fst y ] +id (ΣPropCat C P) = id C +_⋆_ (ΣPropCat C P) = _⋆_ C +⋆IdL (ΣPropCat C P) = ⋆IdL C +⋆IdR (ΣPropCat C P) = ⋆IdR C +⋆Assoc (ΣPropCat C P) = ⋆Assoc C +isSetHom (ΣPropCat C P) = isSetHom C diff --git a/Cubical/Categories/DistLatticeSheaf.agda b/Cubical/Categories/DistLatticeSheaf.agda index c131a1de04..01af11a6c8 100644 --- a/Cubical/Categories/DistLatticeSheaf.agda +++ b/Cubical/Categories/DistLatticeSheaf.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Categories.DistLatticeSheaf where open import Cubical.Foundations.Prelude @@ -44,8 +44,9 @@ module _ (L : DistLattice ℓ) (C : Category ℓ' ℓ'') (T : Terminal C) where 𝟙 : ob C 𝟙 = terminalOb C T - DLCat : Category ℓ ℓ - DLCat = DistLatticeCategory L + private + DLCat : Category ℓ ℓ + DLCat = DistLatticeCategory L open Category DLCat @@ -68,28 +69,26 @@ module _ (L : DistLattice ℓ) (C : Category ℓ' ℓ'') (T : Terminal C) where {- - x ∧ y ----→ y + x ∧ y ----→ x | | | sq | V V - x ----→ x ∨ y + y ----→ x ∨ y -} sq : (x y : L .fst) → hom-∧₂ x y ⋆ hom-∨₂ x y ≡ hom-∧₁ x y ⋆ hom-∨₁ x y sq x y = is-prop-valued (x ∧l y) (x ∨l y) (hom-∧₂ x y ⋆ hom-∨₂ x y) (hom-∧₁ x y ⋆ hom-∨₁ x y) {- - F(x ∨ y) ----→ F(y) + F(x ∨ y) ----→ F(x) | | | Fsq | V V - F(x) ------→ F(x ∧ y) + F(y) ------→ F(x ∧ y) -} Fsq : (F : DLPreSheaf) (x y : L .fst) → F .F-hom (hom-∨₂ x y) ⋆⟨ C ⟩ F .F-hom (hom-∧₂ x y) ≡ F .F-hom (hom-∨₁ x y) ⋆⟨ C ⟩ F .F-hom (hom-∧₁ x y) - Fsq F x y = sym (F-seq F (hom-∨₂ x y) (hom-∧₂ x y)) - ∙∙ cong (F .F-hom) (sq x y) - ∙∙ F-seq F (hom-∨₁ x y) (hom-∧₁ x y) + Fsq F x y = F-square F (sq x y) isDLSheaf : (F : DLPreSheaf) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') isDLSheaf F = (F-ob F 0l ≡ 𝟙) @@ -100,20 +99,81 @@ module _ (L : DistLattice ℓ) (C : Category ℓ' ℓ'') (T : Terminal C) where DLSheaf = Σ[ F ∈ DLPreSheaf ] isDLSheaf F -module Lemma1 (L : DistLattice ℓ) (C : Category ℓ' ℓ'') (T : Terminal C) (L' : ℙ (fst L)) (hB : IsBasis L L') where +module SheafOnBasis (L : DistLattice ℓ) (C : Category ℓ' ℓ'') (T : Terminal C) + (L' : ℙ (fst L)) (hB : IsBasis L L') where - open Category hiding (_⋆_) - open Functor - open DistLatticeStr (snd L) - open IsBasis hB + open Category + open Functor - isDLBasisSheaf : (F : DLPreSheaf L C T) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - isDLBasisSheaf F = (F-ob F 0l ≡ 𝟙 L C T) - × ((x y : L .fst) → x ∈ L' → y ∈ L' → isPullback C _ _ _ (Fsq L C T F x y)) + open DistLatticeStr ⦃...⦄ + open SemilatticeStr ⦃...⦄ + open IsBasis hB - DLBasisSheaf : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - DLBasisSheaf = Σ[ F ∈ DLPreSheaf L C T ] isDLBasisSheaf F + private + DLCat = DistLatticeCategory L + BasisCat = ΣPropCat DLCat L' + DLBasisPreSheaf = Functor (BasisCat ^op) C + + -- to avoid writing 𝟙 L C T + 1c : ob C + 1c = terminalOb C T + + instance + _ = snd L + _ = snd (Basis→MeetSemilattice L L' hB) + + + module condSquare (x y : ob BasisCat) (x∨y∈L' : fst x ∨l fst y ∈ L') where + + private + x∨y : ob BasisCat -- = Σ[ x ∈ L ] (x ∈ L') + x∨y = fst x ∨l fst y , x∨y∈L' + + {- + x ∧ y ----→ x + | | + | sq | + V V + y ----→ x ∨ y + but as a square in BasisCat + -} + Bsq : seq' BasisCat {x = x · y} {y = y} {z = x∨y} (hom-∧₂ L C T (fst x) (fst y)) + (hom-∨₂ L C T (fst x) (fst y)) + ≡ seq' BasisCat {x = x · y} {y = x} {z = x∨y} (hom-∧₁ L C T (fst x) (fst y)) + (hom-∨₁ L C T (fst x) (fst y)) + Bsq = sq L C T (fst x) (fst y) + + {- + F(x ∨ y) ----→ F(x) + | | + | Fsq | + V V + F(y) ------→ F(x ∧ y) + + square in C but now F is only presheaf on BasisCat + -} + BFsq : (F : DLBasisPreSheaf) + → seq' C {x = F .F-ob x∨y} {y = F .F-ob y} {z = F .F-ob (x · y)} + (F .F-hom (hom-∨₂ L C T (fst x) (fst y))) + (F .F-hom (hom-∧₂ L C T (fst x) (fst y))) + ≡ seq' C {x = F .F-ob x∨y} {y = F .F-ob x} {z = F .F-ob (x · y)} + (F .F-hom (hom-∨₁ L C T (fst x) (fst y))) + (F .F-hom (hom-∧₁ L C T (fst x) (fst y))) + BFsq F = F-square F Bsq + + + -- TODO: check that this is equivalent to the functor + -- preserving terminal objects and pullbacks + isDLBasisSheaf : DLBasisPreSheaf → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + isDLBasisSheaf F = ((0∈L' : 0l ∈ L') → F .F-ob (0l , 0∈L') ≡ 1c) + × ((x y : ob BasisCat) (x∨y∈L' : fst x ∨l fst y ∈ L') + → isPullback C _ _ _ (BFsq x y x∨y∈L' F)) + where + open condSquare + + DLBasisSheaf : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + DLBasisSheaf = Σ[ F ∈ DLBasisPreSheaf ] isDLBasisSheaf F -- To prove the statement we probably need that C is: -- 1. univalent diff --git a/Cubical/Categories/Functor/Base.agda b/Cubical/Categories/Functor/Base.agda index 871dd6d72e..e1b3a730f0 100644 --- a/Cubical/Categories/Functor/Base.agda +++ b/Cubical/Categories/Functor/Base.agda @@ -28,6 +28,13 @@ record Functor (C : Category ℓC ℓC') (D : Category ℓD ℓD') : isFaithful = (x y : _) (f g : C [ x , y ]) → F-hom f ≡ F-hom g → f ≡ g isEssentiallySurj = (d : D .ob) → Σ[ c ∈ C .ob ] CatIso D (F-ob c) d + F-square : {x y u v : C .ob} + {f : C [ x , y ]} {g : C [ x , u ]} + {h : C [ u , v ]} {k : C [ y , v ]} + → f ⋆⟨ C ⟩ k ≡ g ⋆⟨ C ⟩ h + → (F-hom f) ⋆⟨ D ⟩ (F-hom k) ≡ (F-hom g) ⋆⟨ D ⟩ (F-hom h) + F-square Csquare = sym (F-seq _ _) ∙∙ cong F-hom Csquare ∙∙ F-seq _ _ + private variable ℓ ℓ' : Level diff --git a/Cubical/Categories/Instances/Algebras.agda b/Cubical/Categories/Instances/Algebras.agda new file mode 100644 index 0000000000..4998c436fd --- /dev/null +++ b/Cubical/Categories/Instances/Algebras.agda @@ -0,0 +1,28 @@ +{-# OPTIONS --safe #-} +module Cubical.Categories.Instances.Algebras where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra + +open import Cubical.Categories.Category + +open Category +open AlgebraHoms + +private + variable + ℓ ℓ' : Level + +module _ (R : Ring ℓ) where + AlgebrasCategory : Category (ℓ-suc (ℓ-max ℓ ℓ')) (ℓ-max ℓ ℓ') + ob AlgebrasCategory = Algebra R _ + Hom[_,_] AlgebrasCategory = AlgebraHom + id AlgebrasCategory {A} = idAlgebraHom A + _⋆_ AlgebrasCategory = compAlgebraHom + ⋆IdL AlgebrasCategory = compIdAlgebraHom + ⋆IdR AlgebrasCategory = idCompAlgebraHom + ⋆Assoc AlgebrasCategory = compAssocAlgebraHom + isSetHom AlgebrasCategory = isSetAlgebraHom _ _ diff --git a/Cubical/Categories/Instances/CommAlgebras.agda b/Cubical/Categories/Instances/CommAlgebras.agda new file mode 100644 index 0000000000..8bb81ec3ec --- /dev/null +++ b/Cubical/Categories/Instances/CommAlgebras.agda @@ -0,0 +1,365 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Categories.Instances.CommAlgebras where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.Instances.Unit + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Limits.Terminal +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Instances.CommRings + +open import Cubical.HITs.PropositionalTruncation + +open Category hiding (_∘_) renaming (_⋆_ to _⋆c_) +open CommAlgebraHoms +open Cospan +open Pullback + +private + variable + ℓ ℓ' ℓ'' : Level + +module _ (R : CommRing ℓ) where + CommAlgebrasCategory : Category (ℓ-suc (ℓ-max ℓ ℓ')) (ℓ-max ℓ ℓ') + ob CommAlgebrasCategory = CommAlgebra R _ + Hom[_,_] CommAlgebrasCategory = CommAlgebraHom + id CommAlgebrasCategory {A} = idCommAlgebraHom A + _⋆c_ CommAlgebrasCategory {A} {B} {C} = compCommAlgebraHom A B C + ⋆IdL CommAlgebrasCategory {A} {B} = compIdCommAlgebraHom {A = A} {B} + ⋆IdR CommAlgebrasCategory {A} {B} = idCompCommAlgebraHom {A = A} {B} + ⋆Assoc CommAlgebrasCategory {A} {B} {C} {D} = compAssocCommAlgebraHom {A = A} {B} {C} {D} + isSetHom CommAlgebrasCategory = isSetAlgebraHom _ _ + + TerminalCommAlgebra : Terminal (CommAlgebrasCategory {ℓ' = ℓ'}) + fst TerminalCommAlgebra = UnitCommAlgebra R + fst (fst (snd TerminalCommAlgebra A)) = λ _ → tt* + snd (fst (snd TerminalCommAlgebra A)) = makeIsAlgebraHom + refl (λ _ _ → refl) (λ _ _ → refl) (λ _ _ → refl) + snd (snd TerminalCommAlgebra A) f = AlgebraHom≡ (funExt (λ _ → refl)) + +module PullbackFromCommRing (R : CommRing ℓ) + (commRingCospan : Cospan (CommRingsCategory {ℓ = ℓ})) + (commRingPB : Pullback _ commRingCospan) + (f₁ : CommRingHom R (commRingPB .pbOb)) + (f₂ : CommRingHom R (commRingCospan .r)) + (f₃ : CommRingHom R (commRingCospan .l)) + (f₄ : CommRingHom R (commRingCospan .m)) + where + + open AlgebraHoms + open CommAlgChar R + open CommAlgebraStr ⦃...⦄ + + private + CommAlgCat = CommAlgebrasCategory {ℓ = ℓ} R {ℓ' = ℓ} + + A = commRingPB .pbOb + B = commRingCospan .r + C = commRingCospan .l + D = commRingCospan .m + + g₁ = commRingPB .pbPr₂ + g₂ = commRingPB .pbPr₁ + g₃ = commRingCospan .s₂ + g₄ = commRingCospan .s₁ + + {- + g₁ + A → B + ⌟ + g₂ ↓ ↓ g₃ + + C → D + g₄ + -} + + open RingHoms + univPropCommRingWithHomHom : (isRHom₁ : isCommRingWithHomHom (A , f₁) (B , f₂) g₁) + (isRHom₂ : isCommRingWithHomHom (A , f₁) (C , f₃) g₂) + (isRHom₃ : isCommRingWithHomHom (B , f₂) (D , f₄) g₃) + (isRHom₄ : isCommRingWithHomHom (C , f₃) (D , f₄) g₄) + (E,f₅ : CommRingWithHom) + (h₂ : CommRingWithHomHom E,f₅ (C , f₃)) + (h₁ : CommRingWithHomHom E,f₅ (B , f₂)) + → g₄ ∘r fst h₂ ≡ g₃ ∘r fst h₁ + → ∃![ h₃ ∈ CommRingWithHomHom E,f₅ (A , f₁) ] + (fst h₂ ≡ g₂ ∘r fst h₃) × (fst h₁ ≡ g₁ ∘r fst h₃) + univPropCommRingWithHomHom isRHom₁ isRHom₂ isRHom₃ isRHom₄ + (E , f₅) (h₂ , comm₂) (h₁ , comm₁) squareComm = + ((h₃ , h₃∘f₅≡f₁) , h₂≡g₂∘h₃ , h₁≡g₁∘h₃) + , λ h₃' → Σ≡Prop (λ _ → isProp× (isSetRingHom _ _ _ _) (isSetRingHom _ _ _ _)) + (Σ≡Prop (λ _ → isSetRingHom _ _ _ _) + (cong fst (commRingPB .univProp h₂ h₁ squareComm .snd + (h₃' .fst .fst , h₃' .snd .fst , h₃' .snd .snd)))) + where + h₃ : CommRingHom E A + h₃ = commRingPB .univProp h₂ h₁ squareComm .fst .fst + h₂≡g₂∘h₃ : h₂ ≡ g₂ ∘r h₃ + h₂≡g₂∘h₃ = commRingPB .univProp h₂ h₁ squareComm .fst .snd .fst + h₁≡g₁∘h₃ : h₁ ≡ g₁ ∘r h₃ + h₁≡g₁∘h₃ = commRingPB .univProp h₂ h₁ squareComm .fst .snd .snd + + -- the crucial obervation: + -- we can apply the universal property to maps R → A + fgSquare : g₄ ∘r f₃ ≡ g₃ ∘r f₂ + fgSquare = isRHom₄ ∙ sym isRHom₃ + + h₃∘f₅≡f₁ : h₃ ∘r f₅ ≡ f₁ + h₃∘f₅≡f₁ = cong fst (isContr→isProp (commRingPB .univProp f₃ f₂ fgSquare) + (h₃ ∘r f₅ , triangle1 , triangle2) (f₁ , (sym isRHom₂) , sym isRHom₁)) + where + triangle1 : f₃ ≡ g₂ ∘r (h₃ ∘r f₅) + triangle1 = sym comm₂ ∙∙ cong (compRingHom f₅) h₂≡g₂∘h₃ ∙∙ sym (compAssocRingHom f₅ h₃ g₂) + + triangle2 : f₂ ≡ g₁ ∘r (h₃ ∘r f₅) + triangle2 = sym comm₁ ∙∙ cong (compRingHom f₅) h₁≡g₁∘h₃ ∙∙ sym (compAssocRingHom f₅ h₃ g₁) + + + + algCospan : (isRHom₁ : isCommRingWithHomHom (A , f₁) (B , f₂) g₁) + (isRHom₂ : isCommRingWithHomHom (A , f₁) (C , f₃) g₂) + (isRHom₃ : isCommRingWithHomHom (B , f₂) (D , f₄) g₃) + (isRHom₄ : isCommRingWithHomHom (C , f₃) (D , f₄) g₄) + → Cospan CommAlgCat + l (algCospan isRHom₁ isRHom₂ isRHom₃ isRHom₄) = toCommAlg (C , f₃) + m (algCospan isRHom₁ isRHom₂ isRHom₃ isRHom₄) = toCommAlg (D , f₄) + r (algCospan isRHom₁ isRHom₂ isRHom₃ isRHom₄) = toCommAlg (B , f₂) + s₁ (algCospan isRHom₁ isRHom₂ isRHom₃ isRHom₄) = toCommAlgebraHom _ _ g₄ isRHom₄ + s₂ (algCospan isRHom₁ isRHom₂ isRHom₃ isRHom₄) = toCommAlgebraHom _ _ g₃ isRHom₃ + + + algPullback : (isRHom₁ : isCommRingWithHomHom (A , f₁) (B , f₂) g₁) + (isRHom₂ : isCommRingWithHomHom (A , f₁) (C , f₃) g₂) + (isRHom₃ : isCommRingWithHomHom (B , f₂) (D , f₄) g₃) + (isRHom₄ : isCommRingWithHomHom (C , f₃) (D , f₄) g₄) + → Pullback CommAlgCat (algCospan isRHom₁ isRHom₂ isRHom₃ isRHom₄) + pbOb (algPullback isRHom₁ isRHom₂ isRHom₃ isRHom₄) = toCommAlg (A , f₁) + pbPr₁ (algPullback isRHom₁ isRHom₂ isRHom₃ isRHom₄) = toCommAlgebraHom _ _ g₂ isRHom₂ + pbPr₂ (algPullback isRHom₁ isRHom₂ isRHom₃ isRHom₄) = toCommAlgebraHom _ _ g₁ isRHom₁ + pbCommutes (algPullback isRHom₁ isRHom₂ isRHom₃ isRHom₄) = + AlgebraHom≡ (cong fst (pbCommutes commRingPB)) + univProp (algPullback isRHom₁ isRHom₂ isRHom₃ isRHom₄) {d = F} h₂' h₁' g₄∘h₂'≡g₃∘h₁' = + (k , kComm₂ , kComm₁) , uniqueness + where + E = fromCommAlg F .fst + f₅ = fromCommAlg F .snd + + h₁ : CommRingHom E B + fst h₁ = fst h₁' + IsRingHom.pres0 (snd h₁) = IsAlgebraHom.pres0 (snd h₁') + IsRingHom.pres1 (snd h₁) = IsAlgebraHom.pres1 (snd h₁') + IsRingHom.pres+ (snd h₁) = IsAlgebraHom.pres+ (snd h₁') + IsRingHom.pres· (snd h₁) = IsAlgebraHom.pres· (snd h₁') + IsRingHom.pres- (snd h₁) = IsAlgebraHom.pres- (snd h₁') + + h₁comm : h₁ ∘r f₅ ≡ f₂ + h₁comm = RingHom≡ (funExt (λ x → IsAlgebraHom.pres⋆ (snd h₁') x 1a + ∙∙ cong (fst f₂ x ·_) (IsAlgebraHom.pres1 (snd h₁')) + ∙∙ ·Rid _)) + where + instance + _ = snd F + _ = snd (toCommAlg (B , f₂)) + + h₂ : CommRingHom E C + fst h₂ = fst h₂' + IsRingHom.pres0 (snd h₂) = IsAlgebraHom.pres0 (snd h₂') + IsRingHom.pres1 (snd h₂) = IsAlgebraHom.pres1 (snd h₂') + IsRingHom.pres+ (snd h₂) = IsAlgebraHom.pres+ (snd h₂') + IsRingHom.pres· (snd h₂) = IsAlgebraHom.pres· (snd h₂') + IsRingHom.pres- (snd h₂) = IsAlgebraHom.pres- (snd h₂') + + h₂comm : h₂ ∘r f₅ ≡ f₃ + h₂comm = RingHom≡ (funExt (λ x → IsAlgebraHom.pres⋆ (snd h₂') x 1a + ∙∙ cong (fst f₃ x ·_) (IsAlgebraHom.pres1 (snd h₂')) + ∙∙ ·Rid _)) + where + instance + _ = snd F + _ = snd (toCommAlg (C , f₃)) + + commRingSquare : g₄ ∘r h₂ ≡ g₃ ∘r h₁ + commRingSquare = RingHom≡ (funExt (λ x → funExt⁻ (cong fst g₄∘h₂'≡g₃∘h₁') x)) + + uniqueH₃ : ∃![ h₃ ∈ CommRingWithHomHom (E , f₅) (A , f₁) ] (h₂ ≡ g₂ ∘r fst h₃) × (h₁ ≡ g₁ ∘r fst h₃) + uniqueH₃ = univPropCommRingWithHomHom isRHom₁ isRHom₂ isRHom₃ isRHom₄ + (E , f₅) (h₂ , h₂comm) (h₁ , h₁comm) commRingSquare + + h₃ : CommRingHom E A + h₃ = uniqueH₃ .fst .fst .fst + + h₃comm = uniqueH₃ .fst .fst .snd + + k : CommAlgebraHom F (toCommAlg (A , f₁)) + fst k = fst h₃ + IsAlgebraHom.pres0 (snd k) = IsRingHom.pres0 (snd h₃) + IsAlgebraHom.pres1 (snd k) = IsRingHom.pres1 (snd h₃) + IsAlgebraHom.pres+ (snd k) = IsRingHom.pres+ (snd h₃) + IsAlgebraHom.pres· (snd k) = IsRingHom.pres· (snd h₃) + IsAlgebraHom.pres- (snd k) = IsRingHom.pres- (snd h₃) + IsAlgebraHom.pres⋆ (snd k) = + λ r y → sym (fst f₁ r · fst h₃ y ≡⟨ cong (_· fst h₃ y) (sym (funExt⁻ (cong fst h₃comm) r)) ⟩ + fst h₃ (fst f₅ r) · fst h₃ y ≡⟨ sym (IsRingHom.pres· (snd h₃) _ _) ⟩ + fst h₃ (fst f₅ r · y) ≡⟨ refl ⟩ + fst h₃ ((r ⋆ 1a) · y) ≡⟨ cong (fst h₃) (⋆-lassoc _ _ _) ⟩ + fst h₃ (r ⋆ (1a · y)) ≡⟨ cong (λ x → fst h₃ (r ⋆ x)) (·Lid y) ⟩ + fst h₃ (r ⋆ y) ∎) + where + instance + _ = snd F + _ = (toCommAlg (A , f₁) .snd) + + kComm₂ : h₂' ≡ toCommAlgebraHom _ _ g₂ isRHom₂ ∘a k + kComm₂ = AlgebraHom≡ (cong fst (uniqueH₃ .fst .snd .fst)) + + kComm₁ : h₁' ≡ toCommAlgebraHom _ _ g₁ isRHom₁ ∘a k + kComm₁ = AlgebraHom≡ (cong fst (uniqueH₃ .fst .snd .snd)) + + uniqueness : (y : Σ[ k' ∈ CommAlgebraHom F (toCommAlg (A , f₁)) ] + (h₂' ≡ toCommAlgebraHom _ _ g₂ isRHom₂ ∘a k') + × (h₁' ≡ toCommAlgebraHom _ _ g₁ isRHom₁ ∘a k')) + → (k , kComm₂ , kComm₁) ≡ y + uniqueness (k' , k'Comm₂ , k'Comm₁) = Σ≡Prop (λ _ → isProp× (isSetAlgebraHom _ _ _ _) + (isSetAlgebraHom _ _ _ _)) + (AlgebraHom≡ (cong (fst ∘ fst ∘ fst) uniqHelper)) + where + h₃' : CommRingHom E A + fst h₃' = fst k' + IsRingHom.pres0 (snd h₃') = IsAlgebraHom.pres0 (snd k') + IsRingHom.pres1 (snd h₃') = IsAlgebraHom.pres1 (snd k') + IsRingHom.pres+ (snd h₃') = IsAlgebraHom.pres+ (snd k') + IsRingHom.pres· (snd h₃') = IsAlgebraHom.pres· (snd k') + IsRingHom.pres- (snd h₃') = IsAlgebraHom.pres- (snd k') + + h₃'IsRHom : h₃' ∘r f₅ ≡ f₁ + h₃'IsRHom = RingHom≡ (funExt (λ x → IsAlgebraHom.pres⋆ (snd k') x 1a + ∙ cong (fst f₁ x ·_) (IsAlgebraHom.pres1 (snd k')) + ∙ ·Rid (fst f₁ x))) + where + instance + _ = snd F + _ = (toCommAlg (A , f₁) .snd) + + h₃'Comm₂ : h₂ ≡ g₂ ∘r h₃' + h₃'Comm₂ = RingHom≡ (cong fst k'Comm₂) + + h₃'Comm₁ : h₁ ≡ g₁ ∘r h₃' + h₃'Comm₁ = RingHom≡ (cong fst k'Comm₁) + + -- basically saying that h₃≡h₃' but we have to give all the commuting triangles + uniqHelper : uniqueH₃ .fst ≡ ((h₃' , h₃'IsRHom) , h₃'Comm₂ , h₃'Comm₁) + uniqHelper = uniqueH₃ .snd ((h₃' , h₃'IsRHom) , h₃'Comm₂ , h₃'Comm₁) + + + +module PreSheafFromUniversalProp (C : Category ℓ ℓ') (P : ob C → Type ℓ) + {R : CommRing ℓ''} (𝓕 : Σ (ob C) P → CommAlgebra R ℓ'') + (uniqueHom : ∀ x y → C [ fst x , fst y ] → isContr (CommAlgebraHom (𝓕 y) (𝓕 x))) + where + + private + ∥P∥ : ℙ (ob C) + ∥P∥ x = ∥ P x ∥ , isPropPropTrunc + ΣC∥P∥Cat = ΣPropCat C ∥P∥ + CommAlgCat = CommAlgebrasCategory {ℓ = ℓ''} R {ℓ' = ℓ''} + + 𝓕UniqueEquiv : (x : ob C) (p q : P x) → isContr (CommAlgebraEquiv (𝓕 (x , p)) (𝓕 (x , q))) + 𝓕UniqueEquiv x = contrCommAlgebraHom→contrCommAlgebraEquiv (curry 𝓕 x) λ p q → uniqueHom _ _ (id C) + + theMap : (x : ob C) → ∥ P x ∥ → CommAlgebra R ℓ'' + theMap x = recPT→CommAlgebra (curry 𝓕 x) (λ p q → 𝓕UniqueEquiv x p q .fst) + λ p q r → 𝓕UniqueEquiv x p r .snd _ + + theAction : (x y : ob C) → C [ x , y ] + → (p : ∥ P x ∥) (q : ∥ P y ∥) → isContr (CommAlgebraHom (theMap y q) (theMap x p)) + theAction _ _ f = elim2 (λ _ _ → isPropIsContr) λ _ _ → uniqueHom _ _ f + + open Functor + universalPShf : Functor (ΣC∥P∥Cat ^op) CommAlgCat + F-ob universalPShf = uncurry theMap + F-hom universalPShf {x = x} {y = y} f = theAction _ _ f (y .snd) (x. snd) .fst + F-id universalPShf {x = x} = theAction (x .fst) (x .fst) (id C) (x .snd) (x .snd) .snd _ + F-seq universalPShf {x = x} {z = z} f g = theAction _ _ (g ⋆⟨ C ⟩ f) (z .snd) (x .snd) .snd _ + + + -- a big transport to help verifying the sheaf property + module toSheaf (x y u v : ob ΣC∥P∥Cat) + {f : C [ v .fst , y . fst ]} {g : C [ v .fst , u .fst ]} + {h : C [ u .fst , x . fst ]} {k : C [ y .fst , x .fst ]} + (Csquare : g ⋆⟨ C ⟩ h ≡ f ⋆⟨ C ⟩ k) + {- + v → y + ↓ ↓ + u → x + -} + (AlgCospan : Cospan CommAlgCat) + (AlgPB : Pullback _ AlgCospan) + (p₁ : AlgPB .pbOb ≡ F-ob universalPShf x) (p₂ : AlgCospan .l ≡ F-ob universalPShf u) + (p₃ : AlgCospan .r ≡ F-ob universalPShf y) (p₄ : AlgCospan .m ≡ F-ob universalPShf v) + where + + private + -- just: 𝓕 k ⋆ 𝓕 f ≡ 𝓕 h ⋆ 𝓕 g + inducedSquare : seq' CommAlgCat {x = F-ob universalPShf x} + {y = F-ob universalPShf u} + {z = F-ob universalPShf v} + (F-hom universalPShf h) (F-hom universalPShf g) + ≡ seq' CommAlgCat {x = F-ob universalPShf x} + {y = F-ob universalPShf y} + {z = F-ob universalPShf v} + (F-hom universalPShf k) (F-hom universalPShf f) + inducedSquare = F-square universalPShf Csquare + + f' = F-hom universalPShf {x = y} {y = v} f + g' = F-hom universalPShf {x = u} {y = v} g + h' = F-hom universalPShf {x = x} {y = u} h + k' = F-hom universalPShf {x = x} {y = y} k + + gPathP : PathP (λ i → CommAlgCat [ p₂ i , p₄ i ]) (AlgCospan .s₁) g' + gPathP = toPathP (sym (theAction _ _ g (v .snd) (u .snd) .snd _)) + + fPathP : PathP (λ i → CommAlgCat [ p₃ i , p₄ i ]) (AlgCospan .s₂) f' + fPathP = toPathP (sym (theAction _ _ f (v .snd) (y .snd) .snd _)) + + kPathP : PathP (λ i → CommAlgCat [ p₁ i , p₃ i ]) (AlgPB .pbPr₂) k' + kPathP = toPathP (sym (theAction _ _ k (y .snd) (x .snd) .snd _)) + + hPathP : PathP (λ i → CommAlgCat [ p₁ i , p₂ i ]) (AlgPB .pbPr₁) h' + hPathP = toPathP (sym (theAction _ _ h (u .snd) (x .snd) .snd _)) + + fgCospan : Cospan CommAlgCat + l fgCospan = F-ob universalPShf u + m fgCospan = F-ob universalPShf v + r fgCospan = F-ob universalPShf y + s₁ fgCospan = g' + s₂ fgCospan = f' + + cospanPath : AlgCospan ≡ fgCospan + l (cospanPath i) = p₂ i + m (cospanPath i) = p₄ i + r (cospanPath i) = p₃ i + s₁ (cospanPath i) = gPathP i + s₂ (cospanPath i) = fPathP i + + squarePathP : PathP (λ i → hPathP i ⋆⟨ CommAlgCat ⟩ gPathP i ≡ kPathP i ⋆⟨ CommAlgCat ⟩ fPathP i) + (AlgPB .pbCommutes) inducedSquare + squarePathP = toPathP (CommAlgCat .isSetHom _ _ _ _) + + abstract + lemma : isPullback CommAlgCat fgCospan {c = F-ob universalPShf x} h' k' inducedSquare + lemma = transport (λ i → isPullback CommAlgCat (cospanPath i) {c = p₁ i} + (hPathP i) (kPathP i) (squarePathP i)) + (AlgPB .univProp) diff --git a/Cubical/Categories/Instances/CommRings.agda b/Cubical/Categories/Instances/CommRings.agda index 1944b479cb..157f4d843e 100644 --- a/Cubical/Categories/Instances/CommRings.agda +++ b/Cubical/Categories/Instances/CommRings.agda @@ -5,6 +5,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure +open import Cubical.Foundations.Powerset open import Cubical.Data.Unit open import Cubical.Data.Sigma @@ -15,15 +16,19 @@ open import Cubical.Algebra.CommRing.FiberedProduct open import Cubical.Algebra.CommRing.Instances.Unit open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Instances.Functors open import Cubical.Categories.Limits.Terminal open import Cubical.Categories.Limits.Pullback +open import Cubical.HITs.PropositionalTruncation + open Category hiding (_∘_) open CommRingHoms private variable - ℓ : Level + ℓ ℓ' ℓ'' : Level CommRingsCategory : Category (ℓ-suc ℓ) ℓ ob CommRingsCategory = CommRing _ @@ -37,7 +42,7 @@ isSetHom CommRingsCategory = isSetRingHom _ _ TerminalCommRing : Terminal {ℓ-suc ℓ-zero} CommRingsCategory fst TerminalCommRing = UnitCommRing -fst (fst (snd TerminalCommRing y)) _ = tt +fst (fst (snd TerminalCommRing y)) _ = tt* snd (fst (snd TerminalCommRing y)) = makeIsRingHom refl (λ _ _ → refl) (λ _ _ → refl) snd (snd TerminalCommRing y) f = RingHom≡ (funExt (λ _ → refl)) diff --git a/Cubical/Categories/Instances/Lattice.agda b/Cubical/Categories/Instances/Lattice.agda index f491afd6ac..3139dcc649 100644 --- a/Cubical/Categories/Instances/Lattice.agda +++ b/Cubical/Categories/Instances/Lattice.agda @@ -10,5 +10,6 @@ open import Cubical.Categories.Instances.Semilattice open Category + LatticeCategory : ∀ {ℓ} (L : Lattice ℓ) → Category ℓ ℓ -LatticeCategory L = SemilatticeCategory (Lattice→JoinSemilattice L) +LatticeCategory L = JoinSemilatticeCategory (Lattice→JoinSemilattice L) diff --git a/Cubical/Categories/Instances/Semilattice.agda b/Cubical/Categories/Instances/Semilattice.agda index 61c03c677d..e698f39ae2 100644 --- a/Cubical/Categories/Instances/Semilattice.agda +++ b/Cubical/Categories/Instances/Semilattice.agda @@ -10,10 +10,16 @@ open import Cubical.Categories.Instances.Poset open Category + module _ {ℓ} (L : Semilattice ℓ) where - -- more convenient than working with meet-semilattices - -- as joins are limits open JoinSemilattice L - SemilatticeCategory : Category ℓ ℓ - SemilatticeCategory = PosetCategory IndPoset + JoinSemilatticeCategory : Category ℓ ℓ + JoinSemilatticeCategory = PosetCategory IndPoset + + +module _ {ℓ} (L : Semilattice ℓ) where + open MeetSemilattice L + + MeetSemilatticeCategory : Category ℓ ℓ + MeetSemilatticeCategory = PosetCategory IndPoset diff --git a/Cubical/Categories/Instances/Sets.agda b/Cubical/Categories/Instances/Sets.agda index ff9483450c..962a0e06f1 100644 --- a/Cubical/Categories/Instances/Sets.agda +++ b/Cubical/Categories/Instances/Sets.agda @@ -6,7 +6,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Data.Unit -open import Cubical.Data.Sigma using (ΣPathP) +open import Cubical.Data.Sigma open import Cubical.Categories.Category open import Cubical.Categories.Functor open import Cubical.Categories.NaturalTransformation @@ -75,62 +75,18 @@ Iso→CatIso is .cInv = is .inv Iso→CatIso is .sec = funExt λ b → is .rightInv b -- is .rightInv Iso→CatIso is .ret = funExt λ b → is .leftInv b -- is .rightInv - - -- SET is complete --- notes: --- didn't need to restrict to *finite* diagrams , why is that required in Set theoretic? --- didn't use coinduction here because Agda didn't like me referencing 'cone' frome 'up' (termination check) - -open NatTrans - -isCompleteSET : ∀ {ℓJ ℓJ'} → complete' {ℓJ = ℓJ} {ℓJ'} (SET (ℓ-max ℓJ ℓJ')) -isCompleteSET J K = record - { head = head' - ; islim = record { cone = cone' ; up = up' } } - where - -- the limit is defined as the Set of all cones with head Unit - head' = Cone K (Unit* , isOfHLevelLift 2 isSetUnit) , isSetNatTrans - - -- the legs are defined by taking a cone to its component at j - cone' : Cone K head' - cone' .N-ob j μ = (μ ⟦ j ⟧) tt* - -- Naturality follows from naturality of the Unit cone - cone' .N-hom {x = i} {j} f - = funExt λ μ → (μ ⟦ j ⟧) tt* - ≡[ i ]⟨ (μ .N-hom f i) tt* ⟩ - (K ⟪ f ⟫) ((μ ⟦ i ⟧) tt*) - ∎ - - -- Given another cone α, we want a unique function f from α → cone' which factors it - -- factorization property enforces that (cone' ⟦ j ⟧ ● f) ≡ α ⟦ j ⟧ - -- but cone' ⟦ j ⟧ simply takes the jth component the output Cone K Unit from f - -- so this enforces that for all x ∈ A, (f x) ⟦ j ⟧ ≡ α ⟦ j ⟧ x - -- this determines the *only* possible factoring morphism - up' : ∀ {A} (α : Cone K A) → cone' uniquelyFactors α - up' {A} α = (f , fact) , unique - where - f : fst A → Cone K (Unit* , isOfHLevelLift 2 isSetUnit) - f x = natTrans (λ j _ → α .N-ob j x) - (λ {m} {n} f → funExt λ μ i → α .N-hom f i x) - - fact : α ≡ (f ◼ cone') - fact = makeNatTransPath refl -- I LOVE DEFINITIONAL EQUALITY - - unique : (τ : cone' factors α) → (f , fact) ≡ τ - unique (f' , fact') = ΣPathP (f≡f' , fact≡fact') - where - f≡f' : f ≡ f' - f≡f' = funExt λ x → makeNatTransPath (funExt λ _ → sym eq2) - where - -- the factorization property enforces that f' must have the same behavior as f - eq1 : ∀ {x j} → ((cone' ⟦ j ⟧) (f' x)) ≡ (α ⟦ j ⟧) x - eq1 {x} {j} i = ((fact' (~ i)) ⟦ j ⟧) x - - eq2 : ∀ {x j} → (f' x) ⟦ j ⟧ ≡ λ _ → (α ⟦ j ⟧) x -- = (f x) ⟦ j ⟧ - eq2 {x} {j} = funExt λ _ → eq1 - - -- follows from Set having homsets - fact≡fact' : PathP (λ i → α ≡ ((f≡f' i) ◼ cone')) fact fact' - fact≡fact' = isOfHLevel→isOfHLevelDep 1 (λ β → isSetNatTrans α β) fact fact' λ i → (f≡f' i) ◼ cone' +open LimCone +open Cone + +completeSET : ∀ {ℓJ ℓJ'} → Limits {ℓJ} {ℓJ'} (SET (ℓ-max ℓJ ℓJ')) +lim (completeSET J D) = Cone D (Unit* , isOfHLevelLift 2 isSetUnit) , isSetCone D _ +coneOut (limCone (completeSET J D)) j e = coneOut e j tt* +coneOutCommutes (limCone (completeSET J D)) j i e = coneOutCommutes e j i tt* +univProp (completeSET J D) c cc = + uniqueExists + (λ x → cone (λ v _ → coneOut cc v x) (λ e i _ → coneOutCommutes cc e i x)) + (λ _ → funExt (λ _ → refl)) + (λ x → isPropIsConeMor cc (limCone (completeSET J D)) x) + (λ x hx → funExt (λ d → cone≡ λ u → funExt (λ _ → sym (funExt⁻ (hx u) d)))) diff --git a/Cubical/Categories/Limits/Limits.agda b/Cubical/Categories/Limits/Limits.agda index b96f83d70b..5d2b54845d 100644 --- a/Cubical/Categories/Limits/Limits.agda +++ b/Cubical/Categories/Limits/Limits.agda @@ -1,84 +1,139 @@ -{-# OPTIONS --safe #-} +-- Limits. +-- Heavily inspired by https://github.com/UniMath/UniMath/blob/master/UniMath/CategoryTheory/limits/graphs/limits.v +-- (except we're using categories instead of graphs to index limits) +{-# OPTIONS --safe #-} module Cubical.Categories.Limits.Limits where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma.Base + open import Cubical.Categories.Category open import Cubical.Categories.Functor -open import Cubical.Categories.NaturalTransformation -private - variable - ℓJ ℓJ' ℓC ℓC' : Level - ℓ ℓ' ℓ'' : Level +open import Cubical.HITs.PropositionalTruncation.Base + +module _ {ℓJ ℓJ' ℓC ℓC' : Level} {J : Category ℓJ ℓJ'} {C : Category ℓC ℓC'} where + open Category + open Functor + + private + ℓ = ℓ-max (ℓ-max (ℓ-max ℓJ ℓJ') ℓC) ℓC' + + record Cone (D : Functor J C) (c : ob C) : Type (ℓ-max (ℓ-max ℓJ ℓJ') ℓC') where + constructor cone + + field + coneOut : (v : ob J) → C [ c , F-ob D v ] + coneOutCommutes : {u v : ob J} (e : J [ u , v ]) → + coneOut u ⋆⟨ C ⟩ D .F-hom e ≡ coneOut v + + open Cone + + cone≡ : {D : Functor J C} {c : ob C} → {c1 c2 : Cone D c} + → ((v : ob J) → coneOut c1 v ≡ coneOut c2 v) → c1 ≡ c2 + coneOut (cone≡ h i) v = h v i + coneOutCommutes (cone≡ {D} {_} {c1} {c2} h i) {u} {v} e = + isProp→PathP (λ j → isSetHom C (h u j ⋆⟨ C ⟩ D .F-hom e) (h v j)) + (coneOutCommutes c1 e) (coneOutCommutes c2 e) i + + -- TODO: can we automate this a bit? + isSetCone : (D : Functor J C) (c : ob C) → isSet (Cone D c) + isSetCone D c = isSetRetract toConeΣ fromConeΣ (λ _ → refl) + (isSetΣSndProp (isSetΠ (λ _ → isSetHom C)) + (λ _ → isPropImplicitΠ2 (λ _ _ → isPropΠ (λ _ → isSetHom C _ _)))) + where + ConeΣ = Σ[ f ∈ ((v : ob J) → C [ c , F-ob D v ]) ] + ({u v : ob J} (e : J [ u , v ]) → f u ⋆⟨ C ⟩ D .F-hom e ≡ f v) + + toConeΣ : Cone D c → ConeΣ + fst (toConeΣ x) = coneOut x + snd (toConeΣ x) = coneOutCommutes x + + fromConeΣ : ConeΣ → Cone D c + coneOut (fromConeΣ x) = fst x + coneOutCommutes (fromConeΣ x) = snd x + + isConeMor : {c1 c2 : ob C} {D : Functor J C} + (cc1 : Cone D c1) (cc2 : Cone D c2) + → C [ c1 , c2 ] → Type (ℓ-max ℓJ ℓC') + isConeMor cc1 cc2 f = (v : ob J) → f ⋆⟨ C ⟩ coneOut cc2 v ≡ coneOut cc1 v + + isPropIsConeMor : {c1 c2 : ob C} {D : Functor J C} + (cc1 : Cone D c1) (cc2 : Cone D c2) (f : C [ c1 , c2 ]) + → isProp (isConeMor cc1 cc2 f) + isPropIsConeMor cc1 cc2 f = isPropΠ (λ _ → isSetHom C _ _) + + isLimCone : (D : Functor J C) (c0 : ob C) → Cone D c0 → Type ℓ + isLimCone D c0 cc0 = (c : ob C) → (cc : Cone D c) + → ∃![ f ∈ C [ c , c0 ] ] isConeMor cc cc0 f -module _ {J : Category ℓJ ℓJ'} - {C : Category ℓC ℓC'} where + isPropIsLimCone : (D : Functor J C) (c0 : ob C) (cc0 : Cone D c0) + → isProp (isLimCone D c0 cc0) + isPropIsLimCone D c0 cc0 = isPropΠ2 λ _ _ → isProp∃! + + record LimCone (D : Functor J C) : Type ℓ where + constructor climCone + + field + lim : ob C + limCone : Cone D lim + univProp : isLimCone D lim limCone + + limOut : (v : ob J) → C [ lim , D .F-ob v ] + limOut = coneOut limCone + + limOutCommutes : {u v : ob J} (e : J [ u , v ]) + → limOut u ⋆⟨ C ⟩ D .F-hom e ≡ limOut v + limOutCommutes = coneOutCommutes limCone + + limArrow : (c : ob C) → (cc : Cone D c) → C [ c , lim ] + limArrow c cc = univProp c cc .fst .fst + + limArrowCommutes : (c : ob C) → (cc : Cone D c) (u : ob J) + → limArrow c cc ⋆⟨ C ⟩ limOut u ≡ coneOut cc u + limArrowCommutes c cc = univProp c cc .fst .snd + + limArrowUnique : (c : ob C) → (cc : Cone D c) (k : C [ c , lim ]) + → isConeMor cc limCone k → limArrow c cc ≡ k + limArrowUnique c cc k hk = cong fst (univProp c cc .snd (k , hk)) + + -- TODO: define limOfArrows + +-- A category is complete if it has all limits +Limits : {ℓJ ℓJ' ℓC ℓC' : Level} → Category ℓC ℓC' → Type _ +Limits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') → (D : Functor J C) → LimCone D + +hasLimits : {ℓJ ℓJ' ℓC ℓC' : Level} → Category ℓC ℓC' → Type _ +hasLimits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') → (D : Functor J C) → ∥ LimCone D ∥ + +-- Limits of a specific shape J in a category C +LimitsOfShape : {ℓJ ℓJ' ℓC ℓC' : Level} → Category ℓJ ℓJ' → Category ℓC ℓC' → Type _ +LimitsOfShape J C = (D : Functor J C) → LimCone D + + +-- Preservation of limits + +module _ {ℓJ ℓJ' ℓC1 ℓC1' ℓC2 ℓC2' : Level} + {C1 : Category ℓC1 ℓC1'} {C2 : Category ℓC2 ℓC2'} + (F : Functor C1 C2) where open Category open Functor - open NatTrans - - -- functor which sends all objects to c and all - -- morphisms to the identity - constF : (c : C .ob) → Functor J C - constF c .F-ob _ = c - constF c .F-hom _ = C .id - constF c .F-id = refl - constF c .F-seq _ _ = sym (C .⋆IdL _) - - - -- a cone is a natural transformation from the constant functor at x to K - Cone : (K : Functor J C) → C .ob → Type _ - Cone K x = NatTrans (constF x) K - - - module _ {K : Functor J C} where - - -- precomposes a cone with a morphism - _◼_ : ∀ {d c : C .ob} - → (f : C [ d , c ]) - → Cone K c - → Cone K d - (f ◼ μ) .N-ob x = f ⋆⟨ C ⟩ μ ⟦ x ⟧ - (f ◼ μ) .N-hom {x = x} {y} k - = C .id ⋆⟨ C ⟩ (f ⋆⟨ C ⟩ μ ⟦ y ⟧) - ≡⟨ C .⋆IdL _ ⟩ - f ⋆⟨ C ⟩ μ ⟦ y ⟧ - ≡⟨ cong (λ v → f ⋆⟨ C ⟩ v) (sym (C .⋆IdL _)) ⟩ - f ⋆⟨ C ⟩ (C .id ⋆⟨ C ⟩ μ ⟦ y ⟧) - ≡⟨ cong (λ v → f ⋆⟨ C ⟩ v) (μ .N-hom k) ⟩ - f ⋆⟨ C ⟩ (μ ⟦ x ⟧ ⋆⟨ C ⟩ K ⟪ k ⟫) - ≡⟨ sym (C .⋆Assoc _ _ _) ⟩ - f ⋆⟨ C ⟩ μ ⟦ x ⟧ ⋆⟨ C ⟩ K ⟪ k ⟫ - ∎ - - -- μ factors ν if there's a morphism such that the natural transformation - -- from precomposing it with ν gives you back μ - _factors_ : {u v : C .ob} (μ : Cone K u) (ν : Cone K v) → Type _ - _factors_ {u} {v} μ ν = Σ[ mor ∈ C [ v , u ] ] ν ≡ (mor ◼ μ) - - -- μ uniquely factors ν if the factor from above is contractible - _uniquelyFactors_ : {u v : C .ob} (μ : Cone K u) (ν : Cone K v) → Type _ - _uniquelyFactors_ {u} {v} μ ν = isContr (μ factors ν) - - module _ (K : Functor J C) where - - -- a Limit is a cone such that all cones are uniquely factored by it - record isLimit (head : C .ob) : Type (ℓ-max (ℓ-max ℓJ ℓJ') (ℓ-max ℓC ℓC')) where - constructor islimit - field - cone : Cone K head - -- TODO: change this to terminal object in category of Cones? - up : ∀ {v} (ν : Cone K v) → cone uniquelyFactors ν - - record Limit : Type (ℓ-max (ℓ-max ℓJ ℓJ') (ℓ-max ℓC ℓC')) where - field - head : C .ob - islim : isLimit head - --- a Category is complete if it has all limits -complete' : {ℓJ ℓJ' : Level} (C : Category ℓC ℓC') → Type _ -complete' {ℓJ = ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') (K : Functor J C) → Limit K - -complete : (C : Category ℓC ℓC') → Typeω -complete C = ∀ {ℓJ ℓJ'} → complete' {ℓJ = ℓJ} {ℓJ'} C + open Cone + + private + ℓ = ℓ-max (ℓ-max (ℓ-max (ℓ-max (ℓ-max ℓJ ℓJ') ℓC1) ℓC1') ℓC2) ℓC2' + + mapCone : {J : Category ℓJ ℓJ'} (D : Functor J C1) {x : ob C1} (ccx : Cone D x) → Cone (funcComp F D) (F .F-ob x) + coneOut (mapCone D ccx) v = F .F-hom (coneOut ccx v) + coneOutCommutes (mapCone D ccx) e = + sym (F-seq F (coneOut ccx _) (D ⟪ e ⟫)) ∙ cong (F .F-hom) (coneOutCommutes ccx e) + + preservesLimit : {J : Category ℓJ ℓJ'} (D : Functor J C1) + → (L : ob C1) → Cone D L → Type ℓ + preservesLimit D L ccL = + isLimCone D L ccL → isLimCone (funcComp F D) (F .F-ob L) (mapCone D ccL) + + -- TODO: prove that right adjoints preserve limits diff --git a/Cubical/Categories/Limits/Pullback.agda b/Cubical/Categories/Limits/Pullback.agda index 47c665c671..ae1e7a333f 100644 --- a/Cubical/Categories/Limits/Pullback.agda +++ b/Cubical/Categories/Limits/Pullback.agda @@ -6,10 +6,12 @@ open import Cubical.Foundations.HLevels open import Cubical.HITs.PropositionalTruncation.Base open import Cubical.Data.Sigma +open import Cubical.Data.Unit open import Cubical.Categories.Category open import Cubical.Categories.Functor open import Cubical.Categories.Instances.Cospan +open import Cubical.Categories.Limits.Limits private variable @@ -90,11 +92,14 @@ module _ (C : Category ℓ ℓ') where hasPullbacks = ∥ Pullbacks ∥ --- TODO: finish the following show that this definition of Pullback --- is equivalent to the Cospan limit +-- Pullbacks from limits module _ {C : Category ℓ ℓ'} where open Category C open Functor + open Pullback + open LimCone + open Cone + open Cospan Cospan→Func : Cospan C → Functor CospanCat C Cospan→Func (cospan l m r f g) .F-ob ⓪ = l @@ -115,3 +120,34 @@ module _ {C : Category ℓ ℓ'} where Cospan→Func (cospan l m r f g) .F-seq {②} {②} {②} φ ψ = sym (⋆IdL _) Cospan→Func (cospan l m r f g) .F-seq {②} {②} {①} φ ψ = sym (⋆IdL _) Cospan→Func (cospan l m r f g) .F-seq {②} {①} {①} φ ψ = sym (⋆IdR _) + + LimitsOfShapeCospanCat→Pullbacks : LimitsOfShape CospanCat C → Pullbacks C + pbOb (LimitsOfShapeCospanCat→Pullbacks H cspn) = lim (H (Cospan→Func cspn)) + pbPr₁ (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOut (H (Cospan→Func cspn)) ⓪ + pbPr₂ (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOut (H (Cospan→Func cspn)) ② + pbCommutes (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOutCommutes (H (Cospan→Func cspn)) {v = ①} tt + ∙ sym (limOutCommutes (H (Cospan→Func cspn)) tt) + univProp (LimitsOfShapeCospanCat→Pullbacks H cspn) {d = d} h k H' = + uniqueExists (limArrow (H (Cospan→Func cspn)) d cc) + ( sym (limArrowCommutes (H (Cospan→Func cspn)) d cc ⓪) + , sym (limArrowCommutes (H (Cospan→Func cspn)) d cc ②)) + (λ _ → isProp× (isSetHom _ _) (isSetHom _ _)) + λ a' ha' → limArrowUnique (H (Cospan→Func cspn)) d cc a' + (λ { ⓪ → sym (ha' .fst) + ; ① → cong (a' ⋆_) (sym (limOutCommutes (H (Cospan→Func cspn)) {⓪} {①} tt)) + ∙∙ sym (⋆Assoc _ _ _) + ∙∙ cong (_⋆ cspn .s₁) (sym (ha' .fst)) + ; ② → sym (ha' .snd) }) + where + cc : Cone (Cospan→Func cspn) d + coneOut cc ⓪ = h + coneOut cc ① = h ⋆ cspn .s₁ + coneOut cc ② = k + coneOutCommutes cc {⓪} {⓪} e = ⋆IdR h + coneOutCommutes cc {⓪} {①} e = refl + coneOutCommutes cc {①} {①} e = ⋆IdR _ + coneOutCommutes cc {②} {①} e = sym H' + coneOutCommutes cc {②} {②} e = ⋆IdR k + + Limits→Pullbacks : Limits C → Pullbacks C + Limits→Pullbacks H = LimitsOfShapeCospanCat→Pullbacks (H CospanCat) diff --git a/Cubical/Categories/NaturalTransformation/Properties.agda b/Cubical/Categories/NaturalTransformation/Properties.agda index 9e46123041..a79e1cc619 100644 --- a/Cubical/Categories/NaturalTransformation/Properties.agda +++ b/Cubical/Categories/NaturalTransformation/Properties.agda @@ -22,6 +22,7 @@ open NatIso open NatTrans open Category open Functor +open Iso module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where private @@ -44,9 +45,9 @@ module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where module NatTransP where module _ {F G : Functor C D} where - open Iso -- same as Sigma version + NatTransΣ : Type (ℓ-max (ℓ-max ℓC ℓC') ℓD') NatTransΣ = Σ[ ob ∈ ((x : C .ob) → D [(F .F-ob x) , (G .F-ob x)]) ] ({x y : _ } (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (ob y) ≡ (ob x) ⋆ᴰ (G .F-hom f)) @@ -56,7 +57,8 @@ module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where NatTransIsoΣ .rightInv _ = refl NatTransIsoΣ .leftInv _ = refl - NatTrans≡Σ = ua (isoToEquiv NatTransIsoΣ) + NatTrans≡Σ : NatTrans F G ≡ NatTransΣ + NatTrans≡Σ = isoToPath NatTransIsoΣ -- introducing paths NatTrans-≡-intro : ∀ {αo βo : N-ob-Type F G} @@ -66,7 +68,8 @@ module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where → PathP (λ i → ({x y : C .ob} (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (p i y) ≡ (p i x) ⋆ᴰ (G .F-hom f))) αh βh → natTrans {F = F} {G} αo αh ≡ natTrans βo βh NatTrans-≡-intro p q i = natTrans (p i) (q i) - module _ {F G : Functor C D} {α β : NatTrans F G} where + + module _ {F G : Functor C D} {α β : NatTrans F G} where open Iso private αOb = α .N-ob @@ -92,55 +95,8 @@ module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where open NatTransP -- if the target category has hom Sets, then any natural transformation is a set - isSetNatTrans : ∀ {F G : Functor C D} - → isSet (NatTrans F G) - isSetNatTrans {F} {G} α β p1 p2 i = comp (λ i → NTPath≡PathΣ {F = F} {G} {α} {β} (~ i)) - (λ j → λ {(i = i0) → transport-filler NTPath≡PathΣ p1 (~ j) ; - (i = i1) → transport-filler NTPath≡PathΣ p2 (~ j)}) - (p1Σ≡p2Σ i) - where - αOb = α .N-ob - βOb = β .N-ob - αHom = α .N-hom - βHom = β .N-hom - - -- convert to sigmas so we can reason about constituent paths separately - p1Σ : Σ[ p ∈ (αOb ≡ βOb) ] - (PathP (λ i → ({x y : _} (f : _) → F ⟪ f ⟫ ⋆ᴰ (p i y) ≡ (p i x) ⋆ᴰ G ⟪ f ⟫)) - αHom - βHom) - p1Σ = transport NTPath≡PathΣ p1 - - p2Σ : Σ[ p ∈ (αOb ≡ βOb) ] - (PathP (λ i → ({x y : _} (f : _) → F ⟪ f ⟫ ⋆ᴰ (p i y) ≡ (p i x) ⋆ᴰ G ⟪ f ⟫)) - αHom - βHom) - p2Σ = transport NTPath≡PathΣ p2 - - -- type aliases - typeN-ob = (x : C .ob) → D [(F .F-ob x) , (G .F-ob x)] - typeN-hom : typeN-ob → Type _ - typeN-hom ϕ = {x y : C .ob} (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (ϕ y) ≡ (ϕ x) ⋆ᴰ (G .F-hom f) - - -- the Ob function is a set - isSetN-ob : isSet ((x : C .ob) → D [(F .F-ob x) , (G .F-ob x)]) - isSetN-ob = isOfHLevelΠ 2 λ _ → D .isSetHom - - -- the Hom function is a set - isSetN-hom : (ϕ : typeN-ob) → isSet (typeN-hom ϕ) - isSetN-hom γ = isProp→isSet (isPropImplicitΠ2 λ x y → isPropΠ λ f → D .isSetHom _ _) - - -- in fact it's a dependent Set, which we need because N-hom depends on N-ob - isSetN-homP : isOfHLevelDep 2 (λ γ → {x y : C .ob} (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (γ y) ≡ (γ x) ⋆ᴰ (G .F-hom f)) - isSetN-homP = isOfHLevel→isOfHLevelDep 2 isSetN-hom - - -- components of the equality - p1Ob≡p2Ob : fst p1Σ ≡ fst p2Σ - p1Ob≡p2Ob = isSetN-ob _ _ (fst p1Σ) (fst p2Σ) - - p1Hom≡p2Hom : PathP (λ i → PathP (λ j → typeN-hom (p1Ob≡p2Ob i j)) αHom βHom) - (snd p1Σ) (snd p2Σ) - p1Hom≡p2Hom = isSetN-homP _ _ (snd p1Σ) (snd p2Σ) p1Ob≡p2Ob - - p1Σ≡p2Σ : p1Σ ≡ p2Σ - p1Σ≡p2Σ = ΣPathP (p1Ob≡p2Ob , p1Hom≡p2Hom) + isSetNatTrans : {F G : Functor C D} → isSet (NatTrans F G) + isSetNatTrans = + isSetRetract (fun NatTransIsoΣ) (inv NatTransIsoΣ) (leftInv NatTransIsoΣ) + (isSetΣSndProp (isSetΠ (λ _ → isSetHom D)) + (λ _ → isPropImplicitΠ2 (λ _ _ → isPropΠ (λ _ → isSetHom D _ _)))) diff --git a/Cubical/Data/Bool/Base.agda b/Cubical/Data/Bool/Base.agda index a9d4cc7de6..96b0595fbd 100644 --- a/Cubical/Data/Bool/Base.agda +++ b/Cubical/Data/Bool/Base.agda @@ -1,8 +1,6 @@ {-# OPTIONS --safe #-} module Cubical.Data.Bool.Base where -open import Cubical.Core.Everything - open import Cubical.Foundations.Prelude open import Cubical.Data.Empty @@ -10,7 +8,6 @@ open import Cubical.Data.Sum.Base open import Cubical.Data.Unit.Base open import Cubical.Relation.Nullary.Base -open import Cubical.Relation.Nullary.DecidableEq -- Obtain the booleans open import Agda.Builtin.Bool public diff --git a/Cubical/Data/Bool/Properties.agda b/Cubical/Data/Bool/Properties.agda index b521a530e1..4548a7eb48 100644 --- a/Cubical/Data/Bool/Properties.agda +++ b/Cubical/Data/Bool/Properties.agda @@ -13,9 +13,9 @@ open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence open import Cubical.Foundations.Pointed +open import Cubical.Data.Sum open import Cubical.Data.Bool.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sum hiding (rec) +open import Cubical.Data.Empty as Empty open import Cubical.Data.Sigma open import Cubical.HITs.PropositionalTruncation hiding (rec) @@ -74,10 +74,10 @@ false≢true p = subst (λ b → if b then ⊥ else Bool) p true ¬true→false : (x : Bool) → ¬ x ≡ true → x ≡ false ¬true→false false _ = refl -¬true→false true p = rec (p refl) +¬true→false true p = Empty.rec (p refl) ¬false→true : (x : Bool) → ¬ x ≡ false → x ≡ true -¬false→true false p = rec (p refl) +¬false→true false p = Empty.rec (p refl) ¬false→true true _ = refl not≢const : ∀ x → ¬ not x ≡ x @@ -195,9 +195,9 @@ module BoolReflection where categorize P | inspect true false p q = subst (λ b → transport P ≡ transport (⊕-Path b)) (sym p) (notLemma P p q) categorize P | inspect false false p q - = rec (¬transportNot P false (q ∙ sym p)) + = Empty.rec (¬transportNot P false (q ∙ sym p)) categorize P | inspect true true p q - = rec (¬transportNot P false (q ∙ sym p)) + = Empty.rec (¬transportNot P false (q ∙ sym p)) ⊕-complete : ∀ P → P ≡ ⊕-Path (transport P false) ⊕-complete P = isInjectiveTransport (categorize P) @@ -226,24 +226,25 @@ Iso.leftInv IsoBool→∙ (f , p) = ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) --- relation to hProp - +-- import here to avoid conflicts open import Cubical.Data.Unit +-- relation to hProp + BoolProp≃BoolProp* : {a : Bool} → Bool→Type a ≃ Bool→Type* {ℓ} a BoolProp≃BoolProp* {a = true} = Unit≃Unit* -BoolProp≃BoolProp* {a = false} = uninhabEquiv rec rec* +BoolProp≃BoolProp* {a = false} = uninhabEquiv Empty.rec Empty.rec* Bool→PropInj : (a b : Bool) → Bool→Type a ≃ Bool→Type b → a ≡ b Bool→PropInj true true _ = refl -Bool→PropInj true false p = rec (p .fst tt) -Bool→PropInj false true p = rec (invEq p tt) +Bool→PropInj true false p = Empty.rec (p .fst tt) +Bool→PropInj false true p = Empty.rec (invEq p tt) Bool→PropInj false false _ = refl Bool→PropInj* : (a b : Bool) → Bool→Type* {ℓ} a ≃ Bool→Type* {ℓ} b → a ≡ b Bool→PropInj* true true _ = refl -Bool→PropInj* true false p = rec* (p .fst tt*) -Bool→PropInj* false true p = rec* (invEq p tt*) +Bool→PropInj* true false p = Empty.rec* (p .fst tt*) +Bool→PropInj* false true p = Empty.rec* (invEq p tt*) Bool→PropInj* false false _ = refl isPropBool→Prop : {a : Bool} → isProp (Bool→Type a) @@ -264,7 +265,7 @@ DecBool→Prop* {a = false} = no (λ (lift x) → x) Dec→DecBool : {P : Type ℓ} → (dec : Dec P) → P → Bool→Type (Dec→Bool dec) Dec→DecBool (yes p) _ = tt -Dec→DecBool (no ¬p) q = rec (¬p q) +Dec→DecBool (no ¬p) q = Empty.rec (¬p q) DecBool→Dec : {P : Type ℓ} → (dec : Dec P) → Bool→Type (Dec→Bool dec) → P DecBool→Dec (yes p) _ = p @@ -278,7 +279,7 @@ Bool≡BoolDec {a = false} = refl Dec→DecBool* : {P : Type ℓ} → (dec : Dec P) → P → Bool→Type* {ℓ} (Dec→Bool dec) Dec→DecBool* (yes p) _ = tt* -Dec→DecBool* (no ¬p) q = rec (¬p q) +Dec→DecBool* (no ¬p) q = Empty.rec (¬p q) DecBool→Dec* : {P : Type ℓ} → (dec : Dec P) → Bool→Type* {ℓ} (Dec→Bool dec) → P DecBool→Dec* (yes p) _ = p @@ -292,15 +293,15 @@ Bool≡BoolDec* {a = false} = refl Bool→Prop× : (a b : Bool) → Bool→Type (a and b) → Bool→Type a × Bool→Type b Bool→Prop× true true _ = tt , tt -Bool→Prop× true false p = rec p -Bool→Prop× false true p = rec p -Bool→Prop× false false p = rec p +Bool→Prop× true false p = Empty.rec p +Bool→Prop× false true p = Empty.rec p +Bool→Prop× false false p = Empty.rec p Bool→Prop×' : (a b : Bool) → Bool→Type a × Bool→Type b → Bool→Type (a and b) Bool→Prop×' true true _ = tt -Bool→Prop×' true false (_ , p) = rec p -Bool→Prop×' false true (p , _) = rec p -Bool→Prop×' false false (p , _) = rec p +Bool→Prop×' true false (_ , p) = Empty.rec p +Bool→Prop×' false true (p , _) = Empty.rec p +Bool→Prop×' false false (p , _) = Empty.rec p Bool→Prop≃ : (a b : Bool) → Bool→Type a × Bool→Type b ≃ Bool→Type (a and b) Bool→Prop≃ a b = @@ -311,14 +312,14 @@ Bool→Prop⊎ : (a b : Bool) → Bool→Type (a or b) → Bool→Type a ⊎ Boo Bool→Prop⊎ true true _ = inl tt Bool→Prop⊎ true false _ = inl tt Bool→Prop⊎ false true _ = inr tt -Bool→Prop⊎ false false p = rec p +Bool→Prop⊎ false false p = Empty.rec p Bool→Prop⊎' : (a b : Bool) → Bool→Type a ⊎ Bool→Type b → Bool→Type (a or b) Bool→Prop⊎' true true _ = tt Bool→Prop⊎' true false _ = tt Bool→Prop⊎' false true _ = tt -Bool→Prop⊎' false false (inl p) = rec p -Bool→Prop⊎' false false (inr p) = rec p +Bool→Prop⊎' false false (inl p) = Empty.rec p +Bool→Prop⊎' false false (inr p) = Empty.rec p PropBoolP→P : (dec : Dec A) → Bool→Type (Dec→Bool dec) → A PropBoolP→P (yes p) _ = p @@ -335,6 +336,19 @@ Bool≡ false false = true Bool≡≃ : (a b : Bool) → (a ≡ b) ≃ Bool→Type (Bool≡ a b) Bool≡≃ true true = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) -Bool≡≃ true false = uninhabEquiv true≢false rec -Bool≡≃ false true = uninhabEquiv false≢true rec +Bool≡≃ true false = uninhabEquiv true≢false Empty.rec +Bool≡≃ false true = uninhabEquiv false≢true Empty.rec Bool≡≃ false false = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) +open Iso + +-- Bool is equivalent to bi-point type + +Iso-⊤⊎⊤-Bool : Iso (Unit ⊎ Unit) Bool +Iso-⊤⊎⊤-Bool .fun (inl tt) = true +Iso-⊤⊎⊤-Bool .fun (inr tt) = false +Iso-⊤⊎⊤-Bool .inv true = inl tt +Iso-⊤⊎⊤-Bool .inv false = inr tt +Iso-⊤⊎⊤-Bool .leftInv (inl tt) = refl +Iso-⊤⊎⊤-Bool .leftInv (inr tt) = refl +Iso-⊤⊎⊤-Bool .rightInv true = refl +Iso-⊤⊎⊤-Bool .rightInv false = refl diff --git a/Cubical/Data/Empty/Base.agda b/Cubical/Data/Empty/Base.agda index c3bdd2af0c..b126a028f1 100644 --- a/Cubical/Data/Empty/Base.agda +++ b/Cubical/Data/Empty/Base.agda @@ -1,7 +1,6 @@ {-# OPTIONS --safe #-} module Cubical.Data.Empty.Base where -open import Cubical.Core.Everything open import Cubical.Foundations.Prelude private diff --git a/Cubical/Data/Fin/Base.agda b/Cubical/Data/Fin/Base.agda index 2f9b6f5bc8..491f4e0537 100644 --- a/Cubical/Data/Fin/Base.agda +++ b/Cubical/Data/Fin/Base.agda @@ -7,7 +7,7 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Nat using (ℕ; zero; suc ; znots) +open import Cubical.Data.Nat using (ℕ ; zero ; suc ; _+_ ; znots) open import Cubical.Data.Nat.Order open import Cubical.Data.Nat.Order.Recursive using () renaming (_≤_ to _≤′_) open import Cubical.Data.Sigma @@ -104,3 +104,8 @@ any? {n = suc n} {P = P} P? = helper (x , Px) with fsplit x ... | inl x≡0 = inl (subst P (sym x≡0) Px) ... | inr (k , x≡sk) = inr (k , subst P (sym x≡sk) Px) + +FinPathℕ : {n : ℕ} (x : Fin n) (y : ℕ) → fst x ≡ y → Σ[ p ∈ _ ] (x ≡ (y , p)) +FinPathℕ {n = n} x y p = + ((fst (snd x)) , (cong (λ y → fst (snd x) + y) (cong suc (sym p)) ∙ snd (snd x))) + , (Σ≡Prop (λ _ → m≤n-isProp) p) diff --git a/Cubical/Data/Fin/Properties.agda b/Cubical/Data/Fin/Properties.agda index dcd188bf1e..b712e61b7a 100644 --- a/Cubical/Data/Fin/Properties.agda +++ b/Cubical/Data/Fin/Properties.agda @@ -636,3 +636,48 @@ FinData≃Fin N = isoToEquiv (FinDataIsoFin N) FinData≡Fin : (N : ℕ) → FinData N ≡ Fin N FinData≡Fin N = ua (FinData≃Fin N) + +-- decidability of Fin + +DecFin : (n : ℕ) → Dec (Fin n) +DecFin 0 = no ¬Fin0 +DecFin (suc n) = yes fzero + +-- propositional truncation of Fin + +Dec∥Fin∥ : (n : ℕ) → Dec ∥ Fin n ∥ +Dec∥Fin∥ n = Dec∥∥ (DecFin n) + +-- some properties about cardinality + +Fin>0→isInhab : (n : ℕ) → 0 < n → Fin n +Fin>0→isInhab 0 p = Empty.rec (¬-<-zero p) +Fin>0→isInhab (suc n) p = fzero + +Fin>1→hasNonEqualTerm : (n : ℕ) → 1 < n → Σ[ i ∈ Fin n ] Σ[ j ∈ Fin n ] ¬ i ≡ j +Fin>1→hasNonEqualTerm 0 p = Empty.rec (snotz (≤0→≡0 p)) +Fin>1→hasNonEqualTerm 1 p = Empty.rec (snotz (≤0→≡0 (pred-≤-pred p))) +Fin>1→hasNonEqualTerm (suc (suc n)) _ = fzero , fone , fzero≠fone + +isEmpty→Fin≡0 : (n : ℕ) → ¬ Fin n → 0 ≡ n +isEmpty→Fin≡0 0 _ = refl +isEmpty→Fin≡0 (suc n) p = Empty.rec (p fzero) + +isInhab→Fin>0 : (n : ℕ) → Fin n → 0 < n +isInhab→Fin>0 0 i = Empty.rec (¬Fin0 i) +isInhab→Fin>0 (suc n) _ = suc-≤-suc zero-≤ + +hasNonEqualTerm→Fin>1 : (n : ℕ) → (i j : Fin n) → ¬ i ≡ j → 1 < n +hasNonEqualTerm→Fin>1 0 i _ _ = Empty.rec (¬Fin0 i) +hasNonEqualTerm→Fin>1 1 i j p = Empty.rec (p (isContr→isProp isContrFin1 i j)) +hasNonEqualTerm→Fin>1 (suc (suc n)) _ _ _ = suc-≤-suc (suc-≤-suc zero-≤) + +Fin≤1→isProp : (n : ℕ) → n ≤ 1 → isProp (Fin n) +Fin≤1→isProp 0 _ = isPropFin0 +Fin≤1→isProp 1 _ = isContr→isProp isContrFin1 +Fin≤1→isProp (suc (suc n)) p = Empty.rec (¬-<-zero (pred-≤-pred p)) + +isProp→Fin≤1 : (n : ℕ) → isProp (Fin n) → n ≤ 1 +isProp→Fin≤1 0 _ = ≤-solver 0 1 +isProp→Fin≤1 1 _ = ≤-solver 1 1 +isProp→Fin≤1 (suc (suc n)) p = Empty.rec (fzero≠fone (p fzero fone)) diff --git a/Cubical/Data/FinSet/Base.agda b/Cubical/Data/FinSet/Base.agda index 43f2ae0e78..9cf0d4138e 100644 --- a/Cubical/Data/FinSet/Base.agda +++ b/Cubical/Data/FinSet/Base.agda @@ -3,7 +3,7 @@ Definition of finite sets There are may different formulations of finite sets in constructive mathematics, -and we will use Bishop finiteness as is called by some people in the literature. +and we will use Bishop finiteness as is usually called in the literature. -} {-# OPTIONS --safe #-} @@ -13,7 +13,7 @@ module Cubical.Data.FinSet.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.Foundations.Univalence open import Cubical.HITs.PropositionalTruncation as Prop @@ -28,16 +28,6 @@ private ℓ ℓ' ℓ'' : Level A : Type ℓ --- operators to more conveniently compose equivalences - -module _ - {A : Type ℓ}{B : Type ℓ'}{C : Type ℓ''} where - - infixr 30 _⋆_ - - _⋆_ : A ≃ B → B ≃ C → A ≃ C - _⋆_ = compEquiv - -- definition of (Bishop) finite sets isFinSet : Type ℓ → Type ℓ diff --git a/Cubical/Data/FinSet/Cardinality.agda b/Cubical/Data/FinSet/Cardinality.agda index f368568ca7..470291c0bb 100644 --- a/Cubical/Data/FinSet/Cardinality.agda +++ b/Cubical/Data/FinSet/Cardinality.agda @@ -1,6 +1,14 @@ {- -This file contains properties and formulae about cardinality. +Properties and Formulae about Cardinality + +This file contains: +- Relation between abstract properties and cardinality in special cases; +- Combinatorial formulae, namely, cardinality of A+B, A×B, ΣAB, ΠAB, etc; +- A general form of Pigeonhole Principle; +- Maximal value of numerical function on finite sets; +- Set truncation of FinSet is equivalent to ℕ; +- FinProp is equivalent to Bool. -} {-# OPTIONS --safe #-} @@ -12,7 +20,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.Foundations.Equiv.Properties open import Cubical.HITs.PropositionalTruncation as Prop @@ -82,14 +90,16 @@ module _ card>0→isInhab : card X > 0 → ∥ X .fst ∥ card>0→isInhab p = - Prop.rec isPropPropTrunc (λ e → ∣ invEq e (Fin>0 _ p) ∣) (∣≃card∣ X) + Prop.map (λ e → invEq e (Fin>0→isInhab _ p)) (∣≃card∣ X) card>1→hasNonEqualTerm : card X > 1 → ∥ Σ[ a ∈ X .fst ] Σ[ b ∈ X .fst ] ¬ a ≡ b ∥ card>1→hasNonEqualTerm p = - Prop.rec isPropPropTrunc - (λ e → ∣ e .fst (Fin>1 _ p .fst) , e .fst (Fin>1 _ p .snd .fst) , - Fin>1 _ p .snd .snd ∘ invEq (congEquiv e) ∣) - (∣invEquiv∣ (∣≃card∣ X)) + Prop.map + (λ e → + e .fst (Fin>1→hasNonEqualTerm _ p .fst) , + e .fst (Fin>1→hasNonEqualTerm _ p .snd .fst) , + Fin>1→hasNonEqualTerm _ p .snd .snd ∘ invEq (congEquiv e)) + (∣invEquiv∣ (∣≃card∣ X)) card≡1→isContr : card X ≡ 1 → isContr (X .fst) card≡1→isContr p = @@ -98,15 +108,16 @@ module _ card≤1→isProp : card X ≤ 1 → isProp (X .fst) card≤1→isProp p = - Prop.rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 (invEquiv e) (Fin≤1 (card X) p)) (∣≃card∣ X) + Prop.rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 (invEquiv e) (Fin≤1→isProp (card X) p)) (∣≃card∣ X) card≡n : card X ≡ n → ∥ X ≡ 𝔽in n ∥ card≡n {n = n} p = - Prop.rec isPropPropTrunc + Prop.map (λ e → - ∣(λ i → ua e i , - isProp→PathP {B = λ j → isFinSet (ua e j)} - (λ _ → isPropIsFinSet) (X .snd) (𝔽in n .snd) i )∣) + (λ i → + ua e i , + isProp→PathP {B = λ j → isFinSet (ua e j)} + (λ _ → isPropIsFinSet) (X .snd) (𝔽in n .snd) i )) (∣≃card∣ X ⋆̂ ∣ pathToEquiv (cong Fin p) ⋆ invEquiv (𝔽in≃Fin n) ∣) card≡0 : card X ≡ 0 → X ≡ 𝟘 @@ -125,34 +136,34 @@ module _ 1 (FinSet≡ X 𝟙) (isOfHLevel≡ 1 (card≤1→isProp (subst (λ a → a ≤ 1) (sym p) (≤-solver 1 1))) (isPropUnit*))) .fst - (Prop.rec isPropPropTrunc (λ q → ∣ q ∙ 𝔽in1≡𝟙 ∣) (card≡n p)) + (Prop.map (λ q → q ∙ 𝔽in1≡𝟙) (card≡n p)) module _ (X : FinSet ℓ) where isEmpty→card≡0 : ¬ X .fst → card X ≡ 0 isEmpty→card≡0 p = - Prop.rec (isSetℕ _ _) (λ e → sym (emptyFin _ (p ∘ invEq e))) (∣≃card∣ X) + Prop.rec (isSetℕ _ _) (λ e → sym (isEmpty→Fin≡0 _ (p ∘ invEq e))) (∣≃card∣ X) isInhab→card>0 : ∥ X .fst ∥ → card X > 0 - isInhab→card>0 = Prop.rec2 m≤n-isProp (λ p x → nonEmptyFin _ (p .fst x)) (∣≃card∣ X) + isInhab→card>0 = Prop.rec2 m≤n-isProp (λ p x → isInhab→Fin>0 _ (p .fst x)) (∣≃card∣ X) hasNonEqualTerm→card>1 : {a b : X. fst} → ¬ a ≡ b → card X > 1 hasNonEqualTerm→card>1 {a = a} {b = b} q = - Prop.rec m≤n-isProp (λ p → nonEqualTermFin _ (p .fst a) (p .fst b) (q ∘ invEq (congEquiv p))) (∣≃card∣ X) + Prop.rec m≤n-isProp (λ p → hasNonEqualTerm→Fin>1 _ (p .fst a) (p .fst b) (q ∘ invEq (congEquiv p))) (∣≃card∣ X) isContr→card≡1 : isContr (X .fst) → card X ≡ 1 isContr→card≡1 p = cardEquiv X (_ , isFinSetUnit) ∣ isContr→≃Unit p ∣ isProp→card≤1 : isProp (X .fst) → card X ≤ 1 - isProp→card≤1 p = propFin (card X) (Prop.rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 e p) (∣≃card∣ X)) + isProp→card≤1 p = isProp→Fin≤1 (card X) (Prop.rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 e p) (∣≃card∣ X)) {- formulae about cardinality -} -- results to be used in diProp.rect induction on FinSet card𝟘 : card (𝟘 {ℓ}) ≡ 0 -card𝟘 {ℓ = ℓ} = isEmpty→card≡0 (𝟘 {ℓ}) Empty.rec* +card𝟘 {ℓ = ℓ} = isEmpty→card≡0 (𝟘 {ℓ}) (Empty.rec*) card𝟙 : card (𝟙 {ℓ}) ≡ 1 card𝟙 {ℓ = ℓ} = isContr→card≡1 (𝟙 {ℓ}) isContrUnit* @@ -273,11 +284,12 @@ module _ (λ X → isPropΠ3 (λ _ _ _ → isSetℕ _ _)) (λ n f c h → prodConst𝔽in n f c h ∙ (λ i → c ^ card𝔽in {ℓ = ℓ} n (~ i))) X f c h -≡≤ : {m n l k r s : ℕ} → m ≤ n → l ≤ k → r ≡ m + l → s ≡ n + k → r ≤ s -≡≤ {m = m} {l = l} {k = k} p q u v = subst2 (_≤_) (sym u) (sym v) (≤-+ p q) +private + ≡≤ : {m n l k r s : ℕ} → m ≤ n → l ≤ k → r ≡ m + l → s ≡ n + k → r ≤ s + ≡≤ {m = m} {l = l} {k = k} p q u v = subst2 (_≤_) (sym u) (sym v) (≤-+-≤ p q) -≡< : {m n l k r s : ℕ} → m < n → l ≤ k → r ≡ m + l → s ≡ n + k → r < s -≡< {m = m} {l = l} {k = k} p q u v = subst2 (_<_) (sym u) (sym v) (<-+-≤ p q) + ≡< : {m n l k r s : ℕ} → m < n → l ≤ k → r ≡ m + l → s ≡ n + k → r < s + ≡< {m = m} {l = l} {k = k} p q u v = subst2 (_<_) (sym u) (sym v) (<-+-≤ p q) sum≤𝔽in : (n : ℕ)(f g : 𝔽in {ℓ} n .fst → ℕ)(h : (x : 𝔽in n .fst) → f x ≤ g x) → sum (𝔽in n) f ≤ sum (𝔽in n) g sum≤𝔽in 0 f g _ = subst2 (_≤_) (sym (sum𝟘 f)) (sym (sum𝟘 g)) ≤-refl @@ -347,14 +359,14 @@ module _ cardΣ : card (_ , isFinSetΣ X Y) ≡ sum X (λ x → card (Y x)) cardΣ = cardEquiv (_ , isFinSetΣ X Y) (_ , isFinSetΣ X (λ x → Fin (card (Y x)) , isFinSetFin)) - (Prop.rec isPropPropTrunc (λ e → ∣ Σ-cong-equiv-snd e ∣) - (choice X (λ x → Y x .fst ≃ Fin (card (Y x))) (λ x → ∣≃card∣ (Y x)))) + (Prop.map Σ-cong-equiv-snd + (choice X (λ x → Y x .fst ≃ Fin (card (Y x))) (λ x → ∣≃card∣ (Y x)))) cardΠ : card (_ , isFinSetΠ X Y) ≡ prod X (λ x → card (Y x)) cardΠ = cardEquiv (_ , isFinSetΠ X Y) (_ , isFinSetΠ X (λ x → Fin (card (Y x)) , isFinSetFin)) - (Prop.rec isPropPropTrunc (λ e → ∣ equivΠCod e ∣) - (choice X (λ x → Y x .fst ≃ Fin (card (Y x))) (λ x → ∣≃card∣ (Y x)))) + (Prop.map equivΠCod + (choice X (λ x → Y x .fst ≃ Fin (card (Y x))) (λ x → ∣≃card∣ (Y x)))) module _ (X : FinSet ℓ ) @@ -382,14 +394,13 @@ module _ -- the pigeonhole priniple -- a logical lemma -module _ - (X : Type ℓ) - (P : X → Type ℓ') - (Q : X → Type ℓ'') - (r : (x : X) → ¬ (P x) → Q x) where - - ¬ΠQ→¬¬ΣP : ¬ ((x : X) → Q x) → ¬ ¬ (Σ X P) - ¬ΠQ→¬¬ΣP g f = g (λ x → r x (λ p → f (x , p))) +private + ¬ΠQ→¬¬ΣP : (X : Type ℓ) + (P : X → Type ℓ' ) + (Q : X → Type ℓ'') + (r : (x : X) → ¬ (P x) → Q x) + → ¬ ((x : X) → Q x) → ¬ ¬ (Σ X P) + ¬ΠQ→¬¬ΣP _ _ _ r g f = g (λ x → r x (λ p → f (x , p))) module _ (f : X .fst → Y .fst) @@ -412,14 +423,12 @@ module _ pigeonHole = PeirceLaw (isFinSetΣ Y (λ _ → _ , isDecProp→isFinSet m≤n-isProp (≤Dec _ _))) ¬¬pigeonHole -- a special case, proved in Cubical.Data.Fin.Properties --- a technical lemma -module _ - (X : Type ℓ) - (P : X → Type ℓ') where - Σ∥P∥→∥ΣP∥ : Σ X (λ x → ∥ P x ∥) → ∥ Σ X P ∥ - Σ∥P∥→∥ΣP∥ (x , p) = - Prop.rec isPropPropTrunc (λ q → ∣ x , q ∣) p +-- a technical lemma +private + Σ∥P∥→∥ΣP∥ : (X : Type ℓ)(P : X → Type ℓ') + → Σ X (λ x → ∥ P x ∥) → ∥ Σ X P ∥ + Σ∥P∥→∥ΣP∥ _ _ (x , p) = Prop.map (λ q → x , q) p module _ (f : X .fst → Y .fst) @@ -439,7 +448,7 @@ module _ pigeonHole' : ∥ Σ[ x ∈ X .fst ] Σ[ x' ∈ X .fst ] (¬ x ≡ x') × (f x ≡ f x') ∥ pigeonHole' = - Prop.rec isPropPropTrunc (λ p → ∣ nonInj p ∣) + Prop.map nonInj (Prop.rec isPropPropTrunc fiberNonEqualTerm (pigeonHole {X = X} {Y = Y} f 1 (subst (λ a → _ > a) (sym (·-identityˡ _)) p))) @@ -524,7 +533,7 @@ module _ ΣMax⊎-case (x , p) (y , q) (gt r) .snd (inr y') = ≤-trans (q y') (<-weaken r) ∃Max⊎ : ∃Max X (f ∘ inl) → ∃Max Y (f ∘ inr) → ∃Max (X ⊎ Y) f - ∃Max⊎ = Prop.rec2 isPropPropTrunc (λ p q → ∣ ΣMax⊎-case p q (_≟_ _ _) ∣) + ∃Max⊎ = Prop.map2 (λ p q → ΣMax⊎-case p q (_≟_ _ _)) ΣMax𝟙 : (f : 𝟙 {ℓ} .fst → ℕ) → ΣMax _ f ΣMax𝟙 f .fst = tt* diff --git a/Cubical/Data/FinSet/Constructor.agda b/Cubical/Data/FinSet/Constructor.agda index d049250ac9..e30993124c 100644 --- a/Cubical/Data/FinSet/Constructor.agda +++ b/Cubical/Data/FinSet/Constructor.agda @@ -1,6 +1,7 @@ {- -This files contains lots of useful properties about constructions on finite sets +This files contains: +- Facts about constructions on finite sets, especially when they preserve finiteness. -} {-# OPTIONS --safe #-} @@ -9,7 +10,7 @@ module Cubical.Data.FinSet.Constructor where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.Foundations.Univalence open import Cubical.HITs.PropositionalTruncation as Prop diff --git a/Cubical/Data/FinSet/DecidablePredicate.agda b/Cubical/Data/FinSet/DecidablePredicate.agda index 57656c7806..f67d9738bf 100644 --- a/Cubical/Data/FinSet/DecidablePredicate.agda +++ b/Cubical/Data/FinSet/DecidablePredicate.agda @@ -10,7 +10,7 @@ module Cubical.Data.FinSet.DecidablePredicate where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.Foundations.Equiv.Properties open import Cubical.HITs.PropositionalTruncation as Prop diff --git a/Cubical/Data/FinSet/FiniteChoice.agda b/Cubical/Data/FinSet/FiniteChoice.agda index fc933298b0..eb5657278b 100644 --- a/Cubical/Data/FinSet/FiniteChoice.agda +++ b/Cubical/Data/FinSet/FiniteChoice.agda @@ -1,8 +1,7 @@ {- Axiom of Finite Choice - -Yep, it's a theorem actually. +- Yep, it's a theorem actually. -} {-# OPTIONS --safe #-} @@ -11,7 +10,7 @@ module Cubical.Data.FinSet.FiniteChoice where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.HITs.PropositionalTruncation as Prop diff --git a/Cubical/Data/FinSet/FiniteStructure.agda b/Cubical/Data/FinSet/FiniteStructure.agda index fce9dd664b..b50419e3ed 100644 --- a/Cubical/Data/FinSet/FiniteStructure.agda +++ b/Cubical/Data/FinSet/FiniteStructure.agda @@ -10,7 +10,7 @@ module Cubical.Data.FinSet.FiniteStructure where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.HITs.SetTruncation as Set diff --git a/Cubical/Data/FinSet/FiniteType.agda b/Cubical/Data/FinSet/FiniteType.agda index eba37ee21d..a9c07ee248 100644 --- a/Cubical/Data/FinSet/FiniteType.agda +++ b/Cubical/Data/FinSet/FiniteType.agda @@ -33,7 +33,7 @@ open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.DecidablePredicate open import Cubical.Data.FinSet.Constructor -open import Cubical.Data.FinSet.Quotient +open import Cubical.Data.FinSet.Quotients open import Cubical.Data.FinSet.Cardinality open import Cubical.Relation.Nullary diff --git a/Cubical/Data/FinSet/Induction.agda b/Cubical/Data/FinSet/Induction.agda index 9826da7761..ca77b9bc2c 100644 --- a/Cubical/Data/FinSet/Induction.agda +++ b/Cubical/Data/FinSet/Induction.agda @@ -1,6 +1,6 @@ {- -Indutiive eliminators to directly prove properties of all finite sets +Inductive eliminators to establish properties of all finite sets directly -} {-# OPTIONS --safe #-} @@ -9,17 +9,18 @@ module Cubical.Data.FinSet.Induction where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.HITs.SetTruncation as Set -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat + renaming (_+_ to _+ℕ_) hiding (elim) open import Cubical.Data.Unit open import Cubical.Data.Empty as Empty -open import Cubical.Data.Sum as Sum +open import Cubical.Data.Sum open import Cubical.Data.Fin renaming (Fin to Finℕ) open import Cubical.Data.SumFin @@ -120,12 +121,12 @@ module _ elimProp𝔽in 0 = p0 elimProp𝔽in (suc n) = p1 (elimProp𝔽in n) - elimProp' : (X : FinSet ℓ) → P X - elimProp' = elimProp elimProp𝔽in + elimProp𝟙+ : (X : FinSet ℓ) → P X + elimProp𝟙+ = elimProp elimProp𝔽in module _ (p0 : P 𝟘)(p1 : P 𝟙) (p+ : {X Y : FinSet ℓ} → P X → P Y → P (X + Y)) where - elimProp'' : (X : FinSet ℓ) → P X - elimProp'' = elimProp' p0 (λ p → p+ p1 p) + elimProp+ : (X : FinSet ℓ) → P X + elimProp+ = elimProp𝟙+ p0 (λ p → p+ p1 p) diff --git a/Cubical/Data/FinSet/Properties.agda b/Cubical/Data/FinSet/Properties.agda index 3024f0e6fd..c87b4719ad 100644 --- a/Cubical/Data/FinSet/Properties.agda +++ b/Cubical/Data/FinSet/Properties.agda @@ -8,7 +8,7 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.Structure open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.HITs.PropositionalTruncation as Prop @@ -18,6 +18,7 @@ open import Cubical.Data.Bool open import Cubical.Data.Empty as Empty open import Cubical.Data.Sigma +open import Cubical.Data.Fin.Base renaming (Fin to Finℕ) open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base @@ -39,13 +40,14 @@ module _ infixr 30 _⋆̂_ _⋆̂_ : ∥ A ≃ B ∥ → ∥ B ≃ C ∥ → ∥ A ≃ C ∥ - _⋆̂_ = Prop.rec2 isPropPropTrunc (λ p q → ∣ p ⋆ q ∣) + _⋆̂_ = Prop.map2 (_⋆_) module _ {A : Type ℓ}{B : Type ℓ'} where ∣invEquiv∣ : ∥ A ≃ B ∥ → ∥ B ≃ A ∥ - ∣invEquiv∣ = Prop.rec isPropPropTrunc (λ p → ∣ invEquiv p ∣) + ∣invEquiv∣ = Prop.map invEquiv + -- useful implications @@ -77,7 +79,7 @@ isDec→isFinSet∥∥ dec = isDecProp→isFinSet isPropPropTrunc (Dec∥∥ dec isFinSet→Dec∥∥ : isFinSet A → Dec ∥ A ∥ isFinSet→Dec∥∥ h = Prop.rec (isPropDec isPropPropTrunc) - (λ p → EquivPresDec (propTrunc≃ (invEquiv p)) (∥Fin∥ _)) (h .snd) + (λ p → EquivPresDec (propTrunc≃ (invEquiv p)) (Dec∥Fin∥ _)) (h .snd) isFinProp→Dec : isFinSet A → isProp A → Dec A isFinProp→Dec p h = subst Dec (propTruncIdempotent h) (isFinSet→Dec∥∥ p) diff --git a/Cubical/Data/FinSet/Quotient.agda b/Cubical/Data/FinSet/Quotients.agda similarity index 94% rename from Cubical/Data/FinSet/Quotient.agda rename to Cubical/Data/FinSet/Quotients.agda index c9dc2d95f4..f5f5a1c386 100644 --- a/Cubical/Data/FinSet/Quotient.agda +++ b/Cubical/Data/FinSet/Quotients.agda @@ -1,11 +1,17 @@ +{- + +This file contains: +- The quotient of finite sets by decidable equivalence relation is still finite, by using equivalence class. + +-} {-# OPTIONS --safe #-} -module Cubical.Data.FinSet.Quotient where +module Cubical.Data.FinSet.Quotients where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.Foundations.Equiv.Properties open import Cubical.Foundations.Univalence @@ -13,6 +19,7 @@ open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.HITs.SetQuotients as SetQuot open import Cubical.HITs.SetQuotients.EqClass +open import Cubical.Data.Nat open import Cubical.Data.Bool open import Cubical.Data.Sigma diff --git a/Cubical/Data/HomotopyGroup.agda b/Cubical/Data/HomotopyGroup.agda deleted file mode 100644 index 04ab4f13b2..0000000000 --- a/Cubical/Data/HomotopyGroup.agda +++ /dev/null @@ -1,4 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Data.HomotopyGroup where - -open import Cubical.Data.HomotopyGroup.Base public diff --git a/Cubical/Data/HomotopyGroup/Base.agda b/Cubical/Data/HomotopyGroup/Base.agda deleted file mode 100644 index 997b31dfe4..0000000000 --- a/Cubical/Data/HomotopyGroup/Base.agda +++ /dev/null @@ -1,52 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Data.HomotopyGroup.Base where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -import Cubical.Foundations.GroupoidLaws as GL -open import Cubical.Foundations.Pointed - -open import Cubical.Data.Nat -open import Cubical.Algebra.Group - -open import Cubical.Homotopy.Loopspace - -open import Cubical.HITs.SetTruncation as SetTrunc - -π^_ : ∀ {ℓ} → ℕ → Pointed ℓ → Group ℓ -π^_ {ℓ} n p = makeGroup e _⨀_ _⁻¹ isSetSetTrunc assoc rUnit lUnit rCancel lCancel - where - n' : ℕ - n' = suc n - - A : Type ℓ - A = typ ((Ω^ n') p) - - e : ∥ A ∥₂ - e = ∣ pt ((Ω^ n') p) ∣₂ - - _⁻¹ : ∥ A ∥₂ → ∥ A ∥₂ - _⁻¹ = SetTrunc.elim {B = λ _ → ∥ A ∥₂} (λ x → squash₂) λ a → ∣ sym a ∣₂ - - _⨀_ : ∥ A ∥₂ → ∥ A ∥₂ → ∥ A ∥₂ - _⨀_ = SetTrunc.elim2 (λ _ _ → squash₂) λ a₀ a₁ → ∣ a₀ ∙ a₁ ∣₂ - - lUnit : (a : ∥ A ∥₂) → (e ⨀ a) ≡ a - lUnit = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) - (λ a → cong ∣_∣₂ (sym (GL.lUnit a) )) - - rUnit : (a : ∥ A ∥₂) → a ⨀ e ≡ a - rUnit = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) - (λ a → cong ∣_∣₂ (sym (GL.rUnit a) )) - - assoc : (a b c : ∥ A ∥₂) → a ⨀ (b ⨀ c) ≡ (a ⨀ b) ⨀ c - assoc = SetTrunc.elim3 (λ _ _ _ → isProp→isSet (squash₂ _ _)) - (λ a b c → cong ∣_∣₂ (GL.assoc _ _ _)) - - lCancel : (a : ∥ A ∥₂) → ((a ⁻¹) ⨀ a) ≡ e - lCancel = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) - λ a → cong ∣_∣₂ (GL.lCancel _) - - rCancel : (a : ∥ A ∥₂) → (a ⨀ (a ⁻¹)) ≡ e - rCancel = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) - λ a → cong ∣_∣₂ (GL.rCancel _) diff --git a/Cubical/Data/Int/Properties.agda b/Cubical/Data/Int/Properties.agda index 9c8f316c89..ce47d3059c 100644 --- a/Cubical/Data/Int/Properties.agda +++ b/Cubical/Data/Int/Properties.agda @@ -84,6 +84,9 @@ isSetℤ = Discrete→isSet discreteℤ -Involutive (pos n) = cong -_ (-pos n) ∙ -neg n -Involutive (negsuc n) = refl +isEquiv- : isEquiv (-_) +isEquiv- = isoToIsEquiv (iso -_ -_ -Involutive -Involutive) + sucℤ+pos : ∀ n m → sucℤ (m +pos n) ≡ (sucℤ m) +pos n sucℤ+pos zero m = refl sucℤ+pos (suc n) m = cong sucℤ (sucℤ+pos n m) @@ -361,6 +364,10 @@ pos- (suc m) (suc n) = -- multiplication +pos·pos : (n m : ℕ) → pos (n ·ℕ m) ≡ pos n · pos m +pos·pos zero m = refl +pos·pos (suc n) m = pos+ m (n ·ℕ m) ∙ cong (pos m +_) (pos·pos n m) + pos·negsuc : (n m : ℕ) → pos n · negsuc m ≡ - (pos n · pos (suc m)) pos·negsuc zero m = refl pos·negsuc (suc n) m = @@ -444,6 +451,10 @@ private -DistR· : (b c : ℤ) → - (b · c) ≡ b · - c -DistR· b c = cong (-_) (·Comm b c) ∙∙ -DistL· c b ∙∙ ·Comm (- c) b +ℤ·negsuc : (n : ℤ) (m : ℕ) → n · negsuc m ≡ - (n · pos (suc m)) +ℤ·negsuc (pos n) m = pos·negsuc n m +ℤ·negsuc (negsuc n) m = negsuc·negsuc n m ∙ sym (-DistL· (negsuc n) (pos (suc m))) + ·Assoc : (a b c : ℤ) → (a · (b · c)) ≡ ((a · b) · c) ·Assoc (pos zero) b c = refl ·Assoc (pos (suc n)) b c = @@ -473,3 +484,8 @@ abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) ⊎→abs (negsuc n₁) n (inl x₁) = cong abs x₁ ⊎→abs x zero (inr x₁) = cong abs x₁ ⊎→abs x (suc n) (inr x₁) = cong abs x₁ + +abs- : (x : ℤ) → abs (- x) ≡ abs x +abs- (pos zero) = refl +abs- (pos (suc n)) = refl +abs- (negsuc n) = refl diff --git a/Cubical/Data/Nat/Base.agda b/Cubical/Data/Nat/Base.agda index 798728362c..427cdb8a2e 100644 --- a/Cubical/Data/Nat/Base.agda +++ b/Cubical/Data/Nat/Base.agda @@ -59,3 +59,13 @@ isEvenT n = toType (isEven n) isOddT : ℕ → Type isOddT n = isEvenT (suc n) + +isZero : ℕ → Bool +isZero zero = true +isZero (suc n) = false + +-- exponential + +_^_ : ℕ → ℕ → ℕ +m ^ 0 = 1 +m ^ (suc n) = m · m ^ n diff --git a/Cubical/Data/Nat/Mod.agda b/Cubical/Data/Nat/Mod.agda index 4eddabaf49..8911504e25 100644 --- a/Cubical/Data/Nat/Mod.agda +++ b/Cubical/Data/Nat/Mod.agda @@ -4,6 +4,7 @@ module Cubical.Data.Nat.Mod where open import Cubical.Foundations.Prelude open import Cubical.Data.Nat open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty -- Defining x mod 0 to be 0. This way all the theorems below are true -- for n : ℕ instead of n : ℕ₊₁. @@ -88,6 +89,52 @@ mod-idempotent {n = suc n} = ∙∙ mod-rUnit (suc n) x ∙ (cong (_mod (suc n)) (+-comm x (suc n))) +zero-charac : (n : ℕ) → n mod n ≡ 0 +zero-charac zero = refl +zero-charac (suc n) = cong (_mod suc n) (+-comm 0 (suc n)) + ∙∙ modIndStep n 0 + ∙∙ modIndBase n 0 (n , (+-comm n 1)) + +zero-charac-gen : (n x : ℕ) → ((x · n) mod n) ≡ 0 +zero-charac-gen zero x = refl +zero-charac-gen (suc n) zero = refl +zero-charac-gen (suc n) (suc x) = + modIndStep n (x · (suc n)) ∙ zero-charac-gen (suc n) x + +mod·mod≡mod : (n x y : ℕ) + → (x · y) mod n ≡ (((x mod n) · (y mod n)) mod n) +mod·mod≡mod zero _ _ = refl +mod·mod≡mod (suc n) = + +induction n _ + (λ x p → +induction n _ + (λ y q + → cong (modInd n) + (cong₂ _·_ (sym (modIndBase n x p)) (sym (modIndBase n y q)))) + λ y p → + cong (modInd n) (sym (·-distribˡ x (suc n) y)) + ∙∙ mod+mod≡mod (suc n) (x · suc n) (x · y) + ∙∙ cong (λ z → modInd n (z + modInd n (x · y))) + (zero-charac-gen (suc n) x) + ∙∙ mod-idempotent (x · y) + ∙∙ p + ∙ cong (_mod (suc n)) (cong (x mod (suc n) ·_) + (sym (mod-idempotent y) + ∙∙ (λ i → modInd n (mod-rUnit (suc n) 0 i + modInd n y)) + ∙∙ sym (mod+mod≡mod (suc n) (suc n) y)))) + λ x p y → + (sym (cong (_mod (suc n)) (·-distribʳ (suc n) x y)) + ∙∙ mod+mod≡mod (suc n) (suc n · y) (x · y) + ∙∙ (λ i → modInd n ((cong (_mod (suc n)) + (·-comm (suc n) y) ∙ zero-charac-gen (suc n) y) i + + modInd n (x · y))) + ∙ mod-idempotent (x · y)) + ∙∙ p y + ∙∙ cong (_mod (suc n)) (cong (_· y mod (suc n)) + ((sym (mod-idempotent x) + ∙ cong (λ z → (z + x mod (suc n)) mod (suc n)) + (mod-rUnit (suc n) 0)) + ∙ sym (mod+mod≡mod (suc n) (suc n) x))) + mod-rCancel : (n x y : ℕ) → (x + y) mod n ≡ (x + y mod n) mod n mod-rCancel zero x y = refl mod-rCancel (suc n) x = @@ -106,12 +153,6 @@ mod-lCancel n x y = ∙∙ mod-rCancel n y x ∙∙ cong (_mod n) (+-comm y (x mod n)) -zero-charac : (n : ℕ) → n mod n ≡ 0 -zero-charac zero = refl -zero-charac (suc n) = cong (_mod suc n) (+-comm 0 (suc n)) - ∙∙ modIndStep n 0 - ∙∙ modIndBase n 0 (n , (+-comm n 1)) - -- remainder and quotient after division by n -- Again, allowing for 0-division to get nicer syntax remainder_/_ : (x n : ℕ) → ℕ diff --git a/Cubical/Data/Nat/Order.agda b/Cubical/Data/Nat/Order.agda index c2568ac4ab..112d620082 100644 --- a/Cubical/Data/Nat/Order.agda +++ b/Cubical/Data/Nat/Order.agda @@ -98,8 +98,8 @@ pred-≤-pred (k , p) = k , injSuc ((sym (+-suc k _)) ∙ p) l3 : 0 ≡ i l3 = sym (snd (m+n≡0→m≡0×n≡0 l2)) -≤-+ : m ≤ n → l ≤ k → m + l ≤ n + k -≤-+ p q = ≤-trans (≤-+k p) (≤-k+ q) +≤-+-≤ : m ≤ n → l ≤ k → m + l ≤ n + k +≤-+-≤ p q = ≤-trans (≤-+k p) (≤-k+ q) ≤-k+-cancel : k + m ≤ k + n → m ≤ n ≤-k+-cancel {k} {m} (l , p) = l , inj-m+ (sub k m ∙ p) @@ -163,8 +163,8 @@ predℕ-≤-predℕ {suc m} {suc n} ineq = pred-≤-pred ineq <-k+ : m < n → k + m < k + n <-k+ {m} {n} {k} p = subst (λ km → km ≤ k + n) (+-suc k m) (≤-k+ p) -+-<-+ : m < n → k < l → m + k < n + l -+-<-+ m0 : (n : ℕ) → 0 < n → Fin n -Fin>0 0 p = ⊥.rec (¬-<-zero p) -Fin>0 (suc n) p = fzero - -Fin>1 : (n : ℕ) → 1 < n → Σ[ i ∈ Fin n ] Σ[ j ∈ Fin n ] ¬ i ≡ j -Fin>1 0 p = ⊥.rec (snotz (≤0→≡0 p)) -Fin>1 1 p = ⊥.rec (snotz (≤0→≡0 (pred-≤-pred p))) -Fin>1 (suc (suc n)) _ = fzero , fsuc fzero , fzero≠fone - -emptyFin : (n : ℕ) → ¬ Fin n → 0 ≡ n -emptyFin 0 _ = refl -emptyFin (suc n) p = ⊥.rec (p fzero) - -nonEmptyFin : (n : ℕ) → Fin n → 0 < n -nonEmptyFin 0 i = ⊥.rec i -nonEmptyFin (suc n) _ = suc-≤-suc zero-≤ - -nonEqualTermFin : (n : ℕ) → (i j : Fin n) → ¬ i ≡ j → 1 < n -nonEqualTermFin 0 i _ _ = ⊥.rec i -nonEqualTermFin 1 i j p = ⊥.rec (p (isContr→isProp isContrSumFin1 i j)) -nonEqualTermFin (suc (suc n)) _ _ _ = suc-≤-suc (suc-≤-suc zero-≤) - -Fin≤1 : (n : ℕ) → n ≤ 1 → isProp (Fin n) -Fin≤1 0 _ = isProp⊥ -Fin≤1 1 _ = isContr→isProp isContrSumFin1 -Fin≤1 (suc (suc n)) p = ⊥.rec (¬-<-zero (pred-≤-pred p)) - -propFin : (n : ℕ) → isProp (Fin n) → n ≤ 1 -propFin 0 _ = ≤-solver 0 1 -propFin 1 _ = ≤-solver 1 1 -propFin (suc (suc n)) p = ⊥.rec (fzero≠fone (p fzero (fsuc fzero))) +Fin>0→isInhab : (n : ℕ) → 0 < n → Fin n +Fin>0→isInhab 0 p = ⊥.rec (¬-<-zero p) +Fin>0→isInhab (suc n) p = fzero + +Fin>1→hasNonEqualTerm : (n : ℕ) → 1 < n → Σ[ i ∈ Fin n ] Σ[ j ∈ Fin n ] ¬ i ≡ j +Fin>1→hasNonEqualTerm 0 p = ⊥.rec (snotz (≤0→≡0 p)) +Fin>1→hasNonEqualTerm 1 p = ⊥.rec (snotz (≤0→≡0 (pred-≤-pred p))) +Fin>1→hasNonEqualTerm (suc (suc n)) _ = fzero , fsuc fzero , fzero≠fone + +isEmpty→Fin≡0 : (n : ℕ) → ¬ Fin n → 0 ≡ n +isEmpty→Fin≡0 0 _ = refl +isEmpty→Fin≡0 (suc n) p = ⊥.rec (p fzero) + +isInhab→Fin>0 : (n : ℕ) → Fin n → 0 < n +isInhab→Fin>0 0 i = ⊥.rec i +isInhab→Fin>0 (suc n) _ = suc-≤-suc zero-≤ + +hasNonEqualTerm→Fin>1 : (n : ℕ) → (i j : Fin n) → ¬ i ≡ j → 1 < n +hasNonEqualTerm→Fin>1 0 i _ _ = ⊥.rec i +hasNonEqualTerm→Fin>1 1 i j p = ⊥.rec (p (isContr→isProp isContrSumFin1 i j)) +hasNonEqualTerm→Fin>1 (suc (suc n)) _ _ _ = suc-≤-suc (suc-≤-suc zero-≤) + +Fin≤1→isProp : (n : ℕ) → n ≤ 1 → isProp (Fin n) +Fin≤1→isProp 0 _ = isProp⊥ +Fin≤1→isProp 1 _ = isContr→isProp isContrSumFin1 +Fin≤1→isProp (suc (suc n)) p = ⊥.rec (¬-<-zero (pred-≤-pred p)) + +isProp→Fin≤1 : (n : ℕ) → isProp (Fin n) → n ≤ 1 +isProp→Fin≤1 0 _ = ≤-solver 0 1 +isProp→Fin≤1 1 _ = ≤-solver 1 1 +isProp→Fin≤1 (suc (suc n)) p = ⊥.rec (fzero≠fone (p fzero (fsuc fzero))) -- automorphisms of SumFin diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index 54d9509b7e..b2938ee014 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -7,6 +7,7 @@ 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.Data.Nat open import Cubical.Data.Unit.Base @@ -18,6 +19,8 @@ open import Cubical.Foundations.Univalence open import Cubical.Reflection.StrictEquiv +open Iso + private variable ℓ ℓ' : Level @@ -67,11 +70,17 @@ module _ (A : Unit* {ℓ} → Type ℓ') where ΠUnit* : ((x : Unit*) → A x) ≃ A tt* ΠUnit* = isoToEquiv ΠUnit*Iso +fiberUnitIso : {A : Type ℓ} → Iso (fiber (λ (a : A) → tt) tt) A +fun fiberUnitIso = fst +inv fiberUnitIso a = a , refl +rightInv fiberUnitIso _ = refl +leftInv fiberUnitIso _ = refl + isContr→Iso2 : {A : Type ℓ} {B : Type ℓ'} → isContr A → Iso (A → B) B -Iso.fun (isContr→Iso2 iscontr) f = f (fst iscontr) -Iso.inv (isContr→Iso2 iscontr) b _ = b -Iso.rightInv (isContr→Iso2 iscontr) _ = refl -Iso.leftInv (isContr→Iso2 iscontr) f = funExt λ x → cong f (snd iscontr x) +fun (isContr→Iso2 iscontr) f = f (fst iscontr) +inv (isContr→Iso2 iscontr) b _ = b +rightInv (isContr→Iso2 iscontr) _ = refl +leftInv (isContr→Iso2 iscontr) f = funExt λ x → cong f (snd iscontr x) diagonal-unit : Unit ≡ Unit × Unit diagonal-unit = isoToPath (iso (λ x → tt , tt) (λ x → tt) (λ {(tt , tt) i → tt , tt}) λ {tt i → tt}) @@ -93,6 +102,9 @@ isContrUnit* = tt* , λ _ → refl isPropUnit* : ∀ {ℓ} → isProp (Unit* {ℓ}) isPropUnit* _ _ = refl +isSetUnit* : ∀ {ℓ} → isSet (Unit* {ℓ}) +isSetUnit* _ _ _ _ = refl + isOfHLevelUnit* : ∀ {ℓ} (n : HLevel) → isOfHLevel n (Unit* {ℓ}) isOfHLevelUnit* zero = tt* , λ _ → refl isOfHLevelUnit* (suc zero) _ _ = refl diff --git a/Cubical/Displayed/Base.agda b/Cubical/Displayed/Base.agda index 9491845a7b..e91565a559 100644 --- a/Cubical/Displayed/Base.agda +++ b/Cubical/Displayed/Base.agda @@ -40,12 +40,12 @@ record UARel (A : Type ℓA) (ℓ≅A : Level) : Type (ℓ-max ℓA (ℓ-suc ℓ open BinaryRelation -- another constructor for UARel using contractibility of relational singletons -make-𝒮 : {A : Type ℓA} {ℓ≅A : Level} {_≅_ : A → A → Type ℓ≅A} +make-𝒮 : {A : Type ℓA} {_≅_ : A → A → Type ℓ≅A} (ρ : isRefl _≅_) (contrTotal : contrRelSingl _≅_) → UARel A ℓ≅A UARel._≅_ (make-𝒮 {_≅_ = _≅_} _ _) = _≅_ UARel.ua (make-𝒮 {_≅_ = _≅_} ρ c) = contrRelSingl→isUnivalent _≅_ ρ c -record DUARel {A : Type ℓA} {ℓ≅A : Level} (𝒮-A : UARel A ℓ≅A) +record DUARel {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) (ℓ≅B : Level) : Type (ℓ-max (ℓ-max ℓA ℓB) (ℓ-max ℓ≅A (ℓ-suc ℓ≅B))) where no-eta-equality constructor duarel @@ -70,7 +70,7 @@ record DUARel {A : Type ℓA} {ℓ≅A : Level} (𝒮-A : UARel A ℓ≅A) -- total UARel induced by a DUARel -module _ {A : Type ℓA} {ℓ≅A : Level} {𝒮-A : UARel A ℓ≅A} +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : A → Type ℓB} {ℓ≅B : Level} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) where diff --git a/Cubical/Displayed/Function.agda b/Cubical/Displayed/Function.agda index 9132432b3c..4c977fc792 100644 --- a/Cubical/Displayed/Function.agda +++ b/Cubical/Displayed/Function.agda @@ -110,34 +110,6 @@ module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} -- SubstRel on a dependent function type -- from a SubstRel on the domain and SubstRel on the codomain -equivΠ' : ∀ {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} - {B : A → Type ℓB} {B' : A' → Type ℓB'} - (eA : A ≃ A') - (eB : {a : A} {a' : A'} → eA .fst a ≡ a' → B a ≃ B' a') - → ((a : A) → B a) ≃ ((a' : A') → B' a') -equivΠ' {B' = B'} eA eB = isoToEquiv isom - where - open Iso - - isom : Iso _ _ - isom .fun f a' = - eB (secEq eA a') .fst (f (invEq eA a')) - isom .inv f' a = - invEq (eB refl) (f' (eA .fst a)) - isom .rightInv f' = - funExt λ a' → - J (λ a'' p → eB p .fst (invEq (eB refl) (f' (p i0))) ≡ f' a'') - (secEq (eB refl) (f' (eA .fst (invEq eA a')))) - (secEq eA a') - isom .leftInv f = - funExt λ a → - subst - (λ p → invEq (eB refl) (eB p .fst (f (invEq eA (eA .fst a)))) ≡ f a) - (sym (commPathIsEq (eA .snd) a)) - (J (λ a'' p → invEq (eB refl) (eB (cong (eA .fst) p) .fst (f (invEq eA (eA .fst a)))) ≡ f a'') - (retEq (eB refl) (f (invEq eA (eA .fst a)))) - (retEq eA a)) - module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : A → Type ℓB} (𝒮ˢ-B : SubstRel 𝒮-A B) {C : Σ A B → Type ℓC} (𝒮ˢ-C : SubstRel (∫ˢ 𝒮ˢ-B) C) diff --git a/Cubical/Displayed/Properties.agda b/Cubical/Displayed/Properties.agda index 037a20bed0..8d59bf8dd1 100644 --- a/Cubical/Displayed/Properties.agda +++ b/Cubical/Displayed/Properties.agda @@ -21,7 +21,7 @@ private -- induction principles -module _ {A : Type ℓA} {ℓ≅A : Level} (𝒮-A : UARel A ℓ≅A) where +module _ {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) where open UARel 𝒮-A 𝒮-J : {a : A} (P : (a' : A) → (p : a ≡ a') → Type ℓ) @@ -44,7 +44,7 @@ module _ {A : Type ℓA} {ℓ≅A : Level} (𝒮-A : UARel A ℓ≅A) where = 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) + g = 𝒮-J (λ y q → P y (≡→≅ q)) d p -- constructors @@ -81,11 +81,9 @@ module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} -- constructor that reduces univalence further to contractibility of relational singletons 𝒮ᴰ-make-2 : (ρᴰ : {a : A} → isRefl _≅ᴰ⟨ ρ a ⟩_) - (contrTotal : (a : A) → contrRelSingl _≅ᴰ⟨ UARel.ρ 𝒮-A a ⟩_) + (contrTotal : (a : A) → contrRelSingl _≅ᴰ⟨ ρ a ⟩_) → DUARel 𝒮-A B ℓ≅B - DUARel._≅ᴰ⟨_⟩_ (𝒮ᴰ-make-2 ρᴰ contrTotal) = _≅ᴰ⟨_⟩_ - DUARel.uaᴰ (𝒮ᴰ-make-2 ρᴰ contrTotal) - = 𝒮ᴰ-make-aux (contrRelSingl→isUnivalent _ ρᴰ (contrTotal _)) + 𝒮ᴰ-make-2 ρᴰ contrTotal = 𝒮ᴰ-make-1 (contrRelSingl→isUnivalent _ ρᴰ (contrTotal _)) -- relational isomorphisms diff --git a/Cubical/Displayed/Subst.agda b/Cubical/Displayed/Subst.agda index 76f7b08a32..a05712db58 100644 --- a/Cubical/Displayed/Subst.agda +++ b/Cubical/Displayed/Subst.agda @@ -54,7 +54,7 @@ DUARel.uaᴰ (Subst→DUA {𝒮-A = 𝒮-A} {B = B} 𝒮ˢ-B) b p b' = ≃⟨ invEquiv (compPathlEquiv (sym (SubstRel.uaˢ 𝒮ˢ-B p b))) ⟩ subst B (≅→≡ p) b ≡ b' ≃⟨ invEquiv (PathP≃Path (λ i → B (≅→≡ p i)) b b') ⟩ - PathP (λ i → B (UARel.≅→≡ 𝒮-A p i)) b b' + PathP (λ i → B (≅→≡ p i)) b b' ■ where open UARel 𝒮-A diff --git a/Cubical/Experiments/Brunerie.agda b/Cubical/Experiments/Brunerie.agda index b60c7afe83..2f32d81ec0 100644 --- a/Cubical/Experiments/Brunerie.agda +++ b/Cubical/Experiments/Brunerie.agda @@ -5,7 +5,6 @@ open import Cubical.Foundations.Everything open import Cubical.Data.Bool open import Cubical.Data.Nat open import Cubical.Data.Int -open import Cubical.Data.HomotopyGroup open import Cubical.HITs.S1 open import Cubical.HITs.S2 open import Cubical.HITs.S3 @@ -19,10 +18,9 @@ open S¹Hopf -- This code is adapted from examples/brunerie3.ctt on the pi4s3_nobug branch of cubicaltt -Bool∙ S¹∙ S²∙ S³∙ : Pointed₀ +Bool∙ S¹∙ S³∙ : Pointed₀ Bool∙ = (Bool , true) S¹∙ = (S¹ , base) -S²∙ = (S² , base) S³∙ = (S³ , base) ∥_∥₃∙ ∥_∥₄∙ : Pointed₀ → Pointed₀ diff --git a/Cubical/Experiments/ZCohomologyOld/Properties.agda b/Cubical/Experiments/ZCohomologyOld/Properties.agda index 6f92b8e4a1..d231e8a5a8 100644 --- a/Cubical/Experiments/ZCohomologyOld/Properties.agda +++ b/Cubical/Experiments/ZCohomologyOld/Properties.agda @@ -32,7 +32,6 @@ open import Cubical.Data.NatMinusOne open import Cubical.HITs.Pushout open import Cubical.Data.Sum.Base -open import Cubical.Data.HomotopyGroup open import Cubical.Experiments.ZCohomologyOld.KcompPrelims diff --git a/Cubical/Algebra/ZariskiLattice/BasicOpens.agda b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda similarity index 93% rename from Cubical/Algebra/ZariskiLattice/BasicOpens.agda rename to Cubical/Experiments/ZariskiLatticeBasicOpens.agda index d754105660..306f1e6457 100644 --- a/Cubical/Algebra/ZariskiLattice/BasicOpens.agda +++ b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Algebra.ZariskiLattice.BasicOpens where +module Cubical.Experiments.ZariskiLatticeBasicOpens where open import Cubical.Foundations.Prelude @@ -15,7 +15,7 @@ open import Cubical.Functions.FunExtEquiv import Cubical.Data.Empty as ⊥ open import Cubical.Data.Bool -open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) open import Cubical.Data.Sigma.Base @@ -54,13 +54,14 @@ module Presheaf (A' : CommRing ℓ) where ; ·Lid to ·rLid ; ·Rid to ·rRid) open Exponentiation A' open CommRingTheory A' + open InvertingElementsBase A' open isMultClosedSubset open CommAlgebraStr ⦃...⦄ private A = fst A' A[1/_] : A → CommAlgebra A' ℓ - A[1/ x ] = AlgLoc.S⁻¹RAsCommAlg A' ([_ⁿ|n≥0] A' x) (powersFormMultClosedSubset _ _) + A[1/ x ] = AlgLoc.S⁻¹RAsCommAlg A' [ x ⁿ|n≥0] (powersFormMultClosedSubset _) A[1/_]ˣ : (x : A) → ℙ (fst A[1/ x ]) A[1/ x ]ˣ = (CommAlgebra→CommRing A[1/ x ]) ˣ @@ -102,7 +103,7 @@ module Presheaf (A' : CommRing ℓ) where RpropValued : isPropValued R RpropValued x y = isProp× isPropPropTrunc isPropPropTrunc - powerIs≽ : (x a : A) → x ∈ ([_ⁿ|n≥0] A' a) → a ≼ x + powerIs≽ : (x a : A) → x ∈ [ a ⁿ|n≥0] → a ≼ x powerIs≽ x a = map powerIs≽Σ where powerIs≽Σ : Σ[ n ∈ ℕ ] (x ≡ a ^ n) → Σ[ n ∈ ℕ ] Σ[ z ∈ A ] (a ^ n ≡ z ·r x) @@ -123,12 +124,12 @@ module Presheaf (A' : CommRing ℓ) where lemmaΣ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y → y ⋆ 1a ∈ A[1/ x ]ˣ lemmaΣ (n , z , p) = [ z , (x ^ n) , PT.∣ n , refl ∣ ] -- xⁿ≡zy → y⁻¹ ≡ z/xⁿ - , eq/ _ _ ((1r , powersFormMultClosedSubset _ _ .containsOne) + , eq/ _ _ ((1r , powersFormMultClosedSubset _ .containsOne) , (path1 _ _ ∙∙ sym p ∙∙ path2 _)) module ≼PowerToLoc (x y : A) (x≼y : x ≼ y) where private - [yⁿ|n≥0] = [_ⁿ|n≥0] A' y + [yⁿ|n≥0] = [ y ⁿ|n≥0] instance _ = snd A[1/ x ] lemma : ∀ (s : A) → s ∈ [yⁿ|n≥0] → s ⋆ 1a ∈ A[1/ x ]ˣ @@ -143,14 +144,14 @@ module Presheaf (A' : CommRing ℓ) where RCoh a b (a≼b , b≼a) = fst (isContrS₁⁻¹R≡S₂⁻¹R (≼PowerToLoc.lemma _ _ b≼a) (≼PowerToLoc.lemma _ _ a≼b)) where - open AlgLocTwoSubsets A' ([_ⁿ|n≥0] A' a) (powersFormMultClosedSubset _ _) - ([_ⁿ|n≥0] A' b) (powersFormMultClosedSubset _ _) + open AlgLocTwoSubsets A' [ a ⁿ|n≥0] (powersFormMultClosedSubset _) + [ b ⁿ|n≥0] (powersFormMultClosedSubset _) LocPathProp : ∀ a b → isProp (A[1/ a ] ≡ A[1/ b ]) LocPathProp a b = isPropS₁⁻¹R≡S₂⁻¹R where - open AlgLocTwoSubsets A' ([_ⁿ|n≥0] A' a) (powersFormMultClosedSubset _ _) - ([_ⁿ|n≥0] A' b) (powersFormMultClosedSubset _ _) + open AlgLocTwoSubsets A' [ a ⁿ|n≥0] (powersFormMultClosedSubset _) + [ b ⁿ|n≥0] (powersFormMultClosedSubset _) -- The quotient A/R corresponds to the basic opens of the Zariski topology. @@ -223,7 +224,7 @@ module Presheaf (A' : CommRing ℓ) where ρᴰᴬ : (a b : A) → a ≼ b → isContr (CommAlgebraHom A[1/ b ] A[1/ a ]) ρᴰᴬ _ b a≼b = A[1/b]HasUniversalProp _ (≼PowerToLoc.lemma _ _ a≼b) where - open AlgLoc A' ([_ⁿ|n≥0] A' b) (powersFormMultClosedSubset _ _) + open AlgLoc A' [ b ⁿ|n≥0] (powersFormMultClosedSubset _) renaming (S⁻¹RHasAlgUniversalProp to A[1/b]HasUniversalProp) ρᴰᴬId : ∀ (a : A) (r : a ≼ a) → ρᴰᴬ a a r .fst ≡ idAlgHom diff --git a/Cubical/Foundations/Equiv.agda b/Cubical/Foundations/Equiv.agda index 1527a480ec..8d9bd9db71 100644 --- a/Cubical/Foundations/Equiv.agda +++ b/Cubical/Foundations/Equiv.agda @@ -20,7 +20,6 @@ module Cubical.Foundations.Equiv where open import Cubical.Foundations.Function open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv.Base public open import Cubical.Data.Sigma.Base @@ -30,6 +29,8 @@ private ℓ ℓ' ℓ'' : Level A B C D : Type ℓ +infixr 30 _∙ₑ_ + equivIsEquiv : (e : A ≃ B) → isEquiv (equivFun e) equivIsEquiv e = snd e @@ -150,6 +151,8 @@ compEquiv {A = A} {C = C} f g .snd .equiv-proof c = contr }) (g .fst (f .snd .equiv-proof (f .fst a) .snd (a , refl) k .snd j)) +_∙ₑ_ = compEquiv + compEquivIdEquiv : (e : A ≃ B) → compEquiv (idEquiv A) e ≡ e compEquivIdEquiv e = equivEq refl @@ -230,35 +233,42 @@ equiv→Iso h k .Iso.leftInv f = funExt λ a → retEq k _ ∙ cong f (retEq h a equiv→ : (A ≃ B) → (C ≃ D) → (A → C) ≃ (B → D) equiv→ h k = isoToEquiv (equiv→Iso h k) -equivΠ : ∀ {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + +equivΠ' : ∀ {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} {B : A → Type ℓB} {B' : A' → Type ℓB'} (eA : A ≃ A') - (eB : (a : A) → B a ≃ B' (eA .fst a)) + (eB : {a : A} {a' : A'} → eA .fst a ≡ a' → B a ≃ B' a') → ((a : A) → B a) ≃ ((a' : A') → B' a') -equivΠ {B' = B'} eA eB = isoToEquiv isom +equivΠ' {B' = B'} eA eB = isoToEquiv isom where open Iso isom : Iso _ _ isom .fun f a' = - subst B' (secEq eA a') (eB _ .fst (f (invEq eA a'))) + eB (secEq eA a') .fst (f (invEq eA a')) isom .inv f' a = - invEq (eB _) (f' (eA .fst a)) + invEq (eB refl) (f' (eA .fst a)) isom .rightInv f' = funExt λ a' → - cong (subst B' (secEq eA a')) (secEq (eB _) _) - ∙ fromPathP (cong f' (secEq eA a')) + J (λ a'' p → eB p .fst (invEq (eB refl) (f' (p i0))) ≡ f' a'') + (secEq (eB refl) (f' (eA .fst (invEq eA a')))) + (secEq eA a') isom .leftInv f = funExt λ a → - invEq (eB a) (subst B' (secEq eA _) (eB _ .fst (f (invEq eA (eA .fst a))))) - ≡⟨ cong (λ t → invEq (eB a) (subst B' t (eB _ .fst (f (invEq eA (eA .fst a)))))) - (commPathIsEq (snd eA) a) ⟩ - invEq (eB a) (subst B' (cong (eA .fst) (retEq eA a)) (eB _ .fst (f (invEq eA (eA .fst a))))) - ≡⟨ cong (invEq (eB a)) (fromPathP (λ i → eB _ .fst (f (retEq eA a i)))) ⟩ - invEq (eB a) (eB a .fst (f a)) - ≡⟨ retEq (eB _) (f a) ⟩ - f a - ∎ + subst + (λ p → invEq (eB refl) (eB p .fst (f (invEq eA (eA .fst a)))) ≡ f a) + (sym (commPathIsEq (eA .snd) a)) + (J (λ a'' p → invEq (eB refl) (eB (cong (eA .fst) p) .fst (f (invEq eA (eA .fst a)))) ≡ f a'') + (retEq (eB refl) (f (invEq eA (eA .fst a)))) + (retEq eA a)) + +equivΠ : ∀ {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + {B : A → Type ℓB} {B' : A' → Type ℓB'} + (eA : A ≃ A') + (eB : (a : A) → B a ≃ B' (eA .fst a)) + → ((a : A) → B a) ≃ ((a' : A') → B' a') +equivΠ {B = B} {B' = B'} eA eB = equivΠ' eA (λ {a = a} p → J (λ a' p → B a ≃ B' a') (eB a) p) + equivCompIso : (A ≃ B) → (C ≃ D) → Iso (A ≃ C) (B ≃ D) equivCompIso h k .Iso.fun f = compEquiv (compEquiv (invEquiv h) f) k diff --git a/Cubical/Foundations/Equiv/HalfAdjoint.agda b/Cubical/Foundations/Equiv/HalfAdjoint.agda index e589d5d068..1f1167ec30 100644 --- a/Cubical/Foundations/Equiv/HalfAdjoint.agda +++ b/Cubical/Foundations/Equiv/HalfAdjoint.agda @@ -31,40 +31,15 @@ record isHAEquiv {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) : Type rinv : ∀ b → f (g b) ≡ b com : ∀ a → cong f (linv a) ≡ rinv (f a) - -- from redtt's ha-equiv/symm com-op : ∀ b → cong g (rinv b) ≡ linv (g b) - com-op b j i = hcomp (λ k → λ { (i = i0) → linv (g b) (j ∧ (~ k)) - ; (j = i0) → g (rinv b i) - ; (j = i1) → linv (g b) (i ∨ (~ k)) - ; (i = i1) → g b }) - (cap1 j i) - - where cap0 : Square {- (j = i0) -} (λ i → f (g (rinv b i))) - {- (j = i1) -} (λ i → rinv b i) - {- (i = i0) -} (λ j → f (linv (g b) j)) - {- (i = i1) -} (λ j → rinv b j) - - cap0 j i = hcomp (λ k → λ { (i = i0) → com (g b) (~ k) j - ; (j = i0) → f (g (rinv b i)) - ; (j = i1) → rinv b i - ; (i = i1) → rinv b j }) - (rinv (rinv b i) j) - - filler : I → I → A - filler j i = hfill (λ k → λ { (i = i0) → g (rinv b k) - ; (i = i1) → g b }) - (inS (linv (g b) i)) j - - cap1 : Square {- (j = i0) -} (λ i → g (rinv b i)) - {- (j = i1) -} (λ i → g b) - {- (i = i0) -} (λ j → linv (g b) j) - {- (i = i1) -} (λ j → g b) - - cap1 j i = hcomp (λ k → λ { (i = i0) → linv (linv (g b) j) k - ; (j = i0) → linv (g (rinv b i)) k - ; (j = i1) → filler i k - ; (i = i1) → filler j k }) - (g (cap0 j i)) + com-op b = subst (λ b → cong g (rinv b) ≡ linv (g b)) (rinv b) (helper (g b)) + where + helper : ∀ a → cong g (rinv (f a)) ≡ linv (g (f a)) + helper a i j = hcomp (λ k → λ { (i = i0) → g (rinv (f a) (j ∨ ~ k)) + ; (i = i1) → linv (linv a (~ k)) j + ; (j = i0) → g (com a (~ i) (~ k)) + ; (j = i1) → linv a (i ∧ ~ k) }) + (linv a (i ∧ j)) isHAEquiv→Iso : Iso A B Iso.fun isHAEquiv→Iso = f diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index 665cc30d82..3dafd58f2f 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -190,12 +190,20 @@ doubleCompPath-elim' : {ℓ : Level} {A : Type ℓ} {w x y z : A} (p : w ≡ x) (r : y ≡ z) → (p ∙∙ q ∙∙ r) ≡ p ∙ (q ∙ r) doubleCompPath-elim' p q r = (split-leftright' p q r) ∙ (sym (leftright p (q ∙ r))) +cong-∙∙-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y z w : A} + (f : A → B) (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) + → I → I → I → B +cong-∙∙-filler {A = A} f p q r k j i = + hfill ((λ k → λ { (j = i1) → doubleCompPath-filler (cong f p) (cong f q) (cong f r) k i + ; (j = i0) → f (doubleCompPath-filler p q r k i) + ; (i = i0) → f (p (~ k)) + ; (i = i1) → f (r k) })) + (inS (f (q i))) + k + cong-∙∙ : ∀ {B : Type ℓ} (f : A → B) (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → cong f (p ∙∙ q ∙∙ r) ≡ (cong f p) ∙∙ (cong f q) ∙∙ (cong f r) -cong-∙∙ f p q r j i = hcomp (λ k → λ { (j = i0) → f (doubleCompPath-filler p q r k i) - ; (i = i0) → f (p (~ k)) - ; (i = i1) → f (r k) }) - (f (q i)) +cong-∙∙ f p q r j i = cong-∙∙-filler f p q r i1 j i cong-∙ : ∀ {B : Type ℓ} (f : A → B) (p : x ≡ y) (q : y ≡ z) → cong f (p ∙ q) ≡ (cong f p) ∙ (cong f q) diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index 6193637d27..4fe950efbf 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -301,6 +301,9 @@ isGroupoid'→isGroupoid : isGroupoid' A → isGroupoid A isGroupoid'→isGroupoid Agpd' x y p q r s = Agpd' r s refl refl refl refl -- h-level of Σ-types +isProp∃! : isProp (∃! A B) +isProp∃! = isPropIsContr + isContrΣ : isContr A → ((x : A) → isContr (B x)) → isContr (Σ A B) isContrΣ {A = A} {B = B} (a , p) q = let h : (x : A) (y : B x) → (q x) .fst ≡ y @@ -451,6 +454,9 @@ isProp→ pB = isPropΠ λ _ → pB isSetΠ : ((x : A) → isSet (B x)) → isSet ((x : A) → B x) isSetΠ = isOfHLevelΠ 2 +isSetImplicitΠ : (h : (x : A) → isSet (B x)) → isSet ({x : A} → B x) +isSetImplicitΠ h f g F G i j {x} = h x (f {x}) (g {x}) (λ i → F i {x}) (λ i → G i {x}) i j + isSetΠ2 : (h : (x : A) (y : B x) → isSet (C x y)) → isSet ((x : A) (y : B x) → C x y) isSetΠ2 h = isSetΠ λ x → isSetΠ λ y → h x y @@ -593,11 +599,11 @@ isOfHLevel→isOfHLevelDep 0 h {a} = (h a .fst , λ b' p → isProp→PathP (λ i → isContr→isProp (h (p i))) (h a .fst) b') isOfHLevel→isOfHLevelDep 1 h = λ b0 b1 p → isProp→PathP (λ i → h (p i)) b0 b1 isOfHLevel→isOfHLevelDep (suc (suc n)) {A = A} {B} h {a0} {a1} b0 b1 = - isOfHLevel→isOfHLevelDep (suc n) (λ p → helper a1 p b1) + isOfHLevel→isOfHLevelDep (suc n) (λ p → helper p) where - helper : (a1 : A) (p : a0 ≡ a1) (b1 : B a1) → + helper : (p : a0 ≡ a1) → isOfHLevel (suc n) (PathP (λ i → B (p i)) b0 b1) - helper a1 p b1 = J (λ a1 p → ∀ b1 → isOfHLevel (suc n) (PathP (λ i → B (p i)) b0 b1)) + helper p = J (λ a1 p → ∀ b1 → isOfHLevel (suc n) (PathP (λ i → B (p i)) b0 b1)) (λ _ → h _ _ _) p b1 isContrDep→isPropDep : isOfHLevelDep 0 B → isOfHLevelDep 1 B @@ -653,3 +659,15 @@ isOfHLevelΣ' {A = A} {B = B} (suc (suc n)) Alvl Blvl (w , y) (x , z) ΣPathP (λ x → refl) (isOfHLevelΣ' (suc n) (Alvl w x) (Blvl y z)) + +ΣSquareSet : ((x : A) → isSet (B x)) → {u v w x : Σ A B} + → {p : u ≡ v} {q : v ≡ w} {r : x ≡ w} {s : u ≡ x} + → Square (cong fst p) (cong fst r) + (cong fst s) (cong fst q) + → Square p r s q +fst (ΣSquareSet pB sq i j) = sq i j +snd (ΣSquareSet {B = B} pB {p = p} {q = q} {r = r} {s = s} sq i j) = lem i j + where + lem : SquareP (λ i j → B (sq i j)) + (cong snd p) (cong snd r) (cong snd s) (cong snd q) + lem = toPathP (isOfHLevelPathP' 1 (pB _) _ _ _ _) diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index a6a9d874f0..67588c3d72 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -16,8 +16,6 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv.Base -open import Cubical.Foundations.Function - private variable ℓ ℓ' : Level diff --git a/Cubical/Foundations/Path.agda b/Cubical/Foundations/Path.agda index 4452fc272a..ded7242604 100644 --- a/Cubical/Foundations/Path.agda +++ b/Cubical/Foundations/Path.agda @@ -2,6 +2,7 @@ module Cubical.Foundations.Path where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism @@ -85,6 +86,14 @@ PathP≡compPath : ∀ {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) (r → (PathP (λ i → x ≡ q i) p r) ≡ (p ∙ q ≡ r) PathP≡compPath p q r k = PathP (λ i → p i0 ≡ q (i ∨ k)) (λ j → compPath-filler p q k j) r +-- a quick corollary for 3-constant functions +3-ConstantCompChar : {A : Type ℓ} {B : Type ℓ'} (f : A → B) (link : 2-Constant f) + → (∀ x y z → link x y ∙ link y z ≡ link x z) + → 3-Constant f +3-Constant.link (3-ConstantCompChar f link coh₂) = link +3-Constant.coh₁ (3-ConstantCompChar f link coh₂) _ _ _ = + transport⁻ (PathP≡compPath _ _ _) (coh₂ _ _ _) + PathP≡doubleCompPathˡ : ∀ {A : Type ℓ} {w x y z : A} (p : w ≡ y) (q : w ≡ x) (r : y ≡ z) (s : x ≡ z) → (PathP (λ i → p i ≡ s i) q r) ≡ (p ⁻¹ ∙∙ q ∙∙ s ≡ r) PathP≡doubleCompPathˡ p q r s k = PathP (λ i → p (i ∨ k) ≡ s (i ∨ k)) @@ -221,8 +230,8 @@ sym≡cong-sym : ∀ {ℓ} {A : Type ℓ} {x : A} (P : Square (refl {x = x}) ref sym≡cong-sym P = sym-cong-sym≡id (sym P) -- sym induces an equivalence on identity types of paths -symIso : {a b : A} (p q : a ≡ b) → Iso (p ≡ q) (q ≡ p) -symIso p q = iso sym sym (λ _ → refl) λ _ → refl +symIso : {a b : A} → Iso (a ≡ b) (b ≡ a) +symIso = iso sym sym (λ _ → refl) λ _ → refl -- J is an equivalence Jequiv : {x : A} (P : ∀ y → x ≡ y → Type ℓ') → P x refl ≃ (∀ {y} (p : x ≡ y) → P y p) diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index 3b93556344..1cc03d57e3 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -32,6 +32,48 @@ idfun∙ : ∀ {ℓ} (A : Pointed ℓ) → A →∙ A idfun∙ A .fst x = x idfun∙ A .snd = refl +{-Pointed equivalences -} +_≃∙_ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') +A ≃∙ B = Σ[ e ∈ fst A ≃ fst B ] fst e (pt A) ≡ pt B + +{- Underlying pointed map of an equivalence -} +≃∙map : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ≃∙ B → A →∙ B +fst (≃∙map e) = fst (fst e) +snd (≃∙map e) = snd e + +invEquiv∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ≃∙ B → B ≃∙ A +fst (invEquiv∙ x) = invEquiv (fst x) +snd (invEquiv∙ {A = A} x) = + sym (cong (fst (invEquiv (fst x))) (snd x)) ∙ retEq (fst x) (pt A) + +compEquiv∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → A ≃∙ B → B ≃∙ C → A ≃∙ C +fst (compEquiv∙ e1 e2) = compEquiv (fst e1) (fst e2) +snd (compEquiv∙ e1 e2) = cong (fst (fst e2)) (snd e1) ∙ snd e2 + +Equiv∙J : ∀ {ℓ ℓ'} {B : Pointed ℓ} (C : (A : Pointed ℓ) → A ≃∙ B → Type ℓ') + → C B (idEquiv (fst B) , refl) + → {A : _} (e : A ≃∙ B) → C A e +Equiv∙J {ℓ} {ℓ'} {B = B} C ind {A = A} = + uncurry λ e p → help e (pt A) (pt B) p C ind + where + help : ∀ {A : Type ℓ} (e : A ≃ typ B) (a : A) (b : typ B) + → (p : fst e a ≡ b) + → (C : (A : Pointed ℓ) → A ≃∙ (fst B , b) → Type ℓ') + → C (fst B , b) (idEquiv (fst B) , refl) + → C (A , a) (e , p) + help = EquivJ (λ A e → (a : A) (b : typ B) + → (p : fst e a ≡ b) + → (C : (A : Pointed ℓ) → A ≃∙ (fst B , b) → Type ℓ') + → C (fst B , b) (idEquiv (fst B) , refl) + → C (A , a) (e , p)) + λ a b → J (λ b p + → (C : (A : Pointed ℓ) → A ≃∙ (fst B , b) → Type ℓ') + → C (fst B , b) + (idEquiv (fst B) , refl) → + C (typ B , a) (idEquiv (typ B) , p)) + λ _ p → p + ua∙ : ∀ {ℓ} {A B : Pointed ℓ} (e : fst A ≃ fst B) → fst e (snd A) ≡ snd B → A ≡ B fst (ua∙ e p i) = ua e i diff --git a/Cubical/Foundations/Pointed/Homotopy.agda b/Cubical/Foundations/Pointed/Homotopy.agda index 06612f6b2d..8cc1efe8ce 100644 --- a/Cubical/Foundations/Pointed/Homotopy.agda +++ b/Cubical/Foundations/Pointed/Homotopy.agda @@ -61,7 +61,7 @@ module _ {A : Pointed ℓ} {B : typ A → Type ℓ'} {ptB : B (pt A)} where s = g₂ P≡Q : P ≡ Q P≡Q = p ≡ r ∙ s ⁻¹ - ≡⟨ isoToPath (symIso p (r ∙ s ⁻¹)) ⟩ + ≡⟨ isoToPath symIso ⟩ r ∙ s ⁻¹ ≡ p ≡⟨ cong (r ∙ s ⁻¹ ≡_) (rUnit p ∙∙ cong (p ∙_) (sym (rCancel s)) ∙∙ assoc p s (s ⁻¹)) ⟩ r ∙ s ⁻¹ ≡ (p ∙ s) ∙ s ⁻¹ diff --git a/Cubical/Foundations/Pointed/Properties.agda b/Cubical/Foundations/Pointed/Properties.agda index fa8693c720..a96ca1be0a 100644 --- a/Cubical/Foundations/Pointed/Properties.agda +++ b/Cubical/Foundations/Pointed/Properties.agda @@ -90,3 +90,10 @@ const∙ _ B .snd = refl (cong h (cong g f∙) ∙ cong h g∙) ∙ h∙ ≡⟨ cong (λ p → p ∙ h∙) ((cong-∙ h (cong g f∙) g∙) ⁻¹) ⟩ (cong h (cong g f∙ ∙ g∙) ∙ h∙) ∎ ) + +module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where + isInIm∙ : (x : typ B) → Type (ℓ-max ℓ ℓ') + isInIm∙ x = Σ[ z ∈ typ A ] fst f z ≡ x + + isInKer∙ : (x : fst A) → Type ℓ' + isInKer∙ x = fst f x ≡ snd B diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 94da7d6ee0..0f15d63467 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -321,16 +321,31 @@ module _ {A : I → Type ℓ} {x : A i0} {y : A i1} where fromPathP p i = transp (λ j → A (i ∨ j)) i (p i) -- Whiskering a dependent path by a path +-- Double whiskering +_◁_▷_ : ∀ {ℓ} {A : I → Type ℓ} {a₀ a₀' : A i0} {a₁ a₁' : A i1} + → a₀ ≡ a₀' → PathP A a₀' a₁ → a₁ ≡ a₁' + → PathP A a₀ a₁' +(p ◁ P ▷ q) i = + hcomp (λ j → λ {(i = i0) → p (~ j) ; (i = i1) → q j}) (P i) + +doubleWhiskFiller : + ∀ {ℓ} {A : I → Type ℓ} {a₀ a₀' : A i0} {a₁ a₁' : A i1} + → (p : a₀ ≡ a₀') (pq : PathP A a₀' a₁) (q : a₁ ≡ a₁') + → PathP (λ i → PathP A (p (~ i)) (q i)) + pq + (p ◁ pq ▷ q) +doubleWhiskFiller p pq q k i = + hfill (λ j → λ {(i = i0) → p (~ j) ; (i = i1) → q j}) + (inS (pq i)) + k _◁_ : ∀ {ℓ} {A : I → Type ℓ} {a₀ a₀' : A i0} {a₁ : A i1} → a₀ ≡ a₀' → PathP A a₀' a₁ → PathP A a₀ a₁ -(p ◁ q) i = - hcomp (λ j → λ {(i = i0) → p (~ j); (i = i1) → q i1}) (q i) +(p ◁ q) = p ◁ q ▷ refl _▷_ : ∀ {ℓ} {A : I → Type ℓ} {a₀ : A i0} {a₁ a₁' : A i1} → PathP A a₀ a₁ → a₁ ≡ a₁' → PathP A a₀ a₁' -(p ▷ q) i = - hcomp (λ j → λ {(i = i0) → p i0; (i = i1) → q j}) (p i) +p ▷ q = refl ◁ p ▷ q -- Direct definitions of lower h-levels @@ -457,7 +472,7 @@ isPropIsContr (c0 , h0) (c1 , h1) j .snd y i = c0 isContr→isProp : isContr A → isProp A -isContr→isProp (x , p) a b = sym (p a) ∙∙ refl ∙∙ p b +isContr→isProp (x , p) a b = sym (p a) ∙ p b isProp→isSet : isProp A → isSet A isProp→isSet h a b p q j i = diff --git a/Cubical/Foundations/Univalence.agda b/Cubical/Foundations/Univalence.agda index b113d83bd6..142959c56b 100644 --- a/Cubical/Foundations/Univalence.agda +++ b/Cubical/Foundations/Univalence.agda @@ -263,6 +263,14 @@ ua→⁻ {e = e} {f₀ = f₀} {f₁} p a i = }) (p i (transp (λ j → ua e (j ∧ i)) (~ i) a)) +ua→2 : ∀ {ℓ ℓ' ℓ''} {A₀ A₁ : Type ℓ} {e₁ : A₀ ≃ A₁} + {B₀ B₁ : Type ℓ'} {e₂ : B₀ ≃ B₁} + {C : (i : I) → Type ℓ''} + {f₀ : A₀ → B₀ → C i0} {f₁ : A₁ → B₁ → C i1} + → (∀ a b → PathP C (f₀ a b) (f₁ (e₁ .fst a) (e₂ .fst b))) + → PathP (λ i → ua e₁ i → ua e₂ i → C i) f₀ f₁ +ua→2 h = ua→ (ua→ ∘ h) + -- Useful lemma for unfolding a transported function over ua -- If we would have regularity this would be refl transportUAop₁ : ∀ {A B : Type ℓ} → (e : A ≃ B) (f : A → A) (x : B) diff --git a/Cubical/Functions/Morphism.agda b/Cubical/Functions/Morphism.agda index c4886cfdfb..b21e5c34ac 100644 --- a/Cubical/Functions/Morphism.agda +++ b/Cubical/Functions/Morphism.agda @@ -6,7 +6,6 @@ module Cubical.Functions.Morphism where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Function open Iso module ax {ℓ : Level} (A : Type ℓ) (_+A_ : A → A → A) (a₀ : A) where diff --git a/Cubical/HITs/AssocList/Properties.agda b/Cubical/HITs/AssocList/Properties.agda index 2133961615..2de9f19927 100644 --- a/Cubical/HITs/AssocList/Properties.agda +++ b/Cubical/HITs/AssocList/Properties.agda @@ -5,6 +5,7 @@ open import Cubical.HITs.AssocList.Base as AL open import Cubical.Foundations.Everything open import Cubical.Foundations.SIP open import Cubical.HITs.FiniteMultiset as FMS + hiding (_++_; unitl-++; unitr-++; assoc-++; cons-++; comm-++) open import Cubical.Data.Nat using (ℕ; zero; suc; _+_; +-assoc; isSetℕ) open import Cubical.Structures.MultiSet open import Cubical.Relation.Nullary @@ -46,6 +47,41 @@ multi-∷-agg : (x : A) (m n : ℕ) (b : FMSet A) → multi-∷ x m (multi-∷ x multi-∷-agg x zero n b = refl multi-∷-agg x (suc m) n b i = x ∷ (multi-∷-agg x m n b i) + +infixr 30 _++_ + +_++_ : (xs ys : AssocList A) → AssocList A +⟨⟩ ++ ys = ys +(⟨ a , n ⟩∷ xs) ++ ys = ⟨ a , n ⟩∷ (xs ++ ys) +per a b xs i ++ ys = per a b (xs ++ ys) i +agg a m n xs i ++ ys = agg a m n (xs ++ ys) i +del a xs i ++ ys = del a (xs ++ ys) i +trunc xs ys p q i j ++ zs = + trunc (xs ++ zs) (ys ++ zs) (cong (_++ _) p) (cong (_++ _) q) i j + +unitl-++ : (xs : AssocList A) → ⟨⟩ ++ xs ≡ xs +unitl-++ xs = refl + +unitr-++ : (xs : AssocList A) → xs ++ ⟨⟩ ≡ xs +unitr-++ = AL.ElimProp.f (trunc _ _) refl λ _ _ → cong (⟨ _ , _ ⟩∷_) + +assoc-++ : (xs ys zs : AssocList A) → xs ++ (ys ++ zs) ≡ (xs ++ ys) ++ zs +assoc-++ = AL.ElimProp.f (isPropΠ2 (λ _ _ → trunc _ _)) + (λ ys zs → refl) + λ x n p ys zs → cong (⟨ _ , _ ⟩∷_) (p ys zs) + +cons-++ : ∀ x n (xs : AssocList A) → ⟨ x , n ⟩∷ xs ≡ xs ++ (⟨ x , n ⟩∷ ⟨⟩) +cons-++ x n = AL.ElimProp.f (trunc _ _) refl + λ y m p → multiPer _ _ _ _ _ ∙ cong (⟨ _ , _ ⟩∷_) p + +comm-++ : (xs ys : AssocList A) → xs ++ ys ≡ ys ++ xs +comm-++ = AL.ElimProp.f (isPropΠ (λ _ → trunc _ _)) + (sym ∘ unitr-++) + λ x n {xs} p ys → cong (⟨ _ , _ ⟩∷_) (p ys) + ∙ cong (_++ _) (cons-++ x n ys) + ∙ sym (assoc-++ ys _ xs) + + AL→FMS : AssocList A → FMSet A AL→FMS = AL.Rec.f FMS.trunc [] multi-∷ comm multi-∷-agg λ _ _ → refl diff --git a/Cubical/HITs/FreeComMonoids.agda b/Cubical/HITs/FreeComMonoids.agda new file mode 100644 index 0000000000..c90d9c9be8 --- /dev/null +++ b/Cubical/HITs/FreeComMonoids.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.FreeComMonoids where + +open import Cubical.HITs.FreeComMonoids.Base public +open import Cubical.HITs.FreeComMonoids.Properties public diff --git a/Cubical/HITs/FreeComMonoids/Base.agda b/Cubical/HITs/FreeComMonoids/Base.agda new file mode 100644 index 0000000000..cd002b66d3 --- /dev/null +++ b/Cubical/HITs/FreeComMonoids/Base.agda @@ -0,0 +1,73 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.FreeComMonoids.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function + +private variable + ℓ : Level + A : Type ℓ + +data FreeComMonoid (A : Type ℓ) : Type ℓ where + ⟦_⟧ : A → FreeComMonoid A + ε : FreeComMonoid A + _·_ : FreeComMonoid A → FreeComMonoid A → FreeComMonoid A + comm : ∀ x y → x · y ≡ y · x + identityᵣ : ∀ x → x · ε ≡ x + identityₗ : ∀ x → ε · x ≡ x + assoc : ∀ x y z → x · (y · z) ≡ (x · y) · z + trunc : isSet (FreeComMonoid A) + +module Elim {ℓ'} {B : FreeComMonoid A → Type ℓ'} + (⟦_⟧* : (x : A) → B ⟦ x ⟧) + (ε* : B ε) + (_·*_ : ∀ {x y} → B x → B y → B (x · y)) + (comm* : ∀ {x y} → (xs : B x) (ys : B y) + → PathP (λ i → B (comm x y i)) (xs ·* ys) (ys ·* xs)) + (identityᵣ* : ∀ {x} → (xs : B x) + → PathP (λ i → B (identityᵣ x i)) (xs ·* ε*) xs) + (identityₗ* : ∀ {x} → (xs : B x) + → PathP (λ i → B (identityₗ x i)) (ε* ·* xs) xs) + (assoc* : ∀ {x y z} → (xs : B x) (ys : B y) (zs : B z) + → PathP (λ i → B (assoc x y z i)) (xs ·* (ys ·* zs)) ((xs ·* ys) ·* zs)) + (trunc* : ∀ xs → isSet (B xs)) where + + f : (xs : FreeComMonoid A) → B xs + f ⟦ x ⟧ = ⟦ x ⟧* + f ε = ε* + f (xs · ys) = f xs ·* f ys + f (comm xs ys i) = comm* (f xs) (f ys) i + f (identityᵣ xs i) = identityᵣ* (f xs) i + f (identityₗ xs i) = identityₗ* (f xs) i + f (assoc xs ys zs i) = assoc* (f xs) (f ys) (f zs) i + f (trunc xs ys p q i j) = isOfHLevel→isOfHLevelDep 2 trunc* (f xs) (f ys) + (cong f p) (cong f q) (trunc xs ys p q) i j + +module ElimProp {ℓ'} {B : FreeComMonoid A → Type ℓ'} + (BProp : {xs : FreeComMonoid A} → isProp (B xs)) + (⟦_⟧* : (x : A) → B ⟦ x ⟧) + (ε* : B ε) + (_·*_ : ∀ {x y} → B x → B y → B (x · y)) where + + f : (xs : FreeComMonoid A) → B xs + f = Elim.f ⟦_⟧* ε* _·*_ + (λ {x y} xs ys → toPathP (BProp (transport (λ i → B (comm x y i)) (xs ·* ys)) (ys ·* xs))) + (λ {x} xs → toPathP (BProp (transport (λ i → B (identityᵣ x i)) (xs ·* ε*)) xs)) + (λ {x} xs → toPathP (BProp (transport (λ i → B (identityₗ x i)) (ε* ·* xs)) xs)) + (λ {x y z} xs ys zs → toPathP (BProp (transport (λ i → B (assoc x y z i)) (xs ·* (ys ·* zs))) ((xs ·* ys) ·* zs))) + (λ _ → (isProp→isSet BProp)) + +module Rec {ℓ'} {B : Type ℓ'} (BType : isSet B) + (⟦_⟧* : (x : A) → B) + (ε* : B) + (_·*_ : B → B → B) + (comm* : (x y : B) → x ·* y ≡ y ·* x) + (identityᵣ* : (x : B) → x ·* ε* ≡ x) + (identityₗ* : (x : B) → ε* ·* x ≡ x) + (assoc* : (x y z : B) → x ·* (y ·* z) ≡ (x ·* y) ·* z) + where + + f : FreeComMonoid A → B + f = Elim.f ⟦_⟧* ε* _·*_ comm* identityᵣ* identityₗ* assoc* (const BType) diff --git a/Cubical/HITs/FreeComMonoids/Properties.agda b/Cubical/HITs/FreeComMonoids/Properties.agda new file mode 100644 index 0000000000..588aaaa9a4 --- /dev/null +++ b/Cubical/HITs/FreeComMonoids/Properties.agda @@ -0,0 +1,93 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.FreeComMonoids.Properties where + +open import Cubical.Foundations.Everything hiding (assoc; ⟨_⟩) + +open import Cubical.Data.Nat hiding (_·_ ; _^_) + +open import Cubical.HITs.FreeComMonoids.Base as FCM +open import Cubical.HITs.AssocList as AL + +private variable + ℓ : Level + A : Type ℓ + +multi-· : A → ℕ → FreeComMonoid A → FreeComMonoid A +multi-· x zero xs = xs +multi-· x (suc n) xs = ⟦ x ⟧ · multi-· x n xs + +_^_ : A → ℕ → FreeComMonoid A +x ^ n = multi-· x n ε + +^+≡ : ∀ (p : A) m n → ((p ^ m) · (p ^ n)) ≡ (p ^ (m + n)) +^+≡ p zero n = identityₗ _ +^+≡ p (suc m) n = sym (assoc _ _ _) ∙ cong (_ ·_) (^+≡ _ _ _) + +AL→FCM : AssocList A → FreeComMonoid A +AL→FCM = AL.Rec.f trunc ε (λ a n p → (a ^ n) · p) + per* agg* (const identityₗ) + where + per* : ∀ x y (mon : FreeComMonoid A) → + ((⟦ x ⟧ · ε) · ((⟦ y ⟧ · ε) · mon)) ≡ + ((⟦ y ⟧ · ε) · ((⟦ x ⟧ · ε) · mon)) + per* x y mon = + (⟦ x ⟧ · ε) · ((⟦ y ⟧ · ε) · mon) + ≡⟨ assoc _ _ _ ⟩ + ((⟦ x ⟧ · ε) · (⟦ y ⟧ · ε)) · mon + ≡⟨ cong (_· mon) (comm _ _) ⟩ + (((⟦ y ⟧ · ε) · (⟦ x ⟧ · ε)) · mon) + ≡⟨ sym (assoc _ _ _) ⟩ + ((⟦ y ⟧ · ε) · ((⟦ x ⟧ · ε) · mon)) ∎ + + agg* : ∀ a m n mon → + ((a ^ m) · ((a ^ n) · mon)) ≡ ((a ^ (m + n)) · mon) + agg* a m n mon = + ((a ^ m) · ((a ^ n) · mon)) + ≡⟨ assoc _ _ _ ⟩ + (((a ^ m) · (a ^ n)) · mon) + ≡⟨ cong (_· _) (^+≡ _ _ _) ⟩ + ((a ^ (m + n)) · mon) ∎ + +FCM→AL : FreeComMonoid A → AssocList A +FCM→AL = FCM.Rec.f trunc ⟨_⟩ ⟨⟩ _++_ comm-++ unitr-++ (λ _ → refl) assoc-++ + +^-id : (x : A) (m : ℕ) (u : FreeComMonoid A) + → FCM→AL ((x ^ m) · u) ≡ ⟨ x , m ⟩∷ FCM→AL u +^-id x zero u = cong FCM→AL (identityₗ u) ∙ sym (del _ _) +^-id x (suc m) u = + FCM→AL ((⟦ x ⟧ · (x ^ m)) · u) + ≡⟨ cong ⟨ x , 1 ⟩∷_ (^-id x m u) ⟩ + ⟨ x , 1 ⟩∷ ⟨ x , m ⟩∷ FCM→AL u + ≡⟨ agg _ _ _ _ ⟩ + ⟨ x , suc m ⟩∷ FCM→AL u ∎ + +++-· : (x y : AssocList A) + → AL→FCM (x ++ y) ≡ AL→FCM x · AL→FCM y +++-· = AL.ElimProp.f + (isPropΠ (λ _ → trunc _ _)) + (λ _ → sym (identityₗ _)) + λ x n {xs} p ys → + AL→FCM (((⟨ x , n ⟩∷ ⟨⟩) ++ xs) ++ ys) + ≡⟨ cong AL→FCM (cong (_++ ys) (comm-++ (⟨ x , n ⟩∷ ⟨⟩) xs) ∙ sym (assoc-++ xs _ ys)) ⟩ + AL→FCM (xs ++ (⟨ x , n ⟩∷ ys)) + ≡⟨ p _ ⟩ + (AL→FCM xs · ((x ^ n) · AL→FCM ys)) + ≡⟨ assoc (AL→FCM xs) _ _ ⟩ + ((AL→FCM xs · (x ^ n)) · AL→FCM ys) + ≡⟨ cong (_· AL→FCM ys) (comm _ _) ⟩ + ((x ^ n) · AL→FCM xs) · AL→FCM ys ∎ + +AL→FCM∘FCM→AL≡id : section {A = AssocList A} AL→FCM FCM→AL +AL→FCM∘FCM→AL≡id = FCM.ElimProp.f (trunc _ _) (λ x → identityᵣ _ ∙ identityᵣ _) refl + λ {x y} p q → ++-· (FCM→AL x) (FCM→AL y) ∙ cong₂ _·_ p q + +FCM→AL∘AL→FCM≡id : retract {A = AssocList A} AL→FCM FCM→AL +FCM→AL∘AL→FCM≡id = AL.ElimProp.f (trunc _ _) refl + λ x n {xs} p → ^-id x n (AL→FCM xs) ∙ cong (⟨ _ , _ ⟩∷_) p + +AssocList≃FreeComMonoid : AssocList A ≃ FreeComMonoid A +AssocList≃FreeComMonoid = isoToEquiv (iso AL→FCM FCM→AL AL→FCM∘FCM→AL≡id FCM→AL∘AL→FCM≡id) + +AssocList≡FreeComMonoid : AssocList A ≡ FreeComMonoid A +AssocList≡FreeComMonoid = ua AssocList≃FreeComMonoid diff --git a/Cubical/HITs/PropositionalTruncation/Properties.agda b/Cubical/HITs/PropositionalTruncation/Properties.agda index e6b7fed1d5..0594340184 100644 --- a/Cubical/HITs/PropositionalTruncation/Properties.agda +++ b/Cubical/HITs/PropositionalTruncation/Properties.agda @@ -209,6 +209,25 @@ elim→Set {A = A} {P = P} Pset f kf t gk : 2-Constant g gk x y i = transp (λ j → P (squash (squash ∣ x ∣ ∣ y ∣ i) t j)) i0 (kf x y i) +elim2→Set : + {P : ∥ A ∥ → ∥ B ∥ → Type ℓ} + → (∀ t u → isSet (P t u)) + → (f : (x : A) (y : B) → P ∣ x ∣ ∣ y ∣) + → (kf₁ : ∀ x y v → PathP (λ i → P (squash ∣ x ∣ ∣ y ∣ i) ∣ v ∣) (f x v) (f y v)) + → (kf₂ : ∀ x v w → PathP (λ i → P ∣ x ∣ (squash ∣ v ∣ ∣ w ∣ i)) (f x v) (f x w)) + → (sf : ∀ x y v w → SquareP (λ i j → P (squash ∣ x ∣ ∣ y ∣ i) (squash ∣ v ∣ ∣ w ∣ j)) + (kf₂ x v w) (kf₂ y v w) (kf₁ x y v) (kf₁ x y w)) + → (t : ∥ A ∥) → (u : ∥ B ∥) → P t u +elim2→Set {A = A} {B = B} {P = P} Pset f kf₁ kf₂ sf = + elim→Set (λ _ → isSetΠ (λ _ → Pset _ _)) mapHelper squareHelper + where + mapHelper : (x : A) (u : ∥ B ∥) → P ∣ x ∣ u + mapHelper x = elim→Set (λ _ → Pset _ _) (f x) (kf₂ x) + + squareHelper : (x y : A) + → PathP (λ i → (u : ∥ B ∥) → P (squash ∣ x ∣ ∣ y ∣ i) u) (mapHelper x) (mapHelper y) + squareHelper x y i = elim→Set (λ _ → Pset _ _) (λ v → kf₁ x y v i) λ v w → sf x y v w i + RecHProp : (P : A → hProp ℓ) (kP : ∀ x y → P x ≡ P y) → ∥ A ∥ → hProp ℓ RecHProp P kP = rec→Set isSetHProp P kP @@ -468,3 +487,13 @@ RecHSet P 3kP = rec→Gpd (isOfHLevelTypeOfHLevel 2) P 3kP ∥∥-× : {A : Type ℓ} {A′ : Type ℓ′} → ∥ A ∥ × ∥ A′ ∥ ≡ ∥ A × A′ ∥ ∥∥-× = ua ∥∥-×-≃ + +-- using this we get a convenient recursor/eliminator for binary functions into sets +rec2→Set : {A B C : Type ℓ} (Cset : isSet C) + → (f : A → B → C) + → (∀ (a a' : A) (b b' : B) → f a b ≡ f a' b') + → ∥ A ∥ → ∥ B ∥ → C +rec2→Set {A = A} {B = B} {C = C} Cset f fconst = curry (g ∘ ∥∥-×-≃ .fst) + where + g : ∥ A × B ∥ → C + g = rec→Set Cset (uncurry f) λ x y → fconst (fst x) (fst y) (snd x) (snd y) diff --git a/Cubical/HITs/Pushout/Base.agda b/Cubical/HITs/Pushout/Base.agda index 2be7530cd1..5a26c78e7f 100644 --- a/Cubical/HITs/Pushout/Base.agda +++ b/Cubical/HITs/Pushout/Base.agda @@ -4,8 +4,10 @@ module Cubical.HITs.Pushout.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Unit +open import Cubical.Data.Sigma open import Cubical.HITs.Susp.Base @@ -61,3 +63,55 @@ PushoutSusp≃Susp = isoToEquiv PushoutSuspIsoSusp PushoutSusp≡Susp : ∀ {ℓ} {A : Type ℓ} → PushoutSusp A ≡ Susp A PushoutSusp≡Susp = isoToPath PushoutSuspIsoSusp + +-- Generalised pushout, used in e.g. BlakersMassey +data PushoutGen {ℓ₁ ℓ₂ ℓ₃ : Level} {X : Type ℓ₁} {Y : Type ℓ₂} + (Q : X → Y → Type ℓ₃) : Type (ℓ-max (ℓ-max ℓ₁ ℓ₂) ℓ₃) + where + inl : X → PushoutGen Q + inr : Y → PushoutGen Q + push : {x : X} {y : Y} → Q x y → inl x ≡ inr y + +-- The usual pushout is a special case of the above +module _ {ℓ₁ ℓ₂ ℓ₃ : Level} {A : Type ℓ₁} {B : Type ℓ₂} {C : Type ℓ₃} + (f : A → B) (g : A → C) where + open Iso + + doubleFib : B → C → Type _ + doubleFib b c = Σ[ a ∈ A ] (f a ≡ b) × (g a ≡ c) + + PushoutGenFib = PushoutGen doubleFib + + Pushout→PushoutGen : Pushout f g → PushoutGenFib + Pushout→PushoutGen (inl x) = inl x + Pushout→PushoutGen (inr x) = inr x + Pushout→PushoutGen (push a i) = push (a , refl , refl) i + + PushoutGen→Pushout : PushoutGenFib → Pushout f g + PushoutGen→Pushout (inl x) = inl x + PushoutGen→Pushout (inr x) = inr x + PushoutGen→Pushout (push (x , p , q) i) = + ((λ i → inl (p (~ i))) ∙∙ push x ∙∙ (λ i → inr (q i))) i + + IsoPushoutPushoutGen : Iso (Pushout f g) (PushoutGenFib) + fun IsoPushoutPushoutGen = Pushout→PushoutGen + inv IsoPushoutPushoutGen = PushoutGen→Pushout + rightInv IsoPushoutPushoutGen (inl x) = refl + rightInv IsoPushoutPushoutGen (inr x) = refl + rightInv IsoPushoutPushoutGen (push (x , p , q) i) j = lem x p q j i + where + lem : {b : B} {c : C} (x : A) (p : f x ≡ b) (q : g x ≡ c) + → cong Pushout→PushoutGen (cong PushoutGen→Pushout (push (x , p , q))) + ≡ push (x , p , q) + lem {c = c} x = + J (λ b p → (q : g x ≡ c) + → cong Pushout→PushoutGen + (cong PushoutGen→Pushout (push (x , p , q))) + ≡ push (x , p , q)) + (J (λ c q → cong Pushout→PushoutGen + (cong PushoutGen→Pushout (push (x , refl , q))) + ≡ push (x , refl , q)) + (cong (cong Pushout→PushoutGen) (sym (rUnit (push x))))) + leftInv IsoPushoutPushoutGen (inl x) = refl + leftInv IsoPushoutPushoutGen (inr x) = refl + leftInv IsoPushoutPushoutGen (push a i) j = rUnit (push a) (~ j) i diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index 233f0a7bae..e596246468 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -21,6 +21,7 @@ module Cubical.HITs.Pushout.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.GroupoidLaws @@ -507,3 +508,33 @@ module _ {ℓ : Level} {A₁ B₁ C₁ A₂ B₂ C₂ : Type ℓ} pushoutEquiv : Pushout f₁ g₁ ≃ Pushout f₂ g₂ pushoutEquiv = isoToEquiv pushoutIso + +module PushoutDistr {ℓ ℓ' ℓ'' ℓ''' : Level} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + (f : B → A) (g : C → B) (h : C → D) where + PushoutDistrFun : Pushout {C = Pushout g h} f inl → Pushout (f ∘ g) h + PushoutDistrFun (inl x) = inl x + PushoutDistrFun (inr (inl x)) = inl (f x) + PushoutDistrFun (inr (inr x)) = inr x + PushoutDistrFun (inr (push a i)) = push a i + PushoutDistrFun (push a i) = inl (f a) + + PushoutDistrInv : Pushout (f ∘ g) h → Pushout {C = Pushout g h} f inl + PushoutDistrInv (inl x) = inl x + PushoutDistrInv (inr x) = inr (inr x) + PushoutDistrInv (push c i) = (push (g c) ∙ λ j → inr (push c j)) i + + PushoutDistrIso : Iso (Pushout {C = Pushout g h} f inl) (Pushout (f ∘ g) h) + fun PushoutDistrIso = PushoutDistrFun + inv PushoutDistrIso = PushoutDistrInv + rightInv PushoutDistrIso (inl x) = refl + rightInv PushoutDistrIso (inr x) = refl + rightInv PushoutDistrIso (push a i) j = + (cong-∙ (fun PushoutDistrIso) (push (g a)) (λ j → inr (push a j)) + ∙ sym (lUnit _)) j i + leftInv PushoutDistrIso (inl x) = refl + leftInv PushoutDistrIso (inr (inl x)) = push x + leftInv PushoutDistrIso (inr (inr x)) = refl + leftInv PushoutDistrIso (inr (push a i)) j = + compPath-filler' (push (g a)) (λ j → inr (push a j)) (~ j) i + leftInv PushoutDistrIso (push a i) j = push a (i ∧ j) diff --git a/Cubical/HITs/S2/Base.agda b/Cubical/HITs/S2/Base.agda index 0e4a267aa8..6b156b5542 100644 --- a/Cubical/HITs/S2/Base.agda +++ b/Cubical/HITs/S2/Base.agda @@ -3,15 +3,55 @@ module Cubical.HITs.S2.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Pointed data S² : Type₀ where base : S² surf : PathP (λ i → base ≡ base) refl refl -S²ToSetRec : ∀ {ℓ} {A : S² → Type ℓ} +S²∙ : Pointed ℓ-zero +S²∙ = S² , base + +S²ToSetElim : ∀ {ℓ} {A : S² → Type ℓ} → ((x : S²) → isSet (A x)) → A base → (x : S²) → A x -S²ToSetRec set b base = b -S²ToSetRec set b (surf i j) = +S²ToSetElim set b base = b +S²ToSetElim set b (surf i j) = isOfHLevel→isOfHLevelDep 2 set b b {a0 = refl} {a1 = refl} refl refl surf i j + +-- Wedge connectivity lemmas for S² (binary maps 2-groupoids) +wedgeconFunS² : ∀ {ℓ} {P : S² → S² → Type ℓ} + → ((x y : _) → isOfHLevel 4 (P x y)) + → (l : ((x : S²) → P x base)) + → (r : (x : S²) → P base x) + → l base ≡ r base + → (x y : _) → P x y +wedgeconFunS² {P = P} hlev l r p base y = r y +wedgeconFunS² {P = P} hlev l r p (surf i i₁) y = help y i i₁ + where + help : (y : S²) → SquareP (λ i j → P (surf i j) y) + (λ _ → r y) (λ _ → r y) + (λ _ → r y) λ _ → r y + help = + S²ToSetElim (λ _ → isOfHLevelPathP' 2 (isOfHLevelPathP' 3 (hlev _ _) _ _) _ _) + λ w j → hcomp (λ k → λ { (j = i0) → p k + ; (j = i1) → p k + ; (w = i0) → p k + ; (w = i1) → p k}) + (l (surf w j)) + +wedgeconFunS²Id : ∀ {ℓ} {P : S² → S² → Type ℓ} + → (h : ((x y : _) → isOfHLevel 4 (P x y))) + → (l : ((x : S²) → P x base)) + → (r : (x : S²) → P base x) + → (p : l base ≡ r base) + → (x : S²) → wedgeconFunS² h l r p x base ≡ l x +wedgeconFunS²Id h l r p base = sym p +wedgeconFunS²Id h l r p (surf i j) k = + hcomp (λ w → λ {(i = i0) → p (~ k ∧ w) + ; (i = i1) → p (~ k ∧ w) + ; (j = i0) → p (~ k ∧ w) + ; (j = i1) → p (~ k ∧ w) + ; (k = i1) → l (surf i j)}) + (l (surf i j)) diff --git a/Cubical/HITs/SetQuotients/Properties.agda b/Cubical/HITs/SetQuotients/Properties.agda index 1ee6891688..d1864ad4b9 100644 --- a/Cubical/HITs/SetQuotients/Properties.agda +++ b/Cubical/HITs/SetQuotients/Properties.agda @@ -35,8 +35,8 @@ open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂ ; ∣_∣ private variable ℓ ℓ' ℓ'' : Level - A B C : Type ℓ - R S T : A → A → Type ℓ + A B C Q : Type ℓ + R S T W : A → A → Type ℓ elimProp : {P : A / R → Type ℓ} → (∀ x → isProp (P x)) @@ -67,6 +67,14 @@ elimProp3 prop f = elimProp (λ x → isPropΠ2 (prop x)) λ a → elimProp2 (prop [ a ]) (f a) +elimProp4 : {P : A / R → B / S → C / T → Q / W → Type ℓ} + → (∀ x y z t → isProp (P x y z t)) + → (∀ a b c d → P [ a ] [ b ] [ c ] [ d ]) + → ∀ x y z t → P x y z t +elimProp4 prop f = + elimProp (λ x → isPropΠ3 (prop x)) λ a → + elimProp3 (prop [ a ]) (f a) + -- sometimes more convenient: elimContr : {P : A / R → Type ℓ} → (∀ a → isContr (P [ a ])) diff --git a/Cubical/HITs/SetTruncation/Base.agda b/Cubical/HITs/SetTruncation/Base.agda index 44b0ecdac7..033d865063 100644 --- a/Cubical/HITs/SetTruncation/Base.agda +++ b/Cubical/HITs/SetTruncation/Base.agda @@ -9,9 +9,15 @@ This file contains: module Cubical.HITs.SetTruncation.Base where open import Cubical.Core.Primitives +open import Cubical.Foundations.Pointed -- set truncation as a higher inductive type: data ∥_∥₂ {ℓ} (A : Type ℓ) : Type ℓ where ∣_∣₂ : A → ∥ A ∥₂ squash₂ : ∀ (x y : ∥ A ∥₂) (p q : x ≡ y) → p ≡ q + +-- Pointed version +∥_∥₂∙ : ∀ {ℓ} (A : Pointed ℓ) → Pointed ℓ +fst ∥ A ∥₂∙ = ∥ fst A ∥₂ +snd ∥ A ∥₂∙ = ∣ pt A ∣₂ diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 7066a1aa47..b6ed287133 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -17,6 +17,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.Pointed.Base open import Cubical.Data.Sigma open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim) hiding (elim2 ; elim3 ; rec2 ; map) @@ -203,6 +204,11 @@ module rec→Gpd {A : Type ℓ} {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A → map : (A → B) → ∥ A ∥₂ → ∥ B ∥₂ map f = rec squash₂ λ x → ∣ f x ∣₂ +map∙ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (f : A →∙ B) → ∥ A ∥₂∙ →∙ ∥ B ∥₂∙ +fst (map∙ f) = map (fst f) +snd (map∙ f) = cong ∣_∣₂ (snd f) + setTruncUniversal : isSet B → (∥ A ∥₂ → B) ≃ (A → B) setTruncUniversal {B = B} Bset = isoToEquiv (iso (λ h x → h ∣ x ∣₂) (rec Bset) (λ _ → refl) rinv) diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 23424fe57b..91cc38611e 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -17,7 +17,6 @@ open import Cubical.Data.Sigma open import Cubical.HITs.Sn.Base open import Cubical.HITs.Susp open import Cubical.HITs.Truncation --- open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected open import Cubical.HITs.Join open import Cubical.Data.Bool diff --git a/Cubical/HITs/Truncation/Base.agda b/Cubical/HITs/Truncation/Base.agda index 643af521b9..fdd0be75c0 100644 --- a/Cubical/HITs/Truncation/Base.agda +++ b/Cubical/HITs/Truncation/Base.agda @@ -12,6 +12,7 @@ module Cubical.HITs.Truncation.Base where open import Cubical.Data.NatMinusOne open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Pointed open import Cubical.HITs.Sn.Base open import Cubical.Data.Nat.Base open import Cubical.Data.Unit.Base @@ -34,3 +35,9 @@ hLevelTrunc (suc n) A = HubAndSpoke A n ∣_∣ₕ : ∀ {ℓ} {A : Type ℓ} {n : ℕ} → A → ∥ A ∥ n ∣_∣ₕ {n = zero} a = tt* ∣_∣ₕ {n = suc n} a = ∣ a ∣ + +-- Pointed version +hLevelTrunc∙ : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Pointed ℓ +fst (hLevelTrunc∙ n A) = hLevelTrunc n (typ A) +snd (hLevelTrunc∙ zero A) = tt* +snd (hLevelTrunc∙ (suc n) A) = ∣ pt A ∣ₕ diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda index f2094d066e..887dc002f4 100644 --- a/Cubical/HITs/Wedge/Base.agda +++ b/Cubical/HITs/Wedge/Base.agda @@ -6,6 +6,7 @@ open import Cubical.Foundations.Pointed open import Cubical.HITs.Pushout.Base open import Cubical.Data.Sigma open import Cubical.Data.Unit +open import Cubical.Foundations.GroupoidLaws _⋁_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') _⋁_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ _ → ptA) (λ _ → ptB) @@ -44,3 +45,15 @@ snd isContr-Unit⋁Unit (push tt i) j = push tt (i ∧ j) ⋁↪ {B = B} (inl x) = x , pt B ⋁↪ {A = A} (inr x) = pt A , x ⋁↪ {A = A} {B = B} (push a i) = pt A , pt B + +fold⋁ : ∀ {ℓ} {A : Pointed ℓ} → A ⋁ A → typ A +fold⋁ (inl x) = x +fold⋁ (inr x) = x +fold⋁ {A = A} (push a i) = snd A + +id∨→∙id : ∀ {ℓ} {A : Pointed ℓ} → ∨→∙ (idfun∙ A) (idfun∙ A) ≡ (fold⋁ , refl) +id∨→∙id {A = A} = + ΣPathP ((funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j → rUnit (λ _ → pt A) (~ j) i})) + , refl) diff --git a/Cubical/Homotopy/BlakersMassey.agda b/Cubical/Homotopy/BlakersMassey.agda index 77ddeacd11..83e2a1a4f3 100644 --- a/Cubical/Homotopy/BlakersMassey.agda +++ b/Cubical/Homotopy/BlakersMassey.agda @@ -22,13 +22,19 @@ open import Cubical.Foundations.Transport open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties open import Cubical.Foundations.Isomorphism +open Iso open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Data.Nat hiding (elim) open import Cubical.Data.Sigma +open import Cubical.Data.Unit open import Cubical.HITs.Truncation renaming (hLevelTrunc to Trunc) +open import Cubical.HITs.Pushout hiding (PushoutGenFib) open import Cubical.Homotopy.Connected open import Cubical.Homotopy.WedgeConnectivity @@ -48,18 +54,14 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} rightFiber : Y → Type (ℓ-max ℓ₁ ℓ₃) rightFiber y = Σ[ x ∈ X ] Q x y - {- An alternative formulation of pushout with fewer parameters -} + {- We use the alternative formulation of pushout with fewer parameters -} - data Pushout : Type ℓ - where - inl : X → Pushout - inr : Y → Pushout - push : {x : X}{y : Y} → Q x y → inl x ≡ inr y + PushoutQ = PushoutGen Q {- Some preliminary definitions for convenience -} fiberSquare : - {x₀ x₁ : X}{y₀ : Y}{p : Pushout}(q₀₀ : Q x₀ y₀)(q₁₀ : Q x₁ y₀) + {x₀ x₁ : X}{y₀ : Y}{p : PushoutQ}(q₀₀ : Q x₀ y₀)(q₁₀ : Q x₁ y₀) → inl x₁ ≡ p → inl x₀ ≡ p → Type ℓ fiberSquare q₀₀ q₁₀ r' r = PathP (λ i → push q₀₀ (~ i) ≡ r' i) (sym (push q₁₀)) r @@ -68,7 +70,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} → inl x₀ ≡ inr y₁ → Type ℓ fiberSquarePush q₀₀ q₁₀ q₁₁ = fiberSquare q₀₀ q₁₀ (push q₁₁) - fiber' : {x₀ : X}{y₀ : Y}(q₀₀ : Q x₀ y₀){x₁ : X}{p : Pushout} → inl x₁ ≡ p → inl x₀ ≡ p → Type ℓ + fiber' : {x₀ : X}{y₀ : Y}(q₀₀ : Q x₀ y₀){x₁ : X}{p : PushoutQ} → inl x₁ ≡ p → inl x₀ ≡ p → Type ℓ fiber' {y₀ = y₀} q₀₀ {x₁ = x₁} r' r = Σ[ q₁₀ ∈ Q x₁ y₀ ] fiberSquare q₀₀ q₁₀ r' r fiber'Push : {x₀ x₁ : X}{y₀ y₁ : Y}(q₀₀ : Q x₀ y₀)(q₁₁ : Q x₁ y₁) → inl x₀ ≡ inr y₁ → Type ℓ @@ -76,10 +78,10 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} leftCodeExtended : {x₀ : X}{y₀ : Y}(q₀₀ : Q x₀ y₀) - → (x₁ : X){p : Pushout} → inl x₁ ≡ p → inl x₀ ≡ p → Type ℓ + → (x₁ : X){p : PushoutQ} → inl x₁ ≡ p → inl x₀ ≡ p → Type ℓ leftCodeExtended {y₀ = y₀} q₀₀ x₁ r' r = Trunc (m + n) (fiber' q₀₀ r' r) - rightCode : {x₀ : X}(y : Y) → inl x₀ ≡ inr y → Type ℓ + rightCode : {x₀ : X}(y : Y) → Path PushoutQ (inl x₀) (inr y) → Type ℓ rightCode y r = Trunc (m + n) (fiber push r) {- Bunch of coherence data that will be used to construct Code -} @@ -95,7 +97,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} (r : inl x₁ ≡ inr y₁) (p : fiberSquarePush q₁₀ q₁₀ q₁₁ r) where - fiber→[q₀₀=q₁₀]-filler : (i j k : I) → Pushout + fiber→[q₀₀=q₁₀]-filler : (i j k : I) → PushoutQ fiber→[q₀₀=q₁₀]-filler i j k' = hfill (λ k → λ { (i = i0) → push q₁₁ (j ∧ k) ; (i = i1) → p k j @@ -116,7 +118,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} (r : inl x₀ ≡ inr y₀) (p : fiberSquarePush q₀₀ q₁₀ q₁₀ r) where - fiber→[q₁₁=q₁₀]-filler : (i j k : I) → Pushout + fiber→[q₁₁=q₁₀]-filler : (i j k : I) → PushoutQ fiber→[q₁₁=q₁₀]-filler i j k' = hfill (λ k → λ { (i = i0) → push q₀₀ (j ∨ ~ k) ; (i = i1) → p k j @@ -135,7 +137,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} fiber→[q₀₀=q₁₀=q₁₁]-filler : (r : inl x₁ ≡ inr y₀) → (p : fiberSquarePush q₁₀ q₁₀ q₁₀ r) - → (i j k l : I) → Pushout + → (i j k l : I) → PushoutQ fiber→[q₀₀=q₁₀=q₁₁]-filler r p i j k l' = hfill (λ l → λ { (i = i0) → fiber→[q₀₀=q₁₀]-filler q₁₀ r p j k l ; (i = i1) → fiber→[q₁₁=q₁₀]-filler q₁₀ r p j k l @@ -163,7 +165,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} (r : inl x₀ ≡ inr y₁) (p : push q₀₁ ≡ r) where - fiber←[q₁₁=q₀₁]-filler : (i j k : I) → Pushout + fiber←[q₁₁=q₀₁]-filler : (i j k : I) → PushoutQ fiber←[q₁₁=q₀₁]-filler i j k' = hfill (λ k → λ { (i = i0) → push q₀₀ (~ j ∧ k) ; (i = i1) → p k j @@ -184,7 +186,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} (r : inl x₀ ≡ inr y₁) (p : push q₀₁ ≡ r) where - fiber←[q₀₀=q₀₁]-filler : (i j k : I) → Pushout + fiber←[q₀₀=q₀₁]-filler : (i j k : I) → PushoutQ fiber←[q₀₀=q₀₁]-filler i j k' = hfill (λ k → λ { (i = i0) → push q₁₁ (~ j ∨ ~ k) ; (i = i1) → p k j @@ -203,7 +205,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} fiber←[q₀₀=q₀₁=q₁₁]-filler : (r : inl x₀ ≡ inr y₁) → (p : push q₀₁ ≡ r) - → (i j k l : I) → Pushout + → (i j k l : I) → PushoutQ fiber←[q₀₀=q₀₁=q₁₁]-filler r p i j k l' = hfill (λ l → λ { (i = i0) → fiber←[q₁₁=q₀₁]-filler q₀₁ r p j k l ; (i = i1) → fiber←[q₀₀=q₀₁]-filler q₀₁ r p j k l @@ -231,7 +233,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} (r : inl x₁ ≡ inr y₁) (p : fiberSquarePush q₁₀ q₁₀ q₁₁ r) where - fiber→←[q₀₀=q₁₀]-filler : (i j k l : I) → Pushout + fiber→←[q₀₀=q₁₀]-filler : (i j k l : I) → PushoutQ fiber→←[q₀₀=q₁₀]-filler i j k l' = let p' = fiber→[q₀₀=q₁₀] q₁₀ q₁₁ r p .snd in hfill (λ l → λ { (i = i0) → fiber←[q₁₁=q₀₁]-filler q₁₁ q₁₀ r p' j k l @@ -251,7 +253,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} (r : inl x₀ ≡ inr y₀) (p : fiberSquarePush q₀₀ q₁₀ q₁₀ r) where - fiber→←[q₁₁=q₁₀]-filler : (i j k l : I) → Pushout + fiber→←[q₁₁=q₁₀]-filler : (i j k l : I) → PushoutQ fiber→←[q₁₁=q₁₀]-filler i j k l' = let p' = fiber→[q₁₁=q₁₀] q₁₀ q₀₀ r p .snd in hfill (λ l → λ { (i = i0) → fiber←[q₀₀=q₀₁]-filler q₀₀ q₁₀ r p' j k l @@ -293,7 +295,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} (r : inl x₀ ≡ inr y₁) (p : push q₀₁ ≡ r) where - fiber←→[q₁₁=q₀₁]-filler : (i j k l : I) → Pushout + fiber←→[q₁₁=q₀₁]-filler : (i j k l : I) → PushoutQ fiber←→[q₁₁=q₀₁]-filler i j k l' = let p' = fiber←[q₁₁=q₀₁] q₀₁ q₀₀ r p .snd in hfill (λ l → λ { (i = i0) → fiber→[q₀₀=q₁₀]-filler q₀₀ q₀₁ r p' j k l @@ -313,7 +315,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} (r : inl x₀ ≡ inr y₁) (p : push q₀₁ ≡ r) where - fiber←→[q₀₀=q₀₁]-filler : (i j k l : I) → Pushout + fiber←→[q₀₀=q₀₁]-filler : (i j k l : I) → PushoutQ fiber←→[q₀₀=q₀₁]-filler i j k l' = let p' = fiber←[q₀₀=q₀₁] q₀₁ q₁₁ r p .snd in hfill (λ l → λ { (i = i0) → fiber→[q₁₁=q₁₀]-filler q₁₁ q₀₁ r p' j k l @@ -551,7 +553,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} module _ (x₀ : X)(y₀ : Y)(q₀₀ : Q x₀ y₀) where - leftCode' : (x : X){p : Pushout} → inl x ≡ p → inl x₀ ≡ p → Type ℓ + leftCode' : (x : X){p : PushoutQ} → inl x ≡ p → inl x₀ ≡ p → Type ℓ leftCode' x r' = leftCodeExtended q₀₀ x r' leftCode : (x : X) → inl x₀ ≡ inl x → Type ℓ @@ -568,7 +570,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} ; (i = i1) → fiberPath q j }) (leftCode' _ (λ j → push q (i ∧ j))) - Code : (p : Pushout) → inl x₀ ≡ p → Type ℓ + Code : (p : PushoutQ) → inl x₀ ≡ p → Type ℓ Code (inl x) = leftCode x Code (inr y) = rightCode y Code (push q i) = pushCode q i @@ -585,7 +587,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} transpPushCodeβ' y q q' i = transportRefl (left→rightCodeExtended _ _ _ (transpLeftCode y q (transportRefl q' i))) i module _ - {p : Pushout}(r : inl x₀ ≡ p) where + {p : PushoutQ}(r : inl x₀ ≡ p) where fiber-filler : I → Type ℓ fiber-filler i = fiber' q₀₀ (λ j → r (i ∧ j)) (λ j → r (i ∧ j)) @@ -593,7 +595,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} module _ (q : fiberSquare q₀₀ q₀₀ refl refl) where - transpLeftCode-filler : (i j k : I) → Pushout + transpLeftCode-filler : (i j k : I) → PushoutQ transpLeftCode-filler i j k' = hfill (λ k → λ { (i = i0) → push q₀₀ (~ j) ; (i = i1) → r (j ∧ k) @@ -602,7 +604,7 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} (inS (q i j)) k' transpLeftCodeβ' : - {p : Pushout} → (r : inl x₀ ≡ p) → (q : fiberSquare q₀₀ q₀₀ refl refl) + {p : PushoutQ} → (r : inl x₀ ≡ p) → (q : fiberSquare q₀₀ q₀₀ refl refl) → transport (λ i → fiber-filler r i) (q₀₀ , q) ≡ (q₀₀ , λ i j → transpLeftCode-filler r q i j i1) transpLeftCodeβ' r q = J (λ p r → transport (λ i → fiber-filler r i) (q₀₀ , q) ≡ (q₀₀ , λ i j → transpLeftCode-filler r q i j i1)) @@ -625,14 +627,14 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} {- The contractibility of Code -} - centerCode : {p : Pushout} → (r : inl x₀ ≡ p) → Code p r + centerCode : {p : PushoutQ} → (r : inl x₀ ≡ p) → Code p r centerCode r = transport (λ i → Code _ (λ j → r (i ∧ j))) ∣ q₀₀ , (λ i j → push q₀₀ (~ i ∧ ~ j)) ∣ₕ module _ (y : Y)(q : Q x₀ y) where - transp-filler : (i j k : I) → Pushout + transp-filler : (i j k : I) → PushoutQ transp-filler = transpLeftCode-filler (push q) (λ i' j' → push q₀₀ (~ i' ∧ ~ j')) transp-square : fiberSquare q₀₀ q₀₀ (push q) (push q) @@ -672,3 +674,136 @@ module BlakersMassey {ℓ₁ ℓ₂ ℓ₃ : Level} {- The Main Result : Blakers-Massey Homotopy Excision Theorem -} Excision : (x : X)(y : Y) → isConnectedFun (m + n) (push {x = x} {y = y}) Excision x y = excision-helper x (leftConn x .fst) y + + +{- +We also give the following version of the theorem: Given a square + + g + A --------------> C + |\ ↗ | + | \ ↗ | + | \ ↗ | +f | X | inr + | / | + | / | + | / | + v / v + B -----------> Pushout f g + inl + +where X is the pullback of inl and inr + (X := Σ[ (b , c) ∈ B × C ] (inl b ≡ inr c)). + +If f in n-connected and g in m-connected, then the diagonal map +A → X is (n+m)-connected +-} + + +private + shuffleFibIso₁ : + {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : A → B) (g : A → C) (b : B) + → Iso (Σ[ c ∈ C ] Σ[ a ∈ A ] (f a ≡ b) × (g a ≡ c)) + (Σ[ a ∈ A ] ((Σ[ c ∈ C ] (g a ≡ c)) × (f a ≡ b))) + shuffleFibIso₁ f g b = + compIso (invIso Σ-assoc-Iso) + (compIso (Σ-cong-iso-fst Σ-swap-Iso) + (compIso + (Σ-cong-iso-snd (λ y → Σ-swap-Iso)) + (compIso Σ-assoc-Iso + (Σ-cong-iso-snd λ a → invIso Σ-assoc-Iso)))) + + shuffleFibIso₂ : {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : A → B) (g : A → C) (x : _) + → Iso (Σ[ a ∈ A ] ((Σ[ c ∈ C ] (g a ≡ c)) × (f a ≡ x))) + (fiber f x) + shuffleFibIso₂ f g x = Σ-cong-iso-snd + λ a → compIso (Σ-cong-iso-fst + (isContr→Iso (isContrSingl (g a)) + isContrUnit)) + lUnit×Iso + +module BlakersMassey□ {ℓ ℓ' ℓ'' : Level} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : A → B) (g : A → C) (n m : ℕ) + (con-f : isConnectedFun (suc n) f) + (con-g : isConnectedFun (suc m) g) where + + {- Some abbreviations and connectivity -} + private + fib = doubleFib f g + + B-con : (x : B) → isConnected (suc n) (Σ[ c ∈ C ] (fib x c)) + B-con x = + isConnectedRetractFromIso (suc n) + (compIso + (shuffleFibIso₁ f g x) + (shuffleFibIso₂ f g x)) + (con-f x) + + C-con : (c : C) → isConnected (suc m) (Σ[ b ∈ B ] (fib b c)) + C-con c = + isConnectedRetractFromIso (suc m) + (compIso + (compIso (Σ-cong-iso-snd + (λ _ → Σ-cong-iso-snd λ _ → Σ-swap-Iso)) + (shuffleFibIso₁ g f c)) + (shuffleFibIso₂ g f c)) + (con-g c) + + open module BM-f-g = BlakersMassey B C fib {m = n} B-con {n = m} C-con + + fib× : (B × C) → Type _ + fib× (b , c) = fib b c + + PushoutGenPath× : B × C → Type _ + PushoutGenPath× (b , c) = Path (PushoutGen fib) (inl b) (inr c) + + PushoutPath× : B × C → Type _ + PushoutPath× (b , c) = Path (Pushout f g) (inl b) (inr c) + + {- The function in question -} + toPullback : A → Σ (B × C) PushoutPath× + toPullback a = (f a , g a) , push a + + {- We redescribe toPullback as a composition of three maps, + two of which are equivs and one of which is (n+m)-connected -} + Totalfib×→Total : Σ (B × C) fib× → Σ (B × C) PushoutGenPath× + Totalfib×→Total = + TotalFun {A = B × C} {B = fib×} {C = PushoutGenPath×} (λ a → push) + + isConnectedTotalFun : isConnectedFun (n + m) Totalfib×→Total + isConnectedTotalFun = + FunConnected→TotalFunConnected (λ _ → push) (n + m) (uncurry BM-f-g.Excision) + + TotalPathGen×Iso : Iso (Σ (B × C) PushoutGenPath×) (Σ (B × C) PushoutPath×) + TotalPathGen×Iso = + Σ-cong-iso-snd λ x + → congIso (invIso (IsoPushoutPushoutGen f g)) + + Totalfib×Iso : Iso (Σ (B × C) fib×) A + fun Totalfib×Iso ((b , c) , a , p) = a + inv Totalfib×Iso a = (f a , g a) , a , refl , refl + rightInv Totalfib×Iso _ = refl + leftInv Totalfib×Iso ((b , c) , a , (p , q)) i = + ((p i) , (q i)) , (a , ((λ j → p (i ∧ j)) , (λ j → q (i ∧ j)))) + + toPullback' : A → Σ (B × C) PushoutPath× + toPullback' = + (fun TotalPathGen×Iso ∘ Totalfib×→Total) ∘ inv Totalfib×Iso + + toPullback'≡toPullback : toPullback' ≡ toPullback + toPullback'≡toPullback = + funExt λ x → ΣPathP (refl , (sym (rUnit (push x)))) + + isConnected-toPullback : isConnectedFun (n + m) toPullback + isConnected-toPullback = + subst (isConnectedFun (n + m)) toPullback'≡toPullback + (isConnectedComp + (fun TotalPathGen×Iso ∘ Totalfib×→Total) + (inv Totalfib×Iso) (n + m) + (isConnectedComp (fun TotalPathGen×Iso) Totalfib×→Total (n + m) + (isEquiv→isConnected _ (isoToIsEquiv TotalPathGen×Iso) (n + m)) + isConnectedTotalFun) + (isEquiv→isConnected _ (isoToIsEquiv (invIso Totalfib×Iso)) (n + m))) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 3e4b6dbdb5..b6307b8f09 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -44,7 +44,6 @@ isConnectedSubtr {A = A} n m iscon = helper zero iscon = isContrUnit* helper (suc n) iscon = ∣ iscon .fst ∣ , (Trunc.elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ a → cong ∣_∣ (iscon .snd a)) - isConnectedFunSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : HLevel) (f : A → B) → isConnectedFun (m + n) f → isConnectedFun n f @@ -60,7 +59,12 @@ private typeToFiber : ∀ {ℓ} (A : Type ℓ) → A ≡ fiber (λ (x : A) → tt) tt typeToFiber A = isoToPath (typeToFiberIso A) - +isOfHLevelIsConnectedStable : ∀ {ℓ} {A : Type ℓ} (n : ℕ) + → isOfHLevel n (isConnected n A) +isOfHLevelIsConnectedStable {A = A} zero = + (tt* , (λ _ → refl)) , λ _ → refl +isOfHLevelIsConnectedStable {A = A} (suc n) = + isProp→isOfHLevelSuc n isPropIsContr module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) where private @@ -164,6 +168,27 @@ indMapEquiv→conType {A = A} (suc n) BEq = , λ a → equiv-proof (BEq (P tt)) a .fst .snd) tt) +isConnectedComp : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : B → C) (g : A → B) (n : ℕ) + → isConnectedFun n f + → isConnectedFun n g + → isConnectedFun n (f ∘ g) +isConnectedComp {C = C} f g n con-f con-g = + elim.isConnectedPrecompose (f ∘ g) n + λ P → + isEquiv→hasSection + (compEquiv + (_ , elim.isEquivPrecompose f n P con-f) + (_ , elim.isEquivPrecompose g n (λ b → P (f b)) con-g) .snd) + +isEquiv→isConnected : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + (f : A → B) + → isEquiv f + → (n : ℕ) → isConnectedFun n f +isEquiv→isConnected f iseq n b = + isContr→isContr∥ n (equiv-proof iseq b) + + isConnectedPath : ∀ {ℓ} (n : HLevel) {A : Type ℓ} → isConnected (suc n) A → (a₀ a₁ : A) → isConnected n (a₀ ≡ a₁) @@ -198,6 +223,51 @@ isConnectedCong' {x = x} n f conf p = → doubleCompPath-filler (sym p) (cong f q) p i)) (isConnectedCong n f conf) +isConnectedΩ→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (f : A →∙ B) + → isConnectedFun (suc n) (fst f) + → isConnectedFun n (fst (Ω→ f)) +isConnectedΩ→ n f g = + transport (λ i → isConnectedFun n λ b + → doubleCompPath-filler (sym (snd f)) (cong (fst f) b) (snd f) i) + (isConnectedCong n _ g) + +isConnectedΩ^→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n m : ℕ) + → (f : A →∙ B) + → isConnectedFun (suc n) (fst f) + → isConnectedFun ((suc n) ∸ m) (fst (Ω^→ m f)) +isConnectedΩ^→ n zero f conf = conf +isConnectedΩ^→ n (suc zero) f conf = isConnectedΩ→ n f conf +isConnectedΩ^→ {A = A} {B = B} n (suc (suc m)) f conf = + transport (λ i → isConnectedFun (suc n ∸ suc (suc m)) + λ q → doubleCompPath-filler + (sym (snd (Ω^→ (suc m) f))) + (cong (fst (Ω^→ (suc m) f)) q) + (snd (Ω^→ (suc m) f)) i) + (main n m (isConnectedΩ^→ n (suc m) f conf)) + where + open import Cubical.Data.Sum + lem : (n m : ℕ) → ((suc n ∸ m ≡ suc (n ∸ m))) ⊎ (suc n ∸ suc m ≡ 0) + lem n zero = inl refl + lem zero (suc m) = inr refl + lem (suc n) (suc m) = lem n m + + main : (n m : ℕ) + → isConnectedFun (suc n ∸ suc m) (fst (Ω^→ (suc m) f)) + → isConnectedFun (suc n ∸ suc (suc m)) + {A = Path ((Ω^ (suc m)) (_ , pt A) .fst) + refl refl} + (cong (fst (Ω^→ (suc m) f))) + main n m conf with (lem n (suc m)) + ... | (inl x) = + isConnectedCong (n ∸ suc m) (fst (Ω^→ (suc m) f)) + (subst (λ x → isConnectedFun x (fst (Ω^→ (suc m) f))) x conf) + ... | (inr x) = + subst (λ x → isConnectedFun x (cong {x = refl} {y = refl} + (fst (Ω^→ (suc m) f)))) + (sym x) + λ b → tt* , λ _ → refl + isConnectedRetract : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) (g : B → A) @@ -324,3 +394,57 @@ inrConnected {A = A} {B = B} {C = C} n f g iscon = (~ i) (equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .snd i a)) + +{- Given two fibration B , C : A → Type and a family of maps on fibres + f : (a : A) → B a → C a, we have that f a is n-connected for all (a : A) + iff the induced map on totalspaces Σ A B → Σ A C is n-connected -} +module _ {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} + (f : ((a : A) → B a → C a)) + where + open Iso + + TotalFun : Σ A B → Σ A C + TotalFun (a , b) = a , f a b + + fibTotalFun→fibFun : (x : Σ A C) + → Σ[ y ∈ Σ A B ] TotalFun y ≡ x + → Σ[ y ∈ B (fst x) ] f (fst x) y ≡ snd x + fibTotalFun→fibFun x = + uncurry + λ r → J (λ x _ → Σ[ y ∈ B (fst x) ] f (fst x) y ≡ snd x) + ((snd r) , refl) + + fibFun→fibTotalFun : (x : Σ A C) + → Σ[ y ∈ B (fst x) ] f (fst x) y ≡ snd x + → Σ[ y ∈ Σ A B ] TotalFun y ≡ x + fibFun→fibTotalFun x (b , p) = (_ , b) , ΣPathP (refl , p) + + Iso-fibTotalFun-fibFun : (x : Σ A C) + → Iso (Σ[ y ∈ Σ A B ] TotalFun y ≡ x) + (Σ[ y ∈ B (fst x) ] f (fst x) y ≡ snd x) + fun (Iso-fibTotalFun-fibFun x) = fibTotalFun→fibFun x + inv (Iso-fibTotalFun-fibFun x) = fibFun→fibTotalFun x + rightInv (Iso-fibTotalFun-fibFun x) (r , y) j = + transp (λ i → Σ[ b ∈ B (fst x) ] (f (fst x) b ≡ y (i ∨ j))) j + (r , λ i → y (i ∧ j)) + leftInv (Iso-fibTotalFun-fibFun x) = + uncurry λ r + → J (λ x y → inv (Iso-fibTotalFun-fibFun x) + (fun (Iso-fibTotalFun-fibFun x) (r , y)) + ≡ (r , y)) + (cong (fibFun→fibTotalFun (TotalFun r)) + (transportRefl (snd r , refl))) + + TotalFunConnected→FunConnected : (n : ℕ) + → isConnectedFun n TotalFun + → ((a : A) → isConnectedFun n (f a)) + TotalFunConnected→FunConnected n con a r = + isConnectedRetractFromIso n + (invIso (Iso-fibTotalFun-fibFun (a , r))) (con (a , r)) + + FunConnected→TotalFunConnected : (n : ℕ) + → ((a : A) → isConnectedFun n (f a)) + → isConnectedFun n TotalFun + FunConnected→TotalFunConnected n con r = + isConnectedRetractFromIso n + (Iso-fibTotalFun-fibFun r) (con (fst r) (snd r)) diff --git a/Cubical/Homotopy/Group/Base.agda b/Cubical/Homotopy/Group/Base.agda index 3d3c79a681..5862f116df 100644 --- a/Cubical/Homotopy/Group/Base.agda +++ b/Cubical/Homotopy/Group/Base.agda @@ -12,12 +12,17 @@ open import Cubical.Foundations.Path open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function +open import Cubical.Foundations.Transport + open import Cubical.Functions.Morphism open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) open import Cubical.HITs.Sn open import Cubical.HITs.Susp renaming (toSusp to σ) open import Cubical.HITs.S1 @@ -25,6 +30,7 @@ open import Cubical.HITs.S1 open import Cubical.Data.Sigma open import Cubical.Data.Nat open import Cubical.Data.Bool +open import Cubical.Data.Unit open import Cubical.Algebra.Group open import Cubical.Algebra.Semigroup @@ -250,7 +256,7 @@ leftInv (SphereMapΩIso (suc n)) = rightInv IsoΩFunSuspFun In order to show that Ω→SphereMap is an equivalence, we show that it factors Ω→SphereMapSplit₁ ΩSphereMap -Ωⁿ⁺¹(Sⁿ →∙ A) ----------------> Ω (Sⁿ →∙ A) -----------> (Sⁿ⁺¹ →∙ A) +Ωⁿ⁺¹A ----------------> Ω (Sⁿ →∙ A) -----------> (Sⁿ⁺¹ →∙ A) -} Ω→SphereMap-split : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : typ ((Ω^ (suc n)) A)) @@ -633,3 +639,419 @@ fst (π'Gr≅πGr n A) = setTruncIso (IsoSphereMapΩ (suc n)) snd (π'Gr≅πGr n A) = makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) + +{- Proof of πₙ(ΩA) = πₙ₊₁(A) -} +Iso-πΩ-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (π n (Ω A)) (π (suc n) A) +Iso-πΩ-π {A = A} n = setTruncIso (invIso (flipΩIso n)) + +GrIso-πΩ-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupIso (πGr n (Ω A)) (πGr (suc n) A) +fst (GrIso-πΩ-π n) = Iso-πΩ-π _ +snd (GrIso-πΩ-π n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (flipΩIso⁻pres· n p q)) + + +{- Proof that πₙ(A) ≅ πₙ(∥ A ∥ₙ) -} +isContrΩTrunc : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → isContr (typ ((Ω^ n) (hLevelTrunc∙ n A))) +isContrΩTrunc {A = A} zero = isContrUnit* +isContrΩTrunc {A = A} (suc n) = + subst isContr main (isContrΩTrunc {A = Ω A} n) + where + lem₁ : (n : ℕ) → fun (PathIdTruncIso n) (λ _ → ∣ pt A ∣) + ≡ snd (hLevelTrunc∙ n (Ω A)) + lem₁ zero = refl + lem₁ (suc n) = transportRefl ∣ refl ∣ + + lem₂ : hLevelTrunc∙ n (Ω A) ≡ (Ω (hLevelTrunc∙ (suc n) A)) + lem₂ = sym (ua∙ (isoToEquiv (PathIdTruncIso n)) + (lem₁ n)) + + main : (typ ((Ω^ n) (hLevelTrunc∙ n (Ω A)))) + ≡ (typ ((Ω^ suc n) (hLevelTrunc∙ (suc n) A))) + main = (λ i → typ ((Ω^ n) (lem₂ i))) + ∙ sym (isoToPath (flipΩIso n)) + + +mutual + ΩTruncSwitchFun : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) → + (hLevelTrunc∙ (suc (suc m)) ((Ω^ n) A)) + →∙ ((Ω^ n) (hLevelTrunc∙ (suc n + suc m) A)) + ΩTruncSwitchFun {A = A} n m = + ((λ x → transport + (λ i → fst ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) i) A))) + (Iso.fun (ΩTruncSwitch {A = A} n (suc (suc m))) x)) + , cong (transport + (λ i → fst ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) i) A)))) + (ΩTruncSwitch∙ n (suc (suc m))) + ∙ λ j → transp + (λ i → fst ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) (i ∨ j)) A))) + j (snd ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) j) A)))) + + ΩTruncSwitchLem : + ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → Iso + (typ (Ω (hLevelTrunc∙ (suc (suc m)) ((Ω^ n) A)))) + (typ ((Ω^ suc n) (hLevelTrunc∙ (suc n + suc m) A))) + ΩTruncSwitchLem {A = A} n m = + (equivToIso + (Ω→ (ΩTruncSwitchFun n m) .fst + , isEquivΩ→ _ (compEquiv (isoToEquiv (ΩTruncSwitch {A = A} n (suc (suc m)))) + (transportEquiv + (λ i → typ ((Ω^ n) (hLevelTrunc∙ (+-suc n (suc m) i) A)))) .snd))) + + ΩTruncSwitch : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → Iso (hLevelTrunc m (fst ((Ω^ n) A))) + (typ ((Ω^ n) (hLevelTrunc∙ (n + m) A))) + ΩTruncSwitch {A = A} n zero = + equivToIso + (invEquiv + (isContr→≃Unit* + (subst isContr + (λ i → (typ ((Ω^ n) (hLevelTrunc∙ (+-comm zero n i) A)))) + (isContrΩTrunc n)))) + ΩTruncSwitch {A = A} zero (suc m) = idIso + ΩTruncSwitch {A = A} (suc n) (suc m) = + compIso (invIso (PathIdTruncIso _)) + (ΩTruncSwitchLem n m) + + ΩTruncSwitch∙ : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → Iso.fun (ΩTruncSwitch {A = A} n m) (snd (hLevelTrunc∙ m ((Ω^ n) A))) + ≡ pt ((Ω^ n) (hLevelTrunc∙ (n + m) A)) + ΩTruncSwitch∙ {A = A} n zero = + isContr→isProp + ((subst isContr + (λ i → (typ ((Ω^ n) (hLevelTrunc∙ (+-comm zero n i) A)))) + (isContrΩTrunc n))) _ _ + ΩTruncSwitch∙ {A = A} zero (suc m) = refl + ΩTruncSwitch∙ {A = A} (suc n) (suc m) = ∙∙lCancel _ + + +ΩTruncSwitch-hom : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) (p q : _) + → Iso.fun (ΩTruncSwitch {A = A} (suc n) (suc m)) ∣ p ∙ q ∣ + ≡ Iso.fun (ΩTruncSwitch {A = A} (suc n) (suc m)) ∣ p ∣ + ∙ Iso.fun (ΩTruncSwitch {A = A} (suc n) (suc m)) ∣ q ∣ +ΩTruncSwitch-hom {A = A} n m p q = + cong (Iso.fun (ΩTruncSwitchLem {A = A} n m)) + (cong-∙ ∣_∣ₕ p q) + ∙ Ω→pres∙ (ΩTruncSwitchFun n m) (cong ∣_∣ₕ p) (cong ∣_∣ₕ q) + +2TruncΩIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (hLevelTrunc 2 (fst ((Ω^ n) A))) + (typ ((Ω^ n) (hLevelTrunc∙ (2 + n) A))) +2TruncΩIso zero = idIso +2TruncΩIso {A = A} (suc n) = + compIso + (ΩTruncSwitch (suc n) 2) + (pathToIso λ i → typ ((Ω^ suc n) (hLevelTrunc∙ (+-comm (suc n) 2 i) A))) + +hLevΩ+ : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) + → isOfHLevel (m + n) (typ A) → isOfHLevel n (typ ((Ω^ m) A)) +hLevΩ+ n zero p = p +hLevΩ+ {A = A} zero (suc zero) p = refl , λ _ → isProp→isSet p _ _ _ _ +hLevΩ+ {A = A} zero (suc (suc zero)) p = + refl , λ y → isOfHLevelSuc 2 p _ _ _ _ refl y +hLevΩ+ {A = A} zero (suc (suc (suc m))) p = + transport + (λ i → isContr (typ (Ω (ua∙ + (isoToEquiv (flipΩIso {A = A} (suc m))) (flipΩrefl m) (~ i))))) + (hLevΩ+ {A = Ω A} zero (suc (suc m)) + (subst (λ x → isOfHLevel x (typ (Ω A))) + (+-comm zero (suc (suc m))) + (lem (pt A) (pt A)))) + where + lem : isOfHLevel (3 + m) (typ A) + lem = subst (λ x → isOfHLevel x (typ A)) + (λ i → suc (+-comm (2 + m) zero i)) p +hLevΩ+ {A = A} (suc n) (suc m) p = + subst (isOfHLevel (suc n)) + (sym (ua (isoToEquiv (flipΩIso {A = A} m)))) + (hLevΩ+ {A = Ω A} (suc n) m + (isOfHLevelPath' (m + suc n) p _ _)) + +isSetΩTrunc : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (isSet (typ (Ω ((Ω^ n) (hLevelTrunc∙ (suc (suc (suc n))) A))))) +isSetΩTrunc {A = A} zero = isOfHLevelTrunc 3 _ _ +isSetΩTrunc {A = A} (suc n) = + hLevΩ+ 2 (suc (suc n)) + (transport + (λ i → isOfHLevel (+-comm 2 (2 + n) i) (hLevelTrunc (4 + n) (typ A))) + (isOfHLevelTrunc (suc (suc (suc (suc n)))))) + +πTruncIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (π n A) (π n (hLevelTrunc∙ (2 + n) A)) +πTruncIso {A = A} zero = + compIso (invIso (setTruncIdempotentIso squash₂)) + (setTruncIso setTruncTrunc2Iso) +πTruncIso {A = A} (suc n) = + compIso setTruncTrunc2Iso + (compIso + (2TruncΩIso (suc n)) + (invIso (setTruncIdempotentIso (isSetΩTrunc n)))) + +πTruncGroupIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupIso (πGr n A) (πGr n (hLevelTrunc∙ (3 + n) A)) +fst (πTruncGroupIso n) = πTruncIso (suc n) +snd (πTruncGroupIso {A = A} n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ a b + → cong (inv (setTruncIdempotentIso (isSetΩTrunc n))) + (cong + (transport + (λ i → typ ((Ω^ suc n) (hLevelTrunc∙ (+-comm (suc n) 2 i) A)))) + (ΩTruncSwitch-hom n 1 a b) + ∙ transpΩTruncSwitch + (λ w → ((Ω^ n) (hLevelTrunc∙ w A))) (+-comm (suc n) 2) _ _)) + where + transpΩTruncSwitch : ∀ {ℓ} (A : ℕ → Pointed ℓ) {n m : ℕ} + (r : n ≡ m) (p q : typ (Ω (A n))) + → subst (λ n → typ (Ω (A n))) r (p ∙ q) + ≡ subst (λ n → typ (Ω (A n))) r p + ∙ subst (λ n → typ (Ω (A n))) r q + transpΩTruncSwitch A {n = n} = + J (λ m r → (p q : typ (Ω (A n))) + → subst (λ n → typ (Ω (A n))) r (p ∙ q) + ≡ subst (λ n → typ (Ω (A n))) r p + ∙ subst (λ n → typ (Ω (A n))) r q) + λ p q → transportRefl _ ∙ cong₂ _∙_ + (sym (transportRefl p)) (sym (transportRefl q)) + +-- Often, we prefer thinking of Ωⁿ A as (Sⁿ →∙ A). +-- The goal of the following lemmas is to show that the maps +-- Ωⁿ A → Ωⁿ B and Ωⁿ (fib f) →∙ Ωⁿ A get sent to post composition +-- under the equivalence Ωⁿ A as (Sⁿ →∙ A). This also gives a proof +-- that post composition induces a homomorphism of homotopy groups. + +-- The following lemmas is not pretty but very helpful +private + bigLemma : ∀ {ℓ ℓ'} {A₁ B₁ C₁ : Type ℓ} {A₂ B₂ C₂ : Type ℓ'} + (A₁→B₁ : A₁ ≃ B₁) (B₁→C₁ : B₁ ≃ C₁) + (A₂→B₂ : A₂ ≃ B₂) (B₂→C₂ : B₂ ≃ C₂) + (A₁→A₂ : A₁ → A₂) + (B₁→B₂ : B₁ → B₂) + (C₁→C₂ : C₁ → C₂) + → (B₁→B₂ ∘ (fst A₁→B₁)) ≡ (fst A₂→B₂ ∘ A₁→A₂) + → C₁→C₂ ∘ fst B₁→C₁ ≡ fst B₂→C₂ ∘ B₁→B₂ + → C₁→C₂ ∘ fst B₁→C₁ ∘ fst A₁→B₁ + ≡ fst B₂→C₂ ∘ fst A₂→B₂ ∘ A₁→A₂ + bigLemma {B₁ = B₁} {C₁ = C₁} {A₂ = A₂} {B₂ = B₂} {C₂ = C₂} = + EquivJ + (λ A₁ A₁→B₁ → (B₁→C₁ : B₁ ≃ C₁) (A₂→B₂ : A₂ ≃ B₂) + (B₂→C₂ : B₂ ≃ C₂) (A₁→A₂ : A₁ → A₂) (B₁→B₂ : B₁ → B₂) + (C₁→C₂ : C₁ → C₂) → + B₁→B₂ ∘ fst A₁→B₁ ≡ fst A₂→B₂ ∘ A₁→A₂ → + C₁→C₂ ∘ fst B₁→C₁ ≡ fst B₂→C₂ ∘ B₁→B₂ → + C₁→C₂ ∘ fst B₁→C₁ ∘ fst A₁→B₁ ≡ fst B₂→C₂ ∘ fst A₂→B₂ ∘ A₁→A₂) + (EquivJ (λ B₁ B₁→C₁ → (A₂→B₂ : A₂ ≃ B₂) (B₂→C₂ : B₂ ≃ C₂) + (A₁→A₂ : B₁ → A₂) (B₁→B₂ : B₁ → B₂) (C₁→C₂ : C₁ → C₂) → + (B₁→B₂) ≡ (fst A₂→B₂ ∘ A₁→A₂) → + (C₁→C₂ ∘ (fst B₁→C₁)) ≡ (fst B₂→C₂ ∘ (B₁→B₂)) → + (C₁→C₂ ∘ (fst B₁→C₁)) ≡ (fst B₂→C₂ ∘ (fst A₂→B₂ ∘ A₁→A₂))) + (EquivJ (λ A₂ A₂→B₂ → (B₂→C₂ : B₂ ≃ C₂) (A₁→A₂ : C₁ → A₂) + (B₁→B₂ : C₁ → B₂) (C₁→C₂ : C₁ → C₂) → + B₁→B₂ ≡ (fst A₂→B₂ ∘ A₁→A₂) → + (C₁→C₂) ≡ (fst B₂→C₂ ∘ B₁→B₂) → + (C₁→C₂) ≡ fst B₂→C₂ ∘ (fst A₂→B₂ ∘ A₁→A₂)) + (EquivJ (λ B₂ B₂→C₂ → (A₁→A₂ B₁→B₂ : C₁ → B₂) (C₁→C₂ : C₁ → C₂) → + B₁→B₂ ≡ A₁→A₂ → + C₁→C₂ ≡ (fst B₂→C₂ ∘ B₁→B₂) → + C₁→C₂ ≡ (fst B₂→C₂ ∘ A₁→A₂)) + λ _ _ _ p q → q ∙ p))) + +{- +We want to show that the following square +commutes. + + Ωⁿ f +Ωⁿ A ----------→ Ωⁿ B +| | +| | +v f∘_ v +(Sⁿ→∙A) ------> (Sⁿ→∙B) +-} + +Ω^→≈post∘∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (f : A →∙ B) + → Path ((Ω^ (suc n)) A →∙ (S₊∙ (suc n) →∙ B ∙)) + (post∘∙ (S₊∙ (suc n)) f ∘∙ Ω→SphereMap∙ (suc n)) + (Ω→SphereMap∙ (suc n) ∘∙ Ω^→ (suc n) f) +Ω^→≈post∘∙ {A = A} {B = B} zero f = + →∙Homogeneous≡ + (subst isHomogeneous + (ua∙ (Ω→SphereMap 1 , (isEquiv-Ω→SphereMap 1)) + (Ω→SphereMap∙ 1 {A = B} .snd)) + (isHomogeneousPath _ _)) + (funExt λ p → + ΣPathP ((funExt (λ { base → snd f + ; (loop i) j → + doubleCompPath-filler + (sym (snd f)) (cong (fst f) p) (snd f) j i})) + , (sym (lUnit (snd f)) ◁ λ i j → snd f (i ∨ j)))) +Ω^→≈post∘∙ {A = A} {B = B} (suc n) f = + →∙Homogeneous≡ + (subst isHomogeneous + (ua∙ (Ω→SphereMap (2 + n) , (isEquiv-Ω→SphereMap (2 + n))) + (Ω→SphereMap∙ (2 + n) {A = B} .snd)) + (isHomogeneousPath _ _)) + ((funExt λ p + → (λ i → post∘∙ (S₊∙ (2 + n)) f .fst (Ω→SphereMap-split (suc n) p i)) + ∙∙ funExt⁻ + (bigLemma + (Ω→SphereMapSplit₁ (suc n) + , isEquivΩ→ _ (isEquiv-Ω→SphereMap (suc n))) + (ΩSphereMap (suc n) , isoToIsEquiv (invIso (SphereMapΩIso (suc n)))) + (Ω→SphereMapSplit₁ (suc n) + , isEquivΩ→ _ (isEquiv-Ω→SphereMap (suc n))) + (ΩSphereMap (suc n) , isoToIsEquiv (invIso (SphereMapΩIso (suc n)))) + (Ω^→ (2 + n) f .fst) (Ω→ (post∘∙ (S₊∙ (suc n)) f) .fst) + (post∘∙ (S₊∙ (2 + n)) f .fst) + (funExt topSquare) + (sym (funExt (bottomSquare f)))) + p + ∙∙ sym (Ω→SphereMap-split (suc n) (Ω^→ (2 + n) f .fst p)))) + where + topSquare : (p : typ ((Ω^ (2 + n)) A)) + → Path (typ (Ω ((S₊∙ (suc n)) →∙ B ∙))) + ((Ω→ (post∘∙ (S₊∙ (suc n)) f) .fst ∘ Ω→ (Ω→SphereMap∙ (suc n)) .fst) p) + (((Ω→ (Ω→SphereMap∙ (suc n))) .fst ∘ (Ω^→ (suc (suc n)) f .fst)) p) + topSquare p = sym (Ω→∘ (post∘∙ (S₊∙ (suc n)) f) (Ω→SphereMap∙ (suc n)) p) + ∙ (λ i → Ω→ (Ω^→≈post∘∙ {A = A} {B = B} n f i) .fst p) + ∙ Ω→∘ (Ω→SphereMap∙ (suc n)) (Ω^→ (suc n) f) p + + bottomSquare : (f : A →∙ B) (g : typ (Ω (S₊∙ (suc n) →∙ A ∙))) + → Path (S₊∙ (2 + n) →∙ B) + (ΩSphereMap (suc n) (Ω→ (post∘∙ (S₊∙ (suc n)) f) .fst g)) + ((post∘∙ (S₊∙ (2 + n)) f .fst ∘ ΩSphereMap (suc n)) g) + bottomSquare = + →∙J (λ b₀ f → (g : typ (Ω (S₊∙ (suc n) →∙ A ∙))) + → Path (S₊∙ (suc (suc n)) →∙ (fst B , b₀)) + (ΩSphereMap (suc n) (Ω→ (post∘∙ (S₊∙ (suc n)) f) .fst g)) + ((post∘∙ (S₊∙ (suc (suc n))) f .fst ∘ ΩSphereMap (suc n)) g)) + λ f g → ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → lem f g a j i})) + , lUnit refl) + where + lem : (f : typ A → typ B) (g : typ (Ω (S₊∙ (suc n) →∙ A ∙))) + → (a : S₊ (suc n)) + → cong (fst (ΩSphereMap (suc n) + (Ω→ (post∘∙ (S₊∙ (suc n)) (f , refl)) .fst g))) + (merid a) + ≡ cong (fst ((f , refl) ∘∙ ΩSphereMap (suc n) g)) (merid a) + lem f g a = + (λ i → funExt⁻ + (cong-∙∙ fst (sym (snd (post∘∙ (S₊∙ (suc n)) (f , (λ _ → f (snd A)))))) + (cong (fst (post∘∙ (S₊∙ (suc n)) (f , (λ _ → f (snd A))))) g) + (snd (post∘∙ (S₊∙ (suc n)) (f , (λ _ → f (snd A))))) i) a) + ∙ sym (rUnit (λ i → f (fst (g i) a))) + +{- We can use this to define prove that post composition induces a homomorphism +πₙ A → πₙ B-} + +π'∘∙fun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → π' (suc n) A → π' (suc n) B +π'∘∙fun n f = sMap (f ∘∙_) + +GroupHomπ≅π'PathP : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') (n : ℕ) + → GroupHom (πGr n A) (πGr n B) ≡ GroupHom (π'Gr n A) (π'Gr n B) +GroupHomπ≅π'PathP A B n i = + GroupHom (fst (GroupPath _ _) (GroupIso→GroupEquiv (π'Gr≅πGr n A)) (~ i)) + (fst (GroupPath _ _) (GroupIso→GroupEquiv (π'Gr≅πGr n B)) (~ i)) + +πFun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → π (suc n) A → π (suc n) B +πFun n f = sMap (fst (Ω^→ (suc n) f)) + +πHom : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → GroupHom (πGr n A) (πGr n B) +fst (πHom n f) = πFun n f +snd (πHom n f) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (Ω^→pres∙ f n p q)) + +π'∘∙Hom' : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → GroupHom (π'Gr n A) (π'Gr n B) +π'∘∙Hom' {A = A} {B = B} n f = + transport (λ i → GroupHomπ≅π'PathP A B n i) + (πHom n f) + +π'∘∙Hom'≡π'∘∙fun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) → π'∘∙Hom' n f .fst ≡ π'∘∙fun n f +π'∘∙Hom'≡π'∘∙fun n f = + funExt (sElim (λ _ → isSetPathImplicit) + λ g → cong ∣_∣₂ + ((λ i → inv (IsoSphereMapΩ (suc n)) + (transportRefl (Ω^→ (suc n) f .fst + (transportRefl (fun (IsoSphereMapΩ (suc n)) g) i)) i)) + ∙ sym (funExt⁻ (cong fst (Ω^→≈post∘∙ n f)) + (fun (IsoSphereMapΩ (suc n)) g)) + ∙ cong (f ∘∙_) (leftInv (IsoSphereMapΩ (suc n)) g))) + +π'∘∙Hom : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → GroupHom (π'Gr n A) (π'Gr n B) +fst (π'∘∙Hom n f) = sMap (f ∘∙_) +snd (π'∘∙Hom {A = A} {B = B} n f) = isHom∘∙ + where + abstract + isHom∘∙ : IsGroupHom (π'Gr n A .snd) (fst (π'∘∙Hom n f)) (π'Gr n B .snd) + isHom∘∙ = + transport (λ i → IsGroupHom (π'Gr n A .snd) + (π'∘∙Hom'≡π'∘∙fun n f i) + (π'Gr n B .snd)) + (π'∘∙Hom' n f .snd) + +-- post composition with an equivalence induces an +-- isomorphism of homotopy groups +π'eqFun : ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) + → A ≃∙ B + → (π' (suc n) A) → π' (suc n) B +π'eqFun n p = π'∘∙fun n (≃∙map p) + +π'eqFun-idEquiv : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → π'eqFun n (idEquiv (fst A) , (λ _ → pt A)) + ≡ idfun _ +π'eqFun-idEquiv n = + funExt (sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (∘∙-idʳ f)) + +π'eqFunIsEquiv : + ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) + → (e : A ≃∙ B) + → isEquiv (π'eqFun n e) +π'eqFunIsEquiv {B = B} n = + Equiv∙J (λ A e → isEquiv (π'eqFun n e)) + (subst isEquiv (sym (π'eqFun-idEquiv n)) + (idIsEquiv (π' (suc n) B))) + +π'eqFunIsHom : ∀ {ℓ} {A B : Pointed ℓ}(n : ℕ) + → (e : A ≃∙ B) + → IsGroupHom (π'Gr n A .snd) (π'eqFun n e) + (π'Gr n B .snd) +π'eqFunIsHom {B = B} n = + Equiv∙J (λ A e → IsGroupHom (π'Gr n A .snd) (π'eqFun n e) (π'Gr n B .snd)) + (subst (λ x → IsGroupHom (π'Gr n B .snd) x (π'Gr n B .snd)) + (sym (π'eqFun-idEquiv n)) + (makeIsGroupHom λ _ _ → refl)) + +π'Iso : ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) + → A ≃∙ B + → GroupEquiv (π'Gr n A) (π'Gr n B) +fst (fst (π'Iso n e)) = π'eqFun n e +snd (fst (π'Iso n e)) = π'eqFunIsEquiv n e +snd (π'Iso n e) = π'eqFunIsHom n e + +πIso : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (A ≃∙ B) + → (n : ℕ) + → GroupEquiv (πGr n A) (πGr n B) +fst (fst (πIso e n)) = fst (πHom n (≃∙map e)) +snd (fst (πIso e n)) = + isoToIsEquiv + (setTruncIso + (equivToIso (_ , isEquivΩ^→ (suc n) (≃∙map e) (snd (fst e))))) +snd (πIso e n) = snd (πHom n (≃∙map e)) diff --git a/Cubical/Homotopy/Group/LES.agda b/Cubical/Homotopy/Group/LES.agda new file mode 100644 index 0000000000..c3a2f53c52 --- /dev/null +++ b/Cubical/Homotopy/Group/LES.agda @@ -0,0 +1,665 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains: +1. The long exact sequence of loop spaces Ωⁿ (fib f) → Ωⁿ A → Ωⁿ B +2. The long exact sequence of homotopy groups πₙ(fib f) → πₙ A → πₙ B +3. Some lemmas relating the map in the sequence to maps using the + other definition of πₙ (maps from Sⁿ) +-} +module Cubical.Homotopy.Group.LES where + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Group.Base + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open Iso +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function + +open import Cubical.HITs.SetTruncation + renaming (rec to sRec + ; elim to sElim ; elim2 to sElim2 + ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec) + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group + +-- We will need an explicitly defined equivalence +-- (PathP (λ i → p i ≡ y) q q) ≃ (sym q ∙∙ p ∙∙ q ≡ refl) +-- This is given by →∙∙lCancel below +module _ {ℓ : Level} {A : Type ℓ} {x y : A} (p : x ≡ x) (q : x ≡ y) where + →∙∙lCancel-fill : PathP (λ i → p i ≡ y) q q → I → I → I → A + →∙∙lCancel-fill PP k i j = + hfill (λ k → λ {(i = i0) → doubleCompPath-filler (sym q) p q k j + ; (i = i1) → y + ; (j = i0) → q (i ∨ k) + ; (j = i1) → q (i ∨ k)}) + (inS (PP j i)) + k + + ←∙∙lCancel-fill : sym q ∙∙ p ∙∙ q ≡ refl → I → I → I → A + ←∙∙lCancel-fill PP k i j = + hfill (λ k → λ {(i = i0) → q (j ∨ ~ k) + ; (i = i1) → q (j ∨ ~ k) + ; (j = i0) → doubleCompPath-filler (sym q) p q (~ k) i + ; (j = i1) → y}) + (inS (PP j i)) + k + + →∙∙lCancel : PathP (λ i → p i ≡ y) q q → sym q ∙∙ p ∙∙ q ≡ refl + →∙∙lCancel PP i j = →∙∙lCancel-fill PP i1 i j + + ←∙∙lCancel : sym q ∙∙ p ∙∙ q ≡ refl → PathP (λ i → p i ≡ y) q q + ←∙∙lCancel PP i j = ←∙∙lCancel-fill PP i1 i j + + ←∙∙lCancel→∙∙lCancel : (PP : PathP (λ i → p i ≡ y) q q) + → ←∙∙lCancel (→∙∙lCancel PP) ≡ PP + ←∙∙lCancel→∙∙lCancel PP r i j = + hcomp (λ k → λ {(r = i0) → ←∙∙lCancel-fill (→∙∙lCancel PP) k i j + ; (r = i1) → PP i j + ; (j = i0) → doubleCompPath-filler (sym q) p q (~ k ∧ ~ r) i + ; (j = i1) → y + ; (i = i0) → q (j ∨ ~ k ∧ ~ r) + ; (i = i1) → q (j ∨ ~ k ∧ ~ r)}) + (hcomp (λ k → λ {(r = i0) → →∙∙lCancel-fill PP k j i + ; (r = i1) → PP i j + ; (j = i0) → doubleCompPath-filler (sym q) p q (k ∧ ~ r) i + ; (j = i1) → y + ; (i = i0) → q (j ∨ k ∧ ~ r) + ; (i = i1) → q (j ∨ k ∧ ~ r)}) + (PP i j)) + + →∙∙lCancel←∙∙lCancel : (PP : sym q ∙∙ p ∙∙ q ≡ refl) + → →∙∙lCancel (←∙∙lCancel PP) ≡ PP + →∙∙lCancel←∙∙lCancel PP r i j = + hcomp (λ k → λ {(r = i0) → →∙∙lCancel-fill (←∙∙lCancel PP) k i j + ; (r = i1) → PP i j + ; (j = i0) → q (i ∨ k ∨ r) + ; (j = i1) → q (i ∨ k ∨ r) + ; (i = i0) → doubleCompPath-filler (sym q) p q (r ∨ k) j + ; (i = i1) → y}) + (hcomp (λ k → λ {(r = i0) → ←∙∙lCancel-fill PP k j i + ; (r = i1) → PP i j + ; (j = i0) → q (i ∨ r ∨ ~ k) + ; (j = i1) → q (i ∨ r ∨ ~ k) + ; (i = i0) → doubleCompPath-filler (sym q) p q (r ∨ ~ k) j + ; (i = i1) → y}) + (PP i j)) + +←∙∙lCancel-refl-refl : + {ℓ : Level} {A : Type ℓ} {x : A} (p : refl {x = x} ≡ refl) + → ←∙∙lCancel {x = x} {y = x} refl refl (sym (rUnit refl) ∙ p) + ≡ flipSquare p +←∙∙lCancel-refl-refl p k i j = + hcomp (λ r → λ { (i = i0) → p i0 i0 + ; (i = i1) → p i0 i0 + ; (j = i0) → rUnit (λ _ → p i0 i0) (~ r) i + ; (j = i1) → p i0 i0 + ; (k = i0) → ←∙∙lCancel-fill refl refl (sym (rUnit refl) ∙ p) r i j + ; (k = i1) → compPath-filler' (sym (rUnit refl)) p (~ r) j i}) + ((sym (rUnit refl) ∙ p) j i) + +{- We need an iso Ω(fib f) ≅ fib(Ω f) -} +ΩFibreIso : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Iso (typ (Ω (fiber (fst f) (pt B) , (pt A) , snd f))) + (fiber (Ω→ f .fst) refl) +fun (ΩFibreIso f) p = (cong fst p) , + →∙∙lCancel (cong (fst f) (cong fst p)) (snd f) + (cong snd p) +fst (inv (ΩFibreIso f) (p , q) i) = p i +snd (inv (ΩFibreIso f) (p , q) i) = ←∙∙lCancel (cong (fst f) p) (snd f) q i +rightInv (ΩFibreIso f) (p , q) = ΣPathP (refl , →∙∙lCancel←∙∙lCancel _ _ q) +fst (leftInv (ΩFibreIso f) p i j) = fst (p j) +snd (leftInv (ΩFibreIso f) p i j) k = + ←∙∙lCancel→∙∙lCancel _ _ (cong snd p) i j k + +{- Some homomorphism properties of the above iso -} +ΩFibreIsopres∙fst : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → (p q : (typ (Ω (fiber (fst f) (pt B) , (pt A) , snd f)))) + → fst (fun (ΩFibreIso f) (p ∙ q)) + ≡ fst (fun (ΩFibreIso f) p) ∙ fst (fun (ΩFibreIso f) q) +ΩFibreIsopres∙fst f p q = cong-∙ fst p q + +ΩFibreIso⁻pres∙snd : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (f : A →∙ B) (p q : typ (Ω (Ω B))) + → inv (ΩFibreIso f) (refl , (Ω→ f .snd ∙ p ∙ q)) + ≡ inv (ΩFibreIso f) (refl , Ω→ f .snd ∙ p) + ∙ inv (ΩFibreIso f) (refl , Ω→ f .snd ∙ q) +ΩFibreIso⁻pres∙snd {A = A} {B = B}= + →∙J (λ b₀ f → (p q : typ (Ω (Ω (fst B , b₀)))) + → inv (ΩFibreIso f) (refl , (Ω→ f .snd ∙ p ∙ q)) + ≡ inv (ΩFibreIso f) (refl , Ω→ f .snd ∙ p) + ∙ inv (ΩFibreIso f) (refl , Ω→ f .snd ∙ q)) + ind + where + ind : (f : typ A → typ B) (p q : typ (Ω (Ω (fst B , f (pt A))))) + → inv (ΩFibreIso (f , refl)) (refl , (sym (rUnit refl) ∙ p ∙ q)) + ≡ inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ p) + ∙ inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ q) + fst (ind f p q i j) = + (rUnit refl + ∙ sym (cong-∙ fst + (inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ p)) + (inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ q)))) i j + snd (ind f p q i j) k = + hcomp (λ r + → λ {(i = i0) → ←∙∙lCancel-refl-refl (p ∙ q) (~ r) j k -- + ; (i = i1) → + snd (compPath-filler + (inv (ΩFibreIso (f , refl)) + (refl , sym (rUnit refl) ∙ p)) + (inv (ΩFibreIso (f , refl)) + (refl , sym (rUnit refl) ∙ q)) r j) k + ; (j = i0) → f (snd A) + ; (j = i1) → snd (inv (ΩFibreIso (f , refl)) + (refl , sym (rUnit refl) ∙ q) (r ∨ ~ i)) k + ; (k = i0) → main r i j + ; (k = i1) → f (snd A)}) + (hcomp (λ r → λ {(i = i0) → (p ∙ q) k j + ; (i = i1) → ←∙∙lCancel-refl-refl p (~ r) j k + ; (j = i0) → f (snd A) + ; (j = i1) → ←∙∙lCancel-refl-refl q (~ r) (~ i) k + ; (k = i0) → f (pt A) + ; (k = i1) → f (snd A)}) + (hcomp (λ r → λ {(i = i0) → (compPath-filler' p q r) k j + ; (i = i1) → p (k ∨ ~ r) j + ; (j = i0) → f (snd A) + ; (j = i1) → q k (~ i) + ; (k = i0) → p (k ∨ ~ r) j + ; (k = i1) → f (snd A)}) + (q k (~ i ∧ j)))) + where + P = (inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ p)) + Q = (inv (ΩFibreIso (f , refl)) (refl , sym (rUnit refl) ∙ q)) + + main : I → I → I → fst B + main r i j = + hcomp (λ k → λ {(i = i0) → f (snd A) + ; (i = i1) → f (fst (compPath-filler P Q (r ∨ ~ k) j)) + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (r = i0) → f (fst (compPath-filler P Q (i ∧ ~ k) j)) + ; (r = i1) → f ((rUnit refl ∙ sym (cong-∙ fst P Q)) i j)}) + (hcomp (λ k → λ {(i = i0) → f (rUnit (λ _ → pt A) (~ k ∧ r) j) + ; (i = i1) → f (fst ((P ∙ Q) j)) + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (r = i0) → f (fst (compPath-filler P Q i j)) + ; (r = i1) → f ((compPath-filler' (rUnit refl) + (sym (cong-∙ fst P Q)) k) i j)}) + (hcomp (λ k → λ {(i = i0) → f (rUnit (λ _ → pt A) (k ∧ r) j) + ; (i = i1) → f (fst (compPath-filler P Q k j)) + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (r = i0) → f (fst (compPath-filler P Q (i ∧ k) j)) + ; (r = i1) → f ((cong-∙∙-filler fst refl P Q) k (~ i) j)}) + (f (snd A)))) + +ΩFibreIso∙ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Iso.fun (ΩFibreIso f) refl ≡ (refl , (∙∙lCancel (snd f))) +ΩFibreIso∙ {A = A} {B = B} = + →∙J (λ b f → Iso.fun (ΩFibreIso f) refl ≡ (refl , (∙∙lCancel (snd f)))) + λ f → ΣPathP (refl , help f) + where + help : (f : fst A → fst B) → + →∙∙lCancel (λ i → f (snd A)) refl (λ i → refl) + ≡ ∙∙lCancel refl + help f i j r = + hcomp (λ k → λ {(i = i0) → + →∙∙lCancel-fill (λ _ → f (snd A)) refl refl k j r + ; (i = i1) → ∙∙lCancel-fill (λ _ → f (snd A)) j r k + ; (j = i0) → rUnit (λ _ → f (snd A)) k r + ; (j = i1) → f (snd A) + ; (r = i1) → f (snd A) + ; (r = i0) → f (snd A)}) + (f (snd A)) + +ΩFibreIso⁻∙ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Iso.inv (ΩFibreIso f) (refl , (∙∙lCancel (snd f))) ≡ refl +ΩFibreIso⁻∙ f = + cong (Iso.inv (ΩFibreIso f)) (sym (ΩFibreIso∙ f)) ∙ leftInv (ΩFibreIso f) refl + +{- Ωⁿ (fib f) ≃∙ fib (Ωⁿ f) -} +Ω^Fibre≃∙ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → ((Ω^ n) (fiber (fst f) (pt B) , (pt A) , snd f)) + ≃∙ ((fiber (Ω^→ n f .fst) (snd ((Ω^ n) B))) + , (snd ((Ω^ n) A)) , (Ω^→ n f .snd)) +Ω^Fibre≃∙ zero f = (idEquiv _) , refl +Ω^Fibre≃∙ (suc n) f = + compEquiv∙ + (Ω≃∙ (Ω^Fibre≃∙ n f)) + ((isoToEquiv (ΩFibreIso (Ω^→ n f))) , ΩFibreIso∙ (Ω^→ n f)) + +{- Its inverse iso directly defined -} +Ω^Fibre≃∙⁻ : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → ((fiber (Ω^→ n f .fst) (snd ((Ω^ n) B))) + , (snd ((Ω^ n) A)) , (Ω^→ n f .snd)) + ≃∙ ((Ω^ n) (fiber (fst f) (pt B) , (pt A) , snd f)) +Ω^Fibre≃∙⁻ zero f = (idEquiv _) , refl +Ω^Fibre≃∙⁻ (suc n) f = + compEquiv∙ + ((isoToEquiv (invIso (ΩFibreIso (Ω^→ n f)))) + , (ΩFibreIso⁻∙ (Ω^→ n f))) + (Ω≃∙ (Ω^Fibre≃∙⁻ n f)) + +isHomogeneousΩ^→fib : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → isHomogeneous + ((fiber (Ω^→ (suc n) f .fst) (snd ((Ω^ (suc n)) B))) + , (snd ((Ω^ (suc n)) A)) , (Ω^→ (suc n) f .snd)) +isHomogeneousΩ^→fib n f = + subst isHomogeneous (ua∙ ((fst (Ω^Fibre≃∙ (suc n) f))) + (snd (Ω^Fibre≃∙ (suc n) f))) + (isHomogeneousPath _ _) + +Ω^Fibre≃∙sect : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → (≃∙map (Ω^Fibre≃∙⁻ n f) ∘∙ ≃∙map (Ω^Fibre≃∙ n f)) + ≡ idfun∙ _ +Ω^Fibre≃∙sect zero f = ΣPathP (refl , (sym (rUnit refl))) +Ω^Fibre≃∙sect (suc n) f = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt + λ p → cong (fst (fst (Ω≃∙ (Ω^Fibre≃∙⁻ n f)))) + (leftInv (ΩFibreIso (Ω^→ n f)) + ((fst (fst (Ω≃∙ (Ω^Fibre≃∙ n f))) p))) + ∙ sym (Ω→∘ (≃∙map (Ω^Fibre≃∙⁻ n f)) + (≃∙map (Ω^Fibre≃∙ n f)) p) + ∙ (λ i → (Ω→ (Ω^Fibre≃∙sect n f i)) .fst p) + ∙ sym (rUnit p)) + +Ω^Fibre≃∙retr : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → (≃∙map (Ω^Fibre≃∙ n f) ∘∙ ≃∙map (Ω^Fibre≃∙⁻ n f)) + ≡ idfun∙ _ +Ω^Fibre≃∙retr zero f = ΣPathP (refl , (sym (rUnit refl))) +Ω^Fibre≃∙retr (suc n) f = + →∙Homogeneous≡ (isHomogeneousΩ^→fib n f) + (funExt (λ p → + cong (fun (ΩFibreIso (Ω^→ n f))) + ((sym (Ω→∘ (≃∙map (Ω^Fibre≃∙ n f)) + (≃∙map (Ω^Fibre≃∙⁻ n f)) + (inv (ΩFibreIso (Ω^→ n f)) p))) + ∙ (λ i → Ω→ (Ω^Fibre≃∙retr n f i) .fst (inv (ΩFibreIso (Ω^→ n f)) p)) + ∙ sym (rUnit (inv (ΩFibreIso (Ω^→ n f)) p))) + ∙ rightInv (ΩFibreIso (Ω^→ n f)) p)) + +Ω^Fibre≃∙' : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → ((Ω^ n) (fiber (fst f) (pt B) , (pt A) , snd f)) + ≃∙ ((fiber (Ω^→ n f .fst) (snd ((Ω^ n) B))) + , (snd ((Ω^ n) A)) , (Ω^→ n f .snd)) +Ω^Fibre≃∙' zero f = idEquiv _ , refl +Ω^Fibre≃∙' (suc zero) f = + (isoToEquiv (ΩFibreIso (Ω^→ zero f))) , ΩFibreIso∙ (Ω^→ zero f) +Ω^Fibre≃∙' (suc (suc n)) f = + compEquiv∙ + (Ω≃∙ (Ω^Fibre≃∙ (suc n) f)) + ((isoToEquiv (ΩFibreIso (Ω^→ (suc n) f))) , ΩFibreIso∙ (Ω^→ (suc n) f)) + +-- The long exact sequence of loop spaces. +module ΩLES {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where + {- Fibre of f -} + fibf : Pointed _ + fibf = fiber (fst f) (pt B) , (pt A , snd f) + + {- Fibre of Ωⁿ f -} + fibΩ^f : (n : ℕ) → Pointed _ + fst (fibΩ^f n) = fiber (Ω^→ n f .fst) (snd ((Ω^ n) B)) + snd (fibΩ^f n) = (snd ((Ω^ n) A)) , (Ω^→ n f .snd) + + Ω^fibf : (n : ℕ) → Pointed _ + Ω^fibf n = (Ω^ n) fibf + + {- Helper function fib (Ωⁿ f) → Ωⁿ A -} + fibΩ^f→A : (n : ℕ) → fibΩ^f n →∙ (Ω^ n) A + fst (fibΩ^f→A n) = fst + snd (fibΩ^f→A n) = refl + + {- The main function Ωⁿ(fib f) → Ωⁿ A, which is just the composition + Ωⁿ(fib f) ≃ fib (Ωⁿ f) → Ωⁿ A, where the last function is + fibΩ^f→A. Hence most proofs will concern fibΩ^f→A, since it is easier to + work with. -} + Ω^fibf→A : (n : ℕ) → Ω^fibf n →∙ (Ω^ n) A + Ω^fibf→A n = fibΩ^f→A n ∘∙ ≃∙map (Ω^Fibre≃∙ n f) + + {- The function preserves path composition -} + Ω^fibf→A-pres∙ : (n : ℕ) → (p q : Ω^fibf (suc n) .fst) + → Ω^fibf→A (suc n) .fst (p ∙ q) + ≡ Ω^fibf→A (suc n) .fst p + ∙ Ω^fibf→A (suc n) .fst q + Ω^fibf→A-pres∙ n p q = + cong (fst (fibΩ^f→A (suc n))) + (cong (fun (ΩFibreIso (Ω^→ n f))) + (Ω→pres∙ (≃∙map (Ω^Fibre≃∙ n f)) p q)) + ∙ ΩFibreIsopres∙fst (Ω^→ n f) + (fst (Ω→ (≃∙map (Ω^Fibre≃∙ n f))) p) + (fst (Ω→ (≃∙map (Ω^Fibre≃∙ n f))) q) + + {- The function Ωⁿ A → Ωⁿ B -} + A→B : (n : ℕ) → (Ω^ n) A →∙ (Ω^ n) B + A→B n = Ω^→ n f + + {- It preserves path composition -} + A→B-pres∙ : (n : ℕ) → (p q : typ ((Ω^ (suc n)) A)) + → fst (A→B (suc n)) (p ∙ q) + ≡ fst (A→B (suc n)) p ∙ fst (A→B (suc n)) q + A→B-pres∙ n p q = Ω^→pres∙ f n p q + + {- Helper function Ωⁿ⁺¹ B → fib (Ωⁿ f) -} + ΩB→fibΩ^f : (n : ℕ) → ((Ω^ (suc n)) B) →∙ fibΩ^f n + fst (ΩB→fibΩ^f n) x = (snd ((Ω^ n) A)) , (Ω^→ n f .snd ∙ x) + snd (ΩB→fibΩ^f n) = ΣPathP (refl , (sym (rUnit _))) + + {- The main function Ωⁿ⁺¹ B → Ωⁿ (fib f), + factoring through the above function -} + ΩB→Ω^fibf : (n : ℕ) → (Ω^ (suc n)) B →∙ Ω^fibf n + ΩB→Ω^fibf n = + (≃∙map (Ω^Fibre≃∙⁻ n f)) + ∘∙ ΩB→fibΩ^f n + + {- It preserves path composition -} + ΩB→Ω^fibf-pres∙ : (n : ℕ) → (p q : typ ((Ω^ (2 + n)) B)) + → fst (ΩB→Ω^fibf (suc n)) (p ∙ q) + ≡ fst (ΩB→Ω^fibf (suc n)) p ∙ fst (ΩB→Ω^fibf (suc n)) q + ΩB→Ω^fibf-pres∙ n p q = + cong (fst (fst (Ω^Fibre≃∙⁻ (suc n) f))) + refl + ∙ cong (fst (fst (Ω≃∙ (Ω^Fibre≃∙⁻ n f)))) + (cong (fun (invIso (ΩFibreIso (Ω^→ n f)))) + (λ _ → snd ((Ω^ suc n) A) , Ω^→ (suc n) f .snd ∙ p ∙ q)) + ∙ cong (fst (fst (Ω≃∙ (Ω^Fibre≃∙⁻ n f)))) + (ΩFibreIso⁻pres∙snd (Ω^→ n f) p q) + ∙ Ω≃∙pres∙ (Ω^Fibre≃∙⁻ n f) + (inv (ΩFibreIso (Ω^→ n f)) (refl , Ω→ (Ω^→ n f) .snd ∙ p)) + (inv (ΩFibreIso (Ω^→ n f)) (refl , Ω→ (Ω^→ n f) .snd ∙ q)) + + {- Hence we have our sequence + ... → Ωⁿ⁺¹B → Ωⁿ(fib f) → Ωⁿ A → Ωⁿ B → ... (*) + We first prove the exactness properties for the helper functions + ΩB→fibΩ^f and fibΩ^f→A, and then deduce exactness of the whole sequence + by noting that the functions in (*) are just ΩB→fibΩ^f, fibΩ^f→A but + composed with equivalences + -} + private + Im-fibΩ^f→A⊂Ker-A→B : (n : ℕ) (x : _) + → isInIm∙ (fibΩ^f→A n) x + → isInKer∙ (A→B n) x + Im-fibΩ^f→A⊂Ker-A→B n x = + uncurry λ p → J (λ x _ → isInKer∙ (A→B n) x) + (snd p) + + Ker-fibΩ^f→A⊂Im-ΩB→fibΩ^f : (n : ℕ) (x : _) + → isInKer∙ (fibΩ^f→A n) x + → isInIm∙ (ΩB→fibΩ^f n) x + Ker-fibΩ^f→A⊂Im-ΩB→fibΩ^f n x ker = + (sym (Ω^→ n f .snd) + ∙ cong (Ω^→ n f .fst) (sym ker) ∙ snd x) , ΣPathP ((sym ker) , + ((∙assoc (Ω^→ n f .snd) + (sym (Ω^→ n f .snd)) + (sym (cong (Ω^→ n f .fst) ker) ∙ snd x) + ∙∙ cong (_∙ (sym (cong (Ω^→ n f .fst) ker) ∙ snd x)) + (rCancel (Ω^→ n f .snd)) + ∙∙ sym (lUnit (sym (cong (Ω^→ n f .fst) ker) ∙ snd x))) + ◁ (λ i j → compPath-filler' + (cong (Ω^→ n f .fst) (sym ker)) (snd x) (~ i) j))) + + Im-A→B⊂Ker-ΩB→fibΩ^f : (n : ℕ) (x : fst (((Ω^ (suc n)) B))) + → isInIm∙ (A→B (suc n)) x + → isInKer∙ (ΩB→fibΩ^f n) x + Im-A→B⊂Ker-ΩB→fibΩ^f n x = + uncurry λ p + → J (λ x _ → isInKer∙ (ΩB→fibΩ^f n) x) + (ΣPathP (p , (((λ i → (λ j → Ω^→ n f .snd (j ∧ ~ i)) + ∙ ((λ j → Ω^→ n f .snd (~ j ∧ ~ i)) + ∙∙ cong (Ω^→ n f .fst) p + ∙∙ Ω^→ n f .snd)) + ∙ sym (lUnit (cong (Ω^→ n f .fst) p ∙ Ω^→ n f .snd))) + ◁ λ i j → compPath-filler' + (cong (Ω^→ n f .fst) p) (Ω^→ n f .snd) (~ i) j))) + + Ker-ΩB→fibΩ^f⊂Im-A→B : (n : ℕ) (x : fst (((Ω^ (suc n)) B))) + → isInKer∙ (ΩB→fibΩ^f n) x + → isInIm∙ (A→B (suc n)) x + fst (Ker-ΩB→fibΩ^f⊂Im-A→B n x inker) = cong fst inker + snd (Ker-ΩB→fibΩ^f⊂Im-A→B n x inker) = lem + where + lem : fst (A→B (suc n)) (λ i → fst (inker i)) ≡ x + lem i j = + hcomp (λ k → λ { (i = i0) → doubleCompPath-filler + (sym (snd (Ω^→ n f))) + ((λ i → Ω^→ n f .fst (fst (inker i)))) + (snd (Ω^→ n f)) k j + ; (i = i1) → compPath-filler' (Ω^→ n f .snd) x (~ k) j + ; (j = i0) → snd (Ω^→ n f) k + ; (j = i1) → snd (Ω^→ n f) (k ∨ i)}) + (hcomp (λ k → λ { (i = i0) → (snd (inker j)) (~ k) + ; (i = i1) → ((Ω^→ n f .snd) ∙ x) (j ∨ ~ k) + ; (j = i0) → ((Ω^→ n f .snd) ∙ x) (~ k) + ; (j = i1) → snd (Ω^→ n f) (i ∨ ~ k)}) + (snd ((Ω^ n) B))) + + {- Finally, we get exactness of the sequence + we are interested in -} + Im-Ω^fibf→A⊂Ker-A→B : (n : ℕ) (x : _) + → isInIm∙ (Ω^fibf→A n) x + → isInKer∙ (A→B n) x + Im-Ω^fibf→A⊂Ker-A→B n x x₁ = + Im-fibΩ^f→A⊂Ker-A→B n x + (((fst (fst (Ω^Fibre≃∙ n f))) (fst x₁)) + , snd x₁) + + Ker-A→B⊂Im-Ω^fibf→A : (n : ℕ) (x : _) + → isInKer∙ (A→B n) x + → isInIm∙ (Ω^fibf→A n) x + Ker-A→B⊂Im-Ω^fibf→A n x ker = + invEq (fst (Ω^Fibre≃∙ n f)) (x , ker) + , (cong fst (secEq (fst (Ω^Fibre≃∙ n f)) (x , ker))) + + Ker-Ω^fibf→A⊂Im-ΩB→Ω^fibf : (n : ℕ) (x : _) + → isInKer∙ (Ω^fibf→A n) x + → isInIm∙ (ΩB→Ω^fibf n) x + Ker-Ω^fibf→A⊂Im-ΩB→Ω^fibf n x p = + fst r + , cong (fst ((fst (Ω^Fibre≃∙⁻ n f)))) (snd r) + ∙ funExt⁻ (cong fst (Ω^Fibre≃∙sect n f)) x + where + r : isInIm∙ (ΩB→fibΩ^f n) (fst (fst (Ω^Fibre≃∙ n f)) x) + r = Ker-fibΩ^f→A⊂Im-ΩB→fibΩ^f n (fst (fst (Ω^Fibre≃∙ n f)) x) p + + Im-ΩB→Ω^fibf⊂Ker-Ω^fibf→A : (n : ℕ) (x : _) + → isInIm∙ (ΩB→Ω^fibf n) x + → isInKer∙ (Ω^fibf→A n) x + Im-ΩB→Ω^fibf⊂Ker-Ω^fibf→A n x = + uncurry λ p → + J (λ x _ → isInKer∙ (Ω^fibf→A n) x) + (cong (fst (fibΩ^f→A n)) + (funExt⁻ (cong fst (Ω^Fibre≃∙retr n f)) _)) + + Im-A→B⊂Ker-ΩB→Ω^fibf : (n : ℕ) (x : fst (((Ω^ (suc n)) B))) + → isInIm∙ (A→B (suc n)) x + → isInKer∙ (ΩB→Ω^fibf n) x + Im-A→B⊂Ker-ΩB→Ω^fibf n x p = + cong (fst ((fst (Ω^Fibre≃∙⁻ n f)))) + (Im-A→B⊂Ker-ΩB→fibΩ^f n x p) + ∙ snd (Ω^Fibre≃∙⁻ n f) + + Ker-ΩB→Ω^fibf⊂Im-A→B : (n : ℕ) (x : fst (((Ω^ (suc n)) B))) + → isInKer∙ (ΩB→Ω^fibf n) x + → isInIm∙ (A→B (suc n)) x + Ker-ΩB→Ω^fibf⊂Im-A→B n x p = + Ker-ΩB→fibΩ^f⊂Im-A→B n x + (funExt⁻ (cong fst (sym (Ω^Fibre≃∙retr n f))) (ΩB→fibΩ^f n .fst x) + ∙ cong (fst (fst (Ω^Fibre≃∙ n f))) p + ∙ snd (Ω^Fibre≃∙ n f)) + +{- Some useful lemmas for converting the above sequence a +a sequence of homotopy groups -} +module setTruncLemmas {ℓ ℓ' ℓ'' : Level} + {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (n m l : ℕ) + (f : (Ω ((Ω^ n) A)) →∙ (Ω ((Ω^ m) B))) + (g : (Ω ((Ω^ m) B)) →∙ (Ω ((Ω^ l) C))) + (e₁ : IsGroupHom (snd (πGr n A)) (sMap (fst f)) (snd (πGr m B))) + (e₂ : IsGroupHom (snd (πGr m B)) (sMap (fst g)) (snd (πGr l C))) where + + ker⊂im : ((x : typ (Ω ((Ω^ m) B))) → isInKer∙ g x → isInIm∙ f x) + → (x : π (suc m) B) → isInKer (_ , e₂) x → isInIm (_ , e₁) x + ker⊂im ind = + sElim (λ _ → isSetΠ λ _ → isProp→isSet squash) + λ p ker → + pRec squash + (λ ker∙ → ∣ ∣ ind p ker∙ .fst ∣₂ , cong ∣_∣₂ (ind p ker∙ .snd) ∣ ) + (fun PathIdTrunc₀Iso ker) + + im⊂ker : ((x : typ (Ω ((Ω^ m) B))) → isInIm∙ f x → isInKer∙ g x) + → (x : π (suc m) B) → isInIm (_ , e₁) x → isInKer (_ , e₂) x + im⊂ker ind = + sElim (λ _ → isSetΠ λ _ → isSetPathImplicit) + λ p → + pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ λ _ → isSetPathImplicit) + λ a q → pRec (squash₂ _ _) + (λ q → cong ∣_∣₂ (ind p (a , q))) + (fun PathIdTrunc₀Iso q))) + +{- The long exact sequence of homotopy groups -} +module πLES {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where + module Ωs = ΩLES f + open Ωs renaming (A→B to A→B') + + fib = fibf + + fib→A : (n : ℕ) → GroupHom (πGr n fib) (πGr n A) + fst (fib→A n) = sMap (fst (Ω^fibf→A (suc n))) + snd (fib→A n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (Ω^fibf→A-pres∙ n p q)) + + A→B : (n : ℕ) → GroupHom (πGr n A) (πGr n B) + fst (A→B n) = sMap (fst (A→B' (suc n))) + snd (A→B n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ g h → cong ∣_∣₂ (Ω^→pres∙ f n g h)) + + B→fib : (n : ℕ) → GroupHom (πGr (suc n) B) (πGr n fib) + fst (B→fib n) = sMap (fst (ΩB→Ω^fibf (suc n))) + snd (B→fib n) = + makeIsGroupHom + (sElim2 + (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (ΩB→Ω^fibf-pres∙ n p q)) + + Ker-A→B⊂Im-fib→A : (n : ℕ) (x : π (suc n) A) + → isInKer (A→B n) x + → isInIm (fib→A n) x + Ker-A→B⊂Im-fib→A n = + setTruncLemmas.ker⊂im n n n + (Ω^fibf→A (suc n)) (A→B' (suc n)) + (snd (fib→A n)) (snd (A→B n)) + (Ker-A→B⊂Im-Ω^fibf→A (suc n)) + + Im-fib→A⊂Ker-A→B : (n : ℕ) (x : π (suc n) A) + → isInIm (fib→A n) x + → isInKer (A→B n) x + Im-fib→A⊂Ker-A→B n = + setTruncLemmas.im⊂ker n n n + (Ω^fibf→A (suc n)) (A→B' (suc n)) + (snd (fib→A n)) (snd (A→B n)) + (Im-Ω^fibf→A⊂Ker-A→B (suc n)) + + Ker-fib→A⊂Im-B→fib : (n : ℕ) (x : π (suc n) fib) + → isInKer (fib→A n) x + → isInIm (B→fib n) x + Ker-fib→A⊂Im-B→fib n = + setTruncLemmas.ker⊂im (suc n) n n + (ΩB→Ω^fibf (suc n)) (Ω^fibf→A (suc n)) + (snd (B→fib n)) (snd (fib→A n)) + (Ker-Ω^fibf→A⊂Im-ΩB→Ω^fibf (suc n)) + + Im-B→fib⊂Ker-fib→A : (n : ℕ) (x : π (suc n) fib) + → isInIm (B→fib n) x + → isInKer (fib→A n) x + Im-B→fib⊂Ker-fib→A n = + setTruncLemmas.im⊂ker (suc n) n n + (ΩB→Ω^fibf (suc n)) (Ω^fibf→A (suc n)) + (snd (B→fib n)) (snd (fib→A n)) + (Im-ΩB→Ω^fibf⊂Ker-Ω^fibf→A (suc n)) + + Im-A→B⊂Ker-B→fib : (n : ℕ) (x : π (suc (suc n)) B) + → isInIm (A→B (suc n)) x + → isInKer (B→fib n) x + Im-A→B⊂Ker-B→fib n = + setTruncLemmas.im⊂ker (suc n) (suc n) n + (A→B' (suc (suc n))) (ΩB→Ω^fibf (suc n)) + (snd (A→B (suc n))) (snd (B→fib n)) + (Im-A→B⊂Ker-ΩB→Ω^fibf (suc n)) + + Ker-B→fib⊂Im-A→B : (n : ℕ) (x : π (suc (suc n)) B) + → isInKer (B→fib n) x + → isInIm (A→B (suc n)) x + Ker-B→fib⊂Im-A→B n = + setTruncLemmas.ker⊂im (suc n) (suc n) n + (A→B' (suc (suc n))) (ΩB→Ω^fibf (suc n)) + (snd (A→B (suc n))) (snd (B→fib n)) + (Ker-ΩB→Ω^fibf⊂Im-A→B (suc n)) + +{- We prove that the map Ωⁿ(fib f) → Ωⁿ A indeed is just the map +Ωⁿ fst -} +private + Ω^fibf→A-ind : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → ΩLES.Ω^fibf→A f (suc n) ≡ Ω→ (ΩLES.Ω^fibf→A f n) + Ω^fibf→A-ind {A = A} {B = B} n f = + (λ _ → πLES.Ωs.fibΩ^f→A f (suc n) + ∘∙ ≃∙map (Ω^Fibre≃∙ (suc n) f)) + ∙ →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ p → + (λ j → cong fst (Ω→ (≃∙map (Ω^Fibre≃∙ n f)) .fst p)) + ∙ rUnit ((λ i → fst + (Ω→ (≃∙map (Ω^Fibre≃∙ n f)) .fst p i))) + ∙ sym (Ω→∘ (πLES.Ωs.fibΩ^f→A f n) (≃∙map (Ω^Fibre≃∙ n f)) p)) + +Ω^fibf→A≡ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → ΩLES.Ω^fibf→A f n ≡ Ω^→ n (fst , refl) +Ω^fibf→A≡ zero f = ΣPathP (refl , (sym (lUnit refl))) +Ω^fibf→A≡ (suc n) f = Ω^fibf→A-ind n f ∙ cong Ω→ (Ω^fibf→A≡ n f) + +{- We now get a nice characterisation of the functions in the induced LES +of homotopy groups defined using (Sⁿ →∙ A) -} + +π∘∙A→B-PathP : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → PathP (λ i → GroupHomπ≅π'PathP A B n i) + (πLES.A→B f n) + (π'∘∙Hom n f) +π∘∙A→B-PathP n f = + toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) (π'∘∙Hom'≡π'∘∙fun n f)) + +π∘∙fib→A-PathP : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) + → PathP (λ i → GroupHomπ≅π'PathP (ΩLES.fibf f) A n i) + (πLES.fib→A f n) + (π'∘∙Hom n (fst , refl)) +π∘∙fib→A-PathP {A = A} {B = B} n f = + toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (cong (transport + (λ i → (fst (GroupPath _ _) + (GroupIso→GroupEquiv (π'Gr≅πGr n (ΩLES.fibf f))) (~ i)) .fst + → (fst (GroupPath _ _) + (GroupIso→GroupEquiv (π'Gr≅πGr n A)) (~ i)) .fst)) + lem + ∙ π'∘∙Hom'≡π'∘∙fun n (fst , refl))) + where + lem : πLES.fib→A f n .fst ≡ sMap (Ω^→ (suc n) (fst , refl) .fst) + lem = cong sMap (cong fst (Ω^fibf→A≡ (suc n) f)) diff --git a/Cubical/Homotopy/Group/Pi3S2.agda b/Cubical/Homotopy/Group/Pi3S2.agda new file mode 100644 index 0000000000..d265a1cf89 --- /dev/null +++ b/Cubical/Homotopy/Group/Pi3S2.agda @@ -0,0 +1,167 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains: +1. The iso π₃S²≅ℤ +2. A proof that π₃S² is generated by the Hopf map +-} +module Cubical.Homotopy.Group.Pi3S2 where + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Homotopy.Group.LES +open import Cubical.Homotopy.Group.PinSn +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.HopfInvariant.HopfMap +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.HopfInvariant.Homomorphism + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open Iso +open import Cubical.Foundations.Equiv + +open import Cubical.HITs.SetTruncation renaming (elim to sElim) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Int hiding (ℤ) + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact + +TotalHopf→∙S² : (Σ (S₊ 2) S¹Hopf , north , base) →∙ S₊∙ 2 +fst TotalHopf→∙S² = fst +snd TotalHopf→∙S² = refl + +IsoTotalSpaceJoin' : Iso (Σ (S₊ 2) S¹Hopf) (S₊ 3) +IsoTotalSpaceJoin' = compIso hopfS¹.IsoTotalSpaceJoin (IsoSphereJoin 1 1) + +IsoFiberTotalHopfS¹ : Iso (fiber (fst TotalHopf→∙S²) north) S¹ +fun IsoFiberTotalHopfS¹ ((x , y) , z) = subst S¹Hopf z y +inv IsoFiberTotalHopfS¹ x = (north , x) , refl +rightInv IsoFiberTotalHopfS¹ x = refl +leftInv IsoFiberTotalHopfS¹ ((x , y) , z) = + ΣPathP + ((ΣPathP + (sym z , (λ i → transp (λ j → S¹Hopf (z (~ i ∧ j))) i y))) + , (λ j i → z (i ∨ ~ j))) + +IsoFiberTotalHopfS¹∙≡ : + (fiber (fst TotalHopf→∙S²) north , (north , base) , refl) ≡ S₊∙ 1 +IsoFiberTotalHopfS¹∙≡ = ua∙ (isoToEquiv IsoFiberTotalHopfS¹) refl + +private + transportGroupEquiv : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → isEquiv (fst (πLES.A→B f n)) + → isEquiv (fst (π'∘∙Hom n f)) + transportGroupEquiv n f iseq = + transport (λ i → isEquiv (fst (π∘∙A→B-PathP n f i))) iseq + +π₃S²≅π₃TotalHopf : GroupEquiv (πGr 2 (Σ (S₊ 2) S¹Hopf , north , base)) + (πGr 2 (S₊∙ 2)) +fst (fst π₃S²≅π₃TotalHopf) = fst (πLES.A→B TotalHopf→∙S² 2) +snd (fst π₃S²≅π₃TotalHopf) = + SES→isEquiv + (isContr→≡UnitGroup + (subst isContr (cong (π 3) (sym IsoFiberTotalHopfS¹∙≡)) + (∣ refl ∣₂ , (sElim (λ _ → isSetPathImplicit) + (λ p → cong ∣_∣₂ + (isOfHLevelSuc 3 isGroupoidS¹ _ _ _ _ _ _ refl p)))))) + (isContr→≡UnitGroup + (subst isContr (cong (π 2) (sym IsoFiberTotalHopfS¹∙≡)) + (∣ refl ∣₂ , (sElim (λ _ → isSetPathImplicit) (λ p + → cong ∣_∣₂ (isGroupoidS¹ _ _ _ _ refl p)))))) + (πLES.fib→A TotalHopf→∙S² 2) + (πLES.A→B TotalHopf→∙S² 2) + (πLES.B→fib TotalHopf→∙S² 1) + (πLES.Ker-A→B⊂Im-fib→A TotalHopf→∙S² 2) + (πLES.Ker-B→fib⊂Im-A→B TotalHopf→∙S² 1) +snd π₃S²≅π₃TotalHopf = snd (πLES.A→B TotalHopf→∙S² 2) + +π'₃S²≅π'₃TotalHopf : GroupEquiv (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base)) + (π'Gr 2 (S₊∙ 2)) +fst (fst π'₃S²≅π'₃TotalHopf) = fst (π'∘∙Hom 2 TotalHopf→∙S²) +snd (fst π'₃S²≅π'₃TotalHopf) = + transportGroupEquiv 2 TotalHopf→∙S² (π₃S²≅π₃TotalHopf .fst .snd) +snd π'₃S²≅π'₃TotalHopf = snd (π'∘∙Hom 2 TotalHopf→∙S²) + +πS³≅πTotalHopf : (n : ℕ) + → GroupEquiv (π'Gr n (S₊∙ 3)) (π'Gr n (Σ (S₊ 2) S¹Hopf , north , base)) +πS³≅πTotalHopf n = + π'Iso n + ((isoToEquiv + (invIso + (compIso + (hopfS¹.IsoTotalSpaceJoin) + (IsoSphereJoin 1 1)))) + , refl) + +π₃S²≅ℤ : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤ +π₃S²≅ℤ = + compGroupEquiv + (invGroupEquiv + (compGroupEquiv (πS³≅πTotalHopf 2) π'₃S²≅π'₃TotalHopf)) + (GroupIso→GroupEquiv (πₙ'Sⁿ≅ℤ 2)) + +-- We prove that the generator is the Hopf map +π₃TotalHopf-gen' : π' 3 (Σ (Susp S¹) S¹Hopf , north , base) +π₃TotalHopf-gen' = + ∣ inv (compIso (hopfS¹.IsoTotalSpaceJoin) (IsoSphereJoin 1 1)) , refl ∣₂ + +πS³≅πTotalHopf-gen : + fst (fst (πS³≅πTotalHopf 2)) ∣ idfun∙ _ ∣₂ ≡ π₃TotalHopf-gen' +πS³≅πTotalHopf-gen = + cong ∣_∣₂ + (∘∙-idʳ (inv (compIso (hopfS¹.IsoTotalSpaceJoin) + (IsoSphereJoin 1 1)) , refl)) + +πTotalHopf-gen : + gen₁-by (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base)) + ∣ inv (compIso (hopfS¹.IsoTotalSpaceJoin) (IsoSphereJoin 1 1)) , refl ∣₂ +πTotalHopf-gen = + subst (gen₁-by (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base))) + πS³≅πTotalHopf-gen + (Iso-pres-gen₁ (π'Gr 2 (S₊∙ 3)) + (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base)) + ∣ idfun∙ _ ∣₂ + (πₙ'Sⁿ-gen-by-idfun 2) + (GroupEquiv→GroupIso (πS³≅πTotalHopf 2))) + +πTotalHopf≅πS²-gen : + fst (fst π'₃S²≅π'₃TotalHopf) + π₃TotalHopf-gen' + ≡ ∣ HopfMap' , refl ∣₂ +πTotalHopf≅πS²-gen = + cong ∣_∣₂ (ΣPathP (refl , (sym (rUnit refl)))) + +π₂S³-gen-by-HopfMap' : gen₁-by (π'Gr 2 (S₊∙ 2)) ∣ HopfMap' , refl ∣₂ +π₂S³-gen-by-HopfMap' = + subst (gen₁-by (π'Gr 2 (S₊∙ 2))) πTotalHopf≅πS²-gen + (Iso-pres-gen₁ (π'Gr 2 (Σ (S₊ 2) S¹Hopf , north , base)) (π'Gr 2 (S₊∙ 2)) + ∣ inv (compIso (hopfS¹.IsoTotalSpaceJoin) (IsoSphereJoin 1 1)) , refl ∣₂ + πTotalHopf-gen + (GroupEquiv→GroupIso π'₃S²≅π'₃TotalHopf)) + +π₂S³-gen-by-HopfMap : gen₁-by (π'Gr 2 (S₊∙ 2)) ∣ HopfMap ∣₂ +π₂S³-gen-by-HopfMap = + subst (gen₁-by (π'Gr 2 (S₊∙ 2))) + (cong ∣_∣₂ (sym hopfMap≡HopfMap')) + π₂S³-gen-by-HopfMap' + +-- As a consequence, we also get that the Hopf invariant determines +-- an iso π₃S²≅ℤ +hopfInvariantEquiv : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤ +fst (fst hopfInvariantEquiv) = HopfInvariant-π' 0 +snd (fst hopfInvariantEquiv) = + GroupEquivℤ-isEquiv (invGroupEquiv π₃S²≅ℤ) ∣ HopfMap ∣₂ + π₂S³-gen-by-HopfMap (GroupHom-HopfInvariant-π' 0) + (abs→⊎ _ _ HopfInvariant-HopfMap) +snd hopfInvariantEquiv = snd (GroupHom-HopfInvariant-π' 0) diff --git a/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda b/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda new file mode 100644 index 0000000000..b6b71c9ac7 --- /dev/null +++ b/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda @@ -0,0 +1,444 @@ +{- + +The goal of this file is to prove the iso π₄S³≅ℤ/β +where β is a natural number (aka "the Brunerie number", +defined below). + +-} +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Group.Pi4S3.BrunerieNumber where + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.Group.Pi3S2 +open import Cubical.Homotopy.Group.PinSn +open import Cubical.Homotopy.BlakersMassey +open import Cubical.Homotopy.Whitehead +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.LES +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Int + renaming (ℤ to Z ; _·_ to _·Z_ ; _+_ to _+Z_) + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.Pushout +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim ; map to pMap) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 + ; map to sMap) +open import Cubical.HITs.Truncation renaming + (rec to trRec ; elim to trElim ; elim2 to trElim2 ; map to trMap) + +open import Cubical.Algebra.Group + renaming (Unit to UnitGr) +open import Cubical.Algebra.Group.Exact +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Instances.IntMod + +open Iso +open Exact4 + +-- The Brunerie number (see Corollary 3.4.5 in Brunerie's PhD thesis) +Brunerie : ℕ +Brunerie = + abs (HopfInvariant-π' 0 [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π') + +-- First we need to define the following maps. +W : S₊ 3 → (S₊∙ 2 ⋁ S₊∙ 2) +W = joinTo⋁ {A = S₊∙ 1} {B = S₊∙ 1} ∘ Iso.inv (IsoSphereJoin 1 1) + +fold∘W : S₊ 3 → S₊ 2 +fold∘W = fold⋁ ∘ W + +{- +We will now instantiate Blakers Massey to get a square: + + fold∘W + S³ --------------> S² + |\ ↗ | + | \ ↗ | + | \ ↗ | + | X | inr + | / | + | / | + | / | + v / v + 1 -----------> coFib fold∘W + inl + +where X is the pullback of inl and inr and the map S³ → X is +surjective on π₃. This will give us a sequence +π₃ S³ --ᶠ→ π₃ S² → π₃ (coFib fold∘W) → π₂ X ≅ 0, where f is +the function induced by fold∘W. From this, we can deduce that +π₃ (coFib fold∘W) ≅ ℤ/ f(1) where f(1) is interpreted as an +integer via the isos π₃ S³ ≅ π₃ S² ≅ ℤ + +(Recall that π₃(coFib fold∘W) ≅ π₄S³) + +For clarity: +X, above will have two names (via a trivial iso) below: + +TotalPushoutPath× (the version the falls out of BM) +P = fiber inr (Same name as in Brunerie's prop 4.3.4.) +-} + + +-- Before instantiating, we need to show that +-- any map S³ → S² is 0-connected +isConnectedS3→S2 : (f : S₊ 3 → S₊ 2) → isConnectedFun 2 f +isConnectedS3→S2 f p = + trRec (isProp→isOfHLevelSuc 1 isPropIsContr) + (J (λ p _ → isConnected 2 (fiber f p)) + (∣ north , refl ∣ + , (trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (uncurry + (sphereElim 2 + (λ _ → isOfHLevelΠ 3 + λ _ → isOfHLevelPath 3 + (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) + λ p → trRec (isOfHLevelTrunc 2 _ _) + (λ r → cong ∣_∣ₕ (ΣPathP (refl , r))) + (fun (PathIdTruncIso 1) + (isContr→isProp + (isConnectedPath 2 (sphereConnected 2) + (f north) (f north)) ∣ refl ∣ ∣ p ∣))))))) + (fun (PathIdTruncIso 2) + (isContr→isProp (sphereConnected 2) ∣ f north ∣ ∣ p ∣)) + +-- We get our square +module BM-inst = + BlakersMassey□ + (λ _ → tt) fold∘W 3 1 + (λ _ → subst (isConnected 4) + (isoToPath (invIso fiberUnitIso)) + (sphereConnected 3)) + (isConnectedS3→S2 fold∘W) + +open BM-inst + +-- The central types +coFib-fold∘W : Type +coFib-fold∘W = Pushout (λ _ → tt) fold∘W + +coFib-fold∘W∙ : Pointed₀ +coFib-fold∘W∙ = coFib-fold∘W , inl tt + +TotalPushoutPath×∙ : Pointed ℓ-zero +fst TotalPushoutPath×∙ = Σ (Unit × S₊ 2) PushoutPath× +snd TotalPushoutPath×∙ = (tt , north) , push north + +S³→TotalPushoutPath× : S₊ 3 → Σ (Unit × S₊ 2) PushoutPath× +S³→TotalPushoutPath× = toPullback + +private + inr' : S₊ 2 → coFib-fold∘W + inr' = inr + + inr∙ : S₊∙ 2 →∙ coFib-fold∘W∙ + fst inr∙ = inr' + snd inr∙ = sym (push north) + + fiberinr'Iso' : Iso (fiber inr' (inl tt)) + (Σ (Unit × S₊ 2) PushoutPath×) + fiberinr'Iso' = + compIso (Σ-cong-iso-snd (λ x → symIso)) + (Σ-cong-iso-fst (invIso lUnit×Iso)) + + fiberinr'Iso : Iso (fiber inr' (inl tt)) + (Σ (Unit × S₊ 2) PushoutPath×) + fun fiberinr'Iso (x , p) = (tt , x) , (sym p) + inv fiberinr'Iso ((tt , x) , p) = x , (sym p) + rightInv fiberinr'Iso _ = refl + leftInv fiberinr'Iso _ = refl + + P : Pointed₀ + P = (fiber inr' (inl tt) , north , (sym (push north))) + +π'P→π'TotalPath× : (n : ℕ) + → GroupEquiv (π'Gr n TotalPushoutPath×∙) (π'Gr n P) +fst (fst (π'P→π'TotalPath× n)) = + π'eqFun n ((invEquiv (isoToEquiv fiberinr'Iso)) , refl) +snd (fst (π'P→π'TotalPath× n)) = π'eqFunIsEquiv n _ +snd (π'P→π'TotalPath× n) = π'eqFunIsHom n _ + +-- Time to invoke the long exact sequence of homotopy groups on +-- inr : S² → coFib-fold∘W +module LESinst = πLES {A = S₊∙ 2} {B = coFib-fold∘W∙} inr∙ + +-- We instantiate the sequence +-- π₃ P → π₃ S² → π₃ coFib-fold∘W∙ → π₂ P +P→S²→Pushout : + Exact4 (πGr 2 P) (πGr 2 (S₊∙ 2)) (πGr 2 coFib-fold∘W∙) (πGr 1 P) + (LESinst.fib→A 2) + (LESinst.A→B 2) + (LESinst.B→fib 1) +Exact4.ImG→H⊂KerH→L P→S²→Pushout = LESinst.Im-fib→A⊂Ker-A→B 2 +Exact4.KerH→L⊂ImG→H P→S²→Pushout = LESinst.Ker-A→B⊂Im-fib→A 2 +Exact4.ImH→L⊂KerL→R P→S²→Pushout = LESinst.Im-A→B⊂Ker-B→fib 1 +Exact4.KerL→R⊂ImH→L P→S²→Pushout = LESinst.Ker-B→fib⊂Im-A→B 1 + +-- The goal now is to rewrite it as +-- π₃ S³ → π₃ S² → π₃ coFib-fold∘W∙ → Unit using the +-- "functions from spheres"-definition of πₙ. +-- Here, the first map is the one induced by fold∘W. We do this by: +-- (1) showing that π₂ P is trivial +-- (2) extending the sequence by appending surjections +-- π₃ S³ → π₃ TotalPushoutPath×∙ → π₃ P on the left. +-- (3) proving that this new composition is indeed the appropriate map + +-- Step 1: π₂ P is trivial +π₂P≅0 : GroupEquiv (πGr 1 P) UnitGr +π₂P≅0 = compGroupEquiv (πIso (isoToEquiv fiberinr'Iso , refl) 1) + (GroupIso→GroupEquiv + (contrGroupIsoUnit + (isOfHLevelRetractFromIso 0 (invIso iso₂) isContrπ₂S³))) + where + iso₁ : Iso (hLevelTrunc 4 (S₊ 3)) + (hLevelTrunc 4 (Σ (Unit × S₊ 2) PushoutPath×)) + iso₁ = connectedTruncIso 4 S³→TotalPushoutPath× isConnected-toPullback + + iso₂ : Iso (π 2 (hLevelTrunc∙ 4 (S₊∙ 3))) + (π 2 TotalPushoutPath×∙) + iso₂ = + (compIso (setTruncIso + (equivToIso (_ , + (isEquivΩ^→ 2 (fun iso₁ , refl) (isoToIsEquiv iso₁))))) + (invIso (πTruncIso 2))) + + isContrπ₂S³ : isContr (π 2 (hLevelTrunc∙ 4 (S₊∙ 3))) + isContrπ₂S³ = + subst (λ x → isContr (π 2 x)) + (λ i → ((sym (isContr→≡Unit (sphereConnected 3))) i) + , transp (λ j → isContr→≡Unit + (sphereConnected 3) (~ i ∧ j)) i ∣ north ∣) + (∣ refl ∣₂ , sElim (λ _ → isSetPathImplicit) + λ p → cong ∣_∣₂ + (isProp→isSet + (isOfHLevelPath 1 isPropUnit _ _) _ _ _ p)) + +-- Step 2. We transform our original sequence to one for the +-- the "maps from spheres" definition of πₙ and where π₂ P is +-- replaced by the trivial group: +-- π₃ P → π₃ S² → π₃ coFib-fold∘W∙ → 0 +P→S²→Pushout→P' : + Exact4 (π'Gr 2 P) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 coFib-fold∘W∙) UnitGr + (π'∘∙Hom 2 (fst , refl)) + (π'∘∙Hom 2 inr∙) + (→UnitHom _) +P→S²→Pushout→P' = + transportExact4 + (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 2 P))))) + (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 2 (S₊∙ 2)))))) + (sym (GroupPath _ _ .fst ((GroupIso→GroupEquiv (π'Gr≅πGr 2 coFib-fold∘W∙))))) + (sym (GroupPath _ _ .fst π₂P≅0)) + _ _ _ _ _ + P→S²→Pushout + (ΣPathPProp (λ _ → isPropIsGroupHom _ _) + λ i → fst (π∘∙fib→A-PathP 2 inr∙ i)) + (ΣPathPProp (λ _ → isPropIsGroupHom _ _) + λ i → fst (π∘∙A→B-PathP 2 inr∙ i)) + +-- The two surjections in question +π₃S³→π₃P : GroupHom (π'Gr 2 (S₊∙ 3)) (π'Gr 2 TotalPushoutPath×∙) +π₃S³→π₃P = π'∘∙Hom 2 (S³→TotalPushoutPath× , refl) + +TotalPushoutPath×∙→P : TotalPushoutPath×∙ →∙ P -- Surjective, and in particular on π₃ +fst TotalPushoutPath×∙→P x = (snd (fst x)) , (sym (snd x)) +snd TotalPushoutPath×∙→P = refl + +-- This surjectivity is where Blakers-Massey is used +-- In particular, it uses isConnected-toPullback +isSurjective-π₃S³→π₃TotalPushoutPath× : isSurjective π₃S³→π₃P +isSurjective-π₃S³→π₃TotalPushoutPath× = + transport (λ i → isSurjective (transportLem i)) + isSurjective-π₃S³→π₃TotalPushoutPath×' + where + π₃S³→π₃TotalPushoutPath× : GroupHom (πGr 2 (S₊∙ 3)) (πGr 2 TotalPushoutPath×∙) + π₃S³→π₃TotalPushoutPath× = πHom 2 (S³→TotalPushoutPath× , refl) + + isSurjective-π₃S³→π₃TotalPushoutPath×' : isSurjective π₃S³→π₃TotalPushoutPath× + isSurjective-π₃S³→π₃TotalPushoutPath×' = + sElim (λ _ → isProp→isSet squash) + λ p → trRec squash + (λ s → ∣ ∣ fst s ∣₂ , (cong ∣_∣₂ (snd s)) ∣) + (((isConnectedΩ^→ 3 3 (S³→TotalPushoutPath× , refl) + isConnected-toPullback) p) .fst) + + transportLem : PathP (λ i → GroupHomπ≅π'PathP (S₊∙ 3) TotalPushoutPath×∙ 2 i) + π₃S³→π₃TotalPushoutPath× π₃S³→π₃P + transportLem = + toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (π'∘∙Hom'≡π'∘∙fun {A = S₊∙ 3} {B = TotalPushoutPath×∙} + 2 (S³→TotalPushoutPath× , refl))) + +-- We get a sequence on the right form π₃S³ → π₃S² → π₃ Pushout → Unit +S³→S²→Pushout→Unit'' : + Exact4 (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 coFib-fold∘W∙) UnitGr + (compGroupHom π₃S³→π₃P + (compGroupHom + (π'∘∙Hom 2 TotalPushoutPath×∙→P) (π'∘∙Hom 2 (fst , refl)))) + (π'∘∙Hom 2 inr∙) + (→UnitHom (π'Gr 2 coFib-fold∘W∙)) +S³→S²→Pushout→Unit'' = + extendExact4Surjective _ _ _ _ _ _ _ _ _ + isSurjective-π₃S³→π₃TotalPushoutPath× + (extendExact4Surjective _ _ _ _ _ _ _ _ _ + ((sElim (λ _ → isProp→isSet squash) + (λ f → ∣ ∣ (λ x → (tt , fst f x .fst) , sym (fst f x .snd)) + , ΣPathP ((ΣPathP (refl , cong fst (snd f))) + , λ j i → snd f j .snd (~ i)) ∣₂ + , cong ∣_∣₂ (ΣPathP (refl , sym (rUnit _))) ∣))) + P→S²→Pushout→P') + +-- Step 3: We need to show that the map π₃S³ → π₃S² in the above sequence +-- indeed comes from fold∘W +tripleComp≡ : + (compGroupHom π₃S³→π₃P + (compGroupHom + (π'∘∙Hom 2 TotalPushoutPath×∙→P) (π'∘∙Hom 2 (fst , refl)))) + ≡ π'∘∙Hom 2 (fold∘W , refl) +tripleComp≡ = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (ΣPathP (refl , (cong (_∙ refl) + (λ j → cong fst (rUnit (cong (fst TotalPushoutPath×∙→P) + (rUnit (cong S³→TotalPushoutPath× (snd f)) (~ j))) (~ j)))))))) + +-- We finally get the correct sequence +S³→S²→Pushout→Unit : + Exact4 (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 coFib-fold∘W∙) UnitGr + (π'∘∙Hom 2 (fold∘W , refl)) + (π'∘∙Hom 2 inr∙) + (→UnitHom (π'Gr 2 coFib-fold∘W∙)) +S³→S²→Pushout→Unit = + subst + (λ F → Exact4 (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 coFib-fold∘W∙) UnitGr + F (π'∘∙Hom 2 inr∙) + (→UnitHom (π'Gr 2 coFib-fold∘W∙))) + tripleComp≡ + S³→S²→Pushout→Unit'' + +-- We need to throw around some pushouts +module _ where + Pushout-coFibW-fold⋁≃coFib-fold∘W : + Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ ≃ fst coFib-fold∘W∙ + Pushout-coFibW-fold⋁≃coFib-fold∘W = compEquiv + (compEquiv pushoutSwitchEquiv + (isoToEquiv (PushoutDistr.PushoutDistrIso fold⋁ W λ _ → tt))) + pushoutSwitchEquiv + + coFibW≅coFibW' : Pushout W (λ _ → tt) ≃ cofibW S¹ S¹ base base + coFibW≅coFibW' = pushoutEquiv W (λ _ → tt) joinTo⋁ (λ _ → tt) + (isoToEquiv (invIso (IsoSphereJoin 1 1))) + (idEquiv _) + (idEquiv _) + refl + refl + + Pushout-coFibW-fold⋁≃Pushout⋁↪fold⋁ : + Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ + ≃ fst (Pushout⋁↪fold⋁∙ (S₊∙ 2)) + Pushout-coFibW-fold⋁≃Pushout⋁↪fold⋁ = pushoutEquiv inl _ ⋁↪ fold⋁ + (idEquiv _) + (compEquiv coFibW≅coFibW' + (isoToEquiv (invIso (Iso-Susp×Susp-cofibJoinTo⋁ S¹ S¹ base base)))) + (idEquiv _) + (Susp×Susp→cofibW≡ S¹ S¹ base base) + refl + + Pushout-coFibW-fold⋁≃Pushout⋁↪fold⋁∙ : + (Pushout {B = (Pushout W (λ _ → tt))} inl fold⋁ , inr north) + ≃∙ (Pushout⋁↪fold⋁∙ (S₊∙ 2)) + fst Pushout-coFibW-fold⋁≃Pushout⋁↪fold⋁∙ = + Pushout-coFibW-fold⋁≃Pushout⋁↪fold⋁ + snd Pushout-coFibW-fold⋁≃Pushout⋁↪fold⋁∙ = + sym (push (inl north)) + +π₄S³≅π₃coFib-fold∘W∙ : GroupEquiv (π'Gr 3 (S₊∙ 3)) (π'Gr 2 coFib-fold∘W∙) +π₄S³≅π₃coFib-fold∘W∙ = + compGroupEquiv + (GroupIso→GroupEquiv + (compGroupIso + (π'Gr≅πGr 3 (S₊∙ 3)) + (compGroupIso + π₄S³≅π₃PushS² + (invGroupIso (π'Gr≅πGr 2 (Pushout⋁↪fold⋁∙ (S₊∙ 2))))))) + (compGroupEquiv + (invGroupEquiv (π'Iso 2 Pushout-coFibW-fold⋁≃Pushout⋁↪fold⋁∙)) + (π'Iso 2 (Pushout-coFibW-fold⋁≃coFib-fold∘W , sym (push north)))) + +-- We get the iso +-- For type checking reasons, let's first prove it for the abstract +-- definition of ℤ/_ + +-- To get everything on the same form as in Brunerie's thesis, we +-- first need the following: +fold∘W≡Whitehead : + fst (π'∘∙Hom 2 (fold∘W , refl)) ∣ idfun∙ (S₊∙ 3) ∣₂ + ≡ ∣ [ idfun∙ (S₊∙ 2) ∣ idfun∙ (S₊∙ 2) ]₂ ∣₂ +fold∘W≡Whitehead = + pRec (squash₂ _ _) + (cong ∣_∣₂) + (indΠ₃S₂ _ _ + (funExt (λ x → funExt⁻ (sym (cong fst (id∨→∙id {A = S₊∙ 2}))) (W x)))) + where + indΠ₃S₂ : ∀ {ℓ} {A : Pointed ℓ} + → (f g : A →∙ S₊∙ 2) + → fst f ≡ fst g → ∥ f ≡ g ∥ + indΠ₃S₂ {A = A} f g p = + trRec squash + (λ r → ∣ ΣPathP (p , r) ∣) + (isConnectedPathP 1 {A = (λ i → p i (snd A) ≡ north)} + (isConnectedPathSⁿ 1 (fst g (pt A)) north) (snd f) (snd g) .fst ) + +BrunerieIsoAbstract : GroupEquiv (π'Gr 3 (S₊∙ 3)) (abstractℤ/ Brunerie) +BrunerieIsoAbstract = + compGroupEquiv π₄S³≅π₃coFib-fold∘W∙ + (invGroupEquiv + (GroupEquiv-abstractℤ/abs-gen + (π'Gr 2 (S₊∙ 3)) (π'Gr 2 (S₊∙ 2)) (π'Gr 2 coFib-fold∘W∙) + (GroupIso→GroupEquiv (invGroupIso (πₙ'Sⁿ≅ℤ 2))) + (invGroupEquiv hopfInvariantEquiv) + (π'∘∙Hom 2 (fold∘W , refl)) + _ + S³→S²→Pushout→Unit Brunerie main)) + where + mainPath : + fst (π'∘∙Hom 2 (fold∘W , refl)) + (Iso.inv (fst (πₙ'Sⁿ≅ℤ 2)) 1) + ≡ [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π' + mainPath = + cong (fst (π'∘∙Hom 2 (fold∘W , refl))) + (cong (Iso.inv (fst (πₙ'Sⁿ≅ℤ 2))) (sym (πₙ'Sⁿ≅ℤ-idfun∙ 2)) + ∙ (Iso.leftInv (fst (πₙ'Sⁿ≅ℤ 2)) ∣ idfun∙ (S₊∙ 3) ∣₂)) + ∙ fold∘W≡Whitehead + ∙ cong ∣_∣₂ (sym ([]≡[]₂ (idfun∙ (S₊∙ 2)) (idfun∙ (S₊∙ 2)))) + + main : _ ≡ Brunerie + main i = abs (HopfInvariant-π' 0 (mainPath i)) + +-- And, finally, we get the actual iso +-- (as in Corollary 3.4.5 in Brunerie's thesis) +BrunerieIso : GroupEquiv (π'Gr 3 (S₊∙ 3)) (ℤ/ Brunerie) +BrunerieIso = + compGroupEquiv BrunerieIsoAbstract (abstractℤ/≅ℤ Brunerie) diff --git a/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso.agda b/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso.agda new file mode 100644 index 0000000000..fca96aa611 --- /dev/null +++ b/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso.agda @@ -0,0 +1,682 @@ +{-# OPTIONS --safe #-} +{- +This file has been split in two due to slow type checking +combined with insufficient reductions when the +experimental-lossy-unification flag is included. +Part 2: Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 + +The goal of these two files is to show that +π₄(S³) ≅ π₃((S² × S²) ⊔ᴬ S²) where A = S² ∨ S². +This is proved in Brunerie (2016) using the James construction and +via (S² × S²) ⊔ᴬ S² ≃ J₂(S²). + +In this file, we prove it directly using the encode-decode method. For +the statement of the isomorphism, see part 2. +-} +module Cubical.Homotopy.Group.Pi4S3.S3PushoutIso where + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Homotopy.Group.Base +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open Iso +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Morphism + +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 + ; map to sMap) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.S1 hiding (decode ; encode) + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Bool + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + + +open import Cubical.HITs.Join +open import Cubical.HITs.Pushout +open import Cubical.HITs.Wedge +open import Cubical.Homotopy.Freudenthal hiding (Code ; encode) +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation renaming + (rec to trRec ; elim to trElim ; elim2 to trElim2 ; map to trMap) +open import Cubical.Foundations.Function +open import Cubical.HITs.S2 + +Pushout⋁↪fold⋁ : ∀ {ℓ} (A : Pointed ℓ) → Type ℓ +Pushout⋁↪fold⋁ A = Pushout {A = A ⋁ A} ⋁↪ fold⋁ + +Pushout⋁↪fold⋁∙ : ∀ {ℓ} (A : Pointed ℓ) → Pointed ℓ +fst (Pushout⋁↪fold⋁∙ A) = Pushout⋁↪fold⋁ A +snd (Pushout⋁↪fold⋁∙ A) = inl (pt A , pt A) + +-- The type of interest -- ''almost`` equivalent to ΩS³ +Pushout⋁↪fold⋁S₊2 = Pushout {A = S₊∙ 2 ⋁ S₊∙ 2} ⋁↪ fold⋁ +-- Same type, using base/surf definition of S² (easier to work with) +Pushout⋁↪fold⋁S² = Pushout⋁↪fold⋁ S²∙ +-- Truncated version, to be proved equivalent to ∥ ΩS³ ∥₅ +∥Pushout⋁↪fold⋁S²∥₅ = hLevelTrunc 5 Pushout⋁↪fold⋁S² + +Ω∥S³∥₆ = Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ north ∣ + +-- Elimination into sets for Pushout⋁↪fold⋁S² +module _ {ℓ : Level} {P : Pushout⋁↪fold⋁S² → Type ℓ} + (hlev : ((x : Pushout⋁↪fold⋁S²) → isSet (P x))) + (b : P (inl (base , base))) where + private + ×fun : (x y : S²) → P (inl (x , y)) + ×fun = S²ToSetElim (λ _ → isSetΠ (λ _ → hlev _)) + (S²ToSetElim (λ _ → hlev _) b) + + inrfun : (x : S²) → P (inr x) + inrfun = S²ToSetElim (λ _ → hlev _) (subst P (push (inl base)) b) + + inl-pp : (x : S²) → PathP (λ i → P (push (inl x) i)) (×fun x base) (inrfun x) + inl-pp = S²ToSetElim (λ _ → isOfHLevelPathP 2 (hlev _) _ _) + λ j → transp (λ i → P (push (inl base) (i ∧ j))) (~ j) b + + inr-pp : (x : S²) → PathP (λ i → P (push (inr x) i)) (×fun base x) (inrfun x) + inr-pp = S²ToSetElim (λ _ → isOfHLevelPathP 2 (hlev _) _ _) + (transport (λ j → PathP (λ i → P (push (push tt j) i)) (×fun base base) + (inrfun base)) + (inl-pp base)) + + Pushout⋁↪fold⋁S²→Set : (x : Pushout⋁↪fold⋁S²) → P x + Pushout⋁↪fold⋁S²→Set (inl x) = ×fun (fst x) (snd x) + Pushout⋁↪fold⋁S²→Set (inr x) = inrfun x + Pushout⋁↪fold⋁S²→Set (push (inl x) i) = inl-pp x i + Pushout⋁↪fold⋁S²→Set (push (inr x) i) = inr-pp x i + Pushout⋁↪fold⋁S²→Set (push (push a j) i) = help j i + where + help : SquareP (λ j i → P (push (push tt j) i)) + (λ i → inl-pp base i) + (λ i → inr-pp base i) + (λ _ → ×fun base base) + (λ _ → inrfun base) + help = toPathP (isProp→PathP (λ _ → isOfHLevelPathP' 1 (hlev _) _ _) _ _) + +{- A wedge connectivity lemma for Pushout⋁↪fold⋁S². It can be +stated for general dependent types, but it's easier to work with +in the special case of path types -} +module Pushout⋁↪fold⋁S²WedgeCon {ℓ : Level } {A : Type ℓ} + (f g : Pushout⋁↪fold⋁S² → A) + (h : isOfHLevel 5 A) + (lp : (x : S²) → f (inl (x , base)) ≡ g (inl (x , base))) + (rp : (x : S²) → f (inl (base , x)) ≡ g (inl (base , x))) + (r≡l : (x : S²) + → (sym (cong f (push (inl x))) ∙∙ lp x ∙∙ cong g (push (inl x))) + ≡ (sym (cong f (push (inr x))) ∙∙ rp x ∙∙ cong g (push (inr x)))) + where + btm-filler : I → I → I → A + btm-filler k i j = + hfill (λ k → λ {(i = i0) + → doubleCompPath-filler (sym (cong f (push (inl base)))) + (lp base) + (cong g (push (inl base))) (~ k) j + ; (i = i1) + → doubleCompPath-filler (sym (cong f (push (inr base)))) + (rp base) + (cong g (push (inr base))) (~ k) j + ; (j = i0) → f (push (push tt i) (~ k)) + ; (j = i1) → g (push (push tt i) (~ k))}) + (inS (r≡l base i j)) + k + + lp-base≡rp-base : lp base ≡ rp base + lp-base≡rp-base = (λ i j → btm-filler i1 i j) + + f∘inl≡g∘inl : (x y : S²) → f (inl (x , y)) ≡ g (inl (x , y)) + f∘inl≡g∘inl = wedgeconFunS² (λ _ _ → h _ _) lp rp lp-base≡rp-base + + f∘inl≡g∘inl-base : (x : S²) → f∘inl≡g∘inl x base ≡ lp x + f∘inl≡g∘inl-base = wedgeconFunS²Id (λ _ _ → h _ _) lp rp lp-base≡rp-base + + f∘inr≡g∘inr : (x : S²) → f (inr x) ≡ g (inr x) + f∘inr≡g∘inr x = cong f (sym (push (inl x))) ∙∙ lp x ∙∙ cong g (push (inl x)) + + inlfill : (x : S²) → I → I → I → A + inlfill x k i j = + hfill (λ k → λ { (i = i0) → f∘inl≡g∘inl-base x (~ k) j + ; (i = i1) → doubleCompPath-filler + (cong f (sym (push (inl x)))) + (lp x) + (cong g (push (inl x))) k j + ; (j = i0) → f (push (inl x) (i ∧ k)) + ; (j = i1) → g (push (inl x) (i ∧ k))}) + (inS (lp x j)) + k + + inrfill : (x : S²) → I → I → I → A + inrfill x k i j = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler + (cong f (sym (push (inr x)))) + (rp x) + (cong g (push (inr x))) (~ k) j + ; (i = i1) → f∘inr≡g∘inr x j + ; (j = i0) → f (push (inr x) (i ∨ ~ k)) + ; (j = i1) → g (push (inr x) (i ∨ ~ k))}) + (inS (r≡l x (~ i) j)) + k + + main : (x : Pushout⋁↪fold⋁S²) → f x ≡ g x + main (inl x) = f∘inl≡g∘inl (fst x) (snd x) + main (inr x) = f∘inr≡g∘inr x + main (push (inl x) i) j = inlfill x i1 i j + main (push (inr x) i) j = inrfill x i1 i j + main (push (push a i) j) k = + hcomp (λ r → λ {(i = i0) → inlfill base i1 j k + ; (i = i1) → inrfill-side r j k + ; (j = i0) → rp base k + ; (j = i1) → r≡l base (~ r ∧ i) k + ; (k = i0) → f (push (push a i) j) + ; (k = i1) → g (push (push a i) j)}) + (hcomp (λ r → λ {(i = i0) → inlfill base r j k + ; (i = i1) → doubleCompPath-filler + (cong f (sym (push (inr base)))) + (rp base) + (cong g (push (inr base))) (j ∧ r) k + ; (j = i0) → lp-base≡rp-base (r ∨ i) k + ; (j = i1) → inlfill-side r i k + ; (k = i0) → f (push (push a i) (j ∧ r)) + ; (k = i1) → g (push (push a i) (j ∧ r))}) + (lp-base≡rp-base i k)) + where + inlfill-side : I → I → I → A + inlfill-side r i k = + hcomp (λ j → λ { (r = i0) → btm-filler j i k + ; (r = i1) → r≡l base i k + ; (i = i0) → doubleCompPath-filler + (cong f (sym (push (inl base)))) + (lp base) + (cong g (push (inl base))) (r ∨ ~ j) k + ; (i = i1) → doubleCompPath-filler + (cong f (sym (push (inr base)))) + (rp base) + (cong g (push (inr base))) (r ∨ ~ j) k + ; (k = i0) → f (push (push a i) (r ∨ ~ j)) + ; (k = i1) → g (push (push a i) (r ∨ ~ j))}) + (r≡l base i k) + + inrfill-side : I → I → I → A + inrfill-side r j k = + hcomp (λ i → λ { (r = i0) + → (doubleCompPath-filler + (cong f (sym (push (inr base)))) + (rp base) + (cong g (push (inr base)))) (j ∨ ~ i) k + ; (r = i1) → inrfill base i j k + ; (j = i0) → inrfill base i i0 k + ; (j = i1) → r≡l base (~ r) k + ; (k = i0) → f (push (inr base) (j ∨ ~ i)) + ; (k = i1) → g (push (inr base) (j ∨ ~ i))}) + (r≡l base (~ r ∨ ~ j) k) + +{- +Goal: Prove Ω ∥ SuspS² ∥₆ ≃ ∥ Pushout⋁↪fold⋁S² ∥₅ +This will by done via the encode-decode method. For this, we nede a family +of equivalences ∥ Pushout⋁↪fold⋁S² ∥₅ ≃ ∥ Pushout⋁↪fold⋁S² ∥₅, indexed by S². +In order to define this, we need the following cubes/coherences in +∥ Pushout⋁↪fold⋁S² ∥₅: +-} + +→Ω²∥Pushout⋁↪fold⋁S²∥₅ : (x y : S²) + → Square {A = ∥Pushout⋁↪fold⋁S²∥₅} (λ _ → ∣ inl (x , y) ∣ₕ) + (λ _ → ∣ inl (x , y) ∣ₕ) + (λ _ → ∣ inl (x , y) ∣ₕ) + (λ _ → ∣ inl (x , y) ∣ₕ) +→Ω²∥Pushout⋁↪fold⋁S²∥₅ = + wedgeconFunS² + (λ _ _ → isOfHLevelPath 4 (isOfHLevelTrunc 5 _ _) _ _) + (λ x → λ i j → ∣ inl (x , surf i j) ∣ₕ) + (λ x → λ i j → ∣ inl (surf i j , x) ∣ₕ) + λ r i j → ∣ hcomp (λ k → λ { (i = i0) → push (push tt (~ r)) (~ k) + ; (i = i1) → push (push tt (~ r)) (~ k) + ; (j = i1) → push (push tt (~ r)) (~ k) + ; (j = i0) → push (push tt (~ r)) (~ k) + ; (r = i0) → push (inr (surf i j)) (~ k) + ; (r = i1) → push (inl (surf i j)) (~ k)}) + (inr (surf i j)) ∣ₕ + +→Ω²∥Pushout⋁↪fold⋁S²∥₅Id : (x : S²) + → →Ω²∥Pushout⋁↪fold⋁S²∥₅ x base ≡ λ i j → ∣ inl (x , surf i j) ∣ₕ +→Ω²∥Pushout⋁↪fold⋁S²∥₅Id = + wedgeconFunS²Id + (λ _ _ → isOfHLevelPath 4 (isOfHLevelTrunc 5 _ _) _ _) + (λ x → λ i j → ∣ inl (x , surf i j) ∣ₕ) + (λ x → λ i j → ∣ inl (surf i j , x) ∣ₕ) + λ r i j → ∣ hcomp (λ k → λ { (i = i0) → push (push tt (~ r)) (~ k) + ; (i = i1) → push (push tt (~ r)) (~ k) + ; (j = i1) → push (push tt (~ r)) (~ k) + ; (j = i0) → push (push tt (~ r)) (~ k) + ; (r = i0) → push (inr (surf i j)) (~ k) + ; (r = i1) → push (inl (surf i j)) (~ k)}) + (inr (surf i j)) ∣ₕ + +push-inl∙push-inr⁻ : (y : S²) → Path Pushout⋁↪fold⋁S² (inl (y , base)) (inl (base , y)) +push-inl∙push-inr⁻ y i = (push (inl y) ∙ sym (push (inr y))) i + +push-inl∙push-inr⁻∙ : push-inl∙push-inr⁻ base ≡ refl +push-inl∙push-inr⁻∙ = + (λ k i → (push (inl base) ∙ sym (push (push tt (~ k)))) i) + ∙ rCancel (push (inl base)) + +push-inl∙push-inr⁻-filler : I → I → I → I → Pushout⋁↪fold⋁S² +push-inl∙push-inr⁻-filler r i j k = + hfill (λ r → λ { (i = i0) → push-inl∙push-inr⁻∙ (~ r) k + ; (i = i1) → push-inl∙push-inr⁻∙ (~ r) k + ; (j = i0) → push-inl∙push-inr⁻∙ (~ r) k + ; (j = i1) → push-inl∙push-inr⁻∙ (~ r) k + ; (k = i0) → inl (surf i j , base) + ; (k = i1) → inl (surf i j , base)}) + (inS (inl (surf i j , base))) + r + +push-inl∙push-inr⁻-hLevFiller : (y : S²) + → Cube {A = ∥Pushout⋁↪fold⋁S²∥₅} + (λ j k → ∣ push-inl∙push-inr⁻ y k ∣) + (λ j k → ∣ push-inl∙push-inr⁻ y k ∣) + (λ j k → ∣ push-inl∙push-inr⁻ y k ∣) + (λ j k → ∣ push-inl∙push-inr⁻ y k ∣) + (→Ω²∥Pushout⋁↪fold⋁S²∥₅ y (pt (S² , base))) + λ i j → ∣ inl (surf i j , y) ∣ +push-inl∙push-inr⁻-hLevFiller = + S²ToSetElim (λ _ → isOfHLevelPathP' 2 + (isOfHLevelPathP' 3 + (isOfHLevelPathP' 4 (isOfHLevelTrunc 5) _ _) _ _) _ _) + λ i j k → ∣ push-inl∙push-inr⁻-filler i1 i j k ∣ₕ + + +{- +We combine the above definitions. The equivalence +∥Pushout⋁↪fold⋁S²∥₅ ≃ ∥Pushout⋁↪fold⋁S²∥₅ (w.r.t. S²) is induced +by the following function +-} +S²→Pushout⋁↪fold⋁S²↺' : (x : S²) → Pushout⋁↪fold⋁S² → ∥Pushout⋁↪fold⋁S²∥₅ +S²→Pushout⋁↪fold⋁S²↺' base (inl (x , y)) = ∣ inl (x , y) ∣ +S²→Pushout⋁↪fold⋁S²↺' (surf i j) (inl (x , y)) = →Ω²∥Pushout⋁↪fold⋁S²∥₅ x y i j +S²→Pushout⋁↪fold⋁S²↺' base (inr z) = ∣ inl (base , z) ∣ₕ +S²→Pushout⋁↪fold⋁S²↺' (surf i i₁) (inr z) = ∣ inl (surf i i₁ , z) ∣ₕ +S²→Pushout⋁↪fold⋁S²↺' base (push (inl y) k) = + ∣ (push (inl y) ∙ sym (push (inr y))) k ∣ +S²→Pushout⋁↪fold⋁S²↺' (surf i j) (push (inl y) k) = + push-inl∙push-inr⁻-hLevFiller y i j k +S²→Pushout⋁↪fold⋁S²↺' base (push (inr y) i) = ∣ inl (base , y) ∣ +S²→Pushout⋁↪fold⋁S²↺' (surf i j) (push (inr y) k) = ∣ inl (surf i j , y) ∣ +S²→Pushout⋁↪fold⋁S²↺' base (push (push a i₁) i) = ∣ push-inl∙push-inr⁻∙ i₁ i ∣ +S²→Pushout⋁↪fold⋁S²↺' (surf i j) (push (push a k) l) = + ∣ hcomp (λ r → λ { (i = i0) → push-inl∙push-inr⁻∙ (k ∨ ~ r) l + ; (i = i1) → push-inl∙push-inr⁻∙ (k ∨ ~ r) l + ; (j = i0) → push-inl∙push-inr⁻∙ (k ∨ ~ r) l + ; (j = i1) → push-inl∙push-inr⁻∙ (k ∨ ~ r) l + ; (k = i0) → push-inl∙push-inr⁻-filler r i j l + ; (k = i1) → inl (surf i j , base) + ; (l = i0) → inl (surf i j , base) + ; (l = i1) → inl (surf i j , base)}) + (inl (surf i j , base)) ∣ₕ + +{- For easier treatment later, we state its inverse explicitly -} +S²→Pushout⋁↪fold⋁S²↺'⁻ : (x : S²) → Pushout⋁↪fold⋁S² → ∥Pushout⋁↪fold⋁S²∥₅ +S²→Pushout⋁↪fold⋁S²↺'⁻ base y = S²→Pushout⋁↪fold⋁S²↺' base y +S²→Pushout⋁↪fold⋁S²↺'⁻ (surf i j) y = S²→Pushout⋁↪fold⋁S²↺' (surf (~ i) j) y + + +S²→Pushout⋁↪fold⋁S²↺'≡idfun : (x : _) → S²→Pushout⋁↪fold⋁S²↺' base x ≡ ∣ x ∣ +S²→Pushout⋁↪fold⋁S²↺'≡idfun (inl x) = refl +S²→Pushout⋁↪fold⋁S²↺'≡idfun (inr x) = cong ∣_∣ₕ (push (inr x)) +S²→Pushout⋁↪fold⋁S²↺'≡idfun (push (inl x) i) j = + ∣ compPath-filler (push (inl x)) (λ i₁ → push (inr x) (~ i₁)) (~ j) i ∣ₕ +S²→Pushout⋁↪fold⋁S²↺'≡idfun (push (inr x) i) j = ∣ push (inr x) (j ∧ i) ∣ +S²→Pushout⋁↪fold⋁S²↺'≡idfun (push (push a i) j) k = + ∣ coh-lem {A = Pushout⋁↪fold⋁S²} + (push (inl base)) (push (inr base)) (λ k → push (push tt k)) i j k ∣ₕ + where + coh-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p q : x ≡ y) (r : p ≡ q) + → Cube {A = A} + (λ j k → compPath-filler p (sym q) (~ k) j) (λ k j → q (k ∧ j)) + (λ i k → x) (λ i k → q k) + (((λ k → p ∙ sym (r (~ k))) ∙ rCancel p)) + r + coh-lem {A = A} {x = x} {y = y} = + J (λ y p → (q : x ≡ y) (r : p ≡ q) + → Cube {A = A} + (λ j k → compPath-filler p (sym q) (~ k) j) (λ k j → q (k ∧ j)) + (λ i k → x) (λ i k → q k) + (((λ k → p ∙ sym (r (~ k))) ∙ rCancel p)) + r) + λ q → J (λ q r → Cube (λ j k → compPath-filler refl (λ i → q (~ i)) (~ k) j) + (λ k j → q (k ∧ j)) (λ i k → x) (λ i k → q k) + ((λ k → refl ∙ (λ i → r (~ k) (~ i))) ∙ rCancel refl) r) + ((λ z j k → lUnit (sym (rUnit (λ _ → x))) z k j) + ◁ (λ i j k → (refl ∙ (λ i₁ → rUnit (λ _ → x) (~ i₁))) (k ∨ i) j)) + +S²→Pushout⋁↪fold⋁S²↺ : (x : S²) → ∥Pushout⋁↪fold⋁S²∥₅ → ∥Pushout⋁↪fold⋁S²∥₅ +S²→Pushout⋁↪fold⋁S²↺ x = trRec (isOfHLevelTrunc 5) (S²→Pushout⋁↪fold⋁S²↺' x) + +isEq-S²→Pushout⋁↪fold⋁S²↺' : (x : _) → isEquiv (S²→Pushout⋁↪fold⋁S²↺ x) +isEq-S²→Pushout⋁↪fold⋁S²↺' = + S²ToSetElim (λ _ → isProp→isSet (isPropIsEquiv _)) + (subst isEquiv (funExt id≡) (idIsEquiv _)) + where + id≡ : (x : _) → x ≡ S²→Pushout⋁↪fold⋁S²↺ base x + id≡ = trElim (λ _ → isOfHLevelPath 5 (isOfHLevelTrunc 5) _ _) + (sym ∘ S²→Pushout⋁↪fold⋁S²↺'≡idfun) + +S²→Pushout⋁↪fold⋁S²↺⁻¹ : (x : S²) → ∥Pushout⋁↪fold⋁S²∥₅ → ∥Pushout⋁↪fold⋁S²∥₅ +S²→Pushout⋁↪fold⋁S²↺⁻¹ x = trRec (isOfHLevelTrunc 5) (S²→Pushout⋁↪fold⋁S²↺'⁻ x) + +{- S²→Pushout⋁↪fold⋁S²↺⁻¹ is a section of S²→Pushout⋁↪fold⋁S²↺ -} +secS²→Pushout⋁↪fold⋁S²↺ : (x : S²) (y : ∥Pushout⋁↪fold⋁S²∥₅) + → S²→Pushout⋁↪fold⋁S²↺ x (S²→Pushout⋁↪fold⋁S²↺⁻¹ x y) ≡ y +secS²→Pushout⋁↪fold⋁S²↺ x = + trElim (λ _ → isOfHLevelPath 5 (isOfHLevelTrunc 5) _ _) + (h x) + where + h : (x : S²) (a : Pushout⋁↪fold⋁S²) + → S²→Pushout⋁↪fold⋁S²↺ x + (S²→Pushout⋁↪fold⋁S²↺⁻¹ x ∣ a ∣) ≡ ∣ a ∣ + h base a = cong (S²→Pushout⋁↪fold⋁S²↺ base) + (S²→Pushout⋁↪fold⋁S²↺'≡idfun a) + ∙ S²→Pushout⋁↪fold⋁S²↺'≡idfun a + h (surf i j) a k = main a k i j + where + side : (a : Pushout⋁↪fold⋁S²) → _ + side a = cong (S²→Pushout⋁↪fold⋁S²↺ base) (S²→Pushout⋁↪fold⋁S²↺'≡idfun a) + ∙ S²→Pushout⋁↪fold⋁S²↺'≡idfun a + + refl-b : Path ∥Pushout⋁↪fold⋁S²∥₅ _ _ + refl-b = (refl {x = ∣ inl (base , base) ∣ₕ}) + + main : (a : Pushout⋁↪fold⋁S²) + → Cube (λ i j → S²→Pushout⋁↪fold⋁S²↺ (surf i j) + (S²→Pushout⋁↪fold⋁S²↺⁻¹ (surf i j) ∣ a ∣)) + (λ _ _ → ∣ a ∣) + (λ i _ → side a i) (λ i _ → side a i) + (λ i _ → side a i) (λ i _ → side a i) + main = + Pushout⋁↪fold⋁S²→Set (λ _ → isOfHLevelPathP' 2 + (isOfHLevelPathP' 3 (isOfHLevelTrunc 5 _ _) _ _) _ _) + λ k i j → + hcomp (λ r → λ { (i = i0) → rUnit refl-b r k + ; (i = i1) → rUnit refl-b r k + ; (j = i0) → rUnit refl-b r k + ; (j = i1) → rUnit refl-b r k + ; (k = i0) → →Ω²∥Pushout⋁↪fold⋁S²∥₅Id + (surf (~ i) j) (~ r) i j + ; (k = i1) → μ base base}) + (μ-coh k i j) + where + μ : (x y : S²) → ∥Pushout⋁↪fold⋁S²∥₅ + μ x y = ∣ inl (x , y) ∣ₕ + + μUnit : (x : S²) → μ x base ≡ μ base x + μUnit base = refl + μUnit (surf i j) k = ∣ + hcomp (λ r → λ {(i = i0) → push (push tt k) (~ r) + ; (i = i1) → push (push tt k) (~ r) + ; (j = i0) → push (push tt k) (~ r) + ; (j = i1) → push (push tt k) (~ r) + ; (k = i0) → push (inl (surf i j)) (~ r) + ; (k = i1) → push (inr (surf i j)) (~ r)}) + (inr (surf i j)) ∣ₕ + + μ-coh : Path (Square {A = ∥Pushout⋁↪fold⋁S²∥₅} + (λ _ → ∣ inl (base , base) ∣) (λ _ → ∣ inl (base , base) ∣) + (λ _ → ∣ inl (base , base) ∣) (λ _ → ∣ inl (base , base) ∣)) + (λ i j → ∣ inl (surf (~ i) j , surf i j) ∣ₕ) + refl + μ-coh = + (cong₂Funct (λ (x y : (Path S² base base)) → cong₂ μ x y) (sym surf) surf + ∙ cong (_∙ cong (cong (μ base)) surf) (λ i → cong (cong (λ x → μUnit x i)) (sym surf)) + ∙ lCancel (cong (cong (μ base)) surf)) + +retrS²→Pushout⋁↪fold⋁S²↺ : (x : S²) (y : ∥Pushout⋁↪fold⋁S²∥₅) + → S²→Pushout⋁↪fold⋁S²↺⁻¹ x (S²→Pushout⋁↪fold⋁S²↺ x y) ≡ y +retrS²→Pushout⋁↪fold⋁S²↺ x = + trElim (λ _ → isOfHLevelPath 5 (isOfHLevelTrunc 5) _ _) + (main x) + where + main : (x : S²) (a : Pushout⋁↪fold⋁S²) + → S²→Pushout⋁↪fold⋁S²↺⁻¹ x (S²→Pushout⋁↪fold⋁S²↺ x ∣ a ∣) ≡ ∣ a ∣ + main base a = secS²→Pushout⋁↪fold⋁S²↺ base ∣ a ∣ + main (surf i j) a l = secS²→Pushout⋁↪fold⋁S²↺ (surf (~ i) j) ∣ a ∣ l + +∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ : (x : S²) + → Iso ∥Pushout⋁↪fold⋁S²∥₅ ∥Pushout⋁↪fold⋁S²∥₅ +fun (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ x) = S²→Pushout⋁↪fold⋁S²↺ x +inv (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ x) = S²→Pushout⋁↪fold⋁S²↺⁻¹ x +rightInv (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ x) = secS²→Pushout⋁↪fold⋁S²↺ x +leftInv (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ x) = retrS²→Pushout⋁↪fold⋁S²↺ x + +PreCode : (x : Susp S²) → Type +PreCode north = ∥Pushout⋁↪fold⋁S²∥₅ +PreCode south = ∥Pushout⋁↪fold⋁S²∥₅ +PreCode (merid a i) = isoToPath (∥Pushout⋁↪fold⋁S²∥₅≃∥Pushout⋁↪fold⋁S²∥₅ a) i + +hLevPreCode : (x : Susp S²) → isOfHLevel 5 (PreCode x) +hLevPreCode = + suspToPropElim base (λ _ → isPropIsOfHLevel 5) (isOfHLevelTrunc 5) + +Code : (hLevelTrunc 6 (Susp S²)) → Type ℓ-zero +Code = + fst ∘ trRec {B = TypeOfHLevel ℓ-zero 5} (isOfHLevelTypeOfHLevel 5) + λ x → (PreCode x) , (hLevPreCode x) + +private + cong-coherence : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (r : refl ≡ p) + → cong (p ∙_) (sym r) ∙ sym (rUnit p) + ≡ cong (_∙ p) (sym r) ∙ sym (lUnit p) + cong-coherence p = J (λ p r → cong (p ∙_) (sym r) ∙ sym (rUnit p) + ≡ cong (_∙ p) (sym r) ∙ sym (lUnit p)) refl + +{- The function Pushout⋁↪fold⋁S² → Ω∥S³∥₆ will be the obvious one -} +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ : Pushout⋁↪fold⋁S² → Ω∥S³∥₆ +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (inl x) = cong ∣_∣ₕ (σ S²∙ (fst x) ∙ σ S²∙ (snd x)) +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (inr x) = cong ∣_∣ₕ (σ S²∙ x) +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (push (inl x) i) j = + ∣ (cong (σ S²∙ x ∙_) (rCancel (merid base)) ∙ sym (rUnit (σ S²∙ x))) i j ∣ₕ +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (push (inr x) i) j = + ∣ (cong (_∙ σ S²∙ x) (rCancel (merid base)) ∙ sym (lUnit (σ S²∙ x))) i j ∣ₕ +Pushout⋁↪fold⋁S²→Ω∥S³∥₆ (push (push a k) i) j = + ∣ cong-coherence (σ S²∙ base) (sym (rCancel (merid base))) k i j ∣ₕ + +{- Truncated version (the equivalence) -} +∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ : ∥Pushout⋁↪fold⋁S²∥₅ → Ω∥S³∥₆ +∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ = + trRec (isOfHLevelTrunc 6 _ _) Pushout⋁↪fold⋁S²→Ω∥S³∥₆ + +decode' : (x : Susp S²) → Code ∣ x ∣ + → Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ x ∣ +decode' north p = ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ p +decode' south p = ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ p ∙ cong ∣_∣ₕ (merid base) +decode' (merid a i) = main a i + where + baseId : (x : Pushout⋁↪fold⋁S²) + → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ (S²→Pushout⋁↪fold⋁S²↺' base x) + ≡ ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ ∣ x ∣ + baseId x = + cong ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ (S²→Pushout⋁↪fold⋁S²↺'≡idfun x) + + mainLemma : (a : S²) (x : Pushout⋁↪fold⋁S²) + → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ (S²→Pushout⋁↪fold⋁S²↺'⁻ a x) + ∙ (λ i → ∣ merid a i ∣) + ≡ ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ ∣ x ∣ ∙ cong ∣_∣ₕ (merid base) + mainLemma base x = cong (_∙ cong ∣_∣ₕ (merid base)) (baseId x) + mainLemma (surf i j) x k = surf-filler x k i j + where + ∙ΩbaseId : + (q : typ (Ω (Susp∙ (typ S²∙)))) → q ≡ q ∙ σ S²∙ base + ∙ΩbaseId q = rUnit q ∙ cong (q ∙_) (sym (rCancel (merid base))) + + cubeCoherence : + ∀ {ℓ} {A : Type ℓ} {x : A} {p : x ≡ x} + → (refl ≡ p) + → (q : Square {A = x ≡ x} (λ _ → p) (λ _ → p) (λ _ → p) (λ _ → p)) + → Cube {A = Path A x x} + (λ i j → q i j ∙ q (~ i) j) (λ _ _ → p ∙ p) + (λ k j → p ∙ p) (λ _ _ → p ∙ p) + (λ _ _ → p ∙ p) λ _ _ → p ∙ p + cubeCoherence {A = A} {x = x} {p = p} = + J (λ p _ → (q : Square {A = x ≡ x} + (λ _ → p) (λ _ → p) (λ _ → p) (λ _ → p)) + → Cube {A = Path A x x} + (λ i j → q i j ∙ q (~ i) j) (λ _ _ → p ∙ p) + (λ k j → p ∙ p) (λ _ _ → p ∙ p) + (λ _ _ → p ∙ p) λ _ _ → p ∙ p) + (λ q → + (cong₂Funct (λ (x y : Path (x ≡ x) refl refl) → cong₂ _∙_ x y) q (sym q) + ∙ transport + (λ i → cong (λ (p : (refl {x = x}) ≡ refl) k → rUnit (p k) i) q + ∙ cong (λ (p : (refl {x = x}) ≡ refl) k → lUnit (p k) i) (sym q) + ≡ refl {x = refl {x = lUnit (refl {x = x}) i}}) + (rCancel q))) + + surf-side : (i j r l : I) → hLevelTrunc 6 (Susp S²) + surf-side i j r l = + ((cong ∣_∣ₕ (∙ΩbaseId (σ S²∙ (surf (~ i) j)) r)) + ∙ cong ∣_∣ₕ (compPath-filler (merid (surf i j)) + (sym (merid base)) (~ r))) l + + side : (r l : I) → hLevelTrunc 6 (Susp S²) + side r l = + ((cong ∣_∣ₕ (∙ΩbaseId (σ S²∙ base) r)) + ∙ cong ∣_∣ₕ (compPath-filler (merid base) + (sym (merid base)) (~ r))) l + + surf-filler : (x : Pushout⋁↪fold⋁S²) + → Cube {A = Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ south ∣} + (λ i j → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ + (S²→Pushout⋁↪fold⋁S²↺' (surf (~ i) j) x) + ∙ (λ k → ∣ merid (surf i j) k ∣)) + (λ _ _ → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ ∣ x ∣ + ∙ cong ∣_∣ₕ (merid base)) + (λ k j → baseId x k ∙ (λ i₂ → ∣ merid base i₂ ∣)) + (λ k j → baseId x k ∙ (λ i₂ → ∣ merid base i₂ ∣)) + (λ k i → baseId x k ∙ (λ i₂ → ∣ merid base i₂ ∣)) + (λ k i → baseId x k ∙ (λ i₂ → ∣ merid base i₂ ∣)) + surf-filler = + Pushout⋁↪fold⋁S²→Set + (λ _ → isOfHLevelPathP' 2 + (isOfHLevelPathP' 3 (isOfHLevelTrunc 6 _ _ _ _) _ _) _ _) + (λ k i j l → hcomp (λ r + → λ {(i = i0) → surf-side i0 i0 r l + ; (i = i1) → surf-side i0 i0 r l + ; (j = i0) → surf-side i0 i0 r l + ; (j = i1) → surf-side i0 i0 r l + ; (k = i0) → surf-side i j r l + ; (k = i1) → surf-side i0 i0 r l + ; (l = i0) → ∣ north ∣ₕ + ; (l = i1) → ∣ merid base r ∣ₕ}) + (cubeCoherence {p = cong ∣_∣ₕ (σ (S²∙) base)} + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))) + (λ i j k → ∣ σ S²∙ (surf (~ i) j) k ∣ₕ) k i j l)) + + main : (a : S²) + → PathP (λ i → Code ∣ merid a i ∣ + → Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ merid a i ∣) + ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ + λ x → ∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ x + ∙ cong ∣_∣ₕ (merid base) + main a = + toPathP (funExt + (trElim (λ _ → isOfHLevelSuc 4 (isOfHLevelTrunc 6 _ _ _ _)) + (λ x + → (λ j → transp (λ i → Path + (hLevelTrunc 6 (Susp S²)) ∣ north ∣ ∣ merid a (i ∨ j) ∣) j + (compPath-filler + (∥Pushout⋁↪fold⋁S²∥₅→Ω∥S³∥₆ + (S²→Pushout⋁↪fold⋁S²↺'⁻ a (transportRefl x j))) + (cong ∣_∣ₕ (merid a)) j)) + ∙ mainLemma a x))) + +decode : (x : hLevelTrunc 6 (Susp S²)) + → Code x → Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ x +decode = + trElim (λ _ → isOfHLevelΠ 6 λ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) + decode' + +decode∙ : decode ∣ north ∣ ∣ inl (base , base) ∣ ≡ refl +decode∙ = + cong (cong ∣_∣ₕ) (cong (λ x → x ∙ x) (rCancel (merid base)) ∙ sym (rUnit refl)) + +encode : (x : hLevelTrunc 6 (Susp S²)) + → Path (hLevelTrunc 6 (Susp S²)) ∣ north ∣ x → Code x +encode x p = subst Code p ∣ inl (base , base) ∣ + +encode-decode : (x : hLevelTrunc 6 (Susp S²)) + → (p : ∣ north ∣ ≡ x) → decode x (encode x p) ≡ p +encode-decode x = J (λ x p → decode x (encode x p) ≡ p) + (cong (decode ∣ north ∣) (transportRefl ∣ inl (base , base) ∣ₕ) + ∙ decode∙) + +decode-encode : (p : Code ∣ north ∣) → encode ∣ north ∣ (decode ∣ north ∣ p) ≡ p +decode-encode = + trElim (λ _ → isOfHLevelPath 5 (isOfHLevelTrunc 5) _ _) + (Pushout⋁↪fold⋁S²WedgeCon.main ((λ a → encode ∣ north ∣ (decode ∣ north ∣ ∣ a ∣))) ∣_∣ₕ + (isOfHLevelTrunc 5) + (λ x → cong (encode ∣ north ∣) (cong (cong ∣_∣ₕ) (cong (σ S²∙ x ∙_) + (rCancel (merid base)) ∙ sym (rUnit (σ S²∙ x)))) + ∙ surf-filler x) + (λ x → (cong (encode ∣ north ∣) (cong (cong ∣_∣ₕ) (cong (_∙ σ S²∙ x) + (rCancel (merid base)) ∙ sym (lUnit (σ S²∙ x)))) + ∙ surf-filler x + ∙ cong ∣_∣ₕ (push (inl x)) ∙ cong ∣_∣ₕ (sym (push (inr x))))) + λ x → lem (encode ∣ north ∣) + (cong (decode ∣ north ∣) (cong ∣_∣ₕ (push (inl x)))) + ((cong (decode ∣ north ∣) (cong ∣_∣ₕ (push (inr x))))) + (surf-filler x) (cong ∣_∣ₕ (push (inl x))) (cong ∣_∣ₕ (sym (push (inr x))))) + where + subber = transport (λ j → Code ∣ merid base (~ j) ∣ₕ) + + lem : ∀ {ℓ} {A B : Type ℓ} {x x' y : A} {l w u : B} + (f : A → B) (p : x ≡ y) (p' : x' ≡ y) (q : f y ≡ l) + (r : l ≡ w) (r' : w ≡ u) + → cong f (sym p) ∙∙ cong f p ∙ q ∙∙ r + ≡ (cong f (sym p') ∙∙ (cong f p' ∙ q ∙ (r ∙ r')) ∙∙ sym r') + lem {x = x} {x' = x'} {y = y'} {l = l} {w = w} {u = u} f p p' q r r' = + (λ i → (λ j → f (p (~ j ∨ i))) ∙∙ (λ j → f (p (j ∨ i))) ∙ q ∙∙ r) + ∙∙ (λ i → (λ j → f (p' (~ j ∨ ~ i))) ∙∙ (λ j → f (p' (j ∨ ~ i))) + ∙ compPath-filler q r i ∙∙ λ j → r (i ∨ j)) + ∙∙ λ i → cong f (sym p') ∙∙ cong f p' ∙ q + ∙ compPath-filler r r' i ∙∙ λ j → r' (~ j ∧ i) + + mainId : (x : S²) + → (S²→Pushout⋁↪fold⋁S²↺' x (inl (base , base))) ≡ ∣ inl (x , base) ∣ₕ + mainId base = refl + mainId (surf i i₁) = refl + + surf-filler : (x : S²) + → encode ∣ north ∣ (λ i → ∣ σ S²∙ x i ∣) ≡ ∣ inl (x , base) ∣ + surf-filler x = + (λ i → transp (λ j → Code (∣ merid base (i ∧ ~ j) ∣ₕ)) (~ i) + (encode ∣ merid base i ∣ + λ j → ∣ compPath-filler + (merid x) (sym (merid base)) (~ i) j ∣ₕ)) + ∙ cong subber + (transportRefl (S²→Pushout⋁↪fold⋁S²↺' x (inl (base , base))) + ∙ mainId x) + +IsoΩ∥SuspS²∥₅∥Pushout⋁↪fold⋁S²∥₅ : + Iso (hLevelTrunc 5 (typ (Ω (Susp∙ S²)))) + (hLevelTrunc 5 Pushout⋁↪fold⋁S²) +IsoΩ∥SuspS²∥₅∥Pushout⋁↪fold⋁S²∥₅ = + compIso (invIso (PathIdTruncIso _)) is + where + is : Iso _ _ + fun is = encode ∣ north ∣ + inv is = decode ∣ north ∣ + rightInv is = decode-encode + leftInv is = encode-decode ∣ north ∣ diff --git a/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso2.agda b/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso2.agda new file mode 100644 index 0000000000..4329e05a97 --- /dev/null +++ b/Cubical/Homotopy/Group/Pi4S3/S3PushoutIso2.agda @@ -0,0 +1,64 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 where + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group + +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Pushout +open import Cubical.HITs.Truncation renaming + (rec to trRec ; elim to trElim ; elim2 to trElim2 ; map to trMap) +open import Cubical.HITs.S2 + +≡→Pushout⋁↪fold⋁≡ : ∀ {ℓ} {A B : Pointed ℓ} + → (A ≡ B) → Pushout⋁↪fold⋁∙ A ≡ Pushout⋁↪fold⋁∙ B +≡→Pushout⋁↪fold⋁≡ = cong Pushout⋁↪fold⋁∙ + +private + ∙≃→π≅ : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) + → (e : typ A ≃ typ B) + → fst e (pt A) ≡ pt B + → GroupEquiv (πGr n A) (πGr n B) + ∙≃→π≅ {A = A} {B = B} n e = + EquivJ (λ A e → (a : A) → fst e a ≡ pt B + → GroupEquiv (πGr n (A , a)) (πGr n B)) + (λ b p → J (λ b p → GroupEquiv (πGr n (typ B , b)) (πGr n B)) + idGroupEquiv (sym p)) + e (pt A) + +{- π₄(S³) ≅ π₃((S² × S²) ⊔ᴬ S²) + where A = S² ∨ S² -} +π₄S³≅π₃PushS² : + GroupIso (πGr 3 (S₊∙ 3)) + (πGr 2 (Pushout⋁↪fold⋁∙ (S₊∙ 2))) +π₄S³≅π₃PushS² = + compGroupIso + (GroupEquiv→GroupIso + (∙≃→π≅ 3 (compEquiv (isoToEquiv (invIso IsoS³S3)) S³≃SuspS²) refl)) + (compGroupIso + (invGroupIso (GrIso-πΩ-π 2)) + (compGroupIso + (πTruncGroupIso 2) + (compGroupIso + (GroupEquiv→GroupIso + (∙≃→π≅ {B = _ , ∣ inl (base , base) ∣ₕ} + 2 (isoToEquiv IsoΩ∥SuspS²∥₅∥Pushout⋁↪fold⋁S²∥₅) encode∙)) + (compGroupIso + (invGroupIso (πTruncGroupIso 2)) + (GroupEquiv→GroupIso (invEq (GroupPath _ _) + (cong (πGr 2) + (cong Pushout⋁↪fold⋁∙ (ua∙ S²≃SuspS¹ refl))))))))) + where + encode∙ : encode ∣ north ∣ refl ≡ ∣ inl (base , base) ∣ + encode∙ = transportRefl _ diff --git a/Cubical/Homotopy/Group/Pi4S3/Summary.agda b/Cubical/Homotopy/Group/Pi4S3/Summary.agda new file mode 100644 index 0000000000..7cfbd97bd7 --- /dev/null +++ b/Cubical/Homotopy/Group/Pi4S3/Summary.agda @@ -0,0 +1,98 @@ +{- + +This file contains a summary of the proof that π₄(S³) ≡ ℤ/2ℤ + +The --experimental-lossy-unification flag is used to speed up type checking. +The file still type checks without it, but it's a lot slower (about 10 times). + +-} +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Group.Pi4S3.Summary where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed + +open import Cubical.Data.Nat.Base +open import Cubical.Data.Sigma.Base + +open import Cubical.HITs.Sn +open import Cubical.HITs.SetTruncation + +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.HopfInvariant.Homomorphism +open import Cubical.Homotopy.HopfInvariant.HopfMap +open import Cubical.Homotopy.HopfInvariant.Brunerie +open import Cubical.Homotopy.Whitehead +open import Cubical.Homotopy.Group.Base hiding (π) +open import Cubical.Homotopy.Group.Pi3S2 +open import Cubical.Homotopy.Group.Pi4S3.BrunerieNumber + +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Instances.Bool +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.GroupPath +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Instances.Int +open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Algebra.Group.ZAction + + +-- Homotopy groups (shifted version of π'Gr to get nicer numbering) +π : ℕ → Pointed₀ → Group₀ +π n X = π'Gr (predℕ n) X + + +-- Nicer notation for the spheres (as pointed types) +𝕊² 𝕊³ : Pointed₀ +𝕊² = S₊∙ 2 +𝕊³ = S₊∙ 3 + + +-- The Brunerie number; defined in Cubical.Homotopy.Group.Pi4S3.BrunerieNumber +-- as "abs (HopfInvariant-π' 0 ([ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×))" +β : ℕ +β = Brunerie + + +-- The connection to π₄(S³) is then also proved in the BrunerieNumber +-- file following Corollary 3.4.5 in Guillaume Brunerie's PhD thesis. +βSpec : GroupEquiv (π 4 𝕊³) (ℤ/ β) +βSpec = BrunerieIso + + +-- Ideally one could prove that β is 2 by normalization, but this does +-- not seem to terminate before we run out of memory. To try normalize +-- this use "C-u C-c C-n β≡2" (which normalizes the term, ignoring +-- abstract's). So instead we prove this by hand as in the second half +-- of Guillaume's thesis. +β≡2 : β ≡ 2 +β≡2 = Brunerie≡2 + + +-- This involves a lot of theory, for example that π₃(S²) ≃ ℤ where +-- the underlying map is induced by the Hopf invariant (which involves +-- the cup product on cohomology). +_ : GroupEquiv (π 3 𝕊²) ℤ +_ = hopfInvariantEquiv + +-- Which is a consequence of the fact that π₃(S²) is generated by the +-- Hopf map. +_ : gen₁-by (π 3 𝕊²) ∣ HopfMap ∣₂ +_ = π₂S³-gen-by-HopfMap + +-- etc. For more details see the proof of "Brunerie≡2". + + +-- Combining all of this gives us the desired equivalence of groups: +π₄S³≃ℤ/2ℤ : GroupEquiv (π 4 𝕊³) (ℤ/ 2) +π₄S³≃ℤ/2ℤ = subst (GroupEquiv (π 4 𝕊³)) (cong ℤ/_ β≡2) βSpec + + +-- By the SIP this induces an equality of groups: +π₄S³≡ℤ/2ℤ : π 4 𝕊³ ≡ ℤ/ 2 +π₄S³≡ℤ/2ℤ = GroupPath _ _ .fst π₄S³≃ℤ/2ℤ + + +-- As a sanity check we also establish the equality with Bool: +π₄S³≡Bool : π 4 𝕊³ ≡ Bool +π₄S³≡Bool = π₄S³≡ℤ/2ℤ ∙ GroupPath _ _ .fst (GroupIso→GroupEquiv ℤ/2≅Bool) diff --git a/Cubical/Homotopy/Group/PinSn.agda b/Cubical/Homotopy/Group/PinSn.agda new file mode 100644 index 0000000000..15dbfccf8c --- /dev/null +++ b/Cubical/Homotopy/Group/PinSn.agda @@ -0,0 +1,334 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains: +1. An iso πₙ'(Sⁿ) ≅ ℤ, where πₙ'(Sⁿ) = ∥ Sⁿ →∙ Sⁿ ∥₀ +2. A proof that idfun∙ : Sⁿ →∙ Sⁿ is the generator of πₙ'(Sⁿ) +-} +module Cubical.Homotopy.Group.PinSn where + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.SuspensionMap +open import Cubical.Homotopy.Connected + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism + +open import Cubical.HITs.SetTruncation + renaming (elim to sElim ; elim2 to sElim2 + ; map to sMap) +open import Cubical.HITs.Truncation renaming + (elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Int renaming (_+_ to _+ℤ_) + +open import Cubical.ZCohomology.Properties + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.ZAction + +open Iso + +-- The goal is to prove that πₙSⁿ ≅ ℤ. This is of course trivial, given +-- that ΩⁿK(ℤ,n) ≅ ℤ is already proved. However, we would like to do +-- this for πₙSⁿ defined via (Sⁿ →∙ Sⁿ) and prove that the generator +-- of this group is idfun∙ : Sⁿ → Sⁿ. + +private + suspMapS^ : (n : ℕ) → (S₊∙ (suc n) →∙ S₊∙ (suc n)) + → S₊∙ (2 + n) →∙ S₊∙ (2 + n) + suspMapS^ n = suspMap {A = S₊∙ (suc n)} n + +is2ConnectedSuspMap : (n : ℕ) → isConnectedFun 2 (suspMapS^ (suc n)) +is2ConnectedSuspMap n = + isConnectedFunSubtr 2 n _ + (subst (λ x → isConnectedFun x (suspMapS^ (suc n))) + (subtrLem n (suc (suc n)) ∙ +-comm 2 n) + (isConnectedSuspMap (suc n) (suc n))) + where + subtrLem : (n m : ℕ) → (n + m ∸ n) ≡ m + subtrLem zero m = refl + subtrLem (suc n) m = subtrLem n m + +-- From the connectedness of the suspension map, +-- we get an iso πₙSⁿ ≅ πₙ₊₁(Sⁿ⁺¹) for n ≥ 2. +SphereSuspIso : (n : ℕ) + → Iso (π' (2 + n) (S₊∙ (2 + n))) (π' (3 + n) (S₊∙ (3 + n))) +SphereSuspIso n = + compIso setTruncTrunc2Iso + (compIso + (connectedTruncIso 2 + (suspMap {A = S₊∙ (suc (suc n))} (suc n)) (is2ConnectedSuspMap n)) + (invIso (setTruncTrunc2Iso))) + +SphereSuspGrIso : (n : ℕ) + → GroupIso (π'Gr (suc n) (S₊∙ (2 + n))) (π'Gr (2 + n) (S₊∙ (3 + n))) +fst (SphereSuspGrIso n) = SphereSuspIso n +snd (SphereSuspGrIso n) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → IsGroupHom.pres· (suspMapπ'Hom (suc n) .snd) ∣ f ∣₂ ∣ g ∣₂) + +private + stLoop₁ : π 2 (S₊∙ 2) + stLoop₁ = ∣ sym (rCancel (merid base)) + ∙∙ (λ i → merid (loop i) ∙ sym (merid base)) + ∙∙ rCancel (merid base) ∣₂ + + stLoop₁flip : π 2 (S₊∙ 2) + stLoop₁flip = sMap flipSquare stLoop₁ + +flipSquareIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupIso (πGr (suc n) A) (πGr (suc n) A) +fun (fst (flipSquareIso n)) = sMap flipSquare +inv (fst (flipSquareIso n)) = sMap flipSquare +rightInv (fst (flipSquareIso n)) = + sElim (λ _ → isSetPathImplicit) λ _ → refl +leftInv (fst (flipSquareIso n)) = + sElim (λ _ → isSetPathImplicit) λ _ → refl +snd (flipSquareIso n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ + ((sym (sym≡flipSquare (f ∙ g)) + ∙∙ symDistr f g + ∙∙ cong₂ _∙_ (sym≡flipSquare g) (sym≡flipSquare f) + ∙ EH n (flipSquare g) (flipSquare f)))) + +-- We now want to prove π₁S¹≅π₂S². This will be done by passing through +-- The loop space defintion of the groups +π'₂S²≅π₂S² : GroupIso (π'Gr 1 (S₊∙ 2)) (πGr 1 (S₊∙ 2)) +π'₂S²≅π₂S² = π'Gr≅πGr 1 (S₊∙ 2) + +{- We now define the iso π₂S² ≅ π₁S¹ -} +π₂S²≅π₁S¹ : GroupIso (πGr 1 (S₊∙ 2)) (πGr 0 (S₊∙ 1)) +fst π₂S²≅π₁S¹ = + compIso setTruncTrunc2Iso + (compIso + (compIso (invIso (PathIdTruncIso 2)) + (compIso (congIso (invIso (PathIdTruncIso 3))) + (compIso + (congIso (invIso (Iso-Kn-ΩKn+1 1))) + (PathIdTruncIso 2)))) + (invIso setTruncTrunc2Iso)) +snd π₂S²≅π₁S¹ = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → + cong (inv setTruncTrunc2Iso) + (cong (fun (PathIdTruncIso 2)) + (cong (cong (ΩKn+1→Kn 1)) + (cong (cong (inv (PathIdTruncIso 3))) + (cong (inv (PathIdTruncIso 2)) + (refl {x = ∣ f ∙ g ∣}) + ∙ cong-∙ ∣_∣ₕ f g) + ∙ cong-∙ (inv (PathIdTruncIso 3)) (cong ∣_∣ₕ f) (cong ∣_∣ₕ g)) + ∙ cong-∙ (ΩKn+1→Kn 1) + (cong (inv (PathIdTruncIso 3)) (cong ∣_∣ₕ f)) + (cong (inv (PathIdTruncIso 3)) (cong ∣_∣ₕ g))) + ∙ PathIdTruncIsoFunct 1 + (cong (ΩKn+1→Kn 1) (λ i → inv (PathIdTruncIso 3) ∣ f i ∣ₕ)) + (cong (ΩKn+1→Kn 1) (λ i → inv (PathIdTruncIso 3) ∣ g i ∣ₕ))) + ∙ setTruncTrunc2IsoFunct + ((fun (PathIdTruncIso 2)) + (cong (ΩKn+1→Kn 1) (λ i → inv (PathIdTruncIso 3) ∣ f i ∣ₕ))) + ((fun (PathIdTruncIso 2)) + (cong (ΩKn+1→Kn 1) (λ i → inv (PathIdTruncIso 3) ∣ g i ∣ₕ)))) + where + setTruncTrunc2IsoFunct : ∀ {ℓ} {A : Type ℓ} {x : A} + (p q : hLevelTrunc 2 (x ≡ x)) + → inv setTruncTrunc2Iso + (Cubical.HITs.Truncation.map2 _∙_ p q) + ≡ ·π 0 (inv setTruncTrunc2Iso p) (inv setTruncTrunc2Iso q) + setTruncTrunc2IsoFunct = + trElim2 (λ _ _ → isSetPathImplicit) λ _ _ → refl + +π₂'S²≅π₁'S¹ : GroupIso (π'Gr 1 (S₊∙ 2)) (π'Gr 0 (S₊∙ 1)) +π₂'S²≅π₁'S¹ = + compGroupIso (π'Gr≅πGr 1 (S₊∙ 2)) + (compGroupIso (flipSquareIso 0) + (compGroupIso π₂S²≅π₁S¹ + (invGroupIso (π'Gr≅πGr 0 (S₊∙ 1))))) + +-- We get our iso +πₙ'Sⁿ≅ℤ : (n : ℕ) → GroupIso (π'Gr n (S₊∙ (suc n))) ℤGroup +πₙ'Sⁿ≅ℤ zero = + compGroupIso (π'Gr≅πGr zero (S₊∙ 1)) + ((compIso (setTruncIdempotentIso (isGroupoidS¹ _ _)) ΩS¹Isoℤ) + , makeIsGroupHom (sElim2 (λ _ _ → isProp→isSet (isSetℤ _ _)) + winding-hom)) +πₙ'Sⁿ≅ℤ (suc zero) = compGroupIso π₂'S²≅π₁'S¹ (πₙ'Sⁿ≅ℤ zero) +πₙ'Sⁿ≅ℤ (suc (suc n)) = + compGroupIso (invGroupIso (SphereSuspGrIso n)) (πₙ'Sⁿ≅ℤ (suc n)) + +-- Same thing for the loop space definition +πₙSⁿ≅ℤ : (n : ℕ) → GroupIso (πGr n (S₊∙ (suc n))) ℤGroup +πₙSⁿ≅ℤ n = + compGroupIso (invGroupIso (π'Gr≅πGr n (S₊∙ (suc n)))) (πₙ'Sⁿ≅ℤ n) + +-- The goal now is to show that πₙ'Sⁿ≅ℤ takes idfun∙ : Sⁿ → Sⁿ to 1. +-- For this, we need a bunch of identities: +private + Isoπ₁S¹ℤ = (compIso (setTruncIdempotentIso (isGroupoidS¹ _ _)) ΩS¹Isoℤ) + + π'₂S²≅π₂S²⁻-stLoop' : inv (fst (π'₂S²≅π₂S²)) stLoop₁flip ≡ ∣ idfun∙ _ ∣₂ + π'₂S²≅π₂S²⁻-stLoop' = + cong ∣_∣₂ (ΣPathP ((funExt + (λ { north → refl + ; south → merid base + ; (merid base i) j → + hcomp (λ k + → λ {(i = i0) → north + ; (i = i1) → merid base (j ∧ k) + ; (j = i0) → rUnit (λ _ → north) k i + ; (j = i1) → merid base (i ∧ k)}) + north + ; (merid (loop k) i) j + → hcomp (λ r + → λ {(i = i0) → north + ; (i = i1) → merid base (j ∧ r) + ; (j = i0) → + rUnit (funExt⁻ (cong fst (cong (Ω→SphereMap 1) + (flipSquare (sym (rCancel (merid base)) + ∙∙ (λ i₁ → merid (loop i₁) + ∙ sym (merid base)) + ∙∙ rCancel (merid base))))) (loop k)) r i + ; (j = i1) → lem₂ r i k}) + (((sym (rCancel (merid base)) ∙∙ + (λ i₁ → merid (loop i₁) ∙ sym (merid base)) ∙∙ + rCancel (merid base))) k i)})) , refl)) + where + genBot+side : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) + → Cube {A = A} (λ j r → x) (λ j r → p (~ j ∨ r)) + (λ i r → p i) (λ i r → p (i ∧ r)) + (λ i j → p (i ∧ ~ j)) (λ i j → p i) + × Cube {A = A} (λ j r → p (~ j ∨ r)) (λ j r → p (r ∧ j)) + (λ k r → p (~ k)) (λ k r → p r) + (λ k j → p (~ k ∧ ~ j)) λ k j → p (j ∨ ~ k) + genBot+side {A = A} {x = x} = + J (λ y p → Cube {A = A} (λ j r → x) (λ j r → p (~ j ∨ r)) + (λ i r → p i) (λ i r → p (i ∧ r)) + (λ i j → p (i ∧ ~ j)) (λ i j → p i) + × Cube {A = A} (λ j r → p (~ j ∨ r)) (λ j r → p (r ∧ j)) + (λ k r → p (~ k)) (λ k r → p r) + (λ k j → p (~ k ∧ ~ j)) λ k j → p (j ∨ ~ k)) + (refl , refl) + + lem₁ : I → I → I → S₊ 2 + lem₁ j i r = + hcomp (λ k → λ {(i = i0) → north + ; (i = i1) → genBot+side (merid base) .snd k j r + ; (j = i0) → compPath-filler + (merid base) (sym (merid base)) k i + ; (j = i1) → merid base (i ∧ r) + ; (r = i0) → rCancel-filler (merid base) k j i + ; (r = i1) → compPath-filler + (merid base) (sym (merid base)) (~ j ∧ k) i}) + (genBot+side (merid base) .fst i j r) + + lem₂ : I → I → I → S₊ 2 + lem₂ r i k = + hcomp (λ j → λ {(i = i0) → north + ; (i = i1) → merid base (r ∧ j) + ; (r = i0) → doubleCompPath-filler + (sym (rCancel (merid base))) + (λ i₁ → merid (loop i₁) ∙ sym (merid base)) + (rCancel (merid base)) j k i + ; (r = i1) → compPath-filler + (merid (loop k)) (sym (merid base)) (~ j) i + ; (k = i0) → lem₁ j i r + ; (k = i1) → lem₁ j i r}) + ((merid (loop k) ∙ sym (merid base)) i) + + π₂S²≅π₁S¹-stLoop : fun (fst π₂S²≅π₁S¹) stLoop₁ ≡ ∣ loop ∣₂ + π₂S²≅π₁S¹-stLoop = + sym (leftInv Isoπ₁S¹ℤ + (fun (fst π₂S²≅π₁S¹) stLoop₁)) + ∙∙ cong (inv Isoπ₁S¹ℤ) compute + ∙∙ leftInv (compIso (setTruncIdempotentIso (isGroupoidS¹ _ _)) ΩS¹Isoℤ) + ∣ loop ∣₂ + where + compute : fun Isoπ₁S¹ℤ (fun (fst π₂S²≅π₁S¹) stLoop₁) + ≡ fun Isoπ₁S¹ℤ ∣ loop ∣₂ + compute = refl + + π₂'S²≅π₁'S¹≡ : fun (fst π₂'S²≅π₁'S¹) ∣ idfun∙ _ ∣₂ ≡ ∣ idfun∙ _ ∣₂ + π₂'S²≅π₁'S¹≡ = + lem₁ ∣ idfun∙ _ ∣₂ stLoop₁ + (sym π'₂S²≅π₂S²⁻-stLoop') + (cong (inv (fst (π'Gr≅πGr zero (S₊∙ 1)))) π₂S²≅π₁S¹-stLoop + ∙ lem₂) + where + lem₁ : (x : _) (y : π 2 (S₊∙ 2)) + → (x ≡ inv (fst π'₂S²≅π₂S²) (fun (fst (flipSquareIso 0)) y)) + → inv (fst (π'Gr≅πGr zero (S₊∙ 1))) (fun (fst π₂S²≅π₁S¹) y) + ≡ ∣ idfun∙ _ ∣₂ + → fun (fst π₂'S²≅π₁'S¹) x ≡ ∣ idfun∙ _ ∣₂ + lem₁ x y p q = + cong (fun (fst π₂'S²≅π₁'S¹)) p + ∙∙ (λ i → inv (fst (π'Gr≅πGr zero (S₊∙ (suc zero)))) (fun (fst π₂S²≅π₁S¹) + (fun (fst (flipSquareIso zero)) + (rightInv + (fst (π'Gr≅πGr (suc zero) (S₊∙ (suc (suc zero))))) + (inv (fst (flipSquareIso zero)) y) i) + ))) + ∙∙ cong (inv (fst (π'Gr≅πGr zero (S₊∙ (suc zero))))) + (cong (fun (fst π₂S²≅π₁S¹)) + (rightInv (fst (flipSquareIso zero)) y)) + ∙ q + + lem₂ : inv (fst (π'Gr≅πGr zero (S₊∙ 1))) ∣ loop ∣₂ ≡ ∣ idfun∙ _ ∣₂ + lem₂ = + cong ∣_∣₂ (ΣPathP (funExt (λ { base → refl ; (loop i) → refl}) , refl)) + + suspPresIdfun : (n : ℕ) → suspMap n (idfun∙ (S₊∙ (suc n))) ≡ idfun∙ _ + suspPresIdfun n = + ΣPathP ((funExt + (λ { north → refl + ; south → merid (ptSn (suc n)) + ; (merid a i) j + → compPath-filler (merid a) (sym (merid (ptSn (suc n)))) (~ j) i})) + , refl) + + suspPresIdfun2 : (n : ℕ) + → fun (fst (invGroupIso (SphereSuspGrIso n))) + ∣ idfun∙ (S₊∙ (suc (suc (suc n)))) ∣₂ + ≡ ∣ idfun∙ _ ∣₂ + suspPresIdfun2 n = + sym (cong (fun (fst (invGroupIso (SphereSuspGrIso n)))) + (cong ∣_∣₂ (suspPresIdfun (suc n)))) + ∙ leftInv (SphereSuspIso n) ∣ idfun∙ _ ∣₂ + +-- We finally have our results +πₙ'Sⁿ≅ℤ-idfun∙ : (n : ℕ) + → fun (fst (πₙ'Sⁿ≅ℤ n)) ∣ idfun∙ _ ∣₂ ≡ (pos (suc zero)) +πₙ'Sⁿ≅ℤ-idfun∙ zero = refl +πₙ'Sⁿ≅ℤ-idfun∙ (suc zero) = speedUp ∣ idfun∙ _ ∣₂ π₂'S²≅π₁'S¹≡ + where + speedUp : (x : _) -- stated like this for faster type checking + → fun (fst π₂'S²≅π₁'S¹) x ≡ ∣ idfun∙ _ ∣₂ + → fun (fst (πₙ'Sⁿ≅ℤ (suc zero))) x ≡ pos (suc zero) + speedUp x p = cong (fun (fst (πₙ'Sⁿ≅ℤ zero))) p +πₙ'Sⁿ≅ℤ-idfun∙ (suc (suc n)) = + cong (fun (fst (πₙ'Sⁿ≅ℤ (suc n)))) (suspPresIdfun2 n) + ∙ πₙ'Sⁿ≅ℤ-idfun∙ (suc n) + +πₙ'Sⁿ-gen-by-idfun : (n : ℕ) → gen₁-by (π'Gr n (S₊∙ (suc n))) ∣ idfun∙ _ ∣₂ +πₙ'Sⁿ-gen-by-idfun n = + subst (gen₁-by (π'Gr n (S₊∙ (suc n)))) + (sym (cong (inv (fst (πₙ'Sⁿ≅ℤ n))) (πₙ'Sⁿ≅ℤ-idfun∙ n)) + ∙ leftInv (fst (πₙ'Sⁿ≅ℤ n)) ∣ idfun∙ _ ∣₂) + (Iso-pres-gen₁ ℤGroup (π'Gr n (S₊∙ (suc n))) + (pos (suc zero)) + (λ h → h , (sym (·Comm h (pos 1)) ∙ ℤ·≡· h (pos 1))) + (invGroupIso (πₙ'Sⁿ≅ℤ n))) diff --git a/Cubical/Homotopy/Group/S3.agda b/Cubical/Homotopy/Group/S3.agda deleted file mode 100644 index 41757fc04f..0000000000 --- a/Cubical/Homotopy/Group/S3.agda +++ /dev/null @@ -1,182 +0,0 @@ -{-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Homotopy.Group.S3 where - -{- -This file contains a summary of what remains for π₄S³≅ℤ/2 to be proved. -See the module π₄S³ at the end of this file. --} - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Function - -open import Cubical.Data.Nat -open import Cubical.Data.Sum -open import Cubical.Data.Sigma -open import Cubical.Data.Int - renaming (_·_ to _·ℤ_ ; _+_ to _+ℤ_) - -open import Cubical.Homotopy.Group.Base -open import Cubical.Homotopy.HopfInvariant.Base -open import Cubical.Homotopy.HopfInvariant.Homomorphism -open import Cubical.Homotopy.HopfInvariant.HopfMap -open import Cubical.Homotopy.Whitehead -open import Cubical.Algebra.Group.Instances.IntMod -open import Cubical.Foundations.Isomorphism - -open import Cubical.HITs.Sn -open import Cubical.HITs.SetTruncation - -open import Cubical.Algebra.Group - renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) -open import Cubical.Algebra.Group.ZAction - - -[_]× : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} - → π' (suc n) X × π' (suc m) X → π' (suc (n + m)) X -[_]× (f , g) = [ f ∣ g ]π' - --- Some type abbreviations (unproved results) -π₃S²-gen : Type -π₃S²-gen = gen₁-by (π'Gr 2 (S₊∙ 2)) ∣ HopfMap ∣₂ - -π₄S³≅ℤ/something : GroupEquiv ℤGroup (π'Gr 2 (S₊∙ 2)) - → Type -π₄S³≅ℤ/something eq = - GroupIso (π'Gr 3 (S₊∙ 3)) - (ℤ/ abs (invEq (fst eq) - [ ∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂ ]×)) - -miniLem₁ : Type -miniLem₁ = (g : ℤ) → gen₁-by ℤGroup g → (g ≡ 1) ⊎ (g ≡ -1) - -miniLem₂ : Type -miniLem₂ = (ϕ : GroupEquiv ℤGroup ℤGroup) (g : ℤ) - → (abs g ≡ abs (fst (fst ϕ) g)) - --- some minor group lemmas -groupLem-help : miniLem₁ → (g : ℤ) → - gen₁-by ℤGroup g → - (ϕ : GroupHom ℤGroup ℤGroup) → - (fst ϕ g ≡ pos 1) ⊎ (fst ϕ g ≡ negsuc 0) - → isEquiv (fst ϕ) -groupLem-help grlem1 g gen ϕ = main (grlem1 g gen) - where - isEquiv- : isEquiv (-_) - isEquiv- = isoToIsEquiv (iso -_ -_ -Involutive -Involutive) - - lem : fst ϕ (pos 1) ≡ pos 1 → fst ϕ ≡ idfun _ - lem p = funExt lem2 - where - lem₁ : (x₁ : ℕ) → fst ϕ (pos x₁) ≡ idfun ℤ (pos x₁) - lem₁ zero = IsGroupHom.pres1 (snd ϕ) - lem₁ (suc zero) = p - lem₁ (suc (suc n)) = - IsGroupHom.pres· (snd ϕ) (pos (suc n)) 1 - ∙ cong₂ _+ℤ_ (lem₁ (suc n)) p - - lem2 : (x₁ : ℤ) → fst ϕ x₁ ≡ idfun ℤ x₁ - lem2 (pos n) = lem₁ n - lem2 (negsuc zero) = - IsGroupHom.presinv (snd ϕ) 1 ∙ cong (λ x → pos 0 - x) p - lem2 (negsuc (suc n)) = - (cong (fst ϕ) (sym (+Comm (pos 0) (negsuc (suc n)))) - ∙ IsGroupHom.presinv (snd ϕ) (pos (suc (suc n)))) - ∙∙ +Comm (pos 0) _ - ∙∙ cong (-_) (lem₁ (suc (suc n))) - - lem₂ : fst ϕ (negsuc 0) ≡ pos 1 → fst ϕ ≡ -_ - lem₂ p = funExt lem2 - where - s = IsGroupHom.presinv (snd ϕ) (negsuc 0) - ∙∙ +Comm (pos 0) _ - ∙∙ cong -_ p - - lem2 : (n : ℤ) → fst ϕ n ≡ - n - lem2 (pos zero) = IsGroupHom.pres1 (snd ϕ) - lem2 (pos (suc zero)) = s - lem2 (pos (suc (suc n))) = - IsGroupHom.pres· (snd ϕ) (pos (suc n)) 1 - ∙ cong₂ _+ℤ_ (lem2 (pos (suc n))) s - lem2 (negsuc zero) = p - lem2 (negsuc (suc n)) = - IsGroupHom.pres· (snd ϕ) (negsuc n) (negsuc 0) - ∙ cong₂ _+ℤ_ (lem2 (negsuc n)) p - - main : (g ≡ pos 1) ⊎ (g ≡ negsuc 0) - → (fst ϕ g ≡ pos 1) ⊎ (fst ϕ g ≡ negsuc 0) - → isEquiv (fst ϕ) - main (inl p) = - J (λ g p → (fst ϕ g ≡ pos 1) - ⊎ (fst ϕ g ≡ negsuc 0) → isEquiv (fst ϕ)) - (λ { (inl x) → subst isEquiv (sym (lem x)) (snd (idEquiv _)) - ; (inr x) → subst isEquiv - (sym (lem₂ (IsGroupHom.presinv (snd ϕ) (pos 1) - ∙ (cong (λ x → pos 0 - x) x)))) - isEquiv- }) - (sym p) - main (inr p) = - J (λ g p → (fst ϕ g ≡ pos 1) - ⊎ (fst ϕ g ≡ negsuc 0) → isEquiv (fst ϕ)) - (λ { (inl x) → subst isEquiv (sym (lem₂ x)) isEquiv- - ; (inr x) → subst isEquiv - (sym (lem ( - IsGroupHom.presinv (snd ϕ) (negsuc 0) - ∙ cong (λ x → pos 0 - x) x))) - (snd (idEquiv _))}) - (sym p) - -groupLem : {G : Group₀} - → miniLem₁ - → GroupEquiv ℤGroup G - → (g : fst G) - → gen₁-by G g - → (ϕ : GroupHom G ℤGroup) - → (fst ϕ g ≡ 1) ⊎ (fst ϕ g ≡ -1) - → isEquiv (fst ϕ) -groupLem {G = G} s = - GroupEquivJ - (λ G _ → (g : fst G) - → gen₁-by G g - → (ϕ : GroupHom G ℤGroup) - → (fst ϕ g ≡ 1) ⊎ (fst ϕ g ≡ -1) - → isEquiv (fst ϕ)) - (groupLem-help s) - --- summary -module π₄S³ - (mini-lem₁ : miniLem₁) - (mini-lem₂ : miniLem₂) - (ℤ≅π₃S² : GroupEquiv ℤGroup (π'Gr 2 (S₊∙ 2))) - (gen-by-HopfMap : π₃S²-gen) - (π₄S³≅ℤ/whitehead : π₄S³≅ℤ/something ℤ≅π₃S²) - (hopfWhitehead : - abs (HopfInvariant-π' 0 - ([ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂) ]×)) - ≡ 2) - where - π₄S³ = π'Gr 3 (S₊∙ 3) - - hopfInvariantEquiv : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤGroup - fst (fst hopfInvariantEquiv) = HopfInvariant-π' 0 - snd (fst hopfInvariantEquiv) = - groupLem mini-lem₁ ℤ≅π₃S² ∣ HopfMap ∣₂ - gen-by-HopfMap - (GroupHom-HopfInvariant-π' 0) - (abs→⊎ _ _ HopfInvariant-HopfMap) - snd hopfInvariantEquiv = snd (GroupHom-HopfInvariant-π' 0) - - lem : ∀ {G : Group₀} (ϕ ψ : GroupEquiv ℤGroup G) (g : fst G) - → abs (invEq (fst ϕ) g) ≡ abs (invEq (fst ψ) g) - lem = - GroupEquivJ - (λ G ϕ → (ψ : GroupEquiv ℤGroup G) (g : fst G) - → abs (invEq (fst ϕ) g) ≡ abs (invEq (fst ψ) g)) - λ ψ → mini-lem₂ (invGroupEquiv ψ) - - main : GroupIso π₄S³ (ℤ/ 2) - main = subst (GroupIso π₄S³) - (cong (ℤ/_) (lem ℤ≅π₃S² (invGroupEquiv (hopfInvariantEquiv)) _ - ∙ hopfWhitehead)) - π₄S³≅ℤ/whitehead diff --git a/Cubical/Homotopy/HopfInvariant/Brunerie.agda b/Cubical/Homotopy/HopfInvariant/Brunerie.agda new file mode 100644 index 0000000000..dd06d87ba0 --- /dev/null +++ b/Cubical/Homotopy/HopfInvariant/Brunerie.agda @@ -0,0 +1,399 @@ +{- +This file contains a proof of the fact that the Brunerie number, +i.e. absolute value of the Hopf invariant of [e , e] : π₃S², is 2. +Here, e is the generator of π₂S² and [_,_] denotes the Whitehead +product. The proof follows Proposition 5.4.4. in Brunerie (2016) +closely, but, for simplicity, considers only the case n = 2. +-} + +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.HopfInvariant.Brunerie where + +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.Group.Pi4S3.BrunerieNumber +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso +open import Cubical.Homotopy.Whitehead + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.SphereProduct +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.RingStructure.GradedCommutativity + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat hiding (_+_) +open import Cubical.Data.Unit +open import Cubical.Data.Sum renaming (rec to ⊎rec) + +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Join +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.Truncation +open import Cubical.HITs.SetTruncation + renaming (elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (map to pMap ; rec to pRec) + +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact + + +open Iso +open IsGroupHom + +-- Some abstract versions of imported lemmas/definitions from +-- ZCohomology.Groups.SphereProduct for faster type checking. +abstract + H²-genₗabs : coHom 2 (S₊ 2 × S₊ 2) + H²-genₗabs = H²-S²×S²-genₗ + + H²-genᵣabs : coHom 2 (S₊ 2 × S₊ 2) + H²-genᵣabs = H²-S²×S²-genᵣ + + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs : (n m : ℕ) + → GroupIso (coHomGr ((suc n) +' (suc m)) + (S₊ (suc n) × S₊ (suc m))) + ℤGroup + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs = Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ + + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs≡ : (n m : ℕ) + → Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs n m ≡ Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ n m + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs≡ n m = refl + + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-⌣ : + fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1)) + (H²-S²×S²-genₗ ⌣ H²-S²×S²-genᵣ) + ≡ 1 + Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-⌣ = + fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1)) (H²-S²×S²-genₗ ⌣ H²-S²×S²-genᵣ) + ≡⟨ cong (fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1))) (sym H²-S²≅H⁴-S²×S²⌣) ⟩ + fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1)) (fun (fst (H²-S²≅H⁴-S²×S²)) ∣ ∣_∣ₕ ∣₂) + ≡⟨ speedUp ∣_∣ₕ ⟩ + fun (fst (Hⁿ-Sⁿ≅ℤ 1)) ∣ ∣_∣ₕ ∣₂ + ≡⟨ refl ⟩ -- Computation! :-) + 1 ∎ + where + speedUp : (f : _) → + fun (fst (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1)) (fun (fst (H²-S²≅H⁴-S²×S²)) ∣ f ∣₂) + ≡ fun (fst (Hⁿ-Sⁿ≅ℤ 1)) ∣ f ∣₂ + speedUp f i = + fun (fst (Hⁿ-Sⁿ≅ℤ 1)) (leftInv (fst H²-S²≅H⁴-S²×S²) ∣ f ∣₂ i) + +-- Some abbreviations +private + inl' : S₊ 2 × S₊ 2 → Pushout⋁↪fold⋁ (S₊∙ 2) + inl' = inl + + qHom : GroupHom (coHomGr 4 (Pushout⋁↪fold⋁ (S₊∙ 2))) + (coHomGr 4 (S₊ 2 × S₊ 2)) + qHom = coHomMorph 4 inl' + + qHomGen : (n : ℕ) → + GroupHom (coHomGr n (Pushout⋁↪fold⋁ (S₊∙ 2))) + (coHomGr n (S₊ 2 × S₊ 2)) + qHomGen n = coHomMorph n inl' + +-- The type C and generator α, β in dim 2 and 4 respectively +-- Recall, the goal is to prove that α ⌣ α = ±2 β + CHopf : Type + CHopf = HopfInvariantPush 0 fold∘W + + Hopfαfold∘W = Hopfα 0 (fold∘W , refl) + Hopfβfold∘W = Hopfβ 0 (fold∘W , refl) + +-- Rewriting CHopf as our favourite pushout +-- S²×S² ← S²∨S² → S² + CHopfIso : Iso CHopf (Pushout⋁↪fold⋁ (S₊∙ 2)) + CHopfIso = + compIso (invIso (equivToIso + (compEquiv (compEquiv pushoutSwitchEquiv + (isoToEquiv (PushoutDistr.PushoutDistrIso fold⋁ W λ _ → tt))) + pushoutSwitchEquiv))) + (equivToIso Pushout-coFibW-fold⋁≃Pushout⋁↪fold⋁) + +-- Cohomology group version of the Iso + coHomCHopfIso : (n : ℕ) + → GroupIso (coHomGr n CHopf) (coHomGr n (Pushout⋁↪fold⋁ (S₊∙ 2))) + coHomCHopfIso n = invGroupIso (coHomIso n CHopfIso) + + +-- We instantiate Mayer-Vietoris for the pushout +module MV-⋁↪-fold⋁ = MV _ _ (S₊∙ 2 ⋁ S₊∙ 2) ⋁↪ fold⋁ + +-- This give us an iso H⁴(S²×S² ← S²∨S² → S²) ≅ H⁴(S²×S²) +isEquiv-qHom : GroupEquiv (coHomGr 4 (Pushout⋁↪fold⋁ (S₊∙ 2))) + (coHomGr 4 (S₊ 2 × S₊ 2)) +fst (fst isEquiv-qHom) = qHom .fst +snd (fst isEquiv-qHom) = + subst isEquiv + (funExt (sElim (λ _ → isSetPathImplicit) (λ _ → refl))) + (×UnitEquiv (isoToPath + (invIso + (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))))) + _ isEquiv-i) + where + ×UnitEquiv : {A B C : Type} + → Unit ≡ C + → (f : A → B × C) + → isEquiv f + → isEquiv (fst ∘ f) + ×UnitEquiv {A = A} {B = B} = + J (λ C _ → (f : A → B × C) → isEquiv f → isEquiv (fst ∘ f)) + λ f eq → record { equiv-proof = + λ b → ((fst (fst (equiv-proof eq (b , tt)))) + , cong fst (fst (equiv-proof eq (b , tt)) .snd)) + , λ y → ΣPathP ((cong fst (equiv-proof eq (b , tt) + .snd ((fst y) + , ΣPathP ((snd y) , refl)))) + , λ i j → equiv-proof eq (b , tt) .snd ((fst y) + , ΣPathP ((snd y) , refl)) i .snd j .fst) } + + isEquiv-i : isEquiv (fst (MV-⋁↪-fold⋁.i 4)) + isEquiv-i = + SES→isEquiv + (isContr→≡UnitGroup + (isOfHLevelRetractFromIso 0 + (compIso + (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 2)) + (compIso + (prodIso (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) + (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p)))) + rUnit×Iso)) + isContrUnit)) + ((isContr→≡UnitGroup + (isOfHLevelRetractFromIso 0 + (compIso + (fst (Hⁿ-⋁ (S₊∙ 2) (S₊∙ 2) 3)) + (compIso + (prodIso (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) + (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p)))) + rUnit×Iso)) + isContrUnit))) + (MV-⋁↪-fold⋁.d 3) + (MV-⋁↪-fold⋁.i 4) + (MV-⋁↪-fold⋁.Δ 4) + (MV-⋁↪-fold⋁.Ker-i⊂Im-d 3) + (MV-⋁↪-fold⋁.Ker-Δ⊂Im-i 4) +snd isEquiv-qHom = qHom .snd + + +-- The goal now is reducing α ⌣ α = ±2 β to gₗ ⌣ gᵣ = e for +-- gₗ, gᵣ generators of H²(S²×S²) ≅ ℤ × ℤ and e generator of +-- H⁴(S²×S²) ≅ ℤ. This essentially just elementary linear algebra at +-- this point. We do it for an arbitrary (well-behaved) iso +-- H⁴(S²×S²) ≅ ℤ in order to speed up type checking. +module BrunerieNumLem + (is : GroupIso (coHomGr 4 (S₊ 2 × S₊ 2)) ℤGroup) + (isEq : (fun (fst is) (H²-S²×S²-genₗ ⌣ H²-S²×S²-genᵣ) ≡ 1)) where + + x = H²-S²×S²-genₗ + y = H²-S²×S²-genᵣ + + α = Hopfαfold∘W + β = Hopfβfold∘W + + α' : coHom 2 (Pushout⋁↪fold⋁ (S₊∙ 2)) + α' = fun (fst (coHomCHopfIso 2)) α + + β' : coHom 4 (Pushout⋁↪fold⋁ (S₊∙ 2)) + β' = fun (fst (coHomCHopfIso 4)) β + + rewriteEquation : + (α' ⌣ α' ≡ β' +ₕ β') ⊎ (α' ⌣ α' ≡ -ₕ (β' +ₕ β')) + → (α ⌣ α ≡ β +ₕ β) ⊎ (α ⌣ α ≡ -ₕ (β +ₕ β)) + rewriteEquation (inl x) = + inl ((λ i → leftInv (fst (coHomCHopfIso 2)) α (~ i) + ⌣ leftInv (fst (coHomCHopfIso 2)) α (~ i)) + ∙∙ cong (inv (fst (coHomCHopfIso 4))) x + ∙∙ leftInv (fst (coHomCHopfIso 4)) (β +ₕ β)) + rewriteEquation (inr x) = + inr ((λ i → leftInv (fst (coHomCHopfIso 2)) α (~ i) + ⌣ leftInv (fst (coHomCHopfIso 2)) α (~ i)) + ∙∙ cong (inv (fst (coHomCHopfIso 4))) x + ∙∙ leftInv (fst (coHomCHopfIso 4)) + (-ₕ (β +ₕ β))) + + rewriteEquation2 : (qHom .fst β' ≡ x ⌣ y) ⊎ (qHom .fst β' ≡ -ₕ (x ⌣ y)) + rewriteEquation2 = + ⊎rec + (λ p → inl (sym (leftInv (fst is) (qHom .fst β')) + ∙∙ cong (inv (fst is)) (p ∙ sym isEq) + ∙∙ leftInv (fst is) (x ⌣ y))) + (λ p → inr (sym (leftInv (fst is) (qHom .fst β')) + ∙∙ cong (inv (fst is)) + (p ∙ sym (cong (GroupStr.inv (snd ℤGroup)) isEq)) + ∙∙ (presinv (invGroupIso is .snd) (fun (fst is) (x ⌣ y)) + ∙ cong -ₕ_ (leftInv (fst is) (x ⌣ y))))) + eqs + where + grIso : GroupEquiv (coHomGr 4 (HopfInvariantPush 0 fold∘W)) ℤGroup + grIso = compGroupEquiv (GroupIso→GroupEquiv (coHomCHopfIso 4)) + (compGroupEquiv + isEquiv-qHom + (GroupIso→GroupEquiv is)) + + eqs : (fst (fst grIso) β ≡ 1) ⊎ (fst (fst grIso) β ≡ -1) + eqs = groupEquivPresGen _ (GroupIso→GroupEquiv (Hopfβ-Iso 0 (fold∘W , refl))) + β (inl (Hopfβ↦1 0 (fold∘W , refl))) grIso + + qpres⌣ : (x y : coHom 2 _) + → fst qHom (x ⌣ y) ≡ fst (qHomGen 2) x ⌣ fst (qHomGen 2) y + qpres⌣ = sElim2 (λ _ _ → isSetPathImplicit) λ _ _ → refl + + α'↦x+y : fst (qHomGen 2) α' ≡ x +ₕ y + α'↦x+y = lem ((coHomFun 2 (inv CHopfIso) α)) refl + where + lem : (x' : coHom 2 (Pushout⋁↪fold⋁ (S₊∙ 2))) + → coHomFun 2 inr x' ≡ ∣ ∣_∣ ∣₂ + → fst (qHomGen 2) x' ≡ x +ₕ y + lem = sElim (λ _ → isSetΠ λ _ → isSetPathImplicit) + λ f p → Cubical.HITs.PropositionalTruncation.rec + (squash₂ _ _) + (λ r → cong ∣_∣₂ (funExt (uncurry + (wedgeconFun 1 1 (λ _ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + (λ x → cong f (push (inr x)) ∙∙ funExt⁻ r x ∙∙ refl) + ((λ x → cong f (push (inl x)) ∙∙ funExt⁻ r x ∙∙ sym (rUnitₖ 2 ∣ x ∣ₕ))) + (cong (_∙∙ funExt⁻ r north ∙∙ refl) + (cong (cong f) λ j i → push (push tt j) i)))))) + (fun PathIdTrunc₀Iso p) + + + mainEq : ((fst qHom) (α' ⌣ α') ≡ qHom .fst (β' +ₕ β')) + ⊎ ((fst qHom) (α' ⌣ α') ≡ qHom .fst (-ₕ (β' +ₕ β'))) + mainEq = + ⊎rec + (λ id → inl (lem₁ ∙ lem₂ + ∙ cong (λ x → x +ₕ x) (sym id) + ∙ sym (pres· (snd qHom) β' β'))) + (λ id → inr (lem₁ ∙ lem₂ + ∙ ((sym (distLem (x ⌣ y)) + ∙ cong -ₕ_ (cong (λ x → x +ₕ x) (sym id))) + ∙ cong (-ₕ_) (pres· (snd qHom) β' β')) + ∙ sym (presinv (snd qHom) (β' +ₕ β')))) + rewriteEquation2 + where + triv⌣ : (a : S¹) + → cong₂ (_⌣ₖ_ {n = 2} {m = 2}) (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (merid a)) + ≡ λ _ → ∣ north ∣ₕ + triv⌣ a = + cong₂Funct (_⌣ₖ_ {n = 2} {m = 2}) + (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (merid a)) + ∙ sym (rUnit λ j → _⌣ₖ_ {n = 2} {m = 2} ∣ merid a j ∣ₕ ∣ north ∣) + ∙ (λ i j → ⌣ₖ-0ₖ 2 2 ∣ merid a j ∣ₕ i) + + distLem : (x : coHom 4 (S₊ 2 × S₊ 2)) → -ₕ ((-ₕ x) +ₕ (-ₕ x)) ≡ x +ₕ x + distLem = + sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (funExt λ x → cong -ₖ_ (sym (-distrₖ 4 (f x) (f x))) + ∙ -ₖ^2 (f x +ₖ f x)) + + x⌣x≡0 : x ⌣ x ≡ 0ₕ 4 + x⌣x≡0 = + cong ∣_∣₂ + (funExt (uncurry λ { north y → refl + ; south y → refl + ; (merid a i) y j → triv⌣ a j i})) + + y⌣y≡0 : y ⌣ y ≡ 0ₕ 4 + y⌣y≡0 = + cong ∣_∣₂ + (funExt (uncurry λ { x north → refl + ; x south → refl + ; x (merid a i) j → triv⌣ a j i})) + + -ₕId : (x : coHom 4 (S₊ 2 × S₊ 2)) → (-ₕ^ 2 · 2) x ≡ x + -ₕId = + sElim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ (funExt λ x → -ₖ-gen-inl-left 2 2 tt (inl tt) (f x)) + + y⌣x≡x⌣y : y ⌣ x ≡ x ⌣ y + y⌣x≡x⌣y = + y ⌣ x ≡⟨ gradedComm-⌣ 2 2 y x ⟩ + (-ₕ^ 2 · 2) (transport refl (x ⌣ y)) ≡⟨ -ₕId (transport refl (x ⌣ y)) ⟩ + transport refl (x ⌣ y) ≡⟨ transportRefl (x ⌣ y) ⟩ + x ⌣ y ∎ + + lem₂ : (x +ₕ y) ⌣ (x +ₕ y) ≡ (x ⌣ y) +ₕ (x ⌣ y) + lem₂ = + (x +ₕ y) ⌣ (x +ₕ y) ≡⟨ leftDistr-⌣ 2 2 (x +ₕ y) x y ⟩ + ((x +ₕ y) ⌣ x) +ₕ ((x +ₕ y) ⌣ y) ≡⟨ cong₂ _+ₕ_ (rightDistr-⌣ 2 2 x y x) (rightDistr-⌣ 2 2 x y y) ⟩ + ((x ⌣ x +ₕ y ⌣ x)) +ₕ (x ⌣ y +ₕ y ⌣ y) ≡⟨ cong₂ _+ₕ_ (cong (_+ₕ y ⌣ x) x⌣x≡0 ∙ lUnitₕ 4 (y ⌣ x)) + (cong (x ⌣ y +ₕ_) y⌣y≡0 ∙ rUnitₕ 4 (x ⌣ y)) ⟩ + y ⌣ x +ₕ x ⌣ y ≡⟨ cong (_+ₕ (x ⌣ y)) y⌣x≡x⌣y ⟩ + ((x ⌣ y) +ₕ (x ⌣ y)) ∎ + + lem₁ : (fst qHom) (α' ⌣ α') ≡ (x +ₕ y) ⌣ (x +ₕ y) + lem₁ = + fst qHom (α' ⌣ α') ≡⟨ refl ⟩ + fst (qHomGen 2) α' ⌣ fst (qHomGen 2) α' ≡⟨ cong (λ x → x ⌣ x) α'↦x+y ⟩ + ((x +ₕ y) ⌣ (x +ₕ y)) ∎ + + main⊎ : (HopfInvariant 0 (fold∘W , refl) ≡ 2) + ⊎ (HopfInvariant 0 (fold∘W , refl) ≡ -2) + main⊎ = + ⊎rec (λ p → inl (lem₁ + ∙ cong (fun (fst (Hopfβ-Iso 0 (fold∘W , refl)))) p + ∙ pres· (Hopfβ-Iso 0 (fold∘W , refl) .snd) β β + ∙ cong (λ x → x + x) (Hopfβ↦1 0 (fold∘W , refl)))) + (λ p → inr (lem₁ + ∙ cong (fun (fst (Hopfβ-Iso 0 (fold∘W , refl)))) p + ∙ presinv (Hopfβ-Iso 0 (fold∘W , refl) .snd) (β +ₕ β) + ∙ cong (GroupStr.inv (snd ℤGroup)) + (pres· (Hopfβ-Iso 0 (fold∘W , refl) .snd) β β + ∙ cong (λ x → x + x) (Hopfβ↦1 0 (fold∘W , refl))))) + lem₂ + where + lem₁ : HopfInvariant 0 (fold∘W , refl) + ≡ fun (fst (Hopfβ-Iso 0 (fold∘W , refl))) (α ⌣ α) + lem₁ = + cong (fun (fst (Hopfβ-Iso 0 (fold∘W , refl)))) (transportRefl (α ⌣ α)) + + lem₂ : (α ⌣ α ≡ β +ₕ β) ⊎ (α ⌣ α ≡ -ₕ (β +ₕ β)) + lem₂ = + rewriteEquation + (⊎rec (λ p → inl (sym (retEq (fst isEquiv-qHom) (α' ⌣ α')) + ∙∙ cong (invEq (fst isEquiv-qHom)) p + ∙∙ retEq (fst isEquiv-qHom) (β' +ₕ β'))) + (λ p → inr ((sym (retEq (fst isEquiv-qHom) (α' ⌣ α')) + ∙∙ cong (invEq (fst isEquiv-qHom)) p + ∙∙ retEq (fst isEquiv-qHom) (-ₕ (β' +ₕ β'))))) + mainEq) + + main : abs (HopfInvariant 0 (fold∘W , refl)) ≡ 2 + main = ⊎→abs _ 2 main⊎ + +-- We instantiate the module +Brunerie'≡2 : abs (HopfInvariant 0 (fold∘W , refl)) ≡ 2 +Brunerie'≡2 = BrunerieNumLem.main (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-abs 1 1) Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ-⌣ + +-- We rewrite the it slightly, to get the definition of the Brunerie +-- number in Brunerie (2016) +Brunerie'≡Brunerie : [ ∣ idfun∙ (S₊∙ 2) ∣₂ ∣ ∣ idfun∙ (S₊∙ 2) ∣₂ ]π' ≡ ∣ fold∘W , refl ∣₂ +Brunerie'≡Brunerie = + cong ∣_∣₂ ([]≡[]₂ (idfun∙ (S₊∙ 2)) (idfun∙ (S₊∙ 2)) ) + ∙ sym fold∘W≡Whitehead + ∙ cong ∣_∣₂ (∘∙-idˡ (fold∘W , refl)) + +-- And we get the main result +Brunerie≡2 : Brunerie ≡ 2 +Brunerie≡2 = + cong abs (cong (HopfInvariant-π' 0) Brunerie'≡Brunerie) + ∙ Brunerie'≡2 diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 9c3ef34ee1..05f722d1e4 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -8,6 +8,7 @@ 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 @@ -55,6 +56,107 @@ snd (Ω→ {A = A} {B = B} (f , p)) = ∙∙lCancel p }) (g .fst (doubleCompPath-filler (sym (f .snd)) (cong (f .fst) p) (f .snd) k i)) +Ω→∘∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (g : B →∙ C) (f : A →∙ B) + → Ω→ (g ∘∙ f) ≡ (Ω→ g ∘∙ Ω→ f) +Ω→∘∙ g f = →∙Homogeneous≡ (isHomogeneousPath _ _) (funExt (Ω→∘ g f)) + +Ω→const : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → Ω→ {A = A} {B = B} ((λ _ → pt B) , refl) ≡ ((λ _ → refl) , refl) +Ω→const = →∙Homogeneous≡ (isHomogeneousPath _ _) (funExt λ _ → sym (rUnit _)) + +{- Ω→ is a homomorphism -} +Ω→pres∙filler : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → (p q : typ (Ω A)) + → I → I → I → fst B +Ω→pres∙filler f p q i j k = + hfill + (λ k → λ + { (i = i0) → doubleCompPath-filler (sym (snd f)) (cong (fst f) (p ∙ q)) (snd f) k j + ; (i = i1) → + (doubleCompPath-filler + (sym (snd f)) (cong (fst f) p) (snd f) k + ∙ doubleCompPath-filler + (sym (snd f)) (cong (fst f) q) (snd f) k) j + ; (j = i0) → snd f k + ; (j = i1) → snd f k}) + (inS (cong-∙ (fst f) p q i j)) + k + +Ω→pres∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → (p q : typ (Ω A)) + → fst (Ω→ f) (p ∙ q) ≡ fst (Ω→ f) p ∙ fst (Ω→ f) q +Ω→pres∙ f p q i j = Ω→pres∙filler f p q i j i1 + +Ω→pres∙reflrefl : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Ω→pres∙ {A = A} {B = B} f refl refl + ≡ cong (fst (Ω→ f)) (sym (rUnit refl)) + ∙ snd (Ω→ f) + ∙ rUnit _ + ∙ cong₂ _∙_ (sym (snd (Ω→ f))) (sym (snd (Ω→ f))) +Ω→pres∙reflrefl {A = A} {B = B} = + →∙J (λ b₀ f → Ω→pres∙ {A = A} {B = (fst B , b₀)} f refl refl + ≡ cong (fst (Ω→ f)) (sym (rUnit refl)) + ∙ snd (Ω→ f) + ∙ rUnit _ + ∙ cong₂ _∙_ (sym (snd (Ω→ f))) (sym (snd (Ω→ f)))) + λ f → lem f + ∙ cong (cong (fst (Ω→ (f , refl))) (sym (rUnit refl)) ∙_) + (((lUnit (cong₂ _∙_ (sym (snd (Ω→ (f , refl)))) + (sym (snd (Ω→ (f , refl)))))) + ∙ cong (_∙ (cong₂ _∙_ (sym (snd (Ω→ (f , refl)))) + (sym (snd (Ω→ (f , refl)))))) + (sym (rCancel (snd (Ω→ (f , refl)))))) + ∙ sym (assoc (snd (Ω→ (f , refl))) + (sym (snd (Ω→ (f , refl)))) + (cong₂ _∙_ (sym (snd (Ω→ (f , refl)))) + (sym (snd (Ω→ (f , refl))))))) + where + lem : (f : fst A → fst B) → Ω→pres∙ (f , refl) (λ _ → snd A) (λ _ → snd A) ≡ + (λ i → fst (Ω→ (f , refl)) (rUnit (λ _ → snd A) (~ i))) ∙ + (λ i → snd (Ω→ (f , refl)) (~ i) ∙ snd (Ω→ (f , refl)) (~ i)) + lem f k i j = + hcomp (λ r → λ { (i = i0) → doubleCompPath-filler + refl (cong f ((λ _ → pt A) ∙ refl)) refl (r ∨ k) j + ; (i = i1) → (∙∙lCancel (λ _ → f (pt A)) (~ r) + ∙ ∙∙lCancel (λ _ → f (pt A)) (~ r)) j + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (k = i0) → Ω→pres∙filler {A = A} {B = fst B , f (pt A)} + (f , refl) refl refl i j r + ; (k = i1) → compPath-filler + ((λ i → fst (Ω→ (f , refl)) + (rUnit (λ _ → snd A) (~ i)))) + ((λ i → snd (Ω→ (f , refl)) (~ i) + ∙ snd (Ω→ (f , refl)) (~ i))) r i j}) + (hcomp (λ r → λ { (i = i0) → doubleCompPath-filler refl (cong f (rUnit (λ _ → pt A) r)) refl k j + ; (i = i1) → rUnit (λ _ → f (pt A)) (r ∨ k) j + ; (j = i0) → f (snd A) + ; (j = i1) → f (snd A) + ; (k = i0) → cong-∙∙-filler f (λ _ → pt A) (λ _ → pt A) (λ _ → pt A) r i j + ; (k = i1) → fst (Ω→ (f , refl)) (rUnit (λ _ → snd A) (~ i ∧ r)) j}) + (rUnit (λ _ → f (pt A)) k j)) + +{- Ω^→ is homomorphism -} +Ω^→pres∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → (n : ℕ) + → (p q : typ ((Ω^ (suc n)) A)) + → fst (Ω^→ (suc n) f) (p ∙ q) + ≡ fst (Ω^→ (suc n) f) p ∙ fst (Ω^→ (suc n) f) q +Ω^→pres∙ {A = A} {B = B} f n p q = Ω→pres∙ (Ω^→ n f) p q + +Ω^→∘∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} (n : ℕ) + (g : B →∙ C) (f : A →∙ B) + → Ω^→ n (g ∘∙ f) ≡ (Ω^→ n g ∘∙ Ω^→ n f) +Ω^→∘∙ zero g f = refl +Ω^→∘∙ (suc n) g f = cong Ω→ (Ω^→∘∙ n g f) ∙ Ω→∘∙ (Ω^→ n g) (Ω^→ n f) + +Ω^→const : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → Ω^→ {A = A} {B = B} n ((λ _ → pt B) , refl) + ≡ ((λ _ → snd ((Ω^ n) B)) , refl) +Ω^→const zero = refl +Ω^→const (suc n) = cong Ω→ (Ω^→const n) ∙ Ω→const + isEquivΩ→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → (f : (A →∙ B)) → isEquiv (fst f) → isEquiv (Ω→ f .fst) @@ -65,6 +167,39 @@ isEquivΩ→ {B = (B , b)} = λ eqf → subst isEquiv (funExt (rUnit ∘ cong f)) (isoToIsEquiv (congIso (equivToIso (f , eqf)))) +isEquivΩ^→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (f : A →∙ B) + → isEquiv (fst f) + → isEquiv (Ω^→ n f .fst) +isEquivΩ^→ zero f iseq = iseq +isEquivΩ^→ (suc n) f iseq = isEquivΩ→ (Ω^→ n f) (isEquivΩ^→ n f iseq) + +Ω≃∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (e : A ≃∙ B) + → (Ω A) ≃∙ (Ω B) +fst (fst (Ω≃∙ e)) = fst (Ω→ (fst (fst e) , snd e)) +snd (fst (Ω≃∙ e)) = isEquivΩ→ (fst (fst e) , snd e) (snd (fst e)) +snd (Ω≃∙ e) = snd (Ω→ (fst (fst e) , snd e)) + +Ω≃∙pres∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (e : A ≃∙ B) + → (p q : typ (Ω A)) + → fst (fst (Ω≃∙ e)) (p ∙ q) + ≡ fst (fst (Ω≃∙ e)) p + ∙ fst (fst (Ω≃∙ e)) q +Ω≃∙pres∙ e p q = Ω→pres∙ (fst (fst e) , snd e) p q + +Ω^≃∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (e : A ≃∙ B) + → ((Ω^ n) A) ≃∙ ((Ω^ n) B) +Ω^≃∙ zero e = e +fst (fst (Ω^≃∙ (suc n) e)) = + fst (Ω→ (fst (fst (Ω^≃∙ n e)) , snd (Ω^≃∙ n e))) +snd (fst (Ω^≃∙ (suc n) e)) = + isEquivΩ→ (fst (fst (Ω^≃∙ n e)) , snd (Ω^≃∙ n e)) (snd (fst (Ω^≃∙ n e))) +snd (Ω^≃∙ (suc n) e) = + snd (Ω→ (fst (fst (Ω^≃∙ n e)) , snd (Ω^≃∙ n e))) + ΩfunExtIso : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') → Iso (typ (Ω (A →∙ B ∙))) (A →∙ Ω B) fst (fun (ΩfunExtIso A B) p) x = funExt⁻ (cong fst p) x @@ -137,7 +272,7 @@ EH-gen-r {A = A} n {x = x} {y = y} β i j z = ; (z = i1) → y i1}) (((λ j → refl ∙ β (j ∧ i)) ∙ λ j → refl ∙ β (i ∨ j)) j z) -{- characerisations of EH α β when α or β is refl -} +{- characterisations of EH α β when α or β is refl -} EH-α-refl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (α : typ ((Ω^ (2 + n)) A)) → EH n α refl ≡ sym (rUnit α) ∙ lUnit α diff --git a/Cubical/Homotopy/WedgeConnectivity.agda b/Cubical/Homotopy/WedgeConnectivity.agda index 2a8f451fb2..27f1795c8d 100644 --- a/Cubical/Homotopy/WedgeConnectivity.agda +++ b/Cubical/Homotopy/WedgeConnectivity.agda @@ -2,10 +2,8 @@ module Cubical.Homotopy.WedgeConnectivity where open import Cubical.Foundations.Everything -open import Cubical.Data.HomotopyGroup open import Cubical.Data.Nat open import Cubical.Data.Sigma -open import Cubical.HITs.Nullification open import Cubical.HITs.Susp open import Cubical.HITs.Truncation as Trunc open import Cubical.Homotopy.Connected diff --git a/Cubical/Relation/Binary/Base.agda b/Cubical/Relation/Binary/Base.agda index 19d9264eec..e476fa5fe2 100644 --- a/Cubical/Relation/Binary/Base.agda +++ b/Cubical/Relation/Binary/Base.agda @@ -99,7 +99,7 @@ module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : Rel A A ℓ') where Iso.rightInv i = J (λ y p → cong fst (h aρa (y , J Q (ρ a) p)) ≡ p) (J (λ q _ → cong fst (h aρa (a , q)) ≡ refl) (J (λ α _ → cong fst α ≡ refl) refl - (isContr→isProp (isProp→isContrPath h aρa aρa) refl (h aρa aρa))) + (isProp→isSet h _ _ refl (h _ _))) (sym (JRefl Q (ρ a)))) Iso.leftInv i r = J (λ w β → J Q (ρ a) (cong fst β) ≡ snd w) (JRefl Q (ρ a)) (h aρa (a' , r)) diff --git a/Cubical/Relation/Nullary/DecidablePropositions.agda b/Cubical/Relation/Nullary/DecidablePropositions.agda new file mode 100644 index 0000000000..e7c31443aa --- /dev/null +++ b/Cubical/Relation/Nullary/DecidablePropositions.agda @@ -0,0 +1,18 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Nullary.DecidablePropositions where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Relation.Nullary.Base +open import Cubical.Relation.Nullary.Properties + +private + variable + ℓ : Level + +DecProp : (ℓ : Level) → Type (ℓ-suc ℓ) +DecProp ℓ = Σ[ P ∈ hProp ℓ ] Dec (P .fst) + +isSetDecProp : isSet (DecProp ℓ) +isSetDecProp = isOfHLevelΣ 2 isSetHProp (λ P → isProp→isSet (isPropDec (P .snd))) diff --git a/Cubical/ZCohomology/GroupStructure.agda b/Cubical/ZCohomology/GroupStructure.agda index 06abedb3fd..43aff7ccfd 100644 --- a/Cubical/ZCohomology/GroupStructure.agda +++ b/Cubical/ZCohomology/GroupStructure.agda @@ -670,6 +670,16 @@ snd (coHomMorph n f) = makeIsGroupHom (helper n) helper (suc zero) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl helper (suc (suc n)) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl +coHomIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) → Iso A B + → GroupIso (coHomGr n B) (coHomGr n A) +fun (fst (coHomIso n is)) = fst (coHomMorph n (fun is)) +inv' (fst (coHomIso n is)) = fst (coHomMorph n (inv' is)) +rightInv (fst (coHomIso n is)) = + sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (leftInv is x)) +leftInv (fst (coHomIso n is)) = + sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (rightInv is x)) +snd (coHomIso n is) = snd (coHomMorph n (fun is)) + -- Alternative definition of cohomology using ΩKₙ instead. Useful for breaking proofs of group isos -- up into smaller parts coHomGrΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Group ℓ diff --git a/Cubical/ZCohomology/Groups/SphereProduct.agda b/Cubical/ZCohomology/Groups/SphereProduct.agda new file mode 100644 index 0000000000..f00a5c9d4f --- /dev/null +++ b/Cubical/ZCohomology/Groups/SphereProduct.agda @@ -0,0 +1,351 @@ +{- +This file shows Hⁿ⁺ᵐ(Sⁿ×Sᵐ)≅ℤ for n, m ≥ 1. This can be done in several +ways (e.g. Mayer-Vietoris). We give the iso as explicitly +as possible. The idea: +Step 1: Hⁿ⁺ᵐ(Sⁿ×Sᵐ)≅H¹⁺ⁿ⁺ᵐ(S¹⁺ⁿ×Sᵐ) (i.e. some form of suspension +axiom). +Step 2: H¹⁺ᵐ(S¹×Sᵐ)≅ℤ +Step 3: Conclude Hⁿ⁺ᵐ(Sⁿ×Sᵐ)≅ℤ by induction on n. + +The iso as defined here has the advantage that it looks a lot like +the cup product, making it trivial to show that e.g. gₗ ⌣ gᵣ ≡ e, where +gₗ, gᵣ are the two generators of H²(S²×S²) and e is the generator of +H⁴(S²×S²) -- this of interest since it is used in π₄S³≅ℤ/2. +-} + +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.SphereProduct where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Unit + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation +open import Cubical.HITs.SetTruncation + renaming (elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (map to pMap ; rec to pRec) + +open import Cubical.Algebra.Group + +open import Cubical.Relation.Nullary + +open import Cubical.Homotopy.Loopspace + +open Iso + +private + ¬lem : (n m : ℕ) → ¬ suc (n + m) ≡ m + ¬lem n zero = snotz + ¬lem n (suc m) p = + ¬lem n m (sym (+-suc n m) ∙ cong predℕ p) + +{- Step 1: Hⁿ⁺ᵐ(Sⁿ×Sᵐ)≅H¹⁺ⁿ⁺ᵐ(S¹⁺ⁿ×Sᵐ) -} +-- The functions back and forth (suspension on the left component) +↓Sⁿ×Sᵐ→Kₙ₊ₘ : (n m : ℕ) + → (S₊ (suc (suc n)) → S₊ (suc m) → coHomK (suc (suc ((suc n) + m)))) + → S₊ (suc n) → S₊ (suc m) → coHomK (suc (suc (n + m))) +↓Sⁿ×Sᵐ→Kₙ₊ₘ n m f x y = + ΩKn+1→Kn ((suc ((suc n) + m))) + (sym (rCancelₖ _ (f north y)) + ∙∙ cong (λ x → f x y -ₖ f north y) (merid x ∙ sym (merid (ptSn (suc n)))) + ∙∙ rCancelₖ _ (f north y)) + +↑Sⁿ×Sᵐ→Kₙ₊ₘ : (n m : ℕ) + → (S₊ (suc n) → S₊ (suc m) → coHomK (suc (suc (n + m)))) + → (S₊ (suc (suc n)) → S₊ (suc m) → coHomK (suc (suc ((suc n) + m)))) +↑Sⁿ×Sᵐ→Kₙ₊ₘ n m f north y = 0ₖ _ +↑Sⁿ×Sᵐ→Kₙ₊ₘ n m f south y = 0ₖ _ +↑Sⁿ×Sᵐ→Kₙ₊ₘ n m f (merid a i) y = + Kn→ΩKn+1 _ (f a y) i + +∥HⁿSᵐPath∥ : (n m : ℕ) → (f : S₊ (suc m) → coHomK (suc n)) + → ¬ (n ≡ m) + → ∥ f ≡ (λ _ → 0ₖ (suc n)) ∥ +∥HⁿSᵐPath∥ n m f p = + fun PathIdTrunc₀Iso + (isContr→isProp + (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Sᵐ≅0 n m p)) isContrUnit) + ∣ f ∣₂ (0ₕ _)) + +-- The iso +×leftSuspensionIso : (n m : ℕ) + → GroupIso (coHomGr (suc (((suc n) + m))) + (S₊ (suc n) × S₊ (suc m))) + (coHomGr (suc (suc ((suc n) + m))) + (S₊ (suc (suc n)) × S₊ (suc m))) +fun (fst (×leftSuspensionIso n m)) = + sMap (uncurry ∘ ↑Sⁿ×Sᵐ→Kₙ₊ₘ n m ∘ curry) +inv (fst (×leftSuspensionIso n m)) = + sMap ((uncurry ∘ ↓Sⁿ×Sᵐ→Kₙ₊ₘ n m ∘ curry)) +rightInv (fst (×leftSuspensionIso n m)) = + sElim (λ _ → isSetPathImplicit) + λ f → inv PathIdTrunc₀Iso + (pRec squash + (uncurry (λ g p + → pMap (λ gid → funExt λ {(x , y) + → (λ i → uncurry (↑Sⁿ×Sᵐ→Kₙ₊ₘ n m + (↓Sⁿ×Sᵐ→Kₙ₊ₘ n m (curry (p (~ i))))) (x , y)) + ∙∙ main g gid x y + ∙∙ funExt⁻ p (x , y)}) + (∥Path∥ g))) + (rewrte f)) + where + lem : (n m : ℕ) → ¬ suc (suc (n + m)) ≡ m + lem n zero = snotz + lem n (suc m) p = lem n m (cong suc (sym (+-suc n m)) ∙ cong predℕ p) + + charac-fun : + (S₊ (suc n) → S₊ (suc m) + → typ (Ω (coHomK-ptd (suc (suc (suc n + m)))))) + → S₊ (suc (suc n)) × S₊ (suc m) → coHomK (suc (suc (suc n + m))) + charac-fun g (north , y) = 0ₖ _ + charac-fun g (south , y) = 0ₖ _ + charac-fun g (merid c i , y) = g c y i + + rewrte : (f : S₊ (suc (suc n)) × S₊ (suc m) + → coHomK (suc (suc (suc n + m)))) + → ∃[ g ∈ (S₊ (suc n) → S₊ (suc m) + → typ (Ω (coHomK-ptd (suc (suc (suc n + m)))))) ] + charac-fun g ≡ f + rewrte f = + pMap (λ p + → (λ x y → sym (funExt⁻ p y) + ∙∙ (λ i → f ((merid x ∙ sym (merid (ptSn (suc n)))) i , y)) + ∙∙ funExt⁻ p y) + , funExt λ { (north , y) → sym (funExt⁻ p y) + ; (south , y) → sym (funExt⁻ p y) + ∙ λ i → f (merid (ptSn (suc n)) i , y) + ; (merid a i , y) j + → hcomp (λ k → λ { (i = i0) → p (~ j ∧ k) y + ; (i = i1) + → compPath-filler' + (sym (funExt⁻ p y)) + (λ i → f (merid (ptSn (suc n)) i , y)) k j + ; (j = i1) → f (merid a i , y) }) + (f (compPath-filler (merid a) + (sym (merid (ptSn (suc n)))) (~ j) i + , y))}) + (∥HⁿSᵐPath∥ _ _ (λ x → f (north , x)) + (lem n m)) + + ∥Path∥ : + (g : S₊ (suc n) → S₊ (suc m) + → typ (Ω (coHomK-ptd (suc (suc (suc n + m)))))) + → ∥ (g (ptSn _)) ≡ (λ _ → refl) ∥ + ∥Path∥ g = + fun PathIdTrunc₀Iso + (isContr→isProp + (isOfHLevelRetractFromIso 0 + ((invIso (fst (coHom≅coHomΩ _ (S₊ (suc m)))))) + (0ₕ _ , sElim (λ _ → isProp→isSet (squash₂ _ _)) + λ f → pRec (squash₂ _ _) (sym ∘ cong ∣_∣₂) + (∥HⁿSᵐPath∥ _ _ f (¬lem n m)))) + ∣ g (ptSn (suc n)) ∣₂ ∣ (λ _ → refl) ∣₂) + + main : (g : _) → (g (ptSn _)) ≡ (λ _ → refl) + → (x : S₊ (suc (suc n))) (y : S₊ (suc m)) + → uncurry (↑Sⁿ×Sᵐ→Kₙ₊ₘ n m + (↓Sⁿ×Sᵐ→Kₙ₊ₘ n m (curry (charac-fun g)))) (x , y) + ≡ charac-fun g (x , y) + main g gid north y = refl + main g gid south y = refl + main g gid (merid a i) y j = helper j i + where + helper : (λ i → uncurry (↑Sⁿ×Sᵐ→Kₙ₊ₘ n m + (↓Sⁿ×Sᵐ→Kₙ₊ₘ n m (curry (charac-fun g)))) + (merid a i , y)) + ≡ g a y + helper = (λ i → rightInv (Iso-Kn-ΩKn+1 _) + ((sym (rCancel≡refl _ i) + ∙∙ cong-∙ (λ x → rUnitₖ _ (charac-fun g (x , y)) i) + (merid a) (sym (merid (ptSn (suc n)))) i + ∙∙ rCancel≡refl _ i)) i) + ∙∙ sym (rUnit _) + ∙∙ (cong (g a y ∙_) (cong sym (funExt⁻ gid y)) + ∙ sym (rUnit (g a y))) +leftInv (fst (×leftSuspensionIso n m)) = + sElim (λ _ → isSetPathImplicit) + λ f → pRec (squash₂ _ _) + (λ id + → cong ∣_∣₂ + (funExt (λ {(x , y) → (λ i → ΩKn+1→Kn _ (sym (rCancel≡refl _ i) + ∙∙ ((λ j → rUnitₖ _ + (↑Sⁿ×Sᵐ→Kₙ₊ₘ n m (curry f) + ((merid x ∙ sym (merid (ptSn (suc n)))) j) y) i)) + ∙∙ rCancel≡refl _ i)) + ∙ (λ i → ΩKn+1→Kn _ (rUnit + (cong-∙ (λ r → ↑Sⁿ×Sᵐ→Kₙ₊ₘ n m (curry f) r y) + (merid x) (sym (merid (ptSn (suc n)))) i) (~ i))) + ∙ cong (ΩKn+1→Kn _) (cong (Kn→ΩKn+1 _ (f (x , y)) ∙_) + (cong sym (cong (Kn→ΩKn+1 _) + (funExt⁻ id y) ∙ (Kn→ΩKn+10ₖ _))) + ∙ sym (rUnit _)) + ∙ leftInv (Iso-Kn-ΩKn+1 _) (f (x , y))}))) + (∥HⁿSᵐPath∥ (suc n + m) m (λ x → f (ptSn _ , x)) + (¬lem n m)) +snd (×leftSuspensionIso n m) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ + (funExt λ { (north , y) → refl + ; (south , y) → refl + ; (merid a i , y) j + → (sym (∙≡+₂ _ (Kn→ΩKn+1 _ (f (a , y))) + (Kn→ΩKn+1 _ (g (a , y)))) + ∙ sym (Kn→ΩKn+1-hom _ + (f (a , y)) (g (a , y)))) (~ j) i})) + +{- Step 2: H¹⁺ᵐ(S¹×Sᵐ)≅ℤ -} + +-- The functions back and forth +Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ : (m : ℕ) + → (S₊ (suc m) → coHomK (suc m)) + → S₊ (suc zero) → S₊ (suc m) → coHomK (suc (suc m)) +Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m f base y = 0ₖ _ +Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m f (loop i) y = Kn→ΩKn+1 (suc m) (f y) i + +Hⁿ-Sⁿ→Hⁿ-S¹×Sⁿ : (m : ℕ) + → (S₊ (suc zero) → S₊ (suc m) → coHomK (suc (suc m))) + → (S₊ (suc m) → coHomK (suc m)) +Hⁿ-Sⁿ→Hⁿ-S¹×Sⁿ m f x = + ΩKn+1→Kn (suc m) + (sym (rCancelₖ _ (f base x)) + ∙∙ ((λ i → f (loop i) x -ₖ f base x)) + ∙∙ rCancelₖ _ (f base x)) + +Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ : (m : ℕ) + → GroupIso (coHomGr (suc m) (S₊ (suc m))) + (coHomGr (suc (suc m)) (S₊ (suc zero) × S₊ (suc m))) +fun (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = sMap (uncurry ∘ Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m) +inv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = sMap (Hⁿ-Sⁿ→Hⁿ-S¹×Sⁿ m ∘ curry) +rightInv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = + sElim (λ _ → isSetPathImplicit) + λ f → inv PathIdTrunc₀Iso + (pMap (uncurry (λ g p + → funExt λ {(x , y) + → (λ i → uncurry + (Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m + (Hⁿ-Sⁿ→Hⁿ-S¹×Sⁿ m (curry (p i)))) (x , y)) + ∙∙ main g x y + ∙∙ sym (funExt⁻ p (x , y))})) + (rewrte f)) + where + characFun : (x : S₊ (suc m) → coHomK (suc m)) + → S₊ 1 × S₊ (suc m) → coHomK (suc (suc m)) + characFun f (base , y) = 0ₖ _ + characFun f (loop i , y) = Kn→ΩKn+1 _ (f y) i + + main : (g : _) (x : S¹) (y : S₊ (suc m)) + → uncurry (Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m (Hⁿ-Sⁿ→Hⁿ-S¹×Sⁿ m (curry (characFun g)))) + (x , y) + ≡ characFun g (x , y) + main g base y = refl + main g (loop i) y j = help j i + where + help : cong (λ x → Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m + (Hⁿ-Sⁿ→Hⁿ-S¹×Sⁿ m (curry (characFun g))) x y) loop + ≡ Kn→ΩKn+1 _ (g y) + help = rightInv (Iso-Kn-ΩKn+1 (suc m)) + (sym (rCancelₖ _ (0ₖ _)) + ∙∙ ((λ i → Kn→ΩKn+1 _ (g y) i -ₖ 0ₖ _)) + ∙∙ rCancelₖ _ (0ₖ _)) + ∙∙ (λ i → sym (rCancel≡refl _ i) + ∙∙ (λ j → rUnitₖ _ (Kn→ΩKn+1 _ (g y) j) i) + ∙∙ rCancel≡refl _ i) + ∙∙ sym (rUnit _) + + lem : (m : ℕ) → ¬ suc m ≡ m + lem zero = snotz + lem (suc m) p = lem m (cong predℕ p) + + rewrte : (f : S₊ 1 × S₊ (suc m) → coHomK (suc (suc m))) + → ∃[ g ∈ (S₊ (suc m) → coHomK (suc m)) ] f ≡ characFun g + rewrte f = + pMap (λ p → + (λ x → ΩKn+1→Kn _ + (sym (funExt⁻ p x) ∙∙ (λ i → f (loop i , x)) ∙∙ funExt⁻ p x)) + , funExt λ { (base , y) → funExt⁻ p y + ; (loop i , x) j → + hcomp (λ k → λ {(i = i0) → funExt⁻ p x j + ; (i = i1) → funExt⁻ p x j + ; (j = i0) → f (loop i , x) + ; (j = i1) → + rightInv (Iso-Kn-ΩKn+1 (suc m)) + (sym (funExt⁻ p x) + ∙∙ (λ i → f (loop i , x)) + ∙∙ funExt⁻ p x) (~ k) i}) + (doubleCompPath-filler + (sym (funExt⁻ p x)) + (λ i → f (loop i , x)) + (funExt⁻ p x) j i)}) + (∥HⁿSᵐPath∥ (suc m) m (λ x → f (base , x)) (lem m)) +leftInv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = + sElim (λ _ → isSetPathImplicit) + λ f + → cong ∣_∣₂ (funExt λ x + → cong (ΩKn+1→Kn (suc m)) + ((λ i → sym (rCancel≡refl _ i) + ∙∙ cong (λ z → rUnitₖ _ (Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m f z x) i) loop + ∙∙ rCancel≡refl _ i) ∙ sym (rUnit (Kn→ΩKn+1 (suc m) (f x)))) + ∙ leftInv (Iso-Kn-ΩKn+1 _) (f x)) +snd (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m) = + makeIsGroupHom + (sElim2 + (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ + (funExt λ { (base , y) → refl + ; (loop i , y) j → + (sym (∙≡+₂ _ (Kn→ΩKn+1 _ (f y)) (Kn→ΩKn+1 _ (g y))) + ∙ sym (Kn→ΩKn+1-hom _ (f y) (g y))) (~ j) i})) + +Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ : (n m : ℕ) + → GroupIso (coHomGr ((suc n) +' (suc m)) + (S₊ (suc n) × S₊ (suc m))) + ℤ +Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ zero m = + compGroupIso (invGroupIso (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) (Hⁿ-Sⁿ≅ℤ m) +Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ (suc n) m = + compGroupIso + (invGroupIso (×leftSuspensionIso n m)) + (Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ n m) + +-- Proof that ⌣ respects generators for H²×H²→H⁴ +-- Todo: generalise + +H²-S²×S²-genₗ : coHom 2 (S₊ 2 × S₊ 2) +H²-S²×S²-genₗ = ∣ (λ x → ∣ fst x ∣) ∣₂ + +H²-S²×S²-genᵣ : coHom 2 (S₊ 2 × S₊ 2) +H²-S²×S²-genᵣ = ∣ (λ x → ∣ snd x ∣) ∣₂ + +H²-S²≅H⁴-S²×S² : GroupIso (coHomGr 2 (S₊ 2)) (coHomGr 4 (S₊ 2 × S₊ 2)) +H²-S²≅H⁴-S²×S² = (compGroupIso (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ 1) (×leftSuspensionIso 0 1)) + +H²-S²≅H⁴-S²×S²⌣ : + fun (fst (H²-S²≅H⁴-S²×S²)) ∣ ∣_∣ₕ ∣₂ ≡ H²-S²×S²-genₗ ⌣ H²-S²×S²-genᵣ +H²-S²≅H⁴-S²×S²⌣ = + cong ∣_∣₂ (funExt (uncurry + (λ { north y → refl + ; south y → refl + ; (merid a i) y j → Kn→ΩKn+1 3 (lem a y j) i}))) + where + lem : (a : S¹) (y : S₊ 2) + → Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ 1 ∣_∣ₕ a y ≡ _⌣ₖ_ {n = 1} {m = 2} ∣ a ∣ₕ ∣ y ∣ₕ + lem base y = refl + lem (loop i) y = refl diff --git a/Cubical/ZCohomology/Groups/Unit.agda b/Cubical/ZCohomology/Groups/Unit.agda index 6c0350f737..ce5b70ec43 100644 --- a/Cubical/ZCohomology/Groups/Unit.agda +++ b/Cubical/ZCohomology/Groups/Unit.agda @@ -1,22 +1,29 @@ {-# OPTIONS --safe #-} module Cubical.ZCohomology.Groups.Unit where +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism + open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.Properties open import Cubical.ZCohomology.GroupStructure + open import Cubical.HITs.Sn -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism + open import Cubical.HITs.Susp -open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) -open import Cubical.HITs.Nullification +open import Cubical.HITs.Truncation +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) + open import Cubical.Data.Int hiding (ℤ ; _+_ ; +Comm) open import Cubical.Data.Nat -open import Cubical.HITs.Truncation -open import Cubical.Homotopy.Connected open import Cubical.Data.Unit + +open import Cubical.Homotopy.Connected + open import Cubical.Data.Sigma open import Cubical.Algebra.Group renaming (Unit to UnitGroup) @@ -71,10 +78,10 @@ rightInv (fst (Hⁿ-contrType≅0 n contr)) _ = refl leftInv (fst (Hⁿ-contrType≅0 {A = A} n contr)) _ = isOfHLevelSuc 0 helper _ _ where helper : isContr (coHom (suc n) A) - helper = (Iso.inv (Hⁿ-contrTypeIso n contr) (0ₕ (suc n))) - , λ y → cong (Iso.inv (Hⁿ-contrTypeIso n contr)) - (isOfHLevelSuc 0 (isContrHⁿ-Unit n) (0ₕ (suc n)) (Iso.fun (Hⁿ-contrTypeIso n contr) y)) - ∙ Iso.leftInv (Hⁿ-contrTypeIso n contr) y + helper = (inv (Hⁿ-contrTypeIso n contr) (0ₕ (suc n))) + , λ y → cong (inv (Hⁿ-contrTypeIso n contr)) + (isOfHLevelSuc 0 (isContrHⁿ-Unit n) (0ₕ (suc n)) (fun (Hⁿ-contrTypeIso n contr) y)) + ∙ leftInv (Hⁿ-contrTypeIso n contr) y snd (Hⁿ-contrType≅0 n contr) = makeIsGroupHom λ _ _ → refl isContr-HⁿRed-Unit : (n : ℕ) → isContr (coHomRed n (Unit , tt)) @@ -83,3 +90,12 @@ snd (isContr-HⁿRed-Unit n) = sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ {(f , p) → cong ∣_∣₂ (ΣPathP (funExt (λ _ → sym p) , λ i j → p (~ i ∨ j)))} + +×rUnitCohomIso : ∀ {ℓ} {A : Type ℓ} (n : ℕ) + → GroupIso (×coHomGr (suc n) A Unit) (coHomGr (suc n) A) +fun (fst (×rUnitCohomIso n)) = fst +inv (fst (×rUnitCohomIso n)) x = x , 0ₕ (suc n) +rightInv (fst (×rUnitCohomIso n)) _ = refl +leftInv (fst (×rUnitCohomIso n)) x = + ΣPathP (refl , isContr→isProp (isContrHⁿ-Unit n) (0ₕ (suc n)) (snd x)) +snd (×rUnitCohomIso n) = makeIsGroupHom λ _ _ → refl diff --git a/Cubical/ZCohomology/RingStructure/RingLaws.agda b/Cubical/ZCohomology/RingStructure/RingLaws.agda index 3abbbe9036..701b805d05 100644 --- a/Cubical/ZCohomology/RingStructure/RingLaws.agda +++ b/Cubical/ZCohomology/RingStructure/RingLaws.agda @@ -1,5 +1,6 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.ZCohomology.RingStructure.RingLaws where + open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Properties @@ -668,4 +669,19 @@ rUnit⌣ (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → cong ∣_∣₂ (funExt λ x → -Distᵣ n m (f x) (g x)) +-ₕDistₗ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) + (x : coHom n A) (y : coHom m A) → (-ₕ (x ⌣ y)) ≡ (-ₕ x) ⌣ y +-ₕDistₗ n m = + sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → cong ∣_∣₂ (funExt λ x → -Distₗ n m (f x) (g x)) + +-ₕDistₗᵣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) + (x : coHom n A) (y : coHom m A) → (-ₕ x) ⌣ (-ₕ y) ≡ x ⌣ y +-ₕDistₗᵣ n m x y = + sym (-ₕDistₗ n m x (-ₕ y)) + ∙∙ cong -ₕ_ (sym (-ₕDistᵣ n m x y)) + ∙∙ sElim2 {C = λ x y → (-ₕ (-ₕ (x ⌣ y))) ≡ x ⌣ y} + (λ _ _ → isSetPathImplicit) + (λ f g → cong ∣_∣₂ (funExt λ _ → -ₖ^2 _)) x y + -- TODO : Graded ring structure diff --git a/INSTALL.md b/INSTALL.md index bfcdb881ec..35bc9310b0 100644 --- a/INSTALL.md +++ b/INSTALL.md @@ -1,7 +1,7 @@ Installation of agda/cubical ============================ -The cubical library should compile on the latest development version +The cubical library should compile on the latest official release of Agda: https://github.com/agda/agda @@ -83,10 +83,9 @@ Once this works go to a suitable directory and run ``` > git clone https://github.com/agda/cubical > cd cubical -> make ``` -This should compile all of the agda/cubical files. To test that it +You can additionally run `make` to typecheck all of the agda/cubical files, otherwise they will be checked as needed. To test that it works in emacs run ``` @@ -164,10 +163,9 @@ Once this works go to a suitable directory and run ``` > git clone https://github.com/agda/cubical > cd cubical -> make ``` -This should compile all of the agda/cubical files. To test that it +You can additionally run `make` to typecheck all of the agda/cubical files, otherwise they will be checked as needed. To test that it works in emacs run ``` @@ -248,10 +246,9 @@ Once this works go to a suitable directory and run ``` > git clone https://github.com/agda/cubical > cd cubical -> make ``` -This should compile all of the agda/cubical files. To test that it +You can additionally run `make` to typecheck all of the agda/cubical files, otherwise they will be checked as needed. To test that it works in emacs run ``` diff --git a/README.md b/README.md index 89e281fd6f..9ba817846e 100644 --- a/README.md +++ b/README.md @@ -1,25 +1,27 @@ A standard library for Cubical Agda =================================== -This library compiles with the master branch of the development -version of [Agda](https://github.com/agda/agda/). For detailed install +This library compiles with the latest official release of +[Agda](https://github.com/agda/agda/). For detailed install instructions see the [INSTALL](https://github.com/agda/cubical/blob/master/INSTALL.md) file. -If you want to use Agda 2.6.2 instead of the latest development version, you +The source code has a glorious clickable [rendered version](https://agda.github.io/cubical/Cubical.README.html). + +If you want to use Agda 2.6.2 instead of the latest release version, you can check out the tag `v0.3` of this library. -If you want to use Agda 2.6.1.3 instead of the latest development version, you +If you want to use Agda 2.6.1.3 instead of the latest release version, you can check out the tag `v0.2` of this library. -If you want to use Agda 2.6.0.1 instead of the latest development version, you +If you want to use Agda 2.6.0.1 instead of the latest release version, you can check out the tag `v0.1` of this library. For some introductory lecture notes see the material for the Cubical Agda course of the [EPIT 2021 spring school](https://github.com/HoTT/EPIT-2020/blob/main/04-cubical-type-theory/). -For a paper on with detials about Cubical Agda, see [Cubical Agda: a dependently typed +For a paper on with details about Cubical Agda, see [Cubical Agda: a dependently typed programming language with univalence and higher inductive types](https://dl.acm.org/doi/10.1145/3341691) by Andrea Vezzosi, Anders Mörtberg, and Andreas Abel. From 2abee65a80dc7fb9d6450995bf013995d4b62f49 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Tue, 15 Mar 2022 22:27:44 +0800 Subject: [PATCH 14/24] comment --- Cubical/Data/FinSet/DecidablePredicate.agda | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Cubical/Data/FinSet/DecidablePredicate.agda b/Cubical/Data/FinSet/DecidablePredicate.agda index f67d9738bf..63ed9aaf71 100644 --- a/Cubical/Data/FinSet/DecidablePredicate.agda +++ b/Cubical/Data/FinSet/DecidablePredicate.agda @@ -1,6 +1,8 @@ {- -This files contains lots of useful properties about decidable predicates on finite sets +This files contains: +- An alternative formulation of decidable propositions; +- Lots of useful properties about (this) decidable predicates on finite sets. -} {-# OPTIONS --safe #-} From 41a494930bf43f82834d4841397863998b27c376 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Wed, 16 Mar 2022 10:34:28 +0800 Subject: [PATCH 15/24] comment --- Cubical/Data/FinSet/Base.agda | 1 + Cubical/Data/FinSet/FiniteStructure.agda | 15 ++++++++++++--- Cubical/Data/FinSet/FiniteType.agda | 14 ++++++++++---- 3 files changed, 23 insertions(+), 7 deletions(-) diff --git a/Cubical/Data/FinSet/Base.agda b/Cubical/Data/FinSet/Base.agda index 9cf0d4138e..40dce36580 100644 --- a/Cubical/Data/FinSet/Base.agda +++ b/Cubical/Data/FinSet/Base.agda @@ -30,6 +30,7 @@ private -- definition of (Bishop) finite sets +-- this definition makes cardinality computation more efficient isFinSet : Type ℓ → Type ℓ isFinSet A = Σ[ n ∈ ℕ ] ∥ A ≃ Fin n ∥ diff --git a/Cubical/Data/FinSet/FiniteStructure.agda b/Cubical/Data/FinSet/FiniteStructure.agda index b50419e3ed..dc89514f38 100644 --- a/Cubical/Data/FinSet/FiniteStructure.agda +++ b/Cubical/Data/FinSet/FiniteStructure.agda @@ -1,7 +1,13 @@ {- -Definition and properties of finite sets equipped with finite structures, -namely the type of structures over a given finte set is a finite set. +Finite Structures over Finite Set + +In short, the type of structures should be finite set itself. + +This file contains: +- Definition and properties of finite sets equipped with finite structures; +- The type of finitely-structured finite sets is Rijke finite, + so that we can count their number up to equivalence/isomorphism. -} {-# OPTIONS --safe #-} @@ -36,7 +42,7 @@ private FinSetWithStr : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') FinSetWithStr ℓ S = Σ[ X ∈ FinSet ℓ ] S X .fst --- type finite sets of a given cardinal +-- type of finite sets with a fixed cardinal FinSetOfCard : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) FinSetOfCard ℓ n = Σ[ X ∈ FinSet ℓ ] (card X ≡ n) @@ -63,6 +69,9 @@ isFinTypeFinSetOfCard .fst = isPathConnected→isFinType0 isPathConnectedFinSetO isFinTypeFinSetOfCard .snd X Y = isFinSet→isFinType 0 (EquivPresIsFinSet (FinSet≡ _ _ ⋆ FinSetOfCard≡ _ _) (isFinSetType≡Eff (X .fst) (Y .fst))) +-- the type of finitely-structured finite sets is Rijke finite +-- in particular, we can count the number of them up to equivalence + isFinTypeFinSetWithStrOfCard : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') (n : ℕ) → isFinType 0 (FinSetWithStrOfCard ℓ S n) isFinTypeFinSetWithStrOfCard ℓ S n = diff --git a/Cubical/Data/FinSet/FiniteType.agda b/Cubical/Data/FinSet/FiniteType.agda index a9c07ee248..c63c59b41d 100644 --- a/Cubical/Data/FinSet/FiniteType.agda +++ b/Cubical/Data/FinSet/FiniteType.agda @@ -1,12 +1,18 @@ {- -Definition and properties of finite types +Finite Types -This file formalize the notion of Rijke finite type, which is a direct generalization -of Bishop finite set. Basically, a type is (Rijke) n-finite if all its i-th order +This file formalize the notion of Rijke finite type, +which is a direct generalization of Bishop finite set. +Basically, a type is (Rijke) n-finite if its i-th order homotopy groups πᵢ for i ≤ n are Bishop finite. -https://github.com/EgbertRijke/OEIS-A000001 +Referrences: + https://github.com/EgbertRijke/OEIS-A000001 + +This file contains: +- The definition and basic properties of Rijke finite types; +- Rijke finiteness is closed under forming Σ-type. -} {-# OPTIONS --safe #-} From c2c7f9f26af8254454fbc54629e57bd5f5d8c9f2 Mon Sep 17 00:00:00 2001 From: ecavallo Date: Fri, 25 Mar 2022 19:04:46 +0100 Subject: [PATCH 16/24] simplify --- Cubical/Data/FinSet/DecidablePredicate.agda | 4 + Cubical/Data/FinSet/FiniteType.agda | 15 +- Cubical/HITs/SetTruncation/Fibers.agda | 312 ++++++-------------- 3 files changed, 107 insertions(+), 224 deletions(-) diff --git a/Cubical/Data/FinSet/DecidablePredicate.agda b/Cubical/Data/FinSet/DecidablePredicate.agda index 63ed9aaf71..0dc80a6caa 100644 --- a/Cubical/Data/FinSet/DecidablePredicate.agda +++ b/Cubical/Data/FinSet/DecidablePredicate.agda @@ -49,6 +49,10 @@ isPropIsDecProp p q = DecProp : (ℓ : Level) → Type (ℓ-suc ℓ) DecProp ℓ = Σ[ P ∈ Type ℓ ] isDecProp P +isDecPropRespectEquiv : {P : Type ℓ} {Q : Type ℓ'} + → P ≃ Q → isDecProp Q → isDecProp P +isDecPropRespectEquiv e (t , e') = t , e ⋆ e' + module _ (X : Type ℓ)(p : isFinOrd X) where diff --git a/Cubical/Data/FinSet/FiniteType.agda b/Cubical/Data/FinSet/FiniteType.agda index c63c59b41d..3b8b64416a 100644 --- a/Cubical/Data/FinSet/FiniteType.agda +++ b/Cubical/Data/FinSet/FiniteType.agda @@ -20,6 +20,7 @@ This file contains: module Cubical.Data.FinSet.FiniteType 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.Isomorphism @@ -103,7 +104,7 @@ isPathConnected→isFinType0 p = isContr→isFinSet p module _ (X : Type ℓ ) - (Y : Type ℓ')(h : isFinType 1 Y) + (Y : Type ℓ') (h : isFinType 1 Y) (f : X → Y) (q : (y : Y) → isFinType 0 (fiber f y)) where @@ -113,19 +114,21 @@ module _ module _ (y : Y) where - isDecPropFiberRel : (x x' : ∥ fiber f y ∥₂) → isDecProp (fiberRel2 _ _ f y x x') - isDecPropFiberRel x x' = isDecProp∃ (_ , h .snd y y) (λ _ → _ , isDecProp≡ (_ , q y) _ _) + isDecPropFiberRel : (x x' : ∥ fiber f y ∥₂) → isDecProp (fiberRel f y x x') + isDecPropFiberRel x x' = + isDecPropRespectEquiv (fiberRel1≃2 f y x x') + (isDecProp∃ (_ , h .snd y y) (λ _ → _ , isDecProp≡ (_ , q y) _ _)) isFinSetFiber∥∥₂' : isFinSet (fiber ∥f∥₂ ∣ y ∣₂) isFinSetFiber∥∥₂' = - EquivPresIsFinSet (∥fiber∥₂/Rel≃fiber∥∥₂ _ _ f y) - (isFinSetQuot (_ , q y) (fiberRel2 _ _ _ _) (isEquivRelFiberRel _ _ _ _) isDecPropFiberRel) + EquivPresIsFinSet (∥fiber∥₂/R≃fiber∥∥₂ f y) + (isFinSetQuot (_ , q y) (fiberRel f y) (isEquivRelFiberRel f y) isDecPropFiberRel) isFinSetFiber∥∥₂ : (y : ∥ Y ∥₂) → isFinSet (fiber ∥f∥₂ y) isFinSetFiber∥∥₂ = Set.elim (λ _ → isProp→isSet isPropIsFinSet) isFinSetFiber∥∥₂' isFinType0Total : isFinType 0 X - isFinType0Total = isFinSetTotal _ (_ , h .fst) ∥f∥₂ isFinSetFiber∥∥₂ + isFinType0Total = isFinSetTotal ∥ X ∥₂ (∥ Y ∥₂ , h .fst) ∥f∥₂ isFinSetFiber∥∥₂ module _ (X : FinType ℓ 1) diff --git a/Cubical/HITs/SetTruncation/Fibers.agda b/Cubical/HITs/SetTruncation/Fibers.agda index 856046c9ad..f15acfa5f8 100644 --- a/Cubical/HITs/SetTruncation/Fibers.agda +++ b/Cubical/HITs/SetTruncation/Fibers.agda @@ -19,6 +19,8 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv +open import Cubical.Data.Sigma + open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.HITs.SetTruncation as Set open import Cubical.HITs.SetQuotients as SetQuot @@ -30,190 +32,78 @@ private ℓ ℓ' : Level module _ - (X : Type ℓ ) - (Y : Type ℓ') + {X : Type ℓ} + {Y : Type ℓ'} (f : X → Y) where private ∥f∥₂ : ∥ X ∥₂ → ∥ Y ∥₂ ∥f∥₂ = Set.map f - module _ - (y : Y) where + module _ (y : Y) where open Iso - fiber→fiber∥∥₂ : fiber f y → fiber ∥f∥₂ ∣ y ∣₂ - fiber→fiber∥∥₂ (x , p) = ∣ x ∣₂ , cong ∣_∣₂ p - isSetFiber∥∥₂ : isSet (fiber ∥f∥₂ ∣ y ∣₂) isSetFiber∥∥₂ = isOfHLevelΣ 2 squash₂ (λ _ → isProp→isSet (squash₂ _ _)) - ∥fiber∥₂→fiber∥∥₂ : ∥ fiber f y ∥₂ → fiber ∥f∥₂ ∣ y ∣₂ - ∥fiber∥₂→fiber∥∥₂ = Set.rec isSetFiber∥∥₂ fiber→fiber∥∥₂ - - ∣transport∣ : ∥ y ≡ y ∥₂ → ∥ fiber f y ∥₂ → ∥ fiber f y ∥₂ - ∣transport∣ = Set.rec2 squash₂ (λ s (x , q) → ∣ x , q ∙ s ∣₂) + fiberRel : ∥ fiber f y ∥₂ → ∥ fiber f y ∥₂ → Type ℓ + fiberRel a b = Set.map fst a ≡ Set.map fst b - fiberRel2 : ∥ fiber f y ∥₂ → ∥ fiber f y ∥₂ → Type (ℓ-max ℓ ℓ') - fiberRel2 x x' = - ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s x ≡ x' ∥ - - fiberRel' : fiber f y → fiber f y → Type (ℓ-max ℓ ℓ') - fiberRel' (x , p) (x' , p') = - Σ[ s ∈ y ≡ y ] Σ[ q ∈ x ≡ x' ] PathP (λ i → f (q i) ≡ s i) p p' - - feq' : (a b : fiber f y)(r : fiberRel' a b) → fiber→fiber∥∥₂ a ≡ fiber→fiber∥∥₂ b - feq' (x , p) (x' , p') (s , q , t) i .fst = ∣ q i ∣₂ - feq' (x , p) (x' , p') (s , q , t) i .snd j = - hcomp (λ k → λ { (i = i0) → ∣ p j ∣₂ - ; (i = i1) → ∣ p' j ∣₂ - ; (j = i0) → ∣ f (q i) ∣₂ - ; (j = i1) → squash₂ ∣ y ∣₂ ∣ y ∣₂ (cong ∣_∣₂ s) refl k i }) - (∣ t i j ∣₂) - - fiberRel : ∥ fiber f y ∥₂ → ∥ fiber f y ∥₂ → Type (ℓ-max ℓ ℓ') - fiberRel a b = Set.rec2 isSetHProp (λ a b → ∥ fiberRel' a b ∥ , isPropPropTrunc) a b .fst - - isPropFiberRel : (x x' : ∥ fiber f y ∥₂) → isProp (fiberRel x x') - isPropFiberRel a b = Set.rec2 isSetHProp (λ a b → ∥ fiberRel' a b ∥ , isPropPropTrunc) a b .snd - - fiberRel'→fiberRel2' : ((x , p) x' : fiber f y) → fiberRel' (x , p) x' → Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' - fiberRel'→fiberRel2' _ _ (_ , _ , t) .fst i = t i i1 - fiberRel'→fiberRel2' _ _ (_ , q , _) .snd i .fst = q i - fiberRel'→fiberRel2' (x , p) (x' , p') (s , q , t) .snd i .snd j = - hcomp (λ k → λ { (i = i0) → compPath-filler p s k j - ; (i = i1) → p' j - ; (j = i0) → f (q i) - ; (j = i1) → s (i ∨ k) }) - (t i j) - - truncRel : ((x , p) x' : fiber f y) → Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' - → ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s ∣ x , p ∣₂ ≡ ∣ x' ∣₂ ∥ - truncRel _ _ (s , p) = ∣ ∣ s ∣₂ , cong ∣_∣₂ p ∣ - - truncRel` : ((x , p) x' : fiber f y) → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ - → ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s ∣ x , p ∣₂ ≡ ∣ x' ∣₂ ∥ - truncRel` x x' = Prop.rec isPropPropTrunc (λ r → truncRel x x' r) - - fiberRel→fiberRel2' : (x x' : fiber f y) → fiberRel ∣ x ∣₂ ∣ x' ∣₂ → fiberRel2 ∣ x ∣₂ ∣ x' ∣₂ - fiberRel→fiberRel2' x x' = Prop.rec isPropPropTrunc (λ r → truncRel x x' (fiberRel'→fiberRel2' x x' r)) - - fiberRel→fiberRel2 : (x x' : ∥ fiber f y ∥₂) → fiberRel x x' → fiberRel2 x x' - fiberRel→fiberRel2 = - Set.elim2 (λ _ _ → isProp→isSet (isPropΠ (λ _ → isPropPropTrunc))) - (λ x x' → fiberRel→fiberRel2' x x') - - truncRel'' : ((x , p) x' : fiber f y) → (s : y ≡ y) → ∣ x , p ∙ s ∣₂ ≡ ∣ x' ∣₂ - → ∥ (x , p ∙ s) ≡ x' ∥ - truncRel'' _ _ _ q = Set.PathIdTrunc₀Iso .fun q - - truncRel''' : ((x , p) x' : fiber f y) → (s : ∥ y ≡ y ∥₂) → ∣transport∣ s ∣ x , p ∣₂ ≡ ∣ x' ∣₂ - → ∥ Σ[ s ∈ y ≡ y ] ∥ (x , p ∙ s) ≡ x' ∥ ∥ - truncRel''' x x' = - Set.elim (λ _ → isProp→isSet (isPropΠ (λ _ → isPropPropTrunc))) - (λ s' p' → ∣ s' , truncRel'' x x' s' p' ∣) - - helper-fun : ((x , p) x' : fiber f y) → (s : y ≡ y) - → ∥ (x , p ∙ s) ≡ x' ∥ → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ - helper-fun _ _ s = Prop.rec isPropPropTrunc (λ q → ∣ s , q ∣) - - helper-fun' : ((x , p) x' : fiber f y) → ∥ Σ[ s ∈ y ≡ y ] ∥ (x , p ∙ s) ≡ x' ∥ ∥ - → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ - helper-fun' _ _ = - Prop.rec isPropPropTrunc (λ (s , q) → helper-fun _ _ s q) - - truncRel' : ((x , p) x' : fiber f y) → ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s ∣ x , p ∣₂ ≡ ∣ x' ∣₂ ∥ - → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ - truncRel' _ _ = Prop.rec isPropPropTrunc (λ (s , q) → helper-fun' _ _ (truncRel''' _ _ s q)) - - fiberRel2'→fiberRel' : ((x , p) x' : fiber f y) → (s : y ≡ y)(t : (x , p ∙ s) ≡ x') → fiberRel' (x , p) x' - fiberRel2'→fiberRel' _ _ s _ .fst = s - fiberRel2'→fiberRel' _ _ _ t .snd .fst i = t i .fst - fiberRel2'→fiberRel' (x , p) (x' , p') s t .snd .snd i j = - hcomp (λ k → λ { (i = i0) → compPath-filler p s (~ k) j - ; (i = i1) → t k .snd j - ; (j = i0) → f (t (i ∧ k) .fst) - ; (j = i1) → s (i ∨ (~ k)) }) - ((p ∙ s) j) - - fiberRel2'→fiberRel'-helper : ((x , p) x' : fiber f y) - → ∥ Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ x' ∥ → fiberRel ∣ x , p ∣₂ ∣ x' ∣₂ - fiberRel2'→fiberRel'-helper _ _ = - Prop.rec isPropPropTrunc (λ (s , p) → ∣ fiberRel2'→fiberRel' _ _ s p ∣) - - fiberRel2→fiberRel : (x x' : ∥ fiber f y ∥₂) → fiberRel2 x x' → fiberRel x x' - fiberRel2→fiberRel = - Set.elim2 (λ x x' → isProp→isSet (isPropΠ (λ _ → isPropFiberRel x x'))) - (λ _ _ r → fiberRel2'→fiberRel'-helper _ _ (truncRel' _ _ r)) - - fiberRel≃fiberRel2 : (x x' : ∥ fiber f y ∥₂) → fiberRel x x' ≃ fiberRel2 x x' - fiberRel≃fiberRel2 x x' = - propBiimpl→Equiv (isPropFiberRel x x') isPropPropTrunc - (fiberRel→fiberRel2 _ _) (fiberRel2→fiberRel _ _) - - feq'' : (a b : fiber f y)(r : fiberRel ∣ a ∣₂ ∣ b ∣₂) → ∥fiber∥₂→fiber∥∥₂ ∣ a ∣₂ ≡ ∥fiber∥₂→fiber∥∥₂ ∣ b ∣₂ - feq'' a b r = Prop.rec (isSetFiber∥∥₂ _ _) (feq' a b) r - - feq : (a b : ∥ fiber f y ∥₂)(r : fiberRel a b) → ∥fiber∥₂→fiber∥∥₂ a ≡ ∥fiber∥₂→fiber∥∥₂ b - feq = Set.elim2 (λ _ _ → isProp→isSet (isPropΠ (λ _ → isSetFiber∥∥₂ _ _))) feq'' + private + proj : ∥ fiber f y ∥₂ / fiberRel → ∥ X ∥₂ + proj = SetQuot.rec squash₂ (Set.map fst) (λ _ _ p → p) ∥fiber∥₂/R→fiber∥∥₂ : ∥ fiber f y ∥₂ / fiberRel → fiber ∥f∥₂ ∣ y ∣₂ ∥fiber∥₂/R→fiber∥∥₂ = SetQuot.rec isSetFiber∥∥₂ ∥fiber∥₂→fiber∥∥₂ feq + where + fiber→fiber∥∥₂ : fiber f y → fiber ∥f∥₂ ∣ y ∣₂ + fiber→fiber∥∥₂ (x , p) = ∣ x ∣₂ , cong ∣_∣₂ p - fun1 : (x x' : X)(q : x ≡ x')(p : f x ≡ y)(p' : f x' ≡ y) → fiberRel' (x , p) (x' , p') - fun1 x x' q p p' .fst = (sym p) ∙∙ cong f q ∙∙ p' - fun1 x x' q p p' .snd .fst = q - fun1 x x' q p p' .snd .snd i j = doubleCompPath-filler (sym p) (cong f q) p' j i - - fun2'' : (x : X) → (p p' : f x ≡ y) → fiberRel ∣ x , p ∣₂ ∣ x , p' ∣₂ - fun2'' x p p' = ∣ fun1 x x refl p p' ∣ + ∥fiber∥₂→fiber∥∥₂ : ∥ fiber f y ∥₂ → fiber ∥f∥₂ ∣ y ∣₂ + ∥fiber∥₂→fiber∥∥₂ = Set.rec isSetFiber∥∥₂ fiber→fiber∥∥₂ - fun2' : (x : X) → (p : ∥ f x ≡ y ∥) → ∥ fiber f y ∥₂ / fiberRel - fun2' x = Prop.rec→Set squash/ (λ p → [ ∣ x , p ∣₂ ]) (λ p p' → eq/ _ _ (fun2'' x p p')) + feq : (a b : ∥ fiber f y ∥₂) (r : fiberRel a b) → ∥fiber∥₂→fiber∥∥₂ a ≡ ∥fiber∥₂→fiber∥∥₂ b + feq = + Set.elim2 (λ _ _ → isProp→isSet (isPropΠ (λ _ → isSetFiber∥∥₂ _ _))) λ _ _ p → + ΣPathP (p , isSet→isSet' squash₂ _ _ _ _) - fun3' : (x : X) → (p : ∥f∥₂ ∣ x ∣₂ ≡ ∣ y ∣₂) → ∥ fiber f y ∥₂ / fiberRel - fun3' x p = fun2' x (Set.PathIdTrunc₀Iso .fun p) - - fun3 : (x : ∥ X ∥₂) → (p : ∥f∥₂ x ≡ ∣ y ∣₂) → ∥ fiber f y ∥₂ / fiberRel - fun3 = Set.elim (λ _ → isSetΠ (λ _ → squash/)) fun3' + mereFiber→∥fiber∥₂/R : (x : X) → ∥ f x ≡ y ∥ → ∥ fiber f y ∥₂ / fiberRel + mereFiber→∥fiber∥₂/R x = Prop.rec→Set squash/ (λ p → [ ∣ _ , p ∣₂ ]) (λ _ _ → eq/ _ _ refl) fiber∥∥₂→∥fiber∥₂/R : fiber ∥f∥₂ ∣ y ∣₂ → ∥ fiber f y ∥₂ / fiberRel - fiber∥∥₂→∥fiber∥₂/R (x , p) = fun3 x p - - leftHtpy : (x : fiber f y) → fiber∥∥₂→∥fiber∥₂/R (fiber→fiber∥∥₂ x) ≡ [ ∣ x ∣₂ ] - leftHtpy (x , p) = eq/ _ _ (fun2'' x _ _) + fiber∥∥₂→∥fiber∥₂/R = + uncurry + (Set.elim (λ _ → isSetΠ λ _ → squash/) λ x p → + mereFiber→∥fiber∥₂/R x (PathIdTrunc₀Iso .fun p)) + + ∥fiber∥₂/R→fiber∥∥₂→fst : (q : ∥ fiber f y ∥₂ / fiberRel) → ∥fiber∥₂/R→fiber∥∥₂ q .fst ≡ proj q + ∥fiber∥₂/R→fiber∥∥₂→fst = + SetQuot.elimProp (λ _ → squash₂ _ _) + (Set.elim (λ _ → isProp→isSet (squash₂ _ _)) (λ _ → refl)) + + fiber∥∥₂→∥fiber∥₂/R→proj : (x : fiber ∥f∥₂ ∣ y ∣₂) → proj (fiber∥∥₂→∥fiber∥₂/R x) ≡ x .fst + fiber∥∥₂→∥fiber∥₂/R→proj = + uncurry + (Set.elim (λ _ → isSetΠ λ _ → isProp→isSet (squash₂ _ _)) λ x p → + Prop.elim {P = λ t → proj (mereFiber→∥fiber∥₂/R x t) ≡ ∣ x ∣₂} + (λ _ → squash₂ _ _) + (λ _ → refl) + (PathIdTrunc₀Iso .fun p)) ∥fiber∥₂/R→fiber∥∥₂→∥fiber∥₂/R : (x : ∥ fiber f y ∥₂ / fiberRel) → fiber∥∥₂→∥fiber∥₂/R (∥fiber∥₂/R→fiber∥∥₂ x) ≡ x - ∥fiber∥₂/R→fiber∥∥₂→∥fiber∥₂/R = SetQuot.elimProp (λ _ → squash/ _ _) (Set.elim (λ _ → isProp→isSet (squash/ _ _)) leftHtpy) - - rightHtpy : - (x : X) → (p : f x ≡ y) → ∥fiber∥₂/R→fiber∥∥₂ (fun2' x ∣ p ∣) ≡ (∣ x ∣₂ , (Set.PathIdTrunc₀Iso .inv ∣ p ∣)) - rightHtpy x p i .fst = ∣ x ∣₂ - rightHtpy x p i .snd = - isOfHLevelRespectEquiv 1 (invEquiv (isoToEquiv (PathIdTrunc₀Iso))) isPropPropTrunc - (∥fiber∥₂/R→fiber∥∥₂ (fun2' x ∣ p ∣) .snd) (Set.PathIdTrunc₀Iso .inv ∣ p ∣) i - - rightHtpy' : - (x : X) → (p : ∥ f x ≡ y ∥) → ∥fiber∥₂/R→fiber∥∥₂ (fun2' x p) ≡ (∣ x ∣₂ , (Set.PathIdTrunc₀Iso .inv p)) - rightHtpy' x = Prop.elim (λ _ → isSetFiber∥∥₂ _ _) (rightHtpy x) - - rightHtpy'' : - (x : X) → (p : ∥f∥₂ ∣ x ∣₂ ≡ ∣ y ∣₂) → ∥fiber∥₂/R→fiber∥∥₂ (fun3' x p) ≡ (∣ x ∣₂ , p) - rightHtpy'' x p i .fst = rightHtpy' x (Set.PathIdTrunc₀Iso .fun p) i .fst - rightHtpy'' x p i .snd = - hcomp (λ j → λ { (i = i0) → rightHtpy' x (Set.PathIdTrunc₀Iso .fun p) i0 .snd - ; (i = i1) → Set.PathIdTrunc₀Iso .leftInv p j }) - (rightHtpy' x (Set.PathIdTrunc₀Iso .fun p) i .snd) - - rightHtpy''' : - (x : ∥ X ∥₂) → (p : ∥f∥₂ x ≡ ∣ y ∣₂) → ∥fiber∥₂/R→fiber∥∥₂ (fiber∥∥₂→∥fiber∥₂/R (x , p)) ≡ (x , p) - rightHtpy''' = Set.elim (λ _ → isSetΠ (λ _ → isProp→isSet (isSetFiber∥∥₂ _ _))) rightHtpy'' + ∥fiber∥₂/R→fiber∥∥₂→∥fiber∥₂/R = + SetQuot.elimProp (λ _ → squash/ _ _) + (Set.elim (λ _ → isProp→isSet (squash/ _ _)) + (λ _ → eq/ _ _ refl)) fiber∥∥₂→∥fiber∥₂/R→fiber∥∥₂ : (x : fiber ∥f∥₂ ∣ y ∣₂) → ∥fiber∥₂/R→fiber∥∥₂ (fiber∥∥₂→∥fiber∥₂/R x) ≡ x - fiber∥∥₂→∥fiber∥₂/R→fiber∥∥₂ (x , p) = rightHtpy''' x p + fiber∥∥₂→∥fiber∥₂/R→fiber∥∥₂ x = + Σ≡Prop + (λ _ → squash₂ _ _) + (∥fiber∥₂/R→fiber∥∥₂→fst (fiber∥∥₂→∥fiber∥₂/R x) ∙ fiber∥∥₂→∥fiber∥₂/R→proj x) Iso-∥fiber∥₂/R-fiber∥∥₂ : Iso (∥ fiber f y ∥₂ / fiberRel) (fiber ∥f∥₂ ∣ y ∣₂) Iso-∥fiber∥₂/R-fiber∥∥₂ .fun = ∥fiber∥₂/R→fiber∥∥₂ @@ -226,69 +116,55 @@ module _ ∥fiber∥₂/R≃fiber∥∥₂ : ∥ fiber f y ∥₂ / fiberRel ≃ fiber ∥f∥₂ ∣ y ∣₂ ∥fiber∥₂/R≃fiber∥∥₂ = isoToEquiv Iso-∥fiber∥₂/R-fiber∥∥₂ - ∥fiber∥₂/Rel≃∥fiber∥₂/R : ∥ fiber f y ∥₂ / fiberRel2 ≃ ∥ fiber f y ∥₂ / fiberRel - ∥fiber∥₂/Rel≃∥fiber∥₂/R = - isoToEquiv (relBiimpl→TruncIso - (λ {a} {b} → fiberRel2→fiberRel a b) - (λ {a} {b} → fiberRel→fiberRel2 a b)) - - ∥fiber∥₂/Rel≃fiber∥∥₂ : ∥ fiber f y ∥₂ / fiberRel2 ≃ fiber ∥f∥₂ ∣ y ∣₂ - ∥fiber∥₂/Rel≃fiber∥∥₂ = compEquiv ∥fiber∥₂/Rel≃∥fiber∥₂/R ∥fiber∥₂/R≃fiber∥∥₂ - - -- the relation is equivalence relation + -- the relation is an equivalence relation open BinaryRelation + open isEquivRel - isReflFiberRel' : (x : fiber f y) → fiberRel2 ∣ x ∣₂ ∣ x ∣₂ - isReflFiberRel' (x , q) = ∣ ∣ refl ∣₂ , (λ i → ∣ x , rUnit q (~ i) ∣₂) ∣ - - isReflFiberRel : isRefl fiberRel2 - isReflFiberRel = Set.elim (λ _ → isProp→isSet isPropPropTrunc) isReflFiberRel' - - isSymFiberRel'' : ((x , p) (x' , p') : fiber f y) - → Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ (x' , p') → Σ[ s' ∈ y ≡ y ] (x' , p' ∙ s') ≡ (x , p) - isSymFiberRel'' _ _ (s , _) .fst = sym s - isSymFiberRel'' (x , p) (x' , p') (s , q) .snd i .fst = q (~ i) .fst - isSymFiberRel'' (x , p) (x' , p') (s , q) .snd i .snd j = - hcomp (λ k → λ { (i = i0) → compPath-filler p' (sym s) k j - ; (i = i1) → compPath-filler p s (~ k) j - ; (j = i0) → f (q (~ i) .fst) - ; (j = i1) → s (~ k) }) - (q (~ i) .snd j) - - isSymFiberRel' : (x x' : fiber f y) - → fiberRel2 ∣ x ∣₂ ∣ x' ∣₂ → fiberRel2 ∣ x' ∣₂ ∣ x ∣₂ - isSymFiberRel' _ _ = truncRel` _ _ ∘ Prop.rec isPropPropTrunc (λ r → ∣ isSymFiberRel'' _ _ r ∣) ∘ truncRel' _ _ - - isSymFiberRel : isSym fiberRel2 - isSymFiberRel = Set.elim2 (λ _ _ → isProp→isSet (isPropΠ (λ _ → isPropPropTrunc))) isSymFiberRel' - - isTransFiberRel'' : ((x , p) (x' , p') (x'' , p'') : fiber f y) - → Σ[ s ∈ y ≡ y ] (x , p ∙ s) ≡ (x' , p') → Σ[ s' ∈ y ≡ y ] (x' , p' ∙ s') ≡ (x'' , p'') - → Σ[ s'' ∈ y ≡ y ] (x , p ∙ s'') ≡ (x'' , p'') - isTransFiberRel'' (x , p) (x' , p') (x'' , p'') (s , q) (s' , q') = - let sq : (i j : I) → Y - sq i j = - hcomp (λ k → λ { (i = i0) → p j - ; (i = i1) → fiberRel2'→fiberRel' _ _ s' q' .snd .snd k j - ; (j = i0) → f (compPath-filler (cong fst q) (cong fst q') k i) - ; (j = i1) → compPath-filler s s' k i }) - (fiberRel2'→fiberRel' _ _ s q .snd .snd i j) - in fiberRel'→fiberRel2' _ _ (_ , _ , (λ i j → sq i j)) - - isTransFiberRel' : (x x' x'' : fiber f y) - → fiberRel2 ∣ x ∣₂ ∣ x' ∣₂ → fiberRel2 ∣ x' ∣₂ ∣ x'' ∣₂ → fiberRel2 ∣ x ∣₂ ∣ x'' ∣₂ - isTransFiberRel' _ _ _ a b = - truncRel` _ _ (Prop.rec2 isPropPropTrunc - (λ r r' → ∣ isTransFiberRel'' _ _ _ r r' ∣) (truncRel' _ _ a) (truncRel' _ _ b)) - - isTransFiberRel : isTrans fiberRel2 - isTransFiberRel = - Set.elim3 (λ _ _ _ → isProp→isSet (isPropΠ2 (λ _ _ → isPropPropTrunc))) isTransFiberRel' + isEquivRelFiberRel : isEquivRel fiberRel + isEquivRelFiberRel .reflexive _ = refl + isEquivRelFiberRel .symmetric _ _ = sym + isEquivRelFiberRel .transitive _ _ _ = _∙_ - open isEquivRel + -- alternative characterization of the relation in terms of equality in Y and fiber f y + + ∣transport∣ : ∥ y ≡ y ∥₂ → ∥ fiber f y ∥₂ → ∥ fiber f y ∥₂ + ∣transport∣ = Set.rec2 squash₂ (λ s (x , q) → ∣ x , q ∙ s ∣₂) - isEquivRelFiberRel : isEquivRel fiberRel2 - isEquivRelFiberRel .reflexive = isReflFiberRel - isEquivRelFiberRel .symmetric = isSymFiberRel - isEquivRelFiberRel .transitive = isTransFiberRel + fiberRel2 : (x x' : ∥ fiber f y ∥₂) → Type (ℓ-max ℓ ℓ') + fiberRel2 x x' = ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s x ≡ x' ∥ + + fiberRel2→1 : ∀ x x' → fiberRel2 x x' → fiberRel x x' + fiberRel2→1 = + Set.elim2 (λ _ _ → isSetΠ λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → + Prop.rec (squash₂ _ _) + (uncurry + (Set.elim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → + cong (Set.map fst))) + + fiberRel1→2 : ∀ x x' → fiberRel x x' → fiberRel2 x x' + fiberRel1→2 = + Set.elim2 (λ _ _ → isSetΠ λ _ → isProp→isSet squash) λ a b p → + Prop.rec squash + (λ q → + let + filler = doubleCompPath-filler (sym (a .snd)) (cong f q) (b .snd) + in + ∣ ∣ filler i1 ∣₂ , cong ∣_∣₂ (ΣPathP (q , adjustLemma (flipSquare filler))) ∣) + (PathIdTrunc₀Iso .Iso.fun p) + where + adjustLemma : {x y z w : Y} {p : x ≡ y} {q : x ≡ z} {r : z ≡ w} {s : y ≡ w} + → PathP (λ i → p i ≡ r i) q s + → PathP (λ i → p i ≡ w) (q ∙ r) s + adjustLemma {p = p} {q} {r} {s} α i j = + hcomp + (λ k → λ + { (i = i0) → compPath-filler q r k j + ; (i = i1) → s j + ; (j = i0) → p i + ; (j = i1) → r (i ∨ k)}) + (α i j) + + fiberRel1≃2 : ∀ x x' → fiberRel x x' ≃ fiberRel2 x x' + fiberRel1≃2 _ _ = + propBiimpl→Equiv (squash₂ _ _) squash (fiberRel1→2 _ _) (fiberRel2→1 _ _) From c14406596886946f833fdf9912551809cde70317 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Sat, 26 Mar 2022 16:58:30 +0800 Subject: [PATCH 17/24] some renaming --- Cubical/Data/FinSet/Cardinality.agda | 2 +- .../{Constructor.agda => Constructors.agda} | 2 +- Cubical/Data/FinSet/FiniteStructure.agda | 22 +++++++++---------- Cubical/Data/FinSet/FiniteType.agda | 16 ++++++-------- Cubical/Data/FinSet/Induction.agda | 2 +- Cubical/Data/FinSet/Quotients.agda | 8 +++---- .../Experiments/CountingFiniteStructure.agda | 6 ++--- Cubical/HITs/SetTruncation/Fibers.agda | 2 +- 8 files changed, 29 insertions(+), 31 deletions(-) rename Cubical/Data/FinSet/{Constructor.agda => Constructors.agda} (99%) diff --git a/Cubical/Data/FinSet/Cardinality.agda b/Cubical/Data/FinSet/Cardinality.agda index 470291c0bb..6e28bb92aa 100644 --- a/Cubical/Data/FinSet/Cardinality.agda +++ b/Cubical/Data/FinSet/Cardinality.agda @@ -40,7 +40,7 @@ open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.FiniteChoice -open import Cubical.Data.FinSet.Constructor +open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Induction hiding (_+_) open import Cubical.Relation.Nullary diff --git a/Cubical/Data/FinSet/Constructor.agda b/Cubical/Data/FinSet/Constructors.agda similarity index 99% rename from Cubical/Data/FinSet/Constructor.agda rename to Cubical/Data/FinSet/Constructors.agda index e30993124c..2d91bac40f 100644 --- a/Cubical/Data/FinSet/Constructor.agda +++ b/Cubical/Data/FinSet/Constructors.agda @@ -6,7 +6,7 @@ This files contains: -} {-# OPTIONS --safe #-} -module Cubical.Data.FinSet.Constructor where +module Cubical.Data.FinSet.Constructors where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels diff --git a/Cubical/Data/FinSet/FiniteStructure.agda b/Cubical/Data/FinSet/FiniteStructure.agda index dc89514f38..8c297070bf 100644 --- a/Cubical/Data/FinSet/FiniteStructure.agda +++ b/Cubical/Data/FinSet/FiniteStructure.agda @@ -1,8 +1,7 @@ {- Finite Structures over Finite Set - -In short, the type of structures should be finite set itself. + In short, the type of structures should be finite set itself. This file contains: - Definition and properties of finite sets equipped with finite structures; @@ -27,7 +26,7 @@ open import Cubical.Data.Sigma open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.Induction -open import Cubical.Data.FinSet.Constructor +open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Cardinality open import Cubical.Data.FinSet.FiniteType @@ -39,16 +38,16 @@ private -- type of finite sets with finite structure -FinSetWithStr : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') -FinSetWithStr ℓ S = Σ[ X ∈ FinSet ℓ ] S X .fst +FinSetWithStr : (S : FinSet ℓ → FinSet ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +FinSetWithStr {ℓ = ℓ} S = Σ[ X ∈ FinSet ℓ ] S X .fst -- type of finite sets with a fixed cardinal FinSetOfCard : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) FinSetOfCard ℓ n = Σ[ X ∈ FinSet ℓ ] (card X ≡ n) -FinSetWithStrOfCard : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') (n : ℕ) → Type (ℓ-max (ℓ-suc ℓ) ℓ') -FinSetWithStrOfCard ℓ S n = Σ[ X ∈ FinSetOfCard ℓ n ] S (X .fst) .fst +FinSetWithStrOfCard : (S : FinSet ℓ → FinSet ℓ') (n : ℕ) → Type (ℓ-max (ℓ-suc ℓ) ℓ') +FinSetWithStrOfCard {ℓ = ℓ} S n = Σ[ X ∈ FinSetOfCard ℓ n ] S (X .fst) .fst FinSetOfCard≡ : (X Y : FinSetOfCard ℓ n) → (X .fst ≡ Y .fst) ≃ (X ≡ Y) FinSetOfCard≡ _ _ = Σ≡PropEquiv (λ _ → isSetℕ _ _) @@ -70,9 +69,10 @@ isFinTypeFinSetOfCard .snd X Y = isFinSet→isFinType 0 (EquivPresIsFinSet (FinSet≡ _ _ ⋆ FinSetOfCard≡ _ _) (isFinSetType≡Eff (X .fst) (Y .fst))) -- the type of finitely-structured finite sets is Rijke finite --- in particular, we can count the number of them up to equivalence +-- in particular, we can count their number up to equivalence -isFinTypeFinSetWithStrOfCard : (ℓ : Level) (S : FinSet ℓ → FinSet ℓ') (n : ℕ) - → isFinType 0 (FinSetWithStrOfCard ℓ S n) -isFinTypeFinSetWithStrOfCard ℓ S n = +isFinTypeFinSetWithStrOfCard : + (S : FinSet ℓ → FinSet ℓ') (n : ℕ) + → isFinType 0 (FinSetWithStrOfCard S n) +isFinTypeFinSetWithStrOfCard S n = isFinTypeΣ {n = 0} (_ , isFinTypeFinSetOfCard) (λ X → _ , isFinSet→isFinType 0 (S (X .fst) .snd)) diff --git a/Cubical/Data/FinSet/FiniteType.agda b/Cubical/Data/FinSet/FiniteType.agda index 3b8b64416a..32156bbf8c 100644 --- a/Cubical/Data/FinSet/FiniteType.agda +++ b/Cubical/Data/FinSet/FiniteType.agda @@ -5,7 +5,7 @@ Finite Types This file formalize the notion of Rijke finite type, which is a direct generalization of Bishop finite set. Basically, a type is (Rijke) n-finite if its i-th order -homotopy groups πᵢ for i ≤ n are Bishop finite. +homotopy groups πᵢ are Bishop finite for i ≤ n. Referrences: https://github.com/EgbertRijke/OEIS-A000001 @@ -31,7 +31,7 @@ open import Cubical.Foundations.Path open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.HITs.SetTruncation as Set open import Cubical.HITs.SetTruncation.Fibers -open import Cubical.HITs.SetQuotients as SetQuot +open import Cubical.HITs.SetQuotients as SetQuot open import Cubical.Data.Nat open import Cubical.Data.Sigma @@ -39,7 +39,7 @@ open import Cubical.Data.Sigma open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.DecidablePredicate -open import Cubical.Data.FinSet.Constructor +open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Quotients open import Cubical.Data.FinSet.Cardinality @@ -111,8 +111,7 @@ module _ ∥f∥₂ : ∥ X ∥₂ → ∥ Y ∥₂ ∥f∥₂ = Set.map f - module _ - (y : Y) where + module _ (y : Y) where isDecPropFiberRel : (x x' : ∥ fiber f y ∥₂) → isDecProp (fiberRel f y x x') isDecPropFiberRel x x' = @@ -136,13 +135,12 @@ module _ isFinType0Σ : isFinType 0 (Σ (X .fst) (λ x → Y x .fst)) isFinType0Σ = - isFinType0Total (Σ (X .fst) (λ x → Y x .fst)) (X .fst) (X .snd) - (λ (x , y) → x) (λ x → EquivPresIsFinType 0 (fiberProjEquiv _ _ x) (Y x .snd)) + isFinType0Total (Σ (X .fst) (λ x → Y x .fst)) (X .fst) (X .snd) fst + (λ x → EquivPresIsFinType 0 (fiberProjEquiv _ _ x) (Y x .snd)) -- the closedness result -isFinTypeΣ : - {n : ℕ} +isFinTypeΣ : {n : ℕ} (X : FinType ℓ (1 + n)) (Y : X .fst → FinType ℓ' n) → isFinType n (Σ (X .fst) (λ x → Y x .fst)) diff --git a/Cubical/Data/FinSet/Induction.agda b/Cubical/Data/FinSet/Induction.agda index ca77b9bc2c..e3dfec1643 100644 --- a/Cubical/Data/FinSet/Induction.agda +++ b/Cubical/Data/FinSet/Induction.agda @@ -26,7 +26,7 @@ open import Cubical.Data.Fin renaming (Fin to Finℕ) open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties -open import Cubical.Data.FinSet.Constructor +open import Cubical.Data.FinSet.Constructors private variable diff --git a/Cubical/Data/FinSet/Quotients.agda b/Cubical/Data/FinSet/Quotients.agda index 3a5d0b1488..05f76a9d9c 100644 --- a/Cubical/Data/FinSet/Quotients.agda +++ b/Cubical/Data/FinSet/Quotients.agda @@ -27,7 +27,7 @@ open import Cubical.Data.SumFin open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.DecidablePredicate -open import Cubical.Data.FinSet.Constructor +open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Cardinality open import Cubical.Relation.Nullary @@ -143,7 +143,7 @@ module _ isFinSetQuot : isFinSet (X .fst / R) isFinSetQuot = EquivPresIsFinSet - ( ∥Eff≃∥Dec X _ dec - ⋆ ∥Dec≃∥ _ _ (λ x x' → isDecProp→Dec (dec x x')) - ⋆ invEquiv (equivQuot {ℓ' = ℓ'} _ _ h)) + ( ∥Eff≃∥Dec X _ dec + ⋆ ∥Dec≃∥ _ _ (λ x x' → isDecProp→Dec (dec x x')) + ⋆ invEquiv (equivQuot {ℓ' = ℓ'} _ _ h)) (isFinSet∥Eff X _ dec) diff --git a/Cubical/Experiments/CountingFiniteStructure.agda b/Cubical/Experiments/CountingFiniteStructure.agda index 3913530c12..c872a41be4 100644 --- a/Cubical/Experiments/CountingFiniteStructure.agda +++ b/Cubical/Experiments/CountingFiniteStructure.agda @@ -20,7 +20,7 @@ open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Data.FinSet.DecidablePredicate open import Cubical.Data.FinSet.Induction -open import Cubical.Data.FinSet.Constructor +open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Cardinality open import Cubical.Data.FinSet.FiniteType open import Cubical.Data.FinSet.FiniteStructure @@ -31,8 +31,8 @@ private -- convenient abbreviation -isFinStrCard : (S : FinSet ℓ-zero → FinSet ℓ) (n : ℕ) → isFinType 0 (FinSetWithStrOfCard ℓ-zero S n) -isFinStrCard S n = isFinTypeFinSetWithStrOfCard ℓ-zero S n +isFinStrCard : (S : FinSet ℓ-zero → FinSet ℓ) (n : ℕ) → isFinType 0 (FinSetWithStrOfCard S n) +isFinStrCard S n = isFinTypeFinSetWithStrOfCard S n -- structure that there is no structure diff --git a/Cubical/HITs/SetTruncation/Fibers.agda b/Cubical/HITs/SetTruncation/Fibers.agda index f15acfa5f8..f183bc1eb2 100644 --- a/Cubical/HITs/SetTruncation/Fibers.agda +++ b/Cubical/HITs/SetTruncation/Fibers.agda @@ -32,7 +32,7 @@ private ℓ ℓ' : Level module _ - {X : Type ℓ} + {X : Type ℓ } {Y : Type ℓ'} (f : X → Y) where From 2d85153efdec00a6c21eeced3c89e132f91700f5 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Sat, 26 Mar 2022 17:23:45 +0800 Subject: [PATCH 18/24] modify minors by request --- Cubical/Data/Bool/Properties.agda | 120 ++++++++++---------- Cubical/Data/FinSet/Cardinality.agda | 5 +- Cubical/Data/FinSet/Constructors.agda | 26 ++--- Cubical/Data/FinSet/DecidablePredicate.agda | 10 +- Cubical/Data/FinSet/FiniteType.agda | 5 +- Cubical/Data/FinSet/Induction.agda | 2 +- Cubical/Data/FinSet/Quotients.agda | 12 +- Cubical/Data/SumFin/Properties.agda | 18 +-- 8 files changed, 96 insertions(+), 102 deletions(-) diff --git a/Cubical/Data/Bool/Properties.agda b/Cubical/Data/Bool/Properties.agda index 4548a7eb48..46e47801e5 100644 --- a/Cubical/Data/Bool/Properties.agda +++ b/Cubical/Data/Bool/Properties.agda @@ -235,33 +235,33 @@ BoolProp≃BoolProp* : {a : Bool} → Bool→Type a ≃ Bool→Type* {ℓ} a BoolProp≃BoolProp* {a = true} = Unit≃Unit* BoolProp≃BoolProp* {a = false} = uninhabEquiv Empty.rec Empty.rec* -Bool→PropInj : (a b : Bool) → Bool→Type a ≃ Bool→Type b → a ≡ b -Bool→PropInj true true _ = refl -Bool→PropInj true false p = Empty.rec (p .fst tt) -Bool→PropInj false true p = Empty.rec (invEq p tt) -Bool→PropInj false false _ = refl - -Bool→PropInj* : (a b : Bool) → Bool→Type* {ℓ} a ≃ Bool→Type* {ℓ} b → a ≡ b -Bool→PropInj* true true _ = refl -Bool→PropInj* true false p = Empty.rec* (p .fst tt*) -Bool→PropInj* false true p = Empty.rec* (invEq p tt*) -Bool→PropInj* false false _ = refl - -isPropBool→Prop : {a : Bool} → isProp (Bool→Type a) -isPropBool→Prop {a = true} = isPropUnit -isPropBool→Prop {a = false} = isProp⊥ - -isPropBool→Prop* : {a : Bool} → isProp (Bool→Type* {ℓ} a) -isPropBool→Prop* {a = true} = isPropUnit* -isPropBool→Prop* {a = false} = isProp⊥* - -DecBool→Prop : {a : Bool} → Dec (Bool→Type a) -DecBool→Prop {a = true} = yes tt -DecBool→Prop {a = false} = no (λ x → x) - -DecBool→Prop* : {a : Bool} → Dec (Bool→Type* {ℓ} a) -DecBool→Prop* {a = true} = yes tt* -DecBool→Prop* {a = false} = no (λ (lift x) → x) +Bool→TypeInj : (a b : Bool) → Bool→Type a ≃ Bool→Type b → a ≡ b +Bool→TypeInj true true _ = refl +Bool→TypeInj true false p = Empty.rec (p .fst tt) +Bool→TypeInj false true p = Empty.rec (invEq p tt) +Bool→TypeInj false false _ = refl + +Bool→TypeInj* : (a b : Bool) → Bool→Type* {ℓ} a ≃ Bool→Type* {ℓ} b → a ≡ b +Bool→TypeInj* true true _ = refl +Bool→TypeInj* true false p = Empty.rec* (p .fst tt*) +Bool→TypeInj* false true p = Empty.rec* (invEq p tt*) +Bool→TypeInj* false false _ = refl + +isPropBool→Type : {a : Bool} → isProp (Bool→Type a) +isPropBool→Type {a = true} = isPropUnit +isPropBool→Type {a = false} = isProp⊥ + +isPropBool→Type* : {a : Bool} → isProp (Bool→Type* {ℓ} a) +isPropBool→Type* {a = true} = isPropUnit* +isPropBool→Type* {a = false} = isProp⊥* + +DecBool→Type : {a : Bool} → Dec (Bool→Type a) +DecBool→Type {a = true} = yes tt +DecBool→Type {a = false} = no (λ x → x) + +DecBool→Type* : {a : Bool} → Dec (Bool→Type* {ℓ} a) +DecBool→Type* {a = true} = yes tt* +DecBool→Type* {a = false} = no (λ (lift x) → x) Dec→DecBool : {P : Type ℓ} → (dec : Dec P) → P → Bool→Type (Dec→Bool dec) Dec→DecBool (yes p) _ = tt @@ -271,9 +271,9 @@ DecBool→Dec : {P : Type ℓ} → (dec : Dec P) → Bool→Type (Dec→Bool dec DecBool→Dec (yes p) _ = p Dec≃DecBool : {P : Type ℓ} → (h : isProp P)(dec : Dec P) → P ≃ Bool→Type (Dec→Bool dec) -Dec≃DecBool h dec = propBiimpl→Equiv h isPropBool→Prop (Dec→DecBool dec) (DecBool→Dec dec) +Dec≃DecBool h dec = propBiimpl→Equiv h isPropBool→Type (Dec→DecBool dec) (DecBool→Dec dec) -Bool≡BoolDec : {a : Bool} → a ≡ Dec→Bool (DecBool→Prop {a = a}) +Bool≡BoolDec : {a : Bool} → a ≡ Dec→Bool (DecBool→Type {a = a}) Bool≡BoolDec {a = true} = refl Bool≡BoolDec {a = false} = refl @@ -285,41 +285,41 @@ DecBool→Dec* : {P : Type ℓ} → (dec : Dec P) → Bool→Type* {ℓ} (Dec→ DecBool→Dec* (yes p) _ = p Dec≃DecBool* : {P : Type ℓ} → (h : isProp P)(dec : Dec P) → P ≃ Bool→Type* {ℓ} (Dec→Bool dec) -Dec≃DecBool* h dec = propBiimpl→Equiv h isPropBool→Prop* (Dec→DecBool* dec) (DecBool→Dec* dec) +Dec≃DecBool* h dec = propBiimpl→Equiv h isPropBool→Type* (Dec→DecBool* dec) (DecBool→Dec* dec) -Bool≡BoolDec* : {a : Bool} → a ≡ Dec→Bool (DecBool→Prop* {ℓ} {a = a}) +Bool≡BoolDec* : {a : Bool} → a ≡ Dec→Bool (DecBool→Type* {ℓ} {a = a}) Bool≡BoolDec* {a = true} = refl Bool≡BoolDec* {a = false} = refl -Bool→Prop× : (a b : Bool) → Bool→Type (a and b) → Bool→Type a × Bool→Type b -Bool→Prop× true true _ = tt , tt -Bool→Prop× true false p = Empty.rec p -Bool→Prop× false true p = Empty.rec p -Bool→Prop× false false p = Empty.rec p - -Bool→Prop×' : (a b : Bool) → Bool→Type a × Bool→Type b → Bool→Type (a and b) -Bool→Prop×' true true _ = tt -Bool→Prop×' true false (_ , p) = Empty.rec p -Bool→Prop×' false true (p , _) = Empty.rec p -Bool→Prop×' false false (p , _) = Empty.rec p - -Bool→Prop≃ : (a b : Bool) → Bool→Type a × Bool→Type b ≃ Bool→Type (a and b) -Bool→Prop≃ a b = - propBiimpl→Equiv (isProp× isPropBool→Prop isPropBool→Prop) isPropBool→Prop - (Bool→Prop×' a b) (Bool→Prop× a b) - -Bool→Prop⊎ : (a b : Bool) → Bool→Type (a or b) → Bool→Type a ⊎ Bool→Type b -Bool→Prop⊎ true true _ = inl tt -Bool→Prop⊎ true false _ = inl tt -Bool→Prop⊎ false true _ = inr tt -Bool→Prop⊎ false false p = Empty.rec p - -Bool→Prop⊎' : (a b : Bool) → Bool→Type a ⊎ Bool→Type b → Bool→Type (a or b) -Bool→Prop⊎' true true _ = tt -Bool→Prop⊎' true false _ = tt -Bool→Prop⊎' false true _ = tt -Bool→Prop⊎' false false (inl p) = Empty.rec p -Bool→Prop⊎' false false (inr p) = Empty.rec p +Bool→Type× : (a b : Bool) → Bool→Type (a and b) → Bool→Type a × Bool→Type b +Bool→Type× true true _ = tt , tt +Bool→Type× true false p = Empty.rec p +Bool→Type× false true p = Empty.rec p +Bool→Type× false false p = Empty.rec p + +Bool→Type×' : (a b : Bool) → Bool→Type a × Bool→Type b → Bool→Type (a and b) +Bool→Type×' true true _ = tt +Bool→Type×' true false (_ , p) = Empty.rec p +Bool→Type×' false true (p , _) = Empty.rec p +Bool→Type×' false false (p , _) = Empty.rec p + +Bool→Type×≃ : (a b : Bool) → Bool→Type a × Bool→Type b ≃ Bool→Type (a and b) +Bool→Type×≃ a b = + propBiimpl→Equiv (isProp× isPropBool→Type isPropBool→Type) isPropBool→Type + (Bool→Type×' a b) (Bool→Type× a b) + +Bool→Type⊎ : (a b : Bool) → Bool→Type (a or b) → Bool→Type a ⊎ Bool→Type b +Bool→Type⊎ true true _ = inl tt +Bool→Type⊎ true false _ = inl tt +Bool→Type⊎ false true _ = inr tt +Bool→Type⊎ false false p = Empty.rec p + +Bool→Type⊎' : (a b : Bool) → Bool→Type a ⊎ Bool→Type b → Bool→Type (a or b) +Bool→Type⊎' true true _ = tt +Bool→Type⊎' true false _ = tt +Bool→Type⊎' false true _ = tt +Bool→Type⊎' false false (inl p) = Empty.rec p +Bool→Type⊎' false false (inr p) = Empty.rec p PropBoolP→P : (dec : Dec A) → Bool→Type (Dec→Bool dec) → A PropBoolP→P (yes p) _ = p diff --git a/Cubical/Data/FinSet/Cardinality.agda b/Cubical/Data/FinSet/Cardinality.agda index 6e28bb92aa..f174e06118 100644 --- a/Cubical/Data/FinSet/Cardinality.agda +++ b/Cubical/Data/FinSet/Cardinality.agda @@ -22,6 +22,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Transport open import Cubical.HITs.PropositionalTruncation as Prop open import Cubical.HITs.SetTruncation as Set @@ -86,7 +87,7 @@ module _ card≡0→isEmpty : card X ≡ 0 → ¬ X .fst card≡0→isEmpty p x = - Prop.rec isProp⊥ (λ e → (idfun _) (transport (cong Fin p) (e .fst x))) (∣≃card∣ X) + Prop.rec isProp⊥ (λ e → subst Fin p (e .fst x)) (∣≃card∣ X) card>0→isInhab : card X > 0 → ∥ X .fst ∥ card>0→isInhab p = @@ -104,7 +105,7 @@ module _ card≡1→isContr : card X ≡ 1 → isContr (X .fst) card≡1→isContr p = Prop.rec isPropIsContr - (λ e → isOfHLevelRespectEquiv 0 (invEquiv (e ⋆ pathToEquiv (cong Fin p))) isContrSumFin1) (∣≃card∣ X) + (λ e → isOfHLevelRespectEquiv 0 (invEquiv (e ⋆ substEquiv Fin p)) isContrSumFin1) (∣≃card∣ X) card≤1→isProp : card X ≤ 1 → isProp (X .fst) card≤1→isProp p = diff --git a/Cubical/Data/FinSet/Constructors.agda b/Cubical/Data/FinSet/Constructors.agda index 2d91bac40f..332138f50a 100644 --- a/Cubical/Data/FinSet/Constructors.agda +++ b/Cubical/Data/FinSet/Constructors.agda @@ -85,16 +85,14 @@ module _ (Y : X .fst → FinSet ℓ') where isFinSetΣ : isFinSet (Σ (X .fst) (λ x → Y x .fst)) - isFinSetΣ = - elim2 (λ _ _ → isPropIsFinSet {A = Σ (X .fst) (λ x → Y x .fst)}) - (λ p q → isFinOrd→isFinSet (isFinOrdΣ (X .fst) (_ , p) (λ x → Y x .fst) q)) - (X .snd .snd) (choice X (λ x → isFinOrd (Y x .fst)) (λ x → isFinSet→isFinSet' (Y x .snd))) + isFinSetΣ = rec2 isPropIsFinSet + (λ p q → isFinOrd→isFinSet (isFinOrdΣ (X .fst) (_ , p) (λ x → Y x .fst) q)) + (X .snd .snd) (choice X (λ x → isFinOrd (Y x .fst)) (λ x → isFinSet→isFinSet' (Y x .snd))) isFinSetΠ : isFinSet ((x : X .fst) → Y x .fst) - isFinSetΠ = - elim2 (λ _ _ → isPropIsFinSet {A = ((x : X .fst) → Y x .fst)}) - (λ p q → isFinOrd→isFinSet (isFinOrdΠ (X .fst) (_ , p) (λ x → Y x .fst) q)) - (X .snd .snd) (choice X (λ x → isFinOrd (Y x .fst)) (λ x → isFinSet→isFinSet' (Y x .snd))) + isFinSetΠ = rec2 isPropIsFinSet + (λ p q → isFinOrd→isFinSet (isFinOrdΠ (X .fst) (_ , p) (λ x → Y x .fst) q)) + (X .snd .snd) (choice X (λ x → isFinOrd (Y x .fst)) (λ x → isFinSet→isFinSet' (Y x .snd))) module _ (X : FinSet ℓ) @@ -148,15 +146,11 @@ module _ isFinSet⊎ : isFinSet (X .fst ⊎ Y .fst) isFinSet⊎ = card X + card Y , - elim2 (λ _ _ → isPropPropTrunc {A = _ ≃ Fin (card X + card Y)}) - (λ p q → ∣ isFinOrd⊎ (X .fst) (_ , p) (Y .fst) (_ , q) .snd ∣) - (X .snd .snd) (Y .snd .snd) + map2 (λ p q → isFinOrd⊎ (X .fst) (_ , p) (Y .fst) (_ , q) .snd) (X .snd .snd) (Y .snd .snd) isFinSet× : isFinSet (X .fst × Y .fst) isFinSet× = card X · card Y , - elim2 (λ _ _ → isPropPropTrunc {A = _ ≃ Fin (card X · card Y)}) - (λ p q → ∣ isFinOrd× (X .fst) (_ , p) (Y .fst) (_ , q) .snd ∣) - (X .snd .snd) (Y .snd .snd) + map2 (λ p q → isFinOrd× (X .fst) (_ , p) (Y .fst) (_ , q) .snd) (X .snd .snd) (Y .snd .snd) isFinSet→ : isFinSet (X .fst → Y .fst) isFinSet→ = isFinSetΠ X (λ _ → Y) @@ -175,9 +169,7 @@ module _ isFinSetAut : isFinSet (X .fst ≃ X .fst) isFinSetAut = LehmerCode.factorial (card X) , - Prop.elim (λ _ → isPropPropTrunc {A = _ ≃ Fin _}) - (λ p → ∣ isFinOrd≃ (X .fst) (card X , p) .snd ∣) - (X .snd .snd) + Prop.map (λ p → isFinOrd≃ (X .fst) (card X , p) .snd) (X .snd .snd) isFinSetTypeAut : isFinSet (X .fst ≡ X .fst) isFinSetTypeAut = EquivPresIsFinSet (invEquiv univalence) isFinSetAut diff --git a/Cubical/Data/FinSet/DecidablePredicate.agda b/Cubical/Data/FinSet/DecidablePredicate.agda index 0dc80a6caa..814ba7cf8b 100644 --- a/Cubical/Data/FinSet/DecidablePredicate.agda +++ b/Cubical/Data/FinSet/DecidablePredicate.agda @@ -36,15 +36,15 @@ isDecProp : Type ℓ → Type ℓ isDecProp P = Σ[ t ∈ Bool ] P ≃ Bool→Type t isDecProp→isProp : {P : Type ℓ} → isDecProp P → isProp P -isDecProp→isProp h = isOfHLevelRespectEquiv 1 (invEquiv (h .snd)) isPropBool→Prop +isDecProp→isProp h = isOfHLevelRespectEquiv 1 (invEquiv (h .snd)) isPropBool→Type isDecProp→Dec : {P : Type ℓ} → isDecProp P → Dec P -isDecProp→Dec h = EquivPresDec (invEquiv (h .snd)) DecBool→Prop +isDecProp→Dec h = EquivPresDec (invEquiv (h .snd)) DecBool→Type isPropIsDecProp : {P : Type ℓ} → isProp (isDecProp P) isPropIsDecProp p q = - Σ≡PropEquiv (λ _ → isOfHLevel⁺≃ᵣ 0 isPropBool→Prop) .fst - (Bool→PropInj _ _ (invEquiv (p .snd) ⋆ q .snd)) + Σ≡PropEquiv (λ _ → isOfHLevel⁺≃ᵣ 0 isPropBool→Type) .fst + (Bool→TypeInj _ _ (invEquiv (p .snd) ⋆ q .snd)) DecProp : (ℓ : Level) → Type (ℓ-suc ℓ) DecProp ℓ = Σ[ P ∈ Type ℓ ] isDecProp P @@ -157,7 +157,7 @@ module _ isDecProp× : isDecProp (P .fst × Q .fst) isDecProp× .fst = P .snd .fst and Q .snd .fst - isDecProp× .snd = Σ-cong-equiv (P .snd .snd) (λ _ → Q .snd .snd) ⋆ Bool→Prop≃ _ _ + isDecProp× .snd = Σ-cong-equiv (P .snd .snd) (λ _ → Q .snd .snd) ⋆ Bool→Type×≃ _ _ module _ (X : FinSet ℓ) where diff --git a/Cubical/Data/FinSet/FiniteType.agda b/Cubical/Data/FinSet/FiniteType.agda index 32156bbf8c..7394fbba47 100644 --- a/Cubical/Data/FinSet/FiniteType.agda +++ b/Cubical/Data/FinSet/FiniteType.agda @@ -108,8 +108,9 @@ module _ (f : X → Y) (q : (y : Y) → isFinType 0 (fiber f y)) where - ∥f∥₂ : ∥ X ∥₂ → ∥ Y ∥₂ - ∥f∥₂ = Set.map f + private + ∥f∥₂ : ∥ X ∥₂ → ∥ Y ∥₂ + ∥f∥₂ = Set.map f module _ (y : Y) where diff --git a/Cubical/Data/FinSet/Induction.agda b/Cubical/Data/FinSet/Induction.agda index e3dfec1643..e240fad93c 100644 --- a/Cubical/Data/FinSet/Induction.agda +++ b/Cubical/Data/FinSet/Induction.agda @@ -93,7 +93,7 @@ module _ -- every finite sets are merely equal to some 𝔽in ∣≡𝔽in∣ : (X : FinSet ℓ) → ∥ Σ[ n ∈ ℕ ] X ≡ 𝔽in n ∥ -∣≡𝔽in∣ X = Prop.rec isPropPropTrunc (λ (n , p) → ∣ n , path X (n , p) ∣) (isFinSet→isFinSet' (X .snd)) +∣≡𝔽in∣ X = Prop.map (λ (n , p) → n , path X (n , p)) (isFinSet→isFinSet' (X .snd)) where path : (X : FinSet ℓ) → ((n , _) : isFinOrd (X .fst)) → X ≡ 𝔽in n path X (n , p) i .fst = ua (p ⋆ invEquiv (𝔽in≃Fin n)) i diff --git a/Cubical/Data/FinSet/Quotients.agda b/Cubical/Data/FinSet/Quotients.agda index 05f76a9d9c..d376366b6d 100644 --- a/Cubical/Data/FinSet/Quotients.agda +++ b/Cubical/Data/FinSet/Quotients.agda @@ -53,8 +53,8 @@ module _ isSetℙEff = isSetΠ (λ _ → isSetBool) ℙEff→ℙDec : ℙEff → ℙDec {ℓ' = ℓ'} X - ℙEff→ℙDec f .fst x = Bool→Type* (f x) , isPropBool→Prop* - ℙEff→ℙDec f .snd x = DecBool→Prop* + ℙEff→ℙDec f .fst x = Bool→Type* (f x) , isPropBool→Type* + ℙEff→ℙDec f .snd x = DecBool→Type* Iso-ℙEff-ℙDec : Iso ℙEff (ℙDec {ℓ' = ℓ'} X) Iso-ℙEff-ℙDec .fun = ℙEff→ℙDec @@ -114,11 +114,11 @@ module _ → ((a : X .fst) → Bool→Type* {ℓ = ℓ'} (f a) ≃ ∥ R a x ∥) → (a : X .fst) → f a ≡ dec a x .fst isEqClass→isEqClassEff' f x h a = - Bool→PropInj* _ _ + Bool→TypeInj* _ _ (h a ⋆ propTruncIdempotent≃ (isDecProp→isProp (dec a x)) ⋆ LiftDecProp (dec a x)) - isEqClassRff→isEqClass : (f : ℙEff (X .fst)) → isEqClassEff f ≃ isEqClass {ℓ' = ℓ'} _ R (ℙEff→ℙDec _ f .fst) - isEqClassRff→isEqClass f = + isEqClassEff→isEqClass : (f : ℙEff (X .fst)) → isEqClassEff f ≃ isEqClass {ℓ' = ℓ'} _ R (ℙEff→ℙDec _ f .fst) + isEqClassEff→isEqClass f = propBiimpl→Equiv isPropPropTrunc isPropPropTrunc (Prop.map (λ (x , p) → x , isEqClassEff→isEqClass' f x p)) (Prop.map (λ (x , p) → x , isEqClass→isEqClassEff' f x p)) @@ -127,7 +127,7 @@ module _ _∥Eff_ = Σ[ f ∈ ℙEff (X .fst) ] isEqClassEff f ∥Eff≃∥Dec : _∥Eff_ ≃ _∥Dec_ (X .fst) R (λ x x' → isDecProp→Dec (dec x x')) - ∥Eff≃∥Dec = Σ-cong-equiv (ℙEff≃ℙDec (X .fst)) isEqClassRff→isEqClass + ∥Eff≃∥Dec = Σ-cong-equiv (ℙEff≃ℙDec (X .fst)) isEqClassEff→isEqClass isFinSet∥Eff : isFinSet _∥Eff_ isFinSet∥Eff = isFinSetSub (_ , isFinSetℙEff X) (λ _ → _ , isDecPropIsEqClassEff) diff --git a/Cubical/Data/SumFin/Properties.agda b/Cubical/Data/SumFin/Properties.agda index 51bc8623e8..1eee9c5cf6 100644 --- a/Cubical/Data/SumFin/Properties.agda +++ b/Cubical/Data/SumFin/Properties.agda @@ -96,11 +96,11 @@ SumFin∥∥Dec : (n : ℕ) → Dec ∥ Fin n ∥ SumFin∥∥Dec 0 = no (Prop.rec isProp⊥ (idfun _)) SumFin∥∥Dec (suc n) = yes ∣ inl tt ∣ -isNonZero : ℕ → Bool -isNonZero 0 = false -isNonZero (suc n) = true +ℕ→Bool : ℕ → Bool +ℕ→Bool 0 = false +ℕ→Bool (suc n) = true -SumFin∥∥DecProp : (n : ℕ) → ∥ Fin n ∥ ≃ Bool→Type (isNonZero n) +SumFin∥∥DecProp : (n : ℕ) → ∥ Fin n ∥ ≃ Bool→Type (ℕ→Bool n) SumFin∥∥DecProp 0 = uninhabEquiv (Prop.rec isProp⊥ ⊥.rec) ⊥.rec SumFin∥∥DecProp (suc n) = isContr→≃Unit (inhProp→isContr ∣ inl tt ∣ isPropPropTrunc) @@ -166,7 +166,7 @@ trueForAll (suc n) f = f (inl tt) and trueForAll n (f ∘ inr) SumFin∃→ : (n : ℕ)(f : Fin n → Bool) → Σ _ (Bool→Type ∘ f) → Bool→Type (trueForSome n f) SumFin∃→ 0 _ = ΣEmpty _ .fst SumFin∃→ (suc n) f = - Bool→Prop⊎' _ _ + Bool→Type⊎' _ _ ∘ map-⊎ (ΣUnit (Bool→Type ∘ f ∘ inl) .fst) (SumFin∃→ n (f ∘ inr)) ∘ Σ⊎≃ .fst @@ -175,12 +175,12 @@ SumFin∃← 0 _ = ⊥.rec SumFin∃← (suc n) f = invEq Σ⊎≃ ∘ map-⊎ (invEq (ΣUnit (Bool→Type ∘ f ∘ inl))) (SumFin∃← n (f ∘ inr)) - ∘ Bool→Prop⊎ _ _ + ∘ Bool→Type⊎ _ _ SumFin∃≃ : (n : ℕ)(f : Fin n → Bool) → ∥ Σ (Fin n) (Bool→Type ∘ f) ∥ ≃ Bool→Type (trueForSome n f) SumFin∃≃ n f = - propBiimpl→Equiv isPropPropTrunc isPropBool→Prop - (Prop.rec isPropBool→Prop (SumFin∃→ n f)) + propBiimpl→Equiv isPropPropTrunc isPropBool→Type + (Prop.rec isPropBool→Type (SumFin∃→ n f)) (∣_∣ ∘ SumFin∃← n f) SumFin∀≃ : (n : ℕ)(f : Fin n → Bool) → ((x : Fin n) → Bool→Type (f x)) ≃ Bool→Type (trueForAll n f) @@ -188,7 +188,7 @@ SumFin∀≃ 0 _ = isContr→≃Unit (isContrΠ⊥) SumFin∀≃ (suc n) f = Π⊎≃ ⋆ Σ-cong-equiv (ΠUnit (Bool→Type ∘ f ∘ inl)) (λ _ → SumFin∀≃ n (f ∘ inr)) - ⋆ Bool→Prop≃ _ _ + ⋆ Bool→Type×≃ _ _ -- internal equality From 86e6f239d7ff5f5f31a24b61c079403c2c7f7127 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Sat, 26 Mar 2022 17:40:16 +0800 Subject: [PATCH 19/24] reorganize decidable predicates --- Cubical/Data/FinSet/DecidablePredicate.agda | 65 +++++-------------- Cubical/Data/FinSet/FiniteType.agda | 2 + Cubical/Experiments/Combinatorics.agda | 6 +- .../Nullary/DecidablePropositions.agda | 33 +++++++++- 4 files changed, 54 insertions(+), 52 deletions(-) diff --git a/Cubical/Data/FinSet/DecidablePredicate.agda b/Cubical/Data/FinSet/DecidablePredicate.agda index 814ba7cf8b..65025ff805 100644 --- a/Cubical/Data/FinSet/DecidablePredicate.agda +++ b/Cubical/Data/FinSet/DecidablePredicate.agda @@ -1,8 +1,8 @@ {- This files contains: -- An alternative formulation of decidable propositions; - Lots of useful properties about (this) decidable predicates on finite sets. + (P.S. We use the alternative definition of decidability for computational effectivity.) -} {-# OPTIONS --safe #-} @@ -27,32 +27,13 @@ open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidablePropositions + hiding (DecProp) renaming (DecProp' to DecProp) private variable ℓ ℓ' ℓ'' ℓ''' : Level -isDecProp : Type ℓ → Type ℓ -isDecProp P = Σ[ t ∈ Bool ] P ≃ Bool→Type t - -isDecProp→isProp : {P : Type ℓ} → isDecProp P → isProp P -isDecProp→isProp h = isOfHLevelRespectEquiv 1 (invEquiv (h .snd)) isPropBool→Type - -isDecProp→Dec : {P : Type ℓ} → isDecProp P → Dec P -isDecProp→Dec h = EquivPresDec (invEquiv (h .snd)) DecBool→Type - -isPropIsDecProp : {P : Type ℓ} → isProp (isDecProp P) -isPropIsDecProp p q = - Σ≡PropEquiv (λ _ → isOfHLevel⁺≃ᵣ 0 isPropBool→Type) .fst - (Bool→TypeInj _ _ (invEquiv (p .snd) ⋆ q .snd)) - -DecProp : (ℓ : Level) → Type (ℓ-suc ℓ) -DecProp ℓ = Σ[ P ∈ Type ℓ ] isDecProp P - -isDecPropRespectEquiv : {P : Type ℓ} {Q : Type ℓ'} - → P ≃ Q → isDecProp Q → isDecProp P -isDecPropRespectEquiv e (t , e') = t , e ⋆ e' - module _ (X : Type ℓ)(p : isFinOrd X) where @@ -105,25 +86,17 @@ module _ (P : X .fst → DecProp ℓ') where isFinSetSub : isFinSet (Σ (X .fst) (λ x → P x .fst)) - isFinSetSub = - Prop.elim - (λ _ → isPropIsFinSet) - (λ p → isFinOrd→isFinSet (isFinOrdSub (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd))) - (X .snd .snd) + isFinSetSub = Prop.rec isPropIsFinSet + (λ p → isFinOrd→isFinSet (isFinOrdSub (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd))) + (X .snd .snd) isDecProp∃ : isDecProp ∥ Σ (X .fst) (λ x → P x .fst) ∥ - isDecProp∃ = - Prop.elim - (λ _ → isPropIsDecProp) - (λ p → isDecProp∃' (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd)) - (X .snd .snd) + isDecProp∃ = Prop.rec isPropIsDecProp + (λ p → isDecProp∃' (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd)) (X .snd .snd) isDecProp∀ : isDecProp ((x : X .fst) → P x .fst) - isDecProp∀ = - Prop.elim - (λ _ → isPropIsDecProp) - (λ p → isDecProp∀' (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd)) - (X .snd .snd) + isDecProp∀ = Prop.rec isPropIsDecProp + (λ p → isDecProp∀' (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd)) (X .snd .snd) module _ (X : FinSet ℓ) @@ -146,10 +119,8 @@ module _ (X : FinSet ℓ) where isDecProp≡ : (a b : X .fst) → isDecProp (a ≡ b) - isDecProp≡ a b = - Prop.rec isPropIsDecProp - (λ p → isDecProp≡' (X .fst) (_ , p) a b) - (X .snd .snd) + isDecProp≡ a b = Prop.rec isPropIsDecProp + (λ p → isDecProp≡' (X .fst) (_ , p) a b) (X .snd .snd) module _ (P : DecProp ℓ ) @@ -163,13 +134,9 @@ module _ (X : FinSet ℓ) where isDecProp¬ : isDecProp (¬ (X .fst)) - isDecProp¬ = - Prop.rec isPropIsDecProp - (λ p → isDecProp¬' (X .fst) (_ , p)) - (X .snd .snd) + isDecProp¬ = Prop.rec isPropIsDecProp + (λ p → isDecProp¬' (X .fst) (_ , p)) (X .snd .snd) isDecProp∥∥ : isDecProp ∥ X .fst ∥ - isDecProp∥∥ = - Prop.rec isPropIsDecProp - (λ p → isDecProp∥∥' (X .fst) (_ , p)) - (X .snd .snd) + isDecProp∥∥ = Prop.rec isPropIsDecProp + (λ p → isDecProp∥∥' (X .fst) (_ , p)) (X .snd .snd) diff --git a/Cubical/Data/FinSet/FiniteType.agda b/Cubical/Data/FinSet/FiniteType.agda index 7394fbba47..2f213b464c 100644 --- a/Cubical/Data/FinSet/FiniteType.agda +++ b/Cubical/Data/FinSet/FiniteType.agda @@ -44,6 +44,8 @@ open import Cubical.Data.FinSet.Quotients open import Cubical.Data.FinSet.Cardinality open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidablePropositions + hiding (DecProp) renaming (DecProp' to DecProp) open import Cubical.Relation.Binary private diff --git a/Cubical/Experiments/Combinatorics.agda b/Cubical/Experiments/Combinatorics.agda index 0bce7bd8ce..ecf9ab7646 100644 --- a/Cubical/Experiments/Combinatorics.agda +++ b/Cubical/Experiments/Combinatorics.agda @@ -20,14 +20,16 @@ open import Cubical.Data.Vec open import Cubical.Data.SumFin renaming (Fin to Fin') open import Cubical.Data.FinSet.Base open import Cubical.Data.FinSet.Properties -open import Cubical.Data.FinSet.Constructor +open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Cardinality open import Cubical.Data.FinSet.DecidablePredicate -open import Cubical.Data.FinSet.Quotient +open import Cubical.Data.FinSet.Quotients open import Cubical.HITs.PropositionalTruncation open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidablePropositions + hiding (DecProp) renaming (DecProp' to DecProp) open import Cubical.Relation.Binary open import Cubical.Functions.Embedding diff --git a/Cubical/Relation/Nullary/DecidablePropositions.agda b/Cubical/Relation/Nullary/DecidablePropositions.agda index e7c31443aa..94f2acd2c1 100644 --- a/Cubical/Relation/Nullary/DecidablePropositions.agda +++ b/Cubical/Relation/Nullary/DecidablePropositions.agda @@ -3,16 +3,47 @@ module Cubical.Relation.Nullary.DecidablePropositions where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv renaming (_∙ₑ_ to _⋆_) + +open import Cubical.Data.Bool +open import Cubical.Data.Sigma open import Cubical.Relation.Nullary.Base open import Cubical.Relation.Nullary.Properties private variable - ℓ : Level + ℓ ℓ' : Level DecProp : (ℓ : Level) → Type (ℓ-suc ℓ) DecProp ℓ = Σ[ P ∈ hProp ℓ ] Dec (P .fst) isSetDecProp : isSet (DecProp ℓ) isSetDecProp = isOfHLevelΣ 2 isSetHProp (λ P → isProp→isSet (isPropDec (P .snd))) + +-- the following is an alternative formulation of decidable equality +-- it separates the boolean value from more proof-relevant part +-- so it performs better when doing computation + +isDecProp : Type ℓ → Type ℓ +isDecProp P = Σ[ t ∈ Bool ] P ≃ Bool→Type t + +DecProp' : (ℓ : Level) → Type (ℓ-suc ℓ) +DecProp' ℓ = Σ[ P ∈ Type ℓ ] isDecProp P + +-- properties of the alternative formulation + +isDecProp→isProp : {P : Type ℓ} → isDecProp P → isProp P +isDecProp→isProp h = isOfHLevelRespectEquiv 1 (invEquiv (h .snd)) isPropBool→Type + +isDecProp→Dec : {P : Type ℓ} → isDecProp P → Dec P +isDecProp→Dec h = EquivPresDec (invEquiv (h .snd)) DecBool→Type + +isPropIsDecProp : {P : Type ℓ} → isProp (isDecProp P) +isPropIsDecProp p q = + Σ≡PropEquiv (λ _ → isOfHLevel⁺≃ᵣ 0 isPropBool→Type) .fst + (Bool→TypeInj _ _ (invEquiv (p .snd) ⋆ q .snd)) + +isDecPropRespectEquiv : {P : Type ℓ} {Q : Type ℓ'} + → P ≃ Q → isDecProp Q → isDecProp P +isDecPropRespectEquiv e (t , e') = t , e ⋆ e' From adcadb6f6d9c946233238d1aa3491e6a5d412308 Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Sat, 26 Mar 2022 17:58:20 +0800 Subject: [PATCH 20/24] reorganizing finite types --- Cubical/Data/FinSet/FiniteType.agda | 156 ------------------ Cubical/Data/FinType.agda | 6 + Cubical/Data/FinType/Base.agda | 63 +++++++ .../{FinSet => FinType}/FiniteStructure.agda | 9 +- Cubical/Data/FinType/Properties.agda | 42 +++++ Cubical/Data/FinType/Sigma.agda | 84 ++++++++++ .../Experiments/CountingFiniteStructure.agda | 9 +- 7 files changed, 202 insertions(+), 167 deletions(-) delete mode 100644 Cubical/Data/FinSet/FiniteType.agda create mode 100644 Cubical/Data/FinType.agda create mode 100644 Cubical/Data/FinType/Base.agda rename Cubical/Data/{FinSet => FinType}/FiniteStructure.agda (92%) create mode 100644 Cubical/Data/FinType/Properties.agda create mode 100644 Cubical/Data/FinType/Sigma.agda diff --git a/Cubical/Data/FinSet/FiniteType.agda b/Cubical/Data/FinSet/FiniteType.agda deleted file mode 100644 index 2f213b464c..0000000000 --- a/Cubical/Data/FinSet/FiniteType.agda +++ /dev/null @@ -1,156 +0,0 @@ -{- - -Finite Types - -This file formalize the notion of Rijke finite type, -which is a direct generalization of Bishop finite set. -Basically, a type is (Rijke) n-finite if its i-th order -homotopy groups πᵢ are Bishop finite for i ≤ n. - -Referrences: - https://github.com/EgbertRijke/OEIS-A000001 - -This file contains: -- The definition and basic properties of Rijke finite types; -- Rijke finiteness is closed under forming Σ-type. - --} -{-# OPTIONS --safe #-} - -module Cubical.Data.FinSet.FiniteType 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.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.Properties -open import Cubical.Foundations.Path - -open import Cubical.HITs.PropositionalTruncation as Prop -open import Cubical.HITs.SetTruncation as Set -open import Cubical.HITs.SetTruncation.Fibers -open import Cubical.HITs.SetQuotients as SetQuot - -open import Cubical.Data.Nat -open import Cubical.Data.Sigma - -open import Cubical.Data.FinSet.Base -open import Cubical.Data.FinSet.Properties -open import Cubical.Data.FinSet.DecidablePredicate -open import Cubical.Data.FinSet.Constructors -open import Cubical.Data.FinSet.Quotients -open import Cubical.Data.FinSet.Cardinality - -open import Cubical.Relation.Nullary -open import Cubical.Relation.Nullary.DecidablePropositions - hiding (DecProp) renaming (DecProp' to DecProp) -open import Cubical.Relation.Binary - -private - variable - ℓ ℓ' : Level - n : ℕ - X : Type ℓ - Y : Type ℓ' - --- the (Rijke) finite type - -isFinType : (n : ℕ) → Type ℓ → Type ℓ -isFinType 0 X = isFinSet ∥ X ∥₂ -isFinType (suc n) X = (isFinType 0 X) × ((a b : X) → isFinType n (a ≡ b)) - -isPropIsFinType : isProp (isFinType n X) -isPropIsFinType {n = 0} = isPropIsFinSet -isPropIsFinType {n = suc n} = isProp× isPropIsFinSet (isPropΠ2 (λ _ _ → isPropIsFinType {n = n})) - --- the type of finite types - -FinType : (ℓ : Level)(n : ℕ) → Type (ℓ-suc ℓ) -FinType ℓ n = TypeWithStr ℓ (isFinType n) - --- basic numerical implications - -isFinTypeSuc→isFinType : isFinType (suc n) X → isFinType n X -isFinTypeSuc→isFinType {n = 0} p = p .fst -isFinTypeSuc→isFinType {n = suc n} p .fst = p .fst -isFinTypeSuc→isFinType {n = suc n} p .snd a b = isFinTypeSuc→isFinType {n = n} (p .snd a b) - -isFinType→isFinType0 : isFinType n X → isFinType 0 X -isFinType→isFinType0 {n = 0} p = p -isFinType→isFinType0 {n = suc n} p = p .fst - -isFinTypeSuc→isFinType1 : isFinType (suc n) X → isFinType 1 X -isFinTypeSuc→isFinType1 {n = 0} p = p -isFinTypeSuc→isFinType1 {n = suc n} p .fst = p .fst -isFinTypeSuc→isFinType1 {n = suc n} p .snd a b = isFinType→isFinType0 {n = suc n} (p .snd a b) - --- useful properties - -EquivPresIsFinType : (n : ℕ) → X ≃ Y → isFinType n X → isFinType n Y -EquivPresIsFinType 0 e = EquivPresIsFinSet (isoToEquiv (setTruncIso (equivToIso e))) -EquivPresIsFinType (suc n) e (p , q) .fst = EquivPresIsFinType 0 e p -EquivPresIsFinType (suc n) e (p , q) .snd a b = - EquivPresIsFinType n (invEquiv (congEquiv (invEquiv e))) (q _ _) - -isFinSet→isFinType : (n : ℕ) → isFinSet X → isFinType n X -isFinSet→isFinType 0 p = EquivPresIsFinSet (invEquiv (setTruncIdempotent≃ (isFinSet→isSet p))) p -isFinSet→isFinType (suc n) p .fst = isFinSet→isFinType 0 p -isFinSet→isFinType (suc n) p .snd a b = isFinSet→isFinType n (isFinSet≡ (_ , p) _ _) - -isPathConnected→isFinType0 : isContr ∥ X ∥₂ → isFinType 0 X -isPathConnected→isFinType0 p = isContr→isFinSet p - -{- finite types are closed under forming Σ-types -} - -module _ - (X : Type ℓ ) - (Y : Type ℓ') (h : isFinType 1 Y) - (f : X → Y) - (q : (y : Y) → isFinType 0 (fiber f y)) where - - private - ∥f∥₂ : ∥ X ∥₂ → ∥ Y ∥₂ - ∥f∥₂ = Set.map f - - module _ (y : Y) where - - isDecPropFiberRel : (x x' : ∥ fiber f y ∥₂) → isDecProp (fiberRel f y x x') - isDecPropFiberRel x x' = - isDecPropRespectEquiv (fiberRel1≃2 f y x x') - (isDecProp∃ (_ , h .snd y y) (λ _ → _ , isDecProp≡ (_ , q y) _ _)) - - isFinSetFiber∥∥₂' : isFinSet (fiber ∥f∥₂ ∣ y ∣₂) - isFinSetFiber∥∥₂' = - EquivPresIsFinSet (∥fiber∥₂/R≃fiber∥∥₂ f y) - (isFinSetQuot (_ , q y) (fiberRel f y) (isEquivRelFiberRel f y) isDecPropFiberRel) - - isFinSetFiber∥∥₂ : (y : ∥ Y ∥₂) → isFinSet (fiber ∥f∥₂ y) - isFinSetFiber∥∥₂ = Set.elim (λ _ → isProp→isSet isPropIsFinSet) isFinSetFiber∥∥₂' - - isFinType0Total : isFinType 0 X - isFinType0Total = isFinSetTotal ∥ X ∥₂ (∥ Y ∥₂ , h .fst) ∥f∥₂ isFinSetFiber∥∥₂ - -module _ - (X : FinType ℓ 1) - (Y : X .fst → FinType ℓ' 0) where - - isFinType0Σ : isFinType 0 (Σ (X .fst) (λ x → Y x .fst)) - isFinType0Σ = - isFinType0Total (Σ (X .fst) (λ x → Y x .fst)) (X .fst) (X .snd) fst - (λ x → EquivPresIsFinType 0 (fiberProjEquiv _ _ x) (Y x .snd)) - --- the closedness result - -isFinTypeΣ : {n : ℕ} - (X : FinType ℓ (1 + n)) - (Y : X .fst → FinType ℓ' n) - → isFinType n (Σ (X .fst) (λ x → Y x .fst)) -isFinTypeΣ {n = 0} = isFinType0Σ -isFinTypeΣ {n = suc n} X Y .fst = - isFinType0Σ (_ , isFinTypeSuc→isFinType1 {n = suc n} (X .snd)) - (λ x → _ , isFinType→isFinType0 {n = suc n} (Y x .snd)) -isFinTypeΣ {n = suc n} X Y .snd a b = - EquivPresIsFinType n (ΣPathTransport≃PathΣ a b) - (isFinTypeΣ {n = n} (_ , X .snd .snd _ _) (λ _ → _ , Y _ .snd .snd _ _)) diff --git a/Cubical/Data/FinType.agda b/Cubical/Data/FinType.agda new file mode 100644 index 0000000000..f92d28a66d --- /dev/null +++ b/Cubical/Data/FinType.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.Data.FinType where + +open import Cubical.Data.FinType.Base public +open import Cubical.Data.FinType.Properties public diff --git a/Cubical/Data/FinType/Base.agda b/Cubical/Data/FinType/Base.agda new file mode 100644 index 0000000000..b700e887f3 --- /dev/null +++ b/Cubical/Data/FinType/Base.agda @@ -0,0 +1,63 @@ +{- + +Finite Types + +This file formalize the notion of Rijke finite type, +which is a direct generalization of Bishop finite set. +Basically, a type is (Rijke) n-finite if its i-th order +homotopy groups πᵢ are Bishop finite for i ≤ n. + +Referrences: + https://github.com/EgbertRijke/OEIS-A000001 + +-} +{-# OPTIONS --safe #-} + +module Cubical.Data.FinType.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure + +open import Cubical.HITs.SetTruncation + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.FinSet + +private + variable + ℓ ℓ' : Level + n : ℕ + X : Type ℓ + +-- the (Rijke) finite type + +isFinType : (n : ℕ) → Type ℓ → Type ℓ +isFinType 0 X = isFinSet ∥ X ∥₂ +isFinType (suc n) X = (isFinType 0 X) × ((a b : X) → isFinType n (a ≡ b)) + +isPropIsFinType : isProp (isFinType n X) +isPropIsFinType {n = 0} = isPropIsFinSet +isPropIsFinType {n = suc n} = isProp× isPropIsFinSet (isPropΠ2 (λ _ _ → isPropIsFinType {n = n})) + +-- the type of finite types + +FinType : (ℓ : Level)(n : ℕ) → Type (ℓ-suc ℓ) +FinType ℓ n = TypeWithStr ℓ (isFinType n) + +-- basic numerical implications + +isFinTypeSuc→isFinType : isFinType (suc n) X → isFinType n X +isFinTypeSuc→isFinType {n = 0} p = p .fst +isFinTypeSuc→isFinType {n = suc n} p .fst = p .fst +isFinTypeSuc→isFinType {n = suc n} p .snd a b = isFinTypeSuc→isFinType {n = n} (p .snd a b) + +isFinType→isFinType0 : isFinType n X → isFinType 0 X +isFinType→isFinType0 {n = 0} p = p +isFinType→isFinType0 {n = suc n} p = p .fst + +isFinTypeSuc→isFinType1 : isFinType (suc n) X → isFinType 1 X +isFinTypeSuc→isFinType1 {n = 0} p = p +isFinTypeSuc→isFinType1 {n = suc n} p .fst = p .fst +isFinTypeSuc→isFinType1 {n = suc n} p .snd a b = isFinType→isFinType0 {n = suc n} (p .snd a b) diff --git a/Cubical/Data/FinSet/FiniteStructure.agda b/Cubical/Data/FinType/FiniteStructure.agda similarity index 92% rename from Cubical/Data/FinSet/FiniteStructure.agda rename to Cubical/Data/FinType/FiniteStructure.agda index 8c297070bf..b34df802cd 100644 --- a/Cubical/Data/FinSet/FiniteStructure.agda +++ b/Cubical/Data/FinType/FiniteStructure.agda @@ -11,7 +11,7 @@ This file contains: -} {-# OPTIONS --safe #-} -module Cubical.Data.FinSet.FiniteStructure where +module Cubical.Data.FinType.FiniteStructure where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism @@ -23,12 +23,11 @@ open import Cubical.HITs.SetTruncation as Set open import Cubical.Data.Nat open import Cubical.Data.Sigma -open import Cubical.Data.FinSet.Base -open import Cubical.Data.FinSet.Properties +open import Cubical.Data.FinSet open import Cubical.Data.FinSet.Induction -open import Cubical.Data.FinSet.Constructors open import Cubical.Data.FinSet.Cardinality -open import Cubical.Data.FinSet.FiniteType +open import Cubical.Data.FinType +open import Cubical.Data.FinType.Sigma private variable diff --git a/Cubical/Data/FinType/Properties.agda b/Cubical/Data/FinType/Properties.agda new file mode 100644 index 0000000000..befeba15a6 --- /dev/null +++ b/Cubical/Data/FinType/Properties.agda @@ -0,0 +1,42 @@ +{- + +This file contains: +- Some basic properties of Rijke finite types. + +-} +{-# OPTIONS --safe #-} + +module Cubical.Data.FinType.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties + +open import Cubical.HITs.SetTruncation + +open import Cubical.Data.Nat +open import Cubical.Data.FinSet +open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinType.Base + +private + variable + ℓ ℓ' : Level + n : ℕ + X : Type ℓ + Y : Type ℓ' + +EquivPresIsFinType : (n : ℕ) → X ≃ Y → isFinType n X → isFinType n Y +EquivPresIsFinType 0 e = EquivPresIsFinSet (isoToEquiv (setTruncIso (equivToIso e))) +EquivPresIsFinType (suc n) e (p , q) .fst = EquivPresIsFinType 0 e p +EquivPresIsFinType (suc n) e (p , q) .snd a b = + EquivPresIsFinType n (invEquiv (congEquiv (invEquiv e))) (q _ _) + +isFinSet→isFinType : (n : ℕ) → isFinSet X → isFinType n X +isFinSet→isFinType 0 p = EquivPresIsFinSet (invEquiv (setTruncIdempotent≃ (isFinSet→isSet p))) p +isFinSet→isFinType (suc n) p .fst = isFinSet→isFinType 0 p +isFinSet→isFinType (suc n) p .snd a b = isFinSet→isFinType n (isFinSet≡ (_ , p) _ _) + +isPathConnected→isFinType0 : isContr ∥ X ∥₂ → isFinType 0 X +isPathConnected→isFinType0 p = isContr→isFinSet p diff --git a/Cubical/Data/FinType/Sigma.agda b/Cubical/Data/FinType/Sigma.agda new file mode 100644 index 0000000000..0933eff68d --- /dev/null +++ b/Cubical/Data/FinType/Sigma.agda @@ -0,0 +1,84 @@ +{- + +This file contains: +- Rijke finiteness is closed under forming Σ-type. + +-} +{-# OPTIONS --safe #-} + +module Cubical.Data.FinType.Sigma where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.HITs.SetTruncation as Set +open import Cubical.HITs.SetTruncation.Fibers + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +open import Cubical.Data.FinSet +open import Cubical.Data.FinSet.DecidablePredicate +open import Cubical.Data.FinSet.Constructors +open import Cubical.Data.FinSet.Quotients +open import Cubical.Data.FinType + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidablePropositions + hiding (DecProp) renaming (DecProp' to DecProp) + +{- finite types are closed under forming Σ-types -} +private + variable + ℓ ℓ' : Level + +module _ + (X : Type ℓ ) + (Y : Type ℓ') (h : isFinType 1 Y) + (f : X → Y) + (q : (y : Y) → isFinType 0 (fiber f y)) where + + private + ∥f∥₂ : ∥ X ∥₂ → ∥ Y ∥₂ + ∥f∥₂ = Set.map f + + module _ (y : Y) where + + isDecPropFiberRel : (x x' : ∥ fiber f y ∥₂) → isDecProp (fiberRel f y x x') + isDecPropFiberRel x x' = + isDecPropRespectEquiv (fiberRel1≃2 f y x x') + (isDecProp∃ (_ , h .snd y y) (λ _ → _ , isDecProp≡ (_ , q y) _ _)) + + isFinSetFiber∥∥₂' : isFinSet (fiber ∥f∥₂ ∣ y ∣₂) + isFinSetFiber∥∥₂' = + EquivPresIsFinSet (∥fiber∥₂/R≃fiber∥∥₂ f y) + (isFinSetQuot (_ , q y) (fiberRel f y) (isEquivRelFiberRel f y) isDecPropFiberRel) + + isFinSetFiber∥∥₂ : (y : ∥ Y ∥₂) → isFinSet (fiber ∥f∥₂ y) + isFinSetFiber∥∥₂ = Set.elim (λ _ → isProp→isSet isPropIsFinSet) isFinSetFiber∥∥₂' + + isFinType0Total : isFinType 0 X + isFinType0Total = isFinSetTotal ∥ X ∥₂ (∥ Y ∥₂ , h .fst) ∥f∥₂ isFinSetFiber∥∥₂ + +module _ + (X : FinType ℓ 1) + (Y : X .fst → FinType ℓ' 0) where + + isFinType0Σ : isFinType 0 (Σ (X .fst) (λ x → Y x .fst)) + isFinType0Σ = + isFinType0Total (Σ (X .fst) (λ x → Y x .fst)) (X .fst) (X .snd) fst + (λ x → EquivPresIsFinType 0 (fiberProjEquiv _ _ x) (Y x .snd)) + +-- the closedness result + +isFinTypeΣ : {n : ℕ} + (X : FinType ℓ (1 + n)) + (Y : X .fst → FinType ℓ' n) + → isFinType n (Σ (X .fst) (λ x → Y x .fst)) +isFinTypeΣ {n = 0} = isFinType0Σ +isFinTypeΣ {n = suc n} X Y .fst = + isFinType0Σ (_ , isFinTypeSuc→isFinType1 {n = suc n} (X .snd)) + (λ x → _ , isFinType→isFinType0 {n = suc n} (Y x .snd)) +isFinTypeΣ {n = suc n} X Y .snd a b = + EquivPresIsFinType n (ΣPathTransport≃PathΣ a b) + (isFinTypeΣ {n = n} (_ , X .snd .snd _ _) (λ _ → _ , Y _ .snd .snd _ _)) diff --git a/Cubical/Experiments/CountingFiniteStructure.agda b/Cubical/Experiments/CountingFiniteStructure.agda index c872a41be4..ba1d56845b 100644 --- a/Cubical/Experiments/CountingFiniteStructure.agda +++ b/Cubical/Experiments/CountingFiniteStructure.agda @@ -16,14 +16,11 @@ open import Cubical.Data.Nat open import Cubical.Data.Bool open import Cubical.Data.Sigma -open import Cubical.Data.FinSet.Base -open import Cubical.Data.FinSet.Properties -open import Cubical.Data.FinSet.DecidablePredicate +open import Cubical.Data.FinSet open import Cubical.Data.FinSet.Induction open import Cubical.Data.FinSet.Constructors -open import Cubical.Data.FinSet.Cardinality -open import Cubical.Data.FinSet.FiniteType -open import Cubical.Data.FinSet.FiniteStructure +open import Cubical.Data.FinType +open import Cubical.Data.FinType.FiniteStructure private variable From 9f68f7fc031316529f08d6748a036c44b7972b4f Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Sat, 26 Mar 2022 18:03:17 +0800 Subject: [PATCH 21/24] rename --- Cubical/Data/FinType/Sigma.agda | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Cubical/Data/FinType/Sigma.agda b/Cubical/Data/FinType/Sigma.agda index 0933eff68d..eb3f949c70 100644 --- a/Cubical/Data/FinType/Sigma.agda +++ b/Cubical/Data/FinType/Sigma.agda @@ -27,7 +27,6 @@ open import Cubical.Relation.Nullary open import Cubical.Relation.Nullary.DecidablePropositions hiding (DecProp) renaming (DecProp' to DecProp) -{- finite types are closed under forming Σ-types -} private variable ℓ ℓ' : Level @@ -69,7 +68,7 @@ module _ isFinType0Total (Σ (X .fst) (λ x → Y x .fst)) (X .fst) (X .snd) fst (λ x → EquivPresIsFinType 0 (fiberProjEquiv _ _ x) (Y x .snd)) --- the closedness result +-- the main result isFinTypeΣ : {n : ℕ} (X : FinType ℓ (1 + n)) From fe5adfbf17c450e82dd9f20a7bd9b51d426f984b Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Sat, 26 Mar 2022 18:10:11 +0800 Subject: [PATCH 22/24] modify --- Cubical/Data/SumFin/Properties.agda | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Cubical/Data/SumFin/Properties.agda b/Cubical/Data/SumFin/Properties.agda index 1eee9c5cf6..52eeec872b 100644 --- a/Cubical/Data/SumFin/Properties.agda +++ b/Cubical/Data/SumFin/Properties.agda @@ -92,10 +92,6 @@ SumFin∥∥≃ (suc n) = isContr→≃Unit (inhProp→isContr ∣ inl tt ∣ isPropPropTrunc) ⋆ isContr→≃Unit (isContrUnit) ⋆ invEquiv (⊎-⊥-≃) -SumFin∥∥Dec : (n : ℕ) → Dec ∥ Fin n ∥ -SumFin∥∥Dec 0 = no (Prop.rec isProp⊥ (idfun _)) -SumFin∥∥Dec (suc n) = yes ∣ inl tt ∣ - ℕ→Bool : ℕ → Bool ℕ→Bool 0 = false ℕ→Bool (suc n) = true From 2f1c9028b61e0c5e2293f952223fc476e614ce3d Mon Sep 17 00:00:00 2001 From: KANG Rongji Date: Sat, 26 Mar 2022 18:40:18 +0800 Subject: [PATCH 23/24] typos --- Cubical/Relation/Nullary/DecidablePropositions.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/Relation/Nullary/DecidablePropositions.agda b/Cubical/Relation/Nullary/DecidablePropositions.agda index 94f2acd2c1..69b7ec8506 100644 --- a/Cubical/Relation/Nullary/DecidablePropositions.agda +++ b/Cubical/Relation/Nullary/DecidablePropositions.agda @@ -21,7 +21,7 @@ DecProp ℓ = Σ[ P ∈ hProp ℓ ] Dec (P .fst) isSetDecProp : isSet (DecProp ℓ) isSetDecProp = isOfHLevelΣ 2 isSetHProp (λ P → isProp→isSet (isPropDec (P .snd))) --- the following is an alternative formulation of decidable equality +-- the following is an alternative formulation of decidable propositions -- it separates the boolean value from more proof-relevant part -- so it performs better when doing computation From 6e15312fdda92f082966b899407fbffccffb7f3a Mon Sep 17 00:00:00 2001 From: Evan Cavallo Date: Sat, 26 Mar 2022 11:43:25 +0100 Subject: [PATCH 24/24] typo again ;) --- Cubical/Data/FinSet/Cardinality.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/Data/FinSet/Cardinality.agda b/Cubical/Data/FinSet/Cardinality.agda index f174e06118..4402e7b8b6 100644 --- a/Cubical/Data/FinSet/Cardinality.agda +++ b/Cubical/Data/FinSet/Cardinality.agda @@ -161,7 +161,7 @@ module _ {- formulae about cardinality -} --- results to be used in diProp.rect induction on FinSet +-- results to be used in direct induction on FinSet card𝟘 : card (𝟘 {ℓ}) ≡ 0 card𝟘 {ℓ = ℓ} = isEmpty→card≡0 (𝟘 {ℓ}) (Empty.rec*)