replace Split stuff with bools
This commit is contained in:
parent
f0f49d9abf
commit
6073ab4705
1 changed files with 31 additions and 67 deletions
|
@ -3,84 +3,50 @@ module Quox.Syntax.Term.Split
|
||||||
import Quox.Syntax.Term.Base
|
import Quox.Syntax.Term.Base
|
||||||
import Quox.Syntax.Term.Subst
|
import Quox.Syntax.Term.Subst
|
||||||
|
|
||||||
import Data.So
|
import Quox.No
|
||||||
import Data.Vect
|
import Data.Vect
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
|
|
||||||
|
|
||||||
public export
|
|
||||||
data IsLam : Pred $ Term {} where ItIsLam : IsLam $ Lam x body
|
|
||||||
|
|
||||||
public export %inline
|
public export %inline
|
||||||
isLam : Dec1 IsLam
|
isLam : Term {} -> Bool
|
||||||
isLam (TYPE _) = No $ \case _ impossible
|
isLam (Lam {}) = True
|
||||||
isLam (Pi {}) = No $ \case _ impossible
|
isLam _ = False
|
||||||
isLam (Lam {}) = Yes ItIsLam
|
|
||||||
isLam (Eq {}) = No $ \case _ impossible
|
|
||||||
isLam (DLam {}) = No $ \case _ impossible
|
|
||||||
isLam (E _) = No $ \case _ impossible
|
|
||||||
isLam (CloT {}) = No $ \case _ impossible
|
|
||||||
isLam (DCloT {}) = No $ \case _ impossible
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 NotLam : Pred $ Term {}
|
0 NotLam : Pred $ Term {}
|
||||||
NotLam = Not . IsLam
|
NotLam = No . isLam
|
||||||
|
|
||||||
|
|
||||||
public export
|
|
||||||
data IsDLam : Pred $ Term {} where ItIsDLam : IsDLam $ DLam i body
|
|
||||||
|
|
||||||
public export %inline
|
public export %inline
|
||||||
isDLam : Dec1 IsDLam
|
isDLam : Term {} -> Bool
|
||||||
isDLam (TYPE _) = No $ \case _ impossible
|
isDLam (DLam {}) = True
|
||||||
isDLam (Pi {}) = No $ \case _ impossible
|
isDLam _ = False
|
||||||
isDLam (Eq {}) = No $ \case _ impossible
|
|
||||||
isDLam (Lam {}) = No $ \case _ impossible
|
|
||||||
isDLam (DLam {}) = Yes ItIsDLam
|
|
||||||
isDLam (E _) = No $ \case _ impossible
|
|
||||||
isDLam (CloT {}) = No $ \case _ impossible
|
|
||||||
isDLam (DCloT {}) = No $ \case _ impossible
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 NotDLam : Pred $ Term {}
|
0 NotDLam : Pred $ Term {}
|
||||||
NotDLam = Not . IsDLam
|
NotDLam = No . isDLam
|
||||||
|
|
||||||
|
|
||||||
public export
|
|
||||||
data IsApp : Pred $ Elim {} where ItIsApp : IsApp $ f :@ s
|
|
||||||
|
|
||||||
public export %inline
|
public export %inline
|
||||||
isApp : Dec1 IsApp
|
isApp : Elim {} -> Bool
|
||||||
isApp (F _) = No $ \case _ impossible
|
isApp (_ :@ _) = True
|
||||||
isApp (B _) = No $ \case _ impossible
|
isApp _ = False
|
||||||
isApp (_ :@ _) = Yes ItIsApp
|
|
||||||
isApp (_ :% _) = No $ \case _ impossible
|
|
||||||
isApp (_ :# _) = No $ \case _ impossible
|
|
||||||
isApp (CloE {}) = No $ \case _ impossible
|
|
||||||
isApp (DCloE {}) = No $ \case _ impossible
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 NotApp : Pred $ Elim {}
|
0 NotApp : Pred $ Elim {}
|
||||||
NotApp = Not . IsApp
|
NotApp = No . isApp
|
||||||
|
|
||||||
|
|
||||||
public export
|
|
||||||
data IsDApp : Pred $ Elim {} where ItIsDApp : IsDApp $ f :% d
|
|
||||||
|
|
||||||
public export %inline
|
public export %inline
|
||||||
isDApp : Dec1 IsDApp
|
isDApp : Elim {} -> Bool
|
||||||
isDApp (F _) = No $ \case _ impossible
|
isDApp (_ :% _) = True
|
||||||
isDApp (B _) = No $ \case _ impossible
|
isDApp _ = False
|
||||||
isDApp (_ :@ _) = No $ \case _ impossible
|
|
||||||
isDApp (_ :% _) = Yes ItIsDApp
|
|
||||||
isDApp (_ :# _) = No $ \case _ impossible
|
|
||||||
isDApp (CloE {}) = No $ \case _ impossible
|
|
||||||
isDApp (DCloE {}) = No $ \case _ impossible
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 NotDApp : Pred $ Elim {}
|
0 NotDApp : Pred $ Elim {}
|
||||||
NotDApp = Not . IsDApp
|
NotDApp = No . isDApp
|
||||||
|
|
||||||
|
|
||||||
infixl 9 :@@
|
infixl 9 :@@
|
||||||
|
@ -98,9 +64,9 @@ record GetArgs q d n where
|
||||||
|
|
||||||
export
|
export
|
||||||
getArgs' : Elim q d n -> List (Term q d n) -> GetArgs q d n
|
getArgs' : Elim q d n -> List (Term q d n) -> GetArgs q d n
|
||||||
getArgs' fun args with (isApp fun)
|
getArgs' fun args = case nchoose $ isApp fun of
|
||||||
getArgs' (f :@ a) args | Yes _ = getArgs' f (a :: args)
|
Left y => let f :@ a = fun in getArgs' f (a :: args)
|
||||||
getArgs' fun args | No no = GotArgs {fun, args, notApp = no}
|
Right n => GotArgs {fun, args, notApp = n}
|
||||||
|
|
||||||
||| splits an application into its head and arguments. if it's not an
|
||| splits an application into its head and arguments. if it's not an
|
||||||
||| application then the list is just empty
|
||| application then the list is just empty
|
||||||
|
@ -124,9 +90,9 @@ record GetDArgs q d n where
|
||||||
|
|
||||||
export
|
export
|
||||||
getDArgs' : Elim q d n -> List (Dim d) -> GetDArgs q d n
|
getDArgs' : Elim q d n -> List (Dim d) -> GetDArgs q d n
|
||||||
getDArgs' fun args with (isDApp fun)
|
getDArgs' fun args = case nchoose $ isDApp fun of
|
||||||
getDArgs' (f :% a) args | Yes yes = getDArgs' f (a :: args)
|
Left y => let f :% d = fun in getDArgs' f (d :: args)
|
||||||
getDArgs' fun args | No no = GotDArgs {fun, args, notDApp = no}
|
Right n => GotDArgs {fun, args, notDApp = n}
|
||||||
|
|
||||||
||| splits a dimension application into its head and arguments. if it's not an
|
||| splits a dimension application into its head and arguments. if it's not an
|
||||||
||| d application then the list is just empty
|
||| d application then the list is just empty
|
||||||
|
@ -172,11 +138,10 @@ export
|
||||||
getLams' : forall lams, rest.
|
getLams' : forall lams, rest.
|
||||||
Vect lams Name -> Term q d rest -> (0 eq : lams + n = rest) ->
|
Vect lams Name -> Term q d rest -> (0 eq : lams + n = rest) ->
|
||||||
GetLams q d n
|
GetLams q d n
|
||||||
getLams' xs s eq with (isLam s)
|
getLams' xs s Refl = case nchoose $ isLam s of
|
||||||
getLams' xs (Lam x sbody) Refl | Yes _ =
|
Left y => let Lam x body = s in
|
||||||
let body = assert_smaller (Lam x sbody) sbody.term in
|
getLams' (x :: xs) (assert_smaller s body.term) Refl
|
||||||
getLams' (x :: xs) body Refl
|
Right n => GotLams xs s Refl n
|
||||||
getLams' xs s eq | No no = GotLams xs s eq no
|
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
getLams : Term q d n -> GetLams q d n
|
getLams : Term q d n -> GetLams q d n
|
||||||
|
@ -196,11 +161,10 @@ export
|
||||||
getDLams' : forall lams, rest.
|
getDLams' : forall lams, rest.
|
||||||
Vect lams Name -> Term q rest n -> (0 eq : lams + d = rest) ->
|
Vect lams Name -> Term q rest n -> (0 eq : lams + d = rest) ->
|
||||||
GetDLams q d n
|
GetDLams q d n
|
||||||
getDLams' is s eq with (isDLam s)
|
getDLams' is s Refl = case nchoose $ isDLam s of
|
||||||
getDLams' is (DLam i sbody) Refl | Yes _ =
|
Left y => let DLam i body = s in
|
||||||
let body = assert_smaller (DLam i sbody) sbody.term in
|
getDLams' (i :: is) (assert_smaller s body.term) Refl
|
||||||
getDLams' (i :: is) body Refl
|
Right n => GotDLams is s Refl n
|
||||||
getDLams' is s eq | No no = GotDLams is s eq no
|
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
getDLams : Term q d n -> GetDLams q d n
|
getDLams : Term q d n -> GetDLams q d n
|
||||||
|
|
Loading…
Reference in a new issue