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

198 lines
4.5 KiB
Idris

module Quox.Syntax.Term.Split
import Quox.Syntax.Term.Base
import Quox.Syntax.Term.Subst
import Quox.No
import Data.Vect
%default total
public export %inline
isLam : Term {} -> Bool
isLam (Lam {}) = True
isLam _ = False
public export
0 NotLam : Pred $ Term {}
NotLam = No . isLam
public export %inline
isDLam : Term {} -> Bool
isDLam (DLam {}) = True
isDLam _ = False
public export
0 NotDLam : Pred $ Term {}
NotDLam = No . isDLam
public export %inline
isPair : Term {} -> Bool
isPair (Pair {}) = True
isPair _ = False
public export
0 NotPair : Pred $ Term {}
NotPair = No . isPair
public export %inline
isApp : Elim {} -> Bool
isApp (_ :@ _) = True
isApp _ = False
public export
0 NotApp : Pred $ Elim {}
NotApp = No . isApp
public export %inline
isDApp : Elim {} -> Bool
isDApp (_ :% _) = True
isDApp _ = False
public export
0 NotDApp : Pred $ Elim {}
NotDApp = No . isDApp
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 = case nchoose $ isApp fun of
Left y => let f :@ a = fun in getArgs' f (a :: args)
Right n => GotArgs {fun, args, notApp = n}
||| 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 []
infixl 9 :%%
||| apply multiple dimension arguments at once
public export %inline
(:%%) : Elim q d n -> List (Dim d) -> Elim q d n
f :%% ss = foldl (:%) f ss
public export
record GetDArgs q d n where
constructor GotDArgs
fun : Elim q d n
args : List (Dim d)
0 notDApp : NotDApp fun
export
getDArgs' : Elim q d n -> List (Dim d) -> GetDArgs q d n
getDArgs' fun args = case nchoose $ isDApp fun of
Left y => let f :% d = fun in getDArgs' f (d :: args)
Right n => GotDArgs {fun, args, notDApp = n}
||| splits a dimension application into its head and arguments. if it's not an
||| d application then the list is just empty
export %inline
getDArgs : Elim q d n -> GetDArgs q d n
getDArgs e = getDArgs' 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
infixr 1 :\\%
public export
dabsN : Vect m Name -> Term q (m + d) n -> Term q d n
dabsN {m = 0} [] s = s
dabsN {m = S m} (x :: xs) s = DLam x $ DUsed $ dabsN xs $
rewrite sym $ plusSuccRightSucc m d in s
public export %inline
(:\\%) : Vect m Name -> Term q (m + d) n -> Term q d n
(:\\%) = dabsN
public export
record GetLams q d n where
constructor GotLams
{0 lams, rest : Nat}
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 -> (0 eq : lams + n = rest) ->
GetLams q d n
getLams' xs s Refl = case nchoose $ isLam s of
Left y => let Lam x body = s in
getLams' (x :: xs) (assert_smaller s body.term) Refl
Right n => GotLams xs s Refl n
export %inline
getLams : Term q d n -> GetLams q d n
getLams s = getLams' [] s Refl
public export
record GetDLams q d n where
constructor GotDLams
{0 lams, rest : Nat}
names : Vect lams Name
body : Term q rest n
0 eq : lams + d = rest
0 notDLam : NotDLam body
export
getDLams' : forall lams, rest.
Vect lams Name -> Term q rest n -> (0 eq : lams + d = rest) ->
GetDLams q d n
getDLams' is s Refl = case nchoose $ isDLam s of
Left y => let DLam i body = s in
getDLams' (i :: is) (assert_smaller s body.term) Refl
Right n => GotDLams is s Refl n
export %inline
getDLams : Term q d n -> GetDLams q d n
getDLams s = getDLams' [] s Refl
public export
record GetPairs q d n where
constructor GotPairs
init : List $ Term q d n
last : Term q d n
notPair : NotPair last
export
getPairs : Term q d n -> GetPairs q d n
getPairs t = case nchoose $ isPair t of
Left y =>
let Pair s t = t; GotPairs ts t np = getPairs t in
GotPairs (s :: ts) t np
Right n => GotPairs [] t n