├── .gitignore ├── Coerce.agda ├── Core.agda ├── Data.agda ├── Data ├── Fin.agda ├── Lift.agda ├── List.agda ├── Maybe.agda ├── Sum.agda ├── Vec.agda └── W.agda ├── Examples └── Favourite.agda ├── Function ├── Elim.agda ├── Erase.agda └── Pi.agda ├── Lib ├── Decidable.agda └── Heteroindexed.agda ├── Main.agda ├── Prelude.agda ├── Property └── Eq.agda └── readme.md /.gitignore: -------------------------------------------------------------------------------- 1 | *~ 2 | *.agdai -------------------------------------------------------------------------------- /Coerce.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --no-termination-check #-} 2 | 3 | module OTT.Coerce where 4 | 5 | open import OTT.Core 6 | 7 | coerceFamℕ : ∀ {n m} -> (A : ℕ -> Set) -> ⟦ n ≟ⁿ m ⟧ -> A n -> A m 8 | coerceFamℕ {0} {0} A q x = x 9 | coerceFamℕ {suc _} {suc _} A q x = coerceFamℕ (A ∘ suc) q x 10 | coerceFamℕ {0} {suc _} A () x 11 | coerceFamℕ {suc _} {0} A () x 12 | 13 | coerceFamLevel : ∀ {a b} {α : Level a} {β : Level b} 14 | -> (A : ∀ {a} -> Level a -> Set) -> ⟦ α ≟ˡ β ⟧ -> A α -> A β 15 | coerceFamLevel {α = lzero } {lzero } A q x = x 16 | coerceFamLevel {α = lsuc _} {lsuc _} A q x = coerceFamLevel (A ∘ lsuc) q x 17 | coerceFamLevel {α = lzero } {lsuc _} A () x 18 | coerceFamLevel {α = lsuc _} {lzero } A () x 19 | 20 | meta-eq : ∀ {a b} {β : Level b} -> (α : Level a) -> ⟦ α ≟ˡ β ⟧ -> a ≡ b 21 | meta-eq α q = coerceFamLevel {α = α} (λ {b} _ -> _ ≡ b) q prefl 22 | 23 | Coerce : ∀ {a b} {α : Level a} {β : Level b} -> ⟦ α ≟ˡ β ⟧ -> Univ α -> Univ β 24 | Coerce = coerceFamLevel Univ 25 | 26 | mutual 27 | coerce : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} -> ⟦ A ≈ B ⇒ A ⇒ B ⟧ 28 | coerce {α = lzero } {lzero } (f , _) = f 29 | coerce {α = lsuc _} {lsuc _} q = coerce′ q 30 | coerce {α = lzero } {lsuc _} () 31 | coerce {α = lsuc _} {lzero } () 32 | 33 | coerce′ : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} -> ⟦ A ≃ B ⇒ A ⇒ B ⟧ 34 | coerce′ {A = bot } {bot } q () 35 | coerce′ {A = top } {top } q tt = tt 36 | coerce′ {A = nat } {nat } q n = n 37 | coerce′ {A = enum n₁ } {enum n₂ } q e = coerceFamℕ (Apply Enum) q e 38 | coerce′ {A = univ α₁ } {univ α₂ } q A = Coerce q A 39 | coerce′ {A = σ A₁ B₁ } {σ A₂ B₂ } q p = let q₁ , q₂ = q ; x , y = p in 40 | coerce q₁ x , coerce (q₂ x (coerce q₁ x) (coherence q₁ x)) y 41 | coerce′ {A = π A₁ B₁ } {π A₂ B₂ } q f = let q₁ , q₂ = q in 42 | λ x -> coerce (q₂ (coerce q₁ x) x (coherence q₁ x)) (f (coerce q₁ x)) 43 | coerce′ {A = desc _ α} {desc _ _} q D = let qI , qα = q in coerceDesc qI (meta-eq α qα) D 44 | coerce′ {A = imu _ _ } {imu _ _ } q d = let qD , qi = q in coerceMu qD qi d 45 | coerce′ {A = bot } {top } () 46 | coerce′ {A = bot } {nat } () 47 | coerce′ {A = bot } {enum _ } () 48 | coerce′ {A = bot } {univ _ } () 49 | coerce′ {A = bot } {σ _ _ } () 50 | coerce′ {A = bot } {π _ _ } () 51 | coerce′ {A = bot } {desc _ _} () 52 | coerce′ {A = bot } {imu _ _ } () 53 | coerce′ {A = top } {bot } () 54 | coerce′ {A = top } {nat } () 55 | coerce′ {A = top } {enum _ } () 56 | coerce′ {A = top } {univ _ } () 57 | coerce′ {A = top } {σ _ _ } () 58 | coerce′ {A = top } {π _ _ } () 59 | coerce′ {A = top } {desc _ _} () 60 | coerce′ {A = top } {imu _ _ } () 61 | coerce′ {A = nat } {bot } () 62 | coerce′ {A = nat } {top } () 63 | coerce′ {A = nat } {enum _ } () 64 | coerce′ {A = nat } {univ _ } () 65 | coerce′ {A = nat } {σ _ _ } () 66 | coerce′ {A = nat } {π _ _ } () 67 | coerce′ {A = nat } {desc _ _} () 68 | coerce′ {A = nat } {imu _ _ } () 69 | coerce′ {A = enum _ } {bot } () 70 | coerce′ {A = enum _ } {top } () 71 | coerce′ {A = enum _ } {nat } () 72 | coerce′ {A = enum _ } {univ _ } () 73 | coerce′ {A = enum _ } {σ _ _ } () 74 | coerce′ {A = enum _ } {π _ _ } () 75 | coerce′ {A = enum _ } {desc _ _} () 76 | coerce′ {A = enum _ } {imu _ _ } () 77 | coerce′ {A = univ _ } {bot } () 78 | coerce′ {A = univ _ } {top } () 79 | coerce′ {A = univ _ } {nat } () 80 | coerce′ {A = univ _ } {enum _ } () 81 | coerce′ {A = univ _ } {σ _ _ } () 82 | coerce′ {A = univ _ } {π _ _ } () 83 | coerce′ {A = univ _ } {desc _ _} () 84 | coerce′ {A = univ _ } {imu _ _ } () 85 | coerce′ {A = σ _ _ } {bot } () 86 | coerce′ {A = σ _ _ } {top } () 87 | coerce′ {A = σ _ _ } {nat } () 88 | coerce′ {A = σ _ _ } {enum _ } () 89 | coerce′ {A = σ _ _ } {univ _ } () 90 | coerce′ {A = σ _ _ } {π _ _ } () 91 | coerce′ {A = σ _ _ } {desc _ _} () 92 | coerce′ {A = σ _ _ } {imu _ _ } () 93 | coerce′ {A = π _ _ } {bot } () 94 | coerce′ {A = π _ _ } {top } () 95 | coerce′ {A = π _ _ } {nat } () 96 | coerce′ {A = π _ _ } {enum _ } () 97 | coerce′ {A = π _ _ } {univ _ } () 98 | coerce′ {A = π _ _ } {σ _ _ } () 99 | coerce′ {A = π _ _ } {desc _ _} () 100 | coerce′ {A = π _ _ } {imu _ _ } () 101 | coerce′ {A = desc _ _} {bot } () 102 | coerce′ {A = desc _ _} {top } () 103 | coerce′ {A = desc _ _} {nat } () 104 | coerce′ {A = desc _ _} {enum _ } () 105 | coerce′ {A = desc _ _} {univ _ } () 106 | coerce′ {A = desc _ _} {σ _ _ } () 107 | coerce′ {A = desc _ _} {π _ _ } () 108 | coerce′ {A = desc _ _} {imu _ _ } () 109 | coerce′ {A = imu _ _ } {bot } () 110 | coerce′ {A = imu _ _ } {top } () 111 | coerce′ {A = imu _ _ } {nat } () 112 | coerce′ {A = imu _ _ } {enum _ } () 113 | coerce′ {A = imu _ _ } {univ _ } () 114 | coerce′ {A = imu _ _ } {σ _ _ } () 115 | coerce′ {A = imu _ _ } {π _ _ } () 116 | coerce′ {A = imu _ _ } {desc _ _} () 117 | -- (almost) generated by http://ideone.com/ltrf04 118 | 119 | coerceDesc : ∀ {i₁ i₂ a₁ a₂} {ι₁ : Level i₁} {ι₂ : Level i₂} 120 | {α₁ : Level a₁} {α₂ : Level a₂} {I₁ : Type ι₁} {I₂ : Type ι₂} 121 | -> ⟦ I₁ ≃ I₂ ⟧ -> a₁ ≡ a₂ -> Desc I₁ α₁ -> Desc I₂ α₂ 122 | coerceDesc qI qa (var i) = var (coerce qI i) 123 | coerceDesc qI qa (π {a} {{q}} A D) = 124 | π {{ptrans (pright (pcong (a ⊔ₘ_) qa) q) qa}} A (coerceDesc qI qa ∘ D) 125 | coerceDesc qI qa (D ⊛ E) = coerceDesc qI qa D ⊛ coerceDesc qI qa E 126 | 127 | coerceSem : ∀ {i₁ i₂ a₁ a₂ b₁ b₂} 128 | {ι₁ : Level i₁} {ι₂ : Level i₂} {α₁ : Level a₁} {α₂ : Level a₂} 129 | {β₁ : Level b₁} {β₂ : Level b₂} {I₁ : Type ι₁} {I₂ : Type ι₂} 130 | {B₁ : ⟦ I₁ ⟧ -> Univ β₁} {B₂ : ⟦ I₂ ⟧ -> Univ β₂} 131 | -> (D₁ : Desc I₁ α₁) 132 | -> (D₂ : Desc I₂ α₂) 133 | -> ⟦ D₁ ≅ᵈ D₂ ⟧ 134 | -> ⟦ B₁ ≅ B₂ ⟧ 135 | -> (⟦ D₁ ⟧ᵈ ⟦ B₁ ⟧ⁱ) 136 | -> (⟦ D₂ ⟧ᵈ ⟦ B₂ ⟧ⁱ) 137 | coerceSem (var i₁) (var i₂) qi qB x = coerce (qB i₁ i₂ qi) x 138 | coerceSem (π A₁ D₁) (π A₂ D₂) (qA , qD) qB f = λ x -> 139 | let qA′ = sym A₁ {A₂} qA 140 | x′ = coerce qA′ x 141 | in coerceSem (D₁ x′) (D₂ x) (qD x′ x (sym x (coherence qA′ x))) qB (f x′) 142 | coerceSem (D₁ ⊛ E₁) (D₂ ⊛ E₂) (qD , qE) qB (s , t) = 143 | coerceSem D₁ D₂ qD qB s , coerceSem E₁ E₂ qE qB t 144 | coerceSem (var _) (π _ _) () 145 | coerceSem (var _) (_ ⊛ _) () 146 | coerceSem (π _ _) (var _) () 147 | coerceSem (π _ _) (_ ⊛ _) () 148 | coerceSem (_ ⊛ _) (var _) () 149 | coerceSem (_ ⊛ _) (π _ _) () 150 | 151 | coerceExtend : ∀ {i₁ i₂ a₁ a₂ b₁ b₂} 152 | {ι₁ : Level i₁} {ι₂ : Level i₂} {α₁ : Level a₁} {α₂ : Level a₂} 153 | {β₁ : Level b₁} {β₂ : Level b₂} {I₁ : Type ι₁} {I₂ : Type ι₂} 154 | {B₁ : ⟦ I₁ ⟧ -> Univ β₁} {B₂ : ⟦ I₂ ⟧ -> Univ β₂} {j₁ j₂} 155 | -> (D₁ : Desc I₁ α₁) 156 | -> (D₂ : Desc I₂ α₂) 157 | -> ⟦ D₁ ≅ᵈ D₂ ⟧ 158 | -> ⟦ B₁ ≅ B₂ ⟧ 159 | -> ⟦ j₁ ≅ j₂ ⟧ 160 | -> Extend D₁ ⟦ B₁ ⟧ⁱ j₁ 161 | -> Extend D₂ ⟦ B₂ ⟧ⁱ j₂ 162 | coerceExtend (var i₁) (var i₂) qi qB qj qij = trans i₂ (right i₁ qi qij) qj 163 | coerceExtend (π A₁ D₁) (π A₂ D₂) (qA , qD) qB qi (x , e) = let x′ = coerce qA x in 164 | x′ , coerceExtend (D₁ x) (D₂ x′) (qD x x′ (coherence qA x)) qB qi e 165 | coerceExtend (D₁ ⊛ E₁) (D₂ ⊛ E₂) (qD , qE) qB qi (s , e) = 166 | coerceSem D₁ D₂ qD qB s , coerceExtend E₁ E₂ qE qB qi e 167 | coerceExtend (var _) (π _ _) () 168 | coerceExtend (var _) (_ ⊛ _) () 169 | coerceExtend (π _ _) (var _) () 170 | coerceExtend (π _ _) (_ ⊛ _) () 171 | coerceExtend (_ ⊛ _) (var _) () 172 | coerceExtend (_ ⊛ _) (π _ _) () 173 | 174 | coerceMu : ∀ {i₁ i₂ a₁ a₂} {ι₁ : Level i₁} {ι₂ : Level i₂} {α₁ : Level a₁} {α₂ : Level a₂} 175 | {I₁ : Type ι₁} {I₂ : Type ι₂} {D₁ : Desc I₁ α₁} {D₂ : Desc I₂ α₂} {j₁ j₂} 176 | -> ⟦ D₁ ≊ᵈ D₂ ⟧ -> ⟦ j₁ ≅ j₂ ⟧ -> μ D₁ j₁ -> μ D₂ j₂ 177 | coerceMu {α₁ = lzero } {lzero } qD qj d = proj₁ (qD _ _ qj) d 178 | coerceMu {α₁ = lsuc _} {lsuc _} {D₁ = D₁} {D₂} qD qj (node e) = 179 | node (coerceExtend D₁ D₂ qD (λ _ _ -> _,_ qD) qj e) 180 | coerceMu {α₁ = lzero } {lsuc _} () 181 | coerceMu {α₁ = lsuc _} {lzero } () 182 | 183 | postulate 184 | refl : ∀ {a} {α : Level a} {A : Univ α} -> (x : ⟦ A ⟧) -> ⟦ x ≅ x ⟧ 185 | coherence : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} 186 | -> (q : ⟦ A ≈ B ⟧) -> (x : ⟦ A ⟧) -> ⟦ x ≅ coerce q x ⟧ 187 | cong-≅z : ∀ {a b c} {α : Level a} {β : Level b} {γ : Level c} 188 | {A : Univ α} {B : Univ β} {C : Univ γ} 189 | -> (x : ⟦ A ⟧) {y : ⟦ B ⟧} {z : ⟦ C ⟧} -> (q : ⟦ x ≅ y ⟧) -> ⟦ (x ≅ z) ≈ (y ≅ z)⟧ 190 | huip : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} 191 | -> (x : ⟦ A ⟧) {y : ⟦ B ⟧} -> (q : ⟦ x ≅ y ⟧) -> ⟦ refl x ≅ q ⟧ 192 | 193 | right : ∀ {a b c} {α : Level a} {β : Level b} {γ : Level c} 194 | {A : Univ α} {B : Univ β} {C : Univ γ} 195 | -> (x : ⟦ A ⟧) {y : ⟦ B ⟧} {z : ⟦ C ⟧} -> ⟦ x ≅ y ⟧ -> ⟦ x ≅ z ⟧ -> ⟦ y ≅ z ⟧ 196 | right x q₁ = proj₁ (cong-≅z x q₁) 197 | 198 | trans : ∀ {a b c} {α : Level a} {β : Level b} {γ : Level c} 199 | {A : Univ α} {B : Univ β} {C : Univ γ} 200 | -> (x : ⟦ A ⟧) {y : ⟦ B ⟧} {z : ⟦ C ⟧} -> ⟦ x ≅ y ⟧ -> ⟦ y ≅ z ⟧ -> ⟦ x ≅ z ⟧ 201 | trans x {y} q₁ = proj₂ (cong-≅z x q₁) 202 | 203 | sym : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} 204 | -> (x : ⟦ A ⟧) {y : ⟦ B ⟧} -> ⟦ x ≅ y ⟧ -> ⟦ y ≅ x ⟧ 205 | sym x q = right x q (refl x) 206 | 207 | left : ∀ {a b c} {α : Level a} {β : Level b} {γ : Level c} 208 | {A : Univ α} {B : Univ β} {C : Univ γ} 209 | -> (x : ⟦ A ⟧) {y : ⟦ B ⟧} {z : ⟦ C ⟧} -> ⟦ x ≅ y ⟧ -> ⟦ z ≅ y ⟧ -> ⟦ x ≅ z ⟧ 210 | left x {z = z} q₁ q₂ = trans x q₁ (sym z q₂) 211 | 212 | subst : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {x y} 213 | -> (B : ⟦ A ⟧ -> Univ β) -> ⟦ x ≅ y ⇒ B x ⇒ B y ⟧ 214 | subst B q = coerce (refl B _ _ q) 215 | 216 | subst₂ : ∀ {a b c} {α : Level a} {β : Level b} {γ : Level c} {A : Univ α} 217 | {B : ⟦ A ⟧ -> Univ β} {x₁ x₂} {y₁ : ⟦ B x₁ ⟧} {y₂ : ⟦ B x₂ ⟧} 218 | -> (C : ∀ x -> ⟦ B x ⟧ -> Univ γ) -> ⟦ x₁ ≅ x₂ ⇒ y₁ ≅ y₂ ⇒ C x₁ y₁ ⇒ C x₂ y₂ ⟧ 219 | subst₂ C q₁ q₂ = coerce (refl C _ _ q₁ _ _ q₂) 220 | 221 | J : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {x y : ⟦ A ⟧} 222 | -> (B : (y : ⟦ A ⟧) -> ⟦ x ≅ y ⟧ -> Univ β) 223 | -> ⟦ B _ (refl x) ⟧ 224 | -> (q : ⟦ x ≅ y ⟧) 225 | -> ⟦ B _ q ⟧ 226 | J {x = x} B z q = subst₂ B q (huip x q) z 227 | -------------------------------------------------------------------------------- /Core.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --no-positivity-check --no-termination-check #-} 2 | 3 | module OTT.Core where 4 | 5 | open import OTT.Prelude public 6 | 7 | infixr 1 _&_ 8 | infixr 2 _⇒_ _⇨_ _⊛_ 9 | infix 3 _≟ⁿ_ _≟ˡ_ _≈_ _≃_ _≅_ _≅ᵉ_ _≅ᵈ_ _≊ᵈ_ _≅s_ _≅e_ 10 | 11 | _≤ₘ_ : MetaLevel -> MetaLevel -> Set 12 | a ≤ₘ b = a ⊔ₘ b ≡ b 13 | 14 | data Level : MetaLevel -> Set where 15 | lzero : Level lzeroₘ 16 | lsuc : ∀ {a} -> Level a -> Level (lsucₘ a) 17 | 18 | natToMetaLevel : ℕ -> MetaLevel 19 | natToMetaLevel 0 = lzeroₘ 20 | natToMetaLevel (suc n) = lsucₘ (natToMetaLevel n) 21 | 22 | natToLevel : ∀ n -> Level (natToMetaLevel n) 23 | natToLevel 0 = lzero 24 | natToLevel (suc n) = lsuc (natToLevel n) 25 | 26 | _⊔_ : ∀ {a b} -> Level a -> Level b -> Level (a ⊔ₘ b) 27 | lzero ⊔ lzero = lzero 28 | lzero ⊔ lsuc β = lsuc β 29 | lsuc α ⊔ lzero = lsuc α 30 | lsuc α ⊔ lsuc β = lsuc (α ⊔ β) 31 | 32 | _⊔ₘ₀_ : ∀ {a} -> MetaLevel -> Level a -> MetaLevel 33 | a ⊔ₘ₀ lzero = lzeroₘ 34 | a ⊔ₘ₀ lsuc {b} β = a ⊔ₘ lsucₘ b 35 | 36 | _⊔₀_ : ∀ {a b} -> Level a -> (β : Level b) -> Level (a ⊔ₘ₀ β) 37 | α ⊔₀ lzero = lzero 38 | α ⊔₀ lsuc β = α ⊔ lsuc β 39 | 40 | data Univ : ∀ {a} -> Level a -> Set 41 | 42 | Univ# = Univ ∘ natToLevel 43 | Type# = Univ# ∘ suc 44 | Prop = Univ# 0 45 | Type = λ {a} -> Univ ∘ lsuc {a} 46 | Type₀ = Type# 0 47 | 48 | Type₋₁ : ∀ {a} -> Level a -> Set 49 | Type₋₁ α = Univ (lsuc lzero ⊔ α) 50 | 51 | ⟦_⟧ : ∀ {a} {α : Level a} -> Univ α -> Set 52 | 53 | _≃_ : ∀ {a b} {α : Level a} {β : Level b} -> Univ α -> Univ β -> Prop 54 | _≅_ : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} -> ⟦ A ⟧ -> ⟦ B ⟧ -> Prop 55 | 56 | data Desc {i b} {ι : Level i} (I : Type ι) (β : Level b) : Set where 57 | var : ⟦ I ⟧ -> Desc I β 58 | π : ∀ {a} {α : Level a} .{{_ : a ≤ₘ b}} -> (A : Univ α) -> (⟦ A ⟧ -> Desc I β) -> Desc I β 59 | _⊛_ : Desc I β -> Desc I β -> Desc I β 60 | 61 | ⟦_⟧ᵈ : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} 62 | -> Desc I α -> (⟦ I ⟧ -> Set) -> Set 63 | ⟦ var i ⟧ᵈ B = B i 64 | ⟦ π A D ⟧ᵈ B = ∀ x -> ⟦ D x ⟧ᵈ B 65 | ⟦ D ⊛ E ⟧ᵈ B = ⟦ D ⟧ᵈ B × ⟦ E ⟧ᵈ B 66 | 67 | Extend : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} 68 | -> Desc I α -> (⟦ I ⟧ -> Set) -> ⟦ I ⟧ -> Set 69 | Extend (var i) B j = ⟦ i ≅ j ⟧ 70 | Extend (π A D) B i = ∃ λ x -> Extend (D x) B i 71 | Extend (D ⊛ E) B i = ⟦ D ⟧ᵈ B × Extend E B i 72 | 73 | -- Funnily, Agda treats inductive records and data types differently wrt termination checking. 74 | -- Perhaps it's not clear to Agda that induction is structural because 75 | -- irrefutable pattern matching elaborates into function application. Do we need refutable patterns? 76 | -- record μ {i a} {ι : Level i} {α : Level a} {I : Type ι} (D : Desc I α) j : Set where 77 | -- inductive 78 | -- constructor node 79 | -- field knot : Extend D (μ D) j 80 | -- open μ public 81 | 82 | data μ {i a} {ι : Level i} {α : Level a} {I : Type ι} (D : Desc I α) j : Set where 83 | node : Extend D (μ D) j -> μ D j 84 | 85 | knot : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} {D : Desc I α} {j} 86 | -> μ D j -> Extend D (μ D) j 87 | knot (node e) = e 88 | 89 | _⇨_ : ∀ {i a b} {ι : Level i} {α : Level a} {β : Level b} {I : Type ι} .{{_ : a ≤ₘ b}} 90 | -> Univ α -> Desc I β -> Desc I β 91 | A ⇨ D = π A λ _ -> D 92 | 93 | data Univ where 94 | bot : Prop 95 | top : Prop 96 | nat : Type₀ 97 | enum : ℕ -> Type₀ 98 | univ : ∀ {a} -> (α : Level a) -> Type α 99 | σ : ∀ {a b} {α : Level a} {β : Level b} -> (A : Univ α) -> (⟦ A ⟧ -> Univ β) -> Univ (α ⊔ β) 100 | π : ∀ {a b} {α : Level a} {β : Level b} -> (A : Univ α) -> (⟦ A ⟧ -> Univ β) -> Univ (α ⊔₀ β) 101 | desc : ∀ {a i} {ι : Level i} -> Type ι -> (α : Level a) -> Type α 102 | imu : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} -> Desc I α -> ⟦ I ⟧ -> Univ α 103 | 104 | ⟦_⟧ⁱ : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} -> (⟦ A ⟧ -> Univ β) -> ⟦ A ⟧ -> Set 105 | ⟦ B ⟧ⁱ x = ⟦ B x ⟧ 106 | 107 | ⟦ bot ⟧ = ⊥ 108 | ⟦ top ⟧ = ⊤ 109 | ⟦ nat ⟧ = ℕ 110 | ⟦ enum n ⟧ = Apply Enum n -- `Apply` allows to make `⟦_⟧` constructor-headed. 111 | ⟦ univ α ⟧ = Univ α 112 | ⟦ σ A B ⟧ = ∃ ⟦ B ⟧ⁱ 113 | ⟦ π A B ⟧ = ∀ x -> ⟦ B x ⟧ 114 | ⟦ desc I α ⟧ = Desc I α 115 | ⟦ imu D j ⟧ = μ D j 116 | 117 | pattern prop = univ lzero 118 | pattern type a = univ (lsuc a) 119 | univ# = univ ∘ natToLevel 120 | type# = univ# ∘ suc 121 | 122 | desc₋₁ : ∀ {a i} {ι : Level i} -> Type ι -> (α : Level a) -> Univ _ 123 | desc₋₁ I α = desc I (lsuc lzero ⊔ α) 124 | 125 | _&_ : ∀ {a b} {α : Level a} {β : Level b} -> Univ α -> Univ β -> Univ (α ⊔ β) 126 | A & B = σ A λ _ -> B 127 | 128 | _⇒_ : ∀ {a b} {α : Level a} {β : Level b} -> Univ α -> Univ β -> Univ (α ⊔₀ β) 129 | A ⇒ B = π A λ _ -> B 130 | 131 | _≟ⁿ_ : ℕ -> ℕ -> Prop 132 | 0 ≟ⁿ 0 = top 133 | suc n ≟ⁿ suc m = n ≟ⁿ m 134 | _ ≟ⁿ _ = bot 135 | 136 | _≟ˡ_ : ∀ {a b} -> Level a -> Level b -> Prop 137 | lzero ≟ˡ lzero = top 138 | lsuc α ≟ˡ lsuc β = α ≟ˡ β 139 | _ ≟ˡ _ = bot 140 | 141 | _≅ᵉ_ : ∃ Enum -> ∃ Enum -> Prop 142 | 0 , () ≅ᵉ 0 , () 143 | 1 , tt ≅ᵉ 1 , tt = top 144 | suc (suc n₁) , nothing ≅ᵉ suc (suc n₂) , nothing = top 145 | suc (suc n₁) , just e₁ ≅ᵉ suc (suc n₂) , just e₂ = suc n₁ , e₁ ≅ᵉ suc n₂ , e₂ 146 | _ ≅ᵉ _ = bot 147 | 148 | _≈_ : ∀ {a b} {α : Level a} {β : Level b} -> Univ α -> Univ β -> Prop 149 | _≈_ {α = lzero } {lzero } A₁ A₂ = A₁ ⇒ A₂ & A₂ ⇒ A₁ 150 | _≈_ {α = lsuc _} {lsuc _} A₁ A₂ = A₁ ≃ A₂ 151 | _≈_ _ _ = bot 152 | 153 | _≅ᵈ_ : ∀ {i₁ i₂ a₁ a₂} {ι₁ : Level i₁} {ι₂ : Level i₂} 154 | {α₁ : Level a₁} {α₂ : Level a₂} {I₁ : Type ι₁} {I₂ : Type ι₂} 155 | -> Desc I₁ α₁ -> Desc I₂ α₂ -> Prop 156 | var i₁ ≅ᵈ var i₂ = i₁ ≅ i₂ 157 | π A₁ D₁ ≅ᵈ π A₂ D₂ = A₁ ≈ A₂ & D₁ ≅ D₂ 158 | (D₁ ⊛ E₁) ≅ᵈ (D₂ ⊛ E₂) = D₁ ≅ᵈ D₂ & E₁ ≅ᵈ E₂ 159 | _ ≅ᵈ _ = bot 160 | 161 | _≊ᵈ_ : ∀ {i₁ i₂ a₁ a₂} {ι₁ : Level i₁} {ι₂ : Level i₂} 162 | {α₁ : Level a₁} {α₂ : Level a₂} {I₁ : Type ι₁} {I₂ : Type ι₂} 163 | -> Desc I₁ α₁ -> Desc I₂ α₂ -> Prop 164 | _≊ᵈ_ {α₁ = lzero } {lzero } D₁ D₂ = imu D₁ ≅ imu D₂ 165 | _≊ᵈ_ {α₁ = lsuc _} {lsuc _} D₁ D₂ = D₁ ≅ᵈ D₂ 166 | _≊ᵈ_ _ _ = bot 167 | 168 | bot ≃ bot = top 169 | top ≃ top = top 170 | nat ≃ nat = top 171 | enum n₁ ≃ enum n₂ = n₁ ≟ⁿ n₂ 172 | univ α₁ ≃ univ α₂ = α₁ ≟ˡ α₂ 173 | σ A₁ B₁ ≃ σ A₂ B₂ = A₁ ≈ A₂ & B₁ ≅ B₂ 174 | π A₁ B₁ ≃ π A₂ B₂ = A₂ ≈ A₁ & π A₁ λ x₁ -> π A₂ λ x₂ -> x₂ ≅ x₁ ⇒ B₁ x₁ ≈ B₂ x₂ 175 | desc I₁ α₁ ≃ desc I₂ α₂ = I₁ ≃ I₂ & α₁ ≟ˡ α₂ 176 | imu D₁ j₁ ≃ imu D₂ j₂ = D₁ ≊ᵈ D₂ & j₁ ≅ j₂ 177 | _ ≃ _ = bot 178 | 179 | _≅e_ : ∀ {i₁ i₂ a₁ a₂ b₁ b₂} {ι₁ : Level i₁} {ι₂ : Level i₂} 180 | {α₁ : Level a₁} {α₂ : Level a₂} {β₁ : Level b₁} {β₂ : Level b₂} 181 | {I₁ : Type ι₁} {I₂ : Type ι₂} {B₁ : ⟦ I₁ ⟧ -> Univ β₁} {B₂ : ⟦ I₂ ⟧ -> Univ β₂} {j₁ j₂} 182 | -> (∃ λ (D₁ : Desc I₁ α₁) -> Extend D₁ ⟦ B₁ ⟧ⁱ j₁) 183 | -> (∃ λ (D₂ : Desc I₂ α₂) -> Extend D₂ ⟦ B₂ ⟧ⁱ j₂) 184 | -> Prop 185 | 186 | _≅_ {A = bot } {bot } _ _ = top 187 | _≅_ {A = top } {top } _ _ = top 188 | _≅_ {A = nat } {nat } n₁ n₂ = n₁ ≟ⁿ n₂ 189 | _≅_ {A = enum n₁ } {enum n₂ } e₁ e₂ = n₁ , detag e₁ ≅ᵉ n₂ , detag e₂ 190 | _≅_ {A = univ α₁ } {univ α₂ } A₁ A₂ = A₁ ≈ A₂ 191 | _≅_ {A = σ A₁ B₁ } {σ A₂ B₂ } p₁ p₂ = let x₁ , y₁ = p₁ ; x₂ , y₂ = p₂ in x₁ ≅ x₂ & y₁ ≅ y₂ 192 | _≅_ {A = π A₁ B₁ } {π A₂ B₂ } f₁ f₂ = π A₁ λ x₁ -> π A₂ λ x₂ -> x₁ ≅ x₂ ⇒ f₁ x₁ ≅ f₂ x₂ 193 | _≅_ {A = desc _ _} {desc _ _} D₁ D₂ = D₁ ≅ᵈ D₂ 194 | _≅_ {A = imu D₁ _} {imu D₂ _} d₁ d₂ = D₁ , knot d₁ ≅e D₂ , knot d₂ 195 | _≅_ _ _ = bot 196 | 197 | _≅s_ : ∀ {i₁ i₂ a₁ a₂ b₁ b₂} {ι₁ : Level i₁} {ι₂ : Level i₂} 198 | {α₁ : Level a₁} {α₂ : Level a₂} {β₁ : Level b₁} {β₂ : Level b₂} 199 | {I₁ : Type ι₁} {I₂ : Type ι₂} {B₁ : ⟦ I₁ ⟧ -> Univ β₁} {B₂ : ⟦ I₂ ⟧ -> Univ β₂} 200 | -> (∃ λ (D₁ : Desc I₁ α₁) -> ⟦ D₁ ⟧ᵈ ⟦ B₁ ⟧ⁱ) 201 | -> (∃ λ (D₂ : Desc I₂ α₂) -> ⟦ D₂ ⟧ᵈ ⟦ B₂ ⟧ⁱ) 202 | -> Prop 203 | var i₁ , x₁ ≅s var i₂ , x₂ = x₁ ≅ x₂ 204 | π A₁ D₁ , f₁ ≅s π A₂ D₂ , f₂ = 205 | π A₁ λ x₁ -> π A₂ λ x₂ -> x₁ ≅ x₂ ⇒ D₁ x₁ , f₁ x₁ ≅s D₂ x₂ , f₂ x₂ 206 | (D₁ ⊛ E₁) , s₁ , t₁ ≅s (D₂ ⊛ E₂) , s₂ , t₂ = D₁ , s₁ ≅s D₂ , s₂ & E₁ , t₁ ≅s E₂ , t₂ 207 | _ ≅s _ = bot 208 | 209 | var i₁ , q₁ ≅e var i₂ , q₂ = top 210 | π A₁ D₁ , x₁ , e₁ ≅e π A₂ D₂ , x₂ , e₂ = x₁ ≅ x₂ & D₁ x₁ , e₁ ≅e D₂ x₂ , e₂ 211 | (D₁ ⊛ E₁) , s₁ , e₁ ≅e (D₂ ⊛ E₂) , s₂ , e₂ = D₁ , s₁ ≅s D₂ , s₂ & E₁ , e₁ ≅e E₂ , e₂ 212 | _ ≅e _ = bot 213 | 214 | pattern #₀ p = node (tag nothing , p) 215 | pattern #₁ p = node (tag (just nothing) , p) 216 | pattern #₂ p = node (tag (just (just nothing)) , p) 217 | pattern #₃ p = node (tag (just (just (just nothing))) , p) 218 | pattern #₄ p = node (tag (just (just (just (just nothing)))) , p) 219 | pattern #₅ p = node (tag (just (just (just (just (just nothing))))) , p) 220 | 221 | pattern !#₀ p = node (tag tt , p) 222 | pattern !#₁ p = node (tag (just tt) , p) 223 | pattern !#₂ p = node (tag (just (just tt)) , p) 224 | pattern !#₃ p = node (tag (just (just (just tt))) , p) 225 | pattern !#₄ p = node (tag (just (just (just (just tt)))) , p) 226 | pattern !#₅ p = node (tag (just (just (just (just (just tt))))) , p) 227 | 228 | pattern unit = enum 1 229 | Unit = Apply Enum 1 230 | 231 | pattern triv = tag tt 232 | 233 | pos : ∀ {a} {α : Level a} -> Desc unit α 234 | pos = var triv 235 | 236 | mu : ∀ {a} {α : Level a} -> Desc unit α -> Univ α 237 | mu D = imu D triv 238 | 239 | liftDesc : ∀ {i a b} {ι : Level i} {α : Level a} {β : Level b} {I : Type ι} .{{_ : a ≤ₘ b}} 240 | -> Desc I α -> Desc I β 241 | liftDesc (var i) = var i 242 | liftDesc {b = b} {{q₁}} (π {c} {{q₂}} A D) = 243 | π {{pright (pcong (c ⊔ₘ_) q₁) (ptrans (pcong (b ⊔ₘ_) q₂) q₁)}} A λ x -> liftDesc (D x) 244 | liftDesc (D ⊛ E) = liftDesc D ⊛ liftDesc E 245 | 246 | var-inj : ∀ {i b} {ι : Level i} {I : Type ι} {β : Level b} {j₁ j₂ : ⟦ I ⟧} 247 | -> var {β = β} j₁ ≡ var j₂ -> j₁ ≡ j₂ 248 | var-inj prefl = prefl 249 | 250 | ⊛-inj : ∀ {i b} {ι : Level i} {I : Type ι} {β : Level b} {D₁ D₂ E₁ E₂ : Desc I β} 251 | -> (D₁ ⊛ E₁) ≡ (D₂ ⊛ E₂) -> D₁ ≡ D₂ × E₁ ≡ E₂ 252 | ⊛-inj prefl = prefl , prefl 253 | 254 | node-inj : ∀ {i a} {α : Level a} {ι : Level i} {I : Type ι} 255 | {D : Desc I α} {j} {e₁ e₂ : Extend D (μ D) j} 256 | -> node {D = D} e₁ ≡ node e₂ -> e₁ ≡ e₂ 257 | node-inj prefl = prefl 258 | -------------------------------------------------------------------------------- /Data.agda: -------------------------------------------------------------------------------- 1 | module OTT.Data where 2 | 3 | open import OTT.Data.Lift public 4 | open import OTT.Data.Fin public 5 | open import OTT.Data.Maybe public 6 | open import OTT.Data.Sum public 7 | open import OTT.Data.W public 8 | open import OTT.Data.Vec public 9 | -------------------------------------------------------------------------------- /Data/Fin.agda: -------------------------------------------------------------------------------- 1 | module OTT.Data.Fin where 2 | 3 | open import OTT.Main 4 | 5 | fin : ℕ -> Type₀ 6 | fin = icmu 7 | $ (π nat λ n -> var (suc n)) 8 | ∷ (π nat λ n -> var n ⊛ var (suc n)) 9 | ∷ [] 10 | 11 | Fin : ℕ -> Set 12 | Fin n = ⟦ fin n ⟧ 13 | 14 | pattern fzeroₑ {n} q = #₀ (n , q) 15 | pattern fsucₑ {n} q i = !#₁ (n , i , q) 16 | 17 | fzero : ∀ {n} -> Fin (suc n) 18 | fzero {n} = fzeroₑ {n} (refl (suc n)) 19 | 20 | fsuc : ∀ {n} -> Fin n -> Fin (suc n) 21 | fsuc {n} = fsucₑ {n} (refl (suc n)) 22 | 23 | gelimFin : ∀ {n π} 24 | -> (P : ∀ {n} -> Fin n -> Set π) 25 | -> (∀ {n m} {i : Fin n} -> (q : ⟦ suc n ≅ m ⟧) -> P i -> P {m} (fsucₑ q i)) 26 | -> (∀ {n m} -> (q : ⟦ suc n ≅ m ⟧) -> P {m} (fzeroₑ q)) 27 | -> (i : Fin n) 28 | -> P i 29 | gelimFin P f x = gelim P (fromTuple ((λ _ _ -> x) , (λ _ _ r _ q -> f q r))) 30 | 31 | foldFin : ∀ {n π} {P : Set π} -> (P -> P) -> P -> Fin n -> P 32 | foldFin f x = gelimFin _ (const f) (const x) 33 | 34 | fromFin : ∀ {n} -> Fin n -> ℕ 35 | fromFin = foldFin suc 0 36 | 37 | elimFin′ : ∀ {n π} 38 | -> (P : ∀ n -> Set π) 39 | -> (∀ {n} {i : Fin n} -> P (fromFin i) -> P (suc (fromFin i))) 40 | -> P 0 41 | -> (i : Fin n) 42 | -> P (fromFin i) 43 | elimFin′ P f x = gelimFin (P ∘ fromFin) (λ {n m i} _ -> f {i = i}) (const x) 44 | 45 | elimFin : ∀ {n p} {π : Level p} 46 | -> (P : ∀ {n} -> Fin n -> Univ π) 47 | -> (∀ {n} {i : Fin n} -> ⟦ P i ⇒ P (fsuc i) ⟧) 48 | -> (∀ {n} -> ⟦ P {suc n} fzero ⟧) 49 | -> (i : Fin n) 50 | -> ⟦ P i ⟧ 51 | elimFin P f x = elim P (fromTuple ((λ _ -> x) , (λ _ _ -> f))) 52 | 53 | 54 | 55 | private 56 | module Test where 57 | -- _+ᶠ_ : ∀ {n m} -> (i : Fin n) -> Fin m -> Fin (fromFin i + m) 58 | -- i +ᶠ j = elimFin′ (λ n -> Fin (n + _)) fsuc j i 59 | 60 | _+ᶠ_ : ∀ {n m} -> (i : Fin n) -> Fin m -> Fin (fromFin i + m) 61 | i +ᶠ j = elimFin (λ i -> fin (fromFin i + _)) fsuc j i 62 | 63 | postulate 64 | n m : ℕ 65 | 66 | test : ⟦ fromFin ((Fin (3 + n) ∋ fsuc (fsuc fzero)) +ᶠ (Fin (2 + m) ∋ fsuc fzero)) ≅ 3 ⟧ 67 | test = tt 68 | -------------------------------------------------------------------------------- /Data/Lift.agda: -------------------------------------------------------------------------------- 1 | module OTT.Data.Lift where 2 | 3 | open import OTT.Main 4 | 5 | Lift : ∀ {a b} {α : Level a} -> (β : Level b) -> Univ α -> Univ (α ⊔ β) 6 | Lift β A = mu (A ⇨ pos) 7 | 8 | -- lift : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} -> ⟦ A ⟧ -> ⟦ Lift β A ⟧ 9 | pattern lift x = node (x , tt) 10 | 11 | lower : ∀ {a b} {α : Level a} {A : Univ α} -> (β : Level b) -> ⟦ Lift β A ⟧ -> ⟦ A ⟧ 12 | lower β (lift x) = x 13 | 14 | private 15 | β = natToLevel 2 16 | 17 | lnat : Univ# 2 18 | lnat = Lift β nat 19 | 20 | ln : ⟦ lnat ⟧ 21 | ln = lift 0 22 | 23 | n : ℕ 24 | n = lower β ln 25 | -------------------------------------------------------------------------------- /Data/List.agda: -------------------------------------------------------------------------------- 1 | module OTT.Data.List where 2 | 3 | open import OTT.Core 4 | open import OTT.Coerce 5 | open import OTT.Function.Pi 6 | open import OTT.Function.Elim 7 | 8 | infixr 5 _∷_ 9 | 10 | -- It was, but this breaks inference with Agda 2.5.2. 11 | -- list : ∀ {a} {α : Level a} -> Univ α -> Type₋₁ α 12 | -- list A = mu $ π (enum 2) (fromTuple (pos , (A ⇨ pos ⊛ pos))) 13 | 14 | list : ∀ {a} {α : Level a} -> Univ α -> Type₋₁ α 15 | list A = mu $ π (enum 2) λ 16 | { (tag nothing) -> pos 17 | ; (tag (just _)) -> A ⇨ pos ⊛ pos 18 | } 19 | 20 | List : ∀ {a} {α : Level a} -> Univ α -> Set 21 | List A = ⟦ list A ⟧ 22 | 23 | pattern [] = #₀ tt 24 | pattern _∷_ x xs = !#₁ (x , xs , tt) 25 | 26 | elimList : ∀ {a π} {α : Level a} {A : Univ α} 27 | -> (P : List A -> Set π) -> (∀ {xs} x -> P xs -> P (x ∷ xs)) -> P [] -> ∀ xs -> P xs 28 | elimList P f z = elim′ P (fromTuple (z , λ x xs -> f x)) 29 | 30 | foldList : ∀ {a β} {α : Level a} {A : Univ α} {B : Set β} 31 | -> (⟦ A ⟧ -> B -> B) -> B -> List A -> B 32 | foldList f = elimList _ f 33 | 34 | length : ∀ {a} {α : Level a} {A : Univ α} -> List A -> ℕ 35 | length = foldList (const suc) 0 36 | 37 | fromList : ∀ {a} {α : Level a} {A : Univ α} -> (xs : List A) -> Apply Enum (length xs) -> ⟦ A ⟧ 38 | fromList xs = go xs ∘ detag where 39 | go : ∀ {a} {α : Level a} {A : Univ α} -> (xs : List A) -> Enum (length xs) -> ⟦ A ⟧ 40 | go [] () 41 | go (x ∷ []) tt = x 42 | go (x ∷ y ∷ xs) nothing = x 43 | go (x ∷ y ∷ xs) (just e) = go (y ∷ xs) e 44 | 45 | icmu : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} 46 | -> List (desc I (lsuc α)) -> ⟦ I ⟧ -> Type α 47 | icmu Ds = imu $ π (enum (length Ds)) (fromList Ds) 48 | 49 | cmu : ∀ {a} {α : Level a} -> List (desc unit (lsuc α)) -> Type α 50 | cmu Ds = icmu Ds triv 51 | 52 | icmu₋₁ : ∀ {i a} {ι : Level i} {I : Type ι} 53 | -> (α : Level a) -> List (desc₋₁ I α) -> ⟦ I ⟧ -> Type₋₁ α 54 | icmu₋₁ α Ds = imu $ π (enum (length Ds)) (fromList Ds) 55 | 56 | cmu₋₁ : ∀ {a} -> (α : Level a) -> List (desc₋₁ unit α) -> Type₋₁ α 57 | cmu₋₁ α Ds = icmu₋₁ α Ds triv 58 | -------------------------------------------------------------------------------- /Data/Maybe.agda: -------------------------------------------------------------------------------- 1 | module OTT.Data.Maybe where 2 | 3 | open import OTT.Main hiding (Maybe; nothing; just) 4 | 5 | maybe : ∀ {a} {α : Level a} -> Univ α -> Type₋₁ α 6 | maybe {α = α} A = cmu₋₁ α 7 | $ pos 8 | ∷ (A ⇨ pos) 9 | ∷ [] 10 | 11 | Maybe : ∀ {a} {α : Level a} -> Univ α -> Set 12 | Maybe A = ⟦ maybe A ⟧ 13 | 14 | pattern nothing = #₀ tt 15 | pattern just x = !#₁ (x , tt) 16 | 17 | elimMaybe : ∀ {a π} {α : Level a} {A : Univ α} {P : Maybe A -> Set π} 18 | -> (∀ x -> P (just x)) -> P nothing -> ∀ a -> P a 19 | elimMaybe f z = elim′ _ (fromTuple (z , f)) 20 | 21 | fmap : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} 22 | -> ⟦ (A ⇒ B) ⇒ maybe A ⇒ maybe B ⟧ 23 | fmap f = elimMaybe (just ∘ f) nothing 24 | -------------------------------------------------------------------------------- /Data/Sum.agda: -------------------------------------------------------------------------------- 1 | module OTT.Data.Sum where 2 | 3 | open import OTT.Main 4 | 5 | infixr 3 _⊕_ _⊎_ 6 | 7 | _⊕_ : ∀ {a b} {α : Level a} {β : Level b} -> Univ α -> Univ β -> Type₋₁ (α ⊔ β) 8 | _⊕_ {α = α} {β} A B = cmu₋₁ (α ⊔ β) 9 | $ (A ⇨ pos) 10 | ∷ (B ⇨ pos) 11 | ∷ [] 12 | 13 | _⊎_ : ∀ {a b} {α : Level a} {β : Level b} -> Univ α -> Univ β -> Set 14 | A ⊎ B = ⟦ A ⊕ B ⟧ 15 | 16 | pattern inj₁ x = #₀ (x , tt) 17 | pattern inj₂ y = !#₁ (y , tt) 18 | 19 | [_,_] : ∀ {a b π} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} {P : A ⊎ B -> Set π} 20 | -> (∀ x -> P (inj₁ x)) -> (∀ y -> P (inj₂ y)) -> ∀ s -> P s 21 | [ f , g ] = elim′ _ (fromTuple (f , g)) 22 | 23 | smap : ∀ {a₁ a₂ b₁ b₂} {α₁ : Level a₁} {α₂ : Level a₂} {β₁ : Level b₁} {β₂ : Level b₂} 24 | {A₁ : Univ α₁} {A₂ : Univ α₂} {B₁ : Univ β₁} {B₂ : Univ β₂} 25 | -> ⟦ (A₁ ⇒ A₂) ⇒ (B₁ ⇒ B₂) ⇒ A₁ ⊕ B₁ ⇒ A₂ ⊕ B₂ ⟧ 26 | smap f g = [ inj₁ ∘ f , inj₂ ∘ g ] 27 | -------------------------------------------------------------------------------- /Data/Vec.agda: -------------------------------------------------------------------------------- 1 | module OTT.Data.Vec where 2 | 3 | open import OTT.Main 4 | open import OTT.Data.Fin 5 | 6 | infixr 5 _∷ᵥ_ 7 | 8 | vec : ∀ {a} {α : Level a} -> Univ α -> ℕ -> Type₋₁ α 9 | vec {α = α} A = icmu₋₁ α 10 | $ var 0 11 | ∷ (π nat λ n -> A ⇨ var n ⊛ var (suc n)) 12 | ∷ [] 13 | 14 | Vec : ∀ {a} {α : Level a} -> Univ α -> ℕ -> Set 15 | Vec A n = ⟦ vec A n ⟧ 16 | 17 | pattern vnilₑ q = #₀ q 18 | pattern vconsₑ {n} q x xs = !#₁ (n , x , xs , q) 19 | 20 | []ᵥ : ∀ {a} {α : Level a} {A : Univ α} -> Vec A 0 21 | []ᵥ = vnilₑ (refl 0) 22 | 23 | _∷ᵥ_ : ∀ {n a} {α : Level a} {A : Univ α} -> ⟦ A ⇒ vec A n ⇒ vec A (suc n) ⟧ 24 | _∷ᵥ_ {n} = vconsₑ (refl (suc n)) 25 | 26 | gelimVec : ∀ {n a π} {α : Level a} {A : Univ α} 27 | -> (P : ∀ {n} -> Vec A n -> Set π) 28 | -> (∀ {n m} {xs : Vec A n} 29 | -> (q : ⟦ suc n ≅ m ⟧) -> (x : ⟦ A ⟧) -> P xs -> P {m} (vconsₑ q x xs)) 30 | -> (∀ {m} -> (q : ⟦ 0 ≅ m ⟧) -> P {m} (vnilₑ q)) 31 | -> (xs : Vec A n) 32 | -> P xs 33 | gelimVec P f z = gelim P (fromTuple ((λ _ -> z) , (λ n x xs r _ q -> f q x r))) 34 | 35 | foldVec : ∀ {n a π} {α : Level a} {A : Univ α} {P : Set π} 36 | -> (⟦ A ⟧ -> P -> P) -> P -> Vec A n -> P 37 | foldVec f z = gelimVec _ (const f) (const z) 38 | 39 | fromVec : ∀ {n a} {α : Level a} {A : Univ α} -> Vec A n -> List A 40 | fromVec = foldVec _∷_ [] 41 | 42 | elimVec′ : ∀ {n a π} {α : Level a} {A : Univ α} 43 | -> (P : List A -> Set π) 44 | -> (∀ {n} {xs : Vec A n} -> (x : ⟦ A ⟧) -> P (fromVec xs) -> P (x ∷ fromVec xs)) 45 | -> P [] 46 | -> (xs : Vec A n) 47 | -> P (fromVec xs) 48 | elimVec′ P f z = gelimVec (P ∘ fromVec) (λ {n m xs} _ -> f {xs = xs}) (const z) 49 | 50 | elimVec : ∀ {n a p} {α : Level a} {π : Level p} {A : Univ α} 51 | -> (P : ∀ {n} -> Vec A n -> Univ π) 52 | -> (∀ {n} {xs : Vec A n} -> (x : ⟦ A ⟧) -> ⟦ P xs ⇒ P (x ∷ᵥ xs) ⟧) 53 | -> ⟦ P []ᵥ ⟧ 54 | -> (xs : Vec A n) 55 | -> ⟦ P xs ⟧ 56 | elimVec P f z = elim P (fromTuple (z , λ n x xs -> f x)) 57 | 58 | splitVec : ∀ {n a β} {α : Level a} {A : Univ α} {B : Set β} 59 | -> (⟦ A ⟧ -> Vec A n -> B) -> Vec A (suc n) -> B 60 | splitVec k (vnilₑ ()) 61 | splitVec k (vconsₑ q x xs) = k x (subst (vec _) q xs) 62 | 63 | module LookupViaElim where 64 | -- open import OTT.Property.Eq 65 | 66 | -- elimIndFin : ∀ {n π} 67 | -- -> (P : ∀ n -> Set π) 68 | -- -> (∀ {n} -> P n -> P (suc n)) 69 | -- -> (∀ {n} -> P (suc n)) 70 | -- -> Fin n 71 | -- -> P n 72 | -- elimIndFin P f x = gelimFin (λ {n} _ -> P n) 73 | -- (λ q -> psubst P (observe q) ∘ f) 74 | -- (λ q -> psubst P (observe q) x) 75 | 76 | elimIndFin : ∀ {n p} {π : Level p} 77 | -> (P : ∀ n -> Univ π) 78 | -> (∀ {n} -> ⟦ P n ⇒ P (suc n) ⟧) 79 | -> (∀ {n} -> ⟦ P (suc n) ⟧) 80 | -> Fin n 81 | -> ⟦ P n ⟧ 82 | elimIndFin P f x = gelimFin (λ {n} _ -> ⟦ P n ⟧) 83 | (λ q -> subst P q ∘ f) 84 | (λ q -> subst P q x) 85 | 86 | vlookup : ∀ {n a} {α : Level a} {A : Univ α} -> Fin n -> Vec A n -> ⟦ A ⟧ 87 | vlookup {A = A} = elimIndFin (λ n -> vec A n ⇒ A) 88 | (splitVec ∘ const) 89 | (splitVec const) 90 | 91 | vlookupₑ : ∀ {n m a} {α : Level a} {A : Univ α} -> ⟦ n ≅ m ⇒ fin n ⇒ vec A m ⇒ A ⟧ 92 | vlookupₑ p (fzeroₑ q) (vconsₑ r x xs) = x 93 | vlookupₑ {n} {m} p (fsucₑ {n′} q i) (vconsₑ {m′} r x xs) = 94 | vlookupₑ (left (suc n′) {m} {suc m′} (trans (suc n′) {n} {m} q p) r) i xs 95 | vlookupₑ {n} {m} p (fzeroₑ {n′} q) (vnilₑ r) = 96 | ⊥-elim $ left (suc n′) {m} {0} (trans (suc n′) {n} {m} q p) r 97 | vlookupₑ {n} {m} p (fsucₑ {n′} q i) (vnilₑ r) = 98 | ⊥-elim $ left (suc n′) {m} {0} (trans (suc n′) {n} {m} q p) r 99 | 100 | vlookup : ∀ {n a} {α : Level a} {A : Univ α} -> Fin n -> Vec A n -> ⟦ A ⟧ 101 | vlookup {n} = vlookupₑ (refl n) 102 | -------------------------------------------------------------------------------- /Data/W.agda: -------------------------------------------------------------------------------- 1 | module OTT.Data.W where 2 | 3 | open import OTT.Main 4 | 5 | w : ∀ {a b} {α : Level a} {β : Level b} -> (A : Univ α) -> (⟦ A ⟧ -> Univ β) -> Univ (α ⊔ β) 6 | w A B = mu (π A λ x -> (B x ⇨ pos) ⊛ pos) 7 | 8 | W : ∀ {a b} {α : Level a} {β : Level b} -> (A : Univ α) -> (⟦ A ⟧ -> Univ β) -> Set 9 | W A B = ⟦ w A B ⟧ 10 | 11 | pattern sup x g = node (x , g , tt) 12 | 13 | elimW : ∀ {a b π} {α : Level a} {β : Level b} {A : Univ α} {B : ⟦ A ⟧ -> Univ β} 14 | -> (P : W A B -> Set π) 15 | -> (∀ {x} {g : ⟦ B x ⟧ -> W A B} -> (∀ y -> P (g y)) -> P (sup x g)) 16 | -> ∀ w 17 | -> P w 18 | elimW P h = elim′ P λ _ _ -> h 19 | -------------------------------------------------------------------------------- /Examples/Favourite.agda: -------------------------------------------------------------------------------- 1 | -- An example from https://pigworker.wordpress.com/2015/01/08/observational-type-theory-delivery/ 2 | 3 | module OTT.Examples.Favourite where 4 | 5 | open import OTT.Main 6 | 7 | c42 : ℕ -> ℕ 8 | c42 = const 42 9 | 10 | favourite : (ℕ -> ℕ) -> Prop 11 | favourite = imu (var c42) 12 | 13 | Favourite : (ℕ -> ℕ) -> Set 14 | Favourite f = ⟦ favourite f ⟧ 15 | 16 | pattern faveₑ q = node q 17 | 18 | fave : Favourite c42 19 | fave = faveₑ (refl c42) 20 | 21 | deep : ℕ -> ℕ 22 | deep 0 = 42 23 | deep (suc n) = deep n 24 | 25 | deep-42 : ∀ n -> ⟦ 42 ≟ⁿ deep n ⟧ 26 | deep-42 0 = tt 27 | deep-42 (suc n) = deep-42 n 28 | 29 | fave′ : Favourite deep 30 | fave′ = faveₑ (λ _ n _ -> deep-42 n) 31 | -------------------------------------------------------------------------------- /Function/Elim.agda: -------------------------------------------------------------------------------- 1 | module OTT.Function.Elim where 2 | 3 | open import OTT.Core 4 | open import OTT.Coerce 5 | 6 | Hyp : ∀ {i b γ} {ι : Level i} {β : Level b} {I : Type ι} {B} 7 | -> (∀ {i} -> B i -> Set γ) -> (D : Desc I β) -> ⟦ D ⟧ᵈ B -> Set γ 8 | Hyp C (var i) y = C y 9 | Hyp C (π A D) f = ∀ x -> Hyp C (D x) (f x) 10 | Hyp C (D ⊛ E) (x , y) = Hyp C D x × Hyp C E y 11 | 12 | GElim : ∀ {i b γ} {ι : Level i} {β : Level b} {I : Type ι} {B} 13 | -> (∀ {i} -> B i -> Set γ) -> (D : Desc I β) -> (∀ {j} -> Extend D B j -> B j) -> Set γ 14 | GElim C (var i) k = ∀ j -> (q : ⟦ i ≅ j ⟧) -> C (k q) 15 | GElim C (π A D) k = ∀ x -> GElim C (D x) (k ∘ _,_ x) 16 | GElim C (D ⊛ E) k = ∀ x -> Hyp C D x -> GElim C E (k ∘ _,_ x) 17 | 18 | module _ {i b γ} {ι : Level i} {β : Level b} {I : Type ι} {D : Desc I β} 19 | (C : ∀ {j} -> μ D j -> Set γ) (h : GElim C D node) where 20 | mutual 21 | elimExtend : ∀ {j} 22 | -> (E : Desc I β) {k : ∀ {j} -> Extend E (μ D) j -> μ D j} 23 | -> GElim C E k 24 | -> (e : Extend E (μ D) j) 25 | -> C (k e) 26 | elimExtend (var i) z q = z _ q 27 | elimExtend (π A D) h (x , e) = elimExtend (D x) (h x) e 28 | elimExtend (D ⊛ E) h (r , e) = elimExtend E (h r (hyp D r)) e 29 | 30 | hyp : ∀ E -> (d : ⟦ E ⟧ᵈ (μ D)) -> Hyp C E d 31 | hyp (var i) d = gelim d 32 | hyp (π A D) f = λ x -> hyp (D x) (f x) 33 | hyp (D ⊛ E) (r , s) = hyp D r , hyp E s 34 | 35 | gelim : ∀ {j} -> (d : μ D j) -> C d 36 | gelim (node e) = elimExtend D h e 37 | 38 | Elim : ∀ {i b γ} {ι : Level i} {β : Level b} {I : Type ι} {B} 39 | -> (∀ {i} -> B i -> Set γ) -> (D : Desc I β) -> (∀ {j} -> Extend D B j -> B j) -> Set γ 40 | Elim C (var i) k = C (k (refl i)) 41 | Elim C (π A D) k = ∀ x -> Elim C (D x) (k ∘ _,_ x) 42 | Elim C (D ⊛ E) k = ∀ x -> Hyp C D x -> Elim C E (k ∘ _,_ x) 43 | 44 | embedElim : ∀ {i b c} {ι : Level i} {β : Level b} {γ : Level c} {I : Type ι} 45 | {B : ⟦ I ⟧ -> Univ β} {C : ∀ {i} -> ⟦ B i ⟧ -> Univ γ} 46 | -> (D : Desc I β) {k : ∀ {j} -> Extend D ⟦ B ⟧ⁱ j -> ⟦ B j ⟧} 47 | -> Elim ⟦ C ⟧ⁱ D k 48 | -> GElim ⟦ C ⟧ⁱ D k 49 | embedElim {C = C} (var i) {k} z = λ j q -> J (λ j -> C ∘ k {j}) z q 50 | embedElim (π A D) h = λ x -> embedElim (D x) (h x) 51 | embedElim (D ⊛ E) h = λ x f -> embedElim E (h x f) 52 | 53 | elim : ∀ {i b c} {ι : Level i} {β : Level b} {γ : Level c} {I : Type ι} {D : Desc I β} {j} 54 | -> (C : ∀ {j} -> μ D j -> Univ γ) -> (h : Elim ⟦ C ⟧ⁱ D node) -> (d : μ D j) -> ⟦ C d ⟧ 55 | elim {D = D} C h = gelim ⟦ C ⟧ⁱ (embedElim D h) 56 | 57 | embedElim′ : ∀ {b γ} {β : Level b} {B} {C : B -> Set γ} 58 | -> (D : Desc unit β) {k : ∀ {j} -> Extend D (const B) j -> B} 59 | -> Elim C D k 60 | -> GElim C D k 61 | embedElim′ (var i) z = λ j q -> z 62 | embedElim′ (π A D) h = λ x -> embedElim′ (D x) (h x) 63 | embedElim′ (D ⊛ E) h = λ x f -> embedElim′ E (h x f) 64 | 65 | elim′ : ∀ {b γ} {β : Level b} {D : Desc unit β} {j} 66 | -> (C : μ D triv -> Set γ) -> (h : Elim C D node) -> (d : μ D j) -> C d 67 | elim′ {D = D} C h = gelim C (embedElim′ D h) 68 | -------------------------------------------------------------------------------- /Function/Erase.agda: -------------------------------------------------------------------------------- 1 | module OTT.Function.Erase where 2 | 3 | open import OTT.Core 4 | 5 | data UnitView : ∀ {a} {α : Level a} -> Univ α -> Set where 6 | yes-unit : UnitView unit 7 | no-unit : ∀ {a} {α : Level a} {A : Univ α} -> UnitView A 8 | 9 | unitView : ∀ {a} {α : Level a} -> (A : Univ α) -> UnitView A 10 | unitView unit = yes-unit 11 | unitView A = no-unit 12 | 13 | Erase : ∀ {i b} {ι : Level i} {I : Type ι} {β : Level b} -> Desc I β -> Desc unit β 14 | Erase (var i) = pos 15 | Erase (π A D) = π A (λ x -> Erase (D x)) 16 | Erase (D ⊛ E) = Erase D ⊛ Erase E 17 | 18 | module _ {i b} {ι : Level i} {I : Type ι} {β : Level b} {D : Desc I β} where 19 | mutual 20 | eraseSem : (E : Desc I β) -> ⟦ E ⟧ᵈ (μ D) -> ⟦ Erase E ⟧ᵈ (μ (Erase D)) 21 | eraseSem (var i) d = erase d 22 | eraseSem (π A D) f = λ x -> eraseSem (D x) (f x) 23 | eraseSem (D ⊛ E) (r , s) = eraseSem D r , eraseSem E s 24 | 25 | eraseExtend : ∀ {j} 26 | -> (E : Desc I β) -> Extend E (μ D) j -> Extend (Erase E) (μ (Erase D)) triv 27 | eraseExtend (var i) q = tt 28 | eraseExtend (π A D) (x , e) = x , eraseExtend (D x) e 29 | eraseExtend (D ⊛ E) (r , e) = eraseSem D r , eraseExtend E e 30 | 31 | erase : ∀ {j} -> μ D j -> μ (Erase D) triv 32 | erase (node e) = node (eraseExtend D e) 33 | 34 | module _ {i b} {ι : Level i} {I : Type ι} {β : Level b} {B} where 35 | mutual 36 | eraseConstSem : (D : Desc I β) -> ⟦ D ⟧ᵈ (const B) -> ⟦ Erase D ⟧ᵈ (const B) 37 | eraseConstSem (var i) y = y 38 | eraseConstSem (π A D) f = λ x -> eraseConstSem (D x) (f x) 39 | eraseConstSem (D ⊛ E) (x , y) = eraseConstSem D x , eraseConstSem E y 40 | 41 | eraseConstExtend : ∀ {j} 42 | -> (D : Desc I β) -> Extend D (const B) j -> Extend (Erase D) (const B) triv 43 | eraseConstExtend (var i) q = tt 44 | eraseConstExtend (π A D) (x , e) = x , eraseConstExtend (D x) e 45 | eraseConstExtend (D ⊛ E) (x , e) = eraseConstSem D x , eraseConstExtend E e 46 | 47 | module _ {b} {β : Level b} {B : Unit -> Set} where 48 | mutual 49 | uneraseSem : (D : Desc unit β) -> ⟦ Erase D ⟧ᵈ B -> ⟦ D ⟧ᵈ B 50 | uneraseSem (var i) y = y 51 | uneraseSem (π A D) f = λ x -> uneraseSem (D x) (f x) 52 | uneraseSem (D ⊛ E) (x , y) = uneraseSem D x , uneraseSem E y 53 | 54 | uneraseExtend : ∀ {j} 55 | -> (D : Desc unit β) -> Extend (Erase D) B j -> Extend D B triv 56 | uneraseExtend (var i) q = tt 57 | uneraseExtend (π A D) (x , e) = x , uneraseExtend (D x) e 58 | uneraseExtend (D ⊛ E) (x , e) = uneraseSem D x , uneraseExtend E e 59 | 60 | CheckErase : ∀ {i b} {ι : Level i} {I : Type ι} {β : Level b} -> Desc I β -> Desc unit β 61 | CheckErase {I = I} D with unitView I 62 | ... | yes-unit = D 63 | ... | no-unit = Erase D 64 | 65 | checkErase : ∀ {i b} {ι : Level i} {I : Type ι} {β : Level b} {D : Desc I β} {j} 66 | -> μ D j -> μ (CheckErase D) triv 67 | checkErase {I = I} d with unitView I 68 | ... | yes-unit = d 69 | ... | no-unit = erase d 70 | -------------------------------------------------------------------------------- /Function/Pi.agda: -------------------------------------------------------------------------------- 1 | module OTT.Function.Pi where 2 | 3 | open import OTT.Core 4 | 5 | Tabulate : ∀ {n α} -> (Apply Enum n -> Set α) -> Set α 6 | Tabulate {α = α} F = go _ (F ∘ tag) where 7 | go : ∀ n -> (Enum n -> Set α) -> Set α 8 | go 0 F = ⊤ 9 | go 1 F = F tt 10 | go (suc (suc n)) F = F nothing × go (suc n) (F ∘ just) 11 | 12 | tabulate : ∀ {n α} {F : Apply Enum n -> Set α} -> (∀ e -> F e) -> Tabulate F 13 | tabulate {α = α} f = go _ (f ∘ tag) where 14 | go : ∀ n {F : Enum n -> Set α} -> (∀ e -> F e) -> Tabulate (F ∘ detag) 15 | go 0 f = tt 16 | go 1 f = f tt 17 | go (suc (suc n)) f = f nothing , go (suc n) (f ∘ just) 18 | 19 | lookup : ∀ {n α} {F : Apply Enum n -> Set α} -> (e : Apply Enum n) -> Tabulate F -> F e 20 | lookup {α = α} e xs = go _ (detag e) xs where 21 | go : ∀ n {F : Enum n -> Set α} -> (e : Enum n) -> Tabulate (F ∘ detag) -> F e 22 | go 0 () 23 | go 1 tt x = x 24 | go (suc (suc n)) nothing (x , xs) = x 25 | go (suc (suc n)) (just e) (x , xs) = go (suc n) e xs 26 | 27 | fromTuple : ∀ {n α} {F : Apply Enum n -> Set α} -> Tabulate F -> (e : Apply Enum n) -> F e 28 | fromTuple = flip lookup 29 | 30 | -- I tried to make a type class from this, but failed because of size issues. 31 | -- Is there any reason to bother with `desc`? 32 | Pi : ∀ {a β} {α : Level a} -> (A : Univ α) -> (⟦ A ⟧ -> Set β) -> Set β 33 | Pi bot F = ⊤ 34 | Pi top F = F tt 35 | Pi (enum n) F = Tabulate F 36 | Pi (σ A B) F = Pi A λ x -> Pi (B x) λ y -> F (x , y) 37 | Pi _ F = ∀ {x} -> F x 38 | 39 | lam : ∀ {a β} {α : Level a} {A : Univ α} {B : ⟦ A ⟧ -> Set β} -> (∀ x -> B x) -> Pi A B 40 | lam {A = bot } f = tt 41 | lam {A = top } f = f tt 42 | lam {A = enum n} f = tabulate f 43 | lam {A = σ A B } g = lam λ x -> lam λ y -> g (x , y) 44 | lam {A = nat } f = f _ 45 | lam {A = univ _ } f = f _ 46 | lam {A = π _ _ } f = f _ 47 | lam {A = desc _ _} f = f _ 48 | lam {A = imu _ _ } f = f _ 49 | 50 | apply : ∀ {a β} {α : Level a} {A : Univ α} {F : ⟦ A ⟧ -> Set β} -> Pi A F -> (x : ⟦ A ⟧) -> F x 51 | apply {A = bot } f () 52 | apply {A = top } x tt = x 53 | apply {A = enum n} xs e = lookup e xs 54 | apply {A = σ A B } g (x , y) = apply (apply g x) y 55 | apply {A = nat } y x = y 56 | apply {A = univ _ } y x = y 57 | apply {A = π _ _ } y x = y 58 | apply {A = desc _ _} y x = y 59 | apply {A = imu _ _ } y x = y 60 | -------------------------------------------------------------------------------- /Lib/Decidable.agda: -------------------------------------------------------------------------------- 1 | module OTT.Lib.Decidable where 2 | 3 | open import Relation.Nullary public 4 | open import Relation.Nullary.Decidable public 5 | open import Relation.Binary using (Decidable) public 6 | 7 | open import Function 8 | open import Relation.Nullary 9 | open import Relation.Binary.PropositionalEquality 10 | open import Data.Product 11 | 12 | open import OTT.Lib.Heteroindexed 13 | 14 | infixl 10 _% 15 | infix 3 _#_ 16 | 17 | _% = _∘_ 18 | 19 | IsSet : ∀ {α} -> Set α -> Set α 20 | IsSet A = Decidable {A = A} _≡_ 21 | 22 | _#_ : ∀ {α} {A : Set α} -> A -> A -> Set α 23 | x # y = Dec (x ≡ y) 24 | 25 | delim : ∀ {α π} {A : Set α} {P : Dec A -> Set π} 26 | -> (∀ x -> P (yes x)) -> (∀ c -> P (no c)) -> (d : Dec A) -> P d 27 | delim f g (yes x) = f x 28 | delim f g (no c) = g c 29 | 30 | drec : ∀ {α β} {A : Set α} {B : Set β} -> (A -> B) -> (¬ A -> B) -> Dec A -> B 31 | drec = delim 32 | 33 | dmap : ∀ {α β} {A : Set α} {B : Set β} -> (A -> B) -> (¬ A -> ¬ B) -> Dec A -> Dec B 34 | dmap f g = drec (yes ∘ f) (no ∘ g) 35 | 36 | sumM2 : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 37 | -> (A -> Dec C) -> (B -> Dec C) -> (¬ A -> ¬ B -> Dec C) -> Dec A -> Dec B -> Dec C 38 | sumM2 f g h d e = drec f (λ c -> drec g (h c) e) d 39 | 40 | prodM2 : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 41 | -> (A -> B -> Dec C) -> (¬ A -> Dec C) -> (¬ B -> Dec C) -> Dec A -> Dec B -> Dec C 42 | prodM2 h f g d e = drec (λ x -> drec (h x) g e) f d 43 | 44 | sumF2 : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 45 | -> (A -> C) -> (B -> C) -> (¬ A -> ¬ B -> ¬ C) -> Dec A -> Dec B -> Dec C 46 | sumF2 f g h = sumM2 (yes ∘ f) (yes ∘ g) (no % ∘ h) 47 | 48 | prodF2 : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 49 | -> (A -> B -> C) -> (¬ A -> ¬ C) -> (¬ B -> ¬ C) -> Dec A -> Dec B -> Dec C 50 | prodF2 h f g = prodM2 (yes % ∘ h) (no ∘ f) (no ∘ g) 51 | 52 | dcong : ∀ {α β} {A : Set α} {B : Set β} {x y} 53 | -> (f : A -> B) -> (f x ≡ f y -> x ≡ y) -> x # y -> f x # f y 54 | dcong f inj = dmap (cong f) (_∘ inj) 55 | 56 | dcong₂ : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} {x₁ x₂ y₁ y₂} 57 | -> (f : A -> B -> C) 58 | -> (f x₁ y₁ ≡ f x₂ y₂ -> x₁ ≡ x₂ × y₁ ≡ y₂) 59 | -> x₁ # x₂ 60 | -> y₁ # y₂ 61 | -> f x₁ y₁ # f x₂ y₂ 62 | dcong₂ f inj = prodF2 (cong₂ f) (λ c -> c ∘ proj₁ ∘ inj) (λ c -> c ∘ proj₂ ∘ inj) 63 | 64 | dhcong₂ : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : Set γ} {x₁ x₂ y₁ y₂} 65 | -> (f : ∀ x -> B x -> C) 66 | -> (f x₁ y₁ ≡ f x₂ y₂ -> [ B ] y₁ ≅ y₂) 67 | -> x₁ # x₂ 68 | -> (∀ y₂ -> y₁ # y₂) 69 | -> f x₁ y₁ # f x₂ y₂ 70 | dhcong₂ f inj (yes refl) q = dcong (f _) (homo ∘ inj) (q _) 71 | dhcong₂ f inj (no c) q = no (c ∘ inds ∘ inj) 72 | 73 | ,-inj : ∀ {α β} {A : Set α} {B : A -> Set β} {x₁ x₂} {y₁ : B x₁} {y₂ : B x₂} 74 | -> (x₁ , y₁) ≡ (x₂ , y₂) -> [ B ] y₁ ≅ y₂ 75 | ,-inj refl = irefl 76 | 77 | _<,>ᵈ_ : ∀ {α β} {A : Set α} {B : Set β} {x₁ x₂ : A} {y₁ y₂ : B} 78 | -> x₁ # x₂ -> y₁ # y₂ -> x₁ , y₁ # x₂ , y₂ 79 | _<,>ᵈ_ = dcong₂ _,_ (inds-homo ∘ ,-inj) 80 | 81 | _<,>ᵈⁱ_ : ∀ {α β} {A : Set α} {B : A -> Set β} {x₁ x₂} {y₁ : B x₁} {y₂ : B x₂} 82 | -> x₁ # x₂ -> (∀ y₂ -> y₁ # y₂) -> x₁ , y₁ # x₂ , y₂ 83 | _<,>ᵈⁱ_ = dhcong₂ _,_ ,-inj 84 | -------------------------------------------------------------------------------- /Lib/Heteroindexed.agda: -------------------------------------------------------------------------------- 1 | module OTT.Lib.Heteroindexed where 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | open import Data.Product 5 | 6 | data [_]_≅_ {ι α} {I : Set ι} {i} (A : I -> Set α) (x : A i) : ∀ {j} -> A j -> Set where 7 | irefl : [ A ] x ≅ x 8 | 9 | inds : ∀ {ι α} {I : Set ι} {A : I -> Set α} {i j} {x : A i} {y : A j} 10 | -> [ A ] x ≅ y -> i ≡ j 11 | inds irefl = refl 12 | 13 | homo : ∀ {ι α} {I : Set ι} {A : I -> Set α} {i} {x y : A i} 14 | -> [ A ] x ≅ y -> x ≡ y 15 | homo irefl = refl 16 | 17 | inds-homo : ∀ {ι α} {I : Set ι} {A : Set α} {i j : I} {x y : A} 18 | -> [_]_≅_ {i = i} (λ _ -> A) x {j} y -> i ≡ j × x ≡ y 19 | inds-homo irefl = refl , refl 20 | 21 | data [_][_]_≅_ {ι α β} {I : Set ι} {i} (A : I -> Set α) {x : A i} 22 | (B : ∀ {i} -> A i -> Set β) (y : B x) : ∀ {i} {x : A i} -> B x -> Set where 23 | iirefl : [ A ][ B ] y ≅ y 24 | -------------------------------------------------------------------------------- /Main.agda: -------------------------------------------------------------------------------- 1 | module OTT.Main where 2 | 3 | open import OTT.Core public 4 | open import OTT.Coerce public 5 | open import OTT.Data.List public 6 | open import OTT.Function.Pi public 7 | open import OTT.Function.Elim public 8 | -------------------------------------------------------------------------------- /Prelude.agda: -------------------------------------------------------------------------------- 1 | module OTT.Prelude where 2 | 3 | open import Level 4 | renaming (Level to MetaLevel; zero to lzeroₘ; suc to lsucₘ; _⊔_ to _⊔ₘ_) using () public 5 | open import Function public 6 | open import Relation.Binary.PropositionalEquality as P using (_≡_) 7 | renaming (refl to prefl; trans to ptrans; subst to psubst; cong to pcong; cong₂ to pcong₂) public 8 | open import Data.Empty public 9 | open import Data.Nat.Base hiding (_⊔_; _≟_; erase) public 10 | open import Data.Maybe.Base using (Maybe; nothing; just) public 11 | open import Data.Product hiding (,_) renaming (map to pmap) public 12 | 13 | open import OTT.Lib.Heteroindexed public 14 | open import OTT.Lib.Decidable public 15 | 16 | open import Relation.Nullary 17 | open import Relation.Binary 18 | 19 | infix 4 ,_ 20 | 21 | pattern ,_ y = _ , y 22 | 23 | record ⊤ {α} : Set α where 24 | constructor tt 25 | 26 | ⊤₀ : Set 27 | ⊤₀ = ⊤ 28 | 29 | instance 30 | iprefl : ∀ {α} {A : Set α} {x : A} -> x ≡ x 31 | iprefl = prefl 32 | 33 | ,-inst : ∀ {α β} {A : Set α} {B : A -> Set β} {{x : A}} {{y : B x}} -> Σ A B 34 | ,-inst {{x}} {{y}} = x , y 35 | 36 | pright : ∀ {α} {A : Set α} {x y z : A} -> x ≡ y -> x ≡ z -> y ≡ z 37 | pright prefl prefl = prefl 38 | 39 | hpcong₂ : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : Set γ} {x₁ x₂} {y₁ : B x₁} {y₂ : B x₂} 40 | -> (f : ∀ x -> B x -> C) -> (q : x₁ ≡ x₂) -> psubst B q y₁ ≡ y₂ -> f x₁ y₁ ≡ f x₂ y₂ 41 | hpcong₂ f prefl prefl = prefl 42 | 43 | record Apply {α β} {A : Set α} (B : A -> Set β) x : Set β where 44 | constructor tag 45 | field detag : B x 46 | open Apply public 47 | 48 | tag-inj : ∀ {α β} {A : Set α} {B : A -> Set β} {x} {y₁ y₂ : B x} 49 | -> tag {B = B} y₁ ≡ tag y₂ -> y₁ ≡ y₂ 50 | tag-inj prefl = prefl 51 | 52 | data Match {ι α} {I : Set ι} {i} (A : I -> Set α) (x : A i) : Set (ι ⊔ₘ α) where 53 | matched : ∀ {j} -> (x′ : A j) -> [ A ] x ≅ x′ -> Match A x 54 | 55 | match : ∀ {ι α} {I : Set ι} {i} -> (A : I -> Set α) -> (x : A i) -> Match A x 56 | match A x = matched x irefl 57 | 58 | data IMatch {ι α β} {I : Set ι} {i} (A : I -> Set α) {x : A i} 59 | (B : ∀ {i} -> A i -> Set β) (y : B x) : Set (ι ⊔ₘ α ⊔ₘ β) where 60 | imatched : ∀ {i} {x : A i} -> (y′ : B x) -> [ A ][ B ] y ≅ y′ -> IMatch A B y 61 | 62 | imatch : ∀ {ι α β} {I : Set ι} {i} 63 | -> (A : I -> Set α) {x : A i} -> (B : ∀ {i} -> A i -> Set β) -> (y : B x) -> IMatch A B y 64 | imatch A B y = imatched y iirefl 65 | 66 | module _ {α} {A : Set α} where 67 | open import Relation.Nullary.Decidable 68 | open import Data.Maybe 69 | 70 | toPropEq : {a b : Maybe A} -> a ≡ b -> Eq _≡_ a b 71 | toPropEq {a = nothing} prefl = nothing 72 | toPropEq {a = just _ } prefl = just prefl 73 | 74 | fromPropEq : {a b : Maybe A} -> Eq _≡_ a b -> a ≡ b 75 | fromPropEq (just q) = pcong just q 76 | fromPropEq nothing = prefl 77 | 78 | fromDecPropEq : {a b : Maybe A} -> Dec (Eq _≡_ a b) -> Dec (a ≡ b) 79 | fromDecPropEq = map′ fromPropEq toPropEq where 80 | 81 | decideMaybe : Decidable (_≡_ {A = A}) -> Decidable (_≡_ {A = Maybe A}) 82 | decideMaybe D a b = fromDecPropEq (a ≟ b) where 83 | open DecSetoid (decSetoid (P.decSetoid D)) 84 | 85 | Enum : ℕ -> Set 86 | Enum 0 = ⊥ 87 | Enum 1 = ⊤ 88 | Enum (suc n) = Maybe (Enum n) 89 | 90 | decEnum : ∀ n -> Decidable (_≡_ {A = Enum n}) 91 | decEnum 0 () () 92 | decEnum 1 tt tt = yes prefl 93 | decEnum (suc (suc n)) e₁ e₂ = decideMaybe (decEnum (suc n)) e₁ e₂ 94 | -------------------------------------------------------------------------------- /Property/Eq.agda: -------------------------------------------------------------------------------- 1 | module OTT.Property.Eq where 2 | 3 | import Data.Nat.Base as Nat 4 | 5 | open import OTT.Main 6 | open import OTT.Function.Pi 7 | 8 | infix 4 _≟_ 9 | 10 | module _ where 11 | open import Relation.Binary.PropositionalEquality.TrustMe 12 | 13 | contr : {A : Prop} {x y : ⟦ A ⟧} -> x ≡ y 14 | contr = trustMe 15 | 16 | SemEq : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} -> Desc I α -> Set 17 | SemEq (var i) = ⊤ 18 | SemEq (π A D) = ⊥ 19 | SemEq (D ⊛ E) = SemEq D × SemEq E 20 | 21 | mutual 22 | ExtendEq : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} -> Desc I α -> Set 23 | ExtendEq (var i) = ⊤ 24 | ExtendEq (π A D) = Eq A × Pi A λ x -> ExtendEq (D x) 25 | ExtendEq (D ⊛ E) = SemEq D × ExtendEq E 26 | 27 | Eq : ∀ {a} {α : Level a} -> Univ α -> Set 28 | Eq bot = ⊤ 29 | Eq top = ⊤ 30 | Eq nat = ⊤ 31 | Eq (enum n) = ⊤ 32 | Eq (σ A B) = Eq A × Pi A λ x -> Eq (B x) 33 | Eq (imu D j) = ExtendEq D 34 | Eq _ = ⊥ 35 | 36 | mutual 37 | decSem : ∀ {i a} {ι : Level i} {α : Level a} 38 | {I : Type ι} {D₀ : Desc I α} {{eqD₀ : ExtendEq D₀}} 39 | -> (D : Desc I α) {{eqD : SemEq D}} -> IsSet (⟦ D ⟧ᵈ (μ D₀)) 40 | decSem (var i) d₁ d₂ = decMu d₁ d₂ 41 | decSem (π A D) {{()}} 42 | decSem (D ⊛ E) {{eqD , eqE}} (s₁ , t₁) (s₂ , t₂) = 43 | decSem D {{eqD}} s₁ s₂ <,>ᵈ decSem E {{eqE}} t₁ t₂ 44 | 45 | decExtend : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} 46 | {D₀ : Desc I α} {j} {{eqD₀ : ExtendEq D₀}} 47 | -> (D : Desc I α) {{eqD : ExtendEq D}} -> IsSet (Extend D (μ D₀) j) 48 | decExtend (var i) q₁ q₂ = yes contr 49 | decExtend (π A D) {{eqA , eqD}} (x₁ , e₁) (x₂ , e₂) = 50 | _≟_ {{eqA}} x₁ x₂ <,>ᵈⁱ decExtend (D x₁) {{apply eqD x₁}} e₁ 51 | decExtend (D ⊛ E) {{eqD , eqE}} (s₁ , e₁) (s₂ , e₂) = 52 | decSem D {{eqD}} s₁ s₂ <,>ᵈ decExtend E {{eqE}} e₁ e₂ 53 | 54 | decMu : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} 55 | {D : Desc I α} {j} {{eqD : ExtendEq D}} 56 | -> IsSet (μ D j) 57 | decMu {D = D} (node e₁) (node e₂) = dcong node node-inj (decExtend D e₁ e₂) 58 | 59 | _≟_ : ∀ {a} {α : Level a} {A : Univ α} {{eqA : Eq A}} -> IsSet ⟦ A ⟧ 60 | _≟_ {A = bot } () () 61 | _≟_ {A = top } tt tt = yes prefl 62 | _≟_ {A = nat } n₁ n₂ = n₁ Nat.≟ n₂ 63 | _≟_ {A = enum n } (tag e₁) (tag e₂) = dcong tag tag-inj (decEnum n e₁ e₂) 64 | _≟_ {A = univ α } {{()}} 65 | _≟_ {A = σ A B } {{eqA , eqB}} (x₁ , y₁) (x₂ , y₂) = 66 | _≟_ {{eqA}} x₁ x₂ <,>ᵈⁱ _≟_ {{apply eqB x₁}} y₁ 67 | _≟_ {A = π A B } {{()}} 68 | _≟_ {A = desc I α} {{()}} 69 | _≟_ {A = imu D j } d₁ d₂ = decMu d₁ d₂ 70 | 71 | coerceFamEnum : ∀ {n} {e f : Apply Enum n} -> (A : Apply Enum n -> Set) -> ⟦ e ≅ f ⟧ -> A e -> A f 72 | coerceFamEnum {0} {tag () } {tag () } A q x 73 | coerceFamEnum {1} {tag tt } {tag tt } A q x = x 74 | coerceFamEnum {suc (suc n)} {tag nothing } {tag nothing } A q x = x 75 | coerceFamEnum {suc (suc n)} {tag (just _)} {tag (just _)} A q x = 76 | coerceFamEnum (A ∘ tag ∘ just ∘ detag) q x 77 | coerceFamEnum {suc (suc n)} {tag nothing } {tag (just _)} A () x 78 | coerceFamEnum {suc (suc n)} {tag (just _)} {tag nothing } A () x 79 | 80 | mutual 81 | observeSem : ∀ {i a} {ι : Level i} {α : Level a} 82 | {I : Type ι} {E : Desc I α} {{eqE : ExtendEq E}} 83 | -> (D : Desc I α) {{edD : SemEq D}} 84 | -> (d₁ d₂ : ⟦ D ⟧ᵈ (μ E)) -> ⟦ D , d₁ ≅s D , d₂ ⟧ -> d₁ ≡ d₂ 85 | observeSem (var i) d₁ d₂ q = observe q 86 | observeSem (π A D) {{()}} d₁ d₂ q 87 | observeSem (D ⊛ E) {{eqD , eqE}} (d₁ , e₁) (d₂ , e₂) (qd , qe) = 88 | pcong₂ _,_ (observeSem D {{eqD}} d₁ d₂ qd) (observeSem E {{eqE}} e₁ e₂ qe) 89 | 90 | observeExtend : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} 91 | {E : Desc I α} {j} {{edE : ExtendEq E}} 92 | -> (D : Desc I α) {{edD : ExtendEq D}} 93 | -> (e₁ e₂ : Extend D (μ E) j) -> ⟦ D , e₁ ≅e D , e₂ ⟧ -> e₁ ≡ e₂ 94 | observeExtend (var i) q₁ q₂ q = contr 95 | observeExtend (π A D) {{eqA , eqD}} (x₁ , e₁) (x₂ , e₂) (qx , qe) 96 | rewrite observe {x = x₁} {x₂} {{eqA}} qx = 97 | pcong (_,_ _) (observeExtend (D x₂) {{apply eqD x₂}} e₁ e₂ qe) 98 | observeExtend (D ⊛ E) {{eqD , eqE}} (d₁ , e₁) (d₂ , e₂) (qd , qe) = 99 | pcong₂ _,_ (observeSem D {{eqD}} d₁ d₂ qd) (observeExtend E {{eqE}} e₁ e₂ qe) 100 | 101 | observeMu : ∀ {i a} {ι : Level i} {α : Level a} {I : Type ι} {j} 102 | {D : Desc I α} {d e : μ D j} {{eqD : ExtendEq D}} 103 | -> ⟦ d ≅ e ⟧ -> d ≡ e 104 | observeMu {D = D} {node e₁} {node e₂} q = pcong node (observeExtend D e₁ e₂ q) 105 | 106 | observe : ∀ {a} {α : Level a} {A : Univ α} {x y : ⟦ A ⟧} {{eqA : Eq A}} -> ⟦ x ≅ y ⟧ -> x ≡ y 107 | observe {A = bot } {()} {()} 108 | observe {A = top } q = prefl 109 | observe {A = nat } q = coerceFamℕ (_ ≡_) q prefl 110 | observe {A = enum n } {e₁} {e₂} q = coerceFamEnum (_ ≡_) q prefl 111 | observe {A = univ α } {{()}} 112 | observe {A = σ A B } {x₁ , y₁} {x₂ , y₂} {{eqA , eqB}} (qx , qy) 113 | rewrite observe {x = x₁} {x₂} {{eqA}} qx = pcong (_,_ _) (observe {{apply eqB x₂}} qy) 114 | observe {A = π A B } {{()}} 115 | observe {A = desc I α} {{()}} 116 | observe {A = imu D j } {node e₁} {node e₂} q = observeMu q 117 | 118 | module _ where 119 | open import Relation.Binary.PropositionalEquality.TrustMe 120 | 121 | eobserve : ∀ {a} {α : Level a} {A : Univ α} {x y : ⟦ A ⟧} {{eqA : Eq A}} -> ⟦ x ≅ y ⟧ -> x ≡ y 122 | eobserve = erase ∘ observe 123 | 124 | private 125 | module Test where 126 | open import OTT.Data.Fin 127 | open import OTT.Data.Sum 128 | 129 | ns₁ : List (list nat ⊕ σ nat fin) 130 | ns₁ = inj₁ (1 ∷ 4 ∷ []) ∷ inj₂ (3 , fsuc fzero) ∷ inj₂ (2 , fzero) ∷ [] 131 | 132 | ns₂ : List (list nat ⊕ σ nat fin) 133 | ns₂ = inj₁ (1 ∷ 4 ∷ []) ∷ inj₂ (2 , fsuc fzero) ∷ inj₂ (2 , fzero) ∷ [] 134 | 135 | ns₃ : List (list nat ⊕ σ nat fin) 136 | ns₃ = inj₁ (1 ∷ 4 ∷ []) ∷ inj₂ (3 , fsuc fzero) ∷ [] 137 | 138 | test₁ : (ns₁ ≟ ns₁) ≡ yes prefl 139 | test₁ = prefl 140 | 141 | test₂ : (ns₁ ≟ ns₂) ≡ no _ 142 | test₂ = prefl 143 | 144 | test₃ : (ns₁ ≟ ns₃) ≡ no _ 145 | test₃ = prefl 146 | -------------------------------------------------------------------------------- /readme.md: -------------------------------------------------------------------------------- 1 | # OTT 2 | 3 | It's an implementation of Observational Type Theory as an Agda library. The universe contains `⊤`, `⊥`, natural numbers, Σ- and Π-types, impredicative `Prop` that lies in `Type₀`, infinite non-cumulative hierarchy `Type₀ : Type₁ : Type₂ : ...` and tools for generic programming: descriptions, their fixed points and finite enumerations. 4 | 5 | ## Details 6 | 7 | (The readme is a bit old: a `Level` now fully reifies the corresponding `MetaLevel` and there is also a [generic machinery](https://github.com/effectfully/OTT/blob/master/Function/Elim.agda) which allows to derive intensional-type-theory-like eliminators which is annoying to do by hand when a data type is indexed so you need to explicitly use `J`) 8 | 9 | There are two main kinds of universe levels (there is also yet another one, but it's not important): metatheory levels (the usual `Level` type renamed to `MetaLevel`) and target theory levels, indexed by metalevels: 10 | 11 | ``` 12 | data Level : MetaLevel -> Set where 13 | lzero : Level lzeroₘ 14 | lsuc : ∀ a -> Level (lsucₘ a) 15 | ``` 16 | 17 | This is needed, because `Prop` is represented as `Univ lzero` and `Type a` is represented as `Univ (lsuc a)`, so we need to be able to pattern match on `α` in `Univ α` to decide whether a type is a proposition or a set (because propositions and sets are often handled differently), but metalevels are parametric and hence can't be pattern matched on. 18 | 19 | The form of descriptions used here is described at the bottom (the `CompProp` module) of [9]. However in this development descriptions support full universe polymorphism (which makes them too down-to-earth to be able to perform levitation (not in a straightforward way at least), but that doesn't seem to be a problem). When defining a description we need to make sure that each `π` binds a variable which type lies in a smaller or the same universe than the type of the whole description. I.e. something like 20 | 21 | `π : ∀ {α} -> α ≤ β -> (A : Univ α) -> (⟦ A ⟧ -> Desc I β) -> Desc I β` 22 | 23 | If levels would be just natural numbers, that would work, however forcing a user to write down such boring and verbose proofs (though, we could probably use a solver with some reflection) would be too obtrusive. But `α ≤ β` can be equally written as `α ⊔ β ≡ β` and Agda has a built-in solver for universe levels, so here is the definition of `Desc`: 24 | 25 | ``` 26 | data Desc {i b} (I : Type i) (β : Level b) : Set where 27 | var : ⟦ I ⟧ -> Desc I β 28 | π : ∀ {a} {α : Level a} .{{_ : a ⊔ₘ b ≡ b}} 29 | -> (A : Univ α) -> (⟦ A ⟧ -> Desc I β) -> Desc I β 30 | _⊛_ : Desc I β -> Desc I β -> Desc I β 31 | ``` 32 | 33 | It's possible to avoid the use of instance arguments: 34 | 35 | ``` 36 | data UDesc {i o} (I : Type i) (ω : Level o) : (a : MetaLevel) -> Set where 37 | var[_] : ∀ {a} -> .(o ≡ a) -> ⟦ I ⟧ -> UDesc I ω a 38 | π[_] : ∀ {a b} {α : Level a} 39 | -> .(a ⊔ₘ o ≡ b) -> (A : Univ α) -> (⟦ A ⟧ -> UDesc I ω b) -> UDesc I ω b 40 | _⊛_ : ∀ {a} -> UDesc I ω a -> UDesc I ω a -> UDesc I ω a 41 | 42 | Desc : ∀ {a i} -> Type i -> Level a -> Set 43 | Desc {a} I α = UDesc I α a 44 | ``` 45 | 46 | Here `ω` is the final level and `a` is the inferred level. `Desc` requires these levels to be equal, thereby forcing unification. However this is more complicated and verbose, so I chose the first version. 47 | 48 | We have two notions of type equality: types can be equal either weakly 49 | 50 | ``` 51 | _≈_ : ∀ {a b} {α : Level a} {β : Level b} -> Univ α -> Univ β -> Prop 52 | _≈_ {α = lzero } {lzero } A₁ A₂ = A₁ ⇒ A₂ & A₂ ⇒ A₁ 53 | _≈_ {α = lsuc _} {lsuc _} A₁ A₂ = A₁ ≃ A₂ 54 | _≈_ _ _ = bot 55 | ``` 56 | 57 | or strictly 58 | 59 | ``` 60 | bot ≃ bot = top 61 | top ≃ top = top 62 | σ A₁ B₁ ≃ σ A₂ B₂ = A₁ ≈ A₂ & B₁ ≅ B₂ 63 | π A₁ B₁ ≃ π A₂ B₂ = A₂ ≈ A₁ & π A₁ λ x₁ -> π A₂ λ x₂ -> x₂ ≅ x₁ ⇒ B₁ x₁ ≈ B₂ x₂ 64 | ... 65 | ``` 66 | 67 | Weak equality for propositions is isomorphism, weak equality for sets is strict equality. 68 | 69 | In the same way there are two equalities for descriptions, since a description can lie in `Prop` — this happens when a description contains only propositions and inductive occurrences. 70 | 71 | The straightforward strict equality: 72 | 73 | ``` 74 | _≅ᵈ_ : ∀ {i₁ i₂ a₁ a₂} {α₁ : Level a₁} {α₂ : Level a₂} {I₁ : Type i₁} {I₂ : Type i₂} 75 | -> Desc I₁ α₁ -> Desc I₂ α₂ -> Prop 76 | var i₁ ≅ᵈ var i₂ = i₁ ≅ i₂ 77 | π A₁ D₁ ≅ᵈ π A₂ D₂ = A₁ ≈ A₂ & D₁ ≅ D₂ 78 | (D₁ ⊛ E₁) ≅ᵈ (D₂ ⊛ E₂) = D₁ ≅ᵈ D₂ & E₁ ≅ᵈ E₂ 79 | _ ≅ᵈ _ = bot 80 | ``` 81 | 82 | Described propositions are equal if their fixed points (propositions themselves) are isomorphic: 83 | 84 | ``` 85 | _≊ᵈ_ : ∀ {i₁ i₂ a₁ a₂} {α₁ : Level a₁} {α₂ : Level a₂} {I₁ : Type i₁} {I₂ : Type i₂} 86 | -> Desc I₁ α₁ -> Desc I₂ α₂ -> Prop 87 | _≊ᵈ_ {α₁ = lzero } {lzero } D₁ D₂ = imu D₁ ≅ imu D₂ 88 | _≊ᵈ_ {α₁ = lsuc _} {lsuc _} D₁ D₂ = D₁ ≅ᵈ D₂ 89 | _≊ᵈ_ _ _ = bot 90 | ``` 91 | 92 | Correspondingly, there are two `coerce`s: a one which requires weak equality 93 | 94 | ``` 95 | coerce : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} -> ⟦ A ≈ B ⇒ A ⇒ B ⟧ 96 | coerce {α = lzero } {lzero } (f , _) = f 97 | coerce {α = lsuc _} {lsuc _} q = coerce′ q 98 | coerce {α = lzero } {lsuc _} () 99 | coerce {α = lsuc _} {lzero } () 100 | ``` 101 | 102 | and a one which requires strict 103 | 104 | ``` 105 | coerce′ : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} -> ⟦ A ≃ B ⇒ A ⇒ B ⟧ 106 | coerce′ {A = bot } {bot } q () 107 | coerce′ {A = top } {top } q tt = tt 108 | coerce′ {A = σ A₁ B₁ } {σ A₂ B₂ } q p = let q₁ , q₂ = q ; x , y = p in 109 | coerce q₁ x , coerce (q₂ x (coerce q₁ x) (coherence q₁ x)) y 110 | coerce′ {A = π A₁ B₁ } {π A₂ B₂ } q f = let q₁ , q₂ = q in 111 | λ x -> coerce (q₂ (coerce q₁ x) x (coherence q₁ x)) (f (coerce q₁ x)) 112 | ... 113 | ``` 114 | 115 | `desc` allows to encode inductive data types (including inductive families) in the target theory. `coerce` computes under constructors of data types (see the "5. Full OTT" section of [1] for how this works). Each inductive family has at least two eliminators: one classical and one "up to propositional equality". An example from the `OTT.Data.Fin` module: 116 | 117 | ``` 118 | elimFinₑ : ∀ {n π} 119 | -> (P : ∀ {n} -> Fin n -> Set π) 120 | -> (∀ {n m} {i : Fin n} -> (q : ⟦ suc n ≅ m ⟧) -> P i -> P {m} (fsucₑ q i)) 121 | -> (∀ {n m} -> (q : ⟦ suc n ≅ m ⟧) -> P {m} (fzeroₑ q)) 122 | -> (i : Fin n) 123 | -> P i 124 | elimFinₑ P f x (fzeroₑ q) = x q 125 | elimFinₑ P f x (fsucₑ q i) = f q (elimFinₑ P f x i) 126 | ``` 127 | 128 | `elimFinₑ` is an "up to propositional equality" eliminator. The thing here is that `elimFinₑ` doesn't contain any coercions at all, so its "non-dependent" computational behaviour is the same as the corresponding behaviour of an eliminator in an intensional type theory. It even gives you slightly more: 129 | 130 | ``` 131 | elimFin′ : ∀ {n π} 132 | -> (P : ∀ n -> Set π) 133 | -> (∀ {n} {i : Fin n} -> P (fromFin i) -> P (suc (fromFin i))) 134 | -> P 0 135 | -> (i : Fin n) 136 | -> P (fromFin i) 137 | elimFin′ P f x = elimFinₑ (P ∘ fromFin) (λ {n m i} _ -> f {i = i}) (const x) 138 | ``` 139 | 140 | `elimFin′` doesn't mention `coerce` as well. 141 | 142 | We can recover the usual eliminator with the help from our old friend: 143 | 144 | ``` 145 | J : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {x y : ⟦ A ⟧} 146 | -> (B : (y : ⟦ A ⟧) -> ⟦ x ≅ y ⟧ -> Univ β) 147 | -> ⟦ B _ (refl x) ⟧ 148 | -> (q : ⟦ x ≅ y ⟧) 149 | -> ⟦ B _ q ⟧ 150 | J {x = x} B z q = subst₂ B q (huip x q) z 151 | 152 | elimFin : ∀ {n k} 153 | -> (P : ∀ {n} -> Fin n -> Univ k) 154 | -> (∀ {n} {i : Fin n} -> ⟦ P i ⇒ P (fsuc i) ⟧) 155 | -> (∀ {n} -> ⟦ P {suc n} fzero ⟧) 156 | -> (i : Fin n) 157 | -> ⟦ P i ⟧ 158 | elimFin P f x = elimFinₑ (⟦_⟧ ∘ P) 159 | (λ q r -> J (λ m q -> P {m} (fsucₑ q _)) (f r) q) 160 | (J (λ m q -> P {m} (fzeroₑ q)) x) 161 | ``` 162 | 163 | `subst₂` is defined in terms of `coerce`, so it computes under constructors of data types too, hence classical eliminators have pretty good computational behaviour too. 164 | 165 | A simple test: 166 | 167 | ``` 168 | postulate 169 | n m : ℕ 170 | 171 | test : ⟦ fromFin ((Fin (3 + n) ∋ fsuc (fsuc fzero)) +ᶠ (Fin (2 + m) ∋ fsuc fzero)) ≅ 3 ⟧ 172 | test = tt 173 | ``` 174 | 175 | `n` and `m` are stuck, but the expression reduces properly regardless of whether `_+ᶠ_` is defined in terms of `elimFin′` or `elimFin`. 176 | 177 | Eliminators for inductive data types (not families) are the usual intensional eliminators, e.g. 178 | 179 | ``` 180 | [_,_] : ∀ {a b π} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} {P : A ⊎ B -> Set π} 181 | -> (∀ x -> P (inj₁ x)) -> (∀ y -> P (inj₂ y)) -> ∀ s -> P s 182 | [ f , g ] (inj₁ x) = f x 183 | [ f , g ] (inj₂ y) = g y 184 | 185 | elimW : ∀ {a b π} {α : Level a} {β : Level b} {A : Univ α} {B : ⟦ A ⟧ -> Univ β} 186 | -> (P : W A B -> Set π) 187 | -> (∀ {x} {g : ⟦ B x ⟧ -> W A B} -> (∀ y -> P (g y)) -> P (sup x g)) 188 | -> ∀ w 189 | -> P w 190 | elimW P h (sup x g) = h (λ y -> elimW P h (g y)) 191 | ``` 192 | 193 | An example of generic programming can be found in the `Property/Eq.agda` module which defines this operator: 194 | 195 | `_≟_ : ∀ {a} {α : Level a} {A : Univ α} {{eqA : Eq A}} -> (x y : ⟦ A ⟧) -> Dec (x ≡ y)` 196 | 197 | It can handle numbers, finite sets, sums, products, lists and many other data types. An example: 198 | 199 | ``` 200 | ns : List (list nat ⊕ σ nat fin) 201 | ns = inj₁ (1 ∷ 4 ∷ []) ∷ inj₂ (3 , fsuc fzero) ∷ inj₂ (2 , fzero) ∷ [] 202 | 203 | test : (ns ≟ ns) ≡ yes prefl 204 | test = prefl 205 | ``` 206 | 207 | There is [an alternative encoding](https://github.com/effectfully/random-stuff/blob/master/Desc/IRDesc.agda) which is a modified version of [7]. It's more powerful (it's able to express induction-recursion), but also significantly more complicated: data types must be defined mutually with coercions, which results in a giant mutual block. I didn't try to define equality and coercions for descriptions, but I suspect it's much harder than how it's now. I'll go with the current simple approach. 208 | 209 | ## Not implemented 210 | 211 | - Definitional proof irrelevance. 212 | 213 | - Erasion of stuck coercions between definitionally equal types (that's not my fault, Agda just doesn't have an available definitional equality checker) (note that we have proper eliminators without this tool unlike in OTT with W-types (and they are still improper, see [4])). 214 | 215 | - Codata (is it simply the coinductive counterpart of `μ`?). 216 | 217 | - Quotients. Or maybe we can implement even quotient inductive types ([10])? 218 | 219 | - Pattern matching doesn't work properly, because Agda's unification is nailed to definitional equality and cannot handle fancy propositional equality. See [Pattern matching in Observational Type Theory](http://stackoverflow.com/questions/38957229/pattern-matching-in-observational-type-theory). 220 | 221 | ## A final remark 222 | 223 | Note that it's a library and not a formalization, since termination and positivity checkers are disabled. There are also several postulates (as well as in the original OTT), but nothing unexpected: 224 | 225 | ``` 226 | postulate 227 | refl : ∀ {a} {α : Level a} {A : Univ α} -> (x : ⟦ A ⟧) -> ⟦ x ≅ x ⟧ 228 | coherence : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} 229 | -> (q : ⟦ A ≈ B ⟧) -> (x : ⟦ A ⟧) -> ⟦ x ≅ coerce q x ⟧ 230 | cong-≅z : ∀ {a b c} {α : Level a} {β : Level b} {γ : Level c} 231 | {A : Univ α} {B : Univ β} {C : Univ γ} 232 | -> (x : ⟦ A ⟧) {y : ⟦ B ⟧} {z : ⟦ C ⟧} -> (q : ⟦ x ≅ y ⟧) -> ⟦ (x ≅ z) ≈ (y ≅ z)⟧ 233 | huip : ∀ {a b} {α : Level a} {β : Level b} {A : Univ α} {B : Univ β} 234 | -> (x : ⟦ A ⟧) {y : ⟦ B ⟧} -> (q : ⟦ x ≅ y ⟧) -> ⟦ refl x ≅ q ⟧ 235 | ``` 236 | 237 | ## References 238 | 239 | [1] ["Towards Observational Type Theory"](http://strictlypositive.org/ott.pdf), Thorsten Altenkirch and Conor McBride 240 | 241 | [2] ["Observational Equality, Now!"](http://www.cs.nott.ac.uk/~psztxa/publ/obseqnow.pdf), Thorsten Altenkirch, Conor McBride, Wouter Swierstra 242 | 243 | [3] ["Hier Soir, an OTT Hierarchy"](http://mazzo.li/epilogue/index.html%3Fp=1098.html), Conor McBride 244 | 245 | [4] ["W-types: good news and bad news"](http://mazzo.li/epilogue/index.html%3Fp=324.html), Conor McBride 246 | 247 | [5] ["On Universes in Type Theory"](http://citeseerx.ist.psu.edu/viewdoc/download?doi=10.1.1.221.1318&rep=rep1&type=pdf), Erik Palmgren 248 | 249 | [6] ["Modeling Elimination of Described Types"](http://spire-lang.org/blog/2014/01/15/modeling-elimination-of-described-types/), Larry Diehl 250 | 251 | [7] ["Inductive-Recursive Descriptions"](http://spire-lang.org/blog/2014/08/04/inductive-recursive-descriptions/), Larry Diehl 252 | 253 | [8] ["The Gengtle Art of Levitation"](https://jmchapman.github.io/papers/levitation.pdf), James Chapman, Pierre-Évariste Dagand, Conor McBride, Peter Morris 254 | 255 | [9] ["Descriptions"](http://effectfully.blogspot.ru/2016/04/descriptions.html) 256 | 257 | [10] ["Type Theory in Type Theory using Quotient Inductive Types"](http://www.cs.nott.ac.uk/~psztxa/publ/tt-in-tt.pdf), Thorsten Altenkirch, Ambrus Kaposi. 258 | --------------------------------------------------------------------------------