├── .gitignore ├── README.md ├── common.agda ├── contextualcat.agda ├── contextualcatmor.agda ├── initiality-existence.agda ├── initiality-uniqueness.agda ├── interpretationsubstitution.agda ├── interpretationweakening.agda ├── partialinterpretation.agda ├── quotients.agda ├── reflection.agda ├── rules.agda ├── syntx.agda ├── termmodel-common.agda ├── termmodel-empty.agda ├── termmodel-id.agda ├── termmodel-id2.agda ├── termmodel-nat.agda ├── termmodel-pi.agda ├── termmodel-sig.agda ├── termmodel-sum.agda ├── termmodel-synccat.agda ├── termmodel-unit.agda ├── termmodel-uuel.agda ├── termmodel.agda ├── totality.agda └── typetheory.agda /.gitignore: -------------------------------------------------------------------------------- 1 | bug.agda 2 | *.agdai 3 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | Introduction 2 | ------------ 3 | 4 | This project aims at giving a formalized proof in Agda of a version of the initiality conjecture. 5 | See tag v2.0 for a proof of initiality for a type theory with Pi-types, Sigma-types, natural 6 | numbers, identity types, binary sum types, the empty and unit types, and an infinite hierarchy of 7 | universes. 8 | 9 | This project has been tested to work on Agda 2.6.1 with the `+RTS -M25G -RTS` options (requires 25G 10 | of RAM), although you might need to restart it in the middle… It could work with less RAM, but 11 | taking more time. 12 | 13 | Some design decisions 14 | --------------------- 15 | 16 | The raw syntax of the type theory is defined using de Bruijn indices for the variables, and terms 17 | are fully annotated. Typing rules are mostly standard, except that we never check that a context is 18 | well-typed, instead we only check that the *type* is well-typed when using the variable rule. There 19 | is moreover a little bit of paranoia in some of the rules. 20 | 21 | For the semantics, we state initiality with contextual categories whose additional structure is as 22 | close to the syntax as possible. 23 | 24 | Metatheory 25 | ---------- 26 | 27 | We are using standard Agda with the sort `Prop` of strict propositions (needs the `--prop` flag) 28 | (see https://hal.inria.fr/hal-01859964v2/document). 29 | We are also using a few axioms: 30 | - dependent function extensionality (three axioms, for the cases when the domain is a type, a 31 | proposition, or implicit) 32 | - propositional extensionality, phrased with `Prop` 33 | - existence of quotients for `Prop`-valued equivalence relations, with a definitional computation 34 | rule (defined using rewrite rules, needs the `--rewriting` flag). 35 | 36 | Files 37 | ----- 38 | 39 | - `common.agda`: basic definitions used throughout the development 40 | - `typetheory.agda`: raw syntax of the type theory 41 | - `reflection.agda`: a basic reflection library (used in `syntx.agda`) 42 | - `syntx.agda`: properties of the raw syntax 43 | - `rules.agda`: typing rules and admissible rules 44 | - `contextualcat.agda`: models of the type theory 45 | - `contextualcatmor.agda`: morphisms between models of the type theory 46 | - `quotients.agda`: quotients (postulated) 47 | - `termmodel-common.agda`: preliminary definitions and helper functions for the term model 48 | - `termmodel-synccat.agda`: the contextual category part of the term model 49 | - `termmodel-X.agda`: the part of the term model related to type former X 50 | - `termmodel.agda`: the full term model 51 | - `partialinterpretation.agda`: the partial interpretation function 52 | - `interpretationsubstitution.agda`: commutation of substitution and the interpretation function 53 | - `interpretationweakening.agda`: commutation of weakening and the interpretation function 54 | - `totality.agda`: totality of the partial interpretation function 55 | - `initiality-existence.agda`: existence part of the proof of initiality of the term model 56 | - `initiality-uniqueness.agda`: uniqueness part of the proof of initiality of the term model 57 | -------------------------------------------------------------------------------- /common.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import Agda.Primitive public 4 | open import Agda.Builtin.Nat public renaming (Nat to ℕ) hiding (_==_; _<_) 5 | open import Agda.Builtin.List public 6 | open import Agda.Builtin.Bool public 7 | open import Agda.Builtin.Unit public 8 | 9 | {- Rewriting -} 10 | 11 | abstract 12 | _↦_ : ∀ {l} {A : Set l} → A → A → Set l 13 | a ↦ b = Id a b where 14 | data Id (a : _) : _ → Set _ where 15 | refl : Id a a 16 | {-# BUILTIN REWRITE _↦_ #-} 17 | 18 | 19 | {- Basic records -} 20 | 21 | record Σ (A : Prop) (B : A → Prop) : Prop where 22 | no-eta-equality 23 | constructor _,_ 24 | field 25 | fst : A 26 | snd : B fst 27 | open Σ public 28 | 29 | record ΣS {l} {l'} (A : Set l) (B : A → Prop l') : Set (l ⊔ l') where 30 | constructor _,_ 31 | field 32 | fst : A 33 | snd : B fst 34 | open ΣS public 35 | 36 | infixr 4 _,_ 37 | 38 | record _∧_ (A B : Prop) : Prop where 39 | constructor _,_ 40 | field 41 | fst : A 42 | snd : B 43 | open _∧_ public 44 | 45 | infixr 42 _∧_ 46 | 47 | record _×_ (A B : Set) : Set where 48 | constructor _,_ 49 | field 50 | fst : A 51 | snd : B 52 | open _∧_ public 53 | 54 | infixr 42 _×_ 55 | 56 | record True : Prop where 57 | constructor tt 58 | 59 | 60 | {- Prop-valued equality -} 61 | 62 | data _≡_ {l} {A : Set l} (x : A) : A → Prop l where 63 | instance refl : x ≡ x 64 | {-# BUILTIN EQUALITY _≡_ #-} 65 | 66 | infix 4 _≡_ 67 | 68 | 69 | Σ= : ∀ {l} {l'} {A : Set l} {B : A → Prop l'} {a a' : A} {b : B a} {b' : B a'} → a ≡ a' → _≡_ {A = ΣS _ _} (a , b) (a' , b') 70 | Σ= refl = refl 71 | 72 | ap : ∀ {l l'} {A : Set l} {B : Set l'} (f : A → B) {a b : A} → a ≡ b → f a ≡ f b 73 | ap f refl = refl 74 | 75 | _∙_ : {A : Set} {a b c : A} → a ≡ b → b ≡ c → a ≡ c 76 | refl ∙ refl = refl 77 | 78 | infixr 4 _∙_ 79 | 80 | ! : {A : Set} {a b : A} → a ≡ b → b ≡ a 81 | ! refl = refl 82 | 83 | {- Lifting from Prop to Set -} 84 | 85 | record Box {l} (P : Prop l) : Set l where 86 | constructor box 87 | field 88 | unbox : P 89 | open Box public 90 | 91 | {- Generalized variables -} 92 | 93 | variable 94 | {n n' m k l} : ℕ 95 | 96 | {- Finite sets 97 | 98 | There are two different use cases for finite sets: 99 | - to specify a certain variable in a context of length [n] ([n] many choices) 100 | - to specify where we want to weaken a context of length [n] ([suc n] many choices) 101 | Instead of using [Fin n] and [Fin (suc n)], we define two different datatypes so that we 102 | don’t mix them up. 103 | -} 104 | 105 | data VarPos : ℕ → Set where 106 | last : VarPos (suc n) 107 | prev : VarPos n → VarPos (suc n) 108 | 109 | -- Size of the context before (and excluding) that variable 110 | _-VarPos'_ : (n : ℕ) → VarPos n → ℕ 111 | (suc m) -VarPos' last = m 112 | (suc m) -VarPos' prev k = m -VarPos' k 113 | 114 | -- Size of the context before (and including) that variable 115 | _-VarPos_ : (n : ℕ) → VarPos n → ℕ 116 | n -VarPos k = suc (n -VarPos' k) 117 | 118 | 119 | data WeakPos : ℕ → Set where 120 | last : WeakPos n 121 | prev : WeakPos n → WeakPos (suc n) 122 | 123 | -- Size of the context before that position 124 | _-WeakPos_ : (n : ℕ) → WeakPos n → ℕ 125 | n -WeakPos last = n 126 | suc n -WeakPos prev k = n -WeakPos k 127 | 128 | data _≤WP_ : WeakPos n → WeakPos n → Prop where 129 | last≤ : {k : WeakPos n} → k ≤WP last 130 | prev≤ : {k k' : WeakPos n} → k ≤WP k' → prev k ≤WP prev k' 131 | 132 | -- Every position of length [n] is a position of length [suc n] 133 | injWeakPos : {n : ℕ} → WeakPos n → WeakPos (suc n) 134 | injWeakPos last = last 135 | injWeakPos (prev k) = prev (injWeakPos k) 136 | 137 | -- Position of the new variable in the weakened context 138 | WeakPosToVarPos : {n : ℕ} → WeakPos n → VarPos (suc n) 139 | WeakPosToVarPos last = last 140 | WeakPosToVarPos (prev k) = prev (WeakPosToVarPos k) 141 | 142 | {- Monads -} 143 | 144 | record Monad {ℓ ℓ'} (M : Set ℓ → Set ℓ') : Set (lsuc ℓ ⊔ ℓ') where 145 | field 146 | return : {A : Set ℓ} → A → M A 147 | _>>=_ : {A B : Set ℓ} → M A → (A → M B) → M B 148 | 149 | _>>_ : {A B : Set ℓ} → M A → M B → M B 150 | a >> b = a >>= (λ _ → b) 151 | 152 | open Monad {{…}} public 153 | 154 | {- The partiality monad -} 155 | 156 | record Partial (A : Set) : Set₁ where 157 | field 158 | isDefined : Prop 159 | _$_ : isDefined → A 160 | infix 5 _$_ 161 | open Partial public 162 | 163 | instance 164 | PartialityMonad : Monad Partial 165 | isDefined (return {{ PartialityMonad }} x) = True 166 | return {{ PartialityMonad }} x Partial.$ tt = x 167 | isDefined (_>>=_ {{ PartialityMonad }} a f) = Σ (isDefined a) (λ x → isDefined (f (a $ x))) 168 | _>>=_ {{ PartialityMonad }} a f Partial.$ x = f (a $ fst x) $ snd x 169 | 170 | assume : (P : Prop) → Partial (Box P) 171 | isDefined (assume P) = P 172 | unbox (assume P Partial.$ x) = x 173 | 174 | 175 | {- Helper functions for proof irrelevance -} 176 | 177 | ap-irr : {A C : Set} {B : A → Prop} (f : (a : A) (b : B a) → C) {a a' : A} (p : a ≡ a') {b : B a} {b' : B a'} → f a b ≡ f a' b' 178 | ap-irr f refl = refl 179 | 180 | ap-irr2 : {A D : Set} {B : A → Prop} {C : (a : A) (_ : B a) → Prop} (f : (a : A) (b : B a) (c : C a b) → D) {a a' : A} (p : a ≡ a') {b : B a} {b' : B a'} {c : C a b} {c' : C a' b'} → f a b c ≡ f a' b' c' 181 | ap-irr2 f refl = refl 182 | 183 | ap2-irr : {A C D : Set} {B : A → C → Prop} (f : (a : A) (c : C) (b : B a c) → D) {a a' : A} (p : a ≡ a') {c c' : C} (q : c ≡ c') {b : B a c} {b' : B a' c'} → f a c b ≡ f a' c' b' 184 | ap2-irr f refl refl = refl 185 | 186 | ap3-irr2 : {A B C D : Set} {E : A → B → C → Prop} {F : A → B → C → Prop} (h : (a : A) (b : B) {c : C} (p : E a b c) (q : F a b c) → D) {a a' : A} (p : a ≡ a') {b b' : B} (q : b ≡ b') {c c' : C} (r : c ≡ c') {e : E a b c} {e' : E a' b' c'} {f : F a b c} {f' : F a' b' c'} → h a b {c = c} e f ≡ h a' b' {c = c'} e' f' 187 | ap3-irr2 h refl refl refl = refl 188 | 189 | 190 | {- Axioms -} 191 | 192 | postulate 193 | -- Dependent function extensionality 194 | funext : ∀ {l l'} {A : Set l} {B : A → Set l'} {f g : (a : A) → B a} (h : (x : A) → f x ≡ g x) → f ≡ g 195 | 196 | -- Dependent function extensionality for function with domain Prop, does not seem to follow from [funext] 197 | funextP : ∀ {l l'} {A : Prop l} {B : A → Set l'} {f g : (a : A) → B a} (h : (x : A) → f x ≡ g x) → f ≡ g 198 | 199 | -- Dependent function extensionality for implicit function spaces 200 | funextI : ∀ {l l'} {A : Set l} {B : A → Set l'} {f g : {a : A} → B a} (h : (x : A) → f {x} ≡ g {x}) → (λ {x} → f {x}) ≡ (λ {x} → g {x}) 201 | 202 | -- Propositional extensionality 203 | prop-ext : {A B : Prop} (f : A → B) (g : B → A) → A ≡ B 204 | -------------------------------------------------------------------------------- /contextualcatmor.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import contextualcat 5 | 6 | {- Morphisms of contextual categories -} 7 | 8 | record CCatMor (C D : CCat) : Set where 9 | open CCat+ 10 | field 11 | Ob→ : Ob C n → Ob D n 12 | Mor→ : Mor C n m → Mor D n m 13 | ∂₀→ : {X : Mor C n m} → Ob→ (∂₀ C X) ≡ ∂₀ D (Mor→ X) 14 | ∂₁→ : {X : Mor C n m} → Ob→ (∂₁ C X) ≡ ∂₁ D (Mor→ X) 15 | id→ : {X : Ob C n} → Mor→ (id C X) ≡ id D (Ob→ X) 16 | comp→ : {n m k : ℕ} {g : Mor C m k} {f : Mor C n m} {X : Ob C m} {g₀ : ∂₀ C g ≡ X} {f₁ : ∂₁ C f ≡ X} → Mor→ (comp C g f g₀ f₁) ≡ comp D (Mor→ g) (Mor→ f) (! ∂₀→ ∙ ap Ob→ g₀) (! ∂₁→ ∙ ap Ob→ f₁) 17 | ft→ : {X : Ob C (suc n)} → Ob→ (ft C X) ≡ ft D (Ob→ X) 18 | pp→ : {X : Ob C (suc n)} → Mor→ (pp C X) ≡ pp D (Ob→ X) 19 | star→ : {n m : ℕ} {f : Mor C m n} {X : Ob C (suc n)} {Y : Ob C n} {q : ft C X ≡ Y} {f₁ : ∂₁ C f ≡ Y} → Ob→ (star C f X q f₁) ≡ star D (Mor→ f) (Ob→ X) (! ft→ ∙ ap Ob→ q) (! ∂₁→ ∙ ap Ob→ f₁) 20 | qq→ : {n m : ℕ} {f : Mor C m n} {X : Ob C (suc n)} {Y : Ob C n} {q : ft C X ≡ Y} {f₁ : ∂₁ C f ≡ Y} → Mor→ (qq C f X q f₁) ≡ qq D (Mor→ f) (Ob→ X) (! ft→ ∙ ap Ob→ q) (! ∂₁→ ∙ ap Ob→ f₁) 21 | ss→ : {f : Mor C m (suc n)} → Mor→ (ss C f) ≡ ss D (Mor→ f) 22 | pt→ : Ob→ (pt C) ≡ pt D 23 | ptmor→ : {X : Ob C n} → Mor→ (ptmor C X) ≡ ptmor D (Ob→ X) 24 | 25 | Mor→ₛ : {n : ℕ} {u : Mor C n (suc n)} (uₛ : is-section C u) → is-section D (Mor→ u) 26 | Mor→ₛ uₛ = is-section→ D (! (comp→ ∙ ap-irr-comp D (pp→ ∙ ap (pp D) ∂₁→) refl) ∙ ap Mor→ (is-section= C (! (is-section₀ C uₛ refl)) uₛ refl) ∙ id→ ∙ ap (id D) ∂₀→) 27 | 28 | Mor→₁ : {n : ℕ} {u : Mor C n (suc n)} {X : Ob C (suc n)} (u₁ : ∂₁ C u ≡ X) → ∂₁ D (Mor→ u) ≡ Ob→ X 29 | Mor→₁ u₁ = ! ∂₁→ ∙ ap Ob→ u₁ 30 | 31 | 32 | {- Morphisms of structured contextual categories -} 33 | 34 | record StructuredCCatMor (sC sD : StructuredCCat) : Set where 35 | private 36 | open StructuredCCat 37 | C = ccat sC 38 | D = ccat sD 39 | 40 | field 41 | ccat→ : CCatMor C D 42 | 43 | open CCatMor ccat→ 44 | open CCat+ 45 | 46 | field 47 | UUStr→ : (i : ℕ) (Γ : Ob C n) → Ob→ (UUStr sC i Γ) ≡ UUStr sD i (Ob→ Γ) 48 | ElStr→ : (i : ℕ) (Γ : Ob C n) (v : Mor C n (suc n)) (vₛ : is-section C v) (v₁ : ∂₁ C v ≡ UUStr sC i Γ) 49 | → Ob→ (ElStr sC i Γ v vₛ v₁) ≡ ElStr sD i (Ob→ Γ) (Mor→ v) (Mor→ₛ vₛ) (Mor→₁ v₁ ∙ UUStr→ i Γ) 50 | SumStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc n)) (B= : ft C B ≡ Γ) 51 | → Ob→ (SumStr sC Γ A A= B B=) ≡ SumStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) 52 | PiStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc (suc n))) (B= : ft C B ≡ A) 53 | → Ob→ (PiStr sC Γ A A= B B=) ≡ PiStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) 54 | SigStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc (suc n))) (B= : ft C B ≡ A) 55 | → Ob→ (SigStr sC Γ A A= B B=) ≡ SigStr sD (Ob→ Γ) (Ob→ A)(! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) 56 | EmptyStr→ : (Γ : Ob C n) → Ob→ (EmptyStr sC Γ) ≡ EmptyStr sD (Ob→ Γ) 57 | UnitStr→ : (Γ : Ob C n) → Ob→ (UnitStr sC Γ) ≡ UnitStr sD (Ob→ Γ) 58 | NatStr→ : (Γ : Ob C n) → Ob→ (NatStr sC Γ) ≡ NatStr sD (Ob→ Γ) 59 | IdStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ A) (b : Mor C n (suc n)) (bₛ : is-section C b) (b₁ : ∂₁ C b ≡ A) 60 | → Ob→ (IdStr sC Γ A A= a aₛ a₁ b bₛ b₁) ≡ IdStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁) (Mor→ b) (Mor→ₛ bₛ) (Mor→₁ b₁) 61 | 62 | 63 | uuStr→ : (i : ℕ) (Γ : Ob C n) 64 | → Mor→ (uuStr sC i Γ) ≡ uuStr sD i (Ob→ Γ) 65 | sumStr→ : (i : ℕ) (Γ : Ob C n) (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ UUStr sC i Γ) (b : Mor C n (suc n)) (bₛ : is-section C b) (b₁ : ∂₁ C b ≡ UUStr sC i Γ) 66 | → Mor→ (sumStr sC i Γ a aₛ a₁ b bₛ b₁) ≡ sumStr sD i (Ob→ Γ) (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁ ∙ UUStr→ i Γ) (Mor→ b) (Mor→ₛ bₛ) (Mor→₁ b₁ ∙ UUStr→ i Γ) 67 | inlStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc n)) (B= : ft C B ≡ Γ) (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ A) 68 | → Mor→ (inlStr sC Γ A A= B B= a aₛ a₁) ≡ inlStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁) 69 | inrStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc n)) (B= : ft C B ≡ Γ) (b : Mor C n (suc n)) (bₛ : is-section C b) (b₁ : ∂₁ C b ≡ B) 70 | → Mor→ (inrStr sC Γ A A= B B= b bₛ b₁) ≡ inrStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) (Mor→ b) (Mor→ₛ bₛ) (Mor→₁ b₁) 71 | matchStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc n)) (B= : ft C B ≡ Γ) (P : Ob C (suc (suc n))) (P= : ft C P ≡ SumStr sC Γ A A= B B=) 72 | (da : Mor C (suc n) (suc (suc n))) (daₛ : is-section C da) (da₁ : ∂₁ C da ≡ T-da₁ sC Γ A A= B B= P P=) 73 | (db : Mor C (suc n) (suc (suc n))) (dbₛ : is-section C db) (db₁ : ∂₁ C db ≡ T-db₁ sC Γ A A= B B= P P=) 74 | (u : Mor C n (suc n)) (uₛ : is-section C u) (u₁ : ∂₁ C u ≡ SumStr sC Γ A A= B B=) 75 | → Mor→ (matchStr sC Γ A A= B B= P P= da daₛ da₁ db dbₛ db₁ u uₛ u₁) ≡ matchStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) 76 | (Ob→ P) (! ft→ ∙ ap Ob→ P= ∙ SumStr→ Γ A A= B B=) 77 | (Mor→ da) (Mor→ₛ daₛ) (Mor→₁ da₁ ∙ star→ 78 | ∙ ap-irr-star D (inlStr→ _ _ _ _ _ _ _ _ 79 | ∙ ap-irr-inlStr sD refl (star→ ∙ ap-irr-star D pp→ refl) 80 | (star→ ∙ ap-irr-star D pp→ refl) 81 | (ss→ ∙ ap (ss D) id→)) 82 | (star→ ∙ ap-irr-star D (qq→ ∙ ap-irr-qq D pp→ (SumStr→ Γ A A= B B=)) refl)) 83 | (Mor→ db) (Mor→ₛ dbₛ) (Mor→₁ db₁ ∙ star→ 84 | ∙ ap-irr-star D (inrStr→ _ _ _ _ _ _ _ _ 85 | ∙ ap-irr-inrStr sD refl (star→ ∙ ap-irr-star D pp→ refl) 86 | (star→ ∙ ap-irr-star D pp→ refl) 87 | (ss→ ∙ ap (ss D) id→)) 88 | (star→ ∙ ap-irr-star D (qq→ ∙ ap-irr-qq D pp→ (SumStr→ Γ A A= B B=)) refl)) 89 | (Mor→ u) (Mor→ₛ uₛ) (Mor→₁ u₁ ∙ SumStr→ _ _ _ _ _) 90 | piStr→ : (i : ℕ) (Γ : Ob C n) (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ UUStr sC i Γ) (b : Mor C (suc n) (suc (suc n))) (bₛ : is-section C b) (b₁ : ∂₁ C b ≡ UUStr sC i (ElStr sC i Γ a aₛ a₁)) 91 | → Mor→ (piStr sC i Γ a aₛ a₁ b bₛ b₁) ≡ piStr sD i (Ob→ Γ) (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁ ∙ UUStr→ i Γ) (Mor→ b) (Mor→ₛ bₛ) (Mor→₁ b₁ ∙ UUStr→ i (ElStr sC i Γ a aₛ a₁) ∙ ap (UUStr sD i) (ElStr→ i Γ a aₛ a₁)) 92 | lamStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc (suc n))) (B= : ft C B ≡ A) (u : Mor C (suc n) (suc (suc n))) (uₛ : is-section C u) (u₁ : ∂₁ C u ≡ B) 93 | → Mor→ (lamStr sC Γ A A= B B= u uₛ u₁) ≡ lamStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) (Mor→ u) (Mor→ₛ uₛ) (Mor→₁ u₁) 94 | appStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc (suc n))) (B= : ft C B ≡ A) (f : Mor C n (suc n)) (fₛ : is-section C f) (f₁ : ∂₁ C f ≡ PiStr sC Γ A A= B B=) (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ A) 95 | → Mor→ (appStr sC Γ A A= B B= f fₛ f₁ a aₛ a₁) ≡ appStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) (Mor→ f) (Mor→ₛ fₛ) (Mor→₁ f₁ ∙ PiStr→ Γ A A= B B=) (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁) 96 | sigStr→ : (i : ℕ) (Γ : Ob C n) (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ UUStr sC i Γ) (b : Mor C (suc n) (suc (suc n))) (bₛ : is-section C b) (b₁ : ∂₁ C b ≡ UUStr sC i (ElStr sC i Γ a aₛ a₁)) 97 | → Mor→ (sigStr sC i Γ a aₛ a₁ b bₛ b₁) ≡ sigStr sD i (Ob→ Γ) (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁ ∙ UUStr→ i Γ) (Mor→ b) (Mor→ₛ bₛ) (Mor→₁ b₁ ∙ UUStr→ i (ElStr sC i Γ a aₛ a₁) ∙ ap (UUStr sD i) (ElStr→ i Γ a aₛ a₁)) 98 | pairStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc (suc n))) (B= : ft C B ≡ A) (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ A) (b : Mor C n (suc n)) (bₛ : is-section C b) (b₁ : ∂₁ C b ≡ star C a B B= a₁) 99 | → Mor→ (pairStr sC Γ A A= B B= a aₛ a₁ b bₛ b₁) ≡ pairStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁) (Mor→ b) (Mor→ₛ bₛ) (Mor→₁ b₁ ∙ star→) 100 | pr1Str→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc (suc n))) (B= : ft C B ≡ A) (u : Mor C n (suc n)) (uₛ : is-section C u) (u₁ : ∂₁ C u ≡ SigStr sC Γ A A= B B=) 101 | → Mor→ (pr1Str sC Γ A A= B B= u uₛ u₁) ≡ pr1Str sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) (Mor→ u) (Mor→ₛ uₛ) (Mor→₁ u₁ ∙ SigStr→ Γ A A= B B=) 102 | pr2Str→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (B : Ob C (suc (suc n))) (B= : ft C B ≡ A) (u : Mor C n (suc n)) (uₛ : is-section C u) (u₁ : ∂₁ C u ≡ SigStr sC Γ A A= B B=) 103 | → Mor→ (pr2Str sC Γ A A= B B= u uₛ u₁) ≡ pr2Str sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ B) (! ft→ ∙ ap Ob→ B=) (Mor→ u) (Mor→ₛ uₛ) (Mor→₁ u₁ ∙ SigStr→ Γ A A= B B=) 104 | emptyStr→ : (i : ℕ) (Γ : Ob C n) → Mor→ (emptyStr sC i Γ) ≡ emptyStr sD i (Ob→ Γ) 105 | emptyelimStr→ : (Γ : Ob C n) (A : Ob C (suc (suc n))) (A= : ft C A ≡ EmptyStr sC Γ) (u : Mor C n (suc n)) (uₛ : is-section C u) (u₁ : ∂₁ C u ≡ EmptyStr sC Γ) 106 | → Mor→ (emptyelimStr sC Γ A A= u uₛ u₁) ≡ emptyelimStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A= ∙ EmptyStr→ Γ) (Mor→ u) (Mor→ₛ uₛ) (Mor→₁ u₁ ∙ EmptyStr→ Γ ) 107 | unitStr→ : (i : ℕ) (Γ : Ob C n) → Mor→ (unitStr sC i Γ) ≡ unitStr sD i (Ob→ Γ) 108 | ttStr→ : (Γ : Ob C n) → Mor→ (ttStr sC Γ) ≡ ttStr sD (Ob→ Γ) 109 | unitelimStr→ : (Γ : Ob C n) (A : Ob C (suc (suc n))) (A= : ft C A ≡ UnitStr sC Γ) (dtt : Mor C n (suc n)) (dttₛ : is-section C dtt) (dtt₁ : ∂₁ C dtt ≡ star C (ttStr sC Γ) A A= (ttStr₁ sC)) (u : Mor C n (suc n)) (uₛ : is-section C u) (u₁ : ∂₁ C u ≡ UnitStr sC Γ) → Mor→ (unitelimStr sC Γ A A= dtt dttₛ dtt₁ u uₛ u₁) ≡ unitelimStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A= ∙ UnitStr→ Γ) (Mor→ dtt) (Mor→ₛ dttₛ) (Mor→₁ dtt₁ ∙ star→ ∙ ap-irr-star D (ttStr→ Γ) refl) (Mor→ u) (Mor→ₛ uₛ) (Mor→₁ u₁ ∙ UnitStr→ Γ) 110 | natStr→ : (i : ℕ) (Γ : Ob C n) 111 | → Mor→ (natStr sC i Γ) ≡ natStr sD i (Ob→ Γ) 112 | zeroStr→ : (Γ : Ob C n) 113 | → Mor→ (zeroStr sC Γ) ≡ zeroStr sD (Ob→ Γ) 114 | sucStr→ : (Γ : Ob C n) (u : Mor C n (suc n)) (uₛ : is-section C u) (u₁ : ∂₁ C u ≡ NatStr sC Γ) 115 | → Mor→ (sucStr sC Γ u uₛ u₁) ≡ sucStr sD (Ob→ Γ) (Mor→ u) (Mor→ₛ uₛ) (Mor→₁ u₁ ∙ NatStr→ Γ) 116 | idStr→ : (i : ℕ) (Γ : Ob C n) (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ UUStr sC i Γ) (u : Mor C n (suc n)) (uₛ : is-section C u) (u₁ : ∂₁ C u ≡ ElStr sC i Γ a aₛ a₁) 117 | (v : Mor C n (suc n)) (vₛ : is-section C v) (v₁ : ∂₁ C v ≡ ElStr sC i Γ a aₛ a₁) 118 | → Mor→ (idStr sC i Γ a aₛ a₁ u uₛ u₁ v vₛ v₁) ≡ idStr sD i (Ob→ Γ) (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁ ∙ UUStr→ i Γ) (Mor→ u) (Mor→ₛ uₛ) (Mor→₁ u₁ ∙ ElStr→ i Γ a aₛ a₁) (Mor→ v) (Mor→ₛ vₛ) (Mor→₁ v₁ ∙ ElStr→ i Γ a aₛ a₁) 119 | reflStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ A) 120 | → Mor→ (reflStr sC Γ A A= a aₛ a₁) ≡ reflStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁) 121 | 122 | abstract 123 | T-dS→ : {Γ : Ob C n} {P : Ob C (suc (suc n))} {P= : ft C P ≡ NatStr sC Γ} → Ob→ (T-dS₁ sC Γ P P=) ≡ T-dS₁ sD (Ob→ Γ) (Ob→ P) (! ft→ ∙ ap Ob→ P= ∙ NatStr→ Γ) 124 | T-dS→ {_} {Γ} {P} {P=} = 125 | star→ ∙ ap-irr-star D (sucStr→ _ _ _ _ ∙ ap-irr-sucStr sD refl (ss→ ∙ ap (ss D) pp→)) 126 | (star→ ∙ ap-irr-star D (qq→ ∙ ap-irr-qq D pp→ (star→ ∙ ap-irr-star D (pp→ ∙ ap (pp D) (NatStr→ Γ)) (NatStr→ Γ))) 127 | (star→ ∙ ap-irr-star D (qq→ ∙ ap-irr-qq D (pp→ ∙ ap (pp D) (NatStr→ Γ)) (NatStr→ Γ)) refl)) 128 | 129 | T-ftP→ : {Γ : Ob C n} {A : Ob C (suc n)} {A= : ft C A ≡ Γ} → Ob→ (T-ftP sC Γ A A=) ≡ T-ftP sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) 130 | T-ftP→ {_} {Γ} {A} {A=} = IdStr→ _ _ _ _ _ _ _ _ _ 131 | ∙ ap-irr-IdStr sD (star→ ∙ ap-irr-star D pp→ refl) 132 | (star→ ∙ ap-irr-star D (pp→ ∙ ap (pp D) (star→ ∙ ap-irr-star D pp→ refl)) (star→ ∙ ap-irr-star D pp→ refl)) 133 | (ss→ ∙ ap (ss D) (pp→ ∙ ap (pp D) (star→ ∙ ap-irr-star D pp→ refl))) 134 | (ss→ ∙ ap (ss D) (id→ ∙ ap (id D) (star→ ∙ ap-irr-star D pp→ refl))) 135 | 136 | 137 | T-d₁→ : {Γ : Ob C n} {A : Ob C (suc n)} {A= : ft C A ≡ Γ} {P : Ob C (suc (suc (suc (suc n))))} {P= : ft C P ≡ T-ftP sC Γ A A=} → 138 | Ob→ (T-d₁ sC Γ A A= P P=) ≡ T-d₁ sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ P) (! ft→ ∙ ap Ob→ P= ∙ T-ftP→) 139 | T-d₁→ {_} {Γ} {A} {A=} {P} {P=} = 140 | star→ ∙ ap-irr-star D 141 | (reflStr→ _ _ _ _ _ _ 142 | ∙ ap-irr-reflStr sD refl 143 | (star→ ∙ ap-irr-star D pp→ refl) 144 | (ss→ ∙ ap (ss D) id→)) 145 | (star→ ∙ ap-irr-star D 146 | (qq→ ∙ ap-irr-qq D 147 | (ss→ ∙ ap (ss D) id→) 148 | (star→ ∙ ap-irr-star D 149 | (qq→ ∙ ap-irr-qq D 150 | (ss→ ∙ ap (ss D) id→) 151 | (star→ ∙ ap-irr-star D 152 | (qq→ ∙ ap-irr-qq D pp→ refl) 153 | (star→ ∙ ap-irr-star D pp→ refl))) 154 | (star→ ∙ ap-irr-star D 155 | (qq→ ∙ ap-irr-qq D 156 | (qq→ ∙ ap-irr-qq D pp→ refl) 157 | (star→ ∙ ap-irr-star D pp→ refl)) 158 | T-ftP→))) 159 | (star→ ∙ ap-irr-star D 160 | (qq→ ∙ ap-irr-qq D 161 | (qq→ ∙ ap-irr-qq D 162 | (ss→ ∙ ap (ss D) id→) 163 | (star→ ∙ ap-irr-star D 164 | (qq→ ∙ ap-irr-qq D pp→ refl) 165 | (star→ ∙ ap-irr-star D pp→ refl))) 166 | (star→ ∙ ap-irr-star D 167 | (qq→ ∙ ap-irr-qq D 168 | (qq→ ∙ ap-irr-qq D pp→ refl) 169 | (star→ ∙ ap-irr-star D pp→ refl)) 170 | T-ftP→)) 171 | (star→ ∙ ap-irr-star D 172 | (qq→ ∙ ap-irr-qq D 173 | (qq→ ∙ ap-irr-qq D 174 | (qq→ ∙ ap-irr-qq D pp→ refl) 175 | (star→ ∙ ap-irr-star D pp→ refl)) 176 | T-ftP→) 177 | refl))) 178 | 179 | 180 | record StructuredCCatMor+ (sC sD : StructuredCCat) : Set where 181 | private 182 | open StructuredCCat 183 | C = ccat sC 184 | D = ccat sD 185 | 186 | 187 | field 188 | strucCCat→ : StructuredCCatMor sC sD 189 | 190 | open StructuredCCatMor strucCCat→ 191 | open CCatMor ccat→ 192 | open CCat+ 193 | 194 | field 195 | natelimStr→ : (Γ : Ob C n) (P : Ob C (suc (suc n))) (P= : ft C P ≡ NatStr sC Γ) 196 | (dO : Mor C n (suc n)) (dOₛ : is-section C dO) (dO₁ : ∂₁ C dO ≡ star C (zeroStr sC Γ) P P= (zeroStr₁ sC)) 197 | (dS : Mor C (suc (suc n)) (suc (suc (suc n)))) (dSₛ : is-section C dS) (dS₁ : ∂₁ C dS ≡ T-dS₁ sC Γ P P=) 198 | (u : Mor C n (suc n)) (uₛ : is-section C u) (u₁ : ∂₁ C u ≡ NatStr sC Γ) 199 | → Mor→ (natelimStr sC Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ u uₛ u₁) 200 | ≡ natelimStr sD (Ob→ Γ) (Ob→ P) (! ft→ ∙ ap Ob→ P= ∙ NatStr→ Γ) 201 | (Mor→ dO) (Mor→ₛ dOₛ) (Mor→₁ dO₁ ∙ star→ ∙ ap-irr-star D (zeroStr→ Γ) refl) 202 | (Mor→ dS) (Mor→ₛ dSₛ) (Mor→₁ dS₁ ∙ T-dS→) 203 | (Mor→ u) (Mor→ₛ uₛ) (Mor→₁ u₁ ∙ NatStr→ Γ) 204 | 205 | field 206 | jjStr→ : (Γ : Ob C n) (A : Ob C (suc n)) (A= : ft C A ≡ Γ) (P : Ob C (suc (suc (suc (suc n))))) 207 | (P= : ft C P ≡ CCatwithId.T-ftP (ccatId sC) Γ A A=) 208 | (d : Mor C (suc n) (suc (suc n))) (dₛ : is-section C d) 209 | (d₁ : ∂₁ C d ≡ CCatwithrefl.T-d₁ (ccatrefl sC) Γ A A= P P=) 210 | (a : Mor C n (suc n)) (aₛ : is-section C a) (a₁ : ∂₁ C a ≡ A) 211 | (b : Mor C n (suc n)) (bₛ : is-section C b) (b₁ : ∂₁ C b ≡ A) 212 | (p : Mor C n (suc n)) (pₛ : is-section C p) (p₁ : ∂₁ C p ≡ IdStr sC Γ A A= a aₛ a₁ b bₛ b₁) 213 | → Mor→ (jjStr sC Γ A A= P P= d dₛ d₁ a aₛ a₁ b bₛ b₁ p pₛ p₁) ≡ jjStr sD (Ob→ Γ) (Ob→ A) (! ft→ ∙ ap Ob→ A=) (Ob→ P) 214 | (! ft→ ∙ ap Ob→ P= ∙ T-ftP→) 215 | (Mor→ d) (Mor→ₛ dₛ) (Mor→₁ d₁ ∙ T-d₁→) 216 | (Mor→ a) (Mor→ₛ aₛ) (Mor→₁ a₁) 217 | (Mor→ b) (Mor→ₛ bₛ) (Mor→₁ b₁) 218 | (Mor→ p) (Mor→ₛ pₛ) (Mor→₁ p₁ ∙ IdStr→ Γ A A= a aₛ a₁ b bₛ b₁) 219 | 220 | 221 | module _ {sC sD : StructuredCCat} where 222 | open StructuredCCatMor+ 223 | open StructuredCCatMor 224 | open StructuredCCat 225 | open CCatMor 226 | open CCat 227 | 228 | {- Equalities between morphisms between structured contextual categories -} 229 | 230 | structuredCCatMorEq : {f g : StructuredCCatMor sC sD} 231 | → ({n : ℕ} (X : Ob (ccat sC) n) → Ob→ (ccat→ f) X ≡ Ob→ (ccat→ g) X) 232 | → ({n m : ℕ} (X : Mor (ccat sC) n m) → Mor→ (ccat→ f) X ≡ Mor→ (ccat→ g) X) 233 | → f ≡ g 234 | structuredCCatMorEq h k = lemma (funextI (λ n → funext h)) (funextI (λ m → funextI (λ n → funext k))) where 235 | 236 | lemma : {f g : StructuredCCatMor sC sD} 237 | → ((λ {n} → Ob→ (ccat→ f) {n}) ≡ (λ {n} → Ob→ (ccat→ g) {n})) 238 | → ((λ {n m} → Mor→ (ccat→ f) {n} {m}) ≡ (λ {n m} → Mor→ (ccat→ g) {n} {m})) 239 | → f ≡ g 240 | lemma refl refl = refl 241 | 242 | structuredCCatMor+Eq : {f+ g+ : StructuredCCatMor+ sC sD} (let f = strucCCat→ f+) (let g = strucCCat→ g+) 243 | → ({n : ℕ} (X : Ob (ccat sC) n) → Ob→ (ccat→ f) X ≡ Ob→ (ccat→ g) X) 244 | → ({n m : ℕ} (X : Mor (ccat sC) n m) → Mor→ (ccat→ f) X ≡ Mor→ (ccat→ g) X) 245 | → f+ ≡ g+ 246 | structuredCCatMor+Eq h k = lemma (structuredCCatMorEq h k) where 247 | 248 | lemma : {f+ g+ : StructuredCCatMor+ sC sD} (let f = strucCCat→ f+) (let g = strucCCat→ g+) 249 | → f ≡ g 250 | → f+ ≡ g+ 251 | lemma refl = refl 252 | -------------------------------------------------------------------------------- /initiality-uniqueness.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import reflection hiding (proj) 6 | open import syntx hiding (Mor; getTy) 7 | open import rules 8 | open import contextualcat 9 | open import contextualcatmor 10 | open import quotients 11 | open import termmodel 12 | import partialinterpretation 13 | -- import totality 14 | 15 | module _ (sC : StructuredCCat) where 16 | 17 | open StructuredCCat 18 | open CCatMor 19 | open partialinterpretation sC 20 | -- open import interpretationweakening sC 21 | -- open import interpretationsubstitution sC 22 | -- -- open totality sC 23 | open StructuredCCatMor+ 24 | open StructuredCCatMor 25 | 26 | private 27 | C = ccat sC 28 | 29 | open CCat+ C renaming (id to idC) 30 | 31 | {- Uniqueness of the morphism -} 32 | 33 | module _ (sf+ sg+ : StructuredCCatMor+ strSynCCat sC) where 34 | 35 | private 36 | sf = strucCCat→ sf+ 37 | sg = strucCCat→ sg+ 38 | f = ccat→ sf 39 | g = ccat→ sg 40 | 41 | -- Non-recursive lemmas that are needed for J 42 | 43 | uniqueness-Ty-//-Id : {Γ : Ctx n} (dΓ : ⊢ Γ) {A : TyExpr n} {u v : TmExpr n} (dA : Derivable (Γ ⊢ A)) (du : Derivable (Γ ⊢ u :> A)) (dv : Derivable (Γ ⊢ v :> A)) 44 | (uΓ : Ob→ f (proj (dctx dΓ)) ≡ Ob→ g (proj (dctx dΓ))) 45 | (uA : Ob→ f (proj (dctx (dΓ , dA))) ≡ Ob→ g (proj (dctx (dΓ , dA)))) 46 | (u-u : Mor→ f (proj (dmorTm (dctx dΓ) du)) ≡ Mor→ g (proj (dmorTm (dctx dΓ) du))) 47 | (uv : Mor→ f (proj (dmorTm (dctx dΓ) dv)) ≡ Mor→ g (proj (dmorTm (dctx dΓ) dv))) 48 | → Ob→ f (proj (dctx (dΓ , Id dA du dv))) ≡ Ob→ g (proj (dctx (dΓ , Id dA du dv))) 49 | uniqueness-Ty-//-Id dΓ dA du dv uΓ uA u-u uv = 50 | IdStr→ sf _ _ refl _ dmorTmₛ refl _ dmorTmₛ refl 51 | ∙ ap-irr-IdStr sC uΓ uA u-u uv 52 | ∙ ! (IdStr→ sg _ _ refl _ dmorTmₛ refl _ dmorTmₛ refl) 53 | 54 | uniqueness-Ty-//-Weak : {Γ : Ctx n} (dΓ : ⊢ Γ) {A B : TyExpr n} (dA : Derivable (Γ ⊢ A)) (dB : Derivable (Γ ⊢ B)) 55 | (uΓ : Ob→ f (proj (dctx dΓ)) ≡ Ob→ g (proj (dctx dΓ))) 56 | (uA : Ob→ f (proj (dctx (dΓ , dA))) ≡ Ob→ g (proj (dctx (dΓ , dA)))) 57 | (uB : Ob→ f (proj (dctx (dΓ , dB))) ≡ Ob→ g (proj (dctx (dΓ , dB)))) 58 | → Ob→ f (proj (dctx ((dΓ , dB) , SubstTy dA (WeakMor (idMorDerivable dΓ))))) 59 | ≡ Ob→ g (proj (dctx ((dΓ , dB) , SubstTy dA (WeakMor (idMorDerivable dΓ))))) 60 | uniqueness-Ty-//-Weak dΓ dA dB uΓ uA uB = 61 | star→ f {f = proj (dmor (dctx (dΓ , dB)) (dctx dΓ) (WeakMor (idMorDerivable dΓ)))} {X = proj (dctx (dΓ , dA))} {q = refl} {f₁ = refl} 62 | ∙ ap-irr-star (pp→ f {X = proj (dctx (dΓ , dB))} ∙ ap pp uB ∙ ! (pp→ g {X = proj (dctx (dΓ , dB))})) uA 63 | ∙ ! (star→ g {f = proj (dmor (dctx (dΓ , dB)) (dctx dΓ) (WeakMor (idMorDerivable dΓ)))} {X = proj (dctx (dΓ , dA))} {q = refl} {f₁ = refl}) 64 | 65 | uniqueness-Tm-//-Var : {Γ : Ctx n} (dΓ : ⊢ Γ) (k : VarPos n) 66 | → (uΓ : Ob→ f (proj (dctx dΓ)) ≡ Ob→ g (proj (dctx dΓ))) 67 | → {A : _} (A= : syntx.getTy k Γ ≡ A) 68 | → Mor→ f (proj (dmorTm (dctx dΓ) (congTmTy A= (Var k (getTyDer k dΓ))))) 69 | ≡ Mor→ g (proj (dmorTm (dctx dΓ) (congTmTy A= (Var k (getTyDer k dΓ))))) 70 | uniqueness-Tm-//-Var dΓ@(dΓ' , dA) last uΓ refl = 71 | ap (Mor→ f) (eq (box (CtxRefl dΓ) (CtxRefl dΓ , congTyRefl (WeakTy dA) (! (ap weakenTy ([idMor]Ty _)) ∙ weaken[]Ty _ (idMor _) last)) 72 | (MorRefl (idMorDerivable dΓ , congTmTy! ([idMor]Ty (weakenTy _)) (VarLast dA))))) 73 | ∙ (ss→ f {f = proj (dmor (dctx dΓ) (dctx dΓ) (idMorDerivable dΓ))} 74 | ∙ ap ss (id→ f ∙ ap idC uΓ ∙ ! (id→ g)) 75 | ∙ ! (ss→ g {f = proj (dmor (dctx dΓ) (dctx dΓ) (idMorDerivable dΓ))})) 76 | ∙ ! (ap (Mor→ g) (eq (box (CtxRefl dΓ) (CtxRefl dΓ , congTyRefl (WeakTy dA) (! (ap weakenTy ([idMor]Ty _)) ∙ weaken[]Ty _ (idMor _) last)) 77 | (MorRefl (idMorDerivable dΓ , congTmTy! ([idMor]Ty (weakenTy _)) (VarLast dA)))))) 78 | uniqueness-Tm-//-Var (dΓ , dB) (prev k) uΓ refl = 79 | ap (Mor→ f) (eq (box (CtxRefl (dΓ , dB)) (CtxRefl (dΓ , dB) , congTyRefl (getTyDer (prev k) (dΓ , dB)) (ap weakenTy (! ([idMor]Ty _)) ∙ weaken[]Ty _ (idMor _) last ∙ ap (λ z → syntx.getTy k _ [ z ]Ty) (! (idMor[]Mor _)))) 80 | (MorRefl (idMorDerivable (dΓ , dB)) , congTmRefl (congTm (! ([idMor]Ty _)) refl (Var (prev k) (getTyDer (prev k) (dΓ , dB)))) (ap weakenTm (! ([idMor]Tm (var k))) ∙ weaken[]Tm (var k) (idMor _) last)))) 81 | ∙ ss→ f {f = proj (dmor (dctx (dΓ , dB)) (dctx (dΓ , getTyDer k dΓ)) 82 | (SubstMor (idMorDerivable dΓ) (WeakMor (idMorDerivable dΓ)) , congTm (ap (λ z → syntx.getTy k _ [ z ]Ty) (! (idMor[]Mor _))) refl (SubstTm (Var k (getTyDer k dΓ)) (WeakMor (idMorDerivable dΓ)))))} 83 | ∙ ap ss (comp→ f {g = proj (dmor (dctx dΓ) (dctx (dΓ , getTyDer k dΓ)) (idMorDerivable dΓ , congTm (! ([idMor]Ty _)) refl (Var k (getTyDer k dΓ))))} {f = ppS (proj (dctx (dΓ , dB)))} {g₀ = refl} {f₁ = refl} 84 | ∙ ap-irr-comp (uniqueness-Tm-//-Var dΓ k (ft→ f ∙ ap ft uΓ ∙ ! (ft→ g)) refl) 85 | (pp→ f {X = proj (dctx (dΓ , dB))} 86 | ∙ ap pp uΓ 87 | ∙ ! (pp→ g {X = proj (dctx (dΓ , dB))})) 88 | ∙ ! (comp→ g {g = proj (dmor (dctx dΓ) (dctx (dΓ , getTyDer k dΓ)) (idMorDerivable dΓ , congTm (! ([idMor]Ty _)) refl (Var k (getTyDer k dΓ))))} 89 | {f = ppS (proj (dctx (dΓ , dB)))} {g₀ = refl} {f₁ = refl})) 90 | ∙ ! (ss→ g {f = proj (dmor (dctx (dΓ , dB)) (dctx (dΓ , getTyDer k dΓ)) 91 | (SubstMor (idMorDerivable dΓ) (WeakMor (idMorDerivable dΓ)) , congTm (ap (λ z → syntx.getTy k _ [ z ]Ty) (! (idMor[]Mor _))) refl (SubstTm (Var k (getTyDer k dΓ)) (WeakMor (idMorDerivable dΓ)))))}) 92 | ∙ ! (ap (Mor→ g) (eq (box (CtxRefl (dΓ , dB)) (CtxRefl (dΓ , dB) , congTyRefl (getTyDer (prev k) (dΓ , dB)) (ap weakenTy (! ([idMor]Ty _)) ∙ weaken[]Ty _ (idMor _) last ∙ ap (λ z → syntx.getTy k _ [ z ]Ty) (! (idMor[]Mor _)))) 93 | (MorRefl (idMorDerivable (dΓ , dB)) , congTmRefl (congTm (! ([idMor]Ty _)) refl (Var (prev k) (getTyDer (prev k) (dΓ , dB)))) (ap weakenTm (! ([idMor]Tm (var k))) ∙ weaken[]Tm (var k) (idMor _) last))))) 94 | 95 | -- The actual lemmas that we prove by mutual induction 96 | 97 | uniqueness-Ob-// : (Γ : DCtx n) 98 | → Ob→ f (proj Γ) ≡ Ob→ g (proj Γ) 99 | uniqueness-Ty-// : {Γ : Ctx n} (dΓ : ⊢ Γ) {A : TyExpr n} (dA : Derivable (Γ ⊢ A)) (uΓ : Ob→ f (proj (dctx dΓ)) ≡ Ob→ g (proj (dctx dΓ))) 100 | → Ob→ f (proj (dctx (dΓ , dA))) ≡ Ob→ g (proj (dctx (dΓ , dA))) 101 | uniqueness-Tm-// : {Γ : Ctx n} (dΓ : ⊢ Γ) {A : TyExpr n} {u : TmExpr n} (du : Derivable (Γ ⊢ u :> A)) 102 | (uΓ : Ob→ f (proj (dctx dΓ)) ≡ Ob→ g (proj (dctx dΓ))) 103 | → Mor→ f (proj (dmorTm (dctx dΓ) du)) ≡ Mor→ g (proj (dmorTm (dctx dΓ) du)) 104 | 105 | uniqueness-Ob-// (dctx' {ctx = ◇} tt) = pt→ f ∙ ! (pt→ g) 106 | uniqueness-Ob-// (dctx' {ctx = _ , _} (dΓ , dA)) = uniqueness-Ty-// dΓ dA (uniqueness-Ob-// (dctx dΓ)) 107 | 108 | uniqueness-Ty-// dΓ UU uΓ = 109 | UUStr→ sf _ _ 110 | ∙ ap (UUStr sC _) uΓ 111 | ∙ ! (UUStr→ sg _ _) 112 | uniqueness-Ty-// dΓ (El dv) uΓ = 113 | ElStr→ sf _ _ _ dmorTmₛ refl 114 | ∙ ap-irr-ElStr sC 115 | uΓ 116 | (uniqueness-Tm-// dΓ dv uΓ) 117 | ∙ ! (ElStr→ sg _ _ _ dmorTmₛ refl) 118 | uniqueness-Ty-// dΓ (Sum dA dB) uΓ = 119 | SumStr→ sf _ _ refl _ refl 120 | ∙ ap-irr-SumStr sC 121 | uΓ 122 | (uniqueness-Ty-// dΓ dA uΓ) 123 | (uniqueness-Ty-// dΓ dB uΓ) 124 | ∙ ! (SumStr→ sg _ _ refl _ refl) 125 | uniqueness-Ty-// dΓ (Pi dA dB) uΓ = 126 | PiStr→ sf _ _ refl _ refl 127 | ∙ ap-irr-PiStr sC 128 | uΓ 129 | (uniqueness-Ty-// dΓ dA uΓ) 130 | (uniqueness-Ty-// (dΓ , dA) dB (uniqueness-Ty-// dΓ dA uΓ)) 131 | ∙ ! (PiStr→ sg _ _ refl _ refl) 132 | uniqueness-Ty-// dΓ (Sig dA dB) uΓ = 133 | SigStr→ sf _ _ refl _ refl 134 | ∙ ap-irr-SigStr sC 135 | uΓ 136 | (uniqueness-Ty-// dΓ dA uΓ) 137 | (uniqueness-Ty-// (dΓ , dA) dB (uniqueness-Ty-// dΓ dA uΓ)) 138 | ∙ ! (SigStr→ sg _ _ refl _ refl) 139 | uniqueness-Ty-// dΓ Empty uΓ = 140 | EmptyStr→ sf _ 141 | ∙ ap (EmptyStr sC) uΓ 142 | ∙ ! (EmptyStr→ sg _) 143 | uniqueness-Ty-// dΓ Unit uΓ = 144 | UnitStr→ sf _ 145 | ∙ ap (UnitStr sC) uΓ 146 | ∙ ! (UnitStr→ sg _) 147 | uniqueness-Ty-// dΓ Nat uΓ = 148 | NatStr→ sf _ 149 | ∙ ap (NatStr sC) uΓ 150 | ∙ ! (NatStr→ sg _) 151 | uniqueness-Ty-// dΓ (Id dA du dv) uΓ = uniqueness-Ty-//-Id dΓ dA du dv uΓ (uniqueness-Ty-// dΓ dA uΓ) (uniqueness-Tm-// dΓ du uΓ) (uniqueness-Tm-// dΓ dv uΓ) 152 | 153 | uniqueness-Tm-// dΓ (Var k dA) uΓ = uniqueness-Tm-//-Var dΓ k uΓ refl 154 | 155 | uniqueness-Tm-// {Γ = Γ} dΓ (Conv dA du dA=) uΓ = 156 | ap (Mor→ f) (! (eq (box (CtxRefl dΓ) (CtxRefl dΓ , dA=) (MorRefl (idMorDerivable dΓ , congTm (! ([idMor]Ty _)) refl du))))) 157 | ∙ uniqueness-Tm-// dΓ du uΓ 158 | ∙ ! (ap (Mor→ g) (! (eq (box (CtxRefl dΓ) (CtxRefl dΓ , dA=) (MorRefl (idMorDerivable dΓ , congTm (! ([idMor]Ty _)) refl du)))))) 159 | 160 | uniqueness-Tm-// {Γ = Γ} dΓ {u = uu i} UUUU uΓ = 161 | uuStr→ sf i _ 162 | ∙ ap (uuStr sC i) uΓ 163 | ∙ ! (uuStr→ sg i _) 164 | 165 | uniqueness-Tm-// {Γ = Γ} dΓ {u = sum i a b} (SumUU da db) uΓ = 166 | sumStr→ sf i _ _ dmorTmₛ refl _ dmorTmₛ refl 167 | ∙ ap-irr-sumStr sC 168 | uΓ 169 | (uniqueness-Tm-// dΓ da uΓ) 170 | (uniqueness-Tm-// dΓ db uΓ) 171 | ∙ ! (sumStr→ sg i _ _ dmorTmₛ refl _ dmorTmₛ refl) 172 | uniqueness-Tm-// {Γ = Γ} dΓ {u = inl A B a} (Inl dA dB da) uΓ = 173 | inlStr→ sf _ _ refl _ refl _ dmorTmₛ refl 174 | ∙ ap-irr-inlStr sC 175 | uΓ 176 | (uniqueness-Ty-// dΓ dA uΓ) 177 | (uniqueness-Ty-// dΓ dB uΓ) 178 | (uniqueness-Tm-// dΓ da uΓ) 179 | ∙ ! (inlStr→ sg _ _ refl _ refl _ dmorTmₛ refl) 180 | uniqueness-Tm-// {Γ = Γ} dΓ {u = inr A B b} (Inr dA dB db) uΓ = 181 | inrStr→ sf _ _ refl _ refl _ dmorTmₛ refl 182 | ∙ ap-irr-inrStr sC 183 | uΓ 184 | (uniqueness-Ty-// dΓ dA uΓ) 185 | (uniqueness-Ty-// dΓ dB uΓ) 186 | (uniqueness-Tm-// dΓ db uΓ) 187 | ∙ ! (inrStr→ sg _ _ refl _ refl _ dmorTmₛ refl) 188 | uniqueness-Tm-// {Γ = Γ} dΓ {u = match A B C da db u} (Match dA dB dC dda ddb du) uΓ = 189 | matchStr→ sf _ _ refl _ refl _ refl _ dmorTmₛ refl _ dmorTmₛ refl _ dmorTmₛ refl 190 | ∙ ap-irr-matchStr sC 191 | uΓ 192 | (uniqueness-Ty-// dΓ dA uΓ) 193 | (uniqueness-Ty-// dΓ dB uΓ) 194 | (uniqueness-Ty-// (dΓ , Sum dA dB) dC (uniqueness-Ty-// dΓ (Sum dA dB) uΓ)) 195 | (uniqueness-Tm-// (dΓ , dA) (congTmTy fixTyda dda) (uniqueness-Ty-// dΓ dA uΓ)) 196 | (uniqueness-Tm-// (dΓ , dB) (congTmTy fixTydb ddb) (uniqueness-Ty-// dΓ dB uΓ)) 197 | (uniqueness-Tm-// dΓ du uΓ) 198 | ∙ ! (matchStr→ sg _ _ refl _ refl _ refl _ dmorTmₛ refl _ dmorTmₛ refl _ dmorTmₛ refl) 199 | 200 | uniqueness-Tm-// {Γ = Γ} dΓ {u = pi i a b} (PiUU da db) uΓ = 201 | piStr→ sf i _ _ dmorTmₛ refl _ dmorTmₛ refl 202 | ∙ ap-irr-piStr sC 203 | uΓ 204 | (uniqueness-Tm-// dΓ da uΓ) 205 | (uniqueness-Tm-// (dΓ , El da) db (uniqueness-Ty-// dΓ (El da) uΓ)) 206 | ∙ ! (piStr→ sg i _ _ dmorTmₛ refl _ dmorTmₛ refl) 207 | uniqueness-Tm-// {Γ = Γ} dΓ {u = lam A B u} (Lam dA dB du) uΓ = 208 | let uΓA = uniqueness-Ty-// dΓ dA uΓ in 209 | lamStr→ sf _ _ refl _ refl _ dmorTmₛ refl 210 | ∙ ap-irr-lamStr sC 211 | uΓ 212 | uΓA 213 | (uniqueness-Ty-// (dΓ , dA) dB uΓA) 214 | (uniqueness-Tm-// (dΓ , dA) du uΓA) 215 | ∙ ! (lamStr→ sg _ _ refl _ refl _ dmorTmₛ refl) 216 | uniqueness-Tm-// {Γ = Γ} dΓ {u = app A B f a} (App dA dB df da) uΓ = 217 | let uΓA = uniqueness-Ty-// dΓ dA uΓ in 218 | appStr→ sf _ _ refl _ refl _ dmorTmₛ refl _ dmorTmₛ refl 219 | ∙ ap-irr-appStr sC 220 | uΓ 221 | uΓA 222 | (uniqueness-Ty-// (dΓ , dA) dB uΓA) 223 | (uniqueness-Tm-// dΓ df uΓ) 224 | (uniqueness-Tm-// dΓ da uΓ) 225 | ∙ ! (appStr→ sg _ _ refl _ refl _ dmorTmₛ refl _ dmorTmₛ refl) 226 | 227 | uniqueness-Tm-// {Γ = Γ} dΓ {u = sig i a b} (SigUU da db) uΓ = 228 | sigStr→ sf i _ _ dmorTmₛ refl _ dmorTmₛ refl 229 | ∙ ap-irr-sigStr sC 230 | uΓ 231 | (uniqueness-Tm-// dΓ da uΓ) 232 | (uniqueness-Tm-// (dΓ , El da) db (uniqueness-Ty-// dΓ (El da) uΓ)) 233 | ∙ ! (sigStr→ sg i _ _ dmorTmₛ refl _ dmorTmₛ refl) 234 | uniqueness-Tm-// {Γ = Γ} dΓ {u = pair A B a b} (Pair dA dB da db) uΓ = 235 | let uΓA = uniqueness-Ty-// dΓ dA uΓ in 236 | pairStr→ sf _ _ refl _ refl _ dmorTmₛ refl _ dmorTmₛ refl 237 | ∙ ap-irr-pairStr sC 238 | uΓ 239 | uΓA 240 | (uniqueness-Ty-// (dΓ , dA) dB uΓA) 241 | (uniqueness-Tm-// dΓ da uΓ) 242 | (uniqueness-Tm-// dΓ db uΓ) 243 | ∙ ! (pairStr→ sg _ _ refl _ refl _ dmorTmₛ refl _ dmorTmₛ refl) 244 | uniqueness-Tm-// {Γ = Γ} dΓ {u = pr1 A B u} (Pr1 dA dB du) uΓ = 245 | let uΓA = uniqueness-Ty-// dΓ dA uΓ in 246 | pr1Str→ sf _ _ refl _ refl _ dmorTmₛ refl 247 | ∙ ap-irr-pr1Str sC 248 | uΓ 249 | uΓA 250 | (uniqueness-Ty-// (dΓ , dA) dB uΓA) 251 | (uniqueness-Tm-// dΓ du uΓ) 252 | ∙ ! (pr1Str→ sg _ _ refl _ refl _ dmorTmₛ refl) 253 | uniqueness-Tm-// {Γ = Γ} dΓ {u = pr2 A B u} (Pr2 dA dB du) uΓ = 254 | let uΓA = uniqueness-Ty-// dΓ dA uΓ in 255 | pr2Str→ sf _ _ refl _ refl _ dmorTmₛ refl 256 | ∙ ap-irr-pr2Str sC 257 | uΓ 258 | uΓA 259 | (uniqueness-Ty-// (dΓ , dA) dB uΓA) 260 | (uniqueness-Tm-// dΓ du uΓ) 261 | ∙ ! (pr2Str→ sg _ _ refl _ refl _ dmorTmₛ refl) 262 | 263 | uniqueness-Tm-// {Γ = Γ} dΓ {u = empty i} EmptyUU uΓ = 264 | emptyStr→ sf i _ 265 | ∙ ap (emptyStr sC i) uΓ 266 | ∙ ! (emptyStr→ sg i _) 267 | uniqueness-Tm-// {Γ = Γ} dΓ {u = emptyelim A u} (Emptyelim dA du) uΓ = 268 | emptyelimStr→ sf _ _ refl _ dmorTmₛ refl 269 | ∙ ap-irr-emptyelimStr sC 270 | uΓ 271 | (uniqueness-Ty-// (dΓ , Empty) dA (uniqueness-Ty-// dΓ Empty uΓ)) 272 | (uniqueness-Tm-// dΓ du uΓ) 273 | ∙ ! (emptyelimStr→ sg _ _ refl _ dmorTmₛ refl) 274 | 275 | uniqueness-Tm-// {Γ = Γ} dΓ {u = unit i} UnitUU uΓ = 276 | unitStr→ sf i _ 277 | ∙ ap (unitStr sC i) uΓ 278 | ∙ ! (unitStr→ sg i _) 279 | uniqueness-Tm-// {Γ = Γ} dΓ {u = tt} TT uΓ = 280 | ttStr→ sf _ 281 | ∙ ap (ttStr sC) uΓ 282 | ∙ ! (ttStr→ sg _) 283 | uniqueness-Tm-// {Γ = Γ} dΓ {u = unitelim A dtt u} (Unitelim dA ddtt du) uΓ = 284 | unitelimStr→ sf _ _ refl _ dmorTmₛ refl _ dmorTmₛ refl 285 | ∙ ap-irr-unitelimStr sC 286 | uΓ 287 | (uniqueness-Ty-// (dΓ , Unit) dA (uniqueness-Ty-// dΓ Unit uΓ)) 288 | (uniqueness-Tm-// dΓ ddtt uΓ) 289 | (uniqueness-Tm-// dΓ du uΓ) 290 | ∙ ! (unitelimStr→ sg _ _ refl _ dmorTmₛ refl _ dmorTmₛ refl) 291 | 292 | uniqueness-Tm-// {Γ = Γ} dΓ {u = nat i} NatUU uΓ = 293 | natStr→ sf i _ 294 | ∙ ap (natStr sC i) uΓ 295 | ∙ ! (natStr→ sg i _) 296 | uniqueness-Tm-// {Γ = Γ} dΓ {u = zero} Zero uΓ = 297 | zeroStr→ sf _ 298 | ∙ ap (zeroStr sC) uΓ 299 | ∙ ! (zeroStr→ sg _) 300 | uniqueness-Tm-// {Γ = Γ} dΓ {u = suc u} (Suc du) uΓ = 301 | sucStr→ sf _ _ dmorTmₛ refl 302 | ∙ ap-irr-sucStr sC uΓ 303 | (uniqueness-Tm-// dΓ du uΓ) 304 | ∙ ! (sucStr→ sg _ _ dmorTmₛ refl) 305 | uniqueness-Tm-// {Γ = Γ} dΓ {u = natelim P dO dS u} (Natelim dP ddO ddS du) uΓ = 306 | natelimStr→ sf+ _ _ refl _ dmorTmₛ refl _ dmorTmₛ refl _ dmorTmₛ refl 307 | ∙ ap-irr-natelimStr sC 308 | uΓ 309 | (uniqueness-Ty-// (dΓ , Nat) dP (uniqueness-Ty-// dΓ Nat uΓ)) 310 | (uniqueness-Tm-// dΓ ddO uΓ) 311 | (uniqueness-Tm-// ((dΓ , Nat) , dP) (congTmTy fixSubstTy ddS) (uniqueness-Ty-// (dΓ , Nat) dP (uniqueness-Ty-// dΓ Nat uΓ))) 312 | (uniqueness-Tm-// dΓ du uΓ) 313 | ∙ ! (natelimStr→ sg+ _ _ refl _ dmorTmₛ refl _ dmorTmₛ refl _ dmorTmₛ refl) 314 | 315 | uniqueness-Tm-// {Γ = Γ} dΓ {u = id i a u v} (IdUU da du dv) uΓ = 316 | idStr→ sf i _ _ dmorTmₛ refl _ dmorTmₛ refl _ dmorTmₛ refl 317 | ∙ ap-irr-idStr sC 318 | uΓ 319 | (uniqueness-Tm-// dΓ da uΓ) 320 | (uniqueness-Tm-// dΓ du uΓ) 321 | (uniqueness-Tm-// dΓ dv uΓ) 322 | ∙ ! (idStr→ sg i _ _ dmorTmₛ refl _ dmorTmₛ refl _ dmorTmₛ refl) 323 | uniqueness-Tm-// {Γ = Γ} dΓ {u = refl A u} (Refl dA du) uΓ = 324 | reflStr→ sf _ _ refl _ dmorTmₛ refl 325 | ∙ ap-irr-reflStr sC uΓ 326 | (uniqueness-Ty-// dΓ dA uΓ) 327 | (uniqueness-Tm-// dΓ du uΓ) 328 | ∙ ! (reflStr→ sg _ _ refl _ dmorTmₛ refl) 329 | uniqueness-Tm-// {Γ = Γ} dΓ {u = jj A P d a b p} (JJ dA dP dd da db dp) uΓ = 330 | let uΓA = uniqueness-Ty-// dΓ dA uΓ in 331 | let uΓAwA = uniqueness-Ty-//-Weak dΓ dA dA uΓ uΓA uΓA in 332 | let dwA = SubstTy dA (WeakMor (idMorDerivable dΓ)) in 333 | let dwwA = SubstTy dwA ((WeakMor (WeakMor (idMorDerivable dΓ))) , congTmTy (ap weakenTy weakenTy-to-[]Ty ∙ weaken[]Ty _ _ last) (VarPrevLast dA)) in 334 | let dvpl = congTmTy (ap weakenTy weakenTy-to-[]Ty ∙ weakenTy-to-[]Ty) (VarPrevLast dA) in 335 | let dvl = congTmTy weakenTy-to-[]Ty (VarLast dwA) in 336 | let did = Id dwwA dvpl dvl in 337 | let dP' = congTyCtx (Ctx+= (Ctx+= refl weakenTy-to-[]Ty) (ap-id-Ty (ap weakenTy weakenTy-to-[]Ty ∙ weakenTy-to-[]Ty) refl refl)) dP in 338 | jjStr→ sf+ (proj (dctx dΓ)) (proj (dctx (dΓ , dA))) refl 339 | (proj (dctx {ctx = ((((_ , _) , _) , _) , _)} ((((dΓ , dA) , dwA) , did) , dP'))) refl 340 | (proj (dmor (dctx {ctx = (_ , _)} (dΓ , dA)) (dctx {ctx = ((_ , _) , _)} ((dΓ , dA) , DerTmTy (dΓ , dA) (congTmTy fixTyJJ dd))) {mor = (idMor _ , d)} (idMor+ (dΓ , dA) (congTmTy fixTyJJ dd)))) dmorTmₛ refl 341 | (proj (dmor (dctx {ctx = Γ} dΓ) (dctx {ctx = (_ , _)} (dΓ , dA)) {mor = (idMor _ , a)} (idMor+ dΓ da))) dmorTmₛ refl 342 | (proj (dmor (dctx {ctx = Γ} dΓ) (dctx {ctx = (_ , _)} (dΓ , dA)) {mor = (idMor _ , b)} (idMor+ dΓ db))) dmorTmₛ refl 343 | (proj (dmor (dctx {ctx = Γ} dΓ) (dctx {ctx = (_ , _)} (dΓ , Id dA da db)) {mor = (idMor _ , p)} (idMor+ dΓ dp))) dmorTmₛ refl 344 | ∙ ap-irr-jjStr sC 345 | uΓ 346 | uΓA 347 | (uniqueness-Ty-// (((dΓ , dA) , dwA) , did) dP' 348 | (uniqueness-Ty-//-Id ((dΓ , dA) , dwA) dwwA dvpl dvl uΓAwA 349 | (uniqueness-Ty-//-Weak (dΓ , dA) dwA dwA uΓA uΓAwA uΓAwA) 350 | (uniqueness-Tm-//-Var ((dΓ , dA) , dwA) (prev last) uΓAwA (ap weakenTy weakenTy-to-[]Ty ∙ weakenTy-to-[]Ty)) 351 | (uniqueness-Tm-//-Var ((dΓ , dA) , dwA) last uΓAwA weakenTy-to-[]Ty))) 352 | (uniqueness-Tm-// (dΓ , dA) (congTmTy fixTyJJ dd) uΓA) 353 | (uniqueness-Tm-// dΓ da uΓ) 354 | (uniqueness-Tm-// dΓ db uΓ) 355 | (uniqueness-Tm-// dΓ dp uΓ) 356 | ∙ ! (jjStr→ sg+ _ (proj (dctx (dΓ , dA))) refl 357 | (proj (dctx {ctx = ((((_ , _) , _) , _) , _)} ((((dΓ , dA) , dwA) , did) , dP'))) refl 358 | (proj (dmor (dctx {ctx = (_ , _)} (dΓ , dA)) (dctx {ctx = ((_ , _) , _)} ((dΓ , dA) , DerTmTy (dΓ , dA) (congTmTy fixTyJJ dd))) {mor = (idMor _ , d)} (idMor+ (dΓ , dA) (congTmTy fixTyJJ dd)))) dmorTmₛ refl 359 | (proj (dmor (dctx {ctx = Γ} dΓ) (dctx {ctx = (_ , _)} (dΓ , dA)) {mor = (idMor _ , a)} (idMor+ dΓ da))) dmorTmₛ refl 360 | (proj (dmor (dctx {ctx = Γ} dΓ) (dctx {ctx = (_ , _)} (dΓ , dA)) {mor = (idMor _ , b)} (idMor+ dΓ db))) dmorTmₛ refl 361 | (proj (dmor (dctx {ctx = Γ} dΓ) (dctx {ctx = (_ , _)} (dΓ , Id dA da db)) {mor = (idMor _ , p)} (idMor+ dΓ dp))) dmorTmₛ refl) 362 | 363 | uniqueness-Ob : (X : ObS n) → Ob→ f X ≡ Ob→ g X 364 | uniqueness-Ob = //-elimP uniqueness-Ob-// 365 | 366 | -- showing [δ,t] = qq([δ]) ∘ [id,t] directly via syntactic equality, rather than first via ss-qq. 367 | uniqueness-Mor-// : (δ : DMor n m) → Mor→ f (proj δ) ≡ Mor→ g (proj δ) 368 | uniqueness-Mor-// (dmor' (dctx' dΓ) (dctx' {ctx = ◇} tt) {mor = ◇} tt) = 369 | ptmor→ f {X = proj (dctx dΓ)} 370 | ∙ ap ptmor (uniqueness-Ob-// (dctx dΓ)) 371 | ∙ ! (ptmor→ g) 372 | uniqueness-Mor-// δδ@(dmor' (dctx' dΓ) (dctx' {ctx = Δ , C} (dΔ , dC)) {mor = (δ , u)} (dδ , du)) = 373 | ap (Mor→ f) 374 | (eq (box (CtxRefl dΓ) 375 | (CtxRefl (dΔ , dC)) 376 | (congMorEq refl refl refl (! (Mor+= (weakenMorInsert _ _ _ ∙ [idMor]Mor _) refl)) 377 | (MorRefl (dδ , du))))) 378 | ∙ comp→ f {g₀ = refl} {f₁ = refl} 379 | ∙ ap-irr-comp (qq→ f {q = refl} {f₁ = refl} 380 | ∙ ap-irr-qq (uniqueness-Mor-// (dmor' (dctx' dΓ) (dctx' dΔ) dδ)) 381 | (uniqueness-Ob-// (dctx' (dΔ , dC))) 382 | ∙ ! (qq→ g {q = refl} {f₁ = refl})) 383 | (uniqueness-Tm-// dΓ du (uniqueness-Ob-// (dctx' dΓ))) 384 | ∙ ! (ap (Mor→ g) 385 | (eq (box (CtxRefl dΓ) 386 | (CtxRefl (dΔ , dC)) 387 | (congMorEq refl refl refl (! (Mor+= (weakenMorInsert _ _ _ ∙ [idMor]Mor _) refl)) 388 | (MorRefl (dδ , du))))) 389 | ∙ comp→ g {g₀ = refl} {f₁ = refl}) 390 | 391 | uniqueness-Mor : (X : MorS n m) → Mor→ f X ≡ Mor→ g X 392 | uniqueness-Mor = //-elimP uniqueness-Mor-// 393 | 394 | uniqueness : sf ≡ sg 395 | uniqueness = structuredCCatMorEq uniqueness-Ob uniqueness-Mor 396 | 397 | uniqueness+ : sf+ ≡ sg+ 398 | uniqueness+ = structuredCCatMor+Eq uniqueness-Ob uniqueness-Mor 399 | -------------------------------------------------------------------------------- /partialinterpretation.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import syntx 6 | open import contextualcat 7 | 8 | module _ (sC : StructuredCCat) where 9 | 10 | open StructuredCCat sC 11 | open CCat+ ccat renaming (Mor to MorC; id to idC) 12 | 13 | 14 | {- Partial interpretation of types and terms -} 15 | 16 | ⟦_⟧Ty : TyExpr n → (X : Ob n) → Partial (Ob (suc n)) 17 | ⟦_⟧Tm : TmExpr n → (X : Ob n) → Partial (MorC n (suc n)) 18 | 19 | ⟦ uu i ⟧Ty X = return (UUStr i X) 20 | ⟦ el i v ⟧Ty X = do 21 | [v] ← ⟦ v ⟧Tm X 22 | [v]ₛ ← assume (is-section [v]) 23 | [v]₁ ← assume (∂₁ [v] ≡ UUStr i X) 24 | return (ElStr i X [v] (unbox [v]ₛ) (unbox [v]₁)) 25 | ⟦ sum A B ⟧Ty X = do 26 | [A] ← ⟦ A ⟧Ty X 27 | [A]= ← assume (ft [A] ≡ X) 28 | [B] ← ⟦ B ⟧Ty X 29 | [B]= ← assume (ft [B] ≡ X) 30 | return (SumStr X [A] (unbox [A]=) [B] (unbox [B]=)) 31 | ⟦ pi A B ⟧Ty X = do 32 | [A] ← ⟦ A ⟧Ty X 33 | [A]= ← assume (ft [A] ≡ X) 34 | [B] ← ⟦ B ⟧Ty [A] 35 | [B]= ← assume (ft [B] ≡ [A]) 36 | return (PiStr X [A] (unbox [A]=) [B] (unbox [B]=)) 37 | ⟦ sig A B ⟧Ty X = do 38 | [A] ← ⟦ A ⟧Ty X 39 | [A]= ← assume (ft [A] ≡ X) 40 | [B] ← ⟦ B ⟧Ty [A] 41 | [B]= ← assume (ft [B] ≡ [A]) 42 | return (SigStr X [A] (unbox [A]=) [B] (unbox [B]=)) 43 | ⟦ empty ⟧Ty X = return (EmptyStr X) 44 | ⟦ unit ⟧Ty X = return (UnitStr X) 45 | ⟦ nat ⟧Ty X = return (NatStr X) 46 | ⟦ id A u v ⟧Ty X = do 47 | [A] ← ⟦ A ⟧Ty X 48 | [A]= ← assume (ft [A] ≡ X) 49 | [u] ← ⟦ u ⟧Tm X 50 | [u]ₛ ← assume (is-section [u]) 51 | [u]₁ ← assume (∂₁ [u] ≡ [A]) 52 | [v] ← ⟦ v ⟧Tm X 53 | [v]ₛ ← assume (is-section [v]) 54 | [v]₁ ← assume (∂₁ [v] ≡ [A]) 55 | return (IdStr X [A] (unbox [A]=) [u] (unbox [u]ₛ) (unbox [u]₁) [v] (unbox [v]ₛ) (unbox [v]₁)) 56 | 57 | 58 | ⟦ var k ⟧Tm X = return (varC k X) 59 | ⟦ uu i ⟧Tm X = return (uuStr i X) 60 | ⟦ sum i a b ⟧Tm X = do 61 | [a] ← ⟦ a ⟧Tm X 62 | [a]ₛ ← assume (is-section [a]) 63 | [a]₁ ← assume (∂₁ [a] ≡ UUStr i X) 64 | [b] ← ⟦ b ⟧Tm X 65 | [b]ₛ ← assume (is-section [b]) 66 | [b]₁ ← assume (∂₁ [b] ≡ UUStr i X) 67 | return (sumStr i X [a] (unbox [a]ₛ) (unbox [a]₁) [b] (unbox [b]ₛ) (unbox [b]₁)) 68 | ⟦ inl A B a ⟧Tm X = do 69 | [A] ← ⟦ A ⟧Ty X 70 | [A]= ← assume (ft [A] ≡ X) 71 | [B] ← ⟦ B ⟧Ty X 72 | [B]= ← assume (ft [B] ≡ X) 73 | [a] ← ⟦ a ⟧Tm X 74 | [a]ₛ ← assume (is-section [a]) 75 | [a]₁ ← assume (∂₁ [a] ≡ [A]) 76 | return (inlStr X [A] (unbox [A]=) [B] (unbox [B]=) [a] (unbox [a]ₛ) (unbox [a]₁)) 77 | ⟦ inr A B b ⟧Tm X = do 78 | [A] ← ⟦ A ⟧Ty X 79 | [A]= ← assume (ft [A] ≡ X) 80 | [B] ← ⟦ B ⟧Ty X 81 | [B]= ← assume (ft [B] ≡ X) 82 | [b] ← ⟦ b ⟧Tm X 83 | [b]ₛ ← assume (is-section [b]) 84 | [b]₁ ← assume (∂₁ [b] ≡ [B]) 85 | return (inrStr X [A] (unbox [A]=) [B] (unbox [B]=) [b] (unbox [b]ₛ) (unbox [b]₁)) 86 | ⟦ match A B C da db u ⟧Tm X = do 87 | [A] ← ⟦ A ⟧Ty X 88 | [A]= ← assume (ft [A] ≡ X) 89 | [B] ← ⟦ B ⟧Ty X 90 | [B]= ← assume (ft [B] ≡ X) 91 | [C] ← ⟦ C ⟧Ty (SumStr X [A] (unbox [A]=) [B] (unbox [B]=)) 92 | [C]= ← assume (ft [C] ≡ SumStr X [A] (unbox [A]=) [B] (unbox [B]=)) 93 | [da] ← ⟦ da ⟧Tm [A] 94 | [da]ₛ ← assume (is-section [da]) 95 | [da]₁ ← assume (∂₁ [da] ≡ T-da₁ X [A] (unbox [A]=) [B] (unbox [B]=) [C] (unbox [C]=)) 96 | [db] ← ⟦ db ⟧Tm [B] 97 | [db]ₛ ← assume (is-section [db]) 98 | [db]₁ ← assume (∂₁ [db] ≡ T-db₁ X [A] (unbox [A]=) [B] (unbox [B]=) [C] (unbox [C]=)) 99 | [u] ← ⟦ u ⟧Tm X 100 | [u]ₛ ← assume (is-section [u]) 101 | [u]₁ ← assume (∂₁ [u] ≡ SumStr X [A] (unbox [A]=) [B] (unbox [B]=)) 102 | return (matchStr X [A] (unbox [A]=) [B] (unbox [B]=) [C] (unbox [C]=) [da] (unbox [da]ₛ) (unbox [da]₁) [db] (unbox [db]ₛ) (unbox [db]₁) [u] (unbox [u]ₛ) (unbox [u]₁)) 103 | ⟦ pi i a b ⟧Tm X = do 104 | [a] ← ⟦ a ⟧Tm X 105 | [a]ₛ ← assume (is-section [a]) 106 | [a]₁ ← assume (∂₁ [a] ≡ UUStr i X) 107 | [b] ← ⟦ b ⟧Tm (ElStr i X [a] (unbox [a]ₛ) (unbox [a]₁)) 108 | [b]ₛ ← assume (is-section [b]) 109 | [b]₁ ← assume (∂₁ [b] ≡ UUStr i (ElStr i X [a] (unbox [a]ₛ) (unbox [a]₁))) 110 | return (piStr i X [a] (unbox [a]ₛ) (unbox [a]₁) [b] (unbox [b]ₛ) (unbox [b]₁)) 111 | ⟦ lam A B u ⟧Tm X = do 112 | [A] ← ⟦ A ⟧Ty X 113 | [A]= ← assume (ft [A] ≡ X) 114 | [B] ← ⟦ B ⟧Ty [A] 115 | [B]= ← assume (ft [B] ≡ [A]) 116 | [u] ← ⟦ u ⟧Tm [A] 117 | [u]ₛ ← assume (is-section [u]) 118 | [u]₁ ← assume (∂₁ [u] ≡ [B]) 119 | return (lamStr X [A] (unbox [A]=) [B] (unbox [B]=) [u] (unbox [u]ₛ) (unbox [u]₁)) 120 | ⟦ app A B f a ⟧Tm X = do 121 | [A] ← ⟦ A ⟧Ty X 122 | [A]= ← assume (ft [A] ≡ X) 123 | [B] ← ⟦ B ⟧Ty [A] 124 | [B]= ← assume (ft [B] ≡ [A]) 125 | [f] ← ⟦ f ⟧Tm X 126 | [f]ₛ ← assume (is-section [f]) 127 | [f]₁ ← assume (∂₁ [f] ≡ PiStr X [A] (unbox [A]=) [B] (unbox [B]=)) 128 | [a] ← ⟦ a ⟧Tm X 129 | [a]ₛ ← assume (is-section [a]) 130 | [a]₁ ← assume (∂₁ [a] ≡ [A]) 131 | return (appStr X [A] (unbox [A]=) [B] (unbox [B]=) [f] (unbox [f]ₛ) (unbox [f]₁) [a] (unbox [a]ₛ) (unbox [a]₁)) 132 | ⟦ sig i a b ⟧Tm X = do 133 | [a] ← ⟦ a ⟧Tm X 134 | [a]ₛ ← assume (is-section [a]) 135 | [a]₁ ← assume (∂₁ [a] ≡ UUStr i X) 136 | [b] ← ⟦ b ⟧Tm (ElStr i X [a] (unbox [a]ₛ) (unbox [a]₁)) 137 | [b]ₛ ← assume (is-section [b]) 138 | [b]₁ ← assume (∂₁ [b] ≡ UUStr i (ElStr i X [a] (unbox [a]ₛ) (unbox [a]₁))) 139 | return (sigStr i X [a] (unbox [a]ₛ) (unbox [a]₁) [b] (unbox [b]ₛ) (unbox [b]₁)) 140 | ⟦ pair A B u v ⟧Tm X = do 141 | [A] ← ⟦ A ⟧Ty X 142 | [A]= ← assume (ft [A] ≡ X) 143 | [B] ← ⟦ B ⟧Ty [A] 144 | [B]= ← assume (ft [B] ≡ [A]) 145 | [u] ← ⟦ u ⟧Tm X 146 | [u]ₛ ← assume (is-section [u]) 147 | [u]₁ ← assume (∂₁ [u] ≡ [A]) 148 | [v] ← ⟦ v ⟧Tm X 149 | [v]ₛ ← assume (is-section [v]) 150 | [v]₁ ← assume (∂₁ [v] ≡ star [u] [B] (unbox [B]=) (unbox [u]₁)) 151 | return (pairStr X [A] (unbox [A]=) [B] (unbox [B]=) [u] (unbox [u]ₛ) (unbox [u]₁) [v] (unbox [v]ₛ) (unbox [v]₁)) 152 | ⟦ pr1 A B u ⟧Tm X = do 153 | [A] ← ⟦ A ⟧Ty X 154 | [A]= ← assume (ft [A] ≡ X) 155 | [B] ← ⟦ B ⟧Ty [A] 156 | [B]= ← assume (ft [B] ≡ [A]) 157 | [u] ← ⟦ u ⟧Tm X 158 | [u]ₛ ← assume (is-section [u]) 159 | [u]₁ ← assume (∂₁ [u] ≡ SigStr X [A] (unbox [A]=) [B] (unbox [B]=)) 160 | return (pr1Str X [A] (unbox [A]=) [B] (unbox [B]=) [u] (unbox [u]ₛ) (unbox [u]₁)) 161 | ⟦ pr2 A B u ⟧Tm X = do 162 | [A] ← ⟦ A ⟧Ty X 163 | [A]= ← assume (ft [A] ≡ X) 164 | [B] ← ⟦ B ⟧Ty [A] 165 | [B]= ← assume (ft [B] ≡ [A]) 166 | [u] ← ⟦ u ⟧Tm X 167 | [u]ₛ ← assume (is-section [u]) 168 | [u]₁ ← assume (∂₁ [u] ≡ SigStr X [A] (unbox [A]=) [B] (unbox [B]=)) 169 | return (pr2Str X [A] (unbox [A]=) [B] (unbox [B]=) [u] (unbox [u]ₛ) (unbox [u]₁)) 170 | ⟦ empty i ⟧Tm X = return (emptyStr i X) 171 | ⟦ emptyelim A u ⟧Tm X = do 172 | [A] ← ⟦ A ⟧Ty (EmptyStr X) 173 | [A]= ← assume (ft [A] ≡ EmptyStr X) 174 | [u] ← ⟦ u ⟧Tm X 175 | [u]ₛ ← assume (is-section [u]) 176 | [u]₁ ← assume (∂₁ [u] ≡ EmptyStr X) 177 | return (emptyelimStr X [A] (unbox [A]=) [u] (unbox [u]ₛ) (unbox [u]₁)) 178 | ⟦ unit i ⟧Tm X = return (unitStr i X) 179 | ⟦ tt ⟧Tm X = return (ttStr X) 180 | ⟦ unitelim A dtt u ⟧Tm X = do 181 | [A] ← ⟦ A ⟧Ty (UnitStr X) 182 | [A]= ← assume (ft [A] ≡ UnitStr X) 183 | [dtt] ← ⟦ dtt ⟧Tm X 184 | [dtt]ₛ ← assume (is-section [dtt]) 185 | [dtt]₁ ← assume (∂₁ [dtt] ≡ star (ttStr X) [A] (unbox [A]=) ttStr₁) 186 | [u] ← ⟦ u ⟧Tm X 187 | [u]ₛ ← assume (is-section [u]) 188 | [u]₁ ← assume (∂₁ [u] ≡ UnitStr X) 189 | return (unitelimStr X [A] (unbox [A]=) [dtt] (unbox [dtt]ₛ) (unbox [dtt]₁) [u] (unbox [u]ₛ) (unbox [u]₁)) 190 | ⟦ nat i ⟧Tm X = return (natStr i X) 191 | ⟦ zero ⟧Tm X = return (zeroStr X) 192 | ⟦ suc u ⟧Tm X = do 193 | [u] ← ⟦ u ⟧Tm X 194 | [u]ₛ ← assume (is-section [u]) 195 | [u]₁ ← assume (∂₁ [u] ≡ NatStr X) 196 | return (sucStr X [u] (unbox [u]ₛ) (unbox [u]₁)) 197 | ⟦ natelim P dO dS u ⟧Tm X = do 198 | [P] ← ⟦ P ⟧Ty (NatStr X) 199 | [P]= ← assume (ft [P] ≡ NatStr X) 200 | [dO] ← ⟦ dO ⟧Tm X 201 | [dO]ₛ ← assume (is-section [dO]) 202 | [dO]₁ ← assume (∂₁ [dO] ≡ star (zeroStr X) [P] (unbox [P]=) zeroStr₁) 203 | [dS] ← ⟦ dS ⟧Tm [P] 204 | [dS]ₛ ← assume (is-section [dS]) 205 | [dS]₁ ← assume (∂₁ [dS] ≡ T-dS₁ X [P] (unbox [P]=)) 206 | [u] ← ⟦ u ⟧Tm X 207 | [u]ₛ ← assume (is-section [u]) 208 | [u]₁ ← assume (∂₁ [u] ≡ NatStr X) 209 | return (natelimStr X [P] (unbox [P]=) [dO] (unbox [dO]ₛ) (unbox [dO]₁) [dS] (unbox [dS]ₛ) (unbox [dS]₁) [u] (unbox [u]ₛ) (unbox [u]₁)) 210 | ⟦ id i a u v ⟧Tm X = do 211 | [a] ← ⟦ a ⟧Tm X 212 | [a]ₛ ← assume (is-section [a]) 213 | [a]₁ ← assume (∂₁ [a] ≡ UUStr i X) 214 | [u] ← ⟦ u ⟧Tm X 215 | [u]ₛ ← assume (is-section [u]) 216 | [u]₁ ← assume (∂₁ [u] ≡ ElStr i X [a] (unbox [a]ₛ) (unbox [a]₁)) 217 | [v] ← ⟦ v ⟧Tm X 218 | [v]ₛ ← assume (is-section [v]) 219 | [v]₁ ← assume (∂₁ [v] ≡ ElStr i X [a] (unbox [a]ₛ) (unbox [a]₁)) 220 | return (idStr i X [a] (unbox [a]ₛ) (unbox [a]₁) [u] (unbox [u]ₛ) (unbox [u]₁) [v] (unbox [v]ₛ) (unbox [v]₁)) 221 | ⟦ refl A u ⟧Tm X = do 222 | [A] ← ⟦ A ⟧Ty X 223 | [A]= ← assume (ft [A] ≡ X) 224 | [u] ← ⟦ u ⟧Tm X 225 | [u]ₛ ← assume (is-section [u]) 226 | [u]₁ ← assume (∂₁ [u] ≡ [A]) 227 | return (reflStr X [A] (unbox [A]=) [u] (unbox [u]ₛ) (unbox [u]₁)) 228 | ⟦ jj A P d a b p ⟧Tm X = do 229 | [A] ← ⟦ A ⟧Ty X 230 | [A]= ← assume (ft [A] ≡ X) 231 | [P] ← ⟦ P ⟧Ty (T-ftP X [A] (unbox [A]=)) 232 | [P]= ← assume (ft [P] ≡ _) 233 | [d] ← ⟦ d ⟧Tm [A] 234 | [d]ₛ ← assume (is-section [d]) 235 | [d]₁ ← assume (∂₁ [d] ≡ _) 236 | [a] ← ⟦ a ⟧Tm X 237 | [a]ₛ ← assume (is-section [a]) 238 | [a]₁ ← assume (∂₁ [a] ≡ [A]) 239 | [b] ← ⟦ b ⟧Tm X 240 | [b]ₛ ← assume (is-section [b]) 241 | [b]₁ ← assume (∂₁ [b] ≡ [A]) 242 | [p] ← ⟦ p ⟧Tm X 243 | [p]ₛ ← assume (is-section [p]) 244 | [p]₁ ← assume (∂₁ [p] ≡ IdStr X [A] (unbox [A]=) [a] (unbox [a]ₛ) (unbox [a]₁) [b] (unbox [b]ₛ) (unbox [b]₁)) 245 | return (jjStr X [A] (unbox [A]=) [P] (unbox [P]=) [d] (unbox [d]ₛ) (unbox [d]₁) [a] (unbox [a]ₛ) (unbox [a]₁) [b] (unbox [b]ₛ) (unbox [b]₁) [p] (unbox [p]ₛ) (unbox [p]₁)) 246 | 247 | {- Partial interpretation of contexts and context morphisms -} 248 | 249 | ⟦_⟧Ctx : (Γ : Ctx n) → Partial (Ob n) 250 | ⟦ ◇ ⟧Ctx = return pt 251 | ⟦ Γ , A ⟧Ctx = do 252 | [Γ] ← ⟦ Γ ⟧Ctx 253 | [A] ← ⟦ A ⟧Ty [Γ] 254 | return [A] 255 | 256 | ⟦_⟧Mor : (δ : Mor n m) (X : Ob n) (Y : Ob m) → Partial (MorC n m) 257 | ⟦ ◇ ⟧Mor X Y = return (ptmor X) 258 | ⟦ δ , u ⟧Mor X Y = do 259 | [δ] ← ⟦ δ ⟧Mor X (ft Y) 260 | [u] ← ⟦ u ⟧Tm X 261 | [δ]₁ ← assume (∂₁ [δ] ≡ ft Y) 262 | [u]₁ ← assume (∂₁ [u] ≡ star [δ] Y refl (unbox [δ]₁)) 263 | return (comp (qq [δ] Y refl (unbox [δ]₁)) [u] qq₀ (unbox [u]₁)) 264 | 265 | 266 | {- Basic properties of the partial interpretation functions -} 267 | 268 | ⟦⟧Tmₛ : {X : Ob n} (u : TmExpr n) {uᵈ : isDefined (⟦ u ⟧Tm X)} → is-section (⟦ u ⟧Tm X $ uᵈ) 269 | ⟦⟧Tmₛ (var k) = varCₛ k _ 270 | ⟦⟧Tmₛ (uu i) = uuStrₛ 271 | ⟦⟧Tmₛ (sum i a b) = sumStrₛ 272 | ⟦⟧Tmₛ (inl A B a) = inlStrₛ 273 | ⟦⟧Tmₛ (inr A B b) = inrStrₛ 274 | ⟦⟧Tmₛ (match A B C da db u) = matchStrₛ 275 | ⟦⟧Tmₛ (pi i a b) = piStrₛ 276 | ⟦⟧Tmₛ (lam A B u) = lamStrₛ 277 | ⟦⟧Tmₛ (app A B f a) = appStrₛ 278 | ⟦⟧Tmₛ (sig i a b) = sigStrₛ 279 | ⟦⟧Tmₛ (pair A B u v) = pairStrₛ 280 | ⟦⟧Tmₛ (pr1 A B u) = pr1Strₛ 281 | ⟦⟧Tmₛ (pr2 A B u) = pr2Strₛ 282 | ⟦⟧Tmₛ (empty i) = emptyStrₛ 283 | ⟦⟧Tmₛ (emptyelim A u) = emptyelimStrₛ 284 | ⟦⟧Tmₛ (unit i) = unitStrₛ 285 | ⟦⟧Tmₛ tt = ttStrₛ 286 | ⟦⟧Tmₛ (unitelim A dtt u) = unitelimStrₛ 287 | ⟦⟧Tmₛ (nat i) = natStrₛ 288 | ⟦⟧Tmₛ zero = zeroStrₛ 289 | ⟦⟧Tmₛ (suc u) = sucStrₛ 290 | ⟦⟧Tmₛ (natelim P d0 dS u) = natelimStrₛ 291 | ⟦⟧Tmₛ (id i a u v) = idStrₛ 292 | ⟦⟧Tmₛ (refl A u) = reflStrₛ 293 | ⟦⟧Tmₛ (jj A P d a b p) = jjStrₛ 294 | 295 | ⟦⟧Ty-ft : {X : Ob n} (A : TyExpr n) {Aᵈ : isDefined (⟦ A ⟧Ty X)} → ft (⟦ A ⟧Ty X $ Aᵈ) ≡ X 296 | ⟦⟧Ty-ft (uu i) = UUStr= 297 | ⟦⟧Ty-ft (el i v) = ElStr= 298 | ⟦⟧Ty-ft (sum A B) = SumStr= 299 | ⟦⟧Ty-ft (pi A B) = PiStr= 300 | ⟦⟧Ty-ft (sig A B) = SigStr= 301 | ⟦⟧Ty-ft empty = EmptyStr= 302 | ⟦⟧Ty-ft unit = UnitStr= 303 | ⟦⟧Ty-ft nat = NatStr= 304 | ⟦⟧Ty-ft (id A u v) = IdStr= 305 | 306 | ⟦⟧Tm₀ : {X : Ob n} (u : TmExpr n) {uᵈ : isDefined (⟦ u ⟧Tm X)} → ∂₀ (⟦ u ⟧Tm X $ uᵈ) ≡ X 307 | ⟦⟧Tm₀ (var k) = varC₀ 308 | ⟦⟧Tm₀ (uu i) = uuStr₀ 309 | ⟦⟧Tm₀ (sum i a b) = sumStr₀ 310 | ⟦⟧Tm₀ (inl A B a) = inlStr₀ 311 | ⟦⟧Tm₀ (inr A B b) = inrStr₀ 312 | ⟦⟧Tm₀ (match A B C da db u) = matchStr₀ 313 | ⟦⟧Tm₀ (pi i a b) = piStr₀ 314 | ⟦⟧Tm₀ (lam A B u) = lamStr₀ 315 | ⟦⟧Tm₀ (app A B f a) = appStr₀ 316 | ⟦⟧Tm₀ (sig i a b) = sigStr₀ 317 | ⟦⟧Tm₀ (pair A B u v) = pairStr₀ 318 | ⟦⟧Tm₀ (pr1 A B u) = pr1Str₀ 319 | ⟦⟧Tm₀ (pr2 A B u) = pr2Str₀ 320 | ⟦⟧Tm₀ (empty i) = emptyStr₀ 321 | ⟦⟧Tm₀ (emptyelim A u) = emptyelimStr₀ 322 | ⟦⟧Tm₀ (unit i) = unitStr₀ 323 | ⟦⟧Tm₀ tt = ttStr₀ 324 | ⟦⟧Tm₀ (unitelim A dtt u) = unitelimStr₀ 325 | ⟦⟧Tm₀ (nat i) = natStr₀ 326 | ⟦⟧Tm₀ zero = zeroStr₀ 327 | ⟦⟧Tm₀ (suc u) = sucStr₀ 328 | ⟦⟧Tm₀ (natelim P d0 dS u) = natelimStr₀ 329 | ⟦⟧Tm₀ (id i a u v) = idStr₀ 330 | ⟦⟧Tm₀ (refl A u) = reflStr₀ 331 | ⟦⟧Tm₀ (jj A P d a b p) = jjStr₀ 332 | 333 | ⟦⟧Tm₁-ft : {X : Ob n} (u : TmExpr n) {uᵈ : isDefined (⟦ u ⟧Tm X)} {Z : Ob (suc n)} (u₁ : ∂₁ (⟦ u ⟧Tm X $ uᵈ) ≡ Z) → ft Z ≡ X 334 | ⟦⟧Tm₁-ft u u₁ = ap ft (! u₁) ∙ ! (is-section₀ (⟦⟧Tmₛ u) refl) ∙ ⟦⟧Tm₀ u 335 | 336 | ⟦⟧Mor₀ : {X : Ob n} {Y : Ob m} (δ : Mor n m) {δᵈ : isDefined (⟦ δ ⟧Mor X Y)} → ∂₀ (⟦ δ ⟧Mor X Y $ δᵈ) ≡ X 337 | ⟦⟧Mor₀ ◇ = ptmor₀ 338 | ⟦⟧Mor₀ (δ , u) = comp₀ ∙ ⟦⟧Tm₀ u 339 | 340 | ⟦⟧Mor₁ : {X : Ob n} {Y : Ob m} (δ : Mor n m) {δᵈ : isDefined (⟦ δ ⟧Mor X Y)} → ∂₁ (⟦ δ ⟧Mor X Y $ δᵈ) ≡ Y 341 | ⟦⟧Mor₁ ◇ = ptmor₁ ∙ ! (pt-unique _) 342 | ⟦⟧Mor₁ (δ , u) = comp₁ ∙ qq₁ 343 | 344 | 345 | 346 | {- Transport along the various partial interpretation functions -} 347 | 348 | cong⟦⟧Tyᵈ : {X Y : Ob n} {A : TyExpr n} → X ≡ Y → isDefined (⟦ A ⟧Ty Y) → isDefined (⟦ A ⟧Ty X) 349 | cong⟦⟧Tyᵈ refl Aᵈ = Aᵈ 350 | 351 | cong⟦⟧TyEq : {X Y : Ob n} {A : TyExpr n} (p : X ≡ Y) (w₁ : _) → ⟦ A ⟧Ty Y $ w₁ ≡ ⟦ A ⟧Ty X $ (cong⟦⟧Tyᵈ {A = A} p w₁) 352 | cong⟦⟧TyEq refl _ = refl 353 | 354 | congTyEq⟦⟧Tyᵈ : {X : Ob n} {A B : TyExpr n} → A ≡ B → isDefined (⟦ A ⟧Ty X) → isDefined (⟦ B ⟧Ty X) 355 | congTyEq⟦⟧Tyᵈ refl Aᵈ = Aᵈ 356 | 357 | congTyEq⟦⟧Ty= : {X : Ob n} {A B : TyExpr n} (p : A ≡ B) (w₁ : _) → ⟦ A ⟧Ty X $ w₁ ≡ ⟦ B ⟧Ty X $ (congTyEq⟦⟧Tyᵈ {A = A} p w₁) 358 | congTyEq⟦⟧Ty= refl _ = refl 359 | 360 | cong⟦⟧Mor : {X : Ob n} {Y Y' : Ob m} {δ : Mor n m} → Y ≡ Y' → isDefined (⟦ δ ⟧Mor X Y) → isDefined (⟦ δ ⟧Mor X Y') 361 | cong⟦⟧Mor refl δᵈ = δᵈ 362 | 363 | cong⟦⟧MorEq : {X : Ob n} {Y Y' : Ob m} {δ : Mor n m} {w₁ : _} (p : Y ≡ Y') → ⟦ δ ⟧Mor X Y $ w₁ ≡ ⟦ δ ⟧Mor X Y' $ cong⟦⟧Mor {δ = δ} p w₁ 364 | cong⟦⟧MorEq refl = refl 365 | -------------------------------------------------------------------------------- /quotients.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | 5 | {- PathOver -} 6 | 7 | -- Proof-irrelevant PathOver, over a proof-irrelevant equalities 8 | data PathOver {l l'} {A : Set l} (B : A → Set l') {a : A} : {a' : A} (p : a ≡ a') → B a → B a' → Prop (l ⊔ l') where 9 | reflo : {u : B a} → PathOver B refl u u 10 | 11 | uncurrify : ∀ {l} {l'} {l''} {A : Set l} {B : A → Prop l''} (C : (x : A) → B x → Set l') → ΣS A B → Set l' 12 | uncurrify C (a , b) = C a b 13 | 14 | {- Equivalence relations -} 15 | 16 | record EquivRel (A : Set) : Set₁ where 17 | field 18 | _≃_ : A → A → Prop 19 | ref : (a : A) → a ≃ a 20 | sym : {a b : A} → a ≃ b → b ≃ a 21 | tra : {a b c : A} → a ≃ b → b ≃ c → a ≃ c 22 | open EquivRel {{…}} public 23 | 24 | {- Quotients -} 25 | 26 | postulate 27 | _//_ : (A : Set) (R : EquivRel A) → Set 28 | 29 | module _ {A : Set} {R : EquivRel A} where 30 | 31 | instance 32 | _ = R 33 | 34 | postulate 35 | {- Introduction rules -} 36 | 37 | proj : A → A // R 38 | eq : {a b : A} (r : a ≃ b) → proj a ≡ proj b 39 | 40 | {- Dependent elimination rule -} 41 | 42 | //-elim' : ∀ {l} {B : A // R → Set l} (proj* : (a : A) → B (proj a)) (eq* : {a b : A} (r : a ≃ b) → PathOver B (eq r) (proj* a) (proj* b)) 43 | → (x : A // R) → B x 44 | 45 | {- Reduction rule -} 46 | 47 | //-beta : ∀ {l} {B : A // R → Set l} {proj* : (a : A) → B (proj a)} {eq* : {a b : A} (r : a ≃ b) → PathOver B (eq r) (proj* a) (proj* b)} {a : A} 48 | → //-elim' proj* eq* (proj a) ↦ proj* a 49 | {-# REWRITE //-beta #-} 50 | 51 | //-elim : ∀ {l} {B : A // R → Set l} (proj* : (a : A) → B (proj a)) (eq* : {a b : A} (r : a ≃ b) → PathOver B (eq r) (proj* a) (proj* b)) 52 | → (x : A // R) → B x 53 | //-elim proj* eq* = //-elim' proj* eq* 54 | 55 | {- Lemmas about PathOver -} 56 | 57 | abstract 58 | PathOver-refl-to : ∀ {l l'} {A : Set l} {B : A → Set l'} {a : A} {u u' : B a} 59 | → u ≡ u' 60 | → PathOver B refl u u' 61 | PathOver-refl-to refl = reflo 62 | 63 | PathOver-refl-from : ∀ {l l'} {A : Set l} {B : A → Set l'} {a : A} {u u' : B a} 64 | → PathOver B refl u u' 65 | → u ≡ u' 66 | PathOver-refl-from reflo = refl 67 | 68 | PathOver-Box : ∀ {l l'} {A : Set l} (B : A → Prop l') {a a' : A} (p : a ≡ a') (u : Box (B a)) (u' : Box (B a')) → PathOver (λ x → Box (B x)) p u u' 69 | PathOver-Box B refl u u' = reflo 70 | 71 | PathOver-Cst : ∀ {l l'} {A : Set l} {B : Set l'} {a a' : A} {p : a ≡ a'} {u v : B} 72 | → u ≡ v → PathOver (λ _ → B) p u v 73 | PathOver-Cst {p = refl} refl = reflo 74 | 75 | PathOver-Prop→ : ∀ {l l' l''} {A : Set l} {B : A → Prop l'} {C : A → Set l''} 76 | {a a' : A} {p : a ≡ a'} {u : B a → C a} {u' : B a' → C a'} 77 | → ((y : B a) (y' : B a') → PathOver C p (u y) (u' y')) 78 | → PathOver (λ x → (B x → C x)) p u u' 79 | PathOver-Prop→ {p = refl} f = PathOver-refl-to (funextP λ x → PathOver-refl-from (f x x)) 80 | 81 | PathOver-Prop→Cst : ∀ {l l' l''} {A : Set l} {B : A → Prop l'} {C : Set l''} 82 | {a a' : A} {p : a ≡ a'} {u : B a → C} {u' : B a' → C} 83 | → ((y : B a) (y' : B a') → u y ≡ u' y') 84 | → PathOver (λ x → (B x → C)) p u u' 85 | PathOver-Prop→Cst {p = refl} f = PathOver-refl-to (funextP λ x → f x x) 86 | 87 | PathOver-CstPi : ∀ {l l' l''} {A : Set l} {B : Set l'} {C : A → B → Set l''} 88 | {a a' : A} {p : a ≡ a'} {u : (b : B) → C a b} {u' : (b : B) → C a' b} 89 | → ((y : B) → PathOver (λ x → C x y) p (u y) (u' y)) 90 | → PathOver (λ x → ((y : B) → C x y)) p u u' 91 | PathOver-CstPi {p = refl} f = PathOver-refl-to (funext (λ y → PathOver-refl-from (f y))) 92 | 93 | PathOver-CstPropPi : ∀ {l l' l''} {A : Set l} {B : Prop l'} {C : A → B → Set l''} 94 | {a a' : A} {p : a ≡ a'} {u : (b : B) → C a b} {u' : (b : B) → C a' b} 95 | → ((y : B) → PathOver (λ x → C x y) p (u y) (u' y)) 96 | → PathOver (λ x → ((y : B) → C x y)) p u u' 97 | PathOver-CstPropPi {p = refl} f = PathOver-refl-to (funextP (λ y → PathOver-refl-from (f y))) 98 | 99 | 100 | PathOver-PropPi : ∀ {l l' l''} {A : Set l} {B : A → Prop l'} {C : (a : A) → B a → Set l''} 101 | {a a' : A} {p : a ≡ a'} {u : (b : B a) → C a b} {u' : (b' : B a') → C a' b'} 102 | → ((y : B a) (y' : B a') → PathOver (uncurrify C) (Σ= {b = y} {b' = y'} p) (u y) (u' y')) 103 | → PathOver (λ x → ((y : B x) → C x y)) p u u' 104 | PathOver-PropPi {p = refl} f = PathOver-refl-to (funextP (λ x → PathOver-refl-from (f x x))) 105 | 106 | PathOver-out : ∀ {l l'} {A : Set l} {B : A → Set l'} {a a' : A} {p : a ≡ a'} {u : B a} {u' : B a'} 107 | → PathOver B p u u' → PathOver (λ X → X) (ap B p) u u' 108 | PathOver-out reflo = reflo 109 | 110 | PathOver-in : ∀ {l l'} {A : Set l} {B : A → Set l'} {a a' : A} {p : a ≡ a'} {u : B a} {u' : B a'} 111 | → PathOver (λ X → X) (ap B p) u u' → PathOver B p u u' 112 | PathOver-in {p = refl} reflo = reflo 113 | 114 | {- Elimination rules that we actually use (most of the time) -} 115 | 116 | module _ {A : Set} {R : EquivRel A} where 117 | 118 | private 119 | instance 120 | _ = R 121 | 122 | -- Non-dependent elimination 123 | //-rec : ∀ {l} {B : Set l} (proj* : A → B) (eq* : {a b : A} (r : a ≃ b) → proj* a ≡ proj* b) → A // R → B 124 | //-rec proj* eq* = //-elim proj* (λ r → PathOver-Cst (eq* r)) 125 | 126 | -- Dependent elimination into a Prop 127 | //-elimP : ∀ {l} {B : A // R → Prop l} (proj* : (a : A) → B (proj a)) → (x : A // R) → B x 128 | //-elimP {B = B} proj* x = unbox (//-elim {B = λ x → Box (B x)} (λ a → box (proj* a)) (λ r → PathOver-Box (λ z → B z) (eq r) (box (proj* _)) (box (proj* _))) x) 129 | 130 | -- Dependent elimination in a dependent type of the form x.((y : B) → C x y) with B a Set 131 | //-elim-PiS : ∀ {l l'} {B : Set l} {C : A // R → B → Set l'} (proj* : (a : A) (b : B) → C (proj a) b) (eq* : {a b : A} (r : a ≃ b) (y : B) → PathOver (λ x → C x y) (eq r) (proj* a y) (proj* b y)) → (x : A // R) → (y : B) → C x y 132 | //-elim-PiS proj* eq* = //-elim proj* (λ r → PathOver-CstPi (eq* r)) 133 | 134 | -- Dependent elimination in a dependent type of the form x.(B x → C) with B a Prop 135 | //-elim-PiP : ∀ {l l'} {B : A // R → Prop l} {C : Set l'} (proj* : (a : A) (b : B (proj a)) → C) (eq* : {a a' : A} (r : a ≃ a') (y : B (proj a)) (y' : B (proj a')) → proj* a y ≡ proj* a' y') → (x : A // R) → (y : B x) → C 136 | //-elim-PiP proj* eq* = //-elim proj* (λ r → PathOver-Prop→Cst (eq* r)) 137 | 138 | -- Dependent elimination in a dependent type of the form x.(B x → C x) with B a Prop 139 | //-elim-PiP2 : ∀ {l l'} {B : A // R → Prop l} {C : A // R → Set l'} (proj* : (a : A) (b : B (proj a)) → C (proj a)) (eq* : {a a' : A} (r : a ≃ a') (y : B (proj a)) (y' : B (proj a')) → PathOver C (eq r) (proj* a y) (proj* a' y')) → (x : A // R) → (y : B x) → C x 140 | //-elim-PiP2 proj* eq* = //-elim proj* (λ r → PathOver-Prop→ (eq* r)) 141 | 142 | //-elimP-PiP : ∀ {l l'} {X : Set l} {B : A // R → X → Prop l} {x₀ x₁ : X} {p : x₀ ≡ x₁} {C : Set l'} {u : (x : A // R) → B x x₀ → C} {v : (x : A // R) → B x x₁ → C} 143 | (proj* : (a : A) (y : B (proj a) x₀) (y' : B (proj a) x₁) → u (proj a) y ≡ v (proj a) y') 144 | → (x : A // R) → PathOver (λ y → (B x y → C)) p (u x) (v x) 145 | //-elimP-PiP proj* = //-elimP (λ a → PathOver-Prop→Cst (proj* a)) 146 | 147 | -- Dependent elimination in a dependent type of the form x.((y : B x) → C x y) with B a Prop 148 | //-elim-PiP3 : ∀ {l l'} {B : A // R → Prop l} {C : (x : A // R) → B x → Set l'} (proj* : (a : A) (b : B (proj a)) → C (proj a) b) (eq* : {a a' : A} (r : a ≃ a') (y : B (proj a)) (y' : B (proj a')) → PathOver (uncurrify C) (Σ= (eq r)) (proj* a y) (proj* a' y')) → (x : A // R) → (y : B x) → C x y 149 | //-elim-PiP3 proj* eq* = //-elim proj* (λ r → PathOver-PropPi (eq* r)) 150 | 151 | {- Effectiveness of quotients, this uses propositional extensionality -} 152 | 153 | module _ {A : Set} {R : EquivRel A} where 154 | 155 | private 156 | instance 157 | _ = R 158 | 159 | -- The "Codes" fibration 160 | _≃'_ : (a : A) (c : A // R) → Prop 161 | _≃'_ a = //-rec (λ b → a ≃ b) (λ r → prop-ext (λ z → tra {A = A} z r) (λ z → tra {A = A} z (sym {A = A} r))) 162 | 163 | abstract 164 | reflect' : {a : A} (c : A // R) → proj a ≡ c → a ≃' c 165 | reflect' {a} c refl = ref a 166 | 167 | reflect : {a b : A} → proj {R = R} a ≡ proj b → a ≃ b 168 | reflect {b = b} p = reflect' (proj b) p 169 | 170 | 171 | -------------------------------------------------------------------------------- /reflection.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | -- Add -v tc.unquote:10 to see the code generated by reflection 3 | 4 | open import Agda.Builtin.Reflection public 5 | open import Agda.Builtin.String public renaming (primStringAppend to _++ₛ_) 6 | 7 | open import common 8 | open import typetheory 9 | 10 | instance 11 | TCMonad : Monad {ℓ = lzero} TC 12 | return {{TCMonad}} = returnTC 13 | _>>=_ {{TCMonad}} = bindTC 14 | 15 | 16 | {- Operations on lists -} 17 | 18 | map : {A B : Set} → (f : A → B) → List A → List B 19 | map f [] = [] 20 | map f (a ∷ as) = (f a ∷ map f as) 21 | 22 | _++_ : {A : Set} → List A → List A → List A 23 | [] ++ l = l 24 | (a ∷ as) ++ l = a ∷ (as ++ l) 25 | 26 | concat : {A : Set} → List (List A) → List A 27 | concat [] = [] 28 | concat (l ∷ ls) = l ++ concat ls 29 | 30 | mapTC : {A B : Set} → (f : A → TC B) → List A → TC (List B) 31 | mapTC f [] = return [] 32 | mapTC f (a ∷ as) = do 33 | fa ← f a 34 | fas ← mapTC f as 35 | return (fa ∷ fas) 36 | 37 | reverse : {A : Set} → List A → List A 38 | reverse [] = [] 39 | reverse (x ∷ l) = reverse l ++ (x ∷ []) 40 | 41 | 42 | {- Operations on arginfo -} 43 | 44 | earg : {A : Set} → A → Arg A 45 | earg x = arg (arg-info visible relevant) x 46 | 47 | iarg : {A : Set} → A → Arg A 48 | iarg x = arg (arg-info hidden relevant) x 49 | 50 | 51 | {- Shifting de Bruijn indices -} 52 | 53 | shift : ℕ → Term → Term 54 | shiftArg : ℕ → Arg Term → Arg Term 55 | shiftListArg : ℕ → List (Arg Term) → List (Arg Term) 56 | shiftAbs : ℕ → Abs Term → Abs Term 57 | shiftSort : ℕ → Sort → Sort 58 | shiftClause : ℕ → Clause → Clause 59 | shiftListClause : ℕ → List Clause → List Clause 60 | 61 | shift n (var x args) = var (n + x) (shiftListArg n args) 62 | shift n (con c args) = con c (shiftListArg n args) 63 | shift n (def f args) = def f (shiftListArg n args) 64 | shift n (lam v t) = lam v (shiftAbs n t) 65 | shift n (pat-lam cs args) = pat-lam (shiftListClause n cs) (shiftListArg n args) 66 | shift n (pi a b) = pi (shiftArg n a) (shiftAbs n b) 67 | shift n (agda-sort s) = agda-sort (shiftSort n s) 68 | shift n (lit l) = lit l 69 | shift n (meta x args) = meta x (shiftListArg n args) 70 | shift n unknown = unknown 71 | 72 | shiftArg n (arg i x) = arg i (shift n x) 73 | 74 | shiftListArg n [] = [] 75 | shiftListArg n (a ∷ as) = shiftArg n a ∷ shiftListArg n as 76 | 77 | shiftAbs n (abs s x) = abs s (shift n x) 78 | 79 | shiftSort n (set u) = set (shift n u) 80 | shiftSort n (lit k) = lit k 81 | shiftSort n unknown = unknown 82 | 83 | shiftClause n (clause ps t) = clause ps (shift n t) 84 | shiftClause n (absurd-clause ps) = absurd-clause ps 85 | 86 | shiftListClause n [] = [] 87 | shiftListClause n (c ∷ cs) = shiftClause n c ∷ shiftListClause n cs 88 | 89 | 90 | if_then_else_ : {A : Set} → Bool → A → A → A 91 | if true then t else f = t 92 | if false then t else f = f 93 | 94 | lookup : List (Name × Name) → Name → Name 95 | lookup [] _ = quote lookup 96 | lookup ((a , b) ∷ l) a' = if primQNameEquality a a' then b else lookup l a' 97 | 98 | getNumberOfPi : Type → ℕ 99 | getNumberOfPi (pi _ (abs _ B)) = suc (getNumberOfPi B) 100 | getNumberOfPi _ = 0 101 | 102 | abstract 103 | record ERROR : Set where 104 | 105 | {- Meta-macros -} 106 | 107 | getConsTyTm : TC (List Name × List Name) 108 | getConsTyTm = do 109 | data-type _ consTy ← getDefinition (quote TyExpr) 110 | where _ → typeError (strErr "TyExpr is not a datatype." ∷ []) 111 | data-type _ consTm ← getDefinition (quote TmExpr) 112 | where _ → typeError (strErr "TmExpr is not a datatype." ∷ []) 113 | return (consTy , consTm) 114 | 115 | applyToFresh : (Name → Name → TC ⊤) → String → Name → TC (Name × Name) 116 | applyToFresh f hint s = do 117 | sNew ← freshName (hint ++ₛ primShowQName s) 118 | f s sNew 119 | return (s , sNew) 120 | 121 | pair× : {A B : Set} → A → B → A × B 122 | pair× = _,_ 123 | 124 | listify : List (Name × Name) → Term 125 | listify [] = con (quote []) [] 126 | listify ((s , t) ∷ l) = con (quote _∷_) (earg (def (quote pair×) (earg (lit (name s)) ∷ earg (lit (name t)) ∷ [])) ∷ earg (listify l) ∷ []) 127 | 128 | 129 | iterateExpr : Name → (Name → Name → TC ⊤) → TC ⊤ 130 | iterateExpr s f = do 131 | (consTy , consTm) ← getConsTyTm 132 | list ← mapTC (applyToFresh f "ap-") (consTy ++ consTm) 133 | defineFun s (clause [] (listify list) ∷ []) 134 | 135 | Ty?Tm : Name → Name → Name → Name 136 | Ty?Tm (quote TyExpr) TyFun TmFun = TyFun 137 | Ty?Tm (quote TmExpr) TyFun TmFun = TmFun 138 | Ty?Tm _ _ _ = quote ERROR 139 | 140 | generateClausewise : Name → Name → List (Arg Pattern) → List (Arg Pattern) → (ℕ → Term) → (ℕ → Name → Term → Term) → TC ⊤ 141 | generateClausewise funTy funTm preArgs postArgs varCase TmTyCase = (do 142 | (consTy , consTm) ← getConsTyTm 143 | clausesTy ← mapTC constructClause consTy 144 | defineFun funTy clausesTy 145 | clausesTm ← mapTC constructClause consTm 146 | defineFun funTm clausesTm) where 147 | 148 | constructPattern : Type → List (Arg Pattern) 149 | constructPattern (pi (arg i _) (abs s B)) = 150 | arg i (var s) ∷ constructPattern B 151 | constructPattern A = [] 152 | 153 | constructClause : Name → TC Clause 154 | constructClause c = do 155 | pi _ (abs _ tyC) ← getType c 156 | where _ → typeError (strErr "The constructor should have [n] as a first implicit argument" ∷ []) 157 | let l = getNumberOfPi tyC 158 | let result = if primQNameEquality c (quote TmExpr.var) then 159 | varCase l 160 | else 161 | TmTyCase l c tyC 162 | return (clause (preArgs ++ (earg (con c (constructPattern tyC)) ∷ postArgs)) result) 163 | 164 | -- var+extra : Name → (ℕ → Term) 165 | -- var+extra varCase l = def varCase (earg (var 0 []) ∷ []) 166 | 167 | depth : Arg Term → ℕ 168 | depth (arg _ (con _ (n ∷ _))) = suc (depth n) 169 | depth _ = zero 170 | 171 | makeArgs : (Name → ℕ → ℕ → Arg Term) → (ℕ → Type → List (Arg Term)) 172 | makeArgs body n (pi (arg i (def T (k ∷ _))) (abs s B)) = 173 | body T n (depth k) ∷ makeArgs body (n - 1) B 174 | makeArgs body n (pi (arg (arg-info visible _) _) (abs s B)) = 175 | earg (con (quote _≡_.refl) []) ∷ makeArgs body (n - 1) B 176 | makeArgs body n (pi _ (abs s B)) = makeArgs body (n - 1) B 177 | makeArgs body n _ = [] 178 | 179 | makeArgs' : ℕ → (Name → ℕ → ℕ → Arg Term) → (ℕ → Type → List (Arg Term)) 180 | makeArgs' shift body n (pi (arg i (def T (k ∷ _))) (abs s B)) = 181 | body T n (depth k) ∷ makeArgs' shift body (n - 1) B 182 | makeArgs' shift body n (pi (arg (arg-info visible _) _) (abs s B)) = 183 | earg (var (n + shift) []) ∷ makeArgs' shift body (n - 1) B 184 | makeArgs' shift body n (pi _ (abs s B)) = makeArgs' shift body (n - 1) B 185 | makeArgs' shift body n _ = [] 186 | 187 | {- Generate the ap lemmas for every type/term constructor -} 188 | 189 | half : ℕ → ℕ 190 | half zero = zero 191 | half (suc zero) = zero 192 | half (suc (suc n)) = suc (half n) 193 | 194 | -- Takes a name [s] and a type [A₀ → … → Aₙ → B] and generates [{x₀ y₀ : A₀} (p₀ : x₀ ≡ y₀) … {xₙ yₙ : Aₙ} (pₙ : xₙ ≡ yₙ) → s x₀ … xₙ ≡ s y₀ … yₙ] 195 | generate-type : ℕ → Name → Type → Type 196 | generate-type n s (pi (arg ai A) (abs x B)) = 197 | pi (iarg (shift n A)) (abs x 198 | (pi (iarg (shift (suc n) A)) (abs (x ++ₛ "'") 199 | (pi (earg (def (quote _≡_) (earg (var 1 []) ∷ earg (var 0 []) ∷ []))) (abs (x ++ₛ "⁼") 200 | (generate-type (2 + n) s B)))))) 201 | generate-type n s _ = 202 | def (quote _≡_) (earg (con s (iarg (var (half n * 3) []) ∷ make-args 2 n)) 203 | ∷ earg (con s (iarg (var (half n * 3) []) ∷ make-args 1 n)) ∷ []) where 204 | 205 | make-args : ℕ → ℕ → List (Arg Term) 206 | make-args k zero = [] 207 | make-args k (suc zero) = [] 208 | make-args k (suc (suc n)) = earg (var (half n * 3 + k) []) ∷ make-args k n 209 | 210 | generate-pattern : Type → List (Arg Pattern) 211 | generate-pattern (pi _ (abs _ B)) = earg (con (quote _≡_.refl) []) ∷ generate-pattern B 212 | generate-pattern _ = [] 213 | 214 | generate-ap : Name → Name → TC ⊤ 215 | generate-ap s res = do 216 | pi A (abs y ts) ← getType s 217 | where _ → typeError (strErr "not a Pi" ∷ []) 218 | _ ← declareDef (earg res) (pi A (abs y (generate-type 0 s ts))) 219 | _ ← defineFun res (clause (generate-pattern ts) (con (quote _≡_.refl) []) ∷ []) 220 | return _ 221 | 222 | corresponding-ap : List (Name × Name) 223 | unquoteDef corresponding-ap = iterateExpr corresponding-ap generate-ap 224 | 225 | apify : (Name → ℕ → ℕ → Arg Term) → (ℕ → Name → Term → Term) 226 | apify body l c tyC = def (lookup corresponding-ap c) (makeArgs body (l - 1) tyC) 227 | 228 | 229 | unquoteDecl ap-uu-Tm = generate-ap (quote TmExpr.uu) ap-uu-Tm 230 | unquoteDecl ap-var-Tm = generate-ap (quote TmExpr.var) ap-var-Tm 231 | unquoteDecl ap-sum-Ty = generate-ap (quote TyExpr.sum) ap-sum-Ty 232 | unquoteDecl ap-pi-Ty = generate-ap (quote TyExpr.pi) ap-pi-Ty 233 | unquoteDecl ap-sig-Ty = generate-ap (quote TyExpr.sig) ap-sig-Ty 234 | unquoteDecl ap-sum-Tm = generate-ap (quote TmExpr.sum) ap-sum-Tm 235 | unquoteDecl ap-inl-Tm = generate-ap (quote TmExpr.inl) ap-inl-Tm 236 | unquoteDecl ap-inr-Tm = generate-ap (quote TmExpr.inr) ap-inr-Tm 237 | unquoteDecl ap-match-Tm = generate-ap (quote TmExpr.match) ap-match-Tm 238 | unquoteDecl ap-pi-Tm = generate-ap (quote TmExpr.pi) ap-pi-Tm 239 | unquoteDecl ap-lam-Tm = generate-ap (quote TmExpr.lam) ap-lam-Tm 240 | unquoteDecl ap-app-Tm = generate-ap (quote TmExpr.app) ap-app-Tm 241 | unquoteDecl ap-sig-Tm = generate-ap (quote TmExpr.sig) ap-sig-Tm 242 | unquoteDecl ap-pair-Tm = generate-ap (quote TmExpr.pair) ap-pair-Tm 243 | unquoteDecl ap-pr1-Tm = generate-ap (quote TmExpr.pr1) ap-pr1-Tm 244 | unquoteDecl ap-pr2-Tm = generate-ap (quote TmExpr.pr2) ap-pr2-Tm 245 | unquoteDecl ap-emptyelim-Tm = generate-ap (quote TmExpr.emptyelim) ap-emptyelim-Tm 246 | unquoteDecl ap-unitelim-Tm = generate-ap (quote TmExpr.unitelim) ap-unitelim-Tm 247 | unquoteDecl ap-natelim-Tm = generate-ap (quote TmExpr.natelim) ap-natelim-Tm 248 | unquoteDecl ap-jj-Tm = generate-ap (quote TmExpr.jj) ap-jj-Tm 249 | unquoteDecl ap-el-Ty = generate-ap (quote TyExpr.el) ap-el-Ty 250 | unquoteDecl ap-id-Ty = generate-ap (quote TyExpr.id) ap-id-Ty 251 | unquoteDecl ap-id-Tm = generate-ap (quote TmExpr.id) ap-id-Tm 252 | unquoteDecl ap-suc-Tm = generate-ap (quote TmExpr.suc) ap-suc-Tm 253 | unquoteDecl ap-refl-Tm = generate-ap (quote TmExpr.refl) ap-refl-Tm 254 | 255 | iterate : ℕ → (List (Arg Term) → Term) → Term → Term 256 | iterate 0 f t = t 257 | iterate (suc n) f t = f (earg (iterate n f t) ∷ []) 258 | 259 | 260 | -- Is this the `same'? 261 | -- iterate' : (n : ℕ) {P : ℕ → Set} (f : {k : ℕ} → P k → P (suc k)) → {k : ℕ} → P k → P (n + k) 262 | -- iterate' zero f x = x 263 | -- iterate' (suc n) f {k = k} x = f {k = n + k} (iterate' n (λ {k} x → f {k = k} x) {k = k} x) 264 | -------------------------------------------------------------------------------- /termmodel-common.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import syntx hiding (getTy) 6 | open import rules 7 | open import quotients 8 | 9 | -- open CCat hiding (Mor) renaming (id to idC) 10 | 11 | {- Preliminary definitions -} 12 | 13 | record DCtx (n : ℕ) : Set where 14 | constructor dctx' 15 | field 16 | {ctx} : Ctx n 17 | der : ⊢ ctx 18 | open DCtx public 19 | 20 | record DMor (n m : ℕ) : Set where 21 | constructor dmor' 22 | field 23 | lhs : DCtx n 24 | rhs : DCtx m 25 | {mor} : Mor n m 26 | morDer : ctx lhs ⊢ mor ∷> ctx rhs 27 | open DMor public 28 | 29 | 30 | --The following allows for erasing propositional memory use. At the moment this is necessary to make termmodel-id.agda type check. 31 | 32 | private 33 | postulate 34 | ‗ : ∀ {l} {P : Prop l} → P 35 | 36 | erase : ∀ {l} {P : Prop l} → P → P 37 | erase p = ‗ 38 | 39 | dctx : {ctx : Ctx n} → ⊢ ctx → DCtx n 40 | dctx dΓ = dctx' (erase dΓ) 41 | 42 | dmor : (lhs : DCtx n) (rhs : DCtx m) {mor : Mor n m} → ctx lhs ⊢ mor ∷> ctx rhs → DMor n m 43 | dmor lhs rhs morDer = dmor' (dctx (der lhs)) (dctx (der rhs)) (erase morDer) 44 | 45 | {- 46 | Defining _Ob≃_ as a datatype as follows rather than being equal to ⊢ ctx Γ == ctx Γ' 47 | allows us to have more arguments implicit. 48 | The reason is that if you want to solve the definitional equality 49 | (Γ ≃ Γ') = (Δ ≃ Δ') 50 | you get that Γ = Δ and Γ' = Δ' rather than simply ctx Γ = ctx Δ and ctx Γ' = ctx Δ'. 51 | -} 52 | 53 | data _Ob≃_ (Γ Γ' : DCtx n) : Prop where 54 | box : ⊢ ctx Γ == ctx Γ' → Γ Ob≃ Γ' 55 | 56 | unOb≃ : {Γ Γ' : DCtx n} → Γ Ob≃ Γ' → ⊢ ctx Γ == ctx Γ' 57 | unOb≃ (box x) = x 58 | 59 | 60 | data _Mor≃_ (δ δ' : DMor n m) : Prop where 61 | box : ⊢ ctx (lhs δ) == ctx (lhs δ') → ⊢ ctx (rhs δ) == ctx (rhs δ') → ctx (lhs δ) ⊢ mor δ == mor δ' ∷> ctx (rhs δ) → δ Mor≃ δ' 62 | 63 | unMor≃-lhs : {δ δ' : DMor n m} → δ Mor≃ δ' → ⊢ ctx (lhs δ) == ctx (lhs δ') 64 | unMor≃-lhs (box x _ _) = x 65 | 66 | 67 | unMor≃-rhs : {δ δ' : DMor n m} → δ Mor≃ δ' → ⊢ ctx (rhs δ) == ctx (rhs δ') 68 | unMor≃-rhs (box _ x _) = x 69 | 70 | unMor≃-mor : {δ δ' : DMor n m} → δ Mor≃ δ' → ctx (lhs δ) ⊢ mor δ == mor δ' ∷> ctx (rhs δ) 71 | unMor≃-mor (box _ _ x) = x 72 | 73 | instance 74 | ObEquiv : {n : ℕ} → EquivRel (DCtx n) 75 | EquivRel._≃_ ObEquiv Γ Γ' = Γ Ob≃ Γ' 76 | EquivRel.ref ObEquiv Γ = box (CtxRefl (der Γ)) 77 | EquivRel.sym ObEquiv (box dΓ=) = box (CtxSymm dΓ=) 78 | EquivRel.tra ObEquiv (box dΓ=) (box dΔ=) = box (CtxTran dΓ= dΔ=) 79 | 80 | MorEquiv : {n m : ℕ} → EquivRel (DMor n m) 81 | EquivRel._≃_ MorEquiv δ δ' = δ Mor≃ δ' 82 | EquivRel.ref MorEquiv δ = box (CtxRefl (der (lhs δ))) (CtxRefl (der (rhs δ))) (MorRefl (morDer δ)) 83 | EquivRel.sym MorEquiv (box Γ= Δ= δ=) = box (CtxSymm Γ=) (CtxSymm Δ=) (ConvMorEq (MorSymm (CtxEqCtx1 Γ=) (CtxEqCtx1 Δ=) δ=) Γ= Δ=) 84 | EquivRel.tra MorEquiv (box Γ= Δ= δ=) (box Γ'= Δ'= δ'=) = box (CtxTran Γ= Γ'=) (CtxTran Δ= Δ'=) (MorTran (CtxEqCtx1 Γ=) (CtxEqCtx1 Δ=) δ= (ConvMorEq δ'= (CtxSymm Γ=) (CtxSymm Δ=))) 85 | 86 | 87 | DCtx= : {Γ Γ' : DCtx n} → ctx Γ ≡ ctx Γ' → proj {R = ObEquiv} Γ ≡ proj Γ' 88 | DCtx= {Γ = dctx' dΓ} {Γ' = dctx' dΓ'} refl = refl 89 | 90 | DMor= : {δ δ' : DMor m n} → ctx (lhs δ) ≡ ctx (lhs δ') → ctx (rhs δ) ≡ ctx (rhs δ') → mor δ ≡ mor δ' → proj {R = MorEquiv} δ ≡ proj δ' 91 | DMor= {δ = dmor' (dctx' dΓ) (dctx' dΔ) _} {δ' = dmor' (dctx' dΓ') (dctx' dΔ') _} refl refl refl = refl 92 | 93 | reflectOb : {Γ Γ' : DCtx n} → proj {R = ObEquiv} Γ ≡ proj Γ' → ⊢ ctx Γ == ctx Γ' 94 | reflectOb p = unOb≃ (reflect p) 95 | 96 | idMor+ : {Γ : Ctx n} {A : TyExpr n} {a : TmExpr n} → ⊢ Γ → Derivable (Γ ⊢ a :> A) → Γ ⊢ (idMor n , a) ∷> (Γ , A) 97 | idMor+ dΓ da = (idMorDerivable dΓ , congTm (! ([idMor]Ty _)) refl da) 98 | 99 | idMor+= : {Γ : Ctx n} {A : TyExpr n} {a a' : TmExpr n} → ⊢ Γ → Derivable (Γ ⊢ a == a' :> A) → Γ ⊢ (idMor n , a) == (idMor n , a') ∷> (Γ , A) 100 | idMor+= dΓ da= = (MorRefl (idMorDerivable dΓ) , congTmEqTy (! ([idMor]Ty _)) da=) 101 | 102 | {- helper function to extract data from DCtx/DMor -} 103 | getCtx : (Γ : Ctx (suc n)) → Ctx n 104 | getCtx ((Γ , _)) = Γ 105 | 106 | getdCtx : (Γ : DCtx (suc n)) → ⊢ getCtx (ctx Γ) 107 | getdCtx (dctx' {ctx = (_ , _)} (dΓ , _)) = dΓ 108 | 109 | getTy' : Ctx (suc n) → TyExpr n 110 | getTy' (Δ , B) = B 111 | 112 | getTy : (X : DCtx (suc n)) → TyExpr n 113 | getTy Δ = getTy' (ctx Δ) 114 | 115 | getdTy : (Γ : DCtx (suc n)) → Derivable (getCtx (ctx Γ) ⊢ getTy Γ) 116 | getdTy (dctx' {ctx = (_ , _)}(_ , dA)) = dA 117 | 118 | getTm : (u : DMor m (suc n)) → TmExpr m 119 | getTm u = getRHS (mor u) 120 | 121 | getMor : (a : DMor m (suc n)) → Mor m n 122 | getMor a = getLHS (mor a) 123 | 124 | getdTm : (a : DMor m (suc n)) → Derivable (ctx (lhs a) ⊢ getTm a :> getTy (rhs a) [ getMor a ]Ty) 125 | getdTm (dmor' _ (dctx' {ctx = (_ , _)} _) {mor = (_ , _)} (_ , da)) = da 126 | 127 | getdMor : (a : DMor m (suc n)) → ctx (lhs a) ⊢ getMor a ∷> (getCtx (ctx (rhs a))) 128 | getdMor (dmor' _ (dctx' {ctx = (_ , _)} _) {mor = (_ , _)} (dδ , _)) = dδ 129 | 130 | 131 | 132 | CtxTy=Ctx : {Γ : DCtx n} (A : DCtx (suc n)) (A= : proj {R = ObEquiv} (dctx (getdCtx A)) ≡ proj Γ) → ⊢ ctx Γ , getTy A == ctx A 133 | CtxTy=Ctx {Γ = Γ} A@(dctx' {ctx = (_ , _)} (_ , _)) A= = CtxSymm (reflectOb A=) , TyRefl (ConvTy (getdTy A) (reflectOb A=)) 134 | 135 | CtxTy=Ctx'' : {Γ : DCtx n} (A : DCtx (suc n)) (A= : (dctx (getdCtx A)) ≃ Γ) → ⊢ ctx Γ , getTy A == ctx A 136 | CtxTy=Ctx'' {Γ = Γ} A@(dctx' {ctx = (_ , _)} (_ , _)) A= = CtxSymm (unOb≃ A=) , TyRefl (ConvTy (getdTy A) (unOb≃ A=)) 137 | 138 | CtxTy=Ctx' : (Γ : DCtx (suc n)) → ⊢ (getCtx (ctx Γ) , getTy Γ) == ctx Γ 139 | CtxTy=Ctx' (dctx' {ctx = (_ , _)} dΓ) = CtxRefl dΓ 140 | 141 | Mor=LHSRHS : (δ : DMor m (suc n)) → ctx (lhs δ) ⊢ mor δ == getLHS (mor δ) , getRHS (mor δ) ∷> ctx (rhs δ) 142 | Mor=LHSRHS (dmor' _ (dctx' {ctx = (_ , _)} _) {mor = (_ , _)} (dδ , du)) = MorRefl (dδ , du) 143 | 144 | getCtx= : {Γ Γ' : Ctx (suc n)} (rΓ : ⊢ Γ == Γ') → ⊢ getCtx Γ == getCtx Γ' 145 | getCtx= {Γ = (Γ , A)} {(Γ' , A')} (dΓ= , _) = dΓ= 146 | 147 | getTy= : {Γ Γ' : DCtx (suc n)} (rΓ : Γ ≃ Γ') → Derivable (getCtx (ctx Γ) ⊢ getTy Γ == getTy Γ') 148 | getTy= {Γ = dctx' {ctx = (_ , _)} (dΓ , A)} {dctx' {ctx = (_ , _)} (dΓ' , dA')} (box (dΓ= , dA=)) = ConvTyEq dA= (CtxRefl dΓ) 149 | 150 | dLHS : {Γ : Ctx m} {Δ : DCtx (suc n)} {δ : Mor m (suc n)} → Γ ⊢ δ ∷> ctx Δ → Γ ⊢ getLHS δ ∷> getCtx (ctx Δ) 151 | dLHS {Δ = dctx' {ctx = (_ , _)} (dΔ , dB)} {δ = δ , u} (dδ , du) = dδ 152 | 153 | getLHS= : {Γ : Ctx m} {Δ : DCtx (suc n)} {δ δ' : Mor m (suc n)} → Γ ⊢ δ == δ' ∷> ctx Δ → Γ ⊢ getLHS δ == getLHS δ' ∷> getCtx (ctx Δ) 154 | getLHS= {Δ = dctx' {ctx = (_ , _)} (dΔ , dB)} {δ = (δ , u)} {δ' = (δ' , u')} (dδ= , du=) = dδ= 155 | 156 | getRHS= : {Γ : Ctx m} {Δ : Ctx (suc n)} {δ δ' : Mor m (suc n)} → Γ ⊢ δ == δ' ∷> Δ → Derivable (Γ ⊢ getRHS δ == getRHS δ' :> (getTy' Δ [ getLHS δ ]Ty)) 157 | getRHS= {Δ = (Δ , B)} {δ = (δ , u)} {δ' = (δ' , u')} (dδ= , du=) = du= 158 | 159 | ConvTyDCtxEq : {Γ Δ : Ctx n} {B B' : TyExpr n} → ⊢ Γ == Δ → Derivable (Γ ⊢ B) → B ≡ B' → ⊢ (Γ , B) == (Δ , B') 160 | ConvTyDCtxEq dΓ= dB refl = dΓ= , TyRefl dB 161 | -------------------------------------------------------------------------------- /termmodel-empty.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import reflection hiding (proj) 6 | open import syntx hiding (getTy) 7 | open import rules 8 | open import contextualcat 9 | open import quotients 10 | open import termmodel-common 11 | open import termmodel-synccat 12 | open import termmodel-uuel 13 | 14 | open CCat hiding (Mor) renaming (id to idC) 15 | 16 | {- Empty -} 17 | 18 | EmptyStrS-// : (Γ : DCtx n) → DCtx (suc n) 19 | EmptyStrS-// Γ = dctx (der Γ , Empty) 20 | 21 | EmptyStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') → EmptyStrS-// Γ ≃ EmptyStrS-// Γ' 22 | EmptyStrS-eq rΓ = box (unOb≃ rΓ , EmptyCong) 23 | 24 | EmptyStrS : (Γ : ObS n) → ObS (suc n) 25 | EmptyStrS = //-elim-Ctx (λ Γ → proj (EmptyStrS-// Γ)) 26 | (λ rΓ → proj= (EmptyStrS-eq rΓ)) 27 | 28 | EmptyStr=S : (Γ : ObS n) → ftS (EmptyStrS Γ) ≡ Γ 29 | EmptyStr=S = //-elimP-Ctx (λ Γ → refl) 30 | 31 | 32 | EmptyStrSynCCat : CCatwithEmpty synCCat 33 | CCatwithEmpty.EmptyStr EmptyStrSynCCat = EmptyStrS 34 | CCatwithEmpty.EmptyStr= EmptyStrSynCCat = EmptyStr=S _ 35 | CCatwithEmpty.EmptyStrNat' EmptyStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → refl))) 36 | 37 | {- empty -} 38 | 39 | emptyStrS-// : (i : ℕ) (Γ : DCtx n) → DMor n (suc n) 40 | emptyStrS-// i Γ = dmorTm Γ (EmptyUU {i = i}) 41 | 42 | emptyStrS-eq : (i : ℕ) {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') → emptyStrS-// i Γ ≃ emptyStrS-// i Γ' 43 | emptyStrS-eq i rΓ = dmorTm= dmorTmₛ dmorTmₛ rΓ UUCong EmptyUUCong 44 | 45 | emptyStrS : (i : ℕ) (Γ : ObS n) → MorS n (suc n) 46 | emptyStrS i = //-elim-Ctx (λ Γ → proj (emptyStrS-// i Γ)) 47 | (λ rΓ → proj= (emptyStrS-eq i rΓ)) 48 | 49 | emptyStrₛS : (i : ℕ) (Γ : ObS n) → S.is-section (emptyStrS i Γ) 50 | emptyStrₛS i = //-elimP (λ Γ → dmorTmₛ) 51 | 52 | emptyStr₁S : (i : ℕ) (Γ : ObS n) → S.∂₁ (emptyStrS i Γ) ≡ UUStrS i Γ 53 | emptyStr₁S i = //-elimP (λ Γ → refl) 54 | 55 | emptyStrSynCCat : CCatwithempty synCCat UUStrSynCCat 56 | CCatwithempty.emptyStr emptyStrSynCCat = emptyStrS 57 | CCatwithempty.emptyStrₛ emptyStrSynCCat {Γ = Γ} = emptyStrₛS _ Γ 58 | CCatwithempty.emptyStr₁ emptyStrSynCCat {Γ = Γ} = emptyStr₁S _ Γ 59 | CCatwithempty.emptyStrNat' emptyStrSynCCat = //-elimP (λ g → JforNat (//-elimP-Ctx (λ Γ g₁ → refl))) 60 | 61 | {- emptyelim -} 62 | 63 | emptyelimStrS-// : (Γ : DCtx n) (A : DCtx (suc (suc n))) (A= : ftS (proj A) ≡ EmptyStrS (proj Γ)) (u : DMor n (suc n)) (uₛ : S.is-section (proj u)) (u₁ : S.∂₁ (proj u) ≡ EmptyStrS (proj Γ)) → DMor n (suc n) 64 | emptyelimStrS-// Γ A A= u uₛ u₁ = dmorTm Γ (Emptyelim (dTy A A=) (dTm refl u uₛ u₁)) 65 | 66 | emptyelimStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc (suc n))} (rA : A ≃ A') (A= : ftS (proj A) ≡ EmptyStrS (proj Γ)) (A'= : ftS (proj A') ≡ EmptyStrS (proj Γ')) {u u' : DMor n (suc n)} (ru : u ≃ u') (uₛ : S.is-section (proj u)) (u'ₛ : S.is-section (proj u')) (u₁ : S.∂₁ (proj u) ≡ EmptyStrS (proj Γ)) (u'₁ : S.∂₁ (proj u') ≡ EmptyStrS (proj Γ')) → emptyelimStrS-// Γ A A= u uₛ u₁ ≃ emptyelimStrS-// Γ' A' A'= u' u'ₛ u'₁ 67 | emptyelimStrS-eq {Γ = Γ} {Γ'} rΓ {A} {A'} rA A= A'= ru uₛ u'ₛ u₁ u'₁ = dmorTm= dmorTmₛ dmorTmₛ rΓ 68 | (SubstTyFullEq' (der Γ) (der Γ , Empty) 69 | (dTy= rA A=) 70 | (idMor+= (der Γ) (dTm= refl ru uₛ u'ₛ u₁ u'₁))) 71 | (EmptyelimCong (dTy= rA A=) (dTm= refl ru uₛ u'ₛ u₁ u'₁)) 72 | 73 | emptyelimStrS : (Γ : ObS n) (A : ObS (suc (suc n))) (A= : ftS A ≡ EmptyStrS Γ) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : S.∂₁ u ≡ EmptyStrS Γ) → MorS n (suc n) 74 | emptyelimStrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Tm (λ u uₛ u₁ → proj (emptyelimStrS-// Γ A A= u uₛ u₁)) 75 | (λ ru uₛ u'ₛ u₁ u'₁ → proj= (emptyelimStrS-eq (ref Γ) (ref A) A= A= ru uₛ u'ₛ u₁ u'₁))) 76 | (λ rA A= A'= → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (emptyelimStrS-eq (ref Γ) rA A= A'= (ref u) uₛ uₛ u₁ u₁')))) 77 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (emptyelimStrS-eq rΓ (ref A) A= A=' (ref u) uₛ uₛ u₁ u₁')))) 78 | 79 | emptyelimStrₛS : (Γ : ObS n) (A : ObS (suc (suc n))) (A= : ftS A ≡ EmptyStrS Γ) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : S.∂₁ u ≡ EmptyStrS Γ) → S.is-section (emptyelimStrS Γ A A= u uₛ u₁) 80 | emptyelimStrₛS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ u uₛ u₁ → dmorTmₛ))) 81 | 82 | emptyelimStr₁S : (Γ : ObS n) (A : ObS (suc (suc n))) (A= : ftS A ≡ EmptyStrS Γ) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : S.∂₁ u ≡ EmptyStrS Γ) → S.∂₁ (emptyelimStrS Γ A A= u uₛ u₁) ≡ S.star u A A= u₁ 83 | emptyelimStr₁S = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ u uₛ u₁ → eq (box (CtxSymm (reflectOb (S.is-section₀ uₛ u₁)) , SubstTyMorEq (dTy A A=) (idMor+ (der Γ) (dTm refl u uₛ u₁)) (MorSymm (der Γ) (der Γ , Empty) (morTm=idMorTm refl u uₛ u₁))))))) 84 | 85 | emptyelimStrSynCCat : CCatwithemptyelim synCCat EmptyStrSynCCat 86 | CCatwithemptyelim.emptyelimStr emptyelimStrSynCCat = emptyelimStrS 87 | CCatwithemptyelim.emptyelimStrₛ emptyelimStrSynCCat {Γ = Γ} {A = A} {u = u} = emptyelimStrₛS Γ A _ u _ _ 88 | CCatwithemptyelim.emptyelimStr₁ emptyelimStrSynCCat {Γ = Γ} {A = A} {u = u} = emptyelimStr₁S Γ A _ u _ _ 89 | CCatwithemptyelim.emptyelimStrNat' emptyelimStrSynCCat = //-elimP (λ g → JforNat (//-elimP-Ctx (λ Γ → //-elimP (λ A A= → //-elimP (λ u uₛ u₁ g₁ → up-to-rhsTyEq (ap (_[_]Ty (substTy (getTy A) (getTm u))) (idMor[]Mor (mor g)) ∙ []Ty-substTy)))))) 90 | 91 | {- ElEmpty= -} 92 | 93 | elemptyStrS : (i : ℕ) (Γ : ObS n) → ElStrS i Γ (emptyStrS i Γ) (emptyStrₛS i Γ) (emptyStr₁S i Γ) ≡ EmptyStrS Γ 94 | elemptyStrS i = //-elimP (λ Γ → eq (box (CtxRefl (der Γ) , ElEmpty=))) 95 | -------------------------------------------------------------------------------- /termmodel-id.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import syntx hiding (getTy) 6 | open import rules 7 | open import reflection hiding (proj) 8 | open import contextualcat 9 | open import quotients 10 | open import termmodel-common 11 | open import termmodel-synccat 12 | open import termmodel-uuel 13 | 14 | 15 | {- Id -} 16 | 17 | IdStrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : ∂₁S (proj a) ≡ proj A) (b : DMor n (suc n)) (bₛ : S.is-section (proj b)) (b₁ : ∂₁S (proj b) ≡ proj A) → DCtx (suc n) 18 | IdStrS-// Γ A A= a aₛ a₁ b bₛ b₁ = dctx (der Γ , Id (dTy A A=) (dTm A= a aₛ a₁) (dTm A= b bₛ b₁)) 19 | 20 | 21 | IdStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : _) (A'= : _) {a a' : DMor n (suc n)} (ra : a ≃ a') (aₛ : _) (a'ₛ : _) (a₁ : _) (a'₁ : _) {b b' : DMor n (suc n)} (rb : b ≃ b') (bₛ : _) (b'ₛ : _) (b₁ : _) (b'₁ : _) → IdStrS-// Γ A A= a aₛ a₁ b bₛ b₁ ≃ IdStrS-// Γ' A' A'= a' a'ₛ a'₁ b' b'ₛ b'₁ 22 | IdStrS-eq rΓ rA A= A'= ra aₛ a'ₛ a₁ a'₁ rb bₛ b'ₛ b₁ b'₁ = box (unOb≃ rΓ , IdCong (dTy= rA A=) (dTm= A= ra aₛ a'ₛ a₁ a'₁) (dTm= A= rb bₛ b'ₛ b₁ b'₁)) 23 | 24 | IdStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ A) → ObS (suc n) 25 | IdStrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Tm (λ a aₛ a₁ → //-elim-Tm (λ b bₛ b₁ → proj (IdStrS-// Γ A A= a aₛ a₁ b bₛ b₁)) 26 | (λ rb bₛ b'ₛ b₁ b'₁ → proj= (IdStrS-eq (ref Γ) (ref A) A= A= (ref a) aₛ aₛ a₁ a₁ rb bₛ b'ₛ b₁ b'₁))) 27 | (λ ra aₛ a'ₛ a₁ a'₁ → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= (IdStrS-eq (ref Γ) (ref A) A= A= ra aₛ a'ₛ a₁ a'₁ (ref b) bₛ bₛ b₁ b₁')))) 28 | (λ rA A= A'= → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= (IdStrS-eq (ref Γ) rA A= A'= (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁'))))) 29 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= (IdStrS-eq rΓ (ref A) A= A=' (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁'))))) 30 | 31 | IdStr=S : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ A) → ftS (IdStrS Γ A A= a aₛ a₁ b bₛ b₁) ≡ Γ 32 | IdStr=S = //-elimP-Ctx (λ Γ → //-elimP (λ A A= → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → refl)))) 33 | 34 | IdStrSynCCat : CCatwithId synCCat 35 | CCatwithId.IdStr IdStrSynCCat = IdStrS 36 | CCatwithId.IdStr= IdStrSynCCat = IdStr=S _ _ _ _ _ _ _ _ _ 37 | CCatwithId.IdStrNat' IdStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ g₁ → refl)))))) 38 | 39 | module idS = CCatwithId IdStrSynCCat 40 | 41 | {- id -} 42 | 43 | idStrS-// : (i : ℕ) (Γ : DCtx n) (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : ∂₁S (proj a) ≡ UUStrS i (proj Γ)) (u : DMor n (suc n)) (uₛ : S.is-section (proj u)) (u₁ : ∂₁S (proj u) ≡ ElStrS i (proj Γ) (proj a) aₛ a₁) (v : DMor n (suc n)) (vₛ : S.is-section (proj v)) (v₁ : ∂₁S (proj v) ≡ ElStrS i (proj Γ) (proj a) aₛ a₁) → DMor n (suc n) 44 | idStrS-// i Γ a aₛ a₁ u uₛ u₁ v vₛ v₁ = dmorTm Γ (IdUU (dTm refl a aₛ a₁) (dTm refl u uₛ u₁) (dTm refl v vₛ v₁)) 45 | 46 | idStrS-eq : (i : ℕ) {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {a a' : DMor n (suc n)} (ra : a ≃ a') (aₛ : S.is-section (proj a)) (a'ₛ : S.is-section (proj a')) (a₁ : ∂₁S (proj a) ≡ UUStrS i (proj Γ)) (a'₁ : ∂₁S (proj a') ≡ UUStrS i (proj Γ')) {u u' : DMor n (suc n)} (ru : u ≃ u') (uₛ : S.is-section (proj u)) (u'ₛ : S.is-section (proj u')) (u₁ : ∂₁S (proj u) ≡ ElStrS i (proj Γ) (proj a) aₛ a₁) (u'₁ : ∂₁S (proj u') ≡ ElStrS i (proj Γ') (proj a') a'ₛ a'₁) {v v' : DMor n (suc n)} (rv : v ≃ v') (vₛ : S.is-section (proj v)) (v'ₛ : S.is-section (proj v')) (v₁ : ∂₁S (proj v) ≡ ElStrS i (proj Γ) (proj a) aₛ a₁) (v'₁ : ∂₁S (proj v') ≡ ElStrS i (proj Γ') (proj a') a'ₛ a'₁) → idStrS-// i Γ a aₛ a₁ u uₛ u₁ v vₛ v₁ ≃ idStrS-// i Γ' a' a'ₛ a'₁ u' u'ₛ u'₁ v' v'ₛ v'₁ 47 | idStrS-eq i {Γ} {Γ'} rΓ {a} {a'} ra aₛ a'ₛ a₁ a'₁ {u} {u'} ru uₛ u'ₛ u₁ u'₁ {v} {v'} rv vₛ v'ₛ v₁ v'₁ = 48 | dmorTm= dmorTmₛ dmorTmₛ rΓ UUCong (IdUUCong (dTm= refl ra aₛ a'ₛ a₁ a'₁) 49 | (dTm= refl ru uₛ u'ₛ u₁ u'₁) 50 | (dTm= refl rv vₛ v'ₛ v₁ v'₁)) 51 | 52 | idStrS : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ ElStrS i Γ a aₛ a₁) (v : MorS n (suc n)) (vₛ : S.is-section v) (v₁ : ∂₁S v ≡ ElStrS i Γ a aₛ a₁) → MorS n (suc n) 53 | idStrS i = //-elim-Ctx (λ Γ → //-elim-Tm (λ a aₛ a₁ → //-elim-Tm (λ u uₛ u₁ → //-elim-Tm (λ v vₛ v₁ → proj (idStrS-// i Γ a aₛ a₁ u uₛ u₁ v vₛ v₁)) 54 | (λ rv vₛ v'ₛ v₁ v'₁ → proj= (idStrS-eq i (ref Γ) (ref a) aₛ aₛ a₁ a₁ (ref u) uₛ uₛ u₁ u₁ rv vₛ v'ₛ v₁ v'₁))) 55 | (λ ru uₛ u'ₛ u₁ u'₁ → //-elimP-Tm (λ v vₛ v₁ v₁' → proj= (idStrS-eq i (ref Γ) (ref a) aₛ aₛ a₁ a₁ ru uₛ u'ₛ u₁ u'₁ (ref v) vₛ vₛ v₁ v₁')))) 56 | (λ ra aₛ a'ₛ a₁ a'₁ → //-elimP-Tm (λ u uₛ u₁ u₁' → //-elimP-Tm (λ v vₛ v₁ v₁' → proj= (idStrS-eq i (ref Γ) ra aₛ a'ₛ a₁ a'₁ (ref u) uₛ uₛ u₁ u₁' (ref v) vₛ vₛ v₁ v₁'))))) 57 | (λ rΓ → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ u uₛ u₁ u₁' → //-elimP-Tm (λ v vₛ v₁ v₁' → proj= (idStrS-eq i rΓ (ref a) aₛ aₛ a₁ a₁' (ref u) uₛ uₛ u₁ u₁' (ref v) vₛ vₛ v₁ v₁'))))) 58 | 59 | idStrₛS : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ ElStrS i Γ a aₛ a₁) (v : MorS n (suc n)) (vₛ : S.is-section v) (v₁ : ∂₁S v ≡ ElStrS i Γ a aₛ a₁) → S.is-section (idStrS i Γ a aₛ a₁ u uₛ u₁ v vₛ v₁) 60 | idStrₛS i = //-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ u uₛ u₁ → //-elimP (λ v vₛ v₁ → dmorTmₛ)))) 61 | 62 | idStr₁S : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ ElStrS i Γ a aₛ a₁) (v : MorS n (suc n)) (vₛ : S.is-section v) (v₁ : ∂₁S v ≡ ElStrS i Γ a aₛ a₁) → ∂₁S (idStrS i Γ a aₛ a₁ u uₛ u₁ v vₛ v₁) ≡ UUStrS i Γ 63 | idStr₁S i = //-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ u uₛ u₁ → //-elimP (λ v vₛ v₁ → refl)))) 64 | 65 | idStrSynCCat : CCatwithid synCCat UUStrSynCCat ElStrSynCCat 66 | CCatwithid.idStr idStrSynCCat = idStrS 67 | CCatwithid.idStrₛ idStrSynCCat {Γ = Γ} {a = a} {u = u} {v = v} = idStrₛS _ Γ a _ _ u _ _ v _ _ 68 | CCatwithid.idStr₁ idStrSynCCat {Γ = Γ} {a = a} {u = u} {v = v} = idStr₁S _ Γ a _ _ u _ _ v _ _ 69 | CCatwithid.idStrNat' idStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ u uₛ u₁ → //-elimP (λ v vₛ v₁ g₁ → refl)))))) 70 | 71 | 72 | {- refl -} 73 | 74 | reflStrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : S.ft (proj A) ≡ proj Γ) (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : ∂₁S (proj a) ≡ (proj A)) → DMor n (suc n) 75 | reflStrS-// Γ A A= a aₛ a₁ = dmorTm Γ (Refl (dTy A A=) (dTm A= a aₛ a₁)) 76 | 77 | reflStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : S.ft (proj A) ≡ proj Γ) (A'= : S.ft (proj A') ≡ proj Γ') {a a' : DMor n (suc n)} (ra : a ≃ a') (aₛ : S.is-section (proj a)) (a'ₛ : S.is-section (proj a')) (a₁ : ∂₁S (proj a) ≡ (proj A)) (a'₁ : ∂₁S (proj a') ≡ (proj A')) → reflStrS-// Γ A A= a aₛ a₁ ≃ reflStrS-// Γ' A' A'= a' a'ₛ a'₁ 78 | reflStrS-eq rΓ {A} {A'} rA A= A'= {a} {a'} ra aₛ a'ₛ a₁ a'₁ = dmorTm= dmorTmₛ dmorTmₛ rΓ (IdCong (dTy= rA A=) (dTm= A= ra aₛ a'ₛ a₁ a'₁) (dTm= A= ra aₛ a'ₛ a₁ a'₁)) 79 | (ReflCong (dTy= rA A=) (dTm= A= ra aₛ a'ₛ a₁ a'₁)) 80 | 81 | reflStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) → MorS n (suc n) 82 | reflStrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Tm (λ a aₛ a₁ → proj (reflStrS-// Γ A A= a aₛ a₁)) 83 | (λ ra aₛ a'ₛ a₁ a'₁ → proj= (reflStrS-eq (ref Γ) (ref A) A= A= ra aₛ a'ₛ a₁ a'₁))) 84 | (λ rA A= A'= → //-elimP-Tm (λ a aₛ a₁ a₁' → proj= (reflStrS-eq (ref Γ) rA A= A'= (ref a) aₛ aₛ a₁ a₁')))) 85 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Tm (λ a aₛ a₁ a₁' → proj= (reflStrS-eq rΓ (ref A) A= A=' (ref a) aₛ aₛ a₁ a₁')))) 86 | 87 | reflStrₛS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) → S.is-section (reflStrS Γ A A= a aₛ a₁) 88 | reflStrₛS = //-elimP (λ Γ → //-elimP (λ A A= → (//-elimP (λ a aₛ a₁ → dmorTmₛ)))) 89 | 90 | reflStr₁S : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) → ∂₁S (reflStrS Γ A A= a aₛ a₁) ≡ IdStrS Γ A A= a aₛ a₁ a aₛ a₁ 91 | reflStr₁S = //-elimP (λ Γ → //-elimP (λ A A= → (//-elimP (λ a aₛ a₁ → refl)))) 92 | 93 | reflStrSynCCat : CCatwithrefl IdStrSynCCat 94 | CCatwithrefl.reflStr reflStrSynCCat = reflStrS 95 | CCatwithrefl.reflStrₛ reflStrSynCCat {Γ = Γ} {A = A} {a = a} = reflStrₛS Γ A _ a _ _ 96 | CCatwithrefl.reflStr₁ reflStrSynCCat {Γ = Γ} {A = A} {a = a} = reflStr₁S Γ A _ a _ _ 97 | CCatwithrefl.reflStrNat' reflStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ a aₛ a₁ g₁ → up-to-rhsTyEq (ap (_[_]Ty (id (getTy A) (getTm a) (getTm a))) (idMor[]Mor (mor g)))))))) 98 | 99 | module reflS = CCatwithrefl reflStrSynCCat 100 | 101 | 102 | {- JJ -} 103 | 104 | fixTyJJ : {A : TyExpr n} {P : TyExpr (suc (suc (suc n)))} → subst3Ty (weakenTy' (prev (prev (prev last))) P) (var last) (var last) (refl (weakenTy A) (var last)) ≡ (((P [ weakenMor+ (weakenMor+ (weakenMor+ (weakenMor (idMor n)))) ]Ty) [ weakenMor+ (weakenMor+ (idMor (suc n) , var last)) ]Ty) [ weakenMor+ (idMor (suc n) , var last) ]Ty) [ idMor (suc n) , refl (A [ weakenMor (idMor n) ]Ty) (var last) ]Ty 105 | fixTyJJ = ap-[]Ty weakenTy+++-to-[]Ty (Mor+= (Mor+= (Mor+= (Mor+= (! (weakenMorInsert _ _ _ ∙ weakenMorInsert _ _ _ ∙ weakenMorInsert _ _ _ ∙ idMor[]Mor _ ∙ weakenMorInsert _ _ _ ∙ [idMor]Mor _)) refl) refl) refl) (ap-refl-Tm weakenTy-to-[]Ty refl)) ∙ ! ([]Ty-assoc _ _ _ ∙ []Ty-assoc _ _ _ ) 106 | 107 | 108 | 109 | abstract 110 | dP : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (P : DCtx (suc (suc (suc (suc n))))) (P= : ftS (proj P) ≡ idS.T-ftP (proj Γ) (proj A) A=) 111 | → Derivable ((((ctx Γ , getTy A) , weakenTy (getTy A)) , id (weakenTy (weakenTy (getTy A))) (var (prev last)) (var last)) ⊢ getTy P) 112 | dP Γ A A= P P= = ConvTy (dTy P P=) (CtxSymm (ConvTyDCtxEq (ConvTyDCtxEq (CtxTy=Ctx A A=) 113 | (WeakTy (dTy A A=)) 114 | weakenTy-to-[]Ty) 115 | (Id (WeakTy (WeakTy (dTy A A=))) 116 | (VarPrevLast (dTy A A=)) 117 | (VarLast (WeakTy (dTy A A=)))) 118 | (ap-id-Ty (weakenTy-to-[]Ty ∙ ap-[]Ty weakenTy-to-[]Ty refl) refl refl))) 119 | 120 | T-d₁S : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (P : DCtx (suc (suc (suc (suc n))))) (P= : ftS (proj P) ≡ idS.T-ftP (proj Γ) (proj A) A=) → DCtx (suc (suc n)) 121 | T-d₁S Γ A A= P P= = dctx (der A , ConvTy (congTy fixTyJJ (Subst3Ty (der Γ , dTy A A=) (WeakTy (dP Γ A A= P P=)) (VarLast (dTy A A=)) (congTmTy (weakenTy-to-[]Ty ∙ ! (weakenTyInsert' (prev last) _ (idMor _) (var last) ∙ weakenTyInsert _ _ _)) (VarLast (dTy A A=))) (congTmTy (ap-id-Ty (! (weakenTyInsert' (prev (prev last)) _ ((weakenMor (idMor _) , var last) , var last) (var last) ∙ weakenTyInsert _ _ _ ∙ [idMor]Ty _)) refl refl) (Refl (WeakTy (dTy A A=)) (VarLast (dTy A A=)))))) (CtxTy=Ctx A A=)) 122 | 123 | abstract 124 | reflectd₁ : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (P : DCtx (suc (suc (suc (suc n))))) (P= : ftS (proj P) ≡ idS.T-ftP (proj Γ) (proj A) A=) 125 | (d : DMor (suc n) (suc (suc n))) (d₁ : ∂₁S (proj d) ≡ reflS.T-d₁ (proj Γ) (proj A) A= (proj P) P=) 126 | → dctx (der (∂₁S-// d)) ≃ T-d₁S Γ A A= P P= 127 | reflectd₁ Γ A A= P P= d d₁ = reflect d₁ 128 | 129 | 130 | 131 | jjStrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : S.ft (proj A) ≡ proj Γ) (P : DCtx (suc (suc (suc (suc n))))) (P= : S.ft (proj P) ≡ idS.T-ftP (proj Γ) (proj A) A=) 132 | (d : DMor (suc n) (suc (suc n))) (dₛ : S.is-section (proj d)) 133 | (d₁ : dctx (der (∂₁S-// d)) ≃ T-d₁S Γ A A= P P=) {-⊢ ctx (∂₁S-// d) == (ctx A , ((((getTy' (ctx P) [ weakenMor+ (weakenMor+ (weakenMor+ (weakenMor (idMor n)))) ]Ty) [ weakenMor+ (weakenMor+ (idMor (suc n) , var last)) ]Ty) [ weakenMor+ (idMor (suc n) , var last) ]Ty) [ idMor (suc n) , refl (getTy A [ weakenMor (idMor n) ]Ty) (var last) ]Ty)))-} 134 | (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : S.∂₁ (proj a) ≡ (proj A)) 135 | (b : DMor n (suc n)) (bₛ : S.is-section (proj b)) (b₁ : S.∂₁ (proj b) ≡ (proj A)) 136 | (p : DMor n (suc n)) (pₛ : S.is-section (proj p)) (p₁ : S.∂₁ (proj p) ≡ idS.IdStr (proj Γ) (proj A) A= (proj a) aₛ a₁ (proj b) bₛ b₁) → DMor n (suc n) 137 | jjStrS-// Γ A A= P P= d dₛ d₁ a aₛ a₁ b bₛ b₁ p pₛ p₁ = dmorTm Γ (JJ (dTy A A=) (dP Γ A A= P P=) dd (dTm A= a aₛ a₁) (dTm A= b bₛ b₁) (dTm (IdStr=S (proj Γ) (proj A) A= (proj a) aₛ a₁ (proj b) bₛ b₁) p pₛ p₁)) 138 | where dd : Derivable ((ctx Γ , getTy A) ⊢ getTm d :> subst3Ty (weakenTy' (prev (prev (prev last))) (getTy P)) (var last) (var last) (refl (weakenTy (getTy A)) (var last))) 139 | dd = congTmTy! fixTyJJ (dTm≃ (Ctx≃ft+Ty (reflect A=)) d dₛ d₁) 140 | 141 | abstract 142 | jjStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : ftS (proj A) ≡ proj Γ) (A'= : ftS (proj A') ≡ proj Γ') 143 | {P P' : DCtx (suc (suc (suc (suc n))))} (rP : P ≃ P') (P= : ftS (proj P) ≡ idS.T-ftP (proj Γ) (proj A) A=) (P'= : ftS (proj P') ≡ idS.T-ftP (proj Γ') (proj A') A'=) 144 | {d d' : DMor (suc n) (suc (suc n))} (rd : d ≃ d') (dₛ : S.is-section (proj d)) (d'ₛ : S.is-section (proj d')) 145 | (d₁ : dctx (der (∂₁S-// d)) ≃ T-d₁S Γ A A= P P=) {-⊢ ctx (∂₁S-// d) == (ctx A , ((((getTy' (ctx P) [ weakenMor+ (weakenMor+ (weakenMor+ (weakenMor (idMor n)))) ]Ty) [ weakenMor+ (weakenMor+ (idMor (suc n) , var last)) ]Ty) [ weakenMor+ (idMor (suc n) , var last) ]Ty) [ idMor (suc n) , refl (getTy A [ weakenMor (idMor n) ]Ty) (var last) ]Ty)))-} 146 | (d'₁ : dctx (der (∂₁S-// d')) ≃ T-d₁S Γ' A' A'= P' P'=) {-⊢ ctx (∂₁S-// d') == (ctx A' , ((((getTy' (ctx P') [ weakenMor+ (weakenMor+ (weakenMor+ (weakenMor (idMor n)))) ]Ty) [ weakenMor+ (weakenMor+ (idMor (suc n) , var last)) ]Ty) [ weakenMor+ (idMor (suc n) , var last) ]Ty) [ idMor (suc n) , refl (getTy A' [ weakenMor (idMor n) ]Ty) (var last) ]Ty)))-} 147 | {a a' : DMor n (suc n)} (ra : a ≃ a') (aₛ : S.is-section (proj a)) (a'ₛ : S.is-section (proj a')) (a₁ : S.∂₁ (proj a) ≡ (proj A)) (a'₁ : ∂₁S (proj a') ≡ (proj A')) 148 | {b b' : DMor n (suc n)} (rb : b ≃ b') (bₛ : S.is-section (proj b)) (b'ₛ : S.is-section (proj b')) (b₁ : S.∂₁ (proj b) ≡ (proj A)) (b'₁ : ∂₁S (proj b') ≡ proj A') 149 | {p p' : DMor n (suc n)} (rp : p ≃ p') (pₛ : S.is-section (proj p)) (p'ₛ : S.is-section (proj p')) (p₁ : S.∂₁ (proj p) ≡ proj (IdStrS-// Γ A A= a aₛ a₁ b bₛ b₁)) (p'₁ : ∂₁S (proj p') ≡ proj (IdStrS-// Γ' A' A'= a' a'ₛ a'₁ b' b'ₛ b'₁)) 150 | → jjStrS-// Γ A A= P P= d dₛ d₁ a aₛ a₁ b bₛ b₁ p pₛ p₁ ≃ jjStrS-// Γ' A' A'= P' P'= d' d'ₛ d'₁ a' a'ₛ a'₁ b' b'ₛ b'₁ p' p'ₛ p'₁ 151 | jjStrS-eq {Γ = Γ} {Γ'} rΓ {A = A} {A'} rA A= A'= {P} {P'} rP P= P'= {d} {d'} rd dₛ d'ₛ d₁ d'₁ {a} {a'} ra aₛ a'ₛ a₁ a'₁ {b} {b'} rb bₛ b'ₛ b₁ b'₁ {p} {p'} rp pₛ p'ₛ p₁ p'₁ = 152 | dmorTm= dmorTmₛ dmorTmₛ rΓ (SubstTyFullEq' (der Γ) 153 | (((der Γ , dTy A A=) 154 | , WeakTy (dTy A A=)) 155 | , Id (WeakTy (WeakTy (dTy A A=))) (VarPrevLast (dTy A A=)) (VarLast (WeakTy (dTy A A=)))) 156 | dP= 157 | ((idMor+= (der Γ) (dTm= A= ra aₛ a'ₛ a₁ a'₁) 158 | , congTmEqTy (! (weakenSubstTy (getTy A) (getTm a))) (dTm= A= rb bₛ b'ₛ b₁ b'₁)) 159 | , congTmEqTy! (ap-id-Ty subst2Ty-weakenTy refl refl) 160 | (dTm= (IdStr=S (proj Γ) (proj A) A= (proj a) aₛ a₁ (proj b) bₛ b₁) rp pₛ p'ₛ p₁ p'₁))) 161 | (JJCong (dTy A A=) (dTy= rA A=) dP= dd= 162 | (dTm= A= ra aₛ a'ₛ a₁ a'₁) (dTm= A= rb bₛ b'ₛ b₁ b'₁) 163 | (dTm= (IdStr=S (proj Γ) (proj A) A= (proj a) aₛ a₁ (proj b) bₛ b₁) rp pₛ p'ₛ p₁ p'₁)) 164 | where dP= : Derivable ((((ctx Γ , getTy A) , weakenTy (getTy A)) , id (weakenTy (weakenTy (getTy A))) (var (prev last)) (var last)) ⊢ getTy P == getTy P') 165 | dP= = ConvTyEq (dTy= rP P=) (CtxSymm (ConvTyDCtxEq (ConvTyDCtxEq (CtxTy=Ctx A A=) 166 | (WeakTy (dTy A A=)) 167 | weakenTy-to-[]Ty) 168 | (Id (WeakTy (WeakTy (dTy A A=))) 169 | (VarPrevLast (dTy A A=)) 170 | (VarLast (WeakTy (dTy A A=)))) 171 | (ap-id-Ty (weakenTy-to-[]Ty ∙ ap-[]Ty weakenTy-to-[]Ty refl) refl refl))) 172 | dd= : Derivable ((ctx Γ , getTy A) ⊢ getTm d == getTm d' :> subst3Ty (weakenTy' (prev (prev (prev last))) (getTy P)) (var last) (var last) (refl (weakenTy (getTy A)) (var last))) 173 | dd= = congTmEqTy! fixTyJJ (dTm=≃ (Ctx≃ft+Ty (reflect A=)) rd dₛ d'ₛ d₁ d'₁) 174 | 175 | jjStrS1 : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (P : DCtx (suc (suc (suc (suc n))))) (P= : ftS (proj P) ≡ idS.T-ftP (proj Γ) (proj A) A=) 176 | (d : DMor (suc n) (suc (suc n))) (dₛ : S.is-section (proj d)) 177 | (d₁ : ∂₁S (proj d) ≡ reflS.T-d₁ (proj Γ) (proj A) A= (proj P) P=) 178 | (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : S.∂₁ (proj a) ≡ proj A) 179 | (b : DMor n (suc n)) (bₛ : S.is-section (proj b)) (b₁ : S.∂₁ (proj b) ≡ proj A) 180 | (p : MorS n (suc n)) (pₛ : S.is-section p) (p₁ : S.∂₁ p ≡ IdStrS (proj Γ) (proj A) A= (proj a) aₛ a₁ (proj b) bₛ b₁) → MorS n (suc n) 181 | jjStrS1 Γ A A= P P= d dₛ d₁ a aₛ a₁ b bₛ b₁ = 182 | //-elim-Tm (λ p pₛ p₁ → proj (jjStrS-// Γ A A= P P= d dₛ rd₁ a aₛ a₁ b bₛ b₁ p pₛ p₁)) 183 | (λ rp pₛ p'ₛ p₁ p'₁ → proj= (jjStrS-eq (ref Γ) (ref A) A= A= (ref P) P= P= (ref d) dₛ dₛ rd₁ rd₁ (ref a) aₛ aₛ a₁ a₁ (ref b) bₛ bₛ b₁ b₁ rp pₛ p'ₛ p₁ p'₁)) 184 | where 185 | rd₁ = reflectd₁ Γ A A= P P= d d₁ 186 | 187 | 188 | jjStrS2 : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (P : DCtx (suc (suc (suc (suc n))))) (P= : ftS (proj P) ≡ idS.T-ftP (proj Γ) (proj A) A=) 189 | (d : DMor (suc n) (suc (suc n))) (dₛ : S.is-section (proj d)) (d₁ : ∂₁S (proj d) ≡ reflS.T-d₁ (proj Γ) (proj A) A= (proj P) P=) 190 | (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : S.∂₁ (proj a) ≡ proj A) 191 | (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ proj A) 192 | (p : MorS n (suc n)) (pₛ : S.is-section p) (p₁ : S.∂₁ p ≡ IdStrS (proj Γ) (proj A) A= (proj a) aₛ a₁ b bₛ b₁) → MorS n (suc n) 193 | jjStrS2 Γ A A= P P= d dₛ d₁ a aₛ a₁ = //-elim-Tm (λ b bₛ b₁ → jjStrS1 Γ A A= P P= d dₛ d₁ a aₛ a₁ b bₛ b₁) 194 | (λ rb bₛ b'ₛ b₁ b'₁ → //-elimP-Tm (λ p pₛ p₁ p₁' → proj= (jjStrS-eq (ref Γ) (ref A) A= A= (ref P) P= P= (ref d) dₛ dₛ rd₁ rd₁ (ref a) aₛ aₛ a₁ a₁ rb bₛ b'ₛ b₁ b'₁ (ref p) pₛ pₛ p₁ p₁'))) 195 | where 196 | rd₁ = reflectd₁ Γ A A= P P= d d₁ 197 | 198 | 199 | jjStrS3 : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (P : DCtx (suc (suc (suc (suc n))))) (P= : ftS (proj P) ≡ idS.T-ftP (proj Γ) (proj A) A=) 200 | (d : DMor (suc n) (suc (suc n))) (dₛ : S.is-section (proj d)) (d₁ : ∂₁S (proj d) ≡ reflS.T-d₁ (proj Γ) (proj A) A= (proj P) P=) 201 | (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ proj A) 202 | (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ proj A) 203 | (p : MorS n (suc n)) (pₛ : S.is-section p) (p₁ : S.∂₁ p ≡ IdStrS (proj Γ) (proj A) A= a aₛ a₁ b bₛ b₁) → MorS n (suc n) 204 | jjStrS3 Γ A A= P P= d dₛ d₁ = //-elim-Tm (λ a aₛ a₁ → jjStrS2 Γ A A= P P= d dₛ d₁ a aₛ a₁) 205 | (λ ra aₛ a'ₛ a₁ a'₁ → //-elimP-Tm (λ b bₛ b₁ b₁' → //-elimP-Tm (λ p pₛ p₁ p₁' → proj= (jjStrS-eq (ref Γ) (ref A) A= A= (ref P) P= P= (ref d) dₛ dₛ rd₁ rd₁ ra aₛ a'ₛ a₁ a'₁ (ref b) bₛ bₛ b₁ b₁' (ref p) pₛ pₛ p₁ p₁')))) 206 | where 207 | rd₁ = reflectd₁ Γ A A= P P= d d₁ 208 | 209 | 210 | jjStrS4 : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (P : DCtx (suc (suc (suc (suc n))))) (P= : ftS (proj P) ≡ idS.T-ftP (proj Γ) (proj A) A=) 211 | (d : MorS (suc n) (suc (suc n))) (dₛ : S.is-section d) (d₁ : ∂₁S d ≡ reflS.T-d₁ (proj Γ) (proj A) A= (proj P) P=) 212 | (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ proj A) 213 | (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ proj A) 214 | (p : MorS n (suc n)) (pₛ : S.is-section p) (p₁ : S.∂₁ p ≡ IdStrS (proj Γ) (proj A) A= a aₛ a₁ b bₛ b₁) → MorS n (suc n) 215 | jjStrS4 Γ A A= P P= = //-elim-Tm (λ d dₛ d₁ → jjStrS3 Γ A A= P P= d dₛ d₁) 216 | (λ {d} {d'} rd dₛ d'ₛ d₁ d'₁ → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → //-elimP-Tm (λ p pₛ p₁ p₁' → proj= (jjStrS-eq (ref Γ) (ref A) A= A= (ref P) P= P= rd dₛ d'ₛ (reflectd₁ Γ A A= P P= d d₁) (reflectd₁ Γ A A= P P= d' d'₁) (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁' (ref p) pₛ pₛ p₁ p₁'))))) 217 | 218 | jjStrS5 : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (P : ObS (suc (suc (suc (suc n))))) (P= : ftS P ≡ idS.T-ftP (proj Γ) (proj A) A=) 219 | (d : MorS (suc n) (suc (suc n))) (dₛ : S.is-section d) (d₁ : ∂₁S d ≡ reflS.T-d₁ (proj Γ) (proj A) A= P P=) 220 | (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ proj A) 221 | (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ proj A) 222 | (p : MorS n (suc n)) (pₛ : S.is-section p) (p₁ : S.∂₁ p ≡ IdStrS (proj Γ) (proj A) A= a aₛ a₁ b bₛ b₁) → MorS n (suc n) 223 | jjStrS5 Γ A A= = //-elim-Ty (λ P P= → jjStrS4 Γ A A= P P=) 224 | (λ {P} {P'} rP P= P'= → //-elimP-Tm (λ d dₛ d₁ d₁' → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → //-elimP-Tm (λ p pₛ p₁ p₁' → proj= (jjStrS-eq (ref Γ) (ref A) A= A= rP P= P'= (ref d) dₛ dₛ (reflect d₁) (reflect d₁') (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁' (ref p) pₛ pₛ p₁ p₁')))))) 225 | 226 | jjStrS6 : (Γ : DCtx n) (A : ObS (suc n)) (A= : ftS A ≡ proj Γ) (P : ObS (suc (suc (suc (suc n))))) (P= : ftS P ≡ idS.T-ftP (proj Γ) A A=) 227 | (d : MorS (suc n) (suc (suc n))) (dₛ : S.is-section d) (d₁ : ∂₁S d ≡ reflS.T-d₁ (proj Γ) A A= P P=) 228 | (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ A) 229 | (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ A) 230 | (p : MorS n (suc n)) (pₛ : S.is-section p) (p₁ : S.∂₁ p ≡ IdStrS (proj Γ) A A= a aₛ a₁ b bₛ b₁) → MorS n (suc n) 231 | jjStrS6 Γ = //-elim-Ty (λ A A= → jjStrS5 Γ A A=) 232 | (λ {A} {A'} rA A= A'= → //-elimP-Ty (λ P P= P=' → //-elimP-Tm (λ d dₛ d₁ d₁' → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → //-elimP-Tm (λ p pₛ p₁ p₁' → proj= (jjStrS-eq (ref Γ) rA A= A'= (ref P) P= P=' (ref d) dₛ dₛ (reflect d₁) (reflect d₁') (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁' (ref p) pₛ pₛ p₁ p₁'))))))) 233 | 234 | jjStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (P : ObS (suc (suc (suc (suc n))))) (P= : ftS P ≡ idS.T-ftP Γ A A=) 235 | (d : MorS (suc n) (suc (suc n))) (dₛ : S.is-section d) (d₁ : ∂₁S d ≡ reflS.T-d₁ Γ A A= P P=) 236 | (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ A) 237 | (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ A) 238 | (p : MorS n (suc n)) (pₛ : S.is-section p) (p₁ : S.∂₁ p ≡ IdStrS Γ A A= a aₛ a₁ b bₛ b₁) → MorS n (suc n) 239 | jjStrS = //-elim-Ctx (λ Γ → jjStrS6 Γ) 240 | (λ {Γ} {Γ'} rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Ty (λ P P= P=' → //-elimP-Tm (λ d dₛ d₁ d₁' → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → //-elimP-Tm (λ p pₛ p₁ p₁' → proj= (jjStrS-eq rΓ (ref A) A= A=' (ref P) P= P=' (ref d) dₛ dₛ (reflect d₁) (reflect d₁') (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁' (ref p) pₛ pₛ p₁ p₁')))))))) 241 | -------------------------------------------------------------------------------- /termmodel-id2.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import syntx hiding (getTy) 6 | open import rules 7 | open import reflection hiding (proj) 8 | open import contextualcat 9 | open import quotients 10 | open import termmodel-common 11 | open import termmodel-synccat 12 | open import termmodel-uuel 13 | open import termmodel-id 14 | 15 | 16 | open CCat hiding (Mor) renaming (id to idC) 17 | 18 | abstract 19 | jjStrₛS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (P : ObS (suc (suc (suc (suc n))))) (P= : ftS P ≡ idS.T-ftP Γ A A=) (d : MorS (suc n) (suc (suc n))) (dₛ : S.is-section d) (d₁ : ∂₁S d ≡ reflS.T-d₁ Γ A A= P P=) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ A) (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ A) (p : MorS n (suc n)) (pₛ : S.is-section p) (p₁ : S.∂₁ p ≡ IdStrS Γ A A= a aₛ a₁ b bₛ b₁) → S.is-section (jjStrS Γ A A= P P= d dₛ d₁ a aₛ a₁ b bₛ b₁ p pₛ p₁) 20 | jjStrₛS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ P P= → //-elimP (λ d dₛ d₁ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → //-elimP (λ p pₛ p₁ → dmorTmₛ))))))) 21 | 22 | jjStr₁S : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (P : ObS (suc (suc (suc (suc n))))) (P= : ftS P ≡ idS.T-ftP Γ A A=) (d : MorS (suc n) (suc (suc n))) (dₛ : S.is-section d) (d₁ : ∂₁S d ≡ reflS.T-d₁ Γ A A= P P=) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ A) (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ A) (p : MorS n (suc n)) (pₛ : S.is-section p) (p₁ : S.∂₁ p ≡ IdStrS Γ A A= a aₛ a₁ b bₛ b₁) → ∂₁S (jjStrS Γ A A= P P= d dₛ d₁ a aₛ a₁ b bₛ b₁ p pₛ p₁) ≡ idS.T-jjStr₁ Γ A A= P P= a aₛ a₁ b bₛ b₁ p pₛ p₁ 23 | jjStr₁S = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ P P= → //-elimP (λ d dₛ d₁ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → //-elimP (λ p pₛ p₁ → 24 | eq (box (reflectOb (! (S.is-section₀ pₛ p₁)) 25 | , congTyEq (! subst3Ty=3substTy) refl 26 | (SubstTyFullEq' (der Γ) 27 | (der Γ , Id (dTy A A=) (dTm A= a aₛ a₁) (dTm A= b bₛ b₁)) 28 | (SubstTyFullEq' (der Γ , Id (dTy A A=) (dTm A= a aₛ a₁) (dTm A= b bₛ b₁)) 29 | ((der Γ , dTy A A=) , Id (WeakTy (dTy A A=)) (WeakTm (dTm A= a aₛ a₁)) (VarLast (dTy A A=))) 30 | (SubstTyMorEq' {Δ = ((ctx Γ , getTy A) , weakenTy (getTy A)) , id (weakenTy (weakenTy (getTy A))) (var (prev last)) (var last)} 31 | ((der Γ , dTy A A=) , Id (WeakTy (dTy A A=)) (WeakTm (dTm A= a aₛ a₁)) (VarLast (dTy A A=))) 32 | (((der Γ , dTy A A=) , WeakTy (dTy A A=)) , 33 | Id (WeakTy (WeakTy (dTy A A=))) (VarPrevLast (dTy A A=)) (VarLast (WeakTy (dTy A A=)))) 34 | (ConvTy (dTy P P=) 35 | (CtxSymm (ConvTyDCtxEq (ConvTyDCtxEq (CtxTy=Ctx A A=) 36 | (WeakTy (dTy A A=)) 37 | weakenTy-to-[]Ty) 38 | (Id (WeakTy (WeakTy (dTy A A=))) 39 | (VarPrevLast (dTy A A=)) 40 | (VarLast (WeakTy (dTy A A=)))) 41 | (ap-id-Ty (weakenTy-to-[]Ty ∙ ap-[]Ty weakenTy-to-[]Ty refl) refl refl)))) 42 | (congMorEqCtxEq {Δ = ((ctx Γ , getTy A) , weakenTy (getTy A)) , 43 | id (weakenTy (weakenTy (getTy A))) (var (prev last)) (var last)} 44 | (Ctx+= refl (ap-id-Ty (weakenTyInsert _ _ _ ∙ weakenTyInsert _ _ _ ∙ ! weakenTy-to-[]Ty) refl refl)) 45 | (WeakMor+Eq' (der Γ , dTy A A=) 46 | ((der Γ , dTy A A=) , WeakTy (dTy A A=)) 47 | (Id (WeakTy (WeakTy (dTy A A=))) 48 | (VarPrevLast (dTy A A=)) 49 | (VarLast (WeakTy (dTy A A=)))) 50 | (congMorEqCtxEq (Ctx+= refl (weakenSubstTy (getTy A) _)) 51 | (WeakMor+Eq' (der Γ) 52 | (der Γ , dTy A A=) 53 | (WeakTy (dTy A A=)) 54 | (MorSymm (der Γ) 55 | (der Γ , dTy A A=) 56 | (morTm=idMorTm A= a aₛ a₁))))))) 57 | (congMorEqCtxEq {Δ = (ctx Γ , getTy A) , id (weakenTy (getTy A)) (weakenTm (getTm a)) (var last)} 58 | (Ctx+= refl (ap-id-Ty (weakenSubstTy (getTy A) (getTm b)) 59 | (weakenSubstTm (getTm a) (getTm b)) 60 | refl)) 61 | (WeakMor+Eq' (der Γ) 62 | (der Γ , dTy A A=) 63 | (Id (WeakTy (dTy A A=)) 64 | (WeakTm (dTm A= a aₛ a₁)) 65 | (VarLast (dTy A A=))) 66 | (MorSymm (der Γ) 67 | (der Γ , dTy A A=) 68 | (morTm=idMorTm A= b bₛ b₁))))) 69 | (MorSymm (der Γ) 70 | (der Γ , Id (dTy A A=) (dTm A= a aₛ a₁) (dTm A= b bₛ b₁)) 71 | (morTm=idMorTm refl p pₛ p₁)))))))))))) 72 | 73 | jjStrSynCCat : CCatwithjj synCCat IdStrSynCCat reflStrSynCCat 74 | CCatwithjj.jjStr jjStrSynCCat = jjStrS 75 | CCatwithjj.jjStrₛ jjStrSynCCat {Γ = Γ} {A = A} {P = P} {d = d} {a = a} {b = b} {p = p} = jjStrₛS Γ A _ P _ d _ _ a _ _ b _ _ p _ _ 76 | CCatwithjj.jjStr₁ jjStrSynCCat {Γ = Γ} {A = A} {P = P} {d = d} {a = a} {b = b} {p = p} = jjStr₁S Γ A _ P _ d _ _ a _ _ b _ _ p _ _ 77 | CCatwithjj.jjStrNat' jjStrSynCCat = //-elimP (λ g → //-elimP (λ Δ g₀ → (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ P P= → //-elimP (λ d dₛ d₁ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → //-elimP (λ p pₛ p₁ g₁ → up-to-rhsTyEq' (reflectOb g₀) (ap-[]Ty refl (idMor[]Mor (mor g)) ∙ []Ty-subst3Ty))))))))))) 78 | 79 | {- ElId= -} 80 | 81 | elidStrS : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ ElStrS i Γ a aₛ a₁) 82 | (v : MorS n (suc n)) (vₛ : S.is-section v) (v₁ : ∂₁S v ≡ ElStrS i Γ a aₛ a₁) → ElStrS i Γ (idStrS i Γ a aₛ a₁ u uₛ u₁ v vₛ v₁) (idStrₛS i Γ a aₛ a₁ u uₛ u₁ v vₛ v₁) (idStr₁S i Γ a aₛ a₁ u uₛ u₁ v vₛ v₁) ≡ IdStrS Γ (ElStrS i Γ a aₛ a₁) (ElStr=S i Γ a aₛ a₁) u uₛ u₁ v vₛ v₁ 83 | elidStrS i = //-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ u uₛ u₁ → //-elimP (λ v vₛ v₁ → eq (box (CtxRefl (der Γ) , ElId= (dTm refl a aₛ a₁) (dTm refl u uₛ u₁) (dTm refl v vₛ v₁))))))) 84 | 85 | {- BetaJ -} 86 | 87 | betaIdStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (P : ObS (suc (suc (suc (suc n))))) (P= : ftS P ≡ idS.T-ftP Γ A A=) 88 | (d : MorS (suc n) (suc (suc n))) (dₛ : S.is-section d) (d₁ : ∂₁S d ≡ reflS.T-d₁ Γ A A= P P=) 89 | (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) 90 | → jjStrS Γ A A= P P= d dₛ d₁ a aₛ a₁ a aₛ a₁ (reflStrS Γ A A= a aₛ a₁) (reflStrₛS Γ A A= a aₛ a₁) (reflStr₁S Γ A A= a aₛ a₁) ≡ S.starTm a d (S.is-section₀ dₛ d₁ ∙ reflS.T-d₁=) a₁ 91 | betaIdStrS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ P P= → //-elimP (λ d dₛ d₁ → //-elimP (λ a aₛ a₁ → 92 | let dd = congTmTy! fixTyJJ (dTm≃ (Ctx≃ft+Ty (reflect A=)) d dₛ (reflectd₁ Γ A A= P P= d d₁)) 93 | in up-to-rhsTyEq2' (CtxSymm (reflectOb (S.is-section₀ aₛ a₁ ∙ A=))) 94 | (congTyEq (ap-[]Ty ([idMor]Ty _ ∙ ! fixTyJJ) refl 95 | ∙ []Ty-assoc _ _ _ 96 | ∙ ap-[]Ty refl (Mor+= (Mor+= (Mor+= (Mor+= (weakenMorInsert _ _ _ ∙ ([idMor]Mor _)) refl) refl) refl) 97 | (ap-refl-Tm (weakenTyInsert _ _ _ ∙ [idMor]Ty _) refl)) 98 | ∙ weakenTyInsert' _ _ _ _ ) 99 | ([]Ty-assoc _ _ _ ) 100 | (SubstTyFullEq' (der Γ) (der Γ , dTy A A=) 101 | (SubstTyFullEq' (der Γ , dTy A A=) (der Γ , dTy A A=) 102 | (dTy=≃ (sym (reflectd₁ Γ A A= P P= d d₁)) (Ctx≃ft+Ty (reflect A=))) 103 | (MorSymm (der Γ , dTy A A=) (der Γ , dTy A A=) 104 | (getMor=idMor≃ (Ctx≃ft+Ty (reflect A=)) d dₛ (reflectd₁ Γ A A= P P= d d₁)))) 105 | (MorSymm (der Γ) (der Γ , dTy A A=) 106 | (morTm=idMorTm A= a aₛ a₁)))) 107 | (TmTran' (der Γ) (BetaIdRefl (dTy A A=) (dP Γ A A= P P=) dd (dTm A= a aₛ a₁)) 108 | (congTmEqTy ([]Ty-subst3Ty 109 | ∙ ap-subst3Ty (weakenTyInsertLemma3 last _ _ _ ∙ weakenTyInsert' _ _ _ _ ∙ [idMor]Ty _) 110 | refl refl 111 | (ap-refl-Tm (weakenSubstTy _ _) refl)) 112 | (SubstTmMorEq' (der Γ) (der Γ , dTy A A=) dd (MorSymm (der Γ) (der Γ , dTy A A=) (morTm=idMorTm A= a aₛ a₁)))))))))) 113 | 114 | 115 | 116 | -------------------------------------------------------------------------------- /termmodel-nat.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import syntx hiding (getTy) 6 | open import rules 7 | open import contextualcat 8 | open import quotients 9 | open import termmodel-common 10 | open import termmodel-synccat 11 | open import termmodel-uuel 12 | 13 | open CCat hiding (Mor) renaming (id to idC) 14 | 15 | 16 | {- Nat -} 17 | 18 | NatStrS-// : DCtx n → DCtx (suc n) 19 | NatStrS-// Γ = dctx (der Γ , Nat) 20 | 21 | NatStrS-eq : {Γ Γ' : DCtx n} → Γ ≃ Γ' → NatStrS-// Γ ≃ NatStrS-// Γ' 22 | NatStrS-eq dΓ= = box (unOb≃ dΓ= , NatCong) 23 | 24 | NatStrS : ObS n → ObS (suc n) 25 | NatStrS = //-elim-Ctx (λ Γ → proj (NatStrS-// Γ)) (λ rΓ → proj= (NatStrS-eq rΓ)) 26 | 27 | NatStr=S : (Γ : ObS n) → ftS (NatStrS Γ) ≡ Γ 28 | NatStr=S = //-elimP-Ctx (λ Γ → refl) 29 | 30 | NatStrNatS : (g : MorS n m) (Γ : ObS m) (g₁ : ∂₁S g ≡ Γ) → S.star g (NatStrS Γ) (NatStr=S Γ) g₁ ≡ NatStrS (∂₀S g) 31 | NatStrNatS = //-elimP (λ g → //-elimP (λ Γ g₁ → refl)) 32 | 33 | NatStrSynCCat : CCatwithNat synCCat 34 | CCatwithNat.NatStr NatStrSynCCat = NatStrS 35 | CCatwithNat.NatStr= NatStrSynCCat = NatStr=S _ 36 | CCatwithNat.NatStrNat' NatStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → refl))) 37 | 38 | 39 | {- nat -} 40 | 41 | natStrS-// : (i : ℕ) (Γ : DCtx n) → DMor n (suc n) 42 | natStrS-// i Γ = dmorTm Γ (NatUU {i = i}) 43 | 44 | natStrS-eq : (i : ℕ) {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') → natStrS-// i Γ ≃ natStrS-// i Γ' 45 | natStrS-eq i rΓ = dmorTm= dmorTmₛ dmorTmₛ rΓ UUCong NatUUCong 46 | 47 | natStrS : (i : ℕ) (Γ : ObS n) → MorS n (suc n) 48 | natStrS i = //-elim-Ctx (λ Γ → proj (natStrS-// i Γ)) (λ rΓ → proj= (natStrS-eq i rΓ)) 49 | 50 | natStrₛS : (i : ℕ) (Γ : ObS n) → S.is-section (natStrS i Γ) 51 | natStrₛS i = //-elimP (λ Γ → dmorTmₛ) 52 | 53 | natStr₁S : (i : ℕ) (Γ : ObS n) → ∂₁S (natStrS i Γ) ≡ UUStrS i Γ 54 | natStr₁S i = //-elimP (λ Γ → refl) 55 | 56 | natStrSynCCat : CCatwithnat synCCat UUStrSynCCat 57 | CCatwithnat.natStr natStrSynCCat = natStrS 58 | CCatwithnat.natStrₛ natStrSynCCat {Γ = Γ} = natStrₛS _ Γ 59 | CCatwithnat.natStr₁ natStrSynCCat {Γ = Γ} = natStr₁S _ Γ 60 | CCatwithnat.natStrNat' natStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ g₁ → refl))) 61 | 62 | 63 | {- zero -} 64 | 65 | zeroStrS-// : (Γ : DCtx n) → DMor n (suc n) 66 | zeroStrS-// Γ = dmorTm Γ Zero 67 | 68 | zeroStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') → zeroStrS-// Γ ≃ zeroStrS-// Γ' 69 | zeroStrS-eq rΓ = dmorTm= dmorTmₛ dmorTmₛ rΓ NatCong ZeroCong 70 | 71 | zeroStrS : (Γ : ObS n) → MorS n (suc n) 72 | zeroStrS = //-elim-Ctx (λ Γ → proj (zeroStrS-// Γ)) (λ rΓ → proj= (zeroStrS-eq rΓ)) 73 | 74 | zeroStrₛS : (Γ : ObS n) → S.is-section (zeroStrS Γ) 75 | zeroStrₛS = //-elimP (λ Γ → dmorTmₛ) 76 | 77 | zeroStr₁S : (Γ : ObS n) → ∂₁S (zeroStrS Γ) ≡ NatStrS Γ 78 | zeroStr₁S = //-elimP (λ Γ → refl) 79 | 80 | zeroStrSynCCat : CCatwithzero synCCat NatStrSynCCat 81 | CCatwithzero.zeroStr zeroStrSynCCat = zeroStrS 82 | CCatwithzero.zeroStrₛ zeroStrSynCCat {Γ = Γ} = zeroStrₛS Γ 83 | CCatwithzero.zeroStr₁ zeroStrSynCCat {Γ = Γ} = zeroStr₁S Γ 84 | CCatwithzero.zeroStrNat' zeroStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ g₁ → refl))) 85 | 86 | 87 | {- suc -} 88 | 89 | sucStrS-// : (Γ : DCtx n) (u : DMor n (suc n)) (uₛ : S.is-section (proj u)) (u₁ : ∂₁S (proj u) ≡ NatStrS (proj Γ)) → DMor n (suc n) 90 | sucStrS-// Γ u uₛ u₁ = dmorTm Γ (Suc (dTm refl u uₛ u₁)) 91 | 92 | sucStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {u u' : DMor n (suc n)} (ru : u ≃ u') (uₛ : _) (u'ₛ : _) (u₁ : _) (u'₁ : _) → sucStrS-// Γ u uₛ u₁ ≃ sucStrS-// Γ' u' u'ₛ u'₁ 93 | sucStrS-eq rΓ ru uₛ u'ₛ u₁ u'₁ = dmorTm= dmorTmₛ dmorTmₛ rΓ NatCong (SucCong (dTm= refl ru uₛ u'ₛ u₁ u'₁)) 94 | 95 | sucStrS : (Γ : ObS n) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ NatStrS Γ) → MorS n (suc n) 96 | sucStrS = //-elim-Ctx (λ Γ → //-elim-Tm (λ u uₛ u₁ → proj (sucStrS-// Γ u uₛ u₁)) 97 | (λ ru uₛ u'ₛ u₁ u'₁ → proj= (sucStrS-eq (ref Γ) ru uₛ u'ₛ u₁ u'₁))) 98 | (λ rΓ → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (sucStrS-eq rΓ (ref u) uₛ uₛ u₁ u₁'))) 99 | 100 | sucStrₛS : (Γ : ObS n) (u : MorS n (suc n)) {uₛ : S.is-section u} {u₁ : ∂₁S u ≡ NatStrS Γ} → S.is-section (sucStrS Γ u uₛ u₁) 101 | sucStrₛS = //-elimP (λ Γ → //-elimP (λ u {uₛ} {u₁} → dmorTmₛ)) 102 | 103 | sucStr₁S : (Γ : ObS n) (u : MorS n (suc n)) {uₛ : S.is-section u} {u₁ : ∂₁S u ≡ NatStrS Γ} → ∂₁S (sucStrS Γ u uₛ u₁) ≡ NatStrS Γ 104 | sucStr₁S = //-elimP (λ Γ → //-elimP (λ u → refl)) 105 | 106 | sucStrSynCCat : CCatwithsuc synCCat NatStrSynCCat 107 | CCatwithsuc.sucStr sucStrSynCCat = sucStrS 108 | CCatwithsuc.sucStrₛ sucStrSynCCat {Γ = Γ} {u = u} = sucStrₛS Γ u 109 | CCatwithsuc.sucStr₁ sucStrSynCCat {Γ = Γ} {u = u} = sucStr₁S Γ u 110 | CCatwithsuc.sucStrNat' sucStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ u uₛ u₁ g₁ → refl)))) 111 | 112 | module sucS = CCatwithsuc sucStrSynCCat 113 | 114 | {- natelim -} 115 | 116 | fixSubstTy : {X : TyExpr (suc n)} 117 | → weakenTy' (prev last) (weakenTy' (prev last) X) [ idMor _ , suc (var (prev last)) ]Ty 118 | ≡ ((X [ weakenMor (weakenMor (idMor n)) , var last ]Ty) [ weakenMor (weakenMor (idMor (suc n))) , var last ]Ty) [ idMor (suc (suc n)) , suc (var (prev last)) ]Ty 119 | fixSubstTy = ap-[]Ty (weakenTy+-to-[]Ty ∙ ap-[]Ty weakenTy+-to-[]Ty refl) refl 120 | 121 | 122 | natelimStrS-// : (Γ : DCtx n) (P : DCtx (suc (suc n))) (P= : ftS (proj P) ≡ NatStrS (proj Γ)) 123 | (dO : DMor n (suc n)) (dOₛ : S.is-section (proj dO)) (dO₁ : ∂₁S (proj dO) ≡ S.star (zeroStrS (proj Γ)) (proj P) P= (zeroStr₁S (proj Γ))) 124 | (dS : DMor (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section (proj dS)) 125 | (dS₁ : ∂₁S (proj dS) ≡ sucS.T-dS₁ (proj Γ) (proj P) P=) 126 | (u : DMor n (suc n)) (uₛ : S.is-section (proj u)) (u₁ : ∂₁S (proj u) ≡ NatStrS (proj Γ)) 127 | → DMor n (suc n) 128 | natelimStrS-// Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ u uₛ u₁ = 129 | dmorTm Γ (Natelim (dTy P P=) 130 | (dTm refl dO dOₛ dO₁) 131 | ddS 132 | (dTm refl u uₛ u₁)) 133 | where ddS : Derivable (((ctx Γ , nat) , getTy P) ⊢ getTm dS :> substTy (weakenTy' (prev last) (weakenTy' (prev last) (getTy P))) (suc (var (prev last)))) 134 | ddS = congTmTy! fixSubstTy (dTm≃ (Ctx≃ft+Ty (reflect P=)) dS dSₛ (reflect dS₁)) 135 | 136 | abstract 137 | natelimStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {P P' : DCtx (suc (suc n))} (rP : P ≃ P') (P= : ftS (proj P) ≡ NatStrS (proj Γ)) (P'= : ftS (proj P') ≡ NatStrS (proj Γ')) 138 | {dO dO' : DMor n (suc n)} (rdO : dO ≃ dO') (dOₛ : S.is-section (proj dO)) (dO'ₛ : S.is-section (proj dO')) 139 | (dO₁ : ∂₁S (proj dO) ≡ S.star (zeroStrS (proj Γ)) (proj P) P= (zeroStr₁S (proj Γ))) (dO'₁ : ∂₁S (proj dO') ≡ S.star (zeroStrS (proj Γ')) (proj P') P'= (zeroStr₁S (proj Γ'))) 140 | {dS dS' : DMor (suc (suc n)) (suc (suc (suc n)))} (rdS : dS ≃ dS') (dSₛ : S.is-section (proj dS))(dS'ₛ : S.is-section (proj dS')) 141 | (dS₁ : ∂₁S (proj dS) ≡ sucS.T-dS₁ (proj Γ) (proj P) P=) 142 | (dS'₁ : ∂₁S (proj dS') ≡ sucS.T-dS₁ (proj Γ') (proj P') P'=) 143 | {u u' : DMor n (suc n)} (ru : u ≃ u') (uₛ : S.is-section (proj u)) (u'ₛ : S.is-section (proj u')) 144 | (u₁ : ∂₁S (proj u) ≡ NatStrS (proj Γ)) (u'₁ : ∂₁S (proj u') ≡ NatStrS (proj Γ')) 145 | → natelimStrS-// Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ u uₛ u₁ ≃ natelimStrS-// Γ' P' P'= dO' dO'ₛ dO'₁ dS' dS'ₛ dS'₁ u' u'ₛ u'₁ 146 | natelimStrS-eq {Γ = Γ} {Γ'} rΓ {P} {P'} rP P= P'= {dO} {dO'} rdO dOₛ dO'ₛ dO₁ dO'₁ {dS} {dS'} rdS dSₛ dS'ₛ dS₁ dS'₁ {u} {u'} ru uₛ u'ₛ u₁ u'₁ = 147 | dmorTm= dmorTmₛ dmorTmₛ rΓ (SubstTyFullEq' (der Γ) (der Γ , Nat) (dTy= rP P=) (idMor+= (der Γ) (dTm= refl ru uₛ u'ₛ u₁ u'₁))) 148 | (NatelimCong (dTy P P=) (dTy= rP P=) 149 | (dTm= refl rdO dOₛ dO'ₛ dO₁ dO'₁) 150 | ddS= 151 | (dTm= refl ru uₛ u'ₛ u₁ u'₁)) 152 | where ddS= : Derivable (((ctx Γ , nat) , getTy P) ⊢ getTm dS == getTm dS' :> substTy (weakenTy' (prev last) (weakenTy' (prev last) (getTy P))) (suc (var (prev last)))) 153 | ddS= = congTmEqTy! fixSubstTy (dTm=≃ (Ctx≃ft+Ty (reflect P=)) rdS dSₛ dS'ₛ (reflect dS₁) (reflect dS'₁)) 154 | 155 | natelimStrS1 : (Γ : DCtx n) (P : DCtx (suc (suc n))) (P= : ftS (proj P) ≡ NatStrS (proj Γ)) 156 | (dO : DMor n (suc n)) (dOₛ : S.is-section (proj dO)) (dO₁ : ∂₁S (proj dO) ≡ S.star (zeroStrS (proj Γ)) (proj P) P= (zeroStr₁S (proj Γ))) 157 | (dS : DMor (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section (proj dS)) (dS₁ : ∂₁S (proj dS) ≡ sucS.T-dS₁ (proj Γ) (proj P) P=) 158 | (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ NatStrS (proj Γ)) 159 | → MorS n (suc n) 160 | natelimStrS1 Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ = //-elim-Tm (λ u uₛ u₁ → proj (natelimStrS-// Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ u uₛ u₁)) 161 | (λ ru uₛ u'ₛ u₁ u'₁ → proj= (natelimStrS-eq (ref Γ) (ref P) P= P= (ref dO) dOₛ dOₛ dO₁ dO₁ (ref dS) dSₛ dSₛ dS₁ dS₁ ru uₛ u'ₛ u₁ u'₁)) 162 | 163 | natelimStrS2 : (Γ : DCtx n) (P : DCtx (suc (suc n))) (P= : ftS (proj P) ≡ NatStrS (proj Γ)) 164 | (dO : DMor n (suc n)) (dOₛ : S.is-section (proj dO)) (dO₁ : ∂₁S (proj dO) ≡ S.star (zeroStrS (proj Γ)) (proj P) P= (zeroStr₁S (proj Γ))) 165 | (dS : MorS (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section dS) (dS₁ : ∂₁S dS ≡ sucS.T-dS₁ (proj Γ) (proj P) P=) 166 | (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ NatStrS (proj Γ)) 167 | → MorS n (suc n) 168 | natelimStrS2 Γ P P= dO dOₛ dO₁ = //-elim-Tm (λ dS dSₛ dS₁ → natelimStrS1 Γ P P= dO dOₛ dO₁ dS dSₛ dS₁) 169 | (λ rdS dSₛ dS'ₛ dS₁ dS'₁ → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (natelimStrS-eq (ref Γ) (ref P) P= P= (ref dO) dOₛ dOₛ dO₁ dO₁ rdS dSₛ dS'ₛ dS₁ dS'₁ (ref u) uₛ uₛ u₁ u₁'))) 170 | 171 | natelimStrS3 : (Γ : DCtx n) (P : DCtx (suc (suc n))) (P= : ftS (proj P) ≡ NatStrS (proj Γ)) 172 | (dO : MorS n (suc n)) (dOₛ : S.is-section dO) (dO₁ : ∂₁S dO ≡ S.star (zeroStrS (proj Γ)) (proj P) P= (zeroStr₁S (proj Γ))) 173 | (dS : MorS (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section dS) (dS₁ : ∂₁S dS ≡ sucS.T-dS₁ (proj Γ) (proj P) P=) 174 | (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ NatStrS (proj Γ)) 175 | → MorS n (suc n) 176 | natelimStrS3 Γ P P= = //-elim-Tm (λ dO dOₛ dO₁ → natelimStrS2 Γ P P= dO dOₛ dO₁) 177 | (λ rdO dOₛ dO'ₛ dO₁ dO'₁ → //-elimP-Tm (λ dS dSₛ dS₁ dS₁' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (natelimStrS-eq (ref Γ) (ref P) P= P= rdO dOₛ dO'ₛ dO₁ dO'₁ (ref dS) dSₛ dSₛ dS₁ dS₁' (ref u) uₛ uₛ u₁ u₁')))) 178 | 179 | natelimStrS4 : (Γ : DCtx n) (P : ObS (suc (suc n))) (P= : ftS P ≡ NatStrS (proj Γ)) 180 | (dO : MorS n (suc n)) (dOₛ : S.is-section dO) (dO₁ : ∂₁S dO ≡ S.star (zeroStrS (proj Γ)) P P= (zeroStr₁S (proj Γ))) 181 | (dS : MorS (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section dS) (dS₁ : ∂₁S dS ≡ sucS.T-dS₁ (proj Γ) P P=) 182 | (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ NatStrS (proj Γ)) 183 | → MorS n (suc n) 184 | natelimStrS4 Γ = //-elim-Ty (λ P P= → natelimStrS3 Γ P P=) 185 | (λ rP P= P'= → //-elimP-Tm (λ dO dOₛ dO₁ dO₁' → //-elimP-Tm (λ dS dSₛ dS₁ dS₁' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (natelimStrS-eq (ref Γ) rP P= P'= (ref dO) dOₛ dOₛ dO₁ dO₁' (ref dS) dSₛ dSₛ dS₁ dS₁' (ref u) uₛ uₛ u₁ u₁'))))) 186 | 187 | natelimStrS : (Γ : ObS n) (P : ObS (suc (suc n))) (P= : ftS P ≡ NatStrS Γ) 188 | (dO : MorS n (suc n)) (dOₛ : S.is-section dO) (dO₁ : ∂₁S dO ≡ S.star (zeroStrS Γ) P P= (zeroStr₁S Γ)) 189 | (dS : MorS (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section dS) (dS₁ : ∂₁S dS ≡ sucS.T-dS₁ Γ P P=) 190 | (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ NatStrS Γ) 191 | → MorS n (suc n) 192 | natelimStrS = //-elim-Ctx (λ Γ → natelimStrS4 Γ) 193 | (λ rΓ → //-elimP-Ty (λ P P= P=' → //-elimP-Tm (λ dO dOₛ dO₁ dO₁' → //-elimP-Tm (λ dS dSₛ dS₁ dS₁' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (natelimStrS-eq rΓ (ref P) P= P=' (ref dO) dOₛ dOₛ dO₁ dO₁' (ref dS) dSₛ dSₛ dS₁ dS₁' (ref u) uₛ uₛ u₁ u₁')))))) 194 | 195 | abstract 196 | natelimStrₛS : (Γ : ObS n) (P : ObS (suc (suc n))) (P= : ftS P ≡ NatStrS Γ) 197 | (dO : MorS n (suc n)) (dOₛ : S.is-section dO) (dO₁ : ∂₁S dO ≡ S.star (zeroStrS Γ) P P= (zeroStr₁S Γ)) 198 | (dS : MorS (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section dS) (dS₁ : ∂₁S dS ≡ sucS.T-dS₁ Γ P P=) 199 | (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ NatStrS Γ) → S.is-section (natelimStrS Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ u uₛ u₁) 200 | natelimStrₛS = //-elimP (λ Γ → //-elimP (λ P P= → //-elimP (λ dO dOₛ dO₁ → //-elimP (λ dS dSₛ dS₁ → //-elimP (λ u uₛ u₁ → dmorTmₛ))))) 201 | 202 | natelimStr₁S : (Γ : ObS n) (P : ObS (suc (suc n))) (P= : ftS P ≡ NatStrS Γ) 203 | (dO : MorS n (suc n)) (dOₛ : S.is-section dO) (dO₁ : ∂₁S dO ≡ S.star (zeroStrS Γ) P P= (zeroStr₁S Γ)) 204 | (dS : MorS (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section dS) (dS₁ : ∂₁S dS ≡ sucS.T-dS₁ Γ P P=) 205 | (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ NatStrS Γ) → S.∂₁ (natelimStrS Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ u uₛ u₁) ≡ S.star u P P= u₁ 206 | natelimStr₁S = //-elimP (λ Γ → //-elimP (λ P P= → //-elimP (λ dO dOₛ dO₁ → //-elimP (λ dS dSₛ dS₁ → //-elimP (λ u uₛ u₁ → eq (box (reflectOb (! (S.is-section₀ uₛ u₁)) , SubstTyMorEq' (der Γ) (der Γ , Nat) (dTy P P=) (MorSymm (der Γ) (der Γ , Nat) (morTm=idMorTm (NatStr=S (proj Γ)) u uₛ u₁))))))))) 207 | 208 | natelimStrSynCCat : CCatwithnatelim synCCat NatStrSynCCat zeroStrSynCCat sucStrSynCCat 209 | CCatwithnatelim.natelimStr natelimStrSynCCat = natelimStrS 210 | CCatwithnatelim.natelimStrₛ natelimStrSynCCat {Γ = Γ} {P = P} {dO = dO} {dS = dS} {u = u} = natelimStrₛS Γ P _ dO _ _ dS _ _ u _ _ 211 | CCatwithnatelim.natelimStr₁ natelimStrSynCCat {Γ = Γ} {P = P} {dO = dO} {dS = dS} {u = u} = natelimStr₁S Γ P _ dO _ _ dS _ _ u _ _ 212 | CCatwithnatelim.natelimStrNat' natelimStrSynCCat = //-elimP (λ g → //-elimP (λ Δ g₀ → //-elimP (λ Γ → //-elimP (λ P P= → //-elimP (λ dO dOₛ dO₁ → //-elimP (λ dS dSₛ dS₁ → //-elimP (λ u uₛ u₁ g₁ → up-to-rhsTyEq' (reflectOb g₀) (ap-[]Ty refl (idMor[]Mor (mor g)) ∙ []Ty-substTy)))))))) 213 | 214 | 215 | module natelimS = CCatwithnatelim natelimStrSynCCat 216 | 217 | {- ElNat= -} 218 | 219 | elnatStrS : (i : ℕ) (Γ : ObS n) → ElStrS i Γ (natStrS i Γ) (natStrₛS i Γ) (natStr₁S i Γ) ≡ NatStrS Γ 220 | elnatStrS i = //-elimP (λ Γ → eq (box (CtxRefl (der Γ) , ElNat=))) 221 | 222 | 223 | {- BetaNatZero -} 224 | 225 | betaNatZeroStrS : (Γ : ObS n) (P : ObS (suc (suc n))) (P= : ftS P ≡ NatStrS Γ) 226 | (dO : MorS n (suc n)) (dOₛ : S.is-section dO) (dO₁ : S.∂₁ dO ≡ S.star (zeroStrS Γ) P P= (zeroStr₁S Γ)) 227 | (dS : MorS (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section dS) (dS₁ : S.∂₁ dS ≡ sucS.T-dS₁ Γ P P=) → 228 | natelimStrS Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ (zeroStrS Γ) (zeroStrₛS Γ) (zeroStr₁S Γ) ≡ dO 229 | betaNatZeroStrS = //-elimP (λ Γ → //-elimP (λ P P= → (//-elimP (λ dO dOₛ dO₁ → //-elimP (λ dS dSₛ dS₁ → eq (box 230 | (CtxSymm (reflectOb (S.is-section₀ dOₛ dO₁))) 231 | (CtxSymm (reflectOb dO₁)) 232 | (MorTran (der Γ) (der Γ , SubstTy (dTy P P=) (idMor+ (der Γ) Zero)) 233 | (idMor+= (der Γ) 234 | (let ddS : Derivable (((ctx Γ , nat) , getTy P) ⊢ getTm dS :> substTy (weakenTy' (prev last) (weakenTy' (prev last) (getTy P))) (suc (var (prev last)))) 235 | ddS = congTmTy! fixSubstTy (dTm≃ (Ctx≃ft+Ty (reflect P=)) dS dSₛ (reflect dS₁)) 236 | in BetaNatZero (dTy P P=) (dTm refl dO dOₛ dO₁) ddS)) (MorSymm (der Γ) (der Γ , SubstTy (dTy P P=) (idMor+ (der Γ) Zero)) (morTm=idMorTm refl dO dOₛ dO₁))))))))) 237 | 238 | 239 | {- BetaNatSuc -} 240 | 241 | betaNatSucStrS : (Γ : ObS n) (P : ObS (suc (suc n))) (P= : ftS P ≡ NatStrS Γ) 242 | (dO : MorS n (suc n)) (dOₛ : S.is-section dO) (dO₁ : S.∂₁ dO ≡ S.star (zeroStrS Γ) P P= (zeroStr₁S Γ)) 243 | (dS : MorS (suc (suc n)) (suc (suc (suc n)))) (dSₛ : S.is-section dS) (dS₁ : S.∂₁ dS ≡ sucS.T-dS₁ Γ P P=) 244 | (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ NatStrS Γ) → 245 | natelimStrS Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ (sucStrS Γ u uₛ u₁) (sucStrₛS Γ u) (sucStr₁S Γ u) ≡ natelimS.Tm-substdS Γ P P= dO dOₛ dO₁ dS dSₛ dS₁ u uₛ u₁ 246 | betaNatSucStrS = //-elimP (λ Γ → //-elimP (λ P P= → (//-elimP (λ dO dOₛ dO₁ → //-elimP (λ dS dSₛ dS₁ → //-elimP (λ u uₛ u₁ → 247 | let ddS : Derivable (((ctx Γ , nat) , getTy P) ⊢ getTm dS :> substTy (weakenTy' (prev last) (weakenTy' (prev last) (getTy P))) (suc (var (prev last)))) 248 | ddS = congTmTy! fixSubstTy (dTm≃ (Ctx≃ft+Ty (reflect P=)) dS dSₛ (reflect dS₁)) 249 | in up-to-rhsTyEq2 (congTyEq (ap-[]Ty ([idMor]Ty _ ∙ ! fixSubstTy) refl 250 | ∙ []Ty-assoc _ _ _ 251 | ∙ ap-[]Ty refl (Mor+= (Mor+= (Mor+= (weakenMorInsert _ _ _ ∙ weakenMorInsert _ _ _ ∙ [idMor]Mor _) refl) refl) refl) 252 | ∙ weakenTyInsert' _ _ _ _ ∙ weakenTyInsert' _ _ _ _) 253 | (! (ap-[]Ty (! ([]Ty-assoc _ _ _)) (Mor+= (weakenMorInsert _ _ _ ∙ idMor[]Mor _) refl) 254 | ∙ []Ty-assoc _ _ _ 255 | ∙ ap-[]Ty refl (Mor+= (weakenMorInsert (mor u) _ _ ∙ [idMor]Mor _) refl))) 256 | (SubstTyFullEq' (der Γ) ((der Γ , Nat) , dTy P P=) 257 | (SubstTyFullEq' ((der Γ , Nat) , dTy P P=) ((der Γ , Nat) , dTy P P=) 258 | (dTy=≃ (reflect (! dS₁)) (Ctx≃ft+Ty (reflect P=))) 259 | (MorSymm ((der Γ , Nat) , dTy P P=) 260 | ((der Γ , Nat) , dTy P P=) 261 | (getMor=idMor≃ (Ctx≃ft+Ty (reflect P=)) dS dSₛ (reflect dS₁)))) 262 | (MorSymm (der Γ) (der Γ , Nat) 263 | (morTm=idMorTm refl u uₛ u₁) 264 | , TmRefl (Natelim (dTy P P=) (dTm refl dO dOₛ dO₁) ddS (dTm refl u uₛ u₁))))) 265 | (TmTran' (der Γ) (BetaNatSuc (dTy P P=) (dTm refl dO dOₛ dO₁) ddS (dTm refl u uₛ u₁)) 266 | (congTmEq refl (! ([]Tm-assoc _ _ (getTm dS) ∙ ap-[]Tm {u = getTm dS} refl (Mor+= (weakenMorInsert (mor u) _ _ ∙ [idMor]Mor _) refl))) 267 | ([]Ty-substTy ∙ ap-[]Ty (weakenTyInsert' _ _ _ _ ∙ weakenTyInsert' _ _ _ _ ∙ [idMor]Ty _) refl) 268 | (SubstTmMorEq' (der Γ) ((der Γ , Nat) , dTy P P=) ddS (MorSymm (der Γ) (der Γ , Nat) 269 | (morTm=idMorTm refl u uₛ u₁) 270 | , TmRefl (Natelim (dTy P P=) (dTm refl dO dOₛ dO₁) ddS (dTm refl u uₛ u₁)))))))))))) 271 | 272 | -------------------------------------------------------------------------------- /termmodel-pi.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import reflection hiding (proj) 6 | open import syntx hiding (getTy) 7 | open import rules 8 | open import contextualcat 9 | open import quotients 10 | open import termmodel-common 11 | open import termmodel-synccat 12 | open import termmodel-uuel 13 | 14 | open CCat hiding (Mor) renaming (id to idC) 15 | 16 | 17 | {- Pi -} 18 | 19 | PiStrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (B : DCtx (suc (suc n))) (B= : ftS (proj B) ≡ proj A) → DCtx (suc n) 20 | PiStrS-// Γ A A= B B= = dctx (der Γ , Pi (dTy A A=) (dTy+ A= B B=)) 21 | 22 | PiStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : _) (A'= : _) {B B' : DCtx (suc (suc n))} (rB : B ≃ B') (B= : _) (B'= : _) 23 | → PiStrS-// Γ A A= B B= ≃ PiStrS-// Γ' A' A'= B' B'= 24 | PiStrS-eq rΓ {A = A} rA A= A'= {B = B} rB B= B'= = box (unOb≃ rΓ , PiCong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=)) 25 | 26 | PiStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) → ObS (suc n) 27 | PiStrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Ty (λ B B= → proj (PiStrS-// Γ A A= B B=)) 28 | (λ rB B= B'= → proj= (PiStrS-eq (ref Γ) (ref A) A= A= rB B= B'=))) 29 | (λ rA A= A'= → //-elimP-Ty (λ B B= B=' → proj= (PiStrS-eq (ref Γ) rA A= A'= (ref B) B= B=')))) 30 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Ty (λ B B= B=' → proj= (PiStrS-eq rΓ (ref A) A= A=' (ref B) B= B=')))) 31 | 32 | PiStr=S : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) → ftS (PiStrS Γ A A= B B=) ≡ Γ 33 | PiStr=S = //-elimP-Ctx (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → refl))) 34 | 35 | PiStrSynCCat : CCatwithPi synCCat 36 | CCatwithPi.PiStr PiStrSynCCat = PiStrS 37 | CCatwithPi.PiStr= PiStrSynCCat = PiStr=S _ _ _ _ _ 38 | CCatwithPi.PiStrNat' PiStrSynCCat = //-elimP (λ g → //-elimP (λ Δ g₀ → (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= g₁ → up-to-CtxTyEq (reflectOb g₀) refl)))))) 39 | 40 | 41 | {- pi -} 42 | 43 | piStrS-// : (i : ℕ) (Γ : DCtx n) (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : ∂₁S (proj a) ≡ UUStrS i (proj Γ)) (b : DMor (suc n) (suc (suc n))) (bₛ : S.is-section (proj b)) (b₁ : ∂₁S (proj b) ≡ UUStrS i (ElStrS i (proj Γ) (proj a) aₛ a₁)) → DMor n (suc n) 44 | piStrS-// i Γ a aₛ a₁ b bₛ b₁ = dmorTm Γ (PiUU (dTm refl a aₛ a₁) (dTm refl b bₛ b₁)) 45 | 46 | piStrS-eq : (i : ℕ) {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {a a' : DMor n (suc n)} (ra : a ≃ a') (aₛ : S.is-section (proj a)) (a'ₛ : S.is-section (proj a')) (a₁ : ∂₁S (proj a) ≡ UUStrS i (proj Γ)) (a'₁ : ∂₁S (proj a') ≡ UUStrS i (proj Γ') ) {b b' : DMor (suc n) (suc (suc n))} (rb : b ≃ b') (bₛ : S.is-section (proj b)) (b'ₛ : S.is-section (proj b')) (b₁ : ∂₁S (proj b) ≡ UUStrS i (ElStrS i (proj Γ) (proj a) aₛ a₁)) (b'₁ : ∂₁S (proj b') ≡ UUStrS i (ElStrS i (proj Γ') (proj a') a'ₛ a'₁)) 47 | → piStrS-// i Γ a aₛ a₁ b bₛ b₁ ≃ piStrS-// i Γ' a' a'ₛ a'₁ b' b'ₛ b'₁ 48 | piStrS-eq i rΓ ra aₛ a'ₛ a₁ a'₁ rb bₛ b'ₛ b₁ b'₁ = 49 | dmorTm= dmorTmₛ dmorTmₛ rΓ UUCong (PiUUCong (dTm refl _ aₛ a₁) 50 | (dTm= refl ra aₛ a'ₛ a₁ a'₁) 51 | (dTm= refl rb bₛ b'ₛ b₁ b'₁)) 52 | 53 | 54 | piStrS : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (b : MorS (suc n) (suc (suc n))) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ UUStrS i (ElStrS i Γ a aₛ a₁)) → MorS n (suc n) 55 | piStrS i = //-elim-Ctx (λ Γ → //-elim-Tm (λ a aₛ a₁ → //-elim-Tm (λ b bₛ b₁ → proj {R = MorEquiv} (piStrS-// i Γ a aₛ a₁ b bₛ b₁)) 56 | (λ rb bₛ b'ₛ b₁ b'₁ → proj= {R = MorEquiv} (piStrS-eq i (ref Γ) (ref a) aₛ aₛ a₁ a₁ rb bₛ b'ₛ b₁ b'₁))) 57 | (λ ra aₛ a'ₛ a₁ a'₁ → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= {R = MorEquiv} (piStrS-eq i (ref Γ) ra aₛ a'ₛ a₁ a'₁ (ref b) bₛ bₛ b₁ b₁')))) 58 | (λ rΓ → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= {R = MorEquiv} (piStrS-eq i rΓ (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁')))) 59 | 60 | piStrₛS : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (b : MorS (suc n) (suc (suc n))) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ UUStrS i (ElStrS i Γ a aₛ a₁)) → S.is-section (piStrS i Γ a aₛ a₁ b bₛ b₁) 61 | piStrₛS i = //-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → dmorTmₛ))) 62 | 63 | piStr₁S : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (b : MorS (suc n) (suc (suc n))) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ UUStrS i (ElStrS i Γ a aₛ a₁)) → ∂₁S (piStrS i Γ a aₛ a₁ b bₛ b₁) ≡ UUStrS i Γ 64 | piStr₁S i = //-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → refl))) 65 | 66 | piStrSynCCat : CCatwithpi synCCat UUStrSynCCat ElStrSynCCat 67 | CCatwithpi.piStr piStrSynCCat = piStrS 68 | CCatwithpi.piStrₛ piStrSynCCat {Γ = Γ} {a = a} {b = b} = piStrₛS _ Γ a _ _ b _ _ 69 | CCatwithpi.piStr₁ piStrSynCCat {Γ = Γ} {a = a} {b = b} = piStr₁S _ Γ a _ _ b _ _ 70 | CCatwithpi.piStrNat' piStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ g₁ → refl))))) 71 | 72 | 73 | {- lam -} 74 | 75 | lamStrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (B : DCtx (suc (suc n))) (B= : ftS (proj B) ≡ proj A) (u : DMor (suc n) (suc (suc n))) (uₛ : S.is-section (proj u)) (u₁ : ∂₁S (proj u) ≡ proj B) → DMor n (suc n) 76 | lamStrS-// Γ A A= B B= u uₛ u₁ = dmorTm Γ (Lam (dTy A A=) (dTy+ A= B B=) (dTm+ A= B= u uₛ u₁)) 77 | 78 | lamStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : _) (A'= : _) {B B' : DCtx (suc (suc n))} (rB : B ≃ B') (B= : _) (B'= : _) {u u' : DMor (suc n) (suc (suc n))} (ru : u ≃ u') (uₛ : _) (u'ₛ : _) (u₁ : _) (u'₁ : _) → lamStrS-// Γ A A= B B= u uₛ u₁ ≃ lamStrS-// Γ' A' A'= B' B'= u' u'ₛ u'₁ 79 | lamStrS-eq rΓ {A = A} rA A= A'= rB B= B'= ru uₛ u'ₛ u₁ u'₁ = dmorTm= dmorTmₛ dmorTmₛ rΓ 80 | (PiCong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=)) 81 | (LamCong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=) (dTm+= A= B= ru uₛ u'ₛ u₁ u'₁)) 82 | 83 | lamStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) (u : MorS (suc n) (suc (suc n))) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ B) → MorS n (suc n) 84 | lamStrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Ty (λ B B= → //-elim-Tm (λ u uₛ u₁ → proj (lamStrS-// Γ A A= B B= u uₛ u₁)) 85 | (λ ru uₛ u'ₛ u₁ u'₁ → proj= (lamStrS-eq (ref Γ) (ref A) A= A= (ref B) B= B= ru uₛ u'ₛ u₁ u'₁))) 86 | (λ rB B= B'= → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (lamStrS-eq (ref Γ) (ref A) A= A= rB B= B'= (ref u) uₛ uₛ u₁ u₁')))) 87 | (λ rA A= A'= → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (lamStrS-eq (ref Γ) rA A= A'= (ref B) B= B=' (ref u) uₛ uₛ u₁ u₁'))))) 88 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (lamStrS-eq rΓ (ref A) A= A=' (ref B) B= B=' (ref u) uₛ uₛ u₁ u₁'))))) 89 | 90 | lamStrₛS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) (u : MorS (suc n) (suc (suc n))) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ B) → S.is-section (lamStrS Γ A A= B B= u uₛ u₁) 91 | lamStrₛS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ → dmorTmₛ)))) 92 | 93 | lamStr₁S : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) (u : MorS (suc n) (suc (suc n))) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ B) → ∂₁S (lamStrS Γ A A= B B= u uₛ u₁) ≡ PiStrS Γ A A= B B= 94 | lamStr₁S = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ → refl)))) 95 | 96 | 97 | lamStrSynCCat : CCatwithlam synCCat PiStrSynCCat 98 | CCatwithlam.lamStr lamStrSynCCat = lamStrS 99 | CCatwithlam.lamStrₛ lamStrSynCCat {Γ = Γ} {A = A} {B = B} {u = u} = lamStrₛS Γ A _ B _ u _ _ 100 | CCatwithlam.lamStr₁ lamStrSynCCat {Γ = Γ} {A = A} {B = B} {u = u} = lamStr₁S Γ A _ B _ u _ _ 101 | CCatwithlam.lamStrNat' lamStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ g₁ → up-to-rhsTyEq (ap (_[_]Ty (pi (getTy A) (getTy B))) (idMor[]Mor (mor g))))))))) 102 | 103 | 104 | {- app -} 105 | 106 | appStrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (B : DCtx (suc (suc n))) (B= : ftS (proj B) ≡ proj A) (f : DMor n (suc n)) (fₛ : S.is-section (proj f)) (f₁ : ∂₁S (proj f) ≡ PiStrS (proj Γ) (proj A) A= (proj B) B=) (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : ∂₁S (proj a) ≡ proj A) → DMor n (suc n) 107 | appStrS-// Γ A A= B B= f fₛ f₁ a aₛ a₁ = dmorTm Γ (App (dTy A A=) (dTy+ A= B B=) (dTm refl f fₛ f₁) (dTm A= a aₛ a₁)) 108 | 109 | appStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : ftS (proj A) ≡ proj Γ) (A'= : ftS (proj A') ≡ proj Γ') {B B' : DCtx (suc (suc n))} (rB : B ≃ B') (B= : ftS (proj B) ≡ proj A) (B'= : ftS (proj B') ≡ proj A') {f f' : DMor n (suc n)} (rf : f ≃ f') (fₛ : S.is-section (proj f)) (f'ₛ : S.is-section (proj f')) (f₁ : ∂₁S (proj f) ≡ PiStrS (proj Γ) (proj A) A= (proj B) B=) (f₁' : ∂₁S (proj f') ≡ PiStrS (proj Γ') (proj A') A'= (proj B') B'=) {a a' : DMor n (suc n)} (ra : a ≃ a') (aₛ : S.is-section (proj a)) (a'ₛ : S.is-section (proj a')) (a₁ : ∂₁S (proj a) ≡ proj A) (a'₁ : ∂₁S (proj a') ≡ proj A') 110 | → appStrS-// Γ A A= B B= f fₛ f₁ a aₛ a₁ ≃ appStrS-// Γ' A' A'= B' B'= f' f'ₛ f₁' a' a'ₛ a'₁ 111 | appStrS-eq {Γ = Γ} rΓ {A} rA A= A'= {B} {B'} rB B= B'= rf fₛ f'ₛ f₁ f'₁ ra aₛ a'ₛ a₁ a'₁ = dmorTm= dmorTmₛ dmorTmₛ rΓ 112 | (SubstTyFullEq' (der Γ) (der Γ , dTy A A=) (dTy+= A= rB B=) (idMor+= (der Γ) (dTm= A= ra aₛ a'ₛ a₁ a'₁))) 113 | (AppCong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=) (dTm= refl rf fₛ f'ₛ f₁ f'₁) (dTm= A= ra aₛ a'ₛ a₁ a'₁)) 114 | 115 | 116 | appStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) (f : MorS n (suc n)) (fₛ : S.is-section f) (f₁ : ∂₁S f ≡ PiStrS Γ A A= B B=) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) 117 | → MorS n (suc n) 118 | appStrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Ty (λ B B= → //-elim-Tm (λ f fₛ f₁ → //-elim-Tm (λ a aₛ a₁ → proj (appStrS-// Γ A A= B B= f fₛ f₁ a aₛ a₁)) 119 | (λ ra aₛ a'ₛ a₁ a'₁ → proj= (appStrS-eq (ref Γ) (ref A) A= A= (ref B) B= B= (ref f) fₛ fₛ f₁ f₁ ra aₛ a'ₛ a₁ a'₁))) 120 | (λ rf fₛ f'ₛ f₁ f'₁ → //-elimP-Tm (λ a aₛ a₁ a₁' → proj= (appStrS-eq (ref Γ) (ref A) A= A= (ref B) B= B= rf fₛ f'ₛ f₁ f'₁ (ref a) aₛ aₛ a₁ a₁')))) 121 | (λ rB B= B'= → //-elimP-Tm (λ f fₛ f₁ f₁' → //-elimP-Tm (λ a aₛ a₁ a₁' → proj= (appStrS-eq (ref Γ) (ref A) A= A= rB B= B'= (ref f) fₛ fₛ f₁ f₁' (ref a) aₛ aₛ a₁ a₁'))))) 122 | (λ rA A= A'= → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ f fₛ f₁ f₁' → //-elimP-Tm (λ a aₛ a₁ a₁' → proj= (appStrS-eq (ref Γ) rA A= A'= (ref B) B= B=' (ref f) fₛ fₛ f₁ f₁' (ref a) aₛ aₛ a₁ a₁')))))) 123 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ f fₛ f₁ f₁' → //-elimP-Tm (λ a aₛ a₁ a₁' → proj= (appStrS-eq rΓ (ref A) A= A=' (ref B) B= B=' (ref f) fₛ fₛ f₁ f₁' (ref a) aₛ aₛ a₁ a₁')))))) 124 | 125 | appStrₛS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) (f : MorS n (suc n)) (fₛ : S.is-section f) (f₁ : ∂₁S f ≡ PiStrS Γ A A= B B=) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) 126 | → S.is-section (appStrS Γ A A= B B= f fₛ f₁ a aₛ a₁) 127 | appStrₛS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ f fₛ f₁ → //-elimP (λ a aₛ a₁ → dmorTmₛ))))) 128 | 129 | appStr₁S : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) (f : MorS n (suc n)) (fₛ : S.is-section f) (f₁ : ∂₁S f ≡ PiStrS Γ A A= B B=) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) 130 | → ∂₁S (appStrS Γ A A= B B= f fₛ f₁ a aₛ a₁) ≡ S.star a B B= a₁ 131 | appStr₁S = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ f fₛ f₁ → //-elimP (λ a aₛ a₁ → eq (box (CtxSymm ((reflectOb (S.is-section₀ aₛ a₁ ∙ A=)) , SubstTyMorEq (dTy B B=) (ConvMor (morDer a) (CtxRefl (der (lhs a))) (reflectOb a₁)) (ConvMorEq (morTm=idMorTm' aₛ) (CtxRefl (der (lhs a))) (reflectOb a₁)))))))))) 132 | 133 | appStrSynCCat : CCatwithapp synCCat PiStrSynCCat 134 | CCatwithapp.appStr appStrSynCCat = appStrS 135 | CCatwithapp.appStrₛ appStrSynCCat {Γ = Γ} {A = A} {B = B} {f = f} {a = a} = appStrₛS Γ A _ B _ f _ _ a _ _ 136 | CCatwithapp.appStr₁ appStrSynCCat {Γ = Γ} {A = A} {B = B} {f = f} {a = a} = appStr₁S Γ A _ B _ f _ _ a _ _ 137 | CCatwithapp.appStrNat' appStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ f fₛ f₁ → //-elimP (λ a aₛ a₁ g₁ → up-to-rhsTyEq (ap (_[_]Ty (substTy (getTy B) (getTm a))) (idMor[]Mor (mor g)) ∙ []Ty-substTy)))))))) 138 | 139 | 140 | {- ElPi= -} 141 | 142 | elpiStrS : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (b : MorS (suc n) (suc (suc n))) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ UUStrS i (ElStrS i Γ a aₛ a₁)) 143 | → ElStrS i Γ (piStrS i Γ a aₛ a₁ b bₛ b₁) (piStrₛS i Γ a aₛ a₁ b bₛ b₁) (piStr₁S i Γ a aₛ a₁ b bₛ b₁) ≡ PiStrS Γ (ElStrS i Γ a aₛ a₁) (ElStr=S i Γ a aₛ a₁) (ElStrS i (ElStrS i Γ a aₛ a₁) b bₛ b₁) (ElStr=S i (ElStrS i Γ a aₛ a₁) b bₛ b₁) 144 | elpiStrS i = //-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → eq (box (CtxRefl (der Γ) , ElPi= (dTm refl a aₛ a₁) (dTm refl b bₛ b₁)))))) 145 | 146 | 147 | {- BetaPi -} 148 | 149 | betaPiStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (u : MorS (suc n) (suc (suc n))) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ B) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) 150 | → appStrS Γ A A= B B= (lamStrS Γ A A= B B= u uₛ u₁) (lamStrₛS Γ A A= B B= u uₛ u₁) (lamStr₁S Γ A A= B B= u uₛ u₁) a aₛ a₁ ≡ S.starTm a u (S.is-section₀ uₛ u₁ ∙ B=) a₁ 151 | betaPiStrS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ → //-elimP (λ a aₛ a₁ → eq (box 152 | (CtxSymm (CtxTran (reflectOb (S.is-section₀ aₛ a₁)) (reflectOb A=))) 153 | (CtxSymm (CtxTran (reflectOb (S.is-section₀ aₛ a₁)) (reflectOb A=)) , SubstTyFullEq' (der Γ) (der Γ , dTy A A=) 154 | (dTy+= A= (sym (reflect u₁)) B=) 155 | (MorTran (der Γ) (der Γ , dTy A A=) 156 | (MorSymm (der Γ) (der Γ , (dTy A A=)) (morTm=idMorTm A= a aₛ a₁)) 157 | (congMorEq refl refl (idMor[]Mor _) refl 158 | (SubstMorEq (MorSymm (der Γ , dTy A A=) (der Γ , dTy A A=) 159 | (getMor=idMor (combine A= B B=) u uₛ u₁)) (dMor A= a aₛ a₁))))) 160 | (idMor+= (der Γ) (TmTran (SubstTm (dTm+ A= B= u uₛ u₁) (idMor+ (der Γ) (dTm A= a aₛ a₁))) 161 | (BetaPi (dTy A A=) (dTy+ A= B B=) (dTm+ A= B= u uₛ u₁) (dTm A= a aₛ a₁)) 162 | (SubstTmMorEq (dTm+ A= B= u uₛ u₁) (idMor+ (der Γ) (dTm A= a aₛ a₁)) (MorSymm (der Γ) (der Γ , dTy A A=) (morTm=idMorTm A= a aₛ a₁))))))))))) 163 | 164 | {- EtaPi -} 165 | 166 | etaPiStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (f : MorS n (suc n)) (fₛ : S.is-section f) (f₁ : ∂₁S f ≡ PiStrS Γ A A= B B=) 167 | → f ≡ T-rhsEtaPi PiStrSynCCat lamStrSynCCat appStrSynCCat Γ A A= B B= f fₛ f₁ 168 | etaPiStrS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ f fₛ f₁ → eq (box (reflectOb (S.is-section₀ fₛ f₁)) (reflectOb f₁) (ConvMorEq (MorTran (der Γ) (der Γ , Pi (dTy A A=) (dTy+ A= B B=)) (morTm=idMorTm refl f fₛ f₁) (idMor+= (der Γ) (congTmEq refl (ap-lam-Tm refl refl (ap-app-Tm weakenTy-to-[]Ty weakenTy+-to-[]Ty weakenTm-to-[]Tm refl)) refl (EtaPi (dTy A A=) (dTy+ A= B B=) (dTm refl f fₛ f₁))))) (reflectOb (! (S.is-section₀ fₛ f₁))) (reflectOb (! f₁)))))))) 169 | -------------------------------------------------------------------------------- /termmodel-sig.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import syntx hiding (getTy) 6 | open import rules 7 | open import contextualcat 8 | open import quotients 9 | open import termmodel-common 10 | open import termmodel-synccat 11 | open import termmodel-uuel 12 | 13 | open CCat hiding (Mor) renaming (id to idC) 14 | 15 | 16 | {- Sig -} 17 | 18 | SigStrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : ftS (proj A) ≡ proj Γ) (B : DCtx (suc (suc n))) (B= : ftS (proj B) ≡ proj A) → DCtx (suc n) 19 | SigStrS-// Γ A A= B B= = dctx (der Γ , Sig (dTy A A=) (dTy+ A= B B=)) 20 | 21 | SigStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : _) (A'= : _) {B B' : DCtx (suc (suc n))} (rB : B ≃ B') (B= : _) (B'= : _) 22 | → SigStrS-// Γ A A= B B= ≃ SigStrS-// Γ' A' A'= B' B'= 23 | SigStrS-eq rΓ {A = A} rA A= A'= rB B= B'= = box (unOb≃ rΓ , SigCong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=)) 24 | 25 | SigStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) → ObS (suc n) 26 | SigStrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Ty (λ B B= → proj (SigStrS-// Γ A A= B B=)) 27 | (λ rB B= B'= → proj= (SigStrS-eq (ref Γ) (ref A) A= A= rB B= B'=))) 28 | (λ rA A= A'= → //-elimP-Ty (λ B B= B=' → proj= (SigStrS-eq (ref Γ) rA A= A'= (ref B) B= B=')))) 29 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Ty (λ B B= B=' → proj= (SigStrS-eq rΓ (ref A) A= A=' (ref B) B= B=')))) 30 | 31 | SigStr=S : (Γ : ObS n) (A : ObS (suc n)) (A= : ftS A ≡ Γ) (B : ObS (suc (suc n))) (B= : ftS B ≡ A) → ftS (SigStrS Γ A A= B B=) ≡ Γ 32 | SigStr=S = //-elimP-Ctx (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → refl))) 33 | 34 | SigStrSynCCat : CCatwithSig synCCat 35 | CCatwithSig.SigStr SigStrSynCCat = SigStrS 36 | CCatwithSig.SigStr= SigStrSynCCat = SigStr=S _ _ _ _ _ 37 | CCatwithSig.SigStrNat' SigStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= g₁ → refl))))) 38 | 39 | 40 | {- sig -} 41 | 42 | sigStrS-// : (i : ℕ) (Γ : DCtx n) (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : ∂₁S (proj a) ≡ UUStrS i (proj Γ)) (b : DMor (suc n) (suc (suc n))) (bₛ : S.is-section (proj b)) (b₁ : ∂₁S (proj b) ≡ UUStrS i (ElStrS i (proj Γ) (proj a) aₛ a₁)) → DMor n (suc n) 43 | sigStrS-// i Γ a aₛ a₁ b bₛ b₁ = dmorTm Γ (SigUU (dTm refl a aₛ a₁) (dTm refl b bₛ b₁)) 44 | 45 | sigStrS-eq : (i : ℕ) {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {a a' : DMor n (suc n)} (ra : a ≃ a') (aₛ : _) (a'ₛ : _) (a₁ : _) (a'₁ : _) {b b' : DMor (suc n) (suc (suc n))} (rb : b ≃ b') (bₛ : _) (b'ₛ : _) (b₁ : _) (b'₁ : _) 46 | → sigStrS-// i Γ a aₛ a₁ b bₛ b₁ ≃ sigStrS-// i Γ' a' a'ₛ a'₁ b' b'ₛ b'₁ 47 | sigStrS-eq i rΓ ra aₛ a'ₛ a₁ a'₁ rb bₛ b'ₛ b₁ b'₁ = 48 | dmorTm= dmorTmₛ dmorTmₛ rΓ UUCong (SigUUCong (dTm refl _ aₛ a₁) 49 | (dTm= refl ra aₛ a'ₛ a₁ a'₁) 50 | (dTm= refl rb bₛ b'ₛ b₁ b'₁)) 51 | 52 | 53 | sigStrS : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (b : MorS (suc n) (suc (suc n))) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ UUStrS i (ElStrS i Γ a aₛ a₁)) → MorS n (suc n) 54 | sigStrS i = //-elim-Ctx (λ Γ → //-elim-Tm (λ a aₛ a₁ → //-elim-Tm (λ b bₛ b₁ → proj (sigStrS-// i Γ a aₛ a₁ b bₛ b₁)) 55 | (λ rb bₛ b'ₛ b₁ b'₁ → proj= (sigStrS-eq i (ref Γ) (ref a) aₛ aₛ a₁ a₁ rb bₛ b'ₛ b₁ b'₁))) 56 | (λ ra aₛ a'ₛ a₁ a'₁ → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= (sigStrS-eq i (ref Γ) ra aₛ a'ₛ a₁ a'₁ (ref b) bₛ bₛ b₁ b₁')))) 57 | (λ rΓ → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= (sigStrS-eq i rΓ (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁')))) 58 | 59 | sigStrₛS : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (b : MorS (suc n) (suc (suc n))) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ UUStrS i (ElStrS i Γ a aₛ a₁)) → S.is-section (sigStrS i Γ a aₛ a₁ b bₛ b₁) 60 | sigStrₛS i = //-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → dmorTmₛ))) 61 | 62 | sigStr₁S : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (b : MorS (suc n) (suc (suc n))) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ UUStrS i (ElStrS i Γ a aₛ a₁)) → ∂₁S (sigStrS i Γ a aₛ a₁ b bₛ b₁) ≡ UUStrS i Γ 63 | sigStr₁S i = //-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → refl))) 64 | 65 | sigStrSynCCat : CCatwithsig synCCat UUStrSynCCat ElStrSynCCat 66 | CCatwithsig.sigStr sigStrSynCCat = sigStrS 67 | CCatwithsig.sigStrₛ sigStrSynCCat {Γ = Γ} {a = a} {b = b} = sigStrₛS _ Γ a _ _ b _ _ 68 | CCatwithsig.sigStr₁ sigStrSynCCat {Γ = Γ} {a = a} {b = b} = sigStr₁S _ Γ a _ _ b _ _ 69 | CCatwithsig.sigStrNat' sigStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ g₁ → refl))))) 70 | 71 | 72 | {- pair -} 73 | 74 | pairStrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : S.ft (proj A) ≡ proj Γ) (B : DCtx (suc (suc n))) (B= : S.ft (proj B) ≡ proj A) (a : DMor n (suc n)) (aₛ : S.is-section (proj a)) (a₁ : S.∂₁ (proj a) ≡ proj A) (b : DMor n (suc n)) (bₛ : S.is-section (proj b)) (b₁ : S.∂₁ (proj b) ≡ S.star (proj a) (proj B) B= a₁) → DMor n (suc n) 75 | pairStrS-// Γ A A= B B= a aₛ a₁ b bₛ b₁ = dmorTm Γ (Pair (dTy A A=) (dTy+ A= B B=) (dTm A= a aₛ a₁) (dTmSubst A= B B= a aₛ a₁ b bₛ b₁)) 76 | 77 | 78 | pairStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : S.ft (proj A) ≡ proj Γ) (A'= : S.ft (proj A') ≡ proj Γ') {B B' : DCtx (suc (suc n))} (rB : B ≃ B') (B= : S.ft (proj B) ≡ proj A) (B'= : S.ft (proj B') ≡ proj A') {a a' : DMor n (suc n)} (ra : a ≃ a') (aₛ : S.is-section (proj a)) (a'ₛ : S.is-section (proj a')) (a₁ : S.∂₁ (proj a) ≡ proj A) (a'₁ : S.∂₁ (proj a') ≡ proj A') {b b' : DMor n (suc n)} (rb : b ≃ b') (bₛ : S.is-section (proj b)) (b'ₛ : S.is-section (proj b')) (b₁ : S.∂₁ (proj b) ≡ S.star (proj a) (proj B) B= a₁) (b'₁ : S.∂₁ (proj b') ≡ S.star (proj a') (proj B') B'= a'₁) → pairStrS-// Γ A A= B B= a aₛ a₁ b bₛ b₁ ≃ pairStrS-// Γ' A' A'= B' B'= a' a'ₛ a'₁ b' b'ₛ b'₁ 79 | pairStrS-eq rΓ {A} {A'} rA A= A'= {B} {B'} rB B= B'= {a} {a'} ra aₛ a'ₛ a₁ a'₁ {b} {b'} rb bₛ b'ₛ b₁ b'₁ = 80 | dmorTm= dmorTmₛ dmorTmₛ rΓ (SigCong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=)) 81 | (PairCong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=) (dTm= A= ra aₛ a'ₛ a₁ a'₁) (dTmSubst= A= rB B= B'= ra aₛ a'ₛ a₁ a'₁ rb bₛ b'ₛ b₁ b'₁)) 82 | 83 | pairStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ A) (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ S.star a B B= a₁) → MorS n (suc n) 84 | pairStrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Ty (λ B B= → //-elim-Tm (λ a aₛ a₁ → //-elim-Tm (λ b bₛ b₁ → proj (pairStrS-// Γ A A= B B= a aₛ a₁ b bₛ b₁)) 85 | (λ rb bₛ b'ₛ b₁ b'₁ → proj= (pairStrS-eq (ref Γ) (ref A) A= A= (ref B) B= B= (ref a) aₛ aₛ a₁ a₁ rb bₛ b'ₛ b₁ b'₁))) 86 | (λ ra aₛ a'ₛ a₁ a'₁ → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= (pairStrS-eq (ref Γ) (ref A) A= A= (ref B) B= B= ra aₛ a'ₛ a₁ a'₁ (ref b) bₛ bₛ b₁ b₁')))) 87 | (λ rB B= B'= → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= (pairStrS-eq (ref Γ) (ref A) A= A= rB B= B'= (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁'))))) 88 | (λ rA A= A'= → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= (pairStrS-eq (ref Γ) rA A= A'= (ref B) B= B=' (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁')))))) 89 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ a aₛ a₁ a₁' → //-elimP-Tm (λ b bₛ b₁ b₁' → proj= (pairStrS-eq rΓ (ref A) A= A=' (ref B) B= B=' (ref a) aₛ aₛ a₁ a₁' (ref b) bₛ bₛ b₁ b₁')))))) 90 | 91 | pairStrₛS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ A) (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ S.star a B B= a₁) → S.is-section (pairStrS Γ A A= B B= a aₛ a₁ b bₛ b₁) 92 | pairStrₛS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → dmorTmₛ))))) 93 | 94 | pairStr₁S : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : S.∂₁ a ≡ A) (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : S.∂₁ b ≡ S.star a B B= a₁) → S.∂₁ (pairStrS Γ A A= B B= a aₛ a₁ b bₛ b₁) ≡ SigStrS Γ A A= B B= 95 | pairStr₁S = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → refl))))) 96 | 97 | pairStrSynCCat : CCatwithpair synCCat SigStrSynCCat 98 | CCatwithpair.pairStr pairStrSynCCat = pairStrS 99 | CCatwithpair.pairStrₛ pairStrSynCCat {Γ = Γ} {A = A} {B = B} {a = a} {b = b} = pairStrₛS Γ A _ B _ a _ _ b _ _ 100 | CCatwithpair.pairStr₁ pairStrSynCCat {Γ = Γ} {A = A} {B = B} {a = a} {b = b} = pairStr₁S Γ A _ B _ a _ _ b _ _ 101 | CCatwithpair.pairStrNat' pairStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ g₁ → up-to-rhsTyEq (ap (_[_]Ty (sig (getTy A) (getTy B))) (idMor[]Mor (mor g)))))))))) 102 | 103 | 104 | {- pr1 -} 105 | 106 | pr1StrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : S.ft (proj A) ≡ proj Γ) (B : DCtx (suc (suc n))) (B= : S.ft (proj B) ≡ proj A) (u : DMor n (suc n)) (uₛ : S.is-section (proj u)) (u₁ : ∂₁S (proj u) ≡ SigStrS (proj Γ) (proj A) A= (proj B) B=) → DMor n (suc n) 107 | pr1StrS-// Γ A A= B B= u uₛ u₁ = dmorTm Γ (Pr1 (dTy A A=) (dTy+ A= B B=) (dTm refl u uₛ u₁)) 108 | 109 | 110 | pr1StrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : S.ft (proj A) ≡ proj Γ) (A'= : S.ft (proj A') ≡ proj Γ') {B B' : DCtx (suc (suc n))} (rB : B ≃ B') (B= : S.ft (proj B) ≡ (proj A)) (B'= : S.ft (proj B') ≡ proj A') {u u' : DMor n (suc n)} (ru : u ≃ u') (uₛ : S.is-section (proj u)) (u'ₛ : S.is-section (proj u')) (u₁ : ∂₁S (proj u) ≡ SigStrS (proj Γ) (proj A) A= (proj B) B=) (u'₁ : ∂₁S (proj u') ≡ SigStrS (proj Γ') (proj A') A'= (proj B') B'=) → pr1StrS-// Γ A A= B B= u uₛ u₁ ≃ pr1StrS-// Γ' A' A'= B' B'= u' u'ₛ u'₁ 111 | pr1StrS-eq {Γ = Γ} {Γ'} rΓ {A} {A'} rA A= A'= {B} {B'} rB B= B'= {u} {u'} ru uₛ u'ₛ u₁ u'₁ = 112 | dmorTm= dmorTmₛ dmorTmₛ rΓ (dTy= rA A=) (Pr1Cong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=) (dTm= refl ru uₛ u'ₛ u₁ u'₁)) 113 | 114 | 115 | pr1StrS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ SigStrS Γ A A= B B=) → MorS n (suc n) 116 | pr1StrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Ty (λ B B= → //-elim-Tm (λ u uₛ u₁ → proj (pr1StrS-// Γ A A= B B= u uₛ u₁)) 117 | (λ ru uₛ u'ₛ u₁ u'₁ → proj= (pr1StrS-eq (ref Γ) (ref A) A= A= (ref B) B= B= ru uₛ u'ₛ u₁ u'₁))) 118 | (λ rB B= B'= → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (pr1StrS-eq (ref Γ) (ref A) A= A= rB B= B'= (ref u) uₛ uₛ u₁ u₁')))) 119 | (λ rA A= A'= → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (pr1StrS-eq (ref Γ) rA A= A'= (ref B) B= B=' (ref u) uₛ uₛ u₁ u₁'))))) 120 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (pr1StrS-eq rΓ (ref A) A= A=' (ref B) B= B=' (ref u) uₛ uₛ u₁ u₁'))))) 121 | 122 | pr1StrₛS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ SigStrS Γ A A= B B=) → S.is-section (pr1StrS Γ A A= B B= u uₛ u₁) 123 | pr1StrₛS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ → dmorTmₛ)))) 124 | 125 | pr1Str₁S : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ SigStrS Γ A A= B B=) → S.∂₁ (pr1StrS Γ A A= B B= u uₛ u₁) ≡ A 126 | pr1Str₁S = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ → eq (box (CtxTy=Ctx A A=)))))) 127 | 128 | 129 | pr1StrSynCCat : CCatwithpr1 synCCat SigStrSynCCat 130 | CCatwithpr1.pr1Str pr1StrSynCCat = pr1StrS 131 | CCatwithpr1.pr1Strₛ pr1StrSynCCat {Γ = Γ} {A = A} {B = B} {u = u} = pr1StrₛS Γ A _ B _ u _ _ 132 | CCatwithpr1.pr1Str₁ pr1StrSynCCat {Γ = Γ} {A = A} {B = B} {u = u} = pr1Str₁S Γ A _ B _ u _ _ 133 | CCatwithpr1.pr1StrNat' pr1StrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ g₁ → up-to-rhsTyEq (ap (_[_]Ty (getTy A)) (idMor[]Mor (mor g))))))))) 134 | 135 | 136 | {- pr2 -} 137 | 138 | pr2StrS-// : (Γ : DCtx n) (A : DCtx (suc n)) (A= : S.ft (proj A) ≡ proj Γ) (B : DCtx (suc (suc n))) (B= : S.ft (proj B) ≡ proj A) (u : DMor n (suc n)) (uₛ : S.is-section (proj u)) (u₁ : ∂₁S (proj u) ≡ SigStrS (proj Γ) (proj A) A= (proj B) B=) → DMor n (suc n) 139 | pr2StrS-// Γ A A= B B= u uₛ u₁ = dmorTm Γ (Pr2 (dTy A A=) (dTy+ A= B B=) (dTm refl u uₛ u₁)) 140 | 141 | 142 | pr2StrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc n)} (rA : A ≃ A') (A= : S.ft (proj A) ≡ proj Γ) (A'= : S.ft (proj A') ≡ proj Γ') {B B' : DCtx (suc (suc n))} (rB : B ≃ B') (B= : S.ft (proj B) ≡ (proj A)) (B'= : S.ft (proj B') ≡ proj A') {u u' : DMor n (suc n)} (ru : u ≃ u') (uₛ : S.is-section (proj u)) (u'ₛ : S.is-section (proj u')) (u₁ : ∂₁S (proj u) ≡ SigStrS (proj Γ) (proj A) A= (proj B) B=) (u'₁ : ∂₁S (proj u') ≡ SigStrS (proj Γ') (proj A') A'= (proj B') B'=) → pr2StrS-// Γ A A= B B= u uₛ u₁ ≃ pr2StrS-// Γ' A' A'= B' B'= u' u'ₛ u'₁ 143 | pr2StrS-eq {Γ = Γ} {Γ'} rΓ {A} {A'} rA A= A'= {B} {B'} rB B= B'= {u} {u'} ru uₛ u'ₛ u₁ u'₁ = 144 | dmorTm= dmorTmₛ dmorTmₛ rΓ (SubstTyFullEq' (der Γ) ((der Γ) , (dTy A A=)) (dTy+= A= rB B=) (idMor+= (der Γ) (Pr1Cong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=) (dTm= refl ru uₛ u'ₛ u₁ u'₁)))) 145 | (Pr2Cong (dTy A A=) (dTy= rA A=) (dTy+= A= rB B=) (dTm= refl ru uₛ u'ₛ u₁ u'₁)) 146 | 147 | 148 | pr2StrS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ SigStrS Γ A A= B B=) → MorS n (suc n) 149 | pr2StrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Ty (λ B B= → //-elim-Tm (λ u uₛ u₁ → proj (pr2StrS-// Γ A A= B B= u uₛ u₁)) 150 | (λ ru uₛ u'ₛ u₁ u'₁ → proj= (pr2StrS-eq (ref Γ) (ref A) A= A= (ref B) B= B= ru uₛ u'ₛ u₁ u'₁))) 151 | (λ rB B= B'= → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (pr2StrS-eq (ref Γ) (ref A) A= A= rB B= B'= (ref u) uₛ uₛ u₁ u₁')))) 152 | (λ rA A= A'= → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (pr2StrS-eq (ref Γ) rA A= A'= (ref B) B= B=' (ref u) uₛ uₛ u₁ u₁'))))) 153 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Ty (λ B B= B=' → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (pr2StrS-eq rΓ (ref A) A= A=' (ref B) B= B=' (ref u) uₛ uₛ u₁ u₁'))))) 154 | 155 | pr2StrₛS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ SigStrS Γ A A= B B=) → S.is-section (pr2StrS Γ A A= B B= u uₛ u₁) 156 | pr2StrₛS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ → dmorTmₛ)))) 157 | 158 | pr2Str₁S : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ SigStrS Γ A A= B B=) → S.∂₁ (pr2StrS Γ A A= B B= u uₛ u₁) ≡ S.star (pr1StrS Γ A A= B B= u uₛ u₁) B B= (pr1Str₁S Γ A A= B B= u uₛ u₁) 159 | pr2Str₁S = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ → refl)))) 160 | 161 | 162 | pr2StrSynCCat : CCatwithpr2 synCCat SigStrSynCCat pr1StrSynCCat 163 | CCatwithpr2.pr2Str pr2StrSynCCat = pr2StrS 164 | CCatwithpr2.pr2Strₛ pr2StrSynCCat {Γ = Γ} {A = A} {B = B} {u = u} = pr2StrₛS Γ A _ B _ u _ _ 165 | CCatwithpr2.pr2Str₁ pr2StrSynCCat {Γ = Γ} {A = A} {B = B} {u = u} = pr2Str₁S Γ A _ B _ u _ _ 166 | CCatwithpr2.pr2StrNat' pr2StrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ g₁ → up-to-rhsTyEq (ap (_[_]Ty (getTy B [ idMor _ , pr1 (getTy A) (getTy B) (getTm u) ]Ty)) (idMor[]Mor (mor g)) ∙ []Ty-substTy))))))) 167 | 168 | 169 | {- ElSig= -} 170 | 171 | elsigStrS : (i : ℕ) (Γ : ObS n) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ UUStrS i Γ) (b : MorS (suc n) (suc (suc n))) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ UUStrS i (ElStrS i Γ a aₛ a₁)) 172 | → ElStrS i Γ (sigStrS i Γ a aₛ a₁ b bₛ b₁) (sigStrₛS i Γ a aₛ a₁ b bₛ b₁) (sigStr₁S i Γ a aₛ a₁ b bₛ b₁) ≡ SigStrS Γ (ElStrS i Γ a aₛ a₁) (ElStr=S i Γ a aₛ a₁) (ElStrS i (ElStrS i Γ a aₛ a₁) b bₛ b₁) (ElStr=S i (ElStrS i Γ a aₛ a₁) b bₛ b₁) 173 | elsigStrS i = //-elimP (λ Γ → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → eq (box (CtxRefl (der Γ) , ElSig= (dTm refl a aₛ a₁) (dTm refl b bₛ b₁)))))) 174 | 175 | 176 | {- BetaSig1 -} 177 | 178 | betaSig1StrS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ S.star a B B= a₁) → pr1StrS Γ A A= B B= (pairStrS Γ A A= B B= a aₛ a₁ b bₛ b₁) (pairStrₛS Γ A A= B B= a aₛ a₁ b bₛ b₁) (pairStr₁S Γ A A= B B= a aₛ a₁ b bₛ b₁) ≡ a 179 | betaSig1StrS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → 180 | eq (box (CtxSymm (CtxTran (reflectOb (S.is-section₀ aₛ a₁)) (reflectOb A=))) 181 | (CtxTran (CtxTy=Ctx A A=) (CtxSymm (reflectOb a₁))) 182 | (MorTran (der Γ) (der Γ , dTy A A=) (idMor+= (der Γ) (BetaSig1 (dTy A A=) (dTy+ A= B B=) (dTm A= a aₛ a₁) (dTmSubst A= B B= a aₛ a₁ b bₛ b₁))) 183 | (MorSymm (der Γ) (der Γ , dTy A A=) (morTm=idMorTm A= a aₛ a₁))))))))) 184 | 185 | 186 | {- BetaSig2 -} 187 | 188 | betaSig2StrS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (a : MorS n (suc n)) (aₛ : S.is-section a) (a₁ : ∂₁S a ≡ A) (b : MorS n (suc n)) (bₛ : S.is-section b) (b₁ : ∂₁S b ≡ S.star a B B= a₁) → pr2StrS Γ A A= B B= (pairStrS Γ A A= B B= a aₛ a₁ b bₛ b₁) (pairStrₛS Γ A A= B B= a aₛ a₁ b bₛ b₁) (pairStr₁S Γ A A= B B= a aₛ a₁ b bₛ b₁) ≡ b 189 | betaSig2StrS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ a aₛ a₁ → //-elimP (λ b bₛ b₁ → eq (box 190 | (CtxSymm (CtxTran (reflectOb (S.is-section₀ bₛ b₁)) (CtxTran (reflectOb (S.is-section₀ aₛ a₁)) (reflectOb A=)))) 191 | (CtxTran (CtxRefl (der Γ) , SubstTyMorEq (dTy+ A= B B=) 192 | (idMor+ (der Γ) (Pr1 (dTy A A=) (dTy+ A= B B=) (Pair (dTy A A=) (dTy+ A= B B=) (dTm A= a aₛ a₁) (dTmSubst A= B B= a aₛ a₁ b bₛ b₁)))) 193 | (MorTran (der Γ) (der Γ , dTy A A=) (idMor+= (der Γ) (BetaSig1 (dTy A A=) (dTy+ A= B B=) (dTm A= a aₛ a₁) (dTmSubst A= B B= a aₛ a₁ b bₛ b₁))) 194 | (MorSymm (der Γ) (der Γ , dTy A A=) (morTm=idMorTm A= a aₛ a₁)))) 195 | (CtxTran (CtxSymm (CtxTran (reflectOb (S.is-section₀ aₛ a₁)) (reflectOb A=)) , TyRefl (SubstTy (dTy+ A= B B=) (dMor A= a aₛ a₁))) 196 | (CtxSymm (reflectOb b₁)))) 197 | (MorTran (der Γ) (der Γ , SubstTy (dTy+ A= B B=) (idMor+ (der Γ) (Pr1 (dTy A A=) (dTy+ A= B B=) 198 | (Pair (dTy A A=) (dTy+ A= B B=) (dTm A= a aₛ a₁) (dTmSubst A= B B= a aₛ a₁ b bₛ b₁))))) 199 | (idMor+= (der Γ) (ConvEq (SubstTy (dTy+ A= B B=) 200 | (idMor+ (der Γ) (dTm A= a aₛ a₁))) 201 | (BetaSig2 (dTy A A=) (dTy+ A= B B=) (dTm A= a aₛ a₁) (dTmSubst A= B B= a aₛ a₁ b bₛ b₁)) 202 | (SubstTyMorEq (dTy+ A= B B=) (idMor+ (der Γ) (dTm A= a aₛ a₁)) 203 | (idMor+= (der Γ) (TmSymm (BetaSig1 (dTy A A=) (dTy+ A= B B=) (dTm A= a aₛ a₁) (dTmSubst A= B B= a aₛ a₁ b bₛ b₁))))))) 204 | (MorSymm (der Γ) (der Γ , SubstTy (dTy+ A= B B=) (idMor+ (der Γ) (Pr1 (dTy A A=) (dTy+ A= B B=) (Pair (dTy A A=) (dTy+ A= B B=) (dTm A= a aₛ a₁) (dTmSubst A= B B= a aₛ a₁ b bₛ b₁))))) 205 | (ConvMorEq (morTm=idMorTm {Γ = Γ} (eq (box (CtxTran (reflectOb (S.is-section₀ aₛ a₁)) (reflectOb A=)))) b bₛ b₁) 206 | (CtxRefl (der Γ)) (CtxRefl (der Γ) , SubstTyMorEq (dTy+ A= B B=) (dMor A= a aₛ a₁) 207 | (MorTran (der Γ) (der Γ , dTy A A=) 208 | (morTm=idMorTm A= a aₛ a₁) 209 | (idMor+= (der Γ) (TmSymm (BetaSig1 (dTy A A=) 210 | (dTy+ A= B B=) 211 | (dTm A= a aₛ a₁) 212 | (dTmSubst A= B B= a aₛ a₁ b bₛ b₁))))))))))))))) 213 | 214 | {- EtaSig -} 215 | 216 | etaSigStrS : (Γ : ObS n) (A : ObS (suc n)) (A= : S.ft A ≡ Γ) (B : ObS (suc (suc n))) (B= : S.ft B ≡ A) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : ∂₁S u ≡ SigStrS Γ A A= B B=) 217 | → u ≡ pairStrS Γ A A= B B= (pr1StrS Γ A A= B B= u uₛ u₁) (pr1StrₛS Γ A A= B B= u uₛ u₁) (pr1Str₁S Γ A A= B B= u uₛ u₁) (pr2StrS Γ A A= B B= u uₛ u₁) (pr2StrₛS Γ A A= B B= u uₛ u₁) (pr2Str₁S Γ A A= B B= u uₛ u₁) 218 | etaSigStrS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ B B= → //-elimP (λ u uₛ u₁ → eq (box (reflectOb (S.is-section₀ uₛ u₁)) (reflectOb u₁) (ConvMorEq (MorTran (der Γ) (der Γ , Sig (dTy A A=) (dTy+ A= B B=)) (morTm=idMorTm refl u uₛ u₁) (idMor+= (der Γ) (EtaSig (dTy A A=) (dTy+ A= B B=) (dTm refl u uₛ u₁)))) (reflectOb (! (S.is-section₀ uₛ u₁))) (reflectOb (! u₁)))))))) 219 | -------------------------------------------------------------------------------- /termmodel-unit.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import reflection hiding (proj) 6 | open import syntx hiding (getTy) 7 | open import rules 8 | open import contextualcat 9 | open import quotients 10 | open import termmodel-common 11 | open import termmodel-synccat 12 | open import termmodel-uuel 13 | 14 | open CCat hiding (Mor) renaming (id to idC) 15 | 16 | {- Unit -} 17 | 18 | UnitStrS-// : (Γ : DCtx n) → DCtx (suc n) 19 | UnitStrS-// Γ = dctx (der Γ , Unit) 20 | 21 | UnitStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') → UnitStrS-// Γ ≃ UnitStrS-// Γ' 22 | UnitStrS-eq rΓ = box (unOb≃ rΓ , UnitCong) 23 | 24 | UnitStrS : (Γ : ObS n) → ObS (suc n) 25 | UnitStrS = //-elim-Ctx (λ Γ → proj (UnitStrS-// Γ)) 26 | (λ rΓ → proj= (UnitStrS-eq rΓ)) 27 | 28 | UnitStr=S : (Γ : ObS n) → ftS (UnitStrS Γ) ≡ Γ 29 | UnitStr=S = //-elimP-Ctx (λ Γ → refl) 30 | 31 | 32 | UnitStrSynCCat : CCatwithUnit synCCat 33 | CCatwithUnit.UnitStr UnitStrSynCCat = UnitStrS 34 | CCatwithUnit.UnitStr= UnitStrSynCCat = UnitStr=S _ 35 | CCatwithUnit.UnitStrNat' UnitStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ → refl))) 36 | 37 | {- unit -} 38 | 39 | unitStrS-// : (i : ℕ) (Γ : DCtx n) → DMor n (suc n) 40 | unitStrS-// i Γ = dmorTm Γ (UnitUU {i = i}) 41 | 42 | unitStrS-eq : (i : ℕ) {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') → unitStrS-// i Γ ≃ unitStrS-// i Γ' 43 | unitStrS-eq i rΓ = dmorTm= dmorTmₛ dmorTmₛ rΓ UUCong UnitUUCong 44 | 45 | unitStrS : (i : ℕ) (Γ : ObS n) → MorS n (suc n) 46 | unitStrS i = //-elim-Ctx (λ Γ → proj (unitStrS-// i Γ)) 47 | (λ rΓ → proj= (unitStrS-eq i rΓ)) 48 | 49 | unitStrₛS : (i : ℕ) (Γ : ObS n) → S.is-section (unitStrS i Γ) 50 | unitStrₛS i = //-elimP-Ctx (λ Γ → dmorTmₛ) 51 | 52 | unitStr₁S : (i : ℕ) (Γ : ObS n) → S.∂₁ (unitStrS i Γ) ≡ UUStrS i Γ 53 | unitStr₁S i = //-elimP-Ctx (λ Γ → refl) 54 | 55 | unitStrSynCCat : CCatwithunit synCCat UUStrSynCCat 56 | CCatwithunit.unitStr unitStrSynCCat = unitStrS 57 | CCatwithunit.unitStrₛ unitStrSynCCat {Γ = Γ} = unitStrₛS _ Γ 58 | CCatwithunit.unitStr₁ unitStrSynCCat {Γ = Γ} = unitStr₁S _ Γ 59 | CCatwithunit.unitStrNat' unitStrSynCCat = //-elimP (λ g → JforNat (//-elimP-Ctx (λ Γ g₁ → refl))) 60 | 61 | {- tt -} 62 | 63 | ttStrS-// : (Γ : DCtx n) → DMor n (suc n) 64 | ttStrS-// Γ = dmorTm Γ TT 65 | 66 | ttStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') → ttStrS-// Γ ≃ ttStrS-// Γ' 67 | ttStrS-eq rΓ = dmorTm= dmorTmₛ dmorTmₛ rΓ UnitCong TTCong 68 | 69 | ttStrS : (Γ : ObS n) → MorS n (suc n) 70 | ttStrS = //-elim-Ctx (λ Γ → proj (ttStrS-// Γ)) (λ rΓ → proj= (ttStrS-eq rΓ)) 71 | 72 | ttStrₛS : (Γ : ObS n) → S.is-section (ttStrS Γ) 73 | ttStrₛS = //-elimP (λ Γ → dmorTmₛ) 74 | 75 | ttStr₁S : (Γ : ObS n) → S.∂₁ (ttStrS Γ) ≡ UnitStrS Γ 76 | ttStr₁S = //-elimP-Ctx (λ Γ → refl) 77 | 78 | ttStrSynCCat : CCatwithtt synCCat UnitStrSynCCat 79 | CCatwithtt.ttStr ttStrSynCCat = ttStrS 80 | CCatwithtt.ttStrₛ ttStrSynCCat {Γ = Γ} = ttStrₛS Γ 81 | CCatwithtt.ttStr₁ ttStrSynCCat {Γ = Γ} = ttStr₁S Γ 82 | CCatwithtt.ttStrNat' ttStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ g₁ → refl))) 83 | 84 | 85 | {- unitelim -} 86 | 87 | unitelimStrS-// : (Γ : DCtx n) (A : DCtx (suc (suc n))) (A= : ftS (proj A) ≡ UnitStrS (proj Γ)) (dtt : DMor n (suc n)) (dttₛ : S.is-section (proj dtt)) (dtt₁ : S.∂₁ (proj dtt) ≡ S.star (ttStrS (proj Γ)) (proj A) A= (ttStr₁S (proj Γ))) (u : DMor n (suc n)) (uₛ : S.is-section (proj u)) (u₁ : S.∂₁ (proj u) ≡ UnitStrS (proj Γ)) → DMor n (suc n) 88 | unitelimStrS-// Γ A A= dtt dttₛ dtt₁ u uₛ u₁ = dmorTm Γ (Unitelim (dTy A A=) (dTm refl dtt dttₛ dtt₁) (dTm refl u uₛ u₁)) 89 | 90 | unitelimStrS-eq : {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {A A' : DCtx (suc (suc n))} (rA : A ≃ A') (A= : ftS (proj A) ≡ UnitStrS (proj Γ)) (A'= : ftS (proj A') ≡ UnitStrS (proj Γ')) {dtt dtt' : DMor n (suc n)} (rdtt : dtt ≃ dtt') (dttₛ : S.is-section (proj dtt)) (dtt'ₛ : S.is-section (proj dtt')) (dtt₁ : S.∂₁ (proj dtt) ≡ S.star (ttStrS (proj Γ)) (proj A) A= (ttStr₁S (proj Γ))) (dtt'₁ : S.∂₁ (proj dtt') ≡ S.star (ttStrS (proj Γ')) (proj A') A'= (ttStr₁S (proj Γ'))) {u u' : DMor n (suc n)} (ru : u ≃ u') (uₛ : S.is-section (proj u)) (u'ₛ : S.is-section (proj u')) (u₁ : S.∂₁ (proj u) ≡ UnitStrS (proj Γ)) (u'₁ : S.∂₁ (proj u') ≡ UnitStrS (proj Γ')) → unitelimStrS-// Γ A A= dtt dttₛ dtt₁ u uₛ u₁ ≃ unitelimStrS-// Γ' A' A'= dtt' dtt'ₛ dtt'₁ u' u'ₛ u'₁ 91 | unitelimStrS-eq {Γ = Γ} {Γ'} rΓ rA A= A'= rdtt dttₛ dtt'ₛ dtt₁ dtt'₁ ru uₛ u'ₛ u₁ u'₁ = dmorTm= dmorTmₛ dmorTmₛ rΓ 92 | (SubstTyFullEq' (der Γ) (der Γ , Unit) 93 | (dTy= rA A=) 94 | (idMor+= (der Γ) (dTm= refl ru uₛ u'ₛ u₁ u'₁))) 95 | (UnitelimCong (dTy= rA A=) (dTm= refl rdtt dttₛ dtt'ₛ dtt₁ dtt'₁) (dTm= refl ru uₛ u'ₛ u₁ u'₁)) 96 | 97 | unitelimStrS : (Γ : ObS n) (A : ObS (suc (suc n))) (A= : ftS A ≡ UnitStrS Γ) (dtt : MorS n (suc n)) (dttₛ : S.is-section dtt) (dtt₁ : S.∂₁ dtt ≡ S.star (ttStrS Γ) A A= (ttStr₁S Γ)) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : S.∂₁ u ≡ UnitStrS Γ) → MorS n (suc n) 98 | unitelimStrS = //-elim-Ctx (λ Γ → //-elim-Ty (λ A A= → //-elim-Tm (λ dtt dttₛ dtt₁ → (//-elim-Tm (λ u uₛ u₁ → proj (unitelimStrS-// Γ A A= dtt dttₛ dtt₁ u uₛ u₁)) 99 | (λ ru uₛ u'ₛ u₁ u'₁ → proj= (unitelimStrS-eq (ref Γ) (ref A) A= A= (ref dtt) dttₛ dttₛ dtt₁ dtt₁ ru uₛ u'ₛ u₁ u'₁)))) 100 | (λ rdtt dttₛ dtt'ₛ dtt₁ dtt'₁ → //-elimP-Tm (λ u uₛ u₁ u₁' → proj= (unitelimStrS-eq (ref Γ) (ref A) A= A= rdtt dttₛ dtt'ₛ dtt₁ dtt'₁ (ref u) uₛ uₛ u₁ u₁')))) 101 | (λ rA A= A'= → //-elimP-Tm (λ dtt dttₛ dtt₁ dtt₁' → (//-elimP-Tm (λ u uₛ u₁ u₁' → proj= (unitelimStrS-eq (ref Γ) rA A= A'= (ref dtt) dttₛ dttₛ dtt₁ dtt₁' (ref u) uₛ uₛ u₁ u₁')))))) 102 | (λ rΓ → //-elimP-Ty (λ A A= A=' → //-elimP-Tm (λ dtt dttₛ dtt₁ dtt₁' → (//-elimP-Tm (λ u uₛ u₁ u₁' → proj= (unitelimStrS-eq rΓ (ref A) A= A=' (ref dtt) dttₛ dttₛ dtt₁ dtt₁' (ref u) uₛ uₛ u₁ u₁')))))) 103 | 104 | unitelimStrₛS : (Γ : ObS n) (A : ObS (suc (suc n))) (A= : ftS A ≡ UnitStrS Γ) (dtt : MorS n (suc n)) (dttₛ : S.is-section dtt) (dtt₁ : S.∂₁ dtt ≡ S.star (ttStrS Γ) A A= (ttStr₁S Γ)) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : S.∂₁ u ≡ UnitStrS Γ) → S.is-section (unitelimStrS Γ A A= dtt dttₛ dtt₁ u uₛ u₁) 105 | unitelimStrₛS = //-elimP-Ctx (λ Γ → //-elimP (λ A A= → //-elimP (λ dtt dttₛ dtt₁ → //-elimP (λ u uₛ u₁ → dmorTmₛ)))) 106 | 107 | unitelimStr₁S : (Γ : ObS n) (A : ObS (suc (suc n))) (A= : ftS A ≡ UnitStrS Γ) (dtt : MorS n (suc n)) (dttₛ : S.is-section dtt) (dtt₁ : S.∂₁ dtt ≡ S.star (ttStrS Γ) A A= (ttStr₁S Γ)) (u : MorS n (suc n)) (uₛ : S.is-section u) (u₁ : S.∂₁ u ≡ UnitStrS Γ) → S.∂₁ (unitelimStrS Γ A A= dtt dttₛ dtt₁ u uₛ u₁) ≡ S.star u A A= u₁ 108 | unitelimStr₁S = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ dtt dttₛ dtt₁ → //-elimP (λ u uₛ u₁ → eq (box (CtxSymm (reflectOb (S.is-section₀ uₛ u₁)) , SubstTyMorEq (dTy A A=) (idMor+ (der Γ) (dTm refl u uₛ u₁)) (MorSymm (der Γ) (der Γ , Unit) (morTm=idMorTm refl u uₛ u₁)))))))) 109 | 110 | unitelimStrSynCCat : CCatwithunitelim synCCat UnitStrSynCCat ttStrSynCCat 111 | CCatwithunitelim.unitelimStr unitelimStrSynCCat = unitelimStrS 112 | CCatwithunitelim.unitelimStrₛ unitelimStrSynCCat {Γ = Γ} {A = A} {dtt = dtt} {u = u} = unitelimStrₛS Γ A _ dtt _ _ u _ _ 113 | CCatwithunitelim.unitelimStr₁ unitelimStrSynCCat {Γ = Γ} {A = A} {dtt = dtt} {u = u} = unitelimStr₁S Γ A _ dtt _ _ u _ _ 114 | CCatwithunitelim.unitelimStrNat' unitelimStrSynCCat = //-elimP (λ g → JforNat (//-elimP-Ctx (λ Γ → //-elimP (λ A A= → //-elimP (λ dtt dttₛ dtt₁ → //-elimP (λ u uₛ u₁ g₁ → up-to-rhsTyEq (ap (_[_]Ty (substTy (getTy A) (getTm u))) (idMor[]Mor (mor g)) ∙ []Ty-substTy))))))) 115 | 116 | {- ElUnit= -} 117 | 118 | elunitStrS : (i : ℕ) (Γ : ObS n) → ElStrS i Γ (unitStrS i Γ) (unitStrₛS i Γ) (unitStr₁S i Γ) ≡ UnitStrS Γ 119 | elunitStrS i = //-elimP (λ Γ → eq (box (CtxRefl (der Γ) , ElUnit=))) 120 | 121 | {- BetaUnit -} 122 | 123 | betaUnitStrS : (Γ : ObS n) (A : ObS (suc (suc n))) (A= : ftS A ≡ UnitStrS Γ) (dtt : MorS n (suc n)) (dttₛ : S.is-section dtt) (dtt₁ : S.∂₁ dtt ≡ S.star (ttStrS Γ) A A= (ttStr₁S Γ)) 124 | → unitelimStrS Γ A A= dtt dttₛ dtt₁ (ttStrS Γ) (ttStrₛS Γ) (ttStr₁S Γ) ≡ dtt 125 | betaUnitStrS = //-elimP (λ Γ → //-elimP (λ A A= → //-elimP (λ dtt dttₛ dtt₁ → eq (box (CtxSymm (reflectOb (S.is-section₀ dttₛ dtt₁))) (CtxSymm (reflectOb dtt₁)) (MorTran (der Γ) (der Γ , SubstTy (dTy A A=) (idMor+ (der Γ) TT)) (idMor+= (der Γ) (BetaUnit (dTy A A=) (dTm refl dtt dttₛ dtt₁))) (MorSymm (der Γ) (der Γ , SubstTy (dTy A A=) (idMor+ (der Γ) TT)) (morTm=idMorTm refl dtt dttₛ dtt₁))))))) 126 | -------------------------------------------------------------------------------- /termmodel-uuel.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import syntx 6 | open import rules 7 | open import contextualcat 8 | open import quotients 9 | open import termmodel-common 10 | open import termmodel-synccat 11 | 12 | open CCat hiding (Mor) renaming (id to idC) 13 | 14 | 15 | {- UU -} 16 | 17 | UUStrS-// : (i : ℕ) → DCtx n → DCtx (suc n) 18 | UUStrS-// i Γ = dctx (der Γ , UU {i = i}) 19 | 20 | UUStrS-eq : {i : ℕ} {Γ Γ' : DCtx n} → Γ ≃ Γ' → UUStrS-// i Γ ≃ UUStrS-// i Γ' 21 | UUStrS-eq dΓ= = box (unOb≃ dΓ= , UUCong) 22 | 23 | UUStrS : (i : ℕ) → ObS n → ObS (suc n) 24 | UUStrS i = //-elim-Ctx (λ Γ → proj (UUStrS-// i Γ)) (λ rΓ → proj= (UUStrS-eq rΓ)) 25 | 26 | UUStr=S : (i : ℕ) (Γ : ObS n) → ftS (UUStrS i Γ) ≡ Γ 27 | UUStr=S i = //-elimP-Ctx (λ Γ → refl) 28 | 29 | UUStrSynCCat : CCatwithUU synCCat 30 | CCatwithUU.UUStr UUStrSynCCat = UUStrS 31 | CCatwithUU.UUStr= UUStrSynCCat = UUStr=S _ _ 32 | CCatwithUU.UUStrNat' UUStrSynCCat = //-elimP (λ g → JforNat (//-elimP (λ Γ g₁ → refl))) 33 | 34 | 35 | {- El -} 36 | 37 | ElStrS-// : (i : ℕ) (Γ : DCtx n) (v : DMor n (suc n)) (vₛ : S.is-section (proj v)) (v₁ : ∂₁S (proj v) ≡ UUStrS i (proj Γ)) → DCtx (suc n) 38 | ElStrS-// i Γ v vₛ v₁ = dctx {ctx = _ , _}(der Γ , El (dTm refl v vₛ v₁)) 39 | 40 | ElStrS-eq : {i : ℕ} {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') {v v' : DMor n (suc n)} (rv : v ≃ v') (vₛ : _) (v'ₛ : _) (v₁ : _) (v'₁ : _) → ElStrS-// i Γ v vₛ v₁ ≃ ElStrS-// i Γ' v' v'ₛ v'₁ 41 | ElStrS-eq rΓ rv vₛ v'ₛ v₁ v'₁ = box (unOb≃ rΓ , ElCong (dTm= refl rv vₛ v'ₛ v₁ v'₁)) 42 | 43 | ElStrS : (i : ℕ) (Γ : ObS n) (v : MorS n (suc n)) (vₛ : S.is-section v) (v₁ : ∂₁S v ≡ UUStrS i Γ) → ObS (suc n) 44 | ElStrS i = //-elim-Ctx (λ Γ → //-elim-Tm (λ v vₛ v₁ → proj (ElStrS-// i Γ v vₛ v₁)) 45 | (λ rv vₛ v'ₛ v₁ v'₁ → proj= (ElStrS-eq (ref Γ) rv vₛ v'ₛ v₁ v'₁))) 46 | (λ rΓ → //-elimP-Tm (λ v vₛ v₁ v₁' → proj= (ElStrS-eq rΓ (ref v) vₛ vₛ v₁ v₁'))) 47 | 48 | ElStr=S : (i : ℕ) (Γ : ObS n) (v : MorS n (suc n)) (vₛ : S.is-section v) (v₁ : ∂₁S v ≡ UUStrS i Γ) → ftS (ElStrS i Γ v vₛ v₁) ≡ Γ 49 | ElStr=S i = //-elimP-Ctx (λ Γ → //-elimP (λ v vₛ v₁ → refl)) 50 | 51 | ElStrSynCCat : CCatwithEl synCCat UUStrSynCCat 52 | CCatwithEl.ElStr ElStrSynCCat i Γ v vₛ v₁ = ElStrS i Γ v vₛ v₁ 53 | CCatwithEl.ElStr= ElStrSynCCat = ElStr=S _ _ _ _ _ 54 | CCatwithEl.ElStrNat' ElStrSynCCat = //-elimP (λ g → //-elimP (λ Δ g₀ → //-elimP (λ Γ → //-elimP (λ v vₛ v₁ g₁ → up-to-CtxTyEq (reflectOb g₀) refl)))) 55 | 56 | 57 | {- uu -} 58 | 59 | uuStrS-// : (i : ℕ) (Γ : DCtx n) → DMor n (suc n) 60 | uuStrS-// i Γ = dmorTm Γ (UUUU {i = i}) 61 | 62 | uuStrS-eq : (i : ℕ) {Γ Γ' : DCtx n} (rΓ : Γ ≃ Γ') → uuStrS-// i Γ ≃ uuStrS-// i Γ' 63 | uuStrS-eq i rΓ = dmorTm= dmorTmₛ dmorTmₛ rΓ UUCong UUUUCong 64 | 65 | uuStrS : (i : ℕ) (Γ : ObS n) → MorS n (suc n) 66 | uuStrS i = //-elim-Ctx (λ Γ → proj (uuStrS-// i Γ)) (λ rΓ → proj= (uuStrS-eq i rΓ)) 67 | 68 | uuStrₛS : (i : ℕ) (Γ : ObS n) → S.is-section (uuStrS i Γ) 69 | uuStrₛS i = //-elimP (λ Γ → dmorTmₛ) 70 | 71 | uuStr₁S : (i : ℕ) (Γ : ObS n) → ∂₁S (uuStrS i Γ) ≡ UUStrS (suc i) Γ 72 | uuStr₁S i = //-elimP (λ Γ → refl) 73 | 74 | uuStrSynCCat : CCatwithuu synCCat UUStrSynCCat 75 | CCatwithuu.uuStr uuStrSynCCat = uuStrS 76 | CCatwithuu.uuStrₛ uuStrSynCCat {Γ = Γ} = uuStrₛS _ Γ 77 | CCatwithuu.uuStr₁ uuStrSynCCat {Γ = Γ} = uuStr₁S _ Γ 78 | CCatwithuu.uuStrNat' uuStrSynCCat = //-elimP (λ g → //-elimP(λ Δ g₀ → (//-elimP (λ Γ g₁ → up-to-rhsTyEq' (reflectOb g₀) refl)))) 79 | 80 | 81 | {- ElUU= -} 82 | 83 | eluuStrS : (i : ℕ) (Γ : ObS n) → ElStrS (suc i) Γ (uuStrS i Γ) (uuStrₛS i Γ) (uuStr₁S i Γ) ≡ UUStrS i Γ 84 | eluuStrS i = //-elimP (λ Γ → eq (box (CtxRefl (der Γ) , ElUU=))) 85 | 86 | -------------------------------------------------------------------------------- /termmodel.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | open import typetheory 5 | open import syntx 6 | open import rules 7 | open import contextualcat 8 | open import quotients 9 | 10 | 11 | open import termmodel-common public 12 | open import termmodel-synccat public 13 | open import termmodel-uuel public 14 | open import termmodel-sum public 15 | open import termmodel-pi public 16 | open import termmodel-sig public 17 | open import termmodel-empty public 18 | open import termmodel-unit public 19 | open import termmodel-nat public 20 | open import termmodel-id public 21 | open import termmodel-id2 public 22 | 23 | open CCat hiding (Mor) renaming (id to idC) 24 | 25 | open StructuredCCat 26 | 27 | strSynCCat : StructuredCCat 28 | 29 | ccat strSynCCat = synCCat 30 | 31 | ccatUU strSynCCat = UUStrSynCCat 32 | ccatEl strSynCCat = ElStrSynCCat 33 | ccatSum strSynCCat = SumStrSynCCat 34 | ccatPi strSynCCat = PiStrSynCCat 35 | ccatSig strSynCCat = SigStrSynCCat 36 | ccatEmpty strSynCCat = EmptyStrSynCCat 37 | ccatUnit strSynCCat = UnitStrSynCCat 38 | ccatNat strSynCCat = NatStrSynCCat 39 | ccatId strSynCCat = IdStrSynCCat 40 | 41 | ccatuu strSynCCat = uuStrSynCCat 42 | ccatsum strSynCCat = sumStrSynCCat 43 | ccatinl strSynCCat = inlStrSynCCat 44 | ccatinr strSynCCat = inrStrSynCCat 45 | ccatmatch strSynCCat = matchStrSynCCat 46 | ccatpi strSynCCat = piStrSynCCat 47 | ccatlam strSynCCat = lamStrSynCCat 48 | ccatapp strSynCCat = appStrSynCCat 49 | ccatsig strSynCCat = sigStrSynCCat 50 | ccatpair strSynCCat = pairStrSynCCat 51 | ccatpr1 strSynCCat = pr1StrSynCCat 52 | ccatpr2 strSynCCat = pr2StrSynCCat 53 | ccatempty strSynCCat = emptyStrSynCCat 54 | ccatemptyelim strSynCCat = emptyelimStrSynCCat 55 | ccatunit strSynCCat = unitStrSynCCat 56 | ccattt strSynCCat = ttStrSynCCat 57 | ccatunitelim strSynCCat = unitelimStrSynCCat 58 | ccatnat strSynCCat = natStrSynCCat 59 | ccatzero strSynCCat = zeroStrSynCCat 60 | ccatsuc strSynCCat = sucStrSynCCat 61 | ccatnatelim strSynCCat = natelimStrSynCCat 62 | ccatid strSynCCat = idStrSynCCat 63 | ccatrefl strSynCCat = reflStrSynCCat 64 | ccatjj strSynCCat = jjStrSynCCat 65 | 66 | betaInlStr strSynCCat {Γ = Γ} {A = A} {B = B} {C = C} {da = da} {db = db} {a = a} = betaInlStrS Γ A _ B _ C _ da _ _ db _ _ a _ _ 67 | betaInrStr strSynCCat {Γ = Γ} {A = A} {B = B} {C = C} {da = da} {db = db} {b = b} = betaInrStrS Γ A _ B _ C _ da _ _ db _ _ b _ _ 68 | betaPiStr strSynCCat {Γ = Γ} {A = A} {B = B} {u = u} {a = a} = betaPiStrS Γ A _ B _ u _ _ a _ _ 69 | betaSig1Str strSynCCat {Γ = Γ} {A = A} {B = B} {a = a} {b = b} = betaSig1StrS Γ A _ B _ a _ _ b _ _ 70 | betaSig2Str strSynCCat {Γ = Γ} {A = A} {B = B} {a = a} {b = b} = betaSig2StrS Γ A _ B _ a _ _ b _ _ 71 | betaUnitStr strSynCCat {Γ = Γ} {A = A} {dtt = dtt} = betaUnitStrS Γ A _ dtt _ _ 72 | betaNatZeroStr strSynCCat {Γ = Γ} {P = P} {dO = dO} {dS = dS} = betaNatZeroStrS Γ P _ dO _ _ dS _ _ 73 | betaNatSucStr strSynCCat {Γ = Γ} {P = P} {dO = dO} {dS = dS} {u = u} = betaNatSucStrS Γ P _ dO _ _ dS _ _ u _ _ 74 | betaIdStr strSynCCat {Γ = Γ} {A = A} {P = P} {d = d} {a = a} = betaIdStrS Γ A _ P _ d _ _ a _ _ 75 | 76 | etaSumStr strSynCCat {Γ = Γ} {A = A} {B = B} {u = u} = etaSumStrS Γ A _ B _ u _ _ 77 | etaPiStr strSynCCat {Γ = Γ} {A = A} {B = B} {f = f} {fₛ = fₛ} {f₁ = f₁} = etaPiStrS Γ A _ B _ f fₛ f₁ 78 | etaSigStr strSynCCat {Γ = Γ} {A = A} {B = B} {u = u} = etaSigStrS Γ A _ B _ u _ _ 79 | 80 | eluuStr strSynCCat {Γ = Γ} = eluuStrS _ Γ 81 | elsumStr strSynCCat {Γ = Γ} {a = a} {b = b} = elsumStrS _ Γ a _ _ b _ _ 82 | elpiStr strSynCCat {Γ = Γ} {a = a} {b = b} = elpiStrS _ Γ a _ _ b _ _ 83 | elsigStr strSynCCat {Γ = Γ} {a = a} {b = b} = elsigStrS _ Γ a _ _ b _ _ 84 | elemptyStr strSynCCat {Γ = Γ} = elemptyStrS _ Γ 85 | elunitStr strSynCCat {Γ = Γ} = elunitStrS _ Γ 86 | elnatStr strSynCCat {Γ = Γ} = elnatStrS _ Γ 87 | elidStr strSynCCat {Γ = Γ} {a = a} {u = u} {v = v} = elidStrS _ Γ a _ _ u _ _ v _ _ 88 | 89 | -------------------------------------------------------------------------------- /typetheory.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting --prop #-} 2 | 3 | open import common 4 | 5 | {- Syntax of term- and type-expressions, using de Bruijn indices -} 6 | 7 | data TyExpr : ℕ → Set 8 | data TmExpr : ℕ → Set 9 | 10 | data TyExpr where 11 | uu : (i : ℕ) → TyExpr n 12 | el : (i : ℕ) (v : TmExpr n) → TyExpr n 13 | sum : (A : TyExpr n) (B : TyExpr n) → TyExpr n 14 | pi : (A : TyExpr n) (B : TyExpr (suc n)) → TyExpr n 15 | sig : (A : TyExpr n) (B : TyExpr (suc n)) → TyExpr n 16 | empty : TyExpr n 17 | unit : TyExpr n 18 | nat : TyExpr n 19 | id : (A : TyExpr n) (u v : TmExpr n) → TyExpr n 20 | 21 | 22 | data TmExpr where 23 | var : (x : VarPos n) → TmExpr n 24 | 25 | uu : (i : ℕ) → TmExpr n 26 | 27 | sum : (i : ℕ) (a : TmExpr n) (b : TmExpr n) → TmExpr n 28 | inl : (A : TyExpr n) (B : TyExpr n) (a : TmExpr n) → TmExpr n 29 | inr : (A : TyExpr n) (B : TyExpr n) (b : TmExpr n) → TmExpr n 30 | match : (A : TyExpr n) (B : TyExpr n) (C : TyExpr (suc n)) (da : TmExpr (suc n)) (db : TmExpr (suc n)) (u : TmExpr n) → TmExpr n 31 | 32 | pi : (i : ℕ) (a : TmExpr n) (b : TmExpr (suc n)) → TmExpr n 33 | lam : (A : TyExpr n) (B : TyExpr (suc n)) (u : TmExpr (suc n)) → TmExpr n 34 | app : (A : TyExpr n) (B : TyExpr (suc n)) (f : TmExpr n) (a : TmExpr n) → TmExpr n 35 | 36 | sig : (i : ℕ) (a : TmExpr n) (b : TmExpr (suc n)) → TmExpr n 37 | pair : (A : TyExpr n) (B : TyExpr (suc n)) (a : TmExpr n) (b : TmExpr n) → TmExpr n 38 | pr1 : (A : TyExpr n) (B : TyExpr (suc n)) (u : TmExpr n) → TmExpr n 39 | pr2 : (A : TyExpr n) (B : TyExpr (suc n)) (u : TmExpr n) → TmExpr n 40 | 41 | empty : (i : ℕ) → TmExpr n 42 | emptyelim : (A : TyExpr (suc n)) (u : TmExpr n) → TmExpr n 43 | 44 | unit : (i : ℕ) → TmExpr n 45 | tt : TmExpr n 46 | unitelim : (A : TyExpr (suc n)) (dtt : TmExpr n) (u : TmExpr n) → TmExpr n 47 | 48 | nat : (i : ℕ) → TmExpr n 49 | zero : TmExpr n 50 | suc : (u : TmExpr n) → TmExpr n 51 | natelim : (P : TyExpr (suc n)) (d0 : TmExpr n) (dS : TmExpr (suc (suc n))) (u : TmExpr n) → TmExpr n 52 | 53 | id : (i : ℕ) (a u v : TmExpr n) → TmExpr n 54 | refl : (A : TyExpr n) (u : TmExpr n) → TmExpr n 55 | jj : (A : TyExpr n) (P : TyExpr (suc (suc (suc n)))) (d : TmExpr (suc n)) (a b p : TmExpr n) → TmExpr n 56 | --------------------------------------------------------------------------------