SimpleTypes.agda 4.9 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174
  1. {-
  2. A simple bidirectional type checker for simply typed lambda calculus which is
  3. sound by construction.
  4. -}
  5. module SimpleTypes where
  6. infix 10 _==_
  7. data _==_ {A : Set}(x : A) : A -> Set where
  8. refl : x == x
  9. data Maybe (A : Set) : Set where
  10. nothing : Maybe A
  11. just : A -> Maybe A
  12. data Nat : Set where
  13. zero : Nat
  14. suc : Nat -> Nat
  15. data Fin : Nat -> Set where
  16. fzero : {n : Nat} -> Fin (suc n)
  17. fsuc : {n : Nat} -> Fin n -> Fin (suc n)
  18. data List (A : Set) : Set where
  19. ε : List A
  20. _,_ : List A -> A -> List A
  21. length : forall {A} -> List A -> Nat
  22. length ε = zero
  23. length (xs , x) = suc (length xs)
  24. infixl 25 _,_
  25. -- Raw terms
  26. data Expr : Set where
  27. varʳ : Nat -> Expr
  28. _•ʳ_ : Expr -> Expr -> Expr
  29. ƛʳ_ : Expr -> Expr
  30. infixl 90 _•ʳ_
  31. infix 50 ƛʳ_
  32. -- Types
  33. data Type : Set where
  34. ι : Type
  35. _⟶_ : Type -> Type -> Type
  36. infixr 40 _⟶_
  37. -- Typed terms
  38. Ctx = List Type
  39. data Var : Ctx -> Type -> Set where
  40. vz : forall {Γ τ} -> Var (Γ , τ) τ
  41. vs : forall {Γ τ σ} -> Var Γ τ -> Var (Γ , σ) τ
  42. data Term : Ctx -> Type -> Set where
  43. var : forall {Γ τ} -> Var Γ τ -> Term Γ τ
  44. _•_ : forall {Γ τ σ} -> Term Γ (τ ⟶ σ) -> Term Γ τ -> Term Γ σ
  45. ƛ_ : forall {Γ τ σ} -> Term (Γ , σ) τ -> Term Γ (σ ⟶ τ)
  46. infixl 90 _•_
  47. infix 50 ƛ_
  48. -- Type erasure
  49. ⌊_⌋ˣ : forall {Γ τ} -> Var Γ τ -> Nat
  50. ⌊ vz ⌋ˣ = zero
  51. ⌊ vs x ⌋ˣ = suc ⌊ x ⌋ˣ
  52. ⌊_⌋ : forall {Γ τ} -> Term Γ τ -> Expr
  53. ⌊ var v ⌋ = varʳ ⌊ v ⌋ˣ
  54. ⌊ s • t ⌋ = ⌊ s ⌋ •ʳ ⌊ t ⌋
  55. ⌊ ƛ t ⌋ = ƛʳ ⌊ t ⌋
  56. -- Type equality
  57. infix 30 _≟_
  58. _≟_ : (σ τ : Type) -> Maybe (σ == τ)
  59. ι ≟ ι = just refl
  60. σ₁ ⟶ τ₁ ≟ σ₂ ⟶ τ₂ with σ₁ ≟ σ₂ | τ₁ ≟ τ₂
  61. σ ⟶ τ ≟ .σ ⟶ .τ | just refl | just refl = just refl
  62. _ ⟶ _ ≟ _ ⟶ _ | _ | _ = nothing
  63. _ ≟ _ = nothing
  64. -- The type checked view
  65. -- ok : forall {Γ τ e} -> Check ⌊ e ⌋ -- unsolved metas with no range!
  66. data Check (Γ : Ctx)(τ : Type) : Expr -> Set where
  67. ok : (t : Term Γ τ) -> Check Γ τ ⌊ t ⌋
  68. bad : {e : Expr} -> Check Γ τ e
  69. data Infer (Γ : Ctx) : Expr -> Set where
  70. yes : (τ : Type)(t : Term Γ τ) -> Infer Γ ⌊ t ⌋
  71. no : {e : Expr} -> Infer Γ e
  72. data Lookup (Γ : Ctx) : Nat -> Set where
  73. found : (τ : Type)(x : Var Γ τ) -> Lookup Γ ⌊ x ⌋ˣ
  74. outofscope : {n : Nat} -> Lookup Γ n
  75. lookup : (Γ : Ctx)(n : Nat) -> Lookup Γ n
  76. lookup ε n = outofscope
  77. lookup (Γ , τ) zero = found τ vz
  78. lookup (Γ , σ) (suc n) with lookup Γ n
  79. lookup (Γ , σ) (suc .(⌊ x ⌋ˣ)) | found τ x = found τ (vs x)
  80. lookup (Γ , σ) (suc n) | outofscope = outofscope
  81. infix 20 _⊢_∋_ _⊢_∈
  82. _⊢_∈ : (Γ : Ctx)(e : Expr) -> Infer Γ e
  83. _⊢_∋_ : (Γ : Ctx)(τ : Type)(e : Expr) -> Check Γ τ e
  84. Γ ⊢ ι ∋ ƛʳ e = bad
  85. Γ ⊢ (σ ⟶ τ) ∋ ƛʳ e with Γ , σ ⊢ τ ∋ e
  86. Γ ⊢ (σ ⟶ τ) ∋ ƛʳ .(⌊ t ⌋) | ok t = ok (ƛ t)
  87. Γ ⊢ (σ ⟶ τ) ∋ ƛʳ _ | bad = bad
  88. Γ ⊢ τ ∋ e with Γ ⊢ e ∈
  89. Γ ⊢ τ ∋ .(⌊ t ⌋) | yes σ t with τ ≟ σ
  90. Γ ⊢ τ ∋ .(⌊ t ⌋) | yes .τ t | just refl = ok t
  91. Γ ⊢ τ ∋ .(⌊ t ⌋) | yes σ t | nothing = bad
  92. Γ ⊢ τ ∋ e | no = bad
  93. Γ ⊢ varʳ i ∈ with lookup Γ i
  94. Γ ⊢ varʳ .(⌊ x ⌋ˣ) ∈ | found τ x = yes τ (var x)
  95. Γ ⊢ varʳ _ ∈ | outofscope = no
  96. Γ ⊢ e₁ •ʳ e₂ ∈ with Γ ⊢ e₁ ∈
  97. Γ ⊢ e₁ •ʳ e₂ ∈ | no = no
  98. Γ ⊢ .(⌊ t₁ ⌋) •ʳ e₂ ∈ | yes ι t₁ = no
  99. Γ ⊢ .(⌊ t₁ ⌋) •ʳ e₂ ∈ | yes (σ ⟶ τ) t₁ with Γ ⊢ σ ∋ e₂
  100. Γ ⊢ .(⌊ t₁ ⌋) •ʳ .(⌊ t₂ ⌋) ∈ | yes (σ ⟶ τ) t₁ | ok t₂ = yes τ (t₁ • t₂)
  101. Γ ⊢ .(⌊ t₁ ⌋) •ʳ _ ∈ | yes (σ ⟶ τ) t₁ | bad = no
  102. Γ ⊢ ƛʳ e ∈ = no
  103. -- Proving completeness (for normal terms)
  104. -- Needs magic with
  105. {-
  106. mutual
  107. data Nf : forall {Γ τ} -> Term Γ τ -> Set where
  108. ƛ-nf : forall {Γ σ τ} -> {t : Term (Γ , σ) τ} -> Nf t -> Nf (ƛ t)
  109. ne-nf : forall {Γ τ} -> {t : Term Γ τ} -> Ne t -> Nf t
  110. data Ne : forall {Γ τ} -> Term Γ τ -> Set where
  111. •-ne : forall {Γ σ τ} ->
  112. {t₁ : Term Γ (σ ⟶ τ)} -> Ne t₁ ->
  113. {t₂ : Term Γ σ} -> Nf t₂ -> Ne (t₁ • t₂)
  114. var-ne : forall {Γ τ} -> {x : Var Γ τ} -> Ne (var x)
  115. mutual
  116. complete-check : forall {Γ τ} -> (t : Term Γ τ) -> Nf t ->
  117. Γ ⊢ τ ∋ ⌊ t ⌋ == ok t
  118. complete-check ._ (ƛ-nf t) = {! !}
  119. complete-check _ (ne-nf n) with complete-infer _ n
  120. complete-check t (ne-nf n) | p = {! !}
  121. complete-infer : forall {Γ τ} -> (t : Term Γ τ) -> Ne t ->
  122. Γ ⊢ ⌊ t ⌋ ∈ == yes τ t
  123. complete-infer t ne = {! !}
  124. -}
  125. -- Testing
  126. test1 = ε ⊢ ι ⟶ ι ∋ ƛʳ varʳ zero
  127. test2 = ε , ι , ι ⟶ ι ⊢ varʳ zero •ʳ varʳ (suc zero) ∈