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
|
2023-02-22 01:45:10 -05:00
|
|
|
|
import public Data.SortedMap
|
|
|
|
|
import public Data.SortedSet
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
|
%default total
|
|
|
|
|
|
|
|
|
|
|
2023-01-22 18:53:34 -05:00
|
|
|
|
public export
|
|
|
|
|
0 TermLike : Type
|
|
|
|
|
TermLike = Type -> Nat -> Nat -> Type
|
|
|
|
|
|
|
|
|
|
public export
|
|
|
|
|
0 TSubstLike : Type
|
|
|
|
|
TSubstLike = Type -> Nat -> Nat -> Nat -> Type
|
|
|
|
|
|
2023-02-22 01:45:10 -05:00
|
|
|
|
public export
|
|
|
|
|
0 TagVal : Type
|
|
|
|
|
TagVal = String
|
2023-01-22 18:53:34 -05:00
|
|
|
|
|
2022-04-23 18:21:30 -04:00
|
|
|
|
infixl 8 :#
|
2023-01-20 20:34:28 -05:00
|
|
|
|
infixl 9 :@, :%
|
2022-04-23 18:21:30 -04:00
|
|
|
|
mutual
|
|
|
|
|
public export
|
2023-01-22 18:53:34 -05:00
|
|
|
|
0 TSubst : TSubstLike
|
2023-01-20 20:34:28 -05:00
|
|
|
|
TSubst q d = Subst $ Elim q d
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
|
||| first argument `q` is quantity type;
|
|
|
|
|
||| second argument `d` is dimension scope size;
|
|
|
|
|
||| third `n` is term scope size
|
2022-04-23 18:21:30 -04:00
|
|
|
|
public export
|
2023-01-22 18:53:34 -05:00
|
|
|
|
data Term : TermLike 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-02-22 01:40:19 -05:00
|
|
|
|
Pi : (qty : q) -> (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-02-22 01:40:19 -05:00
|
|
|
|
Lam : (body : ScopeTerm q d n) -> Term q d n
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-01-26 13:54:46 -05:00
|
|
|
|
||| pair type
|
2023-02-22 01:40:19 -05:00
|
|
|
|
Sig : (fst : Term q d n) -> (snd : ScopeTerm q d n) -> Term q d n
|
2023-01-26 13:54:46 -05:00
|
|
|
|
||| pair value
|
|
|
|
|
Pair : (fst, snd : Term q d n) -> Term q d n
|
|
|
|
|
|
2023-02-22 01:45:10 -05:00
|
|
|
|
||| enumeration type
|
|
|
|
|
Enum : (cases : SortedSet TagVal) -> Term q d n
|
|
|
|
|
||| enumeration value
|
|
|
|
|
Tag : (tag : TagVal) -> Term q d n
|
|
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
|
||| equality type
|
2023-02-22 01:40:19 -05:00
|
|
|
|
Eq : (ty : DScopeTerm q d n) -> (l, r : Term q d n) -> Term q d n
|
2023-01-20 20:34:28 -05:00
|
|
|
|
||| equality term
|
2023-02-22 01:40:19 -05:00
|
|
|
|
DLam : (body : DScopeTerm q d n) -> Term q d n
|
2023-01-20 20:34:28 -05:00
|
|
|
|
|
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-22 18:53:34 -05:00
|
|
|
|
data Elim : TermLike 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
|
|
|
|
|
2023-01-26 13:54:46 -05:00
|
|
|
|
||| pair destruction
|
|
|
|
|
|||
|
2023-02-22 01:40:19 -05:00
|
|
|
|
||| `CasePair 𝜋 𝑒 ([𝑟], 𝐴) ([𝑥, 𝑦], 𝑡)` is
|
2023-03-02 13:56:22 -05:00
|
|
|
|
||| `𝐜𝐚𝐬𝐞 𝜋 · 𝑒 𝐫𝐞𝐭𝐮𝐫𝐧 𝑟 ⇒ 𝐴 𝐨𝐟 { (𝑥, 𝑦). 𝑡 }`
|
2023-01-26 13:54:46 -05:00
|
|
|
|
CasePair : (qty : q) -> (pair : Elim q d n) ->
|
2023-02-22 01:40:19 -05:00
|
|
|
|
(ret : ScopeTerm q d n) ->
|
|
|
|
|
(body : ScopeTermN 2 q d n) ->
|
2023-01-26 13:54:46 -05:00
|
|
|
|
Elim q d n
|
|
|
|
|
|
2023-02-22 01:45:10 -05:00
|
|
|
|
||| enum matching
|
|
|
|
|
CaseEnum : (qty : q) -> (tag : Elim q d n) ->
|
|
|
|
|
(ret : ScopeTerm q d n) ->
|
|
|
|
|
(arms : CaseEnumArms q d n) ->
|
|
|
|
|
Elim q d n
|
|
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
|
||| dim application
|
|
|
|
|
(:%) : (fun : Elim q d n) -> (arg : Dim d) -> 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
|
|
|
|
|
|
|
|
|
public export
|
2023-02-22 01:40:19 -05:00
|
|
|
|
0 CaseEnumArms : TermLike
|
|
|
|
|
CaseEnumArms q d n = SortedMap TagVal (Term q d n)
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-02-22 01:40:19 -05:00
|
|
|
|
||| a scoped term with names
|
2022-04-23 18:21:30 -04:00
|
|
|
|
public export
|
2023-02-22 01:40:19 -05:00
|
|
|
|
record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where
|
|
|
|
|
constructor S
|
|
|
|
|
names : Vect s BaseName
|
|
|
|
|
body : ScopedBody s f n
|
2023-01-22 21:22:50 -05:00
|
|
|
|
|
|
|
|
|
public export
|
2023-02-22 01:40:19 -05:00
|
|
|
|
data ScopedBody : Nat -> (Nat -> Type) -> Nat -> Type where
|
|
|
|
|
Y : (body : f (s + n)) -> ScopedBody s f n
|
|
|
|
|
N : (body : f n) -> ScopedBody s f n
|
2023-01-22 21:22:50 -05:00
|
|
|
|
|
|
|
|
|
public export
|
2023-02-22 01:40:19 -05:00
|
|
|
|
0 ScopeTermN, DScopeTermN : Nat -> TermLike
|
|
|
|
|
ScopeTermN s q d n = Scoped s (Term q d) n
|
|
|
|
|
DScopeTermN s q d n = Scoped s (\d => Term q d n) d
|
|
|
|
|
|
|
|
|
|
public export
|
|
|
|
|
0 ScopeTerm, DScopeTerm : TermLike
|
|
|
|
|
ScopeTerm = ScopeTermN 1
|
2023-01-22 21:22:50 -05:00
|
|
|
|
DScopeTerm = DScopeTermN 1
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
|
%name Term s, t, r
|
|
|
|
|
%name Elim e, f
|
2023-02-22 01:40:19 -05:00
|
|
|
|
%name Scoped body
|
|
|
|
|
%name ScopedBody body
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-02-25 09:24:45 -05:00
|
|
|
|
||| scope which ignores all its binders
|
|
|
|
|
public export %inline
|
|
|
|
|
SN : {s : Nat} -> f n -> Scoped s f n
|
|
|
|
|
SN = S (replicate s "_") . N
|
|
|
|
|
|
|
|
|
|
||| scope which uses its binders
|
|
|
|
|
public export %inline
|
|
|
|
|
SY : Vect s BaseName -> f (s + n) -> Scoped s f n
|
|
|
|
|
SY ns = S ns . Y
|
|
|
|
|
|
|
|
|
|
||| more convenient Pi
|
|
|
|
|
public export %inline
|
|
|
|
|
Pi_ : (qty : q) -> (x : BaseName) ->
|
|
|
|
|
(arg : Term q d n) -> (res : Term q d (S n)) -> Term q d n
|
|
|
|
|
Pi_ {qty, x, arg, res} = Pi {qty, arg, res = S [x] $ Y res}
|
|
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
|
||| non dependent function type
|
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
|
2023-02-25 09:24:45 -05:00
|
|
|
|
Arr {qty, arg, res} = Pi {qty, arg, res = SN res}
|
2022-04-27 14:06:39 -04:00
|
|
|
|
|
2023-02-25 09:24:45 -05:00
|
|
|
|
||| more convenient Sig
|
2023-01-20 20:34:28 -05:00
|
|
|
|
public export %inline
|
2023-02-25 09:24:45 -05:00
|
|
|
|
Sig_ : (x : BaseName) -> (fst : Term q d n) ->
|
|
|
|
|
(snd : Term q d (S n)) -> Term q d n
|
|
|
|
|
Sig_ {x, fst, snd} = Sig {fst, snd = S [x] $ Y snd}
|
2023-01-20 20:34:28 -05:00
|
|
|
|
|
2023-02-12 15:30:08 -05:00
|
|
|
|
||| non dependent pair type
|
|
|
|
|
public export %inline
|
|
|
|
|
And : (fst, snd : Term q d n) -> Term q d n
|
2023-02-25 09:24:45 -05:00
|
|
|
|
And {fst, snd} = Sig {fst, snd = SN snd}
|
|
|
|
|
|
|
|
|
|
||| more convenient Eq
|
|
|
|
|
public export %inline
|
|
|
|
|
Eq_ : (i : BaseName) -> (ty : Term q (S d) n) ->
|
|
|
|
|
(l, r : Term q d n) -> Term q d n
|
|
|
|
|
Eq_ {i, ty, l, r} = Eq {ty = S [i] $ Y ty, l, r}
|
|
|
|
|
|
|
|
|
|
||| non dependent equality type
|
|
|
|
|
public export %inline
|
|
|
|
|
Eq0 : (ty, l, r : Term q d n) -> Term q d n
|
|
|
|
|
Eq0 {ty, l, r} = Eq {ty = SN ty, l, r}
|
2023-02-12 15:30:08 -05: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
|