quox/lib/Quox/Syntax/Term/Base.idr

115 lines
3.2 KiB
Idris
Raw Normal View History

2022-04-23 18:21:30 -04:00
module Quox.Syntax.Term.Base
import public Quox.Syntax.Var
import public Quox.Syntax.Shift
import public Quox.Syntax.Subst
import public Quox.Syntax.Universe
import public Quox.Syntax.Qty
import public Quox.Syntax.Dim
import public Quox.Name
2023-01-08 09:43:54 -05:00
-- import public Quox.OPE
2022-04-23 18:21:30 -04:00
import Quox.Pretty
import public Data.DPair
import Data.List
import Data.Maybe
import Data.Nat
import public Data.So
import Data.String
import Data.Vect
%default total
infixl 8 :#
infixl 9 :@
mutual
public export
2023-01-08 14:44:25 -05:00
TSubst : Type -> Nat -> Nat -> Nat -> Type
TSubst q d = Subst (\n => Elim q d n)
2022-04-23 18:21:30 -04:00
||| first argument `d` is dimension scope size, second `n` is term scope size
public export
2023-01-08 14:44:25 -05:00
data Term : (q : Type) -> (d, n : Nat) -> Type where
2022-04-23 18:21:30 -04:00
||| type of types
2023-01-08 14:44:25 -05:00
TYPE : (l : Universe) -> Term q d n
2022-04-23 18:21:30 -04:00
||| function type
2023-01-08 14:44:25 -05:00
Pi : (qty : q) -> (x : Name) ->
(arg : Term q d n) -> (res : ScopeTerm q d n) -> Term q d n
2022-04-23 18:21:30 -04:00
||| function term
2023-01-08 14:44:25 -05:00
Lam : (x : Name) -> (body : ScopeTerm q d n) -> Term q d n
2022-04-23 18:21:30 -04:00
||| elimination
2023-01-08 14:44:25 -05:00
E : (e : Elim q d n) -> Term q d n
2022-04-23 18:21:30 -04:00
||| term closure/suspended substitution
2023-01-08 14:44:25 -05:00
CloT : (tm : Term q d from) -> (th : Lazy (TSubst q d from to)) ->
Term q d to
2022-04-23 18:21:30 -04:00
||| dimension closure/suspended substitution
2023-01-08 14:44:25 -05:00
DCloT : (tm : Term q dfrom n) -> (th : Lazy (DSubst dfrom dto)) ->
Term q dto n
2022-04-23 18:21:30 -04:00
||| first argument `d` is dimension scope size, second `n` is term scope size
public export
2023-01-08 14:44:25 -05:00
data Elim : (q : Type) -> (d, n : Nat) -> Type where
2022-04-23 18:21:30 -04:00
||| free variable
2023-01-08 14:44:25 -05:00
F : (x : Name) -> Elim q d n
2022-04-23 18:21:30 -04:00
||| bound variable
2023-01-08 14:44:25 -05:00
B : (i : Var n) -> Elim q d n
2022-04-23 18:21:30 -04:00
||| term application
2023-01-08 14:44:25 -05:00
(:@) : (fun : Elim q d n) -> (arg : Term q d n) -> Elim q d n
2022-04-23 18:21:30 -04:00
||| type-annotated term
2023-01-08 14:44:25 -05:00
(:#) : (tm, ty : Term q d n) -> Elim q d n
2022-04-23 18:21:30 -04:00
||| term closure/suspended substitution
2023-01-08 14:44:25 -05:00
CloE : (el : Elim q d from) -> (th : Lazy (TSubst q d from to)) ->
Elim q d to
2022-04-23 18:21:30 -04:00
||| dimension closure/suspended substitution
2023-01-08 14:44:25 -05:00
DCloE : (el : Elim q dfrom n) -> (th : Lazy (DSubst dfrom dto)) ->
Elim q dto n
2022-04-23 18:21:30 -04:00
||| a scope with one more bound variable
public export
2023-01-08 14:44:25 -05:00
data ScopeTerm : (q : Type) -> (d, n : Nat) -> Type where
2022-04-23 18:21:30 -04:00
||| variable is used
2023-01-08 14:44:25 -05:00
TUsed : (body : Term q d (S n)) -> ScopeTerm q d n
2022-04-23 18:21:30 -04:00
||| variable is unused
2023-01-08 14:44:25 -05:00
TUnused : (body : Term q d n) -> ScopeTerm q d n
2022-04-23 18:21:30 -04:00
||| a scope with one more bound dimension variable
public export
2023-01-08 14:44:25 -05:00
data DScopeTerm : (q : Type) -> (d, n : Nat) -> Type where
2022-04-23 18:21:30 -04:00
||| variable is used
2023-01-08 14:44:25 -05:00
DUsed : (body : Term q (S d) n) -> DScopeTerm q d n
2022-04-23 18:21:30 -04:00
||| variable is unused
2023-01-08 14:44:25 -05:00
DUnused : (body : Term q d n) -> DScopeTerm q d n
2022-04-23 18:21:30 -04:00
%name Term s, t, r
%name Elim e, f
%name ScopeTerm body
%name DScopeTerm body
2022-04-27 14:06:39 -04:00
public export %inline
2023-01-08 14:44:25 -05:00
Arr : (qty : q) -> (arg, res : Term q d n) -> Term q d n
2022-05-25 10:03:23 -04:00
Arr {qty, arg, res} = Pi {qty, x = "_", arg, res = TUnused res}
2022-04-27 14:06:39 -04:00
2022-04-23 18:21:30 -04:00
||| same as `F` but as a term
public export %inline
2023-01-08 14:44:25 -05:00
FT : Name -> Term q d n
2022-04-23 18:21:30 -04:00
FT = E . F
||| abbreviation for a bound variable like `BV 4` instead of
||| `B (VS (VS (VS (VS VZ))))`
public export %inline
2023-01-08 14:44:25 -05:00
BV : (i : Nat) -> (0 _ : LT i n) => Elim q d n
2022-04-23 18:21:30 -04:00
BV i = B $ V i
||| same as `BV` but as a term
public export %inline
2023-01-08 14:44:25 -05:00
BVT : (i : Nat) -> (0 _ : LT i n) => Term q d n
2022-04-23 18:21:30 -04:00
BVT i = E $ BV i