├── .gitignore ├── Core ├── Element.agda ├── Everything.agda ├── Opetopes.agda ├── OpetopicFamily.agda ├── OpetopicMap.agda ├── OpetopicSigma.agda ├── OpetopicType.agda └── Prelude.agda ├── Experimental ├── Faces.agda ├── FibrantGroupoids.agda ├── FibrantUniverse.agda ├── LessPositions │ ├── OpetopicMap.agda │ ├── OpetopicType.agda │ ├── Shapes.agda │ └── Structures.agda ├── Local │ ├── CartesianProduct.agda │ ├── Category.agda │ ├── OpetopicMap.agda │ ├── OpetopicType.agda │ ├── Shapes.agda │ ├── Structures.agda │ ├── Terminal.agda │ └── Universe.agda ├── NoDecs │ ├── CartesianProduct.agda │ ├── Category.agda │ ├── Groupoid.agda │ ├── NaturalTransform.agda │ ├── Opetopes.agda │ ├── OpetopicFamily.agda │ ├── OpetopicType.agda │ ├── Opposite.agda │ ├── Representables.agda │ ├── Revised │ │ ├── Category.agda │ │ ├── OpetopicType.agda │ │ ├── Shapes.agda │ │ └── Structures.agda │ ├── RevisedOpetopicType.agda │ ├── Shapes.agda │ ├── Sigma.agda │ └── Structures.agda ├── Representables.agda ├── Scratch.agda └── wip │ ├── Faces.agda │ ├── Groupoids.agda │ ├── IndexedOpetopicType.agda │ ├── Join.agda │ ├── OpetopicContext.agda │ ├── OpetopicMap.agda │ ├── OpetopicSigma.agda │ ├── OpetopicTerm.agda │ ├── OpetopicType2.agda │ ├── Representable.agda │ ├── TheUniverse.agda │ └── YonedaExtension.agda ├── Lib ├── FibrantReplacement.agda ├── Groupoid.agda ├── Join.agda ├── Shapes.agda ├── Structures.agda └── Universe.agda ├── Native ├── InductiveOpetopicTypes.agda ├── NativeOpetopicTypes.org ├── Opetopes.agda └── Test.agda ├── README.md └── opetopic-types.agda-lib /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | -------------------------------------------------------------------------------- /Core/Element.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Element.agda - Elements (i.e. global sections) 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Nat 8 | 9 | open import Core.Prelude 10 | open import Core.Opetopes 11 | open import Core.OpetopicType 12 | open import Core.OpetopicFamily 13 | 14 | module Core.Element where 15 | 16 | El : ∀ {n ℓ} (Γ : 𝕆Type n ℓ) → Type ℓ 17 | 18 | Frm-El : ∀ {n ℓ} {Γ : 𝕆Type n ℓ} (σ : El Γ) 19 | → (o : 𝒪 n) → Frm Γ o 20 | 21 | {-# TERMINATING #-} 22 | Cns-El : ∀ {n ℓ} {Γ : 𝕆Type n ℓ} (σ : El Γ) 23 | → {o : 𝒪 n} (ρ : 𝒫 o) 24 | → Cns Γ (Frm-El σ o) ρ 25 | 26 | postulate 27 | 28 | Shp-Frm-Cns : ∀ {ℓ n} (Γ : 𝕆Type n ℓ) (σ : El Γ) 29 | → {o : 𝒪 n} (ρ : 𝒫 o) (p : Pos ρ) 30 | → Frm-El σ (Typ ρ p) ↦ Shp Γ (Cns-El σ ρ) p 31 | {-# REWRITE Shp-Frm-Cns #-} 32 | 33 | η-El : ∀ {ℓ n} (Γ : 𝕆Type n ℓ) (σ : El Γ) 34 | → (o : 𝒪 n) 35 | → Cns-El σ (ηₒ o) ↦ η Γ (Frm-El σ o) 36 | {-# REWRITE η-El #-} 37 | 38 | μ-El : ∀ {n ℓ} (Γ : 𝕆Type n ℓ) (σ : El Γ) 39 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 40 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 41 | → Cns-El σ (μₒ 𝑝 𝑞) ↦ μ Γ (Cns-El σ 𝑝) (λ p → Cns-El σ (𝑞 p)) 42 | {-# REWRITE μ-El #-} 43 | 44 | El {zero} Γ = Lift Unit 45 | El {suc n} (Γₙ , Γₛₙ) = 46 | Σ (El Γₙ) (λ σ → (o : 𝒪 n) → Γₛₙ (Frm-El σ o)) 47 | 48 | Frm-El {zero} σ ● = tt* 49 | Frm-El {suc n} {Γ = Γₙ , Γₛₙ} (σₙ , σₛₙ) (o ∣ ρ) = 50 | (Frm-El σₙ o , σₛₙ o , Cns-El σₙ ρ , (λ p → σₛₙ (Typ ρ p))) 51 | 52 | Cns-El {zero} σ objₒ = tt* 53 | Cns-El {suc n} (σₙ , σₛₙ) {𝑜 ∣ ._} lfₒ = lf (σₛₙ 𝑜) 54 | Cns-El {suc n} (σₙ , σₛₙ) {𝑜 ∣ ._} (ndₒ 𝑝 𝑞 𝑟) = 55 | nd (σₛₙ 𝑜) (Cns-El σₙ 𝑝) (λ p → σₛₙ (Typ 𝑝 p)) 56 | (λ p → Cns-El σₙ (𝑞 p)) (λ p q → σₛₙ (Typ (𝑞 p) q)) 57 | (λ p → Cns-El (σₙ , σₛₙ) (𝑟 p)) 58 | 59 | -- 60 | -- Extracting the fiber of a family at an element ... 61 | -- 62 | 63 | fiber-at : ∀ {n ℓ₀ ℓ₁} {Γ : 𝕆Type n ℓ₀} (σ : El Γ) 64 | → 𝕆Fam Γ ℓ₁ → 𝕆Type n ℓ₁ 65 | 66 | postulate 67 | 68 | frm-ovr : ∀ {n ℓ₀ ℓ₁} {Γ : 𝕆Type n ℓ₀} (σ : El Γ) 69 | → (X : 𝕆Fam Γ ℓ₁) 70 | → {𝑜 : 𝒪 n} (f : Frm (fiber-at σ X) 𝑜) 71 | → Frm↓ X (Frm-El σ 𝑜) 72 | 73 | fiber-at {zero} σ X = lift tt 74 | fiber-at {suc n} (σₙ , σₛₙ) (Xₙ , Xₛₙ) = 75 | fiber-at σₙ Xₙ , λ {𝑜} f → Xₛₙ (frm-ovr σₙ Xₙ f) (σₛₙ 𝑜) 76 | 77 | 78 | -------------------------------------------------------------------------------- /Core/Everything.agda: -------------------------------------------------------------------------------- 1 | module Core.Everything where 2 | 3 | open import Core.Opetopes public 4 | open import Core.OpetopicType public 5 | open import Core.OpetopicFamily public 6 | open import Core.OpetopicMap public 7 | open import Core.OpetopicSigma public 8 | open import Core.Element public 9 | -------------------------------------------------------------------------------- /Core/OpetopicFamily.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicFamily.agda - Dependent families of opetopic types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Empty 7 | open import Cubical.Data.Unit 8 | open import Cubical.Data.Nat 9 | open import Cubical.Data.Sum 10 | 11 | open import Core.Prelude 12 | open import Core.Opetopes 13 | open import Core.OpetopicType 14 | 15 | module Core.OpetopicFamily where 16 | 17 | 𝕆Fam : ∀ {n ℓ₀} (X : 𝕆Type n ℓ₀) 18 | → (ℓ : Level) → Type (ℓ-max ℓ₀ (ℓ-suc ℓ)) 19 | 20 | Frm↓ : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 21 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) → Type ℓ 22 | 23 | Cns↓ : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 24 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} (f↓ : Frm↓ P f) 25 | → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) → Type ℓ 26 | 27 | Shp↓ : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 28 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 29 | → {𝑝 : 𝒫 𝑜} {c : Cns X f 𝑝} (c↓ : Cns↓ P f↓ c) 30 | → (p : Pos 𝑝) → Frm↓ P (Shp X c p) 31 | 32 | η↓ : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 33 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} (f↓ : Frm↓ P f) 34 | → Cns↓ P f↓ (η X f) 35 | 36 | μ↓ : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 37 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 38 | → {𝑝 : 𝒫 𝑜} {c : Cns X f 𝑝} (c↓ : Cns↓ P f↓ c) 39 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 40 | → {d : (p : Pos 𝑝) → Cns X (Shp X c p) (𝑞 p)} 41 | → (d↓ : (p : Pos 𝑝) → Cns↓ P (Shp↓ P c↓ p) (d p)) 42 | → Cns↓ P f↓ (μ X c d) 43 | 44 | postulate 45 | 46 | η↓-shp : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 47 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} (f↓ : Frm↓ P f) 48 | → (p : Pos (ηₒ 𝑜)) 49 | → Shp↓ P (η↓ P f↓) p ↦ f↓ 50 | {-# REWRITE η↓-shp #-} 51 | 52 | μ↓-shp : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 53 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} (f↓ : Frm↓ P f) 54 | → {𝑝 : 𝒫 𝑜} {c : Cns X f 𝑝} (c↓ : Cns↓ P f↓ c) 55 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 56 | → {d : (p : Pos 𝑝) → Cns X (Shp X c p) (𝑞 p)} 57 | → (d↓ : (p : Pos 𝑝) → Cns↓ P (Shp↓ P c↓ p) (d p)) 58 | → (p : Pos (μₒ 𝑝 𝑞)) 59 | → Shp↓ P (μ↓ P c↓ d↓) p ↦ Shp↓ P (d↓ (fstₚ 𝑝 𝑞 p)) (sndₚ 𝑝 𝑞 p) 60 | {-# REWRITE μ↓-shp #-} 61 | 62 | μ↓-unit-r : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 63 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 64 | → {𝑝 : 𝒫 𝑜} {c : Cns X f 𝑝} (c↓ : Cns↓ P f↓ c) 65 | → μ↓ P c↓ (λ p → η↓ P (Shp↓ P c↓ p)) ↦ c↓ 66 | {-# REWRITE μ↓-unit-r #-} 67 | 68 | μ↓-unit-l : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 69 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 70 | → {𝑞 : (p : Pos (ηₒ 𝑜)) → 𝒫 (Typ (ηₒ 𝑜) p)} 71 | → {d : (p : Pos (ηₒ 𝑜)) → Cns X f (𝑞 p)} 72 | → (d↓ : (p : Pos (ηₒ 𝑜)) → Cns↓ P f↓ (d p)) 73 | → μ↓ P (η↓ P f↓) d↓ ↦ d↓ (ηₒ-pos 𝑜) 74 | {-# REWRITE μ↓-unit-l #-} 75 | 76 | μ↓-assoc : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 77 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 78 | → {𝑝 : 𝒫 𝑜} {c : Cns X f 𝑝} (c↓ : Cns↓ P f↓ c) 79 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 80 | → {d : (p : Pos 𝑝) → Cns X (Shp X c p) (𝑞 p)} 81 | → (d↓ : (p : Pos 𝑝) → Cns↓ P (Shp↓ P c↓ p) (d p)) 82 | → {𝑟 : (p : Pos (μₒ 𝑝 𝑞)) → 𝒫 (Typ (μₒ 𝑝 𝑞) p)} 83 | → {e : (p : Pos (μₒ 𝑝 𝑞)) → Cns X (Shp X (μ X c d) p) (𝑟 p)} 84 | → (e↓ : (p : Pos (μₒ 𝑝 𝑞)) → Cns↓ P (Shp↓ P (μ↓ P c↓ d↓) p) (e p)) 85 | → μ↓ P (μ↓ P c↓ d↓) e↓ ↦ μ↓ P c↓ (λ p → μ↓ P (d↓ p) (λ q → e↓ (pairₚ 𝑝 𝑞 p q))) 86 | {-# REWRITE μ↓-assoc #-} 87 | 88 | -- 89 | -- Implementations 90 | -- 91 | 92 | 𝕆Fam {n = zero} X ℓ = Lift Unit 93 | 𝕆Fam {n = suc n} (Xₙ , Xₛₙ) ℓ = 94 | Σ[ Pₙ ∈ 𝕆Fam Xₙ ℓ ] 95 | ({𝑜 : 𝒪 n} {f : Frm Xₙ 𝑜} (f↓ : Frm↓ Pₙ f) → Xₛₙ f → Type ℓ) 96 | 97 | Frm↓ P {●} f = Lift Unit 98 | Frm↓ (Pₙ , Pₛₙ) {𝑜 ∣ 𝑝} (f , x , c , y) = 99 | Σ[ f↓ ∈ Frm↓ Pₙ f ] 100 | Σ[ x↓ ∈ Pₛₙ f↓ x ] 101 | Σ[ c↓ ∈ Cns↓ Pₙ f↓ c ] 102 | ((p : Pos 𝑝) → Pₛₙ (Shp↓ Pₙ c↓ p) (y p)) 103 | 104 | η↓-dec : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type (suc n) ℓ₀} (P : 𝕆Fam X ℓ₁) 105 | → {𝑜 : 𝒪 n} {f : Frm (fst X) 𝑜} {x : snd X f} 106 | → {f↓ : Frm↓ (fst P) f} (x↓ : (snd P) f↓ x) 107 | → (p : Pos (ηₒ 𝑜)) → snd P (Shp↓ (fst P) (η↓ (fst P) f↓) p) (η-dec X x p) 108 | η↓-dec {X = Xₙ , Xₛₙ} (Pₙ , Pₛₙ) {𝑜} {f} {x} {f↓} x↓ = 109 | ηₒ-pos-elim 𝑜 (λ p → Pₛₙ (Shp↓ Pₙ (η↓ Pₙ f↓) p) (η-dec (Xₙ , Xₛₙ) x p)) x↓ 110 | 111 | data LfCns↓ {n ℓ₀ ℓ₁} {X : 𝕆Type (suc n) ℓ₀} (P : 𝕆Fam X ℓ₁) 112 | {𝑜 : 𝒪 n} {f : Frm (fst X) 𝑜} (x : (snd X) f) 113 | : Frm↓ P (f , x , η (fst X) f , const x) → Type ℓ₁ where 114 | 115 | lf↓ : {f↓ : Frm↓ (fst P) f} (x↓ : (snd P) f↓ x) 116 | → LfCns↓ P x (f↓ , x↓ , η↓ (fst P) f↓ , const x↓) 117 | 118 | data NdCns↓ {n ℓ₀ ℓ₁} {X : 𝕆Type (suc n) ℓ₀} (P : 𝕆Fam X ℓ₁) 119 | {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 120 | {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 121 | {𝑟 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p ∣ 𝑞 p)} 122 | {f : Frm (fst X) 𝑜} (x : (snd X) f) (c : Cns (fst X) f 𝑝) 123 | (y : (p : Pos 𝑝) → (snd X) (Shp (fst X) c p)) 124 | (d : (p : Pos 𝑝) → Cns (fst X) (Shp (fst X) c p) (𝑞 p)) 125 | (z : (p : Pos 𝑝) (q : Pos (𝑞 p)) → (snd X) (Shp (fst X) (d p) q)) 126 | (ψ : (p : Pos 𝑝) → Cns X (Shp (fst X) c p , y p , d p , z p) (𝑟 p)) 127 | : Frm↓ P (f , x , μ (fst X) c d , λ p → z (fstₚ 𝑝 𝑞 p) (sndₚ 𝑝 𝑞 p)) → Type ℓ₁ where 128 | 129 | nd↓ : {f↓ : Frm↓ (fst P) f} (x↓ : (snd P) f↓ x) (c↓ : Cns↓ (fst P) f↓ c) 130 | → (y↓ : ((p : Pos 𝑝) → (snd P) (Shp↓ (fst P) c↓ p) (y p))) 131 | → (d↓ : ((p : Pos 𝑝) → Cns↓ (fst P) (Shp↓ (fst P) c↓ p) (d p))) 132 | → (z↓ : ((p : Pos 𝑝) (q : Pos (𝑞 p)) → (snd P) (Shp↓ (fst P) (d↓ p) q) (z p q))) 133 | → (ψ↓ : ((p : Pos 𝑝) → Cns↓ P (Shp↓ (fst P) c↓ p , y↓ p , d↓ p , z↓ p) (ψ p))) 134 | → NdCns↓ P x c y d z ψ (f↓ , x↓ , μ↓ (fst P) c↓ d↓ , λ p → z↓ (fstₚ 𝑝 𝑞 p) (sndₚ 𝑝 𝑞 p)) 135 | 136 | Cns↓ P {●} f c = Lift Unit 137 | Cns↓ P {𝑜 ∣ ._} f {lfₒ} (lf x) = LfCns↓ P x f 138 | Cns↓ P {𝑜 ∣ ._} f {ndₒ 𝑝 𝑞 𝑟} (nd x c y d z ψ) = NdCns↓ P x c y d z ψ f 139 | 140 | Shp↓ P {●} {𝑝 = objₒ} c↓ p = tt* 141 | Shp↓ P {𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} {c = nd x c y d z ψ} (nd↓ x↓ c↓ y↓ d↓ z↓ ψ↓) (inl tt) = _ , x↓ , c↓ , y↓ 142 | Shp↓ P {𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} {c = nd x c y d z ψ} (nd↓ x↓ c↓ y↓ d↓ z↓ ψ↓) (inr (p , q)) = Shp↓ P (ψ↓ p) q 143 | 144 | graft↓ : ∀ {n ℓ₀ ℓ} {Xₙ : 𝕆Type n ℓ₀} {Xₛₙ : {𝑜 : 𝒪 n} (f : Frm Xₙ 𝑜) → Type ℓ₀} 145 | → (Pₙ : 𝕆Fam Xₙ ℓ) (Pₛₙ : {𝑜 : 𝒪 n} {f : Frm Xₙ 𝑜} (f↓ : Frm↓ Pₙ f) → Xₛₙ f → Type ℓ) 146 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {𝑞 : 𝒫 (𝑜 ∣ 𝑝)} 147 | → {𝑟 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 148 | → {𝑡 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p ∣ 𝑟 p)} 149 | → {f : Frm Xₙ 𝑜} {x : Xₛₙ f} {c : Cns Xₙ f 𝑝} 150 | → {y : (p : Pos 𝑝) → Xₛₙ (Shp Xₙ c p)} 151 | → {ψ : Cns (Xₙ , Xₛₙ) (f , x , c , y) 𝑞} 152 | → {d : (p : Pos 𝑝) → Cns Xₙ (Shp Xₙ c p) (𝑟 p)} 153 | → {z : (p : Pos 𝑝) (q : Pos (𝑟 p)) → Xₛₙ (Shp Xₙ (d p) q)} 154 | → {ω : (p : Pos 𝑝) → Cns (Xₙ , Xₛₙ) (Shp Xₙ c p , y p , d p , z p) (𝑡 p)} 155 | → {f↓ : Frm↓ Pₙ f} (x↓ : Pₛₙ f↓ x) (c↓ : Cns↓ Pₙ f↓ c) 156 | → (y↓ : (p : Pos 𝑝) → Pₛₙ (Shp↓ Pₙ c↓ p) (y p)) 157 | → (ψ↓ : Cns↓ (Pₙ , Pₛₙ) (f↓ , x↓ , c↓ , y↓) ψ) 158 | → (d↓ : (p : Pos 𝑝) → Cns↓ Pₙ (Shp↓ Pₙ c↓ p) (d p)) 159 | → (z↓ : (p : Pos 𝑝) (q : Pos (𝑟 p)) → Pₛₙ (Shp↓ Pₙ (d↓ p) q) (z p q)) 160 | → (ω↓ : (p : Pos 𝑝) → Cns↓ (Pₙ , Pₛₙ) (Shp↓ Pₙ c↓ p , y↓ p , d↓ p , z↓ p) (ω p)) 161 | → Cns↓ (Pₙ , Pₛₙ) (f↓ , x↓ , μ↓ Pₙ c↓ d↓ , λ p → z↓ (fstₚ 𝑝 𝑟 p) (sndₚ 𝑝 𝑟 p)) 162 | (graft Xₙ Xₛₙ x c y ψ d z ω) 163 | graft↓ Pₙ Pₛₙ {𝑜} {𝑞 = lfₒ} {ψ = lf x} .x↓ ._ ._ (lf↓ x↓) dd↓ zz↓ ω↓ = ω↓ (ηₒ-pos 𝑜) 164 | graft↓ Pₙ Pₛₙ {𝑜} {𝑞 = ndₒ 𝑝 𝑞 𝑟} {𝑟 = 𝑟𝑟} {ψ = nd x c y d z ψ} .x↓ ._ ._ (nd↓ x↓ c↓ y↓ d↓ z↓ ψ↓) dd↓ zz↓ ω↓ = 165 | let d↓' p = μ↓ Pₙ (d↓ p) (λ q → dd↓ (pairₚ 𝑝 𝑞 p q)) 166 | z↓' p q = zz↓ (pairₚ 𝑝 𝑞 p (fstₚ (𝑞 p) (λ q → 𝑟𝑟 (pairₚ 𝑝 𝑞 p q)) q)) 167 | (sndₚ (𝑞 p) (λ q → 𝑟𝑟 (pairₚ 𝑝 𝑞 p q)) q) 168 | ψ↓' p = graft↓ Pₙ Pₛₙ (y↓ p) (d↓ p) (z↓ p) (ψ↓ p) 169 | (λ q → dd↓ (pairₚ 𝑝 𝑞 p q)) 170 | (λ q → zz↓ (pairₚ 𝑝 𝑞 p q)) 171 | (λ q → ω↓ (pairₚ 𝑝 𝑞 p q)) 172 | in nd↓ x↓ c↓ y↓ d↓' z↓' ψ↓' 173 | 174 | -- η↓ : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 175 | -- → {𝑜 : 𝒪 n} {f : Frm X 𝑜} (f↓ : Frm↓ P f) 176 | -- → Cns↓ P f↓ (η X f) 177 | η↓ P {●} f↓ = tt* 178 | η↓ (Pₙ , Pₛₙ) {𝑜 ∣ 𝑝} (f↓ , x↓ , c↓ , y↓) = 179 | let d↓ p = η↓ Pₙ (Shp↓ Pₙ c↓ p) 180 | z↓ p q = (y↓ p) 181 | ψ↓ p = lf↓ (y↓ p) 182 | in nd↓ x↓ c↓ y↓ d↓ z↓ ψ↓ 183 | 184 | -- μ↓ : ∀ {n ℓ₀ ℓ} {X : 𝕆Type n ℓ₀} (P : 𝕆Fam X ℓ) 185 | -- → {𝑜 : 𝒪 n} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 186 | -- → {𝑝 : 𝒫 𝑜} {c : Cns X f 𝑝} (c↓ : Cns↓ P f↓ c) 187 | -- → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 188 | -- → {d : (p : Pos 𝑝) → Cns X (Shp X c p) (𝑞 p)} 189 | -- → (d↓ : (p : Pos 𝑝) → Cns↓ P (Shp↓ P c↓ p) (d p)) 190 | -- → Cns↓ P f↓ (μ X c d) 191 | μ↓ P {●} c↓ d↓ = tt* 192 | μ↓ P {𝑜 ∣ ._} {𝑝 = lfₒ} {c = lf x} c↓ d↓ = c↓ 193 | μ↓ (Pₙ , Pₛₙ) {𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} {c = nd x c y d z ψ} (nd↓ x↓ c↓ y↓ d↓ z↓ ψ↓) ω↓ = 194 | graft↓ Pₙ Pₛₙ _ c↓ y↓ (ω↓ (inl tt)) d↓ z↓ 195 | (λ p → μ↓ (Pₙ , Pₛₙ) (ψ↓ p) (λ q → ω↓ (inr (p , q)))) 196 | 197 | -- 198 | -- Infinite dimensional families 199 | -- 200 | 201 | record 𝕆Fam∞ {n ℓ₀ ℓ₁} {Xₙ : 𝕆Type n ℓ₀} (X : 𝕆Type∞ Xₙ) 202 | (Pₙ : 𝕆Fam Xₙ ℓ₁) : Type (ℓ-max ℓ₀ (ℓ-suc ℓ₁)) where 203 | coinductive 204 | field 205 | FillTy : {𝑜 : 𝒪 n} {f : Frm Xₙ 𝑜} → Frm↓ Pₙ f → Fill X f → Type ℓ₁ 206 | HomTy : 𝕆Fam∞ (Hom X) (Pₙ , FillTy) 207 | 208 | open 𝕆Fam∞ public 209 | 210 | 211 | -------------------------------------------------------------------------------- /Core/OpetopicMap.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicMap.agda - Maps of opetopic types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Unit 8 | open import Cubical.Data.Nat 9 | 10 | open import Core.Prelude 11 | open import Core.Opetopes 12 | open import Core.OpetopicType 13 | open import Core.OpetopicFamily 14 | 15 | module Core.OpetopicMap where 16 | 17 | infixr 40 _⇒_ 18 | 19 | _⇒_ : ∀ {n ℓ₀ ℓ₁} → 𝕆Type n ℓ₀ → 𝕆Type n ℓ₁ 20 | → Type (ℓ-max ℓ₀ ℓ₁) 21 | 22 | Frm⇒ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} (σ : X ⇒ Y) 23 | → {o : 𝒪 n} → Frm X o → Frm Y o 24 | 25 | Cns⇒ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} (σ : X ⇒ Y) 26 | → {o : 𝒪 n} {𝑝 : 𝒫 o} {f : Frm X o} 27 | → Cns X f 𝑝 → Cns Y (Frm⇒ σ f) 𝑝 28 | 29 | postulate 30 | 31 | Shp-Frm⇒ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} (σ : X ⇒ Y) 32 | → {o : 𝒪 n} {𝑝 : 𝒫 o} {f : Frm X o} (c : Cns X f 𝑝) (p : Pos 𝑝) 33 | → Frm⇒ σ (Shp X c p) ↦ Shp Y (Cns⇒ σ c) p 34 | {-# REWRITE Shp-Frm⇒ #-} 35 | 36 | η⇒ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} (σ : X ⇒ Y) 37 | → {o : 𝒪 n} (f : Frm X o) 38 | → Cns⇒ σ (η X f) ↦ η Y (Frm⇒ σ f) 39 | {-# REWRITE η⇒ #-} 40 | 41 | μ⇒ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} (σ : X ⇒ Y) 42 | → {o : 𝒪 n} {f : Frm X o} 43 | → {𝑝 : 𝒫 o} (c : Cns X f 𝑝) 44 | → {ι : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 45 | → (κ : (p : Pos 𝑝) → Cns X (Shp X c p) (ι p)) 46 | → Cns⇒ σ (μ X c κ) ↦ μ Y (Cns⇒ σ c) (λ p → Cns⇒ σ (κ p)) 47 | {-# REWRITE μ⇒ #-} 48 | 49 | _⇒_ {zero} _ _ = Lift Unit 50 | _⇒_ {suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) = 51 | Σ (Xₙ ⇒ Yₙ) (λ σ → 52 | {o : 𝒪 n} {f : Frm Xₙ o} 53 | → Xₛₙ f → Yₛₙ (Frm⇒ σ f)) 54 | 55 | Frm⇒ σ {●} f = tt* 56 | Frm⇒ (σₙ , σₛₙ) {𝑜 ∣ 𝑝} (f , x , c , y) = 57 | Frm⇒ σₙ f , σₛₙ x , Cns⇒ σₙ c , λ p → σₛₙ (y p) 58 | 59 | Cns⇒ σ {●} c = tt* 60 | Cns⇒ (σₙ , σₛₙ) {𝑜 ∣ ._} {𝑝 = lfₒ} (lf x) = lf (σₛₙ x) 61 | Cns⇒ (σₙ , σₛₙ) {𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} (nd x c y d z ψ) = 62 | nd (σₛₙ x) (Cns⇒ σₙ c) (σₛₙ ∘ y) (Cns⇒ σₙ ∘ d) 63 | (λ p q → σₛₙ (z p q)) (λ p → Cns⇒ (σₙ , σₛₙ) (ψ p)) 64 | 65 | -- 66 | -- The identity substitution 67 | -- 68 | 69 | id-sub : ∀ {n ℓ} (X : 𝕆Type n ℓ) → X ⇒ X 70 | 71 | postulate 72 | 73 | id-frm : ∀ {n ℓ} (X : 𝕆Type n ℓ) 74 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 75 | → Frm⇒ (id-sub X) f ↦ f 76 | {-# REWRITE id-frm #-} 77 | 78 | id-cns : ∀ {n ℓ} (X : 𝕆Type n ℓ) 79 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 80 | → (f : Frm X 𝑜) (c : Cns X f 𝑝) 81 | → Cns⇒ (id-sub X) c ↦ c 82 | {-# REWRITE id-cns #-} 83 | 84 | id-sub {zero} X = tt* 85 | id-sub {suc n} (Xₙ , Xₛₙ) = id-sub Xₙ , λ x → x 86 | 87 | -- 88 | -- Composition 89 | -- 90 | 91 | infixr 30 _⊚_ 92 | 93 | _⊚_ : ∀ {n ℓ₀ ℓ₁ ℓ₂} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} {Z : 𝕆Type n ℓ₂} 94 | → (Y ⇒ Z) → (X ⇒ Y) → (X ⇒ Z) 95 | 96 | postulate 97 | 98 | ⊚-Frm : ∀ {n ℓ₀ ℓ₁ ℓ₂} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} {Z : 𝕆Type n ℓ₂} 99 | → (σ : Y ⇒ Z) (τ : X ⇒ Y) (o : 𝒪 n) (f : Frm X o) 100 | → Frm⇒ (σ ⊚ τ) f ↦ Frm⇒ σ (Frm⇒ τ f) 101 | {-# REWRITE ⊚-Frm #-} 102 | 103 | ⊚-assoc : ∀ {n ℓ₀ ℓ₁ ℓ₂ ℓ₃} 104 | → {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} {Z : 𝕆Type n ℓ₂} {W : 𝕆Type n ℓ₃} 105 | → (σ : Z ⇒ W) (τ : Y ⇒ Z) (γ : X ⇒ Y) 106 | → (σ ⊚ τ) ⊚ γ ↦ σ ⊚ τ ⊚ γ 107 | {-# REWRITE ⊚-assoc #-} 108 | 109 | -- And unit laws ... 110 | 111 | _⊚_ {zero} σ τ = lift tt 112 | _⊚_ {suc n} {X = Xₙ , Xₛₙ} {Yₙ , Yₛₙ} {Zₙ , Zₛₙ} (σₙ , σₛₙ) (τₙ , τₛₙ) = 113 | σₙ ⊚ τₙ , λ x → σₛₙ (τₛₙ x) 114 | 115 | -- 116 | -- Action of substitutions on families 117 | -- 118 | 119 | _[_]ty : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} {Δ : 𝕆Type n ℓ₁} 120 | → (X : 𝕆Fam Δ ℓ₂) (σ : Γ ⇒ Δ) 121 | → 𝕆Fam Γ ℓ₂ 122 | 123 | [_⊙_] : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} {Δ : 𝕆Type n ℓ₁} 124 | → {X : 𝕆Fam Δ ℓ₂} (σ : Γ ⇒ Δ) 125 | → {𝑜 : 𝒪 n} {f : Frm Γ 𝑜} 126 | → Frm↓ (X [ σ ]ty) f 127 | → Frm↓ X (Frm⇒ σ f) 128 | 129 | [_⊙_]c : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} {Δ : 𝕆Type n ℓ₁} 130 | → {X : 𝕆Fam Δ ℓ₂} (σ : Γ ⇒ Δ) 131 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm Γ 𝑜} {c : Cns Γ f 𝑝} 132 | → {f↓ : Frm↓ (X [ σ ]ty) f} (c↓ : Cns↓ (X [ σ ]ty) f↓ c) 133 | → Cns↓ X [ σ ⊙ f↓ ] (Cns⇒ σ c) 134 | 135 | postulate 136 | 137 | Shp-⊙ : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} {Δ : 𝕆Type n ℓ₁} 138 | → {X : 𝕆Fam Δ ℓ₂} (σ : Γ ⇒ Δ) 139 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm Γ 𝑜} {c : Cns Γ f 𝑝} 140 | → {f↓ : Frm↓ (X [ σ ]ty) f} 141 | → (c↓ : Cns↓ (X [ σ ]ty) f↓ c) (p : Pos 𝑝) 142 | → [ σ ⊙ Shp↓ (X [ σ ]ty) c↓ p ] ↦ Shp↓ X [ σ ⊙ c↓ ]c p 143 | {-# REWRITE Shp-⊙ #-} 144 | 145 | η-⊙ : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} {Δ : 𝕆Type n ℓ₁} 146 | → {X : 𝕆Fam Δ ℓ₂} (σ : Γ ⇒ Δ) 147 | → {𝑜 : 𝒪 n} {f : Frm Γ 𝑜} 148 | → (f↓ : Frm↓ (X [ σ ]ty) f) 149 | → [ σ ⊙ η↓ (X [ σ ]ty) f↓ ]c ↦ η↓ X [ σ ⊙ f↓ ] 150 | {-# REWRITE η-⊙ #-} 151 | 152 | μ-⊙ : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} {Δ : 𝕆Type n ℓ₁} 153 | → {X : 𝕆Fam Δ ℓ₂} (σ : Γ ⇒ Δ) 154 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm Γ 𝑜} {c : Cns Γ f 𝑝} 155 | → {f↓ : Frm↓ (X [ σ ]ty) f} (c↓ : Cns↓ (X [ σ ]ty) f↓ c) 156 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 157 | → {d : (p : Pos 𝑝) → Cns Γ (Shp Γ c p) (𝑞 p)} 158 | → (d↓ : (p : Pos 𝑝) → Cns↓ (X [ σ ]ty) (Shp↓ (X [ σ ]ty) c↓ p) (d p)) 159 | → [ σ ⊙ μ↓ (X [ σ ]ty) c↓ d↓ ]c ↦ μ↓ X [ σ ⊙ c↓ ]c (λ p → [ σ ⊙ d↓ p ]c) 160 | {-# REWRITE μ-⊙ #-} 161 | 162 | _[_]ty {zero} X σ = tt* 163 | _[_]ty {suc n} (Xₙ , Xₛₙ) (σₙ , σₛₙ) = 164 | Xₙ [ σₙ ]ty , λ {𝑜} {f} f↓ γ → Xₛₙ [ σₙ ⊙ f↓ ] (σₛₙ γ) 165 | 166 | [_⊙_] σ {●} f↓ = tt* 167 | [_⊙_] (σₙ , σₛₙ) {𝑜 ∣ 𝑝} {f = f , x , c , y} (f↓ , x↓ , c↓ , y↓) = 168 | [ σₙ ⊙ f↓ ] , x↓ , [ σₙ ⊙ c↓ ]c , y↓ 169 | 170 | [_⊙_]c σ {●} c↓ = tt* 171 | [_⊙_]c (σₙ , σₛₙ) {𝑜 ∣ ._} {𝑝 = lfₒ} {c = lf x} (lf↓ x↓) = lf↓ x↓ 172 | [_⊙_]c (σₙ , σₛₙ) {𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} {c = nd x c y d z ψ} (nd↓ x↓ c↓ y↓ d↓ z↓ ψ↓) = 173 | nd↓ x↓ [ σₙ ⊙ c↓ ]c y↓ (λ p → [ σₙ ⊙ d↓ p ]c) z↓ (λ p → [ (σₙ , σₛₙ) ⊙ (ψ↓ p) ]c) 174 | 175 | -- 176 | -- Infinite Dimensional Maps 177 | -- 178 | 179 | record [_⇒_↓_] {n ℓ} {X Y : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (Y∞ : 𝕆Type∞ Y) 180 | (α : X ⇒ Y) : Type ℓ where 181 | coinductive 182 | field 183 | Fill⇒ : {o : 𝒪 n} {f : Frm X o} → Fill X∞ f → Fill Y∞ (Frm⇒ α f) 184 | Hom⇒ : [ Hom X∞ ⇒ Hom Y∞ ↓ α , Fill⇒ ] 185 | 186 | open [_⇒_↓_] public 187 | 188 | -- Pulling back an extension along a substitution 189 | Pb∞ : ∀ {n ℓ} {X : 𝕆Type n ℓ} {Y : 𝕆Type n ℓ} 190 | → (σ : X ⇒ Y) → 𝕆Type∞ Y → 𝕆Type∞ X 191 | Fill (Pb∞ {X = X} {Y} σ Y∞) {𝑜} f = Fill Y∞ (Frm⇒ σ f) 192 | Hom (Pb∞ {X = X} {Y} σ Y∞) = 193 | Pb∞ {X = (X , λ {𝑜} f → Fill Y∞ (Frm⇒ σ f))} 194 | {Y = (Y , Fill Y∞)} (σ , λ x → x) (Hom Y∞) 195 | 196 | -- Pushing forward and extension along a substitution 197 | Pf∞ : ∀ {n ℓ} {X : 𝕆Type n ℓ} {Y : 𝕆Type n ℓ} 198 | → (σ : X ⇒ Y) → 𝕆Type∞ X → 𝕆Type∞ Y 199 | Fill (Pf∞ {X = X} {Y} σ X∞) {o} f = 200 | Σ[ f' ∈ fiber (Frm⇒ σ) f ] Fill X∞ (fst f') 201 | Hom (Pf∞ {X = X} {Y} σ X∞) = Pf∞ {X = (X , Fill X∞)} {Y = (Y , 202 | (λ {o} f → Σ-syntax (fiber (Frm⇒ σ) f) (λ f' → Fill X∞ (fst f'))))} 203 | (σ , λ {𝑜} {f} x → (f , refl) , x) (Hom X∞) 204 | 205 | 206 | -------------------------------------------------------------------------------- /Core/OpetopicSigma.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicSigma.agda - Dependent Sum of Opetopic Types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Nat 8 | 9 | open import Core.Prelude 10 | open import Core.Opetopes 11 | open import Core.OpetopicType 12 | open import Core.OpetopicFamily 13 | open import Core.Element 14 | 15 | module Core.OpetopicSigma where 16 | 17 | Σₒ : ∀ {n ℓ₀ ℓ₁} (X : 𝕆Type n ℓ₀) (P : 𝕆Fam X ℓ₁) 18 | → 𝕆Type n (ℓ-max ℓ₀ ℓ₁) 19 | 20 | -- Action on Frames 21 | fst-frm : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 22 | → {𝑜 : 𝒪 n} → Frm (Σₒ X P) 𝑜 → Frm X 𝑜 23 | 24 | snd-frm : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 25 | → {𝑜 : 𝒪 n} (f : Frm (Σₒ X P) 𝑜) → Frm↓ P (fst-frm f) 26 | 27 | -- Action on Constructors 28 | fst-cns : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 29 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm (Σₒ X P) 𝑜} 30 | → Cns (Σₒ X P) f 𝑝 → Cns X (fst-frm f) 𝑝 31 | 32 | snd-cns : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 33 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm (Σₒ X P) 𝑜} 34 | → (c : Cns (Σₒ X P) f 𝑝) → Cns↓ P (snd-frm f) (fst-cns c) 35 | 36 | record Σ-cell {n ℓ₀ ℓ₁} {Xₙ : 𝕆Type n ℓ₀} {Pₙ : 𝕆Fam Xₙ ℓ₁} 37 | (Xₛₙ : {𝑜 : 𝒪 n} → Frm Xₙ 𝑜 → Type ℓ₀) 38 | (Pₛₙ : {𝑜 : 𝒪 n} {f : Frm Xₙ 𝑜} → Frm↓ Pₙ f → Xₛₙ f → Type ℓ₁) 39 | {𝑜 : 𝒪 n} (f : Frm (Σₒ Xₙ Pₙ) 𝑜) : Type (ℓ-max ℓ₀ ℓ₁) where 40 | constructor _,_ 41 | field 42 | 43 | fstₒ : Xₛₙ (fst-frm f) 44 | sndₒ : Pₛₙ (snd-frm f) fstₒ 45 | 46 | open Σ-cell public 47 | 48 | postulate 49 | 50 | -- Calculation of shapes 51 | fst-shp : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 52 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} (f : Frm (Σₒ X P) 𝑜) 53 | → (c : Cns (Σₒ X P) f 𝑝) (p : Pos 𝑝) 54 | → Shp X (fst-cns c) p ↦ fst-frm (Shp (Σₒ X P) c p) 55 | {-# REWRITE fst-shp #-} 56 | 57 | snd-shp : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 58 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} (f : Frm (Σₒ X P) 𝑜) 59 | → (c : Cns (Σₒ X P) f 𝑝) (p : Pos 𝑝) 60 | → Shp↓ P (snd-cns c) p ↦ snd-frm (Shp (Σₒ X P) c p) 61 | {-# REWRITE snd-shp #-} 62 | 63 | -- Compatibility with η 64 | fst-η : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 65 | → {𝑜 : 𝒪 n} (f : Frm (Σₒ X P) 𝑜) 66 | → fst-cns (η (Σₒ X P) f) ↦ η X (fst-frm f) 67 | {-# REWRITE fst-η #-} 68 | 69 | snd-η : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 70 | → {𝑜 : 𝒪 n} (f : Frm (Σₒ X P) 𝑜) 71 | → snd-cns (η (Σₒ X P) f) ↦ η↓ P (snd-frm f) 72 | {-# REWRITE snd-η #-} 73 | 74 | -- Compaitbility with μ 75 | fst-μ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 76 | → {𝑜 : 𝒪 n} {f : Frm (Σₒ X P) 𝑜} 77 | → {𝑝 : 𝒫 𝑜} (c : Cns (Σₒ X P) f 𝑝) 78 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 79 | → (d : (p : Pos 𝑝) → Cns (Σₒ X P) (Shp (Σₒ X P) c p) (𝑞 p)) 80 | → fst-cns (μ (Σₒ X P) c d) ↦ μ X (fst-cns c) (λ p → fst-cns (d p)) 81 | {-# REWRITE fst-μ #-} 82 | 83 | snd-μ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 84 | → {𝑜 : 𝒪 n} {f : Frm (Σₒ X P) 𝑜} 85 | → {𝑝 : 𝒫 𝑜} (c : Cns (Σₒ X P) f 𝑝) 86 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 87 | → (d : (p : Pos 𝑝) → Cns (Σₒ X P) (Shp (Σₒ X P) c p) (𝑞 p)) 88 | → snd-cns (μ (Σₒ X P) c d) ↦ μ↓ P (snd-cns c) (λ p → snd-cns (d p)) 89 | {-# REWRITE snd-μ #-} 90 | 91 | -- Implementations 92 | Σₒ {zero} X P = tt* 93 | Σₒ {suc n} (Xₙ , Xₛₙ) (Pₙ , Pₛₙ) = 94 | Σₒ Xₙ Pₙ , Σ-cell Xₛₙ Pₛₙ 95 | 96 | fst-frm {𝑜 = ●} f = tt* 97 | fst-frm {𝑜 = 𝑜 ∣ 𝑝} (f , x , c , y) = 98 | fst-frm f , fstₒ x , fst-cns c , λ p → fstₒ (y p) 99 | 100 | snd-frm {𝑜 = ●} f = tt* 101 | snd-frm {𝑜 = 𝑜 ∣ 𝑝} (f , x , c , y) = 102 | snd-frm f , sndₒ x , snd-cns c , λ p → sndₒ (y p) 103 | 104 | fst-cns {𝑜 = ●} c = tt* 105 | fst-cns {𝑜 = 𝑜 ∣ ._} {𝑝 = lfₒ} (lf x) = lf (fstₒ x) 106 | fst-cns {𝑜 = 𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} (nd x c y d z ψ) = 107 | nd (fstₒ x) (fst-cns c) (λ p → fstₒ (y p)) 108 | (λ p → fst-cns (d p)) (λ p q → fstₒ (z p q)) 109 | (λ p → fst-cns (ψ p)) 110 | 111 | snd-cns {𝑜 = ●} c = tt* 112 | snd-cns {𝑜 = 𝑜 ∣ ._} {𝑝 = lfₒ} (lf x) = lf↓ (sndₒ x) 113 | snd-cns {𝑜 = 𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} (nd x c y d z ψ) = 114 | nd↓ (sndₒ x) (snd-cns c) (λ p → sndₒ (y p)) 115 | (λ p → snd-cns (d p)) (λ p q → sndₒ (z p q)) 116 | (λ p → snd-cns (ψ p)) 117 | 118 | -- 119 | -- Pairing 120 | -- 121 | 122 | pair-frm : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 123 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) (f↓ : Frm↓ P f) 124 | → Frm (Σₒ X P) 𝑜 125 | 126 | pair-cns : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 127 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 128 | → (c : Cns X f 𝑝) (c↓ : Cns↓ P f↓ c) 129 | → Cns (Σₒ X P) (pair-frm f f↓) 𝑝 130 | 131 | -- Axioms 132 | postulate 133 | 134 | -- Computation rules 135 | fst-pair-frm-β : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 136 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) (f↓ : Frm↓ P f) 137 | → fst-frm (pair-frm f f↓) ↦ f 138 | {-# REWRITE fst-pair-frm-β #-} 139 | 140 | snd-pair-frm-β : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 141 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) (f↓ : Frm↓ P f) 142 | → snd-frm (pair-frm f f↓) ↦ f↓ 143 | {-# REWRITE snd-pair-frm-β #-} 144 | 145 | pair-frm-η : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 146 | → {𝑜 : 𝒪 n} (f : Frm (Σₒ X P) 𝑜) 147 | → pair-frm (fst-frm f) (snd-frm f) ↦ f 148 | {-# REWRITE pair-frm-η #-} 149 | 150 | fst-pair-cns-β : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 151 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 152 | → (c : Cns X f 𝑝) (c↓ : Cns↓ P f↓ c) 153 | → fst-cns (pair-cns c c↓) ↦ c 154 | {-# REWRITE fst-pair-cns-β #-} 155 | 156 | snd-pair-cns-β : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 157 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 158 | → (c : Cns X f 𝑝) (c↓ : Cns↓ P f↓ c) 159 | → snd-cns (pair-cns c c↓) ↦ c↓ 160 | {-# REWRITE snd-pair-cns-β #-} 161 | 162 | pair-cns-η : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 163 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm (Σₒ X P) 𝑜} 164 | → (c : Cns (Σₒ X P) f 𝑝) 165 | → pair-cns (fst-cns c) (snd-cns c) ↦ c 166 | {-# REWRITE pair-cns-η #-} 167 | 168 | -- Compatibility with Shape 169 | 170 | pair-shp : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 171 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 172 | → (c : Cns X f 𝑝) (c↓ : Cns↓ P f↓ c) (p : Pos 𝑝) 173 | → Shp (Σₒ X P) (pair-cns c c↓) p ↦ pair-frm (Shp X c p) (Shp↓ P c↓ p) 174 | {-# REWRITE pair-shp #-} 175 | 176 | -- pair-shp-exp : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 177 | -- → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} (f : Frm (Σₒ X P) 𝑜) 178 | -- → (c : Cns (Σₒ X P) f 𝑝) (p : Pos 𝑝) 179 | -- → Shp (Σₒ X P) c p ↦ pair-frm (Shp X (fst-cns c) p) (Shp↓ P (snd-cns c) p) 180 | -- {-# REWRITE pair-shp-exp #-} 181 | 182 | -- Compatibility with monad operations 183 | pair-η : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 184 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) (f↓ : Frm↓ P f) 185 | → pair-cns (η X f) (η↓ P f↓) ↦ η (Σₒ X P) (pair-frm f f↓) 186 | {-# REWRITE pair-η #-} 187 | 188 | pair-μ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {P : 𝕆Fam X ℓ₁} 189 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} {f↓ : Frm↓ P f} 190 | → {𝑝 : 𝒫 𝑜} {c : Cns X f 𝑝} (c↓ : Cns↓ P f↓ c) 191 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 192 | → {d : (p : Pos 𝑝) → Cns X (Shp X c p) (𝑞 p)} 193 | → (d↓ : (p : Pos 𝑝) → Cns↓ P (Shp↓ P c↓ p) (d p)) 194 | → pair-cns (μ X c d) (μ↓ P c↓ d↓) ↦ 195 | μ (Σₒ X P) (pair-cns c c↓) (λ p → pair-cns (d p) (d↓ p)) 196 | {-# REWRITE pair-μ #-} 197 | 198 | pair-frm {𝑜 = ●} f f↓ = tt* 199 | pair-frm {𝑜 = 𝑜 ∣ 𝑝} (f , x , c , y) (f↓ , x↓ , c↓ , y↓) = 200 | pair-frm f f↓ , (x , x↓) , pair-cns c c↓ , λ p → (y p , y↓ p) 201 | 202 | pair-cns {𝑜 = ●} c c↓ = tt* 203 | pair-cns {𝑜 = 𝑜 ∣ ._} {𝑝 = lfₒ} (lf x) (lf↓ x↓) = lf (x , x↓) 204 | pair-cns {𝑜 = 𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} (nd x c y d z ψ) (nd↓ x↓ c↓ y↓ d↓ z↓ ψ↓) = 205 | nd (x , x↓) (pair-cns c c↓) (λ p → (y p , y↓ p)) 206 | (λ p → pair-cns (d p) (d↓ p)) (λ p q → (z p q , z↓ p q)) 207 | (λ p → pair-cns (ψ p) (ψ↓ p)) 208 | 209 | -- 210 | -- Testing ... 211 | -- 212 | 213 | 214 | 215 | -------------------------------------------------------------------------------- /Core/OpetopicType.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicType.agda - Opetopic Types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Nat 10 | open import Cubical.Data.Sum 11 | 12 | open import Core.Prelude 13 | open import Core.Opetopes 14 | 15 | module Core.OpetopicType where 16 | 17 | 𝕆Type : ℕ → (ℓ : Level) → Type (ℓ-suc ℓ) 18 | 19 | Frm : ∀ {n ℓ} → 𝕆Type n ℓ → 𝒪 n → Type ℓ 20 | Cns : ∀ {n ℓ} (X : 𝕆Type n ℓ) 21 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 22 | → 𝒫 𝑜 → Type ℓ 23 | Shp : ∀ {n ℓ} (X : 𝕆Type n ℓ) 24 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 25 | → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 26 | → (p : Pos 𝑝) → Frm X (Typ 𝑝 p) 27 | 28 | η : ∀ {n ℓ} (X : 𝕆Type n ℓ) 29 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 30 | → Cns X f (ηₒ 𝑜) 31 | 32 | μ : ∀ {n ℓ} (X : 𝕆Type n ℓ) 33 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 34 | → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 35 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 36 | → (d : (p : Pos 𝑝) → Cns X (Shp X c p) (𝑞 p)) 37 | → Cns X f (μₒ 𝑝 𝑞) 38 | 39 | postulate 40 | 41 | -- shape compatibilities 42 | η-pos-shp : ∀ {n ℓ} (X : 𝕆Type n ℓ) 43 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 44 | → (p : Pos (ηₒ 𝑜)) 45 | → Shp X (η X f) p ↦ f 46 | {-# REWRITE η-pos-shp #-} 47 | 48 | μ-pos-shp : ∀ {n ℓ} (X : 𝕆Type n ℓ) 49 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 50 | → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 51 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 52 | → (d : (p : Pos 𝑝) → Cns X (Shp X c p) (𝑞 p)) 53 | → (p : Pos (μₒ 𝑝 𝑞)) 54 | → Shp X (μ X c d) p ↦ Shp X (d (fstₚ 𝑝 𝑞 p)) (sndₚ 𝑝 𝑞 p) 55 | {-# REWRITE μ-pos-shp #-} 56 | 57 | -- Monad Laws 58 | μ-unit-r : ∀ {n ℓ} (X : 𝕆Type n ℓ) 59 | → {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) 60 | → {f : Frm X 𝑜} (c : Cns X f 𝑝) 61 | → μ X c (λ p → η X (Shp X c p)) ↦ c 62 | {-# REWRITE μ-unit-r #-} 63 | 64 | μ-unit-l : ∀ {n ℓ} (X : 𝕆Type n ℓ) 65 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 66 | → (𝑞 : (p : Pos (ηₒ 𝑜)) → 𝒫 (Typ (ηₒ 𝑜) p)) 67 | → (d : (p : Pos (ηₒ 𝑜)) → Cns X f (𝑞 p)) 68 | → μ X (η X f) d ↦ d (ηₒ-pos 𝑜) 69 | {-# REWRITE μ-unit-l #-} 70 | 71 | μ-assoc : ∀ {n ℓ} (X : 𝕆Type n ℓ) 72 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 73 | → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 74 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 75 | → (d : (p : Pos 𝑝) → Cns X (Shp X c p) (𝑞 p)) 76 | → (𝑟 : (p : Pos (μₒ 𝑝 𝑞)) → 𝒫 (Typ (μₒ 𝑝 𝑞) p)) 77 | → (e : (p : Pos (μₒ 𝑝 𝑞)) → Cns X (Shp X (d (fstₚ 𝑝 𝑞 p)) (sndₚ 𝑝 𝑞 p)) (𝑟 p)) 78 | → μ X (μ X c d) e 79 | ↦ μ X c (λ p → μ X (d p) (λ q → e (pairₚ 𝑝 𝑞 p q))) 80 | {-# REWRITE μ-assoc #-} 81 | 82 | -- 83 | -- Implementation of the Polynomials 84 | -- 85 | 86 | 𝕆Type zero ℓ = Lift Unit 87 | 𝕆Type (suc n) ℓ = Σ (𝕆Type n ℓ) (λ Xₙ → {𝑜 : 𝒪 n} → Frm Xₙ 𝑜 → Type ℓ) 88 | 89 | Frm {zero} X ● = Lift Unit 90 | Frm {suc n} (Xₙ , Xₛₙ) (𝑜 ∣ 𝑝) = 91 | Σ[ f ∈ Frm Xₙ 𝑜 ] 92 | Σ[ x ∈ Xₛₙ f ] 93 | Σ[ c ∈ Cns Xₙ f 𝑝 ] 94 | ((p : Pos 𝑝) → Xₛₙ (Shp Xₙ c p)) 95 | 96 | η-dec : ∀ {n ℓ} (X : 𝕆Type (suc n) ℓ) 97 | → {𝑜 : 𝒪 n} {f : Frm (fst X) 𝑜} 98 | → (x : snd X f) 99 | → (p : Pos (ηₒ 𝑜)) → snd X (Shp (fst X) (η (fst X) f) p) 100 | η-dec (Xₙ , Xₛₙ) {𝑜} {f} x = ηₒ-pos-elim 𝑜 (λ p → Xₛₙ (Shp Xₙ (η Xₙ f) p)) x 101 | 102 | data LfCns {n ℓ} (X : 𝕆Type (suc n) ℓ) {𝑜 : 𝒪 n} : Frm X (𝑜 ∣ ηₒ 𝑜) → Type ℓ where 103 | 104 | lf : {f : Frm (fst X) 𝑜} (x : (snd X) f) 105 | → LfCns X (f , x , η (fst X) f , const x) 106 | 107 | data NdCns {n ℓ} (X : 𝕆Type (suc n) ℓ) 108 | (𝑜 : 𝒪 n) (𝑝 : 𝒫 𝑜) 109 | (𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)) 110 | (𝑟 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p ∣ 𝑞 p)) 111 | 112 | : Frm X (𝑜 ∣ μₒ 𝑝 𝑞) → Type ℓ where 113 | 114 | nd : {f : Frm (fst X) 𝑜} (x : (snd X) f) (c : Cns (fst X) f 𝑝) 115 | → (y : (p : Pos 𝑝) → (snd X) (Shp (fst X) c p)) 116 | → (d : (p : Pos 𝑝) → Cns (fst X) (Shp (fst X) c p) (𝑞 p)) 117 | → (z : (p : Pos 𝑝) (q : Pos (𝑞 p)) → (snd X) (Shp (fst X) (d p) q)) 118 | → (ψ : (p : Pos 𝑝) → Cns X (Shp (fst X) c p , y p , d p , z p) (𝑟 p)) 119 | → NdCns X 𝑜 𝑝 𝑞 𝑟 (f , x , μ (fst X) c d , λ p → z (fstₚ 𝑝 𝑞 p) (sndₚ 𝑝 𝑞 p)) 120 | 121 | Cns X {●} f 𝑝 = Lift Unit 122 | Cns X {𝑜 ∣ ._} f lfₒ = LfCns X f 123 | Cns X {𝑜 ∣ ._} f (ndₒ 𝑝 𝑞 𝑟) = NdCns X 𝑜 𝑝 𝑞 𝑟 f 124 | 125 | Shp X {●} {𝑝 = objₒ} c p = tt* 126 | Shp X {𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} (nd {f} x c y d z ψ) (inl tt) = f , x , c , y 127 | Shp X {𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} (nd {f} x c y d z ψ) (inr (p , q)) = Shp X (ψ p) q 128 | 129 | graft : ∀ {n ℓ} (Xₙ : 𝕆Type n ℓ) (Xₛₙ : {𝑜 : 𝒪 n} → Frm Xₙ 𝑜 → Type ℓ) 130 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {𝑞 : 𝒫 (𝑜 ∣ 𝑝)} 131 | → {𝑟 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 132 | → {𝑡 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p ∣ 𝑟 p)} 133 | → {f : Frm Xₙ 𝑜} (x : Xₛₙ f) (c : Cns Xₙ f 𝑝) 134 | → (y : (p : Pos 𝑝) → Xₛₙ (Shp Xₙ c p)) 135 | → (ψ : Cns (Xₙ , Xₛₙ) (f , x , c , y) 𝑞) 136 | → (d : (p : Pos 𝑝) → Cns Xₙ (Shp Xₙ c p) (𝑟 p)) 137 | → (z : (p : Pos 𝑝) (q : Pos (𝑟 p)) → Xₛₙ (Shp Xₙ (d p) q)) 138 | → (ω : (p : Pos 𝑝) → Cns (Xₙ , Xₛₙ) (Shp Xₙ c p , y p , d p , z p) (𝑡 p)) 139 | → Cns (Xₙ , Xₛₙ) (f , x , μ Xₙ c d , λ p → z (fstₚ 𝑝 𝑟 p) (sndₚ 𝑝 𝑟 p)) (graftₒ 𝑞 𝑡) 140 | graft Xₙ Xₛₙ {𝑜} {𝑞 = lfₒ} .x ._ ._ (lf x) dd zz ψψ = ψψ (ηₒ-pos 𝑜) 141 | graft Xₙ Xₛₙ {𝑜} {𝑞 = ndₒ 𝑝 𝑞 𝑟} {𝑟𝑟} .x ._ ._ (nd x c y d z ψ) dd zz ψψ = 142 | let d' p = μ Xₙ (d p) (λ q → dd (pairₚ 𝑝 𝑞 p q)) 143 | z' p q = zz (pairₚ 𝑝 𝑞 p (fstₚ (𝑞 p) (λ q → 𝑟𝑟 (pairₚ 𝑝 𝑞 p q)) q)) 144 | (sndₚ (𝑞 p) (λ q → 𝑟𝑟 (pairₚ 𝑝 𝑞 p q)) q) 145 | ψ' p = graft Xₙ Xₛₙ (y p) (d p) (z p) (ψ p) 146 | (λ q → dd (pairₚ 𝑝 𝑞 p q)) 147 | (λ q → zz (pairₚ 𝑝 𝑞 p q)) 148 | (λ q → ψψ (pairₚ 𝑝 𝑞 p q)) 149 | in nd x c y d' z' ψ' 150 | 151 | -- η : ∀ {n ℓ} (X : 𝕆Type n ℓ) 152 | -- → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 153 | -- → Cns X f (ηₒ 𝑜) 154 | η X {●} f = tt* 155 | η X {𝑜 ∣ 𝑝} (f , x , c , y) = 156 | let d p = η (fst X) (Shp (fst X) c p) 157 | z p q = y p 158 | ψ p = lf (y p) 159 | in nd x c y d z ψ 160 | 161 | -- μ : ∀ {n ℓ} (X : 𝕆Type n ℓ) 162 | -- → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 163 | -- → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 164 | -- → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 165 | -- → (d : (p : Pos 𝑝) → Cns X (Shp X c p) (𝑞 p)) 166 | -- → Cns X f (μₒ (𝑝 , 𝑞)) 167 | μ X {●} c d = tt* 168 | μ X {𝑜 ∣ ._} {𝑝 = lfₒ} (lf x) ω = lf x 169 | μ X {𝑜 ∣ ._} {𝑝 = ndₒ 𝑝 𝑞 𝑟} (nd x c y d z ψ) ω = 170 | graft (fst X) (snd X) x c y (ω (inl tt)) d z 171 | (λ p → μ X (ψ p) (λ q → ω (inr (p , q)))) 172 | 173 | -- 174 | -- Infinite dimensional opetopic types 175 | -- 176 | 177 | record 𝕆Type∞ {n ℓ} (X : 𝕆Type n ℓ) : Type (ℓ-suc ℓ) where 178 | coinductive 179 | field 180 | Fill : {o : 𝒪 n} → Frm X o → Type ℓ 181 | Hom : 𝕆Type∞ (X , Fill) 182 | 183 | open 𝕆Type∞ public 184 | 185 | -- 186 | -- The skeleton of an infinite opetopic type 187 | --- 188 | 189 | Skeleton : ∀ {ℓ} (X : 𝕆Type∞ {ℓ = ℓ} tt*) 190 | → (n : ℕ) → 𝕆Type n ℓ 191 | 192 | SkeletonExt : ∀ {ℓ} (X : 𝕆Type∞ {ℓ = ℓ} tt*) 193 | → (n : ℕ) → 𝕆Type∞ {ℓ = ℓ} (Skeleton X n) 194 | 195 | Skeleton X zero = lift tt 196 | Skeleton X (suc n) = Skeleton X n , Fill (SkeletonExt X n) 197 | 198 | SkeletonExt X zero = X 199 | SkeletonExt X (suc n) = Hom (SkeletonExt X n) 200 | 201 | -- 202 | -- The terminal opetopic type 203 | -- 204 | 205 | 𝕋Ext : ∀ {n ℓ} (X : 𝕆Type n ℓ) → 𝕆Type∞ X 206 | Fill (𝕋Ext X) _ = Lift Unit 207 | Hom (𝕋Ext X) = 𝕋Ext (X , (λ _ → Lift Unit)) 208 | 209 | 𝕋 : (n : ℕ) {ℓ : Level} → 𝕆Type n ℓ 210 | 𝕋 zero = lift tt 211 | 𝕋 (suc n) = 𝕋 n , λ _ → Lift Unit 212 | 213 | 𝕋∞ : ∀ {ℓ} → 𝕆Type∞ tt* 214 | 𝕋∞ {ℓ} = 𝕋Ext {ℓ = ℓ} tt* 215 | -------------------------------------------------------------------------------- /Core/Prelude.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | 3 | open import Cubical.Core.Everything 4 | 5 | module Core.Prelude where 6 | 7 | infix 10 _↦_ 8 | postulate 9 | _↦_ : ∀ {ℓ} {A : Type ℓ} → A → A → Type ℓ 10 | 11 | {-# BUILTIN REWRITE _↦_ #-} 12 | 13 | -- Inductive identity types. 14 | data Ident {ℓ} (A : Type ℓ) (a : A) : A → Type ℓ where 15 | idp : Ident A a a 16 | -------------------------------------------------------------------------------- /Experimental/Faces.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Everything 2 | open import Cubical.Data.Nat 3 | open import Cubical.Data.Empty 4 | open import Cubical.Data.Sum 5 | 6 | open import Core.Everything 7 | open import Lib.Structures 8 | 9 | module Experimental.Faces where 10 | 11 | data InnerFace : {n : ℕ} → 𝒪 n → ℕ → Type where 12 | as-src : {n : ℕ} (o : 𝒪 n) (p : 𝒫 o) (q : 𝒫 (o ∣ p)) 13 | → Pos q → InnerFace (o ∣ p ∣ q) (suc n) 14 | as-tgt : {n : ℕ} (o : 𝒪 n) (p : 𝒫 o) (q : 𝒫 (o ∣ p)) 15 | → Pos q → InnerFace (o ∣ p ∣ q) n 16 | raise-face : {n m : ℕ} (o : 𝒪 n) (p : 𝒫 o) 17 | → InnerFace o m → InnerFace (o ∣ p) m 18 | 19 | data Face : {n : ℕ} → 𝒪 n → ℕ → Type where 20 | top : {n : ℕ} (o : 𝒪 n) → Face o n 21 | tgt-face : {n : ℕ} (o : 𝒪 (suc n)) → Face o n 22 | init : {n : ℕ} (o : 𝒪 (suc n)) → Face o 0 23 | inner : {n m : ℕ} (o : 𝒪 n) → InnerFace o m → Face o m 24 | 25 | -- obj : 𝒪 0 26 | -- obj = tt 27 | 28 | obj-face : Face obj 0 29 | obj-face = top ● 30 | 31 | -- arrow : 𝒪 1 32 | -- arrow = tt , tt 33 | 34 | -- arrow-src : Face arrow 0 35 | -- arrow-src = init arrow 36 | 37 | -- arrow-tgt : Face arrow 0 38 | -- arrow-tgt = tgt-face arrow 39 | 40 | -- arrow-top : Face arrow 1 41 | -- arrow-top = top arrow 42 | 43 | -- -- 2-simplex : 𝒪 2 44 | -- -- 2-simplex = (tt , tt) , ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt))) 45 | 46 | -- simplex = 2-simplex 47 | 48 | -- simp-init : Face simplex 0 49 | -- simp-init = init simplex 50 | 51 | -- simp-mid : Face simplex 0 52 | -- simp-mid = inner simplex (as-tgt tt tt 53 | -- (ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt)))) (inl tt)) 54 | 55 | -- simp-term : Face simplex 0 56 | -- simp-term = inner simplex (as-tgt tt tt 57 | -- (ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt)))) (inr (tt , inl tt))) 58 | 59 | -- simp-fst : Face simplex 1 60 | -- simp-fst = inner simplex (as-src tt tt 61 | -- (ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt)))) (inl tt)) 62 | 63 | -- simp-snd : Face simplex 1 64 | -- simp-snd = inner simplex (as-src tt tt 65 | -- (ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt)))) (inr (tt , inl tt))) 66 | 67 | -- simp-tgt : Face simplex 1 68 | -- simp-tgt = tgt-face simplex 69 | 70 | -- simp-top : Face simplex 2 71 | -- simp-top = top simplex 72 | 73 | -- 74 | -- Representables 75 | -- 76 | 77 | -- Rep : {n : ℕ} (o : 𝒪 n) → 𝕆 ℓ-zero (S n) 78 | 79 | -- SrcInc' : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → Rep (Typ ρ p) ⇒ fst (Rep (o , ρ)) 80 | -- SrcInc' = {!!} 81 | 82 | -- SrcInc : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → (Rep (Typ ρ p) , cst ∅) ⇒ Rep (o , ρ) 83 | -- TgtInc : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → (Rep o , cst ∅) ⇒ Rep (o , ρ) 84 | 85 | -- -- Hmmm. This is not quite right. It is the top *3* fibrations 86 | -- -- which must change. 87 | 88 | -- -- So, the cells of the first are parameterized by two things 89 | -- -- 1) the positions of ρ (there will be the externals) 90 | -- -- 2) the positions of τ (this will be the internals) 91 | 92 | -- -- In the next fibration, we have exactly the positions of τ and the target. 93 | 94 | -- -- And then we have the unique next dimension. 95 | 96 | -- -- Right. So this should actually lead to another, slight different 97 | -- -- presentation of the faces above. 98 | 99 | -- data ArrowFill : {o : 𝒪 1} → WebFrm tt (cst (⊤ ⊔ ⊤)) (fst o) (snd o) → Set where 100 | -- top : ArrowFill ⟪ tt , tt , inl tt , (λ _ → inr tt) ⟫ 101 | 102 | -- data BoundaryCell {n} (o : 𝒪 n) (ρ : 𝒫 o) : {q : 𝒪 n} → Frm (fst (Rep o)) q → Set where 103 | -- src-cell : (p : Pos ρ) (f : Frm (Rep (Typ ρ p)) {!!}) 104 | -- → BoundaryCell o ρ {!!} 105 | 106 | -- data FillingCell {n} (o : 𝒪 n) (ρ : 𝒫 o) : {q : 𝒪 (S n)} 107 | -- → WebFrm (fst (Rep o)) (BoundaryCell o ρ) (fst q) (snd q) → Set where 108 | 109 | 110 | -- Rep {n = O} tt = tt , cst ⊤ 111 | -- Rep {n = S O} (tt , tt) = 112 | -- (tt , cst (⊤ ⊔ ⊤)) , ArrowFill 113 | 114 | -- Rep {n = S (S n)} ((o , ρ) , τ) = ((fst (fst (Rep (o , ρ))) , {!!}) , {!!}) , {!!} 115 | 116 | -- -- Rep {n = O} o = tt , cst ⊤ 117 | -- -- Rep {n = S n} (o , ρ) = (fst (Rep o) , BoundaryCell o ρ) , FillingCell o ρ 118 | 119 | -- SrcInc = {!!} 120 | -- TgtInc = {!!} 121 | 122 | 123 | -- -- data LfBdry {n : ℕ} (m : 𝒪 n) : {o : 𝒪 (S n)} 124 | -- -- → WebFrm (fst (Rep m)) (snd (Rep m)) (fst o) (snd o) → Set where 125 | -- -- lf-tgt : LfBdry m ⟪ {!!} , {!η (fst (Rep m))!} , {!!} , {!!} ⟫ 126 | 127 | 128 | -- -- Rep {n = S (S n)} ((x , .(ηₒ x)) , lfₒ .x) = (Rep x , LfBdry x) , {!!} 129 | -- -- Rep {n = S (S n)} ((o , .(μₒ ρ δ)) , ndₒ .o ρ δ ε) = {!!} 130 | 131 | 132 | -------------------------------------------------------------------------------- /Experimental/FibrantGroupoids.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Everything 2 | open import Cubical.Data.Nat 3 | open import Cubical.Data.Empty 4 | open import Cubical.Data.Sum 5 | 6 | open import Core.Everything 7 | open import Lib.Structures 8 | open import Lib.Universe 9 | open import Lib.Groupoid 10 | 11 | -- 12 | -- Fibrancy of the opetopic type associated to a type 13 | -- 14 | 15 | module Experimental.FibrantGroupoids {ℓ} (X : Type ℓ) where 16 | 17 | idp : (x : X) → GrpCell X (tt* , reflₒ x ● , tt* , (λ p → reflₒ x ●)) 18 | idp x = reflₒ x (● ∣ objₒ) 19 | 20 | module _ {n 𝑜} where 21 | 22 | to : Σ[ f ∈ Frm (Grp X n) 𝑜 ] GrpCell X f → X 23 | to (.(Frm-El (Pt x) 𝑜) , reflₒ x .𝑜) = x 24 | 25 | from : X → Σ[ f ∈ Frm (Grp X n) 𝑜 ] GrpCell X f 26 | from x = (Frm-El (Pt x) 𝑜) , reflₒ x 𝑜 27 | 28 | to-from : (x : X) → to (from x) ≡ x 29 | to-from x = refl 30 | 31 | from-to : (d : Σ[ f ∈ Frm (Grp X n) 𝑜 ] GrpCell X f) → from (to d) ≡ d 32 | from-to (.(Frm-El (Pt x) 𝑜) , reflₒ x .𝑜) = refl 33 | 34 | -- A frame with a filler is completely determined by a point. 35 | full-frm-equiv : {n : ℕ} (𝑜 : 𝒪 n) → (Σ[ f ∈ Frm (Grp X n) 𝑜 ] GrpCell X f) ≃ X 36 | full-frm-equiv {n} 𝑜 = isoToEquiv (iso to from (to-from {n} {𝑜}) from-to) 37 | 38 | full-frm-fibers-contr : {n : ℕ} (𝑜 : 𝒪 n) (x : X) → isContr (fiber (to {𝑜 = 𝑜}) x) 39 | full-frm-fibers-contr 𝑜 x = equiv-proof (snd (full-frm-equiv 𝑜)) x 40 | 41 | -- Now, what I would like to show is that the top element is the 42 | -- same as the root-element. 43 | bottom-el : GrpCell X tt* → X 44 | bottom-el (reflₒ x .●) = x 45 | 46 | bottom-el-lem : (g : GrpCell X tt*) → reflₒ (bottom-el g) ● ≡ g 47 | bottom-el-lem (reflₒ x .●) = refl 48 | 49 | root-el : {n : ℕ} {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) 50 | → (Σ[ f ∈ Frm (Grp X n) 𝑜 ] 51 | Σ[ c ∈ Cns (Grp X n) f 𝑝 ] 52 | ((p : Pos 𝑝) → GrpCell X (Shp (Grp X n) c p))) → X 53 | root-el {𝑜 = ●} objₒ (f , c , y) = bottom-el (y tt) 54 | root-el {𝑜 = 𝑜 ∣ 𝑝} 𝑞 ((f , x , c , y) , cc , yy) = root-el 𝑝 (f , c , y) 55 | 56 | canon-dec : {n : ℕ} {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) 57 | → X → (Σ[ f ∈ Frm (Grp X n) 𝑜 ] 58 | Σ[ c ∈ Cns (Grp X n) f 𝑝 ] 59 | ((p : Pos 𝑝) → GrpCell X (Shp (Grp X n) c p))) 60 | canon-dec {𝑜 = 𝑜} 𝑝 x = Frm-El (Pt x) 𝑜 , Cns-El (Pt x) 𝑝 , λ p → reflₒ x (Typ 𝑝 p) 61 | 62 | one-dir : {n : ℕ} {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) (x : X) 63 | → root-el 𝑝 (canon-dec 𝑝 x) ≡ x 64 | one-dir {𝑜 = ●} objₒ x = refl 65 | one-dir {𝑜 = 𝑜 ∣ 𝑝} 𝑞 x = one-dir 𝑝 x 66 | 67 | claim : {n : ℕ} {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) 68 | → (f : Frm (Grp X n) 𝑜) (x : GrpCell X f) 69 | → (c : Cns (Grp X n) f 𝑝) 70 | → (y : (p : Pos 𝑝) → GrpCell X (Shp (Grp X n) c p)) 71 | → (α : GrpCell X {𝑜 = 𝑜 ∣ 𝑝} (f , x , c , y)) 72 | → root-el 𝑝 (f , c , y) ≡ to ((f , x , c , y) , α) 73 | claim 𝑝 ._ ._ ._ ._ (reflₒ x ._) = one-dir 𝑝 x 74 | 75 | 76 | 77 | -- So, now this means that the type of these things being 78 | -- equal to a fixed a are equivalent, since composition 79 | -- with this path with induce the equivalence. 80 | 81 | -- But then this will show that the space of all these things 82 | -- together with a fixed equality to an element a is contractible 83 | -- because of the above equivalence. And this should show that 84 | -- for any.. 85 | 86 | -- Okay. So to get better definitional behavior, I think you 87 | -- should rather define the root element locally so that it 88 | -- computes. 89 | 90 | -- root-el : {n : ℕ} {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) 91 | -- → (Σ[ f ∈ Frm (Grp X n) 𝑜 ] 92 | -- Σ[ c ∈ Cns (Grp X n) f 𝑝 ] 93 | -- ((p : Pos 𝑝) → GrpCell X (Shp (Grp X n) c p))) → X 94 | -- root-el objₒ (f , c , y) with y tt 95 | -- root-el objₒ (f , c , y) | reflₒ x .● = x 96 | -- root-el lfₒ (._ , (lf (reflₒ x 𝑜)) , _) = x 97 | -- root-el (ndₒ 𝑝 𝑞 𝑟) (._ , (nd (reflₒ x 𝑜) c y d z ψ) , _) = x 98 | 99 | -- 100 | -- proving there is a unique constructor with decoration 101 | -- 102 | 103 | -- one-dir : {n : ℕ} {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) (x : X) 104 | -- → root-el 𝑝 (canon-dec 𝑝 x) ≡ x 105 | -- one-dir {𝑜 = ●} objₒ x = refl 106 | -- one-dir {𝑜 = 𝑜 ∣ 𝑝} 𝑞 x = one-dir 𝑝 x 107 | 108 | -- one-dir objₒ x = refl 109 | -- one-dir lfₒ x = refl 110 | -- one-dir (ndₒ 𝑝 𝑞 𝑟) x = refl 111 | 112 | -- lemma₁ : (c : GrpCell X tt*) → reflₒ (root-el objₒ (tt* , tt* , const c)) ● ≡ c 113 | -- lemma₁ (reflₒ x .●) = refl 114 | 115 | -- lemma₂ : ∀ {ℓ} {P : ⊥ → Type ℓ} (f g : (p : ⊥) → P p) → f ≡ g 116 | -- lemma₂ {P = P} f g i () 117 | 118 | -- harder-dir : {n : ℕ} {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) 119 | -- → (pd : Σ[ f ∈ Frm (Grp X n) 𝑜 ] 120 | -- Σ[ c ∈ Cns (Grp X n) f 𝑝 ] 121 | -- ((p : Pos 𝑝) → GrpCell X (Shp (Grp X n) c p))) 122 | -- → canon-dec 𝑝 (root-el 𝑝 pd) ≡ pd 123 | -- harder-dir {𝑜 = ●} objₒ (tt* , tt* , y) = 124 | -- λ i → tt* , tt* , λ p → bottom-el-lem (y p) i 125 | -- harder-dir lfₒ (._ , lf (reflₒ x _) , yy) = {!!} 126 | -- harder-dir (ndₒ 𝑝 𝑞 𝑟) (._ , nd (reflₒ x _) c y d z ψ , yy) = {!!} 127 | 128 | -- harder-local : {n : ℕ} {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) (𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)) 129 | -- → (𝑟 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p ∣ 𝑞 p)) 130 | -- → (f : Frm (Grp X n) 𝑜) (x : GrpCell X f) 131 | -- → (c : Cns (Grp X n) f 𝑝) 132 | -- → (y : (p : Pos 𝑝) → GrpCell X (Shp (Grp X n) c p)) 133 | -- → (d : (p : Pos 𝑝) → Cns (Grp X n) (Shp (Grp X n) c p) (𝑞 p)) 134 | -- → (z : (p : Pos 𝑝) (q : Pos (𝑞 p)) → GrpCell X (Shp (Grp X n) (d p) q)) 135 | -- → (ψ : (p : Pos 𝑝) → Cns (Grp X n , GrpCell X) (Shp (Grp X n) c p , y p , d p , z p) (𝑟 p)) 136 | -- → (δ : (p : Unit ⊎ Σ-syntax (Pos 𝑝) (λ p₁ → Pos (𝑟 p₁))) → 137 | -- GrpCell X (Shp (Grp X n , GrpCell X) (nd x c y d z ψ) p)) 138 | -- → (cell : GrpCell X (f , x , c , y)) 139 | -- → (eq : δ (inl tt) ≡ cell) 140 | -- → canon-dec (ndₒ 𝑝 𝑞 𝑟) (root-el (ndₒ 𝑝 𝑞 𝑟) ((f , x , μ (Grp X n) c d , (λ p → z (fstₚ 𝑝 𝑞 p) (sndₚ 𝑝 𝑞 p))) , nd x c y d z ψ , δ)) 141 | -- ≡ ((f , x , μ (Grp X n) c d , (λ p → z (fstₚ 𝑝 𝑞 p) (sndₚ 𝑝 𝑞 p))) , nd x c y d z ψ , δ) 142 | 143 | -- -- implementation of harder-dir 144 | -- harder-dir objₒ (tt* , tt* , δ) = λ i → tt* , tt* , λ p → lemma₁ (δ p) i 145 | 146 | -- harder-dir {suc n} lfₒ (._ , lf (reflₒ x 𝑜) , δ) = λ i → 147 | -- ((Frm-El (Pt x) 𝑜 , 148 | -- reflₒ x 𝑜 , η (Grp X n) (Frm-El (Pt x) 𝑜) , (λ p → reflₒ x 𝑜)) 149 | -- , lf (reflₒ x 𝑜) , lemma₂ (λ p → reflₒ x (Typ lfₒ p)) δ i) 150 | 151 | -- harder-dir (ndₒ 𝑝 𝑞 𝑟) (._ , nd x c y d z ψ , δ) = 152 | -- harder-local 𝑝 𝑞 𝑟 _ x c y d z ψ δ (δ (inl tt)) refl 153 | 154 | -- -- implementation of harder-local 155 | -- harder-local {n} {𝑜} 𝑝 𝑞 𝑟 ._ ._ ._ ._ d z ψ δ (reflₒ x ._) eq = 156 | -- λ i → (Frm-El (Pt x) 𝑜 , reflₒ x 𝑜 , μ (Grp X n) (Cns-El (Pt x) 𝑝) {!!} , {!!}) , 157 | -- nd (reflₒ x 𝑜) (Cns-El (Pt x) 𝑝) (λ p → reflₒ x (Typ 𝑝 p)) {!!} {!!} {!!} , {!!} 158 | 159 | -- -- Need equalities: 160 | -- -- 161 | -- -- d ≡ (λ p → Cns-El (Pt x) (𝑞 p)) 162 | -- -- z ≡ (λ p q → reflₒ x (Typ (𝑞 p) q)) 163 | -- -- ψ = (λ p → Cns-El (Pt x , reflₒ x) (𝑟 p)) 164 | -- -- δ ≡ (λ p → reflₒ x (Typ (ndₒ 𝑝 𝑞 𝑟) p)) 165 | -- -- 166 | 167 | -- where IH-el : (p : Pos 𝑝) → (Σ[ f ∈ Frm (Grp X (suc n)) (Typ 𝑝 p ∣ 𝑞 p) ] 168 | -- Σ[ c ∈ Cns (Grp X (suc n)) f (𝑟 p) ] 169 | -- ((q : Pos (𝑟 p)) → GrpCell X (Shp (Grp X (suc n)) c q))) 170 | -- IH-el p = ((Shp (Grp X _) (Cns-El (Pt x) 𝑝) p , reflₒ x (Typ 𝑝 p) , d p , z p) , ψ p , λ q → δ (inr (p , q))) 171 | 172 | -- IH : (p : Pos 𝑝) → canon-dec (𝑟 p) (root-el (𝑟 p) (IH-el p)) ≡ IH-el p 173 | -- IH p = harder-dir (𝑟 p) ((Frm-El (Pt x) (Typ 𝑝 p) , reflₒ x (Typ 𝑝 p) , d p , z p) , ψ p , λ q → δ (inr (p , q))) 174 | 175 | -- IH-fst : (p : Pos 𝑝) → Shp (Grp X n) (Cns-El (Pt (root-el (𝑟 p) (IH-el p))) 𝑝) p 176 | -- ≡ Shp (Grp X n) (Cns-El (Pt x) 𝑝) p 177 | -- IH-fst p i = fst (fst (IH p i)) 178 | 179 | 180 | -- other-idea : {n : ℕ} (𝑜 : 𝒪 n) (𝑝 : 𝒫 𝑜) 181 | -- → X ≃ (Σ[ f ∈ Frm (Grp X n) 𝑜 ] 182 | -- Σ[ c ∈ Cns (Grp X n) f 𝑝 ] 183 | -- ((p : Pos 𝑝) → GrpCell X (Shp (Grp X n) c p))) 184 | -- other-idea {n} 𝑜 𝑝 = isoToEquiv (iso (canon-dec 𝑝) (root-el 𝑝) 185 | -- (harder-dir 𝑝) (one-dir 𝑝)) 186 | 187 | -- thm : (n : ℕ) → is-fibrant (Grp X (suc (suc n))) 188 | -- thm n {𝑜} {𝑝} {f} c y = {!!} 189 | -------------------------------------------------------------------------------- /Experimental/FibrantUniverse.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Everything 2 | open import Cubical.Data.Nat 3 | 4 | open import Core.Everything 5 | open import Lib.Structures 6 | open import Lib.Universe 7 | 8 | module Experimental.FibrantUniverse where 9 | 10 | -- Right, so this I think is now correct. Just needs to be cleaned up 11 | -- and written in a way which is comprehensible.... 12 | CompRel : ∀ {n ℓ} {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {𝑡 : 𝒫 (𝑜 ∣ 𝑝)} 13 | → (f : Frm (Σₒ (𝒰ₒ n ℓ) (ℱₒ n)) 𝑜) 14 | → (X : Σ-cell (λ f₁ → Frm↓ (𝒱ₒ n) f₁ → Type ℓ) (λ _ → is-fibrant-rel) f) 15 | → (c : Cns (Σₒ (𝒰ₒ n ℓ) (ℱₒ n)) f 𝑝) 16 | → (Y : (p : Pos 𝑝) → 17 | Σ-cell (λ f₁ → Frm↓ (𝒱ₒ n) f₁ → Type ℓ) (λ _ → is-fibrant-rel) 18 | (Shp (Σₒ (𝒰ₒ n ℓ) (ℱₒ n)) c p)) 19 | → (ω : Cns 20 | (Σₒ (𝒰ₒ n ℓ) (ℱₒ n) , 21 | Σ-cell (λ f₁ → Frm↓ (𝒱ₒ n) f₁ → Type ℓ) (λ _ → is-fibrant-rel)) 22 | (f , X , c , Y) 𝑡) 23 | → (R : (p : Pos 𝑡) → 24 | Σ-cell (λ f₁ → Frm↓ (𝒱ₒ n , (λ f↓ X₁ → X₁ f↓)) f₁ → Type ℓ) 25 | (λ _ → is-fibrant-rel) 26 | (Shp 27 | (Σₒ (𝒰ₒ n ℓ) (ℱₒ n) , 28 | Σ-cell (λ f₁ → Frm↓ (𝒱ₒ n) f₁ → Type ℓ) (λ _ → is-fibrant-rel)) 29 | ω p)) 30 | → Frm↓ (𝒱ₒ (suc n)) {𝑜 = 𝑜 ∣ 𝑝} (fst-frm (f , X , c , Y)) 31 | → Type ℓ 32 | CompRel {n} {ℓ} {𝑜 = 𝑜} {𝑝} {𝑡} f X c Y ω R f↓ = 33 | Σ[ ω↓ ∈ Cns↓ (𝒱ₒ (suc n)) f↓ (fst-cns {P = ℱₒ (suc n)} ω) ] 34 | ((p : Pos 𝑡) → fstₒ (R p) {!(Shp↓ (𝒱ₒ (suc n)) ω↓ p)!}) 35 | 36 | 37 | thm : (n : ℕ) (ℓ : Level) → is-fibrant (𝒮ₙ (suc (suc (suc n))) ℓ) 38 | thm n ℓ {𝑜 ∣ 𝑝} {𝑡} {f , X , c , Y} ω R = 39 | ((CompRel f X c Y ω R , {!!}) , {!!}) , {!!} 40 | 41 | 42 | -------------------------------------------------------------------------------- /Experimental/LessPositions/OpetopicMap.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- PositionlessMap.agda - Maps between opetopic types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Nat 8 | 9 | open import Core.Prelude 10 | open import Experimental.LessPositions.OpetopicType 11 | 12 | module Experimental.LessPositions.OpetopicMap where 13 | 14 | _⇒_ : ∀ {n ℓ} → 𝕆Type n ℓ → 𝕆Type n ℓ → Type ℓ 15 | 16 | Frm⇒ : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} → X ⇒ Y → Frm X → Frm Y 17 | 18 | Src⇒ : ∀ {n ℓ} (Xₙ Yₙ : 𝕆Type n ℓ) 19 | → {Xₛₙ : Frm Xₙ → Type ℓ} 20 | → {Yₛₙ : Frm Yₙ → Type ℓ} 21 | → (σₙ : Xₙ ⇒ Yₙ) 22 | → (σₛₙ : (f : Frm Xₙ) → Xₛₙ f → Yₛₙ (Frm⇒ σₙ f )) 23 | → {f : Frm Xₙ} 24 | → Src Xₛₙ f → Src Yₛₙ (Frm⇒ σₙ f) 25 | 26 | Pos⇒ : ∀ {n ℓ} (Xₙ Yₙ : 𝕆Type n ℓ) 27 | → {Xₛₙ : Frm Xₙ → Type ℓ} 28 | → {Yₛₙ : Frm Yₙ → Type ℓ} 29 | → (σₙ : Xₙ ⇒ Yₙ) 30 | → (σₛₙ : (f : Frm Xₙ) → Xₛₙ f → Yₛₙ (Frm⇒ σₙ f )) 31 | → {f : Frm Xₙ} (s : Src Xₛₙ f) 32 | → Pos Xₛₙ s → Pos Yₛₙ (Src⇒ Xₙ Yₙ σₙ σₛₙ s) 33 | 34 | _⇒_ {zero} X Y = Lift Unit 35 | _⇒_ {suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) = Σ[ fun ∈ Xₙ ⇒ Yₙ ] ((f : Frm Xₙ) → Xₛₙ f → Yₛₙ (Frm⇒ fun f)) 36 | 37 | 38 | Frm⇒ {zero} {ℓ} {X} {Y} fun f = tt* 39 | Frm⇒ {suc n} {ℓ} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} (funₙ , funₛₙ) (f , t , s) = 40 | (Frm⇒ funₙ f) , (funₛₙ _ t) , Src⇒ Xₙ Yₙ funₙ funₛₙ s 41 | 42 | Branch⇒ : ∀ {n ℓ} (Xₙ Yₙ : 𝕆Type (suc n) ℓ) 43 | → {Xₛₙ : Frm Xₙ → Type ℓ} 44 | → {Yₛₙ : Frm Yₙ → Type ℓ} 45 | → (σₙ : Xₙ ⇒ Yₙ) 46 | → (σₛₙ : (f : Frm Xₙ) → Xₛₙ f → Yₛₙ (Frm⇒ σₙ f )) 47 | → {f : Frm (fst Xₙ)} 48 | → Branch Xₛₙ f → Branch Yₛₙ (Frm⇒ (σₙ .fst) f) 49 | Branch⇒ X Y {P} {Q} (σ , σ') σ'' [ stm , lvs , br ] = 50 | [ σ' _ stm , Src⇒ (fst X) (fst Y) σ σ' lvs , Src⇒ X Y (σ , σ') σ'' br ] 51 | 52 | 53 | postulate 54 | η⇒ : ∀ {n ℓ} (Xₙ Yₙ : 𝕆Type n ℓ) 55 | → {Xₛₙ : Frm Xₙ → Type ℓ} 56 | → {Yₛₙ : Frm Yₙ → Type ℓ} 57 | → (σₙ : Xₙ ⇒ Yₙ) 58 | → (σₛₙ : (f : Frm Xₙ) → Xₛₙ f → Yₛₙ (Frm⇒ σₙ f)) 59 | → {f : Frm Xₙ} (x : Xₛₙ f) 60 | → Src⇒ Xₙ Yₙ σₙ σₛₙ (η Xₛₙ x) ↦ η Yₛₙ (σₛₙ f x) 61 | {-# REWRITE η⇒ #-} 62 | 63 | μ⇒' : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (σ : X ⇒ Y) 64 | (P P' : Frm X → Type ℓ) 65 | (Q : Frm Y → Type ℓ) 66 | (σ' : (f : Frm X) → P' f → Q (Frm⇒ σ f)) 67 | {f : Frm X} 68 | (s : Src P f) 69 | (t : (p : Pos P s) → Src P' (Typ s p)) 70 | → Src⇒ X Y {P'} {Q} σ σ' (μ P' s t) ↦ μ Q {!Src⇒ X Y σ σ' !} {!!} 71 | -- {-# REWRITE μ⇒ #-} 72 | 73 | -- (Src⇒ X Y σₙ σₛₙ (μ P brs (λ p → η P (stm (brs ⊚ p))))) != 74 | -- (μ Q brs' (λ p → η Q (stm (brs' ⊚ p)))) of type (Src Q (Frm⇒ σₙ f)) 75 | 76 | μ⇒ : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (σ : X ⇒ Y) 77 | → (P : Frm X → Type ℓ) 78 | → (Q : Frm Y → Type ℓ) 79 | → (σ' : (f : Frm X) → P f → Q (Frm⇒ σ f)) 80 | → {f : Frm X} (s : Src P f) 81 | → (t : (p : Pos P s) → Src P (Typ s p)) 82 | → Src⇒ X Y {P} {Q} σ σ' (μ P s t) ↦ Src⇒ X Y {P} {Q} σ σ' s 83 | {-# REWRITE μ⇒ #-} 84 | 85 | μ'⇒ : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (σ : X ⇒ Y) 86 | (P : Frm X → Type ℓ) 87 | (Q : Frm Y → Type ℓ) 88 | (σ' : (f : Frm X) → P f → Q (Frm⇒ σ f)) 89 | {f : Frm X} 90 | (s : Src (Src P) f) 91 | → Src⇒ X Y {Yₛₙ = Q} σ σ' (μ' {Xₛₙ = P} s) ↦ μ' {Xₛₙ = Q} (Src⇒ X Y σ (λ f → Src⇒ X Y σ σ') s) 92 | 93 | Src⇒ {zero} {ℓ} X Y {P} {Q} σₙ σₛₙ {f} = σₛₙ _ 94 | Src⇒ {suc n} {ℓ} X Y {P} {Q} σₙ σₛₙ (lf tgt) = lf (snd σₙ _ tgt) 95 | Src⇒ {suc n} {ℓ} (X , P) (Y , Q) {P'} {Q'} (σₙ , σₛₙ) σₛₛₙ (nd {f = f} tgt brs flr) = {!!} 96 | 97 | where tgt' : Q (Frm⇒ σₙ f) 98 | tgt' = σₛₙ f tgt 99 | 100 | brs' : Src (Branch Q') (Frm⇒ σₙ f) 101 | brs' = Src⇒ X Y {Branch P'} {Branch Q'} σₙ (λ f → Branch⇒ (X , P) (Y , Q) {P'} {Q'} (σₙ , σₛₙ) σₛₛₙ {f = f}) brs 102 | 103 | flr' : Q' (Frm⇒ (σₙ , σₛₙ) (f , tgt , μ P brs (λ p → η P (stm (brs ⊚ p))))) 104 | flr' = σₛₛₙ (f , tgt , μ P brs (λ p → η P (stm (brs ⊚ p)))) flr 105 | 106 | -- So this is the missing equation. Notice that it's just a map, not a full-fledged 107 | -- multiplication. 108 | 109 | -- (Src⇒ X Y σₙ σₛₙ (μ P brs (λ p → η P (stm (brs ⊚ p))))) != 110 | -- (μ Q brs' (λ p → η Q (stm (brs' ⊚ p)))) of type (Src Q (Frm⇒ σₙ f)) 111 | 112 | Pos⇒ = {!!} 113 | -------------------------------------------------------------------------------- /Experimental/LessPositions/Shapes.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --no-positivity-check #-} 2 | -- 3 | -- OpetopicType.agda - Opetopic Types 4 | -- 5 | 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Sigma 9 | 10 | open import Core.Prelude 11 | open import Experimental.LessPositions.OpetopicType 12 | 13 | module Experimental.LessPositions.Shapes where 14 | globe-Frm : ∀ {n ℓ} {Xₙ : 𝕆Type n ℓ} (Xₛₙ : Frm Xₙ → Type ℓ) {f : Frm Xₙ} (x y : Xₛₙ f) → Frm (Xₙ , Xₛₙ) 15 | globe-Frm Xₛₙ {f} x y = f , y , η Xₛₙ x 16 | 17 | _**_ : ∀ {ℓ} → Type ℓ → ℕ → Type ℓ 18 | X ** zero = Lift Unit 19 | X ** suc zero = X 20 | X ** suc (suc n) = X × (X ** (suc n)) 21 | 22 | fstt : ∀ {ℓ n} {X : Type ℓ} → X ** (suc n) → X 23 | fstt {n = zero} x = x 24 | fstt {n = suc n} (x , _) = x 25 | 26 | last : ∀ {ℓ n} {X : Type ℓ} → X ** (suc n) → X 27 | last {n = zero} x = x 28 | last {n = suc n} (fst , snd) = last snd 29 | 30 | path-chain : ∀ {ℓ} {X : Type ℓ} (Y : X → X → Type ℓ) (n : ℕ) (t : X ** (suc n)) → Type ℓ 31 | path-chain Y zero _ = Lift Unit 32 | path-chain Y (suc zero) (x , y) = Y y x 33 | path-chain Y (suc (suc n)) (x , t) = Y (fstt t) x × path-chain Y (suc n) t 34 | 35 | -- Sequences of arrows 36 | n-path : ∀ {ℓ} {X : 𝕆Type 2 ℓ} (n : ℕ) {t : (X .fst .snd tt*) ** suc n} 37 | → path-chain (λ x y → X .snd (tt* , y , x)) n t 38 | → Src (X .snd) (tt* , fstt t , last t) 39 | n-path zero {x} _ = lf x 40 | n-path (suc zero) {y , x} f = nd y [ x , x , lf x ] f 41 | n-path (suc (suc n)) {y , x , t} (f , l) = nd y [ x , last t , (n-path (suc n) l) ] f 42 | 43 | -- Sequences of unary higher cells 44 | n-path' : ∀ {ℓ} {m : ℕ} {X : 𝕆Type m ℓ} (n : ℕ) {P : Frm X → Type ℓ} (Q : Frm (X , P) → Type ℓ) {f : Frm X} 45 | {t : P f ** (suc n)} 46 | → path-chain (λ x y → Q (globe-Frm P x y)) n t 47 | → Src Q (globe-Frm P (last t) (fstt t)) 48 | n-path' zero {P} Q {f} {x} _ = lf x 49 | n-path' (suc zero) {P} Q {f} {y , x} p = nd y (η (Branch Q) [ x , _ , (lf x) ]) p 50 | n-path' (suc (suc n)) {P} Q {f} {y , x , t} (p , l) = nd y (η (Branch Q) [ x , (η P (last t)) , (n-path' (suc n) Q l) ]) p 51 | 52 | module _ (X : 𝕆Type 3 ℓ-zero) where 53 | X₀ = X .fst .fst .snd tt* 54 | X₁ = X .fst .snd 55 | X₂ = X .snd 56 | 57 | hom-Frm : X₀ → X₀ → Frm (X .fst .fst) 58 | hom-Frm x y = tt* , y , x 59 | 60 | hom : X₀ → X₀ → Type 61 | hom x y = X₁ (hom-Frm x y) 62 | 63 | 64 | simplex-Frm : {x y z : X₀} (f : hom x y) (g : hom y z) (h : hom x z) → Frm (X .fst) 65 | simplex-Frm {x} {y} {z} f g h = hom-Frm x z , h , n-path 2 (g , f) -- nd z [ y , x , nd y [ x , x , lf x ] f ] g 66 | 67 | 2-drop-Frm : (x : X₀) (f : hom x x) → Frm (X .fst) 68 | 2-drop-Frm x f = hom-Frm x x , f , lf x 69 | 70 | unitor-Frm : (x y : X₀) (f : hom x x) (g : hom x y) (h : hom x y) 71 | → (Δ : X₂ (simplex-Frm f g h)) 72 | → (Γ : X₂ (2-drop-Frm x f)) 73 | → (O : X₂ (globe-Frm X₁ g h)) 74 | → Frm X 75 | unitor-Frm x y f g h Δ Γ O = _ , O , nd h (nd y [ x , x , nd x [ x , η _ x , lf x ] [ f , lf x , nd f (lf x) Γ ] ] [ g , _ , lf g ]) Δ 76 | 77 | associator1 : (x y z t : X₀) (f : hom x y) (g : hom y z) (h : hom z t) (i : hom x z) (j : hom x t) 78 | → (Δ₁ : X₂ (simplex-Frm f g i)) 79 | → (Δ₂ : X₂ (simplex-Frm i h j)) 80 | → Src X₂ (hom-Frm x t , j , n-path 3 (h , g , f)) --nd t [ _ , _ , (nd _ [ _ , _ , (nd _ [ _ , _ , (lf _) ] f) ] g) ] h) 81 | associator1 x y z t f g h i j Δ₁ Δ₂ = nd j (nd t [ z , x , (nd z [ x , x , (lf x) ] [ i , nd z [ y , x , (nd _ _ _) ] g , nd i (nd z [ y , x , nd y [ x , x , (lf x) ] [ f , _ , lf f ] ] [ g , _ , (lf g) ]) Δ₁ ]) ] [ h , _ , (lf h) ]) Δ₂ 82 | 83 | -------------------------------------------------------------------------------- /Experimental/LessPositions/Structures.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Everything 2 | 3 | open import Core.Prelude 4 | open import Experimental.LessPositions.OpetopicType 5 | open import Experimental.LessPositions.Shapes 6 | 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Sigma 9 | open import Cubical.Foundations.Equiv.Fiberwise 10 | 11 | module Experimental.LessPositions.Structures where 12 | record 𝕆Type∞ {n ℓ} (Xₙ : 𝕆Type n ℓ) : Type (ℓ-suc ℓ) where 13 | coinductive 14 | field 15 | Fill : Frm Xₙ → Type ℓ 16 | Hom : 𝕆Type∞ (Xₙ , Fill) 17 | 18 | open 𝕆Type∞ public 19 | 20 | horn-filler : ∀ {n ℓ} {Xₙ : 𝕆Type n ℓ} {Xₛₙ : Frm Xₙ → Type ℓ} (Xₛₛₙ : Frm (Xₙ , Xₛₙ) → Type ℓ) {f : Frm Xₙ} → Src Xₛₙ f → Type ℓ 21 | horn-filler {n} {ℓ} {Xₙ} {Xₛₙ} Xₛₛₙ {f} s = Σ[ tgt ∈ Xₛₙ f ] Xₛₛₙ (f , tgt , s) 22 | 23 | is-fibrant : ∀ {n ℓ} → 𝕆Type (suc (suc n)) ℓ → Type ℓ 24 | is-fibrant ((Xₙ , Xₛₙ) , Xₛₛₙ) = (f : Frm Xₙ) (s : Src Xₛₙ f) → isContr (horn-filler Xₛₛₙ s) 25 | 26 | isProp-is-fibrant : ∀ {n ℓ} {X : 𝕆Type (suc (suc n)) ℓ} → isProp (is-fibrant X) 27 | isProp-is-fibrant = isPropΠ2 (λ _ _ → isPropIsContr) 28 | 29 | record is-fibrant-ext {n ℓ} {Xₙ : 𝕆Type n ℓ} (X : 𝕆Type∞ Xₙ) : Type ℓ where 30 | coinductive 31 | field 32 | fill-fib : is-fibrant ((Xₙ , (Fill X)) , (Fill (Hom X))) 33 | hom-fib : is-fibrant-ext (Hom X) 34 | 35 | open is-fibrant-ext 36 | 37 | eta-fib-ext : ∀ {m ℓ} {X : 𝕆Type m ℓ} {X∞ : 𝕆Type∞ X} → X∞ ≡ record { Fill = Fill X∞ ; Hom = Hom X∞ } 38 | Fill (eta-fib-ext {X∞ = X∞} i) = Fill X∞ 39 | Hom (eta-fib-ext {X∞ = X∞} i) = Hom X∞ 40 | 41 | isProp-is-fibrant-ext : ∀ {n ℓ} {X : 𝕆Type n ℓ} {X∞ : 𝕆Type∞ X} → isProp (is-fibrant-ext X∞) 42 | fill-fib (isProp-is-fibrant-ext x y i) = isProp-is-fibrant (fill-fib x) (fill-fib y) i 43 | hom-fib (isProp-is-fibrant-ext x y i) = isProp-is-fibrant-ext (hom-fib x) (hom-fib y) i 44 | 45 | 𝕋 : ∀ {ℓ} (n : ℕ) → 𝕆Type n ℓ 46 | 𝕋 zero = tt* 47 | 𝕋 (suc n) = 𝕋 n , λ _ → Lift Unit 48 | 49 | 𝕋Ext : ∀ {n ℓ} (X : 𝕆Type n ℓ) → 𝕆Type∞ X 50 | Fill (𝕋Ext X) = λ _ → Lift Unit 51 | Hom (𝕋Ext X) = 𝕋Ext _ 52 | 53 | is-fib-ext-𝕋Ext : ∀ {n ℓ} {X : 𝕆Type n ℓ} → is-fibrant-ext (𝕋Ext X) 54 | fill-fib is-fib-ext-𝕋Ext f s = (tt* , tt*) , λ (tt* , tt*) → refl 55 | hom-fib is-fib-ext-𝕋Ext = is-fib-ext-𝕋Ext 56 | 57 | 58 | 59 | 60 | 61 | hom→path0 : ∀ {ℓ} {X : 𝕆Type∞ (𝕋 {ℓ} 0)} → is-fibrant-ext X → (x y : X .Fill tt*) → X .Hom .Fill (tt* , y , x) → x ≡ y 62 | hom→path0 {ℓ} {X} fib x y σ = sym (transportRefl x) ∙ cong fst (isContr→isProp (fib .fill-fib tt* x) (Id x x refl) b) where 63 | Id : (x y : X .Fill tt*) → (x ≡ y) → horn-filler (Fill (Hom X)) x 64 | Id x y = J (λ y p → horn-filler (Fill (Hom X)) x) (x , fib .hom-fib .fill-fib (tt* , x , x) (lf x) .fst .fst) 65 | 66 | b : horn-filler (Fill (Hom X)) x 67 | b = y , σ 68 | 69 | Id-cell : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} (t : Fill X∞ f) → Fill (Hom X∞) (globe-Frm _ t t) 70 | Id-cell X∞ fib t = fib .hom-fib .fill-fib _ (lf t) .fst .fst 71 | 72 | path→hom : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} 73 | → t ≡ u → Fill (Hom X∞) (globe-Frm _ t u) 74 | path→hom X∞ fib {t = t} {u = u} = J (λ u p → Fill (Hom X∞) (globe-Frm _ t u)) (Id-cell X∞ fib t) 75 | 76 | hom→path : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} 77 | → Fill (Hom X∞) (globe-Frm _ t u) → t ≡ u 78 | hom→path X∞ fib {f} {t} {u} cell = cong fst (isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (u , cell)) 79 | 80 | module IsoHomPath {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} where 81 | sec : section (path→hom X∞ fib {f} {t} {u}) (hom→path X∞ fib {f} {t} {u}) 82 | sec cell = fromPathP (cong snd (isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (u , cell))) 83 | -- Wow ! Might need to take a look back at this later cause I didn't expect it to be so "simple". 84 | 85 | ret : retract (path→hom X∞ fib {f} {t} {u}) (hom→path X∞ fib {f} {t} {u}) 86 | ret = J (λ u p → hom→path X∞ fib (path→hom X∞ fib p) ≡ p) 87 | (hom→path X∞ fib (path→hom X∞ fib refl) 88 | ≡⟨ cong (hom→path X∞ fib) (transportRefl _) ⟩ 89 | hom→path X∞ fib (Id-cell X∞ fib t) 90 | ≡⟨ (λ i j → fst (test i j)) ⟩ 91 | refl ∎) where 92 | a : horn-filler (Fill (Hom X∞)) (η (Fill X∞) t) 93 | a = (t , Id-cell X∞ fib t) 94 | 95 | test0 : a ≡ a 96 | test0 = isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (t , Id-cell X∞ fib t) 97 | 98 | test : test0 ≡ refl 99 | test = isProp→isSet (isContr→isProp (fib .fill-fib f (η (Fill X∞) t))) (t , Id-cell X∞ fib t) (t , Id-cell X∞ fib t) test0 refl 100 | 101 | record is-n-trunc {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) : Type ℓ where 102 | coinductive 103 | field 104 | hLevel : (f : Frm X) → isOfHLevel n (X∞ .Fill f) 105 | is-trunc-ext : is-n-trunc (predℕ n) (X∞ .Hom) 106 | open is-n-trunc 107 | 108 | src-comp : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) → is-fibrant-ext X∞ → {f : Frm X} → Src (Fill X∞) f → Fill X∞ f 109 | src-comp X∞ fib s = fib .fill-fib _ s .fst .fst 110 | 111 | -- More general version of the equivalence between hom and path, using the fundamental theorem of identity types 112 | cell≃path : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} (s : Src (Fill X∞) f) (t : Fill X∞ f) 113 | → (src-comp X∞ fib s ≡ t) ≃ Fill (Hom X∞) (f , t , s) 114 | cell≃path X∞ fib s t = recognizeId (λ t → (Fill (Hom X∞)) (_ , t , s)) (fib .fill-fib _ s .fst .snd) (fib .fill-fib _ s) t 115 | 116 | isOfHLevelPathPred : ∀ (n : ℕ) {ℓ} {A : Type ℓ} → isOfHLevel n A → {x y : A} → isOfHLevel (predℕ n) (x ≡ y) 117 | isOfHLevelPathPred zero h = isContr→isContrPath h _ _ 118 | isOfHLevelPathPred (suc n) h = isOfHLevelPath' n h _ _ 119 | 120 | is-n-trunc-fib : ∀ {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) → is-fibrant-ext X∞ → ((f : Frm X) → isOfHLevel n (X∞ .Fill f)) → is-n-trunc n X∞ 121 | hLevel (is-n-trunc-fib n {X} X∞ fib h) = h 122 | is-trunc-ext (is-n-trunc-fib n {X} X∞ fib h) = is-n-trunc-fib _ _ (fib .hom-fib) lemma where 123 | lemma : (f : Frm (X , Fill X∞)) → isOfHLevel (predℕ n) (X∞ .Hom .Fill f) 124 | lemma (f , t , s) = isOfHLevelRespectEquiv (predℕ n) (cell≃path X∞ fib s t) (isOfHLevelPathPred n (h f)) 125 | 126 | 𝕆∞Path : ∀ {m ℓ} {X : 𝕆Type m ℓ} (X∞ X∞' : 𝕆Type∞ X) (p : Fill X∞ ≡ Fill X∞') → PathP (λ i → 𝕆Type∞ (X , p i)) (Hom X∞) (Hom X∞') → X∞ ≡ X∞' 127 | Fill (𝕆∞Path X∞ X∞' p q i) = p i 128 | Hom (𝕆∞Path X∞ X∞' p q i) = q i 129 | 130 | -- Trying to prove that fibrant opetopic types with same base-type are equal 131 | lemma0 : ∀ {m ℓ} {X X' : 𝕆Type m ℓ} (p : X ≡ X') 132 | → (P : Frm X → Type ℓ) (P' : Frm X' → Type ℓ) 133 | → (q : PathP (λ i → Frm (p i) → Type ℓ) P P') 134 | → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X' , P')) 135 | → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 136 | → PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill P∞) (Fill P∞') 137 | lemma0 {ℓ = ℓ} {X = X} = 138 | J 139 | (λ X' p → (P : Frm X → Type ℓ) (P' : Frm X' → Type ℓ) 140 | → (q : PathP (λ i → Frm (p i) → Type ℓ) P P') 141 | → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X' , P')) 142 | → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 143 | → PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill P∞) (Fill P∞')) 144 | 145 | λ P _ → J 146 | 147 | (λ P' q → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X , P')) 148 | → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 149 | → PathP (λ i → Frm (refl i , q i) → Type ℓ) (Fill P∞) (Fill P∞')) 150 | 151 | λ P∞ _ fib fib' → funExt λ x → ua (invEquiv (cell≃path _ fib _ _)) ∙ cong (_≡ (fst (snd x))) {!!} ∙ ua (cell≃path _ fib' _ _) 152 | 153 | fibrant→≡ : ∀ {m ℓ} {X X' : 𝕆Type m ℓ} (p : X ≡ X') 154 | → (X∞ : 𝕆Type∞ X) (X∞' : 𝕆Type∞ X') → PathP (λ i → Frm (p i) → Type ℓ) (Fill X∞) (Fill X∞') 155 | → is-fibrant-ext X∞ → is-fibrant-ext X∞' 156 | → PathP (λ i → 𝕆Type∞ (p i)) X∞ X∞' 157 | Fill (fibrant→≡ p X∞ X∞' q fib fib' i) = q i 158 | Hom (fibrant→≡ {ℓ = ℓ} p X∞ X∞' q fib fib' i) = fibrant→≡ (λ i → (p i , q i)) (Hom X∞) (Hom X∞') lemma (hom-fib fib) (hom-fib fib') i where 159 | lemma : PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill (Hom X∞)) (Fill (Hom X∞')) 160 | lemma = lemma0 p _ _ _ (Hom X∞) (Hom X∞') (subst is-fibrant-ext eta-fib-ext fib) (subst is-fibrant-ext eta-fib-ext fib') 161 | 162 | {- 163 | is-n-cat : ∀ {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) → Type ℓ 164 | is-n-cat zero X∞ = is-fibrant-ext X∞ 165 | is-n-cat (suc n) X∞ = Σ (is-n-cat n (Hom X∞)) {!!} 166 | -} 167 | -------------------------------------------------------------------------------- /Experimental/Local/CartesianProduct.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- CartsianProduct.agda - Cartesian product of Opetopic Types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Nat 10 | open import Cubical.Data.Sum 11 | 12 | open import Core.Prelude 13 | open import Experimental.Local.OpetopicType 14 | open import Experimental.Local.OpetopicMap 15 | open import Experimental.Local.Shapes 16 | open import Experimental.Local.Structures 17 | 18 | module Experimental.Local.CartesianProduct where 19 | 20 | prod : ∀ {n ℓ} (X Y : 𝕆Type n ℓ) → 𝕆Type n ℓ 21 | 22 | -- The use of opetopic maps circumvents any need for additional rewrite rules 23 | Fst : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} → prod X Y ⇒ X 24 | Snd : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} → prod X Y ⇒ Y 25 | 26 | record prod-cell {n ℓ} {X Y : 𝕆Type n ℓ} (P : Frm X → Type ℓ) (Q : Frm Y → Type ℓ) (f : Frm (prod X Y)) : Type ℓ where 27 | constructor _∣_ 28 | field 29 | fstₒ : P (Frm⇒ Fst f) 30 | sndₒ : Q (Frm⇒ Snd f) 31 | open prod-cell 32 | 33 | prod {zero} X Y = tt* 34 | prod {suc n} (X , P) (Y , Q) = prod X Y , prod-cell P Q 35 | 36 | Fst {zero} = tt* 37 | Fst {suc n} {ℓ} {X , P} {Y , Q} = Fst , fstₒ 38 | 39 | Snd {zero} = tt* 40 | Snd {suc n} {ℓ} {X , P} {Y , Q} = Snd , sndₒ 41 | 42 | 43 | ---------- Tests 44 | test1 : 𝕆Type 2 ℓ-zero 45 | test1 = (tt* , (λ _ → Unit)) , λ _ → ℕ 46 | 47 | test2 : 𝕆Type 2 ℓ-zero 48 | test2 = prod test1 test1 49 | 50 | 51 | test3 : snd test2 _ -- since we use records everywhere, agda understands there is only one frame in the product 52 | test3 = 3 ∣ 4 53 | 54 | 55 | -- Fibrancy 56 | fst-src : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (P : Frm X → Type ℓ) (Q : Frm Y → Type ℓ) {f : Frm (prod X Y)} → Src (prod-cell P Q) f → Src P (Frm⇒ Fst f) 57 | fst-src {n} {ℓ} {X} {Y} P Q {f} s = Src⇒ {P = prod-cell P Q} s Fst (λ p → fstₒ (s ⊚ p)) 58 | 59 | snd-src : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (P : Frm X → Type ℓ) (Q : Frm Y → Type ℓ) {f : Frm (prod X Y)} → Src (prod-cell P Q) f → Src Q (Frm⇒ Snd f) 60 | snd-src {n} {ℓ} {X} {Y} P Q {f} s = Src⇒ {P = prod-cell P Q} s Snd (λ p → sndₒ (s ⊚ p)) 61 | 62 | -- Since Inhab is defined differently when n=0 or n=(suc k), the following pattern matching is required for agda to type-check 63 | charac-filler-prod : ∀ {n ℓ} {X Y : 𝕆Type (suc n) ℓ} (P : Frm X → Type ℓ) (Q : Frm Y → Type ℓ) 64 | {f : Frm (prod (fst X) (fst Y))} (s : Src (prod-cell (snd X) (snd Y)) f) → 65 | Iso 66 | (horn-filler (prod-cell P Q) s) 67 | (horn-filler P (fst-src (snd X) (snd Y) s) × horn-filler Q (snd-src (snd X) (snd Y) s)) 68 | charac-filler-prod {zero} P Q s = iso g h (λ _ → refl) (λ _ → refl) where 69 | g : _ 70 | g (tgt , cell) = (fstₒ tgt , fstₒ cell) , (sndₒ tgt , sndₒ cell) 71 | h : _ 72 | h ((tgt₁ , cell₁) , (tgt₂ , cell₂)) = (tgt₁ ∣ tgt₂) , (cell₁ ∣ cell₂) 73 | charac-filler-prod {suc n} P Q s = iso g h (λ _ → refl) λ _ → refl where 74 | g : _ 75 | g (tgt , cell) = (fstₒ tgt , fstₒ cell) , (sndₒ tgt , sndₒ cell) 76 | h : _ 77 | h ((tgt₁ , cell₁) , (tgt₂ , cell₂)) = (tgt₁ ∣ tgt₂) , (cell₁ ∣ cell₂) 78 | 79 | is-fib-prod : ∀ {n ℓ} {X Y : 𝕆Type (suc (suc n)) ℓ} → is-fibrant X → is-fibrant Y → is-fibrant (prod X Y) 80 | is-fib-prod {X = (X , P) , P'} {Y = (Y , Q) , Q'} fibX fibY f s = 81 | -- I'm sure there's a proof in the library that isos/equivs respect HLevel, but I can't find it anymore, so instead I used isContrRetract 82 | isContrRetract (Iso.fun charac) (Iso.inv charac) (Iso.leftInv charac) (isContrΣ (fibX _ _) λ _ → fibY _ _) 83 | where 84 | charac = charac-filler-prod P' Q' s 85 | 86 | prod∞ : ∀ {n ℓ} {Xₙ Yₙ : 𝕆Type ℓ n} (X : 𝕆Type∞ Xₙ) (Y : 𝕆Type∞ Yₙ) → 𝕆Type∞ (prod Xₙ Yₙ) 87 | Fill (prod∞ X Y) = prod-cell (Fill X) (Fill Y) 88 | Hom (prod∞ X Y) = prod∞ (Hom X) (Hom Y) 89 | -------------------------------------------------------------------------------- /Experimental/Local/OpetopicMap.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --no-termination-check #-} 2 | -- 3 | -- OpetopicType.agda - Opetopic Types 4 | -- 5 | 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Data.Sigma 8 | open import Cubical.Data.Empty 9 | open import Cubical.Data.Unit 10 | open import Cubical.Data.Nat 11 | open import Cubical.Data.Sum 12 | 13 | open import Core.Prelude 14 | open import Experimental.Local.OpetopicType 15 | 16 | module Experimental.Local.OpetopicMap where 17 | 18 | -- 19 | -- Maps of Opetopic Types 20 | -- 21 | 22 | infixl 50 _⊙_ 23 | 24 | _⇒_ : ∀ {n ℓ₀ ℓ₁} → 𝕆Type n ℓ₀ → 𝕆Type n ℓ₁ → Type (ℓ-max ℓ₀ ℓ₁) 25 | 26 | id-map : ∀ {n ℓ} → (X : 𝕆Type n ℓ) → X ⇒ X 27 | 28 | _⊙_ : ∀ {n ℓ₀ ℓ₁ ℓ₂} {X : 𝕆Type n ℓ₀} 29 | → {Y : 𝕆Type n ℓ₁} {Z : 𝕆Type n ℓ₂} 30 | → Y ⇒ Z → X ⇒ Y → X ⇒ Z 31 | 32 | Frm⇒ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 33 | → (σ : X ⇒ Y) → Frm X → Frm Y 34 | 35 | Src⇒ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 36 | → {P : Frm X → Type ℓ₀} 37 | → {Q : Frm Y → Type ℓ₁} 38 | → {f : Frm X} (s : Src P f) 39 | → (σ : X ⇒ Y) (σ' : (p : Pos P s) → Q (Frm⇒ σ (Typ P s p))) 40 | → Src Q (Frm⇒ σ f) 41 | 42 | Pos⇒ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 43 | → {P : Frm X → Type ℓ₀} 44 | → {Q : Frm Y → Type ℓ₁} 45 | → {f : Frm X} (s : Src P f) 46 | → (σ : X ⇒ Y) (σ' : (p : Pos P s) → Q (Frm⇒ σ (Typ P s p))) 47 | → Pos P s → Pos Q (Src⇒ s σ σ') 48 | 49 | Pos⇐ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 50 | → {P : Frm X → Type ℓ₀} 51 | → {Q : Frm Y → Type ℓ₁} 52 | → {f : Frm X} (s : Src P f) 53 | → (σ : X ⇒ Y) (σ' : (p : Pos P s) → Q (Frm⇒ σ (Typ P s p))) 54 | → Pos Q (Src⇒ s σ σ') → Pos P s 55 | 56 | postulate 57 | 58 | -- 59 | -- Equations for maps 60 | -- 61 | 62 | map-unit-l : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 63 | → (σ : X ⇒ Y) 64 | → id-map Y ⊙ σ ↦ σ 65 | {-# REWRITE map-unit-l #-} 66 | 67 | map-unit-r : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 68 | → (σ : X ⇒ Y) 69 | → σ ⊙ id-map X ↦ σ 70 | {-# REWRITE map-unit-r #-} 71 | 72 | map-assoc : ∀ {n ℓ₀ ℓ₁ ℓ₂ ℓ₃} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 73 | → {Z : 𝕆Type n ℓ₂} {W : 𝕆Type n ℓ₃} 74 | → (ρ : X ⇒ Y) (σ : Y ⇒ Z) (τ : Z ⇒ W) 75 | → τ ⊙ (σ ⊙ ρ) ↦ τ ⊙ σ ⊙ ρ 76 | {-# REWRITE map-assoc #-} 77 | 78 | -- 79 | -- Frame compatibilities 80 | -- 81 | 82 | Frm⇒-id : ∀ {n ℓ} (X : 𝕆Type n ℓ) (f : Frm X) 83 | → Frm⇒ (id-map X) f ↦ f 84 | {-# REWRITE Frm⇒-id #-} 85 | 86 | Frm⇒-⊙ : ∀ {n ℓ₀ ℓ₁ ℓ₂} {X : 𝕆Type n ℓ₀} 87 | → {Y : 𝕆Type n ℓ₁} {Z : 𝕆Type n ℓ₂} 88 | → (σ : X ⇒ Y) (τ : Y ⇒ Z) (f : Frm X) 89 | → Frm⇒ τ (Frm⇒ σ f) ↦ Frm⇒ (τ ⊙ σ) f 90 | {-# REWRITE Frm⇒-⊙ #-} 91 | 92 | -- 93 | -- Src compatibilities 94 | -- 95 | 96 | Src⇒-id : ∀ {n ℓ} (X : 𝕆Type n ℓ) 97 | → (P : Frm X → Type ℓ) 98 | → (f : Frm X) (s : Src P f) 99 | → Src⇒ {Q = P} s (id-map X) (λ p → s ⊚ p) ↦ s 100 | {-# REWRITE Src⇒-id #-} 101 | 102 | -- Src⇒ and composition? 103 | 104 | Src⇒-ν : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 105 | → {P Q : Frm X → Type ℓ₀} 106 | → {R : Frm Y → Type ℓ₁} 107 | → {f : Frm X} (s : Src P f) 108 | → (ϕ : (p : Pos P s) → Q (Typ P s p)) 109 | → (σ : X ⇒ Y) (σ' : (p : Pos Q (ν s ϕ)) → R (Frm⇒ σ (Typ Q (ν s ϕ) p))) 110 | → Src⇒ {Q = R} (ν s ϕ) σ σ' ↦ Src⇒ s σ (λ p → σ' (ν-pos s ϕ p)) 111 | {-# REWRITE Src⇒-ν #-} 112 | 113 | Src⇒-η : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 114 | → {P : Frm X → Type ℓ₀} 115 | → {Q : Frm Y → Type ℓ₁} 116 | → {f : Frm X} (x : P f) 117 | → (σ : X ⇒ Y) (σ' : (p : Pos P (η P x)) → Q (Frm⇒ σ f)) 118 | → Src⇒ (η P x) σ σ' ↦ η Q (σ' (η-pos P x)) 119 | {-# REWRITE Src⇒-η #-} 120 | 121 | Src⇒-μ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 122 | → {P : Frm X → Type ℓ₀} 123 | → {Q : Frm Y → Type ℓ₁} 124 | → {f : Frm X} (s : Src (Src P) f) 125 | → (σ : X ⇒ Y) (σ' : (p : Pos P (μ P s)) → Q (Frm⇒ σ (Typ P (μ P s) p))) 126 | → Src⇒ (μ P s) σ σ' ↦ μ Q (Src⇒ s σ (λ p → Src⇒ (s ⊚ p) σ (λ q → σ' (μ-pos P s p q)))) 127 | {-# REWRITE Src⇒-μ #-} 128 | 129 | -- 130 | -- Pos⇒ definitional inverse 131 | -- 132 | 133 | Pos⇒-β : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 134 | → {P : Frm X → Type ℓ₀} 135 | → {Q : Frm Y → Type ℓ₁} 136 | → {f : Frm X} (s : Src P f) 137 | → (σ : X ⇒ Y) (σ' : (p : Pos P s) → Q (Frm⇒ σ (Typ P s p))) 138 | → (p : Pos P s) 139 | → Pos⇐ {Q = Q} s σ σ' (Pos⇒ {Q = Q} s σ σ' p) ↦ p 140 | {-# REWRITE Pos⇒-β #-} 141 | 142 | Pos⇒-η : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 143 | → {P : Frm X → Type ℓ₀} 144 | → {Q : Frm Y → Type ℓ₁} 145 | → {f : Frm X} (s : Src P f) 146 | → (σ : X ⇒ Y) (σ' : (p : Pos P s) → Q (Frm⇒ σ (Typ P s p))) 147 | → (p : Pos Q (Src⇒ s σ σ')) 148 | → Pos⇒ {Q = Q} s σ σ' (Pos⇐ {Q = Q} s σ σ' p) ↦ p 149 | {-# REWRITE Pos⇒-η #-} 150 | 151 | -- 152 | -- Typing and Inhabitants 153 | -- 154 | 155 | Pos⇐-Typ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 156 | → {P : Frm X → Type ℓ₀} 157 | → {Q : Frm Y → Type ℓ₁} 158 | → {f : Frm X} (s : Src P f) 159 | → (σ : X ⇒ Y) (σ' : (p : Pos P s) → Q (Frm⇒ σ (Typ P s p))) 160 | → (p : Pos Q (Src⇒ s σ σ')) 161 | → Typ Q (Src⇒ s σ σ') p ↦ Frm⇒ σ (Typ P s (Pos⇐ s σ σ' p)) 162 | {-# REWRITE Pos⇐-Typ #-} 163 | 164 | Pos⇐-⊚ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 165 | → {P : Frm X → Type ℓ₀} 166 | → {Q : Frm Y → Type ℓ₁} 167 | → {f : Frm X} (s : Src P f) 168 | → (σ : X ⇒ Y) (σ' : (p : Pos P s) → Q (Frm⇒ σ (Typ P s p))) 169 | → (p : Pos Q (Src⇒ s σ σ')) 170 | → Src⇒ s σ σ' ⊚ p ↦ σ' (Pos⇐ s σ σ' p) 171 | {-# REWRITE Pos⇐-⊚ #-} 172 | 173 | ν-Src⇒ : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 174 | → {P : Frm X → Type ℓ₀} 175 | → {Q R : Frm Y → Type ℓ₁} 176 | → {f : Frm X} (s : Src P f) 177 | → (σ : X ⇒ Y) (σ' : (p : Pos P s) → Q (Frm⇒ σ (Typ P s p))) 178 | → (ϕ : (p : Pos Q (Src⇒ s σ σ')) → R (Typ Q (Src⇒ s σ σ') p)) 179 | → ν {Q = R} (Src⇒ s σ σ') ϕ ↦ Src⇒ s σ (λ p → ϕ (Pos⇒ s σ σ' p)) 180 | {-# REWRITE ν-Src⇒ #-} 181 | 182 | -- Pos⇒ and id-map and ⊙ ? 183 | 184 | _⇒_ {zero} X Y = Lift Unit 185 | _⇒_ {suc n} (X , P) (Y , Q) = 186 | Σ[ σ ∈ X ⇒ Y ] 187 | ({f : Frm X} → P f → Q (Frm⇒ σ f)) 188 | 189 | id-map {zero} X = tt* 190 | id-map {suc n} (X , P) = id-map X , λ p → p 191 | 192 | _⊙_ {zero} {X = X} {Y} {Z} σ τ = tt* 193 | _⊙_ {suc n} {X = X , P} {Y , Q} {Z , R} (σₙ , σₛₙ) (τₙ , τₛₙ) = 194 | σₙ ⊙ τₙ , λ p → σₛₙ (τₛₙ p) 195 | 196 | Frm⇒ {zero} σ f = tt* 197 | Frm⇒ {suc n} {X = X , P} {Y = Y , Q} (σₙ , σₛₙ) (frm , src , tgt) = 198 | Frm⇒ σₙ frm , Src⇒ src σₙ (λ p → σₛₙ (src ⊚ p)) , σₛₙ tgt 199 | 200 | Src⇒-brs : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 201 | → {P : Frm X → Type ℓ₀} {Q : Frm Y → Type ℓ₁} 202 | → {U : Frm (X , P) → Type ℓ₀} {V : Frm (Y , Q) → Type ℓ₁} 203 | → {f : Frm X} (src : Src P f) (tgt : P f) (flr : U (f , src , tgt)) 204 | → (brs : Dec {P = P} (Branch U) src) 205 | → (σₙ : X ⇒ Y) (σₛₙ : {f : Frm X} → P f → Q (Frm⇒ σₙ f)) 206 | → (σ' : (p : PdPos U (nd src tgt flr brs)) 207 | → V (Frm⇒ (σₙ , σₛₙ) (Typ U (nd src tgt flr brs) p))) 208 | → (p : Pos Q (Src⇒ src σₙ (λ p₁ → σₛₙ (src ⊚ p₁)))) 209 | → Branch V (σₛₙ (src ⊚ Pos⇐ src σₙ (λ p₁ → σₛₙ (src ⊚ p₁)) p)) 210 | Src⇒-brs {X = X} {Y} {P} {Q} {U} {V} src tgt flr brs σₙ σₛₙ σ' p = 211 | let p' = Pos⇐ src σₙ (λ p₁ → σₛₙ (src ⊚ p₁)) p 212 | in [ Src⇒ (lvs (brs ⊛ p')) σₙ (λ q → σₛₙ (lvs (brs ⊛ p') ⊚ q)) 213 | , Src⇒ {X = X , P} {Y = Y , Q} {P = U} {Q = V} (br (brs ⊛ p')) (σₙ , σₛₙ) 214 | (λ q → σ' (nd-there p' q)) ] 215 | 216 | Src⇒ {zero} σ σ' s = s tt* 217 | Src⇒ {suc n} {Q = Q} (lf tgt) (σₙ , σₛₙ) σ' = lf (σₛₙ tgt) 218 | Src⇒ {suc n} {X = X , P} {Y = Y , Q} {P = U} {Q = V} (nd src tgt flr brs) (σₙ , σₛₙ) σ' = 219 | nd (Src⇒ src σₙ (λ p → σₛₙ (src ⊚ p))) (σₛₙ tgt) (σ' nd-here) 220 | (λ-dec {P = Q} (Branch V) (Src⇒ src σₙ (λ p → σₛₙ (src ⊚ p))) 221 | (Src⇒-brs src tgt flr brs σₙ σₛₙ σ')) 222 | 223 | Pos⇒ {zero} s σ σ' p = tt* 224 | Pos⇒ {suc n} (nd src tgt flr brs) (σₙ , σₛₙ) σ' nd-here = nd-here 225 | Pos⇒ {suc n} (nd src tgt flr brs) (σₙ , σₛₙ) σ' (nd-there p q) = 226 | nd-there (Pos⇒ src σₙ (λ q → σₛₙ (src ⊚ q)) p) 227 | (Pos⇒ (br (brs ⊛ p)) (σₙ , σₛₙ) (λ q → σ' (nd-there p q)) q) 228 | 229 | Pos⇐ {zero} s σ σ' p = tt* 230 | Pos⇐ {suc n} (nd src tgt flr brs) (σₙ , σₛₙ) σ' nd-here = nd-here 231 | Pos⇐ {suc n} (nd src tgt flr brs) (σₙ , σₛₙ) σ' (nd-there p q) = 232 | let p' = Pos⇐ src σₙ (λ q → σₛₙ (src ⊚ q)) p 233 | q' = Pos⇐ (br (brs ⊛ p')) (σₙ , σₛₙ) (λ q → σ' (nd-there p' q)) q 234 | in nd-there p' q' 235 | 236 | -------------------------------------------------------------------------------- /Experimental/Local/Shapes.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --no-positivity-check #-} 2 | -- 3 | -- Shapes.agda - General shapes in opetopic types 4 | -- 5 | 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Sigma 9 | 10 | open import Experimental.Local.OpetopicType 11 | 12 | module Experimental.Local.Shapes where 13 | 14 | globe-Frm : ∀ {n ℓ} {Xₙ : 𝕆Type n ℓ} (Xₛₙ : Frm Xₙ → Type ℓ) {f : Frm Xₙ} (x y : Xₛₙ f) → Frm (Xₙ , Xₛₙ) 15 | globe-Frm Xₛₙ {f} x y = f , η Xₛₙ x , y 16 | 17 | _**_ : ∀ {ℓ} → Type ℓ → ℕ → Type ℓ 18 | X ** zero = Lift Unit 19 | X ** suc zero = X 20 | X ** suc (suc n) = X × (X ** (suc n)) 21 | 22 | fstt : ∀ {ℓ n} {X : Type ℓ} → X ** (suc n) → X 23 | fstt {n = zero} x = x 24 | fstt {n = suc n} (x , _) = x 25 | 26 | last : ∀ {ℓ n} {X : Type ℓ} → X ** (suc n) → X 27 | last {n = zero} x = x 28 | last {n = suc n} (fst , snd) = last snd 29 | 30 | cons : ∀ {n ℓ} {X : Type ℓ} → X → X ** n → X ** (suc n) 31 | cons {zero} x l = x 32 | cons {suc n} x l = x , l 33 | 34 | path-chain : ∀ {ℓ} {X : Type ℓ} (Y : X → X → Type ℓ) (n : ℕ) (t : X ** (suc n)) → Type ℓ 35 | path-chain Y zero _ = Lift Unit 36 | path-chain Y (suc zero) (y , x) = Y x y 37 | path-chain Y (suc (suc n)) (x , t) = Y (fstt t) x × path-chain Y (suc n) t 38 | 39 | chain-cons : ∀ {ℓ} {X : Type ℓ} (Y : X → X → Type ℓ) {n : ℕ} {y : X} {t : X ** (suc n)} (p : path-chain Y n t) → Y (fstt t) y → path-chain Y (suc n) (y , t) 40 | chain-cons Y {zero} {y} {t} _ f = f 41 | chain-cons Y {suc n} {y} {t} p f = f , p 42 | 43 | -- Sequences of arrows 44 | n-path : ∀ {ℓ} {X₀ : 𝕆Type 1 ℓ} (X₁ : Frm X₀ → Type ℓ) (n : ℕ) {t : (X₀ .snd tt*) ** suc n} 45 | → path-chain (λ x y → X₁ (tt* , x , y)) n t 46 | → Src X₁ (tt* , last t , fstt t) 47 | n-path X₁ zero {x} f = lf x 48 | n-path X₁ (suc zero) {y , x} f = nd x y f [ x , lf x ] 49 | n-path X₁ (suc (suc n)) {y , x , t} (f , l) = nd x y f [ last t , n-path X₁ (suc n) l ] 50 | 51 | -- -- Sequences of unary higher cells 52 | -- n-path' : ∀ {ℓ} {m : ℕ} {X : 𝕆Type m ℓ} (n : ℕ) {P : Frm X → Type ℓ} (Q : Frm (X , P) → Type ℓ) {f : Frm X} 53 | -- {t : P f ** (suc n)} 54 | -- → path-chain (λ x y → Q (globe-Frm P x y)) n t 55 | -- → Src Q (globe-Frm P (last t) (fstt t)) 56 | -- n-path' zero {P} Q {f} {x} _ = lf x 57 | -- n-path' (suc zero) {P} Q {f} {y , x} p = nd y (η (Branch Q) [ x , _ , (lf x) ]) p 58 | -- n-path' (suc (suc n)) {P} Q {f} {y , x , t} (p , l) = nd y (η (Branch Q) [ x , (η P (last t)) , (n-path' (suc n) Q l) ]) p 59 | 60 | module _ {ℓ} (X₀ : 𝕆Type 1 ℓ) where 61 | Obj : Type ℓ 62 | Obj = snd X₀ tt* 63 | 64 | hom-Frm : Obj → Obj → Frm X₀ 65 | hom-Frm x y = (tt* , x , y) 66 | 67 | module _ (X₁ : Frm X₀ → Type ℓ) where 68 | hom : Obj → Obj → Type ℓ 69 | hom x y = X₁ (hom-Frm x y) 70 | 71 | simplex-Frm : {x y z : Obj} (f : hom x y) (g : hom y z) (h : hom x z) → Frm (X₀ , X₁) 72 | simplex-Frm {x} {y} {z} f g h = hom-Frm x z , n-path X₁ 2 (g , f) , h 73 | 74 | 2-drop-Frm : (x : Obj) (f : hom x x) → Frm (X₀ , X₁) 75 | 2-drop-Frm x f = hom-Frm x x , lf x , f 76 | 77 | -- {- 78 | -- src→chain : {f : Frm X₀} → Src X₁ f → Σ[ n ∈ ℕ ] Σ[ t ∈ (Obj ** n) ] (path-chain hom n (cons (snd (snd f)) t)) 79 | -- src→chain (lf tgt) = 0 , tt* , tt* 80 | -- src→chain (nd tgt [ stm₁ , .(η (λ f → snd X₀ tt*) stm₁) , _ ] flr) = ? 81 | -- src→chain (nd tgt [ stm₁ , .(μ (id-map tt*) (Branch X₁) (λ f → snd X₀ tt*) brs (λ p → lvs (brs ⊚ p))) , nd .stm₁ brs flr₁ ] flr) = ? 82 | -- -} 83 | 84 | module _ (X₂ : Frm (X₀ , X₁) → Type ℓ) where 85 | 86 | Null : (x : Obj) (f : hom x x) → Type ℓ 87 | Null x f = X₂ (2-drop-Frm x f) 88 | 89 | 2Glob : {x y : Obj} (f g : hom x y) → Type ℓ 90 | 2Glob {x} {y} f g = X₂ ((tt* , x , y) , (nd x y f [ x , lf x ]) , g) 91 | 92 | Simplex : {x y z : Obj} (f : hom x y) (g : hom y z) 93 | → (h : hom x z) → Type ℓ 94 | Simplex {x} {y} {z} f g h = X₂ (simplex-Frm f g h) 95 | 96 | right-unitor-Src : (x y : Obj) (f : hom y y) (g h : hom x y) 97 | → (Δ : X₂ (simplex-Frm g f h)) 98 | → (Γ : X₂ (2-drop-Frm y f)) 99 | → Src X₂ (globe-Frm X₁ g h) 100 | right-unitor-Src x y f g h Δ Γ = nd (n-path X₁ 2 (f , g)) h Δ ([ lf y , nd (lf y) f Γ tt* ] , [ η X₁ g , (lf g) ] , tt*) 101 | 102 | left-unitor-Src : (x y : Obj) (f : hom x x) (g h : hom x y) 103 | → (Δ : X₂ (simplex-Frm f g h)) 104 | → (Γ : X₂ (2-drop-Frm x f)) 105 | → Src X₂ (globe-Frm X₁ g h) 106 | left-unitor-Src x y f g h Δ Γ = nd (n-path X₁ 2 (g , f)) h Δ ([ η X₁ g , (lf g) ] , [ lf x , nd (lf x) f Γ tt* ] , tt*) 107 | 108 | 109 | -- (h | g | f) → ((h ∘ g) | f) → ((h ∘ g) ∘ f) 110 | right-associator-Src : (x y z w : Obj) (f : hom x y) (g : hom y z) (h : hom z w) (i : hom y w) (j : hom x w) 111 | → (Δ₁ : X₂ (simplex-Frm f i j)) 112 | → (Δ₂ : X₂ (simplex-Frm g h i)) 113 | → Src X₂ (_ , n-path X₁ 3 (h , g , f) , j) 114 | right-associator-Src x y z w f g h i j Δ₁ Δ₂ = nd (n-path X₁ 2 (i , f)) j Δ₁ ( 115 | [ _ {-"n-path X₁ 2 (h , g)" here should work but doesn't-} , (nd (n-path X₁ 2 (h , g)) i Δ₂ ([ _ , lf h ] , [ _ , lf g ] , tt*)) ] , 116 | [ _ , lf f ] , 117 | tt*) 118 | 119 | -- (h | g | f) → (h | (g ∘ f)) → (h ∘ (g ∘ f)) 120 | left-associator-Src : (x y z w : Obj) (f : hom x y) (g : hom y z) (h : hom z w) (i : hom x z) (j : hom x w) 121 | → (Δ₁ : X₂ (simplex-Frm i h j)) 122 | → (Δ₂ : X₂ (simplex-Frm f g i)) 123 | → Src X₂ (_ , n-path X₁ 3 (h , g , f) , j) 124 | left-associator-Src x y z w f g h i j Δ₁ Δ₂ = nd (n-path X₁ 2 (h , i)) j Δ₁ ( 125 | [ _ , lf h ] , 126 | [ _ , nd (n-path X₁ 2 (g , f)) i Δ₂ ([ _ , lf g ] , [ _ , lf f ] , tt*) ] , 127 | tt*) 128 | -------------------------------------------------------------------------------- /Experimental/Local/Structures.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Everything 2 | 3 | open import Core.Prelude 4 | 5 | open import Cubical.Data.Nat 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Foundations.Equiv.Fiberwise 8 | 9 | open import Experimental.Local.OpetopicType 10 | open import Experimental.Local.OpetopicMap 11 | open import Experimental.Local.Shapes 12 | 13 | module Experimental.Local.Structures where 14 | 15 | record 𝕆Type∞ {n ℓ} (Xₙ : 𝕆Type n ℓ) : Type (ℓ-suc ℓ) where 16 | coinductive 17 | field 18 | Fill : Frm Xₙ → Type ℓ 19 | Hom : 𝕆Type∞ (Xₙ , Fill) 20 | open 𝕆Type∞ public 21 | 22 | record Map {n ℓ} {Xₙ Yₙ : 𝕆Type n ℓ} (σ : Xₙ ⇒ Yₙ) (X : 𝕆Type∞ Xₙ) (Y : 𝕆Type∞ Yₙ) : Type ℓ where 23 | coinductive 24 | field 25 | Fill⇒ : {f : Frm Xₙ} → (Fill X) f → (Fill Y) (Frm⇒ σ f) 26 | Hom⇒ : Map (σ , Fill⇒) (Hom X) (Hom Y) 27 | open Map public 28 | 29 | horn-filler : ∀ {n ℓ} {Xₙ : 𝕆Type n ℓ} {Xₛₙ : Frm Xₙ → Type ℓ} (Xₛₛₙ : Frm (Xₙ , Xₛₙ) → Type ℓ) {f : Frm Xₙ} → Src Xₛₙ f → Type ℓ 30 | horn-filler {n} {ℓ} {Xₙ} {Xₛₙ} Xₛₛₙ {f} s = Σ[ tgt ∈ Xₛₙ f ] Xₛₛₙ (f , s , tgt) 31 | 32 | is-fibrant : ∀ {n ℓ} → 𝕆Type (suc (suc n)) ℓ → Type ℓ 33 | is-fibrant ((Xₙ , Xₛₙ) , Xₛₛₙ) = (f : Frm Xₙ) (s : Src Xₛₙ f) → isContr (horn-filler Xₛₛₙ s) 34 | 35 | isProp-is-fibrant : ∀ {n ℓ} {X : 𝕆Type (suc (suc n)) ℓ} → isProp (is-fibrant X) 36 | isProp-is-fibrant = isPropΠ2 (λ _ _ → isPropIsContr) 37 | 38 | record is-fibrant-ext {n ℓ} {Xₙ : 𝕆Type n ℓ} (X : 𝕆Type∞ Xₙ) : Type ℓ where 39 | coinductive 40 | field 41 | fill-fib : is-fibrant ((Xₙ , (Fill X)) , (Fill (Hom X))) 42 | hom-fib : is-fibrant-ext (Hom X) 43 | 44 | open is-fibrant-ext public 45 | 46 | eta-fib-ext : ∀ {m ℓ} {X : 𝕆Type m ℓ} {X∞ : 𝕆Type∞ X} → X∞ ≡ record { Fill = Fill X∞ ; Hom = Hom X∞ } 47 | Fill (eta-fib-ext {X∞ = X∞} i) = Fill X∞ 48 | Hom (eta-fib-ext {X∞ = X∞} i) = Hom X∞ 49 | 50 | isProp-is-fibrant-ext : ∀ {n ℓ} {X : 𝕆Type n ℓ} {X∞ : 𝕆Type∞ X} → isProp (is-fibrant-ext X∞) 51 | fill-fib (isProp-is-fibrant-ext x y i) = isProp-is-fibrant (fill-fib x) (fill-fib y) i 52 | hom-fib (isProp-is-fibrant-ext x y i) = isProp-is-fibrant-ext (hom-fib x) (hom-fib y) i 53 | 54 | -- hom→path0 : ∀ {ℓ} {X : 𝕆Type∞ (𝕋 {ℓ} 0)} → is-fibrant-ext X → (x y : X .Fill tt*) → X .Hom .Fill (tt* , x , y) → x ≡ y 55 | -- hom→path0 {ℓ} {X} fib x y σ = sym (transportRefl x) ∙ cong fst (isContr→isProp (fib .fill-fib tt* x) (Id x x refl) b) where 56 | -- Id : (x y : X .Fill tt*) → (x ≡ y) → horn-filler (Fill (Hom X)) x 57 | -- Id x y = J (λ y p → horn-filler (Fill (Hom X)) x) (x , fib .hom-fib .fill-fib (tt* , x , x) (lf x) .fst .fst) 58 | 59 | -- b : horn-filler (Fill (Hom X)) x 60 | -- b = y , σ 61 | 62 | -- Id-cell : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} (t : Fill X∞ f) → Fill (Hom X∞) (globe-Frm _ t t) 63 | -- Id-cell X∞ fib t = fib .hom-fib .fill-fib _ (lf t) .fst .fst 64 | 65 | -- path→hom : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} 66 | -- → t ≡ u → Fill (Hom X∞) (globe-Frm _ t u) 67 | -- path→hom X∞ fib {t = t} {u = u} = J (λ u p → Fill (Hom X∞) (globe-Frm _ t u)) (Id-cell X∞ fib t) 68 | 69 | -- hom→path : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} 70 | -- → Fill (Hom X∞) (globe-Frm _ t u) → t ≡ u 71 | -- hom→path X∞ fib {f} {t} {u} cell = cong fst (isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (u , cell)) 72 | 73 | -- module IsoHomPath {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} where 74 | -- sec : section (path→hom X∞ fib {f} {t} {u}) (hom→path X∞ fib {f} {t} {u}) 75 | -- sec cell = fromPathP (cong snd (isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (u , cell))) 76 | -- -- Wow ! Might need to take a look back at this later cause I didn't expect it to be so "simple". 77 | 78 | -- ret : retract (path→hom X∞ fib {f} {t} {u}) (hom→path X∞ fib {f} {t} {u}) 79 | -- ret = J (λ u p → hom→path X∞ fib (path→hom X∞ fib p) ≡ p) 80 | -- (hom→path X∞ fib (path→hom X∞ fib refl) 81 | -- ≡⟨ cong (hom→path X∞ fib) (transportRefl _) ⟩ 82 | -- hom→path X∞ fib (Id-cell X∞ fib t) 83 | -- ≡⟨ (λ i j → fst (test i j)) ⟩ 84 | -- refl ∎) where 85 | -- a : horn-filler (Fill (Hom X∞)) (η (Fill X∞) t) 86 | -- a = (t , Id-cell X∞ fib t) 87 | 88 | -- test0 : a ≡ a 89 | -- test0 = isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (t , Id-cell X∞ fib t) 90 | 91 | -- test : test0 ≡ refl 92 | -- test = isProp→isSet (isContr→isProp (fib .fill-fib f (η (Fill X∞) t))) (t , Id-cell X∞ fib t) (t , Id-cell X∞ fib t) test0 refl 93 | 94 | record is-n-trunc {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) : Type ℓ where 95 | coinductive 96 | field 97 | hLevel : (f : Frm X) → isOfHLevel n (X∞ .Fill f) 98 | is-trunc-ext : is-n-trunc (predℕ n) (X∞ .Hom) 99 | open is-n-trunc 100 | 101 | isProp-is-trunc : ∀ {m ℓ n} {X : 𝕆Type m ℓ} {X∞ : 𝕆Type∞ X} → isProp (is-n-trunc n X∞) 102 | hLevel (isProp-is-trunc x y i) f = isPropIsOfHLevel _ (hLevel x f) (hLevel y f) i 103 | is-trunc-ext (isProp-is-trunc x y i) = isProp-is-trunc (is-trunc-ext x) (is-trunc-ext y) i 104 | 105 | src-comp : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) → is-fibrant-ext X∞ → {f : Frm X} → Src (Fill X∞) f → Fill X∞ f 106 | src-comp X∞ fib s = fib .fill-fib _ s .fst .fst 107 | 108 | src-comp-witness : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib-ext : is-fibrant-ext X∞) {f : Frm X} (s : Src (Fill X∞) f) 109 | → Fill (Hom X∞) (f , s , src-comp X∞ fib-ext s) 110 | src-comp-witness X∞ fib s = fib .fill-fib _ s .fst .snd 111 | 112 | -- More general version of the equivalence between hom and path, using the fundamental theorem of identity types 113 | cell≃path : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} (s : Src (Fill X∞) f) (t : Fill X∞ f) 114 | → (src-comp X∞ fib s ≡ t) ≃ Fill (Hom X∞) (f , s , t) 115 | cell≃path X∞ fib s t = recognizeId (λ t → (Fill (Hom X∞)) (_ , s , t)) (fib .fill-fib _ s .fst .snd) (fib .fill-fib _ s) t 116 | 117 | isOfHLevelPathPred : ∀ (n : ℕ) {ℓ} {A : Type ℓ} → isOfHLevel n A → {x y : A} → isOfHLevel (predℕ n) (x ≡ y) 118 | isOfHLevelPathPred zero h = isContr→isContrPath h _ _ 119 | isOfHLevelPathPred (suc n) h = isOfHLevelPath' n h _ _ 120 | 121 | is-n-trunc-fib : ∀ {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) → is-fibrant-ext X∞ → ((f : Frm X) → isOfHLevel n (X∞ .Fill f)) → is-n-trunc n X∞ 122 | hLevel (is-n-trunc-fib n {X} X∞ fib h) = h 123 | is-trunc-ext (is-n-trunc-fib n {X} X∞ fib h) = is-n-trunc-fib _ _ (fib .hom-fib) lemma where 124 | lemma : (f : Frm (X , Fill X∞)) → isOfHLevel (predℕ n) (X∞ .Hom .Fill f) 125 | lemma (f , s , t) = isOfHLevelRespectEquiv (predℕ n) (cell≃path X∞ fib s t) (isOfHLevelPathPred n (h f)) 126 | 127 | total : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) → 𝕆Type∞ {ℓ = ℓ} tt* 128 | total {zero} X∞ = X∞ 129 | total {suc n} {ℓ} {X , P} X∞ = total (record { Fill = P ; Hom = X∞ }) 130 | 131 | 𝕆∞Path : ∀ {m ℓ} {X : 𝕆Type m ℓ} {X∞ X∞' : 𝕆Type∞ X} (p : Fill X∞ ≡ Fill X∞') → PathP (λ i → 𝕆Type∞ (X , p i)) (Hom X∞) (Hom X∞') → X∞ ≡ X∞' 132 | Fill (𝕆∞Path p q i) = p i 133 | Hom (𝕆∞Path p q i) = q i 134 | 135 | 0-trunc-≡ : ∀ {n ℓ} {X : 𝕆Type n ℓ} {X' : 𝕆Type n ℓ} (p : X ≡ X') {X∞ : 𝕆Type∞ X} {X∞' : 𝕆Type∞ X'} 136 | → is-n-trunc 0 X∞ → is-n-trunc 0 X∞' 137 | → PathP (λ i → 𝕆Type∞ (p i)) X∞ X∞' 138 | Fill (0-trunc-≡ p {X∞} {X∞'} trunc trunc' i) = eq i where 139 | eq : PathP (λ i → Frm (p i) → Type _) (Fill X∞) (Fill X∞') 140 | eq = toPathP (funExt λ f → isoToPath (isContr→Iso (trunc .hLevel _) (trunc' .hLevel _))) 141 | Hom (0-trunc-≡ p trunc trunc' i) = 0-trunc-≡ (λ j → p j , Fill (0-trunc-≡ p trunc trunc' j)) (trunc .is-trunc-ext) (trunc' .is-trunc-ext) i 142 | 143 | -- -- Trying to prove that fibrant opetopic types with same base-type are equal 144 | -- {- 145 | -- lemma0 : ∀ {m ℓ} {X X' : 𝕆Type m ℓ} (p : X ≡ X') 146 | -- → (P : Frm X → Type ℓ) (P' : Frm X' → Type ℓ) 147 | -- → (q : PathP (λ i → Frm (p i) → Type ℓ) P P') 148 | -- → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X' , P')) 149 | -- → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 150 | -- → PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill P∞) (Fill P∞') 151 | -- lemma0 {ℓ = ℓ} {X = X} = 152 | -- J 153 | -- (λ X' p → (P : Frm X → Type ℓ) (P' : Frm X' → Type ℓ) 154 | -- → (q : PathP (λ i → Frm (p i) → Type ℓ) P P') 155 | -- → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X' , P')) 156 | -- → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 157 | -- → PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill P∞) (Fill P∞')) 158 | 159 | -- λ P _ → J 160 | 161 | -- (λ P' q → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X , P')) 162 | -- → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 163 | -- → PathP (λ i → Frm (refl i , q i) → Type ℓ) (Fill P∞) (Fill P∞')) 164 | 165 | -- λ P∞ _ fib fib' → funExt λ x → ua (invEquiv (cell≃path _ fib _ _)) ∙ cong (_≡ (snd (snd x))) {!!} ∙ ua (cell≃path _ fib' _ _) 166 | 167 | -- fibrant→≡ : ∀ {m ℓ} {X X' : 𝕆Type m ℓ} (p : X ≡ X') 168 | -- → (X∞ : 𝕆Type∞ X) (X∞' : 𝕆Type∞ X') → PathP (λ i → Frm (p i) → Type ℓ) (Fill X∞) (Fill X∞') 169 | -- → is-fibrant-ext X∞ → is-fibrant-ext X∞' 170 | -- → PathP (λ i → 𝕆Type∞ (p i)) X∞ X∞' 171 | -- Fill (fibrant→≡ p X∞ X∞' q fib fib' i) = q i 172 | -- Hom (fibrant→≡ {ℓ = ℓ} p X∞ X∞' q fib fib' i) = fibrant→≡ (λ i → (p i , q i)) (Hom X∞) (Hom X∞') lemma (hom-fib fib) (hom-fib fib') i where 173 | -- lemma : PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill (Hom X∞)) (Fill (Hom X∞')) 174 | -- lemma = lemma0 p _ _ _ (Hom X∞) (Hom X∞') (subst is-fibrant-ext eta-fib-ext fib) (subst is-fibrant-ext eta-fib-ext fib') 175 | -- -} 176 | -- {- 177 | -- is-n-cat : ∀ {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) → Type ℓ 178 | -- is-n-cat zero X∞ = is-fibrant-ext X∞ 179 | -- is-n-cat (suc n) X∞ = Σ (is-n-cat n (Hom X∞)) {!!} 180 | -- -} 181 | -------------------------------------------------------------------------------- /Experimental/Local/Terminal.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Everything 2 | 3 | open import Core.Prelude 4 | 5 | open import Cubical.Data.Nat 6 | open import Cubical.Data.Sigma 7 | 8 | open import Experimental.Local.OpetopicType 9 | open import Experimental.Local.Structures 10 | 11 | module Experimental.Local.Terminal where 12 | 13 | 𝕋 : ∀ {ℓ} (n : ℕ) → 𝕆Type n ℓ 14 | 𝕋 zero = tt* 15 | 𝕋 (suc n) = 𝕋 n , λ _ → Lift Unit 16 | 17 | 𝕋Ext : ∀ {n ℓ} (X : 𝕆Type n ℓ) → 𝕆Type∞ X 18 | Fill (𝕋Ext X) = λ _ → Lift Unit 19 | Hom (𝕋Ext X) = 𝕋Ext _ 20 | 21 | is-fib-ext-𝕋Ext : ∀ {n ℓ} {X : 𝕆Type n ℓ} → is-fibrant-ext (𝕋Ext X) 22 | fill-fib is-fib-ext-𝕋Ext f s = (tt* , tt*) , λ (tt* , tt*) → refl 23 | hom-fib is-fib-ext-𝕋Ext = is-fib-ext-𝕋Ext 24 | 25 | is-0-trunc-𝕋Ext : ∀ {n ℓ} {X : 𝕆Type n ℓ} → is-n-trunc 0 (𝕋Ext X) 26 | is-n-trunc.hLevel is-0-trunc-𝕋Ext f = tt* , λ _ → refl 27 | is-n-trunc.is-trunc-ext is-0-trunc-𝕋Ext = is-0-trunc-𝕋Ext 28 | -------------------------------------------------------------------------------- /Experimental/NoDecs/CartesianProduct.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- CartsianProduct.agda - Cartesian product of Opetopic Types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Nat 10 | open import Cubical.Data.Sum 11 | 12 | open import Core.Prelude 13 | open import Experimental.NoDecs.OpetopicType 14 | 15 | module Experimental.NoDecs.CartesianProduct where 16 | 17 | prod : ∀ {n ℓ} (X Y : 𝕆Type n ℓ) → 𝕆Type n ℓ 18 | 19 | -- The use of opetopic maps circumvents any need for additional rewrite rules 20 | Fst : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} → prod X Y ⇒ X 21 | Snd : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} → prod X Y ⇒ Y 22 | 23 | record prod-cell {n ℓ} {X Y : 𝕆Type n ℓ} (P : Frm X → Type ℓ) (Q : Frm Y → Type ℓ) (f : Frm (prod X Y)) : Type ℓ where 24 | constructor _∣_ 25 | field 26 | fstₒ : P (Frm⇒ Fst f) 27 | sndₒ : Q (Frm⇒ Snd f) 28 | open prod-cell 29 | 30 | prod {zero} X Y = tt* 31 | prod {suc n} (X , P) (Y , Q) = prod X Y , prod-cell P Q 32 | 33 | Fst {zero} = tt* 34 | Fst {suc n} {ℓ} {X , P} {Y , Q} = Fst , fstₒ 35 | 36 | Snd {zero} = tt* 37 | Snd {suc n} {ℓ} {X , P} {Y , Q} = Snd , sndₒ 38 | 39 | 40 | ---------- Tests 41 | test1 : 𝕆Type 2 ℓ-zero 42 | test1 = (tt* , (λ _ → Unit)) , λ _ → ℕ 43 | 44 | test2 : 𝕆Type 2 ℓ-zero 45 | test2 = prod test1 test1 46 | 47 | open import Experimental.NoDecs.Shapes 48 | 49 | test3 : snd test2 _ -- since we use records everywhere, agda understands there is only one frame in the product 50 | test3 = 3 ∣ 4 51 | 52 | 53 | -- Fibrancy 54 | fst-src : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (P : Frm X → Type ℓ) (Q : Frm Y → Type ℓ) {f : Frm (prod X Y)} → Src (prod-cell P Q) f → Src P (Frm⇒ Fst f) 55 | fst-src {n} {ℓ} {X} {Y} P Q {f} s = Src⇒ Fst (prod-cell P Q) P fstₒ s 56 | 57 | snd-src : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (P : Frm X → Type ℓ) (Q : Frm Y → Type ℓ) {f : Frm (prod X Y)} → Src (prod-cell P Q) f → Src Q (Frm⇒ Snd f) 58 | snd-src {n} {ℓ} {X} {Y} P Q {f} s = Src⇒ Snd (prod-cell P Q) Q sndₒ s 59 | 60 | open import Experimental.NoDecs.Structures 61 | -- Since Inhab is defined differently when n=0 or n=(suc k), the following pattern matching is required for agda to type-check 62 | charac-filler-prod : ∀ {n ℓ} {X Y : 𝕆Type (suc n) ℓ} (P : Frm X → Type ℓ) (Q : Frm Y → Type ℓ) 63 | {f : Frm (prod (fst X) (fst Y))} (s : Src (prod-cell (snd X) (snd Y)) f) → 64 | Iso 65 | (horn-filler (prod-cell P Q) s) 66 | (horn-filler P (fst-src (snd X) (snd Y) s) × horn-filler Q (snd-src (snd X) (snd Y) s)) 67 | charac-filler-prod {zero} P Q s = iso g h (λ _ → refl) (λ _ → refl) where 68 | g : _ 69 | g (tgt , cell) = (fstₒ tgt , fstₒ cell) , (sndₒ tgt , sndₒ cell) 70 | h : _ 71 | h ((tgt₁ , cell₁) , (tgt₂ , cell₂)) = (tgt₁ ∣ tgt₂) , (cell₁ ∣ cell₂) 72 | charac-filler-prod {suc n} P Q s = iso g h (λ _ → refl) λ _ → refl where 73 | g : _ 74 | g (tgt , cell) = (fstₒ tgt , fstₒ cell) , (sndₒ tgt , sndₒ cell) 75 | h : _ 76 | h ((tgt₁ , cell₁) , (tgt₂ , cell₂)) = (tgt₁ ∣ tgt₂) , (cell₁ ∣ cell₂) 77 | 78 | is-fib-prod : ∀ {n ℓ} {X Y : 𝕆Type (suc (suc n)) ℓ} → is-fibrant X → is-fibrant Y → is-fibrant (prod X Y) 79 | is-fib-prod {X = (X , P) , P'} {Y = (Y , Q) , Q'} fibX fibY f s = 80 | -- I'm sure there's a proof in the library that isos/equivs respect HLevel, but I can't find it anymore, so instead I used isContrRetract 81 | isContrRetract (Iso.fun charac) (Iso.inv charac) (Iso.leftInv charac) (isContrΣ (fibX _ _) λ _ → fibY _ _) 82 | where 83 | charac = charac-filler-prod P' Q' s 84 | 85 | prod∞ : ∀ {n ℓ} {Xₙ Yₙ : 𝕆Type ℓ n} (X : 𝕆Type∞ Xₙ) (Y : 𝕆Type∞ Yₙ) → 𝕆Type∞ (prod Xₙ Yₙ) 86 | Fill (prod∞ X Y) = prod-cell (Fill X) (Fill Y) 87 | Hom (prod∞ X Y) = prod∞ (Hom X) (Hom Y) 88 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Category.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Category.agda - Equivalence between categories and 1-truncated (∞,1)-categories 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Nat 7 | open import Cubical.Data.Sigma 8 | open import Cubical.Categories.Category 9 | 10 | open import Experimental.NoDecs.OpetopicType 11 | open import Experimental.NoDecs.Shapes 12 | open import Experimental.NoDecs.Structures 13 | 14 | module Experimental.NoDecs.Category where 15 | 16 | open import Core.Prelude 17 | μ' : ∀ {n ℓ} {X : 𝕆Type n ℓ} (P : Frm X → Type ℓ) {f : Frm X} → Src (Src P) f → Src P f 18 | μ' {X = X} P s = μ (id-map X) (Src P) P s (λ p → s ⊚ p) 19 | 20 | {-μ-factors : ∀ {n ℓ} {X Y : 𝕆Type (suc n) ℓ} (σ : X ⇒ Y) (P : Frm X → Type ℓ) (Q : Frm Y → Type ℓ) {f : Frm X} (s : Src P f) → 21 | (m : (p : Pos P s) → Src Q (Frm⇒ σ (Typ P s p))) → 22 | μ σ P Q s m ≡ μ' Q (map-src σ P (Src Q) s m) 23 | μ-factors σ P Q s m = {!!}-} 24 | 25 | module _ {ℓ} (C : Category ℓ ℓ) where 26 | open Category C renaming (id to C-id ; _⋆_ to _⨀_) 27 | 28 | n-comp : {n : ℕ} {t : ob ** (suc n)} → path-chain Hom[_,_] n t → Hom[ last t , fstt t ] 29 | n-comp {zero} {x} _ = C-id 30 | n-comp {suc zero} {y , x} f = f 31 | n-comp {suc (suc n)} {y , x , t} (f , l) = n-comp {suc n} {x , t} l ⨀ f 32 | 33 | private 34 | X₀ : Frm (𝕋 {ℓ = ℓ} 0) → Type ℓ 35 | X₀ _ = ob 36 | 37 | X₁ : Frm (tt* , X₀) → Type ℓ 38 | X₁ (_ , x , y) = Hom[ x , y ] 39 | 40 | C-src-comp : {f : Frm (tt* , X₀)} → Src X₁ f → X₁ f 41 | C-src-comp (lf x) = C-id {x = x} 42 | C-src-comp (nd tgt [ stm₁ , lvs₁ , br₁ ] flr) = (C-src-comp br₁) ⨀ flr 43 | 44 | X₂ : Frm ((tt* , X₀) , X₁) → Type ℓ 45 | X₂ (f , s , t) = C-src-comp s ≡ t 46 | 47 | -- lemma relating C-src-comp to μ 48 | {-# TERMINATING #-} 49 | pre-big-lemma : {f : Frm (tt* , X₀)} (brs : Src (Src X₁) f) → 50 | C-src-comp (μ (id-map _) (Src X₁) X₁ brs λ p → brs ⊚ p) ≡ 51 | C-src-comp (μ (id-map _) (Src X₁) X₁ brs λ p → η X₁ (C-src-comp (brs ⊚ p))) 52 | pre-big-lemma (lf tgt) = refl 53 | pre-big-lemma (nd _ brs (lf _)) = pre-big-lemma (br brs) ∙ sym (⋆IdR _) 54 | pre-big-lemma (nd tgt brs (nd _ [ stm₁ , _ , br₁ ] flr)) = (cong (_⨀ flr) (pre-big-lemma (nd stm₁ brs br₁))) ∙ ⋆Assoc _ (C-src-comp br₁) flr 55 | 56 | big-lemma : {X₁' : Frm (tt* , X₀) → Type ℓ} {f : Frm (tt* , X₀)} (brs : Src X₁' f) → 57 | (truc : (p : Pos X₁' brs) → Src X₁ (Frm⇒ (id-map _) (Typ X₁' brs p))) → 58 | C-src-comp (μ (id-map _) X₁' X₁ brs truc) ≡ 59 | C-src-comp (μ (id-map _) X₁' X₁ brs λ p → η X₁ (C-src-comp (truc p))) 60 | big-lemma {X₁'} (lf tgt) truc = refl 61 | big-lemma {X₁'} (nd tgt brs flr) truc with (truc nd-here) -- doesn't work as intended 62 | ... | lf .(stm brs) = {!!} 63 | ... | nd _ [ stm₁ , .(stm brs) , br₁ ] flr₁ = {!!} 64 | 65 | {-cong C-src-comp (μ-factors (id-map (tt* , X₀)) X₁' X₁ s truc) ∙ 66 | pre-big-lemma (map-src (id-map (tt* , X₀)) X₁' (Src X₁) s truc) ∙ 67 | {!!} 68 | -- ∙ sym (cong C-src-comp (μ-factors ((id-map (tt* , X₀))) X₁' X₁ s λ p → η X₁ (C-src-comp (truc p))))-} 69 | 70 | {-# TERMINATING #-} 71 | C-2-src-comp : {f : Frm ((tt* , X₀) , X₁)} → Src X₂ f → X₂ f 72 | C-2-src-comp (lf tgt) = ⋆IdL tgt 73 | C-2-src-comp (nd tgt brs flr) = big-lemma brs _ ∙ lemma1 ∙ flr where -- need some kind of lemma relating C-src-comp to μ 74 | test : (p : PdPos (Branch X₂) brs) → Pd X₁ (PdTyp (Branch X₂) brs p) 75 | test p = nd (snd (snd (PdTyp (Branch X₂) brs p))) 76 | [ fst (snd (PdTyp (Branch X₂) brs p)) , 77 | fst (snd (PdTyp (Branch X₂) brs p)) , 78 | lf (fst (snd (PdTyp (Branch X₂) brs p))) ] 79 | (stm (PdInhab (Branch X₂) brs p)) 80 | 81 | test1 : (p : PdPos (Branch X₂) brs) → Pd X₁ (PdTyp (Branch X₂) brs p) 82 | test1 p = η X₁ (stm (PdInhab (Branch X₂) brs p)) 83 | 84 | test2 : (p : PdPos (Branch X₂) brs) → test p ≡ test1 p 85 | test2 p = refl 86 | 87 | test3 : (p : PdPos (Branch X₂) brs) → C-src-comp (lvs (PdInhab (Branch X₂) brs p)) ≡ stm (PdInhab (Branch X₂) brs p) 88 | test3 p = C-2-src-comp (br (PdInhab (Branch X₂) brs p)) 89 | 90 | lemma1 : 91 | C-src-comp (μ (id-map _) (Branch X₂) X₁ brs (λ p → η X₁ (C-src-comp (lvs (PdInhab (Branch X₂) brs p))))) ≡ 92 | C-src-comp (μ (id-map _) (Branch X₂) X₁ brs (λ p → test1 p)) 93 | lemma1 = cong (λ x → C-src-comp (μ (id-map _) (Branch X₂) X₁ brs x)) (funExt λ p → cong (η X₁) (test3 p)) 94 | 95 | 96 | Cat→𝕆Type : 𝕆Type∞ {ℓ = ℓ} tt* 97 | Fill Cat→𝕆Type = X₀ 98 | Fill (Hom Cat→𝕆Type) = X₁ 99 | Fill (Hom (Hom Cat→𝕆Type)) = X₂ 100 | Hom (Hom (Hom Cat→𝕆Type)) = 𝕋Ext _ 101 | 102 | open is-fibrant-ext 103 | 104 | lemma : ∀ {ℓ} {A : Type ℓ} {a : A} → isContr (Σ[ b ∈ A ] a ≡ b) 105 | lemma {a = a} = (a , refl) , (λ (b , p) i → (p i , λ j → p (i ∧ j))) 106 | 107 | module _ (isGroupoidC : isGroupoid ob) where 108 | isInfCatCat : is-fibrant-ext (Hom Cat→𝕆Type) 109 | fill-fib isInfCatCat f s = lemma 110 | hom-fib (hom-fib isInfCatCat) = is-fib-ext-𝕋Ext 111 | fill-fib (hom-fib isInfCatCat) (f , s , t) s' = (C-2-src-comp s' , tt*) , λ y → Σ≡Prop (λ _ _ _ → refl) (isSetHom _ _ _ _) 112 | 113 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Groupoid.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Groupoid.agda - Infinity Groupoids 3 | -- 4 | 5 | open import Cubical.Foundations.Everything renaming (_∙_ to _⊛_ ; rUnit to RUnit ; lUnit to LUnit ; assoc to Assoc) 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Nat 9 | open import Cubical.Data.Sum 10 | 11 | open import Core.Prelude 12 | open import Experimental.NoDecs.OpetopicType 13 | open import Experimental.NoDecs.Structures 14 | 15 | module Experimental.NoDecs.Groupoid where 16 | 17 | Grp : ∀ {ℓ} (X : Type ℓ) → (n : ℕ) → 𝕆Type n ℓ 18 | Pt : ∀ {n ℓ} {X : Type ℓ} (x : X) → 𝕋 {ℓ} n ⇒ (Grp X n) 19 | 20 | data GrpCell {n ℓ} (X : Type ℓ) : Frm (Grp X n) → Type ℓ where 21 | reflₒ : (x : X) {f : Frm (𝕋 n)} → GrpCell X (Frm⇒ (Pt x) f) 22 | 23 | Grp X zero = tt* 24 | Grp X (suc n) = Grp X n , GrpCell X 25 | 26 | Pt {zero} x = tt* 27 | Pt {suc n} x = Pt x , λ _ → reflₒ x 28 | 29 | GrpJ : ∀ {ℓ ℓ'} (X : Type ℓ) {n : ℕ} (P : (f : Frm (Grp X n)) → GrpCell X f → Type ℓ') → 30 | ((x : X) (f : Frm (𝕋 n)) → P (Frm⇒ (Pt x) f) (reflₒ x)) → 31 | {f : Frm (Grp X n)} (t : GrpCell X f) → P f t 32 | GrpJ X P base-case (reflₒ x) = base-case x _ 33 | 34 | 35 | Grp∞ : ∀ {ℓ} (X : Type ℓ) (n : ℕ) → 𝕆Type∞ (Grp X n) 36 | Fill (Grp∞ X n) = GrpCell X 37 | Hom (Grp∞ X n) = Grp∞ X (suc n) 38 | {- 39 | Grp-Comp : ∀ {ℓ} (X : Type ℓ) (n : ℕ) {f : Frm (Grp X (suc n))} (s : Src (GrpCell X) f) → horn-filler {suc n} (GrpCell X) s 40 | Grp-Comp X n (lf (reflₒ x)) = (reflₒ x {_ , η {n} _ _ , tt*}) , reflₒ x {_ , lf tt* , tt*} 41 | Grp-Comp X n (nd tgt brs flr) = {!!} 42 | -} 43 | {-is-fib-Grp : ∀ {ℓ} (X : Type ℓ) (n : ℕ) → is-fibrant-ext (Grp∞ X n) 44 | fill-fib (is-fib-Grp X n) f s = {!GrpJ X ? ? {f} !} 45 | hom-fib (is-fib-Grp X n) = is-fib-Grp X (suc n)-} 46 | 47 | {-GrpJ : ∀ {ℓ ℓ'} (X : Type ℓ) (P : (f : Frm (Grp X 1)) (s : Src (GrpCell X) f) → Type ℓ') → 48 | ((x : X) (f : Frm (𝕋 1)) → P {!!} {!!}) → --P (Frm⇒ (Pt x) f) {!!}) → 49 | {f : Frm (Grp X 1)} (s : Src (GrpCell X) f) → P f s 50 | GrpJ X P base-case {.(_ , η (λ f → GrpCell X tt*) (reflₒ x) , reflₒ x)} (lf (reflₒ x)) = base-case x _ 51 | GrpJ X P base-case (nd (reflₒ x) [ _ , lvs₁ , br₁ ] (reflₒ .x)) = {!!}-} 52 | 53 | module _ {ℓ} (X : Type ℓ) where 54 | -- Alternative presentation of Grp 55 | Grp2 : (n : ℕ) → 𝕆Type n ℓ 56 | Grp2Cell : (n : ℕ) → Frm (Grp2 n) → Type ℓ 57 | Grp2Comp : (n : ℕ) {f : Frm (Grp2 n)} → Src (Grp2Cell n) f → Grp2Cell n f 58 | 59 | Grp2 zero = tt* 60 | Grp2 (suc n) = (Grp2 n) , Grp2Cell n 61 | 62 | Grp2Cell zero f = X 63 | Grp2Cell (suc n) (f , s , t) = Grp2Comp n s ≡ t 64 | 65 | Grp2∞ : (n : ℕ) → 𝕆Type∞ (Grp2 n) 66 | Fill (Grp2∞ n) = Grp2Cell n 67 | Hom (Grp2∞ n) = Grp2∞ (suc n) 68 | 69 | lemma : ∀ {ℓ} {A : Type ℓ} {a : A} → isContr (Σ[ b ∈ A ] a ≡ b) 70 | lemma {a = a} = (a , refl) , (λ (b , p) i → (p i , λ j → p (i ∧ j))) 71 | 72 | is-fib-Grp2 : (n : ℕ) → is-fibrant-ext (Grp2∞ n) 73 | fill-fib (is-fib-Grp2 n) f s = lemma 74 | hom-fib (is-fib-Grp2 n) = is-fib-Grp2 (suc n) 75 | 76 | Grp2η : (n : ℕ) {f : Frm (Grp2 n)} (x : Grp2Cell n f) → Grp2Comp n (η (Grp2Cell n) x) ≡ x 77 | Grp2μ : (n : ℕ) {P : Frm (Grp2 n) → Type ℓ} {f : Frm (Grp2 n)} (brs : Src P f) → 78 | (m : (p : Pos P brs) → Src (Grp2Cell n) (Frm⇒ (id-map _) (Typ P brs p))) → 79 | Grp2Comp n (μ (id-map _) P (Grp2Cell n) brs m) ≡ 80 | Grp2Comp n (μ (id-map _) P (Grp2Cell n) brs λ p → η (Grp2Cell n) (Grp2Comp n (m p))) 81 | 82 | abstract 83 | infixr 30 _∙_ 84 | _∙_ : {l : Level} {A : Type l} {x y z : A} → x ≡ y → y ≡ z → x ≡ z 85 | _∙_ = _⊛_ 86 | 87 | rUnit : {l : Level} {A : Type l} {x y : A} (p : x ≡ y) → p ≡ p ∙ refl 88 | rUnit = RUnit 89 | 90 | lUnit : {l : Level} {A : Type l} {x y : A} (p : x ≡ y) → p ≡ refl ∙ p 91 | lUnit = LUnit 92 | 93 | assoc : {l : Level} {A : Type l} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → p ∙ q ∙ r ≡ (p ∙ q) ∙ r 94 | assoc = Assoc 95 | 96 | {-# TERMINATING #-} 97 | Grp2Comp zero {f} x = x 98 | --Grp2Comp (suc zero) (lf tgt) = refl 99 | --Grp2Comp (suc zero) (nd tgt brs flr) = Grp2Comp 1 (br brs) ∙ flr 100 | Grp2Comp (suc n) (lf tgt) = Grp2η n tgt 101 | Grp2Comp (suc n) (nd tgt brs flr) = (Grp2μ n {Branch (Grp2Cell (suc n))} brs (λ p → lvs (brs ⊚ p)) ∙ ii) ∙ flr where 102 | i : (p : Pos (Branch (Grp2Cell (suc n))) brs) → Grp2Comp n (lvs (brs ⊚ p)) ≡ (stm (brs ⊚ p)) 103 | i p = Grp2Comp (suc n) (br (brs ⊚ p)) 104 | 105 | ii : 106 | Grp2Comp n (μ (id-map (Grp2 n)) (Branch (Grp2Cell (suc n))) (Grp2Cell n) brs 107 | (λ p → η (Grp2Cell n) (Grp2Comp n (lvs (brs ⊚ p))))) 108 | ≡ 109 | Grp2Comp n (μ (id-map (Grp2 n)) (Branch (Grp2Cell (suc n))) (Grp2Cell n) brs 110 | (λ p → η (snd (Grp2 (suc n))) (stm (brs ⊚ p)))) 111 | ii = cong (λ x → Grp2Comp n (μ (id-map (Grp2 n)) (Branch (Grp2Cell (suc n))) (Grp2Cell n) brs x)) 112 | (funExt λ p → cong (η (Grp2Cell n)) (i p)) 113 | 114 | -- Probably need some rewrite rules at some point... 115 | Grp2μ zero {P} brs m = refl 116 | Grp2μ (suc n) {P} brs m = {!!} 117 | 118 | Grp2η zero {f} x = refl 119 | Grp2η (suc zero) {f , s , t} x = sym (lUnit _ ∙ lUnit _ ∙ assoc _ _ _ ∙ refl) 120 | Grp2η (suc (suc n)) {f} x = {!!} 121 | -------------------------------------------------------------------------------- /Experimental/NoDecs/NaturalTransform.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- NaturalTransform.agda - Natural Transformations of ∞-Functors 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Nat 10 | open import Cubical.Data.Sum 11 | open import Cubical.Data.Bool 12 | 13 | open import Core.Prelude 14 | open import Experimental.NoDecs.OpetopicType 15 | open import Experimental.NoDecs.Structures 16 | open import Experimental.NoDecs.CartesianProduct 17 | 18 | module Experimental.NoDecs.NaturalTransform where 19 | 20 | 𝕀∞ : 𝕆Type∞ (𝕋 {ℓ-zero} 0) 21 | Fill 𝕀∞ = λ _ → Bool 22 | Fill (Hom 𝕀∞) (_ , x , y) = x ≤ y 23 | Hom (Hom 𝕀∞) = 𝕋Ext _ 24 | 25 | ≤-trans : {x y z : Bool} → x ≤ y → y ≤ z → x ≤ z 26 | ≤-trans {false} {y} {z} _ _ = tt 27 | ≤-trans {true} {true} {true} _ _ = tt 28 | 29 | ≤-src : {x y : Bool} (s : Src (Fill (Hom 𝕀∞)) (_ , x , y)) → x ≤ y 30 | ≤-src (lf false) = tt 31 | ≤-src (lf true) = tt 32 | ≤-src (nd tgt brs flr) = ≤-trans (≤-src (br brs)) flr 33 | 34 | is∞cat𝕀 : is-fibrant-ext (Hom 𝕀∞) 35 | fill-fib is∞cat𝕀 _ s = isContrΣ (≤-src s , isProp≤ _ _ _) (λ _ → tt* , λ _ → refl) 36 | hom-fib is∞cat𝕀 = is-fib-ext-𝕋Ext 37 | 38 | NatTrans : (X Y : 𝕆Type∞ (𝕋 {ℓ-zero} 0)) → Type 39 | NatTrans X Y = Map tt* (prod∞ X 𝕀∞) Y 40 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Opetopes.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Representables.agda - Representable Opetopic Types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Nat 10 | open import Cubical.Data.Sum 11 | 12 | open import Core.Prelude 13 | open import Experimental.NoDecs.OpetopicType 14 | 15 | module Experimental.NoDecs.Opetopes where 16 | 17 | 18 | 𝕋 : ∀ {ℓ} (n : ℕ) → 𝕆Type n ℓ 19 | 𝕋 zero = tt* 20 | 𝕋 (suc n) = 𝕋 n , λ _ → Lift Unit 21 | 22 | -- 23 | -- Some opetopes 24 | -- 25 | 26 | 𝒪 : ℕ → Type 27 | 𝒪 n = Frm (𝕋 n) 28 | 29 | obj : 𝒪 0 30 | obj = tt* 31 | 32 | arr : 𝒪 1 33 | arr = tt* , tt* , tt* 34 | 35 | drop : 𝒪 2 36 | drop = arr , lf tt* , tt* 37 | 38 | 2-globe : 𝒪 2 39 | 2-globe = arr , nd tt* [ tt* , tt* , lf tt* ] tt* , tt* 40 | 41 | canopy : {n : ℕ} (π : 𝒪 n) (brs : Src (Branch (const Unit*)) π) → Src (const Unit*) π 42 | canopy {n} π brs = μ (id-map (𝕋 n)) (Branch (const Unit*)) (λ _ → Lift Unit) brs (λ p → lvs (brs ⊚ p)) 43 | 44 | -- 45 | -- An Induction Principle 46 | -- 47 | 48 | record IndData {ℓ} (P : (n : ℕ) (π : 𝒪 n) → Type ℓ) : Type ℓ where 49 | field 50 | obj* : P 0 tt* 51 | arr* : P 1 (tt* , tt* , tt*) 52 | glob* : {n : ℕ} (π : 𝒪 n) → P (suc (suc n)) ((π , _ , tt*) , lf tt* , tt*) 53 | ngon* : (χ : Pd (const Unit*) arr) → P 2 (arr , nd tt* [ tt* , tt* , χ ] tt* , tt*) 54 | glob-on-drop* : {n : ℕ} (π : 𝒪 n) → P (3 + n) (((π , _ , tt*) , lf tt* , tt*) , nd tt* (lf tt*) tt* , tt*) 55 | -- root-exposed* : {n : ℕ} (π : 𝒪 n) (hbrs : Src (Branch {X = 𝕋 n} {P = snd (𝕋 (suc n))} (Branch {X = 𝕋 (suc n)} {P = snd (𝕋 (suc (suc n)))} (const Unit*))) π) 56 | -- → P (3 + n) (((π , _ , tt*) , _ , tt*) , nd tt* (nd tt* hbrs [ tt* , η (snd (fst (𝕋 (suc (suc (suc n)))))) tt* , lf tt* ]) tt* , tt*) 57 | 58 | climbing* : {n : ℕ} (π : 𝒪 n) (hbrs : Src (Branch {X = 𝕋 n} {P = snd (𝕋 (suc n))} (Branch {X = 𝕋 (suc n)} {P = snd (𝕋 (suc (suc n)))} (const Unit*))) π) 59 | → (brs : Pd (Branch {X = 𝕋 (suc n)} {P = snd (𝕋 (suc (suc n)))} (const Unit*)) (π , μ (id-map (𝕋 n)) (Branch (Branch (const Unit*))) 60 | (const Unit*) hbrs (λ p → η (const Unit*) {f = (Typ (Branch (Branch (const Unit*))) hbrs p)} (stm (hbrs ⊚ p))) , tt*)) 61 | → (glue : P (3 + n) (((π , μ (id-map (𝕋 n)) (Branch (Branch (const Unit*))) (const Unit*) hbrs 62 | (λ p → η (const Unit*) {f = (Typ (Branch (Branch (const Unit*))) hbrs p)} (stm (hbrs ⊚ p))) , tt*) , _ , tt*) , nd tt* {!!} tt* , tt*)) 63 | → P (3 + n) (((π , _ , tt*) , _ , tt*) , nd tt* (nd tt* hbrs [ tt* , _ , nd tt* brs tt* ]) tt* , tt*) 64 | 65 | 66 | -- Pd (Branch (λ _ → Lift Unit)) (π , _fst_473 , tt*) 67 | 68 | open IndData 69 | 70 | -- Can you say exactly what is the point of this principle? I mean, 71 | -- I guess it's just a way of avoiding (but we'll see about this) 72 | -- direct interaction with μ , η , γ. It expresses the opetopes 73 | -- entirely in terms of some combinators on previous shapes... 74 | 75 | -- Yeah, so I guess what this is actually doing is exposing enough 76 | -- of the structure so that you only have to show compatibility of 77 | -- your induction hypothesis with γ in one particular case. 78 | -- Compatibility with η and μ is taken care of by exposing the 79 | -- cases correctly. 80 | 81 | opetopic-induction : ∀ {ℓ} (P : (n : ℕ) (π : 𝒪 n) → Type ℓ) 82 | → (ihs : IndData P) 83 | → (n : ℕ) (π : 𝒪 n) → P n π 84 | opetopic-induction P ihs zero π = obj* ihs 85 | opetopic-induction P ihs (suc zero) (tt* , tt* , tt*) = arr* ihs 86 | opetopic-induction P ihs (suc (suc n)) ((π , ._ , tt*) , lf .tt* , tt*) = glob* ihs π 87 | opetopic-induction P ihs (suc (suc zero)) ((tt* , ._ , tt*) , nd .tt* [ tt* , tt* , χ ] tt* , tt*) = ngon* ihs χ 88 | opetopic-induction P ihs (suc (suc (suc n))) ((._ , ._ , tt*) , nd .tt* (lf {π} tt*) tt* , tt*) = glob-on-drop* ihs π 89 | opetopic-induction P ihs (suc (suc (suc n))) ((._ , ._ , tt*) , nd .tt* (nd {π} tt* hbrs [ tt* , ._ , lf .tt* ]) tt* , tt*) = {!hbrs!} 90 | -- opetopic-induction P ihs (suc (suc (suc n))) ((._ , ._ , tt*) , nd .tt* (nd {π} tt* hbrs [ tt* , ._ , nd .tt* brs tt* ]) tt* , tt*) = {!!} 91 | opetopic-induction P ihs (suc (suc (suc zero))) ((._ , ._ , tt*) , nd .tt* (nd {π} tt* hbrs [ tt* , ._ , nd .tt* brs tt* ]) tt* , tt*) = {!hbrs!} 92 | opetopic-induction P ihs (suc (suc (suc (suc n)))) ((._ , ._ , tt*) , nd .tt* (nd {π} tt* hbrs [ tt* , ._ , nd .tt* brs tt* ]) tt* , tt*) = {!hbrs!} 93 | 94 | 95 | -- Okay, I think I see. There's one more case: it's the case where 96 | -- the root box in codim 2 is a loop. After this, I feel like you 97 | -- will have exposed enough that you can just directly graft. 98 | 99 | -- So I can't actually match any more. Hmm. But does that mean 100 | -- there is an upper limit to what you can match on for opetopes? 101 | -- That woule be surprising, right? 102 | 103 | -- -- Dimension 1 104 | -- max-frm zero π = tt* 105 | 106 | -- -- Drops (this case is generic. do it first for better computation) 107 | -- max-frm (suc n) ((π , ._ , tt*) , lf .tt* , tt*) = 108 | -- max-frm n (π , _ , tt*) , 109 | -- (η _ {f = max-frm n (π , _ , tt*)} (lf-cell π)) , 110 | -- (lf-cell π) 111 | 112 | -- -- Dimension 2 - ngons for n > 0 113 | -- max-frm (suc zero) ((π , ._ , tt*) , nd .tt* vbr tt* , tt*) = 114 | -- tt* , nd-cell-there tt* vbr tt* (max-frm (suc zero) (_ , br vbr , tt*) .snd .fst) , 115 | -- nd-cell-here tt* vbr 116 | 117 | -- -- Globs on Drops 118 | -- max-frm (suc (suc n)) ((._ , ._ , tt*) , nd .tt* (lf tt*) tt* , tt*) = 119 | -- _ , lf (lf-cell _) , (nd-cell-here _ (lf tt*)) 120 | 121 | -- -- Dimension ≥ 3 - root of source tree is external 122 | -- max-frm (suc (suc n)) ((._ , ._ , tt*) , nd .tt* (nd tt* hbrs [ tt* , ._ , lf .tt* ]) tt* , tt*) = 123 | -- {!!} , {!!} , {!!} 124 | 125 | -- -- Dimension ≥ 3 - climbing the root box 126 | -- max-frm (suc (suc n)) ((._ , ._ , tt*) , nd .tt* (nd tt* hbrs [ tt* , ._ , nd .tt* brs tt* ]) tt* , tt*) = {!!} 127 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Opposite.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Opposite.agda - Opetopic types with 1-arrows reversed 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Nat 10 | open import Cubical.Data.Sum 11 | open import Cubical.Data.Bool 12 | 13 | open import Core.Prelude 14 | open import Experimental.NoDecs.OpetopicType 15 | open import Experimental.NoDecs.Structures 16 | 17 | module Experimental.NoDecs.Opposite where 18 | 19 | Op : ∀ {n ℓ} → 𝕆Type n ℓ → 𝕆Type n ℓ 20 | 21 | postulate 22 | Op-involution : ∀ {n ℓ} (X : 𝕆Type n ℓ) 23 | → Op (Op X) ↦ X 24 | {-# REWRITE Op-involution #-} 25 | 26 | Frm-Op : ∀ {n ℓ} {X : 𝕆Type n ℓ} → Frm X → Frm (Op X) 27 | Src-Op : ∀ {n ℓ} {X : 𝕆Type n ℓ} (P : Frm X → Type ℓ) {f : Frm X} → Src P f → Src (λ fᵒᵖ → P (Frm-Op fᵒᵖ)) (Frm-Op f) 28 | 29 | postulate 30 | Frm-Op-involution : ∀ {n ℓ} {X : 𝕆Type n ℓ} (f : Frm X) 31 | → Frm-Op (Frm-Op f) ↦ f 32 | {-# REWRITE Frm-Op-involution #-} 33 | 34 | Op {zero} X = tt* 35 | Op {suc n} (X , P) = Op X , λ fᵒᵖ → P (Frm-Op fᵒᵖ) 36 | 37 | Frm-Op {zero} f = tt* 38 | Frm-Op {suc zero} {X = X , P} (f , s , t) = tt* , t , s 39 | Frm-Op {suc (suc n)} {X = X , P} (f , s , t) = Frm-Op f , Src-Op P s , t 40 | 41 | {- 42 | Src-Op {zero} {ℓ} {X} P {f} x = x 43 | Src-Op {suc zero} {ℓ} {X} P {f , s , t} x = {!!} 44 | Src-Op {suc (suc n)} {ℓ} {X} P {f} x = {!!} 45 | -} 46 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Representables.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Representables.agda - Representable Opetopic Types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Nat 10 | open import Cubical.Data.Sum 11 | 12 | open import Core.Prelude 13 | open import Experimental.NoDecs.OpetopicType 14 | 15 | module Experimental.NoDecs.Representables where 16 | 17 | 𝕋 : ∀ {ℓ} (n : ℕ) → 𝕆Type n ℓ 18 | 𝕋 zero = tt* 19 | 𝕋 (suc n) = 𝕋 n , λ _ → Lift Unit 20 | 21 | -- 22 | -- Some opetopes 23 | -- 24 | 25 | 𝒪 : ℕ → Type 26 | 𝒪 n = Frm (𝕋 n) 27 | 28 | obj : 𝒪 0 29 | obj = tt* 30 | 31 | arr : 𝒪 1 32 | arr = tt* , tt* , tt* 33 | 34 | drop : 𝒪 2 35 | drop = arr , lf tt* , tt* 36 | 37 | 2-globe : 𝒪 2 38 | 2-globe = arr , nd tt* [ tt* , tt* , lf tt* ] tt* , tt* 39 | 40 | canopy : {n : ℕ} (π : 𝒪 n) (brs : Src (Branch (const Unit*)) π) → Src (const Unit*) π 41 | canopy {n} π brs = μ (id-map (𝕋 n)) (Branch (const Unit*)) (λ _ → Lift Unit) brs (λ p → lvs (brs ⊚ p)) 42 | 43 | -- 44 | -- The codimension 2 part of the representable 45 | -- 46 | 47 | Repr₀ : (n : ℕ) → 𝒪 (suc n) → 𝕆Type n ℓ-zero 48 | max-frm : (n : ℕ) (π : 𝒪 (suc n)) → Frm (Repr₀ n π) 49 | 50 | data Face₀ : {n : ℕ} (π : 𝒪 n) 51 | (σ : Src (const Unit*) π) 52 | (τ : Pd (const Unit*) (π , σ , tt*)) 53 | → Frm (Repr₀ n (π , σ , tt*)) → Type where 54 | 55 | lf-cell : {n : ℕ} (π : 𝒪 n) 56 | → Face₀ π (η (const Unit*) {f = π} tt*) (lf tt*) 57 | (max-frm n (π , η (const Unit*) {f = π} tt* , tt*)) 58 | 59 | nd-cell-there : {n : ℕ} (π : 𝒪 n) (brs : Src (Branch (const Unit*)) π) 60 | → (p : Pos (Branch (const Unit*)) {f = π} brs) 61 | → Face₀ (Typ (Branch (const Unit*)) brs p) (lvs (brs ⊚ p)) (br (brs ⊚ p)) 62 | (max-frm n (Typ (Branch (const Unit*)) brs p , lvs (brs ⊚ p) , tt*)) 63 | → Face₀ π (canopy π brs) (nd tt* brs tt*) 64 | (max-frm n (π , canopy π brs , tt*)) 65 | 66 | nd-cell-here : {n : ℕ} (π : 𝒪 n) (brs : Src (Branch (const Unit*)) π) 67 | → Face₀ π (canopy π brs) (nd tt* brs tt*) 68 | (max-frm n (π , canopy π brs , tt*)) 69 | 70 | Repr₀ zero _ = tt* 71 | Repr₀ (suc n) ((π , σ , tt*) , τ , tt*) = 72 | Repr₀ n (π , σ , tt*) , Face₀ π σ τ 73 | 74 | 75 | -- Dimension 1 76 | max-frm zero π = tt* 77 | 78 | -- Drops (this case is generic. do it first for better computation) 79 | max-frm (suc n) ((π , ._ , tt*) , lf .tt* , tt*) = 80 | max-frm n (π , _ , tt*) , 81 | (η _ {f = max-frm n (π , _ , tt*)} (lf-cell π)) , 82 | (lf-cell π) 83 | 84 | -- Dimension 2 - ngons for n > 0 85 | max-frm (suc zero) ((π , ._ , tt*) , nd .tt* vbr tt* , tt*) = 86 | tt* , nd-cell-there tt* vbr tt* (max-frm (suc zero) (_ , br vbr , tt*) .snd .fst) , 87 | nd-cell-here tt* vbr 88 | 89 | -- Globs on Drops 90 | max-frm (suc (suc n)) ((._ , ._ , tt*) , nd .tt* (lf tt*) tt* , tt*) = 91 | _ , lf (lf-cell _) , (nd-cell-here _ (lf tt*)) 92 | 93 | -- Dimension ≥ 3 - root of source tree is external 94 | max-frm (suc (suc n)) ((._ , ._ , tt*) , nd .tt* (nd tt* hbrs [ tt* , ._ , lf .tt* ]) tt* , tt*) = 95 | {!!} , {!!} , {!!} 96 | 97 | -- Dimension ≥ 3 - climbing the root box 98 | max-frm (suc (suc n)) ((._ , ._ , tt*) , nd .tt* (nd tt* hbrs [ tt* , ._ , nd .tt* brs tt* ]) tt* , tt*) = {!!} 99 | 100 | 101 | 102 | -- brs : Pd (Branch (λ _ → Lift Unit)) 103 | -- (f , 104 | -- μ (id-map (𝕋 n)) (Branch (Branch (λ _ → Lift Unit))) 105 | -- (λ _ → Lift Unit) hbrs (λ p → η (λ _ → Lift Unit) (stm (hbrs ⊚ p))) 106 | -- , tt*) 107 | -- hbrs : Src (Branch (Branch (λ _ → Lift Unit))) f 108 | -- f : Frm (𝕋 n) (not in scope) 109 | -- n : ℕ 110 | 111 | 112 | -- -- Dimension 2 - The output cell is always the nd-cell-here 113 | -- -- constructor, and for the source we recurse, inserting 114 | -- -- the nd-cell-there constructor to pass up the chain. 115 | -- max-frm (suc zero) ((tt* , tt* , tt*) , lf .tt* , tt*) = 116 | -- tt* , lf-cell tt* , lf-cell tt* 117 | -- max-frm (suc zero) ((tt* , tt* , tt*) , nd .tt* vbr tt* , tt*) = 118 | -- tt* , nd-cell-there tt* vbr tt* (max-frm (suc zero) (_ , br vbr , tt*) .snd .fst) , 119 | -- nd-cell-here tt* vbr 120 | 121 | -- -- Dimension ≥ 3 - A Drop 122 | -- max-frm (suc (suc n)) (((π , ρ , tt*) , ._ , tt*) , lf .tt* , tt*) = 123 | -- max-frm (suc n) ((π , ρ , tt*) , η (const Unit*) {f = (π , ρ , tt*)} tt* , tt*) , 124 | -- η _ {f = max-frm (suc n) ((π , ρ , tt*) , η (const Unit*) {f = (π , ρ , tt*)} tt* , tt*)} (lf-cell (π , ρ , tt*)) , 125 | -- lf-cell (π , ρ , tt*) 126 | 127 | -- -- Dimension ≥ 3 - A Glob on a Drop 128 | -- max-frm (suc (suc n)) (((π , ._ , tt*) , ._ , tt*) , nd .tt* (lf .tt*) tt* , tt*) = 129 | -- (max-frm n (π , η (const Unit*) {f = π} tt* , tt*) , η _ {f = max-frm n (π , η (const Unit*) {f = π} tt* , tt*)} (lf-cell π) , lf-cell π) , 130 | -- lf (lf-cell π) , 131 | -- {!nd-cell-here (π , η (λ _ → Lift Unit) tt* , tt*) (lf tt*)!} 132 | 133 | -- -- max-frm (suc n) ((π , η (const Unit*) {f = π} tt* , tt*) , lf tt* , tt*) , 134 | -- -- {!!} , 135 | -- -- nd-cell-here (π , _ , tt*) (lf tt*) 136 | 137 | -- -- Dimension ≥ 3 - Root of desired tree has been reduced to a single node. Append a node to the result of hbrs. 138 | -- max-frm (suc (suc n)) (((π , ._ , tt*) , ._ , tt*) , nd .tt* (nd .tt* hbrs [ tt* , ._ , lf .tt* ]) tt* , tt*) = 139 | -- {!!} , {!!} , {!!} 140 | 141 | -- -- Dimension ≥ 3 - Still climbing the base box. Graft and continue. 142 | -- max-frm (suc (suc n)) (((π , ._ , tt*) , ._ , tt*) , nd .tt* (nd .tt* hbrs [ tt* , ._ , nd .tt* lclhbrs tt* ]) tt* , tt*) = 143 | -- {!!} , {!!} , {!!} 144 | 145 | 146 | 147 | -- -- Dimension 3 - A drop on an arrow 148 | -- max-frm (suc (suc zero)) (((tt* , tt* , tt*) , ._ , tt*) , lf .tt* , tt*) = 149 | -- max-frm (suc zero) 2-globe , η _ {f = max-frm (suc zero) 2-globe} (lf-cell arr) , lf-cell arr 150 | 151 | -- -- Dimension 3 - A globe on a drop 152 | -- max-frm (suc (suc zero)) (((tt* , tt* , tt*) , ._ , tt*) , nd .tt* (lf .tt*) tt* , tt*) = {!!} 153 | 154 | -- -- Dimension 3 - Root is now a single box 155 | -- max-frm (suc (suc zero)) (((tt* , tt* , tt*) , ._ , tt*) , nd .tt* (nd .tt* [ tt* , .tt* , horiz ] [ tt* , ._ , lf .tt* ]) tt* , tt*) = 156 | -- max-frm 1 (arr , canopy (tt* , tt* , tt*) (nd tt* [ tt* , tt* , horiz ] [ tt* , η (snd (fst (𝕋 3))) tt* , lf tt* ]) , tt*) , 157 | -- nd _ [ _ , _ , {!horiz-ih .snd .fst!} ] (nd-cell-there {!!} {!horiz-ih .snd .fst!} {!!} {!!}) , 158 | -- nd-cell-here arr (nd tt* [ tt* , tt* , horiz ] [ tt* , η (snd (fst (𝕋 3))) tt* , lf tt* ]) 159 | 160 | -- where horiz-ih : Frm (Repr₀ (suc (suc zero)) ((arr , canopy arr horiz , tt*) , nd tt* horiz tt* , tt*)) 161 | -- horiz-ih = max-frm (suc (suc zero)) ((arr , canopy arr horiz , tt*) , nd tt* horiz tt* , tt*) 162 | 163 | -- max-frm (suc (suc zero)) (((tt* , tt* , tt*) , ._ , tt*) , nd .tt* (nd .tt* [ tt* , .tt* , horiz ] [ tt* , ._ , nd .tt* hbrs tt* ]) tt* , tt*) = {!!} 164 | 165 | -- Dimension 3 - Non-trivial outgoing branch 166 | -- 167 | -- horiz - a list containing the remaining vertical brances 168 | -- vert - the current vertical branch 169 | -- 170 | -- max-frm (suc (suc zero)) (((tt* , tt* , tt*) , ._ , tt*) , nd .tt* (nd .tt* [ tt* , .tt* , horiz ] vert) tt* , tt*) = 171 | -- max-frm (suc zero) (arr , canopy arr (nd tt* [ tt* , tt* , horiz ] vert) , tt*) , src-pd , tgt-cell 172 | 173 | -- where tgt-cell : Face₀ arr (canopy arr (nd tt* [ tt* , tt* , horiz ] vert)) 174 | -- (nd tt* (nd tt* [ tt* , tt* , horiz ] vert) tt*) 175 | -- (max-frm (suc zero) (arr , canopy arr (nd tt* [ tt* , tt* , horiz ] vert) , tt*)) 176 | -- tgt-cell = nd-cell-here arr (nd tt* [ tt* , tt* , horiz ] vert) 177 | 178 | -- vert-ih : Frm (Repr₀ (suc (suc zero)) ((arr , lvs vert , tt*) , br vert , tt*)) 179 | -- vert-ih = max-frm (suc (suc zero)) ((arr , lvs vert , tt*) , br vert , tt*) 180 | 181 | -- horiz-ih : Frm (Repr₀ (suc (suc zero)) ((arr , canopy arr horiz , tt*) , nd tt* horiz tt* , tt*)) 182 | -- horiz-ih = max-frm (suc (suc zero)) ((arr , canopy arr horiz , tt*) , nd tt* horiz tt* , tt*) 183 | 184 | -- vert-src : Pd (Face₀ arr (lvs vert) (br vert)) (fst vert-ih) 185 | -- vert-src = vert-ih .snd .fst 186 | 187 | -- horiz-src : Pd (Face₀ arr (canopy arr horiz) (nd tt* horiz tt*)) (fst horiz-ih) 188 | -- horiz-src = horiz-ih .snd .fst 189 | 190 | -- -- src-pd : Pd (Face₀ arr (canopy arr (nd tt* [ tt* , tt* , horiz ] vert)) 191 | -- -- (nd tt* (nd tt* [ tt* , tt* , horiz ] vert) tt*)) 192 | -- -- (max-frm 1 (arr , canopy arr (nd tt* [ tt* , tt* , horiz ] vert) , tt*)) 193 | 194 | -- src-pd : Pd (Face₀ arr (γ (const Unit*) (lvs vert) 195 | -- (λ p → map-src tt* (const Unit*) (const Unit*) tt* (const tt*) , 196 | -- μ (tt* , (λ q → q)) (Branch (const Unit*)) (const Unit*) horiz (λ q → lvs (PdInhab (Branch (const Unit*)) horiz q)))) 197 | -- (nd tt* (nd tt* [ tt* , tt* , horiz ] vert) tt*)) 198 | -- (max-frm 1 (arr , canopy arr (nd tt* [ tt* , tt* , horiz ] vert) , tt*)) 199 | 200 | -- src-pd = {!!} 201 | 202 | 203 | -- 204 | -- Next dimension 205 | -- 206 | 207 | -- Repr₁ : (n : ℕ) (π : 𝒪 (suc n)) → Frm (Repr₀ n π) → Type 208 | 209 | -- data Face₁ : {n : ℕ} (π : 𝒪 n) 210 | -- (σ : Src (const Unit*) π) 211 | -- (τ : Pd (const Unit*) (π , σ , tt*)) 212 | -- → Frm (Repr₀ n (π , σ , tt*) , Face₀ π σ τ) → Type where 213 | 214 | -- nd-cell : {n : ℕ} (π : 𝒪 n) (brs : Src (Branch (const Unit*)) π) 215 | -- → Face₁ π (canopy π brs) (nd tt* brs tt*) 216 | -- (max-frm n (π , canopy π brs , tt*) , 217 | -- -- BUG! No, I don't think this is right. It should 218 | -- -- be the *local* canopy in this case. 219 | -- max-frm-canopy n π brs , 220 | -- (nd-cell π brs)) 221 | 222 | -- tgt-cell : {n : ℕ} (π : 𝒪 n) (σ : Src (const Unit*) π) (τ : Pd (const Unit*) (π , σ , tt*)) 223 | -- → Face₁ π σ τ (max-frm (suc n) ((π , σ , tt*) , τ , tt*)) 224 | 225 | -- Repr₁ zero π _ = Unit ⊎ Unit 226 | -- Repr₁ (suc n) ((π , σ , tt*) , τ , tt*) = Face₁ π σ τ 227 | 228 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Revised/Category.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Category.agda - Equivalence between categories and 1-truncated (∞,1)-categories 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Nat 7 | open import Cubical.Data.Sigma 8 | open import Cubical.Categories.Category 9 | 10 | open import Core.Prelude 11 | open import Experimental.NoDecs.Revised.OpetopicType 12 | open import Experimental.NoDecs.Revised.Shapes 13 | open import Experimental.NoDecs.Revised.Structures 14 | 15 | module Experimental.NoDecs.Revised.Category where 16 | 17 | module _ {ℓ} (C : Category ℓ ℓ) where 18 | lemma-that-should-be-refl : ∀ {n ℓ₀} {X : 𝕆Type n ℓ₀} 19 | → (P : Frm X → Type ℓ₀) 20 | → (Q : Frm X → Type ℓ₀) 21 | → (R : Frm X → Type ℓ₀) 22 | → {f : Frm X} (s : Src P f) 23 | → (ϕ : (p : Pos P s) → Q (Typ P s p)) 24 | → (ψ : (q : Pos Q (ν s (id-map X) ϕ)) → R (Typ Q (ν s (id-map X) ϕ) q)) 25 | → ν {Q = R} (ν s (id-map X) ϕ) (id-map X) ψ ≡ ν s (id-map X) (λ p → ψ (ν-pos s (id-map X) ϕ p)) 26 | lemma-that-should-be-refl P Q R s ϕ ψ = refl 27 | 28 | open Category C renaming (id to C-id ; _⋆_ to _⨀_) 29 | 30 | n-comp : {n : ℕ} {t : ob ** (suc n)} → path-chain Hom[_,_] n t → Hom[ last t , fstt t ] 31 | n-comp {zero} {x} _ = C-id 32 | n-comp {suc zero} {y , x} f = f 33 | n-comp {suc (suc n)} {y , x , t} (f , l) = n-comp {suc n} {x , t} l ⨀ f 34 | 35 | private 36 | X₀ : Frm (𝕋 {ℓ = ℓ} 0) → Type ℓ 37 | X₀ _ = ob 38 | 39 | X₁ : Frm (tt* , X₀) → Type ℓ 40 | X₁ (_ , x , y) = Hom[ x , y ] 41 | 42 | C-src-comp : {f : Frm (tt* , X₀)} → Src X₁ f → X₁ f 43 | C-src-comp (lf x) = C-id {x = x} 44 | C-src-comp (nd tgt [ stm₁ , lvs₁ , br₁ ] flr) = (C-src-comp br₁) ⨀ flr 45 | 46 | X₂ : Frm ((tt* , X₀) , X₁) → Type ℓ 47 | X₂ (f , s , t) = C-src-comp s ≡ t 48 | 49 | -- lemma relating C-src-comp to μ 50 | {-# TERMINATING #-} 51 | big-lemma : {f : Frm (tt* , X₀)} (brs : Src (Src X₁) f) → 52 | C-src-comp (μ X₁ brs) ≡ 53 | C-src-comp (ν {Q = X₁} brs (id-map _) λ p → C-src-comp (brs ⊚ p)) 54 | big-lemma (lf tgt) = refl 55 | big-lemma (nd _ brs (lf _)) = big-lemma (br brs) ∙ sym (⋆IdR _) 56 | big-lemma (nd tgt brs (nd _ [ stm₁ , _ , br₁ ] flr)) = (cong (_⨀ flr) (big-lemma (nd stm₁ brs br₁))) ∙ ⋆Assoc _ (C-src-comp br₁) flr 57 | 58 | {-# TERMINATING #-} 59 | C-2-src-comp : {f : Frm ((tt* , X₀) , X₁)} → Src X₂ f → X₂ f 60 | C-2-src-comp (lf tgt) = ⋆IdL tgt 61 | C-2-src-comp (nd tgt brs flr) = 62 | big-lemma (ν brs (id-map (tt* , X₀)) (λ p → lvs (brs ⊚ p))) ∙ 63 | {!cong C-src-comp lemma0!} ∙ IH ∙ flr where -- This last hole should be refl if the rewrite rules worked correctly 64 | IH : 65 | C-src-comp (ν brs (id-map (tt* , X₀)) λ p → C-src-comp (lvs (brs ⊚ p))) ≡ 66 | C-src-comp (understory X₂ brs) 67 | IH i = C-src-comp (ν brs (id-map (tt* , X₀)) (λ p → C-2-src-comp (br (brs ⊚ p)) i)) 68 | 69 | {-aaaah = lemma-that-should-be-refl (Branch X₂) (Src X₁) X₁ brs (λ p → lvs (brs ⊚ p)) (λ p → 70 | C-src-comp 71 | (ν brs (tt* , (λ p₁ → p₁)) 72 | (λ p₁ → lvs (brs ⊚ p₁)) 73 | ⊚ p)) 74 | 75 | lemma0 : 76 | ν (ν {Q = Src X₁} brs (id-map _) (λ p → lvs (brs ⊚ p))) 77 | (id-map (tt* , X₀)) 78 | (λ p → 79 | C-src-comp 80 | (ν brs (id-map (tt* , X₀)) (λ p₁ → lvs (brs ⊚ p₁)) ⊚ p)) 81 | ≡ 82 | ν brs (id-map (tt* , X₀)) λ p → C-src-comp (lvs (brs ⊚ p)) 83 | lemma0 = aaaah ∙ {!!} where 84 | {-test : (p : Pos (Branch X₂) brs) → 85 | (ν {Q = Src X₁} brs (tt* , (λ p₁ → p₁)) 86 | (λ p₁ → lvs (brs ⊚ p₁))) ⊚ 87 | (ν-pos brs (tt* , (λ p₁ → p₁)) 88 | (λ p₁ → lvs (brs ⊚ p₁)) p) ≡ 89 | (lvs (brs ⊚ p)) 90 | test p = {!!}-} 91 | 92 | -} 93 | 94 | 95 | Cat→𝕆Type : 𝕆Type∞ {ℓ = ℓ} tt* 96 | Fill Cat→𝕆Type = X₀ 97 | Fill (Hom Cat→𝕆Type) = X₁ 98 | Fill (Hom (Hom Cat→𝕆Type)) = X₂ 99 | Hom (Hom (Hom Cat→𝕆Type)) = 𝕋Ext _ 100 | 101 | open is-fibrant-ext 102 | 103 | lemma : ∀ {ℓ} {A : Type ℓ} {a : A} → isContr (Σ[ b ∈ A ] a ≡ b) 104 | lemma {a = a} = (a , refl) , (λ (b , p) i → (p i , λ j → p (i ∧ j))) 105 | 106 | module _ (isGroupoidC : isGroupoid ob) where 107 | isInfCatCat : is-fibrant-ext (Hom Cat→𝕆Type) 108 | fill-fib isInfCatCat f s = lemma 109 | hom-fib (hom-fib isInfCatCat) = is-fib-ext-𝕋Ext 110 | fill-fib (hom-fib isInfCatCat) (f , s , t) s' = (C-2-src-comp s' , tt*) , λ y → Σ≡Prop (λ _ _ _ → refl) (isSetHom _ _ _ _) 111 | 112 | 113 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Revised/Shapes.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --no-positivity-check #-} 2 | -- 3 | -- Shapes.agda - General shapes in opetopic types 4 | -- 5 | 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Sigma 9 | 10 | open import Experimental.NoDecs.Revised.OpetopicType 11 | 12 | module Experimental.NoDecs.Revised.Shapes where 13 | globe-Frm : ∀ {n ℓ} {Xₙ : 𝕆Type n ℓ} (Xₛₙ : Frm Xₙ → Type ℓ) {f : Frm Xₙ} (x y : Xₛₙ f) → Frm (Xₙ , Xₛₙ) 14 | globe-Frm Xₛₙ {f} x y = f , η Xₛₙ x , y 15 | 16 | _**_ : ∀ {ℓ} → Type ℓ → ℕ → Type ℓ 17 | X ** zero = Lift Unit 18 | X ** suc zero = X 19 | X ** suc (suc n) = X × (X ** (suc n)) 20 | 21 | fstt : ∀ {ℓ n} {X : Type ℓ} → X ** (suc n) → X 22 | fstt {n = zero} x = x 23 | fstt {n = suc n} (x , _) = x 24 | 25 | last : ∀ {ℓ n} {X : Type ℓ} → X ** (suc n) → X 26 | last {n = zero} x = x 27 | last {n = suc n} (fst , snd) = last snd 28 | 29 | cons : ∀ {n ℓ} {X : Type ℓ} → X → X ** n → X ** (suc n) 30 | cons {zero} x l = x 31 | cons {suc n} x l = x , l 32 | 33 | path-chain : ∀ {ℓ} {X : Type ℓ} (Y : X → X → Type ℓ) (n : ℕ) (t : X ** (suc n)) → Type ℓ 34 | path-chain Y zero _ = Lift Unit 35 | path-chain Y (suc zero) (y , x) = Y x y 36 | path-chain Y (suc (suc n)) (x , t) = Y (fstt t) x × path-chain Y (suc n) t 37 | 38 | chain-cons : ∀ {ℓ} {X : Type ℓ} (Y : X → X → Type ℓ) {n : ℕ} {y : X} {t : X ** (suc n)} (p : path-chain Y n t) → Y (fstt t) y → path-chain Y (suc n) (y , t) 39 | chain-cons Y {zero} {y} {t} _ f = f 40 | chain-cons Y {suc n} {y} {t} p f = f , p 41 | 42 | -- Sequences of arrows 43 | n-path : ∀ {ℓ} {X : 𝕆Type 2 ℓ} (n : ℕ) {t : (X .fst .snd tt*) ** suc n} 44 | → path-chain (λ x y → X .snd (tt* , x , y)) n t 45 | → Src (X .snd) (tt* , last t , fstt t) 46 | n-path zero {x} f = lf x 47 | n-path (suc zero) {y , x} f = nd y [ x , x , lf x ] f 48 | n-path (suc (suc n)) {y , x , t} (f , l) = nd y [ x , last t , (n-path (suc n) l) ] f 49 | 50 | -- Sequences of unary higher cells 51 | n-path' : ∀ {ℓ} {m : ℕ} {X : 𝕆Type m ℓ} (n : ℕ) {P : Frm X → Type ℓ} (Q : Frm (X , P) → Type ℓ) {f : Frm X} 52 | {t : P f ** (suc n)} 53 | → path-chain (λ x y → Q (globe-Frm P x y)) n t 54 | → Src Q (globe-Frm P (last t) (fstt t)) 55 | n-path' zero {P} Q {f} {x} _ = lf x 56 | n-path' (suc zero) {P} Q {f} {y , x} p = nd y (η (Branch Q) [ x , _ , (lf x) ]) p 57 | n-path' (suc (suc n)) {P} Q {f} {y , x , t} (p , l) = nd y (η (Branch Q) [ x , (η P (last t)) , (n-path' (suc n) Q l) ]) p 58 | 59 | 60 | module _ {ℓ} (X₀ : 𝕆Type 1 ℓ) where 61 | Obj : Type ℓ 62 | Obj = snd X₀ tt* 63 | 64 | hom-Frm : Obj → Obj → Frm X₀ 65 | hom-Frm x y = (tt* , x , y) 66 | 67 | module _ (X₁ : Frm X₀ → Type ℓ) where 68 | hom : Obj → Obj → Type ℓ 69 | hom x y = X₁ (hom-Frm x y) 70 | 71 | simplex-Frm : {x y z : Obj} (f : hom x y) (g : hom y z) (h : hom x z) → Frm (X₀ , X₁) 72 | simplex-Frm {x} {y} {z} f g h = hom-Frm x z , n-path 2 (g , f) , h -- nd z [ y , x , nd y [ x , x , lf x ] f ] g 73 | 74 | 2-drop-Frm : (x : Obj) (f : hom x x) → Frm (X₀ , X₁) 75 | 2-drop-Frm x f = hom-Frm x x , lf x , f 76 | 77 | {- 78 | src→chain : {f : Frm X₀} → Src X₁ f → Σ[ n ∈ ℕ ] Σ[ t ∈ (Obj ** n) ] (path-chain hom n (cons (snd (snd f)) t)) 79 | src→chain (lf tgt) = 0 , tt* , tt* 80 | src→chain (nd tgt [ stm₁ , .(η (λ f → snd X₀ tt*) stm₁) , _ ] flr) = ? 81 | src→chain (nd tgt [ stm₁ , .(μ (id-map tt*) (Branch X₁) (λ f → snd X₀ tt*) brs (λ p → lvs (brs ⊚ p))) , nd .stm₁ brs flr₁ ] flr) = ? 82 | -} 83 | 84 | module _ (X₂ : Frm (X₀ , X₁) → Type ℓ) where 85 | Null : (x : Obj) (f : hom x x) → Type ℓ 86 | Null x f = X₂ (2-drop-Frm x f) 87 | 88 | 2Glob : {x y : Obj} (f g : hom x y) → Type ℓ 89 | 2Glob {x} {y} f g = X₂ ((tt* , x , y) , nd y [ x , x , lf x ] f , g) 90 | 91 | Simplex : {x y z : Obj} (f : hom x y) (g : hom y z) 92 | → (h : hom x z) → Type ℓ 93 | Simplex {x} {y} {z} f g h = X₂ (simplex-Frm f g h) 94 | 95 | unitor-Frm : (x y : Obj) (f : hom x x) (g : hom x y) (h : hom x y) 96 | → (Δ : X₂ (simplex-Frm f g h)) 97 | → (Γ : X₂ (2-drop-Frm x f)) 98 | → (O : X₂ (globe-Frm X₁ g h)) 99 | → Frm ((X₀ , X₁) , X₂) 100 | unitor-Frm x y f g h Δ Γ O = _ , nd h (nd y [ x , x , nd x [ x , η _ x , lf x ] [ f , lf x , nd f (lf x) Γ ] ] [ g , _ , lf g ]) Δ , O 101 | 102 | associator1 : (x y z t : Obj) (f : hom x y) (g : hom y z) (h : hom z t) (i : hom x z) (j : hom x t) 103 | → (Δ₁ : X₂ (simplex-Frm f g i)) 104 | → (Δ₂ : X₂ (simplex-Frm i h j)) 105 | → Src X₂ (hom-Frm x t , n-path 3 (h , g , f) , j) --nd t [ _ , _ , (nd _ [ _ , _ , (nd _ [ _ , _ , (lf _) ] f) ] g) ] h) 106 | associator1 x y z t f g h i j Δ₁ Δ₂ = nd j (nd t [ z , x , (nd z [ x , x , (lf x) ] [ i , nd z [ y , x , (nd _ _ _) ] g , nd i (nd z [ y , x , nd y [ x , x , (lf x) ] [ f , _ , lf f ] ] [ g , _ , (lf g) ]) Δ₁ ]) ] [ h , _ , (lf h) ]) Δ₂ 107 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Revised/Structures.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Everything 2 | 3 | open import Core.Prelude 4 | open import Experimental.NoDecs.Revised.OpetopicType 5 | open import Experimental.NoDecs.Revised.Shapes 6 | 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Sigma 9 | open import Cubical.Foundations.Equiv.Fiberwise 10 | 11 | module Experimental.NoDecs.Revised.Structures where 12 | record 𝕆Type∞ {n ℓ} (Xₙ : 𝕆Type n ℓ) : Type (ℓ-suc ℓ) where 13 | coinductive 14 | field 15 | Fill : Frm Xₙ → Type ℓ 16 | Hom : 𝕆Type∞ (Xₙ , Fill) 17 | open 𝕆Type∞ public 18 | 19 | record Map {n ℓ} {Xₙ Yₙ : 𝕆Type n ℓ} (σ : Xₙ ⇒ Yₙ) (X : 𝕆Type∞ Xₙ) (Y : 𝕆Type∞ Yₙ) : Type ℓ where 20 | coinductive 21 | field 22 | Fill⇒ : {f : Frm Xₙ} → (Fill X) f → (Fill Y) (Frm⇒ σ f) 23 | Hom⇒ : Map (σ , Fill⇒) (Hom X) (Hom Y) 24 | open Map public 25 | 26 | horn-filler : ∀ {n ℓ} {Xₙ : 𝕆Type n ℓ} {Xₛₙ : Frm Xₙ → Type ℓ} (Xₛₛₙ : Frm (Xₙ , Xₛₙ) → Type ℓ) {f : Frm Xₙ} → Src Xₛₙ f → Type ℓ 27 | horn-filler {n} {ℓ} {Xₙ} {Xₛₙ} Xₛₛₙ {f} s = Σ[ tgt ∈ Xₛₙ f ] Xₛₛₙ (f , s , tgt) 28 | 29 | is-fibrant : ∀ {n ℓ} → 𝕆Type (suc (suc n)) ℓ → Type ℓ 30 | is-fibrant ((Xₙ , Xₛₙ) , Xₛₛₙ) = (f : Frm Xₙ) (s : Src Xₛₙ f) → isContr (horn-filler Xₛₛₙ s) 31 | 32 | isProp-is-fibrant : ∀ {n ℓ} {X : 𝕆Type (suc (suc n)) ℓ} → isProp (is-fibrant X) 33 | isProp-is-fibrant = isPropΠ2 (λ _ _ → isPropIsContr) 34 | 35 | record is-fibrant-ext {n ℓ} {Xₙ : 𝕆Type n ℓ} (X : 𝕆Type∞ Xₙ) : Type ℓ where 36 | coinductive 37 | field 38 | fill-fib : is-fibrant ((Xₙ , (Fill X)) , (Fill (Hom X))) 39 | hom-fib : is-fibrant-ext (Hom X) 40 | 41 | open is-fibrant-ext public 42 | 43 | eta-fib-ext : ∀ {m ℓ} {X : 𝕆Type m ℓ} {X∞ : 𝕆Type∞ X} → X∞ ≡ record { Fill = Fill X∞ ; Hom = Hom X∞ } 44 | Fill (eta-fib-ext {X∞ = X∞} i) = Fill X∞ 45 | Hom (eta-fib-ext {X∞ = X∞} i) = Hom X∞ 46 | 47 | isProp-is-fibrant-ext : ∀ {n ℓ} {X : 𝕆Type n ℓ} {X∞ : 𝕆Type∞ X} → isProp (is-fibrant-ext X∞) 48 | fill-fib (isProp-is-fibrant-ext x y i) = isProp-is-fibrant (fill-fib x) (fill-fib y) i 49 | hom-fib (isProp-is-fibrant-ext x y i) = isProp-is-fibrant-ext (hom-fib x) (hom-fib y) i 50 | 51 | 𝕋 : ∀ {ℓ} (n : ℕ) → 𝕆Type n ℓ 52 | 𝕋 zero = tt* 53 | 𝕋 (suc n) = 𝕋 n , λ _ → Lift Unit 54 | 55 | 𝕋Ext : ∀ {n ℓ} (X : 𝕆Type n ℓ) → 𝕆Type∞ X 56 | Fill (𝕋Ext X) = λ _ → Lift Unit 57 | Hom (𝕋Ext X) = 𝕋Ext _ 58 | 59 | is-fib-ext-𝕋Ext : ∀ {n ℓ} {X : 𝕆Type n ℓ} → is-fibrant-ext (𝕋Ext X) 60 | fill-fib is-fib-ext-𝕋Ext f s = (tt* , tt*) , λ (tt* , tt*) → refl 61 | hom-fib is-fib-ext-𝕋Ext = is-fib-ext-𝕋Ext 62 | 63 | 64 | 65 | 66 | 67 | hom→path0 : ∀ {ℓ} {X : 𝕆Type∞ (𝕋 {ℓ} 0)} → is-fibrant-ext X → (x y : X .Fill tt*) → X .Hom .Fill (tt* , x , y) → x ≡ y 68 | hom→path0 {ℓ} {X} fib x y σ = sym (transportRefl x) ∙ cong fst (isContr→isProp (fib .fill-fib tt* x) (Id x x refl) b) where 69 | Id : (x y : X .Fill tt*) → (x ≡ y) → horn-filler (Fill (Hom X)) x 70 | Id x y = J (λ y p → horn-filler (Fill (Hom X)) x) (x , fib .hom-fib .fill-fib (tt* , x , x) (lf x) .fst .fst) 71 | 72 | b : horn-filler (Fill (Hom X)) x 73 | b = y , σ 74 | 75 | Id-cell : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} (t : Fill X∞ f) → Fill (Hom X∞) (globe-Frm _ t t) 76 | Id-cell X∞ fib t = fib .hom-fib .fill-fib _ (lf t) .fst .fst 77 | 78 | path→hom : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} 79 | → t ≡ u → Fill (Hom X∞) (globe-Frm _ t u) 80 | path→hom X∞ fib {t = t} {u = u} = J (λ u p → Fill (Hom X∞) (globe-Frm _ t u)) (Id-cell X∞ fib t) 81 | 82 | hom→path : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} 83 | → Fill (Hom X∞) (globe-Frm _ t u) → t ≡ u 84 | hom→path X∞ fib {f} {t} {u} cell = cong fst (isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (u , cell)) 85 | 86 | module IsoHomPath {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} where 87 | sec : section (path→hom X∞ fib {f} {t} {u}) (hom→path X∞ fib {f} {t} {u}) 88 | sec cell = fromPathP (cong snd (isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (u , cell))) 89 | -- Wow ! Might need to take a look back at this later cause I didn't expect it to be so "simple". 90 | 91 | ret : retract (path→hom X∞ fib {f} {t} {u}) (hom→path X∞ fib {f} {t} {u}) 92 | ret = J (λ u p → hom→path X∞ fib (path→hom X∞ fib p) ≡ p) 93 | (hom→path X∞ fib (path→hom X∞ fib refl) 94 | ≡⟨ cong (hom→path X∞ fib) (transportRefl _) ⟩ 95 | hom→path X∞ fib (Id-cell X∞ fib t) 96 | ≡⟨ (λ i j → fst (test i j)) ⟩ 97 | refl ∎) where 98 | a : horn-filler (Fill (Hom X∞)) (η (Fill X∞) t) 99 | a = (t , Id-cell X∞ fib t) 100 | 101 | test0 : a ≡ a 102 | test0 = isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (t , Id-cell X∞ fib t) 103 | 104 | test : test0 ≡ refl 105 | test = isProp→isSet (isContr→isProp (fib .fill-fib f (η (Fill X∞) t))) (t , Id-cell X∞ fib t) (t , Id-cell X∞ fib t) test0 refl 106 | 107 | record is-n-trunc {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) : Type ℓ where 108 | coinductive 109 | field 110 | hLevel : (f : Frm X) → isOfHLevel n (X∞ .Fill f) 111 | is-trunc-ext : is-n-trunc (predℕ n) (X∞ .Hom) 112 | open is-n-trunc 113 | 114 | src-comp : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) → is-fibrant-ext X∞ → {f : Frm X} → Src (Fill X∞) f → Fill X∞ f 115 | src-comp X∞ fib s = fib .fill-fib _ s .fst .fst 116 | 117 | -- More general version of the equivalence between hom and path, using the fundamental theorem of identity types 118 | cell≃path : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} (s : Src (Fill X∞) f) (t : Fill X∞ f) 119 | → (src-comp X∞ fib s ≡ t) ≃ Fill (Hom X∞) (f , s , t) 120 | cell≃path X∞ fib s t = recognizeId (λ t → (Fill (Hom X∞)) (_ , s , t)) (fib .fill-fib _ s .fst .snd) (fib .fill-fib _ s) t 121 | 122 | isOfHLevelPathPred : ∀ (n : ℕ) {ℓ} {A : Type ℓ} → isOfHLevel n A → {x y : A} → isOfHLevel (predℕ n) (x ≡ y) 123 | isOfHLevelPathPred zero h = isContr→isContrPath h _ _ 124 | isOfHLevelPathPred (suc n) h = isOfHLevelPath' n h _ _ 125 | 126 | is-n-trunc-fib : ∀ {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) → is-fibrant-ext X∞ → ((f : Frm X) → isOfHLevel n (X∞ .Fill f)) → is-n-trunc n X∞ 127 | hLevel (is-n-trunc-fib n {X} X∞ fib h) = h 128 | is-trunc-ext (is-n-trunc-fib n {X} X∞ fib h) = is-n-trunc-fib _ _ (fib .hom-fib) lemma where 129 | lemma : (f : Frm (X , Fill X∞)) → isOfHLevel (predℕ n) (X∞ .Hom .Fill f) 130 | lemma (f , s , t) = isOfHLevelRespectEquiv (predℕ n) (cell≃path X∞ fib s t) (isOfHLevelPathPred n (h f)) 131 | 132 | 𝕆∞Path : ∀ {m ℓ} {X : 𝕆Type m ℓ} (X∞ X∞' : 𝕆Type∞ X) (p : Fill X∞ ≡ Fill X∞') → PathP (λ i → 𝕆Type∞ (X , p i)) (Hom X∞) (Hom X∞') → X∞ ≡ X∞' 133 | Fill (𝕆∞Path X∞ X∞' p q i) = p i 134 | Hom (𝕆∞Path X∞ X∞' p q i) = q i 135 | 136 | -- Trying to prove that fibrant opetopic types with same base-type are equal 137 | {- 138 | lemma0 : ∀ {m ℓ} {X X' : 𝕆Type m ℓ} (p : X ≡ X') 139 | → (P : Frm X → Type ℓ) (P' : Frm X' → Type ℓ) 140 | → (q : PathP (λ i → Frm (p i) → Type ℓ) P P') 141 | → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X' , P')) 142 | → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 143 | → PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill P∞) (Fill P∞') 144 | lemma0 {ℓ = ℓ} {X = X} = 145 | J 146 | (λ X' p → (P : Frm X → Type ℓ) (P' : Frm X' → Type ℓ) 147 | → (q : PathP (λ i → Frm (p i) → Type ℓ) P P') 148 | → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X' , P')) 149 | → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 150 | → PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill P∞) (Fill P∞')) 151 | 152 | λ P _ → J 153 | 154 | (λ P' q → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X , P')) 155 | → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 156 | → PathP (λ i → Frm (refl i , q i) → Type ℓ) (Fill P∞) (Fill P∞')) 157 | 158 | λ P∞ _ fib fib' → funExt λ x → ua (invEquiv (cell≃path _ fib _ _)) ∙ cong (_≡ (snd (snd x))) {!!} ∙ ua (cell≃path _ fib' _ _) 159 | 160 | fibrant→≡ : ∀ {m ℓ} {X X' : 𝕆Type m ℓ} (p : X ≡ X') 161 | → (X∞ : 𝕆Type∞ X) (X∞' : 𝕆Type∞ X') → PathP (λ i → Frm (p i) → Type ℓ) (Fill X∞) (Fill X∞') 162 | → is-fibrant-ext X∞ → is-fibrant-ext X∞' 163 | → PathP (λ i → 𝕆Type∞ (p i)) X∞ X∞' 164 | Fill (fibrant→≡ p X∞ X∞' q fib fib' i) = q i 165 | Hom (fibrant→≡ {ℓ = ℓ} p X∞ X∞' q fib fib' i) = fibrant→≡ (λ i → (p i , q i)) (Hom X∞) (Hom X∞') lemma (hom-fib fib) (hom-fib fib') i where 166 | lemma : PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill (Hom X∞)) (Fill (Hom X∞')) 167 | lemma = lemma0 p _ _ _ (Hom X∞) (Hom X∞') (subst is-fibrant-ext eta-fib-ext fib) (subst is-fibrant-ext eta-fib-ext fib') 168 | -} 169 | {- 170 | is-n-cat : ∀ {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) → Type ℓ 171 | is-n-cat zero X∞ = is-fibrant-ext X∞ 172 | is-n-cat (suc n) X∞ = Σ (is-n-cat n (Hom X∞)) {!!} 173 | -} 174 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Shapes.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --no-positivity-check #-} 2 | -- 3 | -- Shapes.agda - General shapes in opetopic types 4 | -- 5 | 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Sigma 9 | 10 | open import Experimental.NoDecs.OpetopicType 11 | 12 | module Experimental.NoDecs.Shapes where 13 | globe-Frm : ∀ {n ℓ} {Xₙ : 𝕆Type n ℓ} (Xₛₙ : Frm Xₙ → Type ℓ) {f : Frm Xₙ} (x y : Xₛₙ f) → Frm (Xₙ , Xₛₙ) 14 | globe-Frm Xₛₙ {f} x y = f , η Xₛₙ x , y 15 | 16 | _**_ : ∀ {ℓ} → Type ℓ → ℕ → Type ℓ 17 | X ** zero = Lift Unit 18 | X ** suc zero = X 19 | X ** suc (suc n) = X × (X ** (suc n)) 20 | 21 | fstt : ∀ {ℓ n} {X : Type ℓ} → X ** (suc n) → X 22 | fstt {n = zero} x = x 23 | fstt {n = suc n} (x , _) = x 24 | 25 | last : ∀ {ℓ n} {X : Type ℓ} → X ** (suc n) → X 26 | last {n = zero} x = x 27 | last {n = suc n} (fst , snd) = last snd 28 | 29 | path-chain : ∀ {ℓ} {X : Type ℓ} (Y : X → X → Type ℓ) (n : ℕ) (t : X ** (suc n)) → Type ℓ 30 | path-chain Y zero _ = Lift Unit 31 | path-chain Y (suc zero) (x , y) = Y y x 32 | path-chain Y (suc (suc n)) (x , t) = Y (fstt t) x × path-chain Y (suc n) t 33 | 34 | -- Sequences of arrows 35 | n-path : ∀ {ℓ} {X : 𝕆Type 2 ℓ} (n : ℕ) {t : (X .fst .snd tt*) ** suc n} 36 | → path-chain (λ x y → X .snd (tt* , x , y)) n t 37 | → Src (X .snd) (tt* , last t , fstt t) 38 | n-path zero {x} f = lf x 39 | n-path (suc zero) {y , x} f = nd y [ x , x , lf x ] f 40 | n-path (suc (suc n)) {y , x , t} (f , l) = nd y [ x , last t , (n-path (suc n) l) ] f 41 | 42 | -- Sequences of unary higher cells 43 | n-path' : ∀ {ℓ} {m : ℕ} {X : 𝕆Type m ℓ} (n : ℕ) {P : Frm X → Type ℓ} (Q : Frm (X , P) → Type ℓ) {f : Frm X} 44 | {t : P f ** (suc n)} 45 | → path-chain (λ x y → Q (globe-Frm P x y)) n t 46 | → Src Q (globe-Frm P (last t) (fstt t)) 47 | n-path' zero {P} Q {f} {x} _ = lf x 48 | n-path' (suc zero) {P} Q {f} {y , x} p = nd y (η (Branch Q) [ x , _ , (lf x) ]) p 49 | n-path' (suc (suc n)) {P} Q {f} {y , x , t} (p , l) = nd y (η (Branch Q) [ x , (η P (last t)) , (n-path' (suc n) Q l) ]) p 50 | 51 | 52 | module _ {ℓ} (X₀ : 𝕆Type 1 ℓ) where 53 | Obj : Type ℓ 54 | Obj = snd X₀ tt* 55 | 56 | hom-Frm : Obj → Obj → Frm X₀ 57 | hom-Frm x y = (tt* , x , y) 58 | 59 | module _ (X₁ : Frm X₀ → Type ℓ) where 60 | hom : Obj → Obj → Type ℓ 61 | hom x y = X₁ (hom-Frm x y) 62 | 63 | simplex-Frm : {x y z : Obj} (f : hom x y) (g : hom y z) (h : hom x z) → Frm (X₀ , X₁) 64 | simplex-Frm {x} {y} {z} f g h = hom-Frm x z , n-path 2 (g , f) , h -- nd z [ y , x , nd y [ x , x , lf x ] f ] g 65 | 66 | 2-drop-Frm : (x : Obj) (f : hom x x) → Frm (X₀ , X₁) 67 | 2-drop-Frm x f = hom-Frm x x , lf x , f 68 | 69 | module _ (X₂ : Frm (X₀ , X₁) → Type ℓ) where 70 | Null : (x : Obj) (f : hom x x) → Type ℓ 71 | Null x f = X₂ (2-drop-Frm x f) 72 | 73 | 2Glob : {x y : Obj} (f g : hom x y) → Type ℓ 74 | 2Glob {x} {y} f g = X₂ ((tt* , x , y) , nd y [ x , x , lf x ] f , g) 75 | 76 | Simplex : {x y z : Obj} (f : hom x y) (g : hom y z) 77 | → (h : hom x z) → Type ℓ 78 | Simplex {x} {y} {z} f g h = X₂ (simplex-Frm f g h) 79 | 80 | unitor-Frm : (x y : Obj) (f : hom x x) (g : hom x y) (h : hom x y) 81 | → (Δ : X₂ (simplex-Frm f g h)) 82 | → (Γ : X₂ (2-drop-Frm x f)) 83 | → (O : X₂ (globe-Frm X₁ g h)) 84 | → Frm ((X₀ , X₁) , X₂) 85 | unitor-Frm x y f g h Δ Γ O = _ , nd h (nd y [ x , x , nd x [ x , η _ x , lf x ] [ f , lf x , nd f (lf x) Γ ] ] [ g , _ , lf g ]) Δ , O 86 | 87 | associator1 : (x y z t : Obj) (f : hom x y) (g : hom y z) (h : hom z t) (i : hom x z) (j : hom x t) 88 | → (Δ₁ : X₂ (simplex-Frm f g i)) 89 | → (Δ₂ : X₂ (simplex-Frm i h j)) 90 | → Src X₂ (hom-Frm x t , n-path 3 (h , g , f) , j) --nd t [ _ , _ , (nd _ [ _ , _ , (nd _ [ _ , _ , (lf _) ] f) ] g) ] h) 91 | associator1 x y z t f g h i j Δ₁ Δ₂ = nd j (nd t [ z , x , (nd z [ x , x , (lf x) ] [ i , nd z [ y , x , (nd _ _ _) ] g , nd i (nd z [ y , x , nd y [ x , x , (lf x) ] [ f , _ , lf f ] ] [ g , _ , (lf g) ]) Δ₁ ]) ] [ h , _ , (lf h) ]) Δ₂ 92 | 93 | 94 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Sigma.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --no-positivity-check #-} 2 | -- 3 | -- Sigma.agda - Dependent sum of opetopic families 4 | -- 5 | 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Data.Sigma 8 | open import Cubical.Data.Empty 9 | open import Cubical.Data.Unit 10 | open import Cubical.Data.Nat 11 | open import Cubical.Data.Sum 12 | 13 | open import Core.Prelude 14 | open import Experimental.NoDecs.OpetopicType 15 | open import Experimental.NoDecs.OpetopicFamily 16 | 17 | module Experimental.NoDecs.Sigma where 18 | 19 | 𝕋↓ : ∀ {n ℓ₀} (X : 𝕆Type n ℓ₀) (ℓ : Level) → 𝕆Fam X ℓ 20 | 𝕋↓ {zero} {ℓ₀} X ℓ = tt* 21 | 𝕋↓ {suc n} {ℓ₀} (X , P) ℓ = (𝕋↓ X ℓ) , (λ f x → Lift Unit) 22 | 23 | 𝕆Σ : ∀ {n ℓ} (X : 𝕆Type n ℓ) (X↓ : 𝕆Fam X ℓ) → 𝕆Type n ℓ 24 | 25 | Fst : ∀ {n ℓ} {X : 𝕆Type n ℓ} {X↓ : 𝕆Fam X ℓ} → 𝕆Σ X X↓ ⇒ X 26 | Snd : ∀ {n ℓ} {X : 𝕆Type n ℓ} {X↓ : 𝕆Fam X ℓ} → 𝕆Σ X X↓ ⇛[ Fst ] X↓ 27 | 28 | -- Not sure why there are positivity issues for Σ-cell here and not for prod-cell in CartesianProduct.agda 29 | record Σ-cell {n ℓ} {X : 𝕆Type n ℓ} {X↓ : 𝕆Fam X ℓ} 30 | (P : Frm X → Type ℓ) 31 | (P↓ : {f : Frm X} (f↓ : Frm↓ X↓ f) → P f → Type ℓ) 32 | (f : Frm (𝕆Σ X X↓)) : Type ℓ where 33 | constructor _∣_ 34 | field 35 | fstₒ : P (Frm⇒ Fst f) 36 | sndₒ : P↓ (Frm⇛ Snd f ) fstₒ 37 | open Σ-cell 38 | 39 | 𝕆Σ {zero} X X↓ = tt* 40 | 𝕆Σ {suc n} (X , P) (X↓ , P↓) = 𝕆Σ X X↓ , Σ-cell P P↓ 41 | 42 | Fst {zero} = tt* 43 | Fst {suc n} = Fst , fstₒ 44 | 45 | Snd {zero} = tt* 46 | Snd {suc n} = Snd , sndₒ 47 | -------------------------------------------------------------------------------- /Experimental/NoDecs/Structures.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Everything 2 | 3 | open import Core.Prelude 4 | open import Experimental.NoDecs.OpetopicType 5 | open import Experimental.NoDecs.Shapes 6 | 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Sigma 9 | open import Cubical.Foundations.Equiv.Fiberwise 10 | 11 | module Experimental.NoDecs.Structures where 12 | 13 | record 𝕆Type∞ {n ℓ} (Xₙ : 𝕆Type n ℓ) : Type (ℓ-suc ℓ) where 14 | coinductive 15 | field 16 | Fill : Frm Xₙ → Type ℓ 17 | Hom : 𝕆Type∞ (Xₙ , Fill) 18 | open 𝕆Type∞ public 19 | 20 | record Map {n ℓ} {Xₙ Yₙ : 𝕆Type n ℓ} (σ : Xₙ ⇒ Yₙ) (X : 𝕆Type∞ Xₙ) (Y : 𝕆Type∞ Yₙ) : Type ℓ where 21 | coinductive 22 | field 23 | Fill⇒ : {f : Frm Xₙ} → (Fill X) f → (Fill Y) (Frm⇒ σ f) 24 | Hom⇒ : Map (σ , Fill⇒) (Hom X) (Hom Y) 25 | open Map public 26 | 27 | horn-filler : ∀ {n ℓ} {Xₙ : 𝕆Type n ℓ} {Xₛₙ : Frm Xₙ → Type ℓ} (Xₛₛₙ : Frm (Xₙ , Xₛₙ) → Type ℓ) {f : Frm Xₙ} → Src Xₛₙ f → Type ℓ 28 | horn-filler {n} {ℓ} {Xₙ} {Xₛₙ} Xₛₛₙ {f} s = Σ[ tgt ∈ Xₛₙ f ] Xₛₛₙ (f , s , tgt) 29 | 30 | is-fibrant : ∀ {n ℓ} → 𝕆Type (suc (suc n)) ℓ → Type ℓ 31 | is-fibrant ((Xₙ , Xₛₙ) , Xₛₛₙ) = (f : Frm Xₙ) (s : Src Xₛₙ f) → isContr (horn-filler Xₛₛₙ s) 32 | 33 | isProp-is-fibrant : ∀ {n ℓ} {X : 𝕆Type (suc (suc n)) ℓ} → isProp (is-fibrant X) 34 | isProp-is-fibrant = isPropΠ2 (λ _ _ → isPropIsContr) 35 | 36 | record is-fibrant-ext {n ℓ} {Xₙ : 𝕆Type n ℓ} (X : 𝕆Type∞ Xₙ) : Type ℓ where 37 | coinductive 38 | field 39 | fill-fib : is-fibrant ((Xₙ , (Fill X)) , (Fill (Hom X))) 40 | hom-fib : is-fibrant-ext (Hom X) 41 | 42 | open is-fibrant-ext public 43 | 44 | eta-fib-ext : ∀ {m ℓ} {X : 𝕆Type m ℓ} {X∞ : 𝕆Type∞ X} → X∞ ≡ record { Fill = Fill X∞ ; Hom = Hom X∞ } 45 | Fill (eta-fib-ext {X∞ = X∞} i) = Fill X∞ 46 | Hom (eta-fib-ext {X∞ = X∞} i) = Hom X∞ 47 | 48 | isProp-is-fibrant-ext : ∀ {n ℓ} {X : 𝕆Type n ℓ} {X∞ : 𝕆Type∞ X} → isProp (is-fibrant-ext X∞) 49 | fill-fib (isProp-is-fibrant-ext x y i) = isProp-is-fibrant (fill-fib x) (fill-fib y) i 50 | hom-fib (isProp-is-fibrant-ext x y i) = isProp-is-fibrant-ext (hom-fib x) (hom-fib y) i 51 | 52 | 𝕋 : ∀ {ℓ} (n : ℕ) → 𝕆Type n ℓ 53 | 𝕋 zero = tt* 54 | 𝕋 (suc n) = 𝕋 n , λ _ → Lift Unit 55 | 56 | 𝕋Ext : ∀ {n ℓ} (X : 𝕆Type n ℓ) → 𝕆Type∞ X 57 | Fill (𝕋Ext X) = λ _ → Lift Unit 58 | Hom (𝕋Ext X) = 𝕋Ext _ 59 | 60 | is-fib-ext-𝕋Ext : ∀ {n ℓ} {X : 𝕆Type n ℓ} → is-fibrant-ext (𝕋Ext X) 61 | fill-fib is-fib-ext-𝕋Ext f s = (tt* , tt*) , λ (tt* , tt*) → refl 62 | hom-fib is-fib-ext-𝕋Ext = is-fib-ext-𝕋Ext 63 | 64 | 65 | 66 | hom→path0 : ∀ {ℓ} {X : 𝕆Type∞ (𝕋 {ℓ} 0)} → is-fibrant-ext X → (x y : X .Fill tt*) → X .Hom .Fill (tt* , x , y) → x ≡ y 67 | hom→path0 {ℓ} {X} fib x y σ = sym (transportRefl x) ∙ cong fst (isContr→isProp (fib .fill-fib tt* x) (Id x x refl) b) where 68 | Id : (x y : X .Fill tt*) → (x ≡ y) → horn-filler (Fill (Hom X)) x 69 | Id x y = J (λ y p → horn-filler (Fill (Hom X)) x) (x , fib .hom-fib .fill-fib (tt* , x , x) (lf x) .fst .fst) 70 | 71 | b : horn-filler (Fill (Hom X)) x 72 | b = y , σ 73 | 74 | Id-cell : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} (t : Fill X∞ f) → Fill (Hom X∞) (globe-Frm _ t t) 75 | Id-cell X∞ fib t = fib .hom-fib .fill-fib _ (lf t) .fst .fst 76 | 77 | path→hom : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} 78 | → t ≡ u → Fill (Hom X∞) (globe-Frm _ t u) 79 | path→hom X∞ fib {t = t} {u = u} = J (λ u p → Fill (Hom X∞) (globe-Frm _ t u)) (Id-cell X∞ fib t) 80 | 81 | hom→path : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} 82 | → Fill (Hom X∞) (globe-Frm _ t u) → t ≡ u 83 | hom→path X∞ fib {f} {t} {u} cell = cong fst (isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (u , cell)) 84 | 85 | module IsoHomPath {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} {t u : Fill X∞ f} where 86 | sec : section (path→hom X∞ fib {f} {t} {u}) (hom→path X∞ fib {f} {t} {u}) 87 | sec cell = fromPathP (cong snd (isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (u , cell))) 88 | -- Wow ! Might need to take a look back at this later cause I didn't expect it to be so "simple". 89 | 90 | ret : retract (path→hom X∞ fib {f} {t} {u}) (hom→path X∞ fib {f} {t} {u}) 91 | ret = J (λ u p → hom→path X∞ fib (path→hom X∞ fib p) ≡ p) 92 | (hom→path X∞ fib (path→hom X∞ fib refl) 93 | ≡⟨ cong (hom→path X∞ fib) (transportRefl _) ⟩ 94 | hom→path X∞ fib (Id-cell X∞ fib t) 95 | ≡⟨ (λ i j → fst (test i j)) ⟩ 96 | refl ∎) where 97 | a : horn-filler (Fill (Hom X∞)) (η (Fill X∞) t) 98 | a = (t , Id-cell X∞ fib t) 99 | 100 | test0 : a ≡ a 101 | test0 = isContr→isProp (fib .fill-fib f (η (Fill X∞) t)) (t , Id-cell X∞ fib t) (t , Id-cell X∞ fib t) 102 | 103 | test : test0 ≡ refl 104 | test = isProp→isSet (isContr→isProp (fib .fill-fib f (η (Fill X∞) t))) (t , Id-cell X∞ fib t) (t , Id-cell X∞ fib t) test0 refl 105 | 106 | record is-n-trunc {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) : Type ℓ where 107 | coinductive 108 | field 109 | hLevel : (f : Frm X) → isOfHLevel n (X∞ .Fill f) 110 | is-trunc-ext : is-n-trunc (predℕ n) (X∞ .Hom) 111 | open is-n-trunc 112 | 113 | src-comp : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) → is-fibrant-ext X∞ → {f : Frm X} → Src (Fill X∞) f → Fill X∞ f 114 | src-comp X∞ fib s = fib .fill-fib _ s .fst .fst 115 | 116 | -- More general version of the equivalence between hom and path, using the fundamental theorem of identity types 117 | cell≃path : ∀ {n ℓ} {X : 𝕆Type n ℓ} (X∞ : 𝕆Type∞ X) (fib : is-fibrant-ext X∞) {f : Frm X} (s : Src (Fill X∞) f) (t : Fill X∞ f) 118 | → (src-comp X∞ fib s ≡ t) ≃ Fill (Hom X∞) (f , s , t) 119 | cell≃path X∞ fib s t = recognizeId (λ t → (Fill (Hom X∞)) (_ , s , t)) (fib .fill-fib _ s .fst .snd) (fib .fill-fib _ s) t 120 | 121 | isOfHLevelPathPred : ∀ (n : ℕ) {ℓ} {A : Type ℓ} → isOfHLevel n A → {x y : A} → isOfHLevel (predℕ n) (x ≡ y) 122 | isOfHLevelPathPred zero h = isContr→isContrPath h _ _ 123 | isOfHLevelPathPred (suc n) h = isOfHLevelPath' n h _ _ 124 | 125 | is-n-trunc-fib : ∀ {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) → is-fibrant-ext X∞ → ((f : Frm X) → isOfHLevel n (X∞ .Fill f)) → is-n-trunc n X∞ 126 | hLevel (is-n-trunc-fib n {X} X∞ fib h) = h 127 | is-trunc-ext (is-n-trunc-fib n {X} X∞ fib h) = is-n-trunc-fib _ _ (fib .hom-fib) lemma where 128 | lemma : (f : Frm (X , Fill X∞)) → isOfHLevel (predℕ n) (X∞ .Hom .Fill f) 129 | lemma (f , s , t) = isOfHLevelRespectEquiv (predℕ n) (cell≃path X∞ fib s t) (isOfHLevelPathPred n (h f)) 130 | 131 | 𝕆∞Path : ∀ {m ℓ} {X : 𝕆Type m ℓ} (X∞ X∞' : 𝕆Type∞ X) (p : Fill X∞ ≡ Fill X∞') → PathP (λ i → 𝕆Type∞ (X , p i)) (Hom X∞) (Hom X∞') → X∞ ≡ X∞' 132 | Fill (𝕆∞Path X∞ X∞' p q i) = p i 133 | Hom (𝕆∞Path X∞ X∞' p q i) = q i 134 | 135 | -- Trying to prove that fibrant opetopic types with same base-type are equal 136 | {- 137 | lemma0 : ∀ {m ℓ} {X X' : 𝕆Type m ℓ} (p : X ≡ X') 138 | → (P : Frm X → Type ℓ) (P' : Frm X' → Type ℓ) 139 | → (q : PathP (λ i → Frm (p i) → Type ℓ) P P') 140 | → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X' , P')) 141 | → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 142 | → PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill P∞) (Fill P∞') 143 | lemma0 {ℓ = ℓ} {X = X} = 144 | J 145 | (λ X' p → (P : Frm X → Type ℓ) (P' : Frm X' → Type ℓ) 146 | → (q : PathP (λ i → Frm (p i) → Type ℓ) P P') 147 | → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X' , P')) 148 | → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 149 | → PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill P∞) (Fill P∞')) 150 | 151 | λ P _ → J 152 | 153 | (λ P' q → (P∞ : 𝕆Type∞ (X , P)) (P∞' : 𝕆Type∞ (X , P')) 154 | → is-fibrant-ext (record {Fill = P ; Hom = P∞}) → is-fibrant-ext (record { Fill = P' ; Hom = P∞' }) 155 | → PathP (λ i → Frm (refl i , q i) → Type ℓ) (Fill P∞) (Fill P∞')) 156 | 157 | λ P∞ _ fib fib' → funExt λ x → ua (invEquiv (cell≃path _ fib _ _)) ∙ cong (_≡ (snd (snd x))) {!!} ∙ ua (cell≃path _ fib' _ _) 158 | 159 | fibrant→≡ : ∀ {m ℓ} {X X' : 𝕆Type m ℓ} (p : X ≡ X') 160 | → (X∞ : 𝕆Type∞ X) (X∞' : 𝕆Type∞ X') → PathP (λ i → Frm (p i) → Type ℓ) (Fill X∞) (Fill X∞') 161 | → is-fibrant-ext X∞ → is-fibrant-ext X∞' 162 | → PathP (λ i → 𝕆Type∞ (p i)) X∞ X∞' 163 | Fill (fibrant→≡ p X∞ X∞' q fib fib' i) = q i 164 | Hom (fibrant→≡ {ℓ = ℓ} p X∞ X∞' q fib fib' i) = fibrant→≡ (λ i → (p i , q i)) (Hom X∞) (Hom X∞') lemma (hom-fib fib) (hom-fib fib') i where 165 | lemma : PathP (λ i → Frm (p i , q i) → Type ℓ) (Fill (Hom X∞)) (Fill (Hom X∞')) 166 | lemma = lemma0 p _ _ _ (Hom X∞) (Hom X∞') (subst is-fibrant-ext eta-fib-ext fib) (subst is-fibrant-ext eta-fib-ext fib') 167 | -} 168 | {- 169 | is-n-cat : ∀ {m ℓ} (n : ℕ) {X : 𝕆Type m ℓ} (X∞ : 𝕆Type∞ X) → Type ℓ 170 | is-n-cat zero X∞ = is-fibrant-ext X∞ 171 | is-n-cat (suc n) X∞ = Σ (is-n-cat n (Hom X∞)) {!!} 172 | -} 173 | -------------------------------------------------------------------------------- /Experimental/Scratch.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Scratch.agda - Random things I'm working on 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Nat 7 | open import Cubical.Data.Sum 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Empty 10 | 11 | open import Core.Prelude 12 | open import Core.Everything 13 | 14 | module Experimental.Scratch where 15 | 16 | data Tr (A : Type) : ℕ → Type where 17 | 18 | obj : A → Tr A 0 19 | 20 | lf : {n : ℕ} 21 | → Tr A (suc n) 22 | 23 | nd : {n : ℕ} → A 24 | → Tr (Tr A (suc n)) n 25 | → Tr A (suc n) 26 | 27 | 28 | -------------------------------------------------------------------------------- /Experimental/wip/Faces.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --rewriting #-} 2 | 3 | open import MiniHoTT 4 | 5 | open import Opetopes 6 | open import OpetopicType 7 | open import OpetopicMap 8 | 9 | module Faces where 10 | 11 | data InnerFace : {n : ℕ} → 𝒪 n → ℕ → Set where 12 | as-src : {n : ℕ} (o : 𝒪 n) (p : 𝒫 o) (q : 𝒫 (o , p)) 13 | → Pos q → InnerFace ((o , p) , q) (S n) 14 | as-tgt : {n : ℕ} (o : 𝒪 n) (p : 𝒫 o) (q : 𝒫 (o , p)) 15 | → Pos q → InnerFace ((o , p) , q) n 16 | raise-face : {n m : ℕ} (o : 𝒪 n) (p : 𝒫 o) 17 | → InnerFace o m → InnerFace (o , p) m 18 | 19 | data Face : {n : ℕ} → 𝒪 n → ℕ → Set where 20 | top : {n : ℕ} (o : 𝒪 n) → Face o n 21 | tgt-face : {n : ℕ} (o : 𝒪 (S n)) → Face o n 22 | init : {n : ℕ} (o : 𝒪 (S n)) → Face o 0 23 | inner : {n m : ℕ} (o : 𝒪 n) → InnerFace o m → Face o m 24 | 25 | -- obj : 𝒪 0 26 | -- obj = tt 27 | 28 | obj-face : Face obj 0 29 | obj-face = top tt 30 | 31 | -- arrow : 𝒪 1 32 | -- arrow = tt , tt 33 | 34 | arrow-src : Face arrow 0 35 | arrow-src = init arrow 36 | 37 | arrow-tgt : Face arrow 0 38 | arrow-tgt = tgt-face arrow 39 | 40 | arrow-top : Face arrow 1 41 | arrow-top = top arrow 42 | 43 | -- 2-simplex : 𝒪 2 44 | -- 2-simplex = (tt , tt) , ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt))) 45 | 46 | simplex = 2-simplex 47 | 48 | simp-init : Face simplex 0 49 | simp-init = init simplex 50 | 51 | simp-mid : Face simplex 0 52 | simp-mid = inner simplex (as-tgt tt tt 53 | (ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt)))) (inl tt)) 54 | 55 | simp-term : Face simplex 0 56 | simp-term = inner simplex (as-tgt tt tt 57 | (ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt)))) (inr (tt , inl tt))) 58 | 59 | simp-fst : Face simplex 1 60 | simp-fst = inner simplex (as-src tt tt 61 | (ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt)))) (inl tt)) 62 | 63 | simp-snd : Face simplex 1 64 | simp-snd = inner simplex (as-src tt tt 65 | (ndₒ tt tt (cst tt) (λ p → ndₒ tt tt (cst tt) (cst (lfₒ tt)))) (inr (tt , inl tt))) 66 | 67 | simp-tgt : Face simplex 1 68 | simp-tgt = tgt-face simplex 69 | 70 | simp-top : Face simplex 2 71 | simp-top = top simplex 72 | 73 | -- 74 | -- Representables 75 | -- 76 | 77 | Rep : {n : ℕ} (o : 𝒪 n) → 𝕆 ℓ-zero (S n) 78 | 79 | SrcInc' : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → Rep (Typ ρ p) ⇒ fst (Rep (o , ρ)) 80 | SrcInc' = {!!} 81 | 82 | SrcInc : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → (Rep (Typ ρ p) , cst ∅) ⇒ Rep (o , ρ) 83 | TgtInc : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → (Rep o , cst ∅) ⇒ Rep (o , ρ) 84 | 85 | -- Hmmm. This is not quite right. It is the top *3* fibrations 86 | -- which must change. 87 | 88 | -- So, the cells of the first are parameterized by two things 89 | -- 1) the positions of ρ (there will be the externals) 90 | -- 2) the positions of τ (this will be the internals) 91 | 92 | -- In the next fibration, we have exactly the positions of τ and the target. 93 | 94 | -- And then we have the unique next dimension. 95 | 96 | -- Right. So this should actually lead to another, slight different 97 | -- presentation of the faces above. 98 | 99 | data ArrowFill : {o : 𝒪 1} → WebFrm tt (cst (⊤ ⊔ ⊤)) (fst o) (snd o) → Set where 100 | top : ArrowFill ⟪ tt , tt , inl tt , (λ _ → inr tt) ⟫ 101 | 102 | data BoundaryCell {n} (o : 𝒪 n) (ρ : 𝒫 o) : {q : 𝒪 n} → Frm (fst (Rep o)) q → Set where 103 | src-cell : (p : Pos ρ) (f : Frm (Rep (Typ ρ p)) {!!}) 104 | → BoundaryCell o ρ {!!} 105 | 106 | data FillingCell {n} (o : 𝒪 n) (ρ : 𝒫 o) : {q : 𝒪 (S n)} 107 | → WebFrm (fst (Rep o)) (BoundaryCell o ρ) (fst q) (snd q) → Set where 108 | 109 | 110 | Rep {n = O} tt = tt , cst ⊤ 111 | Rep {n = S O} (tt , tt) = 112 | (tt , cst (⊤ ⊔ ⊤)) , ArrowFill 113 | 114 | Rep {n = S (S n)} ((o , ρ) , τ) = ((fst (fst (Rep (o , ρ))) , {!!}) , {!!}) , {!!} 115 | 116 | -- Rep {n = O} o = tt , cst ⊤ 117 | -- Rep {n = S n} (o , ρ) = (fst (Rep o) , BoundaryCell o ρ) , FillingCell o ρ 118 | 119 | SrcInc = {!!} 120 | TgtInc = {!!} 121 | 122 | 123 | -- data LfBdry {n : ℕ} (m : 𝒪 n) : {o : 𝒪 (S n)} 124 | -- → WebFrm (fst (Rep m)) (snd (Rep m)) (fst o) (snd o) → Set where 125 | -- lf-tgt : LfBdry m ⟪ {!!} , {!η (fst (Rep m))!} , {!!} , {!!} ⟫ 126 | 127 | 128 | -- Rep {n = S (S n)} ((x , .(ηₒ x)) , lfₒ .x) = (Rep x , LfBdry x) , {!!} 129 | -- Rep {n = S (S n)} ((o , .(μₒ ρ δ)) , ndₒ .o ρ δ ε) = {!!} 130 | 131 | 132 | -------------------------------------------------------------------------------- /Experimental/wip/Groupoids.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --rewriting #-} 2 | 3 | open import MiniHoTT 4 | open import Opetopes 5 | open import OpetopicType 6 | open import OpetopicMap 7 | 8 | module Groupoids where 9 | 10 | IdType : ∀ {ℓ n} (X : Set ℓ) → 𝕆 ℓ n 11 | ΔMap : ∀ {ℓ n} (X : Set ℓ) (x : X) → 𝕋 {ℓ} n ⇒ IdType X 12 | 13 | data Cell {ℓ n} (X : Set ℓ) : {o : 𝒪 n} → Frm (IdType X) o → Set ℓ where 14 | id-cell : (x : X) {o : 𝒪 n} (f : Frm (𝕋 {ℓ} n) o) 15 | → Cell X (Frm⇒ (ΔMap X x) f) 16 | 17 | IdType {n = O} X = tt 18 | IdType {n = S n} X = 19 | IdType X , Cell X 20 | 21 | ΔMap {n = O} X x = tt 22 | ΔMap {n = S n} X x = 23 | ΔMap X x , λ {o} {f} _ → id-cell x f 24 | 25 | -- What will it be like to prove such a thing is fibrant? 26 | 27 | -------------------------------------------------------------------------------- /Experimental/wip/IndexedOpetopicType.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicType.agda - Definition of Opetopic Types Indexed over Opetopes 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Nat 10 | open import Cubical.Data.Sum 11 | 12 | open import Prelude 13 | open import Opetopes 14 | 15 | module IndexedOpetopicType where 16 | 17 | 𝕆 : (ℓ : Level) → ℕ → Type (ℓ-suc ℓ) 18 | 19 | Frm : ∀ {ℓ n} → 𝕆 ℓ n → 𝒪 n → Type ℓ 20 | Cns : ∀ {ℓ n} (X : 𝕆 ℓ n) 21 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 22 | → 𝒫 𝑜 → Type ℓ 23 | Shp : ∀ {ℓ n} (X : 𝕆 ℓ n) 24 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 25 | → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 26 | → (p : Pos 𝑝) → Frm X (Typ 𝑝 p) 27 | 28 | η : ∀ {n ℓ} (X : 𝕆 ℓ n) 29 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 30 | → Cns X f (ηₒ 𝑜) 31 | 32 | {-# TERMINATING #-} 33 | μ : ∀ {n ℓ} (X : 𝕆 ℓ n) 34 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 35 | → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 36 | → {ι : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 37 | → (κ : (p : Pos 𝑝) → Cns X (Shp X c p) (ι p)) 38 | → Cns X f (μₒ 𝑝 ι) 39 | 40 | postulate 41 | 42 | η-pos-shp : ∀ {ℓ n} (X : 𝕆 ℓ n) 43 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 44 | → (p : Pos (ηₒ 𝑜)) 45 | → Shp X (η X f) p ↦ f 46 | {-# REWRITE η-pos-shp #-} 47 | 48 | μ-pos-shp : ∀ {ℓ n} (X : 𝕆 ℓ n) 49 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 50 | → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 51 | → {ι : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 52 | → (κ : (p : Pos 𝑝) → Cns X (Shp X c p) (ι p)) 53 | → (p : Pos (μₒ 𝑝 ι)) 54 | → Shp X (μ X c κ) p ↦ Shp X (κ (μₒ-pos-fst 𝑝 ι p)) (μₒ-pos-snd 𝑝 ι p) 55 | {-# REWRITE μ-pos-shp #-} 56 | 57 | -- Monad Laws 58 | μ-unit-r : ∀ {n ℓ} (X : 𝕆 ℓ n) 59 | → {𝑜 : 𝒪 n} (𝑝 : 𝒫 𝑜) 60 | → {f : Frm X 𝑜} (c : Cns X f 𝑝) 61 | → μ X c (λ p → η X (Shp X c p)) ↦ c 62 | {-# REWRITE μ-unit-r #-} 63 | 64 | μ-unit-l : ∀ {n ℓ} (X : 𝕆 ℓ n) 65 | → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 66 | → (ι : (p : Pos (ηₒ 𝑜)) → 𝒫 (Typ (ηₒ 𝑜) p)) 67 | → (δ : (p : Pos (ηₒ 𝑜)) → Cns X f (ι p)) 68 | → μ X (η X f) δ ↦ δ (ηₒ-pos 𝑜) 69 | {-# REWRITE μ-unit-l #-} 70 | 71 | μ-assoc : ∀ {n ℓ} (X : 𝕆 ℓ n) 72 | → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 73 | → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 74 | → {ι : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 75 | → (κ : (p : Pos 𝑝) → Cns X (Shp X c p) (ι p)) 76 | → (δ : (p : Pos (μₒ 𝑝 ι)) → 𝒫 (Typ (μₒ 𝑝 ι) p)) 77 | → (ε : (p : Pos (μₒ 𝑝 ι)) → Cns X (Shp X (κ (μₒ-pos-fst 𝑝 ι p)) (μₒ-pos-snd 𝑝 ι p)) (δ p)) 78 | → μ X (μ X c κ) ε 79 | ↦ μ X c (λ p → μ X (κ p) (λ q → ε (μₒ-pos 𝑝 ι p q))) 80 | {-# REWRITE μ-assoc #-} 81 | 82 | -- 83 | -- Definition of the Derived Monad 84 | -- 85 | 86 | module _ {ℓ n} (Xₙ : 𝕆 ℓ n) (Xₛₙ : {𝑜 : 𝒪 n} (f : Frm Xₙ 𝑜) → Type ℓ) where 87 | 88 | η-dec : {𝑜 : 𝒪 n} (f : Frm Xₙ 𝑜) (x : Xₛₙ f) 89 | → (p : Pos (ηₒ 𝑜)) → Xₛₙ (Shp Xₙ (η Xₙ f) p) 90 | η-dec {𝑜} f x = ηₒ-pos-elim 𝑜 (λ p → Xₛₙ (Shp Xₙ (η Xₙ f) p)) x 91 | 92 | μ-dec : {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {f : Frm Xₙ 𝑜} (c : Cns Xₙ f 𝑝) 93 | → (ι : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)) 94 | → (δ : (p : Pos 𝑝) → Cns Xₙ (Shp Xₙ c p) (ι p)) 95 | → (ν : (p : Pos 𝑝) (q : Pos (ι p)) → Xₛₙ (Shp Xₙ (δ p) q)) 96 | → (p : Pos (μₒ 𝑝 ι)) → Xₛₙ (Shp Xₙ (μ Xₙ c δ) p) 97 | μ-dec {𝑝 = 𝑝} c ι δ ν p = ν (μₒ-pos-fst 𝑝 ι p) (μₒ-pos-snd 𝑝 ι p) 98 | 99 | {-# NO_POSITIVITY_CHECK #-} 100 | record WebFrm (𝑜 : 𝒪 n) (𝑝 : 𝒫 𝑜) : Type ℓ where 101 | inductive 102 | eta-equality 103 | constructor ⟪_,_,_,_⟫ 104 | field 105 | frm : Frm Xₙ 𝑜 106 | cns : Cns Xₙ frm 𝑝 107 | tgt : Xₛₙ frm 108 | src : (p : Pos 𝑝) → Xₛₙ (Shp Xₙ cns p) 109 | 110 | open WebFrm public 111 | 112 | -- Attempt at a recursive version ... 113 | WebRec : {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} → WebFrm 𝑜 𝑝 → 𝒯r 𝑜 𝑝 → Type ℓ 114 | WebRec ⟪ f , c , x , ν ⟫ (lfₒ 𝑜) = (c , ν) == (η _ f , const x) 115 | WebRec ⟪ f , c , x , ν ⟫ (ndₒ 𝑜 𝑝 δ ε) = {!!} 116 | 117 | data Web : {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} → WebFrm 𝑜 𝑝 → 𝒯r 𝑜 𝑝 → Type ℓ where 118 | 119 | lf : {𝑜 : 𝒪 n} {f : Frm Xₙ 𝑜} (x : Xₛₙ f) 120 | → Web ⟪ f , η Xₙ f , x , const x ⟫ (lfₒ 𝑜) 121 | 122 | nd : {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} (φ : WebFrm 𝑜 𝑝) 123 | → (ι : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)) 124 | → (κ : (p : Pos 𝑝) → 𝒯r (Typ 𝑝 p) (ι p)) 125 | → (δ : (p : Pos 𝑝) → Cns Xₙ (Shp Xₙ (cns φ) p) (ι p)) 126 | → (ν : (p : Pos 𝑝) (q : Pos (ι p)) → Xₛₙ (Shp Xₙ (δ p) q)) 127 | → (ε : (p : Pos 𝑝) → Web ⟪ Shp Xₙ (cns φ) p , δ p , src φ p , ν p ⟫ (κ p)) 128 | → Web ⟪ frm φ , μ Xₙ (cns φ) δ , tgt φ , μ-dec (cns φ) ι δ ν ⟫ (ndₒ 𝑜 𝑝 ι κ) 129 | 130 | WebPos : {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {φ : WebFrm 𝑜 𝑝} {𝑡 : 𝒯r 𝑜 𝑝} (ω : Web φ 𝑡) → Type ℓ 131 | WebPos (lf _) = Lift ⊥ 132 | WebPos (nd {𝑝 = 𝑝} φ ι κ δ ν ε) = Unit ⊎ Σ (Pos 𝑝) (λ p → WebPos (ε p)) 133 | 134 | WebShp : {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {φ : WebFrm 𝑜 𝑝} {𝑡 : 𝒯r 𝑜 𝑝} 135 | → (ω : Web φ 𝑡) (p : 𝒯rPos 𝑡) 136 | → WebFrm (fst (𝒯rTyp 𝑡 p)) (snd (𝒯rTyp 𝑡 p)) 137 | WebShp (nd φ ι κ δ ν ε) (inl tt) = φ 138 | WebShp (nd φ ι κ δ ν ε) (inr (p , q)) = WebShp (ε p) q 139 | 140 | graft : {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} {φ : WebFrm 𝑜 𝑝} {𝑡 : 𝒯r 𝑜 𝑝} (ω : Web φ 𝑡) 141 | → (ι : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)) 142 | → (κ : (p : Pos 𝑝) → 𝒯r (Typ 𝑝 p) (ι p)) 143 | → (δ : (p : Pos 𝑝) → Cns Xₙ (Shp Xₙ (cns φ) p) (ι p)) 144 | → (ν : (p : Pos 𝑝) (q : Pos (ι p)) → Xₛₙ (Shp Xₙ (δ p) q)) 145 | → (ε : (p : Pos 𝑝) → Web ⟪ Shp Xₙ (cns φ) p , δ p , src φ p , ν p ⟫ (κ p)) 146 | → Web ⟪ frm φ , μ Xₙ (cns φ) δ , tgt φ , μ-dec (cns φ) ι δ ν ⟫ (γₒ 𝑡 ι κ) 147 | graft (lf {𝑜} {f} x) ι₁ κ₁ δ₁ ν₁ ε₁ = ε₁ (ηₒ-pos 𝑜) 148 | graft (nd {𝑝 = 𝑝} φ ι κ δ ν ε) ι₁ κ₁ δ₁ ν₁ ε₁ = 149 | let ι' p = μₒ (ι p) (ι-ih p) 150 | δ' p = μ Xₙ (δ p) (δ-ih p) 151 | κ' p = γₒ (κ p) (ι-ih p) (κ-ih p) 152 | ν' p q = ν₁ (μₒ-pos 𝑝 ι p (μₒ-pos-fst (ι p) (ι-ih p) q)) (μₒ-pos-snd (ι p) (ι-ih p) q) 153 | ε' p = graft (ε p) (ι-ih p) (κ-ih p) (δ-ih p) (ν-ih p) (ε-ih p) 154 | in nd φ ι' κ' δ' ν' ε' 155 | 156 | where ι-ih : (p : Pos 𝑝) (q : Pos (ι p)) → 𝒫 (Typ (ι p) q) 157 | ι-ih p q = ι₁ (μₒ-pos 𝑝 ι p q) 158 | 159 | κ-ih : (p : Pos 𝑝) (q : Pos (ι p)) → 𝒯r (Typ (ι p) q) (ι-ih p q) 160 | κ-ih p q = κ₁ (μₒ-pos 𝑝 ι p q) 161 | 162 | δ-ih : (p : Pos 𝑝) (q : Pos (ι p)) → Cns Xₙ (Shp Xₙ (δ p) q) (ι-ih p q) 163 | δ-ih p q = δ₁ (μₒ-pos 𝑝 ι p q) 164 | 165 | ν-ih : (p : Pos 𝑝) (q : Pos (ι p)) (r : Pos (ι-ih p q)) → Xₛₙ (Shp Xₙ (δ-ih p q) r) 166 | ν-ih p q = ν₁ (μₒ-pos 𝑝 ι p q) 167 | 168 | ε-ih : (p : Pos 𝑝) (q : Pos (ι p)) → Web ⟪ Shp Xₙ (δ p) q , δ-ih p q , ν p q , ν-ih p q ⟫ (κ-ih p q) 169 | ε-ih p q = ε₁ (μₒ-pos 𝑝 ι p q) 170 | 171 | -- TODO: Grafting Axioms 172 | 173 | 𝕆 ℓ zero = Lift Unit 174 | 𝕆 ℓ (suc n) = Σ (𝕆 ℓ n) (λ Xₙ → {𝑜 : 𝒪 n} → Frm Xₙ 𝑜 → Type ℓ) 175 | 176 | Frm {n = zero} X tt = Lift Unit 177 | Frm {n = suc n} (Xₙ , Xₛₙ) (𝑜 , 𝑝) = WebFrm Xₙ Xₛₙ 𝑜 𝑝 178 | 179 | Cns {n = zero} _ _ _ = Lift Unit 180 | Cns {n = suc n} (Xₙ , Xₛₙ) {𝑜 , 𝑝} = Web Xₙ Xₛₙ {𝑜} {𝑝} 181 | 182 | Shp {n = zero} _ _ _ = lift tt 183 | Shp {n = suc n} (Xₙ , Xₛₙ) {𝑜 , 𝑝} ω p = WebShp Xₙ Xₛₙ ω p 184 | 185 | -- η : ∀ {n ℓ} (X : 𝕆 ℓ n) 186 | -- → {𝑜 : 𝒪 n} (f : Frm X 𝑜) 187 | -- → Cns X f (ηₒ 𝑜) 188 | η {n = zero} _ _ = lift tt 189 | η {n = suc n} (Xₙ , Xₛₙ) {𝑜 , 𝑝} φ = 190 | let ι p = ηₒ (Typ 𝑝 p) 191 | κ p = lfₒ (Typ 𝑝 p) 192 | δ p = η Xₙ (Shp Xₙ (cns φ) p) 193 | ν p q = src φ p 194 | ε p = lf (src φ p) 195 | in nd φ ι κ δ ν ε 196 | 197 | -- μ : ∀ {n ℓ} (X : 𝕆 ℓ n) 198 | -- → {𝑜 : 𝒪 n} {f : Frm X 𝑜} 199 | -- → {𝑝 : 𝒫 𝑜} (c : Cns X f 𝑝) 200 | -- → {ι : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 201 | -- → (κ : (p : Pos 𝑝) → Cns X (Shp X c p) (ι p)) 202 | -- → Cns X f (μₒ 𝑝 ι) 203 | μ {n = zero} _ _ _ = lift tt 204 | μ {n = suc n} (Xₙ , Xₛₙ) (lf x) θ = lf x 205 | μ {n = suc n} (Xₙ , Xₛₙ) (nd φ ι κ δ ν ε) {ζ} θ = 206 | let ω = θ (inl tt) 207 | κ' p = μₒ (κ p) (λ q → ζ (inr (p , q))) 208 | ε' p = μ (Xₙ , Xₛₙ) (ε p) (λ q → θ (inr (p , q))) 209 | in graft Xₙ Xₛₙ ω ι κ' δ ν ε' 210 | 211 | -- 212 | -- The terminal opetopic type 213 | -- 214 | 215 | 𝕋 : ∀ {ℓ} (n : ℕ) → 𝕆 ℓ n 216 | 𝕋 zero = lift tt 217 | 𝕋 (suc n) = 𝕋 n , λ _ → Lift Unit 218 | -------------------------------------------------------------------------------- /Experimental/wip/Join.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --rewriting #-} 2 | 3 | open import MiniHoTT 4 | open import Opetopes 5 | open import OpetopicType 6 | open import OpetopicMap 7 | 8 | module Join where 9 | 10 | -- The join of opetopic types. 11 | 12 | infixl 50 _★_ 13 | 14 | _★_ : ∀ {ℓ n} (X Y : 𝕆 ℓ n) → 𝕆 ℓ n 15 | 16 | jn-inl : ∀ {ℓ n} (X Y : 𝕆 ℓ n) → X ⇒ X ★ Y 17 | jn-inr : ∀ {ℓ n} (X Y : 𝕆 ℓ n) → Y ⇒ X ★ Y 18 | 19 | -- Hmmm. But in this version, we don't have n fixed. Oh, so you 20 | -- have to change the indexing and give separate definitions in 21 | -- lower dimensions. Okay. So three *different* data type 22 | -- defintiions. 23 | 24 | data JoinCell {ℓ n} (X Y : 𝕆 ℓ n) : {o : 𝒪 n} (f : Frm (X ★ Y) o) → Set ℓ where 25 | 26 | 27 | _★_ {n = O} X Y = tt 28 | _★_ {n = S O} (_ , X₀) (_ , Y₀) = tt , (λ u → X₀ u ⊔ Y₀ u) 29 | _★_ {n = S (S O)} ((_ , X₀) , X₁) ((_ , Y₀) , Y₁) = (tt , (λ u → X₀ u ⊔ Y₀ u)) , {!!} 30 | _★_ {n = S (S (S n))} X Y = {!!} 31 | 32 | jn-inl = {!!} 33 | jn-inr = {!!} 34 | -------------------------------------------------------------------------------- /Experimental/wip/OpetopicMap.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicMap.agda - Maps of opetopic types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Unit 7 | open import Cubical.Data.Nat 8 | 9 | open import Prelude 10 | open import Opetopes 11 | open import OpetopicCtx 12 | 13 | module OpetopicMap where 14 | 15 | infixr 40 _⇒_ 16 | 17 | _⇒_ : ∀ {ℓ₀ ℓ₁ n} → 𝕆 ℓ₀ n → 𝕆 ℓ₁ n 18 | → Type (ℓ-max ℓ₀ ℓ₁) 19 | 20 | Frm⇒ : ∀ {ℓ₀ ℓ₁ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} (α : X ⇒ Y) 21 | → {o : 𝒪 n} → Frm X o → Frm Y o 22 | 23 | Cns⇒ : ∀ {ℓ₀ ℓ₁ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} (α : X ⇒ Y) 24 | → {o : 𝒪 n} {ρ : 𝒫 o} {f : Frm X o} 25 | → Cns X f ρ → Cns Y (Frm⇒ α f) ρ 26 | 27 | postulate 28 | 29 | Shp-Frm⇒ : ∀ {ℓ₀ ℓ₁ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} (α : X ⇒ Y) 30 | → {o : 𝒪 n} {ρ : 𝒫 o} {f : Frm X o} (c : Cns X f ρ) (p : Pos ρ) 31 | → Frm⇒ α (Shp X c p) ↦ Shp Y (Cns⇒ α c) p 32 | {-# REWRITE Shp-Frm⇒ #-} 33 | 34 | η⇒ : ∀ {ℓ₀ ℓ₁ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} (α : X ⇒ Y) 35 | → {o : 𝒪 n} (f : Frm X o) 36 | → Cns⇒ α (η X f) ↦ η Y (Frm⇒ α f) 37 | {-# REWRITE η⇒ #-} 38 | 39 | μ⇒ : ∀ {ℓ₀ ℓ₁ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} (α : X ⇒ Y) 40 | → {o : 𝒪 n} {f : Frm X o} 41 | → {ρ : 𝒫 o} (c : Cns X f ρ) 42 | → {ι : (p : Pos ρ) → 𝒫 (Typ ρ p)} 43 | → (κ : (p : Pos ρ) → Cns X (Shp X c p) (ι p)) 44 | → Cns⇒ α (μ X c κ) ↦ μ Y (Cns⇒ α c) (λ p → Cns⇒ α (κ p)) 45 | {-# REWRITE μ⇒ #-} 46 | 47 | module _ {ℓ₀ ℓ₁ n} {Xₙ : 𝕆 ℓ₀ n} {Yₙ : 𝕆 ℓ₁ n} 48 | (Xₛₙ : {o : 𝒪 n} → Frm Xₙ o → Type ℓ₀) 49 | (Yₛₙ : {o : 𝒪 n} → Frm Yₙ o → Type ℓ₁) 50 | (αₙ : Xₙ ⇒ Yₙ) (αₛₙ : {o : 𝒪 n} {f : Frm Xₙ o} → Xₛₙ f → Yₛₙ (Frm⇒ αₙ f)) where 51 | 52 | WebFrm⇒ : {o : 𝒪 n} {ρ : 𝒫 o} 53 | → WebFrm Xₙ Xₛₙ o ρ → WebFrm Yₙ Yₛₙ o ρ 54 | WebFrm⇒ φ = ⟪ Frm⇒ αₙ (frm φ) , Cns⇒ αₙ (cns φ) , 55 | αₛₙ (tgt φ) , (λ p → αₛₙ (src φ p)) ⟫ 56 | 57 | Web⇒ : {o : 𝒪 n} {ρ : 𝒫 o} 58 | → {φ : WebFrm Xₙ Xₛₙ o ρ} {τ : 𝒯r o ρ} 59 | → Web Xₙ Xₛₙ φ τ → Web Yₙ Yₛₙ (WebFrm⇒ φ) τ 60 | Web⇒ (lf x) = lf (αₛₙ x) 61 | Web⇒ (nd φ ι κ δ ν ε) = 62 | nd (WebFrm⇒ φ) ι κ 63 | (λ p → Cns⇒ αₙ (δ p)) 64 | (λ p q → αₛₙ (ν p q)) 65 | (λ p → Web⇒ (ε p)) 66 | 67 | _⇒_ {n = zero} _ _ = Lift Unit 68 | _⇒_ {n = suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) = 69 | Σ (Xₙ ⇒ Yₙ) (λ α → 70 | {o : 𝒪 n} {f : Frm Xₙ o} 71 | → Xₛₙ f → Yₛₙ (Frm⇒ α f)) 72 | 73 | Frm⇒ {n = zero} α _ = lift tt 74 | Frm⇒ {n = suc n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} (αₙ , αₛₙ) = 75 | WebFrm⇒ Xₛₙ Yₛₙ αₙ αₛₙ 76 | 77 | Cns⇒ {n = zero} _ _ = lift tt 78 | Cns⇒ {n = suc n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} (αₙ , αₛₙ) = 79 | Web⇒ Xₛₙ Yₛₙ αₙ αₛₙ 80 | 81 | -- 82 | -- Composition 83 | -- 84 | 85 | infixr 30 _⊚_ 86 | 87 | _⊚_ : ∀ {ℓ₀ ℓ₁ ℓ₂ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} {Z : 𝕆 ℓ₂ n} 88 | → (Y ⇒ Z) → (X ⇒ Y) → (X ⇒ Z) 89 | 90 | postulate 91 | 92 | ⊚-Frm : ∀ {ℓ₀ ℓ₁ ℓ₂ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} {Z : 𝕆 ℓ₂ n} 93 | → (α : Y ⇒ Z) (β : X ⇒ Y) (o : 𝒪 n) (f : Frm X o) 94 | → Frm⇒ (α ⊚ β) f ↦ Frm⇒ α (Frm⇒ β f) 95 | {-# REWRITE ⊚-Frm #-} 96 | 97 | ⊚-assoc : ∀ {ℓ₀ ℓ₁ ℓ₂ ℓ₃ n} 98 | → {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} {Z : 𝕆 ℓ₂ n} {W : 𝕆 ℓ₃ n} 99 | → (α : Z ⇒ W) (β : Y ⇒ Z) (γ : X ⇒ Y) 100 | → (α ⊚ β) ⊚ γ ↦ α ⊚ β ⊚ γ 101 | {-# REWRITE ⊚-assoc #-} 102 | 103 | -- And unit laws ... 104 | 105 | _⊚_ {n = zero} α β = lift tt 106 | _⊚_ {n = suc n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} {Zₙ , Zₛₙ} (αₙ , αₛₙ) (βₙ , βₛₙ) = 107 | αₙ ⊚ βₙ , λ x → αₛₙ (βₛₙ x) 108 | 109 | -- 110 | -- Equality of functions 111 | -- 112 | 113 | -- _⇒-≡_ : ∀ {ℓ₀ ℓ₁ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} 114 | -- → (X ⇒ Y) → (X ⇒ Y) → Type (ℓ-max ℓ₀ ℓ₁) 115 | 116 | -- postulate 117 | 118 | -- ⇒-≡-to-≡ : ∀ {ℓ₀ ℓ₁ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} 119 | -- → (α : X ⇒ Y) (β : X ⇒ Y) 120 | -- → α ⇒-≡ β → α ≡ β 121 | 122 | -- _⇒-≡_ {n = O} {X} {Y} α β = ⊤ 123 | -- _⇒-≡_ {n = S n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} (αₙ , αₛₙ) (βₙ , βₛₙ) = 124 | -- Σ (αₙ ⇒-≡ βₙ) (λ e → {o : 𝒪 n} {f : Frm Xₙ o} 125 | -- (x : Xₛₙ f) → αₛₙ x ≡ βₛₙ x [ (λ ϕ → Yₛₙ (Frm⇒ ϕ f)) ↓ ⇒-≡-to-≡ αₙ βₙ e ] ) 126 | 127 | -- -- And this last thing is by fun-ext. 128 | -- -- ⇒-≡-to-≡ {n = O} {X} {Y} α β e = refl 129 | -- -- ⇒-≡-to-≡ {n = S n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} (αₙ , αₛₙ) (βₙ , βₛₙ) (eₙ , eₛₙ) = 130 | -- -- po-to-Σ (⇒-≡-to-≡ αₙ βₙ eₙ) {!!} 131 | 132 | -------------------------------------------------------------------------------- /Experimental/wip/OpetopicSigma.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicSigma - Sigma of Opetopic Types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Nat 8 | 9 | open import Prelude 10 | open import Opetopes 11 | open import OpetopicType 12 | open import OpetopicFam 13 | open import OpetopicTerm 14 | open import OpetopicSub 15 | open import OpetopicExt 16 | 17 | module OpetopicSigma where 18 | 19 | Σₒ : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} 20 | → (X : 𝕆Fam Γ ℓ₁) (Y : 𝕆Fam (Ext Γ X) ℓ₂) 21 | → 𝕆Fam Γ (ℓ-max ℓ₀ ℓ₁) 22 | 23 | Frm-fst : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} 24 | → (X : 𝕆Fam Γ ℓ₁) (Y : 𝕆Fam (Ext Γ X) ℓ₂) 25 | → {𝑜 : 𝒪 n} {f : Frm Γ 𝑜} 26 | → Frm↓ (Σₒ X Y) f → Frm↓ X f 27 | 28 | Frm-snd : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} 29 | → (X : 𝕆Fam Γ ℓ₁) (Y : 𝕆Fam (Ext Γ X) ℓ₂) 30 | → {𝑜 : 𝒪 n} {f : Frm Γ 𝑜} 31 | → (f↓ : Frm↓ (Σₒ X Y) f) 32 | → Frm↓ Y {!Frm-fst X Y f↓!} 33 | 34 | Σₒ {zero} X Y = lift tt 35 | Σₒ {suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) = 36 | Σₒ Xₙ Yₙ , λ {f} f↓ γ → Σ[ x ∈ Xₛₙ (Frm-fst Xₙ Yₙ f↓) γ ] {!!} 37 | 38 | Frm-fst X Y f↓ = {!!} 39 | Frm-snd X Y f↓ = {!!} 40 | 41 | -- Hmmm. Name clash with opetopes ... 42 | Σₒ-pair : ∀ {n ℓ₀ ℓ₁ ℓ₂} {Γ : 𝕆Type n ℓ₀} 43 | → {X : 𝕆Fam Γ ℓ₁} {Y : 𝕆Fam (Ext Γ X) ℓ₂} 44 | → (x : 𝕆Term X) (y : 𝕆Term (Y [ ext-sub x ]ty)) 45 | → 𝕆Term (Σₒ X Y) 46 | Σₒ-pair = {!!} 47 | -------------------------------------------------------------------------------- /Experimental/wip/OpetopicTerm.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicType.agda - Definition of Opetopic Types in Context 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Sum 9 | 10 | open import Prelude 11 | open import Opetopes 12 | open import OpetopicType 13 | open import OpetopicFam 14 | 15 | module OpetopicTerm where 16 | 17 | 𝕆Term : ∀ {n ℓ₀ ℓ₁} {Γ : 𝕆Type n ℓ₀} (X : 𝕆Fam Γ ℓ₁) 18 | → Type (ℓ-max ℓ₀ ℓ₁) 19 | 20 | Frm-Tm : ∀ {n ℓ₀ ℓ₁} {Γ : 𝕆Type n ℓ₀} {X : 𝕆Fam Γ ℓ₁} (σ : 𝕆Term X) 21 | → {𝑜 : 𝒪 n} (f : Frm Γ 𝑜) → Frm↓ X f 22 | 23 | -- usual fix 24 | {-# TERMINATING #-} 25 | Cns-Tm : ∀ {n ℓ₀ ℓ₁} {Γ : 𝕆Type n ℓ₀} {X : 𝕆Fam Γ ℓ₁} (σ : 𝕆Term X) 26 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 27 | → {f : Frm Γ 𝑜} (c : Cns Γ f 𝑝) 28 | → Cns↓ X (Frm-Tm σ f) c 29 | 30 | postulate 31 | 32 | Shp-Tm : ∀ {n ℓ₀ ℓ₁} {Γ : 𝕆Type n ℓ₀} {X : 𝕆Fam Γ ℓ₁} (σ : 𝕆Term X) 33 | → {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 34 | → {f : Frm Γ 𝑜} (c : Cns Γ f 𝑝) (p : Pos 𝑝) 35 | → Frm-Tm σ (Shp Γ c p) ↦ Shp↓ X (Cns-Tm σ c) p 36 | {-# REWRITE Shp-Tm #-} 37 | 38 | η-Tm : ∀ {n ℓ₀ ℓ₁} {Γ : 𝕆Type n ℓ₀} {X : 𝕆Fam Γ ℓ₁} (σ : 𝕆Term X) 39 | → {𝑜 : 𝒪 n} (f : Frm Γ 𝑜) 40 | → Cns-Tm σ (η Γ f) ↦ η↓ X (Frm-Tm σ f) 41 | {-# REWRITE η-Tm #-} 42 | 43 | μ-Tm : ∀ {n ℓ₀ ℓ₁} {Γ : 𝕆Type n ℓ₀} {X : 𝕆Fam Γ ℓ₁} (σ : 𝕆Term X) 44 | → {𝑜 : 𝒪 n} {f : Frm Γ 𝑜} 45 | → {𝑝 : 𝒫 𝑜} (c : Cns Γ f 𝑝) 46 | → {𝑞 : (p : Pos 𝑝) → 𝒫 (Typ 𝑝 p)} 47 | → (d : (p : Pos 𝑝) → Cns Γ (Shp Γ c p) (𝑞 p)) 48 | → Cns-Tm σ (μ Γ c d) ↦ μ↓ X (Cns-Tm σ c) (λ p → Cns-Tm σ (d p)) 49 | {-# REWRITE μ-Tm #-} 50 | 51 | 𝕆Term {zero} X = Lift Unit 52 | 𝕆Term {suc n} {Γ = Γₙ , Γₛₙ} (Xₙ , Xₛₙ) = 53 | Σ[ xₙ ∈ 𝕆Term Xₙ ] 54 | ({𝑜 : 𝒪 n} {f : Frm Γₙ 𝑜} (γ : Γₛₙ f) → Xₛₙ (Frm-Tm xₙ f ) γ) 55 | 56 | Frm-Tm {zero} σ f = lift tt 57 | Frm-Tm {suc n} {X = Xₙ , Xₛₙ} (σₙ , σₛₙ) (f , x , c , y) = 58 | Frm-Tm σₙ f , σₛₙ x , Cns-Tm σₙ c , λ p → σₛₙ (y p) 59 | 60 | Cns-Tm {zero} σ c = lift tt 61 | Cns-Tm {suc n} (σₙ , σₛₙ) (lf x) = idp 62 | Cns-Tm {suc n} (σₙ , σₛₙ) (nd x c y d z ψ) = 63 | Cns-Tm σₙ c , σₛₙ ∘ y , Cns-Tm σₙ ∘ d , (λ p q → σₛₙ (z p q)) , 64 | (λ p → Cns-Tm (σₙ , σₛₙ) (ψ p)) , idp 65 | 66 | 67 | 68 | -------------------------------------------------------------------------------- /Experimental/wip/OpetopicType2.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | -- 3 | -- OpetopicTypes2.agda - Basic Constructions on Opetopic Types 4 | -- 5 | 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Data.Sigma 8 | open import Cubical.Data.Unit 9 | open import Cubical.Data.Nat 10 | 11 | open import Prelude 12 | open import Opetopes 13 | open import OpetopicType 14 | open import OpetopicMap 15 | 16 | module OpetopicType2 where 17 | 18 | -- The product of two opetopic types. 19 | infixl 45 _×ₒ_ 20 | 21 | _×ₒ_ : ∀ {ℓ₀ ℓ₁ n} (X : 𝕆 ℓ₀ n) (Y : 𝕆 ℓ₁ n) → 𝕆 (ℓ-max ℓ₀ ℓ₁) n 22 | π₀ : ∀ {ℓ₀ ℓ₁ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} → X ×ₒ Y ⇒ X 23 | π₁ : ∀ {ℓ₀ ℓ₁ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} → X ×ₒ Y ⇒ Y 24 | 25 | _×ₒ_ {n = zero} X Y = lift tt 26 | _×ₒ_ {n = suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) = 27 | Xₙ ×ₒ Yₙ , λ f → Xₛₙ (Frm⇒ π₀ f) × Yₛₙ (Frm⇒ π₁ f) 28 | 29 | π₀ {n = zero} {X} {Y} = lift tt 30 | π₀ {n = suc n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} = 31 | π₀ {X = Xₙ} {Y = Yₙ} , fst 32 | 33 | π₁ {n = zero} {X} {Y} = lift tt 34 | π₁ {n = suc n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} = 35 | π₁ {X = Xₙ} {Y = Yₙ} , snd 36 | 37 | -- Hmm. So you gave the elim principles. But what about the intro? 38 | 39 | -- Frm-pair : ∀ {ℓ₀ ℓ₁ n} (X : 𝕆 ℓ₀ n) (Y : 𝕆 ℓ₁ n) 40 | -- → {o : 𝒪 n} (fX : Frm X o) (fY : Frm Y o) 41 | -- → Frm (X ×ₒ Y) o 42 | -- Frm-pair {n = zero} X Y fx fy = lift tt 43 | -- Frm-pair {n = suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) fx fy = 44 | -- ⟪ Frm-pair Xₙ Yₙ (frm fx) (frm fy) , {!!} , ({!tgt fx!} , {!!}) , {!!} ⟫ 45 | 46 | -- The pullback of two opetopic types 47 | Pb : ∀ {ℓ₀ ℓ₁ ℓ₂ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} {Z : 𝕆 ℓ₂ n} 48 | → (f : X ⇒ Z) (g : Y ⇒ Z) → 𝕆 (ℓ-max (ℓ-max ℓ₀ ℓ₁) ℓ₂) n 49 | 50 | pb-fst : ∀ {ℓ₀ ℓ₁ ℓ₂ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} {Z : 𝕆 ℓ₂ n} 51 | → (f : X ⇒ Z) (g : Y ⇒ Z) → Pb f g ⇒ X 52 | 53 | pb-snd : ∀ {ℓ₀ ℓ₁ ℓ₂ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} {Z : 𝕆 ℓ₂ n} 54 | → (f : X ⇒ Z) (g : Y ⇒ Z) → Pb f g ⇒ Y 55 | 56 | pb-comm : ∀ {ℓ₀ ℓ₁ ℓ₂ n} {X : 𝕆 ℓ₀ n} {Y : 𝕆 ℓ₁ n} {Z : 𝕆 ℓ₂ n} 57 | → (f : X ⇒ Z) (g : Y ⇒ Z) 58 | → f ⊚ pb-fst f g ≡ g ⊚ pb-snd f g 59 | 60 | Pb {n = zero} f g = lift tt 61 | Pb {n = suc n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} {Zₙ , Zₛₙ} (fₙ , fₛₙ) (gₙ , gₛₙ) = 62 | Pb fₙ gₙ , λ {o} φ → 63 | Σ (Xₛₙ (Frm⇒ (pb-fst fₙ gₙ) φ)) (λ x → 64 | Σ (Yₛₙ (Frm⇒ (pb-snd fₙ gₙ) φ)) (λ y → 65 | PathP (λ i → Zₛₙ (Frm⇒ (pb-comm fₙ gₙ i) φ)) (fₛₙ x) (gₛₙ y))) 66 | 67 | pb-fst {n = zero} _ _ = lift tt 68 | pb-fst {n = suc n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} {Zₙ , Zₛₙ} (fₙ , fₛₙ) (gₙ , gₛₙ) = 69 | pb-fst fₙ gₙ , λ { (x , _ , _) → x } 70 | 71 | pb-snd {n = zero} _ _ = lift tt 72 | pb-snd {n = suc n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} {Zₙ , Zₛₙ} (fₙ , fₛₙ) (gₙ , gₛₙ) = 73 | pb-snd fₙ gₙ , λ { (_ , y , _) → y } 74 | 75 | pb-comm {n = zero} _ _ = refl 76 | pb-comm {n = suc n} {Xₙ , Xₛₙ} {Yₙ , Yₛₙ} {Zₙ , Zₛₙ} (fₙ , fₛₙ) (gₙ , gₛₙ) = 77 | λ i → pb-comm fₙ gₙ i , λ {o} {φ} → λ { (x , y , e) → e i} 78 | 79 | -- 80 | -- What is a global section? 81 | -- 82 | 83 | Sec : ∀ {ℓ n} (X : 𝕆 ℓ n) → Type ℓ 84 | 85 | Frm-Sec : ∀ {ℓ n} {X : 𝕆 ℓ n} (σ : Sec X) 86 | → (o : 𝒪 n) → Frm X o 87 | 88 | Cns-Sec : ∀ {ℓ n} {X : 𝕆 ℓ n} (σ : Sec X) 89 | → {o : 𝒪 n} (ρ : 𝒫 o) 90 | → Cns X (Frm-Sec σ o) ρ 91 | 92 | postulate 93 | 94 | Shp-Frm-Cns : ∀ {ℓ n} (X : 𝕆 ℓ n) (σ : Sec X) 95 | → {o : 𝒪 n} (ρ : 𝒫 o) (p : Pos ρ) 96 | → Frm-Sec σ (Typ ρ p) ↦ Shp X (Cns-Sec σ ρ) p 97 | {-# REWRITE Shp-Frm-Cns #-} 98 | 99 | η-Sec : ∀ {ℓ n} (X : 𝕆 ℓ n) (σ : Sec X) 100 | → (o : 𝒪 n) 101 | → η X (Frm-Sec σ o) ↦ Cns-Sec σ (ηₒ o) 102 | {-# REWRITE η-Sec #-} 103 | 104 | Sec {n = zero} X = Lift Unit 105 | Sec {n = suc n} (Xₙ , Xₛₙ) = 106 | Σ (Sec Xₙ) (λ σ → (o : 𝒪 n) → Xₛₙ (Frm-Sec σ o)) 107 | 108 | Frm-Sec {n = zero} σ o = lift tt 109 | Frm-Sec {n = suc n} {Xₙ , Xₛₙ} (σₙ , σₛₙ) (o , ρ) = 110 | ⟪ Frm-Sec σₙ o , Cns-Sec σₙ ρ , σₛₙ o , (λ p → σₛₙ (Typ ρ p)) ⟫ 111 | 112 | Cns-Sec {n = zero} σ ρ = lift tt 113 | Cns-Sec {n = suc n} {Xₙ , Xₛₙ} (σₙ , σₛₙ) {o , .(ηₒ o)} (lfₒ .o) = lf (σₛₙ o) 114 | Cns-Sec {n = suc n} σ {o , .(μₒ ρ δ)} (ndₒ .o ρ δ ε) = {!!} 115 | 116 | -------------------------------------------------------------------------------- /Experimental/wip/Representable.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --rewriting #-} 2 | 3 | open import MiniHoTT 4 | 5 | open import Opetopes 6 | open import OpetopicType 7 | open import OpetopicMap 8 | 9 | module Representable where 10 | 11 | Rep : {n : ℕ} (o : 𝒪 n) → 𝕆 ℓ-zero (S n) 12 | 13 | SrcInc : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → Rep (Typ ρ p) ⇒ fst (Rep (o , ρ)) 14 | TgtInc : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → Rep o ⇒ fst (Rep (o , ρ)) 15 | 16 | 17 | -- Hmmm. This is not quite right. It is the top *3* fibrations 18 | -- which must change. 19 | 20 | -- So, the cells of the first are parameterized by two things 21 | -- 1) the positions of ρ (there will be the externals) 22 | -- 2) the positions of τ (this will be the internals) 23 | 24 | -- In the next fibration, we have exactly the positions of τ and the target. 25 | 26 | -- And then we have the unique next dimension. 27 | 28 | -- Right. So this should actually lead to another, slight different 29 | -- presentation of the faces above. 30 | 31 | data ArrowFill : {o : 𝒪 1} → WebFrm tt (cst (⊤ ⊔ ⊤)) (fst o) (snd o) → Set where 32 | top : ArrowFill ⟪ tt , tt , inl tt , (λ _ → inr tt) ⟫ 33 | 34 | data Pivot {n} (o : 𝒪 n) (ρ : 𝒫 o) (τ : 𝒯r o ρ) : {q : 𝒪 n} → Frm (fst (fst (Rep (o , ρ)))) q → Set where 35 | 36 | -- Does this make sense? Hmmm. But don't you want the frame to be full? 37 | -- So the pair should make this a complete guy. So that looks okay. 38 | ext : (p : Pos ρ) (f : Frm (fst (Rep (Typ ρ p))) (Typ ρ p)) (c : snd (Rep (Typ ρ p)) f) 39 | → Pivot o ρ τ (Frm⇒ (fst (SrcInc o ρ p)) f) 40 | 41 | int : (p : Pos τ) (f : Frm (fst (Rep (fst (𝒯rTyp τ p)))) (fst (𝒯rTyp τ p))) (c : snd (Rep (fst (𝒯rTyp τ p))) f) 42 | → Pivot o ρ τ {!TgtInc o ρ !} 43 | 44 | Rep {n = O} tt = tt , cst ⊤ 45 | Rep {n = S O} (tt , tt) = 46 | (tt , cst (⊤ ⊔ ⊤)) , ArrowFill 47 | 48 | Rep {n = S (S n)} ((o , ρ) , τ) = ((fst (fst (Rep (o , ρ))) , Pivot o ρ τ) , {!!}) , {!!} 49 | 50 | -- Rep {n = O} o = tt , cst ⊤ 51 | -- Rep {n = S n} (o , ρ) = (fst (Rep o) , BoundaryCell o ρ) , FillingCell o ρ 52 | 53 | SrcInc = {!!} 54 | TgtInc = {!!} 55 | 56 | 57 | -- data LfBdry {n : ℕ} (m : 𝒪 n) : {o : 𝒪 (S n)} 58 | -- → WebFrm (fst (Rep m)) (snd (Rep m)) (fst o) (snd o) → Set where 59 | -- lf-tgt : LfBdry m ⟪ {!!} , {!η (fst (Rep m))!} , {!!} , {!!} ⟫ 60 | 61 | 62 | -- Rep {n = S (S n)} ((x , .(ηₒ x)) , lfₒ .x) = (Rep x , LfBdry x) , {!!} 63 | -- Rep {n = S (S n)} ((o , .(μₒ ρ δ)) , ndₒ .o ρ δ ε) = {!!} 64 | 65 | -------------------------------------------------------------------------------- /Experimental/wip/TheUniverse.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- TheUniverse.agda - The Internal Universe in Opetopic Types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Nat 8 | 9 | open import Prelude 10 | open import Opetopes 11 | open import OpetopicType 12 | open import OpetopicType2 13 | open import OpetopicMap 14 | 15 | module TheUniverse where 16 | 17 | 𝒰ₒ : (ℓ : Level) (n : ℕ) → 𝕆 (ℓ-suc ℓ) n 18 | 𝒱ₒ : (ℓ : Level) (n : ℕ) → 𝕆 (ℓ-suc ℓ) n 19 | 𝓊 : (ℓ : Level) (n : ℕ) → 𝒱ₒ ℓ n ⇒ 𝒰ₒ ℓ n 20 | 21 | 𝒰ₒ ℓ zero = lift tt 22 | 𝒰ₒ ℓ (suc n) = 𝒰ₒ ℓ n , λ {o} f → 23 | (f' : Frm (𝒱ₒ ℓ n) o) (e : Frm⇒ (𝓊 ℓ n) f' ≡ f) 24 | → Type ℓ 25 | 26 | 𝒱ₒ ℓ zero = lift tt 27 | 𝒱ₒ ℓ (suc n) = 𝒱ₒ ℓ n , λ {o} f* → 28 | Σ ((f' : Frm (𝒱ₒ ℓ n) o) (e : Frm⇒ (𝓊 ℓ n) f' ≡ Frm⇒ (𝓊 ℓ n) f*) → Type ℓ) (λ F → F f* refl) 29 | 30 | 𝓊 ℓ zero = lift tt 31 | 𝓊 ℓ (suc n) = 𝓊 ℓ n , λ {o} {f} X → fst X 32 | 33 | -- 34 | -- Σ of opetopic types 35 | -- 36 | 37 | Σₒ : ∀ {ℓ₀ ℓ₁ n} (X : 𝕆 ℓ₀ n) (Y : X ⇒ 𝒰ₒ ℓ₁ n) → 𝕆 (ℓ-max ℓ₀ ℓ₁) n 38 | 39 | fstₒ : ∀ {ℓ₀ ℓ₁ n} (X : 𝕆 ℓ₀ n) (Y : X ⇒ 𝒰ₒ ℓ₁ n) 40 | → Σₒ X Y ⇒ X 41 | 42 | {-# TERMINATING #-} 43 | sndₒ : ∀ {ℓ₀ ℓ₁ n} (X : 𝕆 ℓ₀ n) (Y : X ⇒ 𝒰ₒ ℓ₁ n) 44 | → Σₒ X Y ⇒ 𝒱ₒ ℓ₁ n 45 | 46 | commₒ : ∀ {ℓ₀ ℓ₁ n} (X : 𝕆 ℓ₀ n) (Y : X ⇒ 𝒰ₒ ℓ₁ n) 47 | → 𝓊 ℓ₁ n ⊚ sndₒ X Y ≡ Y ⊚ fstₒ X Y 48 | 49 | Σₒ {ℓ₀} {ℓ₁} {zero} X Y = lift tt 50 | Σₒ {ℓ₀} {ℓ₁} {suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) = 51 | Σₒ Xₙ Yₙ , λ φ → Σ (Xₛₙ (Frm⇒ (fstₒ Xₙ Yₙ) φ)) (λ x → 52 | Yₛₙ x (Frm⇒ (sndₒ Xₙ Yₙ) φ) (λ i → Frm⇒ (commₒ Xₙ Yₙ i) φ)) 53 | 54 | fstₒ {n = zero} X Y = lift tt 55 | fstₒ {n = suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) = 56 | fstₒ Xₙ Yₙ , fst 57 | 58 | sndₒ {n = zero} X Y = lift tt 59 | sndₒ {n = suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) = 60 | sndₒ Xₙ Yₙ , λ { {o} {f} (x , y) → (λ φ eq → Yₛₙ x φ (eq ∙ (λ i → Frm⇒ (commₒ Xₙ Yₙ i) f))) , {!!} } 61 | 62 | commₒ = {!!} 63 | 64 | -- 65 | -- Can we do Π ? 66 | -- 67 | 68 | Πₒ : ∀ {ℓ₀ ℓ₁ n} (X : 𝕆 ℓ₀ n) (Y : X ⇒ 𝒰ₒ ℓ₁ n) → 𝕆 (ℓ-max ℓ₀ ℓ₁) n 69 | 70 | evₒ : ∀ {ℓ₀ ℓ₁ n} (X : 𝕆 ℓ₀ n) (Y : X ⇒ 𝒰ₒ ℓ₁ n) 71 | → (Πₒ X Y) ×ₒ X ⇒ 𝒱ₒ ℓ₁ n 72 | 73 | Πₒ {n = zero} X Y = lift tt 74 | Πₒ {n = suc n} (Xₙ , Xₛₙ) (Yₙ , Yₛₙ) = 75 | Πₒ Xₙ Yₙ , λ φ → {o : 𝒪 n} {f : Frm Xₙ o} (x : Xₛₙ f) → Yₛₙ x {!Frm⇒ (evₒ Xₙ Yₙ)!} {!!} 76 | 77 | evₒ = {!!} 78 | 79 | -- 80 | -- A first test: suppose I have an opetopic type. Does this 81 | -- determine a point of the universe? 82 | -- 83 | 84 | -- classify : ∀ {ℓ} (n : ℕ) → 𝕆 ℓ n → 𝕋 {ℓ} n ⇒ 𝒰ₒ {ℓ} n 85 | 86 | -- unclassify : ∀ {ℓ} (n : ℕ) (X : 𝕆 ℓ n) {o : 𝒪 n} 87 | -- → (t : Frm (𝕋 n) o) (f : Frm (𝒱ₒ n) o) 88 | -- → (e : Frm⇒ (𝓊 n) f ≡ Frm⇒ (classify n X) t) 89 | -- → Frm X o 90 | 91 | -- classify O _ = tt 92 | -- classify (S n) (Xₙ , Xₛₙ) = classify n Xₙ , 93 | -- λ {o} {f} _ f' e → Xₛₙ (unclassify n Xₙ f f' e) 94 | 95 | -- unclassify O X t f e = tt 96 | -- unclassify (S n) (Xₙ , Xₛₙ) t f e = {!!} 97 | 98 | -- Yeah, so you need to see that pullbacks can be computed pointwise 99 | -- in trees. But I think characterizing the identity types of frames 100 | -- and trees and so on will be done for a fixed n. So I don't see 101 | -- that this will necessarily have any coherence problems. 102 | 103 | -- 𝒰ₒ' : (ℓ : Level) (n : ℕ) → 𝕆 (ℓ-suc ℓ) n 104 | 105 | -- ℯ𝓁 : (ℓ : Level) (n : ℕ) {o : 𝒪 n} (f : Frm (𝒰ₒ' ℓ n) o) 106 | -- → Type ℓ 107 | 108 | -- ℯ𝓁-cns : (ℓ : Level) (n : ℕ) {o : 𝒪 n} {ρ : 𝒫 o} 109 | -- → (f : Frm (𝒰ₒ' ℓ n) o) (e : ℯ𝓁 ℓ n f) 110 | -- → (c : Cns (𝒰ₒ' ℓ n) f ρ) → Type ℓ 111 | 112 | -- ℯ𝓁-bdy : (ℓ : Level) (n : ℕ) {o : 𝒪 n} {ρ : 𝒫 o} 113 | -- → (f : Frm (𝒰ₒ' ℓ n) o) (e : ℯ𝓁 ℓ n f) 114 | -- → (c : Cns (𝒰ₒ' ℓ n) f ρ) (σ : ℯ𝓁-cns ℓ n f e c) 115 | -- → (p : Pos ρ) → ℯ𝓁 ℓ n (Shp (𝒰ₒ' ℓ n) c p) 116 | 117 | -- 𝒰ₒ' ℓ zero = lift tt 118 | -- 𝒰ₒ' ℓ (suc n) = 𝒰ₒ' ℓ n , λ {o} φ → ℯ𝓁 ℓ n φ → Type ℓ 119 | 120 | -- ℯ𝓁 ℓ zero f = Lift Unit 121 | -- ℯ𝓁 ℓ (suc n) {o , ρ} ⟪ f , c , τ , σ ⟫ = 122 | -- Σ (ℯ𝓁 ℓ n f) (λ e → 123 | -- Σ (ℯ𝓁-cns ℓ n f e c) (λ b → τ e × ((p : Pos ρ) → σ p (ℯ𝓁-bdy ℓ n f e c b p)))) 124 | 125 | -- ℯ𝓁-cns ℓ zero f e c = Lift Unit 126 | -- ℯ𝓁-cns ℓ (suc n) ⟪ f , ._ , X , ._ ⟫ (e₀ , e₁ , l , r) (lf {o} .X) = {!!} -- r (ηₒ-pos o) ≡ l 127 | -- ℯ𝓁-cns ℓ (suc n) ._ (e₀ , e₁ , s , t) (nd φ ι κ δ ν ε) = {!!} 128 | 129 | -- ℯ𝓁-bdy = {!!} 130 | -------------------------------------------------------------------------------- /Experimental/wip/YonedaExtension.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --rewriting --guardedness #-} 2 | 3 | open import MiniHoTT 4 | 5 | open import Opetopes 6 | open import OpetopicType 7 | open import OpetopicMap 8 | 9 | module YonedaExtension where 10 | 11 | record 𝕆∞ {ℓ n} (X : 𝕆 ℓ n) : Set (ℓ-suc ℓ) where 12 | coinductive 13 | field 14 | Fill : {o : 𝒪 n} → Frm X o → Set ℓ 15 | Hom : 𝕆∞ (X , Fill) 16 | 17 | open 𝕆∞ 18 | 19 | record [_⇒_↓_] {ℓ n} {X Y : 𝕆 ℓ n} (X∞ : 𝕆∞ X) (Y∞ : 𝕆∞ Y) (α : X ⇒ Y) : Set ℓ where 20 | coinductive 21 | field 22 | Fill⇒ : {o : 𝒪 n} {f : Frm X o} → Fill X∞ f → Fill Y∞ (Frm⇒ α f) 23 | Hom⇒ : [ Hom X∞ ⇒ Hom Y∞ ↓ α , Fill⇒ ] 24 | 25 | 26 | OType : (ℓ : Level) → Set (ℓ-suc ℓ) 27 | OType ℓ = 𝕆∞ tt 28 | 29 | -- Oh, but maybe this is not general enough for the applications 30 | -- you have in mind... Should D take values in opetopic types, not 31 | -- just truncated ones? 32 | 33 | -- truncated version ... 34 | -- record CoOpetopicDiagram (ℓ : Level) : Set (ℓ-suc ℓ) where 35 | -- field 36 | -- D : {n : ℕ} (o : 𝒪 n) → 𝕆 ℓ (S n) 37 | -- σ : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → D (Typ ρ p) ⇒ fst (D (o , ρ)) 38 | -- τ : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → D o ⇒ fst (D (o , ρ)) 39 | 40 | record CoOpetopicDiagram (ℓ : Level) : Set (ℓ-suc ℓ) where 41 | field 42 | D : {n : ℕ} (o : 𝒪 n) → 𝕆∞ {ℓ} tt 43 | σ : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → [ D (Typ ρ p) ⇒ D (o , ρ) ↓ tt ] 44 | τ : {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) (p : Pos ρ) → [ D o ⇒ D (o , ρ) ↓ tt ] 45 | 46 | open CoOpetopicDiagram 47 | 48 | Right : ∀ {ℓ} (𝒟 : CoOpetopicDiagram ℓ) (X : 𝕆∞ {ℓ} tt) 49 | → (n : ℕ) → 𝕆 ℓ n 50 | 51 | Restrict : ∀ {ℓ} (𝒟 : CoOpetopicDiagram ℓ) (X : 𝕆∞ {ℓ} tt) 52 | → {n : ℕ} (o : 𝒪 n) (ρ : 𝒫 o) → [ D 𝒟 (o , ρ) ⇒ X ↓ tt ] → Frm (Right 𝒟 X n) o 53 | 54 | Right 𝒟 X O = tt 55 | Right 𝒟 X (S O) = {!!} 56 | Right 𝒟 X (S (S n)) = Right 𝒟 X (S n) , {!!} 57 | 58 | Restrict = {!!} 59 | 60 | -- Maybe this sould be like a Π in that it comes with a kind of 61 | -- evaluation function? 62 | 63 | -------------------------------------------------------------------------------- /Lib/FibrantReplacement.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- FibrantReplacement.agda - Consructing the Fibrant Replacement of an Opetopic Type 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Nat 8 | 9 | open import Core.Everything 10 | 11 | module Lib.FibrantReplacement where 12 | 13 | -- 14 | -- The free groupoid generated by an opetopic type 15 | -- 16 | 17 | FreeGrp : ∀ {ℓ} (X : 𝕆Type∞ {ℓ = ℓ} tt*) 18 | → (n : ℕ) → 𝕆Type n ℓ 19 | 20 | data FreeCell {ℓ} (X : 𝕆Type∞ {ℓ = ℓ} tt*) : {n : ℕ} {𝑜 : 𝒪 n} (f : Frm (FreeGrp X n) 𝑜) → Type ℓ 21 | 22 | FreeInc : ∀ {ℓ} (X : 𝕆Type∞ {ℓ = ℓ} tt*) 23 | → (n : ℕ) → Skeleton X n ⇒ FreeGrp X n 24 | 25 | FreeGrp X zero = tt* 26 | FreeGrp X (suc n) = FreeGrp X n , FreeCell X 27 | 28 | data FreeCell {ℓ} X where 29 | 30 | free-in : {n : ℕ} {𝑜 : 𝒪 n} {f : Frm (Skeleton X n) 𝑜} 31 | → (x : Fill (SkeletonExt X n) f) 32 | → FreeCell X (Frm⇒ (FreeInc X n) f) 33 | 34 | free-comp : {n : ℕ} {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 35 | → {f : Frm (FreeGrp X n) 𝑜} (c : Cns (FreeGrp X n) f 𝑝) 36 | → (y : (p : Pos 𝑝) → FreeCell X (Shp (FreeGrp X n) c p)) 37 | → FreeCell X f 38 | 39 | free-fill : {n : ℕ} {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 40 | → {f : Frm (FreeGrp X n) 𝑜} (c : Cns (FreeGrp X n) f 𝑝) 41 | → (y : (p : Pos 𝑝) → FreeCell X (Shp (FreeGrp X n) c p)) 42 | → FreeCell X {suc n} {𝑜 ∣ 𝑝} (f , free-comp c y , c , y) 43 | 44 | free-comp-unique : {n : ℕ} {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 45 | → {f : Frm (FreeGrp X n) 𝑜} (c : Cns (FreeGrp X n) f 𝑝) 46 | → (y : (p : Pos 𝑝) → FreeCell X (Shp (FreeGrp X n) c p)) 47 | → (x : FreeCell X f) (α : FreeCell X {𝑜 = 𝑜 ∣ 𝑝} (f , x , c , y)) 48 | → free-comp c y ≡ x 49 | 50 | free-fill-unique : {n : ℕ} {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 51 | → {f : Frm (FreeGrp X n) 𝑜} (c : Cns (FreeGrp X n) f 𝑝) 52 | → (y : (p : Pos 𝑝) → FreeCell X (Shp (FreeGrp X n) c p)) 53 | → (x : FreeCell X f) (α : FreeCell X {𝑜 = 𝑜 ∣ 𝑝} (f , x , c , y)) 54 | → (λ i → FreeCell X {𝑜 = 𝑜 ∣ 𝑝} (f , free-comp-unique c y x α i , c , y)) 55 | [ free-fill c y ≡ α ] 56 | 57 | FreeInc X zero = tt* 58 | FreeInc X (suc n) = FreeInc X n , free-in 59 | -------------------------------------------------------------------------------- /Lib/Groupoid.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Groupoid.agda - Infinity Groupoids 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Empty 8 | open import Cubical.Data.Nat 9 | open import Cubical.Data.Sum 10 | 11 | open import Core.Everything 12 | open import Lib.Structures 13 | 14 | module Lib.Groupoid where 15 | 16 | Grp : ∀ {ℓ} (X : Type ℓ) → (n : ℕ) → 𝕆Type n ℓ 17 | Pt : ∀ {n ℓ} {X : Type ℓ} (x : X) → El (Grp X n) 18 | 19 | data GrpCell {n ℓ} (X : Type ℓ) : {𝑜 : 𝒪 n} → Frm (Grp X n) 𝑜 → Type ℓ where 20 | reflₒ : (x : X) (𝑜 : 𝒪 n) → GrpCell X (Frm-El (Pt x) 𝑜) 21 | 22 | Grp X zero = tt* 23 | Grp X (suc n) = Grp X n , GrpCell X 24 | 25 | Pt {zero} x = tt* 26 | Pt {suc n} x = Pt {n} x , reflₒ x 27 | 28 | -------------------------------------------------------------------------------- /Lib/Join.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Join.agda - The Join of Opetopic Types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Nat 7 | open import Cubical.Data.Sum 8 | open import Cubical.Data.Empty 9 | 10 | open import Core.Everything 11 | 12 | module Lib.Join where 13 | 14 | Join : ∀ {n ℓ₀ ℓ₁} (X : 𝕆Type n ℓ₀) (Y : 𝕆Type n ℓ₁) 15 | → 𝕆Type n (ℓ-max ℓ₀ ℓ₁) 16 | 17 | join-inl : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 18 | → X ⇒ Join X Y 19 | 20 | join-inr : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type n ℓ₀} {Y : 𝕆Type n ℓ₁} 21 | → Y ⇒ Join X Y 22 | 23 | is-traversing : ∀ {n ℓ₀ ℓ₁} {X : 𝕆Type (2 + n) ℓ₀} {Y : 𝕆Type (2 + n) ℓ₁} 24 | → {𝑜 : 𝒪 (2 + n)} (f : Frm (Join X Y) 𝑜) → Type 25 | 26 | data JoinArrow {ℓ₀ ℓ₁} (X : 𝕆Type 2 ℓ₀) (Y : 𝕆Type 2 ℓ₁) : {𝑜 : 𝒪 1} 27 | → Frm (tt* , λ _ → snd (fst X) tt* ⊎ snd (fst Y) tt*) 𝑜 28 | → Type (ℓ-max ℓ₀ ℓ₁) 29 | 30 | data JoinCell {n ℓ₀ ℓ₁} (X : 𝕆Type (3 + n) ℓ₀) (Y : 𝕆Type (3 + n) ℓ₁) 31 | : {𝑜 : 𝒪 (2 + n)} → Frm (Join (fst X) (fst Y)) 𝑜 → Type (ℓ-max ℓ₀ ℓ₁) 32 | 33 | Join {zero} X Y = tt* 34 | Join {suc zero} X Y = 35 | Join {0} (fst X) (fst Y) , 36 | λ _ → snd X tt* ⊎ snd Y tt* 37 | Join {suc (suc zero)} ((X₀ , X₁) , X₂) ((Y₀ , Y₁) , Y₂) = 38 | Join {1} (X₀ , X₁) (Y₀ , Y₁) , JoinArrow ((X₀ , X₁) , X₂) ((Y₀ , Y₁) , Y₂) 39 | Join {suc (suc (suc n))} X Y = 40 | Join {suc (suc n)} (fst X) (fst Y) , JoinCell X Y 41 | 42 | data JoinArrow {ℓ₀ ℓ₁} X Y where 43 | 44 | jarr-inl : (x₀ : snd (fst X) tt*) (x₁ : snd (fst X) tt*) 45 | → (α : (snd X) {● ∣ objₒ} (tt* , x₀ , tt* , λ _ → x₁)) 46 | → JoinArrow X Y {● ∣ objₒ} (tt* , inl x₀ , tt* , λ _ → inl x₁) 47 | 48 | jarr-inr : (y₀ : snd (fst Y) tt*) (y₁ : snd (fst Y) tt*) 49 | → (β : snd Y {● ∣ objₒ} (tt* , y₀ , tt* , λ _ → y₁)) 50 | → JoinArrow X Y {● ∣ objₒ} (tt* , inr y₀ , tt* , λ _ → inr y₁) 51 | 52 | jarr-inm : (x : snd (fst X) tt*) (y : snd (fst Y) tt*) 53 | → JoinArrow X Y {● ∣ objₒ} (tt* , inl x , tt* , λ _ → inr y) 54 | 55 | data JoinCell {n ℓ₀ ℓ₁} X Y where 56 | 57 | jcell-inl : {𝑜 : 𝒪 (2 + n)} {f : Frm (fst X) 𝑜} (x : snd X f) 58 | → JoinCell X Y (Frm⇒ {Y = Join (fst X) (fst Y)} join-inl f) 59 | 60 | jcell-inr : {𝑜 : 𝒪 (2 + n)} {f : Frm (fst Y) 𝑜} (y : snd Y f) 61 | → JoinCell X Y (Frm⇒ {Y = Join (fst X) (fst Y)} join-inr f) 62 | 63 | jcell-inm : {𝑜 : 𝒪 (2 + n)} (f : Frm (Join (fst X) (fst Y)) 𝑜) 64 | → is-traversing f 65 | → JoinCell X Y f 66 | 67 | is-traversing {𝑜 = ● ∣ objₒ ∣ 𝑞} ((lift .tt , .(inl x₀) , lift .tt , .(λ _ → inl x₁)) , jarr-inl x₀ x₁ α , c , y) = ⊥ 68 | is-traversing {𝑜 = ● ∣ objₒ ∣ 𝑞} ((lift .tt , .(inr y₀) , lift .tt , .(λ _ → inr y₁)) , jarr-inr y₀ y₁ β , c , y) = ⊥ 69 | is-traversing {𝑜 = ● ∣ objₒ ∣ 𝑞} ((lift .tt , .(inl x) , lift .tt , .(λ _ → inr y₁)) , jarr-inm x y₁ , c , y) = Unit 70 | is-traversing {𝑜 = 𝑜 ∣ 𝑝 ∣ 𝑞 ∣ 𝑟} f = is-traversing (fst f) 71 | 72 | join-inl {zero} = tt* 73 | join-inl {suc zero} = tt* , λ { {●} {tt*} x → inl x } 74 | join-inl {suc (suc zero)} {X = (X₀ , X₁) , X₂} {Y = (Y₀ , Y₁) , Y₂} = 75 | join-inl {1} {X = (X₀ , X₁)} {Y = (Y₀ , Y₁)} , 76 | λ { {● ∣ objₒ} {_ , x₀ , _ , x₁} α → jarr-inl x₀ (x₁ tt) α } 77 | join-inl {suc (suc (suc n))} = join-inl {2 + n} , jcell-inl 78 | 79 | join-inr {zero} = tt* 80 | join-inr {suc zero} = tt* , λ { {●} {tt*} y → inr y } 81 | join-inr {suc (suc zero)} {X = (X₀ , X₁) , X₂} {Y = (Y₀ , Y₁) , Y₂} = 82 | join-inr {1} {X = (X₀ , X₁)} {Y = (Y₀ , Y₁)} , 83 | λ { {● ∣ objₒ} {_ , y₀ , _ , y₁} β → jarr-inr y₀ (y₁ tt) β } 84 | join-inr {suc (suc (suc n))} = join-inr {2 + n} , jcell-inr 85 | 86 | -------------------------------------------------------------------------------- /Lib/Shapes.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- Opetopes.agda - Underlying shapes for opetopic types 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Nat 7 | open import Cubical.Data.Sum 8 | open import Cubical.Data.Empty 9 | renaming (elim to ⊥-elim) 10 | 11 | open import Core.Prelude 12 | open import Core.Opetopes 13 | 14 | module Lib.Shapes where 15 | 16 | -- Dim 0 17 | 18 | obj : 𝒪 0 19 | obj = ● 20 | 21 | -- Dim 1 22 | 23 | arrow : 𝒪 1 24 | arrow = ● ∣ objₒ 25 | 26 | -- Dim 2 27 | 28 | n-path : ℕ → 𝒫 (● ∣ objₒ) 29 | n-path zero = lfₒ 30 | n-path (suc n) = ndₒ objₒ (λ _ → objₒ) (λ _ → n-path n) 31 | 32 | n-gon : ℕ → 𝒪 2 33 | n-gon n = ● ∣ objₒ ∣ n-path n 34 | 35 | 2-drop : 𝒪 2 36 | 2-drop = n-gon 0 37 | 38 | 2-globe : 𝒪 2 39 | 2-globe = n-gon 1 40 | 41 | 2-simplex : 𝒪 2 42 | 2-simplex = n-gon 2 43 | 44 | -- Dim 3 45 | 46 | left-unitor-𝑞 : (p : Pos (n-path 2)) → 𝒫 (Typ (n-path 2) p) 47 | left-unitor-𝑞 (inl tt) = n-path 0 48 | left-unitor-𝑞 (inr (tt , inl tt)) = n-path 1 49 | 50 | left-unitor-𝑟 : (p : Pos (n-path 2)) → 𝒫 (Typ (n-path 2) p ∣ left-unitor-𝑞 p) 51 | left-unitor-𝑟 (inl tt) = ndₒ {𝑜 = ● ∣ objₒ} lfₒ ⊥-elim ⊥-elim 52 | left-unitor-𝑟 (inr (tt , inl tt)) = lfₒ 53 | 54 | left-unitor : 𝒫 2-globe 55 | left-unitor = ndₒ (n-path 2) left-unitor-𝑞 left-unitor-𝑟 56 | 57 | left-drop-right-glob : 𝒫 2-globe 58 | left-drop-right-glob = ndₒ (n-path 2) 𝑞 𝑟 59 | 60 | where 𝑞 : (p : Pos (n-path 2)) → 𝒫 (Typ (n-path 2) p) 61 | 𝑞 (inl tt) = n-path 0 62 | 𝑞 (inr (tt , inl tt)) = n-path 1 63 | 64 | 𝑟 : (p : Pos (n-path 2)) → 𝒫 (Typ (n-path 2) p ∣ 𝑞 p) 65 | 𝑟 (inl tt) = ndₒ {𝑜 = ● ∣ objₒ} lfₒ (λ { () }) (λ { () }) 66 | 𝑟 (inr (tt , inl tt)) = ηₒ (n-gon 1) 67 | 68 | assoc-l : 𝒫 (n-gon 3) 69 | assoc-l = ndₒ (n-path 2) 𝑞 𝑟 70 | 71 | where 𝑞 : (p : Pos (n-path 2)) → 𝒫 (Typ (n-path 2) p) 72 | 𝑞 (inl tt) = n-path 2 73 | 𝑞 (inr (tt , inl tt)) = n-path 1 74 | 75 | 𝑟 : (p : Pos (n-path 2)) → 𝒫 (Typ (n-path 2) p ∣ 𝑞 p) 76 | 𝑟 (inl tt) = ndₒ (n-path 2) (λ p → ηₒ (Typ (n-path 2) p)) λ _ → lfₒ 77 | 𝑟 (inr (tt , inl tt)) = lfₒ 78 | 79 | assoc-r : 𝒫 (n-gon 3) 80 | assoc-r = ndₒ (n-path 2) 𝑞 𝑟 81 | 82 | where 𝑞 : (p : Pos (n-path 2)) → 𝒫 (Typ (n-path 2) p) 83 | 𝑞 (inl tt) = n-path 1 84 | 𝑞 (inr (tt , inl tt)) = n-path 2 85 | 86 | 𝑟 : (p : Pos (n-path 2)) → 𝒫 (Typ (n-path 2) p ∣ 𝑞 p) 87 | 𝑟 (inl tt) = lfₒ 88 | 𝑟 (inr (tt , inl tt)) = ndₒ (n-path 2) (λ p → ηₒ (Typ (n-path 2) p)) λ _ → lfₒ 89 | 90 | -------------------------------------------------------------------------------- /Lib/Structures.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicStructures.agda - Definitions of various structures 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Nat 7 | 8 | open import Core.Everything 9 | 10 | module Lib.Structures where 11 | 12 | is-fibrant : ∀ {n ℓ} (X : 𝕆Type (suc (suc n)) ℓ) → Type ℓ 13 | is-fibrant {n} ((Xₙ , Xₛₙ) , Xₛₛₙ) = 14 | {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 15 | {f : Frm Xₙ 𝑜} (c : Cns Xₙ f 𝑝) 16 | (y : (p : Pos 𝑝) → Xₛₙ (Shp Xₙ c p)) 17 | → isContr (Σ[ x ∈ Xₛₙ f ] Xₛₛₙ {𝑜 ∣ 𝑝} (f , x , c , y)) 18 | 19 | record is-fibrant-ext {n ℓ} {Γ : 𝕆Type n ℓ} (Γ∞ : 𝕆Type∞ Γ) : Type ℓ where 20 | coinductive 21 | field 22 | fill-fib : is-fibrant ((Γ , Fill Γ∞) , Fill (Hom Γ∞)) 23 | hom-fib : is-fibrant-ext (Hom Γ∞) 24 | 25 | ∞-Groupoid : (ℓ : Level) → Type (ℓ-suc ℓ) 26 | ∞-Groupoid ℓ = Σ[ X ∈ 𝕆Type∞ (𝕋 0) ] is-fibrant-ext X 27 | 28 | ∞-Category : (ℓ : Level) → Type (ℓ-suc ℓ) 29 | ∞-Category ℓ = Σ[ X ∈ 𝕆Type∞ (𝕋 0) ] is-fibrant-ext (Hom X) 30 | 31 | A∞-space : (ℓ : Level) → Type (ℓ-suc ℓ) 32 | A∞-space ℓ = Σ[ X ∈ 𝕆Type∞ (𝕋 1) ] is-fibrant-ext X 33 | 34 | ∞-Planar-Operad : (ℓ : Level) → Type (ℓ-suc ℓ) 35 | ∞-Planar-Operad ℓ = Σ[ X ∈ 𝕆Type∞ (𝕋 2) ] is-fibrant-ext X 36 | 37 | -- Composition Structures 38 | composition-structure : ∀ {n ℓ} (X : 𝕆Type (suc (suc n)) ℓ) → Type ℓ 39 | composition-structure {n} ((Xₙ , Xₛₙ) , Xₛₛₙ) = 40 | {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 41 | {f : Frm Xₙ 𝑜} (c : Cns Xₙ f 𝑝) 42 | (y : (p : Pos 𝑝) → Xₛₙ (Shp Xₙ c p)) 43 | → Σ[ x ∈ Xₛₙ f ] Xₛₛₙ {𝑜 ∣ 𝑝} (f , x , c , y) 44 | 45 | -- Uniquely composable, aka fibrant, types 46 | uniqueness-structure : ∀ {n ℓ} {X : 𝕆Type (suc (suc n)) ℓ} 47 | → composition-structure X → Type ℓ 48 | uniqueness-structure {n} {X = (Xₙ , Xₛₙ) , Xₛₛₙ} ϕ = 49 | {𝑜 : 𝒪 n} {𝑝 : 𝒫 𝑜} 50 | {f : Frm Xₙ 𝑜} (c : Cns Xₙ f 𝑝) 51 | (y : (p : Pos 𝑝) → Xₛₙ (Shp Xₙ c p)) 52 | (cmp : Xₛₙ f) (fill : Xₛₛₙ (f , cmp , c , y)) 53 | → (cmp , fill) ≡ ϕ c y 54 | -------------------------------------------------------------------------------- /Lib/Universe.agda: -------------------------------------------------------------------------------- 1 | -- 2 | -- OpetopicType.agda - Definition of Opetopic Types in Context 3 | -- 4 | 5 | open import Cubical.Foundations.Everything 6 | open import Cubical.Data.Nat 7 | 8 | open import Core.Everything 9 | open import Lib.Structures 10 | 11 | module Lib.Universe where 12 | 13 | -- The universe as an opetopic type 14 | 𝒰ₒ : (n : ℕ) (ℓ : Level) → 𝕆Type n (ℓ-suc ℓ) 15 | 𝒱ₒ : (n : ℕ) {ℓ : Level} → 𝕆Fam (𝒰ₒ n ℓ) ℓ 16 | 17 | 𝒰ₒ zero ℓ = lift tt 18 | 𝒰ₒ (suc n) ℓ = 𝒰ₒ n ℓ , λ f → Frm↓ (𝒱ₒ n) f → Type ℓ 19 | 20 | 𝒱ₒ zero = lift tt 21 | 𝒱ₒ (suc n) = 𝒱ₒ n , λ f↓ X → X f↓ 22 | 23 | -- 24 | -- Let's get set up for some of the main theorems 25 | -- 26 | 27 | is-fibrant-rel : ∀ {n ℓ} {𝑜 : 𝒪 n} {f : Frm (𝒰ₒ n ℓ) 𝑜} 28 | → (X : Frm↓ (𝒱ₒ n) f → Type ℓ) → Type ℓ 29 | is-fibrant-rel {zero} X = Lift Unit 30 | is-fibrant-rel {suc n} {𝑜 = 𝑜 ∣ 𝑝} {f , Xₙ , c , Yₙ} R = 31 | (f↓ : Frm↓ (𝒱ₒ n) f) (c↓ : Cns↓ (𝒱ₒ n) f↓ c) 32 | (y↓ : (p : Pos 𝑝) → Yₙ p (Shp↓ (𝒱ₒ n) c↓ p)) 33 | → isContr (Σ[ x↓ ∈ Xₙ f↓ ] R (f↓ , x↓ , c↓ , y↓)) 34 | 35 | -- The dependent type of fibrancy 36 | ℱₒ : (n : ℕ) {ℓ : Level} → 𝕆Fam (𝒰ₒ n ℓ) ℓ 37 | ℱₒ zero = tt* 38 | ℱₒ (suc n) = ℱₒ n , λ _ X → is-fibrant-rel X 39 | 40 | -- We can now define the (∞,1)-category of spaces: 41 | 𝒮ₙ : (n : ℕ) (ℓ : Level) → 𝕆Type n (ℓ-suc ℓ) 42 | 𝒮ₙ n ℓ = Σₒ (𝒰ₒ n ℓ) (ℱₒ n) 43 | 44 | -------------------------------------------------------------------------------- /Native/NativeOpetopicTypes.org: -------------------------------------------------------------------------------- 1 | 2 | * Hacking Notes for Native Opetopic Types 3 | 4 | ** Compiling Agda 5 | 6 | make install-bin 7 | 8 | ** Places to change when adding a new builtin 9 | 10 | *** Give a name to the builtin in [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/Syntax/Builtin.hs::builtinNat, builtinSuc, builtinZero, builtinNatPlus, builtinNatMinus,][Syntax/Builtin.hs]] 11 | *** Add to list of primitives in [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/TypeChecking/Monad/Builtin.hs::primInteger, primIntegerPos, primIntegerNegSuc,][Monad/Builtin.hs]] 12 | *** Define the builtin in [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/TypeChecking/Rules/Builtin.hs::, (builtinOpetope |-> BuiltinData (tnat --> tset) \[\])][Rules/Builtin.hs]] 13 | 14 | ** Places to change to add a new primitive 15 | 16 | *** Link the implementation in [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/TypeChecking/Primitive.hs::, "primInhab" |-> primInhab'][Primitive.hs]] 17 | 18 | ** References to various interesting files 19 | 20 | *** Cubical 21 | 22 | **** Cubical primitive [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/TypeChecking/Primitive/Cubical.hs::requireCubical][implementations]] 23 | **** Cubical primitive Agda [[file:~/Code/agda/Agda-2.6.3/src/data/lib/prim/Agda/Primitive/Cubical.agda::{-# OPTIONS --erased-cubical #-}][links]] 24 | **** Implementation of native [[file:~/Code/agda/Agda-2.6.3/src/data/lib/prim/Agda/Builtin/Cubical/Glue.agda::_≃_ : ∀ {ℓ ℓ'} (A : Set ℓ) (B : Set ℓ') → Set (ℓ ⊔ ℓ')][equivalences]] 25 | 26 | This shows how BuiltinUnknown is used for a builtin which can be given 27 | a definition in Agda, but it's type must be known for the whole 28 | signature to work. 29 | 30 | *** Agda 31 | 32 | **** [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/Syntax/Internal.hs::data Term = Var {-# UNPACK #-} !Int Elims -- ^ @x es@ neutral][Internal Syntax]] (i.e Terms) 33 | 34 | **** Smart term constructors: [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/TypeChecking/Primitive/Base.hs::gpi :: (MonadAddContext m, MonadDebug m)][Primitive/Base.hs]] 35 | 36 | These are the main tools for writing internal code as far as I can 37 | see... 38 | 39 | **** Definition of [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/TypeChecking/Names.hs::cl :: Monad m => m a -> NamesT m a][cl]] 40 | 41 | Appears to just be lifting a monadic computation into the names guy... 42 | 43 | **** Definition of Sigma "Kit": [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/TypeChecking/Primitive/Base.hs::data SigmaKit = SigmaKit][sigma kit]] 44 | 45 | Shows how to retrieve constructors and so on. 46 | 47 | **** Level view and an example of using the debug print is [[file:~/Code/agda/Agda-2.6.3/src/full/Agda/TypeChecking/Level.hs::levelView :: PureTCM m => Term -> m Level][here]]. 48 | 49 | Shows how to match on definitions 50 | 51 | ** Things To Do 52 | 53 | *** TODO Builtins should depend on --opetopic-types 54 | 55 | *** Local Signature which needs adding 56 | 57 | **** TODO Branch 58 | **** TODO Tr 59 | **** TODO TrPos 60 | **** TODO Gamma 61 | **** TODO GammaInl 62 | **** TODO GammaInr 63 | **** TODO GammaElim 64 | 65 | -------------------------------------------------------------------------------- /Native/Opetopes.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --opetopic-types --verbose=tc.primitive.optypes:10 #-} 2 | 3 | module Native.Opetopes where 4 | 5 | open import Agda.Primitive public 6 | using ( Level ) 7 | renaming ( lzero to ℓ-zero 8 | ; lsuc to ℓ-suc 9 | ; _⊔_ to ℓ-max 10 | ; Set to Type 11 | ; Setω to Typeω ) 12 | 13 | open import Agda.Builtin.Nat public 14 | using (zero; suc) 15 | renaming (Nat to ℕ) 16 | 17 | open import Agda.Builtin.Sigma public 18 | 19 | -- Σ-types 20 | infix 2 Σ-syntax 21 | 22 | Σ-syntax : ∀ {ℓ ℓ'} (A : Type ℓ) (B : A → Type ℓ') → Type (ℓ-max ℓ ℓ') 23 | Σ-syntax = Σ 24 | 25 | syntax Σ-syntax A (λ x → B) = Σ[ x ∈ A ] B 26 | 27 | -- 28 | -- Universe Polymorphic Unit Type 29 | -- 30 | 31 | record ● {ℓ} : Type ℓ where 32 | constructor ∙ 33 | 34 | {-# BUILTIN POLYUNIT ● #-} 35 | 36 | -- 37 | -- Opetopic Types 38 | -- 39 | 40 | postulate 41 | 42 | 𝕆Type : ℕ → (ℓ : Level) → Type (ℓ-suc ℓ) 43 | 44 | Frm : ∀ {n ℓ} → 𝕆Type n ℓ → Type ℓ 45 | 46 | Src : ∀ {n ℓ} (X : 𝕆Type n ℓ) 47 | → (P : Frm X → Type ℓ) 48 | → Frm X → Type ℓ 49 | 50 | Pos : ∀ {n ℓ} {X : 𝕆Type n ℓ} 51 | → (P : Frm X → Type ℓ) 52 | → {f : Frm X} (s : Src X P f) 53 | → Type ℓ 54 | 55 | Typ : ∀ {n ℓ} {X : 𝕆Type n ℓ} 56 | → (P : Frm X → Type ℓ) 57 | → {f : Frm X} (s : Src X P f) 58 | → (p : Pos P s) → Frm X 59 | 60 | Inhab : ∀ {n ℓ} {X : 𝕆Type n ℓ} 61 | → (P : Frm X → Type ℓ) 62 | → {f : Frm X} (s : Src X P f) 63 | → (p : Pos P s) 64 | → P (Typ P s p) 65 | 66 | {-# BUILTIN OPETOPICTYPE 𝕆Type #-} 67 | {-# BUILTIN FRM Frm #-} 68 | {-# BUILTIN SRC Src #-} 69 | {-# BUILTIN POS Pos #-} 70 | {-# BUILTIN TYP Typ #-} 71 | {-# BUILTIN INHAB Inhab #-} 72 | 73 | -- 74 | -- Maps of Opetopic Types 75 | -- 76 | 77 | infixl 50 _⊙_ 78 | 79 | postulate 80 | 81 | _⇒_ : ∀ {n ℓ} → 𝕆Type n ℓ → 𝕆Type n ℓ → Type ℓ 82 | 83 | id-map : ∀ {n ℓ} → (X : 𝕆Type n ℓ) → X ⇒ X 84 | 85 | _⊙_ : ∀ {n ℓ} {X Y Z : 𝕆Type n ℓ} 86 | → Y ⇒ Z → X ⇒ Y → X ⇒ Z 87 | 88 | Frm⇒ : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} 89 | → (σ : X ⇒ Y) 90 | → Frm X → Frm Y 91 | 92 | {-# BUILTIN OPETOPICMAP _⇒_ #-} 93 | {-# BUILTIN IDMAP id-map #-} 94 | {-# BUILTIN MAPCOMP _⊙_ #-} 95 | {-# BUILTIN FRMMAP Frm⇒ #-} 96 | 97 | -- 98 | -- Monadic Signature 99 | -- 100 | 101 | postulate 102 | 103 | η : ∀ {n ℓ} {X : 𝕆Type n ℓ} 104 | → (P : Frm X → Type ℓ) 105 | → {f : Frm X} (x : P f) 106 | → Src X P f 107 | 108 | η-pos : ∀ {n ℓ} {X : 𝕆Type n ℓ} 109 | → (P : Frm X → Type ℓ) 110 | → {f : Frm X} (x : P f) 111 | → Pos P (η P x) 112 | 113 | μ : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (σ : X ⇒ Y) 114 | → (P : Frm X → Type ℓ) 115 | → (Q : Frm Y → Type ℓ) 116 | → {f : Frm X} (s : Src X P f) 117 | → (ϕ : (p : Pos P s) → Src Y Q (Frm⇒ σ (Typ P s p))) 118 | → Src Y Q (Frm⇒ σ f) 119 | 120 | μ-pos : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (σ : X ⇒ Y) 121 | → (P : Frm X → Type ℓ) 122 | → (Q : Frm Y → Type ℓ) 123 | → {f : Frm X} (s : Src X P f) 124 | → (ϕ : (p : Pos P s) → Src Y Q (Frm⇒ σ (Typ P s p))) 125 | → (p : Pos P s) (q : Pos Q (ϕ p)) 126 | → Pos Q (μ σ P Q s ϕ) 127 | 128 | μ-fst : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (σ : X ⇒ Y) 129 | → (P : Frm X → Type ℓ) 130 | → (Q : Frm Y → Type ℓ) 131 | → {f : Frm X} (s : Src X P f) 132 | → (ϕ : (p : Pos P s) → Src Y Q (Frm⇒ σ (Typ P s p))) 133 | → (p : Pos Q (μ σ P Q s ϕ)) 134 | → Pos P s 135 | 136 | μ-snd : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (σ : X ⇒ Y) 137 | → (P : Frm X → Type ℓ) 138 | → (Q : Frm Y → Type ℓ) 139 | → {f : Frm X} (s : Src X P f) 140 | → (ϕ : (p : Pos P s) → Src Y Q (Frm⇒ σ (Typ P s p))) 141 | → (p : Pos Q (μ σ P Q s ϕ)) 142 | → Pos Q (ϕ (μ-fst σ P Q s ϕ p)) 143 | 144 | {-# BUILTIN UNT η #-} 145 | {-# BUILTIN UNTPOS η-pos #-} 146 | 147 | {-# BUILTIN SUBST μ #-} 148 | {-# BUILTIN SUBSTPOS μ-pos #-} 149 | {-# BUILTIN SUBSTFST μ-fst #-} 150 | {-# BUILTIN SUBSTSND μ-snd #-} 151 | 152 | -- 153 | -- Trees and Their Positions 154 | -- 155 | 156 | module _ {n ℓ} {X : 𝕆Type n ℓ} {P : Frm X → Type ℓ} 157 | (U : Frm {suc n} (X , P) → Type ℓ) where 158 | 159 | data Tr : Frm {suc n} (X , P) → Type ℓ 160 | 161 | {-# NO_POSITIVITY_CHECK #-} 162 | record Branch (f : Frm X) : Type ℓ where 163 | inductive 164 | eta-equality 165 | constructor [_,_,_] 166 | field 167 | stm : P f 168 | lvs : Src X P f 169 | br : Tr (f , lvs , stm) 170 | 171 | open Branch public 172 | 173 | data Tr where 174 | 175 | lf : {f : Frm X} (tgt : P f) 176 | → Tr (f , η P tgt , tgt) 177 | 178 | nd : {f : Frm X} (tgt : P f) 179 | → (brs : Src X Branch f) 180 | → (flr : U (f , μ (id-map X) Branch P brs (λ p → η P (stm (Inhab Branch brs p))) , tgt)) 181 | → Tr (f , μ (id-map X) Branch P brs (λ p → lvs (Inhab Branch brs p)) , tgt) 182 | -------------------------------------------------------------------------------- /Native/Test.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --opetopic-types --verbose=tc.primitive.optypes:10 #-} 2 | 3 | open import Native.Opetopes 4 | 5 | module Native.Test where 6 | 7 | -- 8 | -- Testing 9 | -- 10 | 11 | data _≡_ {i} {A : Type i} (a : A) : A → Type i where 12 | refl : a ≡ a 13 | 14 | {-# BUILTIN EQUALITY _≡_ #-} 15 | 16 | 𝕆Type-zero : (ℓ : Level) 17 | → 𝕆Type zero ℓ ≡ ● 18 | 𝕆Type-zero ℓ = refl 19 | 20 | 𝕆Type-suc : (n : ℕ) (ℓ : Level) 21 | → 𝕆Type (suc n) ℓ ≡ (Σ[ X ∈ 𝕆Type n ℓ ] (Frm X → Type ℓ)) 22 | 𝕆Type-suc n ℓ = refl 23 | 24 | Frm-zero : (ℓ : Level) (X : 𝕆Type zero ℓ) → Frm {zero} X ≡ ● 25 | Frm-zero ℓ X = refl 26 | 27 | Frm-suc : (n : ℕ) (ℓ : Level) (X : 𝕆Type (suc n) ℓ) 28 | → Frm {suc n} X ≡ (Σ[ f ∈ Frm {n} (fst X) ] 29 | Σ[ s ∈ Src {n} (fst X) (snd X) f ] 30 | (snd X) f) 31 | Frm-suc n ℓ X = refl 32 | 33 | ⇒-zero : (ℓ : Level) (X Y : 𝕆Type zero ℓ) 34 | → _⇒_ {zero} X Y ≡ ● 35 | ⇒-zero ℓ X Y = refl 36 | 37 | ⇒-suc : (n : ℕ) (ℓ : Level) (X Y : 𝕆Type (suc n) ℓ) 38 | → _⇒_ {suc n} X Y ≡ (Σ[ σ ∈ (fst X) ⇒(fst Y) ] 39 | ({f : Frm (fst X)} → (snd X) f → (snd Y) (Frm⇒ σ f))) 40 | ⇒-suc n ℓ X Y = refl 41 | 42 | id-map-zero : ∀ {ℓ} (X : 𝕆Type zero ℓ) 43 | → id-map {zero} X ≡ ∙ 44 | id-map-zero X = refl 45 | 46 | id-map-suc : ∀ {n ℓ} → (X : 𝕆Type (suc n) ℓ) 47 | → id-map {suc n} X ≡ (id-map {n} (fst X) , λ {f} p → p) 48 | id-map-suc X = refl 49 | 50 | map-comp-zero : ∀ {ℓ} (X Y Z : 𝕆Type zero ℓ) (σ : _⇒_ {zero} X Y) (τ : _⇒_ {zero} Y Z) 51 | → _⊙_ {zero} τ σ ≡ ∙ 52 | map-comp-zero X Y Z σ τ = refl 53 | 54 | map-comp-suc : ∀ {n ℓ} (X Y Z : 𝕆Type (suc n) ℓ) (σ : _⇒_ {suc n} X Y) (τ : _⇒_ {suc n} Y Z) 55 | → _⊙_ {suc n} {X = X} {Y} {Z} τ σ ≡ (fst τ ⊙ fst σ , λ {f} p → snd τ (snd σ p)) 56 | map-comp-suc X Y Z σ τ = refl 57 | 58 | map-comp-unit-right : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (σ : X ⇒ Y) 59 | → (σ ⊙ id-map X) ≡ σ 60 | map-comp-unit-right σ = refl 61 | 62 | map-comp-unit-left : ∀ {n ℓ} {X Y : 𝕆Type n ℓ} (σ : X ⇒ Y) 63 | → (id-map Y ⊙ σ) ≡ σ 64 | map-comp-unit-left σ = {!refl!} 65 | 66 | Frm⇒-id : ∀ {n ℓ} (X : 𝕆Type n ℓ) (f : Frm X) 67 | → Frm⇒ (id-map X) f ≡ f 68 | Frm⇒-id X f = refl 69 | 70 | Frm⇒-⊙ : ∀ {n ℓ} {X Y Z : 𝕆Type n ℓ} 71 | → (σ : X ⇒ Y) (τ : Y ⇒ Z) (f : Frm X) 72 | → Frm⇒ (τ ⊙ σ) f ≡ Frm⇒ τ (Frm⇒ σ f) 73 | Frm⇒-⊙ σ τ f = refl 74 | 75 | 76 | 77 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Opetopic Types 2 | 3 | This is an Agda library which adds a notion of **opetopic type** to 4 | type theory for use in formalizing higher algebraic structures. 5 | 6 | The overall methodology here is to strictify the equations which make 7 | opetopic types definable. These are essentially the equations of a 8 | polynomial monad. One way to view the setup is that we are defining 9 | an internal model with exactly enough structure to give us the 10 | type theory of opetopic types themsevles. 11 | 12 | The essential idea of this approach is that adding these equations 13 | should be enough to "do" higher algebra. 14 | 15 | The library can accomplish the following definitions: 16 | 17 | 1. Σ and Π of opetopic types 18 | 2. The opetopic type associated to a type 19 | 3. The "fibrant replacement", turning an opetopic type into a type 20 | 4. The join of opetopic types 21 | 5. The definition of the (∞,1)-category of types 22 | 6. A proof of the equivalence between "truncated" (∞,1)-categories and 23 | univalent categories 24 | 25 | -------------------------------------------------------------------------------- /opetopic-types.agda-lib: -------------------------------------------------------------------------------- 1 | name: opetopic-types 2 | include: . 3 | depend: cubical 4 | flags: --cubical --no-import-sorts --rewriting --guardedness 5 | 6 | --------------------------------------------------------------------------------