108 lines
2.7 KiB
Idris
108 lines
2.7 KiB
Idris
module Quox.Syntax.Term.Split
|
|
|
|
import Quox.Syntax.Term.Base
|
|
import Quox.Syntax.Term.Subst
|
|
|
|
import Data.So
|
|
import Data.Vect
|
|
|
|
%default total
|
|
|
|
|
|
public export
|
|
data IsLam : Term {} -> Type where
|
|
ItIsLam : IsLam $ Lam x body
|
|
|
|
public export %inline
|
|
isLam : (s : Term {}) -> Dec (IsLam s)
|
|
isLam (TYPE _) = No $ \case _ impossible
|
|
isLam (Pi {}) = No $ \case _ impossible
|
|
isLam (Lam {}) = Yes ItIsLam
|
|
isLam (E _) = No $ \case _ impossible
|
|
isLam (CloT {}) = No $ \case _ impossible
|
|
isLam (DCloT {}) = No $ \case _ impossible
|
|
|
|
public export %inline
|
|
NotLam : Term {} -> Type
|
|
NotLam = Not . IsLam
|
|
|
|
|
|
public export
|
|
data IsApp : Elim {} -> Type where
|
|
ItIsApp : IsApp $ f :@ s
|
|
|
|
public export %inline
|
|
isApp : (e : Elim {}) -> Dec (IsApp e)
|
|
isApp (F _) = No $ \case _ impossible
|
|
isApp (B _) = No $ \case _ impossible
|
|
isApp (_ :@ _) = Yes ItIsApp
|
|
isApp (_ :# _) = No $ \case _ impossible
|
|
isApp (CloE {}) = No $ \case _ impossible
|
|
isApp (DCloE {}) = No $ \case _ impossible
|
|
|
|
public export
|
|
NotApp : Elim {} -> Type
|
|
NotApp = Not . IsApp
|
|
|
|
|
|
infixl 9 :@@
|
|
||| apply multiple arguments at once
|
|
public export %inline
|
|
(:@@) : Elim q d n -> List (Term q d n) -> Elim q d n
|
|
f :@@ ss = foldl (:@) f ss
|
|
|
|
public export
|
|
record GetArgs q d n where
|
|
constructor GotArgs
|
|
fun : Elim q d n
|
|
args : List (Term q d n)
|
|
0 notApp : NotApp fun
|
|
|
|
export
|
|
getArgs' : Elim q d n -> List (Term q d n) -> GetArgs q d n
|
|
getArgs' fun args with (isApp fun)
|
|
getArgs' (f :@ a) args | Yes yes = getArgs' f (a :: args)
|
|
getArgs' fun args | No no = GotArgs {fun, args, notApp = no}
|
|
|
|
||| splits an application into its head and arguments. if it's not an
|
|
||| application then the list is just empty
|
|
export %inline
|
|
getArgs : Elim q d n -> GetArgs q d n
|
|
getArgs e = getArgs' e []
|
|
|
|
|
|
infixr 1 :\\
|
|
public export
|
|
absN : Vect m Name -> Term q d (m + n) -> Term q d n
|
|
absN {m = 0} [] s = s
|
|
absN {m = S m} (x :: xs) s = Lam x $ TUsed $ absN xs $
|
|
rewrite sym $ plusSuccRightSucc m n in s
|
|
|
|
public export %inline
|
|
(:\\) : Vect m Name -> Term q d (m + n) -> Term q d n
|
|
(:\\) = absN
|
|
|
|
|
|
public export
|
|
record GetLams q d n where
|
|
constructor GotLams
|
|
names : Vect lams Name
|
|
body : Term q d rest
|
|
0 eq : lams + n = rest
|
|
0 notLam : NotLam body
|
|
|
|
export
|
|
getLams' : forall lams, rest.
|
|
Vect lams Name -> Term q d rest -> lams + n = rest -> GetLams q d n
|
|
getLams' xs s eq with (isLam s)
|
|
getLams' xs (Lam x sbody) Refl | Yes ItIsLam =
|
|
let 0 s = Lam x sbody
|
|
body = assert_smaller s $ fromScopeTerm sbody
|
|
in
|
|
getLams' (x :: xs) body Refl
|
|
getLams' xs s eq | No no =
|
|
GotLams xs s eq no
|
|
|
|
export
|
|
getLams : Term q d n -> GetLams q d n
|
|
getLams s = getLams' [] s Refl
|