├── .gitignore ├── LICENSE ├── README.md ├── formaltt.agda-lib └── src ├── Experimental ├── AlgebraFormalization.agda ├── Category.agda ├── Equality.agda ├── Finite.agda ├── Lambda.agda ├── ManySortedAlgebra.agda ├── MultiCategories.agda ├── Quotient.agda └── SimpleTT.agda ├── MultiSorted ├── AlgebraicTheory.agda ├── Completeness.agda ├── Context.agda ├── FromSingleSorted.agda ├── Group.agda ├── Interpretation.agda ├── InterpretationCategory.agda ├── Model.agda ├── ModelCategory.agda ├── Product.agda ├── Substitution.agda ├── SyntacticCategory.agda ├── UniversalInterpretation.agda └── UniversalModel.agda ├── SecondOrder ├── Arity.agda ├── Equality.agda ├── IndexedCategory.agda ├── Instantiation.agda ├── MContext.agda ├── MRelMonMorphism.agda ├── MRelativeMonad.agda ├── MRenaming.agda ├── Metavariable.agda ├── Mslot.agda ├── README.md ├── RelMon.agda ├── RelativeKleisli.agda ├── RelativeMonadMorphism.agda ├── Signature.agda ├── Substitution.agda ├── Term.agda ├── Theory.agda ├── VContext.agda ├── VRelMonMorphism.agda ├── VRelativeMonad.agda └── VRenaming.agda └── SingleSorted ├── AlgebraicTheory.agda ├── Combinators.agda ├── Context.agda ├── Example.agda ├── Group.agda ├── Interpretation.agda ├── Model.agda ├── Power.agda ├── Substitution.agda ├── SyntacticCategory.agda ├── UniversalInterpretation.agda └── UniversalModel.agda /.gitignore: -------------------------------------------------------------------------------- 1 | #* 2 | .#* 3 | *.agdai 4 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | MIT License 2 | 3 | Copyright (c) 2021 Gaïa Loutchmia 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in all 13 | copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE 21 | SOFTWARE. 22 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Formalization of simple type theory 2 | 3 | ## Coding standards 4 | 5 | We collect here some coding standards. 6 | 7 | #### `import` and `open` 8 | 9 | Only `import` what is necessary. Avoid global `open` and prefer local `open` statements. 10 | 11 | #### Standard and categories libraries 12 | 13 | Use the standard and categories libraries instead of reinventing the wheel. 14 | 15 | #### Indentation level and line length 16 | 17 | Use indentation level 2. Instead of 18 | 19 | ``` 20 | Functor-Jⱽ.F-resp-≈ (Monad.F₀ VMonad Γ A) {ψ} {Ω} {Δ} {Ξ} {ρ} {τ} {ι} {μ} ξᵛʳ ξᵐʳ {t} {u} ξ = 21 | begin⟨ Term-setoid Θ _ _ ⟩ 22 | ([ [ (λ x₁ → var-inl x₁) , (λ x₁ → var-inr (ρ x₁)) ]ᵛʳ ]ᵛʳ t) ≈⟨ []ᵛʳ-resp-≈ ξ ⟩ 23 | ([ [ (λ x₁ → var-inl x₁) , (λ x₁ → var-inr (ρ x₁)) ]ᵛʳ ]ᵛʳ u) 24 | ≈⟨ []ᵛʳ-resp-≡ᵛʳ ([,]ᵛʳ-resp-≡ᵛʳ (λ x → refl) λ x → ρ-resp-≡ {ρ = var-inr} (ξᵛʳ x)) ⟩ 25 | ([ [ (λ x₁ → var-inl x₁) , (λ x₁ → var-inr (τ x₁)) ]ᵛʳ ]ᵛʳ u) ∎ 26 | ``` 27 | 28 | you should write 29 | 30 | ``` 31 | Functor-Jⱽ.F-resp-≈ (Monad.F₀ VMonad Γ A) {ψ} {Ω} {Δ} {Ξ} {ρ} {τ} {ι} {μ} ξᵛʳ ξᵐʳ {t} {u} ξ = 32 | begin⟨ Term-setoid Θ _ _ ⟩ 33 | ([ [ (λ x₁ → var-inl x₁) , (λ x₁ → var-inr (ρ x₁)) ]ᵛʳ ]ᵛʳ t) 34 | ≈⟨ []ᵛʳ-resp-≈ ξ ⟩ 35 | ([ [ (λ x₁ → var-inl x₁) , (λ x₁ → var-inr (ρ x₁)) ]ᵛʳ ]ᵛʳ u) 36 | ≈⟨ []ᵛʳ-resp-≡ᵛʳ ([,]ᵛʳ-resp-≡ᵛʳ (λ x → refl) λ x → ρ-resp-≡ {ρ = var-inr} (ξᵛʳ x)) ⟩ 37 | ([ [ (λ x₁ → var-inl x₁) , (λ x₁ → var-inr (τ x₁)) ]ᵛʳ ]ᵛʳ u) ∎ 38 | ``` 39 | 40 | or something similar that doesn't produce exceedingly long lines with unecessary indentation. 41 | 42 | 43 | #### Naming conventions 44 | 45 | 1. Use full names: `RelativeMonad` instead of `RelMon`, `RelativeMorphism` instead of `RelMorph`. Abbreviations should be use very sparingly. 46 | 2. We use subscripts to indicate entities, as follows: 47 | * `ᵛ` for *variable renaming* 48 | * `ᵐ` for *metavariable renaming* 49 | * `ˢ` for *substitition* 50 | * `ⁱ` for *instantiation* 51 | 3. Composition of entities is written as `∘ˣ` where `x` is a supscript indicating the kind, e.g., 52 | `σ ∘ˢ τ` is composition of subtitutions. There are also mixed compositions `ʸ∘ˣ` which compose 53 | an entity of kind `y` with an entitiy of kind `x`. 54 | 4. The action of `f` on a term `t` is written as `[ f ]ˣ t` where `x` is a supscript indicating the kind of `f`. 55 | For example, `[ σ ]ˢ t` is the action of the substitution `σ` on term `t`. 56 | 5. A theorem explaining how an action interacts with composition are named `[∘ˣ]` or `[ˣ∘ʸ]`. 57 | 6. A theorem stating that an action `[]ˣ` respects equality `≈ʸ` are named `[]ˣ-resp-≈ʸ`. 58 | 59 | 60 | ## The categorical structure of second-order syntax 61 | 62 | Given functors `I F : ℂ → 𝔻`, let `RelativeMonad I F` be the category whose objects are 63 | relative monad structures on `(I, F)` and the morphisms are relative monad morphisms between them. 64 | 65 | The ingridients: 66 | * A set of types `Ty`. 67 | * A category `VCtx` of variable contexts and renamings. The category has finite coproducts. 68 | * A category `MCtx` of metavariable contexts and renamings. The category has finite coproducts. 69 | * A coproduct-preserving functor `Jᵛ : VCtx → Set^Ty` where `Jᵛ Γ A` is thought of as variables of type `A` in context `Γ` 70 | * A coproduct-preserving functor `Jᵐ : VCtx → Set^(VCtx × Ty)`, where `Jᵐ Θ (Δ, A)` is thought of as metavariables of arity `Δ` and type `A` in metacontext `Θ`. 71 | * A functor `Term : MCtx → VCtx → Set^Ty`, where `Term Θ Γ A` is thought of as the set of terms of type `A` in metacontext `Θ` and context `Γ`. 72 | * A functor `MCtx → RelativeMonad Jᵛ Tᵛ` where `Tᵛ : MCtx → VCtx → Set^Ty` is `Tᵛ Θ Γ A := Term Θ Γ A`. 73 | * A functor `VCtx → RelativeMonad Jᵐ Tᵐ` where `Tᵐ : VCtx → MCtx → Set^(VCtx × Ty)` is `Tᵐ Θ Γ (Δ, A) = Term Θ (Γ + Δ) A`. 74 | * For every `f : Jⱽ Γ → Tᵛ Θ Γ'` we can define `hat f : Tᵐ Γ → Tᵐ Γ'` and it is a `Jᵐ`-relative monad morphism 75 | * For every `I : Jᵐ Θ → Tᵐ Θ'` we can define `hat I : Tᵛ Θ → Tᵛ Θ'` and it is a `Jᵛ`-relative monad morphism 76 | 77 | The syntactic structure is recovered from the above category-theoretic ingridients as follows: 78 | 79 | * Types are the elements of `Ty`. 80 | * Variable contexts are the objects of `VCtx`, variable renamings are the morphisms. 81 | * Metavariable contexts are the objects of `MCtx`, metavariable renamings are the morphisms. 82 | * `x : A ∈ Γ` is `x ∈ Jⱽ Γ A` 83 | * `M : [ Δ , A ]∈ Θ` is `M ∈ Jᴹ Θ (Δ, A)` 84 | * `Θ ; Γ ⊢ t : A` is `t ∈ Term Θ Γ A` 85 | * a substitution `σ : Γ ⇒ˢ Θ ⊕ Δ` is a morphism `σ : Jᵛ Γ → Tᵛ Θ Δ` in `Set^Ty`. 86 | * the renaming action `[ ρ ]ʳ : Term Θ Γ → Term Θ Δ` is `Term Θ ρ` 87 | * the substitution action `[ σ ]ˢ : Term Θ Γ → Term Θ Δ` is the Kleisli extension of `σ` 88 | * an instantiation `I : Θ ⇒ⁱ Ξ ⊕ Γ` is a morphism `I : Jᵐ Θ → Tᵐ Γ Ξ` in `Set^(VCtx × Ty)` 89 | * the instantiation action `[ I ]ⁱ : Term Θ Γ → Term Ξ Γ` is **(please fill in)** 90 | * extension `⇑ʳ` arises from the coproduct structure on `VCtx` 91 | * extension `⇑ˢ` arises from the coproduct structure on `Terms`(category whose objects are the contexts, and whose morphisms are ) 92 | * extension `⇑ⁱ` arises from **(please fill in)** 93 | 94 | 95 | 96 | ## Outline of the ideas we are pursing 97 | 98 | In this project we are aiming to formalize simple type theories in Agda. We may proceed along two axes, **generality** and **meta-analysis**. 99 | 100 | ### Generality 101 | 102 | Generality is about how large a class of type theories we encompass. There are many ways to measure "generality", but 103 | here is one that ought to work well and allow us to proceed in steps. 104 | 105 | #### Single-sorted algebraic theory 106 | 107 | **Parameterized by:** 108 | 109 | * a family of **term symbols**, each having an arity (a set, or a number) 110 | * a family of **equational axioms** `Γ ⊢ ℓ ≡ r`, see below 111 | 112 | **Expressions:** 113 | 114 | * **terms `s`, `t`:** 115 | 116 | **Judgement forms:** 117 | 118 | | Form | Intent | 119 | |:----:|:------:| 120 | | `Γ ⊢ t` | `t` is a term in context `Γ = (x₁, ..., xᵢ)` | 121 | | `Γ ⊢ s ≡ t` | `s` and `t` are equal terms in context `Γ` | 122 | 123 | **Examples:** monoid, group, (fintarily branching) inductive types, ring, `R`-module (for a fixed `R`). 124 | 125 | 126 | #### Multi-sorted algebraic theory 127 | 128 | Parameterized by: 129 | 130 | * a set of **sorts** 131 | * a family of **term symbols**, each having an arity `(σ₁, ..., σᵢ) → τ` where `σ`'s and `τ` are sorts. 132 | * a family of **equational axioms** `Γ ⊢ ℓ ≡ r : σ`, see below 133 | 134 | 135 | **Expressions:** 136 | 137 | * **terms `s`, `t`** 138 | 139 | **Judgement forms:** 140 | 141 | | Form | Intent | 142 | |:----:|:------:| 143 | | `Γ ⊢ t` | `t` is a term in context `Γ = (x₁, ..., xᵢ)` | 144 | | `Γ ⊢ s ≡ t : σ` | `s` and `t` are equal terms of sort `σ` in context `Γ` | 145 | 146 | **Examples:** module, directed graph, mutually recursive (finitarily branching) inductive types. 147 | 148 | 149 | #### Simple type theory I 150 | 151 | These extend many-sorted algebraic theory in two ways: 152 | 153 | * instead of having an amorphous set of sorts, we have types which are generated using type symbols 154 | * term symbols may bind variables (e.g., `λ`-abstraction) 155 | 156 | Parameterized by: 157 | 158 | * a family of **type symbols**, each having an arity (a number, or a set) 159 | * a family of **term symbols**, each having a formation rule (to be determined precisely what that is) 160 | * a family of **equational axioms** `Γ ⊢ ℓ ≡ r : σ`, see below 161 | 162 | **Syntactic classes:** 163 | 164 | * **types `σ`, `τ`**, generated by type symbols 165 | * **terms `s`, `t`** generated by variables, and term symbols applied to arguments 166 | * **arguments** are types and (possibly abstracted) terms 167 | 168 | **Judgement forms:** 169 | 170 | | Form | Intent | 171 | |:----:|:------:| 172 | | `Γ ⊢ t : τ` | `t` is a term of type `τ` in context `Γ = (x₁:σ₁, …, xᵢ:σᵢ)` | 173 | | `Γ ⊢ s ≡ t : τ` | `s` and `t` are equal terms of type `τ` in context `Γ` | 174 | 175 | 176 | **Example:** untyped λ-calculus, simply typed λ-calculus 177 | 178 | 179 | #### Simple type theory II 180 | 181 | These are like simple type theories, but we also allow types to take term arguments. However, types may not contain any free variables. This allows us to form types that depende on terms, such as a subtype `{ x : τ | ϕ }` where `x : τ ⊢ ϕ : bool` -- here `x` is bound in the type. 182 | 183 | Parameterized by: 184 | 185 | * a family of **type symbols**, each having a formation rule (to be determined) 186 | * a family of **term symbols**, each having a formation rule (to be determined) 187 | * a family of **equational term axioms** `Γ ⊢ ℓ ≡ r : σ`, see below 188 | * a family of **equational type axioms** `Γ ⊢ σ ≡ τ`, see below 189 | 190 | **Syntactic classes:** 191 | 192 | * **types `σ`, `τ`**, generated by type symbols appied to arguments 193 | * **terms `s`, `t`** generated by variables, and term symbols applied to term arguments 194 | * **arguments** are types and (possibly abstracted) terms 195 | 196 | Important: types may *not* contain any free variables! 197 | 198 | 199 | **Judgement forms:** 200 | 201 | | Form | Intent | 202 | |:----:|:------:| 203 | | `σ ≡ τ` | `σ` and `τ` are equal types | 204 | | `Γ ⊢ t : τ` | `t` is a term of type `τ` in context `Γ = (x₁:σ₁, …, xᵢ:σᵢ)` | 205 | | `Γ ⊢ s ≡ t : τ` | `s` and `t` are equal terms of type `τ` in context `Γ` | 206 | 207 | **Example:** internal logic of an elementary topos 208 | 209 | 210 | ### Meta-analysis 211 | 212 | We may pursue several directions of study. 213 | 214 | #### Concreteness of definition 215 | 216 | Definitions can be given at different levels of concreteness. For example, we may have raw terms and types, and a 217 | separate judgement that such raw entities are well formed, or we have a single definition of terms and types that 218 | prevents us from ever generating an ill-typed term. 219 | 220 | 221 | #### Syntactic meta-theorems 222 | 223 | We can prove theorems, such as uniqueness of typing, substitution lemmas, admissibility of substitutions, etc. 224 | 225 | 226 | #### Semantics 227 | 228 | 229 | We can define categorical semantics and show it to be sound. 230 | 231 | #### Semantic completeness 232 | 233 | We can show that the semantics is complete by constructing the initial models. 234 | -------------------------------------------------------------------------------- /formaltt.agda-lib: -------------------------------------------------------------------------------- 1 | name: formaltt -- formalization of type theories 2 | depend: agda-categories 3 | include: src -------------------------------------------------------------------------------- /src/Experimental/AlgebraFormalization.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive 2 | open import Agda.Builtin.Equality renaming (_≡_ to _==_) --(( If I want to rename the built-in equality )) 3 | 4 | module AlgebraFormalization where 5 | 6 | -- Here is an attempt to translate (part of) the Coq formalization. I did not translate the definition of the free algebra, the proof, and the part concerning the groups yet. I hope I did not do too weird things with the levels. 7 | 8 | 9 | -- ** Formalization of single-sorted algebraic theories ** 10 | 11 | -- Function extensionality 12 | postulate 13 | funext : ∀ {X : Set} {Y : X → Set} {f g : ∀ (x : X) → (Y x)} → (∀ (x : X) → ((f x) == (g x))) → (f == g) 14 | -- Operation Signature 15 | record OpSignature {l : Level} : Set (lsuc l) where 16 | field 17 | operation : Set l 18 | arity : operation → Set l 19 | -- I used the things with the levels of the hierachy of types because Agda was unhappy with what I was doing (and I did not want to define it only with Set₀ and Set₁ beacause I wanted to avoid weird things if we then want to define an OpSignature with an operation OpSignature) 20 | 21 | 22 | -- Here is a attempt (partial for the moment) to add context and terms, based on the same principles as in the file "ManySortedAlgebra.agda" 23 | -- Contexts 24 | record Context {l : Level} (Σ : OpSignature {l}) : Set (lsuc l) where 25 | field 26 | var : Set l 27 | 28 | open Context 29 | 30 | -- Terms on an operation signature 31 | data Term {l : Level} {Σ : OpSignature {l}} (Γ : Context Σ) : Set l where 32 | tm-var : ∀ (x : var Γ) → Term Γ 33 | tm-op : ∀ (f : OpSignature.operation Σ) → (OpSignature.arity Σ f → Term Γ) → Term Γ 34 | 35 | substitution : ∀ {l : Level} {Σ : OpSignature {l}} (Γ Δ : Context Σ) → Set l 36 | substitution Γ Δ = ∀ (x : var Γ) → Term Δ 37 | 38 | 39 | -- the action of a substitution on a term 40 | _·_ : ∀ {l : Level} {Σ : OpSignature {l}} {Γ Δ : Context Σ} → substitution Γ Δ → Term Γ → Term Δ 41 | σ · (tm-var x) = σ x 42 | σ · (tm-op f x) = tm-op f (λ i → σ · x i) 43 | 44 | infixr 6 _·_ 45 | 46 | -- composition of substitutions 47 | _○_ : ∀ {l : Level} {Σ : OpSignature {l}} {Γ Δ Θ : Context Σ} → substitution Δ Θ → substitution Γ Δ → substitution Γ Θ 48 | (σ ○ τ) x = σ · τ x 49 | 50 | infixl 7 _○_ 51 | 52 | -- End of the attempt 53 | 54 | 55 | 56 | 57 | 58 | 59 | -- Operation Algebra 60 | record OpAlgebra {l : Level} (S : OpSignature {l}) : Set (lsuc l) where 61 | field 62 | carrier : Set l 63 | op : ∀ (o : OpSignature.operation S) (f : OpSignature.arity S o → carrier) → carrier 64 | 65 | -- Homomorphisms 66 | record Hom {l : Level} {S : OpSignature {l}} (A : OpAlgebra {l} S) (B : OpAlgebra {l} S) : Set (lsuc l) where 67 | field 68 | map : (OpAlgebra.carrier A) → (OpAlgebra.carrier B) 69 | op-commute : ∀ (o : OpSignature.operation S) (args : OpSignature.arity S o → OpAlgebra.carrier A) → (map (OpAlgebra.op A o args) == OpAlgebra.op B o (λ x → map (args x) )) 70 | 71 | -- For the moment I skip the translation of the part concerning the free algebra 72 | 73 | 74 | 75 | 76 | 77 | 78 | 79 | -- Attempt to formalize the Equational theories diffenrently, based on what is done in the file "ManySortedAglebra.agda" 80 | 81 | -- Equational Theory (equations are seen as "in a context" and not with arities anymore) 82 | record EquationalTheory {l : Level} (Σ : OpSignature {l}) : Set (lsuc l) where 83 | field 84 | eq : Set l 85 | eq-ctx : ∀ (ε : eq) → Context {l} Σ 86 | eq-lhs : ∀ (ε : eq) → Term (eq-ctx ε) 87 | eq-rhs : ∀ (ε : eq) → Term (eq-ctx ε) 88 | 89 | open EquationalTheory 90 | 91 | -- Equality of terms 92 | infix 4 _≡_ 93 | 94 | data _≡_ {l : Level} {Σ : OpSignature {l}} {T : EquationalTheory {l} Σ } : {Γ : Context Σ} → Term Γ → Term Γ → Set (lsuc l) where 95 | -- general rules 96 | eq-refl : ∀ {Γ} {t : Term Γ } → t ≡ t 97 | eq-symm : ∀ {Γ} {s t : Term {l} {Σ} Γ } → _≡_ {T = T} s t → t ≡ s 98 | eq-tran : ∀ {Γ} {s t u : Term Γ } → _≡_ {T = T} s t → _≡_ {T = T} t u → s ≡ u 99 | -- congruence rule 100 | eq-congr : ∀ {Γ} {f : OpSignature.operation Σ} (x y : ∀ (i : OpSignature.arity Σ f) → Term Γ) → (∀ i → _≡_ {_} {_} {T} (x i) (y i)) → ((tm-op f x) ≡ (tm-op f y)) 101 | -- equational axiom 102 | eq-axiom : ∀ (ε : eq T) {Δ : Context {l} Σ} (σ : substitution (eq-ctx T ε) Δ) → 103 | ((σ · eq-lhs T ε) ≡ (σ · eq-rhs T ε)) 104 | 105 | 106 | -- composition is functorial 107 | subst-○ : ∀ {l : Level} {Σ : OpSignature {l}} {T : EquationalTheory Σ} {Γ Δ Θ : Context Σ} 108 | (σ : substitution Δ Θ) (τ : substitution Γ Δ) → 109 | ∀ (t : Term Γ ) → _≡_ {T = T} (σ · τ · t) (σ ○ τ · t) 110 | subst-○ σ τ (tm-var x) = eq-refl 111 | subst-○ σ τ (tm-op f x) = eq-congr (λ i → σ · τ · x i) (λ i → σ ○ τ · x i) λ i → subst-○ σ τ (x i) 112 | 113 | -- substitution preserves equality 114 | eq-subst : ∀ {l : Level} {Σ : OpSignature {l}} {T : EquationalTheory Σ} {Γ Δ : Context Σ} (σ : substitution Γ Δ) 115 | {s t : Term Γ} → _≡_ {T = T} s t → _≡_ {T = T} (σ · s) (σ · t) 116 | eq-subst σ eq-refl = eq-refl 117 | eq-subst σ (eq-symm ξ) = eq-symm (eq-subst σ ξ) 118 | eq-subst σ (eq-tran ζ ξ) = eq-tran (eq-subst σ ζ) (eq-subst σ ξ) 119 | eq-subst σ (eq-congr x y ξ) = eq-congr (λ i → σ · x i) (λ i → σ · y i) λ i → eq-subst σ (ξ i) 120 | eq-subst {T = T} σ (eq-axiom ε τ) = 121 | eq-tran (subst-○ σ τ (eq-lhs T ε)) 122 | (eq-tran (eq-axiom ε (σ ○ τ)) (eq-symm (subst-○ σ τ (eq-rhs T ε)))) 123 | 124 | 125 | -- End of the attempt 126 | 127 | 128 | 129 | 130 | 131 | -- Equation Signature 132 | record EqSignature {l : Level} (S : OpSignature {l}) : Set (lsuc l) where 133 | field 134 | eq : Set l 135 | eq-arity : eq → Set l 136 | -- the two sides of the equation 137 | lhs : ∀ {A : OpAlgebra {l} S} {e : eq} → ((eq-arity e) → (OpAlgebra.carrier A)) → (OpAlgebra.carrier A) 138 | rhs : ∀ {A : OpAlgebra {l} S} {e : eq} → ((eq-arity e) → (OpAlgebra.carrier A)) → (OpAlgebra.carrier A) 139 | -- naturality / commutation 140 | lhs-natural : ∀ (A B : OpAlgebra {l} S) (f : Hom A B) (e : eq) (args : eq-arity e → OpAlgebra.carrier A) → ( (Hom.map f) (lhs {A} {e} args) == lhs {B} {e} (λ i → (Hom.map f) (args i))) 141 | rhs-natural : ∀ (A B : OpAlgebra {l} S) (f : Hom A B) (e : eq) (args : eq-arity e → OpAlgebra.carrier A) → ( (Hom.map f) (rhs {A} {e} args) == rhs {B} {e} (λ i → (Hom.map f) (args i))) 142 | 143 | -- Algebra 144 | record Algebra {l : Level} (S : OpSignature {l}) (E : EqSignature {l} S) : Set (lsuc l) where 145 | field 146 | alg : OpAlgebra S 147 | equations : ∀ (e : EqSignature.eq E) (args : EqSignature.eq-arity E e → OpAlgebra.carrier alg) → (EqSignature.rhs E {alg} {e} args == EqSignature.lhs E {alg} {e} args) 148 | 149 | 150 | -- For the moment, I think that we did not talk about terms in contexts, did we ? Should we generalize the things we did in Lambda.agda (using De Bruijn indices for the variables) ? -> I tried things 151 | 152 | 153 | 154 | 155 | 156 | 157 | 158 | -- ** Other things ** 159 | 160 | -- Useful arities 161 | data nullary : Set where 162 | data unary : Set where 163 | only : unary 164 | data binary : Set where 165 | fst : binary 166 | snd : binary 167 | 168 | 169 | -- Definition of groups 170 | data GroupOperation : Set where 171 | one : GroupOperation 172 | mul : GroupOperation 173 | inv : GroupOperation 174 | 175 | GroupArity : (o : GroupOperation) → Set 176 | GroupArity one = nullary 177 | GroupArity mul = binary 178 | GroupArity inv = unary 179 | 180 | GroupOp : OpSignature 181 | GroupOp = record { operation = GroupOperation ; arity = GroupArity} 182 | 183 | 184 | 185 | -------------------------------------------------------------------------------- /src/Experimental/Category.agda: -------------------------------------------------------------------------------- 1 | module Category where 2 | 3 | open import Agda.Primitive 4 | 5 | open import Equality 6 | open import Finite 7 | 8 | record Category ℓ : Set (lsuc ℓ) where 9 | field 10 | obj : Set ℓ 11 | hom : obj → obj → Set ℓ 12 | id-hom : ∀ {x} → hom x x 13 | compose : ∀ {x y z} → hom y z → hom x y → hom x z 14 | compose-id-left : ∀ {x y} {f : hom x y} → compose id-hom f ≡ f 15 | compose-id-right : ∀ {x y} {f : hom x y} → compose f id-hom ≡ f 16 | compose-assoc : ∀ {x y z w} {f : hom x y} {g : hom y z} {h : hom z w} → compose (compose h g) f ≡ compose h (compose g f) 17 | 18 | record FPCategory ℓ : Set (lsuc ℓ) where 19 | open Category 20 | 21 | field 22 | {{category}} : Category ℓ 23 | product : ∀ n (xs : Fin n → obj category) → obj category 24 | project : ∀ n {xs : Fin n → obj category} → ∀ (i : Fin n) → hom category (product n xs) (xs i) 25 | tuple : ∀ n {xs : Fin n → obj category} {y} → ∀ (f : ∀ i → hom category y (xs i)) → hom category y (product n xs) 26 | project-tuple : ∀ {n} {xs : Fin n → obj category} {y} {f : ∀ i → hom category y (xs i)} {j} → 27 | compose category (project n j) (tuple n f) ≡ f j 28 | tuple-project : ∀ {n} {xs : Fin n → obj category} {y} {u v : hom category y (product n xs)} → 29 | (∀ i → compose category (project n i) u ≡ compose category (project n i) v) → u ≡ v 30 | 31 | tuple-hom : ∀ n {xs ys : Fin n → obj category} (f : ∀ i → hom category (xs i) (ys i)) → 32 | hom category (product n xs) (product n ys) 33 | tuple-hom n f = tuple n (λ i → compose category (f i) (project n i)) 34 | 35 | power : ℕ → obj category → obj category 36 | power n x = product n (λ _ → x) 37 | 38 | power-hom : ∀ n {x y : obj category} → hom category x y → hom category (power n x) (power n y) 39 | power-hom n f = tuple-hom n (λ i → f) 40 | 41 | terminal-obj : obj category 42 | terminal-obj = product Z (λ ()) 43 | 44 | terminal-hom : ∀ {x : obj category} → hom category x terminal-obj 45 | terminal-hom = tuple Z (λ ()) 46 | 47 | terminal-eq : ∀ {x : obj category} {f g : hom category x terminal-obj} → f ≡ g 48 | terminal-eq {x} {f} {g} = tuple-project λ () 49 | -------------------------------------------------------------------------------- /src/Experimental/Equality.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive 2 | 3 | module Experimental.Equality where 4 | 5 | data _≡_ {ℓ} {A : Set ℓ} : A → A → Set ℓ where 6 | refl : ∀ {x : A} → x ≡ x 7 | 8 | sym : ∀ {ℓ} {A : Set ℓ} {x y : A} → x ≡ y → y ≡ x 9 | sym refl = refl 10 | 11 | tran : ∀ {ℓ} {A : Set ℓ} {x y z : A} → x ≡ y → y ≡ z → x ≡ z 12 | tran refl q = q 13 | 14 | ap : ∀ {ℓ k} {A : Set ℓ} {B : Set k} {x y : A} {f : A → B} → x ≡ y → f x ≡ f y 15 | ap refl = refl 16 | 17 | transport : ∀ {ℓ k} {A : Set ℓ} (B : A → Set k) {x y : A} → x ≡ y → B x → B y 18 | transport _ refl u = u 19 | -------------------------------------------------------------------------------- /src/Experimental/Finite.agda: -------------------------------------------------------------------------------- 1 | module Experimental.Finite where 2 | 3 | data ℕ : Set where 4 | Z : ℕ 5 | S : ℕ → ℕ 6 | 7 | data Fin : ℕ → Set where 8 | Fin-S : ∀ {n} → Fin n → Fin (S n) 9 | Fin-Z : ∀ {n} → Fin (S n) 10 | -------------------------------------------------------------------------------- /src/Experimental/Lambda.agda: -------------------------------------------------------------------------------- 1 | module Lambda where 2 | 3 | -- The typed λ-calculus 4 | 5 | -- The type of types 6 | data ty : Set where 7 | ι : ty -- the base type 8 | _⇒_ : ty → ty → ty -- function type 9 | _×_ : ty → ty → ty -- product type 10 | 11 | infixr 7 _⇒_ 12 | infixl 8 _×_ 13 | 14 | -- A typing context, there are no variable names because we use de Bruijn indices. 15 | data ctx : Set where 16 | • : ctx -- empty context 17 | _,_ : ctx → ty → ctx -- context extension 18 | 19 | infixr 5 _,_ 20 | 21 | -- A variable is a natural number (de Bruijn index, position in the context), 22 | -- but at the same time a proof that a type appears in the context. 23 | -- We write Γ ∋ A for the type of variables in Γ of type A. 24 | data _∋_ : (Γ : ctx) (B : ty) → Set where 25 | Z : {Γ : ctx} {B : ty} → (Γ , B) ∋ B -- zero index 26 | S : {Γ : ctx} {A B : ty} → Γ ∋ B → (Γ , A) ∋ B -- successor index 27 | 28 | infix 4 _∋_ 29 | 30 | -- The type of terms in context Γ of type A 31 | data tm (Γ : ctx) : (A : ty) → Set where 32 | tm-var : {A : ty} → Γ ∋ A → tm Γ A 33 | tm-λ : {B : ty} {A : ty} → tm (Γ , A) B → tm Γ (A ⇒ B) 34 | tm-app : {A B : ty} → tm Γ (A ⇒ B) → tm Γ A → tm Γ B 35 | tm-pair : {A B : ty} → tm Γ A → tm Γ B → tm Γ (A × B) 36 | tm-fst : {A B : ty} → tm Γ (A × B) → tm Γ A 37 | tm-snd : {A B : ty} → tm Γ (A × B) → tm Γ B 38 | 39 | -- We need serveral boring auxiliary functions whose 40 | -- purpose is to define substitution 41 | 42 | -- A variable renaming is a type-preserving map from variables in Γ to variables in Δ 43 | -- (In Agda "renaming" is a reserved word.) 44 | variable-renaming = λ Γ Δ → ∀ {A} → Γ ∋ A → Δ ∋ A 45 | 46 | -- we may extend a renaming by one more variable (which does not get renamed) 47 | extend-renaming : ∀ {Γ Δ A} → variable-renaming Γ Δ → variable-renaming (Γ , A) (Δ , A) 48 | extend-renaming ρ Z = Z 49 | extend-renaming ρ (S x) = S (ρ x) 50 | 51 | -- the action of a renaming on a term 52 | term-rename : ∀ {Γ Δ} → variable-renaming Γ Δ → (∀ {A} → tm Γ A → tm Δ A) 53 | term-rename ρ (tm-var x) = tm-var (ρ x) 54 | term-rename ρ (tm-λ t) = tm-λ (term-rename (extend-renaming ρ) t) 55 | term-rename ρ (tm-app s t) = tm-app (term-rename ρ s) (term-rename ρ t) 56 | term-rename ρ (tm-pair s t) = tm-pair (term-rename ρ s) (term-rename ρ t) 57 | term-rename ρ (tm-fst t) = tm-fst (term-rename ρ t) 58 | term-rename ρ (tm-snd t) = tm-snd (term-rename ρ t) 59 | 60 | -- a special kind of renaming is weakening by a variable, which we write as ↑ 61 | ↑ : ∀ {Γ A B} → tm Γ A → tm (Γ , B) A 62 | ↑ = term-rename S 63 | 64 | -- a substitution from Γ to Δ takes variables in Γ to terms in Δ 65 | substitution = λ Γ Δ → (∀ {A} → Γ ∋ A → tm Δ A) 66 | 67 | extend-substutition : ∀ {Γ Δ A} → substitution Γ Δ → substitution (Γ , A) (Δ , A) 68 | extend-substutition σ Z = tm-var Z 69 | extend-substutition σ (S x) = ↑ (σ x) 70 | 71 | -- The action of a substitution on a term 72 | term-substitute : ∀ {Γ Δ} → substitution Γ Δ → ∀ {A} → tm Γ A → tm Δ A 73 | term-substitute σ (tm-var x) = σ x 74 | term-substitute σ (tm-λ t) = tm-λ (term-substitute (extend-substutition σ) t) 75 | term-substitute σ (tm-app s t) = tm-app (term-substitute σ s) (term-substitute σ t) 76 | term-substitute σ (tm-pair s t) = tm-pair (term-substitute σ s) (term-substitute σ t) 77 | term-substitute σ (tm-fst t) = tm-fst (term-substitute σ t) 78 | term-substitute σ (tm-snd t) = tm-snd (term-substitute σ t) 79 | 80 | -- Auxiliary substitution that replaces just the 0-th variable 81 | subst-Z : ∀ {Γ A} → tm Γ A → substitution (Γ , A) Γ 82 | subst-Z t Z = t 83 | subst-Z T (S x) = tm-var x 84 | 85 | -- A common kind of substitution only replaces the 0-th variable with a term 86 | -- and leaves all the others intact, so we define its action as a shorthand 87 | _[_] : ∀ {Γ A B} → tm (Γ , A) B → tm Γ A → tm Γ B 88 | _[_] {Γ} {A} {B} s t = term-substitute (subst-Z t) s where 89 | 90 | infix 5 _[_] 91 | 92 | -- Judgemental equality 93 | data _≡_ {Γ : ctx} : {A : ty} (t s : tm Γ A) → Set where 94 | -- general rules 95 | eq-refl : {A : ty} {t : tm Γ A} → t ≡ t 96 | eq-tran : {A : ty} {t s u : tm Γ A} → t ≡ s → s ≡ u → t ≡ u 97 | eq-sym : {A : ty} {t s : tm Γ A} → t ≡ s → s ≡ t 98 | -- congruence rules 99 | eq-congr-app : ∀ {A B} {t₁ t₂ : tm Γ (A ⇒ B)} {s₁ s₂ : tm Γ A} → 100 | t₁ ≡ t₂ → s₁ ≡ s₂ → tm-app t₁ s₁ ≡ tm-app t₂ s₂ 101 | eq-congr-λ : ∀ {A B} {t₁ t₂ : tm (Γ , A) B} → t₁ ≡ t₂ → tm-λ t₁ ≡ tm-λ t₂ 102 | eq-congr-pair : ∀ {A B} {t₁ t₂ : tm Γ (A ⇒ B)} {s₁ s₂ : tm Γ A} → 103 | t₁ ≡ t₂ → s₁ ≡ s₂ → tm-pair t₁ s₁ ≡ tm-pair t₂ s₂ 104 | eq-congr-fst : ∀ {A B} {s t : tm Γ (A × B)} → s ≡ t → tm-fst s ≡ tm-fst t 105 | eq-congr-snd : ∀ {A B} {s t : tm Γ (A × B)} → s ≡ t → tm-snd s ≡ tm-snd t 106 | -- function type rules 107 | eq-β : ∀ {A B} {t : tm (Γ , A) B} {s : tm Γ A} → tm-app (tm-λ t) s ≡ t [ s ] 108 | eq-⇒ : ∀ {A B} {s t : tm Γ (A ⇒ B)} → 109 | tm-app (↑ s) (tm-var Z) ≡ tm-app (↑ t) (tm-var Z) 110 | → s ≡ t 111 | -- product rules 112 | eq-fst : ∀ {A B} {u : tm Γ A} {v : tm Γ B} → tm-fst (tm-pair u v) ≡ u 113 | eq-snd : ∀ {A B} {u : tm Γ A} {v : tm Γ B} → tm-snd (tm-pair u v) ≡ v 114 | eq-× : ∀{A B} {s t : tm Γ (A × B)} → tm-fst s ≡ tm-fst t → tm-snd s ≡ tm-snd t → s ≡ t 115 | 116 | infix 4 _≡_ 117 | 118 | -- "η-equivalence" for products 119 | eq-pair-η : ∀ {Γ A B} {s : tm Γ (A × B)} → tm-pair (tm-fst s) (tm-snd s) ≡ s 120 | eq-pair-η = eq-× eq-fst eq-snd 121 | 122 | -- subst-↑ : ∀ {Γ A B} (t : tm Γ A) (s : tm Γ B) → (( (↑ t) [ s ] ) ≡ t) 123 | -- subst-↑ (tm-var x) s = eq-refl 124 | -- subst-↑ (tm-λ t) s = eq-congr-λ {!!} 125 | -- subst-↑ (tm-app t₁ t₂) s = eq-congr-app (subst-↑ t₁ s) (subst-↑ t₂ s) 126 | -- I commented this lemma that we tried to prove this morning because we did not achieve to do it yet, and Agda doesn't like the fact that I did not detail the other case of terms yet 127 | 128 | 129 | -- Example: the identity function 130 | -- Note that we actually define a family of terms, indexed by a context 131 | -- and a type, but the family is constant, i.e., it is always the same term. 132 | tm-id : ∀ {Γ A} → tm Γ (A ⇒ A) 133 | tm-id = tm-λ (tm-var Z) 134 | 135 | -- Appying the identity function twice does nothing 136 | app-id-id : ∀ {Γ A} {t : tm Γ A} → tm-app tm-id (tm-app tm-id t) ≡ t 137 | -- app-id-id = eq-tran (eq-congr-app eq-refl eq-β) eq-β 138 | app-id-id = eq-tran eq-β eq-β 139 | 140 | -- Eta-rule 141 | 142 | eq-η : ∀ {Γ A B} {t : tm Γ (A ⇒ B)} → tm-λ (tm-app (↑ t) (tm-var Z)) ≡ t 143 | eq-η = eq-⇒(eq-tran eq-β (eq-congr-app {!!} eq-refl)) 144 | 145 | -- natural numbers 146 | data N : Set where 147 | zero : N 148 | succ : N → N 149 | 150 | -- church numerals 151 | 152 | tm-numeral : ∀ {Γ} → N → tm Γ ((ι ⇒ ι) ⇒ (ι ⇒ ι)) 153 | tm-numeral zero = tm-id 154 | tm-numeral (succ n) = tm-λ (tm-λ (tm-app (tm-app (tm-numeral n) (tm-var (S Z))) (tm-app (tm-var (S Z)) (tm-var Z)))) 155 | -------------------------------------------------------------------------------- /src/Experimental/ManySortedAlgebra.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive 2 | 3 | module ManySortedAlgebra where 4 | 5 | -- a many sorted signature 6 | record Signature {l : Level} : Set (lsuc l) where 7 | field 8 | sort : Set l -- sorts 9 | op : Set l -- operations 10 | arg : op → Set l 11 | op-sort : op → sort -- the sort of the operation 12 | arg-sort : ∀ {f} → arg f → sort -- the sorts of arguments 13 | 14 | open Signature 15 | 16 | -- we allow general contexts in which there are arbitrarily many variables, 17 | -- which makes things easier 18 | record Context {l : Level} (Σ : Signature {l}) : Set (lsuc l) where 19 | field 20 | var : Set l 21 | var-sort : var → sort Σ 22 | 23 | open Context 24 | 25 | -- terms over a signature in a context of a given sort 26 | data Term {l : Level} {Σ : Signature {l}} (Γ : Context Σ) : sort Σ → Set l where 27 | tm-var : ∀ (x : var Γ) → Term Γ (var-sort Γ x) 28 | tm-op : ∀ (f : op Σ) → (∀ (i : arg Σ f) → Term Γ (arg-sort Σ i)) → Term Γ (op-sort Σ f) 29 | 30 | substitution : ∀ {l : Level} {Σ : Signature {l}} (Γ Δ : Context Σ) → Set l 31 | substitution Γ Δ = ∀ (x : var Γ) → Term Δ (var-sort Γ x) 32 | 33 | -- the action of a substitution on a term 34 | _·_ : ∀ {l : Level} {Σ : Signature {l}} {Γ Δ : Context Σ} → substitution Γ Δ → ∀ {A} → Term Γ A → Term Δ A 35 | σ · (tm-var x) = σ x 36 | σ · (tm-op f x) = tm-op f (λ i → σ · x i) 37 | 38 | infixr 6 _·_ 39 | 40 | -- composition of substitutions 41 | _○_ : ∀ {l : Level} {Σ : Signature {l}} {Γ Δ Θ : Context Σ} → substitution Δ Θ → substitution Γ Δ → substitution Γ Θ 42 | (σ ○ τ) x = σ · τ x 43 | 44 | infixl 7 _○_ 45 | 46 | -- an equational theory is a family of equations over a given sort 47 | record EquationalTheory {l : Level} (Σ : Signature {l}) : Set (lsuc l) where 48 | field 49 | eq : Set l 50 | eq-ctx : ∀ (ε : eq) → Context {l} Σ 51 | eq-sort : ∀ (ε : eq) → sort Σ 52 | eq-lhs : ∀ (ε : eq) → Term (eq-ctx ε) (eq-sort ε) 53 | eq-rhs : ∀ (ε : eq) → Term (eq-ctx ε) (eq-sort ε) 54 | 55 | open EquationalTheory 56 | 57 | infix 4 _≡_ 58 | 59 | -- the remaining judgement form is equality 60 | data _≡_ {l : Level} {Σ : Signature {l}} {T : EquationalTheory {l} Σ } : {Γ : Context Σ} → {S : sort Σ} → Term Γ S → Term Γ S → Set (lsuc l) where 61 | -- general rules 62 | eq-refl : ∀ {Γ} {S : sort Σ} {t : Term Γ S} → t ≡ t 63 | eq-symm : ∀ {Γ} {S : sort Σ} {s t : Term {l} {Σ} Γ S} → _≡_ {T = T} s t → t ≡ s 64 | eq-tran : ∀ {Γ} {S : sort Σ} {s t u : Term Γ S} → _≡_ {T = T} s t → _≡_ {T = T} t u → s ≡ u 65 | -- congruence rule 66 | eq-congr : ∀ {Γ} {f : op Σ} (x y : ∀ (i : arg Σ f) → Term Γ (arg-sort Σ i)) → 67 | (∀ i → _≡_ {_} {_} {T} (x i) (y i)) → tm-op f x ≡ tm-op f y 68 | -- equational axiom 69 | eq-axiom : ∀ (ε : eq T) {Δ : Context Σ} (σ : substitution (eq-ctx T ε) Δ) → 70 | σ · eq-lhs T ε ≡ σ · eq-rhs T ε 71 | 72 | -- composition is functorial 73 | subst-○ : ∀ {l : Level} {Σ : Signature {l}} {T : EquationalTheory Σ} {Γ Δ Θ : Context Σ} 74 | (σ : substitution Δ Θ) (τ : substitution Γ Δ) → 75 | ∀ {A} (t : Term Γ A) → _≡_ {T = T} (σ · τ · t) (σ ○ τ · t) 76 | subst-○ σ τ (tm-var x) = eq-refl 77 | subst-○ σ τ (tm-op f x) = eq-congr (λ i → σ · τ · x i) (λ i → σ ○ τ · x i) λ i → subst-○ σ τ (x i) 78 | 79 | -- substitution preserves equality 80 | eq-subst : ∀ {l : Level} {Σ : Signature {l}} {T : EquationalTheory Σ} {Γ Δ : Context Σ} {S : sort Σ} (σ : substitution Γ Δ) 81 | {s t : Term Γ S} → _≡_ {T = T} s t → _≡_ {T = T} (σ · s) (σ · t) 82 | eq-subst σ eq-refl = eq-refl 83 | eq-subst σ (eq-symm ξ) = eq-symm (eq-subst σ ξ) 84 | eq-subst σ (eq-tran ζ ξ) = eq-tran (eq-subst σ ζ) (eq-subst σ ξ) 85 | eq-subst σ (eq-congr x y ξ) = eq-congr (λ i → σ · x i) (λ i → σ · y i) λ i → eq-subst σ (ξ i) 86 | eq-subst {T = T} σ (eq-axiom ε τ) = 87 | eq-tran (subst-○ σ τ (eq-lhs T ε)) 88 | (eq-tran (eq-axiom ε (σ ○ τ)) (eq-symm (subst-○ σ τ (eq-rhs T ε)))) 89 | -------------------------------------------------------------------------------- /src/Experimental/MultiCategories.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive 2 | open import Agda.Builtin.Equality renaming (_≡_ to _==_) --(( If I want to rename the built-in equality )) 3 | 4 | -- What follows is an attempt to formalize multicategories. For this purpose, we also need to define lists, since the source of a multimap in a multicategory is a list. 5 | -- For the moment, I do not try to prove that the lists an the associated concatenation form a monoid (because I do not know if this is useful or not). 6 | 7 | module MultiCategories where 8 | 9 | -- ** Function extensionality ** 10 | postulate 11 | funext : ∀ {X : Set} {Y : X → Set} {f g : ∀ (x : X) → (Y x)} → (∀ (x : X) → ((f x) == (g x))) → (f == g) 12 | 13 | 14 | 15 | -- ** Lists ** 16 | 17 | -- We first define lists 18 | data List {l : Level} (A : Set l) : Set l where 19 | [] : List A 20 | _::_ : A → List A → List A 21 | 22 | infixr 30 _::_ 23 | 24 | open List 25 | 26 | -- We define equality on lists, that extends the built-in equality (since for the moment, we dont need other definitions of equality → but maybe we could do something more general by asking the "equality" as a parameter ?) 27 | data _≡ᴸ_ {l : Level} {A : Set l} : (u v : List {l} A) → Set (lsuc l) where 28 | eq-nil : [] ≡ᴸ [] 29 | eq-cons : ∀ (x y : A) (su sv : List A) → (x == y) → (su ≡ᴸ sv) → ( (x :: su) ≡ᴸ (y :: sv)) 30 | 31 | -- We define the mapping of lists 32 | list-map : ∀ {l : Level} {A B : Set l} → (f : A → B) → List A → List B 33 | list-map f [] = [] 34 | list-map f (x :: u) = ((f x) :: (list-map f u)) 35 | 36 | -- We define the concatenation of lists 37 | _++_ : ∀ {l}{A : Set l} → List A → List A → List A 38 | [] ++ ys = ys 39 | (x :: xs) ++ ys = x :: (xs ++ ys) 40 | 41 | -- We define a fold function 42 | fold : ∀ {l : Level} {A B : Set l} → (A → B → B) → B → List A → B 43 | fold _⊗_ e [] = e 44 | fold _⊗_ e (x :: xs) = x ⊗ (fold _⊗_ e xs) 45 | 46 | -- We define a flatten function 47 | flatten : ∀ {l : Level} {A : Set l} → (u : List ( List A)) → (List A) 48 | flatten [] = [] 49 | flatten (x :: xs) = x ++ (flatten xs) 50 | -- (I wanted to do this with the above fold function, which would be more elegant, but I don't know why, I miserably failed at it - must be tired) 51 | 52 | -- We define a function that takes a list of functions and a list of arguments and apply the functions point to point 53 | list-apply : ∀ {l : Level} {A B : Set l} → (list-f : List (A → B)) → (list-arg : List A) → List B 54 | list-apply [] [] = [] 55 | list-apply (f :: fs) [] = [] 56 | list-apply [] (x :: xs) = [] 57 | list-apply (f :: fs) (x :: xs) = (f x) :: (list-apply fs xs) 58 | -- The two cases in the middle should be forbidden, but I don't know how to do this 59 | 60 | 61 | -- ** Multicategories ** 62 | 63 | -- -- We first define the multimaps on a set 64 | -- data Multimap {l : Level} {O : Setl l} : Set l where 65 | -- map : () 66 | 67 | record MultiCategory {l : Level} : Set (lsuc l) where 68 | field 69 | -- Objects and maps 70 | object : Set l 71 | multimap : Set l 72 | sources : ∀ (m : multimap) → List object 73 | target : ∀ (m : multimap) → object 74 | -- Composition and associated equations / conditions 75 | _∘_ : ∀ (f : multimap) → (list : List multimap) → multimap 76 | plug : ∀ {g : multimap} → {f : multimap} → {list : List multimap} → {g == (f ∘ list)} → (sources f) ≡ᴸ (list-map target list) 77 | dom-comp : ∀ {f : multimap} → {list : List multimap} → ((flatten (list-map sources list)) ≡ᴸ (sources (f ∘ list))) 78 | comp-codom : ∀ (f : multimap) → (list : List multimap) → (target f == target (f ∘ list)) 79 | -- Identities and associated equations / conditions 80 | id : ∀ (o : object) → multimap 81 | id-dom-codom : ∀ (o : object) → (sources (id o) == o :: [] ) 82 | id-codom : ∀ (o : object) → (target (id o) == o) 83 | id-left : ∀ {o : object} {f : multimap} → {f == id o} → (list : List multimap) → ( (f ∘ list) == f) 84 | id-rigth : ∀ {f : multimap} {list : List multimap} → {list ≡ᴸ (list-map id (sources f))} → ((f ∘ list) == f) 85 | -- Associativity 86 | assoc : ∀ {f : multimap} {list-g : List multimap} {list-h : List (List multimap)} → (f ∘ (list-apply (list-map _∘_ list-g) list-h)) == ( (f ∘ list-g) ∘ (flatten list-h)) 87 | 88 | open MultiCategory 89 | 90 | -- List over a list 91 | data ListOver {l : Level} {A : Set l} (B : A → Set l) : List A → Set l where 92 | [[]] : ListOver B [] 93 | _:::_ : ∀ {x xs} → (y : B x) → (ys : ListOver B xs) → ListOver B (x :: xs) 94 | 95 | infixr 25 _:::_ 96 | 97 | over-map : ∀ {l : Level} {A : Set l} {B : A → Set l} {xs} {C : Set l} → (∀ {x} → B x → C) → ListOver B xs → List C 98 | over-map f [[]] = [] 99 | over-map f (y ::: ys) = f y :: over-map f ys 100 | 101 | over-over-map : ∀ {l : Level} {A : Set l} {B : A → Set l} {xs} {C : A → Set l} → (∀ {x} → B x → C x) → ListOver B xs → ListOver C xs 102 | over-over-map f [[]] = [[]] 103 | over-over-map f (y ::: ys) = f y ::: over-over-map f ys 104 | 105 | over-lift : ∀ {l : Level} {A : Set l} (list : List A) → ListOver (λ x → A) list 106 | over-lift [] = [[]] 107 | over-lift (y :: ys) = y ::: (over-lift ys) 108 | 109 | over-flatten : ∀ {l : Level} {A B : Set l} {list : List A} (list-ov : ListOver (λ x → List B) list) → List B 110 | over-flatten [[]] = [] 111 | over-flatten (x ::: xs) = x ++ (over-flatten xs) 112 | 113 | -- Dependent sum 114 | record Σ {l} (A : Set l) (B : A → Set l) : Set l where 115 | constructor ⟨_,_⟩ 116 | field 117 | π₁ : A 118 | π₂ : B π₁ 119 | 120 | open Σ 121 | 122 | -- Shortcuts to map the projections on lists when the dependent sum is a "product" 123 | list-π₁ : ∀ {l : Level} {A : Set l} {B} (list : List ( Σ {l} A B)) → List A 124 | list-π₁ list = list-map π₁ list 125 | 126 | list-π₂ : ∀ {l : Level} {A C : Set l} (list : List ( Σ {l} A ( λ x → C))) → List C 127 | list-π₂ list = list-map π₂ list 128 | 129 | 130 | 131 | -- A more dependent attempt at multicategories 132 | record MultiCategory2 {l : Level} : Set (lsuc l) where 133 | field 134 | object : Set l 135 | multimap : List object → object → Set l 136 | 𝟙 : ∀ {x} → multimap (x :: []) x 137 | _•_ : ∀ {ys x} → multimap ys x → ∀ (gs : ListOver (λ y → Σ (List object) (λ zs → multimap zs y)) ys) → 138 | multimap (flatten (over-map π₁ gs)) x 139 | -- Another attempt to define multimaps, "putting the dependance elsewhere" 140 | _●_ : ∀ {x : object} {ys : List (Σ (List object) (λ x → object))} → multimap (list-π₂ ys) x → ListOver (λ y → multimap (π₁ y) (π₂ y)) ys → multimap (flatten (list-π₁ ys)) x 141 | -- here complications start 142 | -- 𝟙-left : ∀ {ys x} → (f : multimap ys x) → 𝟙 • (⟨ ys , f ⟩ ::: [[]]) == f 143 | -- 𝟙-right : ∀ {ys x} → (f : multimap ys x) → f • (over-over-map ? (over-lift ys)) == f 144 | --(ListOver ( λ y → ⟨ x ::: [[]] , 𝟙 ⟩ ) ys) == f 145 | -- Attempt with the alternative composition 146 | 𝟙-left-● : ∀ {x ys} → (f : multimap ys x) → 𝟙 ● (f ::: [[]]) == f -- Here it seems that we have a lemma to prove to say that ys = (ys ++ flatten (list-map π₁ [])). Do we do it locally or would it be useful to have a more genral lemma ? 147 | 148 | -- Here I tried to fix 𝟙-left and to define 𝟙-right, but I did not manage to do it, it both cases. Maybe we should revise the definition of _•_ ? 149 | -- Agda seems to struggle with the fact that thnigs that should be equal are not equal by definition (conversion/reduction problems). Maybe there are some lemmas to prove here. 150 | -- Also, I do not understand why we use the "over-map" : it would feel more natural to me if, once we lift a list to a dependent one, and use dependant lists, we only use dependeant lists, that's why I defined over-lift and over-over-map. (I also defined an "over-flatten" but don't know if it's useful) 151 | -------------------------------------------------------------------------------- /src/Experimental/Quotient.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive 2 | 3 | open import Equality 4 | 5 | module Quotient where 6 | 7 | is-prop : ∀ {ℓ} (A : Set ℓ) → Set ℓ 8 | is-prop A = ∀ (x y : A) → x ≡ y 9 | 10 | record is-equivalence {ℓ k} {A : Set ℓ} (E : A → A → Set k) : Set (ℓ ⊔ k) where 11 | field 12 | equiv-prop : ∀ {x y} → is-prop (E x y) 13 | equiv-refl : ∀ {x} → E x x 14 | equiv-symm : ∀ {x y} → E x y → E y x 15 | equiv-tran : ∀ {x y z} → E x y → E y z → E x z 16 | 17 | record Equivalence {ℓ} (A : Set ℓ) : Set (lsuc ℓ) where 18 | field 19 | equiv-relation : A → A → Set ℓ 20 | equiv-is-equivalence : is-equivalence equiv-relation 21 | 22 | is-set : ∀ {ℓ} (A : Set ℓ) → Set ℓ 23 | is-set A = ∀ (x y : A) (p q : x ≡ y) → p ≡ q 24 | 25 | record HSet ℓ : Set (lsuc ℓ) where 26 | field 27 | hset-type : Set ℓ 28 | hset-is-set : is-set hset-type 29 | 30 | record Quotient {ℓ} (A : Set ℓ) (E : Equivalence A) : Set (lsuc ℓ) where 31 | open Equivalence 32 | 33 | field 34 | quot-type : Set ℓ 35 | quot-class : ∀ (x : A) → quot-type 36 | quot-≡ : ∀ {x y : A} → equiv-relation E x y → quot-class x ≡ quot-class y 37 | quot-is-set : is-set quot-type 38 | quot-elim : 39 | ∀ (B : ∀ quot-type → Set ℓ) 40 | (f : ∀ (x : A) → B (quot-class x)) → 41 | (∀ (x y : A) {ε : equiv-relation E x y} → transport B (quot-≡ ε) (f x) ≡ f y) → 42 | ∀ (ξ : quot-type) → B ξ 43 | quot-compute : 44 | ∀ (B : ∀ quot-type → Set ℓ) → 45 | (f : ∀ (x : A) → B (quot-class x)) → 46 | (η : ∀ (x y : A) {ε : equiv-relation E x y} → transport B (quot-≡ ε) (f x) ≡ f y) → 47 | (quot-elim ? quot-class η) 48 | -------------------------------------------------------------------------------- /src/Experimental/SimpleTT.agda: -------------------------------------------------------------------------------- 1 | module SimpleTT where 2 | 3 | data 𝟘 : Set where 4 | 5 | absurd : ∀ {ℓ} {A : Set ℓ} → 𝟘 → A 6 | absurd () 7 | 8 | data _⁺ (A : Set) : Set where 9 | Z : A ⁺ 10 | S : A → A ⁺ 11 | 12 | record Signature : Set₁ where 13 | field 14 | ty-op : Set 15 | tm-op : Set 16 | ty-ty-arg : ty-op → Set 17 | -- ty-tm-arg : ty-op → Set 18 | -- tm-ty-arg : tm-op → Set 19 | tm-tm-arg : tm-op → Set 20 | 21 | open Signature 22 | 23 | data Theory : Set₁ 24 | thy-signature : Theory → Signature 25 | data Ty (T : Theory) : Set 26 | data Tm (T : Theory) (A : Ty T) : Set 27 | 28 | data Theory where 29 | thy-empty : Theory 30 | th-type : ∀ (T : Theory) → (α : Set) → (α → Ty T) → Theory 31 | th-term : ∀ (T : Theory) → (α : Set) → (α → Ty T) → Ty T → Theory 32 | 33 | thy-signature thy-empty = 34 | record { ty-op = 𝟘 35 | ; tm-op = 𝟘 36 | ; ty-ty-arg = absurd 37 | ; tm-tm-arg = absurd 38 | } 39 | 40 | thy-signature (th-type T α x) = 41 | record { ty-op = (ty-op Σ) ⁺ 42 | ; tm-op = tm-op Σ 43 | ; ty-ty-arg = t 44 | ; tm-tm-arg = tm-tm-arg Σ 45 | } 46 | where 47 | Σ = thy-signature T 48 | t : ∀ (f : (ty-op Σ) ⁺) → Set 49 | t Z = α 50 | t (S f) = ty-ty-arg (thy-signature T) f 51 | 52 | thy-signature (th-term T α ts t) = 53 | record { ty-op = ty-op Σ 54 | ; tm-op = (tm-op Σ) ⁺ 55 | ; ty-ty-arg = ty-ty-arg Σ 56 | ; tm-tm-arg = tm 57 | } 58 | where 59 | Σ = thy-signature T 60 | tm : ∀ (f : (tm-op Σ) ⁺) → Set 61 | tm Z = {!!} 62 | tm (S x) = {!!} 63 | 64 | data Ty T where 65 | ty-expr : ∀ (f : ty-op (thy-signature T)) → (ty-ty-arg (thy-signature T) f → Ty T) → Ty T 66 | 67 | data Tm T A where 68 | -------------------------------------------------------------------------------- /src/MultiSorted/AlgebraicTheory.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary using (Setoid) 3 | 4 | import MultiSorted.Context as Context 5 | 6 | module MultiSorted.AlgebraicTheory where 7 | 8 | -- an algebraic signature 9 | record Signature {𝓈 ℴ} : Set (lsuc (𝓈 ⊔ ℴ)) where 10 | field 11 | sort : Set 𝓈 -- sorts 12 | oper : Set ℴ -- operations 13 | 14 | open Context sort public 15 | 16 | -- Arity and arguments 17 | Arity : Set 𝓈 18 | Arity = Context 19 | 20 | field 21 | oper-arity : oper → Arity -- the arity of an operation (with the sorts of the arguments) 22 | oper-sort : oper → sort -- the sort of an operation 23 | 24 | arg : Arity → Set 25 | arg = var 26 | 27 | arg-sort : ∀ (f : oper) → arg (oper-arity f) → sort 28 | arg-sort f = sort-of (oper-arity f) 29 | 30 | -- terms in a context of a given sort 31 | data Term (Γ : Context) : ∀ (A : sort) → Set (lsuc ℴ) where 32 | tm-var : ∀ (x : var Γ) → Term Γ (sort-of Γ x) 33 | tm-oper : ∀ (f : oper) → (∀ (i : arg (oper-arity f)) → Term Γ (arg-sort f i)) → Term Γ (oper-sort f) 34 | 35 | -- Substitutions (definitions - some useful properties are in another file) 36 | _⇒s_ : ∀ (Γ Δ : Context) → Set (lsuc ℴ) 37 | Γ ⇒s Δ = ∀ (x : var Δ) → Term Γ (sort-of Δ x) 38 | 39 | infix 4 _⇒s_ 40 | 41 | -- identity substitution 42 | id-s : ∀ {Γ : Context} → Γ ⇒s Γ 43 | id-s = tm-var 44 | 45 | -- the action of a substitution on a term (contravariant) 46 | _[_]s : ∀ {Γ Δ : Context} {A : sort} → Term Δ A → Γ ⇒s Δ → Term Γ A 47 | (tm-var x) [ σ ]s = σ x 48 | (tm-oper f ts) [ σ ]s = tm-oper f (λ i → ts i [ σ ]s) 49 | 50 | infixr 6 _[_]s 51 | 52 | -- composition of substitutions 53 | _∘s_ : ∀ {Γ Δ Θ : Context} → Δ ⇒s Θ → Γ ⇒s Δ → Γ ⇒s Θ 54 | (σ ∘s ρ) x = σ x [ ρ ]s 55 | 56 | infixl 7 _∘s_ 57 | 58 | -- Equations 59 | record Equation {s o} (Σ : Signature {s} {o} ) : Set (lsuc (s ⊔ o)) where 60 | constructor make-eq 61 | field 62 | eq-ctx : Signature.Context Σ 63 | eq-sort : Signature.sort Σ 64 | eq-lhs : Signature.Term Σ eq-ctx eq-sort 65 | eq-rhs : Signature.Term Σ eq-ctx eq-sort 66 | 67 | infix 5 make-eq 68 | 69 | syntax make-eq Γ A s t = Γ ∥ s ≈ t ⦂ A 70 | -- Theory 71 | -- an equational theory is a family of axioms over a given sort 72 | record Theory ℓ {𝓈 ℴ} (Σ : Signature {𝓈} {ℴ}) : Set (lsuc (ℓ ⊔ 𝓈 ⊔ ℴ)) where 73 | open Signature Σ public 74 | field 75 | ax : Set ℓ -- the axioms 76 | ax-eq : ax → Equation Σ 77 | 78 | ax-ctx : ax → Context 79 | ax-ctx ε = Equation.eq-ctx (ax-eq ε) 80 | 81 | ax-sort : ax → sort 82 | ax-sort ε = Equation.eq-sort (ax-eq ε) 83 | 84 | ax-lhs : ∀ (ε : ax) → Term (ax-ctx ε) (ax-sort ε) 85 | ax-lhs ε = Equation.eq-lhs (ax-eq ε) 86 | 87 | ax-rhs : ∀ (ε : ax) → Term (ax-ctx ε) (ax-sort ε) 88 | ax-rhs ε = Equation.eq-rhs (ax-eq ε) 89 | 90 | -- equality of terms 91 | infix 4 ⊢_ 92 | 93 | data ⊢_ : Equation Σ → Set (lsuc (ℓ ⊔ 𝓈 ⊔ ℴ)) where 94 | -- general rules 95 | eq-refl : ∀ {Γ A} {t : Term Γ A} → ⊢ Γ ∥ t ≈ t ⦂ A 96 | eq-symm : ∀ {Γ A} {s t : Term Γ A} → ⊢ Γ ∥ s ≈ t ⦂ A → ⊢ Γ ∥ t ≈ s ⦂ A 97 | eq-tran : ∀ {Γ A} {s t u : Term Γ A} → ⊢ Γ ∥ s ≈ t ⦂ A → ⊢ Γ ∥ t ≈ u ⦂ A → ⊢ Γ ∥ s ≈ u ⦂ A 98 | -- congruence rule 99 | eq-congr : ∀ {Γ} {f : oper} {xs ys : ∀ (i : arg (oper-arity f)) → Term Γ (sort-of (oper-arity f) i)} → 100 | (∀ i → ⊢ Γ ∥ (xs i) ≈ (ys i) ⦂ (sort-of (oper-arity f) i)) → ⊢ Γ ∥ (tm-oper f xs) ≈ (tm-oper f ys) ⦂ (oper-sort f) 101 | -- equational axiom 102 | eq-axiom : ∀ (ε : ax) {Γ : Context} (σ : Γ ⇒s ax-ctx ε) → 103 | ⊢ Γ ∥ (ax-lhs ε [ σ ]s) ≈ (ax-rhs ε [ σ ]s) ⦂ (ax-sort ε) 104 | 105 | -- the action of the identity substitution is the identity 106 | id-action : ∀ {Γ : Context} {A} {a : Term Γ A} → (⊢ Γ ∥ a ≈ (a [ id-s ]s) ⦂ A) 107 | id-action {a = tm-var a} = eq-refl 108 | id-action {a = tm-oper f x} = eq-congr (λ i → id-action {a = x i}) 109 | 110 | eq-axiom-id : ∀ (ε : ax) → ⊢ (ax-ctx ε ∥ ax-lhs ε ≈ ax-rhs ε ⦂ (ax-sort ε)) 111 | eq-axiom-id ε = eq-tran id-action (eq-tran (eq-axiom ε id-s) (eq-symm id-action)) 112 | 113 | eq-setoid : ∀ (Γ : Context) (A : sort) → Setoid (lsuc ℴ) (lsuc (ℓ ⊔ ℴ ⊔ 𝓈)) 114 | eq-setoid Γ A = 115 | record 116 | { Carrier = Term Γ A 117 | ; _≈_ = λ s t → (⊢ Γ ∥ s ≈ t ⦂ A) 118 | ; isEquivalence = 119 | record 120 | { refl = eq-refl 121 | ; sym = eq-symm 122 | ; trans = eq-tran 123 | } 124 | } 125 | -------------------------------------------------------------------------------- /src/MultiSorted/Completeness.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_; lsuc; lzero) 2 | 3 | import Categories.Category as Category 4 | import Categories.Category.Cartesian as Cartesian 5 | 6 | import MultiSorted.Model as Model 7 | import MultiSorted.Interpretation as Interpretation 8 | import MultiSorted.UniversalModel as UniversalModel 9 | import MultiSorted.SyntacticCategory as SyntacticCategory 10 | import MultiSorted.UniversalModel as UniversalModel 11 | 12 | open import MultiSorted.AlgebraicTheory 13 | 14 | module MultiSorted.Completeness 15 | {ℓt} 16 | {𝓈 ℴ} 17 | {Σ : Signature {𝓈} {ℴ}} 18 | (T : Theory ℓt Σ) where 19 | 20 | open Theory T 21 | open UniversalModel T 22 | 23 | -- An equation is semantically valid when it holds in all models 24 | valid : ∀ (ε : Equation Σ) → Set (lsuc (lsuc ℓt ⊔ lsuc 𝓈 ⊔ lsuc ℴ)) 25 | valid ε = ∀ {𝒞 : Category.Category 𝓈 (lsuc ℴ) (lsuc (ℓt ⊔ 𝓈 ⊔ ℴ))} 26 | {cartesian-𝒞 : Cartesian.Cartesian 𝒞} 27 | {I : Interpretation.Interpretation Σ cartesian-𝒞} 28 | (M : Model.Model T I) → Interpretation.Interpretation.⊨_ I ε 29 | 30 | -- Completeness: semantic validity implies provability 31 | valid-⊢ : ∀ (ε : Equation Σ) → valid ε → ⊢ ε 32 | valid-⊢ ε v = universality ε (v 𝒰) 33 | -------------------------------------------------------------------------------- /src/MultiSorted/Context.agda: -------------------------------------------------------------------------------- 1 | module MultiSorted.Context {s} (Sort : Set s) where 2 | 3 | -- An attempt to define more structured context 4 | -- that directly support the cartesian structure 5 | 6 | data Context : Set s where 7 | ctx-empty : Context 8 | ctx-slot : Sort → Context 9 | ctx-concat : Context → Context → Context 10 | 11 | -- the variables in a context 12 | data var : Context → Set where 13 | var-var : ∀ {A} → var (ctx-slot A) 14 | var-inl : ∀ {Γ Δ} → var Γ → var (ctx-concat Γ Δ) 15 | var-inr : ∀ {Γ Δ} → var Δ → var (ctx-concat Γ Δ) 16 | 17 | sort-of : ∀ (Γ : Context) → var Γ → Sort 18 | sort-of (ctx-slot A) _ = A 19 | sort-of (ctx-concat Γ Δ) (var-inl x) = sort-of Γ x 20 | sort-of (ctx-concat Γ Δ) (var-inr x) = sort-of Δ x 21 | 22 | -- It is absurd to have a variable in the empty context 23 | ctx-empty-absurd : ∀ {ℓ} {P : var ctx-empty → Set ℓ} (x : var ctx-empty) → P x 24 | ctx-empty-absurd () 25 | -------------------------------------------------------------------------------- /src/MultiSorted/FromSingleSorted.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; subst) 3 | 4 | open import Relation.Binary 5 | 6 | import MultiSorted.Context as MSC 7 | import MultiSorted.AlgebraicTheory as MST 8 | import SingleSorted.AlgebraicTheory 9 | 10 | module MultiSorted.FromSingleSorted 11 | (Σ : SingleSorted.AlgebraicTheory.Signature) 12 | (𝒯 : SingleSorted.AlgebraicTheory.Theory lzero Σ) 13 | where 14 | 15 | module SS where 16 | open SingleSorted.AlgebraicTheory public 17 | open SingleSorted.AlgebraicTheory.Signature Σ public 18 | 19 | data 𝒜 : Set where 20 | A : 𝒜 21 | 22 | single-sort : ∀ (X : 𝒜) → X ≡ A 23 | single-sort A = refl 24 | 25 | 26 | -- We have to transform the following constructions from the single sorted to the multi sorted setting 27 | -- contexts, variables, operations, terms, equations, theory 28 | 29 | 30 | singleSortedToMultiSortedContext : ∀ {𝓈} {Sort : Set 𝓈} (X : Sort) → SS.Context → MSC.Context Sort 31 | singleSortedToMultiSortedContext _ SS.ctx-empty = MSC.Context.ctx-empty 32 | singleSortedToMultiSortedContext X SS.ctx-slot = MSC.Context.ctx-slot X 33 | singleSortedToMultiSortedContext X (SS.ctx-concat Γ Δ) = 34 | MSC.Context.ctx-concat (singleSortedToMultiSortedContext X Γ) (singleSortedToMultiSortedContext X Δ) 35 | 36 | multiSortedToSingleSortedContext : ∀ {𝓈} {Sort : Set 𝓈} (X : Sort) → MSC.Context Sort → SS.Context 37 | multiSortedToSingleSortedContext X MSC.ctx-empty = SS.ctx-empty 38 | multiSortedToSingleSortedContext X (MSC.ctx-slot x) = SS.ctx-slot 39 | multiSortedToSingleSortedContext X (MSC.ctx-concat Γ Δ) = 40 | SS.ctx-concat (multiSortedToSingleSortedContext X Γ) (multiSortedToSingleSortedContext X Δ) 41 | 42 | 43 | S : MST.Signature 44 | S = record { sort = 𝒜 45 | ; oper = SS.Signature.oper Σ 46 | ; oper-arity = λ{ f → singleSortedToMultiSortedContext A (SS.Signature.oper-arity Σ f) } 47 | ; oper-sort = λ{ f → A } 48 | } 49 | 50 | module MS where 51 | open MST.Signature S public 52 | 53 | 54 | singleSortedToMultiSortedVar : ∀ {Γ : SS.Context} 55 | → SS.var Γ → MS.var (singleSortedToMultiSortedContext A Γ) 56 | singleSortedToMultiSortedVar SS.var-var = MS.var-var 57 | singleSortedToMultiSortedVar (SS.var-inl x) = MS.var-inl (singleSortedToMultiSortedVar x) 58 | singleSortedToMultiSortedVar (SS.var-inr x) = MS.var-inr (singleSortedToMultiSortedVar x) 59 | 60 | 61 | coerce : ∀ {Γ : MS.Context} {X} {Y} → MS.Term Γ X → MS.Term Γ Y 62 | coerce {Γ} {A} {A} t = subst (MS.Term Γ) refl t 63 | 64 | argToArg : ∀ {Δ : SS.Context} → MS.arg (singleSortedToMultiSortedContext A Δ) → SS.arg Δ 65 | argToArg {SS.ctx-slot} i = SS.var-var 66 | argToArg {SS.ctx-concat Δ Γ} (MSC.var-inl i) = SS.var-inl (argToArg i) 67 | argToArg {SS.ctx-concat Δ Γ} (MSC.var-inr i) = SS.var-inr (argToArg i) 68 | 69 | singleSortedToMultiSortedTerm : ∀ {Γ : SS.Context} 70 | → SS.Term Γ 71 | → MS.Term (singleSortedToMultiSortedContext A Γ) A 72 | singleSortedToMultiSortedTerm {Γ} (SS.tm-var x) = coerce (MS.tm-var (singleSortedToMultiSortedVar x)) 73 | singleSortedToMultiSortedTerm (SS.tm-oper f ts) = MS.tm-oper f 74 | (λ i → coerce (singleSortedToMultiSortedTerm (ts (argToArg i)))) 75 | 76 | singleSortedToMultiSortedEquation : SS.Equation → MST.Equation S 77 | singleSortedToMultiSortedEquation eq = 78 | MST.make-eq 79 | (singleSortedToMultiSortedContext A (SS.Equation.eq-ctx eq)) 80 | A 81 | (singleSortedToMultiSortedTerm (SS.Equation.eq-lhs eq)) 82 | (singleSortedToMultiSortedTerm (SS.Equation.eq-rhs eq)) 83 | 84 | 𝓣 : MST.Theory lzero S 85 | 𝓣 = record { ax = SS.Theory.ax 𝒯 86 | ; ax-eq = λ{ a → singleSortedToMultiSortedEquation (SS.Theory.ax-eq 𝒯 a)} 87 | } 88 | -------------------------------------------------------------------------------- /src/MultiSorted/Group.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality using (_≡_; refl) 3 | open import SingleSorted.Substitution 4 | open import Data.Nat using (ℕ; zero; suc) 5 | import MultiSorted.Context as Context 6 | 7 | module MultiSorted.Group where 8 | 9 | data 𝒜 : Set₀ where 10 | A : 𝒜 11 | 12 | single-sort : ∀ (X : 𝒜) → X ≡ A 13 | single-sort A = refl 14 | 15 | open import MultiSorted.AlgebraicTheory 16 | 17 | open import MultiSorted.AlgebraicTheory 18 | data GroupOp : Set where 19 | e : GroupOp 20 | inv : GroupOp 21 | mul : GroupOp 22 | 23 | ctx : ∀ (n : ℕ) → Context.Context 𝒜 24 | ctx zero = Context.ctx-empty 25 | ctx (suc n) = Context.ctx-concat (ctx n) (Context.ctx-slot A) 26 | 27 | -- the signature of the theory of small groups 28 | -- has one constant, one unary operation, one binary operation 29 | Σ : Signature {lzero} {lzero} 30 | Σ = record { sort = 𝒜 31 | ; oper = GroupOp 32 | ; oper-arity = λ{ e → ctx 0 ; inv → ctx 1 ; mul → ctx 2} 33 | ; oper-sort = λ{ e → A ; inv → A ; mul → A} 34 | } 35 | 36 | open Signature Σ 37 | 38 | singleton-context : (var (ctx-slot A)) → var (ctx-concat ctx-empty (ctx-slot A)) 39 | singleton-context (var-var {A}) = var-inr (var-var {A}) 40 | 41 | single-sort-context : ∀ {Γ : Context} (x : var Γ) → sort-of Γ x ≡ A 42 | single-sort-context {Γ} x = single-sort (Context.sort-of 𝒜 Γ x) 43 | 44 | single-sort-terms : ∀ {X : 𝒜} {Γ : Context} → Term Γ X ≡ Term Γ A 45 | single-sort-terms {A} = refl 46 | 47 | σ : ∀ {Γ : Context} {t : Term Γ A} → Γ ⇒s (ctx 1) 48 | σ {Γ} {t} = λ{ (var-inr var-var) → t} 49 | 50 | δ : ∀ {Γ : Context} {t : Term Γ A} {s : Term Γ A} → Γ ⇒s (ctx 2) 51 | δ {Γ} {t} {s} = sub 52 | where 53 | sub : Γ ⇒s (ctx 2) 54 | sub (var-inl x) rewrite (single-sort-terms {(sort-of (ctx 2) (var-inl x))} {Γ}) = t 55 | sub (var-inr y) rewrite (single-sort-terms {(sort-of (ctx 2) (var-inr y))} {Γ}) = s 56 | 57 | -- helper functions for creating terms 58 | e' : ∀ {Γ : Context} → Term Γ A 59 | e' {Γ} = tm-oper e λ() 60 | 61 | _∗_ : ∀ {Γ} → Term Γ A → Term Γ A → Term Γ A 62 | t ∗ s = tm-oper mul λ{ xs → δ {t = t} {s = s} xs} 63 | 64 | _ⁱ : ∀ {Γ : Context} → Term Γ A → Term Γ A 65 | t ⁱ = tm-oper inv λ{ x → σ {t = t} x} 66 | 67 | -- _∗_ : ∀ {Γ} → Term Γ A → Term Γ A → Term Γ A 68 | -- _∗_ {Γ} t s = tm-oper mul λ{ (var-inl i) → {!!} ; (var-inr i) → {!!}} 69 | 70 | -- _ⁱ : ∀ {Γ : Context} → Term Γ A → Term Γ A 71 | -- t ⁱ = tm-oper inv λ{ x → t } 72 | 73 | infixl 5 _∗_ 74 | infix 6 _ⁱ 75 | 76 | _ : Term (ctx 2) A 77 | _ = tm-var (var-inl (var-inr var-var)) ∗ tm-var (var-inr var-var) 78 | 79 | 80 | -- group equations 81 | data GroupEq : Set where 82 | mul-assoc e-left e-right inv-left inv-right : GroupEq 83 | 84 | mul-assoc-eq : Equation Σ 85 | e-left-eq : Equation Σ 86 | e-right-eq : Equation Σ 87 | inv-left-eq : Equation Σ 88 | inv-right-eq : Equation Σ 89 | 90 | mul-assoc-eq = record { eq-ctx = ctx 3 91 | ; eq-lhs = x ∗ y ∗ z 92 | ; eq-rhs = x ∗ (y ∗ z) 93 | } 94 | where 95 | x : Term (ctx 3) A 96 | y : Term (ctx 3) A 97 | z : Term (ctx 3) A 98 | x = tm-var (var-inl (var-inl (var-inr var-var))) 99 | y = tm-var (var-inl (var-inr var-var)) 100 | z = tm-var (var-inr var-var) 101 | 102 | e-left-eq = record { eq-ctx = ctx 1 ; eq-lhs = e' ∗ x ; eq-rhs = x } 103 | where 104 | x : Term (ctx 1) A 105 | x = tm-var (var-inr var-var) 106 | 107 | e-right-eq = record { eq-ctx = ctx 1 ; eq-lhs = x ∗ e' ; eq-rhs = x } 108 | where 109 | x : Term (ctx 1) A 110 | x = tm-var (var-inr var-var) 111 | 112 | 113 | inv-left-eq = record { eq-ctx = ctx 1 ; eq-lhs = x ⁱ ∗ x ; eq-rhs = e' } 114 | where 115 | x : Term (ctx 1) A 116 | x = tm-var (var-inr var-var) 117 | 118 | inv-right-eq = record { eq-ctx = ctx 1 ; eq-lhs = x ∗ x ⁱ ; eq-rhs = e' } 119 | where 120 | x : Term (ctx 1) A 121 | x = tm-var (var-inr var-var) 122 | 123 | 124 | 𝒢 : Theory lzero Σ 125 | 𝒢 = record { ax = GroupEq 126 | ; ax-eq = λ{ mul-assoc → mul-assoc-eq 127 | ; e-left → e-left-eq 128 | ; e-right → e-right-eq 129 | ; inv-left → inv-left-eq 130 | ; inv-right → inv-right-eq 131 | } 132 | } 133 | -------------------------------------------------------------------------------- /src/MultiSorted/Interpretation.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_) 2 | 3 | import Categories.Category as Category 4 | import Categories.Category.Cartesian as Cartesian 5 | 6 | open import MultiSorted.AlgebraicTheory 7 | import MultiSorted.Product as Product 8 | 9 | module MultiSorted.Interpretation 10 | {o ℓ e} 11 | {𝓈 ℴ} 12 | (Σ : Signature {𝓈} {ℴ}) 13 | {𝒞 : Category.Category o ℓ e} 14 | (cartesian-𝒞 : Cartesian.Cartesian 𝒞) where 15 | open Signature Σ 16 | open Category.Category 𝒞 17 | 18 | -- An interpretation of Σ in 𝒞 19 | record Interpretation : Set (o ⊔ ℓ ⊔ e ⊔ 𝓈 ⊔ ℴ) where 20 | 21 | field 22 | interp-sort : sort → Obj 23 | interp-ctx : Product.Producted 𝒞 {Σ = Σ} interp-sort 24 | interp-oper : ∀ (f : oper) → Product.Producted.prod interp-ctx (oper-arity f) ⇒ interp-sort (oper-sort f) 25 | 26 | open Product.Producted interp-ctx 27 | 28 | -- the interpretation of a term 29 | interp-term : ∀ {Γ : Context} {A} → Term Γ A → prod Γ ⇒ interp-sort A 30 | interp-term (tm-var x) = π x 31 | interp-term (tm-oper f ts) = interp-oper f ∘ tuple (oper-arity f) (λ i → interp-term (ts i)) 32 | 33 | -- the interpretation of a substitution 34 | interp-subst : ∀ {Γ Δ} → Γ ⇒s Δ → prod Γ ⇒ prod Δ 35 | interp-subst {Γ} {Δ} σ = tuple Δ λ i → interp-term (σ i) 36 | 37 | -- the equality of interpretations 38 | ⊨_ : (ε : Equation Σ) → Set e 39 | open Equation 40 | ⊨ ε = interp-term (eq-lhs ε) ≈ interp-term (eq-rhs ε) 41 | 42 | -- interpretation commutes with substitution 43 | open HomReasoning 44 | 45 | interp-[]s : ∀ {Γ Δ} {A} {t : Term Δ A} {σ : Γ ⇒s Δ} → 46 | interp-term (t [ σ ]s) ≈ interp-term t ∘ interp-subst σ 47 | interp-[]s {Γ} {Δ} {A} {tm-var x} {σ} = ⟺ (project {Γ = Δ}) 48 | interp-[]s {Γ} {Δ} {A} {tm-oper f ts} {σ} = (∘-resp-≈ʳ 49 | (tuple-cong 50 | {fs = λ i → interp-term (ts i [ σ ]s)} 51 | {gs = λ z → interp-term (ts z) ∘ interp-subst σ} 52 | (λ i → interp-[]s {t = ts i} {σ = σ}) 53 | ○ (∘-distribʳ-tuple 54 | {Γ = oper-arity f} 55 | {fs = λ z → interp-term (ts z)} 56 | {g = interp-subst σ}))) 57 | ○ (Equiv.refl ○ sym-assoc) 58 | 59 | -- -- Every signature has the trivial interpretation 60 | 61 | open Product 𝒞 62 | 63 | Trivial : Interpretation 64 | Trivial = 65 | let open Cartesian.Cartesian cartesian-𝒞 in 66 | record 67 | { interp-sort = (λ _ → ⊤) 68 | ; interp-ctx = StandardProducted (λ _ → ⊤) cartesian-𝒞 69 | ; interp-oper = λ f → ! } 70 | 71 | record _⇒I_ (I J : Interpretation) : Set (o ⊔ ℓ ⊔ e ⊔ 𝓈 ⊔ ℴ) where 72 | open Interpretation 73 | open Producted 74 | 75 | field 76 | hom-morphism : ∀ {A} → interp-sort I A ⇒ interp-sort J A 77 | hom-commute : 78 | ∀ (f : oper) → 79 | hom-morphism ∘ interp-oper I f ≈ 80 | interp-oper J f ∘ tuple (interp-ctx J) (oper-arity f) (λ i → hom-morphism ∘ π (interp-ctx I) i) 81 | 82 | infix 4 _⇒I_ 83 | 84 | -- The identity homomorphism 85 | id-I : ∀ {A : Interpretation} → A ⇒I A 86 | id-I {A} = 87 | let open Interpretation A in 88 | let open HomReasoning in 89 | let open Producted interp-sort in 90 | record 91 | { hom-morphism = id 92 | ; hom-commute = λ f → 93 | begin 94 | (id ∘ interp-oper f) ≈⟨ identityˡ ⟩ 95 | interp-oper f ≈˘⟨ identityʳ ⟩ 96 | (interp-oper f ∘ id) ≈˘⟨ refl⟩∘⟨ unique interp-ctx (λ i → identityʳ ○ ⟺ identityˡ) ⟩ 97 | (interp-oper f ∘ 98 | Product.Producted.tuple interp-ctx (oper-arity f) 99 | (λ i → id ∘ Product.Producted.π interp-ctx i)) ∎ 100 | } 101 | 102 | -- Compositon of homomorphisms 103 | _∘I_ : ∀ {A B C : Interpretation} → B ⇒I C → A ⇒I B → A ⇒I C 104 | _∘I_ {A} {B} {C} ϕ ψ = 105 | let open _⇒I_ in 106 | record { hom-morphism = hom-morphism ϕ ∘ hom-morphism ψ 107 | ; hom-commute = 108 | let open Interpretation in 109 | let open Producted in 110 | let open HomReasoning in 111 | λ f → 112 | begin 113 | (((hom-morphism ϕ) ∘ hom-morphism ψ) ∘ interp-oper A f) ≈⟨ assoc ⟩ 114 | (hom-morphism ϕ ∘ hom-morphism ψ ∘ interp-oper A f) ≈⟨ (refl⟩∘⟨ hom-commute ψ f) ⟩ 115 | (hom-morphism ϕ ∘ 116 | interp-oper B f ∘ 117 | tuple (interp-ctx B) (oper-arity f) 118 | (λ i → hom-morphism ψ ∘ π (interp-ctx A) i)) ≈˘⟨ assoc ⟩ 119 | ((hom-morphism ϕ ∘ interp-oper B f) ∘ 120 | tuple (interp-ctx B) (oper-arity f) 121 | (λ i → hom-morphism ψ ∘ π (interp-ctx A) i)) ≈⟨ (hom-commute ϕ f ⟩∘⟨refl) ⟩ 122 | ((interp-oper C f ∘ 123 | tuple (interp-ctx C) (oper-arity f) 124 | (λ i → hom-morphism ϕ ∘ π (interp-ctx B) i)) 125 | ∘ 126 | tuple (interp-ctx B) (oper-arity f) 127 | (λ i → hom-morphism ψ ∘ π (interp-ctx A) i)) ≈⟨ assoc ⟩ 128 | (interp-oper C f ∘ 129 | tuple (interp-ctx C) (oper-arity f) 130 | (λ i → hom-morphism ϕ ∘ π (interp-ctx B) i) 131 | ∘ 132 | tuple (interp-ctx B) (oper-arity f) 133 | (λ i → hom-morphism ψ ∘ π (interp-ctx A) i)) ≈⟨ (refl⟩∘⟨ ⟺ (∘-distribʳ-tuple (interp-sort C) (interp-ctx C))) ⟩ 134 | (interp-oper C f ∘ 135 | tuple (interp-ctx C) (oper-arity f) 136 | (λ x → 137 | (hom-morphism ϕ ∘ π (interp-ctx B) x) ∘ 138 | tuple (interp-ctx B) (oper-arity f) 139 | (λ i → hom-morphism ψ ∘ π (interp-ctx A) i))) ≈⟨ (refl⟩∘⟨ tuple-cong (interp-sort C) (interp-ctx C) λ i → assoc) ⟩ 140 | (interp-oper C f ∘ 141 | tuple (interp-ctx C) (oper-arity f) 142 | (λ z → 143 | hom-morphism ϕ ∘ 144 | π (interp-ctx B) z ∘ 145 | tuple (interp-ctx B) (oper-arity f) 146 | (λ i → hom-morphism ψ ∘ π (interp-ctx A) i))) ≈⟨ (refl⟩∘⟨ tuple-cong (interp-sort C) (interp-ctx C) λ i → refl⟩∘⟨ project (interp-ctx B)) ⟩ 147 | (interp-oper C f ∘ 148 | tuple (interp-ctx C) (oper-arity f) 149 | (λ z → hom-morphism ϕ ∘ hom-morphism ψ ∘ π (interp-ctx A) z)) ≈⟨ (refl⟩∘⟨ tuple-cong (interp-sort C) (interp-ctx C) λ i → sym-assoc) ⟩ 150 | (interp-oper C f ∘ 151 | tuple (interp-ctx C) (oper-arity f) 152 | (λ z → (hom-morphism ϕ ∘ hom-morphism ψ) ∘ π (interp-ctx A) z)) ∎} 153 | -------------------------------------------------------------------------------- /src/MultiSorted/Model.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_) 2 | 3 | import Categories.Category as Category 4 | import Categories.Category.Cartesian as Cartesian 5 | import MultiSorted.Interpretation as Interpretation 6 | 7 | open import MultiSorted.AlgebraicTheory 8 | open import MultiSorted.Substitution 9 | import MultiSorted.Product as Product 10 | 11 | module MultiSorted.Model {o ℓ e ℓt} 12 | {𝓈 ℴ} 13 | {Σ : Signature {𝓈} {ℴ}} 14 | (T : Theory ℓt Σ) 15 | {𝒞 : Category.Category o ℓ e} 16 | {cartesian-𝒞 : Cartesian.Cartesian 𝒞} where 17 | 18 | -- Model of a theory 19 | 20 | record Is-Model (I : Interpretation.Interpretation Σ cartesian-𝒞) : Set (ℓt ⊔ o ⊔ ℓ ⊔ e) where 21 | 22 | open Theory T 23 | open Category.Category 𝒞 24 | open Interpretation.Interpretation I 25 | open HomReasoning 26 | 27 | field 28 | model-eq : ∀ (ε : ax) → ⊨ ax-eq (ε) 29 | 30 | -- Soundness of semantics 31 | module _ where 32 | open Product.Producted interp-ctx 33 | 34 | -- first we show that substitution preserves validity 35 | model-resp-[]s : ∀ {Γ Δ} {A} {u v : Term Γ A} {σ : Δ ⇒s Γ} → 36 | interp-term u ≈ interp-term v → interp-term (u [ σ ]s) ≈ interp-term (v [ σ ]s) 37 | model-resp-[]s {u = u} {v = v} {σ = σ} ξ = 38 | begin 39 | interp-term (u [ σ ]s) ≈⟨ interp-[]s {t = u} ⟩ 40 | (interp-term u ∘ interp-subst σ) ≈⟨ ξ ⟩∘⟨refl ⟩ 41 | (interp-term v ∘ interp-subst σ) ≈˘⟨ interp-[]s {t = v} ⟩ 42 | interp-term (v [ σ ]s) ∎ 43 | 44 | -- the soundness statement 45 | ⊢-⊨ : ∀ {ε : Equation Σ} → ⊢ ε → ⊨ ε 46 | ⊢-⊨ eq-refl = Equiv.refl 47 | ⊢-⊨ (eq-symm ξ) = ⟺ (⊢-⊨ ξ) 48 | ⊢-⊨ (eq-tran ξ θ) = ⊢-⊨ ξ ○ ⊢-⊨ θ 49 | ⊢-⊨ (eq-congr ξ) = ∘-resp-≈ʳ (unique λ i → project ○ ⟺ (⊢-⊨ (ξ i))) 50 | ⊢-⊨ (eq-axiom ε σ) = model-resp-[]s {u = ax-lhs ε} {v = ax-rhs ε} (model-eq ε) 51 | 52 | -- Every theory has the trivial model, whose carrier is the terminal object 53 | Trivial : Is-Model (Interpretation.Trivial Σ cartesian-𝒞) 54 | Trivial = 55 | let open Cartesian.Cartesian cartesian-𝒞 in 56 | record { model-eq = λ ε → !-unique₂ } 57 | -------------------------------------------------------------------------------- /src/MultiSorted/ModelCategory.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_ ; lsuc ; Level) 2 | 3 | import Categories.Category as Category 4 | import Categories.Category.Cartesian as Cartesian 5 | open import Categories.Object.Terminal using (Terminal) 6 | open import Categories.Object.Product using (Product) 7 | 8 | open import MultiSorted.AlgebraicTheory 9 | open import MultiSorted.Substitution 10 | import MultiSorted.Product as Product 11 | import MultiSorted.Interpretation as Interpretation 12 | import MultiSorted.Model as Is-Model 13 | import MultiSorted.InterpretationCategory as InterpretationCategory 14 | 15 | module MultiSorted.ModelCategory 16 | {ℓt o ℓ e} 17 | {𝓈 ℴ} 18 | {Σ : Signature {𝓈} {ℴ}} 19 | (T : Theory ℓt Σ) 20 | {𝒞 : Category.Category o ℓ e} 21 | (cartesian-𝒞 : Cartesian.Cartesian 𝒞) where 22 | open Signature Σ 23 | open Theory T 24 | open Category.Category 𝒞 25 | open Interpretation Σ cartesian-𝒞 26 | open Is-Model {o = o} {ℓ = ℓ} {e = e} {Σ = Σ} T 27 | 28 | -- Useful shortcuts for levels 29 | ℓℴ : Level 30 | ℓℴ = lsuc (o ⊔ ℓ ⊔ e ⊔ 𝓈 ⊔ ℴ ⊔ ℓt) 31 | 32 | ℓ𝒽 : Level 33 | ℓ𝒽 = ℓ ⊔ o ⊔ e ⊔ 𝓈 ⊔ ℴ 34 | 35 | ℓ𝓇 : Level 36 | ℓ𝓇 = e ⊔ 𝓈 37 | 38 | -- New definition of models (as a set, not only a property of interpretations) 39 | record Model : Set ℓℴ where 40 | field 41 | interpretation : Interpretation 42 | is-model : Is-Model.Is-Model T interpretation 43 | 44 | open Model 45 | 46 | -- Homomorphisms of models 47 | _⇒M_ : ∀ (M N : Model) → Set ℓ𝒽 48 | _⇒M_ M N = (interpretation M) ⇒I (interpretation N) 49 | 50 | -- Equality of homomorphisms of models (the same as for the interpretations) 51 | _≈M_ : ∀ {M N : Model} → M ⇒M N → M ⇒M N → Set ℓ𝓇 52 | _≈M_ {M} {N} ϕ ψ = 53 | let open _⇒I_ in 54 | ∀ A → (hom-morphism ϕ {A}) ≈ (hom-morphism ψ) 55 | 56 | -- The identity morphism on models 57 | id-M : (M : Model) → M ⇒M M 58 | id-M = λ M → id-I {interpretation M} 59 | 60 | -- Composition of morphisms of Models 61 | _∘M_ : ∀ {M N O : Model} → N ⇒M O → M ⇒M N → M ⇒M O 62 | _∘M_ ϕ ψ = ϕ ∘I ψ 63 | 64 | 65 | -- The category of Models of Σ in 𝒞 66 | ℳ : Category.Category ℓℴ ℓ𝒽 ℓ𝓇 67 | ℳ = record 68 | { Obj = Model 69 | ; _⇒_ = _⇒M_ 70 | ; _≈_ = λ {M} {N} ϕ ψ → _≈M_ {M} {N} ϕ ψ 71 | ; id = λ {M} → id-M M 72 | ; _∘_ = λ {M} {N} {O} ϕ ψ → _∘M_ {M} {N} {O} ϕ ψ 73 | ; assoc = λ A → assoc -- λ A → assoc 74 | ; sym-assoc = λ A → sym-assoc 75 | ; identityˡ = λ A → identityˡ 76 | ; identityʳ = λ A → identityʳ 77 | ; identity² = λ A → identity² 78 | ; equiv = record { refl = λ A → Equiv.refl 79 | ; sym = λ p A → Equiv.sym (p A) 80 | ; trans = λ p₁ p₂ A → Equiv.trans (p₁ A) (p₂ A) 81 | } 82 | ; ∘-resp-≈ = λ p₁ p₂ A → ∘-resp-≈ (p₁ A) (p₂ A) 83 | } 84 | 85 | -- The product of "Model proofs" 86 | 87 | module _ (M N : Model) where 88 | open Product.Producted 89 | open HomReasoning 90 | open InterpretationCategory 91 | open Cartesian.Cartesian cartesian-𝒞 92 | open Interpretation.Interpretation 93 | open import Categories.Object.Product.Morphisms {o} {ℓ} {e} 𝒞 94 | open Equation 95 | 96 | -- A proof that an axiom holds in a product interpretation amounts to a apir of proofs that the axiom holds in each model 97 | is-model-pairs : ∀ ε → (interp-term (interpretation M) (Equation.eq-lhs (ax-eq ε)) ⁂ interp-term (interpretation N) (Equation.eq-lhs (ax-eq ε))) 98 | ≈ (interp-term (interpretation M) (Equation.eq-rhs (ax-eq ε)) ⁂ interp-term (interpretation N) (Equation.eq-rhs (ax-eq ε))) → 99 | Interpretation.interp-term (A×B-ℐ𝓃𝓉 Σ cartesian-𝒞 (interpretation M) (interpretation N)) (Equation.eq-lhs (ax-eq ε)) 100 | ≈ Interpretation.interp-term (A×B-ℐ𝓃𝓉 Σ cartesian-𝒞 (interpretation M) (interpretation N)) (Equation.eq-rhs (ax-eq ε)) 101 | is-model-pairs ε p = 102 | begin 103 | Interpretation.interp-term 104 | (A×B-ℐ𝓃𝓉 Σ cartesian-𝒞 (interpretation M) (interpretation N)) 105 | (Equation.eq-lhs (ax-eq ε)) ≈⟨ ⟺ 106 | (Cartesian.Cartesian.unique cartesian-𝒞 107 | (natural-π₁ Σ cartesian-𝒞 {I = interpretation M} {interpretation N} (ax-lhs ε)) 108 | (natural-π₂ Σ cartesian-𝒞 {I = interpretation M} {interpretation N} (ax-lhs ε))) ⟩ 109 | product.⟨ 110 | Interpretation.interp-term (interpretation M) (eq-lhs (ax-eq ε)) ∘ 111 | π₁ 112 | , 113 | Interpretation.interp-term (interpretation N) (eq-lhs (ax-eq ε)) ∘ 114 | π₂ 115 | ⟩ ≈⟨ ⟨⟩-cong₂ (∘-resp-≈ˡ (Is-Model.model-eq (is-model M) ε)) (∘-resp-≈ˡ (Is-Model.model-eq (is-model N) ε)) ⟩ 116 | product.⟨ 117 | Interpretation.interp-term (interpretation M) (eq-rhs (ax-eq ε)) ∘ 118 | π₁ 119 | , 120 | Interpretation.interp-term (interpretation N) (eq-rhs (ax-eq ε)) ∘ 121 | π₂ 122 | ⟩ ≈⟨ Cartesian.Cartesian.unique cartesian-𝒞 123 | (natural-π₁ Σ cartesian-𝒞 {I = interpretation M} {interpretation N} (ax-rhs ε)) 124 | (natural-π₂ Σ cartesian-𝒞 {I = interpretation M} {interpretation N} (ax-rhs ε)) ⟩ 125 | Interpretation.interp-term 126 | (A×B-ℐ𝓃𝓉 Σ cartesian-𝒞 (interpretation M) (interpretation N)) 127 | (eq-rhs (ax-eq ε)) ∎ 128 | 129 | -- The proof that the product interpetation of two models is a model 130 | is-model-product : Is-Model (A×B-ℐ𝓃𝓉 Σ cartesian-𝒞 (interpretation M) (interpretation N)) 131 | Is-Model.model-eq is-model-product ε = 132 | begin 133 | Interpretation.interp-term 134 | (A×B-ℐ𝓃𝓉 Σ cartesian-𝒞 (interpretation M) (interpretation N)) 135 | (Equation.eq-lhs (ax-eq ε)) ≈⟨ is-model-pairs ε (⁂-cong₂ (Is-Model.model-eq (is-model M) ε) (Is-Model.model-eq (is-model N) ε)) ⟩ 136 | Interpretation.interp-term 137 | (A×B-ℐ𝓃𝓉 Σ cartesian-𝒞 (interpretation M) (interpretation N)) 138 | (Equation.eq-rhs (ax-eq ε)) ∎ 139 | 140 | 141 | 142 | 143 | -- The product of ℐ𝓃𝓉 carries over the models : the product of two models is a model 144 | module _ (M N : Model) where 145 | open Product.Producted 146 | open HomReasoning 147 | open InterpretationCategory 148 | A×B-ℳ : Model 149 | A×B-ℳ = record 150 | { interpretation = A×B-ℐ𝓃𝓉 Σ cartesian-𝒞 (interpretation M) (interpretation N) 151 | ; is-model = is-model-product M N 152 | } 153 | 154 | 155 | -- The cartesian structure of the category of models 156 | 157 | module _ {M N : Model} where 158 | import Categories.Object.Product.Core 159 | open Categories.Object.Product.Core.Product 160 | open InterpretationCategory Σ cartesian-𝒞 161 | private 162 | UM UN : Interpretation 163 | UM = interpretation M 164 | UN = interpretation N 165 | UM×UN : Product ℐ𝓃𝓉 UM UN 166 | UM×UN = product-ℐ𝓃𝓉 167 | product-ℳ : Product ℳ M N 168 | -- Structure 169 | A×B product-ℳ = A×B-ℳ M N 170 | π₁ product-ℳ = π₁ UM×UN 171 | π₂ product-ℳ = π₂ UM×UN 172 | ⟨_,_⟩ product-ℳ = ⟨_,_⟩ UM×UN 173 | -- Properties 174 | project₁ product-ℳ {O} {h} {i} = project₁ UM×UN {interpretation O} {h} {i} 175 | project₂ product-ℳ {O} {h} {i} = project₂ UM×UN {interpretation O} {h} {i} 176 | unique product-ℳ {O} {h} {i} {j} = unique UM×UN {interpretation O} {h} {i} {j} 177 | -------------------------------------------------------------------------------- /src/MultiSorted/Product.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_) 2 | 3 | import Categories.Category as Category 4 | import Categories.Category.Cartesian as Cartesian 5 | 6 | open import MultiSorted.AlgebraicTheory 7 | 8 | -- Finite products indexed by contexts 9 | module MultiSorted.Product 10 | {o ℓ e} 11 | (𝒞 : Category.Category o ℓ e) 12 | {𝓈 ℴ} 13 | {Σ : Signature {𝓈} {ℴ}} 14 | (interp-sort : Signature.sort Σ → Category.Category.Obj 𝒞) where 15 | 16 | open Signature Σ 17 | open Category.Category 𝒞 18 | open HomReasoning 19 | 20 | interp-sort-var : {Γ : Context} → var Γ → Obj 21 | interp-sort-var {Γ} x = interp-sort (sort-of Γ x) 22 | 23 | record Producted : Set (o ⊔ ℓ ⊔ e ⊔ 𝓈) where 24 | 25 | field 26 | prod : Context → Obj 27 | π : ∀ {Γ} (x : var Γ) → prod Γ ⇒ interp-sort-var x 28 | tuple : ∀ Γ {B} → ((x : var Γ) → B ⇒ interp-sort-var x) → B ⇒ prod Γ 29 | project : ∀ {Γ} {B} {x : var Γ} {fs : (y : var Γ) → B ⇒ interp-sort-var y} → π x ∘ tuple Γ fs ≈ fs x 30 | unique : ∀ {Γ} {B} {fs : (x : var Γ) → B ⇒ interp-sort-var x} {g : B ⇒ prod Γ} → (∀ i → π i ∘ g ≈ fs i) → tuple Γ fs ≈ g 31 | 32 | η : ∀ {Γ} → tuple Γ π ≈ id 33 | η = unique (λ i → identityʳ) 34 | 35 | ! : ∀ {B : Obj} → B ⇒ prod ctx-empty 36 | ! {B} = tuple ctx-empty ctx-empty-absurd 37 | 38 | !-unique : ∀ {B : Obj} {f : B ⇒ prod ctx-empty} → ! ≈ f 39 | !-unique {f = f} = unique ctx-empty-absurd 40 | 41 | !-unique₂ : ∀ {B : Obj} {f g : B ⇒ prod ctx-empty} → f ≈ g 42 | !-unique₂ = (⟺ !-unique) ○ !-unique 43 | 44 | tuple-cong : ∀ {B : Obj} {Γ} {fs gs : (x : var Γ) → B ⇒ interp-sort-var x} → (∀ i → fs i ≈ gs i) → 45 | tuple Γ fs ≈ tuple Γ gs 46 | tuple-cong ξ = unique (λ i → project ○ ⟺ (ξ i)) 47 | 48 | ∘-distribʳ-tuple : ∀ {B C} {Γ} {fs : (x : var Γ) → B ⇒ interp-sort-var x} {g : C ⇒ B} → 49 | tuple Γ (λ x → fs x ∘ g) ≈ tuple Γ fs ∘ g 50 | ∘-distribʳ-tuple = unique (λ i → ⟺ assoc ○ ∘-resp-≈ˡ project) 51 | 52 | -- A cartesian category has a standard products structure (which we need not use) 53 | module _ (cartesian-𝒞 : Cartesian.Cartesian 𝒞) where 54 | open Cartesian.Cartesian cartesian-𝒞 55 | 56 | standard-prod : Context → Obj 57 | standard-prod ctx-empty = ⊤ 58 | standard-prod (ctx-slot A) = interp-sort A 59 | standard-prod (ctx-concat Γ Δ) = standard-prod Γ × standard-prod Δ 60 | 61 | standard-π : ∀ {Γ} (x : var Γ) → standard-prod Γ ⇒ interp-sort-var x 62 | standard-π var-var = id 63 | standard-π (var-inl i) = standard-π i ∘ π₁ 64 | standard-π (var-inr i) = standard-π i ∘ π₂ 65 | 66 | standard-tuple : ∀ Γ {B} → ((x : var Γ) → B ⇒ interp-sort-var x) → B ⇒ standard-prod Γ 67 | standard-tuple ctx-empty fs = ! 68 | standard-tuple (ctx-slot _) fs = fs var-var 69 | standard-tuple (ctx-concat Γ Δ) fs = ⟨ standard-tuple Γ (λ i → fs (var-inl i)) , standard-tuple Δ (λ j → fs (var-inr j)) ⟩ 70 | 71 | standard-project : ∀ {Γ} {B} {x : var Γ} {fs : (x : var Γ) → B ⇒ interp-sort-var x} → 72 | standard-π x ∘ standard-tuple Γ fs ≈ fs x 73 | standard-project {x = var-var} = identityˡ 74 | standard-project {x = var-inl x} = assoc ○ ((∘-resp-≈ʳ project₁) ○ standard-project {x = x}) 75 | standard-project {x = var-inr x} = assoc ○ ((∘-resp-≈ʳ project₂) ○ standard-project {x = x}) 76 | 77 | standard-unique : ∀ {Γ} {B} {fs : (x : var Γ) → B ⇒ interp-sort-var x} {g : B ⇒ standard-prod Γ} → (∀ i → standard-π i ∘ g ≈ fs i) → 78 | standard-tuple Γ fs ≈ g 79 | standard-unique {ctx-empty} ξ = !-unique _ 80 | standard-unique {ctx-slot _} ξ = ⟺ (ξ var-var) ○ identityˡ 81 | standard-unique {ctx-concat Γ Δ} {fs = fs} ξ = 82 | unique 83 | (⟺ (standard-unique (λ x → sym-assoc ○ (ξ (var-inl x))))) 84 | (⟺ (standard-unique (λ y → sym-assoc ○ (ξ (var-inr y))))) 85 | 86 | StandardProducted : Producted 87 | StandardProducted = 88 | record 89 | { prod = standard-prod 90 | ; π = standard-π 91 | ; tuple = standard-tuple 92 | ; project = λ {Γ} → standard-project {Γ} 93 | ; unique = standard-unique } 94 | -------------------------------------------------------------------------------- /src/MultiSorted/Substitution.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lsuc; _⊔_) 2 | open import Data.Fin using (Fin) 3 | 4 | open import MultiSorted.AlgebraicTheory 5 | 6 | module MultiSorted.Substitution {ℓ} {𝓈 ℴ} {Σ : Signature {𝓈} {ℴ}} (T : Theory ℓ Σ) where 7 | 8 | open Theory T 9 | 10 | -- an equality is preserved by the action of the identity 11 | eq-id-action : ∀ {Γ : Context} {A} {u v : Term Γ A} → (⊢ Γ ∥ (u [ id-s ]s) ≈ (v [ id-s ]s) ⦂ A) → (⊢ Γ ∥ u ≈ v ⦂ A) 12 | eq-id-action {u = u} {v = v} p = eq-tran (id-action {a = u}) (eq-tran p (eq-symm (id-action {a = v}))) 13 | 14 | -- equality of substitutions 15 | _≈s_ : ∀ {Γ Δ : Context} → Γ ⇒s Δ → Γ ⇒s Δ → Set (lsuc (ℓ ⊔ 𝓈 ⊔ ℴ)) 16 | _≈s_ {Γ} {Δ} σ ρ = ∀ x → ⊢ Γ ∥ σ x ≈ ρ x ⦂ sort-of Δ x 17 | 18 | infix 5 _≈s_ 19 | 20 | -- reflexivity of the equality of substitutions 21 | refl-subst : ∀ {γ δ : Context} {f : γ ⇒s δ} → f ≈s f 22 | refl-subst = λ x → eq-refl 23 | 24 | -- symmetry of the equality of substitutions 25 | symm-subst : ∀ {Γ Δ : Context} {f g : Γ ⇒s Δ} → f ≈s g → g ≈s f 26 | symm-subst {Γ} {Δ} {f} {g} p = λ x → eq-symm (p x) 27 | 28 | -- transitivity of the equality of substitutions 29 | trans-subst : ∀ {Γ Δ : Context} {f g h : Γ ⇒s Δ} → f ≈s g → g ≈s h → f ≈s h 30 | trans-subst {Γ} {Δ} {f} {g} {h} p q = λ x → eq-tran (p x) (q x) 31 | 32 | -- neutrality of tm-var 33 | tm-var-id : ∀ {Γ} {A} {t : Term Γ A} → ⊢ Γ ∥ t [ id-s ]s ≈ t ⦂ A 34 | tm-var-id {t = tm-var x} = eq-refl 35 | tm-var-id {t = tm-oper f ts} = eq-congr (λ i → tm-var-id) 36 | 37 | -- any two substitutions into the empty context are equal 38 | empty-ctx-subst-unique : ∀ {Γ : Context} {σ ρ : Γ ⇒s ctx-empty} → σ ≈s ρ 39 | empty-ctx-subst-unique () 40 | 41 | -- composition of substitutions is functorial 42 | subst-∘s : ∀ {Γ Δ Θ} {ρ : Δ ⇒s Γ} {σ : Θ ⇒s Δ} {A} (t : Term Γ A) → 43 | ⊢ Θ ∥ (t [ ρ ]s) [ σ ]s ≈ t [ ρ ∘s σ ]s ⦂ A 44 | subst-∘s (tm-var x) = eq-refl 45 | subst-∘s (tm-oper f ts) = eq-congr (λ i → subst-∘s (ts i)) 46 | 47 | -- substitution preserves equality 48 | eq-subst : ∀ {Γ Δ : Context} {σ : Γ ⇒s Δ} {A} {u v : Term Δ A} → 49 | ⊢ Δ ∥ u ≈ v ⦂ A → ⊢ Γ ∥ u [ σ ]s ≈ v [ σ ]s ⦂ A 50 | eq-subst eq-refl = eq-refl 51 | eq-subst (eq-symm ξ) = eq-symm (eq-subst ξ) 52 | eq-subst (eq-tran ζ ξ) = eq-tran (eq-subst ζ) (eq-subst ξ) 53 | eq-subst (eq-congr ξ) = eq-congr (λ i → eq-subst (ξ i)) 54 | eq-subst {σ = σ} (eq-axiom ε ρ) = eq-tran (subst-∘s (ax-lhs ε)) (eq-tran (eq-axiom ε (ρ ∘s σ)) (eq-symm (subst-∘s (ax-rhs ε)))) 55 | 56 | -- equivalent substitutions act the same on terms 57 | equiv-subst : ∀ {Γ Δ : Context} {σ τ : Γ ⇒s Δ} {A} (u : Term Δ A) → 58 | σ ≈s τ → ⊢ Γ ∥ u [ σ ]s ≈ u [ τ ]s ⦂ A 59 | equiv-subst (tm-var x) ξ = ξ x 60 | equiv-subst (tm-oper f ts) ξ = eq-congr (λ i → equiv-subst (ts i) ξ) 61 | 62 | -- equivalent substitution preserve equality 63 | equiv-eq-subst : ∀ {Γ Δ : Context} {σ τ : Γ ⇒s Δ} {A} {u v : Term Δ A} → 64 | σ ≈s τ → ⊢ Δ ∥ u ≈ v ⦂ A → ⊢ Γ ∥ u [ σ ]s ≈ v [ τ ]s ⦂ A 65 | equiv-eq-subst {u = u} p eq-refl = equiv-subst u p 66 | equiv-eq-subst p (eq-symm q) = eq-symm (equiv-eq-subst (symm-subst p) q) 67 | equiv-eq-subst p (eq-tran q r) = eq-tran (eq-subst q) (equiv-eq-subst p r) 68 | equiv-eq-subst p (eq-congr ξ) = eq-congr λ i → equiv-eq-subst p (ξ i) 69 | equiv-eq-subst {σ = σ} {τ = τ} p (eq-axiom ε θ) = eq-tran (eq-subst (eq-axiom ε θ)) (equiv-subst (ax-rhs ε [ θ ]s) p) 70 | 71 | -- the pairing of two substitutions 72 | ⟨_,_⟩s : ∀ {Γ Δ Θ} (σ : Γ ⇒s Δ) (ρ : Γ ⇒s Θ) → Γ ⇒s ctx-concat Δ Θ 73 | ⟨ σ , ρ ⟩s (var-inl x) = σ x 74 | ⟨ σ , ρ ⟩s (var-inr y) = ρ y 75 | 76 | -- composition of substitutions preserves equality 77 | ∘s-resp-≈s : ∀ {Γ Δ Θ} {σ₁ σ₂ : Γ ⇒s Δ} {τ₁ τ₂ : Δ ⇒s Θ} → 78 | τ₁ ≈s τ₂ → σ₁ ≈s σ₂ → τ₁ ∘s σ₁ ≈s τ₂ ∘s σ₂ 79 | ∘s-resp-≈s ξ ζ z = equiv-eq-subst ζ (ξ z) 80 | 81 | -- the action of a substitution on an equation 82 | open Equation 83 | _[_]s-eq : ∀ (ε : Equation Σ) {Γ : Context} ( σ : Γ ⇒s (eq-ctx ε)) → Equation Σ 84 | _[_]s-eq ε {Γ} σ = Γ ∥ ((eq-lhs ε) [ σ ]s) ≈ ((eq-rhs ε) [ σ ]s) ⦂ (eq-sort ε) 85 | -------------------------------------------------------------------------------- /src/MultiSorted/SyntacticCategory.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | import Relation.Binary.Reasoning.Setoid as SetoidR 5 | 6 | import Categories.Category as Category 7 | import Categories.Category.Cartesian as Cartesian 8 | open import Categories.Object.Terminal using (Terminal) 9 | open import Categories.Object.Product using (Product) 10 | 11 | open import MultiSorted.AlgebraicTheory 12 | import MultiSorted.Substitution as Substitution 13 | import MultiSorted.Product as Product 14 | 15 | module MultiSorted.SyntacticCategory 16 | {ℓt} 17 | {𝓈 ℴ} 18 | {Σ : Signature {𝓈} {ℴ}} 19 | (T : Theory ℓt Σ) where 20 | 21 | open Theory T 22 | open Substitution T 23 | 24 | -- The syntactic category 25 | 26 | 𝒮 : Category.Category 𝓈 (lsuc ℴ) (lsuc (ℓt ⊔ 𝓈 ⊔ ℴ)) 27 | 𝒮 = 28 | record 29 | { Obj = Context 30 | ; _⇒_ = _⇒s_ 31 | ; _≈_ = _≈s_ 32 | ; id = id-s 33 | ; _∘_ = _∘s_ 34 | ; assoc = λ {_ _ _ _ _ _ σ} x → subst-∘s (σ x) 35 | ; sym-assoc = λ {_ _ _ _ _ _ σ} x → eq-symm (subst-∘s (σ x)) 36 | ; identityˡ = λ x → eq-refl 37 | ; identityʳ = λ {A B f} x → tm-var-id 38 | ; identity² = λ x → eq-refl 39 | ; equiv = record 40 | { refl = λ x → eq-refl 41 | ; sym = λ ξ y → eq-symm (ξ y) 42 | ; trans = λ ζ ξ y → eq-tran (ζ y) (ξ y)} 43 | ; ∘-resp-≈ = ∘s-resp-≈s 44 | } 45 | 46 | -- We use the product structure which gives back the context directly 47 | prod-𝒮 : Context → Context 48 | prod-𝒮 Γ = Γ 49 | 50 | π-𝒮 : ∀ {Γ} (x : var Γ) → Γ ⇒s ctx-slot (sort-of Γ x) 51 | π-𝒮 x _ = tm-var x 52 | 53 | tuple-𝒮 : ∀ Γ {Δ} → (∀ (x : var Γ) → Δ ⇒s ctx-slot (sort-of Γ x)) → Δ ⇒s Γ 54 | tuple-𝒮 Γ ts = λ x → ts x var-var 55 | 56 | project-𝒮 : ∀ {Γ Δ} {x : var Γ} {ts : ∀ (y : var Γ) → Δ ⇒s ctx-slot (sort-of Γ y)} → 57 | π-𝒮 x ∘s tuple-𝒮 Γ ts ≈s ts x 58 | project-𝒮 {Γ} {Δ} {x} {ts} var-var = eq-refl 59 | 60 | unique-𝒮 : ∀ {Γ Δ} {ts : ∀ (x : var Γ) → Δ ⇒s ctx-slot (sort-of Γ x)} {g : Δ ⇒s Γ} → 61 | (∀ x → π-𝒮 x ∘s g ≈s ts x) → tuple-𝒮 Γ ts ≈s g 62 | unique-𝒮 ξ x = eq-symm (ξ x var-var) 63 | 64 | 65 | producted-𝒮 : Product.Producted 𝒮 {Σ = Σ} ctx-slot 66 | producted-𝒮 = 67 | record 68 | { prod = prod-𝒮 69 | ; π = π-𝒮 70 | ; tuple = tuple-𝒮 71 | ; project = λ {Γ Δ x ts} → project-𝒮 {ts = ts} 72 | ; unique = unique-𝒮 73 | } 74 | 75 | -- The terminal object is the empty context 76 | ⊤-𝒮 : Context 77 | ⊤-𝒮 = ctx-empty 78 | 79 | !-𝒮 : ∀ {Γ} → Γ ⇒s ⊤-𝒮 80 | !-𝒮 () 81 | 82 | !-unique-𝒮 : ∀ {Γ} (σ : Γ ⇒s ⊤-𝒮) → !-𝒮 ≈s σ 83 | !-unique-𝒮 σ () 84 | 85 | terminal-𝒮 : Terminal 𝒮 86 | terminal-𝒮 = 87 | record 88 | { ⊤ = ⊤-𝒮 89 | ; ⊤-is-terminal = 90 | record { ! = !-𝒮 91 | ; !-unique = !-unique-𝒮 } } 92 | 93 | -- Binary product is context contatenation 94 | product-𝒮 : ∀ {Γ Δ} → Product 𝒮 Γ Δ 95 | product-𝒮 {Γ} {Δ} = 96 | record 97 | { A×B = ctx-concat Γ Δ 98 | ; π₁ = λ x → tm-var (var-inl x) 99 | ; π₂ = λ x → tm-var (var-inr x) 100 | ; ⟨_,_⟩ = ⟨_,_⟩s 101 | ; project₁ = λ x → eq-refl 102 | ; project₂ = λ x → eq-refl 103 | ; unique = λ {Θ σ σ₁ σ₂} ξ₁ ξ₂ z → u Θ σ σ₁ σ₂ ξ₁ ξ₂ z 104 | } 105 | where u : ∀ Θ (σ : Θ ⇒s ctx-concat Γ Δ) (σ₁ : Θ ⇒s Γ) (σ₂ : Θ ⇒s Δ) → 106 | ((λ x → σ (var-inl x)) ≈s σ₁) → ((λ y → σ (var-inr y)) ≈s σ₂) → ⟨ σ₁ , σ₂ ⟩s ≈s σ 107 | u Θ σ σ₁ σ₂ ξ₁ ξ₂ (var-inl z) = eq-symm (ξ₁ z) 108 | u Θ σ σ₁ σ₂ ξ₁ ξ₂ (var-inr z) = eq-symm (ξ₂ z) 109 | 110 | -- The cartesian structure of the syntactic category 111 | cartesian-𝒮 : Cartesian.Cartesian 𝒮 112 | cartesian-𝒮 = 113 | record 114 | { terminal = terminal-𝒮 115 | ; products = record { product = product-𝒮 } } 116 | -------------------------------------------------------------------------------- /src/MultiSorted/UniversalInterpretation.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --allow-unsolved-metas #-} 2 | 3 | open import MultiSorted.AlgebraicTheory 4 | 5 | import MultiSorted.Interpretation as Interpretation 6 | import MultiSorted.SyntacticCategory as SyntacticCategory 7 | import MultiSorted.Substitution as Substitution 8 | import Agda.Builtin.Equality as Eq 9 | open import Relation.Binary.PropositionalEquality using (_≡_; refl ; cong) 10 | 11 | 12 | module MultiSorted.UniversalInterpretation 13 | {ℓt} 14 | {𝓈 ℴ} 15 | {Σ : Signature {𝓈} {ℴ}} 16 | (T : Theory ℓt Σ) where 17 | 18 | open Theory T 19 | open Substitution T 20 | open SyntacticCategory T 21 | 22 | -- The universal interpretation in the syntactic category 23 | interp-oper-𝒮 : ∀ (f : oper) → (prod-𝒮 (oper-arity f)) ⇒s (ctx-slot (oper-sort f)) 24 | interp-oper-𝒮 f _ = tm-oper f λ i → tm-var i 25 | 26 | ℐ : Interpretation.Interpretation Σ cartesian-𝒮 27 | ℐ = 28 | record 29 | { interp-sort = ctx-slot 30 | ; interp-ctx = producted-𝒮 31 | ; interp-oper = interp-oper-𝒮 32 | } 33 | 34 | open Interpretation.Interpretation ℐ 35 | 36 | interp-term-self : ∀ {Γ} {A} (t : Term Γ A) (y : var (interp-sort A)) → ⊢ Γ ∥ (interp-term t y) ≈ t ⦂ A 37 | interp-term-self (tm-var x) _ = eq-refl 38 | interp-term-self (tm-oper f xs) _ = eq-congr (λ i → interp-term-self (xs i) var-var) 39 | -------------------------------------------------------------------------------- /src/MultiSorted/UniversalModel.agda: -------------------------------------------------------------------------------- 1 | import Relation.Binary.Reasoning.Setoid as SetoidR 2 | open import MultiSorted.AlgebraicTheory 3 | 4 | import MultiSorted.Interpretation as Interpretation 5 | import MultiSorted.Model as Model 6 | import MultiSorted.UniversalInterpretation as UniversalInterpretation 7 | import MultiSorted.Substitution as Substitution 8 | import MultiSorted.SyntacticCategory as SyntacticCategory 9 | 10 | module MultiSorted.UniversalModel 11 | {ℓt} 12 | {𝓈 ℴ} 13 | {Σ : Signature {𝓈} {ℴ}} 14 | (T : Theory ℓt Σ) where 15 | 16 | open Theory T 17 | open Substitution T 18 | open UniversalInterpretation T 19 | open Interpretation.Interpretation ℐ 20 | open SyntacticCategory T 21 | 22 | 𝒰 : Model.Is-Model T ℐ 23 | 𝒰 = 24 | record 25 | { model-eq = λ ε var-var → 26 | let open SetoidR (eq-setoid (ax-ctx ε) (sort-of (ctx-slot (ax-sort ε)) var-var)) in 27 | begin 28 | interp-term (ax-lhs ε) var-var ≈⟨ interp-term-self (ax-lhs ε) var-var ⟩ 29 | ax-lhs ε ≈⟨ id-action ⟩ 30 | ax-lhs ε [ id-s ]s ≈⟨ eq-axiom ε id-s ⟩ 31 | ax-rhs ε [ id-s ]s ≈˘⟨ id-action ⟩ 32 | ax-rhs ε ≈˘⟨ interp-term-self (ax-rhs ε) var-var ⟩ 33 | interp-term (ax-rhs ε) var-var ∎ 34 | } 35 | 36 | 37 | -- The universal model is universal 38 | universality : ∀ (ε : Equation Σ) → ⊨ ε → ⊢ ε 39 | universality ε p = 40 | let open Equation in 41 | let open SetoidR (eq-setoid (eq-ctx ε) (eq-sort ε)) in 42 | (begin 43 | eq-lhs ε ≈˘⟨ interp-term-self (eq-lhs ε) var-var ⟩ 44 | interp-term (eq-lhs ε) var-var ≈⟨ p var-var ⟩ 45 | interp-term (eq-rhs ε) var-var ≈⟨ interp-term-self (eq-rhs ε) var-var ⟩ 46 | eq-rhs ε ∎) 47 | -------------------------------------------------------------------------------- /src/SecondOrder/Arity.agda: -------------------------------------------------------------------------------- 1 | import Data.Empty 2 | import Data.Unit 3 | import Data.Bool 4 | import Data.Nat 5 | import Data.Fin 6 | 7 | module SecondOrder.Arity where 8 | 9 | -- A notion of arity is given by a set of possible arities, and a mapping which to each arity assings a set of 10 | -- argument positions. 11 | 12 | record Arity : Set₁ where 13 | field 14 | arity : Set -- the set of permissible arities, e.g., ℕ for finitary arities 15 | arg : arity → Set -- every arity gives a set of argument (position), e.g. Fin 16 | 17 | -- finitary arities 18 | arity-finite : Arity 19 | arity-finite = record { arity = Data.Nat.ℕ ; arg = Data.Fin.Fin } 20 | 21 | module Arity012 where 22 | -- For example, in algebra we quite often only consider constants, unary and binary 23 | -- operations. Thus we would only need three arities. 24 | 25 | data arity012 : Set where 26 | Constant Unary Binary : arity012 27 | 28 | arg012 : arity012 → Set 29 | arg012 Constant = Data.Empty.⊥ 30 | arg012 Unary = Data.Unit.⊤ 31 | arg012 Binary = Data.Bool.Bool 32 | 33 | arity-012 : Arity 34 | arity-012 = record { arity = arity012 ; arg = arg012 } 35 | -------------------------------------------------------------------------------- /src/SecondOrder/Equality.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality 3 | open import Relation.Binary using (Setoid) 4 | 5 | import SecondOrder.Arity 6 | import SecondOrder.Signature 7 | import SecondOrder.Metavariable 8 | import SecondOrder.VRenaming 9 | import SecondOrder.Term 10 | import SecondOrder.Substitution 11 | import SecondOrder.Instantiation 12 | import SecondOrder.Theory 13 | 14 | module SecondOrder.Equality 15 | {ℓ ℓa} 16 | {𝔸 : SecondOrder.Arity.Arity} 17 | {Σ : SecondOrder.Signature.Signature ℓ 𝔸} 18 | (𝕋 : SecondOrder.Theory.Theory Σ ℓa) 19 | where 20 | 21 | open SecondOrder.Metavariable Σ 22 | open SecondOrder.Renaming Σ 23 | open SecondOrder.Term Σ 24 | open SecondOrder.Substitution Σ 25 | open SecondOrder.Signature.Signature Σ 26 | open SecondOrder.Instantiation Σ 27 | open SecondOrder.Theory.Theory 𝕋 28 | 29 | record Equation : Set (lsuc ℓ) where 30 | constructor make-eq 31 | field 32 | eq-mv-ctx : MContext -- metavariable context of an equation 33 | eq-ctx : VContext -- variable context of an equation 34 | eq-sort : sort -- sort of an equation 35 | eq-lhs : Term eq-mv-ctx eq-ctx eq-sort -- left-hand side 36 | eq-rhs : Term eq-mv-ctx eq-ctx eq-sort -- right-hand side 37 | 38 | infix 5 make-eq 39 | 40 | syntax make-eq Θ Γ A s t = Θ ⊕ Γ ∥ s ≋ t ⦂ A 41 | 42 | -- Instantiate an axiom of 𝕋 to an equation 43 | instantiate-axiom : ∀ (ε : ax) {Θ Γ} (I : ax-mv-ctx ε ⇒ⁱ Θ ⊕ Γ) → Equation 44 | instantiate-axiom ε {Θ} {Γ} I = 45 | Θ ⊕ Γ ∥ [ I ]ⁱ (ax-lhs ε) ≋ [ I ]ⁱ (ax-rhs ε) ⦂ ax-sort ε 46 | 47 | 48 | -- The equality judgement 49 | 50 | infix 4 ⊢_ 51 | 52 | data ⊢_ : Equation → Set (lsuc (ℓ ⊔ ℓa)) where 53 | -- general rules 54 | eq-refl : ∀ {Θ Γ A} {t : Term Θ Γ A} → ⊢ Θ ⊕ Γ ∥ t ≋ t ⦂ A 55 | eq-symm : ∀ {Θ Γ A} {s t : Term Θ Γ A} → ⊢ Θ ⊕ Γ ∥ s ≋ t ⦂ A → ⊢ Θ ⊕ Γ ∥ t ≋ s ⦂ A 56 | eq-trans : ∀ {Θ Γ A} {s t u : Term Θ Γ A} → ⊢ Θ ⊕ Γ ∥ s ≋ t ⦂ A → ⊢ Θ ⊕ Γ ∥ t ≋ u ⦂ A → ⊢ Θ ⊕ Γ ∥ s ≋ u ⦂ A 57 | 58 | -- Congruence rule for operations 59 | -- the premises are: an operation f, two sets of arguments xs, ys of f that give 60 | -- for each argument of f a term in the extended context with the arguments that f binds 61 | -- such that xᵢ ≋ yᵢ for each i ∈ oper-arity f 62 | -- then f xs ≋ f ys (in the appropriate context) 63 | eq-oper : ∀ {Γ Θ} {f : oper} 64 | {xs ys : ∀ (i : oper-arg f) → Term Θ (Γ ,, arg-bind f i) (arg-sort f i)} 65 | → (∀ i → ⊢ Θ ⊕ (Γ ,, arg-bind f i) ∥ (xs i) ≋ (ys i) ⦂ (arg-sort f i)) 66 | → ⊢ Θ ⊕ Γ ∥ (tm-oper f xs) ≋ (tm-oper f ys) ⦂ (oper-sort f) 67 | -- Congruence rule for metavariables 68 | -- the permises are: a meta-variable M, and two sets of arguments of the appropriate 69 | -- sorts and arities to apply M, such that xᵢ ≋ yᵢ 70 | -- then M xs ≋ M ys 71 | eq-meta : ∀ {Γ Θ} {Γᴹ Aᴹ} {M : [ Γᴹ , Aᴹ ]∈ Θ} {xs ys : ∀ {B : sort} (i : B ∈ Γᴹ) → Term Θ Γ B} 72 | → (∀ {B : sort} (i : B ∈ Γᴹ) 73 | → ⊢ Θ ⊕ Γ ∥ (xs i) ≋ (ys i) ⦂ B) 74 | → ⊢ Θ ⊕ Γ ∥ (tm-meta M xs) ≋ (tm-meta M ys) ⦂ Aᴹ 75 | -- equational axiom 76 | eq-axiom : ∀ (ε : ax) {Θ Γ} (I : ax-mv-ctx ε ⇒ⁱ Θ ⊕ Γ) → ⊢ instantiate-axiom ε I 77 | 78 | -- Syntactically equal terms are judgementally equal 79 | ≈-≋ : ∀ {Θ Γ A} {s t : Term Θ Γ A} → s ≈ t → ⊢ Θ ⊕ Γ ∥ s ≋ t ⦂ A 80 | ≈-≋ (≈-≡ refl) = eq-refl 81 | ≈-≋ (≈-meta ξ) = eq-meta (λ i → ≈-≋ (ξ i)) 82 | ≈-≋ (≈-oper ξ) = eq-oper (λ i → ≈-≋ (ξ i)) 83 | 84 | 85 | -- terms and judgemental equality form a setoid 86 | eq-setoid : ∀ (Γ : VContext) (Θ : MContext) (A : sort) → Setoid ℓ (lsuc (ℓ ⊔ ℓa)) 87 | eq-setoid Γ Θ A = 88 | record 89 | { Carrier = Term Θ Γ A 90 | ; _≈_ = λ s t → ⊢ Θ ⊕ Γ ∥ s ≋ t ⦂ A 91 | ; isEquivalence = 92 | record 93 | { refl = eq-refl 94 | ; sym = eq-symm 95 | ; trans = eq-trans 96 | } 97 | } 98 | 99 | -- judgemental equality of substitutions 100 | _≋ˢ_ : ∀ {Θ Γ Δ} (σ τ : Θ ⊕ Γ ⇒ˢ Δ) → Set (lsuc (ℓ ⊔ ℓa)) 101 | _≋ˢ_ {Θ} {Γ} {Δ} σ τ = ∀ {A} (x : A ∈ Γ) → ⊢ Θ ⊕ Δ ∥ σ x ≋ τ x ⦂ A 102 | 103 | ≈ˢ-≋ˢ : ∀ {Θ Γ Δ} {σ τ : Θ ⊕ Γ ⇒ˢ Δ} → σ ≈ˢ τ → σ ≋ˢ τ 104 | ≈ˢ-≋ˢ ξ = λ x → ≈-≋ (ξ x) 105 | 106 | -- judgemental equality of metavariable instatiations 107 | _≋ⁱ_ : ∀ {Θ Ξ Γ} (I J : Θ ⇒ⁱ Ξ ⊕ Γ) → Set (lsuc (ℓ ⊔ ℓa)) 108 | _≋ⁱ_ {Θ} {Ξ} {Γ} I J = ∀ {Γᴹ Aᴹ} (M : [ Γᴹ , Aᴹ ]∈ Θ) → ⊢ Ξ ⊕ (Γ ,, Γᴹ) ∥ I M ≋ J M ⦂ Aᴹ 109 | -------------------------------------------------------------------------------- /src/SecondOrder/IndexedCategory.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive 2 | open import Categories.Category 3 | 4 | module SecondOrder.IndexedCategory where 5 | 6 | IndexedCategory : ∀ {i o l e} (I : Set i) (𝒞 : Category o l e) → Category (i ⊔ o) (i ⊔ l) (i ⊔ e) 7 | IndexedCategory I 𝒞 = 8 | let open Category 𝒞 in 9 | record 10 | { Obj = I → Obj 11 | ; _⇒_ = λ A B → ∀ i → A i ⇒ B i 12 | ; _≈_ = λ f g → ∀ i → f i ≈ g i 13 | ; id = λ i → id 14 | ; _∘_ = λ f g i → f i ∘ g i 15 | ; assoc = λ i → assoc 16 | ; sym-assoc = λ i → sym-assoc 17 | ; identityˡ = λ i → identityˡ 18 | ; identityʳ = λ i → identityʳ 19 | ; identity² = λ i → identity² 20 | ; equiv = record { refl = λ i → Equiv.refl 21 | ; sym = λ ξ i → Equiv.sym (ξ i) ; trans = λ ζ ξ i → Equiv.trans (ζ i) (ξ i) } 22 | ; ∘-resp-≈ = λ ζ ξ i → ∘-resp-≈ (ζ i) (ξ i) 23 | } 24 | -------------------------------------------------------------------------------- /src/SecondOrder/Instantiation.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; cong; subst) 3 | 4 | import SecondOrder.Arity 5 | import SecondOrder.Signature 6 | import SecondOrder.Metavariable 7 | import SecondOrder.VRenaming 8 | import SecondOrder.MRenaming 9 | import SecondOrder.Term 10 | import SecondOrder.Substitution 11 | import SecondOrder.RelativeMonadMorphism 12 | 13 | module SecondOrder.Instantiation 14 | {ℓ} 15 | {𝔸 : SecondOrder.Arity.Arity} 16 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 17 | where 18 | 19 | open SecondOrder.Signature.Signature Σ 20 | open SecondOrder.Metavariable Σ 21 | open SecondOrder.Term Σ 22 | open SecondOrder.VRenaming Σ 23 | open SecondOrder.MRenaming Σ 24 | open SecondOrder.Substitution Σ 25 | open import SecondOrder.RelativeMonadMorphism 26 | 27 | -- metavariable instantiation 28 | 29 | _⇒ⁱ_⊕_ : MContext → MContext → VContext → Set ℓ 30 | Θ ⇒ⁱ Ξ ⊕ Γ = ∀ {Γᴹ Aᴹ} (M : [ Γᴹ , Aᴹ ]∈ Θ) → Term Ξ (Γ ,, Γᴹ) Aᴹ 31 | 32 | 33 | -- instantiation respects propositionnal equality 34 | 35 | I-resp-≡ : ∀ {Θ Ω Γ Δ A} {M N : [ Δ , A ]∈ Θ} {I : Θ ⇒ⁱ Ω ⊕ Γ} → M ≡ N → I M ≡ I N 36 | I-resp-≡ refl = refl 37 | 38 | -- syntactic equality of instantiations 39 | 40 | infix 5 _≈ⁱ_ 41 | 42 | _≈ⁱ_ : ∀ {Θ Ψ Γ} (I J : Θ ⇒ⁱ Ψ ⊕ Γ) → Set ℓ 43 | _≈ⁱ_ {Θ} I J = ∀ {Γᴹ Aᴹ} (M : [ Γᴹ , Aᴹ ]∈ Θ) → I M ≈ J M 44 | 45 | -- equality of instantiations is an equivalence relation 46 | 47 | ≈ⁱ-refl : ∀ {Θ Ψ Γ} {I : Θ ⇒ⁱ Ψ ⊕ Γ} → I ≈ⁱ I 48 | ≈ⁱ-refl M = ≈-refl 49 | 50 | ≈ⁱ-sym : ∀ {Θ Ψ Γ} {I J : Θ ⇒ⁱ Ψ ⊕ Γ} → I ≈ⁱ J → J ≈ⁱ I 51 | ≈ⁱ-sym ξ M = ≈-sym (ξ M) 52 | 53 | ≈ⁱ-trans : ∀ {Θ Ψ Γ} {I J K : Θ ⇒ⁱ Ψ ⊕ Γ} → I ≈ⁱ J → J ≈ⁱ K → I ≈ⁱ K 54 | ≈ⁱ-trans ζ ξ M = ≈-trans (ζ M) (ξ M) 55 | 56 | -- extension of an instantiation 57 | 58 | ⇑ⁱ : ∀ {Θ Ψ Γ Δ} → Θ ⇒ⁱ Ψ ⊕ Γ → Θ ⇒ⁱ Ψ ⊕ (Γ ,, Δ) 59 | ⇑ⁱ I M = [ ⇑ᵛ var-inl ]ᵛ I M 60 | 61 | -- extension of instantiations preserve equality of instantiations 62 | 63 | ⇑ⁱ-resp-≈ⁱ : ∀ {Θ Ψ Γ Δ} (I J : Θ ⇒ⁱ Ψ ⊕ Γ) → (I ≈ⁱ J) → (⇑ⁱ {Δ = Δ} I ≈ⁱ ⇑ⁱ J) 64 | ⇑ⁱ-resp-≈ⁱ I J ξ M = []ᵛ-resp-≈ (ξ M) 65 | 66 | -- action of a metavariable instantiation on terms 67 | 68 | infixr 6 [_]ⁱ_ 69 | 70 | [_]ⁱ_ : ∀ {Θ ψ Γ} → ψ ⇒ⁱ Θ ⊕ Γ → ∀ {A} → Term ψ Γ A → Term Θ Γ A 71 | [ I ]ⁱ (tm-var x) = tm-var x 72 | [ I ]ⁱ (tm-meta M ts) = [ [ idˢ , (λ i → [ I ]ⁱ ts i) ]ˢ ]ˢ (I M) 73 | [ I ]ⁱ (tm-oper f es) = tm-oper f λ i → [ ⇑ⁱ I ]ⁱ es i 74 | 75 | -- instantiation preserves equality of terms 76 | 77 | []ⁱ-resp-≈ : ∀ {Θ Ξ Γ} (I : Ξ ⇒ⁱ Θ ⊕ Γ) → ∀ {A} {t u : Term Ξ Γ A} → 78 | t ≈ u → [ I ]ⁱ t ≈ [ I ]ⁱ u 79 | []ⁱ-resp-≈ I (≈-≡ refl) = ≈-refl 80 | []ⁱ-resp-≈ I (≈-meta {M = M} ξ) = []ˢ-resp-≈ˢ (I M) ([,]ˢ-resp-≈ˢ (λ x → ≈-refl) λ i → []ⁱ-resp-≈ I (ξ i)) 81 | []ⁱ-resp-≈ I (≈-oper ξ) = ≈-oper λ i → []ⁱ-resp-≈ (⇑ⁱ I) (ξ i) 82 | 83 | -- action preserves equality of instantiation 84 | 85 | []ⁱ-resp-≈ⁱ : ∀ {Θ Ξ Γ} {I J : Ξ ⇒ⁱ Θ ⊕ Γ} → ∀ {A} (t : Term Ξ Γ A) → 86 | I ≈ⁱ J → [ I ]ⁱ t ≈ [ J ]ⁱ t 87 | []ⁱ-resp-≈ⁱ (tm-var x) ξ = ≈-refl 88 | []ⁱ-resp-≈ⁱ (tm-meta M ts) ξ = 89 | []ˢ-resp-≈ˢ-≈ ([,]ˢ-resp-≈ˢ (λ x → ≈-refl) λ i → []ⁱ-resp-≈ⁱ (ts i) ξ) (ξ M) 90 | []ⁱ-resp-≈ⁱ {I = I} {J = J} (tm-oper f es) ξ = ≈-oper λ i → []ⁱ-resp-≈ⁱ (es i) (⇑ⁱ-resp-≈ⁱ I J ξ) 91 | 92 | -- generically applied metavariable 93 | 94 | tm-meta-generic : ∀ {Θ} {Γ} {Γᴹ Aᴹ} (M : [ Γᴹ , Aᴹ ]∈ Θ) → Term Θ (Γ ,, Γᴹ) Aᴹ 95 | tm-meta-generic M = tm-meta M λ i → tm-var (var-inr i) 96 | 97 | -- the action of an instantiation on a generically applied metavariable 98 | 99 | []ⁱ-generic : ∀ {Θ Ξ} {Γ} {I : Θ ⇒ⁱ Ξ ⊕ Γ} {Γᴹ Aᴹ} {M : [ Γᴹ , Aᴹ ]∈ Θ} → 100 | [ ⇑ⁱ {Δ = Γᴹ} I ]ⁱ tm-meta-generic M ≈ I M 101 | []ⁱ-generic {Θ = Θ} {Ξ = Ξ} {Γ = Γ} {I = I} {Γᴹ = Γᴹ} {M = M} = 102 | ≈-trans 103 | (≈-sym ([ˢ∘ᵛ] (I M))) 104 | (≈ˢ-idˢ-[]ˢ (λ { (var-inl _) → ≈-refl ; (var-inr _) → ≈-refl})) 105 | 106 | 107 | -- Interactions involving instantiations 108 | 109 | -- the identity metavariable instantiation 110 | 111 | idⁱ : ∀ {Θ Γ} → Θ ⇒ⁱ Θ ⊕ Γ 112 | idⁱ M = tm-meta-generic M 113 | 114 | -- composition of metavariable instantiations 115 | 116 | infixl 6 _∘ⁱ_ 117 | 118 | _∘ⁱ_ : ∀ {Θ Ξ Ω Γ} → Ξ ⇒ⁱ Ω ⊕ Γ → Θ ⇒ⁱ Ξ ⊕ Γ → (Θ ⇒ⁱ Ω ⊕ Γ) 119 | (I ∘ⁱ J) M = [ ⇑ⁱ I ]ⁱ J M 120 | 121 | -- composition of a renaming and an instantiation 122 | 123 | _ᵛ∘ⁱ_ : ∀ {Θ ψ Γ Δ} → Γ ⇒ᵛ Δ → Θ ⇒ⁱ ψ ⊕ Γ → Θ ⇒ⁱ ψ ⊕ Δ 124 | (ρ ᵛ∘ⁱ I) M = [ ⇑ᵛ ρ ]ᵛ I M 125 | 126 | -- composition of a substitution and an instantiation 127 | 128 | _ˢ∘ⁱ_ : ∀ {Θ ψ Γ Δ} → ψ ⊕ Γ ⇒ˢ Δ → Θ ⇒ⁱ ψ ⊕ Γ → Θ ⇒ⁱ ψ ⊕ Δ 129 | (σ ˢ∘ⁱ I) M = [ ⇑ˢ σ ]ˢ I M 130 | 131 | -- composition of an instantiation and a substitution 132 | 133 | _ⁱ∘ˢ_ : ∀ {Θ ψ Γ Δ} → Θ ⇒ⁱ ψ ⊕ Δ → Θ ⊕ Γ ⇒ˢ Δ → ψ ⊕ Γ ⇒ˢ Δ 134 | (I ⁱ∘ˢ σ) x = [ I ]ⁱ σ x 135 | 136 | -- composition of a renaming and an instantiation preerve equality of instantiations 137 | 138 | [ᵛ∘ⁱ]-resp-≈ⁱ : ∀ {Θ ψ Γ Δ} (ρ : Γ ⇒ᵛ Δ) (I J : Θ ⇒ⁱ ψ ⊕ Γ) → (I ≈ⁱ J) → (ρ ᵛ∘ⁱ I) ≈ⁱ (ρ ᵛ∘ⁱ J) 139 | [ᵛ∘ⁱ]-resp-≈ⁱ σ I J ξ M = []ᵛ-resp-≈ (ξ M) 140 | 141 | -- composition of a renaming and an instantiation preserve equality of instantiations 142 | 143 | [ˢ∘ⁱ]-resp-≈ⁱ : ∀ {Θ ψ Γ Δ} (σ : ψ ⊕ Γ ⇒ˢ Δ) (I J : Θ ⇒ⁱ ψ ⊕ Γ) → (I ≈ⁱ J) → (σ ˢ∘ⁱ I) ≈ⁱ (σ ˢ∘ⁱ J) 144 | [ˢ∘ⁱ]-resp-≈ⁱ σ I J ξ M = []ˢ-resp-≈ (⇑ˢ σ) (ξ M) 145 | 146 | 147 | -- Actions correspondig to the interactions 148 | 149 | -- the action of the identity 150 | 151 | [idⁱ] : ∀ {Θ Γ A Δ} {t : Term Θ (Γ ,, Δ) A} → [ idⁱ ]ⁱ t ≈ t 152 | [idⁱ] {t = tm-var x} = ≈-refl 153 | [idⁱ] {t = tm-meta M ts} = ≈-meta (λ i → [idⁱ]) 154 | [idⁱ] {t = tm-oper f es} = ≈-oper (λ i → [idⁱ]) 155 | 156 | -- extension commutes with composition of renaming and substitution 157 | 158 | ⇑ⁱ-resp-ᵛ∘ⁱ : ∀ {Θ ψ} {Γ Δ Ξ} {I : Θ ⇒ⁱ ψ ⊕ Γ} {ρ : Γ ⇒ᵛ Δ} 159 | → ⇑ⁱ {Δ = Ξ} (ρ ᵛ∘ⁱ I) ≈ⁱ ⇑ᵛ ρ ᵛ∘ⁱ ⇑ⁱ I 160 | ⇑ⁱ-resp-ᵛ∘ⁱ {ρ = ρ} M = ≈-trans (≈-sym [∘ᵛ]) (≈-trans ([]ᵛ-resp-≡ᵛ λ {(var-inl _) → refl ; (var-inr x) → refl}) [∘ᵛ]) 161 | 162 | -- the action of the composition of an instantiation and a renaming 163 | 164 | [ᵛ∘ⁱ] : ∀ {Θ Ψ Γ Δ A} {ρ : Γ ⇒ᵛ Δ} {I : Θ ⇒ⁱ Ψ ⊕ Γ} (t : Term Θ Γ A) → 165 | [ ρ ]ᵛ [ I ]ⁱ t ≈ [ ρ ᵛ∘ⁱ I ]ⁱ [ ρ ]ᵛ t 166 | [ᵛ∘ⁱ] (tm-var x) = ≈-refl 167 | [ᵛ∘ⁱ] {I = I} (tm-meta M ts) = 168 | ≈-trans 169 | (≈-sym ([ᵛ∘ˢ] (I M))) 170 | (≈-trans 171 | ([]ˢ-resp-≈ˢ (I M) (λ { (var-inl _) → ≈-refl ; (var-inr j) → [ᵛ∘ⁱ] (ts j)})) 172 | ([ˢ∘ᵛ] (I M))) 173 | [ᵛ∘ⁱ] {ρ = ρ} {I = I} (tm-oper f es) = 174 | ≈-oper λ i → ≈-trans ([ᵛ∘ⁱ] (es i)) ([]ⁱ-resp-≈ⁱ ([ ⇑ᵛ ρ ]ᵛ es i) (≈ⁱ-sym (⇑ⁱ-resp-ᵛ∘ⁱ {I = I}))) 175 | 176 | 177 | -- extension commutes with composition 178 | 179 | ⇑ⁱ-resp-∘ⁱ : ∀ {Θ Ξ Ω} {Γ Δ} {I : Θ ⇒ⁱ Ξ ⊕ Γ} {J : Ξ ⇒ⁱ Ω ⊕ Γ} → 180 | ⇑ⁱ {Δ = Δ} (J ∘ⁱ I) ≈ⁱ ⇑ⁱ J ∘ⁱ ⇑ⁱ I 181 | ⇑ⁱ-resp-∘ⁱ {I = I} {J = J} M = 182 | ≈-trans 183 | ([ᵛ∘ⁱ] (I M)) 184 | ([]ⁱ-resp-≈ⁱ 185 | ([ ⇑ᵛ var-inl ]ᵛ I M) 186 | (λ N → ≈-trans 187 | (≈-sym [∘ᵛ]) 188 | (≈-trans 189 | ([]ᵛ-resp-≡ᵛ (λ { (var-inl x) → refl ; (var-inr x) → refl })) 190 | [∘ᵛ]))) 191 | 192 | 193 | ⇑ˢ-resp-ⁱ∘ˢ : ∀ {Θ ψ Γ Δ Ξ} → {I : Θ ⇒ⁱ ψ ⊕ Δ} → {σ : Θ ⊕ Γ ⇒ˢ Δ} 194 | → ⇑ˢ {Ξ = Ξ} (I ⁱ∘ˢ σ) ≈ˢ ⇑ⁱ I ⁱ∘ˢ ⇑ˢ σ 195 | ⇑ˢ-resp-ⁱ∘ˢ {σ = σ} (var-inl x) = [ᵛ∘ⁱ] (σ x) 196 | ⇑ˢ-resp-ⁱ∘ˢ (var-inr x) = ≈-refl 197 | 198 | -- interaction lemma 199 | []ⁱ-[]ˢ : ∀ {Θ Ψ Γ Δ A} {I : Θ ⇒ⁱ Ψ ⊕ Δ} {σ : Θ ⊕ Γ ⇒ˢ Δ} {ρ : Δ ⇒ᵛ Γ} (t : Term Θ Γ A) → 200 | σ ˢ∘ᵛ ρ ≈ˢ idˢ → ([ I ]ⁱ ([ σ ]ˢ t)) ≈ ([ I ⁱ∘ˢ σ ]ˢ [ ρ ᵛ∘ⁱ I ]ⁱ t) 201 | []ⁱ-[]ˢ (tm-var x) ξ = ≈-refl 202 | []ⁱ-[]ˢ {I = I} {σ = σ} {ρ = ρ} (tm-meta M ts) ξ = 203 | ≈-trans 204 | ([]ˢ-resp-≈ˢ (I M) 205 | (λ { (var-inl i) → []ⁱ-resp-≈ I (≈-sym (ξ i)) ; (var-inr j) → []ⁱ-[]ˢ (ts j) ξ})) 206 | (≈-trans 207 | ([ˢ∘ᵛ] (I M)) 208 | ([∘ˢ] ((ρ ᵛ∘ⁱ I) M))) 209 | []ⁱ-[]ˢ {I = I} {σ = σ} {ρ = ρ} (tm-oper f es) ξ = 210 | ≈-oper λ i → 211 | ≈-trans 212 | ([]ⁱ-[]ˢ {σ = ⇑ˢ σ} {ρ = ⇑ᵛ ρ} (es i) 213 | (≈ˢ-trans 214 | (≈ˢ-sym (⇑ˢ-resp-ˢ∘ᵛ {ρ = ρ} {σ = σ})) 215 | (≈ˢ-trans (⇑ˢ-resp-≈ˢ ξ) ⇑ˢ-resp-idˢ))) 216 | ([]ˢ-resp-≈ˢ-≈ 217 | {σ = ⇑ⁱ I ⁱ∘ˢ ⇑ˢ σ } 218 | {τ = ⇑ˢ (I ⁱ∘ˢ σ)} 219 | {t = ([ ⇑ᵛ ρ ᵛ∘ⁱ ⇑ⁱ I ]ⁱ es i)} 220 | {u = ([ ⇑ⁱ (ρ ᵛ∘ⁱ I) ]ⁱ es i)} 221 | (≈ˢ-sym ⇑ˢ-resp-ⁱ∘ˢ) 222 | ([]ⁱ-resp-≈ⁱ (es i) (≈ⁱ-sym (⇑ⁱ-resp-ᵛ∘ⁱ {I = I})))) 223 | 224 | -- the action of a composition is functorial 225 | 226 | [∘ⁱ] : ∀ {Θ Ξ Ω Γ} → {I : Θ ⇒ⁱ Ξ ⊕ Γ} → {J : Ξ ⇒ⁱ Ω ⊕ Γ} → 227 | ∀ {A} → ∀ (t : Term Θ Γ A) → [ J ∘ⁱ I ]ⁱ t ≈ [ J ]ⁱ [ I ]ⁱ t 228 | [∘ⁱ] (tm-var x) = ≈-refl 229 | [∘ⁱ] {Θ = Θ} {Ξ = Ξ} {Γ = Γ} {I = I} {J = J} (tm-meta {Γᴹ = Γᴹ} M ts) = 230 | ≈-trans 231 | ([]ˢ-resp-≈ˢ 232 | ([ ⇑ⁱ J ]ⁱ (I M)) 233 | ([,]ˢ-resp-≈ˢ (λ x → ≈-refl) (λ i → [∘ⁱ] {I = I} {J = J} (ts i)))) 234 | (≈-trans 235 | ([]ˢ-resp-≈ˢ {τ = λ i → [ J ]ⁱ var-or-ts i} ([ ⇑ⁱ J ]ⁱ (I M)) 236 | λ {(var-inl x) → ≈-refl ; (var-inr y) → ≈-refl}) 237 | (≈-trans 238 | (≈-sym ([]ⁱ-[]ˢ {I = J} {σ = var-or-ts} {ρ = var-inl} (I M) λ x → ≈-refl)) 239 | ([]ⁱ-resp-≈ J 240 | ([]ˢ-resp-≈ˢ (I M) λ { (var-inl x) → ≈-refl ; (var-inr x) → ≈-refl})))) 241 | where 242 | var-or-ts : ∀ {A} → A ∈ (Γ ,, Γᴹ) → Term Ξ Γ A 243 | var-or-ts (var-inl x) = tm-var x 244 | var-or-ts (var-inr y) = [ I ]ⁱ ts y 245 | 246 | [∘ⁱ] {I = I} {J = J} (tm-oper f es) = 247 | ≈-oper (λ i → ≈-trans ([]ⁱ-resp-≈ⁱ (es i) (⇑ⁱ-resp-∘ⁱ {I = I})) ([∘ⁱ] (es i))) 248 | 249 | 250 | -- The category of metacontext and instantiations 251 | 252 | module _ {Γ} where 253 | 254 | open import Categories.Category 255 | 256 | Metacontexts-Insts : Category ℓ ℓ ℓ 257 | Metacontexts-Insts = 258 | record 259 | { Obj = MContext --VContext 260 | ; _⇒_ = _⇒ⁱ_⊕ Γ 261 | ; _≈_ = _≈ⁱ_ 262 | ; id = idⁱ 263 | ; _∘_ = _∘ⁱ_ 264 | ; assoc = λ {Θ} {Ω} {ψ} {Ξ} {K} {J} {I} {Γᴹ} {Aᴹ} M → 265 | ≈-trans 266 | ([]ⁱ-resp-≈ⁱ (K M) (⇑ⁱ-resp-∘ⁱ {Δ = Γᴹ} {I = J} {J = I})) 267 | ([∘ⁱ] (K M)) 268 | ; sym-assoc = λ {Θ} {Ω} {ψ} {Ξ} {K} {J} {I} {Γᴹ} {Aᴹ} M → ≈-sym 269 | ( ≈-trans 270 | ([]ⁱ-resp-≈ⁱ (K M) (⇑ⁱ-resp-∘ⁱ {Δ = Γᴹ} {I = J} {J = I})) 271 | ([∘ⁱ] (K M))) 272 | ; identityˡ = λ M → [idⁱ] 273 | ; identityʳ = λ {A} {B} {I} M → 274 | ≈-trans 275 | (≈-trans 276 | (≈-sym ([ˢ∘ᵛ] (I M))) 277 | ([]ˢ-resp-≈ˢ (I M) λ x → idʳ x)) 278 | [idˢ] 279 | ; identity² = λ M → ≈-refl 280 | ; equiv = record 281 | { refl = λ {I} → ≈ⁱ-refl {I = I} 282 | ; sym = ≈ⁱ-sym 283 | ; trans = ≈ⁱ-trans } 284 | ; ∘-resp-≈ = λ {Θ} {Ω} {ψ} {I} {J} {K} {L} ξ₁ ξ₂ M → 285 | ≈-trans 286 | ([]ⁱ-resp-≈ (⇑ⁱ I) (ξ₂ M)) 287 | ([]ⁱ-resp-≈ⁱ (L M) (⇑ⁱ-resp-≈ⁱ I J ξ₁)) 288 | } 289 | 290 | where 291 | idʳ : ∀ {Θ Δ A} (x : A ∈ (Γ ,, Δ)) → 292 | [ tm-var , (λ i → tm-var (var-inr i)) ]ˢ 293 | ([ (λ x₁ → var-inl {Γ = Γ ,, Δ} (var-inl x₁)) , (λ x₁ → var-inr x₁) ]ᵛ x) 294 | ≈ tm-var {Θ = Θ} x 295 | idʳ (var-inl x) = ≈-refl 296 | idʳ (var-inr x) = ≈-refl 297 | -------------------------------------------------------------------------------- /src/SecondOrder/MContext.agda: -------------------------------------------------------------------------------- 1 | module SecondOrder.MContext {ℓ} (arity : Set ℓ) (sort : Set ℓ) where 2 | 3 | data MContext : Set ℓ where 4 | ctx-empty : MContext 5 | ctx-slot : arity → sort → MContext 6 | _,,_ : MContext → MContext → MContext 7 | 8 | infixl 5 _,,_ 9 | 10 | infix 4 [_,_]∈_ 11 | 12 | -- the meta-variables of a given type in a context 13 | data [_,_]∈_ (Λ : arity) (A : sort) : MContext → Set ℓ where 14 | var-slot : [ Λ , A ]∈ ctx-slot Λ A 15 | var-inl : ∀ {Θ ψ} (x : [ Λ , A ]∈ Θ) → [ Λ , A ]∈ Θ ,, ψ 16 | var-inr : ∀ {Θ ψ} (y : [ Λ , A ]∈ ψ) → [ Λ , A ]∈ Θ ,, ψ 17 | -------------------------------------------------------------------------------- /src/SecondOrder/MRelMonMorphism.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; cong; cong₂; subst; setoid) 3 | open import Data.Product using (_×_; Σ; _,_; proj₁; proj₂; zip; map; <_,_>; swap) 4 | import Function.Equality 5 | open import Relation.Binary using (Setoid) 6 | import Relation.Binary.Reasoning.Setoid as SetoidR 7 | 8 | import SecondOrder.Arity 9 | import SecondOrder.Signature 10 | import SecondOrder.Metavariable 11 | import SecondOrder.VRenaming 12 | import SecondOrder.MRenaming 13 | import SecondOrder.Term 14 | import SecondOrder.Substitution 15 | import SecondOrder.RelativeMonadMorphism 16 | import SecondOrder.Instantiation 17 | import SecondOrder.IndexedCategory 18 | import SecondOrder.RelativeKleisli 19 | import SecondOrder.Mslot 20 | import SecondOrder.MRelativeMonad 21 | import SecondOrder.VRelativeMonad 22 | 23 | 24 | module SecondOrder.MRelMonMorphism 25 | {ℓ} 26 | {𝔸 : SecondOrder.Arity.Arity} 27 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 28 | where 29 | 30 | open SecondOrder.Signature.Signature Σ 31 | open SecondOrder.Metavariable Σ 32 | open SecondOrder.Term Σ 33 | open SecondOrder.VRenaming Σ 34 | open SecondOrder.MRenaming Σ 35 | open SecondOrder.Mslot Σ 36 | open SecondOrder.Substitution Σ 37 | open SecondOrder.Instantiation Σ 38 | open SecondOrder.MRelativeMonad Σ 39 | open SecondOrder.VRelativeMonad Σ 40 | open SecondOrder.RelativeMonadMorphism 41 | 42 | module MRelMonMorph {Γ Γ′ : VContext} where 43 | open MTerm {Γ} renaming (MMonad to M_Γ) 44 | open MTerm {Γ′} renaming (MMonad to M_Γ′) 45 | 46 | 47 | Fᴹ : ∀ (ρ : Γ ⇒ᵛ Γ′) → RMonadMorph (M_Γ) (M_Γ′) 48 | Fᴹ ρ = record 49 | { morph = λ Γᴹ Aᴹ → 50 | record 51 | { _⟨$⟩_ = [ ⇑ᵛ ρ ]ᵛ_ 52 | ; cong = λ M≈N → []ᵛ-resp-≈ M≈N 53 | } 54 | ; law-unit = λ Γᴹ Aᴹ M≡N → ≈-≡ 55 | (I-resp-≡ 56 | {I = λ M → tm-meta M λ i → tm-var (var-inr i) } 57 | M≡N) 58 | ; law-extend = λ {Θ} {Ω} {k} Γᴹ Aᴹ {s} {t} s≈t → 59 | ≈-trans 60 | ([ᵛ∘ⁱ] s) 61 | (≈-trans 62 | ([]ⁱ-resp-≈ 63 | (⇑ᵛ ρ ᵛ∘ⁱ ⇑ⁱ (λ {Λ} {B} → 64 | Function.Equality._⟨$⟩_ (k Λ B))) 65 | ([]ᵛ-resp-≈ s≈t)) 66 | ([]ⁱ-resp-≈ⁱ ([ (⇑ᵛ ρ) ]ᵛ t) λ M → 67 | (≈-trans 68 | (≈-sym [∘ᵛ]) 69 | (≈-trans 70 | ([]ᵛ-resp-≡ᵛ extend-aux) 71 | [∘ᵛ])))) 72 | } 73 | 74 | where 75 | extend-aux : ∀ {Ξ Λ} → ⇑ᵛ {Ξ = Ξ} (⇑ᵛ {Ξ = Λ} ρ) ∘ᵛ (⇑ᵛ var-inl) ≡ᵛ ⇑ᵛ var-inl ∘ᵛ (⇑ᵛ ρ) 76 | extend-aux (var-inl x) = refl 77 | extend-aux (var-inr x) = refl 78 | 79 | 80 | 81 | 82 | 83 | 84 | 85 | 86 | 87 | 88 | 89 | 90 | 91 | 92 | 93 | 94 | 95 | 96 | 97 | 98 | -- Attempt with substitutions instead of renamings : not sure it works 99 | -- hat : (f : (Θ : MContext) → Θ ⊕ Γ ⇒ˢ Γ') → RMonadMorph M_Γ M_Γ' 100 | -- hat f = 101 | -- let open Function.Equality using (_⟨$⟩_) renaming (cong to func-cong) in 102 | -- record 103 | -- { morph = λ {Θ} Δ A → record { _⟨$⟩_ = λ t → [ ⇑ˢ (f Θ) ]ˢ t ; cong = []ˢ-resp-≈ (⇑ˢ (f Θ))} 104 | -- ; law-unit = λ Δ A {t} {s} t≡s → ≈-≡ (cong tm-meta-generic t≡s) 105 | -- ; law-extend = λ {Θ} {Ψ} {I} Δ A {t} {s} t≈s → 106 | -- let open SetoidR (Term-setoid Ψ (Γ' ,, Δ) A) in 107 | -- begin 108 | -- ([ ⇑ˢ (f Ψ) ]ˢ ([ ⇑ⁱ (λ {Λ} {B} M → (I Λ B ⟨$⟩ M) ) ]ⁱ t) ) 109 | -- ≈⟨ {!!} ⟩ -- this is the crucial step 110 | -- ([ ⇑ⁱ (λ {Λ} {B} M → [ ⇑ˢ (f Ψ) ]ˢ (I Λ B ⟨$⟩ M)) ]ⁱ ([ ⇑ˢ (f Θ) ]ˢ t)) 111 | -- ≈⟨ []ⁱ-resp-≈ (⇑ⁱ (λ {Λ} {B} M → [ ⇑ˢ (f Ψ) ]ˢ (I Λ B ⟨$⟩ M))) ([]ˢ-resp-≈ (⇑ˢ (f Θ)) t≈s) ⟩ 112 | -- ([ ⇑ⁱ (λ {Λ} {B} M → [ ⇑ˢ (f Ψ) ]ˢ (I Λ B ⟨$⟩ M)) ]ⁱ ([ ⇑ˢ (f Θ) ]ˢ s)) 113 | -- ∎ 114 | -- } 115 | -------------------------------------------------------------------------------- /src/SecondOrder/MRelativeMonad.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; cong; cong₂; subst; setoid) 3 | open import Data.Product using (_×_; Σ; _,_; proj₁; proj₂; zip; map; <_,_>; swap) 4 | import Function.Equality 5 | open import Relation.Binary using (Setoid) 6 | import Relation.Binary.Reasoning.Setoid as SetoidR 7 | 8 | import Categories.Category 9 | import Categories.Functor 10 | import Categories.Category.Instance.Setoids 11 | import Categories.Monad.Relative 12 | import Categories.Category.Equivalence 13 | import Categories.Category.Cocartesian 14 | import Categories.Category.Construction.Functors 15 | import Categories.Category.Product 16 | import Categories.NaturalTransformation 17 | import Categories.NaturalTransformation.NaturalIsomorphism 18 | 19 | import SecondOrder.Arity 20 | import SecondOrder.Signature 21 | import SecondOrder.Metavariable 22 | import SecondOrder.VRenaming 23 | import SecondOrder.MRenaming 24 | import SecondOrder.Term 25 | import SecondOrder.Substitution 26 | import SecondOrder.RelativeMonadMorphism 27 | import SecondOrder.Instantiation 28 | import SecondOrder.IndexedCategory 29 | import SecondOrder.RelativeKleisli 30 | import SecondOrder.Mslot 31 | 32 | 33 | module SecondOrder.MRelativeMonad 34 | {ℓ} 35 | {𝔸 : SecondOrder.Arity.Arity} 36 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 37 | where 38 | 39 | open SecondOrder.Signature.Signature Σ 40 | open SecondOrder.Metavariable Σ 41 | open SecondOrder.Term Σ 42 | open SecondOrder.VRenaming Σ 43 | open SecondOrder.MRenaming Σ 44 | open SecondOrder.Mslot Σ 45 | open SecondOrder.Substitution Σ 46 | open SecondOrder.Instantiation Σ 47 | open import SecondOrder.RelativeMonadMorphism 48 | open Categories.Category 49 | open Categories.Functor using (Functor) 50 | open Categories.NaturalTransformation renaming (id to idNt) 51 | open Categories.NaturalTransformation.NaturalIsomorphism renaming (refl to reflNt; sym to symNt; trans to transNt) 52 | open Categories.Category.Construction.Functors 53 | open Categories.Category.Instance.Setoids 54 | open Categories.Category.Product 55 | open Function.Equality using () renaming (setoid to Π-setoid) 56 | open SecondOrder.IndexedCategory 57 | 58 | module MTerm {Γ : VContext} where 59 | open Category 60 | open NaturalTransformation 61 | open Functor 62 | open Categories.Monad.Relative 63 | open Function.Equality using () renaming (setoid to Π-setoid) 64 | open Categories.Category.Equivalence using (StrongEquivalence) 65 | open SecondOrder.RelativeKleisli 66 | 67 | MMonad : Monad Mslots 68 | MMonad = 69 | let open Function.Equality using (_⟨$⟩_) renaming (cong to func-cong) in 70 | record 71 | -- The object mapping (which is also a functor) 72 | { F₀ = λ Θ Δ A → Term-setoid Θ (Γ ,, Δ) A 73 | -- The unit of the relative monad 74 | ; unit = λ {Θ} Δ A → 75 | record 76 | { _⟨$⟩_ = λ M → tm-meta-generic M 77 | ; cong = λ {M} {N} M≡N → ≈-≡ (cong tm-meta-generic M≡N) 78 | } 79 | -- The extension in the rel monad 80 | ; extend = λ {Θ} {Ψ} I Δ A → 81 | record 82 | { _⟨$⟩_ = λ t → [ ⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M) ]ⁱ t 83 | ; cong = λ {t} {s} t≈s → 84 | let open SetoidR (Term-setoid Ψ (Γ ,, Δ) A) in 85 | begin 86 | ([ ⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M) ]ⁱ t) 87 | ≈⟨ []ⁱ-resp-≈ (⇑ⁱ (λ {Λ} {B} → _⟨$⟩_ (I Λ B))) t≈s ⟩ 88 | ([ ⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M) ]ⁱ s) 89 | ∎ 90 | } 91 | -- This is the law that says: I* ∘ η = I 92 | ; identityʳ = λ {Θ} {Ψ} {I} Δ A {t} {s} t≈s 93 | → ≈-trans ([]ⁱ-generic {Θ} {Ψ} {Γ} {λ {Λ} {B} M → I Λ B ⟨$⟩ M}) (func-cong (I Δ A) t≈s) 94 | -- This is the law that says η* = id 95 | ; identityˡ = λ {Θ} Δ A {t} {s} t≈s → 96 | let open SetoidR (Term-setoid Θ (Γ ,, Δ) A) in 97 | begin 98 | ([ ⇑ⁱ (λ {Λ} {B} M → tm-meta-generic M) ]ⁱ t) ≈⟨ []ⁱ-resp-≈ (⇑ⁱ tm-meta-generic) t≈s ⟩ 99 | ([ ⇑ⁱ (λ {Λ} {B} M → tm-meta-generic M) ]ⁱ s) ≈⟨ [idⁱ] ⟩ 100 | s 101 | ∎ 102 | -- This is the law that says: (J* ∘ I)* = J* ∘ I* 103 | ; assoc = λ {Θ} {Ψ} {Ξ} {I} {J} Δ A {t} {s} t≈s → 104 | let open SetoidR (Term-setoid Ξ (Γ ,, Δ) A) in 105 | begin 106 | ([ ⇑ⁱ ((λ {Λ} {B} M → J Λ B ⟨$⟩ M) ∘ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M)) ]ⁱ t) 107 | ≈⟨ []ⁱ-resp-≈ⁱ t (⇑ⁱ-resp-∘ⁱ {Θ} {Ψ} {Ξ} {Γ} {Δ} {(λ {Λ} {B} M → I Λ B ⟨$⟩ M)}) ⟩ 108 | ([ (⇑ⁱ (λ {Λ} {B} M → J Λ B ⟨$⟩ M)) ∘ⁱ (⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M)) ]ⁱ t) 109 | ≈⟨ [∘ⁱ] t ⟩ 110 | ([ ⇑ⁱ (λ {Λ} {B} M → J Λ B ⟨$⟩ M) ]ⁱ ([ ⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M) ]ⁱ t)) 111 | ≈⟨ []ⁱ-resp-≈ (⇑ⁱ (λ {Λ} {B} M → J Λ B ⟨$⟩ M)) 112 | ([]ⁱ-resp-≈ (⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M)) t≈s) ⟩ 113 | ([ ⇑ⁱ (λ {Λ} {B} M → J Λ B ⟨$⟩ M) ]ⁱ ([ ⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M) ]ⁱ s)) 114 | ∎ 115 | -- Extension respects equality of instantiations 116 | ; extend-≈ = λ {Θ} {Ψ} {I} {J} I≈ⁱJ Δ A {t} {s} t≈s → 117 | let open SetoidR (Term-setoid Ψ (Γ ,, Δ) A) in 118 | begin 119 | ([ ⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M) ]ⁱ t) ≈⟨ []ⁱ-resp-≈ (⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M)) t≈s ⟩ 120 | (([ ⇑ⁱ (λ {Λ} {B} M → I Λ B ⟨$⟩ M) ]ⁱ s)) 121 | ≈⟨ []ⁱ-resp-≈ⁱ s (⇑ⁱ-resp-≈ⁱ ((λ {Λ} {B} M → I Λ B ⟨$⟩ M)) 122 | ((λ {Λ} {B} M → J Λ B ⟨$⟩ M)) λ M → I≈ⁱJ _ _ refl) ⟩ 123 | ([ ⇑ⁱ (λ {Λ} {B} M → J Λ B ⟨$⟩ M) ]ⁱ s) 124 | ∎ 125 | } 126 | -------------------------------------------------------------------------------- /src/SecondOrder/MRenaming.agda: -------------------------------------------------------------------------------- 1 | open import Level 2 | open import Relation.Binary.PropositionalEquality 3 | open import Relation.Binary using (Setoid) 4 | import Function.Equality 5 | import Relation.Binary.Reasoning.Setoid as SetoidR 6 | 7 | import Categories.Category 8 | import Categories.Functor 9 | import Categories.NaturalTransformation 10 | import Categories.Category.Instance.Setoids 11 | 12 | import Categories.Category.Cocartesian 13 | 14 | import SecondOrder.Arity 15 | import SecondOrder.Signature 16 | import SecondOrder.Metavariable 17 | import SecondOrder.Term 18 | import SecondOrder.VRenaming 19 | 20 | module SecondOrder.MRenaming 21 | {ℓ} 22 | {𝔸 : SecondOrder.Arity.Arity} 23 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 24 | where 25 | 26 | open SecondOrder.Signature.Signature Σ 27 | open SecondOrder.Metavariable Σ 28 | open SecondOrder.Term Σ 29 | open SecondOrder.VRenaming Σ 30 | 31 | -- a metarenaming maps metavariables between contexts in an arity-preserving way 32 | _⇒ᵐ_ : ∀ (Θ ψ : MContext) → Set ℓ 33 | Θ ⇒ᵐ ψ = ∀ {Δ A} → [ Δ , A ]∈ Θ → [ Δ , A ]∈ ψ 34 | 35 | infix 4 _⇒ᵐ_ 36 | 37 | -- equality of metarenamings 38 | 39 | _≡ᵐ_ : ∀ {Θ ψ} (ι μ : Θ ⇒ᵐ ψ) → Set ℓ 40 | _≡ᵐ_ {Θ} ι μ = ∀ {Δ A} (M : [ Δ , A ]∈ Θ) → ι M ≡ μ M 41 | 42 | infixl 3 _≡ᵐ_ 43 | 44 | ≡ᵐ-refl : ∀ {Θ ψ} {ι : Θ ⇒ᵐ ψ} → ι ≡ᵐ ι 45 | ≡ᵐ-refl = λ M → refl 46 | 47 | ≡ᵐ-sym : ∀ {Θ ψ} {ι μ : Θ ⇒ᵐ ψ} 48 | → ι ≡ᵐ μ 49 | → μ ≡ᵐ ι 50 | ≡ᵐ-sym eq M = sym (eq M) 51 | 52 | ≡ᵐ-trans : ∀ {Θ ψ} {ι μ δ : Θ ⇒ᵐ ψ} 53 | → ι ≡ᵐ μ 54 | → μ ≡ᵐ δ 55 | → ι ≡ᵐ δ 56 | ≡ᵐ-trans eq1 eq2 M = trans (eq1 M) (eq2 M) 57 | 58 | -- meta-variable renamings form a setoid 59 | 60 | Mrenaming-setoid : ∀ (Θ ψ : MContext) → Setoid ℓ ℓ 61 | Mrenaming-setoid Θ ψ = 62 | record 63 | { Carrier = Θ ⇒ᵐ ψ 64 | ; _≈_ = λ ι μ → ι ≡ᵐ μ 65 | ; isEquivalence = 66 | record 67 | { refl = λ {ι} M → ≡ᵐ-refl {Θ} {ψ} {ι} M 68 | ; sym = ≡ᵐ-sym 69 | ; trans = ≡ᵐ-trans 70 | } 71 | } 72 | 73 | -- the identity renaming 74 | 75 | idᵐ : ∀ {Θ : MContext} → Θ ⇒ᵐ Θ 76 | idᵐ M = M 77 | 78 | -- equal metavariable renaming act the same on metavariables 79 | 80 | 81 | -- composition of renamings 82 | _∘ᵐ_ : ∀ {Θ ψ Ω} → ψ ⇒ᵐ Ω → Θ ⇒ᵐ ψ → Θ ⇒ᵐ Ω 83 | (ι ∘ᵐ μ) M = ι (μ M) 84 | 85 | infix 7 _∘ᵐ_ 86 | 87 | -- composition respects equality 88 | ∘ᵐ-resp-≡ᵐ : ∀ {Γ Δ Ξ} {τ₁ τ₂ : Δ ⇒ᵐ Ξ} {σ₁ σ₂ : Γ ⇒ᵐ Δ} → 89 | τ₁ ≡ᵐ τ₂ → σ₁ ≡ᵐ σ₂ → τ₁ ∘ᵐ σ₁ ≡ᵐ τ₂ ∘ᵐ σ₂ 90 | ∘ᵐ-resp-≡ᵐ {τ₁ = τ₁} {σ₂ = σ₂} ζ ξ x = trans (cong τ₁ (ξ x)) (ζ (σ₂ x)) 91 | 92 | -- the identity is the unit 93 | 94 | identity-leftᵐ : ∀ {Γ Δ} {ρ : Γ ⇒ᵐ Δ} → idᵐ ∘ᵐ ρ ≡ᵐ ρ 95 | identity-leftᵐ ρ = refl 96 | 97 | identity-rightᵐ : ∀ {Γ Δ} {ρ : Γ ⇒ᵐ Δ} → ρ ∘ᵐ idᵐ ≡ᵐ ρ 98 | identity-rightᵐ ρ = refl 99 | 100 | -- composition is associative 101 | 102 | assocᵐ : ∀ {Γ Δ Ξ Ψ} {τ : Γ ⇒ᵐ Δ} {ρ : Δ ⇒ᵐ Ξ} {σ : Ξ ⇒ᵐ Ψ} → 103 | (σ ∘ᵐ ρ) ∘ᵐ τ ≡ᵐ σ ∘ᵐ (ρ ∘ᵐ τ) 104 | assocᵐ x = refl 105 | 106 | sym-assocᵐ : ∀ {Γ Δ Ξ Ψ} {τ : Γ ⇒ᵐ Δ} {ρ : Δ ⇒ᵐ Ξ} {σ : Ξ ⇒ᵐ Ψ} → 107 | σ ∘ᵐ (ρ ∘ᵐ τ) ≡ᵐ (σ ∘ᵐ ρ) ∘ᵐ τ 108 | sym-assocᵐ x = refl 109 | 110 | -- contexts and renamings form a category 111 | module _ where 112 | open Categories.Category 113 | 114 | MContexts : Category ℓ ℓ ℓ 115 | MContexts = 116 | record 117 | { Obj = MContext 118 | ; _⇒_ = _⇒ᵐ_ 119 | ; _≈_ = _≡ᵐ_ 120 | ; id = idᵐ 121 | ; _∘_ = _∘ᵐ_ 122 | ; assoc = λ {_} {_} {_} {_} {f} {g} {h} {_} → assocᵐ {τ = f} {ρ = g} {σ = h} 123 | ; sym-assoc = λ {_} {_} {_} {_} {f} {g} {h} {_} → sym-assocᵐ {τ = f} {ρ = g} {σ = h} 124 | ; identityˡ = λ x → refl 125 | ; identityʳ = λ x → refl 126 | ; identity² = λ x → refl 127 | ; equiv = record { refl = λ {ι} {_} → ≡ᵐ-refl {ι = ι} ; sym = ≡ᵐ-sym ; trans = ≡ᵐ-trans } 128 | ; ∘-resp-≈ = ∘ᵐ-resp-≡ᵐ 129 | } 130 | 131 | 132 | -- the coproduct structure of the category 133 | module _ where 134 | 135 | infixl 7 [_,_]ᵐ 136 | 137 | [_,_]ᵐ : ∀ {Γ Δ Ξ} → Γ ⇒ᵐ Ξ → Δ ⇒ᵐ Ξ → Γ ,, Δ ⇒ᵐ Ξ 138 | [ σ , τ ]ᵐ (var-inl x) = σ x 139 | [ σ , τ ]ᵐ (var-inr y) = τ y 140 | 141 | inlᵐ : ∀ {Γ Δ} → Γ ⇒ᵐ Γ ,, Δ 142 | inlᵐ = var-inl 143 | 144 | inrᵐ : ∀ {Γ Δ} → Δ ⇒ᵐ Γ ,, Δ 145 | inrᵐ = var-inr 146 | 147 | uniqueᵐ : ∀ {Γ Δ Ξ} {τ : Γ ,, Δ ⇒ᵐ Ξ} {ρ : Γ ⇒ᵐ Ξ} {σ : Δ ⇒ᵐ Ξ} 148 | → τ ∘ᵐ inlᵐ ≡ᵐ ρ 149 | → τ ∘ᵐ inrᵐ ≡ᵐ σ 150 | → [ ρ , σ ]ᵐ ≡ᵐ τ 151 | uniqueᵐ ξ ζ (var-inl x) = sym (ξ x) 152 | uniqueᵐ ξ ζ (var-inr y) = sym (ζ y) 153 | 154 | MContext-+ : Categories.Category.Cocartesian.BinaryCoproducts MContexts 155 | MContext-+ = 156 | record { 157 | coproduct = 158 | λ {Γ Δ} → 159 | record 160 | { A+B = Γ ,, Δ 161 | ; i₁ = inlᵐ 162 | ; i₂ = inrᵐ 163 | ; [_,_] = [_,_]ᵐ 164 | ; inject₁ = λ x → refl 165 | ; inject₂ = λ x → refl 166 | ; unique = uniqueᵐ 167 | } 168 | } 169 | 170 | open Categories.Category.Cocartesian.BinaryCoproducts MContext-+ 171 | 172 | -- the renaming from the empty context 173 | 174 | inᵐ : ∀ {Γ} → ctx-empty ⇒ᵐ Γ 175 | inᵐ () 176 | 177 | -- extension of a renaming is summing with identity 178 | ⇑ᵐ : ∀ {Θ Ψ Ω} → Θ ⇒ᵐ Ψ → Θ ,, Ω ⇒ᵐ Ψ ,, Ω 179 | ⇑ᵐ ρ = ρ +₁ idᵐ 180 | 181 | -- a renaming can also be extended on the right 182 | ᵐ⇑ᵐ : ∀ {Θ Ψ} → Θ ⇒ᵐ Ψ → ∀ {Ω} → Ω ,, Θ ⇒ᵐ Ω ,, Ψ 183 | ᵐ⇑ᵐ ρ = idᵐ +₁ ρ 184 | 185 | 186 | -- the action of a metavariable renaming on terms 187 | infix 6 [_]ᵐ_ 188 | 189 | [_]ᵐ_ : ∀ {Θ Ψ Γ A} → Θ ⇒ᵐ Ψ → Term Θ Γ A → Term Ψ Γ A 190 | [ ι ]ᵐ (tm-var x) = tm-var x 191 | [ ι ]ᵐ (tm-meta M ts) = tm-meta (ι M) (λ i → [ ι ]ᵐ ts i) 192 | [ ι ]ᵐ (tm-oper f es) = tm-oper f (λ i → [ ι ]ᵐ es i) 193 | 194 | -- The sum of identities is an identity 195 | idᵐ+idᵐ : ∀ {Θ Ψ} → idᵐ {Θ = Θ} +₁ idᵐ {Θ = Ψ} ≡ᵐ idᵐ {Θ = Θ ,, Ψ} 196 | idᵐ+idᵐ (var-inl x) = refl 197 | idᵐ+idᵐ (var-inr y) = refl 198 | 199 | -- The action of a renaming respects equality of terms 200 | []ᵐ-resp-≈ : ∀ {Θ Ψ Γ A} {s t : Term Θ Γ A} {ι : Θ ⇒ᵐ Ψ} → s ≈ t → [ ι ]ᵐ s ≈ [ ι ]ᵐ t 201 | []ᵐ-resp-≈ (≈-≡ refl) = ≈-≡ refl 202 | []ᵐ-resp-≈ (≈-meta ξ) = ≈-meta (λ i → []ᵐ-resp-≈ (ξ i)) 203 | []ᵐ-resp-≈ (≈-oper ξ) = ≈-oper (λ i → []ᵐ-resp-≈ (ξ i)) 204 | 205 | -- The action of a renaming respects equality of renamings 206 | []ᵐ-resp-≡ᵐ : ∀ {Θ Ψ Γ A} {ι μ : Θ ⇒ᵐ Ψ} {t : Term Θ Γ A} → ι ≡ᵐ μ → [ ι ]ᵐ t ≈ [ μ ]ᵐ t 207 | []ᵐ-resp-≡ᵐ {t = tm-var x} ξ = ≈-≡ refl 208 | []ᵐ-resp-≡ᵐ {Θ} {Ψ} {Γ} {A} {ι = ι} {μ = μ} {t = tm-meta M ts} ξ = 209 | let open SetoidR (Term-setoid Ψ Γ A) in 210 | begin 211 | tm-meta (ι M) (λ i → [ ι ]ᵐ ts i) ≈⟨ ≈-meta (λ i → []ᵐ-resp-≡ᵐ ξ) ⟩ 212 | tm-meta (ι M) (λ i → [ μ ]ᵐ ts i) ≈⟨ ≈-≡ ((cong λ N → tm-meta N (λ i → [ μ ]ᵐ ts i)) (ξ M)) ⟩ 213 | tm-meta (μ M) (λ i → [ μ ]ᵐ ts i) ≈⟨ ≈-≡ refl ⟩ 214 | tm-meta (μ M) (λ i → [ μ ]ᵐ ts i) 215 | ∎ 216 | []ᵐ-resp-≡ᵐ {t = tm-oper f es} ξ = ≈-oper λ i → []ᵐ-resp-≡ᵐ ξ 217 | 218 | -- The action of the identity is trival 219 | [idᵐ] : ∀ {Θ Γ A} {t : Term Θ Γ A} → [ idᵐ ]ᵐ t ≈ t 220 | [idᵐ] {t = tm-var x} = ≈-refl 221 | [idᵐ] {t = tm-meta M ts} = ≈-meta λ i → [idᵐ] 222 | [idᵐ] {t = tm-oper f es} = ≈-oper λ i → [idᵐ] 223 | 224 | -- Extension respects composition 225 | ⇑ᵐ-resp-∘ᵐ : ∀ {Γ Δ Ξ Ψ} {ρ : Γ ⇒ᵐ Δ} {τ : Δ ⇒ᵐ Ξ} → ⇑ᵐ {Ω = Ψ} (τ ∘ᵐ ρ) ≡ᵐ (⇑ᵐ τ) ∘ᵐ (⇑ᵐ ρ) 226 | ⇑ᵐ-resp-∘ᵐ (var-inl x) = refl 227 | ⇑ᵐ-resp-∘ᵐ (var-inr y) = refl 228 | 229 | ᵐ⇑ᵐ-resp-∘ᵐ : ∀ {Θ Ψ Ω Ξ} {ρ : Θ ⇒ᵐ Ψ} {τ : Ψ ⇒ᵐ Ω} 230 | → ᵐ⇑ᵐ {Θ} {Ω} (τ ∘ᵐ ρ) {Ξ} ≡ᵐ (ᵐ⇑ᵐ τ) ∘ᵐ (ᵐ⇑ᵐ ρ) 231 | ᵐ⇑ᵐ-resp-∘ᵐ (var-inl M) = refl 232 | ᵐ⇑ᵐ-resp-∘ᵐ (var-inr N) = refl 233 | 234 | -- Extension of the identity renaming is the identity 235 | ⇑ᵐ-resp-idᵐ : ∀ {Θ Ψ} → (⇑ᵐ {Θ} {Θ} {Ψ}) (idᵐ {Θ}) ≡ᵐ idᵐ 236 | ⇑ᵐ-resp-idᵐ (var-inl M) = refl 237 | ⇑ᵐ-resp-idᵐ (var-inr N) = refl 238 | 239 | ᵐ⇑ᵐ-resp-idᵐ : ∀ {Θ Ψ} → (ᵐ⇑ᵐ {Θ} {Θ}) (idᵐ {Θ}) {Ψ} ≡ᵐ idᵐ 240 | ᵐ⇑ᵐ-resp-idᵐ (var-inl M) = refl 241 | ᵐ⇑ᵐ-resp-idᵐ (var-inr N) = refl 242 | 243 | -- Extension preserves equality of metavariable renamings 244 | ᵐ⇑ᵐ-resp-≡ᵐ : ∀ {Θ Ψ Ω} {ι μ : Θ ⇒ᵐ Ψ} → ι ≡ᵐ μ → ᵐ⇑ᵐ ι {Ω} ≡ᵐ ᵐ⇑ᵐ μ 245 | ᵐ⇑ᵐ-resp-≡ᵐ ι≡μ (var-inl M) = refl 246 | ᵐ⇑ᵐ-resp-≡ᵐ {ι = ι} ι≡μ (var-inr N) = cong (inrᵐ) (ι≡μ N) 247 | 248 | ⇑ᵐ-resp-≡ᵐ : ∀ {Θ Ψ Ω} {ι μ : Θ ⇒ᵐ Ψ} → ι ≡ᵐ μ → ⇑ᵐ {Ω = Ω} ι ≡ᵐ ⇑ᵐ μ 249 | ⇑ᵐ-resp-≡ᵐ ι≡μ (var-inl M) = cong var-inl (ι≡μ M) 250 | ⇑ᵐ-resp-≡ᵐ {ι = ι} ι≡μ (var-inr N) = refl 251 | 252 | 253 | -- The action of a renaming is functorial 254 | [∘ᵐ] : ∀ {Θ Ψ Ω Γ} {ι : Θ ⇒ᵐ Ψ} {μ : Ψ ⇒ᵐ Ω} {A} {t : Term Θ Γ A} 255 | → [ μ ∘ᵐ ι ]ᵐ t ≈ [ μ ]ᵐ ([ ι ]ᵐ t) 256 | [∘ᵐ] {t = tm-var x} = ≈-refl 257 | [∘ᵐ] {t = tm-meta M ts} = ≈-meta (λ i → [∘ᵐ]) 258 | [∘ᵐ] {t = tm-oper f es} = ≈-oper (λ i → [∘ᵐ]) 259 | 260 | ᵐ∘ᵛ≈ᵛ∘ᵐ : ∀ {Θ Ψ Γ Δ A} {ρ : Γ ⇒ᵛ Δ} {ι : Θ ⇒ᵐ Ψ} {t : Term Θ Γ A} 261 | → [ ι ]ᵐ ([ ρ ]ᵛ t) ≈ [ ρ ]ᵛ ([ ι ]ᵐ t) 262 | ᵐ∘ᵛ≈ᵛ∘ᵐ {t = tm-var x} = ≈-refl 263 | ᵐ∘ᵛ≈ᵛ∘ᵐ {t = tm-meta M ts} = ≈-meta (λ i → ᵐ∘ᵛ≈ᵛ∘ᵐ {t = ts i}) 264 | ᵐ∘ᵛ≈ᵛ∘ᵐ {t = tm-oper f es} = ≈-oper (λ i → ᵐ∘ᵛ≈ᵛ∘ᵐ {t = es i}) 265 | 266 | split-sum : ∀ {Θ Ψ Ξ Ω} {ι : Θ ⇒ᵐ Ψ} {μ : Ξ ⇒ᵐ Ω} 267 | → (μ +₁ ι) ≡ᵐ (⇑ᵐ μ) ∘ᵐ (ᵐ⇑ᵐ ι) 268 | split-sum (var-inl M) = refl 269 | split-sum (var-inr N) = refl 270 | 271 | split-sum2 : ∀ {Θ Ψ Ξ Ω} {ι : Θ ⇒ᵐ Ψ} {μ : Ξ ⇒ᵐ Ω} 272 | → (μ +₁ ι) ≡ᵐ (ᵐ⇑ᵐ ι) ∘ᵐ (⇑ᵐ μ) 273 | split-sum2 (var-inl M) = refl 274 | split-sum2 (var-inr N) = refl 275 | 276 | ⇑-resp-+ : ∀ {Θ Ψ Ξ Ω Γ A} {ι : Θ ⇒ᵐ Ψ} {μ : Ξ ⇒ᵐ Ω} {t : Term (Ξ + Θ) Γ A} 277 | → [ (⇑ᵐ μ) ]ᵐ ([ (ᵐ⇑ᵐ ι) ]ᵐ t) ≈ [ (ᵐ⇑ᵐ ι) ]ᵐ ([ (⇑ᵐ μ) ]ᵐ t) 278 | ⇑-resp-+ {Θ} {Ψ} {Ξ} {Ω} {Γ} {A} {ι} {μ} {t = t} = 279 | let open SetoidR (Term-setoid (Ω ,, Ψ) Γ A) in 280 | begin 281 | [ ⇑ᵐ μ ]ᵐ ([ ᵐ⇑ᵐ ι ]ᵐ t) ≈⟨ ≈-sym [∘ᵐ] ⟩ 282 | [ (⇑ᵐ μ) ∘ᵐ (ᵐ⇑ᵐ ι) ]ᵐ t ≈⟨ ≈-sym ([]ᵐ-resp-≡ᵐ split-sum) ⟩ 283 | [ (μ +₁ ι) ]ᵐ t ≈⟨ []ᵐ-resp-≡ᵐ split-sum2 ⟩ 284 | [(ᵐ⇑ᵐ ι) ∘ᵐ (⇑ᵐ μ) ]ᵐ t ≈⟨ [∘ᵐ] ⟩ 285 | [ ᵐ⇑ᵐ ι ]ᵐ ([ ⇑ᵐ μ ]ᵐ t) 286 | ∎ 287 | 288 | ∘ᵐ-resp-⇑ : ∀ {Θ Ψ Ξ Ω} {ι : Θ ⇒ᵐ Ψ} {μ : Ψ ⇒ᵐ Ω} 289 | → ⇑ᵐ {Ω = Ξ} (μ ∘ᵐ ι) ≡ᵐ ⇑ᵐ μ ∘ᵐ ⇑ᵐ ι 290 | ∘ᵐ-resp-⇑ (var-inl M) = refl 291 | ∘ᵐ-resp-⇑ (var-inr N) = refl 292 | 293 | ∘ᵐ-resp-⇑-term : ∀ {Θ Ψ Ξ Ω Γ A} {ι : Θ ⇒ᵐ Ψ} {μ : Ψ ⇒ᵐ Ω} {t : Term (Θ ,, Ξ) Γ A} 294 | → [ ⇑ᵐ {Ω = Ξ} (μ ∘ᵐ ι) ]ᵐ t ≈ [ ⇑ᵐ μ ]ᵐ ([ ⇑ᵐ ι ]ᵐ t) 295 | ∘ᵐ-resp-⇑-term {Θ} {Ψ} {Ξ} {Ω} {Γ} {A} {ι} {μ} {t = t} = 296 | let open SetoidR (Term-setoid (Ω ,, Ξ) Γ A) in 297 | begin 298 | [ ⇑ᵐ {Ω = Ξ} (μ ∘ᵐ ι) ]ᵐ t ≈⟨ []ᵐ-resp-≡ᵐ ∘ᵐ-resp-⇑ ⟩ 299 | [ ⇑ᵐ μ ∘ᵐ ⇑ᵐ ι ]ᵐ t ≈⟨ [∘ᵐ] ⟩ 300 | [ ⇑ᵐ μ ]ᵐ ([ ⇑ᵐ ι ]ᵐ t) 301 | ∎ 302 | 303 | ∘ᵐ-resp-ᵐ⇑ : ∀ {Θ Ψ Ξ Ω} {ι : Θ ⇒ᵐ Ψ} {μ : Ψ ⇒ᵐ Ω} 304 | → ᵐ⇑ᵐ (μ ∘ᵐ ι) {Ω = Ξ} ≡ᵐ ᵐ⇑ᵐ μ ∘ᵐ ᵐ⇑ᵐ ι 305 | ∘ᵐ-resp-ᵐ⇑ (var-inl M) = refl 306 | ∘ᵐ-resp-ᵐ⇑ (var-inr N) = refl 307 | 308 | ∘ᵐ-resp-ᵐ⇑-term : ∀ {Θ Ψ Ξ Ω Γ A} {ι : Θ ⇒ᵐ Ψ} {μ : Ψ ⇒ᵐ Ω} {t : Term (Ξ ,, Θ) Γ A} 309 | → [ ᵐ⇑ᵐ (μ ∘ᵐ ι) {Ω = Ξ} ]ᵐ t ≈ [ ᵐ⇑ᵐ μ ]ᵐ ([ ᵐ⇑ᵐ ι ]ᵐ t) 310 | ∘ᵐ-resp-ᵐ⇑-term {Θ} {Ψ} {Ξ} {Ω} {Γ} {A} {ι} {μ} {t = t} = 311 | let open SetoidR (Term-setoid (Ξ ,, Ω) Γ A) in 312 | begin 313 | [ ᵐ⇑ᵐ (μ ∘ᵐ ι) {Ω = Ξ} ]ᵐ t ≈⟨ []ᵐ-resp-≡ᵐ ∘ᵐ-resp-ᵐ⇑ ⟩ 314 | [ ᵐ⇑ᵐ μ ∘ᵐ ᵐ⇑ᵐ ι ]ᵐ t ≈⟨ [∘ᵐ] ⟩ 315 | [ ᵐ⇑ᵐ μ ]ᵐ ([ ᵐ⇑ᵐ ι ]ᵐ t) 316 | ∎ 317 | 318 | 319 | vr-comm-mr : ∀ {Θ Ψ Γ Δ A} {ρ : Γ ⇒ᵛ Δ} {ι : Θ ⇒ᵐ Ψ} {t : Term Θ Γ A} 320 | → [ ι ]ᵐ ([ ρ ]ᵛ t) ≈ [ ρ ]ᵛ ([ ι ]ᵐ t) 321 | vr-comm-mr {t = tm-var x} = ≈-refl 322 | vr-comm-mr {t = tm-meta M ts} = ≈-meta (λ i → vr-comm-mr) 323 | vr-comm-mr {t = tm-oper f es} = ≈-oper (λ i → vr-comm-mr) 324 | 325 | mr-comm-vr : ∀ {Θ Ψ Γ Δ A} {ρ : Γ ⇒ᵛ Δ} {ι : Θ ⇒ᵐ Ψ} {t : Term Θ Γ A} 326 | → [ ρ ]ᵛ ([ ι ]ᵐ t) ≈ [ ι ]ᵐ ([ ρ ]ᵛ t) 327 | mr-comm-vr {t = tm-var x} = ≈-refl 328 | mr-comm-vr {t = tm-meta M ts} = ≈-meta (λ i → mr-comm-vr) 329 | mr-comm-vr {t = tm-oper f es} = ≈-oper (λ i → mr-comm-vr) 330 | 331 | module _ {Θ Ψ : MContext} {A : sort} where 332 | open Categories.Category 333 | open Categories.Category.Instance.Setoids 334 | open Categories.Functor 335 | open Categories.NaturalTransformation 336 | 337 | MRenaming-NT : ∀ (ι : Θ ⇒ᵐ Ψ) → NaturalTransformation (Term-Functor {Θ} {A}) (Term-Functor {Ψ} {A}) 338 | MRenaming-NT ι = 339 | record 340 | { η = λ Γ → record { _⟨$⟩_ = [ ι ]ᵐ_ ; cong = []ᵐ-resp-≈ } 341 | ; commute = λ ρ t≈s → ≈-trans ([]ᵐ-resp-≈ ([]ᵛ-resp-≈ t≈s)) (ᵐ∘ᵛ≈ᵛ∘ᵐ) 342 | ; sym-commute = λ ρ t≈s → ≈-trans (≈-sym ᵐ∘ᵛ≈ᵛ∘ᵐ) ([]ᵐ-resp-≈ ([]ᵛ-resp-≈ t≈s)) 343 | } 344 | -------------------------------------------------------------------------------- /src/SecondOrder/Metavariable.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | 3 | import SecondOrder.Arity 4 | import SecondOrder.MContext 5 | import SecondOrder.Signature 6 | 7 | module SecondOrder.Metavariable 8 | {ℓ} 9 | {𝔸 : SecondOrder.Arity.Arity} 10 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 11 | where 12 | 13 | open SecondOrder.Signature.Signature Σ 14 | 15 | open SecondOrder.MContext VContext sort public 16 | -------------------------------------------------------------------------------- /src/SecondOrder/Mslot.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --allow-unsolved-metas #-} 2 | 3 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 4 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; cong; subst; setoid) 5 | open import Data.Product using (_×_; Σ; _,_; proj₁; proj₂; zip; map; <_,_>; swap) 6 | import Function.Equality 7 | -- open import Relation.Binary using (Setoid) 8 | -- import Relation.Binary.Reasoning.Setoid as SetoidR 9 | 10 | import Categories.Category 11 | import Categories.Functor 12 | import Categories.Category.Instance.Setoids 13 | import Categories.Monad.Relative 14 | import Categories.Category.Equivalence 15 | import Categories.Category.Cocartesian 16 | import Categories.Category.Construction.Functors 17 | import Categories.Category.Product 18 | import Categories.NaturalTransformation 19 | import Categories.NaturalTransformation.NaturalIsomorphism 20 | 21 | import SecondOrder.Arity 22 | import SecondOrder.Signature 23 | import SecondOrder.Metavariable 24 | import SecondOrder.VRenaming 25 | import SecondOrder.MRenaming 26 | import SecondOrder.Term 27 | import SecondOrder.Substitution 28 | import SecondOrder.RelativeMonadMorphism 29 | import SecondOrder.Instantiation 30 | import SecondOrder.IndexedCategory 31 | import SecondOrder.RelativeKleisli 32 | 33 | 34 | module SecondOrder.Mslot 35 | {ℓ} 36 | {𝔸 : SecondOrder.Arity.Arity} 37 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 38 | where 39 | 40 | open SecondOrder.Signature.Signature Σ 41 | open SecondOrder.Metavariable Σ 42 | open SecondOrder.Term Σ 43 | open SecondOrder.VRenaming Σ 44 | open SecondOrder.MRenaming Σ 45 | -- open SecondOrder.Substitution Σ 46 | -- open import SecondOrder.RelativeMonadMorphism 47 | -- open SecondOrder.Instantiation 48 | open Categories.Category 49 | open Categories.Functor using (Functor) 50 | open Categories.NaturalTransformation renaming (id to idNt) 51 | open Categories.NaturalTransformation.NaturalIsomorphism renaming (refl to reflNt; sym to symNt; trans to transNt) 52 | open Categories.Category.Construction.Functors 53 | open Categories.Category.Instance.Setoids 54 | open Categories.Category.Product 55 | open Function.Equality using () renaming (setoid to Π-setoid) 56 | open SecondOrder.IndexedCategory 57 | -- open import SecondOrder.RelativeKleisli 58 | 59 | MTele = MContexts 60 | VTele = VContexts 61 | 62 | -- the codomain category of the MSlots functor. It should be equivalent to the 63 | -- functor category [ MTele x VTele , < Setoid >ₛₒᵣₜ ] 64 | -- objects are functors, which are really pairs of functions, one on objects 65 | -- one on morphisms 66 | -- morphisms in this category are natural transformations 67 | 68 | module _ where 69 | open Category 70 | open NaturalTransformation 71 | open Function.Equality renaming (_∘_ to _∙_) 72 | 73 | Mslots : Functor MContexts (IndexedCategory VContext (IndexedCategory sort (Setoids ℓ ℓ))) 74 | Mslots = 75 | let open Categories.NaturalTransformation in 76 | let open NaturalTransformation in 77 | let open Relation.Binary.PropositionalEquality.≡-Reasoning in 78 | record 79 | { F₀ = λ Θ Γ A → setoid ([ Γ , A ]∈ Θ) 80 | ; F₁ = λ ι Γ A → record { _⟨$⟩_ = λ M → ι M ; cong = λ M≡N → cong ι M≡N } 81 | ; identity = λ {Θ} Γ A {M} {N} M≡N → cong idᵐ M≡N 82 | ; homomorphism = λ {Θ} {Ψ} {Ξ} {ι} {μ} Γ A M≡N → cong (μ ∘ᵐ ι) M≡N 83 | ; F-resp-≈ = λ {Θ} {Ψ} {ι} {μ} ι≡ᵐμ Γ A {M} {N} M≡N → 84 | begin 85 | ι M ≡⟨ cong ι M≡N ⟩ 86 | ι N ≡⟨ ι≡ᵐμ N ⟩ 87 | μ N 88 | ∎ 89 | } 90 | -------------------------------------------------------------------------------- /src/SecondOrder/README.md: -------------------------------------------------------------------------------- 1 | This folder contains the formalization of second-order multi-sorted theories, where by *second-order* we mean that term 2 | symbols may bind variables. 3 | -------------------------------------------------------------------------------- /src/SecondOrder/RelMon.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_; lsuc) 2 | 3 | open import Categories.Category 4 | open import Categories.Functor 5 | import Categories.Category.Cartesian as Cartesian 6 | open import Categories.Monad.Relative 7 | 8 | import SecondOrder.RelativeKleisli 9 | open import SecondOrder.RelativeMonadMorphism 10 | 11 | -- The category of relative monads and relative monad morphisms 12 | 13 | module SecondOrder.RelMon 14 | {o l e o' l' e'} 15 | {𝒞 : Category o l e} 16 | {𝒟 : Category o' l' e'} 17 | {J : Functor 𝒞 𝒟} 18 | where 19 | 20 | RelMon : Category (o ⊔ o' ⊔ l' ⊔ e') (lsuc o ⊔ lsuc l' ⊔ lsuc e') (o ⊔ e') 21 | RelMon = 22 | let open Category 𝒟 renaming (id to id_D; identityˡ to identˡ; identityʳ to identʳ; assoc to ass) in 23 | let open RMonadMorph in 24 | let open Monad in 25 | let open HomReasoning in 26 | record 27 | { Obj = Monad J 28 | ; _⇒_ = λ M N → RMonadMorph M N 29 | ; _≈_ = λ {M} {N} φ ψ → ∀ X → morph φ {X} ≈ morph ψ {X} 30 | ; id = λ {M} → 31 | record 32 | { morph = λ {X} → id_D {F₀ M X} 33 | ; law-unit = λ {X} → identˡ 34 | ; law-extend = λ {X} {Y} {f} → 35 | begin 36 | (id_D ∘ extend M f) ≈⟨ identˡ ⟩ 37 | extend M f ≈⟨ extend-≈ M {k = f} {h = id_D ∘ f} (⟺ identˡ) ⟩ 38 | extend M (id_D ∘ f) ≈⟨ ⟺ identʳ ⟩ 39 | (extend M (id_D ∘ f) ∘ id_D) 40 | ∎ 41 | } 42 | ; _∘_ = λ {M} {N} {L} φ ψ → 43 | record 44 | { morph = λ {X} → (morph φ {X}) ∘ (morph ψ {X}) 45 | ; law-unit = λ {X} → 46 | begin 47 | ((morph φ ∘ morph ψ) ∘ unit M) ≈⟨ ass ⟩ 48 | (morph φ ∘ (morph ψ ∘ unit M)) ≈⟨ refl⟩∘⟨ law-unit ψ ⟩ 49 | (morph φ ∘ unit N) ≈⟨ law-unit φ ⟩ 50 | unit L 51 | ∎ 52 | ; law-extend = λ {X} {Y} {f} → 53 | begin 54 | (morph φ ∘ morph ψ) ∘ extend M f ≈⟨ ass ⟩ 55 | morph φ ∘ (morph ψ ∘ extend M f) ≈⟨ refl⟩∘⟨ law-extend ψ ⟩ 56 | morph φ ∘ (extend N (morph ψ ∘ f) ∘ morph ψ) ≈⟨ ⟺ (ass) ⟩ 57 | (morph φ ∘ extend N (morph ψ ∘ f)) ∘ morph ψ ≈⟨ (law-extend φ ⟩∘⟨refl) ⟩ 58 | ((extend L (morph φ ∘ morph ψ ∘ f)) ∘ morph φ) ∘ morph ψ ≈⟨ ass ⟩ 59 | (extend L (morph φ ∘ morph ψ ∘ f)) ∘ morph φ ∘ morph ψ 60 | ≈⟨ ( extend-≈ L (⟺ (ass)) ⟩∘⟨refl) ⟩ 61 | extend L ((morph φ ∘ morph ψ) ∘ f) ∘ morph φ ∘ morph ψ 62 | ∎ 63 | } 64 | ; assoc = λ X → ass 65 | ; sym-assoc = λ X → ⟺ (ass) 66 | ; identityˡ = λ X → identˡ 67 | ; identityʳ = λ X → identʳ 68 | ; identity² = λ X → identˡ 69 | ; equiv = record 70 | { refl = λ X → Equiv.refl 71 | ; sym = λ φ≈ψ X → ⟺ (φ≈ψ X) 72 | ; trans = λ φ≈ψ ψ≈θ X → φ≈ψ X ○ ψ≈θ X 73 | } 74 | ; ∘-resp-≈ = λ φ≈ψ θ≈ω X → φ≈ψ X ⟩∘⟨ θ≈ω X 75 | } 76 | -------------------------------------------------------------------------------- /src/SecondOrder/RelativeKleisli.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_) 2 | open import Categories.Category 3 | open import Categories.Functor 4 | open import Categories.Monad.Relative 5 | 6 | module SecondOrder.RelativeKleisli 7 | {o l e o' l' e'} 8 | {𝒞 : Category o l e} 9 | {𝒟 : Category o' l' e'} 10 | {J : Functor 𝒞 𝒟} 11 | (M : Monad J) 12 | where 13 | 14 | Kleisli : Category o l' e' 15 | Kleisli = 16 | let open Category in 17 | let open Functor in 18 | record 19 | { Obj = Obj 𝒞 20 | ; _⇒_ = λ A B → 𝒟 [ F₀ J A , Monad.F₀ M B ] 21 | ; _≈_ = _≈_ 𝒟 22 | ; id = Monad.unit M 23 | ; _∘_ = λ f g → _∘_ 𝒟 (Monad.extend M f) g 24 | ; assoc = λ {A} {B} {C} {D} {f} {g} {h} 25 | → Equiv.trans 𝒟 (∘-resp-≈ˡ 𝒟 (Monad.assoc M)) (assoc 𝒟) 26 | ; sym-assoc = λ {A} {B} {C} {D} {f} {g} {h} 27 | → Equiv.sym 𝒟 (Equiv.trans 𝒟 (∘-resp-≈ˡ 𝒟 (Monad.assoc M)) (assoc 𝒟)) 28 | ; identityˡ = λ {A} {B} {f} 29 | → Equiv.trans 𝒟 (∘-resp-≈ˡ 𝒟 (Monad.identityˡ M)) (identityˡ 𝒟) 30 | ; identityʳ = λ {A} {B} {f} 31 | → Monad.identityʳ M 32 | ; identity² = λ {A} → Equiv.trans 𝒟 ((∘-resp-≈ˡ 𝒟 (Monad.identityˡ M))) (identityˡ 𝒟) 33 | ; equiv = equiv 𝒟 34 | ; ∘-resp-≈ = λ p₁ p₂ → ∘-resp-≈ 𝒟 (Monad.extend-≈ M p₁) p₂ 35 | } 36 | -------------------------------------------------------------------------------- /src/SecondOrder/RelativeMonadMorphism.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; setoid; cong; trans) 3 | import Function.Equality 4 | open import Relation.Binary using (Setoid) 5 | 6 | open import Categories.Category 7 | open import Categories.Functor 8 | open import Categories.Category.Instance.Setoids 9 | open import Categories.Monad.Relative 10 | open import Categories.Category.Equivalence 11 | open import Categories.Category.Cocartesian 12 | 13 | module SecondOrder.RelativeMonadMorphism 14 | {o ℓ e o′ ℓ′ e′ : Level} 15 | {C : Category o ℓ e} 16 | {D : Category o′ ℓ′ e′} 17 | where 18 | 19 | record RMonadMorph {J : Functor C D} (M M’ : Monad J) : Set (lsuc (o ⊔ ℓ′ ⊔ e′)) 20 | where 21 | open Category D 22 | open Monad 23 | field 24 | morph : ∀ {X} → (F₀ M X) ⇒ (F₀ M’ X) 25 | law-unit : ∀ {X} → morph ∘ (unit M {X}) ≈ unit M’ {X} 26 | law-extend : ∀ {X Y} {k : (Functor.F₀ J X) ⇒ (F₀ M Y)} 27 | → (morph {Y}) ∘ (extend M k) 28 | ≈ (extend M’ ((morph {Y}) ∘ k)) ∘ (morph {X}) 29 | 30 | -- here, the equality used is the equivalence relation between morphisms of the category D 31 | -------------------------------------------------------------------------------- /src/SecondOrder/Signature.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | 3 | import SecondOrder.Arity 4 | import SecondOrder.VContext 5 | 6 | module SecondOrder.Signature 7 | ℓ 8 | (𝔸 : SecondOrder.Arity.Arity) 9 | where 10 | 11 | open SecondOrder.Arity.Arity 𝔸 12 | 13 | -- a second-order algebraic signature 14 | record Signature : Set (lsuc ℓ) where 15 | 16 | -- a signature consists of a set of sorts and a set of operations 17 | -- e.g. sorts A, B, C, ... and operations f, g, h 18 | field 19 | sort : Set ℓ -- sorts 20 | oper : Set ℓ -- operations 21 | 22 | open SecondOrder.VContext sort public 23 | 24 | -- each operation has arity and a sort (the sort of its codomain) 25 | field 26 | oper-arity : oper → arity -- the arity of an operation 27 | oper-sort : oper → sort -- the operation sort 28 | arg-sort : ∀ (f : oper) → arg (oper-arity f) → sort -- the sorts of arguments 29 | -- a second order operation can bind some free variables that occur as its arguments 30 | -- e.g. the lambda operation binds one type and one term 31 | arg-bind : ∀ (f : oper) → arg (oper-arity f) → VContext -- the argument bindings 32 | 33 | -- the arguments of an operation 34 | oper-arg : oper → Set 35 | oper-arg f = arg (oper-arity f) 36 | -------------------------------------------------------------------------------- /src/SecondOrder/Substitution.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; setoid; cong; trans) 3 | import Function.Equality 4 | open import Relation.Binary using (Setoid) 5 | 6 | import Categories.Category 7 | import Categories.Functor 8 | import Categories.Category.Instance.Setoids 9 | import Categories.Monad.Relative 10 | import Categories.Category.Equivalence 11 | import Categories.Category.Cocartesian 12 | 13 | import SecondOrder.Arity 14 | import SecondOrder.Signature 15 | import SecondOrder.Metavariable 16 | import SecondOrder.VRenaming 17 | import SecondOrder.MRenaming 18 | import SecondOrder.Term 19 | import SecondOrder.IndexedCategory 20 | import SecondOrder.RelativeKleisli 21 | 22 | module SecondOrder.Substitution 23 | {ℓ} 24 | {𝔸 : SecondOrder.Arity.Arity} 25 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 26 | where 27 | 28 | open SecondOrder.Signature.Signature Σ 29 | open SecondOrder.Metavariable Σ 30 | open SecondOrder.Term Σ 31 | open SecondOrder.MRenaming Σ 32 | open SecondOrder.VRenaming Σ 33 | 34 | -- substitution 35 | 36 | infix 4 _⊕_⇒ˢ_ 37 | 38 | _⊕_⇒ˢ_ : ∀ (Θ : MContext) (Γ Δ : VContext) → Set ℓ 39 | Θ ⊕ Γ ⇒ˢ Δ = ∀ {A} (x : A ∈ Γ) → Term Θ Δ A 40 | 41 | -- substitution preserves propositionnal equality 42 | 43 | σ-resp-≡ : ∀ {Θ Γ Δ A} {x y : A ∈ Γ} {σ : Θ ⊕ Γ ⇒ˢ Δ} → x ≡ y → σ x ≡ σ y 44 | σ-resp-≡ refl = refl 45 | 46 | -- syntactic equality of substitutions 47 | 48 | infix 5 _≈ˢ_ 49 | 50 | _≈ˢ_ : ∀ {Θ} {Γ Δ} (σ τ : Θ ⊕ Γ ⇒ˢ Δ) → Set ℓ 51 | _≈ˢ_ {Θ} {Γ} σ τ = ∀ {A} (x : A ∈ Γ) → σ x ≈ τ x 52 | 53 | -- equality of substitutions form a setoid 54 | ≈ˢ-refl : ∀ {Θ Γ Δ} {σ : Θ ⊕ Γ ⇒ˢ Δ} 55 | → σ ≈ˢ σ 56 | ≈ˢ-refl x = ≈-refl 57 | 58 | ≈ˢ-sym : ∀ {Θ Γ Δ} {σ τ : Θ ⊕ Γ ⇒ˢ Δ} 59 | → σ ≈ˢ τ 60 | → τ ≈ˢ σ 61 | ≈ˢ-sym eq x = ≈-sym (eq x) 62 | 63 | ≈ˢ-trans : ∀ {Θ Γ Δ} {σ τ μ : Θ ⊕ Γ ⇒ˢ Δ} 64 | → σ ≈ˢ τ → τ ≈ˢ μ 65 | → σ ≈ˢ μ 66 | ≈ˢ-trans eq1 eq2 x = ≈-trans (eq1 x) (eq2 x) 67 | 68 | substitution-setoid : ∀ (Γ Δ : VContext) (Θ : MContext) → Setoid ℓ ℓ 69 | substitution-setoid Γ Δ Θ = 70 | record 71 | { Carrier = Θ ⊕ Γ ⇒ˢ Δ 72 | ; _≈_ = λ σ τ → σ ≈ˢ τ 73 | ; isEquivalence = 74 | record 75 | { refl = λ {σ} x → ≈ˢ-refl {σ = σ} x 76 | ; sym = ≈ˢ-sym 77 | ; trans = ≈ˢ-trans 78 | } 79 | } 80 | 81 | congˢ : ∀ {Θ} {Γ Δ} {A} {σ τ : Θ ⊕ Γ ⇒ˢ Δ} {x : A ∈ Γ} → σ ≈ˢ τ → σ x ≈ τ x 82 | congˢ {x = x} eq = eq x 83 | 84 | congˢ-var : ∀ {Θ} {Γ Δ} {A} {σ : Θ ⊕ Γ ⇒ˢ Δ} {x y : A ∈ Γ} → x ≡ y → σ x ≈ σ y 85 | congˢ-var refl = ≈-refl 86 | 87 | -- extension of a substitution 88 | 89 | ⇑ˢ : ∀ {Θ Γ Δ Ξ} → Θ ⊕ Γ ⇒ˢ Δ → Θ ⊕ (Γ ,, Ξ) ⇒ˢ (Δ ,, Ξ) 90 | ⇑ˢ σ (var-inl x) = [ var-inl ]ᵛ σ x 91 | ⇑ˢ σ (var-inr y) = tm-var (var-inr y) 92 | 93 | -- extension respects equality of substitutions 94 | 95 | ⇑ˢ-resp-≈ˢ : ∀ {Θ Γ Δ Ξ} {σ τ : Θ ⊕ Γ ⇒ˢ Δ} → σ ≈ˢ τ → ⇑ˢ {Ξ = Ξ} σ ≈ˢ ⇑ˢ {Ξ = Ξ} τ 96 | ⇑ˢ-resp-≈ˢ ξ (var-inl x) = []ᵛ-resp-≈ (ξ x) 97 | ⇑ˢ-resp-≈ˢ ξ (var-inr y) = ≈-refl 98 | 99 | -- the compositions of renamings and substitutions 100 | 101 | infixr 6 _ᵛ∘ˢ_ 102 | 103 | _ᵛ∘ˢ_ : ∀ {Θ} {Γ Δ Ξ} (ρ : Δ ⇒ᵛ Ξ) (σ : Θ ⊕ Γ ⇒ˢ Δ) → Θ ⊕ Γ ⇒ˢ Ξ 104 | (ρ ᵛ∘ˢ σ) x = [ ρ ]ᵛ (σ x) 105 | 106 | infixl 6 _ˢ∘ᵛ_ 107 | 108 | _ˢ∘ᵛ_ : ∀ {Θ} {Γ Δ Ξ} (σ : Θ ⊕ Δ ⇒ˢ Ξ) (ρ : Γ ⇒ᵛ Δ) → Θ ⊕ Γ ⇒ˢ Ξ 109 | (σ ˢ∘ᵛ ρ) x = σ (ρ x) 110 | 111 | -- the composition of metarenamings and substitutions 112 | 113 | infixr 6 _ᵐ∘ˢ_ 114 | 115 | _ᵐ∘ˢ_ : ∀ {Θ Ω} {Γ Δ} (μ : Θ ⇒ᵐ Ω) (σ : Θ ⊕ Γ ⇒ˢ Δ) → Ω ⊕ Γ ⇒ˢ Δ 116 | (μ ᵐ∘ˢ σ) x = [ μ ]ᵐ σ x 117 | 118 | -- infixl 6 _ˢ∘ᵐ_ 119 | 120 | -- _ˢ∘ᵐ_ : ∀ {Ω Θ} {Γ Δ} (σ : Ω ⊕ Γ ⇒ˢ Δ) (μ : Θ ⇒ᵐ Ω) → Ω ⊕ Γ ⇒ˢ Δ 121 | -- (σ ˢ∘ᵐ μ) x = {![_]ˢ_!} 122 | 123 | 124 | -- extension commutes with renaming action 125 | 126 | ⇑ˢ-resp-ˢ∘ᵛ : ∀ {Θ} {Γ Δ Ξ Ψ} {ρ : Γ ⇒ᵛ Δ} {σ : Θ ⊕ Δ ⇒ˢ Ξ} → ⇑ˢ {Ξ = Ψ} (σ ˢ∘ᵛ ρ) ≈ˢ ⇑ˢ σ ˢ∘ᵛ ⇑ᵛ ρ 127 | ⇑ˢ-resp-ˢ∘ᵛ (var-inl x) = ≈-refl 128 | ⇑ˢ-resp-ˢ∘ᵛ (var-inr x) = ≈-refl 129 | 130 | -- the action of a substitution on a term 131 | 132 | infix 6 [_]ˢ_ 133 | 134 | [_]ˢ_ : ∀ {Θ Γ Δ A} → Θ ⊕ Γ ⇒ˢ Δ → Term Θ Γ A → Term Θ Δ A 135 | [ σ ]ˢ (tm-var x) = σ x 136 | [ σ ]ˢ (tm-meta M ts) = tm-meta M (λ i → [ σ ]ˢ ts i) 137 | [ σ ]ˢ (tm-oper f es) = tm-oper f (λ i → [ ⇑ˢ σ ]ˢ es i) 138 | 139 | -- composition of substitutions 140 | 141 | infixl 7 _∘ˢ_ 142 | _∘ˢ_ : ∀ {Θ} {Γ Δ Ξ} → Θ ⊕ Δ ⇒ˢ Ξ → Θ ⊕ Γ ⇒ˢ Δ → Θ ⊕ Γ ⇒ˢ Ξ 143 | (σ ∘ˢ τ) x = [ σ ]ˢ τ x 144 | 145 | -- substitution action respects equality of terms 146 | 147 | []ˢ-resp-≈ : ∀ {Θ} {Γ Δ} {A} (σ : Θ ⊕ Γ ⇒ˢ Δ) {t u : Term Θ Γ A} → t ≈ u → [ σ ]ˢ t ≈ [ σ ]ˢ u 148 | []ˢ-resp-≈ σ (≈-≡ refl) = ≈-refl 149 | []ˢ-resp-≈ σ (≈-meta ξ) = ≈-meta (λ i → []ˢ-resp-≈ σ (ξ i)) 150 | []ˢ-resp-≈ σ (≈-oper ξ) = ≈-oper (λ i → []ˢ-resp-≈ (⇑ˢ σ) (ξ i)) 151 | 152 | -- substitution action respects equality of substitutions 153 | 154 | []ˢ-resp-≈ˢ : ∀ {Θ} {Γ Δ} {A} {σ τ : Θ ⊕ Γ ⇒ˢ Δ} (t : Term Θ Γ A) → σ ≈ˢ τ → [ σ ]ˢ t ≈ [ τ ]ˢ t 155 | []ˢ-resp-≈ˢ (tm-var x) ξ = ξ x 156 | []ˢ-resp-≈ˢ (tm-meta M ts) ξ = ≈-meta (λ i → []ˢ-resp-≈ˢ (ts i) ξ) 157 | []ˢ-resp-≈ˢ (tm-oper f es) ξ = ≈-oper (λ i → []ˢ-resp-≈ˢ (es i) (⇑ˢ-resp-≈ˢ ξ)) 158 | 159 | -- substitution actions respects both equalities 160 | 161 | []ˢ-resp-≈ˢ-≈ : ∀ {Θ} {Γ Δ} {A} {σ τ : Θ ⊕ Γ ⇒ˢ Δ} {t u : Term Θ Γ A} → σ ≈ˢ τ → t ≈ u → [ σ ]ˢ t ≈ [ τ ]ˢ u 162 | []ˢ-resp-≈ˢ-≈ {τ = τ} {t = t} ζ ξ = ≈-trans ([]ˢ-resp-≈ˢ t ζ) ([]ˢ-resp-≈ τ ξ) 163 | 164 | -- identity substitution 165 | 166 | idˢ : ∀ {Θ Γ} → Θ ⊕ Γ ⇒ˢ Γ 167 | idˢ = tm-var 168 | 169 | -- extension preserves identity 170 | 171 | ⇑ˢ-resp-idˢ : ∀ {Θ} {Γ Δ} → ⇑ˢ idˢ ≈ˢ idˢ {Θ = Θ} {Γ = Γ ,, Δ} 172 | ⇑ˢ-resp-idˢ (var-inl x) = ≈-refl 173 | ⇑ˢ-resp-idˢ (var-inr y) = ≈-refl 174 | 175 | -- the identity substution acts trivially 176 | 177 | [idˢ] : ∀ {Θ} {Γ} {A} {t : Term Θ Γ A} → [ idˢ ]ˢ t ≈ t 178 | [idˢ] {t = tm-var x} = ≈-refl 179 | [idˢ] {t = tm-meta M ts} = ≈-meta (λ i → [idˢ]) 180 | [idˢ] {t = tm-oper f es} = ≈-oper (λ i → ≈-trans ([]ˢ-resp-≈ˢ (es i) ⇑ˢ-resp-idˢ) [idˢ]) 181 | 182 | -- the identity substitution preserves equality of terms 183 | [idˢ]-resp-≈ : ∀ {Θ} {Γ} {A} {t s : Term Θ Γ A} → t ≈ s → [ idˢ ]ˢ t ≈ s 184 | [idˢ]-resp-≈ t≈s = ≈-trans ([]ˢ-resp-≈ idˢ t≈s) [idˢ] 185 | 186 | 187 | -- if a substiution is equal to the identity then it acts trivially 188 | ≈ˢ-idˢ-[]ˢ : ∀ {Θ} {Γ} {A} {σ : Θ ⊕ Γ ⇒ˢ Γ} {t : Term Θ Γ A} → σ ≈ˢ idˢ → [ σ ]ˢ t ≈ t 189 | ≈ˢ-idˢ-[]ˢ {t = t} ξ = ≈-trans ([]ˢ-resp-≈ˢ t ξ) [idˢ] 190 | 191 | -- interaction of extension and right renaming action 192 | 193 | [⇑ˢ∘ᵛ] : ∀ {Θ} {A} {Γ Δ Ξ Ψ} {σ : Θ ⊕ Δ ⇒ˢ Ξ} {ρ : Γ ⇒ᵛ Δ} (t : Term Θ (Γ ,, Ψ) A) → 194 | [ ⇑ˢ (σ ˢ∘ᵛ ρ) ]ˢ t ≈ [ ⇑ˢ σ ]ˢ [ ⇑ᵛ ρ ]ᵛ t 195 | [⇑ˢ∘ᵛ] (tm-var (var-inl x)) = ≈-refl 196 | [⇑ˢ∘ᵛ] (tm-var (var-inr x)) = ≈-refl 197 | [⇑ˢ∘ᵛ] (tm-meta M ts) = ≈-meta (λ i → [⇑ˢ∘ᵛ] (ts i)) 198 | [⇑ˢ∘ᵛ] (tm-oper f es) = ≈-oper (λ i → ≈-trans ([]ˢ-resp-≈ˢ (es i) (⇑ˢ-resp-≈ˢ ⇑ˢ-resp-ˢ∘ᵛ)) ([⇑ˢ∘ᵛ] (es i))) 199 | 200 | -- interaction of extension and left renaming action 201 | 202 | ⇑ˢ-resp-ᵛ∘ˢ : ∀ {Θ} {Γ Δ Ξ Ψ} {σ : Θ ⊕ Γ ⇒ˢ Δ} {ρ : Δ ⇒ᵛ Ξ} → 203 | ⇑ˢ {Ξ = Ψ} (ρ ᵛ∘ˢ σ) ≈ˢ ⇑ᵛ ρ ᵛ∘ˢ ⇑ˢ σ 204 | ⇑ˢ-resp-ᵛ∘ˢ (var-inl x) = ≈-trans (≈-sym [∘ᵛ]) (≈-trans ([]ᵛ-resp-≡ᵛ (λ _ → refl)) [∘ᵛ]) 205 | ⇑ˢ-resp-ᵛ∘ˢ (var-inr y) = ≈-refl 206 | 207 | [⇑ᵛ∘ˢ] : ∀ {Θ} {A} {Γ Δ Ξ Ψ} {σ : Θ ⊕ Γ ⇒ˢ Δ} {ρ : Δ ⇒ᵛ Ξ} (t : Term Θ (Γ ,, Ψ) A) → 208 | [ ⇑ˢ (ρ ᵛ∘ˢ σ) ]ˢ t ≈ [ ⇑ᵛ ρ ]ᵛ ([ ⇑ˢ σ ]ˢ t) 209 | [⇑ᵛ∘ˢ] (tm-var x) = ⇑ˢ-resp-ᵛ∘ˢ x 210 | [⇑ᵛ∘ˢ] (tm-meta M ts) = ≈-meta (λ i → [⇑ᵛ∘ˢ] (ts i)) 211 | [⇑ᵛ∘ˢ] (tm-oper f es) = ≈-oper (λ i → ≈-trans ([]ˢ-resp-≈ˢ (es i) (⇑ˢ-resp-≈ˢ ⇑ˢ-resp-ᵛ∘ˢ)) ([⇑ᵛ∘ˢ] (es i))) 212 | 213 | -- action of the composition of a renaming and a substitution 214 | 215 | [ᵛ∘ˢ] : ∀ {Θ} {A} {Γ Δ Ξ} {ρ : Δ ⇒ᵛ Ξ} {σ : Θ ⊕ Γ ⇒ˢ Δ} (t : Term Θ Γ A) → 216 | [ ρ ᵛ∘ˢ σ ]ˢ t ≈ [ ρ ]ᵛ [ σ ]ˢ t 217 | [ᵛ∘ˢ] (tm-var x) = ≈-refl 218 | [ᵛ∘ˢ] (tm-meta M ts) = ≈-meta (λ i → [ᵛ∘ˢ] (ts i)) 219 | [ᵛ∘ˢ] (tm-oper f es) = ≈-oper (λ i → [⇑ᵛ∘ˢ] (es i)) 220 | 221 | -- action of the composition of a substitution and a renaming 222 | 223 | [ˢ∘ᵛ] : ∀ {Θ} {A} {Γ Δ Ξ} {σ : Θ ⊕ Δ ⇒ˢ Ξ} {ρ : Γ ⇒ᵛ Δ} (t : Term Θ Γ A) → 224 | [ σ ˢ∘ᵛ ρ ]ˢ t ≈ [ σ ]ˢ [ ρ ]ᵛ t 225 | [ˢ∘ᵛ] (tm-var x) = ≈-refl 226 | [ˢ∘ᵛ] (tm-meta M ts) = ≈-meta (λ i → [ˢ∘ᵛ] (ts i)) 227 | [ˢ∘ᵛ] (tm-oper f es) = ≈-oper (λ i → [⇑ˢ∘ᵛ] (es i)) 228 | 229 | 230 | -- the action of renamings and metarenamings don't interfere with each other 231 | 232 | []ᵛ-[]ᵐ : ∀ {Θ Ω Γ Δ A} (ρ : Γ ⇒ᵛ Δ) (μ : Θ ⇒ᵐ Ω) (t : Term Θ Γ A) → [ ρ ]ᵛ ([ μ ]ᵐ t) ≈ [ μ ]ᵐ ([ ρ ]ᵛ t) 233 | []ᵛ-[]ᵐ ρ μ (tm-var x) = ≈-refl 234 | []ᵛ-[]ᵐ ρ μ (tm-meta M ts) = ≈-meta (λ i → []ᵛ-[]ᵐ ρ μ (ts i)) 235 | []ᵛ-[]ᵐ ρ μ (tm-oper f es) = ≈-oper (λ i → []ᵛ-[]ᵐ [ (λ x → var-inl (ρ x)) , (λ x → var-inr x) ]ᵛ μ (es i)) 236 | 237 | -- substitution extension interacts nicely with the composition of a renaming and a substitution 238 | 239 | ⇑ˢ-ᵐ∘ˢ : ∀ {Θ Ω Γ Δ Ξ} (μ : Θ ⇒ᵐ Ω) (σ : Θ ⊕ Γ ⇒ˢ Δ) → ⇑ˢ {Ξ = Ξ} (μ ᵐ∘ˢ σ) ≈ˢ μ ᵐ∘ˢ ⇑ˢ σ 240 | ⇑ˢ-ᵐ∘ˢ μ σ (var-inl x) = []ᵛ-[]ᵐ var-inl μ (σ x) 241 | ⇑ˢ-ᵐ∘ˢ μ σ (var-inr x) = ≈-refl 242 | 243 | -- action of the composition of a metarenaming and a substitution 244 | 245 | [ᵐ∘ˢ] : ∀ {Θ Ω} {A} {Γ Δ} {μ : Θ ⇒ᵐ Ω} {σ : Θ ⊕ Γ ⇒ˢ Δ} (t : Term Θ Γ A) → 246 | [ μ ᵐ∘ˢ σ ]ˢ ( [ μ ]ᵐ t) ≈ [ μ ]ᵐ [ σ ]ˢ t 247 | [ᵐ∘ˢ] (tm-var x) = ≈-refl 248 | [ᵐ∘ˢ] {μ = μ} {σ = σ} (tm-meta M ts) = ≈-meta λ i → ([ᵐ∘ˢ] (ts i)) 249 | [ᵐ∘ˢ] {μ = μ} {σ = σ} (tm-oper f es) = ≈-oper λ i → 250 | ≈-trans 251 | ([]ˢ-resp-≈ˢ ([ μ ]ᵐ (es i)) (⇑ˢ-ᵐ∘ˢ μ σ)) 252 | ([ᵐ∘ˢ] (es i)) 253 | 254 | 255 | -- composition commutes with extension 256 | 257 | ⇑ˢ-resp-∘ˢ : ∀ {Θ} {Γ Δ Ξ Ψ} {σ : Θ ⊕ Γ ⇒ˢ Δ} {τ : Θ ⊕ Δ ⇒ˢ Ξ} → 258 | ⇑ˢ {Ξ = Ψ} (τ ∘ˢ σ) ≈ˢ ⇑ˢ τ ∘ˢ ⇑ˢ σ 259 | ⇑ˢ-resp-∘ˢ {σ = σ} {τ = τ} (var-inl x) = ≈-trans (≈-sym ([ᵛ∘ˢ] (σ x))) ([ˢ∘ᵛ] (σ x)) 260 | ⇑ˢ-resp-∘ˢ (var-inr y) = ≈-refl 261 | 262 | -- substitition action is functorial 263 | 264 | [∘ˢ] : ∀ {Θ} {Γ Δ Ξ} {A} {σ : Θ ⊕ Γ ⇒ˢ Δ} {τ : Θ ⊕ Δ ⇒ˢ Ξ} (t : Term Θ Γ A) → 265 | [ τ ∘ˢ σ ]ˢ t ≈ [ τ ]ˢ ([ σ ]ˢ t) 266 | [∘ˢ] (tm-var x) = ≈-refl 267 | [∘ˢ] (tm-meta M ts) = ≈-meta (λ i → [∘ˢ] (ts i)) 268 | [∘ˢ] (tm-oper f es) = ≈-oper (λ i → ≈-trans ([]ˢ-resp-≈ˢ (es i) ⇑ˢ-resp-∘ˢ) ([∘ˢ] (es i))) 269 | 270 | -- Terms form a relative monad 271 | 272 | module _ where 273 | open Categories.Category 274 | open Categories.Functor using (Functor) 275 | open Categories.Category.Instance.Setoids 276 | open Categories.Monad.Relative 277 | open Function.Equality using () renaming (setoid to Π-setoid) 278 | open Categories.Category.Equivalence using (StrongEquivalence) 279 | open import SecondOrder.IndexedCategory 280 | open import SecondOrder.RelativeKleisli 281 | 282 | -- The embedding of contexts into setoids indexed by sorts 283 | 284 | slots : Functor VContexts (IndexedCategory sort (Setoids ℓ ℓ)) 285 | slots = record 286 | { F₀ = λ Γ A → setoid (A ∈ Γ) 287 | ; F₁ = λ ρ A → record { _⟨$⟩_ = ρ ; cong = cong ρ } 288 | ; identity = λ A ξ → ξ 289 | ; homomorphism = λ {_} {_} {_} {ρ} {σ} A {_} {_} ξ → cong σ (cong ρ ξ) 290 | ; F-resp-≈ = λ ξ A ζ → trans (ξ _) (cong _ ζ) 291 | } 292 | 293 | 294 | module _ {Θ : MContext} where 295 | open Categories.Category 296 | open Categories.Functor using (Functor) 297 | open Categories.Category.Instance.Setoids 298 | open Categories.Monad.Relative 299 | open Function.Equality using () renaming (setoid to Π-setoid) 300 | open Categories.Category.Equivalence using (StrongEquivalence) 301 | open import SecondOrder.IndexedCategory 302 | open import SecondOrder.RelativeKleisli 303 | 304 | 305 | -- The relative monad of terms over contexts 306 | 307 | Term-Monad : Monad slots 308 | Term-Monad = 309 | let open Function.Equality using (_⟨$⟩_) renaming (cong to func-cong) in 310 | record 311 | { F₀ = Term-setoid Θ 312 | ; unit = λ A → record { _⟨$⟩_ = idˢ ; cong = λ ξ → ≈-≡ (cong idˢ ξ) } 313 | ; extend = λ σ A → record { _⟨$⟩_ = [ (σ _ ⟨$⟩_) ]ˢ_ ; cong = []ˢ-resp-≈ (σ _ ⟨$⟩_)} 314 | ; identityʳ = λ {_} {_} {σ} A {_} {_} ξ → func-cong (σ A) ξ 315 | ; identityˡ = λ A → ≈-trans [idˢ] 316 | ; assoc = λ {_} {_} {_} {σ} {ρ} A {_} {t} ξ → ≈-trans ([]ˢ-resp-≈ _ ξ) ([∘ˢ] t) 317 | ; extend-≈ = λ {Γ} {Δ} {σ} {ρ} ζ B {s} {t} ξ → []ˢ-resp-≈ˢ-≈ (λ x → ζ _ refl) ξ 318 | } 319 | 320 | 321 | -- the category of contexts and substitutions 322 | 323 | -- we show below that the category of contexts and substitiions is equivalent 324 | -- to the Kleisli category for the Term relative monad. However, we define 325 | -- the category of contexts and substitutions directly, as that it is easier 326 | -- to work with it 327 | 328 | Terms : Category ℓ ℓ ℓ 329 | Terms = 330 | record 331 | { Obj = VContext 332 | ; _⇒_ = Θ ⊕_⇒ˢ_ 333 | ; _≈_ = _≈ˢ_ 334 | ; id = idˢ 335 | ; _∘_ = _∘ˢ_ 336 | ; assoc = λ {Γ} {Δ} {Ξ} {Ψ} {σ} {τ} {ψ} {A} x → [∘ˢ] (σ x) 337 | ; sym-assoc = λ {Γ} {Δ} {Ξ} {Ψ} {σ} {τ} {ψ} {A} x → ≈-sym ([∘ˢ] (σ x)) 338 | ; identityˡ = λ x → [idˢ] 339 | ; identityʳ = λ x → ≈-refl 340 | ; identity² = λ x → ≈-refl 341 | ; equiv = record { refl = λ {σ} {A} → ≈ˢ-refl {σ = σ} ; sym = ≈ˢ-sym ; trans = ≈ˢ-trans } 342 | ; ∘-resp-≈ = λ f≈ˢg g≈ˢi x → []ˢ-resp-≈ˢ-≈ f≈ˢg (g≈ˢi x) 343 | } 344 | 345 | Terms-is-Kleisli : StrongEquivalence Terms (Kleisli Term-Monad) 346 | Terms-is-Kleisli = 347 | record 348 | { F = record 349 | { F₀ = λ Γ → Γ 350 | ; F₁ = λ σ A → record { _⟨$⟩_ = λ x → σ x ; cong = λ i≡j → ≈-≡ (cong σ i≡j) } 351 | ; identity = λ A eq → ≈-≡ (cong idˢ eq) 352 | ; homomorphism = λ {Γ} {Δ} {Ξ} {σ} {τ} A eq → ≈-≡ (cong (λ x → [ τ ]ˢ σ x) eq) 353 | ; F-resp-≈ = λ {Γ} {Δ} {σ} {τ} hom_eq A eq 354 | → ≈-trans (congˢ hom_eq) (≈-≡ (cong τ eq)) 355 | } 356 | ; G = 357 | let open Function.Equality using (_⟨$⟩_) renaming (cong to func-cong) in 358 | record 359 | { F₀ = λ Γ → Γ 360 | ; F₁ = λ {Γ} {Δ} σ {A} → λ x → σ A ⟨$⟩ x 361 | ; identity = λ x → ≈-refl 362 | ; homomorphism = λ x → ≈-refl 363 | ; F-resp-≈ = λ {Γ} {Δ} {σ} {τ} σ≈τ {A} x → σ≈τ A refl 364 | } 365 | ; weak-inverse = 366 | let open Function.Equality using (_⟨$⟩_) renaming (cong to func-cong) in 367 | record 368 | { F∘G≈id = 369 | record 370 | { F⇒G = 371 | record 372 | { η = λ Γ A → record { _⟨$⟩_ = idˢ 373 | ; cong = λ i≡j → ≈-≡ (cong idˢ i≡j) 374 | } 375 | ; commute = λ σ A x≡y → [idˢ]-resp-≈ (≈-≡ (cong (λ x → σ A ⟨$⟩ x) x≡y)) 376 | ; sym-commute = λ σ A x≡y 377 | → ≈-sym ([idˢ]-resp-≈ (≈-≡ (cong (λ x → σ A ⟨$⟩ x ) (sym x≡y)))) 378 | } 379 | ; F⇐G = 380 | record 381 | { η = λ Γ A → record { _⟨$⟩_ = idˢ 382 | ; cong = λ i≡j → ≈-≡ (cong idˢ i≡j) 383 | } 384 | ; commute = λ σ A x≡y → [idˢ]-resp-≈ (≈-≡ (cong (λ x → σ A ⟨$⟩ x) x≡y)) 385 | ; sym-commute = λ σ A x≡y 386 | → ≈-sym ([idˢ]-resp-≈ (≈-≡ (cong (λ x → σ A ⟨$⟩ x ) (sym x≡y)))) 387 | } 388 | ; iso = λ Γ → record { isoˡ = λ A x≡y → ≈-≡ (cong tm-var x≡y) 389 | ; isoʳ = λ A x≡y → ≈-≡ (cong tm-var x≡y) 390 | } 391 | } 392 | ; G∘F≈id = 393 | record 394 | { F⇒G = 395 | record 396 | { η = λ Γ x → tm-var x 397 | ; commute = λ σ x → [idˢ] 398 | ; sym-commute = λ σ x → ≈-sym [idˢ] 399 | } 400 | ; F⇐G = 401 | record 402 | { η = λ Γ x → tm-var x 403 | ; commute = λ σ x → [idˢ] 404 | ; sym-commute = λ σ x → ≈-sym [idˢ] 405 | } 406 | ; iso = λ Γ → record { isoˡ = λ x → ≈-refl 407 | ; isoʳ = λ x → ≈-refl 408 | } 409 | } 410 | } 411 | } 412 | 413 | -- the binary coproduct structure on Terms 414 | 415 | infixl 7 [_,_]ˢ 416 | 417 | [_,_]ˢ : ∀ {Γ Δ Ξ} (σ : Θ ⊕ Γ ⇒ˢ Ξ) (τ : Θ ⊕ Δ ⇒ˢ Ξ) → Θ ⊕ (Γ ,, Δ) ⇒ˢ Ξ 418 | [ σ , τ ]ˢ (var-inl x) = σ x 419 | [ σ , τ ]ˢ (var-inr y) = τ y 420 | 421 | inlˢ : ∀ {Γ Δ} → Θ ⊕ Γ ⇒ˢ Γ ,, Δ 422 | inlˢ x = tm-var (var-inl x) 423 | 424 | inrˢ : ∀ {Γ Δ} → Θ ⊕ Δ ⇒ˢ Γ ,, Δ 425 | inrˢ y = tm-var (var-inr y) 426 | 427 | [,]ˢ-resp-≈ˢ : ∀ {Γ Δ Ξ} {σ₁ σ₂ : Θ ⊕ Γ ⇒ˢ Ξ} {τ₁ τ₂ : Θ ⊕ Δ ⇒ˢ Ξ} → 428 | σ₁ ≈ˢ σ₂ → τ₁ ≈ˢ τ₂ → [ σ₁ , τ₁ ]ˢ ≈ˢ [ σ₂ , τ₂ ]ˢ 429 | [,]ˢ-resp-≈ˢ ζ ξ (var-inl x) = ζ x 430 | [,]ˢ-resp-≈ˢ ζ ξ (var-inr y) = ξ y 431 | 432 | uniqueˢ : ∀ {Γ Δ Ξ} {τ : Θ ⊕ Γ ,, Δ ⇒ˢ Ξ} {ρ : Θ ⊕ Γ ⇒ˢ Ξ} {σ : Θ ⊕ Δ ⇒ˢ Ξ} 433 | → τ ∘ˢ inlˢ ≈ˢ ρ 434 | → τ ∘ˢ inrˢ ≈ˢ σ 435 | → [ ρ , σ ]ˢ ≈ˢ τ 436 | uniqueˢ ξ ζ (var-inl x) = ≈-sym (ξ x) 437 | uniqueˢ ξ ζ (var-inr y) = ≈-sym (ζ y) 438 | 439 | Terms-Coproduct : Categories.Category.Cocartesian.BinaryCoproducts Terms 440 | Terms-Coproduct = 441 | let open Function.Equality using (_⟨$⟩_) renaming (cong to func-cong) in 442 | record { 443 | coproduct = 444 | λ {Γ Δ} → 445 | record 446 | { A+B = Γ ,, Δ 447 | ; i₁ = inlˢ 448 | ; i₂ = inrˢ 449 | ; [_,_] = [_,_]ˢ 450 | ; inject₁ = λ x → ≈-≡ refl 451 | ; inject₂ = λ x → ≈-≡ refl 452 | ; unique = λ {Ξ} {h} p₁ p₂ → uniqueˢ {τ = h} p₁ p₂ 453 | } 454 | } 455 | 456 | open Categories.Category.Cocartesian.BinaryCoproducts Terms-Coproduct 457 | 458 | -- the sum of subsitutions 459 | 460 | infixl 6 _+ˢ_ 461 | 462 | _+ˢ_ : ∀ {Γ₁ Γ₂ Δ₁ Δ₂} (σ : Θ ⊕ Γ₁ ⇒ˢ Δ₁) (τ : Θ ⊕ Γ₂ ⇒ˢ Δ₂) → Θ ⊕ Γ₁ ,, Γ₂ ⇒ˢ Δ₁ ,, Δ₂ 463 | σ +ˢ τ = σ +₁ τ 464 | 465 | -- reassociations of context sums 466 | 467 | assoc-l : ∀ {Γ Δ Ξ} → Θ ⊕ (Γ ,, Δ) ,, Ξ ⇒ˢ Γ ,, (Δ ,, Ξ) 468 | assoc-l (var-inl (var-inl x)) = tm-var (var-inl x) 469 | assoc-l (var-inl (var-inr y)) = tm-var (var-inr (var-inl y)) 470 | assoc-l (var-inr z) = tm-var (var-inr (var-inr z)) 471 | 472 | assoc-r : ∀ {Γ Δ Ξ} → Θ ⊕ Γ ,, (Δ ,, Ξ) ⇒ˢ (Γ ,, Δ) ,, Ξ 473 | assoc-r (var-inl x) = tm-var (var-inl (var-inl x)) 474 | assoc-r (var-inr (var-inl y)) = tm-var (var-inl (var-inr y)) 475 | assoc-r (var-inr (var-inr z)) = tm-var (var-inr z) 476 | 477 | assoc-lr : ∀ {Γ Δ Ξ} → assoc-l {Γ = Γ} {Δ = Δ} {Ξ = Ξ} ∘ˢ assoc-r {Γ = Γ} {Δ = Δ} {Ξ = Ξ} ≈ˢ idˢ 478 | assoc-lr (var-inl x) = ≈-refl 479 | assoc-lr (var-inr (var-inl y)) = ≈-refl 480 | assoc-lr (var-inr (var-inr y)) = ≈-refl 481 | 482 | assoc-rl : ∀ {Γ Δ Ξ} → assoc-r {Γ = Γ} {Δ = Δ} {Ξ = Ξ} ∘ˢ assoc-l {Γ = Γ} {Δ = Δ} {Ξ = Ξ} ≈ˢ idˢ 483 | assoc-rl (var-inl (var-inl x)) = ≈-refl 484 | assoc-rl (var-inl (var-inr x)) = ≈-refl 485 | assoc-rl (var-inr z) = ≈-refl 486 | 487 | -- summing with the empty context is the unit 488 | 489 | sum-ctx-empty-r : ∀ {Γ} → Θ ⊕ Γ ,, ctx-empty ⇒ˢ Γ 490 | sum-ctx-empty-r (var-inl x) = tm-var x 491 | 492 | sum-ctx-empty-l : ∀ {Γ} → Θ ⊕ ctx-empty ,, Γ ⇒ˢ Γ 493 | sum-ctx-empty-l (var-inr x) = tm-var x 494 | -------------------------------------------------------------------------------- /src/SecondOrder/Term.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality 3 | open import Relation.Binary using (Setoid) 4 | 5 | import SecondOrder.Arity 6 | import SecondOrder.Signature 7 | import SecondOrder.Metavariable 8 | 9 | module SecondOrder.Term 10 | {ℓ} 11 | {𝔸 : SecondOrder.Arity.Arity} 12 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 13 | where 14 | 15 | open SecondOrder.Signature.Signature Σ 16 | open SecondOrder.Metavariable Σ 17 | 18 | -- The term judgement 19 | data Term (Θ : MContext) : ∀ (Γ : VContext) (A : sort) → Set ℓ 20 | 21 | Arg : ∀ (Θ : MContext) (Γ : VContext) (A : sort) (Δ : VContext) → Set ℓ 22 | Arg Θ Γ A Δ = Term Θ (Γ ,, Δ) A 23 | 24 | data Term Θ where 25 | tm-var : ∀ {Γ} {A} (x : A ∈ Γ) → Term Θ Γ A 26 | tm-meta : ∀ {Γᴹ A} {Γ} (M : [ Γᴹ , A ]∈ Θ) 27 | (ts : ∀ {B} (i : B ∈ Γᴹ) → Term Θ Γ B) 28 | → Term Θ Γ A 29 | tm-oper : ∀ {Γ} (f : oper) (es : ∀ (i : oper-arg f) → Arg Θ Γ (arg-sort f i) (arg-bind f i)) 30 | → Term Θ Γ (oper-sort f) 31 | 32 | -- Syntactic equality of terms 33 | 34 | infix 4 _≈_ 35 | 36 | data _≈_ {Θ : MContext} : ∀ {Γ : VContext} {A : sort} → Term Θ Γ A → Term Θ Γ A → Set ℓ where 37 | ≈-≡ : ∀ {Γ A} {t u : Term Θ Γ A} (ξ : t ≡ u) → t ≈ u 38 | ≈-meta : ∀ {Γ} {α A} {M : [ α , A ]∈ Θ} {ts us : ∀ {B} (i : B ∈ α) → Term Θ Γ B} 39 | (ξ : ∀ {B} i → ts {B} i ≈ us {B} i) → tm-meta M ts ≈ tm-meta M us 40 | ≈-oper : ∀ {Γ} {f : oper} {ds es : ∀ (i : oper-arg f) → Arg Θ Γ (arg-sort f i) (arg-bind f i)} 41 | (ξ : ∀ i → ds i ≈ es i) → tm-oper f ds ≈ tm-oper f es 42 | 43 | 44 | -- Syntactic equality of terms is an equivalence relation 45 | 46 | ≈-refl : ∀ {Θ Γ A} {t : Term Θ Γ A} → t ≈ t 47 | ≈-refl = ≈-≡ refl 48 | 49 | ≈-sym : ∀ {Θ Γ A} {t u : Term Θ Γ A} → t ≈ u → u ≈ t 50 | ≈-sym (≈-≡ refl) = ≈-≡ refl 51 | ≈-sym (≈-meta ξ) = ≈-meta λ i → ≈-sym (ξ i) 52 | ≈-sym (≈-oper ξ) = ≈-oper (λ i → ≈-sym (ξ i)) 53 | 54 | ≈-trans : ∀ {Θ Γ A} {t u v : Term Θ Γ A} → t ≈ u → u ≈ v → t ≈ v 55 | ≈-trans (≈-≡ refl) ξ = ξ 56 | ≈-trans (≈-meta ζ) (≈-≡ refl) = ≈-meta ζ 57 | ≈-trans (≈-meta ζ) (≈-meta ξ) = ≈-meta (λ i → ≈-trans (ζ i) (ξ i)) 58 | ≈-trans (≈-oper ζ) (≈-≡ refl) = ≈-oper ζ 59 | ≈-trans (≈-oper ζ) (≈-oper ξ) = ≈-oper (λ i → ≈-trans (ζ i) (ξ i)) 60 | 61 | Term-setoid : ∀ (Θ : MContext) (Γ : VContext) (A : sort) → Setoid ℓ ℓ 62 | Term-setoid Θ Γ A = 63 | record 64 | { Carrier = Term Θ Γ A 65 | ; _≈_ = _≈_ 66 | ; isEquivalence = record { refl = ≈-refl ; sym = ≈-sym ; trans = ≈-trans } } 67 | -------------------------------------------------------------------------------- /src/SecondOrder/Theory.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | 3 | import SecondOrder.Arity 4 | import SecondOrder.Signature 5 | import SecondOrder.Metavariable 6 | import SecondOrder.VRenaming 7 | import SecondOrder.Term 8 | 9 | module SecondOrder.Theory 10 | {ℓ} 11 | {𝔸 : SecondOrder.Arity.Arity} 12 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 13 | where 14 | 15 | open SecondOrder.Metavariable Σ public 16 | open SecondOrder.Term Σ public 17 | open SecondOrder.Signature.Signature Σ public 18 | open SecondOrder.VRenaming Σ 19 | 20 | record Axiom : Set ℓ where 21 | constructor make-ax 22 | field 23 | ax-mv-ctx : MContext -- metavariable context of an equation 24 | ax-sort : sort -- sort of an equation 25 | ax-lhs : Term ax-mv-ctx ctx-empty ax-sort -- left-hand side 26 | ax-rhs : Term ax-mv-ctx ctx-empty ax-sort -- right-hand side 27 | 28 | record Theory ℓa : Set (lsuc (ℓ ⊔ ℓa)) where 29 | field 30 | ax : Set ℓa -- the axioms 31 | ax-eq : ax → Axiom -- each axiom has a corresponding Axiom 32 | 33 | ax-mv-ctx : ax → MContext -- the meta-context of each axiom 34 | ax-mv-ctx ε = Axiom.ax-mv-ctx (ax-eq ε) 35 | 36 | ax-sort : ax → sort -- the sort of each axiom 37 | ax-sort ε = Axiom.ax-sort (ax-eq ε) 38 | 39 | -- the left- and right-hand side of each axiom s ≈ t, promoted to any context 40 | ax-lhs : ∀ (ε : ax) {Γ} → Term (ax-mv-ctx ε) Γ (ax-sort ε) 41 | ax-lhs ε = [ inʳ ]ʳ Axiom.ax-lhs (ax-eq ε) 42 | 43 | ax-rhs : ∀ (ε : ax) {Γ} → Term (ax-mv-ctx ε) Γ (ax-sort ε) 44 | ax-rhs ε = [ inʳ ]ʳ Axiom.ax-rhs (ax-eq ε) 45 | -------------------------------------------------------------------------------- /src/SecondOrder/VContext.agda: -------------------------------------------------------------------------------- 1 | module SecondOrder.VContext {ℓ} (sort : Set ℓ) where 2 | 3 | -- a context is a binary tree whose leaves are labeled with sorts 4 | data VContext : Set ℓ where 5 | ctx-empty : VContext 6 | ctx-slot : sort → VContext 7 | _,,_ : VContext → VContext → VContext 8 | 9 | infixl 5 _,,_ 10 | 11 | infix 4 _∈_ 12 | 13 | -- the variables of a given type in a context 14 | data _∈_ (A : sort) : VContext → Set ℓ where 15 | var-slot : A ∈ ctx-slot A 16 | var-inl : ∀ {Γ Δ} (x : A ∈ Γ) → A ∈ Γ ,, Δ 17 | var-inr : ∀ {Γ Δ} (y : A ∈ Δ) → A ∈ Γ ,, Δ 18 | -------------------------------------------------------------------------------- /src/SecondOrder/VRelMonMorphism.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; cong; cong₂; subst; setoid) 3 | open import Data.Product using (_×_; Σ; _,_; proj₁; proj₂; zip; map; <_,_>; swap) 4 | import Function.Equality 5 | open import Relation.Binary using (Setoid) 6 | import Relation.Binary.Reasoning.Setoid as SetoidR 7 | 8 | import Categories.Category 9 | import Categories.Functor 10 | import Categories.Category.Instance.Setoids 11 | import Categories.Monad.Relative 12 | import Categories.Category.Equivalence 13 | import Categories.Category.Cocartesian 14 | import Categories.Category.Construction.Functors 15 | import Categories.Category.Product 16 | import Categories.NaturalTransformation 17 | import Categories.NaturalTransformation.NaturalIsomorphism 18 | 19 | import SecondOrder.Arity 20 | import SecondOrder.Signature 21 | import SecondOrder.Metavariable 22 | import SecondOrder.VRenaming 23 | import SecondOrder.MRenaming 24 | import SecondOrder.Term 25 | import SecondOrder.Substitution 26 | import SecondOrder.RelativeMonadMorphism 27 | import SecondOrder.Instantiation 28 | import SecondOrder.IndexedCategory 29 | import SecondOrder.RelativeKleisli 30 | import SecondOrder.Mslot 31 | import SecondOrder.MRelativeMonad 32 | import SecondOrder.VRelativeMonad 33 | 34 | module SecondOrder.VRelMonMorphism 35 | {ℓ} 36 | {𝔸 : SecondOrder.Arity.Arity} 37 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 38 | where 39 | 40 | open SecondOrder.RelativeMonadMorphism 41 | open SecondOrder.Metavariable Σ 42 | open SecondOrder.VRelativeMonad Σ 43 | open SecondOrder.Instantiation Σ 44 | open SecondOrder.MRenaming Σ 45 | open SecondOrder.VRenaming Σ 46 | open SecondOrder.Term Σ 47 | open SecondOrder.Substitution Σ 48 | 49 | 50 | -- In this file, the goal is to show that given two variable relative monads 51 | -- on different metacontexts, metarenaming from one of the metacontexts to the other 52 | -- we can define a relative monad morphism between the two variable relative monads 53 | 54 | Fⱽ : ∀ (Θ Θ′ : MContext) (μ : Θ ⇒ᵐ Θ′) → RMonadMorph (VMonad {Θ}) (VMonad {Θ′}) 55 | Fⱽ Θ Θ′ μ = record 56 | { morph = λ A → 57 | record 58 | { _⟨$⟩_ = [ μ ]ᵐ_ 59 | ; cong = λ s≈t → []ᵐ-resp-≈ s≈t 60 | } 61 | ; law-unit = λ A x≡y → ≈-≡ (σ-resp-≡ {σ = tm-var} x≡y) 62 | ; law-extend = λ {Γ} {Δ} {k} A {s} {t} s≈t → 63 | ≈-trans 64 | (≈-sym ([ᵐ∘ˢ] s)) 65 | (≈-trans 66 | ([]ˢ-resp-≈ˢ ([ μ ]ᵐ s) λ x → ≈-refl ) 67 | ([]ˢ-resp-≈ (λ {B} x → [ μ ]ᵐ (k B Function.Equality.⟨$⟩ x)) ([]ᵐ-resp-≈ s≈t))) 68 | } 69 | -------------------------------------------------------------------------------- /src/SecondOrder/VRelativeMonad.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_ ;Level) 2 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; setoid; cong; trans) 3 | import Function.Equality 4 | open import Relation.Binary using (Setoid) 5 | 6 | import Categories.Category 7 | import Categories.Functor 8 | import Categories.Category.Instance.Setoids 9 | import Categories.Monad.Relative 10 | import Categories.Category.Equivalence 11 | import Categories.Category.Cocartesian 12 | import Categories.Category.Construction.Functors 13 | import Categories.NaturalTransformation.Equivalence 14 | import Relation.Binary.Core 15 | 16 | import SecondOrder.Arity 17 | import SecondOrder.Signature 18 | import SecondOrder.Metavariable 19 | import SecondOrder.VRenaming 20 | import SecondOrder.MRenaming 21 | import SecondOrder.Term 22 | import SecondOrder.IndexedCategory 23 | import SecondOrder.RelativeKleisli 24 | import SecondOrder.Substitution 25 | 26 | module SecondOrder.VRelativeMonad 27 | {ℓ} 28 | {𝔸 : SecondOrder.Arity.Arity} 29 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 30 | where 31 | 32 | open SecondOrder.Signature.Signature Σ 33 | open SecondOrder.Metavariable Σ 34 | open SecondOrder.Term Σ 35 | open SecondOrder.VRenaming Σ 36 | open SecondOrder.MRenaming Σ 37 | open SecondOrder.Substitution Σ 38 | 39 | 40 | -- TERMS FORM A RELATIVE MONAD 41 | -- (FOR A FUNCTOR WHOSE DOMAIN IS THE 42 | -- CATEGORY OF VARIABLE CONTEXTS AND RENAMINGS) 43 | 44 | module _ where 45 | open Categories.Category 46 | open Categories.Functor using (Functor) 47 | open Categories.Category.Instance.Setoids 48 | open import SecondOrder.IndexedCategory 49 | 50 | -- The embedding of contexts into setoids indexed by sorts 51 | -- (seen as a "variable" functor) 52 | 53 | Jⱽ : Functor VContexts (IndexedCategory sort (Setoids ℓ ℓ)) 54 | Jⱽ = slots 55 | 56 | -- The relative monad over Jⱽ 57 | 58 | module _ {Θ : MContext} where 59 | open Categories.Category 60 | open Categories.Functor using (Functor) 61 | open Categories.Category.Instance.Setoids 62 | open Categories.Category.Category (Setoids ℓ ℓ) 63 | open Categories.Monad.Relative 64 | open Function.Equality using () renaming (setoid to Π-setoid) 65 | open Categories.Category.Equivalence using (StrongEquivalence) 66 | open import SecondOrder.IndexedCategory 67 | open import SecondOrder.RelativeKleisli 68 | open Function.Equality using (_⟨$⟩_) renaming (cong to func-cong) 69 | open import Relation.Binary.Reasoning.MultiSetoid 70 | 71 | to-σ : ∀ {Γ Δ} → (IndexedCategory sort (Setoids ℓ ℓ) Category.⇒ Functor.₀ Jⱽ Γ) (λ A → Term-setoid Θ Δ A) → (Θ ⊕ Γ ⇒ˢ Δ) 72 | to-σ ℱ {A} x = ℱ A ⟨$⟩ x 73 | 74 | VMonad : Monad Jⱽ 75 | VMonad = record 76 | { F₀ = λ Γ A → Term-setoid Θ Γ A 77 | ; unit = λ A → 78 | record 79 | { _⟨$⟩_ = λ x → tm-var x 80 | ; cong = λ ξ → ≈-≡ (σ-resp-≡ {σ = tm-var} ξ) } 81 | ; extend = λ ℱ A → 82 | record 83 | { _⟨$⟩_ = [ to-σ ℱ ]ˢ_ 84 | ; cong = λ ξ → []ˢ-resp-≈ (λ {B} z → ℱ B ⟨$⟩ z) ξ } 85 | ; identityʳ = λ {Γ} {Δ} {ℱ} A {x} {y} ξ → func-cong (ℱ A) ξ 86 | ; identityˡ = λ A ξ → ≈-trans [idˢ] ξ 87 | ; assoc = λ {Γ} {Δ} {Ξ} {k} {l} A {x} {y} ξ → 88 | begin⟨ Term-setoid Θ _ _ ⟩ 89 | ( ([ (λ {B} r → [ (λ {A = B} z → l B ⟨$⟩ z) ]ˢ (k B ⟨$⟩ r)) ]ˢ x) ) ≈⟨ []ˢ-resp-≈ ((λ {B} r → [ (λ {A = B} z → l B ⟨$⟩ z) ]ˢ (k B ⟨$⟩ r))) ξ ⟩ 90 | ([ (λ {B} r → [ (λ {A = A₁} z → l A₁ ⟨$⟩ z) ]ˢ (k B ⟨$⟩ r)) ]ˢ y) ≈⟨ [∘ˢ] y ⟩ 91 | (([ (λ {B} z → l B ⟨$⟩ z) ]ˢ ([ (λ {B} z → k B ⟨$⟩ z) ]ˢ y))) ∎ 92 | ; extend-≈ = λ {Γ} {Δ} {k} {h} ξˢ A {x} {y} ξ → 93 | begin⟨ Term-setoid Θ _ _ ⟩ 94 | ([ (λ {B} z → k B ⟨$⟩ z) ]ˢ x) ≈⟨ []ˢ-resp-≈ ((λ {B} z → k B ⟨$⟩ z)) ξ ⟩ 95 | ([ (λ {B} z → k B ⟨$⟩ z) ]ˢ y) ≈⟨ []ˢ-resp-≈ˢ y (λ {B} z → ξˢ B refl) ⟩ 96 | ([ (λ {B} z → h B ⟨$⟩ z) ]ˢ y) ∎ 97 | } 98 | -------------------------------------------------------------------------------- /src/SecondOrder/VRenaming.agda: -------------------------------------------------------------------------------- 1 | open import Level 2 | open import Relation.Binary.PropositionalEquality 3 | open import Relation.Binary using (Setoid) 4 | import Function.Equality 5 | import Relation.Binary.Reasoning.Setoid as SetoidR 6 | 7 | import Categories.Category 8 | import Categories.Functor 9 | import Categories.Category.Instance.Setoids 10 | 11 | import Categories.Category.Cocartesian 12 | 13 | import SecondOrder.Arity 14 | import SecondOrder.Signature 15 | import SecondOrder.Metavariable 16 | import SecondOrder.Term 17 | 18 | module SecondOrder.VRenaming 19 | {ℓ} 20 | {𝔸 : SecondOrder.Arity.Arity} 21 | (Σ : SecondOrder.Signature.Signature ℓ 𝔸) 22 | where 23 | 24 | open SecondOrder.Signature.Signature Σ 25 | open SecondOrder.Metavariable Σ 26 | open SecondOrder.Term Σ 27 | 28 | -- a renaming maps variables between contexts in a type-preserving way 29 | _⇒ᵛ_ : ∀ (Γ Δ : VContext) → Set ℓ 30 | Γ ⇒ᵛ Δ = ∀ {A} → A ∈ Γ → A ∈ Δ 31 | 32 | infix 4 _⇒ᵛ_ 33 | 34 | -- equality of renamings 35 | 36 | _≡ᵛ_ : ∀ {Γ Δ} (σ τ : Γ ⇒ᵛ Δ) → Set ℓ 37 | _≡ᵛ_ {Γ} σ τ = ∀ {A} (x : A ∈ Γ) → σ x ≡ τ x 38 | 39 | infixl 3 _≡ᵛ_ 40 | 41 | ≡ᵛ-refl : ∀ {Γ Δ} {ρ : Γ ⇒ᵛ Δ} → ρ ≡ᵛ ρ 42 | ≡ᵛ-refl = λ x → refl 43 | 44 | ≡ᵛ-sym : ∀ {Γ Δ} {ρ ν : Γ ⇒ᵛ Δ} 45 | → ρ ≡ᵛ ν 46 | → ν ≡ᵛ ρ 47 | ≡ᵛ-sym eq x = sym (eq x) 48 | 49 | ≡ᵛ-trans : ∀ {Γ Δ} {ρ ν γ : Γ ⇒ᵛ Δ} 50 | → ρ ≡ᵛ ν 51 | → ν ≡ᵛ γ 52 | → ρ ≡ᵛ γ 53 | ≡ᵛ-trans eq1 eq2 x = trans (eq1 x) (eq2 x) 54 | 55 | -- renamings form a setoid 56 | 57 | renaming-setoid : ∀ (Γ Δ : VContext) → Setoid ℓ ℓ 58 | renaming-setoid Γ Δ = 59 | record 60 | { Carrier = Γ ⇒ᵛ Δ 61 | ; _≈_ = λ ρ ν → ρ ≡ᵛ ν 62 | ; isEquivalence = 63 | record 64 | { refl = λ {ρ} x → ≡ᵛ-refl {Γ} {Δ} {ρ} x 65 | ; sym = ≡ᵛ-sym 66 | ; trans = ≡ᵛ-trans 67 | } 68 | } 69 | 70 | -- renaming preserves equality of variables 71 | ρ-resp-≡ : ∀ {Γ Δ A} {x y : A ∈ Γ} {ρ : Γ ⇒ᵛ Δ} → x ≡ y → ρ x ≡ ρ y 72 | ρ-resp-≡ refl = refl 73 | 74 | -- the identity renaming 75 | 76 | idᵛ : ∀ {Γ : VContext} → Γ ⇒ᵛ Γ 77 | idᵛ x = x 78 | 79 | -- composition of renamings 80 | _∘ᵛ_ : ∀ {Γ Δ Ξ} → Δ ⇒ᵛ Ξ → Γ ⇒ᵛ Δ → Γ ⇒ᵛ Ξ 81 | (σ ∘ᵛ ρ) x = σ (ρ x) 82 | 83 | infix 7 _∘ᵛ_ 84 | 85 | -- composition respects equality 86 | ∘ᵛ-resp-≡ᵛ : ∀ {Γ Δ Ξ} {τ₁ τ₂ : Δ ⇒ᵛ Ξ} {σ₁ σ₂ : Γ ⇒ᵛ Δ} → 87 | τ₁ ≡ᵛ τ₂ → σ₁ ≡ᵛ σ₂ → τ₁ ∘ᵛ σ₁ ≡ᵛ τ₂ ∘ᵛ σ₂ 88 | ∘ᵛ-resp-≡ᵛ {τ₁ = τ₁} {σ₂ = σ₂} ζ ξ x = trans (cong τ₁ (ξ x)) (ζ (σ₂ x)) 89 | 90 | -- the identity is the unit 91 | 92 | identity-leftᵛ : ∀ {Γ Δ} {ρ : Γ ⇒ᵛ Δ} → idᵛ ∘ᵛ ρ ≡ᵛ ρ 93 | identity-leftᵛ ρ = refl 94 | 95 | identity-rightᵛ : ∀ {Γ Δ} {ρ : Γ ⇒ᵛ Δ} → ρ ∘ᵛ idᵛ ≡ᵛ ρ 96 | identity-rightᵛ ρ = refl 97 | 98 | -- composition is associative 99 | 100 | assocᵛ : ∀ {Γ Δ Ξ Ψ} {τ : Γ ⇒ᵛ Δ} {ρ : Δ ⇒ᵛ Ξ} {σ : Ξ ⇒ᵛ Ψ} → 101 | (σ ∘ᵛ ρ) ∘ᵛ τ ≡ᵛ σ ∘ᵛ (ρ ∘ᵛ τ) 102 | assocᵛ x = refl 103 | 104 | sym-assocᵛ : ∀ {Γ Δ Ξ Ψ} {τ : Γ ⇒ᵛ Δ} {ρ : Δ ⇒ᵛ Ξ} {σ : Ξ ⇒ᵛ Ψ} → 105 | σ ∘ᵛ (ρ ∘ᵛ τ) ≡ᵛ (σ ∘ᵛ ρ) ∘ᵛ τ 106 | sym-assocᵛ x = refl 107 | 108 | -- contexts and renamings form a category 109 | module _ where 110 | open Categories.Category 111 | 112 | VContexts : Category ℓ ℓ ℓ 113 | VContexts = 114 | record 115 | { Obj = VContext 116 | ; _⇒_ = _⇒ᵛ_ 117 | ; _≈_ = _≡ᵛ_ 118 | ; id = idᵛ 119 | ; _∘_ = _∘ᵛ_ 120 | ; assoc = λ {_} {_} {_} {_} {f} {g} {h} {_} → assocᵛ {τ = f} {ρ = g} {σ = h} 121 | ; sym-assoc = λ {_} {_} {_} {_} {f} {g} {h} {_} → sym-assocᵛ {τ = f} {ρ = g} {σ = h} 122 | ; identityˡ = λ x → refl 123 | ; identityʳ = λ x → refl 124 | ; identity² = λ x → refl 125 | ; equiv = record { refl = λ {ρ} {_} → ≡ᵛ-refl {ρ = ρ} ; sym = ≡ᵛ-sym ; trans = ≡ᵛ-trans } 126 | ; ∘-resp-≈ = ∘ᵛ-resp-≡ᵛ 127 | } 128 | 129 | 130 | -- the coproduct structure of the category 131 | module _ where 132 | 133 | infixl 7 [_,_]ᵛ 134 | 135 | [_,_]ᵛ : ∀ {Γ Δ Ξ} → Γ ⇒ᵛ Ξ → Δ ⇒ᵛ Ξ → Γ ,, Δ ⇒ᵛ Ξ 136 | [ σ , τ ]ᵛ (var-inl x) = σ x 137 | [ σ , τ ]ᵛ (var-inr y) = τ y 138 | 139 | [,]ᵛ-resp-≡ᵛ : ∀ {Γ Δ Ξ} {ρ₁ ρ₂ : Γ ⇒ᵛ Ξ} {τ₁ τ₂ : Δ ⇒ᵛ Ξ} → ρ₁ ≡ᵛ ρ₂ → τ₁ ≡ᵛ τ₂ → [ ρ₁ , τ₁ ]ᵛ ≡ᵛ [ ρ₂ , τ₂ ]ᵛ 140 | [,]ᵛ-resp-≡ᵛ pρ pτ (var-inl x) = pρ x 141 | [,]ᵛ-resp-≡ᵛ pρ pτ (var-inr x) = pτ x 142 | 143 | inlᵛ : ∀ {Γ Δ} → Γ ⇒ᵛ Γ ,, Δ 144 | inlᵛ = var-inl 145 | 146 | inrᵛ : ∀ {Γ Δ} → Δ ⇒ᵛ Γ ,, Δ 147 | inrᵛ = var-inr 148 | 149 | uniqueᵛ : ∀ {Γ Δ Ξ} {τ : Γ ,, Δ ⇒ᵛ Ξ} {ρ : Γ ⇒ᵛ Ξ} {σ : Δ ⇒ᵛ Ξ} 150 | → τ ∘ᵛ inlᵛ ≡ᵛ ρ 151 | → τ ∘ᵛ inrᵛ ≡ᵛ σ 152 | → [ ρ , σ ]ᵛ ≡ᵛ τ 153 | uniqueᵛ ξ ζ (var-inl x) = sym (ξ x) 154 | uniqueᵛ ξ ζ (var-inr y) = sym (ζ y) 155 | 156 | uniqueᵛ² : ∀ {Γ Δ Ξ} {τ σ : Γ ,, Δ ⇒ᵛ Ξ} 157 | → τ ∘ᵛ inlᵛ ≡ᵛ σ ∘ᵛ inlᵛ 158 | → τ ∘ᵛ inrᵛ ≡ᵛ σ ∘ᵛ inrᵛ 159 | → τ ≡ᵛ σ 160 | uniqueᵛ² ξ ζ (var-inl x) = ξ x 161 | uniqueᵛ² ξ ζ (var-inr y) = ζ y 162 | 163 | Context-+ : Categories.Category.Cocartesian.BinaryCoproducts VContexts 164 | Context-+ = 165 | record { 166 | coproduct = 167 | λ {Γ Δ} → 168 | record 169 | { A+B = Γ ,, Δ 170 | ; i₁ = inlᵛ 171 | ; i₂ = inrᵛ 172 | ; [_,_] = [_,_]ᵛ 173 | ; inject₁ = λ x → refl 174 | ; inject₂ = λ x → refl 175 | ; unique = uniqueᵛ 176 | } 177 | } 178 | 179 | open Categories.Category.Cocartesian.BinaryCoproducts Context-+ 180 | 181 | -- the renaming from the empty context 182 | 183 | inᵛ : ∀ {Γ} → ctx-empty ⇒ᵛ Γ 184 | inᵛ () 185 | 186 | -- extension of a renaming is summing with identity 187 | ⇑ᵛ : ∀ {Γ Δ Ξ} → Γ ⇒ᵛ Δ → Γ ,, Ξ ⇒ᵛ Δ ,, Ξ 188 | ⇑ᵛ ρ = ρ +₁ idᵛ 189 | 190 | -- a renaming can also be extended on the right 191 | ʳ⇑ᵛ : ∀ {Γ Δ} → Γ ⇒ᵛ Δ → ∀ {Ξ} → Ξ ,, Γ ⇒ᵛ Ξ ,, Δ 192 | ʳ⇑ᵛ ρ = idᵛ +₁ ρ 193 | 194 | -- right extension of renamings commutes with right injection 195 | ʳ⇑ᵛ-comm-inrᵛ : ∀ {Γ Δ Ξ} (ρ : Γ ⇒ᵛ Δ) 196 | → (ʳ⇑ᵛ ρ {Ξ = Ξ}) ∘ᵛ (inrᵛ {Δ = Γ}) ≡ᵛ inrᵛ ∘ᵛ ρ 197 | ʳ⇑ᵛ-comm-inrᵛ ρ var-slot = refl 198 | ʳ⇑ᵛ-comm-inrᵛ ρ (var-inl x) = refl 199 | ʳ⇑ᵛ-comm-inrᵛ ρ (var-inr x) = refl 200 | 201 | -- left extension of renamings commutes with left injection 202 | ⇑ᵛ-comm-inlᵛ : ∀ {Γ Δ Ξ} (ρ : Γ ⇒ᵛ Δ) → (⇑ᵛ {Ξ = Ξ} ρ) ∘ᵛ (inlᵛ {Δ = Ξ}) ≡ᵛ inlᵛ ∘ᵛ ρ 203 | ⇑ᵛ-comm-inlᵛ ρ var-slot = refl 204 | ⇑ᵛ-comm-inlᵛ ρ (var-inl x) = refl 205 | ⇑ᵛ-comm-inlᵛ ρ (var-inr x) = refl 206 | 207 | -- the action of a renaming on terms 208 | module _ {Θ : MContext} where 209 | 210 | infix 6 [_]ᵛ_ 211 | 212 | [_]ᵛ_ : ∀ {Γ Δ A} → Γ ⇒ᵛ Δ → Term Θ Γ A → Term Θ Δ A 213 | [ ρ ]ᵛ (tm-var x) = tm-var (ρ x) 214 | [ ρ ]ᵛ (tm-meta M ts) = tm-meta M (λ i → [ ρ ]ᵛ (ts i)) 215 | [ ρ ]ᵛ (tm-oper f es) = tm-oper f (λ i → [ ⇑ᵛ ρ ]ᵛ (es i)) 216 | 217 | -- The sum of identities is an identity 218 | idᵛ+idᵛ : ∀ {Γ Δ} → idᵛ {Γ = Γ} +₁ idᵛ {Γ = Δ} ≡ᵛ idᵛ {Γ = Γ ,, Δ} 219 | idᵛ+idᵛ (var-inl x) = refl 220 | idᵛ+idᵛ (var-inr y) = refl 221 | 222 | -- The action of a renaming respects equality of terms 223 | []ᵛ-resp-≈ : ∀ {Θ Γ Δ A} {s t : Term Θ Γ A} {ρ : Γ ⇒ᵛ Δ} → s ≈ t → [ ρ ]ᵛ s ≈ [ ρ ]ᵛ t 224 | []ᵛ-resp-≈ (≈-≡ refl) = ≈-≡ refl 225 | []ᵛ-resp-≈ (≈-meta ξ) = ≈-meta (λ i → []ᵛ-resp-≈ (ξ i)) 226 | []ᵛ-resp-≈ (≈-oper ξ) = ≈-oper (λ i → []ᵛ-resp-≈ (ξ i)) 227 | 228 | -- The action of a renaming respects equality of renamings 229 | []ᵛ-resp-≡ᵛ : ∀ {Θ Γ Δ A} {ρ τ : Γ ⇒ᵛ Δ} {t : Term Θ Γ A} → ρ ≡ᵛ τ → [ ρ ]ᵛ t ≈ [ τ ]ᵛ t 230 | []ᵛ-resp-≡ᵛ {t = tm-var x} ξ = ≈-≡ (cong tm-var (ξ x)) 231 | []ᵛ-resp-≡ᵛ {t = tm-meta M ts} ξ = ≈-meta (λ i → []ᵛ-resp-≡ᵛ ξ) 232 | []ᵛ-resp-≡ᵛ {t = tm-oper f es} ξ = ≈-oper (λ i → []ᵛ-resp-≡ᵛ (+₁-cong₂ ξ ≡ᵛ-refl)) 233 | 234 | -- The action of the identity is trival 235 | [idᵛ] : ∀ {Θ Γ A} {t : Term Θ Γ A} → [ idᵛ ]ᵛ t ≈ t 236 | [idᵛ] {t = tm-var x} = ≈-refl 237 | [idᵛ] {t = tm-meta M ts} = ≈-meta λ i → [idᵛ] 238 | [idᵛ] {t = tm-oper f es} = ≈-oper λ i → ≈-trans ([]ᵛ-resp-≡ᵛ idᵛ+idᵛ) [idᵛ] 239 | 240 | -- Extension respects composition 241 | ⇑ᵛ-resp-∘ᵛ : ∀ {Γ Δ Ξ Ψ} {ρ : Γ ⇒ᵛ Δ} {τ : Δ ⇒ᵛ Ξ} → ⇑ᵛ {Ξ = Ψ} (τ ∘ᵛ ρ) ≡ᵛ (⇑ᵛ τ) ∘ᵛ (⇑ᵛ ρ) 242 | ⇑ᵛ-resp-∘ᵛ (var-inl x) = refl 243 | ⇑ᵛ-resp-∘ᵛ (var-inr y) = refl 244 | 245 | -- Right extension respects composition 246 | ʳ⇑ᵛ-resp-∘ᵛ : ∀ {Γ Δ Ξ Ψ} {ρ : Γ ⇒ᵛ Δ} {τ : Δ ⇒ᵛ Ξ} → ʳ⇑ᵛ (τ ∘ᵛ ρ) {Ξ = Ψ} ≡ᵛ (ʳ⇑ᵛ τ) ∘ᵛ (ʳ⇑ᵛ ρ) 247 | ʳ⇑ᵛ-resp-∘ᵛ (var-inl x) = refl 248 | ʳ⇑ᵛ-resp-∘ᵛ (var-inr y) = refl 249 | 250 | -- The action of a renaming is functorial 251 | [∘ᵛ] : ∀ {Θ Γ Δ Ξ} {ρ : Γ ⇒ᵛ Δ} {τ : Δ ⇒ᵛ Ξ} {A} {t : Term Θ Γ A} 252 | → [ τ ∘ᵛ ρ ]ᵛ t ≈ [ τ ]ᵛ ([ ρ ]ᵛ t) 253 | [∘ᵛ] {t = tm-var x} = ≈-refl 254 | [∘ᵛ] {t = tm-meta M ts} = ≈-meta (λ i → [∘ᵛ]) 255 | [∘ᵛ] {t = tm-oper f es} = ≈-oper (λ i → ≈-trans ([]ᵛ-resp-≡ᵛ ⇑ᵛ-resp-∘ᵛ) [∘ᵛ]) 256 | 257 | ∘ᵛ-resp-ʳ⇑ᵛ : ∀ {Γ Δ Ξ Λ} {ρ : Γ ⇒ᵛ Δ} {τ : Δ ⇒ᵛ Ξ} 258 | → ʳ⇑ᵛ (τ ∘ᵛ ρ) {Ξ = Λ} ≡ᵛ ʳ⇑ᵛ τ ∘ᵛ ʳ⇑ᵛ ρ 259 | ∘ᵛ-resp-ʳ⇑ᵛ (var-inl x) = refl 260 | ∘ᵛ-resp-ʳ⇑ᵛ (var-inr y) = refl 261 | 262 | ∘ᵛ-resp-ʳ⇑ᵛ-term : ∀ {Θ Γ Δ Ξ Λ A} {ρ : Γ ⇒ᵛ Δ} {τ : Δ ⇒ᵛ Ξ} {t : Term Θ (Λ ,, Γ) A} 263 | → [ ʳ⇑ᵛ (τ ∘ᵛ ρ) {Ξ = Λ} ]ᵛ t ≈ [ ʳ⇑ᵛ τ ]ᵛ ([ ʳ⇑ᵛ ρ ]ᵛ t) 264 | ∘ᵛ-resp-ʳ⇑ᵛ-term {Θ} {Γ} {Δ} {Ξ} {Λ} {A} {ρ} {τ} {t = t} = 265 | let open SetoidR (Term-setoid Θ (Λ ,, Ξ) A) in 266 | begin 267 | [ ʳ⇑ᵛ (τ ∘ᵛ ρ) {Ξ = Λ} ]ᵛ t ≈⟨ []ᵛ-resp-≡ᵛ ∘ᵛ-resp-ʳ⇑ᵛ ⟩ 268 | [ ʳ⇑ᵛ τ ∘ᵛ ʳ⇑ᵛ ρ ]ᵛ t ≈⟨ [∘ᵛ] ⟩ 269 | [ ʳ⇑ᵛ τ ]ᵛ ([ ʳ⇑ᵛ ρ ]ᵛ t) 270 | ∎ 271 | 272 | ʳ⇑ᵛ-comm-inrᵛ-term : ∀ {Θ Γ Δ Ξ A} {ρ : Γ ⇒ᵛ Δ} {t : Term Θ Γ A} 273 | → ([ ʳ⇑ᵛ ρ {Ξ = Ξ} ]ᵛ ([ inrᵛ {Δ = Γ} ]ᵛ t)) ≈ ([ inrᵛ ]ᵛ ([ ρ ]ᵛ t)) 274 | ʳ⇑ᵛ-comm-inrᵛ-term {Θ} {Γ} {Δ} {Ξ} {A} {ρ} {t = t} = 275 | let open SetoidR (Term-setoid Θ (Ξ ,, Δ) A) in 276 | begin 277 | [ ʳ⇑ᵛ ρ ]ᵛ ([ inrᵛ ]ᵛ t) ≈⟨ ≈-sym [∘ᵛ] ⟩ 278 | [ ʳ⇑ᵛ ρ ∘ᵛ var-inr ]ᵛ t ≈⟨ []ᵛ-resp-≡ᵛ (ʳ⇑ᵛ-comm-inrᵛ ρ) ⟩ 279 | [ inrᵛ ∘ᵛ ρ ]ᵛ t ≈⟨ [∘ᵛ] ⟩ 280 | [ inrᵛ ]ᵛ ([ ρ ]ᵛ t) 281 | ∎ 282 | 283 | -- Forming terms over a given metacontext and sort is functorial in the context 284 | module _ {Θ : MContext} {A : sort} where 285 | open Categories.Functor 286 | open Categories.Category.Instance.Setoids 287 | 288 | Term-Functor : Functor VContexts (Setoids ℓ ℓ) 289 | Term-Functor = 290 | record 291 | { F₀ = λ Γ → Term-setoid Θ Γ A 292 | ; F₁ = λ ρ → record { _⟨$⟩_ = [ ρ ]ᵛ_ ; cong = []ᵛ-resp-≈ } 293 | ; identity = ≈-trans [idᵛ] 294 | ; homomorphism = λ ξ → ≈-trans ([]ᵛ-resp-≈ ξ) [∘ᵛ] 295 | ; F-resp-≈ = λ ζ ξ → ≈-trans ([]ᵛ-resp-≡ᵛ ζ) ([]ᵛ-resp-≈ ξ) 296 | } 297 | -------------------------------------------------------------------------------- /src/SingleSorted/AlgebraicTheory.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 2 | open import Agda.Builtin.Nat 3 | open import Relation.Binary.PropositionalEquality using (_≡_; refl) 4 | open import Data.Fin 5 | 6 | open import Relation.Binary 7 | 8 | module SingleSorted.AlgebraicTheory where 9 | 10 | open import SingleSorted.Context public 11 | 12 | Arity : Set 13 | Arity = Context 14 | 15 | arg : Arity → Set 16 | arg = var 17 | 18 | -- an algebraic signature 19 | record Signature : Set₁ where 20 | field 21 | oper : Set -- operations 22 | oper-arity : oper → Arity -- the arity of an operation 23 | 24 | -- terms over a signature in a context of a given sort 25 | data Term (Γ : Context) : Set where 26 | tm-var : var Γ → Term Γ 27 | tm-oper : ∀ (f : oper) → (arg (oper-arity f) → Term Γ) → Term Γ 28 | 29 | -- Substitutions (definitions - some useful properties are in another file) 30 | _⇒s_ : ∀ (Γ Δ : Context) → Set 31 | Γ ⇒s Δ = var Δ → Term Γ 32 | 33 | infix 4 _⇒s_ 34 | 35 | -- identity substitution 36 | id-s : ∀ {Γ : Context} → Γ ⇒s Γ 37 | id-s = tm-var 38 | 39 | -- the action of a substitution on a term (contravariant) 40 | _[_]s : ∀ {Γ Δ : Context} → Term Δ → Γ ⇒s Δ → Term Γ 41 | (tm-var x) [ σ ]s = σ x 42 | (tm-oper f x) [ σ ]s = tm-oper f (λ i → x i [ σ ]s) 43 | 44 | infixr 6 _[_]s 45 | 46 | -- composition of substitutions 47 | _∘s_ : ∀ {Γ Δ Θ : Context} → Δ ⇒s Θ → Γ ⇒s Δ → Γ ⇒s Θ 48 | (σ ∘s ρ) x = σ x [ ρ ]s 49 | 50 | infixl 7 _∘s_ 51 | 52 | -- Axioms are equations in context 53 | record Equation : Set where 54 | field 55 | eq-ctx : Context 56 | eq-lhs : Term eq-ctx 57 | eq-rhs : Term eq-ctx 58 | 59 | -- Theory 60 | -- an equational theory is a family of equations over a given sort 61 | record Theory ℓ (Σ : Signature) : Set (lsuc ℓ) where 62 | open Signature Σ public 63 | 64 | field 65 | ax : Set ℓ -- the axiom names 66 | ax-eq : ax → Equation -- the axioms 67 | 68 | ax-ctx : ax → Context 69 | ax-ctx ε = Equation.eq-ctx (ax-eq ε) 70 | 71 | ax-lhs : ∀ (ε : ax) → Term (ax-ctx ε) 72 | ax-lhs ε = Equation.eq-lhs (ax-eq ε) 73 | 74 | ax-rhs : ∀ (ε : ax) → Term (ax-ctx ε) 75 | ax-rhs ε = Equation.eq-rhs (ax-eq ε) 76 | 77 | infix 4 _⊢_≈_ 78 | 79 | -- equality of terms 80 | data _⊢_≈_ : (Γ : Context) → Term Γ → Term Γ → Set (lsuc ℓ) where 81 | -- general rules 82 | eq-refl : ∀ {Γ} {t : Term Γ} → Γ ⊢ t ≈ t 83 | eq-symm : ∀ {Γ} {s t : Term Γ} → Γ ⊢ s ≈ t → Γ ⊢ t ≈ s 84 | eq-tran : ∀ {Γ} {s t u : Term Γ} → Γ ⊢ s ≈ t → Γ ⊢ t ≈ u → Γ ⊢ s ≈ u 85 | -- congruence rule 86 | eq-congr : ∀ {Γ} {f : oper} {xs ys : arg (oper-arity f) → Term Γ} → 87 | (∀ i → Γ ⊢ xs i ≈ ys i) → Γ ⊢ tm-oper f xs ≈ tm-oper f ys 88 | -- equational axiom 89 | eq-axiom : ∀ (ε : ax) {Γ : Context} (σ : Γ ⇒s (ax-ctx ε)) → 90 | Γ ⊢ ax-lhs ε [ σ ]s ≈ ax-rhs ε [ σ ]s 91 | 92 | ≡-⊢-≈ : {Γ : Context} {s t : Term Γ} → s ≡ t → Γ ⊢ s ≈ t 93 | ≡-⊢-≈ refl = eq-refl 94 | 95 | -- the action of the identity substitution is the identity 96 | id-action : ∀ {Γ : Context} {a : Term Γ} → (Γ ⊢ a ≈ (a [ id-s ]s)) 97 | id-action {a = tm-var a} = eq-refl 98 | id-action {a = tm-oper f x} = eq-congr (λ i → id-action {a = x i}) 99 | 100 | eq-axiom-id : ∀ (ε : ax) → ax-ctx ε ⊢ ax-lhs ε ≈ ax-rhs ε 101 | eq-axiom-id ε = eq-tran id-action (eq-tran (eq-axiom ε id-s) (eq-symm id-action)) 102 | 103 | eq-setoid : ∀ (Γ : Context) → Setoid lzero (lsuc ℓ) 104 | eq-setoid Γ = 105 | record 106 | { Carrier = Term Γ 107 | ; _≈_ = λ s t → (Γ ⊢ s ≈ t) 108 | ; isEquivalence = 109 | record 110 | { refl = eq-refl 111 | ; sym = eq-symm 112 | ; trans = eq-tran 113 | } 114 | } 115 | -------------------------------------------------------------------------------- /src/SingleSorted/Combinators.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; subst) 2 | open import Data.Sum 3 | 4 | import SingleSorted.AlgebraicTheory as SS 5 | 6 | module SingleSorted.Combinators where 7 | 8 | 9 | module Sum {𝓈} (Σ₁ Σ₂ : SS.Signature) (T₁ : SS.Theory 𝓈 Σ₁) (T₂ : SS.Theory 𝓈 Σ₂) where 10 | 11 | -- disjoint sum of signatures 12 | S : SS.Signature 13 | S = record { oper = SS.Signature.oper Σ₁ ⊎ SS.Signature.oper Σ₂ 14 | ; oper-arity = [ SS.Signature.oper-arity Σ₁ , SS.Signature.oper-arity Σ₂ ] 15 | } 16 | 17 | inj-term-l : ∀ {Γ : SS.Context} → SS.Signature.Term Σ₁ Γ → SS.Signature.Term S Γ 18 | inj-term-l {Γ} (SS.Signature.tm-var x) = SS.Signature.tm-var x 19 | inj-term-l {Γ} (SS.Signature.tm-oper f ts) = SS.Signature.tm-oper (inj₁ f) λ{ i → inj-term-l (ts i)} 20 | 21 | inj-term-r : ∀ {Γ : SS.Context} → SS.Signature.Term Σ₂ Γ → SS.Signature.Term S Γ 22 | inj-term-r {Γ} (SS.Signature.tm-var x) = SS.Signature.tm-var x 23 | inj-term-r {Γ} (SS.Signature.tm-oper f ts) = SS.Signature.tm-oper (inj₂ f) λ{ i → inj-term-r (ts i)} 24 | 25 | coerce₁ : SS.Signature.Equation Σ₁ → SS.Signature.Equation S 26 | coerce₁ eq = record { eq-ctx = SS.Signature.Equation.eq-ctx eq 27 | ; eq-lhs = inj-term-l (SS.Signature.Equation.eq-lhs eq) 28 | ; eq-rhs = inj-term-l (SS.Signature.Equation.eq-rhs eq) 29 | } 30 | 31 | coerce₂ : SS.Signature.Equation Σ₂ → SS.Signature.Equation S 32 | coerce₂ eq = record { eq-ctx = SS.Signature.Equation.eq-ctx eq 33 | ; eq-lhs = inj-term-r (SS.Signature.Equation.eq-lhs eq) 34 | ; eq-rhs = inj-term-r (SS.Signature.Equation.eq-rhs eq) 35 | } 36 | 37 | -- define a theory with the set of axioms a union of the axioms of both theories 38 | T : SS.Theory 𝓈 S 39 | T = record { ax = SS.Theory.ax T₁ ⊎ SS.Theory.ax T₂ 40 | ; ax-eq = [ (λ a → coerce₁ (SS.Theory.ax-eq T₁ a)) , (λ a → coerce₂ (SS.Theory.ax-eq T₂ a)) ] 41 | } 42 | -------------------------------------------------------------------------------- /src/SingleSorted/Context.agda: -------------------------------------------------------------------------------- 1 | module SingleSorted.Context where 2 | 3 | -- An attempt to define more structured context 4 | -- that directly support the cartesian structure 5 | 6 | data Context : Set where 7 | ctx-empty : Context 8 | ctx-slot : Context 9 | ctx-concat : Context → Context → Context 10 | 11 | -- the variables in a context 12 | data var : Context → Set where 13 | var-var : var ctx-slot 14 | var-inl : ∀ {Γ Δ} → var Γ → var (ctx-concat Γ Δ) 15 | var-inr : ∀ {Γ Δ} → var Δ → var (ctx-concat Γ Δ) 16 | 17 | -- It is absurd to have a variable in the empty context 18 | ctx-empty-absurd : ∀ {ℓ} {A : Set ℓ} → var ctx-empty → A 19 | ctx-empty-absurd () 20 | -------------------------------------------------------------------------------- /src/SingleSorted/Example.agda: -------------------------------------------------------------------------------- 1 | module SingleSorted.Example where 2 | 3 | import Relation.Binary.PropositionalEquality as Eq 4 | open Eq using (_≡_; refl) 5 | open import Data.Product using (_×_; proj₁; proj₂; <_,_>; ∃; ∃-syntax; _,_) 6 | import Function using (_∘_) 7 | open import SingleSorted.AlgebraicTheory 8 | open import Categories.Category.Instance.Sets 9 | open import Categories.Category.Cartesian 10 | open import Agda.Primitive 11 | open import Categories.Category 12 | 13 | module Sets₀ where 14 | 𝒮 : Category (lsuc lzero) lzero lzero 15 | 𝒮 = Sets lzero 16 | 17 | open Sets₀ 18 | open Category 𝒮 19 | 20 | -- Function extensionality 21 | postulate 22 | funext : ∀ {X : Set} {Y : X → Set} {f g : ∀ (x : X) → (Y x)} → (∀ (x : X) → ((f x) ≡ (g x))) → (f ≡ g) 23 | 24 | 25 | record singleton : Set where 26 | constructor ⋆ 27 | 28 | η-singleton : ∀ (x : singleton) → ⋆ ≡ x 29 | η-singleton ⋆ = refl 30 | 31 | prod-elem-structure : ∀ {A B : Set} {x : A × B} → ∃[ a ] ∃[ b ] (a , b) ≡ x 32 | prod-elem-structure {A} {B} {x} = proj₁ x Data.Product., (proj₂ x) , refl 33 | 34 | first-factor : ∀ {X A B : Obj} {f : X ⇒ A} {g : X ⇒ B} → proj₁ ∘ < f , g > ≡ f 35 | first-factor {X} {A} {B} {f} {g} = funext λ x → refl 36 | 37 | second-factor : ∀ {X A B : Obj} {f : X ⇒ A} {g : X ⇒ B} → proj₂ ∘ < f , g > ≡ g 38 | second-factor {X} {A} {B} {f} {g} = funext λ{ x → refl} 39 | 40 | unique-map-into-product : ∀ {X A B : Set} {h : X → A × B} {f : X → A} {g : X → B} {x : X} 41 | → (proj₁ ∘ h) x ≡ f x 42 | → (proj₂ ∘ h) x ≡ g x 43 | -------------------- 44 | → < f , g > x ≡ h x 45 | unique-map-into-product {X} {A} {B} {h} {f} {g} {x} eq1 eq2 46 | rewrite first-factor {X} {A} {B} {f} {g} | second-factor {X} {A} {B} {f} {g} 47 | = {!!} 48 | 49 | 50 | cartesian-Sets : Cartesian 𝒮 51 | cartesian-Sets = 52 | record 53 | { terminal = record 54 | { ⊤ = singleton 55 | ; ⊤-is-terminal = record 56 | { ! = λ{ x → ⋆} 57 | ; !-unique = λ{ f {x} → η-singleton (f x)} 58 | } 59 | } 60 | ; products = record 61 | { product = λ {A} {B} → 62 | record 63 | { A×B = A × B 64 | ; π₁ = proj₁ 65 | ; π₂ = proj₂ 66 | ; ⟨_,_⟩ = <_,_> 67 | ; project₁ = λ{ → refl} 68 | ; project₂ = λ{ → refl} 69 | ; unique = λ{ {X} {h} {f} {g} eq1 eq2 {x} → {!!}} 70 | } 71 | 72 | } 73 | } 74 | 75 | 76 | -------------------------------------------------------------------------------- /src/SingleSorted/Group.agda: -------------------------------------------------------------------------------- 1 | open import SingleSorted.AlgebraicTheory 2 | open import Agda.Primitive using (lzero; lsuc; _⊔_) 3 | open import Relation.Binary.PropositionalEquality using (_≡_; refl) 4 | open import SingleSorted.Substitution 5 | open import Data.Nat using (ℕ; zero; suc; _+_; _*_) 6 | 7 | module SingleSorted.Group where 8 | 9 | data GroupOp : Set where 10 | e : GroupOp 11 | inv : GroupOp 12 | mul : GroupOp 13 | 14 | _ : Context 15 | _ = ctx-empty 16 | 17 | _ : Context 18 | _ = ctx-concat ctx-slot ctx-empty 19 | 20 | _ : var ctx-slot 21 | _ = var-var 22 | 23 | _ : var (ctx-concat ctx-slot ctx-slot) 24 | _ = var-inl var-var 25 | 26 | _ : var (ctx-concat ctx-slot ctx-slot) 27 | _ = var-inr var-var 28 | 29 | ctx-1 : Context 30 | ctx-1 = ctx-slot 31 | 32 | ctx-2 : Context 33 | ctx-2 = ctx-concat ctx-1 ctx-1 34 | 35 | ctx : ∀ (n : ℕ) → Context 36 | ctx zero = ctx-empty 37 | ctx (suc n) = ctx-concat (ctx n) ctx-slot 38 | 39 | -- the signature of the theory of groups 40 | -- has one constant, one unary operation, one binary operation 41 | Σ : Signature 42 | Σ = record { oper = GroupOp ; oper-arity = λ{ e → ctx-empty ; inv → ctx 1 ; mul → ctx 2} } 43 | 44 | open Signature Σ 45 | 46 | -- some example terms 47 | _ : Term ctx-1 48 | _ = tm-var var-var 49 | 50 | _ : Term ctx-2 51 | _ = tm-var (var-inr var-var) 52 | 53 | _ : Term ctx-2 54 | _ = tm-var (var-inr var-var) 55 | 56 | 57 | -- helper functions for creating terms 58 | e' : ∀ {Γ : Context} → Term Γ 59 | e' {Γ} = tm-oper e λ() 60 | 61 | -- inv' : ∀ {Γ : Context} → Term Γ → Term Γ 62 | -- inv' x = tm-oper inv λ{ _ → x} 63 | 64 | -- mul' : ∀ {Γ : Context} → Term Γ → Term Γ → Term Γ 65 | -- mul' x y = tm-oper mul λ{ (var-inl _) → x ; (var-inr _) → y} 66 | 67 | concat-empty : var (ctx-concat ctx-empty ctx-slot) → (var ctx-slot) 68 | concat-empty (var-inr x) = x 69 | 70 | 71 | x*y : Term ctx-2 72 | x*y = tm-oper mul λ{ (var-inl x) → tm-var (var-inl (concat-empty x)) ; (var-inr y) → tm-var (var-inr y)} 73 | 74 | -- concat-empty-idʳ : ctx-concat ctx-empty ctx-slot ≡ ctx-slot 75 | -- concat-empty-idʳ = {!!} 76 | 77 | singleton-context : (var ctx-slot) → var (ctx-concat ctx-empty ctx-slot) 78 | singleton-context (var-var) = var-inr var-var 79 | 80 | σ : ∀ {Γ : Context} {t : Term Γ} → Γ ⇒s (ctx 1) 81 | σ {Γ} {t} = λ{ (var-inr var-var) → t} 82 | 83 | δ : ∀ {Γ : Context} {t : Term Γ} {s : Term Γ} → Γ ⇒s (ctx 2) 84 | δ {Γ} {t} {s} = λ{ (var-inl x) → t ; (var-inr y) → s} 85 | 86 | _∗_ : ∀ {Γ} → Term Γ → Term Γ → Term Γ 87 | t ∗ s = tm-oper mul λ{ xs → δ {t = t} {s = s} xs} 88 | 89 | _ⁱ : ∀ {Γ : Context} → Term Γ → Term Γ 90 | t ⁱ = tm-oper inv λ{ x → σ {t = t} x} 91 | 92 | -- _∗_ : ∀ {Γ} → Term Γ → Term Γ → Term Γ 93 | -- t ∗ s = tm-oper mul λ{ (var-inl x) → t ; (var-inr args) → s} 94 | 95 | -- _ⁱ : ∀ {Γ : Context} → Term Γ → Term Γ 96 | -- t ⁱ = tm-oper inv λ{ x → t } 97 | 98 | infixl 5 _∗_ 99 | infix 6 _ⁱ 100 | 101 | _ : Term (ctx 2) 102 | _ = tm-var (var-inl (var-inr var-var)) ∗ tm-var (var-inr var-var) 103 | 104 | _ : Term (ctx 1) 105 | _ = e' ∗ a 106 | where 107 | a : Term (ctx 1) 108 | a = tm-var (var-inr var-var) 109 | 110 | -- group equations 111 | data GroupEq : Set where 112 | mul-assoc e-left e-right inv-left inv-right : GroupEq 113 | 114 | mul-assoc-ax : Equation 115 | e-left-ax : Equation 116 | e-right-ax : Equation 117 | inv-left-ax : Equation 118 | inv-right-ax : Equation 119 | 120 | mul-assoc-ax = record { eq-ctx = ctx 3 121 | ; eq-lhs = x ∗ y ∗ z 122 | ; eq-rhs = x ∗ (y ∗ z) 123 | } 124 | where 125 | x : Term (ctx 3) 126 | y : Term (ctx 3) 127 | z : Term (ctx 3) 128 | x = tm-var (var-inl (var-inl (var-inr var-var))) 129 | y = tm-var (var-inl (var-inr var-var)) 130 | z = tm-var (var-inr var-var) 131 | 132 | e-left-ax = record { eq-ctx = ctx 1 ; eq-lhs = e' ∗ x ; eq-rhs = x } 133 | where 134 | x : Term (ctx 1) 135 | x = tm-var (var-inr var-var) 136 | 137 | e-right-ax = record { eq-ctx = ctx 1 ; eq-lhs = x ∗ e' ; eq-rhs = x } 138 | where 139 | x : Term (ctx 1) 140 | x = tm-var (var-inr var-var) 141 | 142 | 143 | inv-left-ax = record { eq-ctx = ctx 1 ; eq-lhs = x ⁱ ∗ x ; eq-rhs = e' } 144 | where 145 | x : Term (ctx 1) 146 | x = tm-var (var-inr var-var) 147 | 148 | inv-right-ax = record { eq-ctx = ctx 1 ; eq-lhs = x ∗ x ⁱ ; eq-rhs = e' } 149 | where 150 | x : Term (ctx 1) 151 | x = tm-var (var-inr var-var) 152 | 153 | 154 | 𝒢 : Theory lzero Σ 155 | 𝒢 = record { ax = GroupEq 156 | ; ax-eq = λ{ mul-assoc → mul-assoc-ax 157 | ; e-left → e-left-ax 158 | ; e-right → e-right-ax 159 | ; inv-left → inv-left-ax 160 | ; inv-right → inv-right-ax 161 | } 162 | } 163 | -------------------------------------------------------------------------------- /src/SingleSorted/Interpretation.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_) 2 | 3 | import Categories.Category as Category 4 | import Categories.Category.Cartesian as Cartesian 5 | 6 | open import SingleSorted.AlgebraicTheory 7 | import SingleSorted.Power as Power 8 | 9 | module SingleSorted.Interpretation 10 | {o ℓ e} 11 | (Σ : Signature) 12 | {𝒞 : Category.Category o ℓ e} 13 | (cartesian-𝒞 : Cartesian.Cartesian 𝒞) where 14 | open Signature Σ 15 | open Category.Category 𝒞 16 | open Power 𝒞 17 | 18 | -- An interpretation of Σ in 𝒞 19 | record Interpretation : Set (o ⊔ ℓ ⊔ e) where 20 | 21 | field 22 | interp-carrier : Obj 23 | interp-pow : Powered interp-carrier 24 | interp-oper : ∀ (f : oper) → Powered.pow interp-pow (oper-arity f) ⇒ interp-carrier 25 | 26 | open Powered interp-pow 27 | 28 | -- the interpretation of a term 29 | interp-term : ∀ {Γ : Context} → Term Γ → (pow Γ) ⇒ interp-carrier 30 | interp-term (tm-var x) = π x 31 | interp-term (tm-oper f ts) = interp-oper f ∘ tuple (oper-arity f) (λ i → interp-term (ts i)) 32 | 33 | -- the interpretation of a context 34 | interp-ctx : Context → Obj 35 | interp-ctx Γ = pow Γ 36 | 37 | -- the interpretation of a substitution 38 | interp-subst : ∀ {Γ Δ} → Γ ⇒s Δ → interp-ctx Γ ⇒ interp-ctx Δ 39 | interp-subst {Γ} {Δ} σ = tuple Δ λ i → interp-term (σ i) 40 | 41 | -- interpretation commutes with substitution 42 | open HomReasoning 43 | 44 | interp-[]s : ∀ {Γ Δ} {t : Term Δ} {σ : Γ ⇒s Δ} → 45 | interp-term (t [ σ ]s) ≈ interp-term t ∘ interp-subst σ 46 | interp-[]s {Γ} {Δ} {tm-var x} {σ} = ⟺ (project {Γ = Δ}) 47 | interp-[]s {Γ} {Δ} {tm-oper f ts} {σ} = (∘-resp-≈ʳ 48 | (tuple-cong 49 | {fs = λ i → interp-term (ts i [ σ ]s)} 50 | {gs = λ z → interp-term (ts z) ∘ interp-subst σ} 51 | (λ i → interp-[]s {t = ts i} {σ = σ}) 52 | ○ (∘-distribʳ-tuple 53 | {Γ = oper-arity f} 54 | {fs = λ z → interp-term (ts z)} 55 | {g = interp-subst σ}))) 56 | ○ (Equiv.refl ○ sym-assoc) 57 | 58 | -- -- Every signature has the trivial interpretation 59 | 60 | Trivial : Interpretation 61 | Trivial = 62 | let open Cartesian.Cartesian cartesian-𝒞 in 63 | record 64 | { interp-carrier = ⊤ 65 | ; interp-pow = StandardPowered cartesian-𝒞 ⊤ 66 | ; interp-oper = λ f → ! } 67 | 68 | record HomI (A B : Interpretation) : Set (o ⊔ ℓ ⊔ e) where 69 | open Interpretation 70 | open Powered 71 | 72 | field 73 | hom-morphism : interp-carrier A ⇒ interp-carrier B 74 | hom-commute : 75 | ∀ (f : oper) → 76 | hom-morphism ∘ interp-oper A f ≈ 77 | interp-oper B f ∘ tuple (interp-pow B) (oper-arity f) (λ i → hom-morphism ∘ π (interp-pow A) i) 78 | 79 | -- The identity homomorphism 80 | IdI : ∀ (A : Interpretation) → HomI A A 81 | IdI A = 82 | let open Interpretation A in 83 | let open HomReasoning in 84 | let open Powered interp-pow in 85 | record 86 | { hom-morphism = id 87 | ; hom-commute = 88 | λ f → 89 | begin 90 | (id ∘ interp-oper f) ≈⟨ identityˡ ⟩ 91 | interp-oper f ≈˘⟨ identityʳ ⟩ 92 | (interp-oper f ∘ id) ≈˘⟨ refl⟩∘⟨ unique (λ i → identityʳ ○ ⟺ identityˡ) ⟩ 93 | (interp-oper f ∘ tuple (oper-arity f) (λ i → id ∘ π i)) ∎ 94 | } 95 | 96 | -- Compositon of homomorphisms 97 | _∘I_ : ∀ {A B C : Interpretation} → HomI B C → HomI A B → HomI A C 98 | _∘I_ {A} {B} {C} ϕ ψ = 99 | let open HomI in 100 | record { hom-morphism = hom-morphism ϕ ∘ hom-morphism ψ 101 | ; hom-commute = 102 | let open Interpretation in 103 | let open Powered in 104 | let open HomReasoning in 105 | λ f → 106 | begin 107 | (((hom-morphism ϕ) ∘ hom-morphism ψ) ∘ interp-oper A f) ≈⟨ assoc ⟩ 108 | (hom-morphism ϕ ∘ hom-morphism ψ ∘ interp-oper A f) ≈⟨ (refl⟩∘⟨ hom-commute ψ f) ⟩ 109 | (hom-morphism ϕ ∘ 110 | interp-oper B f ∘ 111 | tuple (interp-pow B) (oper-arity f) 112 | (λ i → hom-morphism ψ ∘ π (interp-pow A) i)) ≈˘⟨ assoc ⟩ 113 | ((hom-morphism ϕ ∘ interp-oper B f) ∘ 114 | tuple (interp-pow B) (oper-arity f) 115 | (λ i → hom-morphism ψ ∘ π (interp-pow A) i)) ≈⟨ (hom-commute ϕ f ⟩∘⟨refl) ⟩ 116 | ((interp-oper C f ∘ 117 | tuple (interp-pow C) (oper-arity f) 118 | (λ i → hom-morphism ϕ ∘ π (interp-pow B) i)) 119 | ∘ 120 | tuple (interp-pow B) (oper-arity f) 121 | (λ i → hom-morphism ψ ∘ π (interp-pow A) i)) ≈⟨ assoc ⟩ 122 | (interp-oper C f ∘ 123 | tuple (interp-pow C) (oper-arity f) 124 | (λ i → hom-morphism ϕ ∘ π (interp-pow B) i) 125 | ∘ 126 | tuple (interp-pow B) (oper-arity f) 127 | (λ i → hom-morphism ψ ∘ π (interp-pow A) i)) ≈⟨ (refl⟩∘⟨ ⟺ (∘-distribʳ-tuple (interp-pow C))) ⟩ 128 | (interp-oper C f ∘ 129 | tuple (interp-pow C) (oper-arity f) 130 | (λ x → 131 | (hom-morphism ϕ ∘ π (interp-pow B) x) ∘ 132 | tuple (interp-pow B) (oper-arity f) 133 | (λ i → hom-morphism ψ ∘ π (interp-pow A) i))) ≈⟨ (refl⟩∘⟨ tuple-cong (interp-pow C) λ i → assoc) ⟩ 134 | (interp-oper C f ∘ 135 | tuple (interp-pow C) (oper-arity f) 136 | (λ z → 137 | hom-morphism ϕ ∘ 138 | π (interp-pow B) z ∘ 139 | tuple (interp-pow B) (oper-arity f) 140 | (λ i → hom-morphism ψ ∘ π (interp-pow A) i))) ≈⟨ (refl⟩∘⟨ tuple-cong (interp-pow C) λ i → refl⟩∘⟨ project (interp-pow B)) ⟩ 141 | (interp-oper C f ∘ 142 | tuple (interp-pow C) (oper-arity f) 143 | (λ z → hom-morphism ϕ ∘ hom-morphism ψ ∘ π (interp-pow A) z)) ≈⟨ (refl⟩∘⟨ tuple-cong (interp-pow C) λ i → sym-assoc) ⟩ 144 | (interp-oper C f ∘ 145 | tuple (interp-pow C) (oper-arity f) 146 | (λ z → (hom-morphism ϕ ∘ hom-morphism ψ) ∘ π (interp-pow A) z)) ∎} 147 | -------------------------------------------------------------------------------- /src/SingleSorted/Model.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_) 2 | 3 | import Categories.Category as Category 4 | import Categories.Category.Cartesian as Cartesian 5 | import SingleSorted.Interpretation as Interpretation 6 | 7 | open import SingleSorted.AlgebraicTheory 8 | open import SingleSorted.Substitution 9 | import SingleSorted.Power as Power 10 | 11 | module SingleSorted.Model {o ℓ e ℓt} 12 | {Σ : Signature} 13 | (T : Theory ℓt Σ) 14 | {𝒞 : Category.Category o ℓ e} 15 | (cartesian-𝒞 : Cartesian.Cartesian 𝒞) where 16 | 17 | -- Model of a theory 18 | 19 | record Model (I : Interpretation.Interpretation Σ cartesian-𝒞) : Set (ℓt ⊔ o ⊔ ℓ ⊔ e) where 20 | 21 | open Theory T 22 | open Category.Category 𝒞 23 | open Interpretation.Interpretation I 24 | open HomReasoning 25 | 26 | field 27 | model-eq : ∀ (ε : ax) → interp-term (ax-lhs ε) ≈ interp-term (ax-rhs ε) 28 | 29 | -- Soundness of semantics 30 | module _ where 31 | open Power.Powered interp-pow 32 | 33 | -- first we show that substitution preserves validity 34 | model-resp-[]s : ∀ {Γ Δ} {u v : Term Γ} {σ : Δ ⇒s Γ} → 35 | interp-term u ≈ interp-term v → interp-term (u [ σ ]s) ≈ interp-term (v [ σ ]s) 36 | model-resp-[]s {u = u} {v = v} {σ = σ} ξ = 37 | begin 38 | interp-term (u [ σ ]s) ≈⟨ interp-[]s {t = u} ⟩ 39 | (interp-term u ∘ interp-subst σ) ≈⟨ ξ ⟩∘⟨refl ⟩ 40 | (interp-term v ∘ interp-subst σ) ≈˘⟨ interp-[]s {t = v} ⟩ 41 | interp-term (v [ σ ]s) ∎ 42 | 43 | -- the soundness statement 44 | model-⊢-≈ : ∀ {Γ} {s t : Term Γ} → Γ ⊢ s ≈ t → interp-term s ≈ interp-term t 45 | model-⊢-≈ eq-refl = Equiv.refl 46 | model-⊢-≈ (eq-symm ξ) = ⟺ (model-⊢-≈ ξ) 47 | model-⊢-≈ (eq-tran ξ θ) = (model-⊢-≈ ξ) ○ (model-⊢-≈ θ) 48 | model-⊢-≈ (eq-congr ξ) = ∘-resp-≈ʳ (unique (λ i → project ○ model-⊢-≈ (eq-symm (ξ i)))) 49 | model-⊢-≈ (eq-axiom ε σ) = model-resp-[]s {u = ax-lhs ε} {v = ax-rhs ε} (model-eq ε) 50 | 51 | -- Every theory has the trivial model, whose carrier is the terminal object 52 | Trivial : Model (Interpretation.Trivial Σ cartesian-𝒞) 53 | Trivial = 54 | let open Cartesian.Cartesian cartesian-𝒞 in 55 | record { model-eq = λ ε → !-unique₂ } 56 | -------------------------------------------------------------------------------- /src/SingleSorted/Power.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (_⊔_) 2 | 3 | import Categories.Category as Category 4 | import Categories.Category.Cartesian as Cartesian 5 | 6 | open import SingleSorted.Context 7 | 8 | module SingleSorted.Power 9 | {o ℓ e} 10 | (𝒞 : Category.Category o ℓ e) where 11 | 12 | open Category.Category 𝒞 13 | open HomReasoning 14 | 15 | record Powered (A : Obj) : Set (o ⊔ ℓ ⊔ e) where 16 | 17 | field 18 | pow : Context → Obj 19 | π : ∀ {Γ} → var Γ → pow Γ ⇒ A 20 | tuple : ∀ Γ {B} → (var Γ → B ⇒ A) → B ⇒ pow Γ 21 | project : ∀ {Γ} {B} {x : var Γ} {fs : var Γ → B ⇒ A} → π x ∘ tuple Γ fs ≈ fs x 22 | unique : ∀ {Γ} {B} {fs : var Γ → B ⇒ A} {g : B ⇒ pow Γ} → (∀ i → π i ∘ g ≈ fs i) → tuple Γ fs ≈ g 23 | 24 | η : ∀ {Γ} → tuple Γ π ≈ id 25 | η = unique (λ i → identityʳ) 26 | 27 | ! : ∀ {B : Obj} → B ⇒ pow ctx-empty 28 | ! = tuple ctx-empty ctx-empty-absurd 29 | 30 | !-unique : ∀ {B : Obj} {f : B ⇒ pow ctx-empty} → ! ≈ f 31 | !-unique = unique λ i → ctx-empty-absurd i 32 | 33 | !-unique₂ : ∀ {B : Obj} {f g : B ⇒ pow ctx-empty} → f ≈ g 34 | !-unique₂ = (⟺ !-unique) ○ !-unique 35 | 36 | tuple-cong : ∀ {B : Obj} {Γ} {fs gs : var Γ → B ⇒ A} → (∀ i → fs i ≈ gs i) → 37 | tuple Γ fs ≈ tuple Γ gs 38 | tuple-cong ξ = unique (λ i → project ○ ⟺ (ξ i)) 39 | 40 | ∘-distribʳ-tuple : ∀ {B C} {Γ} {fs : var Γ → B ⇒ A} {g : C ⇒ B} → 41 | tuple Γ (λ x → fs x ∘ g) ≈ tuple Γ fs ∘ g 42 | ∘-distribʳ-tuple = unique (λ i → ⟺ assoc ○ ∘-resp-≈ˡ project) 43 | 44 | -- A cartesian category has a standard power structure (which we need not use) 45 | module _ (cartesian-𝒞 : Cartesian.Cartesian 𝒞) (A : Obj) where 46 | open Cartesian.Cartesian cartesian-𝒞 47 | 48 | standard-pow : Context → Obj 49 | standard-pow ctx-empty = ⊤ 50 | standard-pow ctx-slot = A 51 | standard-pow (ctx-concat Γ Δ) = standard-pow Γ × standard-pow Δ 52 | 53 | standard-π : ∀ {Γ} → var Γ → standard-pow Γ ⇒ A 54 | standard-π var-var = id 55 | standard-π (var-inl i) = standard-π i ∘ π₁ 56 | standard-π (var-inr i) = standard-π i ∘ π₂ 57 | 58 | standard-tuple : ∀ Γ {B} → (var Γ → B ⇒ A) → B ⇒ standard-pow Γ 59 | standard-tuple ctx-empty fs = ! 60 | standard-tuple ctx-slot fs = fs var-var 61 | standard-tuple (ctx-concat Γ Δ) fs = ⟨ standard-tuple Γ (λ i → fs (var-inl i)) , standard-tuple Δ (λ j → fs (var-inr j)) ⟩ 62 | 63 | standard-project : ∀ {Γ} {B} {x : var Γ} {fs : var Γ → B ⇒ A} → 64 | standard-π x ∘ standard-tuple Γ fs ≈ fs x 65 | standard-project {x = var-var} = identityˡ 66 | standard-project {x = var-inl x} = assoc ○ ((∘-resp-≈ʳ project₁) ○ standard-project {x = x}) 67 | standard-project {x = var-inr x} = assoc ○ ((∘-resp-≈ʳ project₂) ○ standard-project {x = x}) 68 | 69 | standard-unique : ∀ {Γ} {B} {fs : var Γ → B ⇒ A} {g : B ⇒ standard-pow Γ} → (∀ i → standard-π i ∘ g ≈ fs i) → 70 | standard-tuple Γ fs ≈ g 71 | standard-unique {ctx-empty} ξ = !-unique _ 72 | standard-unique {ctx-slot} ξ = ⟺ (ξ var-var) ○ identityˡ 73 | standard-unique {ctx-concat Γ Δ} {fs = fs} ξ = 74 | unique 75 | (⟺ (standard-unique (λ x → sym-assoc ○ (ξ (var-inl x))))) 76 | (⟺ (standard-unique (λ y → sym-assoc ○ (ξ (var-inr y))))) 77 | 78 | StandardPowered : Powered A 79 | StandardPowered = 80 | record 81 | { pow = standard-pow 82 | ; π = standard-π 83 | ; tuple = standard-tuple 84 | ; project = λ {Γ} → standard-project {Γ} 85 | ; unique = standard-unique } 86 | -------------------------------------------------------------------------------- /src/SingleSorted/Substitution.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lsuc; _⊔_) 2 | open import Data.Fin using (Fin) 3 | 4 | open import SingleSorted.AlgebraicTheory 5 | 6 | module SingleSorted.Substitution {ℓ} {Σ : Signature} (T : Theory ℓ Σ) where 7 | 8 | open Theory T 9 | 10 | -- an equality is preserved by the action of the identity 11 | eq-id-action : ∀ {Γ : Context} {u v : Term Γ} → (Γ ⊢ (u [ id-s ]s) ≈ (v [ id-s ]s)) → (Γ ⊢ u ≈ v) 12 | eq-id-action {u = u} {v = v} p = eq-tran (id-action {a = u}) (eq-tran p (eq-symm (id-action {a = v}))) 13 | 14 | -- equality of substitutions 15 | _≈s_ : ∀ {Γ Δ : Context} → Γ ⇒s Δ → Γ ⇒s Δ → Set (lsuc ℓ) 16 | _≈s_ {Γ = Γ} σ ρ = ∀ x → Γ ⊢ σ x ≈ ρ x 17 | 18 | infix 5 _≈s_ 19 | 20 | -- reflexivity of the equality of substitutions 21 | refl-subst : ∀ {Γ Δ : Context} {f : Γ ⇒s Δ} → f ≈s f 22 | refl-subst = λ x → eq-refl 23 | 24 | -- symmetry of the equality of substitutions 25 | symm-subst : ∀ {Γ Δ : Context} {f g : Γ ⇒s Δ} → f ≈s g → g ≈s f 26 | symm-subst {Γ} {Δ} {f} {g} p = λ x → eq-symm (p x) 27 | 28 | -- transitivity of the equality of substitutions 29 | trans-subst : ∀ {Γ Δ : Context} {f g h : Γ ⇒s Δ} → f ≈s g → g ≈s h → f ≈s h 30 | trans-subst {Γ} {Δ} {f} {g} {h} p q = λ x → eq-tran (p x) (q x) 31 | 32 | -- neutrality of tm-var 33 | tm-var-id : ∀ {Γ : Context} {x : Term Γ} → Γ ⊢ x [ id-s ]s ≈ x 34 | tm-var-id {x = tm-var x} = eq-refl 35 | tm-var-id {x = tm-oper f x} = eq-congr (λ i → tm-var-id) 36 | 37 | -- any two substitutions into the empty context are equal 38 | empty-ctx-subst-unique : ∀ {Γ : Context} {σ ρ : Γ ⇒s ctx-empty} → σ ≈s ρ 39 | empty-ctx-subst-unique () 40 | 41 | -- composition of substitutions is functorial 42 | subst-∘s : ∀ {Γ Δ Θ} {ρ : Δ ⇒s Γ} {σ : Θ ⇒s Δ} (t : Term Γ) → Θ ⊢ (t [ ρ ]s) [ σ ]s ≈ t [ ρ ∘s σ ]s 43 | subst-∘s (tm-var x) = eq-refl 44 | subst-∘s (tm-oper f ts) = eq-congr (λ i → subst-∘s (ts i)) 45 | 46 | -- substitution preserves equality 47 | eq-subst : ∀ {Γ Δ : Context} {σ : Γ ⇒s Δ} {u v : Term Δ} → Δ ⊢ u ≈ v → Γ ⊢ u [ σ ]s ≈ v [ σ ]s 48 | eq-subst eq-refl = eq-refl 49 | eq-subst (eq-symm ξ) = eq-symm (eq-subst ξ) 50 | eq-subst (eq-tran ζ ξ) = eq-tran (eq-subst ζ) (eq-subst ξ) 51 | eq-subst (eq-congr ξ) = eq-congr (λ i → eq-subst (ξ i)) 52 | eq-subst {σ = σ} (eq-axiom ε ρ) = eq-tran (subst-∘s (ax-lhs ε)) (eq-tran (eq-axiom ε (ρ ∘s σ)) (eq-symm (subst-∘s (ax-rhs ε)))) 53 | 54 | -- equivalent substitutions act the same on terms 55 | equiv-subst : ∀ {Γ Δ : Context} {σ τ : Γ ⇒s Δ} u → σ ≈s τ → Γ ⊢ u [ σ ]s ≈ u [ τ ]s 56 | equiv-subst (tm-var x) ξ = ξ x 57 | equiv-subst (tm-oper f ts) ξ = eq-congr (λ i → equiv-subst (ts i) ξ) 58 | 59 | -- equivalent substitution preserve equality 60 | equiv-eq-subst : ∀ {Γ Δ : Context} {σ τ : Γ ⇒s Δ} {u v : Term Δ} → σ ≈s τ → Δ ⊢ u ≈ v → Γ ⊢ u [ σ ]s ≈ v [ τ ]s 61 | equiv-eq-subst {u = u} p eq-refl = equiv-subst u p 62 | equiv-eq-subst p (eq-symm q) = eq-symm (equiv-eq-subst (symm-subst p) q) 63 | equiv-eq-subst p (eq-tran q r) = eq-tran (eq-subst q) (equiv-eq-subst p r) 64 | equiv-eq-subst p (eq-congr ξ) = eq-congr λ i → equiv-eq-subst p (ξ i) 65 | equiv-eq-subst {σ = σ} {τ = τ} p (eq-axiom ε θ) = eq-tran (eq-subst (eq-axiom ε θ)) (equiv-subst (ax-rhs ε [ θ ]s) p) 66 | 67 | -- the pairing of two substitutions 68 | ⟨_,_⟩s : ∀ {Γ Δ Θ} (σ : Γ ⇒s Δ) (ρ : Γ ⇒s Θ) → Γ ⇒s (ctx-concat Δ Θ) 69 | ⟨ σ , ρ ⟩s (var-inl x) = σ x 70 | ⟨ σ , ρ ⟩s (var-inr y) = ρ y 71 | 72 | -- composition of substitutions preserves equality 73 | ∘s-resp-≈s : ∀ {Γ Δ Θ} {σ₁ σ₂ : Γ ⇒s Δ} {τ₁ τ₂ : Δ ⇒s Θ} → 74 | τ₁ ≈s τ₂ → σ₁ ≈s σ₂ → τ₁ ∘s σ₁ ≈s τ₂ ∘s σ₂ 75 | ∘s-resp-≈s ξ ζ z = equiv-eq-subst ζ (ξ z) 76 | -------------------------------------------------------------------------------- /src/SingleSorted/SyntacticCategory.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (lzero; lsuc) 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | import Relation.Binary.Reasoning.Setoid as SetoidR 5 | 6 | import Categories.Category as Category 7 | import Categories.Category.Cartesian as Cartesian 8 | open import Categories.Object.Terminal using (Terminal) 9 | open import Categories.Object.Product using (Product) 10 | 11 | open import SingleSorted.AlgebraicTheory 12 | import SingleSorted.Substitution as Substitution 13 | import SingleSorted.Power as Power 14 | 15 | module SingleSorted.SyntacticCategory 16 | {ℓt} 17 | {Σ : Signature} 18 | (T : Theory ℓt Σ) where 19 | 20 | open Theory T 21 | open Substitution T 22 | 23 | -- The syntactic category 24 | 25 | 𝒮 : Category.Category lzero lzero (lsuc ℓt) 26 | 𝒮 = 27 | record 28 | { Obj = Context 29 | ; _⇒_ = _⇒s_ 30 | ; _≈_ = _≈s_ 31 | ; id = id-s 32 | ; _∘_ = _∘s_ 33 | ; assoc = λ {_ _ _ _ _ _ σ} x → subst-∘s (σ x) 34 | ; sym-assoc = λ {_ _ _ _ _ _ σ} x → eq-symm (subst-∘s (σ x)) 35 | ; identityˡ = λ x → eq-refl 36 | ; identityʳ = λ {A B f} x → tm-var-id 37 | ; identity² = λ x → eq-refl 38 | ; equiv = record 39 | { refl = λ x → eq-refl 40 | ; sym = λ ξ y → eq-symm (ξ y) 41 | ; trans = λ ζ ξ y → eq-tran (ζ y) (ξ y)} 42 | ; ∘-resp-≈ = ∘s-resp-≈s 43 | } 44 | 45 | 46 | -- We use the power structure which gives back the context directly 47 | power-𝒮 : Power.Powered 𝒮 ctx-slot 48 | power-𝒮 = 49 | record 50 | { pow = λ Γ → Γ 51 | ; π = λ x _ → tm-var x 52 | ; tuple = λ Γ {Δ} ts x → ts x var-var 53 | ; project = λ {Γ} {Δ} {x} {fs} y → ≡-⊢-≈ (cong₂ fs refl var-var-unique) 54 | ; unique = λ {Δ} {fs} {σ} {ts} ξ x → eq-symm (ξ x var-var) 55 | } 56 | where var-var-unique : ∀ {x : var ctx-slot} → var-var ≡ x 57 | var-var-unique {var-var} = refl 58 | 59 | -- The terminal object is the empty context 60 | terminal-𝒮 : Terminal 𝒮 61 | terminal-𝒮 = 62 | record 63 | { ⊤ = ctx-empty 64 | ; ⊤-is-terminal = 65 | record { ! = ctx-empty-absurd 66 | ; !-unique = λ σ x → ctx-empty-absurd x } } 67 | 68 | -- Binary product is context contatenation 69 | product-𝒮 : ∀ {Γ Δ} → Product 𝒮 Γ Δ 70 | product-𝒮 {Γ} {Δ} = 71 | record 72 | { A×B = ctx-concat Γ Δ 73 | ; π₁ = λ x → tm-var (var-inl x) 74 | ; π₂ = λ x → tm-var (var-inr x) 75 | ; ⟨_,_⟩ = ⟨_,_⟩s 76 | ; project₁ = λ x → eq-refl 77 | ; project₂ = λ x → eq-refl 78 | ; unique = λ {Θ σ σ₁ σ₂} ξ₁ ξ₂ z → u Θ σ σ₁ σ₂ ξ₁ ξ₂ z 79 | } 80 | where u : ∀ Θ (σ : Θ ⇒s (ctx-concat Γ Δ)) (σ₁ : Θ ⇒s Γ) (σ₂ : Θ ⇒s Δ) → 81 | ((λ x → σ (var-inl x)) ≈s σ₁) → ((λ y → σ (var-inr y)) ≈s σ₂) → ⟨ σ₁ , σ₂ ⟩s ≈s σ 82 | u Θ σ σ₁ σ₂ ξ₁ ξ₂ (var-inl z) = eq-symm (ξ₁ z) 83 | u Θ σ σ₁ σ₂ ξ₁ ξ₂ (var-inr z) = eq-symm (ξ₂ z) 84 | 85 | -- The cartesian structure of the syntactic category 86 | cartesian-𝒮 : Cartesian.Cartesian 𝒮 87 | cartesian-𝒮 = 88 | record 89 | { terminal = terminal-𝒮 90 | ; products = record { product = product-𝒮 } } 91 | -------------------------------------------------------------------------------- /src/SingleSorted/UniversalInterpretation.agda: -------------------------------------------------------------------------------- 1 | open import SingleSorted.AlgebraicTheory 2 | 3 | import SingleSorted.Interpretation as Interpretation 4 | import SingleSorted.SyntacticCategory as SyntacticCategory 5 | import SingleSorted.Substitution as Substitution 6 | 7 | 8 | module SingleSorted.UniversalInterpretation 9 | {ℓt} 10 | {Σ : Signature} 11 | (T : Theory ℓt Σ) where 12 | 13 | open Theory T 14 | open Substitution T 15 | open SyntacticCategory T 16 | 17 | -- The universal interpretation in the syntactic category 18 | ℐ : Interpretation.Interpretation Σ cartesian-𝒮 19 | ℐ = 20 | record 21 | { interp-carrier = ctx-slot 22 | ; interp-pow = power-𝒮 23 | ; interp-oper = λ f var-var → tm-oper f (λ i → tm-var i) 24 | } 25 | 26 | open Interpretation.Interpretation ℐ 27 | 28 | -- A term is essentially interpreted by itself 29 | interp-term-self : ∀ {Γ} (t : Term Γ) y → Γ ⊢ interp-term t y ≈ t 30 | interp-term-self (tm-var x) _ = eq-refl 31 | interp-term-self (tm-oper f xs) y = eq-congr (λ i → interp-term-self (xs i) var-var) 32 | -------------------------------------------------------------------------------- /src/SingleSorted/UniversalModel.agda: -------------------------------------------------------------------------------- 1 | import Relation.Binary.Reasoning.Setoid as SetoidR 2 | open import SingleSorted.AlgebraicTheory 3 | 4 | import SingleSorted.Interpretation as Interpretation 5 | import SingleSorted.Model as Model 6 | import SingleSorted.UniversalInterpretation as UniversalInterpretation 7 | import SingleSorted.Substitution as Substitution 8 | import SingleSorted.SyntacticCategory as SyntacticCategory 9 | 10 | module SingleSorted.UniversalModel 11 | {ℓt} 12 | {Σ : Signature} 13 | (T : Theory ℓt Σ) where 14 | 15 | open Theory T 16 | open Substitution T 17 | open UniversalInterpretation T 18 | open Interpretation.Interpretation ℐ 19 | open SyntacticCategory T 20 | 21 | 𝒰 : Model.Model T cartesian-𝒮 ℐ 22 | 𝒰 = 23 | record 24 | { model-eq = λ ε var-var → 25 | let open SetoidR (eq-setoid (ax-ctx ε)) in 26 | begin 27 | interp-term (ax-lhs ε) var-var ≈⟨ interp-term-self (ax-lhs ε) var-var ⟩ 28 | ax-lhs ε ≈⟨ id-action ⟩ 29 | ax-lhs ε [ id-s ]s ≈⟨ eq-axiom ε id-s ⟩ 30 | ax-rhs ε [ id-s ]s ≈˘⟨ id-action ⟩ 31 | ax-rhs ε ≈˘⟨ interp-term-self (ax-rhs ε) var-var ⟩ 32 | interp-term (ax-rhs ε) var-var ∎ } 33 | --------------------------------------------------------------------------------