├── .gitignore ├── LICENSE.txt ├── README.md ├── agda ├── BoolBench.agda ├── Brunerie.agda ├── Brunerie2.agda ├── BrunerieExperiments2.agda ├── HLevelsTemp.agda └── S1Bench.agda ├── cctt.el ├── cubicaltt └── brunerie_appendix.ctt ├── hott2023 ├── EasyChair3.5 │ ├── chairEC.pdf │ ├── debug.pdf │ ├── easychair.bib │ ├── easychair.cls │ ├── easychair.pdf │ ├── easychair.tex │ ├── epic.pdf │ ├── logoEC.pdf │ └── throneEC.jpg ├── HoTT-2023_abstract_23.pdf ├── easychair.cls ├── hott-2023-template.aux ├── hott-2023-template.bbl ├── hott-2023-template.blg ├── hott-2023-template.fdb_latexmk ├── hott-2023-template.fls ├── hott-2023-template.log ├── hott-2023-template.out ├── hott-2023-template.pdf ├── hott-2023-template.tex ├── prez │ ├── AndrasKovacs.pdf │ ├── brunerie.png │ ├── prez.tex │ ├── references.bib │ └── unglues.png └── references.bib ├── newnotes ├── Connections.txt ├── notes.txt └── todo.txt ├── oldnotes ├── CHMGlueNotes.txt ├── HITs.txt ├── Notes.agda ├── Notes.hs ├── Notes2.agda ├── TranspGlue.agda ├── TranspGlueIssue.agda ├── newnotes.txt ├── notes.txt └── prez.txt ├── package.yaml ├── references ├── brunerie_nums.pdf ├── cctt_investigations.pdf ├── coefcom.pdf ├── extension.pdf └── lattice_impl.pdf ├── rules.txt ├── src ├── Common.hs ├── Conversion.hs ├── Core.hs ├── CoreTypes.hs ├── Cubical.hs ├── Cubical │ ├── Cofibration.hs │ ├── Interval.hs │ └── Substitution.hs ├── Data │ └── LvlMap.hs ├── ElabState.hs ├── Elaboration.hs ├── Errors.hs ├── Lexer.hs ├── Main.hs ├── MainInteraction.hs ├── Parser.hs ├── Presyntax.hs ├── Pretty.hs ├── Quotation.hs └── Statistics.hs ├── stack.yaml ├── tests ├── Russell.agda ├── Russell.cctt ├── RussellNoPath.cctt ├── basics.cctt ├── bench │ ├── basics.cctt │ ├── bool.cctt │ ├── int.cctt │ ├── nat.cctt │ ├── negPath.cctt │ └── s1.cctt ├── bool.cctt ├── brunerie.cctt ├── brunerie_james.cctt ├── brunerie_james_revised.cctt ├── entanglement.cctt ├── hedberg.cctt ├── importtest │ ├── Foo │ │ ├── Bar │ │ │ ├── bar.cctt │ │ │ └── bar2.cctt │ │ └── foo.cctt │ ├── a.cctt │ └── foo.cctt ├── int.cctt ├── new_brunerie.cctt ├── new_brunerie_adjusthcom.cctt ├── new_brunerie_adjusthcom_1sym.cctt ├── new_brunerie_adjusthcom_2sym.cctt ├── tests │ ├── basics.cctt │ ├── bool.cctt │ ├── gluetests_ghcom.cctt │ ├── gluetests_no_ghcom.cctt │ ├── hittests.cctt │ ├── indtests.cctt │ ├── indtests2.cctt │ ├── nat.cctt │ ├── notes.cctt │ ├── oldbrunerie-shrink.cctt │ ├── recHITboundary.cctt │ ├── test.cctt │ └── yaccttprelude.cctt └── yaccttprelude.cctt ├── tutorial.cctt └── types23 ├── bench.txt ├── brunerie.png ├── prez.aux ├── prez.fdb_latexmk ├── prez.fls ├── prez.log ├── prez.nav ├── prez.out ├── prez.pdf ├── prez.snm ├── prez.tex ├── prez.toc ├── references.bib └── unglues.png /.gitignore: -------------------------------------------------------------------------------- 1 | *~ 2 | *.hi 3 | *.o 4 | /dist 5 | *.agdai 6 | *# 7 | TAGS 8 | .stack-work/ 9 | *.lock 10 | -------------------------------------------------------------------------------- /LICENSE.txt: -------------------------------------------------------------------------------- 1 | Copyright 2022 András Kovács 2 | 3 | Permission is hereby granted, free of charge, to any person obtaining a copy of 4 | this software and associated documentation files (the "Software"), to deal in 5 | the Software without restriction, including without limitation the rights to 6 | use, copy, modify, merge, publish, distribute, sublicense, and/or sell copies of 7 | the Software, and to permit persons to whom the Software is furnished to do so, 8 | subject to the following conditions: 9 | 10 | The above copyright notice and this permission notice shall be included in all 11 | copies or substantial portions of the Software. 12 | 13 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 14 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, FITNESS 15 | FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE AUTHORS OR 16 | COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER LIABILITY, WHETHER 17 | IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, OUT OF OR IN 18 | CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE SOFTWARE. 19 | -------------------------------------------------------------------------------- /agda/BoolBench.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --cubical #-} 2 | 3 | open import Cubical.Foundations.Everything 4 | open import Cubical.Data.Nat 5 | open import Cubical.Data.Bool 6 | 7 | 8 | BoolPathN : ℕ → Bool ≡ Bool 9 | BoolPathN n = iter n (λ p → p ∙ notEq) notEq 10 | 11 | fun : ℕ → Bool → Bool 12 | fun n x = transport (BoolPathN n) x 13 | 14 | val : ℕ → Bool 15 | val n = transport (BoolPathN n) true 16 | 17 | test : val 1000 ≡ false 18 | test = refl 19 | -------------------------------------------------------------------------------- /agda/Brunerie.agda: -------------------------------------------------------------------------------- 1 | -- Mostly self contained definitions of the numbers from: https://arxiv.org/abs/2302.00151 2 | {-# OPTIONS --safe --cubical #-} 3 | module Brunerie where 4 | 5 | open import Cubical.Core.Primitives 6 | 7 | open import Cubical.Foundations.Prelude using 8 | (_,_ ; fst ; snd ; _≡_ ; cong ; _∙_ ; refl ; isProp→isSet ; J ; sym ; isProp ; isSet ; isGroupoid ; is2Groupoid) 9 | open import Cubical.Foundations.Equiv using 10 | (_≃_ ; isEquiv ; isPropIsEquiv ; idIsEquiv ; idEquiv) 11 | open import Cubical.Foundations.HLevels using 12 | (hSet ; hGroupoid ; isOfHLevelTypeOfHLevel ; isOfHLevelPath ; isOfHLevelΠ ; isOfHLevel→isOfHLevelDep) 13 | open import Cubical.Foundations.Univalence using 14 | (Glue ; ua) 15 | open import Cubical.Data.Int using 16 | (ℤ ; pos ; neg ; isSetℤ ; sucPathℤ) 17 | open import Cubical.Foundations.CartesianKanOps 18 | 19 | private variable ℓ ℓ' : Level 20 | 21 | -- S1 22 | data S¹ : Type₀ where 23 | base : S¹ 24 | loop : base ≡ base 25 | 26 | helix : S¹ → Type₀ 27 | helix base = ℤ 28 | helix (loop i) = sucPathℤ i 29 | 30 | rotLoop : (a : S¹) → a ≡ a 31 | rotLoop base = loop 32 | rotLoop (loop i) j = 33 | hcomp (λ k → λ { (i = i0) → loop (j ∨ ~ k) 34 | ; (i = i1) → loop (j ∧ k) 35 | ; (j = i0) → loop (i ∨ ~ k) 36 | ; (j = i1) → loop (i ∧ k)}) base 37 | 38 | _·_ : S¹ → S¹ → S¹ 39 | base · x = x 40 | (loop i) · x = rotLoop x i 41 | 42 | isPropFamS¹ : (P : S¹ → Type ℓ) (pP : (x : S¹) → isProp (P x)) (b0 : P base) → 43 | PathP (λ i → P (loop i)) b0 b0 44 | isPropFamS¹ P pP b0 i = pP (loop i) (coe0→i (λ k → P (loop k)) i b0) 45 | (coe1→i (λ k → P (loop k)) i b0) i 46 | 47 | rotIsEquiv : (a : S¹) → isEquiv (a ·_) 48 | rotIsEquiv base = idIsEquiv S¹ 49 | rotIsEquiv (loop i) = isPropFamS¹ (λ x → isEquiv (x ·_)) 50 | (λ x → isPropIsEquiv (x ·_)) 51 | (idIsEquiv _) i 52 | 53 | -- S2 54 | data S² : Type₀ where 55 | base : S² 56 | surf : PathP (λ i → base ≡ base) refl refl 57 | 58 | S²ToSetElim : {A : S² → Type ℓ} 59 | → ((x : S²) → isSet (A x)) 60 | → A base 61 | → (x : S²) → A x 62 | S²ToSetElim set b base = b 63 | S²ToSetElim {A = A} set b (surf i j) = 64 | isOfHLevel→isOfHLevelDep 2 {S²}{λ z → A z} set {base}{base} b b {a0 = refl} {a1 = refl} refl refl surf i j 65 | 66 | 67 | -- Join 68 | data join (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') where 69 | inl : A → join A B 70 | inr : B → join A B 71 | push : ∀ a b → inl a ≡ inr b 72 | 73 | -- SetTruncation 74 | data ∥_∥₀ (A : Type ℓ) : Type ℓ where 75 | ∣_∣₀ : A → ∥ A ∥₀ 76 | squash₀ : ∀ (x y : ∥ A ∥₀) (p q : x ≡ y) → p ≡ q 77 | 78 | rec₀ : {A : Type ℓ} {B : Type ℓ'} → isSet B → (A → B) → ∥ A ∥₀ → B 79 | rec₀ Bset f ∣ x ∣₀ = f x 80 | rec₀ Bset f (squash₀ x y p q i j) = 81 | Bset (rec₀ Bset f x) (rec₀ Bset f y) (cong (rec₀ Bset f) p) (cong (rec₀ Bset f) q) i j 82 | 83 | 84 | -- GroupoidTruncation 85 | data ∥_∥₁ (A : Type ℓ) : Type ℓ where 86 | ∣_∣₁ : A → ∥ A ∥₁ 87 | squash₁ : ∀ (x y : ∥ A ∥₁) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s 88 | 89 | rec₁ : {A : Type ℓ} {B : Type ℓ'} → isGroupoid B → (A → B) → ∥ A ∥₁ → B 90 | rec₁ gB f ∣ x ∣₁ = f x 91 | rec₁ gB f (squash₁ x y p q r s i j k) = 92 | gB (rec₁ gB f x) (rec₁ gB f y) 93 | (λ n → rec₁ gB f (r i0 n)) 94 | (λ n → rec₁ gB f (r i1 n)) 95 | (λ m n → rec₁ gB f (r m n)) (λ m n → rec₁ gB f (s m n)) i j k 96 | 97 | 98 | -- 2GroupoidTruncation 99 | data ∥_∥₂ (A : Type ℓ) : Type ℓ where 100 | ∣_∣₂ : A → ∥ A ∥₂ 101 | squash₂ : ∀ (x y : ∥ A ∥₂) (p q : x ≡ y) (r s : p ≡ q) (t u : r ≡ s) → t ≡ u 102 | 103 | rec₂ : ∀ {A : Type ℓ} {B : Type ℓ'} → is2Groupoid B → (A → B) → ∥ A ∥₂ → B 104 | rec₂ gB f ∣ x ∣₂ = f x 105 | rec₂ gB f (squash₂ a b p q r s t u i j k l) = 106 | gB (rec₂ gB f a) 107 | (rec₂ gB f b) 108 | (λ o → rec₂ gB f (p o)) 109 | (λ o → rec₂ gB f (q o)) 110 | (λ n o → rec₂ gB f (t i0 n o)) 111 | (λ n o → rec₂ gB f (t i1 n o)) 112 | (λ m n o → rec₂ gB f (t m n o)) 113 | (λ m n o → rec₂ gB f (u m n o)) 114 | i j k l 115 | 116 | elim₂ : {A : Type ℓ} {B : ∥ A ∥₂ → Type ℓ} 117 | (bG : (x : ∥ A ∥₂) → is2Groupoid (B x)) 118 | (f : (x : A) → B ∣ x ∣₂) (x : ∥ A ∥₂) → B x 119 | elim₂ bG f ∣ x ∣₂ = f x 120 | elim₂ {A = A}{B} bG f (squash₂ a b p q r s u v i j k l) = 121 | isOfHLevel→isOfHLevelDep 4 {A = ∥ A ∥₂}{λ z → B z} bG 122 | {a} 123 | {b} 124 | (elim₂ {A = A}{B} bG f a) 125 | (elim₂ {A = A}{B} bG f b) 126 | {p} 127 | {q} 128 | (λ l₁ → elim₂ bG f (p l₁)) 129 | (λ l₁ → elim₂ bG f (q l₁)) 130 | {u i0} 131 | {u i1} 132 | (λ k₁ l₁ → elim₂ bG f (u i0 k₁ l₁)) 133 | (λ k₁ l₁ → elim₂ bG f (u i1 k₁ l₁)) 134 | {u} 135 | {v} 136 | (λ j k l → elim₂ bG f (u j k l)) (λ j k l → elim₂ bG f (v j k l)) 137 | (squash₂ a b p q r s u v) 138 | i j k l 139 | 140 | 141 | -- Susp 142 | data Susp (A : Type ℓ) : Type ℓ where 143 | north : Susp A 144 | south : Susp A 145 | merid : (a : A) → north ≡ south 146 | 147 | 148 | -- Pointed 149 | Pointed₀ : Type₁ 150 | Pointed₀ = Σ[ X ∈ Type₀ ] X 151 | 152 | S¹∙ S²∙ : Pointed₀ 153 | S¹∙ = (S¹ , base) 154 | S²∙ = (S² , base) 155 | 156 | Susp∙ : Type₀ → Pointed₀ 157 | Susp∙ A = Susp A , north 158 | 159 | ∥_∥₁∙ ∥_∥₂∙ : Pointed₀ → Pointed₀ 160 | ∥ A , a ∥₁∙ = ∥ A ∥₁ , ∣ a ∣₁ 161 | ∥ A , a ∥₂∙ = ∥ A ∥₂ , ∣ a ∣₂ 162 | 163 | Ω Ω² : Pointed₀ → Pointed₀ 164 | Ω (_ , a) = ((a ≡ a) , refl) 165 | Ω² p = Ω (Ω p) 166 | 167 | -- The maps 168 | σ : S² → Ω (Susp∙ S²) .fst 169 | σ x = merid x ∙ sym (merid base) 170 | 171 | S¹×S¹→S² : S¹ → S¹ → S² 172 | S¹×S¹→S² base y = base 173 | S¹×S¹→S² (loop i) base = base 174 | S¹×S¹→S² (loop i) (loop j) = surf i j 175 | 176 | f7 : Ω (Susp∙ S²) .fst → ∥ S² ∥₂ 177 | f7 = encode north 178 | where 179 | _+₂_ : ∥ S² ∥₂ → ∥ S² ∥₂ → ∥ S² ∥₂ 180 | _+₂_ = elim₂ (λ _ → isOfHLevelΠ 4 λ _ → squash₂) 181 | λ { base x → x 182 | ; (surf i j) x → surfc x i j} 183 | where 184 | surfc : (x : ∥ S² ∥₂) → Ω² (∥ S² ∥₂ , x) .fst 185 | surfc = elim₂ (λ _ → isOfHLevelPath 4 (isOfHLevelPath 4 squash₂ _ _) _ _) 186 | (S²ToSetElim (λ _ → squash₂ _ _ _ _) λ i j → ∣ surf i j ∣₂) 187 | 188 | ∥S²∥₂≃∥S²∥₂ : (x : S²) → ∥ S² ∥₂ ≃ ∥ S² ∥₂ 189 | fst (∥S²∥₂≃∥S²∥₂ x) = ∣ x ∣₂ +₂_ 190 | snd (∥S²∥₂≃∥S²∥₂ x) = help x 191 | where 192 | help : (x : _) → isEquiv (λ y → ∣ x ∣₂ +₂ y) 193 | help = S²ToSetElim (λ _ → isProp→isSet (isPropIsEquiv _)) (idEquiv _ .snd) 194 | 195 | Code : Susp S² → Type₀ 196 | Code north = ∥ S² ∥₂ 197 | Code south = ∥ S² ∥₂ 198 | Code (merid a i) = ua (∥S²∥₂≃∥S²∥₂ a) i 199 | 200 | encode : (x : Susp S²) → north ≡ x → Code x 201 | encode x = J (λ x p → Code x) ∣ base ∣₂ 202 | 203 | g8 : Ω² ∥ S²∙ ∥₂∙ .fst → Ω ∥ S¹∙ ∥₁∙ .fst 204 | g8 p i = encodeTruncS² (p i) 205 | where 206 | HopfS² : S² → Type₀ 207 | HopfS² base = S¹ 208 | HopfS² (surf i j) = Glue S¹ (λ { (i = i0) → _ , idEquiv S¹ 209 | ; (i = i1) → _ , idEquiv S¹ 210 | ; (j = i0) → _ , idEquiv S¹ 211 | ; (j = i1) → _ , _ , rotIsEquiv (loop i) } ) 212 | 213 | codeS² : S² → hGroupoid _ 214 | codeS² s = ∥ HopfS² s ∥₁ , squash₁ 215 | 216 | codeTruncS² : ∥ S² ∥₂ → hGroupoid _ 217 | codeTruncS² = rec₂ (isOfHLevelTypeOfHLevel 3) codeS² 218 | 219 | encodeTruncS² : Ω ∥ S²∙ ∥₂∙ .fst → ∥ S¹ ∥₁ 220 | encodeTruncS² p = coe0→1 (λ i → codeTruncS² (p i) .fst) ∣ base ∣₁ 221 | 222 | g9 : Ω ∥ S¹∙ ∥₁∙ .fst → ∥ ℤ ∥₀ 223 | g9 p = coe0→1 (λ i → codeTruncS¹ (p i) .fst) ∣ pos 0 ∣₀ 224 | where 225 | codeS¹ : S¹ → hSet _ 226 | codeS¹ s = ∥ helix s ∥₀ , squash₀ 227 | 228 | codeTruncS¹ : ∥ S¹ ∥₁ → hSet _ 229 | codeTruncS¹ = rec₁ (isOfHLevelTypeOfHLevel 2) codeS¹ 230 | 231 | -- Use trick to eliminate away the truncation last 232 | g10 : ∥ ℤ ∥₀ → ℤ 233 | g10 = rec₀ isSetℤ (λ x → x) 234 | 235 | -- TODO: define η₁ and η₂ and some more maps 236 | 237 | -- Original η₃ as in the paper 238 | η₃ : join S¹ S¹ → Susp S² 239 | η₃ (inl x) = north 240 | η₃ (inr x) = north 241 | η₃ (push a b i) = 242 | (sym (σ (S¹×S¹→S² a b)) ∙ sym (σ (S¹×S¹→S² a b))) i 243 | 244 | -- Dropping the syms makes it compute fast (why?! maybe some slow inverse map gets applied with the sym?) 245 | η₃' : join S¹ S¹ → Susp S² 246 | η₃' (inl x) = north 247 | η₃' (inr x) = north 248 | η₃' (push a b i) = 249 | (σ (S¹×S¹→S² a b) ∙ σ (S¹×S¹→S² a b)) i 250 | 251 | -- Remark: dropping only one sym does not seem to make it fast enough 252 | 253 | -- Brunerie numbers 254 | 255 | β₃ : ℤ 256 | β₃ = g10 (g9 (g8 λ i j → f7 λ k → η₃ (push (loop i) (loop j) k))) 257 | 258 | -- This does not terminate fast 259 | -- β₃≡-2 : β₃ ≡ -2 260 | -- β₃≡-2 = refl 261 | 262 | β₃' : ℤ 263 | β₃' = g10 (g9 (g8 λ i j → f7 λ k → η₃' (push (loop i) (loop j) k))) 264 | 265 | -- This terminates fast 266 | β₃'≡-2 : β₃' ≡ -2 267 | β₃'≡-2 = refl 268 | 269 | β₃'-pos : ℤ 270 | β₃'-pos = g10 (g9 (g8 λ i j → f7 λ k → η₃' (push (loop (~ i)) (loop j) k))) 271 | 272 | brunerie'≡2 : β₃'-pos ≡ pos 2 273 | brunerie'≡2 = refl 274 | -------------------------------------------------------------------------------- /agda/Brunerie2.agda: -------------------------------------------------------------------------------- 1 | -- Mostly self contained definitions of the numbers from: https://arxiv.org/abs/2302.00151 2 | {-# OPTIONS --safe --cubical #-} 3 | module Brunerie2 where 4 | 5 | open import Cubical.Core.Primitives 6 | 7 | open import Cubical.Foundations.Prelude using 8 | (_,_ ; fst ; snd ; _≡_ ; cong ; _∙_ ; refl ; isProp→isSet ; J ; sym ; isProp ; isSet ; isGroupoid ; is2Groupoid; 9 | doubleComp-faces) 10 | open import Cubical.Foundations.Equiv using 11 | (_≃_ ; isEquiv ; isPropIsEquiv ; idIsEquiv ; idEquiv) 12 | open import Cubical.Foundations.HLevels using 13 | (hSet ; hGroupoid ; isOfHLevelTypeOfHLevel ; isOfHLevelPath ; isOfHLevelΠ ; isOfHLevel→isOfHLevelDep ; is2GroupoidΠ) 14 | open import Cubical.Foundations.Univalence using 15 | (Glue ; ua) 16 | open import Cubical.Data.Int using 17 | (ℤ ; pos ; neg ; isSetℤ ; sucPathℤ) 18 | open import Cubical.Foundations.CartesianKanOps 19 | 20 | private variable ℓ ℓ' ℓ'' : Level 21 | 22 | -- S1 23 | data S¹ : Type₀ where 24 | base : S¹ 25 | loop : base ≡ base 26 | 27 | helix : S¹ → Type₀ 28 | helix base = ℤ 29 | helix (loop i) = sucPathℤ i 30 | 31 | rotLoop : (a : S¹) → a ≡ a 32 | rotLoop base = loop 33 | rotLoop (loop i) j = 34 | hcomp (λ k → λ { (i = i0) → loop (j ∨ ~ k) 35 | ; (i = i1) → loop (j ∧ k) 36 | ; (j = i0) → loop (i ∨ ~ k) 37 | ; (j = i1) → loop (i ∧ k)}) base 38 | 39 | _·_ : S¹ → S¹ → S¹ 40 | base · x = x 41 | (loop i) · x = rotLoop x i 42 | 43 | isPropFamS¹ : (P : S¹ → Type ℓ) (pP : (x : S¹) → isProp (P x)) (b0 : P base) → 44 | PathP (λ i → P (loop i)) b0 b0 45 | isPropFamS¹ P pP b0 i = pP (loop i) (coe0→i (λ k → P (loop k)) i b0) 46 | (coe1→i (λ k → P (loop k)) i b0) i 47 | 48 | rotIsEquiv : (a : S¹) → isEquiv (a ·_) 49 | rotIsEquiv base = idIsEquiv S¹ 50 | rotIsEquiv (loop i) = isPropFamS¹ (λ x → isEquiv (x ·_)) 51 | (λ x → isPropIsEquiv (x ·_)) 52 | (idIsEquiv _) i 53 | 54 | -- S2 55 | data S² : Type₀ where 56 | base : S² 57 | surf : PathP (λ i → base ≡ base) refl refl 58 | 59 | S²ToSetElim : {A : S² → Type ℓ} 60 | → ((x : S²) → isSet (A x)) 61 | → A base 62 | → (x : S²) → A x 63 | S²ToSetElim set b base = b 64 | S²ToSetElim set b (surf i j) = 65 | isOfHLevel→isOfHLevelDep 2 set b b {a0 = refl} {a1 = refl} refl refl surf i j 66 | 67 | -- Join 68 | data join (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') where 69 | inl : A → join A B 70 | inr : B → join A B 71 | push : ∀ a b → inl a ≡ inr b 72 | 73 | 74 | -- SetTruncation 75 | data ∥_∥₀ (A : Type ℓ) : Type ℓ where 76 | ∣_∣₀ : A → ∥ A ∥₀ 77 | squash₀ : ∀ (x y : ∥ A ∥₀) (p q : x ≡ y) → p ≡ q 78 | 79 | rec₀ : {A : Type ℓ} {B : Type ℓ'} → isSet B → (A → B) → ∥ A ∥₀ → B 80 | rec₀ Bset f ∣ x ∣₀ = f x 81 | rec₀ Bset f (squash₀ x y p q i j) = 82 | Bset _ _ (cong (rec₀ Bset f) p) (cong (rec₀ Bset f) q) i j 83 | 84 | 85 | -- GroupoidTruncation 86 | data ∥_∥₁ (A : Type ℓ) : Type ℓ where 87 | ∣_∣₁ : A → ∥ A ∥₁ 88 | squash₁ : ∀ (x y : ∥ A ∥₁) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s 89 | 90 | rec₁ : {A : Type ℓ} {B : Type ℓ'} → isGroupoid B → (A → B) → ∥ A ∥₁ → B 91 | rec₁ gB f ∣ x ∣₁ = f x 92 | rec₁ gB f (squash₁ x y p q r s i j k) = 93 | gB _ _ _ _ (λ m n → rec₁ gB f (r m n)) (λ m n → rec₁ gB f (s m n)) i j k 94 | 95 | 96 | -- 2GroupoidTruncation 97 | data ∥_∥₂ (A : Type ℓ) : Type ℓ where 98 | ∣_∣₂ : A → ∥ A ∥₂ 99 | squash₂ : ∀ (x y : ∥ A ∥₂) (p q : x ≡ y) (r s : p ≡ q) (t u : r ≡ s) → t ≡ u 100 | 101 | rec₂ : ∀ {A : Type ℓ} {B : Type ℓ'} → is2Groupoid B → (A → B) → ∥ A ∥₂ → B 102 | rec₂ gB f ∣ x ∣₂ = f x 103 | rec₂ gB f (squash₂ _ _ _ _ _ _ t u i j k l) = 104 | gB _ _ _ _ _ _ (λ m n o → rec₂ gB f (t m n o)) (λ m n o → rec₂ gB f (u m n o)) 105 | i j k l 106 | 107 | -- rec₂Bin : ∀ {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → is2Groupoid C → (A → B → C) → ∥ A ∥₂ → ∥ B ∥₂ → C 108 | -- rec₂Bin gB f = rec₂ (is2GroupoidΠ λ _ → gB) λ x → rec₂ gB (f x) 109 | 110 | elim₂ : {A : Type ℓ} {B : ∥ A ∥₂ → Type ℓ} 111 | (bG : (x : ∥ A ∥₂) → is2Groupoid (B x)) 112 | (f : (x : A) → B ∣ x ∣₂) (x : ∥ A ∥₂) → B x 113 | elim₂ bG f ∣ x ∣₂ = f x 114 | elim₂ bG f (squash₂ x y p q r s u v i j k l) = 115 | isOfHLevel→isOfHLevelDep 4 bG _ _ _ _ _ _ 116 | (λ j k l → elim₂ bG f (u j k l)) (λ j k l → elim₂ bG f (v j k l)) 117 | (squash₂ x y p q r s u v) 118 | i j k l 119 | 120 | 121 | -- Susp 122 | data Susp (A : Type ℓ) : Type ℓ where 123 | north : Susp A 124 | south : Susp A 125 | merid : (a : A) → north ≡ south 126 | 127 | 128 | -- Pointed 129 | Pointed₀ : Type₁ 130 | Pointed₀ = Σ[ X ∈ Type₀ ] X 131 | 132 | S¹∙ S²∙ : Pointed₀ 133 | S¹∙ = (S¹ , base) 134 | S²∙ = (S² , base) 135 | 136 | Susp∙ : Type₀ → Pointed₀ 137 | Susp∙ A = Susp A , north 138 | 139 | ∥_∥₁∙ ∥_∥₂∙ : Pointed₀ → Pointed₀ 140 | ∥ A , a ∥₁∙ = ∥ A ∥₁ , ∣ a ∣₁ 141 | ∥ A , a ∥₂∙ = ∥ A ∥₂ , ∣ a ∣₂ 142 | 143 | Ω Ω² : Pointed₀ → Pointed₀ 144 | Ω (_ , a) = ((a ≡ a) , refl) 145 | Ω² p = Ω (Ω p) 146 | 147 | -- The maps 148 | σ : S² → Ω (Susp∙ S²) .fst 149 | σ x = merid x ∙ sym (merid base) 150 | 151 | S¹×S¹→S² : S¹ → S¹ → S² 152 | S¹×S¹→S² base y = base 153 | S¹×S¹→S² (loop i) base = base 154 | S¹×S¹→S² (loop i) (loop j) = surf i j 155 | 156 | f7 : Ω (Susp∙ S²) .fst → ∥ S² ∥₂ 157 | f7 p = coe0→1 (λ i → Code (p i)) ∣ base ∣₂ 158 | where 159 | _+₂_ : ∥ S² ∥₂ → ∥ S² ∥₂ → ∥ S² ∥₂ 160 | _+₂_ = rec₂ {A = S²}{∥ S² ∥₂ → ∥ S² ∥₂} (is2GroupoidΠ {A = ∥ S² ∥₂}{_}{λ _ → ∥ S² ∥₂} λ _ → squash₂) 161 | (λ { base x → x 162 | ; (surf i j) x → surfc x i j}) 163 | where 164 | surfc : (x : ∥ S² ∥₂) → refl {x = x} ≡ refl {x = x} 165 | surfc = elim₂ {A = S²}{λ x → refl {x = x} ≡ refl} 166 | (λ x → isOfHLevelPath {A = x ≡ x} 4 (isOfHLevelPath {A = ∥ S² ∥₂} 4 squash₂ x x) refl refl) 167 | (S²ToSetElim {A = λ x → refl {x = ∣ x ∣₂} ≡ refl} 168 | (λ x → squash₂ ∣ x ∣₂ ∣ x ∣₂ refl refl) 169 | (λ i j → ∣ surf i j ∣₂)) 170 | 171 | ∥S²∥₂≃∥S²∥₂ : (x : S²) → ∥ S² ∥₂ ≃ ∥ S² ∥₂ 172 | fst (∥S²∥₂≃∥S²∥₂ x) = ∣ x ∣₂ +₂_ 173 | snd (∥S²∥₂≃∥S²∥₂ x) = help x 174 | where 175 | help : (x : S²) → isEquiv {A = ∥ S² ∥₂}{∥ S² ∥₂} (∣ x ∣₂ +₂_) 176 | help = S²ToSetElim {A = λ x → isEquiv (_+₂_ ∣ x ∣₂)} 177 | (λ x → isProp→isSet {A = isEquiv {A = ∥ S² ∥₂}{∥ S² ∥₂} (_+₂_ ∣ x ∣₂)} 178 | (isPropIsEquiv {A = ∥ S² ∥₂}{B = ∥ S² ∥₂} (_+₂_ ∣ x ∣₂))) 179 | (idIsEquiv ∥ S² ∥₂) 180 | 181 | Code : Susp S² → Type₀ 182 | Code north = ∥ S² ∥₂ 183 | Code south = ∥ S² ∥₂ 184 | Code (merid a i) = ua (∥S²∥₂≃∥S²∥₂ a) i 185 | 186 | g8 : Ω² ∥ S²∙ ∥₂∙ .fst → Ω ∥ S¹∙ ∥₁∙ .fst 187 | g8 p i = coe0→1 (λ j → codeTruncS² (p i j) .fst) ∣ base ∣₁ 188 | where 189 | HopfS² : S² → Type₀ 190 | HopfS² base = S¹ 191 | HopfS² (surf i j) = Glue S¹ (λ { (i = i0) → S¹ , idEquiv S¹ 192 | ; (i = i1) → S¹ , idEquiv S¹ 193 | ; (j = i0) → S¹ , idEquiv S¹ 194 | ; (j = i1) → S¹ , (loop i) ·_ , rotIsEquiv (loop i) } ) 195 | 196 | codeTruncS² : ∥ S² ∥₂ → hGroupoid ℓ-zero 197 | codeTruncS² = rec₂ {A = S²}{hGroupoid ℓ-zero} 198 | (isOfHLevelTypeOfHLevel 3) 199 | (λ s → ∥ HopfS² s ∥₁ , squash₁) 200 | 201 | codeTruncS¹ : ∥ S¹ ∥₁ → hSet _ 202 | codeTruncS¹ = rec₁ {A = S¹}{hSet _} (isOfHLevelTypeOfHLevel 2) (λ s → ∥ helix s ∥₀ , squash₀) 203 | 204 | g9 : Ω ∥ S¹∙ ∥₁∙ .fst → ∥ ℤ ∥₀ 205 | g9 p = coe0→1 (λ i → codeTruncS¹ (p i) .fst) ∣ pos 0 ∣₀ 206 | 207 | -- Use trick to eliminate away the truncation last 208 | g10 : ∥ ℤ ∥₀ → ℤ 209 | g10 = rec₀ {A = ℤ}{ℤ} isSetℤ (λ x → x) 210 | 211 | -- TODO: define η₁ and η₂ and some more maps 212 | 213 | -- Original η₃ as in the paper 214 | η₃ : join S¹ S¹ → Susp S² 215 | η₃ (inl x) = north 216 | η₃ (inr x) = north 217 | η₃ (push a b i) = 218 | (sym (σ (S¹×S¹→S² a b)) ∙ sym (σ (S¹×S¹→S² a b))) i 219 | 220 | -- Dropping the syms makes it compute fast (why?! maybe some slow inverse map gets applied with the sym?) 221 | η₃' : join S¹ S¹ → Susp S² 222 | η₃' (inl x) = north 223 | η₃' (inr x) = north 224 | η₃' (push a b i) = 225 | (σ (S¹×S¹→S² a b) ∙ σ (S¹×S¹→S² a b)) i 226 | 227 | -- Remark: dropping only one sym does not seem to make it fast enough 228 | 229 | -- Brunerie numbers 230 | 231 | β₃ : ℤ 232 | β₃ = g10 (g9 (g8 λ i j → f7 λ k → η₃ (push (loop i) (loop j) k))) 233 | 234 | -- This does not terminate fast 235 | -- β₃≡-2 : β₃ ≡ -2 236 | -- β₃≡-2 = refl 237 | 238 | β₃' : ℤ 239 | β₃' = g10 (g9 (g8 λ i j → f7 λ k → η₃' (push (loop i) (loop j) k))) 240 | 241 | ---------------------------------------------------------------------------------------------------- 242 | 243 | v1 : (i j k : I) → Susp S² 244 | v1 i j k = η₃' (push (loop i) (loop j) k) 245 | 246 | data List (A : Set) : Set where 247 | nil : List A 248 | cons : A → List A → List A 249 | 250 | bnd3 : {A : Set}(f : I → I → I → A) → List A 251 | bnd3 f = 252 | cons (f i0 i0 i0) ( 253 | cons (f i0 i0 i1) ( 254 | cons (f i0 i1 i0) ( 255 | cons (f i0 i1 i1) ( 256 | cons (f i1 i0 i0) ( 257 | cons (f i1 i0 i1) ( 258 | cons (f i1 i1 i0) ( 259 | cons (f i1 i1 i1) nil))))))) 260 | 261 | bnd2 : {A : Set}(f : I → I → A) → List A 262 | bnd2 f = 263 | cons (f i0 i0) ( 264 | cons (f i0 i1) ( 265 | cons (f i1 i0) ( 266 | cons (f i1 i1) nil))) 267 | 268 | v2 : (i j : I) → ∥ S² ∥₂ 269 | v2 i j = f7 λ k → η₃' (push (loop i) (loop j) k) 270 | 271 | -- nf doesn't terminate 272 | v3 : (Ω ∥ S¹∙ ∥₁∙) .fst 273 | v3 = g8 (λ i j → f7 λ k → η₃' (push (loop i) (loop j) k)) 274 | 275 | v4 = g9 v3 -- ∣ ℤ.negsuc 1 ∣₀ 276 | 277 | v5 = g10 v4 278 | 279 | foo : ℤ 280 | foo = -2 281 | 282 | -------------------------------------------------------------------------------- 283 | 284 | -- This terminates fast 285 | β₃'≡-2 : β₃' ≡ -2 286 | β₃'≡-2 = refl 287 | 288 | β₃'-pos : ℤ 289 | β₃'-pos = g10 (g9 (g8 λ i j → f7 λ k → η₃' (push (loop (~ i)) (loop j) k))) 290 | 291 | brunerie'≡2 : β₃'-pos ≡ pos 2 292 | brunerie'≡2 = refl 293 | -------------------------------------------------------------------------------- /agda/HLevelsTemp.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | module HLevelsTemp where 3 | 4 | open import Cubical.Foundations.Prelude 5 | open import Cubical.Foundations.Equiv 6 | open import Cubical.Foundations.Univalence 7 | open import Cubical.Foundations.CartesianKanOps 8 | 9 | open import Cubical.Data.Sigma 10 | 11 | hSet hGroupoid : ∀ ℓ → Type (ℓ-suc ℓ) 12 | hSet ℓ = Σ[ A ∈ Type ℓ ] isSet A 13 | hGroupoid ℓ = Σ[ A ∈ Type ℓ ] isGroupoid A 14 | 15 | private 16 | variable 17 | ℓ : Level 18 | A : Type ℓ 19 | B : A → Type ℓ 20 | 21 | isPropIsSet : isProp (isSet A) 22 | isPropIsSet {A = A} f g i a b = isPropIsProp {A = a ≡ b} (f a b) (g a b) i 23 | 24 | isPropIsGroupoid : isProp (isGroupoid A) 25 | isPropIsGroupoid {A = A} f g i a b = isPropIsSet {A = a ≡ b} (f a b) (g a b) i 26 | 27 | isPropΣ : isProp A → ((x : A) → isProp (B x)) → isProp (Σ A B) 28 | isPropΣ pA pB t u = Σ≡Prop pB (pA (t .fst) (u .fst)) 29 | 30 | isPropRetract : {B : Type ℓ} (f : A → B) (g : B → A) (h : (x : A) → g (f x) ≡ x) → isProp B → isProp A 31 | isPropRetract f g h p x y i = 32 | hcomp (λ j → λ { (i = i0) → h x j 33 | ; (i = i1) → h y j}) 34 | (g (p (f x) (f y) i)) 35 | 36 | isSetRetract : {B : Type ℓ} (f : A → B) (g : B → A) (h : (x : A) → g (f x) ≡ x) → isSet B → isSet A 37 | isSetRetract f g h set x y p q i j = 38 | hcomp (λ k → λ { (i = i0) → h (p j) k 39 | ; (i = i1) → h (q j) k 40 | ; (j = i0) → h x k 41 | ; (j = i1) → h y k}) 42 | (g (set (f x) (f y) (cong f p) (cong f q) i j)) 43 | 44 | isGroupoidRetract : {B : Type ℓ} (f : A → B) (g : B → A) (h : (x : A) → g (f x) ≡ x) → isGroupoid B → isGroupoid A 45 | isGroupoidRetract f g h grp x y p q P Q i j k = 46 | hcomp ((λ l → λ { (i = i0) → h (P j k) l 47 | ; (i = i1) → h (Q j k) l 48 | ; (j = i0) → h (p k) l 49 | ; (j = i1) → h (q k) l 50 | ; (k = i0) → h x l 51 | ; (k = i1) → h y l})) 52 | (g (grp (f x) (f y) (cong f p) (cong f q) (cong (cong f) P) (cong (cong f) Q) i j k)) 53 | 54 | congFst⁻ : (pB : (x : A) → isProp (B x)) (t u : Σ A B) → t .fst ≡ u .fst → t ≡ u 55 | congFst⁻ {B = B} pB t u q i = (q i) , 56 | hcomp (λ j → λ { (i = i0) → pB (t .fst) (t .snd) (t .snd) j 57 | ; (i = i1) → pB (u .fst) (coe0→1 (λ k → B (q k)) (t .snd)) (u .snd) j }) 58 | (coe0→i (λ x → B (q x)) i (t .snd)) 59 | 60 | congFst⁻-congFst : (pB : (x : A) → isProp (B x)) (t u : Σ A B) → (p : t ≡ u) → congFst⁻ pB t u (cong fst p) ≡ p 61 | congFst⁻-congFst {B = B} pB t u p j i = 62 | (p i .fst) , 63 | (hcomp {A = B (p i .fst)} 64 | (λ k → λ { (i = i0) → pB (t .fst) (t .snd) (t .snd) k 65 | ; (i = i1) → pB (u .fst) (coe0→1 (λ k → B (p k .fst)) (t .snd)) (u .snd) k 66 | ; (j = i1) → pB (p i .fst) (coe0→i (λ k → B (p k .fst)) i (t .snd)) (p i .snd) k }) 67 | (coe0→i (λ x → B (p x .fst)) i (t .snd))) 68 | 69 | isSetSndProp : (pB : (x : A) → isProp (B x)) (sA : (t u : Σ A B) → isProp (t .fst ≡ u .fst)) → isSet (Σ A B) 70 | isSetSndProp pB sA t u = 71 | isPropRetract {A = t ≡ u}{fst t ≡ fst u} (cong fst) (congFst⁻ pB t u) (congFst⁻-congFst pB t u) (sA t u) 72 | 73 | isGroupoidSndProp : (pB : (x : A) → isProp (B x)) → (sA : (t u : Σ A B) → isSet (t .fst ≡ u .fst)) → isGroupoid (Σ A B) 74 | isGroupoidSndProp pB sA t u = 75 | isSetRetract {A = t ≡ u}{fst t ≡ fst u} (cong fst) (congFst⁻ pB t u) (congFst⁻-congFst pB t u) (sA t u) 76 | 77 | is2GroupoidSndProp : (pB : (x : A) → isProp (B x)) → (sA : (t u : Σ A B) → isGroupoid (t .fst ≡ u .fst)) → is2Groupoid (Σ A B) 78 | is2GroupoidSndProp pB sA t u = isGroupoidRetract {A = t ≡ u}{B = fst t ≡ fst u} 79 | (cong fst) (congFst⁻ pB t u) (congFst⁻-congFst pB t u) (sA t u) 80 | 81 | isSetΠ : ((x : A) → isSet (B x)) → isSet ((x : A) → B x) 82 | isSetΠ h x y p q i j z = h z (x z) (y z) (λ k → p k z) (λ k → q k z) i j 83 | 84 | setPath : (A B : Type ℓ) (sB : isSet B) → isSet (A ≡ B) 85 | setPath A B sB = 86 | isSetRetract {A = A ≡ B}{B = A ≃ B} 87 | (pathToEquiv {A = A}{B = B}) 88 | (ua {A = A}{B}) 89 | (ua-pathToEquiv {A = A}{B}) 90 | (isSetSndProp {A = A → B}{B = isEquiv {A = A}{B}} 91 | (isPropIsEquiv {A = A}{B = B}) 92 | (λ e1 e2 → isSetΠ {A = A}{B = λ _ → B} (λ _ → sB) (e1 .fst) (e2 .fst))) 93 | 94 | isGroupoid-hSet : isGroupoid (hSet ℓ) 95 | isGroupoid-hSet = 96 | isGroupoidSndProp {A = Type _}{B = isSet} 97 | (λ A → isPropIsSet {A = A}) 98 | (λ t u → setPath (t .fst) (u .fst) (u .snd)) 99 | 100 | 101 | -- Next result 102 | isGroupoidΠ : ((x : A) → isGroupoid (B x)) → isGroupoid ((x : A) → B x) 103 | isGroupoidΠ h x y p q r s i j k z = h z (x z) (y z) (λ k → p k z) (λ k → q k z) (λ k l → r k l z) (λ k l → s k l z) i j k 104 | 105 | groupoidPath : (A B : Type ℓ) (sB : isGroupoid B) → isGroupoid (A ≡ B) 106 | groupoidPath A B sB = 107 | isGroupoidRetract {A = A ≡ B}{B = A ≃ B} 108 | (pathToEquiv {A = A}{B = B}) 109 | (ua {A = A}{B = B}) 110 | (ua-pathToEquiv {A = A}{B = B}) 111 | (isGroupoidSndProp {A = A → B}{B = isEquiv} 112 | isPropIsEquiv 113 | (λ e1 e2 → isGroupoidΠ {A = A}{B = λ _ → B} (λ _ → sB) (e1 .fst) (e2 .fst))) 114 | 115 | is2Groupoid-hGroupoid : is2Groupoid (hGroupoid ℓ) 116 | is2Groupoid-hGroupoid = 117 | is2GroupoidSndProp {A = Type _}{B = isGroupoid} 118 | (λ A → isPropIsGroupoid {A = A}) 119 | (λ t u → groupoidPath (t .fst) (u .fst) (u .snd)) 120 | 121 | 122 | -------------------------------------------------------------------------------- 123 | -------------------------------------------------------------------------------- /agda/S1Bench.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --cubical #-} 2 | 3 | module S1Bench where 4 | 5 | open import Cubical.Foundations.Prelude 6 | open import Cubical.HITs.S1.Base 7 | open import Cubical.Data.Nat 8 | open import Cubical.Data.Int 9 | open import Cubical.Data.Unit 10 | 11 | forceℕ : ℕ → Unit 12 | forceℕ zero = tt 13 | forceℕ (suc n) = forceℕ n 14 | 15 | forceℤ : ℤ → Unit 16 | forceℤ (pos n) = forceℕ n 17 | forceℤ (negsuc n) = forceℕ n 18 | 19 | test1 : forceℤ (winding (intLoop (pos 1000))) ≡ tt 20 | test1 = refl 21 | 22 | -- test2 = forceℤ (winding (intLoop (neg 10))) 23 | -- test3 = winding (intLoop (pos 10) ∙ intLoop (neg 10)) 24 | -------------------------------------------------------------------------------- /cctt.el: -------------------------------------------------------------------------------- 1 | ;;; cctt.el --- Mode for the cctt programming language -*- lexical-binding: t -*- 2 | ;; URL: https://github.com/AndrasKovacs/cctt 3 | ;; Package-version: 1.0 4 | ;; Package-Requires: ((emacs "24.1") (cl-lib "0.5")) 5 | ;; Keywords: languages 6 | 7 | ;;; Commentary: 8 | ;; This package provides a major mode for editing proofs or programs 9 | ;; in cctt, an implementation of a cartesian cubical type theory. 10 | 11 | (require 'comint) 12 | (require 'cl-lib) 13 | 14 | ;;;; Customization options 15 | 16 | (defgroup cctt nil "Options for cctt-mode" 17 | :group 'languages 18 | :prefix 'cctt- 19 | :tag "cctt") 20 | 21 | ;;;; Syntax 22 | 23 | (defvar cctt-keywords 24 | '("inductive" "higher" "import" "let" "module") 25 | "Keywords.") 26 | 27 | (defvar cctt-operations 28 | '("Glue" "glue" "unglue" "hcom" "com" "coe" "refl" "U" "I" "ap" "case") 29 | "Operations.") 30 | 31 | (defvar cctt-special 32 | '("undefined") 33 | "Special operators.") 34 | 35 | (defvar cctt-keywords-regexp 36 | (regexp-opt cctt-keywords 'words) 37 | "Regexp that recognizes keywords.") 38 | 39 | (defvar cctt-operations-regexp 40 | (regexp-opt cctt-operations 'words) 41 | "Regexp that recognizes operations.") 42 | 43 | (defvar cctt-operators-regexp 44 | (regexp-opt '(":" "->" "→" "∙" "=" ":=" "|" "\\" "λ" "*" "×" "_" "@" "." "⁻¹" "[" "]") t) 45 | "Regexp that recognizes operators.") 46 | 47 | (defvar cctt-special-regexp 48 | (regexp-opt cctt-special 'words) 49 | "Regexp that recognizes special operators.") 50 | 51 | (defvar cctt-def-regexp "^[[:word:]'-]+" 52 | "Regexp that recognizes the beginning of a definition.") 53 | 54 | (defvar cctt-font-lock-keywords 55 | `((,cctt-keywords-regexp . font-lock-keyword-face) 56 | (,cctt-operations-regexp . font-lock-builtin-face) 57 | (,cctt-operators-regexp . font-lock-variable-name-face) 58 | (,cctt-special-regexp . font-lock-warning-face) 59 | (,cctt-def-regexp . font-lock-function-name-face)) 60 | "Font-lock information, assigning each class of keyword a face.") 61 | 62 | (defvar cctt-syntax-table 63 | (let ((st (make-syntax-table))) 64 | (modify-syntax-entry ?\{ "(}1nb" st) 65 | (modify-syntax-entry ?\} "){4nb" st) 66 | (modify-syntax-entry ?- "_ 123" st) 67 | (modify-syntax-entry ?\n ">" st) 68 | (modify-syntax-entry ?\\ "." st) 69 | st) 70 | "Syntax table with Haskell-style comments.") 71 | 72 | (defun cctt-indent-line () 73 | "Indent current line." 74 | (insert-char ?\s 2)) 75 | 76 | ;;;###autoload 77 | (define-derived-mode cctt-mode prog-mode 78 | "cctt" 79 | "Major mode for editing cctt files." 80 | 81 | :syntax-table cctt-syntax-table 82 | 83 | ;; Make comment-dwim do the right thing for Cubical 84 | (set (make-local-variable 'comment-start) "--") 85 | (set (make-local-variable 'comment-end) "") 86 | 87 | ;; Code for syntax highlighting 88 | (setq font-lock-defaults '(cctt-font-lock-keywords)) 89 | 90 | ;; Indentation 91 | (setq indent-line-function 'cctt-indent-line) 92 | 93 | ;; Clear memory 94 | (setq cctt-keywords-regexp nil) 95 | (setq cctt-operators-regexp nil) 96 | (setq cctt-special-regexp nil)) 97 | 98 | ;;;###autoload 99 | (add-to-list 'auto-mode-alist '("\\.cctt\\'" . cctt-mode)) 100 | 101 | (provide 'cctt) 102 | -------------------------------------------------------------------------------- /hott2023/EasyChair3.5/chairEC.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/EasyChair3.5/chairEC.pdf -------------------------------------------------------------------------------- /hott2023/EasyChair3.5/debug.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/EasyChair3.5/debug.pdf -------------------------------------------------------------------------------- /hott2023/EasyChair3.5/easychair.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/EasyChair3.5/easychair.pdf -------------------------------------------------------------------------------- /hott2023/EasyChair3.5/epic.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/EasyChair3.5/epic.pdf -------------------------------------------------------------------------------- /hott2023/EasyChair3.5/logoEC.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/EasyChair3.5/logoEC.pdf -------------------------------------------------------------------------------- /hott2023/EasyChair3.5/throneEC.jpg: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/EasyChair3.5/throneEC.jpg -------------------------------------------------------------------------------- /hott2023/HoTT-2023_abstract_23.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/HoTT-2023_abstract_23.pdf -------------------------------------------------------------------------------- /hott2023/hott-2023-template.aux: -------------------------------------------------------------------------------- 1 | \relax 2 | \providecommand\hyper@newdestlabel[2]{} 3 | \providecommand\HyperFirstAtBeginDocument{\AtBeginDocument} 4 | \HyperFirstAtBeginDocument{\ifx\hyper@anchor\@undefined 5 | \global\let\oldcontentsline\contentsline 6 | \gdef\contentsline#1#2#3#4{\oldcontentsline{#1}{#2}{#3}} 7 | \global\let\oldnewlabel\newlabel 8 | \gdef\newlabel#1#2{\newlabelxx{#1}#2} 9 | \gdef\newlabelxx#1#2#3#4#5#6{\oldnewlabel{#1}{{#2}{#3}}} 10 | \AtEndDocument{\ifx\hyper@anchor\@undefined 11 | \let\contentsline\oldcontentsline 12 | \let\newlabel\oldnewlabel 13 | \fi} 14 | \fi} 15 | \global\let\hyper@last\relax 16 | \gdef\HyperFirstAtBeginDocument#1{#1} 17 | \providecommand\HyField@AuxAddToFields[1]{} 18 | \providecommand\HyField@AuxAddToCoFields[2]{} 19 | \bibstyle{alphaurl} 20 | \citation{abcfhl} 21 | \citation{ctt-normalization} 22 | \bibdata{references} 23 | \bibcite{abcfhl}{ABC{$^{+}$}21} 24 | \bibcite{ctt-normalization}{SA21} 25 | \newlabel{LastPage}{{}{2}{}{page.2}{}} 26 | \xdef\lastpage@lastpage{2} 27 | \xdef\lastpage@lastpageHy{2} 28 | \gdef \@abspage@last{2} 29 | -------------------------------------------------------------------------------- /hott2023/hott-2023-template.bbl: -------------------------------------------------------------------------------- 1 | \newcommand{\etalchar}[1]{$^{#1}$} 2 | \begin{thebibliography}{ABC{\etalchar{+}}21} 3 | 4 | \bibitem[ABC{\etalchar{+}}21]{abcfhl} 5 | Carlo Angiuli, Guillaume Brunerie, Thierry Coquand, Robert Harper, 6 | Kuen{-}Bang~Hou (Favonia), and Daniel~R. Licata. 7 | \newblock Syntax and models of cartesian cubical type theory. 8 | \newblock {\em Math. Struct. Comput. Sci.}, 31(4):424--468, 2021. 9 | \newblock \href {https://doi.org/10.1017/S0960129521000347} 10 | {\path{doi:10.1017/S0960129521000347}}. 11 | 12 | \bibitem[SA21]{ctt-normalization} 13 | Jonathan Sterling and Carlo Angiuli. 14 | \newblock Normalization for cubical type theory. 15 | \newblock In {\em 36th Annual {ACM/IEEE} Symposium on Logic in Computer 16 | Science, {LICS} 2021, Rome, Italy, June 29 - July 2, 2021}, pages 1--15. 17 | {IEEE}, 2021. 18 | \newblock \href {https://doi.org/10.1109/LICS52264.2021.9470719} 19 | {\path{doi:10.1109/LICS52264.2021.9470719}}. 20 | 21 | \end{thebibliography} 22 | -------------------------------------------------------------------------------- /hott2023/hott-2023-template.blg: -------------------------------------------------------------------------------- 1 | This is BibTeX, Version 0.99d (TeX Live 2022/dev/Debian) 2 | Capacity: max_strings=200000, hash_size=200000, hash_prime=170003 3 | The top-level auxiliary file: hott-2023-template.aux 4 | The style file: alphaurl.bst 5 | Reallocated glb_str_ptr (elt_size=4) to 20 items from 10. 6 | Reallocated global_strs (elt_size=200001) to 20 items from 10. 7 | Reallocated glb_str_end (elt_size=4) to 20 items from 10. 8 | Reallocated singl_function (elt_size=4) to 100 items from 50. 9 | Reallocated singl_function (elt_size=4) to 100 items from 50. 10 | Database file #1: references.bib 11 | Reallocated wiz_functions (elt_size=4) to 6000 items from 3000. 12 | You've used 2 entries, 13 | 3307 wiz_defined-function locations, 14 | 702 strings with 6129 characters, 15 | and the built_in function-call counts, 1365 in all, are: 16 | = -- 142 17 | > -- 62 18 | < -- 1 19 | + -- 21 20 | - -- 21 21 | * -- 97 22 | := -- 226 23 | add.period$ -- 9 24 | call.type$ -- 2 25 | change.case$ -- 16 26 | chr.to.int$ -- 2 27 | cite$ -- 2 28 | duplicate$ -- 36 29 | empty$ -- 105 30 | format.name$ -- 22 31 | if$ -- 320 32 | int.to.chr$ -- 1 33 | int.to.str$ -- 0 34 | missing$ -- 2 35 | newline$ -- 16 36 | num.names$ -- 6 37 | pop$ -- 20 38 | preamble$ -- 1 39 | purify$ -- 18 40 | quote$ -- 0 41 | skip$ -- 81 42 | stack$ -- 0 43 | substring$ -- 58 44 | swap$ -- 12 45 | text.length$ -- 1 46 | text.prefix$ -- 0 47 | top$ -- 0 48 | type$ -- 16 49 | warning$ -- 0 50 | while$ -- 11 51 | width$ -- 4 52 | write$ -- 34 53 | -------------------------------------------------------------------------------- /hott2023/hott-2023-template.out: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/hott-2023-template.out -------------------------------------------------------------------------------- /hott2023/hott-2023-template.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/hott-2023-template.pdf -------------------------------------------------------------------------------- /hott2023/hott-2023-template.tex: -------------------------------------------------------------------------------- 1 | % This is a template for submissions to HoTT 2023. 2 | % The file was created on the basis of easychair.tex,v 3.5 2017/03/15 3 | % 4 | % Before submitting, rename the resulting pdf file to "yourname.pdf" 5 | % 6 | \documentclass[letterpaper]{easychair} 7 | \usepackage{doc} 8 | \usepackage{amsfonts} 9 | \usepackage{bm} 10 | 11 | \newcommand{\easychair}{\textsf{easychair}} 12 | \newcommand{\miktex}{MiK{\TeX}} 13 | \newcommand{\texniccenter}{{\TeX}nicCenter} 14 | \newcommand{\makefile}{\texttt{Makefile}} 15 | \newcommand{\latexeditor}{LEd} 16 | 17 | \newcommand{\msf}[1]{\mathsf{#1}} 18 | \renewcommand{\mit}[1]{\mathit{#1}} 19 | \newcommand{\mbb}[1]{\mathbb{#1}} 20 | \newcommand{\mbf}[1]{\mathbf{#1}} 21 | \newcommand{\bs}[1]{\boldsymbol{#1}} 22 | \newcommand{\Nat}{\mbb{N}} 23 | \newcommand{\suc}{\msf{suc}} 24 | 25 | \bibliographystyle{alphaurl} 26 | 27 | % \authorrunning{} has to be set for the shorter version of the authors' names; 28 | % otherwise a warning will be rendered in the running heads. When processed by 29 | % EasyChair, this command is mandatory: a document without \authorrunning 30 | % will be rejected by EasyChair 31 | \authorrunning{} 32 | 33 | % \titlerunning{} has to be set to either the main title or its shorter 34 | % version for the running heads. When processed by 35 | % EasyChair, this command is mandatory: a document without \titlerunning 36 | % will be rejected by EasyChair 37 | \titlerunning{} 38 | 39 | % 40 | \title{Efficient Evaluation for Cubical Type Theories 41 | % \thanks{Your acknowledgements.} 42 | } 43 | 44 | \author{ 45 | András Kovács 46 | } 47 | 48 | \institute{ 49 | Eötvös Loránd University\\ 50 | \email{kovacsandras@inf.elte.hu} 51 | } 52 | 53 | \begin{document} 54 | \maketitle 55 | 56 | Currently, cubical type theories are the only known systems which support 57 | computational univalence. We can use computation in these systems to shortcut 58 | some proofs, by appealing to definitional equality of sides of 59 | equations. However, efficiency issues in existing implementations often preclude 60 | such computational proofs, or it takes a large amount of effort to find 61 | definitions which are feasible to compute. In this abstract we investigate the 62 | efficiency of the ABCFHL \cite{abcfhl} Cartesian cubical type theory with 63 | separate homogeneous composition ($\msf{hcom}$) and coercion ($\msf{coe}$), 64 | although most of our findings transfer to other systems. 65 | 66 | \subsubsection*{Cubical normalization-by-evaluation} 67 | 68 | In variants of non-cubical Martin-Löf type theory, definitional equalities are 69 | specified by reference to a substitution operation on terms. However, well-known 70 | efficient implementations do not actually use term substitution. Instead, 71 | \emph{normalization-by-evaluation} (NbE) is used, which corresponds to certain 72 | \emph{environment machines} from a more operational point of view. In these 73 | setups, there is a distinction between syntactic terms and semantic 74 | values. Terms are viewed as immutable program code that supports evaluation into 75 | the semantic domain but no other operations. 76 | 77 | In contrast, in cubical type theories interval substitution is an essential 78 | component of computation which seemingly cannot be removed from the 79 | semantics. Most existing implementations use NbE for ordinary non-cubical 80 | computation, but also include interval substitution as an operation that acts on 81 | semantic values. Unfortunately, a naive combination of NbE and interval 82 | substitution performs poorly, as it destroys the implicit sharing of work 83 | and structure which underlies the efficiency of NbE in the first place. We 84 | propose a restructured cubical NbE which handles interval substitution more 85 | gracefully. The basic operations are the following. 86 | 87 | \begin{enumerate} 88 | \item \emph{Evaluation} maps from syntax to semantics like before, but it 89 | additionally takes as input an interval environment and a cofibration. 90 | \item \emph{Interval substitution} acts on values, but it has trivial cost by itself; 91 | it only shallowly stores an explicit substitution. 92 | \item \emph{Forcing} computes a value to weak head form by sufficiently computing 93 | previously stored delayed substitutions. 94 | \end{enumerate} 95 | 96 | 97 | On canonical values, forcing simply pushes substitutions further down, incurring 98 | minimal cost. But on neutral values, since neutrals are not stable under 99 | substitution, forcing has to potentially perform arbitrary computation. Here we 100 | take a hint from the formal cubical NbE by Sterling and Angiuli 101 | \cite{ctt-normalization}, by annotating neutral values with stability 102 | information. This allows us to quickly determine whether a neutral value is 103 | stable under a given substitution. When it is stable, forcing does not have to 104 | look inside it. 105 | 106 | It turns out that there is only a single computation rule in the ABCFHL theory 107 | which can trigger interval substitution with significant cost: the coercion rule 108 | for the $\msf{Glue}$ type former. In every other case, only a \emph{weakening} 109 | substitution may be created, but all neutral values are stable under weakening, 110 | so forcing by weakening always has a trivial cost. 111 | 112 | \subsubsection*{Using canonicity in closed evaluation} 113 | 114 | In non-cubical type theories, evaluation of closed terms can be more efficient 115 | than that of open terms. For instance, when we evaluate an 116 | $\msf{if\!-\!then\!-\!else}$ expression, we know that exactly one branch will be 117 | taken. In open evaluation, the $\msf{Bool}$ scrutinee may be neutral, in which 118 | case both branches may have to be evaluated. 119 | 120 | In the cubical setting, \emph{systems of partial values} can be viewed as 121 | branching structures which make case distinctions on cofibrations. Importantly, 122 | there are computation rules which scrutinize \emph{all} components of a cubical 123 | system. These are precisely the homogeneous composition rules ($\msf{hcom}$) for 124 | strict inductive types. For example: 125 | \[ \msf{hcom}^{r\to r'}_{\mbb{N}}\,[\psi \mapsto i.\,\suc\,t]\,(\suc\,b) = 126 | \suc\,(\msf{hcom}^{r\to r'}_{\mbb{N}}\,[\psi \mapsto i.\,t]\,b) \] 127 | When we only have interval variables and a cofibration in the context, we do not 128 | have to compute every system component to check for $\suc$. In this case, which 129 | we may call ``closed cubical'', we can use the canonicity property of the 130 | theory. Here $\suc\,b$ in the $\msf{hcom}$ base implies that every system 131 | component is headed by $\suc$ as well. Hence, we can use the following rule 132 | instead: 133 | \[ \msf{hcom}^{r\to r'}_{\mbb{N}}\,[\psi \mapsto i.\,t]\,(\suc\,b) = 134 | \suc\,(\msf{hcom}^{r\to r'}_{\mbb{N}}\,[\psi \mapsto i.\,\msf{pred}\,t]\,b) \] 135 | Here, $\msf{pred}$ is a metatheoretic function which takes the predecessor of a 136 | value which is already known to be definitionally $\suc$. The revised rule 137 | assumes nothing about the shape of $t$ on the left hand side, so we can compute 138 | $\msf{pred}$ lazily in the output. These lazy projections work analogously for 139 | all non-higher inductive types. For higher-inductive types, $\msf{hcom}$ is a 140 | canonical value, so there is no efficiency issue to begin with. 141 | 142 | \subsubsection*{Summary} 143 | 144 | \begin{itemize} 145 | \item Costly interval substitution can only arise from computing with $\msf{Glue}$ types. 146 | \item In closed cubical evaluation, no computation rule forces all components of a system. 147 | \end{itemize} 148 | We are optimistic that an implementation with these properties would yield significant 149 | performance improvements. We are currently in the process of developing this system 150 | and adapting existing benchmarks to it. 151 | 152 | \bibliography{references} 153 | 154 | \end{document} 155 | -------------------------------------------------------------------------------- /hott2023/prez/AndrasKovacs.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/prez/AndrasKovacs.pdf -------------------------------------------------------------------------------- /hott2023/prez/brunerie.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/prez/brunerie.png -------------------------------------------------------------------------------- /hott2023/prez/unglues.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/hott2023/prez/unglues.png -------------------------------------------------------------------------------- /newnotes/Connections.txt: -------------------------------------------------------------------------------- 1 | 2 | Adding Connections 3 | ------------------------------------------------------------ 4 | -------------------------------------------------------------------------------- /newnotes/notes.txt: -------------------------------------------------------------------------------- 1 | 2 | Syntactic context: 3 | 4 | ψ, α, Γ (ivars, cofib, vars) 5 | 6 | Semantic context is a "morphism" into a cxt 7 | 8 | src: (ψ₀, α₀, Γ₀) 9 | tgt: (ψ₁, α₁, Γ₁) 10 | 11 | consisting of: 12 | 13 | (σ : ISub ψ₀ ψ₁) × (f : α₀ ⊢ α₁[σ]) × (δ : Sub Γ₀ (Γ₁[σ, f])) 14 | 15 | We store: 16 | - ψ₀ because because we'll have to get fresh ivars for weakening in 17 | filling. 18 | - σ 19 | - α₀, but not "f", because 20 | 21 | For "σ", we store σ itself , 22 | 23 | For the "f" component, we only want to store ψ₀. 24 | 25 | 26 | We want to distinguish ISub-s and cofibs 27 | - it might not be possible to unify their representation 28 | - we want to handle them differently: 29 | - ISub-s should be explicitly stored and pushed into sub-values in forcing 30 | - cofibs should not be explicitly stored, we should just force w.r.t. cofib weakening 31 | when it's needed, in neutral cases 32 | 33 | Principles: 34 | - only explicit isub are stored 35 | - eval does not force anything, just builds neutrals 36 | 37 | Closures: 38 | - just capture all 4 eval params plus a term 39 | 40 | Neutrals: 41 | - set of blocking ivars 42 | 43 | -------------------------------------------------------------------------------- 44 | 45 | eval : (σ : ISub ψ₀ ψ₁)(f : α₀ ⊢ α₁[σ])(δ : Env Γ₀ (Γ₁[σ, f])) 46 | → Tm (ψ₁,α₁,Γ₁) A 47 | → Val (ψ₀,α₀,Γ₀) A[σ,f,δ] 48 | 49 | eval ψ σ α δ = \case 50 | var x -> lookup x δ 51 | app t u -> case eval ψ σ α δ t of 52 | VLam t -> apply t (eval ψ σ α δ u) 53 | ne t is -> ne (app t (eval ψ σ α δ u)) is 54 | lam x t -> VLam (x, ψ, σ, α, δ, t) 55 | plam x t -> VPlam (x, ψ, σ, α, δ, t) 56 | papp t i -> case eval ψ σ α δ t of 57 | VPlam t -> papply t (eval Ψ σ α i) 58 | ne t is -> 59 | 60 | 61 | isub : Val (ψ,α,Γ) → (σ : ISub ψ' ψ) → Val (ψ',α[σ],Γ[σ]) 62 | isub v σ = case v of 63 | ISub δ v -> ISub (δ ∘ σ) v 64 | v -> ISub σ v 65 | 66 | force : ∀ α → Val (ψ,α,Γ) → Val (ψ,α,Γ) 67 | force α v = case v of 68 | ISub σ v -> forceWithSub σ α v 69 | 70 | -- we have to check cofib weakening! 71 | ne v is σ α' 72 | | α == α' -> ne v is σ α' 73 | | α∘α is injective renaming on is -> ne v is σ α 74 | | otherwise -> forceNe σ α v 75 | 76 | v -> v 77 | 78 | forceNe : (σ : ISub ψ' ψ)(α' : Cof ψ') → Ne (Ψ,α,Γ) A → Val (ψ',α',Γ[σ]) A[σ] 79 | forceNe σ α = \case 80 | TODO 81 | 82 | forceWithSub (σ : ISub ψ' ψ) (α' : Cof ψ') → Val (ψ,α,Γ) A → Val (ψ',α',Γ[σ]) A[σ] 83 | forceWithSub σ α = \case 84 | ISub{} -> impossible 85 | Ne v is σ' α' -> let 86 | 87 | 88 | 89 | 90 | freshvar ψ α Γ = ne (var x) [] α (id : ISub ψ ψ) : Ne (ψ,α,Γ,x:A) A 91 | -------------------------------------------------------------------------------- /newnotes/todo.txt: -------------------------------------------------------------------------------- 1 | 2 | - Fix outstanding bug 3 | - Get rid of case body tags in conversion, simply use pointer equality instead 4 | (I see no reason why this wouldn't work) 5 | - Perhaps experiment with pointer equality shortcutting in other closures 6 | -------------------------------------------------------------------------------- /oldnotes/CHMGlueNotes.txt: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/oldnotes/CHMGlueNotes.txt -------------------------------------------------------------------------------- /oldnotes/HITs.txt: -------------------------------------------------------------------------------- 1 | 2 | Parameterized HITs 3 | 4 | Signature: (fibrant-only) parameters, list of constructors 5 | Constructor: telescope with mixed I-s and fibrant types (under params) 6 | + system with component type exactly (Sort params) 7 | HCom: trivial, returns FHCom (which is canonical!) 8 | Coercion: 9 | - A boundary is coherent iff it contains no elimination (app, papp, lapp, proj, case) 10 | - Coherent coercion is the same as indtype coercion 11 | - Incoherent coercion: 12 | - coerce fields as usual, accummulate the line env + isub 13 | - evaluate boundary in the line env. This coeFill-s all 14 | field vars, from "r" to "i". The coe from "i" to "r'" each component 15 | on the outside. 16 | 17 | con fs : A ps [α. t] 18 | 19 | coe r r' (i. A (ps i)) (con fs) = 20 | hcom r r' (A (ps r')) [⟦α⟧ i. coe i r' (j. A (ps j)) ⟦t⟧] (con (coe fs)) 21 | 22 | QUESTION: 23 | - recompute boundary from syntax on every forcing, or store the boundary? 24 | - kinda bad either way 25 | 1. recompute. Then we have to store Tm typarams, which will be normalized 26 | due to bidirectional checking. TODO: allow explicit param notation. 27 | 2. store. Then we have to store the normalized boundary, again due to 28 | bidirectional checking. In coercion we get the typaram filler from the coe type arg, 29 | so we don't have to store typarams in HDCon. 30 | 31 | In any case, coercion *must* recompute the boundary, because we need a 32 | filler for the boundary adjustment. 33 | 34 | WAIT A MINUTE 35 | BOUNDARY FORCING ONLY REQUIRES COMPUTING THE COFS! 36 | THAT MAKES FALSE POSITIVE FORCING HELLA CHEAP 37 | 38 | Let's go with option 1. 39 | We'd like to evaluate boundary components lazily. 40 | We don't need the NeSys part of the VSys result. TODO: specialize evalSys to not return NeSys. 41 | let's call it evalBoundary. 42 | 43 | ------------------------------------------------------------ 44 | 45 | higher inductive S1 := base | loop (i : I)[i=0. base; i=1.base]; 46 | 47 | higher inductive SetTrunc (A : U) := 48 | SEmbed A 49 | | STrunc (x y : SetTrunc A)(p q : x = y) (i j : I)[i=0. p j; i=1. q j]; 50 | 51 | (\i j. STrunc x y p q i j) : p = q 52 | 53 | 54 | Tm 55 | - HTyCon HTyConInfo TyParams 56 | - HDCon HDConInfo HDSpine 57 | 58 | Val 59 | - VHTyCon HTyConInfo Env 60 | - VHDCon HDConInfo VHDSpine IVarSet (similar to GlueTy, neutral but not in Ne! 61 | 62 | - Case and Split are reused. 63 | - FHCom is not distinguished from ordinary Tm/Ne HCom 64 | - We can always just look at the base type 65 | 66 | Don't forget action of coe on fhcom! 67 | 68 | ------------------------------------------------------------ 69 | 70 | Case elaboration: 71 | 72 | case t (x. B) [ 73 | con fs. rhs 74 | ] 75 | 76 | - lookup arg telescope of fs 77 | 78 | - starting from typarams env, push field vars to cxt, create LHS val, check rhs with 79 | B (rhsVal) 80 | 81 | (almost the same as with indtypes up to this point) 82 | 83 | - lookup syntactic con boundary, evaluate it in paramEnv + idSub 84 | during eval, apply (case _ (x. B) prevcases) to each boundary component, where 85 | (prevcases) are the previously checked cases. 86 | 87 | - Check compatibility of rhs with the evaluated boundary system 88 | 89 | Case evaluation: 90 | 91 | - action on constructors exactly the same as in indtypes 92 | 93 | - fhcom: 94 | case _ (x.B x) cases : (x : A ps) → B x 95 | 96 | base : A ps 97 | whole thing : B (fhcom ...) 98 | 99 | case base B cases : B base 100 | α i ⊢ case (t i) B cases : B (t i) 101 | 102 | Case (fhcom r r' (A ps) [α i. t i] base) (x. B x) cases = 103 | 104 | com r r' (i. B (fhcom r i (A ps) [α i. t i] base)) [α i. case (t i) B cases] 105 | (case base B cases) = 106 | 107 | let B* = λ i. B (fhcom r i (A ps) [α j. t j] base); 108 | hcom r r' (B* r') 109 | [α i. coe i r' B* (case (t i) B cases)] 110 | (coe r r' B* (case base B cases)) 111 | 112 | HDCon evaluation: 113 | - eval fields, push field values to paramenv + sub 114 | - eval boundary in paramenv + sub 115 | - total: return val 116 | - neutral: return neutral 117 | -------------------------------------------------------------------------------- /oldnotes/Notes.hs: -------------------------------------------------------------------------------- 1 | 2 | {-# language Strict, LambdaCase, BangPatterns, ViewPatterns #-} 3 | 4 | type Name = String 5 | 6 | data ITm = IVar Int | IZero | IOne deriving Show 7 | 8 | data Cof = COr Cof Cof | CAll Name Cof | CEq ITm ITm deriving Show 9 | 10 | data System = SEmpty | SCons ITm ITm Tm System deriving Show 11 | 12 | type Ty = Tm 13 | 14 | data Tm 15 | = LocalVar Int 16 | | Let Name Ty Tm Tm 17 | 18 | | Pi Name Tm Tm 19 | | App Tm Tm 20 | | Lam Name Tm 21 | 22 | | Univ 23 | 24 | | Path Name Ty Tm Tm 25 | | LamP Name Tm 26 | | AppP Tm Tm Tm ITm 27 | | Coe Name ITm ITm Ty Tm 28 | | HCom Name ITm ITm Ty System Tm 29 | deriving Show 30 | 31 | {- 32 | Plan: 33 | - delayed interval subst, no interval closures, eager subst composition, 34 | no interval vals 35 | - closures & values for ordinary values 36 | - glued eval for top defs, also probably useful for interval subst (because top defs are closed) 37 | - probably we need fresh vars and disjoint subst for intervals 38 | 39 | - we need to restrart reductions of neutrals on interval subst 40 | - how to represent blocked eval? 41 | * problem is that some things can block on multiple things: 42 | - AppP : path & arg 43 | - hcom : src/tgt & system 44 | * only definitely non-Ivar-headed neutrals are stable under I-subst 45 | 46 | * for efficient forcing we definitely need to gather up *all* possible blocking 47 | points in neutrals! 48 | 49 | E.g. take a spine which has a path application *anywhere* inside. Then 50 | if the outermost path application is forced by 0/1 arg, all inner blocking 51 | points are just dropped! 52 | 53 | 54 | -} 55 | 56 | type VTy = Val 57 | 58 | data Head 59 | data Spine 60 | 61 | data Val = 62 | VNe Head Spine 63 | 64 | 65 | 66 | 67 | -- = VVar Int 68 | 69 | -- | VPi Name Val Val 70 | -- | VApp Val Val 71 | -- | VLam Name Val 72 | 73 | -- | VUniv 74 | 75 | -- | VPath Name Ty Val Val 76 | -- | VLamP Name Val 77 | -- | VAppP Val Val Val ITm 78 | 79 | -- | VCofSplit Cof Val Cof Val 80 | -- | VCoe Name ITm ITm Ty Val 81 | -- | VHCom Name ITm ITm Ty Cof Val Val 82 | -- deriving Show 83 | -------------------------------------------------------------------------------- /oldnotes/Notes2.agda: -------------------------------------------------------------------------------- 1 | 2 | 3 | {-# OPTIONS --cubical --show-irrelevant --type-in-type #-} 4 | 5 | module Notes2 where 6 | 7 | open import Cubical.Foundations.Everything 8 | open import Agda.Builtin.Cubical.Path 9 | open import Agda.Builtin.Cubical.Glue 10 | 11 | -- comp from hcomp & transp 12 | comp' : {ℓ : I → Level} (A : ∀ i → Type (ℓ i)) {φ : I} → 13 | (a : ∀ i → Partial φ (A i)) 14 | → Sub (A i0) φ (a i0) 15 | → Sub (A i1) φ (a i1) 16 | comp' {ℓ} A {φ} a a₀ = 17 | inS (hcomp {ℓ i1}{A i1}{φ} 18 | (λ i p → transp (λ j → A (i ∨ j)) i (a i p)) 19 | (transp A i0 (outS a₀))) 20 | 21 | 22 | -- transp from comp 23 | transp' : 24 | ∀ {ℓ : I → Level} 25 | (A₀ : Type (ℓ i0)) 26 | (φ : I) 27 | (A : ∀ i → Sub (Type (ℓ i)) φ λ _ → A₀) 28 | → (a₀ : outS (A i0)) 29 | → Sub (outS (A i1)) φ λ {(φ = i1) → a₀} 30 | transp' {ℓ} A₀ φ A a₀ = inS (comp (λ i → outS (A i)) (λ {i (φ = i1) → a₀}) a₀) 31 | 32 | 33 | -- comp from hcomp & transport? I can't see how 34 | comp'' : {ℓ : I → Level} (A : ∀ i → Type (ℓ i)) {φ : I} → 35 | (a : ∀ i → Partial φ (A i)) 36 | → Sub (A i0) φ (a i0) 37 | → Sub (A i1) φ (a i1) 38 | comp'' {ℓ} A {φ} a a₀ = 39 | inS (hcomp {ℓ i1}{A i1}{φ} 40 | (λ i p → transp (λ j → A (i ∨ j)) i (a i p)) 41 | (transport (λ i → A i) (outS a₀))) 42 | 43 | -- The point of generalized transp is to allow splitting comp to transp and hcomp 44 | -- which makes it possible to have HIT-s? Because they need formal hcomp, and possibly 45 | -- formal transp? 46 | 47 | -- transpⁱ (Glue [α ↦ (T, e)] A) ψ 48 | 49 | -- only reduce if `α`, `T` and `e` are all constant in `i` under `ψ`. 50 | -------------------------------------------------------------------------------- /oldnotes/TranspGlue.agda: -------------------------------------------------------------------------------- 1 | 2 | {-# OPTIONS --cubical --show-irrelevant --type-in-type #-} 3 | 4 | module TranspGlue where 5 | 6 | open import Cubical.Foundations.Everything 7 | open import Agda.Builtin.Cubical.Path 8 | open import Agda.Builtin.Cubical.Glue 9 | 10 | module Mod (A₀ A₁ : Type)(A≡ : A₀ ≡ A₁) (α : I) (e : ∀ i → A₀ ≃ A≡ i) where 11 | 12 | G : I → Type 13 | G i = Glue (A≡ i) {α} (λ _ → A₀ , e i) 14 | 15 | -- G is constantly A₀ under α 16 | constantG : ∀ i → Sub Type α λ _ → A₀ 17 | constantG i = inS (G i) 18 | 19 | -- so transport works 20 | transpG : G i0 → G i1 21 | transpG = transp G α 22 | 23 | -- But the definition below computes to an ill-typed term. 24 | -- The computation rule assumes that A≡ i1 is definitionally equal to A≡ i0. 25 | f : G i0 → A₁ 26 | f g₀ = unglue α (transpG g₀) 27 | 28 | unglue g₀ : A₀ 29 | 30 | `a₀` is `unglue g₀ : A₀` 31 | 32 | A≡ i 33 | 34 | `α` (`psi` in the source) 35 | -------------------------------------------------------------------------------- /oldnotes/TranspGlueIssue.agda: -------------------------------------------------------------------------------- 1 | 2 | {-# OPTIONS --cubical --show-irrelevant #-} 3 | 4 | module TranspGlueIssue where 5 | 6 | open import Cubical.Foundations.Everything 7 | open import Cubical.Core.Everything 8 | 9 | open import Cubical.Relation.Nullary.DecidableEq 10 | open import Cubical.Relation.Nullary.Base 11 | open import Cubical.Data.Empty 12 | open import Cubical.Data.Unit 13 | 14 | the : ∀{l}(A : Type l) → A → A 15 | the A x = x 16 | 17 | data Bool₀ : Set where true false : Bool₀ 18 | data Bool₁ : Set where true false : Bool₁ 19 | 20 | true≢false : the Bool₁ true ≡ false → ⊥ 21 | true≢false p = subst (λ {true → Unit; false → ⊥}) p tt 22 | 23 | Bool₁-discrete : Discrete Bool₁ 24 | Bool₁-discrete true true = yes refl 25 | Bool₁-discrete true false = no true≢false 26 | Bool₁-discrete false true = no (true≢false ∘ _⁻¹) 27 | Bool₁-discrete false false = yes refl 28 | 29 | Bool₁-set : isSet Bool₁ 30 | Bool₁-set = Discrete→isSet Bool₁-discrete 31 | 32 | f : Bool₀ → Bool₁ 33 | f true = true 34 | f false = false 35 | 36 | f-equiv : isEquiv f 37 | fst (fst (equiv-proof f-equiv true)) = true 38 | snd (fst (equiv-proof f-equiv true)) = refl 39 | snd (equiv-proof f-equiv true) (true , p) i = (true , Bool₁-set _ _ refl p i) 40 | snd (equiv-proof f-equiv true) (false , p) = rec (true≢false (p ⁻¹)) 41 | fst (fst (equiv-proof f-equiv false)) = false 42 | snd (fst (equiv-proof f-equiv false)) = refl 43 | snd (equiv-proof f-equiv false) (true , p) = rec (true≢false p) 44 | snd (equiv-proof f-equiv false) (false , p) i = false , Bool₁-set _ _ refl p i 45 | 46 | module _ (α : I) where 47 | 48 | lem : ∀ i → Bool₀ ≃ ua (f , f-equiv) i 49 | lem i = {!!} 50 | 51 | G : I → Type 52 | G i = Glue (ua (f , f-equiv) i) {α} (λ _ → Bool₀ , lem i) 53 | 54 | transpG : G i0 → G i1 55 | transpG g₀ = transp G α g₀ 56 | 57 | foo = {!transpG!} 58 | 59 | 60 | bluh = transport (λ i → (ua (f , f-equiv)⁻¹) i) 61 | 62 | x = {!ua (f , f-equiv)!} 63 | 64 | -- x = {!bluh!} 65 | 66 | -- -- module _ (α : I) where 67 | 68 | -- -- A : Bool₀ ≡ Bool₁ 69 | -- -- A = ua (f , f-equiv) 70 | 71 | -- -- T : Type 72 | -- -- T = Bool₀ 73 | 74 | -- -- e : PathP (λ i → T ≃ A i) (idEquiv Bool₀) (f , f-equiv) 75 | -- -- e = {!!} 76 | 77 | -- -- Te : ∀ i → Σ Type λ T → T ≃ A i 78 | -- -- Te i = Bool₀ , e i 79 | 80 | -- -- G : I → Type 81 | -- -- G i = Glue (A i) {α} (λ _ → Te i) 82 | 83 | -- -- constantG : PartialP {ℓ-suc ℓ-zero} α λ {(α = i1) → G i0 ≡ G i1} 84 | -- -- constantG (α = i1) = refl 85 | 86 | -- -- foo = {!!} 87 | 88 | -- module Mod (A₀ A₁ : Type)(A≃ : A₀ ≃ A₁) (α : I) (e : ∀ i → A₀ ≃ ua A≃ i) where 89 | 90 | -- G : I → Type 91 | -- G i = Glue (ua A≃ i) {α} (λ _ → A₀ , e i) 92 | 93 | -- -- G is constantly A₀ under α 94 | -- constantG : ∀ i → Sub Type α λ {(α = i1) → A₀} 95 | -- constantG i = inS (G i) 96 | 97 | -- -- so transport works 98 | -- transpG : G i0 → G i1 99 | -- transpG = transp G α 100 | 101 | -- -- A₀ and A₁ are definitionally distinct. However, the computation rule for 102 | -- -- transpG assumes that they are the same! 103 | 104 | -- module _ (t : PartialP α λ _ → A₀)(a₀ : Sub A₀ α λ p → e i0 .fst (t p)) where 105 | 106 | -- g₀ : G i0 107 | -- g₀ = glue t (outS a₀) 108 | 109 | -- g₁ : G i1 110 | -- g₁ = {!transpG g₀!} 111 | 112 | 113 | 114 | -- -- glue : ∀ {A : Type} {φ : I} 115 | -- -- → {T : Partial φ Type} 116 | -- -- → {e : PartialP φ (λ p → T p ≃ A)} 117 | -- -- → (t : PartialP φ T) 118 | -- -- → (a : Sub A φ (λ p → e p .fst (t p))) 119 | -- -- → Glue A (λ p → T p , e p) 120 | 121 | record Prod (A B : Type) : Type where 122 | constructor prod 123 | field 124 | fst : A 125 | snd : B 126 | open Prod 127 | 128 | transpProd : 129 | ∀ α (A₀ B₀ : Type) 130 | (A : I → Sub Type α λ _ → A₀) 131 | (B : I → Sub Type α λ _ → B₀) 132 | (ab₀ : Prod (outS (A i0)) (outS (B i0))) 133 | → transp (λ i → Prod (outS (A i)) (outS (B i))) α ab₀ ≡ 134 | prod (transp (λ i → outS (A i)) α (fst ab₀)) 135 | (transp (λ i → outS (B i)) α (snd ab₀)) 136 | transpProd α A₀ B₀ A B ab₀ = refl 137 | -------------------------------------------------------------------------------- /oldnotes/newnotes.txt: -------------------------------------------------------------------------------- 1 | 2 | - Anders note (coe Glue) 3 | - ghcomp in CCTT 4 | - add negation to cofs 5 | - alternative version by Evan? 6 | 7 | - ghcomp i=1 8 | - transp 9 | - hcom 10 | 11 | transpⁱ A ϕ t 12 | ϕ,i ⊢ A ≡ A[i/0] 13 | 14 | transpⁱ (Glue [α ↦ (T, f)] B) ϕ t 15 | 16 | transpⁱ ((x : A) → B) ϕ t 17 | = ... 18 | 19 | compⁱ 20 | -------------------------------------------------------------------------------- /oldnotes/prez.txt: -------------------------------------------------------------------------------- 1 | 2 | -------------------------------------------------------------------------------- /package.yaml: -------------------------------------------------------------------------------- 1 | 2 | name: cctt 3 | version: 0.1.0.0 4 | github: "AndrasKovacs/cctt" 5 | license: MIT 6 | author: "András Kovács" 7 | maintainer: "puttamalac@gmail.com" 8 | copyright: "2022 András Kovács" 9 | category: Language 10 | description: "Experiments on high-performance evaluation for cubical type theories" 11 | 12 | flags: 13 | dump: 14 | description: dump core, stg and cmm to files 15 | default: false 16 | manual: true 17 | llvm: 18 | description: use llvm for codegen 19 | default: false 20 | manual: true 21 | debug: 22 | description: turn on debug info printing 23 | default: false 24 | manual: true 25 | emptyhcomstats: 26 | description: turn on empty hcom counting 27 | default: false 28 | manual: true 29 | 30 | when: 31 | - condition: flag(dump) 32 | ghc-options: 33 | - -ddump-simpl 34 | - -ddump-stg-final 35 | - -ddump-cmm 36 | - -dsuppress-all 37 | - -dno-suppress-type-signatures 38 | - -ddump-to-file 39 | - condition: flag(llvm) 40 | ghc-options: 41 | - -fllvm 42 | - condition: flag(debug) 43 | cpp-options: 44 | - -DDEBUG 45 | - condition: flag(emptyhcomstats) 46 | cpp-options: 47 | - -DEMPTYHCOMSTATS 48 | 49 | default-extensions: 50 | - AllowAmbiguousTypes 51 | - BangPatterns 52 | - BlockArguments 53 | - CPP 54 | - ConstraintKinds 55 | - DataKinds 56 | - DerivingVia 57 | - EmptyDataDeriving 58 | - ExplicitNamespaces 59 | - FlexibleContexts 60 | - FlexibleInstances 61 | - FunctionalDependencies 62 | - GADTs 63 | - ImplicitParams 64 | - InstanceSigs 65 | - KindSignatures 66 | - LambdaCase 67 | - MagicHash 68 | - MultiParamTypeClasses 69 | - OverloadedStrings 70 | - PackageImports 71 | - PatternSynonyms 72 | - PolyKinds 73 | - QualifiedDo 74 | - RankNTypes 75 | - ScopedTypeVariables 76 | - StandaloneDeriving 77 | - StandaloneKindSignatures 78 | - Strict 79 | - TemplateHaskell 80 | - TypeApplications 81 | - TypeFamilies 82 | - TypeFamilyDependencies 83 | - UnboxedSums 84 | - UnboxedTuples 85 | - UndecidableInstances 86 | - UnicodeSyntax 87 | - ViewPatterns 88 | 89 | ghc-options: 90 | - -O2 91 | - -fdicts-strict 92 | - -fmax-worker-args=18 93 | - -fspecialise-aggressively 94 | - -fexpose-all-unfoldings 95 | - -fplugin StrictImplParams 96 | - -fworker-wrapper-cbv 97 | - -rtsopts 98 | - -threaded 99 | - -with-rtsopts= "-N4 -A8M -qb0" 100 | - -Wall 101 | - -Wno-name-shadowing 102 | - -Wno-missing-signatures 103 | - -Wno-unused-do-bind 104 | - -Wno-unused-matches 105 | - -Wno-missing-pattern-synonym-signatures 106 | - -fmax-relevant-binds=3 107 | - -fmax-valid-hole-fits=0 108 | 109 | dependencies: 110 | - base >= 4.7 && < 5 111 | - bytestring 112 | - containers 113 | - directory 114 | - filepath 115 | - flatparse 116 | - ghc-prim 117 | - microlens 118 | - microlens-th 119 | - primdata 120 | - split 121 | - strict-impl-params 122 | - template-haskell 123 | - time 124 | 125 | source-dirs: src 126 | 127 | executables: 128 | cctt: 129 | main: Main.hs 130 | when: 131 | - condition: false 132 | other-modules: Paths_cubeval 133 | -------------------------------------------------------------------------------- /references/brunerie_nums.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/references/brunerie_nums.pdf -------------------------------------------------------------------------------- /references/cctt_investigations.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/references/cctt_investigations.pdf -------------------------------------------------------------------------------- /references/coefcom.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/references/coefcom.pdf -------------------------------------------------------------------------------- /references/extension.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/references/extension.pdf -------------------------------------------------------------------------------- /references/lattice_impl.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/references/lattice_impl.pdf -------------------------------------------------------------------------------- /rules.txt: -------------------------------------------------------------------------------- 1 | THIS IS ALSO PRETTY OLD NOW 2 | 3 | 4 | TODO: 5 | - harmonize notation with implementation everywhere 6 | - add newest coe Glue (old/notes), look at hcom fib unfolding optimization 7 | - write hcom Glue 8 | 9 | -------------------------------------------------------------------------------- 10 | 11 | Interval vars: i,j,k 12 | 13 | Cofibrations: 14 | - A cof is a list of atomic cofs, viewed as a *conjunction*. 15 | - Atomic cofs: i=0 | i=1 | i=j 16 | - Cof vars: α, β, γ 17 | 18 | Context: 19 | - Γ + interval cxt + one cofibration 20 | - In notation, we can just extend Γ with either i:I or α:Cof, 21 | but formally we mean extending the appropriate part of the cxt. 22 | Extending the cof is just conjunction. 23 | 24 | Systems: 25 | - list of (cof, term) pairs, can be empty 26 | - notation: [α₀ ↦ t₀, α₁ ↦ t₁, ... αᵢ ↦ tᵢ] 27 | - typing: 28 | 29 | Γ ⊢ A : U Γ, αᵢ ⊢ tᵢ : A Γ, αᵢ ∧ αⱼ ⊢ tᵢ ≡ tⱼ 30 | ─────────────────────────────────────────────────── 31 | Γ ⊢ [αᵢ ↦ tᵢ] is a system 32 | 33 | Cubical extension judgement: 34 | 35 | Γ ⊢ t : A[α ↦ u] means Γ ⊢ t : A and Γ,α ⊢ t ≡ u 36 | 37 | Coercion 38 | 39 | Γ,i ⊢ A : U Γ ⊢ t : A r 40 | ──────────────────────────────────── 41 | Γ ⊢ coeⁱ r r' A t : (A r') [r=r' ↦ t] 42 | 43 | Homogeneous composition 44 | 45 | 46 | Γ ⊢ α cof Γ ⊢ A : U Γ, α, i ⊢ t : A Γ ⊢ b : A Γ, α ⊢ t r ≡ b 47 | ────────────────────────────────────────────────────────────────────── 48 | Γ ⊢ hcomⁱ r r' A [α ↦ t] b : A [r=r' ↦ b, α ↦ t r') 49 | 50 | Composition (derived) 51 | 52 | Γ ⊢ α cof Γ, i ⊢ A : U Γ, α, i ⊢ t : A Γ ⊢ b : A r Γ, α ⊢ t r ≡ b 53 | ──────────────────────────────────────────────────────────────────────── 54 | Γ ⊢ comⁱ r r' A [α ↦ t] b : (A r') [r=r' ↦ b, α ↦ t r'] 55 | com r r' (i. A i) [α i. t] b := 56 | hcom r r' (A r') [α i. coe i r' (j. A j) t] (coe r r' (i. A i) b) 57 | 58 | 59 | -- filling 60 | -------------------------------------------------------------------------------- 61 | 62 | Γ, i ⊢ coeFillⁱ r A t : A [i=r ↦ t, i=r' ↦ coeⁱ r r' A t ] 63 | coeFillⁱ r A t := coe r i A t 64 | 65 | Γ, i ⊢ coeFill⁻¹ⁱ r A t : A [i=r ↦ coe r r' A t, i=r' ↦ t] 66 | coeFill⁻¹ⁱ r A t := coe i r' A t 67 | 68 | 69 | -- Equivalences 70 | -------------------------------------------------------------------------------- 71 | 72 | isEquiv : (A → B) → U 73 | isEquiv f := 74 | (f⁻¹ : B → A) 75 | × (linv : ∀ a → f⁻¹ (f a) = a) 76 | × (rinv : ∀ b → f (f⁻¹ b) = b) 77 | × (coh : ∀ a → 78 | Pathⁱ (f (linv a i) = f a) (rinv (f a)) (refl (f a))) 79 | 80 | idIsEquiv : (A : U) → isEquiv (λ (a:A). a) 81 | _⁻¹ = λ a. a 82 | linv = λ a i. a 83 | rinv = λ b i. b 84 | coh = λ a i j. a 85 | 86 | isEquivCoe : (Γ, i ⊢ A : U) (r r' : I) → Γ ⊢ isEquiv (coeⁱ r r' A : Ar → Ar') 87 | isEquivCoe A r r' = 88 | 89 | _⁻¹ := coeⁱ r' r A 90 | 91 | linvFill : ∀ s a → a = coeⁱ s r A (coeⁱ r s A a) 92 | linvFill s a = λ j. hcomᵏ r s (A r) [j=0 ↦ coeⁱ k r A (coeⁱ r k A a), j=1 ↦ a] a 93 | 94 | linv := linvFill r' 95 | 96 | rinvFill : ∀ s b → coeⁱ s r' A (coeⁱ r' s A b) = b 97 | rinvFill s b = λ j. hcomᵏ r' s (A r') [j=0 ↦ coeⁱ k r' A (coeⁱ r' k A b), j=1 ↦ b] b 98 | 99 | rinv := rinvFill r 100 | 101 | coh : ∀ a → PathPⁱ (f (linv a i) = f a) 102 | (refl (f a)) (rinv (f a))) 103 | 104 | coh = TODO 105 | 106 | -------------------------------------------------------------------------------- 107 | 108 | Glue 109 | 110 | Γ ⊢ B : U Γ, α ⊢ eqv : (A : U) × (f : A → B) × isEquiv f 111 | ────────────────────────────────────────────────────────── 112 | Γ ⊢ Glue [α ↦ eqv] B : U 113 | Γ, α ⊢ Glue [α ↦ eqv] B ≡ A 114 | 115 | 116 | -------------------------------------------------------------------------------- 117 | 118 | isEquiv : (A → B) → U 119 | isEquiv A B f := 120 | (g : B → A) 121 | × (linv : (x : A) → Path A (g (f x)) x) 122 | × (rinv : (x : B) → Path B (f (g x)) x) 123 | × (coh : (x : A) → 124 | PathP i (Path B (f (linv x {x}{g (f x)} i)) (f x)) 125 | (rinv (f x)) 126 | (refl B (f x))) 127 | 128 | coeIsEquiv : (A : I → U) → (r r' : I) → isEquiv (coeⁱ r r' A : A r → A r') 129 | coeIsEquiv A r r' = 130 | _⁻¹ := λ x. coe r' r A x 131 | linv := λ x. λ j. hcom r r' (A r) [j=0 k. coe k r A (coe r k A x); j=1 k. x] x 132 | rinv := λ x. λ j. hcom r' r (A r') [j=0 k. coe k r' A (coe r' k A x); j=1 k. x] x 133 | coh := TODO 134 | 135 | -- coh : 136 | 137 | -- PathP i (Path B (f (linv x {x}{g (f x)} i)) (f x)) 138 | -- (rinv (f x))) 139 | -- (λ _. coe r r' A 140 | 141 | 142 | 143 | -------------------------------------------------------------------------------- 144 | 145 | 146 | coeⁱ r r' ((a : A) × B a) t = 147 | (coeⁱ r r' A t.1, coeⁱ r r' (B (coeFillⁱ r r' A t.1)) t.2) 148 | 149 | coeⁱ r r' ((a : A) → B a) t = 150 | (λ (a' : A r'). coeⁱ r r' (B (coeFill⁻¹ⁱ r r' A a')) (t (coeⁱ r' r A a'))) 151 | 152 | coeⁱ r r' (Pathʲ A t u) p = 153 | (λ j. comⁱ r r' (A i j) [j=0 ↦ t i, j=1 ↦ u i] (p @ j)) 154 | : Pathʲ (A[i↦r']) (t[i↦r']) (u[i↦r']) 155 | 156 | coe r r' (i. t i ={j. A i j} u i) p = 157 | λ {t r'}{u r'} j. com r r' (i. A i j) [j=0 i. t i; j=1 i. u i] (p j) 158 | 159 | coeⁱ r r' ((j : I) → A j) p = 160 | (λ j. coeⁱ r r' (i. A i j) (p @ j)) 161 | 162 | coeⁱ r r' ℕ t = t 163 | coeⁱ r r' U t = t 164 | coeⁱ r r' (Glue [α ↦ (T, f)] A) gr = TODO 165 | 166 | hcom r r' ((a : A) × B a) [α i. t i] b = 167 | ( hcom r r' A [α i. (t i).1] b.1 168 | , com r r' (i. B (hcom r i A [α j. (t j).1] b.1)) [α i. (t i).2] b.2) 169 | 170 | hcomⁱ r r' ((a : A) → B a) [α ↦ t] b = 171 | λ a. hcomⁱ r r' (B a) [α ↦ t i a] (b a) 172 | 173 | hcom r r' (lhs ={j.A j} rhs) [α i. t i] base = 174 | λ j. hcom r r' (A j) [j=0 i. lhs i; j=1 i. rhs i; α i. t i j] (base j) 175 | 176 | 177 | hcom r r' ((i : I) → B i) [α j. t] b = 178 | (λ arg. hcom r r' (A arg) [α i. ↦ t arg] (base arg)) 179 | 180 | hcomⁱ r r' ℕ [α ↦ zero] zero = zero 181 | hcomⁱ r r' ℕ [α ↦ suc t] (suc b) = suc (hcomⁱ r r' ℕ [α ↦ t] b) 182 | hcomⁱ r r' U [α ↦ t] b = Glue [α ↦ (t r', (coeⁱ r' r (t i), _)), r=r' ↦ (b, idEqv)] b 183 | 184 | hcomⁱ r r' (Glue [α ↦ (T, f)] A) [β ↦ t] gr = 185 | glue [α ↦ hcomⁱ r r' T [β ↦ t] gr] 186 | (hcomⁱ r r' A [β ↦ unglue t, α ↦ f (hfillⁱ r r' T [β ↦ t] gr)] (unglue gr)) 187 | 188 | 189 | -- System 190 | -------------------------------------------------------------------------------- 191 | - CCTT, coe, hcom, no cof disjunction 192 | - no (∀i.α), in coeGlue we compute ∀i.α extended systems on the stop 193 | - parameterized (strict) inductives, HITs, all single-sorted 194 | - no indexed inductives 195 | 196 | -- Eval 197 | -------------------------------------------------------------------------------- 198 | 199 | - Defunctionalized closures 200 | - Lambdas, path lambdas are closures 201 | - CBV except for system components which are lazy 202 | - binders that we have to match on are not closures 203 | - hcom and coe types are not closures 204 | - system components are not closures 205 | - delayed isubst 206 | - no sharing of isubst forcing computation 207 | - eager isubst composition: isubst of isubst collapses 208 | - in the closed eval case, hcom can lazily peel off strict inductive constructors 209 | (because of canonicity!!!) 210 | 211 | -- Neutrals 212 | -------------------------------------------------------------------------------- 213 | 214 | - Neutrals store a bitmask of blocking ivars + delayed isubst 215 | - we collect blocking ivars during eval/forcing 216 | - forcing a neutral: 217 | - if forcing sub is an injective renaming on blocking vars, then remain blocked 218 | - otherwise force the whole thing 219 | 220 | NOTE: 221 | - Right now I don't want to collect more precise blocking info, e.g. blocking cofs. 222 | (potential TODO) 223 | - In open eval, if I get to actually force the whole hcom system because of 224 | a strict inductive hcom rule, I gather all base+tube blocking ivars 225 | 226 | 227 | -- Delayed unfoldings 228 | -------------------------------------------------------------------------------- 229 | -------------------------------------------------------------------------------- /src/Common.hs: -------------------------------------------------------------------------------- 1 | module Common ( 2 | module Common 3 | , module Control.Monad 4 | , module Data.Bits 5 | , module Lens.Micro 6 | , module Lens.Micro.TH 7 | , module Data.IORef 8 | , module GHC.Exts 9 | , module GHC.Word 10 | , module Text.Show 11 | , module Data.Foldable 12 | , FP.utf8ToStr 13 | , FP.strToUtf8 14 | , runIO 15 | , trace 16 | , traceM 17 | , traceShow 18 | , traceShowM) where 19 | 20 | import Control.Monad 21 | import Data.Bits 22 | import Data.List 23 | import Data.Time.Clock 24 | import GHC.Exts hiding (lazy, fromList, toList) 25 | import Lens.Micro 26 | import Lens.Micro.TH 27 | import Data.IORef 28 | import Data.Foldable 29 | import Debug.Trace (trace, traceM, traceShow, traceShowM) 30 | import Data.Flat 31 | import IO (runIO) 32 | import GHC.IO 33 | import GHC.Word 34 | import Text.Show 35 | 36 | import qualified Data.ByteString.Char8 as B 37 | import qualified Data.ByteString.Internal as B 38 | import qualified FlatParse.Stateful as FP 39 | 40 | -- Debug printing, toggled by "debug" cabal flag 41 | -------------------------------------------------------------------------------- 42 | 43 | #define DEBUG 44 | 45 | #ifdef DEBUG 46 | import GHC.Stack 47 | #endif 48 | 49 | #ifdef DEBUG 50 | type Dbg = HasCallStack 51 | 52 | debug :: [String] -> IO () 53 | debug strs = putStrLn (intercalate " | " strs ++ " END") 54 | 55 | debugging :: IO () -> IO () 56 | debugging act = act 57 | {-# inline debugging #-} 58 | #else 59 | type Dbg = () :: Constraint 60 | 61 | debug :: [String] -> IO () 62 | debug strs = pure () 63 | {-# inline debug #-} 64 | 65 | debugging :: IO () -> IO () 66 | debugging _ = pure () 67 | {-# inline debugging #-} 68 | #endif 69 | 70 | debug' :: [String] -> IO () 71 | debug' strs = putStrLn (intercalate " | " strs ++ " END") 72 | 73 | debugging' :: IO () -> IO () 74 | debugging' act = act 75 | {-# inline debugging' #-} 76 | 77 | -------------------------------------------------------------------------------- 78 | 79 | noinlineRunIO :: IO a -> a 80 | noinlineRunIO (IO f) = runRW# (\s -> case f s of (# _, a #) -> a) 81 | {-# noinline noinlineRunIO #-} 82 | 83 | data Name 84 | = NSpan {-# unpack #-} Span 85 | | NGeneric B.ByteString 86 | | N_ 87 | deriving (Eq, Ord) 88 | 89 | instance Show Name where 90 | show (NSpan x) = show x 91 | show (NGeneric x) = B.unpack x 92 | show N_ = "_" 93 | 94 | nameToBs :: Name -> B.ByteString 95 | nameToBs = \case 96 | NSpan x -> spanToBs x 97 | NGeneric x -> x 98 | N_ -> "_" 99 | 100 | a_ = NGeneric "a" 101 | b_ = NGeneric "b" 102 | c_ = NGeneric "c" 103 | d_ = NGeneric "d" 104 | e_ = NGeneric "e" 105 | f_ = NGeneric "f" 106 | g_ = NGeneric "g" 107 | h_ = NGeneric "h" 108 | i_ = NGeneric "i" 109 | j_ = NGeneric "j" 110 | k_ = NGeneric "k" 111 | l_ = NGeneric "l" 112 | m_ = NGeneric "m" 113 | n_ = NGeneric "n" 114 | o_ = NGeneric "o" 115 | p_ = NGeneric "p" 116 | q_ = NGeneric "q" 117 | r_ = NGeneric "r" 118 | s_ = NGeneric "s" 119 | t_ = NGeneric "t" 120 | u_ = NGeneric "u" 121 | x_ = NGeneric "x" 122 | y_ = NGeneric "y" 123 | z_ = NGeneric "z" 124 | linv_ = NGeneric "linv" 125 | rinv_ = NGeneric "rinv" 126 | coh_ = NGeneric "coh" 127 | ty_ = NGeneric "Ty_" 128 | 129 | ---------------------------------------------------------------------------------------------------- 130 | 131 | uf :: Dbg => a 132 | uf = undefined 133 | {-# noinline uf #-} 134 | 135 | impossible :: Dbg => a 136 | impossible = error "impossible" 137 | {-# noinline impossible #-} 138 | 139 | ctz :: Word -> Word 140 | ctz (W# n) = W# (ctz# n) 141 | {-# inline ctz #-} 142 | 143 | clz :: Word -> Word 144 | clz (W# n) = W# (clz# n) 145 | {-# inline clz #-} 146 | 147 | i2w :: Int -> Word 148 | i2w (I# n) = W# (int2Word# n) 149 | {-# inline i2w #-} 150 | 151 | w2i :: Word -> Int 152 | w2i (W# n) = I# (word2Int# n) 153 | {-# inline w2i #-} 154 | 155 | ($$!) :: (a -> b) -> a -> b 156 | ($$!) f a = f $! a 157 | {-# inline ($$!) #-} 158 | infixl 9 $$! 159 | 160 | infixl 4 <*!> 161 | (<*!>) :: Monad m => m (a -> b) -> m a -> m b 162 | (<*!>) mf ma = do 163 | f <- mf 164 | a <- ma 165 | pure $! f a 166 | {-# inline (<*!>) #-} 167 | 168 | infixr 4 // 169 | (//) :: a -> b -> (a, b) 170 | a // b = (a, b) 171 | {-# inline (//) #-} 172 | 173 | ptrEq :: a -> a -> Bool 174 | ptrEq x y = isTrue# (reallyUnsafePtrEquality# x y) 175 | {-# inline ptrEq #-} 176 | 177 | -- De Bruijn indices and levels 178 | -------------------------------------------------------------------------------- 179 | 180 | newtype Ix = Ix {unIx :: Word} 181 | deriving (Eq, Ord, Show, Num, Enum, Bits, Integral, Real) via Word 182 | 183 | newtype Lvl = Lvl {unLvl :: Word} 184 | deriving (Eq, Ord, Show, Num, Enum, Bits, Integral, Real) via Word 185 | 186 | newtype IVar = IVar# {unIVar :: Word} 187 | deriving (Eq, Ord, Show, Num, Enum, Bits, Integral, Real, Flat) via Word 188 | 189 | class HasDom s where 190 | dom :: s -> IVar 191 | setDom :: IVar -> s -> s 192 | 193 | class HasCod s where 194 | cod :: s -> IVar 195 | setCod :: IVar -> s -> s 196 | 197 | class Lift a where 198 | lift :: a -> a 199 | 200 | lvlToIx :: Lvl -> Lvl -> Ix 201 | lvlToIx (Lvl envl) (Lvl x) = Ix (envl - x - 1) 202 | {-# inline lvlToIx #-} 203 | 204 | ixToLvl :: Lvl -> Ix -> Lvl 205 | ixToLvl (Lvl envl) (Ix x) = Lvl (envl - x - 1) 206 | {-# inline ixToLvl #-} 207 | 208 | type LvlArg = (?lvl :: Lvl) 209 | 210 | 211 | -- Not printing stuff 212 | -------------------------------------------------------------------------------- 213 | 214 | newtype DontShow a = DontShow {unDontShow :: a} deriving Eq 215 | 216 | instance Show (DontShow a) where 217 | showsPrec _ _ x = x 218 | 219 | -------------------------------------------------------------------------------- 220 | 221 | data Box a = Box ~a deriving Show 222 | 223 | -- Time measurement 224 | -------------------------------------------------------------------------------- 225 | 226 | -- | Time an IO computation. Result is forced to whnf. 227 | timed :: IO a -> IO (a, NominalDiffTime) 228 | timed a = do 229 | t1 <- getCurrentTime 230 | res <- a 231 | t2 <- getCurrentTime 232 | let diff = diffUTCTime t2 t1 233 | pure (res, diff) 234 | {-# inline timed #-} 235 | 236 | -- | Time a lazy pure value. Result is forced to whnf. 237 | timedPure :: a -> IO (a, NominalDiffTime) 238 | timedPure ~a = do 239 | t1 <- getCurrentTime 240 | let res = a 241 | t2 <- getCurrentTime 242 | let diff = diffUTCTime t2 t1 243 | pure (res, diff) 244 | {-# noinline timedPure #-} 245 | 246 | -- | Time a lazy pure value. Result is forced to whnf. 247 | timedPure_ :: a -> IO NominalDiffTime 248 | timedPure_ ~a = do 249 | t1 <- getCurrentTime 250 | seq a $ do 251 | t2 <- getCurrentTime 252 | let diff = diffUTCTime t2 t1 253 | pure diff 254 | {-# noinline timedPure_ #-} 255 | 256 | 257 | -- Source descriptors, positions, spans 258 | -------------------------------------------------------------------------------- 259 | 260 | -- | Describes a source bytestring such that a position can point into it. 261 | data Src 262 | = SrcFile FilePath B.ByteString -- ^ Concrete source file. 263 | | SrcImpossible -- ^ Impossible case just for killing unboxing. 264 | 265 | instance Show Src where 266 | show (SrcFile fp _) = "File " ++ fp 267 | show SrcImpossible = impossible 268 | 269 | srcToBs :: Src -> B.ByteString 270 | srcToBs (SrcFile _ bs) = bs 271 | srcToBs SrcImpossible = impossible 272 | 273 | -- | Equality of bytestrings by reference, used for sanity checks. 274 | samePtr :: B.ByteString -> B.ByteString -> Bool 275 | samePtr x y = case B.toForeignPtr x of 276 | (x, _, _) -> case B.toForeignPtr y of 277 | (y, _, _) -> x == y 278 | {-# inline samePtr #-} 279 | 280 | -- | Equality of sources only checks that underlying bytestrings have the same 281 | -- underlying data. 282 | instance Eq Src where 283 | SrcFile _ s == SrcFile _ s' = samePtr s s' 284 | _ == _ = impossible 285 | 286 | -- | Source position. We decorate raw terms with this. 287 | data Pos = Pos Src FP.Pos 288 | deriving Show via DontShow Pos 289 | deriving Eq 290 | 291 | data ParsedPos = ParsedPos FilePath String Int Int 292 | 293 | parsePos :: Pos -> ParsedPos 294 | parsePos (Pos src p) = case src of 295 | SrcImpossible -> impossible 296 | SrcFile path src -> case FP.posLineCols src [p] of 297 | [(l, c)] -> let f = FP.utf8ToStr src in ParsedPos path f l c 298 | _ -> impossible 299 | 300 | instance Show ParsedPos where 301 | show (ParsedPos path _ linum colnum) = 302 | path ++ ":" ++ show (linum + 1) ++ ":" ++ show colnum 303 | 304 | rawPos :: Pos -> FP.Pos 305 | rawPos (Pos _ p) = p 306 | 307 | initialPos :: Src -> Pos 308 | initialPos src = Pos src (FP.Pos 0) 309 | 310 | instance Ord Pos where 311 | compare (Pos src i) (Pos src' i') 312 | | src == src' = compare i i' 313 | | otherwise = impossible 314 | 315 | (<) (Pos src i) (Pos src' i') 316 | | src == src' = i < i' 317 | | otherwise = impossible 318 | 319 | (<=) (Pos src i) (Pos src' i') 320 | | src == src' = i <= i' 321 | | otherwise = impossible 322 | 323 | -- | Source span. The left position must not be larger than the right one. 324 | data Span = Span# Src FP.Pos FP.Pos 325 | -- deriving Show via DontShow Span 326 | 327 | instance Show Span where 328 | show = spanToString 329 | 330 | pattern Span :: Pos -> Pos -> Span 331 | pattern Span x y <- ((\(Span# src x y) -> (Pos src x, Pos src y)) -> (x, y)) where 332 | Span (Pos src x) (Pos src' y) 333 | | src == src' && x <= y = Span# src x y 334 | | otherwise = impossible 335 | {-# complete Span #-} 336 | 337 | spanToBs :: Span -> B.ByteString 338 | spanToBs (Span (Pos src i) (Pos _ j)) = 339 | let bstr = srcToBs src 340 | i' = B.length bstr - coerce i -- Pos counts backwards from the end of the string 341 | j' = B.length bstr - coerce j 342 | in B.take (j' - i') (B.drop i' bstr) 343 | 344 | instance Eq Span where 345 | x == y = spanToBs x == spanToBs y 346 | 347 | instance Ord Span where 348 | compare x y = compare (spanToBs x) (spanToBs y) 349 | 350 | spanToString :: Span -> String 351 | spanToString s = FP.utf8ToStr (spanToBs s) 352 | -------------------------------------------------------------------------------- /src/Conversion.hs: -------------------------------------------------------------------------------- 1 | 2 | module Conversion where 3 | 4 | import Common 5 | import Core 6 | import CoreTypes 7 | import Cubical 8 | 9 | -- Note: neutral inputs (NeSys, Ne, NeSysHCom) are assumed to be forced 10 | -- other things are not! 11 | -- Also: neutral inputs may have different types! 12 | ---------------------------------------------------------------------------------------------------- 13 | 14 | class Conv a where 15 | conv :: NCofArg => DomArg => a -> a -> Bool 16 | 17 | instance Conv Val where 18 | conv t t' = case frcFull t // frcFull t' of 19 | 20 | -- rigid match 21 | (VNe n , VNe n' ) -> conv n n' 22 | (VGlueTy a sys , VGlueTy a' sys' ) -> conv a a' && conv sys sys' 23 | (VPi a b , VPi a' b' ) -> conv a a' && conv b b' 24 | (VLam t , VLam t' ) -> conv t t' 25 | (VPath a t u , VPath a' t' u' ) -> conv a a' && conv t t' && conv u u' 26 | (VPLam _ _ t , VPLam _ _ t' ) -> conv t t' 27 | -- (VSg a b , VSg a' b' ) -> b^.name == b'^.name && conv a a' && conv b b' 28 | (VSg a b , VSg a' b' ) -> conv a a' && conv b b' 29 | -- (VWrap x a , VWrap x' a' ) -> x == x' && conv a a' 30 | (VWrap x a , VWrap x' a' ) -> conv a a' 31 | (VPair _ t u , VPair _ t' u' ) -> conv t t' && conv u u' 32 | (VU , VU ) -> True 33 | (VLine a , VLine a' ) -> conv a a' 34 | (VLLam t , VLLam t' ) -> conv t t' 35 | (VTyCon inf ts , VTyCon inf' ts' ) -> inf^.tyId == inf'^.tyId && conv ts ts' 36 | (VHTyCon inf ts , VHTyCon inf' ts' ) -> inf^.tyId == inf'^.tyId && conv ts ts' 37 | (VDCon inf sp , VDCon inf' sp' ) -> inf^.conId == inf'^.conId && conv sp sp' 38 | (VGlue t eqs sys , VGlue t' eqs' sys' ) -> conv t t' && conv eqs eqs' && conv sys sys' 39 | 40 | (VHDCon i _ fs s , VHDCon i' _ fs' s' ) -> 41 | i^.conId == i'^.conId && conv fs fs' && conv s s' 42 | 43 | (VHCom r1 r2 a t b , VHCom r1' r2' a' t' b') -> 44 | conv r1 r1' && conv r2 r2' && conv a a' && conv t t' && conv b b' 45 | 46 | -- Holes convert to anything 47 | (VHole{}, _) -> True 48 | (_, VHole{}) -> True 49 | 50 | -- Alternative: track hole evaluation context 51 | -- (VHole _ p s e , VHole _ p' s' e' ) -> p == p' && conv s s' && conv e e' 52 | 53 | -- eta 54 | (VLam t , t' ) -> fresh \x -> conv (t ∙ x) (t' ∙ x) 55 | (t , VLam t' ) -> fresh \x -> conv (t ∙ x) (t' ∙ x) 56 | (VPair x t u , t' ) -> conv t (proj1 x t') && conv u (proj2 x t') 57 | (t , VPair x t' u' ) -> conv (proj1 x t) t' && conv (proj2 x t) u' 58 | (VPLam l r t , t' ) -> freshI \i -> conv (t ∙ i) (papp l r t' i) 59 | (t , VPLam l r t' ) -> freshI \i -> conv (papp l r t i) (t' ∙ i) 60 | (VLLam t , t' ) -> freshI \i -> conv (t ∙ i) (lapp t' i) 61 | (t , VLLam t' ) -> freshI \i -> conv (lapp t i) (t' ∙ i) 62 | (VPack x t , t' ) -> conv t (unpack x t') 63 | (t , VPack x t' ) -> conv (unpack x t) t' 64 | (VGlue b sys fib , t' ) -> conv b (ungluen t' sys) 65 | (t , VGlue b' sys' fib' ) -> conv (ungluen t sys') b' 66 | 67 | (VSub{}, _ ) -> impossible 68 | (_ , VSub{}) -> impossible 69 | (VUnf{}, _ ) -> impossible 70 | (_ , VUnf{}) -> impossible 71 | _ -> False 72 | 73 | instance Conv Ne where 74 | 75 | conv t t' = case unSubNe t // unSubNe t' of 76 | (NLocalVar x , NLocalVar x' ) -> x == x' 77 | (NDontRecurse inf, NDontRecurse inf' ) -> inf^.recId == inf'^.recId 78 | (NApp t u , NApp t' u' ) -> conv t t' && conv u u' 79 | (NPApp l r t i , NPApp l' r' t' i' ) -> conv t t' && conv i i' 80 | (NProj1 n _ , NProj1 n' _ ) -> conv n n' 81 | (NProj2 n _ , NProj2 n' _ ) -> conv n n' 82 | (NLApp t i , NLApp t' i' ) -> conv t t' && conv i i' 83 | 84 | (NCoe r1 r2 a t, NCoe r1' r2' a' t') -> 85 | conv r1 r1' && conv r2 r2' && conv a a' && conv t t' 86 | 87 | -- the types of the "a"-s can be different Glue-s a priori, that's 88 | -- why we first convert sys-es. Neutrals of different types can 89 | -- be converted, and then the type equality is part of the output, 90 | -- but the "a" here aren't neutrals! 91 | (NUnglue a sys, NUnglue a' sys' ) -> conv sys sys' && conv a a' 92 | 93 | (NCase t b cs , NCase t' b' cs' ) -> conv t t' && conv b b' && convCases cs cs' 94 | (NHCase t b cs, NHCase t' b' cs') -> conv t t' && conv b b' && convHCases cs cs' 95 | 96 | (NUnpack n _ , NUnpack n' _ ) -> conv n n' 97 | 98 | (NSub{} , _ ) -> impossible 99 | (t , NSub{} ) -> impossible 100 | _ -> False 101 | 102 | convCases' :: NCofArg => DomArg => RecurseArg => Sub -> Env -> Cases -> Sub -> Env -> Cases -> Bool 103 | convCases' sub env cs sub' env' cs' = case (cs, cs') of 104 | (CSNil , CSNil ) -> True 105 | (CSCons x xs t cs , CSCons _ _ t' cs') -> 106 | (case pushVars env xs of 107 | (!env, !dom) -> case pushVars env' xs of 108 | (!env', !_) -> 109 | let ?dom = dom in 110 | let v = (let ?env = env ; ?sub = sub in eval t ) in 111 | let v' = (let ?env = env'; ?sub = sub' in eval t') in 112 | conv v v') 113 | && 114 | convCases' sub env cs sub' env' cs' 115 | _ -> 116 | impossible 117 | 118 | convHCases' :: NCofArg => DomArg => RecurseArg => Sub -> Env -> HCases -> Sub -> Env -> HCases -> Bool 119 | convHCases' sub env cs sub' env' cs' = case (cs, cs') of 120 | (HCSNil , HCSNil ) -> True 121 | (HCSCons x xs is t cs , HCSCons _ _ _ t' cs') -> 122 | (case pushVars env xs of 123 | (!env, !dom) -> case pushVars env' xs of 124 | (!env', !_) -> case pushIVars sub is of 125 | (!sub, !cof) -> case pushIVars sub' is of 126 | (!sub', _) -> 127 | let ?dom = dom in 128 | let ?cof = cof in 129 | let v = (let ?env = env ; ?sub = sub in eval t ) in 130 | let v' = (let ?env = env'; ?sub = sub' in eval t') in 131 | conv v v') 132 | && 133 | convHCases' sub env cs sub' env' cs' 134 | _ -> 135 | impossible 136 | 137 | convCases :: NCofArg => DomArg => EvalClosure Cases -> EvalClosure Cases -> Bool 138 | convCases (EC sub env _ cs) (EC sub' env' _ cs') = case (wkSub sub, wkSub sub') of 139 | (!sub, !sub') -> 140 | -- approximate conversion, based on case body 141 | (ptrEq cs cs' && conv sub sub' && conv env env') 142 | || 143 | (let ?recurse = DontRecurse 144 | in convCases' sub env cs sub' env' cs') 145 | 146 | convHCases :: NCofArg => DomArg => EvalClosure HCases -> EvalClosure HCases -> Bool 147 | convHCases (EC sub env _ cs) (EC sub' env' _ cs') = case (wkSub sub, wkSub sub') of 148 | (!sub, !sub') -> 149 | -- approximate conversion, based on case body 150 | (ptrEq cs cs' && conv sub sub' && conv env env') 151 | || 152 | (let ?recurse = DontRecurse 153 | in convHCases' sub env cs sub' env' cs') 154 | 155 | instance Conv NeCof where 156 | conv (NCEq i j) (NCEq i' j') = conv i i' && conv j j' 157 | 158 | instance Conv VDSpine where 159 | conv sp sp' = case (sp, sp') of 160 | (VDNil , VDNil ) -> True 161 | (VDCons t sp , VDCons t' sp' ) -> conv t t' && conv sp sp' 162 | _ -> impossible 163 | 164 | instance Conv a => Conv [a] where 165 | conv as as' = case (as, as') of 166 | ([], []) -> True 167 | (a:as, a':as') -> conv a a' && conv as as' 168 | _ -> False 169 | {-# inline conv #-} 170 | 171 | instance Conv (BindI Val) where 172 | conv t t' = freshI \i -> conv (t ∙ i) (t' ∙ i); {-# inline conv #-} 173 | 174 | instance Conv (BindILazy Val) where 175 | conv t t' = freshI \i -> conv (t ∙ i) (t' ∙ i); {-# inline conv #-} 176 | 177 | instance Conv a => Conv (BindCofLazy a) where 178 | conv t t' = conv (t^.binds) (t'^.binds) 179 | && assumeCof (t^.binds) (conv (t^.body) (t'^.body)) 180 | {-# inline conv #-} 181 | 182 | instance Conv a => Conv (BindCof a) where 183 | conv t t' = conv (t^.binds) (t'^.binds) 184 | && assumeCof (t^.binds) (conv (t^.body) (t'^.body)) 185 | {-# inline conv #-} 186 | 187 | instance Conv I where 188 | conv r r' = frc r == frc r'; {-# inline conv #-} 189 | 190 | instance Conv Sub where 191 | conv (Sub _ _ s) (Sub _ _ s') = go s s' where 192 | go :: NCofArg => IList -> IList -> Bool 193 | go s s' = case (s, s') of 194 | (ILNil , ILNil ) -> True 195 | (ILDef is i, ILDef is' i') -> go is is' && conv i i' 196 | _ -> False 197 | 198 | instance Conv NeSys where 199 | conv t t' = case (t, t') of 200 | (NSEmpty , NSEmpty ) -> True 201 | (NSCons t sys , NSCons t' sys' ) -> conv t t' && conv sys sys' 202 | _ -> False 203 | 204 | instance Conv NeSysHCom where 205 | conv t t' = case (t, t') of 206 | (NSHEmpty , NSHEmpty ) -> True 207 | (NSHCons t sys , NSHCons t' sys' ) -> conv t t' && conv sys sys' 208 | _ -> False 209 | 210 | instance Conv NamedClosure where 211 | conv t t' = fresh \x -> conv (t ∙ x) (t' ∙ x); {-# inline conv #-} 212 | 213 | instance Conv NamedIClosure where 214 | conv t t' = freshI \i -> conv (t ∙ i) (t' ∙ i); {-# inline conv #-} 215 | 216 | instance Conv NeCof' where 217 | conv (NeCof' _ c) (NeCof' _ c') = conv c c' 218 | {-# inline conv #-} 219 | 220 | instance Conv Env where 221 | conv e e' = case (e, e') of 222 | (ENil , ENil ) -> True 223 | (EDef e t , EDef e' t' ) -> conv e e' && conv t t' 224 | _ -> False 225 | -------------------------------------------------------------------------------- /src/Cubical.hs: -------------------------------------------------------------------------------- 1 | 2 | module Cubical ( 3 | module Cubical.Interval 4 | , module Cubical.Substitution 5 | , module Cubical.Cofibration 6 | ) where 7 | 8 | import Cubical.Interval 9 | import Cubical.Substitution 10 | import Cubical.Cofibration 11 | -------------------------------------------------------------------------------- /src/Cubical/Cofibration.hs: -------------------------------------------------------------------------------- 1 | 2 | module Cubical.Cofibration where 3 | 4 | import Common 5 | import Cubical.Interval 6 | import Cubical.Substitution 7 | import Statistics (bumpMaxIVar) 8 | 9 | -- import qualified Data.IVarSet as IS 10 | 11 | ---------------------------------------------------------------------------------------------------- 12 | 13 | data Cof = CEq I I deriving Show 14 | data NeCof = NCEq I I deriving Show 15 | 16 | -- | Substitution which maps each ivar to its smallest 17 | -- representative ivar in the equivalence class. 18 | data NCof = NCof IVar IList 19 | 20 | type NCofArg = (?cof :: NCof) 21 | 22 | data NeCof' = NeCof' { 23 | neCof'Extended :: NCof 24 | , neCof'Extra :: {-# unpack #-} NeCof} 25 | 26 | -- TODO: unbox 27 | data VCof 28 | = VCTrue 29 | | VCFalse 30 | | VCNe {-# unpack #-} NeCof' 31 | 32 | makeFields ''NeCof' 33 | 34 | instance HasDom NCof where 35 | dom (NCof d is) = d; {-# inline dom #-} 36 | setDom i (NCof d is) = NCof i (dropIList (d - i) is); {-# inline setDom #-} 37 | 38 | instance Show NCof where 39 | show nc = "[" ++ go nc "" ++ "]" where 40 | sep = \case ILNil -> id 41 | _ -> (" | "++) 42 | go (NCof d is) = case is of 43 | ILNil -> id 44 | ILDef is i -> go (NCof (d - 1) is).sep is.((show (d - 1)++" = "++show i)++) 45 | 46 | instance Show NeCof' where 47 | show (NeCof' c c') = show (c, c') 48 | 49 | deriving instance Show VCof 50 | 51 | instance Lift NCof where 52 | lift (NCof d is) = runIO (bumpMaxIVar d >> (pure $! (NCof (d + 1) (ILDef is (IVar d))))) 53 | -- lift (NCof d is) = NCof (d + 1) (ILDef is (IVar d)) 54 | {-# inline lift #-} 55 | 56 | lookupNCof :: IVar -> NCof -> I 57 | lookupNCof i (NCof d is) = lookupIList (d - i - 1) is 58 | {-# inline lookupNCof #-} 59 | 60 | appNCof :: NCof -> I -> I 61 | appNCof ncof i = matchIVar i (\i -> lookupNCof i ncof) i 62 | {-# inline appNCof #-} 63 | 64 | orient :: (IVar, IVar) -> (IVar, IVar) 65 | orient (i, j) | i < j = (i, j) 66 | | True = (j, i) 67 | 68 | solve :: NCof -> IVar -> I -> NCof 69 | solve (NCof d is) i j = NCof d (go d is i j) where 70 | go d is i j = case is of 71 | ILNil -> impossible 72 | ILDef is i' | d - 1 == i -> ILDef is j 73 | | i' == IVar i -> ILDef (go (d - 1) is i j) j 74 | | True -> ILDef (go (d - 1) is i j) i' 75 | 76 | -- | Equate forced I-s. 77 | eq# :: NCofArg => I -> I -> VCof 78 | eq# i j = case (i, j) of 79 | (i, j) | i == j -> VCTrue 80 | (IVar i, IVar j) -> case orient (i, j) of 81 | (i', j') -> VCNe (NeCof' (solve ?cof j' (IVar i')) (NCEq (IVar i) (IVar j))) 82 | (IVar i, j ) -> VCNe (NeCof' (solve ?cof i j) (NCEq (IVar i) j)) 83 | (i , IVar j) -> VCNe (NeCof' (solve ?cof j i) (NCEq (IVar j) i)) 84 | _ -> VCFalse 85 | 86 | eq :: NCofArg => I -> I -> VCof 87 | eq i j = eq# (appNCof ?cof i) (appNCof ?cof j) 88 | {-# inline eq #-} 89 | 90 | evalI :: NCofArg => SubArg => I -> I 91 | evalI i = appNCof ?cof (sub i) 92 | {-# inline evalI #-} 93 | 94 | eqS :: SubArg => NCofArg => I -> I -> VCof 95 | eqS i j = eq# (evalI i) (evalI j) 96 | {-# inline eqS #-} 97 | 98 | -- | Extend with a forced neutral NeCof. Error if non-neutral. 99 | conjNeCof :: NCofArg => NeCof -> NCof 100 | conjNeCof = \case 101 | NCEq i j -> case (i, j) of 102 | (i , j ) | i == j -> impossible 103 | (IVar i, IVar j) -> case orient (i, j) of (i, j) -> solve ?cof j (IVar i) 104 | (IVar i, j ) -> solve ?cof i j 105 | (i , IVar j) -> solve ?cof j i 106 | (i, j) -> impossible 107 | 108 | assumeNeCof :: NeCof -> (NCofArg => a) -> (NCofArg => a) 109 | assumeNeCof nc act = let ?cof = conjNeCof nc in act 110 | {-# inline assumeNeCof #-} 111 | 112 | idNCof :: IVar -> NCof 113 | idNCof i = NCof i (go 0 ILNil) where 114 | go j acc | i == j = acc 115 | go j acc = go (j + 1) (ILDef acc (IVar j)) 116 | 117 | emptyNCof :: NCof 118 | emptyNCof = NCof 0 ILNil 119 | 120 | appNCofToSub :: NCof -> Sub -> Sub 121 | appNCofToSub nc (Sub d c is) = Sub d c (go nc is) where 122 | go nc ILNil = ILNil 123 | go nc (ILDef is i) = ILDef (go nc is) (appNCof nc i) 124 | 125 | wkSub :: NCofArg => Sub -> Sub 126 | wkSub s = setDom (dom ?cof) s 127 | {-# inline wkSub #-} 128 | 129 | evalCof :: NCofArg => SubArg => Cof -> VCof 130 | evalCof (CEq i j) = eqS i j 131 | {-# inline evalCof #-} 132 | -------------------------------------------------------------------------------- /src/Cubical/Interval.hs: -------------------------------------------------------------------------------- 1 | 2 | module Cubical.Interval where 3 | 4 | import GHC.Exts 5 | import Common 6 | 7 | -- Interval expressions are 64 bits: I0 is 0, I1 is 1, any other value N 8 | -- represents IVar (n - 2). 9 | 10 | -- We use 64-bit bitsets of IVar-s. 11 | maxIVar :: IVar 12 | maxIVar = 63 13 | {-# inline maxIVar #-} 14 | 15 | -- data I = I0 | I1 | IVar IVar 16 | newtype I = I {unI :: Word} 17 | deriving Eq 18 | 19 | -- | Unsafely flip an I which is known to be 0 or 1. 20 | flip# :: I -> I 21 | flip# (I w) = I (1 - w) 22 | {-# inline flip# #-} 23 | 24 | unpackI# :: I -> (# (# #) | (# #) | Word# #) 25 | unpackI# (I (W# x)) = case x of 26 | 0## -> (# (# #) | | #) -- I0 27 | 1## -> (# | (# #) | #) -- I1 28 | x -> (# | | x `minusWord#` 2## #) -- IVar 29 | {-# inline unpackI# #-} 30 | 31 | pattern I0 :: I 32 | pattern I0 <- (unpackI# -> (# (# #) | | #)) where I0 = I 0 33 | 34 | pattern I1 :: I 35 | pattern I1 <- (unpackI# -> (# | (# #) | #)) where I1 = I 1 36 | 37 | pattern IVar :: IVar -> I 38 | pattern IVar x <- (unpackI# -> (# | | (W# -> (IVar# -> x)) #)) where 39 | IVar (IVar# x) = I (x + 2) 40 | {-# complete I0, I1, IVar #-} 41 | 42 | matchIVar :: I -> (IVar -> a) -> a -> a 43 | matchIVar (I x) f ~a | x >= 2 = f (IVar# (x - 2)) 44 | | True = a 45 | {-# inline matchIVar #-} 46 | 47 | instance Show I where 48 | showsPrec _ I0 acc = 'I':'0':acc 49 | showsPrec _ I1 acc = 'I':'1':acc 50 | showsPrec d (IVar x) acc = showParen (d > 10) (("IVar " ++).showsPrec 11 x) acc 51 | -------------------------------------------------------------------------------- /src/Cubical/Substitution.hs: -------------------------------------------------------------------------------- 1 | 2 | 3 | module Cubical.Substitution where 4 | 5 | import Data.Foldable 6 | import qualified Data.Array.LI as ALI 7 | 8 | import Common 9 | import Cubical.Interval 10 | 11 | ---------------------------------------------------------------------------------------------------- 12 | 13 | data IList = ILNil | ILDef IList I 14 | deriving (Eq, Show) 15 | 16 | data Sub = Sub# Word32 Word32 IList 17 | deriving Eq 18 | 19 | instance Show Sub where 20 | show (Sub d c is) = show (d, c, subToList (Sub d c is)) 21 | 22 | pattern Sub :: IVar -> IVar -> IList -> Sub 23 | pattern Sub d c is <- Sub# (fromIntegral -> d) (fromIntegral -> c) is where 24 | Sub d c is = Sub# (fromIntegral d) (fromIntegral c) is 25 | {-# complete Sub #-} 26 | 27 | instance HasDom Sub where 28 | dom (Sub d c s) = d; {-# inline dom #-} 29 | setDom i (Sub _ c s) = Sub i c s; {-# inline setDom #-} 30 | 31 | instance HasCod Sub where 32 | cod (Sub d c s) = c; {-# inline cod #-} 33 | setCod i (Sub d c is) = Sub d i (dropIList (c - i) is); {-# inline setCod #-} 34 | 35 | type SubArg = (?sub :: Sub) 36 | 37 | mkIdSub :: IVar -> Sub 38 | mkIdSub i = Sub i i (go 0 ILNil) where 39 | go j acc | i == j = acc 40 | go j acc = go (j + 1) (ILDef acc (IVar j)) 41 | 42 | mkIdSubs :: IVar -> [Sub] 43 | mkIdSubs i = go i (case mkIdSub i of Sub _ _ is -> is) [] where 44 | go i is acc = case is of 45 | ILNil -> acc 46 | ILDef is j -> go (i - 1) is (Sub (i-1) (i-1) is:acc) 47 | 48 | idCacheSize :: IVar 49 | idCacheSize = 1000 50 | 51 | idSubs :: ALI.Array Sub 52 | idSubs = ALI.fromList $ mkIdSubs idCacheSize 53 | {-# noinline idSubs #-} 54 | 55 | idSub :: IVar -> Sub 56 | idSub i | i < idCacheSize = idSubs ALI.! fromIntegral i 57 | | True = mkIdSub i 58 | {-# inline idSub #-} 59 | 60 | lookupIList :: IVar -> IList -> I 61 | lookupIList i is = case (i, is) of 62 | (0, ILDef _ i) -> i 63 | (i, ILDef is _) -> lookupIList (i - 1) is 64 | _ -> impossible 65 | 66 | lookupSub :: IVar -> Sub -> I 67 | lookupSub i (Sub d c s) = lookupIList (c - i - 1) s 68 | {-# inline lookupSub #-} 69 | 70 | emptySub :: IVar -> Sub 71 | emptySub d = Sub d 0 ILNil 72 | {-# inline emptySub #-} 73 | 74 | subToList :: Sub -> [I] 75 | subToList (Sub _ _ is) = go is [] where 76 | go ILNil acc = acc 77 | go (ILDef is i) acc = go is (i:acc) 78 | 79 | extError :: I -> IVar -> a 80 | extError i d = error $ "ext: " ++ show i ++ " " ++ show d 81 | {-# noinline extError #-} 82 | 83 | -- | Extend a sub with an expression. Domain doesn't change. 84 | ext :: Sub -> I -> Sub 85 | ext (Sub d c is) i = 86 | let i' = matchIVar i (\case i | i < d -> IVar i; i -> extError (IVar i) d) i 87 | in Sub d (c + 1) (ILDef is i') 88 | {-# inline ext #-} 89 | 90 | instance Lift Sub where 91 | lift (Sub d c is) = Sub (d + 1) (c + 1) (ILDef is (IVar d)) 92 | {-# inline lift #-} 93 | 94 | subFromList :: [I] -> Sub 95 | subFromList is = 96 | let dom = case [x | IVar x <- is] of [] -> 0; is -> maximum is + 1 97 | in foldl' ext (emptySub dom) is 98 | 99 | dropIList :: IVar -> IList -> IList 100 | dropIList i is = case i // is of 101 | (0, is) -> is 102 | (i, ILDef is _) -> dropIList (i - 1) is 103 | _ -> impossible 104 | 105 | class SubAction a where 106 | sub :: SubArg => a -> a 107 | 108 | appSub :: SubAction a => Sub -> a -> a 109 | appSub s a = let ?sub = s in sub a 110 | {-# inline appSub #-} 111 | 112 | instance SubAction I where 113 | sub i = matchIVar i (\i -> lookupSub i ?sub) i 114 | {-# inline sub #-} 115 | 116 | -- | Substitution composition. 117 | instance SubAction Sub where 118 | sub s@(Sub d c is) 119 | | d > cod ?sub = error ("SUB " ++ show s ++ " " ++ show ?sub) 120 | | True = Sub (dom ?sub) c (go is ?sub) where 121 | go :: IList -> Sub -> IList 122 | go is s = case is of 123 | ILNil -> ILNil 124 | ILDef is i -> ILDef (go is s) (appSub s i) 125 | 126 | pushSub :: Sub -> Sub -> Sub 127 | pushSub (Sub d c is) (Sub _ c' is') = Sub d (c + c') (go is is') where 128 | go is ILNil = is 129 | go is (ILDef is' i) = ILDef (go is is') i 130 | -------------------------------------------------------------------------------- /src/Data/LvlMap.hs: -------------------------------------------------------------------------------- 1 | 2 | module Data.LvlMap where 3 | 4 | import qualified Data.IntMap.Strict as IM 5 | import Common 6 | 7 | newtype Map a = Map {unMap :: IM.IntMap a} 8 | deriving (Eq, Show, Semigroup, Monoid) via IM.IntMap a 9 | 10 | infixl 9 ! 11 | (!) :: Map a -> Lvl -> a 12 | (!) m x = coerce m IM.! fromIntegral x; {-# inline (!) #-} 13 | 14 | lookup :: Lvl -> Map a -> Maybe a 15 | lookup x m = IM.lookup (fromIntegral x) (coerce m); {-# inline lookup #-} 16 | 17 | insert :: Lvl -> a -> Map a -> Map a 18 | insert x a m = coerce (IM.insert (fromIntegral x) a (coerce m)); {-# inline insert #-} 19 | 20 | delete :: Lvl -> Map a -> Map a 21 | delete x m = Map (IM.delete (fromIntegral x) (unMap m)); {-# inline delete #-} 22 | 23 | update :: (a -> Maybe a) -> Lvl -> Map a -> Map a 24 | update f x m = coerce (IM.update f (fromIntegral x) (coerce m)); {-# inline update #-} 25 | 26 | adjust :: (a -> a) -> Lvl -> Map a -> Map a 27 | adjust f x m = coerce (IM.adjust f (fromIntegral x) (coerce m)); {-# inline adjust #-} 28 | 29 | elems :: Map a -> [a] 30 | elems x = IM.elems (coerce x); {-# inline elems #-} 31 | 32 | foldrWithKey' :: (Lvl -> a -> b -> b) -> b -> Map a -> b 33 | foldrWithKey' f b as = IM.foldrWithKey' (\l a b -> f (fromIntegral l) a b) b (unMap as) 34 | {-# inline foldrWithKey' #-} 35 | -------------------------------------------------------------------------------- /src/ElabState.hs: -------------------------------------------------------------------------------- 1 | 2 | module ElabState where 3 | 4 | import qualified Data.Map.Strict as M 5 | import qualified Data.Set as S 6 | import Data.IORef 7 | import Data.Time.Clock 8 | 9 | import Common 10 | import CoreTypes 11 | import Cubical 12 | 13 | import qualified Presyntax as P 14 | import qualified Data.LvlMap as LM 15 | import qualified Data.ByteString.Char8 as B 16 | 17 | data TopEntry 18 | = TEDef DefInfo 19 | | TERec (Maybe RecInfo) 20 | | TETyCon TyConInfo 21 | | TEDCon DConInfo 22 | | TEHTyCon HTyConInfo 23 | | TEHDCon HDConInfo 24 | deriving Show 25 | 26 | data Locals 27 | = LNil 28 | | LBind Locals Name ~VTy ~Ty 29 | | LBindI Locals Name 30 | | LCof Locals NeCof 31 | deriving Show 32 | 33 | data RevLocals 34 | = RLNil 35 | | RLBind Name ~VTy ~Ty RevLocals 36 | | RLBindI Name RevLocals 37 | | RLCof NeCof RevLocals 38 | deriving Show 39 | 40 | revLocals :: Locals -> RevLocals 41 | revLocals = go RLNil where 42 | go acc LNil = acc 43 | go acc (LBind ls x a qa) = go (RLBind x a qa acc) ls 44 | go acc (LBindI ls r) = go (RLBindI r acc) ls 45 | go acc (LCof ls c) = go (RLCof c acc) ls 46 | 47 | lookupLocalType :: Ix -> Locals -> Box VTy 48 | lookupLocalType x ls = case (x, ls) of 49 | (0, LBind _ _ a _ ) -> Box a 50 | (x, LBind ls _ _ _) -> lookupLocalType (x - 1) ls 51 | (x, LBindI ls _ ) -> lookupLocalType x ls 52 | (x, LCof ls _ ) -> lookupLocalType x ls 53 | _ -> impossible 54 | 55 | type PosArg = (?srcPos :: Box Pos) 56 | type LocalsArg = (?locals :: Locals) 57 | type Elab a = LocalsArg => NCofArg => DomArg => EnvArg => PosArg => a 58 | 59 | data PrintingOpts = PrintingOpts { 60 | printingOptsVerbose :: Bool 61 | , printingOptsErrPrinting :: Bool 62 | , printingOptsPrintNextUnfolding :: Bool 63 | , printingOptsShowHoleCxts :: Bool } 64 | deriving Show 65 | 66 | data LoadState = LoadState { 67 | loadStateLoadedFiles :: S.Set FilePath 68 | , loadStateCurrentPath :: FilePath -- path of file being currently processed 69 | 70 | , loadStateBasePath :: FilePath -- Base path of the module hierarchy. 71 | 72 | , loadStateLoadCycle :: [FilePath] -- Chain of imports going back to original input file to cctt. 73 | -- We check for import cycles on this. 74 | , loadStateCurrentSrc :: B.ByteString 75 | } deriving Show 76 | 77 | data HCaseBoundaryCheck where 78 | HCBC :: (LocalsArg, NCofArg, DomArg, EnvArg, PosArg) => 79 | Env -> [HDConInfo] -> NamedClosure -> HCases -> HCaseBoundaryCheck 80 | deriving Show via DontShow HCaseBoundaryCheck 81 | 82 | data State = State { 83 | stateTop :: M.Map B.ByteString TopEntry 84 | , stateTop' :: LM.Map TopEntry 85 | , stateLvl :: Lvl 86 | , stateLoadState :: LoadState 87 | , stateParsingTime :: NominalDiffTime 88 | , statePrintingOpts :: PrintingOpts 89 | , stateHCaseBoundaryChecks :: [HCaseBoundaryCheck] 90 | , stateUnfolding :: Bool 91 | } deriving Show 92 | 93 | makeFields ''LoadState 94 | makeFields ''State 95 | makeFields ''PrintingOpts 96 | 97 | defaultPrintingOpts :: PrintingOpts 98 | defaultPrintingOpts = PrintingOpts False False False True 99 | 100 | initLoadState :: LoadState 101 | initLoadState = LoadState mempty mempty mempty mempty mempty 102 | 103 | initState :: State 104 | initState = State mempty mempty 0 initLoadState 0 defaultPrintingOpts [] False 105 | 106 | stateRef :: IORef State 107 | stateRef = runIO $ newIORef initState 108 | {-# noinline stateRef #-} 109 | 110 | getState :: IO State 111 | getState = readIORef stateRef 112 | 113 | putState :: State -> IO () 114 | putState = writeIORef stateRef 115 | 116 | modState :: (State -> State) -> IO () 117 | modState = modifyIORef' stateRef 118 | 119 | resetElabState :: IO () 120 | resetElabState = putState initState 121 | 122 | withTopElab :: Elab (IO a) -> IO a 123 | withTopElab act = do 124 | st <- getState 125 | let ls = st^.loadState 126 | let ?locals = LNil 127 | ?cof = emptyNCof 128 | ?dom = 0 129 | ?env = ENil 130 | ?srcPos = Box $! initialPos (SrcFile (ls^.currentPath) (ls^.currentSrc)) 131 | act 132 | {-# inline withTopElab #-} 133 | 134 | -- | Bind a fibrant variable. 135 | bind :: Name -> VTy -> Ty -> Elab (Val -> a) -> Elab a 136 | bind x ~a ~qa act = 137 | let v = vVar ?dom in 138 | let ?dom = ?dom + 1 139 | ?env = EDef ?env v 140 | ?locals = LBind ?locals x a qa in 141 | let _ = ?dom; _ = ?env in 142 | act v 143 | {-# inline bind #-} 144 | 145 | -- | Define a fibrant variable. 146 | define :: Name -> VTy -> Ty -> Val -> Elab a -> Elab a 147 | define x ~a ~qa ~v act = 148 | let ?env = EDef ?env v 149 | ?dom = ?dom + 1 150 | ?locals = LBind ?locals x a qa in 151 | let _ = ?env; _ = ?dom in 152 | act 153 | {-# inline define #-} 154 | 155 | bindToName :: P.Bind -> Name 156 | bindToName = \case 157 | P.BName x -> NSpan x 158 | P.BDontBind _ -> N_ 159 | 160 | -- | Bind an ivar. 161 | bindI :: Name -> Elab (IVar -> a) -> Elab a 162 | bindI x act = 163 | let fresh = dom ?cof in 164 | if fresh == maxIVar then error "RAN OUT OF IVARS IN ELABORATION" else 165 | let ?cof = lift ?cof in 166 | let ?locals = LBindI ?locals x in 167 | seq ?cof $ act fresh 168 | {-# inline bindI #-} 169 | 170 | bindCof :: NeCof' -> Elab a -> Elab a 171 | bindCof (NeCof' cof c) act = 172 | let ?cof = cof 173 | ?locals = LCof ?locals c in 174 | act; {-# inline bindCof #-} 175 | 176 | isNameUsed :: Elab (B.ByteString -> Bool) 177 | isNameUsed x = go ?locals x where 178 | go :: Locals -> B.ByteString -> Bool 179 | go LNil _ = False 180 | go (LBind ls x' _ _) x = x == nameToBs x' || go ls x 181 | go (LBindI ls x') x = x == nameToBs x' || go ls x 182 | go (LCof ls _) x = go ls x 183 | 184 | -- | Try to pick an informative fresh ivar name. 185 | pickIVarName :: Elab Name 186 | pickIVarName 187 | | not (isNameUsed "i") = i_ 188 | | not (isNameUsed "j") = j_ 189 | | not (isNameUsed "k") = k_ 190 | | not (isNameUsed "l") = l_ 191 | | not (isNameUsed "m") = m_ 192 | | True = i_ 193 | -------------------------------------------------------------------------------- /src/Errors.hs: -------------------------------------------------------------------------------- 1 | 2 | module Errors where 3 | 4 | import Control.Exception 5 | import Data.List 6 | 7 | import Common 8 | import CoreTypes 9 | import ElabState 10 | import Cubical 11 | import Pretty 12 | 13 | data Error 14 | = NameNotInScope Span 15 | | LocalLvlNotInScope 16 | | TopLvlNotInScope 17 | | UnexpectedI 18 | | UnexpectedIType 19 | | ExpectedI 20 | | ExpectedPiPathLine ~Tm 21 | | ExpectedSg ~Tm 22 | | ExpectedGlueTy ~Tm 23 | | ExpectedPathLine ~Tm 24 | | ExpectedPath ~Tm 25 | | ExpectedPathULineU ~Tm 26 | | ExpectedNonDepFun ~Tm 27 | | CantInferLam 28 | | CantInferPair 29 | | CantInfer 30 | | CantInferGlueTm 31 | | CantConvert ~Tm ~Tm 32 | | CantConvertCof Cof Cof 33 | | CantConvertEndpoints ~Tm ~Tm 34 | | CantConvertReflEndpoints ~Tm ~Tm 35 | | CantConvertExInf ~Tm ~Tm 36 | | GlueTmSystemMismatch Sys 37 | | TopShadowing Pos 38 | | NonNeutralCofInSystem 39 | | NoSuchField Name ~Tm 40 | | CantInferHole 41 | | CantInferConParamTy String 42 | | ImportCycle FilePath [FilePath] 43 | | CantOpenFile FilePath 44 | | GenericError String 45 | | ExpectedInductiveType ~Tm 46 | | CaseMismatch 47 | | UnderAppliedCon String 48 | | OverAppliedCon String 49 | deriving Show 50 | 51 | instance Exception Error where 52 | 53 | data ErrInCxt where 54 | ErrInCxt :: (LocalsArg, NCofArg, DomArg, EnvArg, PosArg) => Error -> ErrInCxt 55 | 56 | instance Show ErrInCxt where 57 | show (ErrInCxt e) = show e 58 | 59 | instance Exception ErrInCxt 60 | 61 | err :: Elab (Error -> IO a) 62 | err e = throwIO $ ErrInCxt e 63 | 64 | showError :: PrettyArgs (Error -> String) 65 | showError = \case 66 | UnderAppliedCon nm -> 67 | "Constructor " ++ show nm ++ " applied to too few arguments" 68 | 69 | OverAppliedCon nm -> 70 | "Constructor " ++ show nm ++ " applied to too many arguments" 71 | 72 | CaseMismatch -> 73 | "Case expressions must cover all constructors of an inductive type in the order of declaration" 74 | 75 | GenericError msg -> 76 | msg 77 | 78 | CantConvertExInf t u -> 79 | "Can't convert expected type\n\n" ++ 80 | " " ++ pretty u ++ "\n\n" ++ 81 | "and inferred type\n\n" ++ 82 | " " ++ pretty t 83 | 84 | CantInfer -> 85 | "Can't infer type for expression" 86 | 87 | CantConvert t u -> 88 | "Can't convert\n\n" ++ 89 | " " ++ pretty u ++ "\n\n" ++ 90 | "and\n\n" ++ 91 | " " ++ pretty t 92 | 93 | CantConvertEndpoints t u -> 94 | "Can't convert expected path endpoint\n\n" ++ 95 | " " ++ pretty u ++ "\n\n" ++ 96 | "to the inferred endpoint\n\n" ++ 97 | " " ++ pretty t 98 | 99 | CantConvertReflEndpoints t u -> 100 | "Can't convert endpoints when checking \"refl\":\n\n" ++ 101 | " " ++ pretty u ++ "\n\n" ++ 102 | " " ++ pretty t 103 | 104 | CantConvertCof c1 c2 -> 105 | "Can't convert expected path endpoint\n\n" ++ 106 | " " ++ pretty c1 ++ "\n\n" ++ 107 | "to\n\n" ++ 108 | " " ++ pretty c2 109 | 110 | NameNotInScope x -> 111 | "Name not in scope: " ++ show x 112 | 113 | LocalLvlNotInScope -> 114 | "Local De Bruijn level not in scope" 115 | 116 | TopLvlNotInScope -> 117 | "Top-level De Bruijn level not in scope" 118 | 119 | TopShadowing pos -> 120 | "Duplicate top-level name\n" 121 | ++ "Previously defined at: " ++ show (parsePos pos) 122 | 123 | UnexpectedI -> 124 | "Unexpected interval expression" 125 | 126 | ExpectedI -> 127 | "Expected an interval expression" 128 | 129 | ExpectedPath a -> 130 | "Expected a path type, inferred\n\n" ++ 131 | " " ++ pretty a 132 | 133 | ExpectedPiPathLine a -> 134 | "Expected a function, line or path type, inferred\n\n" ++ 135 | " " ++ pretty a 136 | 137 | ExpectedSg a -> 138 | "Expected a sigma type, inferred\n\n" ++ 139 | " " ++ pretty a 140 | 141 | ExpectedGlueTy a -> 142 | "Expected a glue type, inferred\n\n" ++ 143 | " " ++ pretty a 144 | 145 | ExpectedPathLine a -> 146 | "Expected a path type or a line type, inferred\n\n" ++ 147 | " " ++ pretty a 148 | 149 | ExpectedPathULineU a -> 150 | "Expected a path type or a line type in U, inferred\n\n" ++ 151 | " " ++ pretty a 152 | 153 | ExpectedNonDepFun a -> 154 | "Expected a non-dependent function type, inferred\n\n" ++ 155 | " " ++ pretty a 156 | 157 | CantInferLam -> 158 | "Can't infer type for lambda expression" 159 | 160 | CantInferPair -> 161 | "Can't infer type for pair expression" 162 | 163 | CantInferConParamTy nm -> 164 | "Can't infer type for data constructor " ++ show nm ++ " which has type parameters" 165 | 166 | CantInferGlueTm -> 167 | "Can't infer type for glue expression" 168 | 169 | GlueTmSystemMismatch sys -> 170 | "Can't match glue system with expected Glue type system\n\n" ++ 171 | " " ++ pretty sys 172 | 173 | UnexpectedIType -> 174 | "The type of intervals I can be only used in function domains" 175 | 176 | NonNeutralCofInSystem -> 177 | "Only neutral cofibrations are allowed in systems" 178 | 179 | NoSuchField x a -> 180 | "Field projection " ++ show x ++ " is not supported by inferred type:\n\n" ++ 181 | " " ++ pretty a 182 | 183 | CantInferHole -> 184 | "Can't infer type for hole" 185 | 186 | ImportCycle path cycle -> 187 | "Imports form a cycle:\n\n" 188 | ++ intercalate " -> " (dropWhile (/=path) cycle ++ [path]) 189 | 190 | CantOpenFile path -> 191 | "Can't open file: " ++ path 192 | 193 | ExpectedInductiveType a -> 194 | "Expected an inductive type, inferred\n\n" ++ 195 | " " ++ pretty a 196 | 197 | displayErrInCxt :: ErrInCxt -> IO () 198 | displayErrInCxt (ErrInCxt e) = withPrettyArgs do 199 | modState (printingOpts.errPrinting .~ True) 200 | 201 | let Box pos = ?srcPos 202 | ParsedPos path file linum colnum = parsePos pos 203 | lnum = show (linum + 1) 204 | lpad = map (const ' ') lnum 205 | 206 | putStrLn ("\nERROR " ++ path ++ ":" ++ lnum ++ ":" ++ show colnum) 207 | putStrLn (lpad ++ " |") 208 | putStrLn (lnum ++ " | " ++ (lines file !! linum)) 209 | putStrLn (lpad ++ " | " ++ replicate colnum ' ' ++ "^") 210 | putStrLn (showError e) 211 | putStrLn "" 212 | -------------------------------------------------------------------------------- /src/Lexer.hs: -------------------------------------------------------------------------------- 1 | {-# options_ghc -Wno-x-partial #-} -- usage of head in prettyError 2 | 3 | module Lexer where 4 | 5 | import FlatParse.Stateful 6 | hiding (Parser, runParser, string, char, cut, Pos, Span, getPos, err) 7 | 8 | import qualified FlatParse.Stateful as FP 9 | import qualified Data.ByteString.Char8 as B 10 | import Language.Haskell.TH 11 | 12 | import Data.String 13 | import qualified Data.Set as S 14 | import Data.Char 15 | 16 | import Common 17 | import Presyntax 18 | 19 | ---------------------------------------------------------------------------------------------------- 20 | 21 | data Expected 22 | = Lit String -- ^ Expected a concrete `String`. 23 | | Msg String -- ^ An arbitrary error message. 24 | deriving (Eq, Show, Ord) 25 | 26 | instance IsString Expected where fromString = Lit 27 | 28 | data Error' 29 | = Precise Expected -- ^ Expected exactly something. 30 | | Imprecise [String] -- ^ Expected one of multiple things. 31 | deriving Show 32 | 33 | data Error 34 | = Error {-# unpack #-} Pos Error' 35 | | DontUnboxError 36 | deriving Show 37 | 38 | merge :: Error -> Error -> Error 39 | merge err@(Error p e) err'@(Error p' e') 40 | | p == p', Imprecise ss <- e, Imprecise ss' <- e' = err' 41 | | otherwise = err 42 | merge _ _ = impossible 43 | {-# noinline merge #-} 44 | 45 | type Parser = FP.Parser Src Error 46 | 47 | prettyError :: Src -> Error -> String 48 | prettyError src DontUnboxError = impossible 49 | prettyError src (Error pos e) = 50 | 51 | let bstr = srcToBs src 52 | ls = FP.linesUtf8 bstr 53 | (l, c) = head $ posLineCols bstr [rawPos pos] 54 | line = if 0 <= l && l < length ls then ls !! l else "" 55 | linum = show (l+1) 56 | lpad = map (const ' ') linum 57 | 58 | expected (Lit s) = "expected " ++ s 59 | expected (Msg s) = s 60 | 61 | err (Precise exp) = expected exp 62 | err (Imprecise exps) = imprec $ S.toList $ S.fromList exps 63 | 64 | imprec :: [String] -> String 65 | imprec [] = impossible 66 | imprec [s] = "expected " ++ s 67 | imprec ss = "expected one of the following:\n" ++ unlines (map (" - "++) ss) 68 | 69 | showsrc (SrcFile path _) = path 70 | showsrc SrcImpossible = impossible 71 | 72 | in "parse error:\n"++ 73 | showsrc src ++ ":" ++ show l ++ ":" ++ show c ++ ":\n" ++ 74 | lpad ++ "|\n" ++ 75 | linum ++ "| " ++ line ++ "\n" ++ 76 | lpad ++ "| " ++ replicate c ' ' ++ "^\n" ++ 77 | err e 78 | 79 | getPos :: Parser Pos 80 | getPos = Pos <$> ask <*> FP.getPos 81 | {-# inline getPos #-} 82 | 83 | -- | Throw an error. 84 | err :: Error' -> Parser a 85 | err err = do 86 | pos <- getPos 87 | FP.err $ Error pos err 88 | 89 | -- | Imprecise cut: we slap a list of expected things on inner errors. 90 | cut :: Parser a -> [String] -> Parser a 91 | cut p exps = do 92 | pos <- getPos 93 | cutting p (Error pos (Imprecise exps)) merge 94 | 95 | -- | Precise cut: we propagate at most a single expected thing. 96 | pcut :: Parser a -> Expected -> Parser a 97 | pcut p exp = do 98 | pos <- getPos 99 | cutting p (Error pos (Precise exp)) merge 100 | 101 | runParser :: Parser a -> Src -> Result Error a 102 | runParser p src = FP.runParser p src 0 (srcToBs src) 103 | 104 | -- | Run parser, print pretty error on failure. 105 | testParser :: Show a => Parser a -> String -> IO () 106 | testParser p str = case SrcFile "(interactive)" (strToUtf8 str) of 107 | b -> case runParser p b of 108 | Err e -> putStrLn $ prettyError b e 109 | OK a _ _ -> print a 110 | Fail -> putStrLn "parse error" 111 | 112 | dummy :: Parser () 113 | dummy = pure () 114 | {-# inline dummy #-} 115 | 116 | ws :: Parser () 117 | ws = $(switch [| case _ of 118 | " " -> dummy >> ws 119 | "\n" -> dummy >> ws 120 | "\t" -> dummy >> ws 121 | "\r" -> dummy >> ws 122 | "--" -> dummy >> lineComment 123 | "{-" -> dummy >> multilineComment 124 | _ -> pure () |]) 125 | 126 | -- | Parse a line comment. 127 | lineComment :: Parser () 128 | lineComment = 129 | withOption anyWord8 130 | (\case 10 -> ws 131 | _ -> lineComment) 132 | (pure ()) 133 | 134 | -- | Parse a potentially nested multiline comment. 135 | multilineComment :: Parser () 136 | multilineComment = go (1 :: Int) where 137 | go 0 = ws 138 | go n = $(switch [| case _ of 139 | "\n" -> dummy >> go n 140 | "-}" -> dummy >> go (n - 1) 141 | "{-" -> dummy >> go (n + 1) 142 | _ -> FP.branch anyChar (go n) (pure ()) |]) 143 | 144 | token :: Parser a -> Parser a 145 | token p = p <* ws 146 | {-# inline token #-} 147 | 148 | spanOfToken :: Parser a -> Parser Span 149 | spanOfToken p = do 150 | src <- ask 151 | FP.Span x y <- FP.spanOf p <* ws 152 | pure $ Span# src x y 153 | {-# inline spanOfToken #-} 154 | 155 | -- | Read a starting character of an identifier. 156 | identStartChar :: Parser Char 157 | identStartChar = fusedSatisfy 158 | isLatinLetter 159 | (\c -> isGreekLetter c || isLetter c) 160 | isLetter 161 | isLetter 162 | 163 | -- | Read a non-starting character of an identifier. 164 | identChar :: Parser Char 165 | identChar = fusedSatisfy 166 | (\c -> isLatinLetter c || FP.isDigit c || c == '\'' || c == '-') 167 | (\c -> isGreekLetter c || isAlphaNum c) 168 | isAlphaNum 169 | isAlphaNum 170 | 171 | inlineIdentChar :: Parser Char 172 | inlineIdentChar = fusedSatisfy 173 | (\c -> isLatinLetter c || FP.isDigit c || c == '\'' || c == '-') 174 | (\c -> isGreekLetter c || isAlphaNum c) 175 | isAlphaNum 176 | isAlphaNum 177 | {-# inline inlineIdentChar #-} 178 | 179 | -- | Parse a non-keyword string, return the `Span` of the symbol. 180 | sym :: String -> Q Exp 181 | sym str = [| spanOfToken $(FP.string str) |] 182 | 183 | -- | Parse a non-keyword string, throw precise error on failure, return the `Span` of the symbol 184 | sym' :: String -> Q Exp 185 | sym' str = 186 | [| spanOfToken ($(FP.string str) `pcut` Lit str) |] 187 | 188 | -- | Parse a keyword string, return the `Span`. 189 | kw :: String -> Q Exp 190 | kw str = 191 | [| spanOfToken ($(FP.string str) `notFollowedBy` identChar) |] 192 | 193 | -- | Parse a keyword string, throw precise error on failure, return the `Span`. 194 | kw' :: String -> Q Exp 195 | kw' str = 196 | [| spanOfToken (($(FP.string str) `notFollowedBy` identChar) `pcut` Lit str) |] 197 | 198 | -- | Raw non-token parser that reads any keyword. 199 | anyKw :: Parser () 200 | anyKw = $(switch [| case _ of 201 | "Glue" -> eof 202 | "I" -> eof 203 | "U" -> eof 204 | "ap" -> eof 205 | "case" -> eof 206 | "coe" -> eof 207 | "com" -> eof 208 | "inductive" -> eof 209 | "higher" -> eof 210 | "glue" -> eof 211 | "module" -> eof 212 | "hcom" -> eof 213 | "import" -> eof 214 | "let" -> eof 215 | "refl" -> eof 216 | "unglue" -> eof 217 | "λ" -> eof |]) 218 | 219 | manyIdentChars :: Parser () 220 | manyIdentChars = skipMany identChar 221 | 222 | scanIdent :: Parser () 223 | scanIdent = identStartChar >> skipMany inlineIdentChar 224 | 225 | withSpan :: Parser a -> Parser (a, Span) 226 | withSpan a = FP.withSpan a \a (FP.Span x y) -> do 227 | src <- ask 228 | pure (a, Span# src x y) 229 | 230 | identBase :: Parser Presyntax.Name 231 | identBase = FP.withSpan scanIdent \_ span@(FP.Span x y) -> do 232 | fails $ inSpan span anyKw 233 | ws 234 | src <- ask 235 | pure $ Span# src x y 236 | 237 | -- | Parse an identifier. 238 | ident :: Parser Presyntax.Name 239 | ident = identBase 240 | {-# inline ident #-} 241 | 242 | -- | Parse an identifier. 243 | ident' :: Parser Presyntax.Name 244 | ident' = identBase `pcut` Lit "identifier" 245 | {-# inline ident' #-} 246 | 247 | test = FP.runParser traceRest () 0 (B.pack "") 248 | -------------------------------------------------------------------------------- /src/Main.hs: -------------------------------------------------------------------------------- 1 | 2 | module Main where 3 | 4 | import MainInteraction 5 | import Control.Exception 6 | import Statistics 7 | 8 | main :: IO () 9 | main = do 10 | catch mainInteraction \case 11 | UserInterrupt -> renderStats 12 | e -> throw e 13 | -------------------------------------------------------------------------------- /src/MainInteraction.hs: -------------------------------------------------------------------------------- 1 | 2 | module MainInteraction where 3 | 4 | import qualified Data.Map.Strict as M 5 | import System.Environment 6 | import System.Exit 7 | import qualified FlatParse.Stateful as FP 8 | import Text.Read (readMaybe) 9 | import Data.Maybe 10 | 11 | import Common 12 | import CoreTypes 13 | import Core 14 | import Cubical 15 | import ElabState 16 | import Elaboration 17 | import Pretty 18 | import Quotation 19 | import Statistics 20 | 21 | helpMsg = unlines [ 22 | "usage: cctt [nf ] [trace ] [trace-to-end ] [ty ] [elab] [verbose] [unfold] [no-hole-cxts]" 23 | ,"" 24 | ,"Checks . Options:" 25 | ," nf Prints the normal form of top-level definition ." 26 | ," Implies the \"unfold\" option." 27 | ," trace Interactively steps through head unfoldings of a value." 28 | ," ty Prints the type of the top-level definition ." 29 | ," elab Prints the elaboration output." 30 | ," verbose Prints path endpoints, hcom types and hole source positions explicitly." 31 | ," Verbose output can be always re-checked by cctt." 32 | ," unfold Option to immediately unfold all top definitions during evaluation." 33 | ," May increase performance significantly." 34 | ," no-hole-cxt Turns off printing local contexts of holes." 35 | ] 36 | 37 | elabPath :: FilePath -> String -> IO () 38 | elabPath path args = mainWith (pure $ path : words args) 39 | 40 | mainInteraction :: IO () 41 | mainInteraction = mainWith getArgs 42 | 43 | frc0 :: Val -> Val 44 | frc0 = let ?cof = emptyNCof; ?dom = 0 in frc 45 | {-# inline frc0 #-} 46 | 47 | quote0Frozen :: Frozen -> Tm 48 | quote0Frozen f = 49 | let ?cof = emptyNCof; ?dom = 0; ?trace = True; ?opt = QDontUnfold in quoteFrozen f 50 | {-# inline quote0Frozen #-} 51 | 52 | quote0 :: Quote a b => a -> b 53 | quote0 = let ?cof = emptyNCof; ?dom = 0; ?opt = QDontUnfold in quote 54 | {-# inline quote0 #-} 55 | 56 | tracingCommands :: IO () 57 | tracingCommands = do 58 | putStrLn "Tracing commands:" 59 | putStrLn " ENTER proceed one step" 60 | putStrLn " skip ENTER print final value" 61 | putStrLn " skip ENTER skip given number of steps" 62 | putStrLn " trace ENTER print all steps until final value" 63 | putStrLn " trace ENTER proceed given number of steps" 64 | putStrLn " help ENTER print this message" 65 | 66 | traceUnfold :: Val -> Int -> Int -> Bool -> IO () 67 | traceUnfold v batch i silent = case frc0 v of 68 | VUnf f v v' -> do 69 | if batch <= 0 then do 70 | putStrLn $ pretty0 $ quote0Frozen v 71 | putStrLn $ "STEP " ++ show (i + 1) ++ ", NEXT UNFOLDING: " ++ show (f^.name) 72 | let getCommand = do 73 | l <- words <$> getLine 74 | let retry = do 75 | putStrLn "Invalid input" 76 | tracingCommands 77 | getCommand 78 | case l of 79 | [] -> 80 | traceUnfold v' 0 (i + 1) False 81 | ["trace"] -> 82 | traceUnfold v' maxBound (i + 1) False 83 | ["trace", n] -> case readMaybe n of 84 | Just n | n > 0 -> traceUnfold v' (n - 1) (i + 1) False 85 | _ -> retry 86 | ["skip"] -> do 87 | traceUnfold v' maxBound (i + 1) True 88 | ["skip", n] -> case readMaybe n of 89 | Just n | n > 0 -> traceUnfold v' (n - 1) (i + 1) True 90 | _ -> retry 91 | ["help"] -> do 92 | tracingCommands 93 | getCommand 94 | _ -> 95 | retry 96 | getCommand 97 | else do 98 | unless silent $ do 99 | putStrLn $ pretty0 $ quote0Frozen v 100 | putStrLn $ "STEP " ++ show (i + 1) 101 | putStrLn "" 102 | traceUnfold v' (batch - 1) (i + 1) silent 103 | v -> do 104 | putStrLn $ pretty0 $ quote0 v 105 | putStrLn $ "FINISHED AT STEP " ++ show (i + 1) 106 | 107 | parseArgs :: [String] 108 | -> IO (FilePath, Maybe String, Maybe String, Maybe String, Bool, Bool, Bool, Bool) 109 | parseArgs args = do 110 | let exit = putStrLn helpMsg >> exitSuccess 111 | (path, args) <- case args of 112 | path:args -> pure (path, args) 113 | _ -> exit 114 | (printnf, args) <- case args of 115 | "nf":def:args -> pure (Just def, args) 116 | args -> pure (Nothing, args) 117 | (trace, args) <- case args of 118 | "trace":def:args -> pure (Just def, args) 119 | args -> pure (Nothing, args) 120 | (printty, args) <- case args of 121 | "ty":def:args -> pure (Just def, args) 122 | args -> pure (Nothing, args) 123 | (elab, args) <- case args of 124 | "elab":args -> pure (True, args) 125 | args -> pure (False, args) 126 | (verbose, args) <- case args of 127 | "verbose":args -> pure (True, args) 128 | args -> pure (False, args) 129 | (unfold, args) <- case args of 130 | "unfold":args -> pure (True, args) 131 | args -> pure (False, args) 132 | holeCxts <- case args of 133 | "no-hole-cxts":[] -> pure False 134 | [] -> pure True 135 | _ -> exit 136 | pure (path, printnf, trace, printty, elab, verbose, unfold, holeCxts) 137 | 138 | mainWith :: IO [String] -> IO () 139 | mainWith getArgs = do 140 | resetElabState 141 | resetStats 142 | (path, printnf, trace, printty, printelab, verbosity, unfold, holeCxts) <- parseArgs =<< getArgs 143 | 144 | unfold <- pure $ isJust printnf || unfold 145 | 146 | modState $ 147 | (printingOpts %~ (verbose .~ verbosity) 148 | . (showHoleCxts .~ holeCxts)) 149 | . (unfolding .~ unfold) 150 | 151 | 152 | (_, !totaltime) <- timed (elaborate path) 153 | st <- getState 154 | putStrLn "" 155 | putStrLn ("parsing time: " ++ show (st^.parsingTime)) 156 | putStrLn ("elaboration time: " ++ show (totaltime - st^.parsingTime)) 157 | putStrLn ("checked " ++ show (st^.lvl) ++ " definitions") 158 | 159 | when printelab do 160 | putStrLn ("\n------------------------------------------------------------") 161 | putStrLn ("------------------------------------------------------------") 162 | putStrLn $ pretty0 (st^.top') 163 | 164 | case printnf of 165 | Just x -> do 166 | (!nf, !nftime) <- case M.lookup (FP.strToUtf8 x) (st^.top) of 167 | Just (TEDef i) -> do 168 | let ?env = ENil; ?cof = emptyNCof; ?dom = 0 169 | timedPure (quoteUnfold (i^.defVal)) 170 | _ -> do 171 | putStrLn $ "No top-level definition with name: " ++ show x 172 | exitSuccess 173 | putStrLn ("------------------------------------------------------------") 174 | putStrLn ("------------------------------------------------------------") 175 | putStrLn "" 176 | putStrLn $ "Normal form of " ++ x ++ ":\n\n" ++ pretty0 nf 177 | putStrLn "" 178 | putStrLn $ "Normalized in " ++ show nftime 179 | Nothing -> 180 | pure () 181 | 182 | case trace of 183 | Just x -> do 184 | v <- case M.lookup (FP.strToUtf8 x) (st^.top) of 185 | Just (TEDef i) -> 186 | pure $! i^.defVal 187 | _ -> do 188 | putStrLn $ "No top-level definition with name: " ++ show x 189 | exitSuccess 190 | putStrLn "" 191 | tracingCommands 192 | putStrLn "" 193 | traceUnfold v 0 0 False 194 | Nothing -> 195 | pure () 196 | 197 | case printty of 198 | Just x -> do 199 | case M.lookup (FP.strToUtf8 x) (st^.top) of 200 | Just (TEDef i) -> do 201 | putStrLn "" 202 | putStrLn $ show (i^.pos) 203 | putStrLn $ x ++ " : " ++ pretty0 (i^.defTy) 204 | _ -> do 205 | putStrLn $ "No top-level definition with name: " ++ show x 206 | exitSuccess 207 | Nothing -> 208 | pure () 209 | 210 | renderStats 211 | -------------------------------------------------------------------------------- /src/Presyntax.hs: -------------------------------------------------------------------------------- 1 | {-# options_ghc -funbox-strict-fields #-} 2 | 3 | module Presyntax where 4 | 5 | import Common hiding (Name) 6 | 7 | type Ty = Tm 8 | type Name = Span 9 | 10 | data Bind 11 | = BName Name 12 | | BDontBind Pos 13 | deriving Show 14 | 15 | data LamAnnot 16 | = LANone 17 | | LAAnn Ty 18 | | LADesugared Ty 19 | deriving Show 20 | 21 | data Tm 22 | = Ident Name 23 | | LocalLvl Pos Lvl Pos 24 | | TopLvl Pos Lvl (Maybe Lvl) Pos -- @@n or @@n.m (the latter identifies a data constructor) 25 | | I Pos 26 | | ILvl Pos IVar Pos 27 | | I0 Pos 28 | | I1 Pos 29 | | Let Pos Name (Maybe Ty) Tm Tm 30 | | Pi Pos Bind Ty Ty 31 | | App Tm Tm 32 | | PApp Pos Tm Tm Tm Tm -- explicit endpoints: t {l}{r} u 33 | | Lam Pos Bind LamAnnot Tm 34 | | PLam Pos Tm Tm Bind Tm -- explicit endpoints: λ {l}{r} x. t 35 | 36 | | Parens Pos Tm Pos 37 | 38 | | Sg Pos Bind Ty Ty 39 | | Pair Tm Tm 40 | | Proj1 Tm Pos 41 | | Proj2 Tm Pos 42 | | ProjField Tm Name 43 | 44 | | Wrap Pos Bind Ty Pos 45 | | Hole Pos (Maybe Bind) 46 | 47 | | Case Pos Tm (Maybe (Bind, Ty)) [CaseItem] Pos 48 | | Split Pos [CaseItem] Pos 49 | 50 | | U Pos 51 | | Path Tm Tm -- x = y 52 | | DepPath BindMaybe Tm Tm -- x ={i.A} y | x ={p} y 53 | 54 | | Coe Pos Tm Tm BindMaybe Tm -- coe r r' i A t 55 | | HCom Pos Tm Tm (Maybe Ty) SysHCom Tm -- hcom r r' {A} [α x. t] u 56 | | Com Pos Tm Tm BindMaybe SysHCom Tm -- com r r' i A [α x. t] u 57 | 58 | | GlueTy Pos Ty Sys Pos -- Glue A [α. B] 59 | | GlueTm Pos Tm (Maybe Sys) Sys Pos -- glue a {} 60 | | Unglue Pos Tm 61 | 62 | | Refl Pos Pos 63 | | Sym Tm Pos 64 | | Trans Tm Tm 65 | | Ap Pos Tm Tm 66 | deriving Show 67 | 68 | data Cof = CEq Tm Tm 69 | deriving Show 70 | 71 | data Sys = SEmpty | SCons Cof Tm Sys 72 | deriving Show 73 | 74 | data BindMaybe = BMBind Bind Tm | BMDontBind Tm 75 | deriving Show 76 | 77 | data SysHCom = SHEmpty | SHCons Pos Cof BindMaybe SysHCom 78 | deriving Show 79 | 80 | type Constructor = (Name, [(Bind, Ty)]) 81 | type HConstructor = (Name, [(Bind, Ty)], Maybe Sys) 82 | type CaseItem = (Bind, [Bind], Tm) 83 | 84 | data Top 85 | = TDef Pos Name (Maybe Ty) Tm Top 86 | | TData Pos Name [(Bind, Ty)] [Constructor] Top 87 | | THData Pos Name [(Bind, Ty)] [HConstructor] Top 88 | | TImport Pos String Top 89 | | TEmpty 90 | deriving Show 91 | 92 | class Spanned a where 93 | spanOf :: a -> Span 94 | spanOf a = Span (leftPos a) (rightPos a) 95 | 96 | leftPos :: a -> Pos 97 | leftPos a = case spanOf a of Span p _ -> p 98 | 99 | rightPos :: a -> Pos 100 | rightPos a = case spanOf a of Span _ p -> p 101 | 102 | instance Spanned Bind where 103 | spanOf (BDontBind p) = Span p p 104 | spanOf (BName x) = x 105 | 106 | instance Spanned Name where 107 | spanOf x = x 108 | 109 | instance Spanned Tm where 110 | leftPos = \case 111 | Ident x -> leftPos x 112 | LocalLvl l _ _ -> l 113 | TopLvl l _ _ _ -> l 114 | I l -> l 115 | ILvl l _ _ -> l 116 | I0 l -> l 117 | I1 l -> l 118 | Let l _ _ _ _ -> l 119 | Pi l _ _ _ -> l 120 | App t _ -> leftPos t 121 | PApp l _ _ _ _ -> l 122 | Lam l _ _ _ -> l 123 | PLam l _ _ _ _ -> l 124 | Parens l _ _ -> l 125 | Sg l _ _ _ -> l 126 | Pair t _ -> leftPos t 127 | Proj1 t _ -> leftPos t 128 | Proj2 t _ -> leftPos t 129 | ProjField t _ -> leftPos t 130 | Wrap l _ _ _ -> l 131 | Hole l _ -> l 132 | Case l _ _ _ _ -> l 133 | Split l _ _ -> l 134 | U l -> l 135 | Path t _ -> leftPos t 136 | DepPath _ t _ -> leftPos t 137 | Coe l _ _ _ _ -> l 138 | HCom l _ _ _ _ _ -> l 139 | Com l _ _ _ _ _ -> l 140 | GlueTy l _ _ _ -> l 141 | GlueTm l _ _ _ _ -> l 142 | Unglue l _ -> l 143 | Refl l _ -> l 144 | Sym t _ -> leftPos t 145 | Trans t _ -> leftPos t 146 | Ap l _ _ -> l 147 | 148 | rightPos = \case 149 | Ident x -> rightPos x 150 | LocalLvl _ _ r -> r 151 | TopLvl _ _ _ r -> r 152 | I r -> r 153 | ILvl _ _ r -> r 154 | I0 r -> r 155 | I1 r -> r 156 | Let _ _ _ _ u -> rightPos u 157 | Pi _ _ _ b -> rightPos b 158 | App _ t -> rightPos t 159 | PApp _ _ _ _ t -> rightPos t 160 | Lam _ _ _ t -> rightPos t 161 | PLam _ _ _ _ t -> rightPos t 162 | Parens _ _ r -> r 163 | Sg _ _ _ b -> rightPos b 164 | Pair _ t -> rightPos t 165 | Proj1 _ r -> r 166 | Proj2 _ r -> r 167 | ProjField _ r -> rightPos r 168 | Wrap _ _ _ r -> r 169 | Hole p x -> maybe p rightPos x 170 | Case _ _ _ _ r -> r 171 | Split _ _ r -> r 172 | U r -> r 173 | Path _ t -> rightPos t 174 | DepPath _ _ t -> rightPos t 175 | Coe _ _ _ _ t -> rightPos t 176 | HCom _ _ _ _ _ t -> rightPos t 177 | Com _ _ _ _ _ t -> rightPos t 178 | GlueTy _ _ _ r -> r 179 | GlueTm _ _ _ _ r -> r 180 | Unglue _ t -> rightPos t 181 | Refl _ r -> r 182 | Sym _ r -> r 183 | Trans _ t -> rightPos t 184 | Ap _ _ t -> rightPos t 185 | 186 | instance (Spanned a, Spanned b) => Spanned (a, b) where 187 | leftPos (x, y) = leftPos x 188 | rightPos (x, y) = rightPos y 189 | 190 | instance Spanned Pos where 191 | leftPos x = x 192 | rightPos x = x 193 | 194 | instance (Spanned a, Spanned b, Spanned c) => Spanned (a, b, c) where 195 | leftPos (x, y, z) = leftPos x 196 | rightPos (x, y, z) = rightPos z 197 | 198 | instance Spanned a => Spanned [a] where 199 | leftPos [] = impossible 200 | leftPos (x:_) = leftPos x 201 | rightPos [] = impossible 202 | rightPos [x] = rightPos x 203 | rightPos (_:xs) = rightPos xs 204 | 205 | instance Spanned BindMaybe where 206 | leftPos (BMBind x t) = leftPos x 207 | leftPos (BMDontBind t) = leftPos t 208 | rightPos (BMBind x t) = rightPos t 209 | rightPos (BMDontBind t) = rightPos t 210 | 211 | instance Spanned Cof where 212 | leftPos (CEq t _) = leftPos t 213 | rightPos (CEq _ t) = rightPos t 214 | 215 | unParens :: Tm -> Tm 216 | unParens (Parens _ t _) = unParens t 217 | unParens t = t 218 | -------------------------------------------------------------------------------- /src/Quotation.hs: -------------------------------------------------------------------------------- 1 | 2 | module Quotation where 3 | 4 | import Cubical 5 | import Common 6 | import CoreTypes 7 | import Core 8 | 9 | -- Note: neutral inputs (NeSys, Ne, NeSysHCom) are assumed to be forced 10 | -- other things are not! 11 | ---------------------------------------------------------------------------------------------------- 12 | 13 | data QOpt = QUnfold | QDontUnfold 14 | 15 | type QuoteOpt = (?opt :: QOpt) 16 | 17 | class Quote a b | a -> b where 18 | quote :: NCofArg => DomArg => QuoteOpt => a -> b 19 | 20 | quoteNoUnfold :: Quote a b => NCofArg => DomArg => a -> b 21 | quoteNoUnfold = let ?opt = QDontUnfold in quote 22 | {-# inline quoteNoUnfold #-} 23 | 24 | quoteUnfold :: Quote a b => NCofArg => DomArg => a -> b 25 | quoteUnfold = let ?opt = QUnfold in quote 26 | {-# inline quoteUnfold #-} 27 | 28 | instance Quote I I where 29 | quote i = frc i; {-# inline quote #-} 30 | 31 | instance Quote Ne Tm where 32 | -- forced input 33 | quote n = case unSubNe n of 34 | NLocalVar x -> LocalVar (lvlToIx ?dom x) 35 | NSub n s -> impossible 36 | NApp t u -> App (quote t) (quote u) 37 | NPApp t l r i -> PApp (quote t) (quote l) (quote r) (quote i) 38 | NProj1 t x -> Proj1 (quote t) x 39 | NProj2 t x -> Proj2 (quote t) x 40 | NUnpack t x -> Unpack (quote t) x 41 | NCoe r r' a t -> Coe (quote r) (quote r') (a^.name) (quote a) (quote t) 42 | NUnglue t sys -> Unglue (quote t) (quote sys) 43 | NLApp t i -> LApp (quote t) (quote i) 44 | NDontRecurse x -> RecursiveCall x 45 | NCase t b cs -> Case (quote t) (b^.name) (quote b) (quoteCases cs) 46 | NHCase t b cs -> HCase (quote t) (b^.name) (quote b) (quoteHCases cs) 47 | 48 | quoteFrozen :: NCofArg => DomArg => QuoteOpt => (?trace :: Bool) => Frozen -> Tm 49 | quoteFrozen = go where 50 | go :: NCofArg => DomArg => QuoteOpt => (?trace :: Bool) => Frozen -> Tm 51 | go t = case unSubFrozen t of 52 | f@(FTopVar inf) -> case ?trace of 53 | True -> TopVar inf PrintTrace 54 | _ -> TopVar inf DontPrintTrace 55 | FSub f s -> impossible 56 | FApp t u -> App (go t) (quote u) 57 | FPApp l r t i -> PApp (quote l) (quote r) (go t) (quote i) 58 | FLApp l i -> LApp (go l) (quote i) 59 | FProj1 x t -> Proj1 (go t) x 60 | FProj2 x t -> Proj2 (go t) x 61 | FUnpack t x -> Unpack (go t) x 62 | 63 | FCoeTy r r' a t | quote r == quote r' -> quote t 64 | FCoeTy r r' a t -> Coe (quote r) (quote r') (a^.name) (freshI \i -> go (a ∙ i)) (quote t) 65 | 66 | FCoeVal r r' a t | quote r == quote r' -> go t 67 | FCoeVal r r' a t -> Coe (quote r) (quote r') (a^.name) (quote a) (go t) 68 | 69 | -- we have to handle system forcing, because the the rest of cctt can 70 | -- only handle systems containing neutral cofibrations! 71 | FHComTy r r' a sys t | quote r == quote r' -> quote t 72 | FHComTy r r' a sys t -> case frc sys of 73 | VSHTotal t' -> quote (t' ∙ r') 74 | VSHNe sys -> HCom (quote r) (quote r') (go a) (quote sys) (quote t) 75 | 76 | FHComVal r r' a sys t | quote r == quote r' -> go t 77 | FHComVal r r' a sys t -> case frc sys of 78 | VSHTotal t' -> quote (t' ∙ r') 79 | VSHNe sys -> HCom (quote r) (quote r') (quote a) (quote sys) (go t) 80 | 81 | FUnglue t sys -> case frc sys of 82 | VSTotal teqv -> App (Proj1 (Proj2 (quote teqv) ty_) f_) (go t) 83 | VSNe sys -> Unglue (go t) (quote sys) 84 | 85 | FCase_ t b cs -> Case (go t) (b^.name) (quote b) (quoteCases cs) 86 | FHCase_ t b cs -> HCase (go t) (b^.name) (quote b) (quoteHCases cs) 87 | 88 | instance Quote Frozen Tm where 89 | quote t = let ?trace = False in quoteFrozen t 90 | 91 | instance Quote Val Tm where 92 | quote t = case frc t of 93 | VSub{} -> impossible 94 | VUnf _ v v' -> case ?opt of QUnfold -> quote v' 95 | _ -> quote v 96 | VNe n -> quote n 97 | VGlueTy a sys -> GlueTy (quote a) (quote sys) 98 | VGlue t eqs sys -> Glue (quote t) (quote eqs) (quote sys) 99 | VPi a b -> Pi (b^.name) (quote a) (quote b) 100 | VLam t -> Lam (t^.name) (quote t) 101 | VPath a lhs rhs -> Path (a^.name) (quote a) (quote lhs) (quote rhs) 102 | VPLam lhs rhs t -> PLam (quote lhs) (quote rhs) (t^.name) (quote t) 103 | VSg a b -> Sg (b^.name) (quote a) (quote b) 104 | VPair x t u -> Pair x (quote t) (quote u) 105 | VWrap x a -> Wrap x (quote a) 106 | VPack x t -> Pack x (quote t) 107 | VU -> U 108 | VHole h -> Hole h 109 | VLine t -> Line (t^.name) (quote t) 110 | VLLam t -> LLam (t^.name) (quote t) 111 | VTyCon x ts -> TyCon x (quote ts) 112 | VDCon ci sp -> DCon ci (quote sp) 113 | VHTyCon i ts -> HTyCon i (quote ts) 114 | VHDCon i ps fs s -> HDCon i (quoteParams ps) (quote fs) s 115 | VHCom r r' a sys t -> HCom (quote r) (quote r') (quote a) (quote sys) (quote t) 116 | 117 | -------------------------------------------------------------------------------- 118 | 119 | quoteCases' :: QuoteOpt => EvalArgs (Cases -> Cases) 120 | quoteCases' = \case 121 | CSNil -> CSNil 122 | CSCons x xs body cs -> 123 | CSCons x xs 124 | (let (!env, !dom) = pushVars ?env xs in 125 | let ?env = env; ?dom = dom in 126 | quote (eval body)) 127 | (quoteCases' cs) 128 | 129 | -- We don't do recursive unfolding under Case binders 130 | quoteCases :: NCofArg => DomArg => QuoteOpt => EvalClosure Cases -> Cases 131 | quoteCases (EC sub env _ cs) = 132 | let ?sub = wkSub sub; ?env = env; ?recurse = DontRecurse in quoteCases' cs 133 | {-# inline quoteCases #-} 134 | 135 | quoteHCases' :: QuoteOpt => EvalArgs (HCases -> HCases) 136 | quoteHCases' = \case 137 | HCSNil -> HCSNil 138 | HCSCons x xs is body cs -> 139 | HCSCons x xs is 140 | (let (!env, !dom) = pushVars ?env xs in 141 | let ?env = env; ?dom = dom in 142 | freshIVarsS is (quote (eval body))) 143 | (quoteHCases' cs) 144 | 145 | -- We don't do recursive unfolding under Case binders 146 | quoteHCases :: NCofArg => DomArg => QuoteOpt => EvalClosure HCases -> HCases 147 | quoteHCases (EC sub env _ cs) = 148 | let ?sub = wkSub sub; ?env = env; ?recurse = DontRecurse in quoteHCases' cs 149 | {-# inline quoteHCases #-} 150 | 151 | quoteParamsNoUnfold :: NCofArg => DomArg => Env -> LazySpine 152 | quoteParamsNoUnfold = let ?opt = QDontUnfold in quoteParams 153 | {-# inline quoteParamsNoUnfold #-} 154 | 155 | quoteParams :: NCofArg => DomArg => QuoteOpt => Env -> LazySpine 156 | quoteParams = go LSPNil where 157 | go acc = \case 158 | ENil -> acc 159 | EDef env v -> go (LSPCons (quote v) acc) env 160 | 161 | -------------------------------------------------------------------------------- 162 | 163 | -- instance Quote VSys Sys where 164 | -- quote = \case 165 | -- VSTo 166 | 167 | instance Quote VDSpine Spine where 168 | quote = \case 169 | VDNil -> SPNil 170 | VDCons t sp -> SPCons (quote t) (quote sp) 171 | 172 | instance Quote a b => Quote (BindCofLazy a) b where 173 | quote t = assumeCof (t^.binds) (quote (t^.body)); {-# inline quote #-} 174 | 175 | instance Quote a b => Quote [a] [b] where 176 | quote = \case 177 | [] -> [] 178 | a:as -> (:) $$! quote a $$! quote as 179 | {-# inline quote #-} 180 | 181 | instance Quote a b => Quote (BindCof a) b where 182 | quote t = assumeCof (t^.binds) (quote (t^.body)); {-# inline quote #-} 183 | 184 | instance Quote NeCof Cof where 185 | quote (NCEq i j) = CEq (quote i) (quote j) 186 | 187 | instance (SubAction a, Quote a b) => Quote (BindI a) b where 188 | quote t = freshI \i -> quote (t ∙ i); {-# inline quote #-} 189 | 190 | instance (SubAction a, Quote a b) => Quote (BindILazy a) b where 191 | quote t = freshI \i -> quote (t ∙ i); {-# inline quote #-} 192 | 193 | instance Quote NeSys Sys where 194 | quote = \case 195 | NSEmpty -> SEmpty 196 | NSCons t sys -> SCons (quote (t^.binds)) (quote t) (quote sys) 197 | 198 | instance Quote NeSysHCom SysHCom where 199 | quote = \case 200 | NSHEmpty -> SHEmpty 201 | NSHCons t sys -> SHCons (quote (t^.binds)) (t^.body.name) (quote t) (quote sys) 202 | 203 | instance Quote NamedClosure Tm where 204 | quote t = fresh \x -> quote (t ∙ x); {-# inline quote #-} 205 | 206 | instance Quote NamedIClosure Tm where 207 | quote t = freshI \i -> quote (t ∙ i); {-# inline quote #-} 208 | 209 | instance Quote Env Spine where 210 | quote = go SPNil where 211 | go acc = \case 212 | ENil -> acc 213 | EDef env v -> go (SPCons (quote v) acc) env 214 | -------------------------------------------------------------------------------- /src/Statistics.hs: -------------------------------------------------------------------------------- 1 | 2 | module Statistics where 3 | 4 | import Common 5 | import qualified Data.Ref.F as RF 6 | 7 | -------------------------------------------------------------------------------- 8 | 9 | hcomStat :: RF.Ref Int 10 | hcomStat = runIO $ RF.new 0 11 | {-# noinline hcomStat #-} 12 | 13 | emptyhcomStat :: RF.Ref Int 14 | emptyhcomStat = runIO $ RF.new 0 15 | {-# noinline emptyhcomStat #-} 16 | 17 | maxIVarStat :: RF.Ref IVar 18 | maxIVarStat = runIO $ RF.new 0 19 | {-# noinline maxIVarStat #-} 20 | 21 | blockStat :: RF.Ref Int 22 | blockStat = runIO $ RF.new 0 23 | {-# noinline blockStat #-} 24 | 25 | unblockStat :: RF.Ref Int 26 | unblockStat = runIO $ RF.new 0 27 | {-# noinline unblockStat #-} 28 | 29 | ------------------------------------------------------------ 30 | 31 | -- bumpHCom :: IO () 32 | -- bumpHCom = pure () 33 | 34 | -- bumpMaxIVar :: IVar -> IO () 35 | -- bumpMaxIVar i = pure () 36 | 37 | -- bumpBlock :: IO () 38 | -- bumpBlock = pure () 39 | 40 | -- bumpUnblock :: IO () 41 | -- bumpUnblock = pure () 42 | 43 | -- bumpEmptyHCom :: IO () 44 | -- bumpEmptyHCom = pure () 45 | 46 | -- bumpHCom' :: Bool -> IO () 47 | -- bumpHCom' isEmpty = pure () 48 | 49 | bumpHCom :: IO () 50 | bumpHCom = RF.modify hcomStat (+1) 51 | 52 | bumpMaxIVar :: IVar -> IO () 53 | bumpMaxIVar i = do 54 | m <- RF.read maxIVarStat 55 | RF.write maxIVarStat (max i m) 56 | 57 | bumpBlock :: IO () 58 | bumpBlock = RF.modify blockStat (+1) 59 | 60 | bumpUnblock :: IO () 61 | bumpUnblock = RF.modify unblockStat (+1) 62 | 63 | bumpEmptyHCom :: IO () 64 | bumpEmptyHCom = RF.modify emptyhcomStat (+1) 65 | 66 | bumpHCom' :: Bool -> IO () 67 | bumpHCom' isEmpty = do 68 | bumpHCom 69 | #ifdef EMPTYHCOMSTATS 70 | if isEmpty then bumpEmptyHCom else pure () 71 | #endif 72 | 73 | {-# inline bumpHCom #-} 74 | {-# inline bumpMaxIVar #-} 75 | {-# inline bumpBlock #-} 76 | {-# inline bumpUnblock #-} 77 | {-# inline bumpEmptyHCom #-} 78 | {-# inline bumpHCom' #-} 79 | 80 | renderStats :: IO () 81 | renderStats = do 82 | hcs <- RF.read hcomStat 83 | ehcs <- RF.read emptyhcomStat 84 | maxi <- RF.read maxIVarStat 85 | bs <- RF.read blockStat 86 | ubs <- RF.read unblockStat 87 | putStrLn $ "Total hcom calls: " ++ show hcs 88 | #ifdef EMPTYHCOMSTATS 89 | putStrLn $ "Non-diagonal empty hcom calls: " ++ show ehcs 90 | putStrLn $ "Empty hcom ratio: " ++ show (fromIntegral ehcs / (fromIntegral hcs :: Double)) 91 | #endif 92 | putStrLn $ "Largest interval scope size: " ++ show maxi 93 | -- putStrLn $ "Total neutral forcings: " ++ show (bs + ubs) 94 | -- putStrLn $ "Of which blocked: " ++ show bs 95 | 96 | resetStats :: IO () 97 | resetStats = do 98 | RF.write hcomStat 0 99 | RF.write emptyhcomStat 0 100 | RF.write maxIVarStat 0 101 | RF.write blockStat 0 102 | RF.write unblockStat 0 103 | -------------------------------------------------------------------------------- /stack.yaml: -------------------------------------------------------------------------------- 1 | 2 | # Note: stack tracing setup: 3 | # 1. disable strict impl params plugin in package.yaml 4 | # 2. enable profiling here 5 | # 3. stack clean, stack build 6 | # 4. customize-group > haskell > haskell-interactive > stack ghci args 7 | # add the following args 8 | # --ghci-options="-fexternal-interpreter" 9 | # --ghci-options="-prof" 10 | # --ghci-options="-fprof-auto-calls" 11 | # 5. reload ghci 12 | 13 | # NOTE: compiled cctt doesn't work in backtrace mode! It segfaults for 14 | # who knows what reason. 15 | 16 | resolver: nightly-2024-04-08 17 | 18 | packages: 19 | - . 20 | 21 | ghc-options: 22 | "$everything": -split-sections 23 | 24 | # build: 25 | # library-profiling: true 26 | # executable-profiling: true 27 | 28 | extra-deps: 29 | - git: https://github.com/AndrasKovacs/ghc-strict-implicit-params.git 30 | commit: fdef75e73c3354bd39b641e6b452b7087839e09b 31 | - git: https://github.com/AndrasKovacs/primdata 32 | commit: cfd5a67cc6448c862590e901f3843e782c83fde7 33 | -------------------------------------------------------------------------------- /tests/Russell.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | 3 | module Russell where 4 | 5 | open import Relation.Binary.PropositionalEquality 6 | open import Data.Product 7 | 8 | data ⊥ : Set where 9 | 10 | ¬ : Set → Set 11 | ¬ A = A → ⊥ 12 | 13 | -- a model of set theory, uses Set : Set 14 | data U : Set where 15 | set : (I : Set) → (I → U) → U 16 | 17 | -- a set is regular if it doesn't contain itself 18 | regular : U → Set 19 | regular (set I f) = (i : I) → ¬ (f i ≡ set I f) 20 | 21 | -- Russell's set: the set of all regular sets 22 | R : U 23 | R = set (Σ U regular) proj₁ 24 | 25 | -- R is not regular 26 | R-nonreg : ¬ (regular R) 27 | R-nonreg reg = reg (R , reg) refl 28 | 29 | -- R is regular 30 | R-reg : regular R 31 | R-reg (x , reg) p = subst regular p reg (x , reg) p 32 | 33 | -- contradiction 34 | 35 | absurd : ⊥ 36 | absurd = R-nonreg R-reg 37 | -------------------------------------------------------------------------------- /tests/Russell.cctt: -------------------------------------------------------------------------------- 1 | 2 | inductive Empty := ; 3 | 4 | inductive SET := set (X : U) (f : X → SET); 5 | 6 | regular : SET → U := λ[ 7 | set X f. (x : X) → f x = set X f → Empty ]; 8 | 9 | R : SET := 10 | set ((S : SET) × regular S) (λ p. p.1); 11 | 12 | R-nonreg : regular R → Empty := 13 | λ reg. reg (R , reg) (λ _. R); 14 | 15 | R-reg : regular R := 16 | λ x p. coe 0 1 (i. regular (p i)) x.2 x p; 17 | 18 | false : Empty 19 | := R-nonreg R-reg; 20 | -------------------------------------------------------------------------------- /tests/RussellNoPath.cctt: -------------------------------------------------------------------------------- 1 | 2 | inductive Empty := ; 3 | 4 | inductive SET := set (X : U) (f : X → SET); 5 | 6 | Eq : SET → SET → U 7 | := λ x y. (P : SET → U) → P x → P y; 8 | 9 | rfl : (x : SET) → Eq x x 10 | := λ x P px. px; 11 | 12 | regular : SET → U := λ[ 13 | set X f. (x : X) → Eq (f x) (set X f) → Empty ]; 14 | 15 | R : SET := 16 | set ((S : SET) × regular S) (λ p. p.1); 17 | 18 | R-nonreg : regular R → Empty := 19 | λ reg. reg (R , reg) (rfl R); 20 | 21 | R-reg : regular R := 22 | λ x p. p regular x.2 x p; 23 | 24 | false : Empty 25 | := R-nonreg R-reg; 26 | -------------------------------------------------------------------------------- /tests/basics.cctt: -------------------------------------------------------------------------------- 1 | 2 | inductive Top := tt; 3 | inductive Bot :=; 4 | not (A : U) := A → Bot; 5 | 6 | inductive Sum (A B : U) := left A | right B; 7 | 8 | Dec (A : U) := Sum A (not A); 9 | 10 | -------------------------------------------------------------------------------- 11 | 12 | isLogEquiv (A B : U)(f : A → B) : U := (B → A); 13 | 14 | isSection (A B : U)(f : A → B) : U := (g : B → A) × (linv : (x : A) → x = g (f x)); 15 | 16 | isRetraction (A B : U)(f : A → B) : U := (g : B → A) × (rinv : (x : B) → f (g x) = x); 17 | 18 | isIso (A B : U)(f : A → B) : U := 19 | (g : B → A) 20 | × (linv : (x : A) → g (f x) = x) 21 | × (rinv : (x : B) → f (g x) = x); 22 | 23 | isEquiv (A B : U)(f : A → B) : U := 24 | (g : B → A) 25 | × (linv : (x : A) → g (f x) = x) 26 | × (rinv : (x : B) → f (g x) = x) 27 | × (coh : (x : A) → (rinv (f x)) ={i. f (linv x i) = f x} refl); 28 | 29 | coeIsEquiv (A : I → U)(r r' : I) : isEquiv (A r) (A r') (λ x. coe r r' A x) := 30 | let f' (i : I)(x : A r) : A i := coe r i A x; 31 | let g' (i : I)(x : A i) : A r := coe i r A x; 32 | let linv' (i : I)(x : A r)(j : I) : A r := hcom r i (A r) [j=0 k. coe k r A (coe r k A x); j=1 k. x] x; 33 | let rinv' (i : I)(x : A i)(j : I) : A i := hcom i r (A i) [j=0 k. coe k i A (coe i k A x); j=1 k. x] x; 34 | 35 | ( λ x. g' r' x 36 | , λ x j. linv' r' x j 37 | , λ x j. rinv' r' x j 38 | , λ x l k. com r r' A 39 | [k=0 i. f' i (linv' i x l) 40 | ;k=1 i. f' i x 41 | ;l=0 i. rinv' i (f' i x) k 42 | ;l=1 i. f' i x] 43 | x 44 | ); 45 | 46 | equiv (A B : U) : U := 47 | (f : A → B) × isEquiv A B f; 48 | 49 | iso (A B : U) : U := (f : A → B) × isIso A B f; 50 | 51 | idEquiv (A : U) : equiv A A := 52 | (λ x. x, λ x. x, λ x _. x, λ x _. x, λ x _ _. x); 53 | 54 | ua (A B : U)(f : equiv A B) : A = B := 55 | λ i. Glue B [i=0. (A, f); i=1. (B, idEquiv B)]; 56 | 57 | J (A : U) (a : A) (C : (x : A) → a = x → U) 58 | (d : C a (λ _. a))(x : A) (p : a = x) : C x p 59 | := let sq (i j : I) : A := hcom 0 j [i=0 _. a; i=1. p] a; 60 | coe 0 1 (i. C (sq i 1) (λ j. sq i j)) d; 61 | 62 | JEq (A : U) (a : A) (C : (x : A) → a = x → U) (d : C a (λ _. a)) 63 | : J A a C d a (λ _. a) = d := 64 | 65 | let sq (i j : I) : A := hcom 0 j [i=0 _. a; i=1 _. a] a; 66 | let cube (i j k : I) : A := hcom 0 j [k=0 x. sq i x; 67 | k=1 _. a; 68 | i=0 _. a; 69 | i=1 _. a] a; 70 | 71 | λ k. com 0 1 (i. C (cube i 1 k) (λ j. cube i j k)) 72 | [k=0 i. coe 0 i (i. C (sq i 1) (λ j. sq i j)) d; 73 | k=1 _. d] 74 | d; 75 | 76 | fiber (A B : U) (f : A → B) (b : B) : U 77 | := (x : A) * f x = b; 78 | 79 | fiberRefl (A B : U) (f : A → B) (a : A) : fiber A B f (f a) 80 | := (a , refl); 81 | 82 | contractIsoFiber (A B : U) (is : iso A B) (a : A) 83 | (fib : fiber A B is.f (is.f a)) 84 | : fib = fiberRefl A B is.f a 85 | := 86 | let sq (j k : I) : A := 87 | hcom k j [k=0 j. is.g (fib.2 j); k=1 _. fib.1] (is.linv (fib.1) k); 88 | let sq2 (i k : I) : A := 89 | hcom 0 k 90 | [i=0. sq 1; 91 | i=1. is.linv a] 92 | (is.g (is.f a)); 93 | λ i. 94 | (sq2 i 1, 95 | λ j. 96 | let aux : A := 97 | hcom j 0 98 | [i=0. sq j; 99 | i=1. is.linv a; 100 | j=1. sq2 i] 101 | (is.linv (sq2 i 1) j); 102 | hcom 0 1 103 | [i=0. is.rinv (fib.2 j); 104 | i=1. is.rinv (is.f a); 105 | j=0. is.rinv (is.f (sq2 i 1)); 106 | j=1. is.rinv (is.f a)] 107 | (is.f aux)); 108 | 109 | isoToEquiv (A B : U) (is : iso A B) : equiv A B := 110 | is.f 111 | , is.g 112 | , λ a i. (contractIsoFiber A B is a (is.g (is.f a), is.rinv (is.f a)) i).1 113 | , is.rinv 114 | , λ a i. (contractIsoFiber A B is a (is.g (is.f a), is.rinv (is.f a)) i).2 115 | ; 116 | 117 | isContr (A : U) : U 118 | := (x : A) * ((y : A) → x = y); 119 | 120 | hasContrFibers (A B : U) (f : A → B) : U := (b : B) → isContr (fiber A B f b); 121 | 122 | connAnd (A : U)(a b : A)(p : a = b) : (λ _. a) ={i. a = p i} p 123 | := 124 | let sq (l k : I) : A := hcom 0 l [k=0 _. a; k=1. p] a; 125 | λ i j. hcom 0 1 [i=0 _. a; i=1. sq j; j=0 _. a; j=1. sq i] a; 126 | 127 | connAndRefl (A : U)(a : A) : connAnd A a a refl = refl 128 | := 129 | λ m. 130 | let sq (l k : I) : A := hcom 0 l [k=0 _. a; k=1 _. a; m=1 _. a] a; 131 | λ i j. hcom 0 1 [i=0 _. a; i=1. sq j; j=0 _. a; j=1. sq i; m=1 _. a] a; 132 | 133 | isEquivToContrFibers (A B : U) (f : A → B) (e : isEquiv A B f) : hasContrFibers A B f 134 | := 135 | λ b. 136 | ((e.g b, e.rinv b), 137 | λ fib i. 138 | let a : A := fib.1; 139 | let p : f a = b := fib.2; 140 | let domFiller (k : I) : A := 141 | hcom 0 k 142 | [i=0 k. e.g (p k); 143 | i=1 _. a] 144 | (e.linv a i); 145 | (domFiller 1, 146 | λ j. 147 | hcom 0 1 148 | [i=0 k. e.rinv (p k) j; 149 | i=1. connAnd B (f a) b p j; 150 | j=0 k. f (domFiller k); 151 | j=1. p] 152 | (e.coh a i j))); 153 | 154 | contrFibersToIsEquiv (A B : U) (f : A → B) (c : hasContrFibers A B f) : isEquiv A B f 155 | := 156 | λ b. (c b).1.1, 157 | λ a i. ((c (f a)).2 (a, refl) i).1, 158 | λ b. (c b).1.2, 159 | λ a i. ((c (f a)).2 (a, refl) i).2; 160 | 161 | isEquivRetractContrFibers (A B : U) (f : A → B) (e : isEquiv A B f) 162 | : contrFibersToIsEquiv A B f (isEquivToContrFibers A B f e) = e 163 | := 164 | λ l. 165 | let domFiller (a : A) (i : I) (k : I) : A := 166 | hcom 0 k 167 | [i=0 _. e.g (f a); 168 | i=1 _. a; 169 | l=1 _. e.linv a i] 170 | (e.linv a i); 171 | (e.g, 172 | λ a i. domFiller a i 1, 173 | e.rinv, 174 | λ a i j. 175 | hcom 0 1 176 | [i=0 _. e.rinv (f a) j; 177 | i=1. connAndRefl B (f a) l j; 178 | j=0 k. f (domFiller a i k); 179 | j=1 _. f a; 180 | l=1 _. e.coh a i j] 181 | (e.coh a i j)); 182 | 183 | -------------------------------------------------------------------------------- 184 | -------------------------------------------------------------------------------- /tests/bench/basics.cctt: -------------------------------------------------------------------------------- 1 | 2 | 3 | inductive Top := tt; 4 | inductive Bot :=; 5 | not (A : U) := A → Bot; 6 | 7 | inductive Sum (A B : U) := left A | right B; 8 | 9 | Dec (A : U) := Sum A (not A); 10 | 11 | -------------------------------------------------------------------------------- 12 | 13 | isLogEquiv (A B : U)(f : A → B) : U := (B → A); 14 | 15 | isSection (A B : U)(f : A → B) : U := (g : B → A) × (linv : (x : A) → x = g (f x)); 16 | 17 | isRetraction (A B : U)(f : A → B) : U := (g : B → A) × (rinv : (x : B) → f (g x) = x); 18 | 19 | isIso (A B : U)(f : A → B) : U := 20 | (g : B → A) 21 | × (linv : (x : A) → g (f x) = x) 22 | × (rinv : (x : B) → f (g x) = x); 23 | 24 | isEquiv (A B : U)(f : A → B) : U := 25 | (g : B → A) 26 | × (linv : (x : A) → g (f x) = x) 27 | × (rinv : (x : B) → f (g x) = x) 28 | × (coh : (x : A) → (rinv (f x)) ={i. f (linv x i) = f x} refl); 29 | 30 | coeIsEquiv (A : I → U)(r r' : I) : isEquiv (A r) (A r') (λ x. coe r r' A x) := 31 | let f' (i : I)(x : A r) : A i := coe r i A x; 32 | let g' (i : I)(x : A i) : A r := coe i r A x; 33 | let linv' (i : I)(x : A r)(j : I) : A r := hcom r i (A r) [j=0 k. coe k r A (coe r k A x); j=1 k. x] x; 34 | let rinv' (i : I)(x : A i)(j : I) : A i := hcom i r (A i) [j=0 k. coe k i A (coe i k A x); j=1 k. x] x; 35 | 36 | ( λ x. g' r' x 37 | , λ x j. linv' r' x j 38 | , λ x j. rinv' r' x j 39 | , λ x l k. com r r' A 40 | [k=0 i. f' i (linv' i x l) 41 | ;k=1 i. f' i x 42 | ;l=0 i. rinv' i (f' i x) k 43 | ;l=1 i. f' i x] 44 | x 45 | ); 46 | 47 | equiv (A B : U) : U := 48 | (f : A → B) × isEquiv A B f; 49 | 50 | iso (A B : U) : U := (f : A → B) × isIso A B f; 51 | 52 | idEquiv (A : U) : equiv A A := 53 | (λ x. x, λ x. x, λ x _. x, λ x _. x, λ x _ _. x); 54 | 55 | ua (A B : U)(f : equiv A B) : A = B := 56 | λ i. Glue B [i=0. (A, f); i=1. (B, idEquiv B)]; 57 | 58 | J (A : U) (a : A) (C : (x : A) → a = x → U) 59 | (d : C a (λ _. a))(x : A) (p : a = x) : C x p 60 | := let sq (i j : I) : A := hcom 0 j [i=0 _. a; i=1. p] a; 61 | coe 0 1 (i. C (sq i 1) (λ j. sq i j)) d; 62 | 63 | JEq (A : U) (a : A) (C : (x : A) → a = x → U) (d : C a (λ _. a)) 64 | : J A a C d a (λ _. a) = d := 65 | 66 | let sq (i j : I) : A := hcom 0 j [i=0 _. a; i=1 _. a] a; 67 | let cube (i j k : I) : A := hcom 0 j [k=0 x. sq i x; 68 | k=1 _. a; 69 | i=0 _. a; 70 | i=1 _. a] a; 71 | 72 | λ k. com 0 1 (i. C (cube i 1 k) (λ j. cube i j k)) 73 | [k=0 i. coe 0 i (i. C (sq i 1) (λ j. sq i j)) d; 74 | k=1 _. d] 75 | d; 76 | 77 | fiber (A B : U) (f : A → B) (b : B) : U 78 | := (x : A) * f x = b; 79 | 80 | fiberRefl (A B : U) (f : A → B) (a : A) : fiber A B f (f a) 81 | := (a , refl); 82 | 83 | contractIsoFiber (A B : U) (is : iso A B) (a : A) 84 | (fib : fiber A B is.f (is.f a)) 85 | : fib = fiberRefl A B is.f a 86 | := 87 | let sq (j k : I) : A := 88 | hcom k j [k=0 j. is.g (fib.2 j); k=1 _. fib.1] (is.linv (fib.1) k); 89 | let sq2 (i k : I) : A := 90 | hcom 0 k 91 | [i=0. sq 1; 92 | i=1. is.linv a] 93 | (is.g (is.f a)); 94 | λ i. 95 | (sq2 i 1, 96 | λ j. 97 | let aux : A := 98 | hcom j 0 99 | [i=0. sq j; 100 | i=1. is.linv a; 101 | j=1. sq2 i] 102 | (is.linv (sq2 i 1) j); 103 | hcom 0 1 104 | [i=0. is.rinv (fib.2 j); 105 | i=1. is.rinv (is.f a); 106 | j=0. is.rinv (is.f (sq2 i 1)); 107 | j=1. is.rinv (is.f a)] 108 | (is.f aux)); 109 | 110 | isoToEquiv (A B : U) (is : iso A B) : equiv A B := 111 | is.f 112 | , is.g 113 | , λ a i. (contractIsoFiber A B is a (is.g (is.f a), is.rinv (is.f a)) i).1 114 | , is.rinv 115 | , λ a i. (contractIsoFiber A B is a (is.g (is.f a), is.rinv (is.f a)) i).2 116 | ; 117 | 118 | isContr (A : U) : U 119 | := (x : A) * ((y : A) → x = y); 120 | 121 | hasContrFibers (A B : U) (f : A → B) : U := (b : B) → isContr (fiber A B f b); 122 | 123 | connAnd (A : U)(a b : A)(p : a = b) : (λ _. a) ={i. a = p i} p 124 | := 125 | let sq (l k : I) : A := hcom 0 l [k=0 _. a; k=1. p] a; 126 | λ i j. hcom 0 1 [i=0 _. a; i=1. sq j; j=0 _. a; j=1. sq i] a; 127 | 128 | connAndRefl (A : U)(a : A) : connAnd A a a refl = refl 129 | := 130 | λ m. 131 | let sq (l k : I) : A := hcom 0 l [k=0 _. a; k=1 _. a; m=1 _. a] a; 132 | λ i j. hcom 0 1 [i=0 _. a; i=1. sq j; j=0 _. a; j=1. sq i; m=1 _. a] a; 133 | 134 | isEquivToContrFibers (A B : U) (f : A → B) (e : isEquiv A B f) : hasContrFibers A B f 135 | := 136 | λ b. 137 | ((e.g b, e.rinv b), 138 | λ fib i. 139 | let a : A := fib.1; 140 | let p : f a = b := fib.2; 141 | let domFiller (k : I) : A := 142 | hcom 0 k 143 | [i=0 k. e.g (p k); 144 | i=1 _. a] 145 | (e.linv a i); 146 | (domFiller 1, 147 | λ j. 148 | hcom 0 1 149 | [i=0 k. e.rinv (p k) j; 150 | i=1. connAnd B (f a) b p j; 151 | j=0 k. f (domFiller k); 152 | j=1. p] 153 | (e.coh a i j))); 154 | 155 | contrFibersToIsEquiv (A B : U) (f : A → B) (c : hasContrFibers A B f) : isEquiv A B f 156 | := 157 | λ b. (c b).1.1, 158 | λ a i. ((c (f a)).2 (a, refl) i).1, 159 | λ b. (c b).1.2, 160 | λ a i. ((c (f a)).2 (a, refl) i).2; 161 | 162 | isEquivRetractContrFibers (A B : U) (f : A → B) (e : isEquiv A B f) 163 | : contrFibersToIsEquiv A B f (isEquivToContrFibers A B f e) = e 164 | := 165 | λ l. 166 | let domFiller (a : A) (i : I) (k : I) : A := 167 | hcom 0 k 168 | [i=0 _. e.g (f a); 169 | i=1 _. a; 170 | l=1 _. e.linv a i] 171 | (e.linv a i); 172 | (e.g, 173 | λ a i. domFiller a i 1, 174 | e.rinv, 175 | λ a i j. 176 | hcom 0 1 177 | [i=0 _. e.rinv (f a) j; 178 | i=1. connAndRefl B (f a) l j; 179 | j=0 k. f (domFiller a i k); 180 | j=1 _. f a; 181 | l=1 _. e.coh a i j] 182 | (e.coh a i j)); 183 | 184 | -------------------------------------------------------------------------------- 185 | -------------------------------------------------------------------------------- /tests/bench/bool.cctt: -------------------------------------------------------------------------------- 1 | import basics; 2 | 3 | inductive Bool := false | true; 4 | 5 | falseNeqTrue : not (false = true) := 6 | λ p. coe 0 1 (i. case (p i) [false. Top; true. Bot]) tt; 7 | 8 | trueNeqFalse : not (true = false) := 9 | λ p. coe 0 1 (i. case (p i) [false. Bot; true. Top]) tt; 10 | 11 | boolDec : (b1 b2 : Bool) → Dec (b1 = b2) := λ[ 12 | false. λ[false. left refl ; true. right falseNeqTrue]; 13 | true. λ[false. right trueNeqFalse; true. left refl ] 14 | ]; 15 | 16 | negBool : Bool → Bool := λ[false. true; true. false]; 17 | 18 | negIso : iso Bool Bool := 19 | negBool 20 | , negBool 21 | , λ[false. refl; true. refl] 22 | , λ[false. refl; true. refl]; 23 | 24 | negPath : Bool = Bool := ua Bool Bool (isoToEquiv Bool Bool negIso); -------------------------------------------------------------------------------- /tests/bench/int.cctt: -------------------------------------------------------------------------------- 1 | 2 | import basics; 3 | import nat; 4 | 5 | inductive ℤ := pos ℕ | neg ℕ; 6 | 7 | forceℤ : ℤ → Top := λ[ 8 | pos n. forceℕ n; 9 | neg n. forceℕ n 10 | ]; 11 | 12 | zeroℤ := pos zero; 13 | 14 | predℤ : ℤ → ℤ := λ[ 15 | pos u. case u [zero. neg zero; suc n. pos n]; 16 | neg v. neg (suc v)]; 17 | 18 | pos' : ℕ → ℤ := λ n. pos n; 19 | neg' : ℕ → ℤ := λ n. neg n; 20 | 21 | sucℤ : ℤ → ℤ := λ[ 22 | pos u. pos (suc u); 23 | neg v. case v [zero. pos zero; suc n. neg n]]; 24 | 25 | predsucℤ : (x : ℤ) → predℤ (sucℤ x) = x := λ[ 26 | pos u. refl; 27 | neg v. case v (v. predℤ (sucℤ (neg v)) = neg v) [ 28 | zero. refl; 29 | suc n. refl] 30 | ]; 31 | 32 | sucpredℤ : (x : ℤ) → sucℤ (predℤ x) = x := λ[ 33 | pos u. case u (u. sucℤ (predℤ (pos u)) = pos u) [ 34 | zero. refl; 35 | suc n. refl]; 36 | neg v. refl 37 | ]; 38 | 39 | sucℤIso : iso ℤ ℤ := 40 | λ x. sucℤ x 41 | , λ x. predℤ x 42 | , predsucℤ 43 | , sucpredℤ; 44 | 45 | sucPathℤ : ℤ = ℤ := 46 | ua ℤ ℤ (isoToEquiv ℤ ℤ sucℤIso); -------------------------------------------------------------------------------- /tests/bench/nat.cctt: -------------------------------------------------------------------------------- 1 | 2 | import basics; 3 | 4 | inductive ℕ := zero | suc ℕ; 5 | 6 | forceℕ : ℕ → Top := λ[ 7 | zero. tt; 8 | suc n. forceℕ n 9 | ]; 10 | 11 | iter (A : U)(f : A → A)(a : A)(n : ℕ) : A := case n [ 12 | zero. a; 13 | suc n. f (iter A f a n)]; 14 | 15 | add (x y : ℕ) : ℕ := case x [ 16 | zero . y; 17 | suc x. suc (add x y)]; 18 | 19 | suc' (x : ℕ) := suc x; 20 | 21 | add0 : (x : ℕ) → add x zero = x := λ[ 22 | zero. refl; 23 | suc x. λ i. suc (add0 x i)]; 24 | 25 | addSuc : (x y : ℕ) → add x (suc y) = suc (add x y) := λ[ 26 | zero. λ y i. suc y; 27 | suc x. λ y i. suc (addSuc x y i)]; 28 | 29 | addAssoc : (x y z : ℕ) → add x (add y z) = add (add x y) z := λ[ 30 | zero. λ y z i. add y z; 31 | suc x. λ y z i. suc (addAssoc x y z i)]; 32 | 33 | addComm : (x y : ℕ) → add x y = add y x := λ[ 34 | zero. λ y. (add0 y)⁻¹; 35 | suc x. λ y. ap suc' (addComm x y) ∙ (addSuc y x)⁻¹]; 36 | 37 | mul (x y : ℕ) : ℕ := case x [ 38 | zero. zero; 39 | suc x. add y (mul x y)]; 40 | 41 | n1 := suc zero; 42 | n2 := suc n1; 43 | n3 := suc n2; 44 | n4 := suc n3; 45 | n5 := suc n4; 46 | n6 := suc n5; 47 | n7 := suc n6; 48 | n8 := suc n7; 49 | n9 := suc n8; 50 | n10 := suc n9; 51 | n11 := suc n10; 52 | n12 := suc n11; 53 | n13 := suc n12; 54 | n14 := suc n13; 55 | n15 := suc n14; 56 | n16 := suc n15; 57 | n17 := suc n16; 58 | n18 := suc n17; 59 | n19 := suc n18; 60 | n20 := suc n19; 61 | n30 := add n10 n20; 62 | n40 := add n10 n30; 63 | n50 := add n10 n40; 64 | n100 := mul n10 n10; 65 | n200 := add n100 n100; 66 | n250 := add n200 n50; 67 | n300 := add n100 n200; 68 | n400 := add n100 n300; 69 | n500 := add n100 n400; 70 | n750 := add n250 n500; 71 | n1k := mul n10 n100; 72 | n2k := add n1k n1k; 73 | n3k := add n1k n2k; 74 | n4k := add n1k n3k; 75 | n5k := add n1k n4k; 76 | n10k := mul n10 n1k; 77 | n100k := mul n10 n10k; 78 | n200k := add n100k n100k; 79 | n500k := mul n5 n100k; 80 | n1m := mul n1k n1k; 81 | n2m := add n1m n1m; 82 | n5m := mul n5 n1m; 83 | -------------------------------------------------------------------------------- /tests/bench/negPath.cctt: -------------------------------------------------------------------------------- 1 | 2 | import basics; 3 | import nat; 4 | import bool; 5 | 6 | BoolPathN (n : ℕ) := 7 | iter (Bool = Bool) (λ p. p ∙ negPath) negPath n; 8 | 9 | fun (n : ℕ) (x : Bool) : Bool := coe 0 1 (BoolPathN n) x; 10 | val (n : ℕ) : Bool := coe 0 1 (BoolPathN n) true; 11 | 12 | test := val n1m; -- 4.8 s 13 | -- test := val n1m; 14 | -- test : val n1m = false := refl; 15 | -------------------------------------------------------------------------------- /tests/bench/s1.cctt: -------------------------------------------------------------------------------- 1 | 2 | import basics; 3 | import nat; 4 | import int; 5 | 6 | higher inductive S¹ := 7 | base 8 | | loop (i : I)[i=0. base; i=1. base]; 9 | 10 | helix : S¹ → U := λ[ 11 | base. ℤ; 12 | loop i. sucPathℤ i 13 | ]; 14 | 15 | loop' : base = base := 16 | λ i. loop i; 17 | 18 | ΩS¹ : U := 19 | base = base; 20 | 21 | encode : (x : S¹) → base = x → helix x := 22 | λ x p. coe 0 1 (i. helix (p i)) (pos zero); 23 | 24 | winding : ΩS¹ → ℤ := 25 | encode base; 26 | 27 | intLoop : ℤ → ΩS¹ := λ[ 28 | pos n. case n [zero. refl; suc n. intLoop (pos n) ∙ loop']; 29 | neg n. case n [zero. loop'⁻¹; suc n. intLoop (neg n) ∙ loop'⁻¹] 30 | ]; 31 | 32 | intLoop' : ℤ → ΩS¹ := λ[ 33 | pos n. case n [zero. refl; suc n. loop' ∙ intLoop (pos n)]; 34 | neg n. case n [zero. loop'⁻¹; suc n. loop'⁻¹ ∙ intLoop (neg n)] 35 | ]; 36 | 37 | toℕ : ℤ → ℕ := λ[ 38 | pos n. n; 39 | neg n. n 40 | ]; 41 | 42 | test0 := forceℤ (winding (intLoop' (pos n1k))); -- 0.018 43 | test1 := forceℤ (winding (intLoop' (pos n1m))); -- 6.3 44 | test2 := forceℤ (winding (intLoop' (neg n100k))); -- 0.8 45 | test3 := winding (intLoop (pos n100k) ∙ intLoop (neg n100k)); --1.7 46 | -------------------------------------------------------------------------------- /tests/bool.cctt: -------------------------------------------------------------------------------- 1 | 2 | import basics; 3 | 4 | inductive Bool := false | true; 5 | 6 | falseNeqTrue : not (false = true) := 7 | λ p. coe 0 1 (i. case (p i) [false. Top; true. Bot]) tt; 8 | 9 | trueNeqFalse : not (true = false) := 10 | λ p. coe 0 1 (i. case (p i) [false. Bot; true. Top]) tt; 11 | 12 | boolDec : (b1 b2 : Bool) → Dec (b1 = b2) := λ[ 13 | false. λ[false. left refl ; true. right falseNeqTrue]; 14 | true. λ[false. right trueNeqFalse; true. left refl ] 15 | ]; 16 | 17 | negBool : Bool → Bool := λ[false. true; true. false]; 18 | 19 | negIso : iso Bool Bool := 20 | negBool 21 | , negBool 22 | , λ[false. refl; true. refl] 23 | , λ[false. refl; true. refl]; 24 | 25 | negEq : Bool = Bool := 26 | ua Bool Bool (isoToEquiv Bool Bool negIso); 27 | 28 | comp (A : U)(x y z : A)(p : x = y)(q : y = z) := p ∙ q; 29 | 30 | inductive Nat := zero | suc Nat; 31 | 32 | nsym (A : U)(x : A)(p : x = x)(n : Nat) : x = x := 33 | case n [zero. p; suc n. nsym A x p⁻¹ n]; 34 | 35 | n5 := suc (suc (suc (suc (suc zero)))); 36 | 37 | add (n m : Nat) : Nat := case n [ 38 | zero. m; 39 | suc n. suc (add n m)]; 40 | 41 | n10 := add n5 n5; 42 | n15 := add n10 n5; 43 | n20 := add n10 n10; 44 | 45 | ncomp (p : Bool = Bool) (n : Nat) : Bool = Bool := case n [ 46 | zero. p; 47 | suc n. p ∙ ncomp p n]; 48 | 49 | bigpath := negEq ∙ negEq ∙ negEq ∙ 50 | negEq ∙ negEq ∙ negEq ∙ 51 | negEq ∙ negEq ∙ negEq ∙ 52 | negEq ∙ negEq; 53 | 54 | test := coe 0 1 bigpath false; 55 | -------------------------------------------------------------------------------- /tests/entanglement.cctt: -------------------------------------------------------------------------------- 1 | -- "square induction" to derive splitting below in a braindead way 2 | isContr (A : U) : U := (fst : A) × ((y : A) → fst = y); 3 | 4 | sigPath (A : U) (B : A -> U) (u v : (x : A) * B x) 5 | (p : u.1 = v.1) (q : u.2 ={i. B (p i)} v.2) 6 | : u = v 7 | := λ i. (p i, q i); 8 | 9 | higher inductive interval := seg (i : I) []; 10 | 11 | intervalContr : isContr interval 12 | := (seg 0, λ [ seg i. λ j. hcom 1 i [j=0 _. seg 0; j=1 k. seg k] (seg j) ]); 13 | 14 | sqq : U := interval × interval; 15 | 16 | sqqContr : isContr sqq 17 | := ((seg 0, seg 0), λ x. sigPath interval (λ _. interval) (seg 0, seg 0) x (intervalContr.2 x.1) (intervalContr.2 x.2)); 18 | 19 | toSqq (A : U) (f : I → I → A) (x : sqq) : A := 20 | case x.1 [ seg i. case x.2 [ seg j. f i j ] ]; 21 | 22 | ofSqq (A : U) (f : sqq → A) (i j : I) : A := 23 | f (seg i, seg j); 24 | 25 | lemSqq (A : U) (f : sqq → A) : (λ _. f (seg 0, seg 0)) ={_. sqq → A} f 26 | := λ t x. f (sqqContr.2 x t); 27 | 28 | sqInd (A : U) (P : (I → I → A) → U) 29 | (p : I → I → A) (h : P (λ _ _. p 0 0)) : P p 30 | := coe 0 1 (t. P (ofSqq A (lemSqq A (toSqq A p) t))) h; 31 | 32 | 33 | 34 | -- Now, we will split a square along the main diagonal into 'left' and 35 | -- 'right' halves 36 | 37 | -- I apologize, but I think of squares like this: 38 | -- 39 | -- ---> i 40 | -- | 41 | -- v 42 | -- j 43 | 44 | left (A : U) (p : I → I → A) : p 0 0 = p 0 1 := λ j. p 0 j; 45 | leftInv (A : U) (p : I → I → A) : p 0 1 = p 0 0 := (left A p)⁻¹; 46 | down (A : U) (p : I → I → A) : p 0 1 = p 1 1 := λ i. p i 1; 47 | 48 | up (A : U) (p : I → I → A) : p 0 0 = p 1 0 := λ i. p i 0; 49 | right (A : U) (p : I → I → A) : p 1 0 = p 1 1 := λ j. p 1 j; 50 | rightInv (A : U) (p : I → I → A) : p 1 1 = p 1 0 := (right A p)⁻¹; 51 | 52 | diag (A : U) (p : I → I → A) : p 0 0 = p 1 1 := λ j. p j j; 53 | 54 | splitLeft (A : U) (p : I → I → A) : refl ={i. leftInv A p i = down A p i} diag A p := 55 | sqInd A (λ p. refl ={i. leftInv A p i = down A p i} diag A p) 56 | p 57 | (λ i j. hcom j 1 [i=0 _. p 0 0; i=1 _. p 0 0] (p 0 0)); 58 | 59 | splitRight (A : U) (p : I → I → A) : diag A p ={i. up A p i = rightInv A p i} refl := 60 | sqInd A (λ p. diag A p ={i. up A p i = rightInv A p i} refl) 61 | p 62 | (λ i j. hcom 0 j [i=0 _. p 0 0; i=1 _. p 0 0] (p 0 0)); 63 | 64 | 65 | 66 | -- Now, a weird example using S² 67 | 68 | higher inductive S² := 69 | base2 70 | | loop2 (i j : I) [i=0. base2 71 | ;i=1. base2 72 | ;j=0. base2 73 | ;j=1. base2]; 74 | 75 | ΩS² : U := base2 = base2; 76 | 77 | Ω²S² : U := refl ={_. ΩS²} refl; 78 | 79 | symRefl : ΩS² := (λ _. base2) ⁻¹; 80 | 81 | tweakRight : symRefl ={_. ΩS²} refl := 82 | λ i j. hcom i 1 [j=0 _. base2; j=1 _. base2] base2; 83 | 84 | tweakLeft : refl ={_. ΩS²} symRefl := 85 | λ i j. hcom 0 i [j=0 _. base2; j=1 _. base2] base2; 86 | 87 | weird1 : Ω²S² := 88 | λ i j. 89 | hcom 0 1 90 | [ i=0 k. base2 91 | ; i=1 k. base2 92 | ; j=0 k. base2 93 | ; j=1 k. loop2 i k 94 | ; i=j k. base2 95 | ] 96 | base2; 97 | 98 | weird1Left : symRefl ={_. ΩS²} refl := 99 | λ i j. splitLeft S² (λ i j. weird1 i j) j i; 100 | 101 | weird1Right : refl ={_. ΩS²} symRefl := 102 | λ i j. splitRight S² (λ i j. weird1 i j) j i; 103 | 104 | -- I've computed the elements below to integers elsewhere; todo 105 | -- compute them here 106 | 107 | -- this is +/- 1 108 | uhoh1Left : Ω²S² := tweakLeft ∙ weird1Left; 109 | -- this is 0 110 | uhoh1Right : Ω²S² := weird1Right ∙ tweakRight; 111 | 112 | loop2Inv : Ω²S² := (λ i j. loop2 i j) ⁻¹; 113 | 114 | -- this is propositionally equal to weird1 (todo, but I confirmed by 115 | -- computations elsewhere in cctt, and also in redtt) 116 | weird2 : Ω²S² := 117 | λ i j. 118 | hcom 0 1 119 | [ i=0 k. base2 120 | ; i=1 k. base2 121 | ; j=0 k. loop2Inv i k 122 | ; j=1 k. base2 123 | ; i=j k. base2 124 | ] 125 | base2; 126 | 127 | weird2Left : symRefl ={_. ΩS²} refl := 128 | λ i j. splitLeft S² (λ i j. weird2 i j) j i; 129 | 130 | weird2Right : refl ={_. ΩS²} symRefl := 131 | λ i j. splitRight S² (λ i j. weird2 i j) j i; 132 | 133 | -- this is 0 134 | uhoh2Left : Ω²S² := tweakLeft ∙ weird2Left; 135 | -- this is +/- 1 136 | uhoh2Right : Ω²S² := weird2Right ∙ tweakRight; 137 | 138 | -- So, that's weird! We have weird1 = weird2, yet we were able to 139 | -- distinguish them in some sense, determining which side of the 140 | -- diagonal the loop2 is on. For weird1, we can see it's on the 141 | -- "left", because uhoh1Left is nontrivial, but for weird2, it's on 142 | -- the "right", because uhoh1Right is nontrivial. 143 | 144 | -- But, it seems this does not lead to a contradiction, because we are 145 | -- not able to define a splitting map Ω²S² → Ω²S² × Ω²S² which would 146 | -- fail to be continuous -- we relied on the fact that the diagonals 147 | -- of the weird elements of Ω²S² are _judgmentally_ equal to refl... 148 | 149 | -- But... what if we had extension types? 150 | -------------------------------------------------------------------------------- /tests/hedberg.cctt: -------------------------------------------------------------------------------- 1 | 2 | import yaccttprelude; 3 | 4 | hedbergLemma (A : U)(a : A) (f : (x : A) → a = x → a = x) : 5 | (b : A) (p : a = b) → Square A a a a b refl p (f a refl) (f b p) := 6 | J A a 7 | (λ b p. Square A a a a b refl p (f a refl) (f b p)) 8 | refl; 9 | 10 | hedbergStable (A : U) (a b : A) (h : (x : A) -> stable (a = x))(p q : a = b) : p = q := 11 | 12 | let rem1 (x : A) : exConst (a = x) := stableConst (a = x) (h x); 13 | 14 | let f (x : A) : a = x → a = x := (rem1 x).1; 15 | 16 | let fIsConst (x : A) : const (a = x) (f x) := (rem1 x).2; 17 | 18 | let rem4 : Square A a a b b refl refl (f b p) (f b q) := fIsConst b p q; 19 | 20 | let r : a = a := f a refl; 21 | 22 | let rem2 (i j : I) := hedbergLemma A a f b p i j; 23 | let rem3 (i j : I) := hedbergLemma A a f b q i j; 24 | 25 | λ j i. hcom 0 1 [j=0. rem2 i 26 | ; j=1. rem3 i 27 | ; i=0. r 28 | ; i=1. rem4 j] a; 29 | 30 | hedbergS (A:U) (h : (a x : A) -> stable (a = x)) : set A := 31 | λ a b. hedbergStable A a b (h a); 32 | 33 | hedberg (A : U) (h : discrete A) : set A := 34 | λ a b. hedbergStable A a b (λ b. decStable (Path A a b) (h a b)); 35 | -------------------------------------------------------------------------------- /tests/importtest/Foo/Bar/bar.cctt: -------------------------------------------------------------------------------- 1 | 2 | module Foo.Bar.bar; 3 | 4 | import Foo.foo; 5 | import foo; -------------------------------------------------------------------------------- /tests/importtest/Foo/Bar/bar2.cctt: -------------------------------------------------------------------------------- 1 | 2 | module Foo.Bar.bar2; 3 | 4 | crab := U; -------------------------------------------------------------------------------- /tests/importtest/Foo/foo.cctt: -------------------------------------------------------------------------------- 1 | 2 | module Foo.foo; 3 | 4 | import Foo.Bar.bar2; -------------------------------------------------------------------------------- /tests/importtest/a.cctt: -------------------------------------------------------------------------------- 1 | 2 | module a; -------------------------------------------------------------------------------- /tests/importtest/foo.cctt: -------------------------------------------------------------------------------- 1 | 2 | module foo; 3 | 4 | import a; 5 | 6 | frog := U; -------------------------------------------------------------------------------- /tests/int.cctt: -------------------------------------------------------------------------------- 1 | 2 | import yaccttprelude; 3 | import hedberg; 4 | 5 | inductive nat := zero | suc nat; 6 | 7 | pred : nat → nat := λ[ 8 | zero. zero; 9 | suc n. n]; 10 | 11 | inductive Z := pos nat | neg nat; 12 | 13 | zeroZ := pos zero; 14 | 15 | predZ : Z → Z := λ[ 16 | pos u. case u [zero. neg zero; suc n. pos n]; 17 | neg v. neg (suc v)]; 18 | 19 | sucZ : Z → Z := λ[ 20 | pos u. pos (suc u); 21 | neg v. case v [zero. pos zero; suc n. neg n]]; 22 | 23 | predsucZ : (x : Z) → predZ (sucZ x) = x := λ[ 24 | pos u. refl; 25 | neg v. case v (v. predZ (sucZ (neg v)) = neg v) [ 26 | zero. refl; 27 | suc n. refl 28 | ] 29 | ]; 30 | 31 | sucpredZ : (x : Z) → sucZ (predZ x) = x := λ[ 32 | pos u. case u (u. sucZ (predZ (pos u)) = pos u) [ 33 | zero. refl; 34 | suc n. refl 35 | ]; 36 | neg v. refl 37 | ]; 38 | 39 | caseNat (A : U) (x y : A) : nat → A := λ[ 40 | zero. x; 41 | suc _. y]; 42 | 43 | caseDNat (P : nat → U) (z : P zero) (s : (n : nat) → P (suc n)) : (n : nat) → P n := λ[ 44 | zero. z; 45 | suc n. s n]; 46 | 47 | znots (n : nat) : not (zero = suc n) := 48 | λ h. subst nat (caseNat U nat N0) zero (suc n) h zero; 49 | 50 | discreteNat : discrete nat := λ[ 51 | zero. λ[zero. inl refl; 52 | suc m. inr (λ p. case (znots m p) [])]; 53 | suc n. λ[zero. inr (λ p. case (znots n p⁻¹) []); 54 | suc m. case (discreteNat n m) [ 55 | inl p. inl (ap (λ x. suc x) p); 56 | inr p. inr (λ q. case (p (ap pred q)) []) 57 | ]] 58 | ]; 59 | 60 | posNotneg (a b : nat) : not (pos a = neg b) := 61 | λ h. subst Z (λ[pos _. Unit; neg _. N0]) (pos a) (neg b) h tt; 62 | 63 | injPos (a b : nat) (h : pos a = pos b) : a = b := 64 | subst Z (λ[pos c. a = c; neg _. N0]) (pos a) (pos b) h refl; 65 | 66 | injNeg (a b : nat) (h : neg a = neg b) : a = b := 67 | subst Z (λ[pos _. N0; neg c. a = c]) (neg a) (neg b) h refl; 68 | 69 | discreteZ : discrete Z := λ[ 70 | pos a. λ[ 71 | pos a'. case (discreteNat a a') [ 72 | inl p. inl (ap (λ x. pos x) p); 73 | inr p. inr (λ q. p (injPos a a' q)) 74 | ]; 75 | neg b. inr (λ p. case (posNotneg a b p) []) 76 | ]; 77 | neg b. λ[ 78 | pos a. inr (λ p. case (posNotneg a b p⁻¹) []); 79 | neg b'. case (discreteNat b b') [ 80 | inl p. inl (ap (λ x. neg x) p); 81 | inr p. inr (λ q. p (injNeg b b' q)) 82 | ] 83 | ] 84 | ]; 85 | 86 | ZSet : set Z := hedberg Z discreteZ; 87 | 88 | 89 | 90 | -- ctr (y : Z) : fiber Z Z sucZ y := (predZ y, sucpredZ y); 91 | 92 | -- gengoal (A : U)(hA : set A) (x y : A) (p : y = x) : 93 | -- (z : A)(r : y = z)(q : z = x) → p ={i. r i = x} q := 94 | -- J A y (λ z r. (q : z = x) → p ={i. r i = x} q) 95 | -- (hA y x p); 96 | -------------------------------------------------------------------------------- /tests/tests/basics.cctt: -------------------------------------------------------------------------------- 1 | 2 | 3 | inductive Top := tt; 4 | inductive Bot :=; 5 | not (A : U) := A → Bot; 6 | 7 | inductive Sum (A B : U) := left A | right B; 8 | 9 | Dec (A : U) := Sum A (not A); 10 | 11 | -------------------------------------------------------------------------------- 12 | 13 | isLogEquiv (A B : U)(f : A → B) : U := (B → A); 14 | 15 | isSection (A B : U)(f : A → B) : U := (g : B → A) × (linv : (x : A) → x = g (f x)); 16 | 17 | isRetraction (A B : U)(f : A → B) : U := (g : B → A) × (rinv : (x : B) → f (g x) = x); 18 | 19 | isIso (A B : U)(f : A → B) : U := 20 | (g : B → A) 21 | × (linv : (x : A) → g (f x) = x) 22 | × (rinv : (x : B) → f (g x) = x); 23 | 24 | isEquiv (A B : U)(f : A → B) : U := 25 | (g : B → A) 26 | × (linv : (x : A) → g (f x) = x) 27 | × (rinv : (x : B) → f (g x) = x) 28 | × (coh : (x : A) → (rinv (f x)) ={i. f (linv x i) = f x} refl); 29 | 30 | coeIsEquiv (A : I → U)(r r' : I) : isEquiv (A r) (A r') (λ x. coe r r' A x) := 31 | let f' (i : I)(x : A r) : A i := coe r i A x; 32 | let g' (i : I)(x : A i) : A r := coe i r A x; 33 | let linv' (i : I)(x : A r)(j : I) : A r := hcom r i (A r) [j=0 k. coe k r A (coe r k A x); j=1 k. x] x; 34 | let rinv' (i : I)(x : A i)(j : I) : A i := hcom i r (A i) [j=0 k. coe k i A (coe i k A x); j=1 k. x] x; 35 | 36 | ( λ x. g' r' x 37 | , λ x j. linv' r' x j 38 | , λ x j. rinv' r' x j 39 | , λ x l k. com r r' A 40 | [k=0 i. f' i (linv' i x l) 41 | ;k=1 i. f' i x 42 | ;l=0 i. rinv' i (f' i x) k 43 | ;l=1 i. f' i x] 44 | x 45 | ); 46 | 47 | equiv (A B : U) : U := 48 | (f : A → B) × isEquiv A B f; 49 | 50 | iso (A B : U) : U := (f : A → B) × isIso A B f; 51 | 52 | idEquiv (A : U) : equiv A A := 53 | (λ x. x, λ x. x, λ x _. x, λ x _. x, λ x _ _. x); 54 | 55 | ua (A B : U)(f : equiv A B) : A = B := 56 | λ i. Glue B [i=0. (A, f); i=1. (B, idEquiv B)]; 57 | 58 | J (A : U) (a : A) (C : (x : A) → a = x → U) 59 | (d : C a (λ _. a))(x : A) (p : a = x) : C x p 60 | := let sq (i j : I) : A := hcom 0 j [i=0 _. a; i=1. p] a; 61 | coe 0 1 (i. C (sq i 1) (λ j. sq i j)) d; 62 | 63 | JEq (A : U) (a : A) (C : (x : A) → a = x → U) (d : C a (λ _. a)) 64 | : J A a C d a (λ _. a) = d := 65 | 66 | let sq (i j : I) : A := hcom 0 j [i=0 _. a; i=1 _. a] a; 67 | let cube (i j k : I) : A := hcom 0 j [k=0 x. sq i x; 68 | k=1 _. a; 69 | i=0 _. a; 70 | i=1 _. a] a; 71 | 72 | λ k. com 0 1 (i. C (cube i 1 k) (λ j. cube i j k)) 73 | [k=0 i. coe 0 i (i. C (sq i 1) (λ j. sq i j)) d; 74 | k=1 _. d] 75 | d; 76 | 77 | fiber (A B : U) (f : A → B) (b : B) : U 78 | := (x : A) * f x = b; 79 | 80 | fiberRefl (A B : U) (f : A → B) (a : A) : fiber A B f (f a) 81 | := (a , refl); 82 | 83 | contractIsoFiber (A B : U) (is : iso A B) (a : A) 84 | (fib : fiber A B is.f (is.f a)) 85 | : fib = fiberRefl A B is.f a 86 | := 87 | let sq (j k : I) : A := 88 | hcom k j [k=0 j. is.g (fib.2 j); k=1 _. fib.1] (is.linv (fib.1) k); 89 | let sq2 (i k : I) : A := 90 | hcom 0 k 91 | [i=0. sq 1; 92 | i=1. is.linv a] 93 | (is.g (is.f a)); 94 | λ i. 95 | (sq2 i 1, 96 | λ j. 97 | let aux : A := 98 | hcom j 0 99 | [i=0. sq j; 100 | i=1. is.linv a; 101 | j=1. sq2 i] 102 | (is.linv (sq2 i 1) j); 103 | hcom 0 1 104 | [i=0. is.rinv (fib.2 j); 105 | i=1. is.rinv (is.f a); 106 | j=0. is.rinv (is.f (sq2 i 1)); 107 | j=1. is.rinv (is.f a)] 108 | (is.f aux)); 109 | 110 | isoToEquiv (A B : U) (is : iso A B) : equiv A B := 111 | is.f 112 | , is.g 113 | , λ a i. (contractIsoFiber A B is a (is.g (is.f a), is.rinv (is.f a)) i).1 114 | , is.rinv 115 | , λ a i. (contractIsoFiber A B is a (is.g (is.f a), is.rinv (is.f a)) i).2 116 | ; 117 | 118 | isContr (A : U) : U 119 | := (x : A) * ((y : A) → x = y); 120 | 121 | hasContrFibers (A B : U) (f : A → B) : U := (b : B) → isContr (fiber A B f b); 122 | 123 | connAnd (A : U)(a b : A)(p : a = b) : (λ _. a) ={i. a = p i} p 124 | := 125 | let sq (l k : I) : A := hcom 0 l [k=0 _. a; k=1. p] a; 126 | λ i j. hcom 0 1 [i=0 _. a; i=1. sq j; j=0 _. a; j=1. sq i] a; 127 | 128 | connAndRefl (A : U)(a : A) : connAnd A a a refl = refl 129 | := 130 | λ m. 131 | let sq (l k : I) : A := hcom 0 l [k=0 _. a; k=1 _. a; m=1 _. a] a; 132 | λ i j. hcom 0 1 [i=0 _. a; i=1. sq j; j=0 _. a; j=1. sq i; m=1 _. a] a; 133 | 134 | isEquivToContrFibers (A B : U) (f : A → B) (e : isEquiv A B f) : hasContrFibers A B f 135 | := 136 | λ b. 137 | ((e.g b, e.rinv b), 138 | λ fib i. 139 | let a : A := fib.1; 140 | let p : f a = b := fib.2; 141 | let domFiller (k : I) : A := 142 | hcom 0 k 143 | [i=0 k. e.g (p k); 144 | i=1 _. a] 145 | (e.linv a i); 146 | (domFiller 1, 147 | λ j. 148 | hcom 0 1 149 | [i=0 k. e.rinv (p k) j; 150 | i=1. connAnd B (f a) b p j; 151 | j=0 k. f (domFiller k); 152 | j=1. p] 153 | (e.coh a i j))); 154 | 155 | contrFibersToIsEquiv (A B : U) (f : A → B) (c : hasContrFibers A B f) : isEquiv A B f 156 | := 157 | λ b. (c b).1.1, 158 | λ a i. ((c (f a)).2 (a, refl) i).1, 159 | λ b. (c b).1.2, 160 | λ a i. ((c (f a)).2 (a, refl) i).2; 161 | 162 | isEquivRetractContrFibers (A B : U) (f : A → B) (e : isEquiv A B f) 163 | : contrFibersToIsEquiv A B f (isEquivToContrFibers A B f e) = e 164 | := 165 | λ l. 166 | let domFiller (a : A) (i : I) (k : I) : A := 167 | hcom 0 k 168 | [i=0 _. e.g (f a); 169 | i=1 _. a; 170 | l=1 _. e.linv a i] 171 | (e.linv a i); 172 | (e.g, 173 | λ a i. domFiller a i 1, 174 | e.rinv, 175 | λ a i j. 176 | hcom 0 1 177 | [i=0 _. e.rinv (f a) j; 178 | i=1. connAndRefl B (f a) l j; 179 | j=0 k. f (domFiller a i k); 180 | j=1 _. f a; 181 | l=1 _. e.coh a i j] 182 | (e.coh a i j)); 183 | 184 | -------------------------------------------------------------------------------- 185 | -------------------------------------------------------------------------------- /tests/tests/bool.cctt: -------------------------------------------------------------------------------- 1 | import basics; 2 | 3 | inductive Bool := false | true; 4 | 5 | falseNeqTrue : not (false = true) := 6 | λ p. coe 0 1 (i. case (p i) [false. Top; true. Bot]) tt; 7 | 8 | trueNeqFalse : not (true = false) := 9 | λ p. coe 0 1 (i. case (p i) [false. Bot; true. Top]) tt; 10 | 11 | boolDec : (b1 b2 : Bool) → Dec (b1 = b2) := λ[ 12 | false. λ[false. left refl ; true. right falseNeqTrue]; 13 | true. λ[false. right trueNeqFalse; true. left refl ] 14 | ]; 15 | 16 | negBool : Bool → Bool := λ[false. true; true. false]; 17 | 18 | negIso : iso Bool Bool := 19 | negBool 20 | , negBool 21 | , λ[false. refl; true. refl] 22 | , λ[false. refl; true. refl]; 23 | 24 | negPath : Bool = Bool := ua Bool Bool (isoToEquiv Bool Bool negIso); -------------------------------------------------------------------------------- /tests/tests/hittests.cctt: -------------------------------------------------------------------------------- 1 | 2 | import basics; 3 | import bool; 4 | 5 | higher inductive S¹ := 6 | base 7 | | loop (i : I) [i=0. base; i=1. base]; 8 | 9 | loop' : base = base := λ i. loop i; 10 | 11 | test1 := base; 12 | 13 | test2 (i : I) : S¹ := loop i; 14 | 15 | test3 : loop 0 = base := refl; 16 | test4 : loop 1 = base := refl; 17 | 18 | f1 (x : S¹)(A : U)(p : A = A) : U := case x [ 19 | base. A; 20 | loop i. p i]; 21 | 22 | F : S¹ → U := λ[ 23 | base. Bool; 24 | loop i. negPath i]; 25 | 26 | loopNeqRefl (p : loop' = refl) : Bot := 27 | let lem1 : ap F loop' = refl := ap (λ q. ap F q) p; 28 | let lem2 : false = true := ap (λ p. coe 0 1 p true) lem1; 29 | falseNeqTrue lem2; 30 | 31 | higher inductive List (A : U) := 32 | nil 33 | | cons A (List A); 34 | 35 | test5 (A : I → U)(i j : I)(xs : List (A i)) : 36 | coe i j (k. List (A k)) nil = nil := 37 | refl; 38 | 39 | test6 (A : I → U)(i j : I)(x : A i)(xs : List (A i)) : 40 | coe i j (k. List (A k)) (cons x xs) = cons (coe i j A x) (coe i j (k. List (A k)) xs) := 41 | refl; 42 | 43 | higher inductive Pushout (A B C : U)(f : A → B)(g : A → C) := 44 | inl B 45 | | inr C 46 | | quot (a : A)(i : I) [i=0. inl (f a); i=1. inr (g a)]; 47 | 48 | test7 (A B C : U)(f : A → B)(g : A → C) (a : A)(i : I) : Pushout A B C f g 49 | := quot a i; 50 | 51 | higher inductive PropTrunc (A : U) := 52 | pt-embed A 53 | | pt-quot (x y : PropTrunc A)(i : I)[i=0. x; i=1. y]; 54 | 55 | test8 (A : U)(x : A)(p : pt-embed x ={_. PropTrunc A} pt-embed x) 56 | := p⁻¹; 57 | 58 | test9 (A : I → U)(a : A 0) := coe 0 1 (i. PropTrunc (A i)) (pt-embed a); 59 | 60 | test10 (A : I → U)(x y : PropTrunc (A 0))(j : I) : PropTrunc (A 1) := 61 | coe 0 1 (i. PropTrunc (A i)) (pt-quot x y j); 62 | 63 | test10Nf (A : I → U)(x y : PropTrunc (A 0))(j : I) : PropTrunc (A 1) := 64 | pt-quot (coe 0 1 (i. PropTrunc (A i)) x) (coe 0 1 (i. PropTrunc (A i)) y) j; 65 | 66 | test10OK : test10 = test10Nf := refl; 67 | 68 | higher inductive SetTrunc (A : U) := 69 | st-embed A 70 | | st-quot (x y : SetTrunc A)(p q : x = y)(i j : I) 71 | [i=0. p j; i=1. q j; j=0. x; j=1. y]; 72 | 73 | test11 (A : I → U)(x y : SetTrunc (A 0))(p q : x = y)(j k : I) 74 | : SetTrunc (A 1) := 75 | coe 0 1 (i. SetTrunc (A i)) (st-quot x y p q j k); 76 | 77 | test11Nf (A : I → U)(x y : SetTrunc (A 0))(p q : x = y)(j k : I) : SetTrunc (A 1) := 78 | hcom 1 0 (SetTrunc (A 1)) [j = 0 i. hcom 0 i (SetTrunc (A 1)) 79 | [k = 0 i. coe @2 1 (i. SetTrunc (A i)) (coe i @2 (i. SetTrunc (A i)) (coe 0 i 80 | (i. SetTrunc (A i)) x)); k = 1 i. coe @2 1 (i. SetTrunc (A i)) (coe i @2 81 | (i. SetTrunc (A i)) (coe 0 i (i. SetTrunc (A i)) y))] (coe i 1 (i. SetTrunc (A 82 | i)) (coe 0 i (i. SetTrunc (A i)) (p {x} {y} k))); j = 1 i. hcom 0 i (SetTrunc (A 83 | 1)) [k = 0 i. coe @2 1 (i. SetTrunc (A i)) (coe i @2 (i. SetTrunc (A i)) (coe 0 84 | i (i. SetTrunc (A i)) x)); k = 1 i. coe @2 1 (i. SetTrunc (A i)) (coe i @2 85 | (i. SetTrunc (A i)) (coe 0 i (i. SetTrunc (A i)) y))] (coe i 1 (i. SetTrunc (A 86 | i)) (coe 0 i (i. SetTrunc (A i)) (q {x} {y} k))); k = 0 i. coe i 1 (i. SetTrunc 87 | (A i)) (coe 0 i (i. SetTrunc (A i)) x); k = 1 i. coe i 1 (i. SetTrunc (A i)) 88 | (coe 0 i (i. SetTrunc (A i)) y)] (st-quot (coe 0 1 (i. SetTrunc (A i)) x) (coe 0 89 | 1 (i. SetTrunc (A i)) y) (λ {coe 0 1 (i. SetTrunc (A i)) x} {coe 0 1 90 | (i. SetTrunc (A i)) y} _. hcom 0 1 (SetTrunc (A 1)) [@2 = 0 i. coe i 1 91 | (i. SetTrunc (A i)) (coe 0 i (i. SetTrunc (A i)) x); @2 = 1 i. coe i 1 92 | (i. SetTrunc (A i)) (coe 0 i (i. SetTrunc (A i)) y)] (coe 0 1 (i. SetTrunc (A 93 | i)) (p {x} {y} @2))) (λ {coe 0 1 (i. SetTrunc (A i)) x} {coe 0 1 (i. SetTrunc (A 94 | i)) y} _. hcom 0 1 (SetTrunc (A 1)) [@2 = 0 i. coe i 1 (i. SetTrunc (A i)) (coe 95 | 0 i (i. SetTrunc (A i)) x); @2 = 1 i. coe i 1 (i. SetTrunc (A i)) (coe 0 i 96 | (i. SetTrunc (A i)) y)] (coe 0 1 (i. SetTrunc (A i)) (q {x} {y} @2))) j k); 97 | 98 | test11OK : test11 = test11Nf := refl; 99 | 100 | 101 | -- Simple incoherent HIT examples 102 | 103 | higher inductive CohTest (b : Bool) := 104 | ct-quot (f : I → CohTest b)(i : I) [i=0. f 0; i=1. f 1]; 105 | 106 | test12 (b : I → Bool)(f : I → CohTest (b 0))(i : I) : CohTest (b 1) := 107 | coe 0 1 (i. CohTest (b i)) (ct-quot f i); 108 | 109 | test12NF (b : I → Bool)(f : I → CohTest (b 0))(i : I) : CohTest (b 1) := 110 | hcom 1 0 (CohTest (b 1)) [i = 0 i. coe i 1 (i. CohTest (b i)) (coe 0 i (i. CohTest (b i)) (f 0)); i = 1 i. coe i 1 (i. CohTest (b i)) (coe 0 i (i. CohTest (b i)) (f 1))] (ct-quot (λ _. coe 0 1 (i. CohTest (b i)) (f @1)) i); 111 | 112 | test12NFOK : test12 = test12NF := refl; 113 | 114 | -------------------------------------------------------------------------------- 115 | 116 | higher inductive CohTest2 (A : U) := 117 | ct2-quot (f : A → CohTest2 A)(x : A)(i : I)[i=0. f x]; 118 | 119 | test13 (A : I → U)(f : A 0 → CohTest2 (A 0))(x : A 0)(i : I) := 120 | coe 0 1 (i. CohTest2 (A i)) (ct2-quot f x i); 121 | 122 | test13NF (A : I → U)(f : A 0 → CohTest2 (A 0))(x : A 0)(i : I) : CohTest2 (A 1) := 123 | hcom 1 0 (CohTest2 (A 1)) [i = 0 i. coe i 1 (i. CohTest2 (A i)) (coe 124 | 0 i (j. CohTest2 (A j)) (f (coe i 0 A (coe 0 i A x))))] (ct2-quot (λ _. coe 0 1 125 | (j. CohTest2 (A j)) (f (coe 1 0 A @3))) (coe 0 1 A x) i); 126 | 127 | test13OK : test13 = test13NF := refl; 128 | 129 | 130 | -- U-Quotient 131 | -------------------------------------------------------------------------------- 132 | 133 | higher inductive Quot (A : U)(R : A → A → U) := 134 | q-emb A 135 | | q-quot (x y : A) (R x y) (i : I) [i=0. q-emb x; i=1. q-emb y]; 136 | 137 | test14 (A : I → U)(R : (i : I) → A i → A i → U) 138 | (x y : A 0)(r : R 0 x y)(i : I) : Quot (A 1) (R 1) := 139 | coe 0 1 (i. Quot (A i) (R i)) (q-quot x y r i); 140 | -- OK (coherent) 141 | 142 | test14NF (A : I → U)(R : (i : I) → A i → A i → U) 143 | (x y : A 0)(r : R 0 x y)(i : I) : Quot (A 1) (R 1) := 144 | q-quot (coe 0 1 A x) (coe 0 1 A y) (coe 0 1 (i. R i (coe 0 i A x) (coe 0 i A y)) r) i; 145 | 146 | test14OK : test14 = test14NF := refl; 147 | 148 | -- TODO: more precise coherence checking! 149 | -- For example: applications to vars with closed inductive types are OK 150 | -- How to do it more generally? 151 | 152 | -------------------------------------------------------------------------------- 153 | 154 | higher inductive Coeq (A B : U) (f g : A → B) := 155 | coeq-emb B 156 | | coeq-quot (a : A)(i : I)[i=0. coeq-emb (f a); i=1. coeq-emb (g a)]; 157 | 158 | test15 (A B : I → U)(f g : (i : I) → A i → B i)(a : A 0)(i : I) : Coeq (A 1) (B 1) (f 1) (g 1) := 159 | coe 0 1 (i. Coeq (A i) (B i) (f i) (g i)) (coeq-quot a i); 160 | 161 | test15NF (A B : I → U)(f g : (i : I) → A i → B i)(a : A 0)(i : I) : Coeq (A 1) (B 1) (f 1) (g 1) := 162 | hcom 1 0 (Coeq (A 1) (B 1) (f 1) (g 1)) [i = 0 i. coeq-emb (coe i 1 B (f i (coe 163 | 0 i A a))); i = 1 i. coeq-emb (coe i 1 B (g i (coe 0 i A a)))] (coeq-quot (coe 0 164 | 1 A a) i); 165 | 166 | test15NFOK : test15 = test15NF := refl; 167 | -------------------------------------------------------------------------------- /tests/tests/indtests.cctt: -------------------------------------------------------------------------------- 1 | 2 | 3 | -------------------------------------------------------------------------------- 4 | 5 | inductive Foo := R | G | B; 6 | foo : Foo := G; 7 | 8 | inductive ℕ := zero | suc ℕ; 9 | 10 | czero (n : ℕ) : ℕ := 11 | case n (_. ℕ) [ 12 | zero. zero; 13 | suc n. czero n]; 14 | 15 | add : ℕ → ℕ → ℕ := λ[ 16 | zero. λ x. x; 17 | suc n. λ x. suc (add n x)]; 18 | 19 | five := suc (suc (suc (suc (suc zero)))); 20 | 21 | ten := add five five; 22 | 23 | mul (n m : ℕ) : ℕ := 24 | case n [zero. zero; suc n. add m (mul n m)]; 25 | 26 | mul' : ℕ → ℕ → ℕ := λ[ 27 | zero. λ _. zero; 28 | suc n. λ m. add m (mul' n m)]; 29 | 30 | n100 := mul ten ten; 31 | n1k := mul' ten n100; 32 | n100k := mul n1k n100; 33 | n10m := mul n100 n100k; 34 | 35 | inductive Bool := true | false; 36 | 37 | even : ℕ → Bool := λ[ 38 | zero. true; 39 | suc n. case (even n) [true. false; false. true]]; 40 | 41 | testn1k := even n1k; 42 | testn100k := even n100k; 43 | testn10m := even n10m; 44 | 45 | -------------------------------------------------------------------------------- 46 | 47 | inductive List (A : U) := nil | cons A (List A); 48 | 49 | map (A B : U) (f : A → B)(as : List A) : List B := 50 | case as (_. List B) [ 51 | nil. nil; 52 | cons a as. cons (f a) (map A B f as)]; 53 | 54 | foldr (A B : U)(f : A → B → B)(b : B)(as : List A) : B := 55 | case as [nil. b; 56 | cons a as. f a (foldr A B f b as)]; 57 | 58 | foldr' : (A B : U) → (A → B → B) → B → List A → B := 59 | λ A B f b. λ[nil. b; cons a as. f a (foldr' A B f b as)]; 60 | 61 | list1 : List ℕ := cons ten (cons ten (cons ten nil)); 62 | list2 := map ℕ ℕ (λ _. zero) list1; 63 | 64 | -------------------------------------------------------------------------------- 65 | 66 | -- No positivity check. 67 | inductive Bad := mkBad (Bad → Bad); 68 | 69 | inductive Top := tt; 70 | 71 | inductive Bot :=; 72 | 73 | -- makes more sense to just write "case p []" inline. 74 | Bot-elim (A : U)(p : Bot) : A := case p []; 75 | 76 | inductive W (A : U)(B : A → U) := sup (a : A)(f : B a → W A B); 77 | 78 | inductive O := ozero | osuc O | olim (ℕ → O); 79 | 80 | fromℕ : ℕ → O := λ[ 81 | zero. ozero; 82 | suc n. osuc (fromℕ n)]; 83 | 84 | oten := fromℕ ten; 85 | 86 | ω := olim fromℕ; 87 | 88 | iter (A : U)(n : ℕ) (f : A → A) (a : A) : A := case n (_. A) [ 89 | zero. a; 90 | suc n. f (iter A n f a)]; 91 | 92 | -- fast-growing function 93 | fgf (o : O) (n : ℕ) : ℕ := case o [ 94 | ozero . suc n; 95 | osuc o. iter ℕ n (fgf o) n; 96 | olim f. fgf (f n) n]; 97 | 98 | fω2 := fgf ω (suc (suc zero)); 99 | 100 | -------------------------------------------------------------------------------- 101 | 102 | -- non-uniform parameters are allowed 103 | inductive Vec (A : U) (n : ℕ) := 104 | vnil (n = zero) 105 | | vcons (m : ℕ) (n = suc m) A (Vec A m) ; 106 | 107 | vmap (A B : U)(n : ℕ)(f : A → B)(as : Vec A n) : Vec B n := 108 | case as (_. Vec B n) [ 109 | vnil p. vnil p; 110 | vcons m p a as. vcons m p (f a) (vmap A B m f as)]; 111 | 112 | v1 : Vec ℕ (suc (suc zero)) := 113 | vcons (suc zero) refl ten (vcons zero refl ten (vnil refl)); 114 | 115 | v2 := vmap ℕ ℕ (suc (suc zero)) (add ten) v1; 116 | 117 | 118 | -- Recursive identities (will be needed in HITs) 119 | -------------------------------------------------------------------------------- 120 | 121 | inductive RecId := 122 | Base 123 | | Path (Base = Base); -- later constructors can refer to earlier ones 124 | 125 | recid1 := Path refl⁻¹; 126 | 127 | -------------------------------------------------------------------------------- 128 | 129 | not (A : U) := A → Bot; 130 | 131 | inductive Sum (A B : U) := Left A | Right B; 132 | Dec (A : U) := Sum A (not A); 133 | 134 | the (A : U)(x : A) := x; 135 | 136 | ℕ-Elim : (P : ℕ → U)(pz : P zero)(ps : (n : ℕ) → P n → P (suc n))(n : ℕ) → P n := 137 | λ P pz ps. λ[ 138 | zero. pz; 139 | suc n. ps n (ℕ-Elim P pz ps n)]; 140 | 141 | zeroNotSuc (n : ℕ)(p : zero = suc n) : Bot := 142 | let f : ℕ → U := λ[zero. Top; suc _. Bot]; 143 | coe 0 1 (ap f p) tt; 144 | 145 | pred : ℕ → ℕ := λ[zero. zero; suc n. n]; 146 | 147 | ℕ-Dec : (n m : ℕ) → Dec (n = m) := 148 | λ [zero. λ [zero. Left refl; 149 | suc m. Right (zeroNotSuc m)]; 150 | suc n. λ [zero. Right (λ p. zeroNotSuc n p⁻¹); 151 | suc m. case (ℕ-Dec n m) (_. Dec (suc n = suc m)) [ 152 | Left p. Left (λ i. suc (p i)); 153 | Right p. Right (λ q. p (ap pred q))]]]; 154 | 155 | test := ℕ-Dec (suc zero) zero; 156 | 157 | ℕ-DecNF : (n m : ℕ) → Dec (n = m) := 158 | λ n. case n (n. (m : ℕ) → Sum (n ={_. ℕ} m) (n ={_. ℕ} m → Bot)) [zero. λ 159 | m. case m (m. Sum (zero ={_. ℕ} m) (zero ={_. ℕ} m → Bot)) [zero. Left (λ {zero} 160 | {zero} _. zero); suc m. Right (λ p. coe 0 1 (i. case (p {zero} {suc m} i) (_. U) 161 | [zero. Top; suc _. Bot]) tt)]; suc n. λ m. case m (m. Sum (suc n ={_. ℕ} m) (suc 162 | n ={_. ℕ} m → Bot)) [zero. Right (λ p. coe 0 1 (i. case (hcom 0 1 ℕ [i = 0 j. p 163 | {suc n} {zero} j; i = 1 _. suc n] (suc n)) (_. U) [zero. Top; suc _. Bot]) tt); 164 | suc m. case (ℕ-Dec n m) (_. Sum (suc n ={_. ℕ} suc m) (suc n ={_. ℕ} suc m → 165 | Bot)) [Left p. Left (λ {suc n} {suc m} i. suc (p {n} {m} i)); Right p. Right (λ 166 | q. p (λ {n} {m} i. case (q {suc n} {suc m} i) (_. ℕ) [zero. zero; suc 167 | n. n]))]]]; 168 | 169 | -------------------------------------------------------------------------------- 170 | -------------------------------------------------------------------------------- /tests/tests/indtests2.cctt: -------------------------------------------------------------------------------- 1 | 2 | inductive ℕ := zero | suc ℕ; 3 | 4 | nf1 := hcom 0 1 [] zero; 5 | nf2 := hcom 0 1 [] (suc zero); 6 | nf3 := coe 0 1 (_. ℕ) (suc zero); 7 | nf4 := hcom 0 1 [] (suc (suc (suc zero))); 8 | 9 | nf5 (i j : I) (x : ℕ) := hcom i j [] (suc (suc (suc x))); 10 | nf6 (i j k : I) (x : ℕ) := hcom i j [k = 0 _. suc (suc (suc x))] (suc (suc (suc x))); 11 | nf7 (i j k : I) := hcom i j [k = 0 _. suc (suc (suc zero))] (suc (suc (suc zero))); 12 | 13 | t1 (r r' : I)(x : ℕ) : coe r r' (_. ℕ) x = x := refl; 14 | 15 | t2 (r r' : I) : hcom r r' [] zero = zero := refl; 16 | 17 | t3 : hcom 0 1 [] zero = zero := refl; 18 | 19 | t4 (r r' : I) (x : ℕ) : hcom r r' [] (suc x) = suc (hcom r r' [] x) := refl; 20 | 21 | t5 (x : ℕ) : hcom 0 1 [] (suc x) = suc (hcom 0 1 [] x) := refl; 22 | 23 | t6 (r r' k : I) (x : ℕ) : hcom r r' [k = 0 j. suc x] (suc x) = suc (hcom r r' [k = 0 _. x] x) := refl; 24 | 25 | -------------------------------------------------------------------------------- 26 | 27 | 28 | inductive Unit := tt; 29 | inductive Fam (x : Unit) := fam; 30 | 31 | inductive Sg (A : U)(B : A → U) := pair (a : A) (b : B a); 32 | 33 | t7 (A : I → U)(B : (i : I) → A i → U)(a : A 0)(b : B 0 a) : 34 | coe 0 1 (i. Sg (A i) (B i)) (pair a b) 35 | = pair (coe 0 1 A a) (coe 0 1 (i. B i (coe 0 i A a)) b) := refl; 36 | 37 | nf8 (A : I → U)(B : (i : I) → A i → U)(a : A 0)(b : B 0 a) := 38 | coe 0 1 (i. Sg (A i) (B i)) (pair a b); 39 | 40 | t8 (A : I → U)(B : (i : I) → A i → U)(a : A 0)(b : B 0 a) : 41 | coe 0 1 (i. Sg (A i) (B i)) (pair a b) 42 | = pair (coe 0 1 A a) (coe 0 1 (i. B i (coe 0 i A a)) b) := refl; 43 | 44 | nf9 (A : U)(B : A → U) (a : I → A)(b : (i : I) → B (a i))(i : I) := 45 | hcom 0 1 (Sg A B) [i = 0 j. pair (a j) (b j)] (pair (a 0) (b 0)); 46 | 47 | nf10 48 | (A : U)(B : A → U) (a : I → A)(b : (i : I) → B (a i))(i : I) 49 | (p : pair (a 0) (b 0) ={_. Sg A B} pair (a 1) (b 1)) := 50 | hcom 0 1 (Sg A B) [i = 0. p] (pair (a 0) (b 0)); 51 | 52 | nf11 (i : I) := hcom 0 1 (Sg Unit (λ x. Fam x)) [i = 0 j. pair tt fam] (pair tt fam); 53 | 54 | -------------------------------------------------------------------------------- 55 | -------------------------------------------------------------------------------- /tests/tests/nat.cctt: -------------------------------------------------------------------------------- 1 | 2 | inductive ℕ := zero | suc ℕ; 3 | 4 | iter (A : U)(f : A → A)(a : A)(n : ℕ) : A := case n [ 5 | zero. a; 6 | suc n. f (iter A f a n)]; 7 | 8 | add (x y : ℕ) : ℕ := case x [ 9 | zero . y; 10 | suc x. suc (add x y)]; 11 | 12 | suc' (x : ℕ) := suc x; 13 | 14 | add0 : (x : ℕ) → add x zero = x := λ[ 15 | zero. refl; 16 | suc x. λ i. suc (add0 x i)]; 17 | 18 | addSuc : (x y : ℕ) → add x (suc y) = suc (add x y) := λ[ 19 | zero. λ y i. suc y; 20 | suc x. λ y i. suc (addSuc x y i)]; 21 | 22 | addAssoc : (x y z : ℕ) → add x (add y z) = add (add x y) z := λ[ 23 | zero. λ y z i. add y z; 24 | suc x. λ y z i. suc (addAssoc x y z i)]; 25 | 26 | addComm : (x y : ℕ) → add x y = add y x := λ[ 27 | zero. λ y. (add0 y)⁻¹; 28 | suc x. λ y. ap suc' (addComm x y) ∙ (addSuc y x)⁻¹]; 29 | 30 | mul (x y : ℕ) : ℕ := case x [ 31 | zero. zero; 32 | suc x. add y (mul x y)]; 33 | 34 | n1 := suc zero; 35 | n2 := suc n1; 36 | n3 := suc n2; 37 | n4 := suc n3; 38 | n5 := suc n4; 39 | n6 := suc n5; 40 | n7 := suc n6; 41 | n8 := suc n7; 42 | n9 := suc n8; 43 | n10 := suc n9; 44 | n11 := suc n10; 45 | n12 := suc n11; 46 | n13 := suc n12; 47 | n14 := suc n13; 48 | n15 := suc n14; 49 | n16 := suc n15; 50 | n17 := suc n16; 51 | n18 := suc n17; 52 | n19 := suc n18; 53 | n20 := suc n19; 54 | n100 := mul n10 n10; 55 | n1k := mul n10 n100; 56 | n2k := add n1k n1k; 57 | n10k := mul n10 n1k; 58 | n100k := mul n10 n10k; 59 | n200k := add n100k n100k; 60 | n500k := mul n5 n100k; 61 | n1m := mul n1k n1k; 62 | n2m := add n1m n1m; 63 | n5m := mul n5 n1m; 64 | -------------------------------------------------------------------------------- /tests/tests/notes.cctt: -------------------------------------------------------------------------------- 1 | 2 | {- 3 | - Homotopy type theory 2023 conference 4 | - kollab: - Anders Mörtberg 5 | - Evan Cavallo 6 | + Tom Jack 7 | 8 | 9 | Mi a téma? 10 | 11 | - Cubical type theory (sok féle Cubical TT) 12 | - implementációja HoTT-nak 13 | 14 | - Cél: implementáció gyorsabb legyen 15 | 16 | - Mit szeretnénk gyorsítani: 17 | 18 | - fő felh terület: 19 | - matematikai formalizálás 20 | - homotópia-elmélet: 21 | - "gyengébb" topológia, folytonosan transzformálható dolgok 22 | azonosnak vannak véve (fánk = bögre/csésze) 23 | 24 | - N-dimenziós gömbök tulajdonságai 25 | - loop space: veszünk egy pontot és a hurkokat azon a ponton 26 | - hurkok körön ~ egész számok halmaza 27 | 0 egy helyben állunk 28 | +1 óramutatóval megegyező kör 29 | -1 másik irányba kör 30 | (bijekció a kettő között) 31 | - N-edik loop space: 32 | - 2-ik loop space: hurkok közötti transzformációk 33 | - 3-ik loop space: ... 34 | 35 | - példa: 3d gömb felület, azon veszünk egy pontot 36 | - 2-ik loop space: triviális (1 elemű halmaz) 37 | 38 | - Halmazelmélet: 39 | - primitív objektum: halmaz ("well-founded" fák) 40 | tartalmazás reláció nem mehet végtelenül lefelé 41 | 42 | - HoTT: 43 | - primitív objektum: 44 | tér, ahol folytonos transzformáció alatt minden invariáns 45 | - speciális eset: diszkrét tér ~ halmaz 46 | 47 | - Kör mint tér HoTT-ban: 48 | 49 | - Agda szintaxis: 50 | 51 | data Circle : Type where 52 | base : Circle -- "bázis" pont 53 | loop : base = base -- út ami base és base között van 54 | 55 | Kör := a "legkisebb" tér, amiben van egy pont és egy hurok a ponton 56 | ("legkisebb": csak az van, amit specifikáltunk, és semmi más) 57 | 58 | "higher inductive" definíció 59 | 60 | induktív definíció: 61 | data Nat : Set where 62 | zero : Nat 63 | suc : Nat → Nat 64 | 65 | Nat := szabadon generált halmaz, amiben van egy elem (zero), és 66 | egy endofüggvény (rákövetkezés) 67 | 68 | n, m : Nat 69 | n = m 70 | 71 | 72 | CTT-ben formalizálták: π₄(S³) -- (4D gömbfelület 4-edik loop space-je) 73 | 74 | - klasszikus homotópia elmélet: nem konstruktív 75 | 76 | - létezik egy egész szám B, amire igaz, hogy π₄(S³) ≈ ℤ/B 77 | B definíciója konstruktív, ha futtatjuk, akkor azt szeretnénk, hogy 78 | B = 2 79 | 80 | 81 | -} 82 | -------------------------------------------------------------------------------- /tests/tests/oldbrunerie-shrink.cctt: -------------------------------------------------------------------------------- 1 | isIso (A B : U)(f : A → B) : U := 2 | (g : B → A) 3 | × (linv : (x : A) → g (f x) = x) 4 | × (rinv : (x : B) → f (g x) = x); 5 | 6 | isEquiv (A B : U)(f : A → B) : U := 7 | (g : B → A) 8 | × (linv : (x : A) → g (f x) = x) 9 | × (rinv : (x : B) → f (g x) = x) 10 | × (coh : (x : A) → (rinv (f x)) ={i. f (linv x i) = f x} refl); 11 | 12 | equiv (A B : U) : U := 13 | (f : A → B) × isEquiv A B f; 14 | 15 | iso (A B : U) : U := (f : A → B) × isIso A B f; 16 | 17 | idEquiv (A : U) : equiv A A := 18 | (λ x. x, λ x. x, λ x _. x, λ x _. x, λ x _ _. x); 19 | 20 | 21 | ---------------------------------------------------------------------------------------------------- 22 | 23 | higher inductive join (A B : U) := 24 | inl A 25 | | inr B 26 | | push (a : A) (b : B) (i : I) [ i = 0. inl a 27 | ; i = 1. inr b ]; 28 | 29 | higher inductive S1 := 30 | base1 31 | | loop1 (i : I) [i=0. base1; i=1. base1]; 32 | 33 | higher inductive S2 := 34 | base2 35 | | loop2 (i j : I) [i=0. base2 36 | ;i=1. base2 37 | ;j=0. base2 38 | ;j=1. base2]; 39 | 40 | higher inductive S3 := 41 | base3 42 | | loop3 (i j k : I) [ i=0. base3 43 | ;i=1. base3 44 | ;j=0. base3 45 | ;j=1. base3 46 | ;k=0. base3 47 | ;k=1. base3]; 48 | 49 | rotLoop : (a : S1) → a = a := 50 | λ a. case a (a. a ={_. S1} a) [ 51 | base1. λ i. loop1 i 52 | ; loop1 i. λ j. hcom 0 1 S1 [ 53 | i = 0 k. hcom 1 j S1 [k = 0 _. base1; k = 1 l. loop1 l] base1 54 | ; i = 1 k. hcom 0 j S1 [k = 0 _. base1; k = 1 l. loop1 l] base1 55 | ; j = 0 k. hcom 1 i S1 [k = 0 _. base1; k = 1 l. loop1 l] base1 56 | ; j = 1 k. hcom 0 i S1 [k = 0 _. base1; k = 1 l. loop1 l] base1] base1]; 57 | 58 | unrotLoop (a : S1) : a = a := (rotLoop a)⁻¹; 59 | 60 | higher inductive PostTotalHopf := 61 | b (x : S1) 62 | | l (x : S1) (i j : I) 63 | [ i=0. b x 64 | ; i=1. b x 65 | ; j=0. b (unrotLoop x i) 66 | ; j=1. b x ]; 67 | 68 | the (A : U)(x : A) := x; 69 | 70 | test0To4 : 71 | (λ _ _. base2) ={_. (λ _. base2) ={_. base2 ={_. S2} base2} (λ _. base2)} (λ _ _. base2) := 72 | λ i j k. hcom 1 0 S2 [ 73 | i = 1 l. hcom 0 1 S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] (loop2 k i)] base2 74 | ; j = 0 l. hcom 0 i S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] (loop2 k i)] base2 75 | ; j = 1 l. hcom 0 i S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] (loop2 k i)] base2 76 | ; i = 0 l. hcom 1 0 S2 [ 77 | j = 0 i. hcom 0 i S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] base2] base2 78 | ; j = 1 i. hcom 0 i S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] base2] base2 79 | ; l = 0 _. base2 80 | ; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. loop2 @1 j] base2] 81 | (hcom 0 1 S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] base2] base2) 82 | ; k = 0 l. hcom 1 i S2 [ 83 | j = 0 i. hcom 0 i S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] base2] base2 84 | ; j = 1 i. hcom 0 i S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] base2] base2 85 | ; l = 0 _. base2 86 | ; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. loop2 @1 j] base2] 87 | (hcom 0 1 S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] base2] base2) 88 | ; k = 1 l. hcom 1 i S2 [ 89 | j = 0 i. hcom 0 i S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] base2] base2 90 | ; j = 1 i. hcom 0 i S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] base2] base2 91 | ; l = 0 _. base2 92 | ; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. loop2 @1 j] base2] 93 | (hcom 0 1 S2 [l = 0 _. base2; l = 1 i. hcom 0 1 S2 [i = 0 _. base2; i = 1 j. base2] base2] base2)] 94 | (hcom 0 1 S2 [i = 0 _. base2; i = 1 j. loop2 @1 j] (loop2 k i)); 95 | 96 | f5 : 97 | (λ _ _. base2) ={_. (λ _. base2) ={_. base2 ={_. S2} base2} (λ _. base2)} (λ _ _. base2) 98 | 99 | → (λ _ _. inl base1) ={_. (λ _. inl base1) ={_. inl base1 ={_. join S1 S1} inl base1} (λ _. inl base1)} 100 | 101 | (λ _ _. inl base1) := 102 | 103 | λ p i j k. 104 | case 105 | (case (p i j k) 106 | (x. (case x (_. U) [ 107 | base2. S1; 108 | loop2 i j. Glue S1 [ 109 | i = 0. S1, idEquiv S1; 110 | i = 1. S1, idEquiv S1; 111 | j = 0. S1, ?_; 112 | j = 1. S1, idEquiv S1]]) → PostTotalHopf) 113 | 114 | [base2. λ y. b y 115 | ;loop2 i j. λ y. hcom 0 1 PostTotalHopf 116 | [i = 0 _. b y; 117 | i = 1 _. b y; 118 | j = 0 k. b ?_; 119 | j = 1 _. b y 120 | ] 121 | (l (unglue y) i j)] 122 | ?_) 123 | (_. join S1 S1) 124 | [b x. inl x; l x i j. ?_]; 125 | 126 | test0To5 := f5 test0To4; 127 | -------------------------------------------------------------------------------- /tests/tests/recHITboundary.cctt: -------------------------------------------------------------------------------- 1 | 2 | 3 | isProp (A : U) : U := (a b : A) → a = b; 4 | isSet (A : U) : U := (a b : A) → isProp (a = b); 5 | 6 | higher inductive pTrunc (A : U) := 7 | pinc A 8 | | pquot (x y : pTrunc A)(i : I) [i=0. x; i=1. y]; 9 | 10 | recpt (A B : U) (Bprop : isProp B) (f : A → B) : (x : pTrunc A) → B := λ[ 11 | pinc x. f x; 12 | pquot x y i. Bprop (recpt A B Bprop f x) (recpt A B Bprop f y) i 13 | ]; 14 | 15 | higher inductive sTrunc (A : U) := 16 | sinc (a : A) 17 | | ssquash (a b : sTrunc A) (p q : a = b) (i j : I) 18 | [i=0. p j 19 | ;i=1. q j 20 | ;j=0. a 21 | ;j=1. b]; 22 | 23 | rec0 (A B : U) (Bset : isSet B) (f : A → B) : (x : sTrunc A) → B := λ[ 24 | sinc x. f x; 25 | ssquash x y p q i j. 26 | Bset (rec0 A B Bset f x) (rec0 A B Bset f y) 27 | (ap (rec0 A B Bset f) p) 28 | (ap (rec0 A B Bset f) q) 29 | i j 30 | ]; 31 | 32 | inductive bool := true | false; 33 | 34 | higher inductive Int := left | right | seg (i : I)[i=0. left; i=1.right]; 35 | 36 | -- bad : Int → bool := λ[ 37 | -- left. true; 38 | -- right. false; 39 | -- seg i. true 40 | -- ]; 41 | -------------------------------------------------------------------------------- /tests/tests/test.cctt: -------------------------------------------------------------------------------- 1 | 2 | foo (A : U)(f : A → A)(x : A) : A := f x; 3 | -------------------------------------------------------------------------------- /types23/bench.txt: -------------------------------------------------------------------------------- 1 | 2 | 3 | BoolNeg 4 | 5 | Agda 6 | 7 | 100 0.29 8 | 250 0.97 9 | 500 3.36 10 | 750 7.07 11 | 1000 12.57 12 | 1 million N/A 13 | 14 | 15 | cctt 16 | 17 | 100 0.00041 18 | 250 0.00095 19 | 500 0.0019 20 | 750 0.0030 21 | 1000 0.0047 22 | 1 million 5.65 23 | 24 | 25 | winding 26 | 27 | Agda 28 | 29 | 100 0.34 30 | 250 1.89 31 | 500 5.643 32 | 750 10.37 33 | 1000 18.52 34 | 1 million N/A 35 | 36 | cctt 37 | 38 | 39 | 100 0.0005 40 | 250 0.0012 41 | 500 0.0023 42 | 750 0.0043 43 | 1000 0.0059 44 | 1 million 7.98 45 | -------------------------------------------------------------------------------- /types23/brunerie.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/types23/brunerie.png -------------------------------------------------------------------------------- /types23/prez.aux: -------------------------------------------------------------------------------- 1 | \relax 2 | \providecommand\hyper@newdestlabel[2]{} 3 | \providecommand\HyperFirstAtBeginDocument{\AtBeginDocument} 4 | \HyperFirstAtBeginDocument{\ifx\hyper@anchor\@undefined 5 | \global\let\oldcontentsline\contentsline 6 | \gdef\contentsline#1#2#3#4{\oldcontentsline{#1}{#2}{#3}} 7 | \global\let\oldnewlabel\newlabel 8 | \gdef\newlabel#1#2{\newlabelxx{#1}#2} 9 | \gdef\newlabelxx#1#2#3#4#5#6{\oldnewlabel{#1}{{#2}{#3}}} 10 | \AtEndDocument{\ifx\hyper@anchor\@undefined 11 | \let\contentsline\oldcontentsline 12 | \let\newlabel\oldnewlabel 13 | \fi} 14 | \fi} 15 | \global\let\hyper@last\relax 16 | \gdef\HyperFirstAtBeginDocument#1{#1} 17 | \providecommand\HyField@AuxAddToFields[1]{} 18 | \providecommand\HyField@AuxAddToCoFields[2]{} 19 | \bibstyle{alpha} 20 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{1}{1/1}{}{0}}} 21 | \@writefile{nav}{\headcommand {\beamer@framepages {1}{1}}} 22 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{2}{2/9}{}{0}}} 23 | \@writefile{nav}{\headcommand {\beamer@framepages {2}{9}}} 24 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{3}{10/13}{}{0}}} 25 | \@writefile{nav}{\headcommand {\beamer@framepages {10}{13}}} 26 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{4}{14/15}{}{0}}} 27 | \@writefile{nav}{\headcommand {\beamer@framepages {14}{15}}} 28 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{5}{16/16}{}{0}}} 29 | \@writefile{nav}{\headcommand {\beamer@framepages {16}{16}}} 30 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{6}{17/18}{}{0}}} 31 | \@writefile{nav}{\headcommand {\beamer@framepages {17}{18}}} 32 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{7}{19/23}{}{0}}} 33 | \@writefile{nav}{\headcommand {\beamer@framepages {19}{23}}} 34 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{8}{24/25}{}{0}}} 35 | \@writefile{nav}{\headcommand {\beamer@framepages {24}{25}}} 36 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{9}{26/29}{}{0}}} 37 | \@writefile{nav}{\headcommand {\beamer@framepages {26}{29}}} 38 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{10}{30/31}{}{0}}} 39 | \@writefile{nav}{\headcommand {\beamer@framepages {30}{31}}} 40 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{11}{32/35}{}{0}}} 41 | \@writefile{nav}{\headcommand {\beamer@framepages {32}{35}}} 42 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{12}{36/40}{}{0}}} 43 | \@writefile{nav}{\headcommand {\beamer@framepages {36}{40}}} 44 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{13}{41/44}{}{0}}} 45 | \@writefile{nav}{\headcommand {\beamer@framepages {41}{44}}} 46 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{14}{45/47}{}{0}}} 47 | \@writefile{nav}{\headcommand {\beamer@framepages {45}{47}}} 48 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{15}{48/49}{}{0}}} 49 | \@writefile{nav}{\headcommand {\beamer@framepages {48}{49}}} 50 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{16}{50/50}{}{0}}} 51 | \@writefile{nav}{\headcommand {\beamer@framepages {50}{50}}} 52 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{17}{51/51}{}{0}}} 53 | \@writefile{nav}{\headcommand {\beamer@framepages {51}{51}}} 54 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{18}{52/52}{}{0}}} 55 | \@writefile{nav}{\headcommand {\beamer@framepages {52}{52}}} 56 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{19}{53/53}{}{0}}} 57 | \@writefile{nav}{\headcommand {\beamer@framepages {53}{53}}} 58 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{20}{54/56}{}{0}}} 59 | \@writefile{nav}{\headcommand {\beamer@framepages {54}{56}}} 60 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{21}{57/60}{}{0}}} 61 | \@writefile{nav}{\headcommand {\beamer@framepages {57}{60}}} 62 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{22}{61/62}{}{0}}} 63 | \@writefile{nav}{\headcommand {\beamer@framepages {61}{62}}} 64 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{23}{63/69}{}{0}}} 65 | \@writefile{nav}{\headcommand {\beamer@framepages {63}{69}}} 66 | \@writefile{nav}{\headcommand {\slideentry {0}{0}{24}{70/73}{}{0}}} 67 | \@writefile{nav}{\headcommand {\beamer@framepages {70}{73}}} 68 | \@writefile{nav}{\headcommand {\beamer@partpages {1}{73}}} 69 | \@writefile{nav}{\headcommand {\beamer@subsectionpages {1}{73}}} 70 | \@writefile{nav}{\headcommand {\beamer@sectionpages {1}{73}}} 71 | \@writefile{nav}{\headcommand {\beamer@documentpages {73}}} 72 | \@writefile{nav}{\headcommand {\gdef \inserttotalframenumber {24}}} 73 | \gdef \@abspage@last{73} 74 | -------------------------------------------------------------------------------- /types23/prez.nav: -------------------------------------------------------------------------------- 1 | \headcommand {\slideentry {0}{0}{1}{1/1}{}{0}} 2 | \headcommand {\beamer@framepages {1}{1}} 3 | \headcommand {\slideentry {0}{0}{2}{2/9}{}{0}} 4 | \headcommand {\beamer@framepages {2}{9}} 5 | \headcommand {\slideentry {0}{0}{3}{10/13}{}{0}} 6 | \headcommand {\beamer@framepages {10}{13}} 7 | \headcommand {\slideentry {0}{0}{4}{14/15}{}{0}} 8 | \headcommand {\beamer@framepages {14}{15}} 9 | \headcommand {\slideentry {0}{0}{5}{16/16}{}{0}} 10 | \headcommand {\beamer@framepages {16}{16}} 11 | \headcommand {\slideentry {0}{0}{6}{17/18}{}{0}} 12 | \headcommand {\beamer@framepages {17}{18}} 13 | \headcommand {\slideentry {0}{0}{7}{19/23}{}{0}} 14 | \headcommand {\beamer@framepages {19}{23}} 15 | \headcommand {\slideentry {0}{0}{8}{24/25}{}{0}} 16 | \headcommand {\beamer@framepages {24}{25}} 17 | \headcommand {\slideentry {0}{0}{9}{26/29}{}{0}} 18 | \headcommand {\beamer@framepages {26}{29}} 19 | \headcommand {\slideentry {0}{0}{10}{30/31}{}{0}} 20 | \headcommand {\beamer@framepages {30}{31}} 21 | \headcommand {\slideentry {0}{0}{11}{32/35}{}{0}} 22 | \headcommand {\beamer@framepages {32}{35}} 23 | \headcommand {\slideentry {0}{0}{12}{36/40}{}{0}} 24 | \headcommand {\beamer@framepages {36}{40}} 25 | \headcommand {\slideentry {0}{0}{13}{41/44}{}{0}} 26 | \headcommand {\beamer@framepages {41}{44}} 27 | \headcommand {\slideentry {0}{0}{14}{45/47}{}{0}} 28 | \headcommand {\beamer@framepages {45}{47}} 29 | \headcommand {\slideentry {0}{0}{15}{48/49}{}{0}} 30 | \headcommand {\beamer@framepages {48}{49}} 31 | \headcommand {\slideentry {0}{0}{16}{50/50}{}{0}} 32 | \headcommand {\beamer@framepages {50}{50}} 33 | \headcommand {\slideentry {0}{0}{17}{51/51}{}{0}} 34 | \headcommand {\beamer@framepages {51}{51}} 35 | \headcommand {\slideentry {0}{0}{18}{52/52}{}{0}} 36 | \headcommand {\beamer@framepages {52}{52}} 37 | \headcommand {\slideentry {0}{0}{19}{53/53}{}{0}} 38 | \headcommand {\beamer@framepages {53}{53}} 39 | \headcommand {\slideentry {0}{0}{20}{54/56}{}{0}} 40 | \headcommand {\beamer@framepages {54}{56}} 41 | \headcommand {\slideentry {0}{0}{21}{57/60}{}{0}} 42 | \headcommand {\beamer@framepages {57}{60}} 43 | \headcommand {\slideentry {0}{0}{22}{61/62}{}{0}} 44 | \headcommand {\beamer@framepages {61}{62}} 45 | \headcommand {\slideentry {0}{0}{23}{63/69}{}{0}} 46 | \headcommand {\beamer@framepages {63}{69}} 47 | \headcommand {\slideentry {0}{0}{24}{70/73}{}{0}} 48 | \headcommand {\beamer@framepages {70}{73}} 49 | \headcommand {\beamer@partpages {1}{73}} 50 | \headcommand {\beamer@subsectionpages {1}{73}} 51 | \headcommand {\beamer@sectionpages {1}{73}} 52 | \headcommand {\beamer@documentpages {73}} 53 | \headcommand {\gdef \inserttotalframenumber {24}} 54 | -------------------------------------------------------------------------------- /types23/prez.out: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/types23/prez.out -------------------------------------------------------------------------------- /types23/prez.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/types23/prez.pdf -------------------------------------------------------------------------------- /types23/prez.snm: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/types23/prez.snm -------------------------------------------------------------------------------- /types23/prez.toc: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/types23/prez.toc -------------------------------------------------------------------------------- /types23/unglues.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/AndrasKovacs/cctt/08917c63189862bd537d0648ecbe5c1e46551651/types23/unglues.png --------------------------------------------------------------------------------