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

247 lines
5.8 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 Quox.Context
2022-04-23 18:21:30 -04:00
2023-02-11 12:14:12 -05:00
import public Quox.No
import public Data.Vect
2022-04-23 18:21:30 -04:00
%default total
public export %inline
2023-01-22 21:39:46 -05:00
isLam : Term {} -> Bool
2023-02-22 01:40:19 -05:00
isLam (Lam _) = True
isLam _ = False
2022-04-23 18:21:30 -04:00
2023-01-20 20:34:28 -05:00
public export
0 NotLam : Pred $ Term {}
2023-01-22 21:39:46 -05:00
NotLam = No . isLam
2022-04-23 18:21:30 -04:00
2023-01-20 20:34:28 -05:00
public export %inline
2023-01-22 21:39:46 -05:00
isDLam : Term {} -> Bool
2023-02-22 01:40:19 -05:00
isDLam (DLam _) = True
isDLam _ = False
2023-01-20 20:34:28 -05:00
public export
0 NotDLam : Pred $ Term {}
2023-01-22 21:39:46 -05:00
NotDLam = No . isDLam
2023-01-20 20:34:28 -05:00
2023-01-26 13:54:46 -05:00
public export %inline
isPair : Term {} -> Bool
isPair (Pair {}) = True
isPair _ = False
public export
0 NotPair : Pred $ Term {}
NotPair = No . isPair
2022-04-23 18:21:30 -04:00
public export %inline
2023-01-22 21:39:46 -05:00
isApp : Elim {} -> Bool
isApp (_ :@ _) = True
isApp _ = False
2022-04-23 18:21:30 -04:00
public export
2023-01-20 20:34:28 -05:00
0 NotApp : Pred $ Elim {}
2023-01-22 21:39:46 -05:00
NotApp = No . isApp
2022-04-23 18:21:30 -04:00
2023-01-20 20:34:28 -05:00
public export %inline
2023-01-22 21:39:46 -05:00
isDApp : Elim {} -> Bool
isDApp (_ :% _) = True
isDApp _ = False
2023-01-20 20:34:28 -05:00
public export
0 NotDApp : Pred $ Elim {}
2023-01-22 21:39:46 -05:00
NotDApp = No . isDApp
2023-01-20 20:34:28 -05:00
2022-04-23 18:21:30 -04:00
infixl 9 :@@
||| apply multiple arguments at once
public export %inline
2023-04-01 13:16:43 -04:00
(:@@) : Elim d n -> List (Term d n) -> Elim d n
2022-04-23 18:21:30 -04:00
f :@@ ss = foldl (:@) f ss
public export
2023-04-01 13:16:43 -04:00
record GetArgs d n where
2022-04-23 18:21:30 -04:00
constructor GotArgs
2023-04-01 13:16:43 -04:00
fun : Elim d n
args : List (Term d n)
2022-04-23 18:21:30 -04:00
0 notApp : NotApp fun
mutual
export %inline
2023-04-01 13:16:43 -04:00
getArgs' : Elim d n -> List (Term d n) -> GetArgs d n
getArgs' fun0 args =
let Element fun nc = pushSubsts fun0 in
getArgsNc (assert_smaller fun0 fun) args
private
2023-04-01 13:16:43 -04:00
getArgsNc : (e : Elim d n) -> (0 nc : NotClo e) =>
List (Term d n) -> GetArgs d n
getArgsNc 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}
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.
||| looks through substitutions for applications.
2022-04-23 18:21:30 -04:00
export %inline
2023-04-01 13:16:43 -04:00
getArgs : Elim d n -> GetArgs d n
2022-04-23 18:21:30 -04:00
getArgs e = getArgs' e []
2023-01-20 20:34:28 -05:00
infixl 9 :%%
||| apply multiple dimension arguments at once
public export %inline
2023-04-01 13:16:43 -04:00
(:%%) : Elim d n -> List (Dim d) -> Elim d n
2023-01-20 20:34:28 -05:00
f :%% ss = foldl (:%) f ss
public export
2023-04-01 13:16:43 -04:00
record GetDArgs d n where
2023-01-20 20:34:28 -05:00
constructor GotDArgs
2023-04-01 13:16:43 -04:00
fun : Elim d n
2023-01-20 20:34:28 -05:00
args : List (Dim d)
0 notDApp : NotDApp fun
mutual
export %inline
2023-04-01 13:16:43 -04:00
getDArgs' : Elim d n -> List (Dim d) -> GetDArgs d n
getDArgs' fun0 args =
let Element fun nc = pushSubsts fun0 in
getDArgsNc (assert_smaller fun0 fun) args
private
2023-04-01 13:16:43 -04:00
getDArgsNc : (e : Elim d n) -> (0 nc : NotClo e) =>
List (Dim d) -> GetDArgs d n
getDArgsNc 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}
2023-01-20 20:34:28 -05:00
||| splits a dimension application into its head and arguments. if it's not an
||| d application then the list is just empty
export %inline
2023-04-01 13:16:43 -04:00
getDArgs : Elim d n -> GetDArgs d n
2023-01-20 20:34:28 -05:00
getDArgs e = getDArgs' e []
2022-04-23 18:21:30 -04:00
infixr 1 :\\
public export
2023-04-01 13:16:43 -04:00
absN : NContext m -> Term d (m + n) -> Term d n
absN [<] s = s
absN (xs :< x) s = absN xs $ Lam $ SY [< x] s
2023-01-08 14:44:25 -05:00
public export %inline
2023-04-01 13:16:43 -04:00
(:\\) : NContext m -> Term d (m + n) -> Term d n
2023-01-08 14:44:25 -05:00
(:\\) = absN
2022-04-23 18:21:30 -04:00
2023-01-20 20:34:28 -05:00
infixr 1 :\\%
public export
2023-04-01 13:16:43 -04:00
dabsN : NContext m -> Term (m + d) n -> Term d n
dabsN [<] s = s
dabsN (xs :< x) s = dabsN xs $ DLam $ SY [< x] s
2023-01-20 20:34:28 -05:00
public export %inline
2023-04-01 13:16:43 -04:00
(:\\%) : NContext m -> Term (m + d) n -> Term d n
2023-01-20 20:34:28 -05:00
(:\\%) = dabsN
2022-04-23 18:21:30 -04:00
public export
2023-04-01 13:16:43 -04:00
record GetLams d n where
2022-04-23 18:21:30 -04:00
constructor GotLams
2023-01-20 20:34:28 -05:00
{0 lams, rest : Nat}
names : NContext lams
2023-04-01 13:16:43 -04:00
body : Term d rest
2022-04-23 18:21:30 -04:00
0 eq : lams + n = rest
0 notLam : NotLam body
mutual
export %inline
getLams' : forall lams, rest.
2023-04-01 13:16:43 -04:00
NContext lams -> Term d rest -> (0 eq : lams + n = rest) ->
GetLams d n
getLams' xs s0 eq =
let Element s nc = pushSubsts s0 in
getLamsNc xs (assert_smaller s0 s) eq
private
getLamsNc : forall lams, rest.
NContext lams ->
2023-04-01 13:16:43 -04:00
(t : Term d rest) -> (0 nc : NotClo t) =>
(0 eq : lams + n = rest) ->
2023-04-01 13:16:43 -04:00
GetLams d n
getLamsNc xs s Refl = case nchoose $ isLam s of
Left y => let Lam (S [< x] body) = s in
getLams' (xs :< x) (assert_smaller s body.term) Refl
Right n => GotLams xs s Refl n
2023-01-08 14:44:25 -05:00
2023-01-20 20:34:28 -05:00
export %inline
2023-04-01 13:16:43 -04:00
getLams : Term d n -> GetLams d n
getLams s = getLams' [<] s Refl
2023-01-20 20:34:28 -05:00
public export
2023-04-01 13:16:43 -04:00
record GetDLams d n where
2023-01-20 20:34:28 -05:00
constructor GotDLams
{0 lams, rest : Nat}
names : NContext lams
2023-04-01 13:16:43 -04:00
body : Term rest n
2023-01-20 20:34:28 -05:00
0 eq : lams + d = rest
0 notDLam : NotDLam body
mutual
export %inline
getDLams' : forall lams, rest.
2023-04-01 13:16:43 -04:00
NContext lams -> Term rest n -> (0 eq : lams + d = rest) ->
GetDLams d n
getDLams' xs s0 eq =
let Element s nc = pushSubsts s0 in
getDLamsNc xs (assert_smaller s0 s) eq
private
getDLamsNc : forall lams, rest.
NContext lams ->
2023-04-01 13:16:43 -04:00
(t : Term rest n) -> (0 nc : NotClo t) =>
(0 eq : lams + d = rest) ->
2023-04-01 13:16:43 -04:00
GetDLams d n
getDLamsNc is s Refl = case nchoose $ isDLam s of
Left y => let DLam (S [< i] body) = s in
getDLams' (is :< i) (assert_smaller s body.term) Refl
Right n => GotDLams is s Refl n
2023-01-20 20:34:28 -05:00
export %inline
2023-04-01 13:16:43 -04:00
getDLams : Term d n -> GetDLams d n
getDLams s = getDLams' [<] s Refl
2023-01-26 13:54:46 -05:00
public export
2023-04-01 13:16:43 -04:00
record GetPairs d n where
2023-01-26 13:54:46 -05:00
constructor GotPairs
2023-04-01 13:16:43 -04:00
init : SnocList $ Term d n
last : Term d n
2023-01-26 13:54:46 -05:00
notPair : NotPair last
mutual
export %inline
2023-04-01 13:16:43 -04:00
getPairs' : SnocList (Term d n) -> Term d n -> GetPairs d n
getPairs' ss t0 =
let Element t nc = pushSubsts t0 in getPairsNc ss (assert_smaller t0 t)
private
2023-04-01 13:16:43 -04:00
getPairsNc : SnocList (Term d n) ->
(t : Term d n) -> (0 nc : NotClo t) =>
GetPairs d n
getPairsNc ss t = case nchoose $ isPair t of
Left y => let Pair s t = t in
getPairs' (ss :< s) t
Right n => GotPairs ss t n
2023-01-26 13:54:46 -05:00
export
2023-04-01 13:16:43 -04:00
getPairs : Term d n -> GetPairs d n
getPairs = getPairs' [<]