├── readme.md └── src └── STLC ├── Context.agda ├── Context ├── Core.agda └── Properties.agda ├── Experimental ├── BadIdea.agda ├── Safe.agda └── Unsafe.agda ├── Lib ├── MaybeElim.agda ├── Membership.agda ├── NoEtaProduct.agda ├── Prelude.agda └── Sets.agda ├── M ├── Main.agda ├── Term.agda └── Typecheck.agda ├── Main.agda ├── NbE ├── Main.agda ├── NF.agda └── Read.agda ├── Semantics ├── Core.agda └── Eval.agda ├── Term.agda ├── Term ├── Core.agda └── Syntax.agda ├── Tests ├── Church.agda └── Combs.agda ├── Type.agda └── Type ├── Core.agda ├── Properties.agda └── Unify.agda /readme.md: -------------------------------------------------------------------------------- 1 | # STLC-in-Agda 2 | 3 | ## A quick taste 4 | 5 | Getting a universe polymorphic Agda function from a pure lambda term: 6 | 7 | ``` 8 | phoenix : Syntax⁽⁾ 9 | phoenix = 4 # λ a b c d → a · (b · d) · (c · d) 10 | 11 | phoenixᵗ : Term ((b ⇒ c ⇒ d) ⇒ (a ⇒ b) ⇒ (a ⇒ c) ⇒ a ⇒ d) 12 | phoenixᵗ = term⁻ phoenix 13 | 14 | liftM2 : ∀ {α β γ δ} {A : Set α} {B : Set β} {C : Set γ} {D : Set δ} 15 | -> ((B -> C -> D) -> (A -> B) -> (A -> C) -> A -> D) 16 | liftM2 = eval phoenixᵗ 17 | ``` 18 | 19 | `C-c C-n liftM2` gives `λ {.α} {.β} {.γ} {.δ} {.A} {.B} {.C} {.D} x x₁ x₂ x₃ → 20 | x (x₁ x₃) (x₂ x₃)`. 21 | 22 | ## Overview 23 | 24 | This is simply typed lambda calculus with type variables in Agda. We have raw [Syntax](src/STLC/Term/Syntax.agda): 25 | 26 | ``` 27 | data Syntax n : Set where 28 | var : Fin n -> Syntax n 29 | ƛ_ : Syntax (suc n) -> Syntax n 30 | _·_ : Syntax n -> Syntax n -> Syntax n 31 | ``` 32 | 33 | [typed terms](src/STLC/Term/Core.agda): 34 | 35 | ``` 36 | data _⊢_ {n} Γ : Type n -> Set where 37 | var : ∀ {σ} -> σ ∈ Γ -> Γ ⊢ σ 38 | ƛ_ : ∀ {σ τ} -> Γ ▻ σ ⊢ τ -> Γ ⊢ σ ⇒ τ 39 | _·_ : ∀ {σ τ} -> Γ ⊢ σ ⇒ τ -> Γ ⊢ σ -> Γ ⊢ τ 40 | ``` 41 | 42 | and [a mapping](src/STLC/M/Main.agda) from the former to the latter via a type-safe version of algorithm M: 43 | 44 | ``` 45 | M : ∀ {n l} -> (Γ : Con n l) -> Syntax l -> (σ : Type n) 46 | -> Maybe (∃ λ m -> ∃ λ (Ψ : Subst n m) -> mapᶜ (apply Ψ) Γ ⊢ apply Ψ σ) 47 | ``` 48 | 49 | `M` receives a context, a term and a type, and checks, whether there is a substitution that allows to typify the term in this context and with this type, after the substitution is applied to them. `M` uses rewrite rules under the hood — this simplifies the definition a lot. 50 | 51 | There is an [NbE](src/STLC/NbE/Main.agda), which uses the traversal from [5]. 52 | 53 | There is [a part](src/STLC/NbE/Read.agda) of the liftable terms approach to NbE (described in [4]), which is used to coerce Agda's lambda terms to their first-order counterparts. 54 | 55 | There is [a universe polymorphic eval](src/STLC/Semantics/Eval.agda). 56 | 57 | Using all this we can, for example, make an Agda lambda term universe polymorphic: 58 | 59 | ``` 60 | mono-app : {A B : Set} -> (A -> B) -> A -> B 61 | mono-app f x = f x 62 | 63 | poly-app : ∀ {α β} {A : Set α} {B : Set β} -> (A -> B) -> A -> B 64 | poly-app = eval (read (inst 2 λ A B -> mono-app {A} {B})) 65 | ``` 66 | 67 | ## Details 68 | 69 | I turned off the eta rule for products, because Agda lacks sharing. My Agda is out of date a bit, so I redefined `Σ` as a `data` instead of placing `{-# NO_ETA Σ #-}` somewhere. Algorithm M is still incredibly slow, and the current `_>>=ᵀ_` eats twice the memory comparing to the previous version, which was more performant, but less usable and inference-friendly, which is important because types signatures are quite verbose. 70 | 71 | The [Main](src/STLC/Main.agda) module contains 72 | 73 | ``` 74 | on-typed : ∀ {α} {A : ∀ {n} {σ : Type n} -> Term⁽⁾ σ -> Set α} 75 | -> (f : ∀ {n} {σ : Type n} -> (t : Term⁽⁾ σ) -> A t) -> ∀ e -> _ 76 | on-typed f e = fromJustᵗ $ infer e >>=ᵗ f ∘ thicken ∘ core ∘ proj₂ ∘ proj₂ 77 | 78 | typed = on-typed $ id 79 | term = on-typed $ λ t {m Δ} -> generalize {m} Δ t 80 | term⁻ = on-typed $ λ {n} t {Δ} -> generalize {n} Δ t 81 | normᵖ = on-typed $ pure ∘ erase ∘ norm 82 | ``` 83 | 84 | `on-typed` receives a function `f` and a term, tries to typify the term, makes type variables consecutive and strengthened, applies `f` to the result and removes `just`, when inferring is successful, or returns `lift tt` otherwise. 85 | 86 | `_>>=ᵗ_` constructs values of type `mx >>=ᵀ B`, which are "either `nothing` or `B x`". The idea is described [here](http://stackoverflow.com/questions/31105947/eliminating-a-maybe-at-the-type-level) (I changed the definition a bit though). 87 | 88 | `thicken` is defined in terms of `enumerate` described [here](http://stackoverflow.com/questions/33345899/how-to-enumerate-the-elements-of-a-list-by-fins-in-linear-time). 89 | 90 | `generalize` substitutes type variables with universally quantified types with universally quantified upper bounds for variables and prepends a universally quantified context: 91 | 92 | ``` 93 | app : ε {2} ⊢ (Var zero ⇒ Var (suc zero)) ⇒ Var zero ⇒ Var (suc zero) 94 | app = ƛ ƛ var (vs vz) · var vz 95 | 96 | gapp : ∀ {n σ τ} {Γ : Con n} -> Γ ⊢ (σ ⇒ τ) ⇒ σ ⇒ τ 97 | gapp = generalize _ app 98 | ``` 99 | 100 | `normᵖ` normalizes a pure lambda term whenever it's typeable. Uses NbE under the hood. 101 | 102 | ## Unsafeties 103 | 104 | Two simple lemmas are used for for rewriting via the `REWRITE` pragma: 105 | 106 | ```agda 107 | apply-apply : ∀ {n m p} {Φ : Subst m p} {Ψ : Subst n m} σ 108 | -> apply Φ (apply Ψ σ) ≡ apply (Φ ∘ˢ Ψ) σ 109 | apply-apply (Var i) = refl 110 | apply-apply (σ ⇒ τ) = cong₂ _⇒_ (apply-apply σ) (apply-apply τ) 111 | 112 | mapᶜ-mapᶜ : ∀ {n m p l} {g : Type m -> Type p} {f : Type n -> Type m} (Γ : Con n l) 113 | -> mapᶜ g (mapᶜ f Γ) ≡ mapᶜ (g ∘ f) Γ 114 | mapᶜ-mapᶜ ε = refl 115 | mapᶜ-mapᶜ (Γ ▻ σ) = cong (_▻ _) (mapᶜ-mapᶜ Γ) 116 | ``` 117 | 118 | In one of the previous version I avoided rewriting, but the code was unreadable. 119 | 120 | Unification is postulated to be `TERMINATING`: 121 | 122 | ```agda 123 | {-# TERMINATING #-} 124 | unify : ∀ {n} -> (σ τ : Type n) -> Maybe (∃ λ (Ψ : Subst n n) -> apply Ψ σ ≡ apply Ψ τ) 125 | ``` 126 | 127 | It can be proved so using the techniques from [6]. 128 | 129 | There are highly unsafe things in the [STLC.Experimental.Unsafe](src/STLC/Experimental/Unsafe.agda) module, but they are not used in the Algotihm M itself, only to define `on-typed` and its derivatives presented above. 130 | 131 | ## References 132 | 133 | 1. Martin Grabmüller. [Algorithm W Step by Step](https://github.com/wh5a/Algorithm-W-Step-By-Step) 134 | 2. Oukseh Lee, Kwangkeun Yi. [A Generalized Let-Polymorphic Type Inference Algorithm](http://citeseerx.ist.psu.edu/viewdoc/summary?doi=10.1.1.41.6832) 135 | 3. Dr. Gergő Érdi. [Compositional Type Checking](http://gergo.erdi.hu/projects/tandoori/Tandoori-Compositional-Typeclass.pdf) 136 | 4. Andreas Abel. [Normalization by Evaluation: Dependent Types and Impredicativity](http://www2.tcs.ifi.lmu.de/~abel/habil.pdf) 137 | 5. Guillaume Allais. [Type and Scope Preserving Semantics](https://github.com/gallais/type-scope-semantics) 138 | 6. Conor McBride. [First-Order Unification by Structural Recursion](http://citeseerx.ist.psu.edu/viewdoc/download;jsessionid=3E26A845A6124F33E00E24E3D1C6036C?doi=10.1.1.25.1516&rep=rep1&type=pdf) -------------------------------------------------------------------------------- /src/STLC/Context.agda: -------------------------------------------------------------------------------- 1 | module STLC.Context where 2 | 3 | open import STLC.Type public 4 | open import STLC.Context.Core public 5 | open import STLC.Context.Properties public 6 | -------------------------------------------------------------------------------- /src/STLC/Context/Core.agda: -------------------------------------------------------------------------------- 1 | module STLC.Context.Core where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Type 5 | 6 | infixl 5 _▻_ _▻▻_ 7 | infix 4 _⊂[_]_ _⊆_ 8 | 9 | data Con n : Set where 10 | ε : Con n 11 | _▻_ : Con n -> Type n -> Con n 12 | 13 | data _⊂[_]_ {n} Γ σ : Con n -> Set where 14 | vtop : Γ ⊂[ σ ] Γ ▻ σ 15 | vskip : ∀ {Δ τ} -> Γ ⊂[ σ ] Δ -> Γ ⊂[ σ ] Δ ▻ τ 16 | 17 | data _⊆_ {n} : Con n -> Con n -> Set where 18 | stop : ∀ {Γ} -> Γ ⊆ Γ 19 | skip : ∀ {Γ Δ σ} -> Γ ⊆ Δ -> Γ ⊆ Δ ▻ σ 20 | keep : ∀ {Γ Δ σ} -> Γ ⊆ Δ -> Γ ▻ σ ⊆ Δ ▻ σ 21 | 22 | Links : Set₁ 23 | Links = ∀ {n} -> Con n -> Type n -> Set 24 | 25 | Renames : Links -> Set 26 | Renames _∙_ = ∀ {n σ} {Γ Δ : Con n} -> Γ ⊆ Δ -> Γ ∙ σ -> Δ ∙ σ 27 | 28 | _∸>_ : Links -> Links -> Set 29 | _∙_ ∸> _◆_ = ∀ {n σ} {Γ : Con n} -> Γ ∙ σ -> Γ ◆ σ 30 | 31 | foldrᶜ : ∀ {α n} {A : Set α} -> (Type n -> A -> A) -> A -> Con n -> A 32 | foldrᶜ f z ε = z 33 | foldrᶜ f z (Γ ▻ σ) = f σ (foldrᶜ f z Γ) 34 | 35 | mapᶜ : ∀ {n m} -> (Type n -> Type m) -> Con n -> Con m 36 | mapᶜ f = foldrᶜ (λ σ Γ -> Γ ▻ f σ) ε 37 | 38 | lenᶜ : ∀ {n} -> Con n -> ℕ 39 | lenᶜ = foldrᶜ (const suc) 0 40 | 41 | _▻▻_ : ∀ {n} -> Con n -> Con n -> Con n 42 | _▻▻_ = foldrᶜ (flip _▻_) 43 | 44 | top : ∀ {n σ} {Γ : Con n} -> Γ ⊆ Γ ▻ σ 45 | top = skip stop 46 | 47 | full : ∀ {n} {Γ : Con n} -> ε ⊆ Γ 48 | full {Γ = ε} = stop 49 | full {Γ = Γ ▻ σ} = skip full 50 | 51 | _∘ᵉ_ : ∀ {n} {Γ Δ Ξ : Con n} -> Δ ⊆ Ξ -> Γ ⊆ Δ -> Γ ⊆ Ξ 52 | stop ∘ᵉ ι = ι 53 | skip κ ∘ᵉ ι = skip (κ ∘ᵉ ι) 54 | keep κ ∘ᵉ stop = keep κ 55 | keep κ ∘ᵉ skip ι = skip (κ ∘ᵉ ι) 56 | keep κ ∘ᵉ keep ι = keep (κ ∘ᵉ ι) 57 | -------------------------------------------------------------------------------- /src/STLC/Context/Properties.agda: -------------------------------------------------------------------------------- 1 | module STLC.Context.Properties where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Type 5 | open import STLC.Context.Core 6 | 7 | ▻-inj : ∀ {n σ₁ σ₂} {Γ₁ Γ₂ : Con n} -> Γ₁ ▻ σ₁ ≡ Γ₂ ▻ σ₂ -> Γ₁ ≡ Γ₂ × σ₁ ≡ σ₂ 8 | ▻-inj refl = refl , refl 9 | 10 | instance 11 | Con-DecEq : ∀ {n} -> DecEq (Con n) 12 | Con-DecEq = record { _≟_ = _≟ᶜ_ } where 13 | infix 4 _≟ᶜ_ 14 | _≟ᶜ_ : ∀ {n} -> Decidable (_≡_ {A = Con n}) 15 | ε ≟ᶜ ε = yes refl 16 | Γ ▻ σ ≟ᶜ Δ ▻ τ = dcong₂ _▻_ ▻-inj (Γ ≟ᶜ Δ) (σ ≟ τ) 17 | ε ≟ᶜ Δ ▻ τ = no λ() 18 | Γ ▻ σ ≟ᶜ ε = no λ() 19 | 20 | foldrᶜ-mapᶜ : ∀ {α n m} {A : Set α} {g : Type m -> A -> A} {f : Type n -> Type m} {z} Γ 21 | -> foldrᶜ g z (mapᶜ f Γ) ≡ foldrᶜ (g ∘ f) z Γ 22 | foldrᶜ-mapᶜ ε = refl 23 | foldrᶜ-mapᶜ {g = g} {f} (Γ ▻ σ) = cong (g (f σ)) (foldrᶜ-mapᶜ Γ) 24 | 25 | ⊂-▻ : ∀ {n σ τ} {Γ Δ : Con n} -> Γ ⊂[ σ ] Δ ▻ τ -> Γ ⊂[ σ ] Δ ⊎ Γ ≡ Δ × σ ≡ τ 26 | ⊂-▻ vtop = inj₂ (refl , refl) 27 | ⊂-▻ (vskip p) = inj₁ p 28 | 29 | _⊂?_ : ∀ {n} {σ : Type n} -> Decidable _⊂[ σ ]_ 30 | _⊂?_ Γ ε = no λ() 31 | _⊂?_ {σ = σ} Γ (Δ ▻ τ) = 32 | ddprod (yes % ∘ vtop′) (λ c -> r (c ∘ proj₂)) (λ c -> r (c ∘ proj₁)) (σ ≟ τ) (Γ ≟ Δ) where 33 | r = λ c₁ -> drec (yes ∘ vskip) (λ c₂ -> no ([ c₂ , c₁ ] ∘ ⊂-▻)) (Γ ⊂? Δ) 34 | 35 | vtop′ : ∀ {n σ τ} {Γ Δ : Con n} -> σ ≡ τ -> Γ ≡ Δ -> Γ ⊂[ σ ] Δ ▻ τ 36 | vtop′ refl refl = vtop 37 | -------------------------------------------------------------------------------- /src/STLC/Experimental/BadIdea.agda: -------------------------------------------------------------------------------- 1 | module STLC.Experimental.BadIdea where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Context 5 | open import STLC.Term 6 | 7 | data SubstsIn {n} : Con n -> ℕ -> Set where 8 | øˢ : ∀ {m} -> SubstsIn ε m 9 | _▷ˢ_ : ∀ {m Γ σ} -> SubstsIn Γ m -> SubstIn σ m -> SubstsIn (Γ ▻ σ) m 10 | 11 | runSubstsIn : ∀ {n m} {Γ : Con n} -> SubstsIn Γ m -> Con m 12 | runSubstsIn øˢ = ε 13 | runSubstsIn (ρ ▷ˢ Ψ) = runSubstsIn ρ ▻ runSubstIn Ψ 14 | 15 | -- Yeah. 16 | specialize⁽⁾ : ∀ {n m σ} {Γ : Con n} 17 | -> (ρ : SubstsIn Γ m) -> (Ψ : SubstIn σ m) -> Γ ⊢ σ -> runSubstsIn ρ ⊢ runSubstIn Ψ 18 | specialize⁽⁾ ρ Ψ (var v) = {!!} 19 | specialize⁽⁾ ρ Ψ (ƛ b) = ƛ (specialize⁽⁾ (ρ ▷ˢ (Ψ ∘ vl)) (Ψ ∘ vr) b) 20 | specialize⁽⁾ ρ Ψ (f · x) = specialize⁽⁾ ρ {!!} f · specialize⁽⁾ ρ {!!} x 21 | -------------------------------------------------------------------------------- /src/STLC/Experimental/Safe.agda: -------------------------------------------------------------------------------- 1 | module STLC.Experimental.Safe where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Lib.Membership 5 | open import STLC.Type.Core 6 | open import STLC.Type.Properties 7 | 8 | infix 4 _∈ᵗ_ 9 | 10 | data _∈ᵗ_ {n} i : Type n -> Set where 11 | vz : i ∈ᵗ Var i 12 | vl : ∀ {σ τ} -> i ∈ᵗ σ -> i ∈ᵗ σ ⇒ τ 13 | vr : ∀ {σ τ} -> i ∈ᵗ τ -> i ∈ᵗ σ ⇒ τ 14 | 15 | SubstIn : ∀ {n} -> Type n -> ℕ -> Set 16 | SubstIn σ m = ∀ {i} -> i ∈ᵗ σ -> Type m 17 | 18 | runSubstIn : ∀ {n m} {σ : Type n} -> SubstIn σ m -> Type m 19 | runSubstIn {σ = Var i} Ψ = Ψ vz 20 | runSubstIn {σ = σ ⇒ τ} Ψ = runSubstIn (Ψ ∘ vl) ⇒ runSubstIn (Ψ ∘ vr) 21 | 22 | ∈ᵗ-Var : ∀ {n} {i j : Fin n} -> i ∈ᵗ Var j -> i ≡ j 23 | ∈ᵗ-Var vz = refl 24 | 25 | ∈ᵗ-⇒ : ∀ {n i} {σ τ : Type n} -> i ∈ᵗ σ ⇒ τ -> i ∈ᵗ σ ⊎ i ∈ᵗ τ 26 | ∈ᵗ-⇒ (vl v) = inj₁ v 27 | ∈ᵗ-⇒ (vr v) = inj₂ v 28 | 29 | _∈ᵗ?_ : ∀ {n} -> (i : Fin n) -> (σ : Type n) -> Dec (i ∈ᵗ σ) 30 | i ∈ᵗ? Var j = dJ (λ i j -> i ∈ᵗ Var j) (yes vz) (λ c -> no (c ∘ ∈ᵗ-Var)) (i ≟ j) 31 | i ∈ᵗ? (σ ⇒ τ) = dsum vl vr (λ c₁ c₂ -> [ c₁ , c₂ ]′ ∘ ∈ᵗ-⇒) (i ∈ᵗ? σ) (i ∈ᵗ? τ) 32 | 33 | ∈ᵗ-ftv-all : ∀ {n i} {σ : Type n} -> i ∈ᵗ σ -> i ∈ ftv-all σ 34 | ∈ᵗ-ftv-all vz = here 35 | ∈ᵗ-ftv-all (vl v) = ∈-++₁ (∈ᵗ-ftv-all v) 36 | ∈ᵗ-ftv-all {σ = σ ⇒ τ} (vr v) = ∈-++₂ (ftv-all σ) (∈ᵗ-ftv-all v) 37 | 38 | ∈ᵗ-ftv : ∀ {n i} {σ : Type n} -> i ∈ᵗ σ -> i ∈ ftv σ 39 | ∈ᵗ-ftv = ∈-nub ∘ ∈ᵗ-ftv-all 40 | 41 | thickenⱽ≢nothing : ∀ {n i} {σ : Type n} -> i ∈ᵗ σ -> thickenⱽ i σ ≢ nothing 42 | thickenⱽ≢nothing {σ = σ} v p with λ q -> lookup-for≢nothing (map swap (enumerate (ftv σ))) q p 43 | ... | r rewrite sym (map-compose {g = proj₁} {swap} (enumerate (ftv σ))) 44 | | map-cong {f = proj₁ ∘ swap} {proj₂} proj₁-swap (enumerate (ftv σ)) 45 | | enumerated (ftv σ) = r (∈ᵗ-ftv v) 46 | 47 | fromInFtv : ∀ {n m} -> (σ : Type n) -> SubstInFtv σ m -> SubstIn σ m 48 | fromInFtv _ Ψ = Ψ _ ∘ ∈ᵗ-ftv 49 | 50 | thickenˢ : ∀ {n} -> (σ : Type n) -> SubstIn σ (length (ftv σ)) 51 | thickenˢ σ {i} v = inspectMaybe (thickenⱽ i σ) (λ i _ -> Var i) (⊥-elim ∘ thickenⱽ≢nothing v) 52 | -------------------------------------------------------------------------------- /src/STLC/Experimental/Unsafe.agda: -------------------------------------------------------------------------------- 1 | module STLC.Experimental.Unsafe where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Type.Core 5 | open import STLC.Context.Core 6 | open import STLC.Term.Core 7 | open import STLC.Experimental.Safe 8 | 9 | module _ where 10 | private postulate undefined : ∀ {α} {A : Set α} -> A 11 | 12 | UnsafeAssociate : ∀ {α β} {A : Set α} {B : Set β} {{_ : DecEq A}} 13 | -> List A -> ((A -> B) -> Set β) -> Set β 14 | UnsafeAssociate [] C = C undefined 15 | UnsafeAssociate (x ∷ xs) C = ∀ {y} -> 16 | UnsafeAssociate xs λ f -> C λ x' -> if x == x' then y else f x' 17 | 18 | unsafeAssociate : ∀ {α β} {A : Set α} {B : Set β} {{_ : DecEq A}} 19 | {C : (A -> B) -> Set β} 20 | -> ∀ xs -> (∀ f -> C f) -> UnsafeAssociate xs C 21 | unsafeAssociate [] c = c undefined 22 | unsafeAssociate (x ∷ xs) c {y} = unsafeAssociate xs λ f -> c λ x' -> if x == x' then y else f x' 23 | 24 | unsafeToSubst : ∀ {n m} {σ : Type n} -> SubstIn σ m -> Subst n m 25 | unsafeToSubst {σ = σ} Ψ = λ i -> drec Ψ undefined (i ∈ᵗ? σ) 26 | 27 | thicken : ∀ {n σ} {Γ : Con n} -> Γ ⊢ σ -> _ 28 | thicken {σ = σ} = specialize (unsafeToSubst (thickenˢ σ)) 29 | 30 | generalize : ∀ {m n} {σ : Type n} 31 | -> (Γ : Con m) -> Term⁽⁾ σ -> UnsafeAssociate (ftv σ) λ (Ψ : Subst n m) -> Γ ⊢ apply Ψ σ 32 | generalize {σ = σ} _ t = unsafeAssociate (ftv σ) (wk ∘ flip specialize t) 33 | 34 | -- -- Just a bit safer, but way more inefficient. 35 | -- generalize : ∀ {m n} {σ : Type n} 36 | -- -> (Γ : Con m) -> Term⁽⁾ σ -> 37 | -- Associate (ftv σ) λ (Ψ : SubstInFtv σ m) -> 38 | -- Γ ⊢ apply (unsafeToSubst (fromInFtv σ Ψ)) σ 39 | -- generalize {σ = σ} _ t = associate (ftv σ) λ Ψ -> wk (specialize (unsafeToSubst (fromInFtv σ Ψ)) t) 40 | -------------------------------------------------------------------------------- /src/STLC/Lib/MaybeElim.agda: -------------------------------------------------------------------------------- 1 | module STLC.Lib.MaybeElim where 2 | 3 | open import Level 4 | open import Function 5 | open import Relation.Binary.PropositionalEquality 6 | open import Data.Unit.Base 7 | open import Data.Bool.Base 8 | open import Data.Maybe.Base 9 | 10 | infixl 1 _>>=ᵀ_ _>>=ᵗ_ _>>=⊤_ 11 | infixl 4 _<$>ᵗ_ 12 | infixr 1 _>=>ᵗ_ 13 | infixr 10 _<∘>ᵗ_ 14 | 15 | data _>>=ᵀ_ {α β} {A : Set α} : (mx : Maybe A) -> (∀ x -> mx ≡ just x -> Set β) -> Set β where 16 | nothingᵗ : ∀ {B} -> nothing >>=ᵀ B 17 | justᵗ : ∀ {x B} -> B x refl -> just x >>=ᵀ B 18 | 19 | -- We could write (Set (maybe (const β) zero mx)), 20 | -- but I don't want to introduce dependency at the type level. 21 | FromJustᵗ : ∀ {α β} {A : Set α} {mx : Maybe A} {B : ∀ x -> mx ≡ just x -> Set β} 22 | -> mx >>=ᵀ B -> Set β 23 | FromJustᵗ nothingᵗ = Lift ⊤ 24 | FromJustᵗ (justᵗ {x} {B} y) = B x refl 25 | 26 | fromJustᵗ : ∀ {α β} {A : Set α} {mx : Maybe A} {B : ∀ x -> mx ≡ just x -> Set β} 27 | -> (yᵗ : mx >>=ᵀ B) -> FromJustᵗ yᵗ 28 | fromJustᵗ nothingᵗ = _ 29 | fromJustᵗ (justᵗ y) = y 30 | 31 | _>>=ᵗ_ : ∀ {α β} {A : Set α} {B : A -> Set β} 32 | -> (mx : Maybe A) -> (∀ x -> B x) -> mx >>=ᵀ λ x _ -> B x 33 | nothing >>=ᵗ f = nothingᵗ 34 | just x >>=ᵗ f = justᵗ (f x) 35 | 36 | _>>=⊤_ : ∀ {α β} {A : Set α} {B : A -> Set β} 37 | -> (mx : Maybe A) -> (g : ∀ x -> B x) -> FromJustᵗ (mx >>=ᵗ g) 38 | mx >>=⊤ g = fromJustᵗ $ mx >>=ᵗ g 39 | 40 | runᵗ : ∀ {α β} {A : Set α} {mx : Maybe A} {B : ∀ x -> mx ≡ just x -> Set β} {x} 41 | -> mx >>=ᵀ B -> (p : mx ≡ just x) -> B x p 42 | runᵗ (justᵗ y) refl = y 43 | 44 | _<$>ᵗ_ : ∀ {α β γ} {A : Set α} {mx : Maybe A} 45 | {B : ∀ x -> mx ≡ just x -> Set β} 46 | {C : ∀ {x} {p : mx ≡ just x} -> B x p -> Set γ} 47 | -> (∀ {x} {p : mx ≡ just x} -> (y : B x p) -> C y) 48 | -> (yᵗ : mx >>=ᵀ B) 49 | -> mx >>=ᵀ λ _ -> C ∘ runᵗ yᵗ 50 | g <$>ᵗ nothingᵗ = nothingᵗ 51 | g <$>ᵗ justᵗ y = justᵗ (g y) 52 | 53 | _>=>ᵗ_ : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : ∀ {x} -> B x -> Set γ} 54 | -> (f : ∀ x -> Maybe (B x)) 55 | -> (∀ {x} -> (y : B x) -> C y) 56 | -> ∀ x -> f x >>=ᵀ λ y _ -> C y 57 | (f >=>ᵗ g) x = f x >>=ᵗ g 58 | 59 | _<∘>ᵗ_ : ∀ {α β γ δ} {A : Set α} {B : A -> Set β} {f : ∀ x -> Maybe (B x)} 60 | {C : ∀ x y -> f x ≡ just y -> Set γ} 61 | {D : ∀ {x y} {p : f x ≡ just y} -> C x y p -> Set δ} 62 | -> (∀ {x y} {p : f x ≡ just y} -> (z : C x y p) -> D z) 63 | -> (g : ∀ x -> f x >>=ᵀ C x) 64 | -> ∀ x -> f x >>=ᵀ λ _ -> D ∘ runᵗ (g x) 65 | (h <∘>ᵗ g) x = h <$>ᵗ g x 66 | -------------------------------------------------------------------------------- /src/STLC/Lib/Membership.agda: -------------------------------------------------------------------------------- 1 | module STLC.Lib.Membership where 2 | 3 | open import STLC.Lib.Prelude 4 | 5 | infix 4 _∈_ _∉_ 6 | 7 | data _∈_ {α} {A : Set α} (x : A) : List A -> Set α where 8 | here : ∀ {xs} -> x ∈ x ∷ xs 9 | there : ∀ {y xs} -> x ∈ xs -> x ∈ y ∷ xs 10 | 11 | split-∈ : ∀ {α β} {A : Set α} {B : A -> Set β} {xs : List A} {x y} 12 | -> (∀ x -> x ∈ xs -> B x) -> B y -> x ∈ y ∷ xs -> B x 13 | split-∈ f z here = z 14 | split-∈ f z (there p) = f _ p 15 | 16 | _∉_ : ∀ {α} {A : Set α} -> A -> List A -> Set α 17 | x ∉ xs = ¬ (x ∈ xs) 18 | 19 | _∘∉_ : ∀ {α} {A : Set α} {xs : List A} {x y} -> x ≢ y -> x ∉ xs -> x ∉ y ∷ xs 20 | _∘∉_ p q here = p refl 21 | _∘∉_ p q (there r) = q r 22 | 23 | ∈-++₁ : ∀ {α} {A : Set α} {xs ys : List A} {x} 24 | -> x ∈ xs -> x ∈ xs ++ ys 25 | ∈-++₁ here = here 26 | ∈-++₁ (there p) = there (∈-++₁ p) 27 | 28 | ∈-++₂ : ∀ {α} {A : Set α} {ys : List A} {x} 29 | -> (xs : List A) -> x ∈ ys -> x ∈ xs ++ ys 30 | ∈-++₂ [] p = p 31 | ∈-++₂ (x ∷ xs) p = there (∈-++₂ xs p) 32 | 33 | _∈?_ : ∀ {α} {A : Set α} {{_ : DecEq A}} -> Decidable (_∈_ {A = A}) 34 | y ∈? [] = no λ() 35 | y ∈? (x ∷ xs) = dJ (λ y x -> y ∈ x ∷ xs) 36 | (yes here) 37 | (λ c -> dmap there (c ∘∉_) (y ∈? xs)) 38 | (y ≟ x) 39 | 40 | ≢-∈-delete : ∀ {α} {A : Set α} {{_ : DecEq A}} {x y} {xs : List A} 41 | -> x ≢ y -> x ∈ xs -> x ∈ delete y xs 42 | ≢-∈-delete {y = y} {x ∷ xs} c here with x ≟ y 43 | ... | yes p = ⊥-elim (c p) 44 | ... | no d with deletem y xs 45 | ... | nothing = here 46 | ... | just _ = here 47 | ≢-∈-delete {y = y} {z ∷ xs} c (there p) with z ≟ y 48 | ... | yes _ = p 49 | ... | no _ with ≢-∈-delete c p 50 | ... | r with deletem y xs 51 | ... | nothing = there r 52 | ... | just _ = there r 53 | 54 | ∈-nub : ∀ {α} {A : Set α} {{_ : DecEq A}} {x} {xs : List A} -> x ∈ xs -> x ∈ nub xs 55 | ∈-nub here = here 56 | ∈-nub {x = x} (there {y} p) with x ≟ y 57 | ... | yes q rewrite q = here 58 | ... | no c = there (≢-∈-delete c (∈-nub p)) 59 | 60 | lookup-for≢nothing : ∀ {α β} {A : Set α} {B : Set β} {{_ : DecEq A}} {x} 61 | -> (xs : List (A × B)) -> x ∈ map proj₁ xs -> lookup-for x xs ≢ nothing 62 | lookup-for≢nothing [] () 63 | lookup-for≢nothing ((x , z) ∷ xs) here q rewrite ≟-refl x = case q of λ() 64 | lookup-for≢nothing {x = x} ((y , z) ∷ xs) (there p) q with y ≟ x 65 | ... | yes _ = case q of λ() 66 | ... | no c = lookup-for≢nothing xs p q 67 | 68 | Associate : ∀ {α β} {A : Set α} {B : A -> Set β} 69 | -> (xs : List A) -> ((∀ x -> x ∈ xs -> B x) -> Set β) -> Set β 70 | Associate [] C = C (λ _ ()) 71 | Associate (x ∷ xs) C = ∀ {y} -> Associate xs λ f -> C λ _ -> split-∈ f y 72 | 73 | associate : ∀ {α β} {A : Set α} {B : A -> Set β} 74 | -> ∀ xs -> {C : (∀ x -> x ∈ xs -> B x) -> Set β} -> (∀ f -> C f) -> Associate xs C 75 | associate [] c = c (λ _ ()) 76 | associate (x ∷ xs) c {y} = associate xs λ f -> c λ _ -> split-∈ f y 77 | -------------------------------------------------------------------------------- /src/STLC/Lib/NoEtaProduct.agda: -------------------------------------------------------------------------------- 1 | module STLC.Lib.NoEtaProduct where 2 | 3 | open import Function 4 | open import Level 5 | open import Relation.Nullary 6 | 7 | infixr 4 _,_ _,′_ 8 | infix 4 ,_ 9 | infixr 2 _×_ _-×-_ _-,-_ 10 | 11 | ------------------------------------------------------------------------ 12 | -- Definition 13 | 14 | data Σ {α β} (A : Set α) (B : A -> Set β) : Set (α ⊔ β) where 15 | _,_ : ∀ x -> B x -> Σ A B 16 | 17 | proj₁ : ∀ {α β} {A : Set α} {B : A -> Set β} -> Σ A B -> A 18 | proj₁ (x , y) = x 19 | 20 | proj₂ : ∀ {α β} {A : Set α} {B : A -> Set β} -> (p : Σ A B) -> B (proj₁ p) 21 | proj₂ (x , y) = y 22 | 23 | -- The syntax declaration below is attached to Σ-syntax, to make it 24 | -- easy to import Σ without the special syntax. 25 | 26 | infix 2 Σ-syntax 27 | 28 | Σ-syntax : ∀ {a b} (A : Set a) → (A → Set b) → Set (a ⊔ b) 29 | Σ-syntax = Σ 30 | 31 | syntax Σ-syntax A (λ x → B) = Σ[ x ∈ A ] B 32 | 33 | ∃ : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b) 34 | ∃ = Σ _ 35 | 36 | ∄ : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b) 37 | ∄ P = ¬ ∃ P 38 | 39 | ∃₂ : ∀ {a b c} {A : Set a} {B : A → Set b} 40 | (C : (x : A) → B x → Set c) → Set (a ⊔ b ⊔ c) 41 | ∃₂ C = ∃ λ a → ∃ λ b → C a b 42 | 43 | _×_ : ∀ {a b} (A : Set a) (B : Set b) → Set (a ⊔ b) 44 | A × B = Σ[ x ∈ A ] B 45 | 46 | _,′_ : ∀ {a b} {A : Set a} {B : Set b} → A → B → A × B 47 | _,′_ = _,_ 48 | 49 | ------------------------------------------------------------------------ 50 | -- Unique existence 51 | 52 | -- Parametrised on the underlying equality. 53 | 54 | ∃! : ∀ {a b ℓ} {A : Set a} → 55 | (A → A → Set ℓ) → (A → Set b) → Set (a ⊔ b ⊔ ℓ) 56 | ∃! _≈_ B = ∃ λ x → B x × (∀ {y} → B y → x ≈ y) 57 | 58 | ------------------------------------------------------------------------ 59 | -- Functions 60 | 61 | -- Sometimes the first component can be inferred. 62 | 63 | ,_ : ∀ {a b} {A : Set a} {B : A → Set b} {x} → B x → ∃ B 64 | , y = _ , y 65 | 66 | <_,_> : ∀ {a b c} {A : Set a} {B : A → Set b} {C : ∀ {x} → B x → Set c} 67 | (f : (x : A) → B x) → ((x : A) → C (f x)) → 68 | ((x : A) → Σ (B x) C) 69 | < f , g > x = (f x , g x) 70 | 71 | map : ∀ {a b p q} 72 | {A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q} → 73 | (f : A → B) → (∀ {x} → P x → Q (f x)) → 74 | Σ A P → Σ B Q 75 | map f g (x , y) = (f x , g y) 76 | 77 | zip : ∀ {a b c p q r} 78 | {A : Set a} {B : Set b} {C : Set c} 79 | {P : A → Set p} {Q : B → Set q} {R : C → Set r} → 80 | (_∙_ : A → B → C) → 81 | (∀ {x y} → P x → Q y → R (x ∙ y)) → 82 | Σ A P → Σ B Q → Σ C R 83 | zip _∙_ _∘_ (a , p) (b , q) = ((a ∙ b) , (p ∘ q)) 84 | 85 | swap : ∀ {a b} {A : Set a} {B : Set b} → A × B → B × A 86 | swap (x , y) = (y , x) 87 | 88 | _-×-_ : ∀ {a b i j} {A : Set a} {B : Set b} → 89 | (A → B → Set i) → (A → B → Set j) → (A → B → Set _) 90 | f -×- g = f -[ _×_ ]- g 91 | 92 | _-,-_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → 93 | (A → B → C) → (A → B → D) → (A → B → C × D) 94 | f -,- g = f -[ _,_ ]- g 95 | 96 | curry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} → 97 | ((p : Σ A B) → C p) → 98 | ((x : A) → (y : B x) → C (x , y)) 99 | curry f x y = f (x , y) 100 | 101 | uncurry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} → 102 | ((x : A) → (y : B x) → C (x , y)) → 103 | ((p : Σ A B) → C p) 104 | uncurry f (x , y) = f x y 105 | 106 | uncurry′ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 107 | (A → B → C) → (A × B → C) 108 | uncurry′ = uncurry 109 | -------------------------------------------------------------------------------- /src/STLC/Lib/Prelude.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | 3 | module STLC.Lib.Prelude where 4 | 5 | open import Level renaming (zero to lzero; suc to lsuc) public 6 | open import Function hiding (_∋_) public 7 | open import Relation.Nullary public 8 | open import Relation.Nullary.Decidable hiding (map) public 9 | open import Relation.Binary hiding (_⇒_) public 10 | open import Relation.Binary.PropositionalEquality hiding ([_]) public 11 | open import Data.Empty public 12 | open import Data.Unit.Base hiding (_≤_; module _≤_) public 13 | open import Data.Bool.Base hiding (_≟_) public 14 | open import Data.Nat.Base hiding (_≟_; _⊔_; _^_; erase; Ordering; compare) public 15 | open import Data.Fin hiding (_≟_; _+_; _<_; _≤_; _≤?_; fold; lift; pred) public 16 | open import Data.Sum hiding (swap) renaming (map to smap) public 17 | open import Data.Maybe.Base hiding (map) public 18 | open import Data.List.Base hiding ([_]; lookup; zip; fromMaybe; tabulate) public 19 | open import Data.List.Properties public 20 | open import Data.Vec using (Vec; []; _∷_; lookup; tabulate) public 21 | open import Data.Vec.N-ary renaming (_$ⁿ_ to _$ᵗⁿ_) public 22 | 23 | open import STLC.Lib.NoEtaProduct renaming (map to pmap) public 24 | 25 | {-# BUILTIN REWRITE _≡_ #-} 26 | 27 | import Data.Fin.Properties as FinP 28 | open import Category.Monad 29 | import Data.Maybe as Maybe 30 | 31 | private open module Dummy {α} = RawMonad {α} Maybe.monad hiding (pure; zipWith) public 32 | 33 | record Wrap {α} (A : Set α) : Set α where 34 | constructor wrap 35 | field unwrap : A 36 | open Wrap public 37 | 38 | infixl 10 _% 39 | infixl 2 _>>>_ 40 | 41 | _% = _∘_ 42 | 43 | _>>>_ : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : ∀ {x} -> B x -> Set γ} 44 | -> (f : (x : A) -> B x) -> (∀ {x} -> (y : B x) -> C y) -> ∀ x -> C (f x) 45 | f >>> g = g ∘ f 46 | 47 | right : ∀ {α} {A : Set α} {x y z : A} -> x ≡ z -> y ≡ z -> x ≡ y 48 | right p q rewrite q = p 49 | 50 | fromMaybe : ∀ {α} {A : Set α} -> A -> Maybe A -> A 51 | fromMaybe = maybe id 52 | 53 | first : ∀ {α β γ} {A : Set α} {B : Set β} {C : A -> Set γ} 54 | -> (∀ x -> C x) -> (p : A × B) -> C (proj₁ p) × B 55 | first f (x , y) = f x , y 56 | 57 | second : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : A -> Set γ} 58 | -> (∀ {x} -> B x -> C x) -> Σ A B -> Σ A C 59 | second g (x , y) = , g y 60 | 61 | inspectMaybe : ∀ {α β} {A : Set α} {B : Maybe A -> Set β} 62 | -> (mx : Maybe A) 63 | -> (∀ x -> mx ≡ just x -> B (just x)) 64 | -> (mx ≡ nothing -> B nothing) 65 | -> B mx 66 | inspectMaybe nothing g f = f refl 67 | inspectMaybe (just x) g f = g x refl 68 | 69 | delim : ∀ {α β} {A : Set α} {B : Dec A -> Set β} 70 | -> (∀ x -> B (yes x)) -> (∀ c -> B (no c)) -> (d : Dec A) -> B d 71 | delim f g (yes x) = f x 72 | delim f g (no c) = g c 73 | 74 | drec : ∀ {α β} {A : Set α} {B : Set β} -> (A -> B) -> (¬ A -> B) -> Dec A -> B 75 | drec = delim 76 | 77 | dmap : ∀ {α β} {A : Set α} {B : Set β} -> (A -> B) -> (¬ A -> ¬ B) -> Dec A -> Dec B 78 | dmap f g = drec (yes ∘ f) (no ∘ g) 79 | 80 | ddsum : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 81 | -> (A -> Dec C) -> (B -> Dec C) -> (¬ A -> ¬ B -> Dec C) -> Dec A -> Dec B -> Dec C 82 | ddsum f g h d₁ d₂ = drec f (λ c -> drec g (h c) d₂) d₁ 83 | 84 | ddprod : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 85 | -> (A -> B -> Dec C) -> (¬ A -> Dec C) -> (¬ B -> Dec C) -> Dec A -> Dec B -> Dec C 86 | ddprod h f g d₁ d₂ = drec (λ x -> drec (h x) g d₂) f d₁ 87 | 88 | dsum : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 89 | -> (A -> C) -> (B -> C) -> (¬ A -> ¬ B -> ¬ C) -> Dec A -> Dec B -> Dec C 90 | dsum f g h = ddsum (yes ∘ f) (yes ∘ g) (no % ∘ h) 91 | 92 | dprod : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 93 | -> (A -> B -> C) -> (¬ A -> ¬ C) -> (¬ B -> ¬ C) -> Dec A -> Dec B -> Dec C 94 | dprod h f g = ddprod (yes % ∘ h) (no ∘ f) (no ∘ g) 95 | 96 | dcong : ∀ {α β} {A : Set α} {B : Set β} {x y} 97 | -> (f : A -> B) 98 | -> (∀ {x y} -> f x ≡ f y -> x ≡ y) 99 | -> Dec (x ≡ y) 100 | -> Dec (f x ≡ f y) 101 | dcong f inj = dmap (cong f) (_∘ inj) 102 | 103 | dcong₂ : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} {x y v w} 104 | -> (f : A -> B -> C) 105 | -> (∀ {x y} -> f x v ≡ f y w -> x ≡ y × v ≡ w) 106 | -> Dec (x ≡ y) 107 | -> Dec (v ≡ w) 108 | -> Dec (f x v ≡ f y w) 109 | dcong₂ f inj = dprod (cong₂ f) (λ c -> c ∘ proj₁ ∘ inj) (λ c -> c ∘ proj₂ ∘ inj) 110 | 111 | dJ : ∀ {α β} {A : Set α} {x y} 112 | -> (B : A -> A -> Set β) 113 | -> Dec (B x x) 114 | -> (x ≢ y -> Dec (B x y)) 115 | -> Dec (x ≡ y) 116 | -> Dec (B x y) 117 | dJ B y f (yes refl) = y 118 | dJ B y f (no c) = f c 119 | 120 | record DecEq {α} (A : Set α) : Set α where 121 | infix 5 _≟_ _==_ 122 | field _≟_ : Decidable (_≡_ {A = A}) 123 | 124 | _==_ : A -> A -> Bool 125 | n == m = ⌊ n ≟ m ⌋ 126 | 127 | ≟-refl : ∀ x -> x ≟ x ≡ yes refl 128 | ≟-refl x with x ≟ x 129 | ... | yes refl = refl 130 | ... | no c = ⊥-elim (c refl) 131 | 132 | ≢-no : ∀ {x y} -> x ≢ y -> ∃ λ c -> x ≟ y ≡ no c 133 | ≢-no {x} {y} c with x ≟ y 134 | ... | yes p = ⊥-elim (c p) 135 | ... | no d = d , refl 136 | open DecEq {{...}} public 137 | 138 | instance 139 | DecEq-Fin : ∀ {n} -> DecEq (Fin n) 140 | DecEq-Fin = record { _≟_ = FinP._≟_ } 141 | 142 | lookup-for : ∀ {α β} {A : Set α} {B : Set β} {{_ : DecEq A}} 143 | -> A -> List (A × B) -> Maybe B 144 | lookup-for x [] = nothing 145 | lookup-for x ((y , z) ∷ xs) = if y == x then just z else lookup-for x xs 146 | 147 | -- It should be (List A ⊎ List A). 148 | deletem : ∀ {α} {A : Set α} {{_ : DecEq A}} -> A -> List A -> Maybe (List A) 149 | deletem x [] = nothing 150 | deletem x (y ∷ xs) = if y == x then just xs else (y ∷_) <$> deletem x xs 151 | 152 | delete : ∀ {α} {A : Set α} {{_ : DecEq A}} -> A -> List A -> List A 153 | delete x xs = fromMaybe xs (deletem x xs) 154 | 155 | nub : ∀ {α} {A : Set α} {{_ : DecEq A}} -> List A -> List A 156 | nub = foldr (λ x r -> x ∷ delete x r) [] 157 | 158 | proj₁-swap : ∀ {α β} {A : Set α} {B : Set β} -> (p : A × B) -> proj₁ (swap p) ≡ proj₂ p 159 | proj₁-swap (_ , _) = refl 160 | 161 | module Enumerate where 162 | go : ∀ {α} {A : Set α} 163 | -> (k : ℕ -> ℕ) 164 | -> (∀ {n} -> Fin (k n) -> Fin (k (suc n))) 165 | -> (∀ {n} -> Fin (k (suc n))) 166 | -> (xs : List A) 167 | -> List (Fin (k (length xs)) × A) 168 | go k s i [] = [] 169 | go k s i (x ∷ xs) = (i , x) ∷ go (k ∘ suc) s (s i) xs 170 | 171 | enumerate : ∀ {α} {A : Set α} -> (xs : List A) -> List (Fin (length xs) × A) 172 | enumerate = go id suc zero 173 | 174 | goed : ∀ {α} {A : Set α} {k : ℕ -> ℕ} 175 | {s : ∀ {n} -> Fin (k n) -> Fin (k (suc n))} 176 | {i : ∀ {n} -> Fin (k (suc n))} 177 | -> (xs : List A) -> map proj₂ (go k s i xs) ≡ xs 178 | goed [] = refl 179 | goed (x ∷ xs) = cong (x ∷_) (goed xs) 180 | 181 | enumerated : ∀ {α} {A : Set α} -> (xs : List A) -> map proj₂ (enumerate xs) ≡ xs 182 | enumerated = goed 183 | open Enumerate using (enumerate; enumerated) public 184 | 185 | _$ⁿ_ : ∀ {α β n} {A : Set α} {F : N-ary n A (Set β)} -> ∀ⁿ n F -> (xs : Vec _ n) -> F $ᵗⁿ xs 186 | y $ⁿ [] = y 187 | f $ⁿ (x ∷ xs) = f x $ⁿ xs 188 | 189 | _$ⁿʰ_ : ∀ {α β n} {A : Set α} {F : N-ary n A (Set β)} -> ∀ⁿʰ n F -> (xs : Vec _ n) -> F $ᵗⁿ xs 190 | y $ⁿʰ [] = y 191 | y $ⁿʰ (x ∷ xs) = y $ⁿʰ xs 192 | -------------------------------------------------------------------------------- /src/STLC/Lib/Sets.agda: -------------------------------------------------------------------------------- 1 | module STLC.Lib.Sets where 2 | 3 | open import Level as L 4 | open import Function 5 | open import Data.Unit.Base 6 | open import Data.Nat.Base hiding (_^_) 7 | open import Data.Fin 8 | open import Data.Product hiding (map) 9 | open import Data.Vec 10 | 11 | infixl 6 _^_ 12 | 13 | _^_ : ∀ {α} -> Set α -> ℕ -> Set α 14 | A ^ 0 = Lift ⊤ 15 | A ^ suc n = A × A ^ n 16 | 17 | to-^ : ∀ {n α} {A : Set α} -> Vec A n -> A ^ n 18 | to-^ = foldr (_^_ _) _,_ _ 19 | 20 | from-^ : ∀ {n α} {A : Set α} -> A ^ n -> Vec A n 21 | from-^ {0} _ = [] 22 | from-^ {suc _} (x , xs) = x ∷ from-^ xs 23 | 24 | on-^ : ∀ {α β n} {A : Set α} {B : Vec A n -> Set β} -> (∀ xs -> B xs) -> ∀ xs -> B (from-^ xs) 25 | on-^ f = f ∘ from-^ 26 | 27 | mono-^ : ∀ {α n m} {A : Set α} -> (Vec A n -> Vec A m) -> A ^ n -> A ^ m 28 | mono-^ f = to-^ ∘ on-^ f 29 | 30 | _⊔ⁿ_ : ∀ {n} -> Level ^ n -> Level -> Level 31 | _⊔ⁿ_ = on-^ $ flip $ foldr _ L._⊔_ 32 | 33 | Sets : ∀ {n} -> (αs : Level ^ n) -> Set (mono-^ (map L.suc) αs ⊔ⁿ L.zero) 34 | Sets {0} _ = ⊤ 35 | Sets {suc _} (α , αs) = Set α × Sets αs 36 | 37 | Lookup : ∀ {n} {αs : Level ^ n} i -> Sets αs -> Set (on-^ (lookup i) αs) 38 | Lookup zero (A , As) = A 39 | Lookup (suc i) (A , As) = Lookup i As 40 | -------------------------------------------------------------------------------- /src/STLC/M/Main.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | 3 | module STLC.M.Main where 4 | 5 | open import STLC.Lib.Prelude 6 | open import STLC.Term.Syntax 7 | open import STLC.Type 8 | open import STLC.Type.Unify 9 | open import STLC.M.Term 10 | 11 | {-# REWRITE mapᶜ-mapᶜ #-} 12 | 13 | -- `specialize' as a constructor? 14 | M : ∀ {n l} -> (Γ : Con n l) -> Syntax l -> (σ : Type n) 15 | -> Maybe (∃ λ m -> ∃ λ (Ψ : Subst n m) -> mapᶜ (apply Ψ) Γ ⊢ apply Ψ σ) 16 | M Γ (var i) σ = 17 | unify σ (lookupᶜ i Γ) 18 | >>= uncurry λ Ψ p -> 19 | just (, Ψ , coerceBy (sym p) (specialize Ψ (var (lookupᶜ-∈ i Γ)))) 20 | M Γ (ƛ bˢ) σ = 21 | unify (renᵗ 2 σ) (Var zero ⇒ Var (suc zero)) 22 | >>= uncurry λ Ψ p -> 23 | M (mapᶜ (apply (Ψ ∘ raise 2)) Γ ▻ apply Ψ (Var zero)) bˢ (apply Ψ (Var (suc zero))) 24 | >>= proj₂ >>> uncurry λ Φ b -> 25 | just (, Φ ∘ˢ Ψ ∘ raise 2 , coerceBy (cong (apply Φ) (sym p)) (ƛ b)) 26 | M Γ (fˢ · xˢ) σ = 27 | M (mapᶜ (renᵗ 1) Γ) fˢ (Var zero ⇒ renᵗ 1 σ) 28 | >>= proj₂ >>> uncurry λ Ψ f -> 29 | M (mapᶜ (apply (Ψ ∘ raise 1)) Γ) xˢ (apply Ψ (Var zero)) 30 | >>= proj₂ >>> uncurry λ Φ x -> 31 | just (, Φ ∘ˢ Ψ ∘ raise 1 , specialize Φ f · x) 32 | 33 | runM : Syntax⁽⁾ -> _ 34 | runM e = M ε e (Var {1} zero) 35 | -------------------------------------------------------------------------------- /src/STLC/M/Term.agda: -------------------------------------------------------------------------------- 1 | module STLC.M.Term where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Type 5 | import STLC.Term as C 6 | 7 | infixl 5 _▻_ 8 | infix 4 _∈_ _⊢_ 9 | infixr 3 ƛ_ 10 | infixl 6 _·_ 11 | 12 | data Con n : ℕ -> Set where 13 | ε : Con n 0 14 | _▻_ : ∀ {l} -> Con n l -> Type n -> Con n (suc l) 15 | 16 | data _∈_ {n} σ : ∀ {l} -> Con n l -> Set where 17 | vz : ∀ {l} {Γ : Con n l} -> σ ∈ Γ ▻ σ 18 | vs : ∀ {l τ} {Γ : Con n l} -> σ ∈ Γ -> σ ∈ Γ ▻ τ 19 | 20 | data _⊢_ {n l} (Γ : Con n l) : Type n -> Set where 21 | var : ∀ {σ} -> σ ∈ Γ -> Γ ⊢ σ 22 | ƛ_ : ∀ {σ τ} -> Γ ▻ σ ⊢ τ -> Γ ⊢ σ ⇒ τ 23 | _·_ : ∀ {σ τ} -> Γ ⊢ σ ⇒ τ -> Γ ⊢ σ -> Γ ⊢ τ 24 | 25 | coreᶜ : ∀ {n l} -> (Γ : Con n l) -> C.Con n 26 | coreᶜ ε = C.ε 27 | coreᶜ (Γ ▻ σ) = coreᶜ Γ C.▻ σ 28 | 29 | coreᵛ : ∀ {n l σ} {Γ : Con n l} -> σ ∈ Γ -> σ C.∈ coreᶜ Γ 30 | coreᵛ vz = C.vz 31 | coreᵛ (vs v) = C.vs (coreᵛ v) 32 | 33 | core : ∀ {n l σ} {Γ : Con n l} -> Γ ⊢ σ -> coreᶜ Γ C.⊢ σ 34 | core (var v) = C.var (coreᵛ v) 35 | core (ƛ b) = C.ƛ (core b) 36 | core (f · x) = core f C.· core x 37 | 38 | lookupᶜ : ∀ {n l} -> Fin l -> Con n l -> Type n 39 | lookupᶜ zero (Γ ▻ σ) = σ 40 | lookupᶜ (suc i) (Γ ▻ σ) = lookupᶜ i Γ 41 | 42 | mapᶜ : ∀ {n m l} -> (Type n -> Type m) -> Con n l -> Con m l 43 | mapᶜ f ε = ε 44 | mapᶜ f (Γ ▻ σ) = mapᶜ f Γ ▻ f σ 45 | 46 | lookupᶜ-∈ : ∀ {n l} i -> (Γ : Con n l) -> lookupᶜ i Γ ∈ Γ 47 | lookupᶜ-∈ zero (Γ ▻ σ) = vz 48 | lookupᶜ-∈ (suc i) (Γ ▻ σ) = vs (lookupᶜ-∈ i Γ) 49 | 50 | coerceBy : ∀ {n l σ τ} {Γ : Con n l} -> σ ≡ τ -> Γ ⊢ σ -> Γ ⊢ τ 51 | coerceBy refl = id 52 | 53 | coerce : ∀ {n l σ τ} {Γ : Con n l} -> Γ ⊢ σ -> Maybe (Γ ⊢ τ) 54 | coerce {σ = σ} {τ} t = flip coerceBy t <$> decToMaybe (σ ≟ τ) 55 | 56 | specializeᵛ : ∀ {n m l σ} {Γ : Con n l} 57 | -> (Ψ : Subst n m) -> σ ∈ Γ -> apply Ψ σ ∈ mapᶜ (apply Ψ) Γ 58 | specializeᵛ Ψ vz = vz 59 | specializeᵛ Ψ (vs v) = vs (specializeᵛ Ψ v) 60 | 61 | specialize : ∀ {n m l σ} {Γ : Con n l} 62 | -> (Ψ : Subst n m) -> Γ ⊢ σ -> mapᶜ (apply Ψ) Γ ⊢ apply Ψ σ 63 | specialize Ψ (var v) = var (specializeᵛ Ψ v) 64 | specialize Ψ (ƛ b) = ƛ (specialize Ψ b) 65 | specialize Ψ (f · x) = specialize Ψ f · specialize Ψ x 66 | 67 | mapᶜ-mapᶜ : ∀ {n m p l} {g : Type m -> Type p} {f : Type n -> Type m} (Γ : Con n l) 68 | -> mapᶜ g (mapᶜ f Γ) ≡ mapᶜ (g ∘ f) Γ 69 | mapᶜ-mapᶜ ε = refl 70 | mapᶜ-mapᶜ (Γ ▻ σ) = cong (_▻ _) (mapᶜ-mapᶜ Γ) 71 | -------------------------------------------------------------------------------- /src/STLC/M/Typecheck.agda: -------------------------------------------------------------------------------- 1 | module STLC.M.Typecheck where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Term.Syntax 5 | open import STLC.Type 6 | open import STLC.Type.Unify 7 | open import STLC.M.Term 8 | 9 | M : ∀ {n l} -> (Γ : Con n l) -> Syntax l -> (σ : Type n) -> Maybe (∃ λ m -> Subst n m) 10 | M Γ (var i) σ = 11 | unify σ (lookupᶜ i Γ) 12 | >>= uncurry λ Ψ p -> 13 | just (, Ψ) 14 | M Γ (ƛ bˢ) σ = 15 | unify (renᵗ 2 σ) (Var zero ⇒ Var (suc zero)) 16 | >>= uncurry λ Ψ p -> 17 | M (mapᶜ (apply (Ψ ∘ raise 2)) Γ ▻ apply Ψ (Var zero)) bˢ (apply Ψ (Var (suc zero))) 18 | >>= proj₂ >>> λ Φ -> 19 | just (, Φ ∘ˢ Ψ ∘ raise 2) 20 | M Γ (fˢ · xˢ) σ = 21 | M (mapᶜ (renᵗ 1) Γ) fˢ (Var zero ⇒ renᵗ 1 σ) 22 | >>= proj₂ >>> λ Ψ -> 23 | M (mapᶜ (apply (Ψ ∘ raise 1)) Γ) xˢ (apply Ψ (Var zero)) 24 | >>= proj₂ >>> λ Φ -> 25 | just (, Φ ∘ˢ Ψ ∘ raise 1) 26 | 27 | runM : Syntax⁽⁾ -> _ 28 | runM e = M ε e (Var {1} zero) 29 | 30 | mutual 31 | infer : ∀ {n l} -> (Γ : Con n l) -> Syntax l -> Maybe (∃ (Γ ⊢_)) 32 | infer Γ (var v) = just (, var (lookupᶜ-∈ v Γ)) 33 | infer Γ (ƛ b) = nothing 34 | infer Γ (f · x) = infer Γ f >>= λ 35 | { (σ ⇒ τ , f') -> (λ x' -> , f' · x') <$> check Γ x σ 36 | ; (Var i , x') -> nothing 37 | } 38 | 39 | check : ∀ {n l} -> (Γ : Con n l) -> Syntax l -> (σ : Type n) -> Maybe (Γ ⊢ σ) 40 | check Γ (ƛ b) (σ ⇒ τ) = ƛ_ <$> check (Γ ▻ σ) b τ 41 | check Γ e σ = infer Γ e >>= coerce ∘ proj₂ 42 | 43 | typecheck : ∀ {n} -> Syntax⁽⁾ -> (σ : Type n) -> Maybe (ε ⊢ σ) 44 | typecheck = check ε 45 | -------------------------------------------------------------------------------- /src/STLC/Main.agda: -------------------------------------------------------------------------------- 1 | module STLC.Main where 2 | 3 | open import STLC.Lib.Prelude public 4 | open import STLC.Term.Syntax public 5 | open import STLC.Term public 6 | open import STLC.Semantics.Eval using (eval) public 7 | open import STLC.NbE.Main using (norm) public 8 | open import STLC.NbE.Read using (read; inst) public 9 | 10 | open import STLC.Lib.MaybeElim 11 | open import STLC.Lib.Membership 12 | open import STLC.Experimental.Safe 13 | open import STLC.Experimental.Unsafe 14 | open import STLC.M.Term using (core) 15 | 16 | module NF where 17 | open import STLC.M.Typecheck using (runM; typecheck) 18 | 19 | -- If `e` is in NF. 20 | on-typed : ∀ {α} {A : ∀ {n} {σ : Type n} -> Term⁽⁾ σ -> Set α} 21 | -> (f : ∀ {n} {σ : Type n} -> (t : Term⁽⁾ σ) -> A t) -> ∀ e -> _ 22 | on-typed f e = 23 | runM e >>=⊤ proj₂ >>> λ Ψ -> 24 | let σ = apply Ψ (Var zero) in 25 | typecheck e (runSubstIn (thickenˢ σ)) >>=⊤ λ t -> 26 | f (core t) 27 | 28 | open import STLC.M.Main using (runM) 29 | 30 | on-typed : ∀ {α} {A : ∀ {n} {σ : Type n} -> Term⁽⁾ σ -> Set α} 31 | -> (f : ∀ {n} {σ : Type n} -> (t : Term⁽⁾ σ) -> A t) -> ∀ e -> _ 32 | on-typed f e = runM e >>=⊤ f ∘ thicken ∘ core ∘ proj₂ ∘ proj₂ 33 | 34 | typed = on-typed $ id 35 | term = on-typed $ λ t {m Γ} -> generalize {m} Γ t 36 | term⁻ = on-typed $ λ {n} t {Γ} -> generalize {n} Γ t 37 | pnorm = on-typed $ pure ∘ erase ∘ norm 38 | 39 | module Names {m} where 40 | name : ∀ n -> Type (suc n + m) 41 | name = wkᵗ ∘ Var ∘ fromℕ 42 | 43 | a = name 0 44 | b = name 1 45 | c = name 2 46 | d = name 3 47 | e = name 4 48 | f = name 5 49 | g = name 6 50 | h = name 7 51 | i = name 8 52 | j = name 9 53 | k = name 10 54 | open Names public 55 | -------------------------------------------------------------------------------- /src/STLC/NbE/Main.agda: -------------------------------------------------------------------------------- 1 | module STLC.NbE.Main where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Term 5 | open import STLC.Semantics.Core 6 | open import STLC.NbE.NF 7 | 8 | infix 4 _⊨_ _⊨*_ 9 | infixl 6 _·ᵐ_ 10 | 11 | mutual 12 | _⊨_ : Links 13 | Γ ⊨ σ = Γ ⊢ⁿᵉ σ ⊎ Γ ⊨* σ 14 | 15 | _⊨*_ : Links 16 | Γ ⊨* Var i = ⊥ 17 | Γ ⊨* σ ⇒ τ = ∀ {Δ} -> Γ ⊆ Δ -> Δ ⊨ σ -> Δ ⊨ τ 18 | 19 | varᵐ : _∋_ ∸> _⊨_ 20 | varᵐ = inj₁ ∘ varⁿᵉ 21 | 22 | open Kripkable record 23 | { _∙_ = _⊨_ 24 | ; varˢ = varᵐ 25 | } 26 | 27 | renᵐ : Renames _⊨_ 28 | renᵐ ι (inj₁ n) = inj₁ (renⁿᵉ ι n) 29 | renᵐ {σ = Var i} ι (inj₂ ()) 30 | renᵐ {σ = σ ⇒ τ} ι (inj₂ k) = inj₂ (renᵏ ι k) 31 | 32 | readback : _⊨_ ∸> _⊢ⁿᶠ_ 33 | readback (inj₁ n) = neⁿᶠ n 34 | readback {σ = Var i} (inj₂ ()) 35 | readback {σ = σ ⇒ τ} (inj₂ k) = ƛⁿᶠ (readback (apᵏ k)) 36 | 37 | _·ᵐ_ : ∀ {n σ τ} {Γ : Con n} -> Γ ⊨ σ ⇒ τ -> Γ ⊨ σ -> Γ ⊨ τ 38 | inj₁ f ·ᵐ x = inj₁ (f ·ⁿᵉ readback x) 39 | inj₂ k ·ᵐ x = k ·ᵏ x 40 | 41 | open Environment record 42 | { renˢ = renᵐ 43 | } 44 | 45 | open Semantics record 46 | { _◆_ = _⊨_ 47 | ; embˢ = id 48 | ; lamˢ = inj₂ 49 | ; _·ˢ_ = _·ᵐ_ 50 | } 51 | 52 | nf : _⊢_ ∸> _⊢ⁿᶠ_ 53 | nf = readback ∘ eval 54 | 55 | norm : _⊢_ ∸> _⊢_ 56 | norm = embⁿᶠ ∘ nf 57 | -------------------------------------------------------------------------------- /src/STLC/NbE/NF.agda: -------------------------------------------------------------------------------- 1 | module STLC.NbE.NF where 2 | 3 | open import STLC.Term 4 | 5 | infix 4 _⊢ⁿᵉ_ _⊢ⁿᶠ_ 6 | infixl 6 _·ⁿᵉ_ 7 | 8 | mutual 9 | data _⊢ⁿᵉ_ {n} Γ : Type n -> Set where 10 | varⁿᵉ : ∀ {σ} -> σ ∈ Γ -> Γ ⊢ⁿᵉ σ 11 | _·ⁿᵉ_ : ∀ {σ τ} -> Γ ⊢ⁿᵉ σ ⇒ τ -> Γ ⊢ⁿᶠ σ -> Γ ⊢ⁿᵉ τ 12 | 13 | data _⊢ⁿᶠ_ {n} Γ : Type n -> Set where 14 | neⁿᶠ : ∀ {σ} -> Γ ⊢ⁿᵉ σ -> Γ ⊢ⁿᶠ σ 15 | ƛⁿᶠ : ∀ {σ τ} -> Γ ▻ σ ⊢ⁿᶠ τ -> Γ ⊢ⁿᶠ σ ⇒ τ 16 | 17 | mutual 18 | embⁿᵉ : _⊢ⁿᵉ_ ∸> _⊢_ 19 | embⁿᵉ (varⁿᵉ v) = var v 20 | embⁿᵉ (f ·ⁿᵉ x) = embⁿᵉ f · embⁿᶠ x 21 | 22 | embⁿᶠ : _⊢ⁿᶠ_ ∸> _⊢_ 23 | embⁿᶠ (neⁿᶠ x) = embⁿᵉ x 24 | embⁿᶠ (ƛⁿᶠ b) = ƛ (embⁿᶠ b) 25 | 26 | mutual 27 | renⁿᵉ : Renames _⊢ⁿᵉ_ 28 | renⁿᵉ ι (varⁿᵉ v) = varⁿᵉ (renᵛ ι v) 29 | renⁿᵉ ι (f ·ⁿᵉ x) = renⁿᵉ ι f ·ⁿᵉ renⁿᶠ ι x 30 | 31 | renⁿᶠ : Renames _⊢ⁿᶠ_ 32 | renⁿᶠ ι (neⁿᶠ n) = neⁿᶠ (renⁿᵉ ι n) 33 | renⁿᶠ ι (ƛⁿᶠ b) = ƛⁿᶠ (renⁿᶠ (keep ι) b) 34 | -------------------------------------------------------------------------------- /src/STLC/NbE/Read.agda: -------------------------------------------------------------------------------- 1 | module STLC.NbE.Read where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Term 5 | open import STLC.NbE.NF 6 | 7 | Ne : ∀ {n} -> Type n -> Set 8 | Ne σ = Wrap (∀ {Γ} -> Γ ⊢ⁿᵉ σ) 9 | 10 | NF : ∀ {n} -> Type n -> Set 11 | NF σ = ∀ {Γ} -> Γ ⊢ⁿᶠ σ 12 | 13 | ⟦_⟧ : ∀ {n} -> Type n -> Set 14 | ⟦ Var i ⟧ = Ne (Var i) 15 | ⟦ σ ⇒ τ ⟧ = ⟦ σ ⟧ -> ⟦ τ ⟧ 16 | 17 | mutual 18 | ↑ : ∀ {n} {σ : Type n} -> Ne σ -> ⟦ σ ⟧ 19 | ↑ {σ = Var i} n = n 20 | ↑ {σ = σ ⇒ τ} f = λ x -> ↑ (wrap (unwrap f ·ⁿᵉ ↓ x)) 21 | 22 | ↓ : ∀ {n} {σ : Type n} -> ⟦ σ ⟧ -> NF σ 23 | ↓ {σ = Var i} n = neⁿᶠ (unwrap n) 24 | ↓ {σ = σ ⇒ τ} f {Γ} = ƛⁿᶠ (↓ (f (↑ (wrap λ {Δ} -> varⁿᵉ (diff Δ Γ σ))))) where 25 | diff : ∀ Δ Γ σ -> σ ∈ Δ 26 | diff Δ Γ σ = drec ⊂[]-to-∈ impossible (Γ ⊂? Δ) where 27 | postulate impossible : _ 28 | 29 | read : ∀ {n} {σ : Type n} -> ⟦ σ ⟧ -> Term σ 30 | read x = wk (embⁿᶠ (↓ x {ε})) 31 | 32 | inst : ∀ n {F : N-ary n Set Set} -> ∀ⁿ n F -> F $ᵗⁿ tabulate {n} (Ne ∘ Var) 33 | inst n f = f $ⁿ tabulate {n} (Ne ∘ Var) 34 | 35 | instʰ : ∀ n {F : N-ary n Set Set} -> ∀ⁿʰ n F -> F $ᵗⁿ tabulate {n} (Ne ∘ Var) 36 | instʰ n y = y $ⁿʰ tabulate {n} (Ne ∘ Var) 37 | -------------------------------------------------------------------------------- /src/STLC/Semantics/Core.agda: -------------------------------------------------------------------------------- 1 | module STLC.Semantics.Core where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Term 5 | 6 | -- Forgive me this. 7 | record Kripkable : Set₁ where 8 | infix 4 _∙_ 9 | infixr 9 _∘ᵏ_ 10 | 11 | field 12 | _∙_ : Links 13 | varˢ : _∋_ ∸> _∙_ 14 | 15 | Kripke : ∀ {n} -> Con n -> Type n -> Type n -> Set 16 | Kripke Γ σ τ = ∀ {Δ} -> Γ ⊆ Δ -> Δ ∙ σ -> Δ ∙ τ 17 | 18 | apᵏ : ∀ {n σ τ} {Γ : Con n} -> Kripke Γ σ τ -> Γ ▻ σ ∙ τ 19 | apᵏ k = k top (varˢ vz) 20 | 21 | _·ᵏ_ : ∀ {n σ τ} {Γ : Con n} -> Kripke Γ σ τ -> Γ ∙ σ -> Γ ∙ τ 22 | k ·ᵏ t = k stop t 23 | 24 | renᵏ : ∀ {n σ τ} {Γ Δ : Con n} -> Γ ⊆ Δ -> Kripke Γ σ τ -> Kripke Δ σ τ 25 | renᵏ ι k κ = k (κ ∘ᵉ ι) 26 | 27 | _∘ᵏ_ : ∀ {n σ τ ν} {Γ : Con n} -> Kripke Γ τ ν -> Kripke Γ σ τ -> Kripke Γ σ ν 28 | (k₂ ∘ᵏ k₁) ι = k₂ ι ∘ k₁ ι 29 | 30 | record Environment : Set where 31 | infix 4 _⊢*_ 32 | 33 | field 34 | renˢ : Renames _∙_ 35 | 36 | data _⊢*_ {n} Δ : Con n -> Set where 37 | Ø : Δ ⊢* ε 38 | _▷_ : ∀ {Γ σ} -> Δ ⊢* Γ -> Δ ∙ σ -> Δ ⊢* Γ ▻ σ 39 | 40 | lookupᵉ : ∀ {n σ} {Γ Δ : Con n} -> σ ∈ Γ -> Δ ⊢* Γ -> Δ ∙ σ 41 | lookupᵉ vz (ρ ▷ x) = x 42 | lookupᵉ (vs v) (ρ ▷ x) = lookupᵉ v ρ 43 | 44 | renᵉ : ∀ {n} {Γ Δ Ξ : Con n} -> Δ ⊆ Ξ -> Δ ⊢* Γ -> Ξ ⊢* Γ 45 | renᵉ ι Ø = Ø 46 | renᵉ ι (ρ ▷ x) = renᵉ ι ρ ▷ renˢ ι x 47 | 48 | keepᵉ : ∀ {n σ} {Γ Δ : Con n} -> Δ ⊢* Γ -> Δ ▻ σ ⊢* Γ ▻ σ 49 | keepᵉ ρ = renᵉ top ρ ▷ varˢ vz 50 | 51 | idᵉ : ∀ {n} {Γ : Con n} -> Γ ⊢* Γ 52 | idᵉ {Γ = ε} = Ø 53 | idᵉ {Γ = Γ ▻ σ} = keepᵉ idᵉ 54 | 55 | record Semantics : Set₁ where 56 | infix 4 _◆_ 57 | infixl 6 _·ˢ_ 58 | 59 | field 60 | _◆_ : Links 61 | embˢ : _∙_ ∸> _◆_ 62 | lamˢ : ∀ {n σ τ} {Γ : Con n} -> (∀ {Δ} -> Γ ⊆ Δ -> Δ ∙ σ -> Δ ◆ τ) -> Γ ◆ σ ⇒ τ 63 | _·ˢ_ : ∀ {n σ τ} {Γ : Con n} -> Γ ◆ σ ⇒ τ -> Γ ◆ σ -> Γ ◆ τ 64 | 65 | ⟦_⟧ : ∀ {n σ} {Δ Γ : Con n} -> Γ ⊢ σ -> Δ ⊢* Γ -> Δ ◆ σ 66 | ⟦ var v ⟧ ρ = embˢ (lookupᵉ v ρ) 67 | ⟦ ƛ b ⟧ ρ = lamˢ λ ι x -> ⟦ b ⟧ (renᵉ ι ρ ▷ x) 68 | ⟦ f · x ⟧ ρ = ⟦ f ⟧ ρ ·ˢ ⟦ x ⟧ ρ 69 | 70 | eval : _⊢_ ∸> _◆_ 71 | eval t = ⟦ t ⟧ idᵉ 72 | -------------------------------------------------------------------------------- /src/STLC/Semantics/Eval.agda: -------------------------------------------------------------------------------- 1 | module STLC.Semantics.Eval where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Lib.Sets 5 | open import STLC.Term 6 | 7 | ⟦_⟧ᵗˡ : ∀ {n} σ -> Level ^ n -> Level 8 | ⟦ Var i ⟧ᵗˡ αs = on-^ (lookup i) αs 9 | ⟦ σ ⇒ τ ⟧ᵗˡ αs = ⟦ σ ⟧ᵗˡ αs ⊔ ⟦ τ ⟧ᵗˡ αs 10 | 11 | ⟦_⟧ᵗ : ∀ {n} {αs : Level ^ n} σ -> Sets αs -> Set (⟦ σ ⟧ᵗˡ αs) 12 | ⟦ Var i ⟧ᵗ As = Lookup i As 13 | ⟦ σ ⇒ τ ⟧ᵗ As = ⟦ σ ⟧ᵗ As -> ⟦ τ ⟧ᵗ As 14 | 15 | ⟦_⟧ᶜˡ : ∀ {n} -> Con n -> Level ^ n -> Level 16 | ⟦ ε ⟧ᶜˡ αs = lzero 17 | ⟦ Γ ▻ σ ⟧ᶜˡ αs = ⟦ σ ⟧ᵗˡ αs ⊔ ⟦ Γ ⟧ᶜˡ αs 18 | 19 | ⟦_⟧ᶜ : ∀ {n} {αs : Level ^ n} Γ -> Sets αs -> Set (⟦ Γ ⟧ᶜˡ αs) 20 | ⟦ ε ⟧ᶜ As = ⊤ 21 | ⟦ Γ ▻ σ ⟧ᶜ As = ⟦ σ ⟧ᵗ As × ⟦ Γ ⟧ᶜ As 22 | 23 | lookupᵉ : ∀ {n Γ σ} {αs : Level ^ n} {As : Sets αs} -> σ ∈ Γ -> ⟦ Γ ⟧ᶜ As -> ⟦ σ ⟧ᵗ As 24 | lookupᵉ vz (x , ρ) = x 25 | lookupᵉ (vs v) (x , ρ) = lookupᵉ v ρ 26 | 27 | ⟦_⟧ : ∀ {n Γ σ} {αs : Level ^ n} {As : Sets αs} -> Γ ⊢ σ -> ⟦ Γ ⟧ᶜ As -> ⟦ σ ⟧ᵗ As 28 | ⟦ var v ⟧ ρ = lookupᵉ v ρ 29 | ⟦ ƛ b ⟧ ρ = λ x -> ⟦ b ⟧ (x , ρ) 30 | ⟦ f · x ⟧ ρ = ⟦ f ⟧ ρ (⟦ x ⟧ ρ) 31 | 32 | eval : ∀ {n σ} {αs : Level ^ n} {As : Sets αs} -> Term⁽⁾ σ -> ⟦ σ ⟧ᵗ As 33 | eval t = ⟦ t ⟧ _ 34 | -------------------------------------------------------------------------------- /src/STLC/Term.agda: -------------------------------------------------------------------------------- 1 | module STLC.Term where 2 | 3 | open import STLC.Context public 4 | open import STLC.Term.Core public 5 | -------------------------------------------------------------------------------- /src/STLC/Term/Core.agda: -------------------------------------------------------------------------------- 1 | module STLC.Term.Core where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Term.Syntax 5 | open import STLC.Context 6 | 7 | infix 4 _∈_ _∋_ _⊢_ 8 | infix 3 vs_ 9 | infixr 3 ƛ_ 10 | infixl 6 _·_ 11 | 12 | data _∈_ {n} σ : Con n -> Set where 13 | vz : ∀ {Γ} -> σ ∈ Γ ▻ σ 14 | vs_ : ∀ {Γ τ} -> σ ∈ Γ -> σ ∈ Γ ▻ τ 15 | 16 | data _⊢_ {n} Γ : Type n -> Set where 17 | var : ∀ {σ} -> σ ∈ Γ -> Γ ⊢ σ 18 | ƛ_ : ∀ {σ τ} -> Γ ▻ σ ⊢ τ -> Γ ⊢ σ ⇒ τ 19 | _·_ : ∀ {σ τ} -> Γ ⊢ σ ⇒ τ -> Γ ⊢ σ -> Γ ⊢ τ 20 | 21 | Term⁽⁾ : ∀ {n} -> Type n -> Set 22 | Term⁽⁾ σ = ε ⊢ σ 23 | 24 | Term : ∀ {n} -> Type n -> Set 25 | Term σ = ∀ {Γ} -> Γ ⊢ σ 26 | 27 | Term⁺ : ∀ n -> (∀ {m} -> Type (n + m)) -> Set 28 | Term⁺ n σ = ∀ {m Γ} -> Γ ⊢ σ {m} 29 | 30 | _∋_ : Links 31 | _∋_ = flip _∈_ 32 | 33 | ⊂[]-to-∈ : ∀ {n σ} {Γ Δ : Con n} -> Γ ⊂[ σ ] Δ -> σ ∈ Δ 34 | ⊂[]-to-∈ vtop = vz 35 | ⊂[]-to-∈ (vskip p) = vs (⊂[]-to-∈ p) 36 | 37 | ∈-to-Fin : ∀ {n σ} {Γ : Con n} -> σ ∈ Γ -> Fin (lenᶜ Γ) 38 | ∈-to-Fin vz = zero 39 | ∈-to-Fin (vs v) = suc (∈-to-Fin v) 40 | 41 | wkᵛ : ∀ {n σ} {Δ Γ : Con n} -> σ ∈ Γ -> σ ∈ Δ ▻▻ Γ 42 | wkᵛ vz = vz 43 | wkᵛ (vs v) = vs (wkᵛ v) 44 | 45 | renᵛ : Renames _∋_ 46 | renᵛ stop v = v 47 | renᵛ (skip ι) v = vs (renᵛ ι v) 48 | renᵛ (keep ι) vz = vz 49 | renᵛ (keep ι) (vs v) = vs (renᵛ ι v) 50 | 51 | erase : ∀ {n σ} {Γ : Con n} -> Γ ⊢ σ -> Syntax (lenᶜ Γ) 52 | erase (var v) = var (∈-to-Fin v) 53 | erase (ƛ b) = ƛ (erase b) 54 | erase (f · x) = erase f · erase x 55 | 56 | upper : ∀ {n σ} {Γ : Con n} -> Γ ⊢ σ -> ℕ 57 | upper {n} _ = n 58 | 59 | wk : ∀ {n σ} {Δ Γ : Con n} -> Γ ⊢ σ -> Δ ▻▻ Γ ⊢ σ 60 | wk (var v) = var (wkᵛ v) 61 | wk (ƛ b) = ƛ (wk b) 62 | wk (f · x) = wk f · wk x 63 | 64 | ren : Renames _⊢_ 65 | ren ι (var v) = var (renᵛ ι v) 66 | ren ι (ƛ b) = ƛ (ren (keep ι) b) 67 | ren ι (f · x) = ren ι f · ren ι x 68 | 69 | specializeᵛ : ∀ {n m σ} {Γ : Con n} -> (Ψ : Subst n m) -> σ ∈ Γ -> apply Ψ σ ∈ mapᶜ (apply Ψ) Γ 70 | specializeᵛ Ψ vz = vz 71 | specializeᵛ Ψ (vs v) = vs (specializeᵛ Ψ v) 72 | 73 | specialize : ∀ {n m σ} {Γ : Con n} -> (Ψ : Subst n m) -> Γ ⊢ σ -> mapᶜ (apply Ψ) Γ ⊢ apply Ψ σ 74 | specialize Ψ (var v) = var (specializeᵛ Ψ v) 75 | specialize Ψ (ƛ b) = ƛ (specialize Ψ b) 76 | specialize Ψ (f · x) = specialize Ψ f · specialize Ψ x 77 | 78 | widen : ∀ {n σ} {Γ : Con n} m -> Γ ⊢ σ -> _ 79 | widen m = specialize (wkᵗ {m} ∘ Var) 80 | -------------------------------------------------------------------------------- /src/STLC/Term/Syntax.agda: -------------------------------------------------------------------------------- 1 | module STLC.Term.Syntax where 2 | 3 | open import STLC.Lib.Prelude 4 | 5 | infix 3 ƛ_ 6 | infixl 6 _·_ 7 | infix 10 _#_ 8 | 9 | data Syntax n : Set where 10 | var : Fin n -> Syntax n 11 | ƛ_ : Syntax (suc n) -> Syntax n 12 | _·_ : Syntax n -> Syntax n -> Syntax n 13 | 14 | Syntax⁽⁾ : Set 15 | Syntax⁽⁾ = Syntax 0 16 | 17 | wkˢ : ∀ {m n} -> Syntax n -> Syntax (n + m) 18 | wkˢ (var i) = var (inject+ _ i) 19 | wkˢ (ƛ b) = ƛ (wkˢ b) 20 | wkˢ (f · x) = wkˢ f · wkˢ x 21 | 22 | Pure : Set 23 | Pure = ∀ {n} -> Syntax n 24 | 25 | pure : Syntax⁽⁾ -> Pure 26 | pure e = wkˢ e 27 | 28 | shift : ∀ {m} n -> Fin (suc (n + m)) 29 | shift 0 = fromℕ _ 30 | shift (suc n) = inject₁ (shift n) 31 | 32 | Bind : ℕ -> ℕ -> Set 33 | Bind n 0 = Syntax n 34 | Bind n (suc m) = (∀ {p} -> Syntax (suc n + p)) -> Bind (suc n) m 35 | 36 | _#_ : ∀ {n} m -> Bind n m -> Syntax n 37 | _#_ zero b = b 38 | _#_ {n} (suc m) b = ƛ (m # b (var (shift n))) 39 | -------------------------------------------------------------------------------- /src/STLC/Tests/Church.agda: -------------------------------------------------------------------------------- 1 | module STLC.Tests.Church where 2 | 3 | open import STLC.Main hiding (_·_) 4 | open import STLC.Term using (_·_) 5 | 6 | ℕᵗ : ∀ {n} -> Type n -> Type n 7 | ℕᵗ σ = (σ ⇒ σ) ⇒ σ ⇒ σ 8 | 9 | zeroᵗ : ∀ {n} {σ : Type n} -> Term (ℕᵗ σ) 10 | zeroᵗ = ƛ ƛ var vz 11 | 12 | foldᵗ : ∀ {n} {σ : Type n} -> Term ((σ ⇒ σ) ⇒ σ ⇒ ℕᵗ σ ⇒ σ) 13 | foldᵗ = ƛ ƛ ƛ var vz · var (vs vs vz) · var (vs vz) 14 | 15 | sucᵗ : ∀ {n} {σ : Type n} -> Term (ℕᵗ σ ⇒ ℕᵗ σ) 16 | sucᵗ = ƛ ƛ ƛ var (vs vz) · (foldᵗ · (var (vs vz)) · (var vz) · var (vs (vs vz))) 17 | 18 | plusᵗ : ∀ {n} {σ : Type n} -> Term (ℕᵗ (ℕᵗ σ) ⇒ ℕᵗ σ ⇒ ℕᵗ σ) 19 | plusᵗ = ƛ ƛ foldᵗ · sucᵗ · var vz · var (vs vz) 20 | 21 | Listᵗ : ∀ {n} -> Type n -> Type n -> Type n 22 | Listᵗ σ τ = (σ ⇒ τ ⇒ τ) ⇒ τ ⇒ τ 23 | 24 | nilᵗ : ∀ {n} {σ τ : Type n} -> Term (Listᵗ σ τ) 25 | nilᵗ = ƛ ƛ var vz 26 | 27 | foldrᵗ : ∀ {n} {σ τ : Type n} -> Term ((σ ⇒ τ ⇒ τ) ⇒ τ ⇒ Listᵗ σ τ ⇒ τ) 28 | foldrᵗ = ƛ ƛ ƛ var vz · var (vs vs vz) · var (vs vz) 29 | 30 | consᵗ : ∀ {n} {σ τ : Type n} -> Term (σ ⇒ Listᵗ σ τ ⇒ Listᵗ σ τ) 31 | consᵗ = ƛ ƛ ƛ ƛ var (vs vz) · var (vs vs vs vz) 32 | · (foldrᵗ · var (vs vz) · var vz · var (vs (vs vz))) 33 | 34 | sumᵗ : ∀ {n} {σ : Type n} -> Term (Listᵗ (ℕᵗ (ℕᵗ σ)) (ℕᵗ σ) ⇒ ℕᵗ σ) 35 | sumᵗ = ƛ foldrᵗ · plusᵗ · zeroᵗ · var vz 36 | 37 | module Nums {n σ Γ} where 38 | 0ᵗ = zeroᵗ {n} {σ} {Γ} 39 | 1ᵗ = sucᵗ · 0ᵗ 40 | 2ᵗ = sucᵗ · 1ᵗ 41 | 3ᵗ = sucᵗ · 2ᵗ 42 | 4ᵗ = sucᵗ · 3ᵗ 43 | 5ᵗ = sucᵗ · 4ᵗ 44 | 6ᵗ = sucᵗ · 5ᵗ 45 | open Nums 46 | 47 | ex : ∀ {n} {σ τ : Type n} -> Term (Listᵗ (ℕᵗ σ) τ) 48 | ex = consᵗ · 0ᵗ · (consᵗ · 1ᵗ · (consᵗ · 2ᵗ · (consᵗ · 3ᵗ · nilᵗ))) 49 | 50 | test : ∀ {n Γ σ} -> norm (sumᵗ · ex) ≡ norm (6ᵗ {n} {σ} {Γ}) 51 | test = refl 52 | -------------------------------------------------------------------------------- /src/STLC/Tests/Combs.agda: -------------------------------------------------------------------------------- 1 | module STLC.Tests.Combs where 2 | 3 | open import STLC.Main 4 | 5 | I : Pure 6 | I = pure $ 1 # λ x → x 7 | 8 | -- λ {.n} → ƛ var zero 9 | I' : Pure 10 | I' = pnorm $ I · I 11 | 12 | ω : Pure 13 | ω = pure $ 1 # λ x → x · x 14 | 15 | Ωᵀ : Lift ⊤ 16 | Ωᵀ = term $ ω · ω 17 | 18 | applicator : ∀ {n} {a b : Type n} -> Term ((a ⇒ b) ⇒ a ⇒ b) 19 | applicator = term $ 2 # λ a b → a · b 20 | 21 | applicator⁺ : Term⁺ 2 ((a ⇒ b) ⇒ a ⇒ b) 22 | applicator⁺ = term $ 2 # λ a b → a · b 23 | 24 | applicator⁻ : Term ((b ⇒ a) ⇒ b ⇒ a) 25 | applicator⁻ = term⁻ $ 2 # λ a b → a · b 26 | 27 | dollar : ∀ {α β} {A : Set α} {B : Set β} -> (A -> B) -> A -> B 28 | dollar = eval applicator⁻ 29 | 30 | same-applicator : ∀ {n} {a b : Type n} -> Term ((a ⇒ b) ⇒ a ⇒ b) 31 | same-applicator = read _$′_ 32 | 33 | mono-app : {A B : Set} -> (A -> B) -> A -> B 34 | mono-app f x = f x 35 | 36 | poly-app : ∀ {α β} {A : Set α} {B : Set β} -> (A -> B) -> A -> B 37 | poly-app = eval $ read $ inst 2 λ A B -> mono-app {A} {B} 38 | 39 | applicator-speсialized : Term⁺ 3 (((b ⇒ c) ⇒ a) ⇒ (b ⇒ c) ⇒ a) 40 | applicator-speсialized = term $ 2 # λ a b → a · b 41 | 42 | -- Why the hell this type checks forever? 43 | -- applicator-generic : ∀ {a b} -> Term ((a ⇒ b) ⇒ a ⇒ b) 44 | -- applicator-generic = term⁻ $ 2 # λ a b → a · b 45 | 46 | -- applicator-generic-specialized : ∀ {a} -> Term ((a ⇒ a) ⇒ a ⇒ a) 47 | -- applicator-generic-specialized = applicator-generic 48 | 49 | cardinal : Term ((a ⇒ b ⇒ c) ⇒ b ⇒ a ⇒ c) 50 | cardinal = term⁻ $ 3 # λ a b c → a · c · b 51 | 52 | owl : Term (((a ⇒ b) ⇒ a) ⇒ (a ⇒ b) ⇒ b) 53 | owl = term⁻ $ 2 # λ a b → b · (a · b) 54 | 55 | quacky : Term (a ⇒ (a ⇒ b) ⇒ (b ⇒ c) ⇒ c) 56 | quacky = term⁻ $ 3 # λ a b c → c · (b · a) 57 | 58 | psi : Term ((b ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ a ⇒ c) 59 | psi = term⁻ $ 4 # λ a b c d → a · (b · c) · (b · d) 60 | 61 | phoenix : Term ((b ⇒ c ⇒ d) ⇒ (a ⇒ b) ⇒ (a ⇒ c) ⇒ a ⇒ d) 62 | phoenix = term⁻ $ 4 # λ a b c d → a · (b · d) · (c · d) 63 | 64 | liftM2 : ∀ {α β γ δ} {A : Set α} {B : Set β} {C : Set γ} {D : Set δ} 65 | -> ((B -> C -> D) -> (A -> B) -> (A -> C) -> A -> D) 66 | liftM2 = eval phoenix 67 | 68 | -- With the old version consumed ~550 MB and typechecked in several minutes on my slow machine. 69 | -- eaglebald : Term ((e ⇒ f ⇒ g) ⇒ (a ⇒ b ⇒ e) ⇒ a ⇒ b ⇒ (c ⇒ d ⇒ f) ⇒ c ⇒ d ⇒ g) 70 | -- eaglebald = term⁻ $ 7 # λ a b c d e f g → a · (b · c · d) · (e · f · g) 71 | -------------------------------------------------------------------------------- /src/STLC/Type.agda: -------------------------------------------------------------------------------- 1 | module STLC.Type where 2 | 3 | open import STLC.Type.Core public 4 | open import STLC.Type.Properties public 5 | -------------------------------------------------------------------------------- /src/STLC/Type/Core.agda: -------------------------------------------------------------------------------- 1 | module STLC.Type.Core where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Lib.Membership 5 | 6 | infixr 6 _⇒_ 7 | infixr 9 _∘ˢ_ 8 | 9 | data Type n : Set where 10 | Var : Fin n -> Type n 11 | _⇒_ : Type n -> Type n -> Type n 12 | 13 | Type⁽⁾ : Set 14 | Type⁽⁾ = Type 0 15 | 16 | ftv-all : ∀ {n} -> Type n -> List (Fin n) 17 | ftv-all (Var i) = i ∷ [] 18 | ftv-all (σ ⇒ τ) = ftv-all σ ++ ftv-all τ 19 | 20 | ftv : ∀ {n} -> Type n -> List (Fin n) 21 | ftv = nub ∘ ftv-all 22 | 23 | Subst : ℕ -> ℕ -> Set 24 | Subst n m = Fin n -> Type m 25 | 26 | SubstInFtv : ∀ {n} -> Type n -> ℕ -> Set 27 | SubstInFtv σ m = ∀ i -> i ∈ ftv σ -> Type m 28 | 29 | [_/_] : ∀ {n} -> Fin n -> Type n -> Fin n -> Type n 30 | [ i / σ ] j = drec (const σ) (const (Var j)) (i ≟ j) 31 | 32 | -- Generalize `Type' to `ITree`, make it an instance of `IMonad' and `REWRITE' by the monad laws? 33 | apply : ∀ {n m} -> Subst n m -> Type n -> Type m 34 | apply Ψ (Var i) = Ψ i 35 | apply Ψ (σ ⇒ τ) = apply Ψ σ ⇒ apply Ψ τ 36 | 37 | _∘ˢ_ : ∀ {n m p} -> Subst m p -> Subst n m -> Subst n p 38 | Φ ∘ˢ Ψ = apply Φ ∘ Ψ 39 | 40 | wkᵗ : ∀ {m n} -> Type n -> Type (n + m) 41 | wkᵗ = apply (Var ∘ inject+ _) 42 | 43 | renᵗ : ∀ {n} m -> Type n -> Type (m + n) 44 | renᵗ m = apply (Var ∘ raise m) 45 | 46 | thickenⱽ : ∀ {n} -> Fin n -> (σ : Type n) -> Maybe (Fin (length (ftv σ))) 47 | thickenⱽ i = lookup-for i ∘ map swap ∘ enumerate ∘ ftv 48 | -------------------------------------------------------------------------------- /src/STLC/Type/Properties.agda: -------------------------------------------------------------------------------- 1 | module STLC.Type.Properties where 2 | 3 | open import STLC.Lib.Prelude 4 | open import STLC.Lib.Membership 5 | open import STLC.Type.Core 6 | 7 | Var-inj : ∀ {n} {i j : Fin n} -> Var i ≡ Var j -> i ≡ j 8 | Var-inj refl = refl 9 | 10 | ⇒-inj : ∀ {n} {σ₁ σ₂ τ₁ τ₂ : Type n} -> σ₁ ⇒ τ₁ ≡ σ₂ ⇒ τ₂ -> σ₁ ≡ σ₂ × τ₁ ≡ τ₂ 11 | ⇒-inj refl = refl , refl 12 | 13 | instance 14 | Type-DecEq : ∀ {n} -> DecEq (Type n) 15 | Type-DecEq = record { _≟_ = _≟ᵗ_ } where 16 | infix 4 _≟ᵗ_ 17 | _≟ᵗ_ : ∀ {n} -> Decidable (_≡_ {A = Type n}) 18 | Var i ≟ᵗ Var j = dcong Var Var-inj (i ≟ j) 19 | σ₁ ⇒ τ₁ ≟ᵗ σ₂ ⇒ τ₂ = dcong₂ _⇒_ ⇒-inj (σ₁ ≟ᵗ σ₂) (τ₁ ≟ᵗ τ₂) 20 | Var i ≟ᵗ σ₂ ⇒ τ₂ = no λ() 21 | σ₁ ⇒ τ₁ ≟ᵗ Var j = no λ() 22 | 23 | apply-apply : ∀ {n m p} {Φ : Subst m p} {Ψ : Subst n m} σ 24 | -> apply Φ (apply Ψ σ) ≡ apply (Φ ∘ˢ Ψ) σ 25 | apply-apply (Var i) = refl 26 | apply-apply (σ ⇒ τ) = cong₂ _⇒_ (apply-apply σ) (apply-apply τ) 27 | 28 | sub-self : ∀ {n σ} -> (i : Fin n) -> [ i / σ ] i ≡ σ 29 | sub-self i rewrite ≟-refl i = refl 30 | 31 | non-sub : ∀ {n σ τ} {i : Fin n} -> i ∉ ftv-all σ -> apply [ i / τ ] σ ≡ σ 32 | non-sub {σ = Var j} {i = i} c with i ≟ j 33 | ... | yes p rewrite p = ⊥-elim (c here) 34 | ... | no d = refl 35 | non-sub {σ = σ ⇒ τ} c = cong₂ _⇒_ (non-sub (c ∘ ∈-++₁)) (non-sub (c ∘ ∈-++₂ (ftv-all σ))) 36 | 37 | sub : ∀ {n} i -> (τ : Type n) -> Maybe (∃ λ (Ψ : Subst n n) -> apply Ψ (Var i) ≡ apply Ψ τ) 38 | sub i σ = drec (const nothing) 39 | (λ c -> just ([ i / σ ] , right (sub-self i) (non-sub c))) 40 | (i ∈? ftv-all σ) 41 | -------------------------------------------------------------------------------- /src/STLC/Type/Unify.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | 3 | module STLC.Type.Unify where 4 | 5 | open import STLC.Lib.Prelude 6 | open import STLC.Type.Core 7 | open import STLC.Type.Properties 8 | 9 | {-# REWRITE apply-apply #-} 10 | 11 | {-# TERMINATING #-} 12 | unify : ∀ {n} -> (σ τ : Type n) -> Maybe (∃ λ (Ψ : Subst n n) -> apply Ψ σ ≡ apply Ψ τ) 13 | unify (Var i) (Var j) = drec (λ p -> just (Var , cong Var p)) (const (sub i (Var j))) (i ≟ j) 14 | unify (Var i) τ = sub i τ 15 | unify σ (Var j) = second sym <$> sub j σ 16 | unify (σ₁ ⇒ τ₁) (σ₂ ⇒ τ₂) = 17 | unify σ₁ σ₂ >>= uncurry λ Ψ p -> 18 | unify (apply Ψ τ₁) (apply Ψ τ₂) >>= uncurry λ Φ q -> 19 | just (Φ ∘ˢ Ψ , cong₂ _⇒_ (cong (apply Φ) p) q) 20 | --------------------------------------------------------------------------------