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

109 lines
2.7 KiB
Idris
Raw Normal View History

2022-04-23 18:21:30 -04:00
module Quox.Syntax.Term.Split
import Quox.Syntax.Term.Base
import Quox.Syntax.Term.Subst
import Data.So
import Data.Vect
%default total
2023-01-08 14:44:25 -05:00
public export
data IsLam : Term {} -> Type where
ItIsLam : IsLam $ Lam x body
2022-04-23 18:21:30 -04:00
public export %inline
2023-01-08 14:44:25 -05:00
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
2022-04-23 18:21:30 -04:00
2023-01-08 14:44:25 -05:00
public export %inline
NotLam : Term {} -> Type
NotLam = Not . IsLam
2022-04-23 18:21:30 -04:00
2023-01-08 14:44:25 -05:00
public export
data IsApp : Elim {} -> Type where
ItIsApp : IsApp $ f :@ s
2022-04-23 18:21:30 -04:00
public export %inline
2023-01-08 14:44:25 -05:00
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
2022-04-23 18:21:30 -04:00
public export
2023-01-08 14:44:25 -05:00
NotApp : Elim {} -> Type
NotApp = Not . IsApp
2022-04-23 18:21:30 -04:00
infixl 9 :@@
||| apply multiple arguments at once
public export %inline
2023-01-08 14:44:25 -05:00
(:@@) : Elim q d n -> List (Term q d n) -> Elim q d n
2022-04-23 18:21:30 -04:00
f :@@ ss = foldl (:@) f ss
public export
2023-01-08 14:44:25 -05:00
record GetArgs q d n where
2022-04-23 18:21:30 -04:00
constructor GotArgs
2023-01-08 14:44:25 -05:00
fun : Elim q d n
args : List (Term q d n)
2022-04-23 18:21:30 -04:00
0 notApp : NotApp fun
export
2023-01-08 14:44:25 -05:00
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}
2022-04-23 18:21:30 -04:00
||| splits an application into its head and arguments. if it's not an
||| application then the list is just empty
export %inline
2023-01-08 14:44:25 -05:00
getArgs : Elim q d n -> GetArgs q d n
2022-04-23 18:21:30 -04:00
getArgs e = getArgs' e []
infixr 1 :\\
public export
2023-01-08 14:44:25 -05:00
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
2022-04-23 18:21:30 -04:00
public export
2023-01-08 14:44:25 -05:00
record GetLams q d n where
2022-04-23 18:21:30 -04:00
constructor GotLams
names : Vect lams Name
2023-01-08 14:44:25 -05:00
body : Term q d rest
2022-04-23 18:21:30 -04:00
0 eq : lams + n = rest
0 notLam : NotLam body
2023-01-08 14:44:25 -05:00
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