make overloaded reduce stuff into interfaces
this is kinda a pain so i might change it back i guess
This commit is contained in:
parent
56791e286d
commit
cb5bd6c98c
1 changed files with 165 additions and 175 deletions
|
@ -9,120 +9,129 @@ import Data.Maybe
|
||||||
%default total
|
%default total
|
||||||
|
|
||||||
|
|
||||||
namespace Elim
|
|
||||||
public export %inline
|
|
||||||
isClo : Elim {} -> Bool
|
|
||||||
isClo (CloE {}) = True
|
|
||||||
isClo (DCloE {}) = True
|
|
||||||
isClo _ = False
|
|
||||||
|
|
||||||
public export
|
|
||||||
0 NotClo : Pred $ Elim {}
|
|
||||||
NotClo = No . isClo
|
|
||||||
|
|
||||||
namespace Term
|
|
||||||
public export %inline
|
|
||||||
isClo : Term {} -> Bool
|
|
||||||
isClo (CloT {}) = True
|
|
||||||
isClo (DCloT {}) = True
|
|
||||||
isClo (E e) = isClo e
|
|
||||||
isClo _ = False
|
|
||||||
|
|
||||||
public export
|
|
||||||
0 NotClo : Pred $ Term {}
|
|
||||||
NotClo = No . isClo
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 NonCloElim : TermLike
|
0 CloTest : TermLike -> Type
|
||||||
NonCloElim q d n = Subset (Elim q d n) NotClo
|
CloTest tm = forall q, d, n. tm q d n -> Bool
|
||||||
|
|
||||||
|
interface PushSubsts (0 tm : TermLike) (0 isClo : CloTest tm) | tm where
|
||||||
|
pushSubstsWith : DSubst dfrom dto -> TSubst q dto from to ->
|
||||||
|
tm q dfrom from -> Subset (tm q dto to) (No . isClo)
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 NonCloTerm : TermLike
|
0 NotClo : {isClo : CloTest tm} -> PushSubsts tm isClo => Pred (tm q d n)
|
||||||
NonCloTerm q d n = Subset (Term q d n) NotClo
|
NotClo = No . isClo
|
||||||
|
|
||||||
|
public export
|
||||||
|
0 NonClo : (tm : TermLike) -> {isClo : CloTest tm} ->
|
||||||
|
PushSubsts tm isClo => TermLike
|
||||||
|
NonClo tm q d n = Subset (tm q d n) NotClo
|
||||||
|
|
||||||
public export %inline
|
public export %inline
|
||||||
ncloT : (t : Term q d n) -> (0 nc : NotClo t) => NonCloTerm q d n
|
nclo : {isClo : CloTest tm} -> (0 _ : PushSubsts tm isClo) =>
|
||||||
ncloT t = Element t nc
|
(t : tm q d n) -> (0 nc : NotClo t) => NonClo tm q d n
|
||||||
|
nclo t = Element t nc
|
||||||
|
|
||||||
public export %inline
|
parameters {0 isClo : CloTest tm} {auto _ : PushSubsts tm isClo}
|
||||||
ncloE : (e : Elim q d n) -> (0 nc : NotClo e) => NonCloElim q d n
|
||| if the input term has any top-level closures, push them under one layer of
|
||||||
ncloE e = Element e nc
|
||| syntax
|
||||||
|
export %inline
|
||||||
|
pushSubsts : tm q d n -> NonClo tm q d n
|
||||||
|
pushSubsts s = pushSubstsWith id id s
|
||||||
|
|
||||||
|
export %inline
|
||||||
|
pushSubstsWith' : DSubst dfrom dto -> TSubst q dto from to ->
|
||||||
|
tm q dfrom from -> tm q dto to
|
||||||
|
pushSubstsWith' th ph x = fst $ pushSubstsWith th ph x
|
||||||
|
|
||||||
mutual
|
mutual
|
||||||
namespace Term
|
public export
|
||||||
||| if the input term has any top-level closures, push them under one layer of
|
isCloT : CloTest Term
|
||||||
||| syntax
|
isCloT (CloT {}) = True
|
||||||
export %inline
|
isCloT (DCloT {}) = True
|
||||||
pushSubsts : Term q d n -> NonCloTerm q d n
|
isCloT (E e) = isCloE e
|
||||||
pushSubsts s = pushSubstsWith id id s
|
isCloT _ = False
|
||||||
|
|
||||||
export
|
public export
|
||||||
pushSubstsWith : DSubst dfrom dto -> TSubst q dto from to ->
|
isCloE : CloTest Elim
|
||||||
Term q dfrom from -> NonCloTerm q dto to
|
isCloE (CloE {}) = True
|
||||||
|
isCloE (DCloE {}) = True
|
||||||
|
isCloE _ = False
|
||||||
|
|
||||||
|
mutual
|
||||||
|
export
|
||||||
|
PushSubsts Term Reduce.isCloT where
|
||||||
pushSubstsWith th ph (TYPE l) =
|
pushSubstsWith th ph (TYPE l) =
|
||||||
ncloT $ TYPE l
|
nclo $ TYPE l
|
||||||
pushSubstsWith th ph (Pi qty x a body) =
|
pushSubstsWith th ph (Pi qty x a body) =
|
||||||
ncloT $ Pi qty x (subs a th ph) (subs body th ph)
|
nclo $ Pi qty x (subs a th ph) (subs body th ph)
|
||||||
pushSubstsWith th ph (Lam x body) =
|
pushSubstsWith th ph (Lam x body) =
|
||||||
ncloT $ Lam x $ subs body th ph
|
nclo $ Lam x $ subs body th ph
|
||||||
pushSubstsWith th ph (Sig x a b) =
|
pushSubstsWith th ph (Sig x a b) =
|
||||||
ncloT $ Sig x (subs a th ph) (subs b th ph)
|
nclo $ Sig x (subs a th ph) (subs b th ph)
|
||||||
pushSubstsWith th ph (Pair s t) =
|
pushSubstsWith th ph (Pair s t) =
|
||||||
ncloT $ Pair (subs s th ph) (subs t th ph)
|
nclo $ Pair (subs s th ph) (subs t th ph)
|
||||||
pushSubstsWith th ph (Eq i ty l r) =
|
pushSubstsWith th ph (Eq i ty l r) =
|
||||||
ncloT $ Eq i (subs ty th ph) (subs l th ph) (subs r th ph)
|
nclo $ Eq i (subs ty th ph) (subs l th ph) (subs r th ph)
|
||||||
pushSubstsWith th ph (DLam i body) =
|
pushSubstsWith th ph (DLam i body) =
|
||||||
ncloT $ DLam i $ subs body th ph
|
nclo $ DLam i $ subs body th ph
|
||||||
pushSubstsWith th ph (E e) =
|
pushSubstsWith th ph (E e) =
|
||||||
let Element e nc = pushSubstsWith th ph e in ncloT $ E e
|
let Element e nc = pushSubstsWith th ph e in nclo $ E e
|
||||||
pushSubstsWith th ph (CloT s ps) =
|
pushSubstsWith th ph (CloT s ps) =
|
||||||
pushSubstsWith th (comp th ps ph) s
|
pushSubstsWith th (comp th ps ph) s
|
||||||
pushSubstsWith th ph (DCloT s ps) =
|
pushSubstsWith th ph (DCloT s ps) =
|
||||||
pushSubstsWith (ps . th) ph s
|
pushSubstsWith (ps . th) ph s
|
||||||
|
|
||||||
namespace Elim
|
export
|
||||||
||| if the input elimination has any top-level closures, push them under one
|
PushSubsts Elim Reduce.isCloE where
|
||||||
||| layer of syntax
|
|
||||||
export %inline
|
|
||||||
pushSubsts : Elim q d n -> NonCloElim q d n
|
|
||||||
pushSubsts e = pushSubstsWith id id e
|
|
||||||
|
|
||||||
export
|
|
||||||
pushSubstsWith : DSubst dfrom dto -> TSubst q dto from to ->
|
|
||||||
Elim q dfrom from -> NonCloElim q dto to
|
|
||||||
pushSubstsWith th ph (F x) =
|
pushSubstsWith th ph (F x) =
|
||||||
ncloE $ F x
|
nclo $ F x
|
||||||
pushSubstsWith th ph (B i) =
|
pushSubstsWith th ph (B i) =
|
||||||
let res = ph !! i in
|
let res = ph !! i in
|
||||||
case nchoose $ isClo res of
|
case nchoose $ isCloE res of
|
||||||
Left yes => assert_total pushSubsts res
|
Left yes => assert_total pushSubsts res
|
||||||
Right no => Element res no
|
Right no => Element res no
|
||||||
pushSubstsWith th ph (f :@ s) =
|
pushSubstsWith th ph (f :@ s) =
|
||||||
ncloE $ subs f th ph :@ subs s th ph
|
nclo $ subs f th ph :@ subs s th ph
|
||||||
pushSubstsWith th ph (CasePair pi p x r y z b) =
|
pushSubstsWith th ph (CasePair pi p x r y z b) =
|
||||||
ncloE $ CasePair pi (subs p th ph) x (subs r th ph) y z (subs b th ph)
|
nclo $ CasePair pi (subs p th ph) x (subs r th ph) y z (subs b th ph)
|
||||||
pushSubstsWith th ph (f :% d) =
|
pushSubstsWith th ph (f :% d) =
|
||||||
ncloE $ subs f th ph :% (d // th)
|
nclo $ subs f th ph :% (d // th)
|
||||||
pushSubstsWith th ph (s :# a) =
|
pushSubstsWith th ph (s :# a) =
|
||||||
ncloE $ subs s th ph :# subs a th ph
|
nclo $ subs s th ph :# subs a th ph
|
||||||
pushSubstsWith th ph (CloE e ps) =
|
pushSubstsWith th ph (CloE e ps) =
|
||||||
pushSubstsWith th (comp th ps ph) e
|
pushSubstsWith th (comp th ps ph) e
|
||||||
pushSubstsWith th ph (DCloE e ps) =
|
pushSubstsWith th ph (DCloE e ps) =
|
||||||
pushSubstsWith (ps . th) ph e
|
pushSubstsWith (ps . th) ph e
|
||||||
|
|
||||||
|
|
||||||
parameters (th : DSubst dfrom dto) (ph : TSubst q dto from to)
|
|
||||||
namespace Term
|
|
||||||
public export %inline
|
|
||||||
pushSubstsWith' : Term q dfrom from -> Term q dto to
|
|
||||||
pushSubstsWith' s = (pushSubstsWith th ph s).fst
|
|
||||||
|
|
||||||
namespace Elim
|
public export
|
||||||
public export %inline
|
0 RedexTest : TermLike -> Type
|
||||||
pushSubstsWith' : Elim q dfrom from -> Elim q dto to
|
RedexTest tm = forall q, d, n, g. Definitions' q g -> tm q d n -> Bool
|
||||||
pushSubstsWith' e = (pushSubstsWith th ph e).fst
|
|
||||||
|
public export
|
||||||
|
interface Whnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm where
|
||||||
|
whnf : (defs : Definitions' q g) ->
|
||||||
|
tm q d n -> Subset (tm q d n) (No . isRedex defs)
|
||||||
|
|
||||||
|
public export
|
||||||
|
0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> Whnf tm isRedex =>
|
||||||
|
Definitions' q g -> Pred (tm q d n)
|
||||||
|
IsRedex defs = So . isRedex defs
|
||||||
|
NotRedex defs = No . isRedex defs
|
||||||
|
|
||||||
|
public export
|
||||||
|
0 NonRedex : (tm : TermLike) -> {isRedex : RedexTest tm} -> Whnf tm isRedex =>
|
||||||
|
(q : Type) -> (d, n : Nat) -> {g : _} ->
|
||||||
|
(defs : Definitions' q g) -> Type
|
||||||
|
NonRedex tm q d n defs = Subset (tm q d n) (NotRedex defs)
|
||||||
|
|
||||||
|
public export %inline
|
||||||
|
nred : {0 isRedex : RedexTest tm} -> (0 _ : Whnf tm isRedex) =>
|
||||||
|
(t : tm q d n) -> (0 nr : NotRedex defs t) =>
|
||||||
|
NonRedex tm q d n defs
|
||||||
|
nred t = Element t nr
|
||||||
|
|
||||||
|
|
||||||
public export %inline
|
public export %inline
|
||||||
|
@ -150,118 +159,99 @@ isAnn : Elim {} -> Bool
|
||||||
isAnn (_ :# _) = True
|
isAnn (_ :# _) = True
|
||||||
isAnn _ = False
|
isAnn _ = False
|
||||||
|
|
||||||
parameters (defs : Definitions' q g)
|
|
||||||
mutual
|
|
||||||
namespace Elim
|
|
||||||
public export
|
|
||||||
isRedex : Elim q d n -> Bool
|
|
||||||
isRedex (F x) {d, n} = isJust $ lookupElim x defs {d, n}
|
|
||||||
isRedex (B _) = False
|
|
||||||
isRedex (f :@ _) = isRedex f || isLamHead f
|
|
||||||
isRedex (CasePair {pair, _}) = isRedex pair || isPairHead pair
|
|
||||||
isRedex (f :% _) = isRedex f || isDLamHead f
|
|
||||||
isRedex (t :# a) = isE t || isRedex t || isRedex a
|
|
||||||
isRedex (CloE {}) = True
|
|
||||||
isRedex (DCloE {}) = True
|
|
||||||
|
|
||||||
namespace Term
|
mutual
|
||||||
public export
|
|
||||||
isRedex : Term q d n -> Bool
|
|
||||||
isRedex (CloT {}) = True
|
|
||||||
isRedex (DCloT {}) = True
|
|
||||||
isRedex (E e) = isAnn e || isRedex e
|
|
||||||
isRedex _ = False
|
|
||||||
|
|
||||||
namespace Elim
|
|
||||||
public export
|
|
||||||
0 IsRedex, NotRedex : Pred $ Elim q d n
|
|
||||||
IsRedex = So . isRedex
|
|
||||||
NotRedex = No . isRedex
|
|
||||||
|
|
||||||
namespace Term
|
|
||||||
public export
|
|
||||||
0 IsRedex, NotRedex : Pred $ Term q d n
|
|
||||||
IsRedex = So . isRedex
|
|
||||||
NotRedex = No . isRedex
|
|
||||||
|
|
||||||
parameters (q : Type) (d, n : Nat) {g : q -> Type} (defs : Definitions' q g)
|
|
||||||
public export
|
public export
|
||||||
0 NonRedexElim, NonRedexTerm : Type
|
isRedexE : RedexTest Elim
|
||||||
NonRedexElim = Subset (Elim q d n) (NotRedex defs)
|
isRedexE defs (F x) {d, n} =
|
||||||
NonRedexTerm = Subset (Term q d n) (NotRedex defs)
|
isJust $ lookupElim x defs {d, n}
|
||||||
|
isRedexE _ (B _) = False
|
||||||
|
isRedexE defs (f :@ _) =
|
||||||
|
isRedexE defs f || isLamHead f
|
||||||
|
isRedexE defs (CasePair {pair, _}) =
|
||||||
|
isRedexE defs pair || isPairHead pair
|
||||||
|
isRedexE defs (f :% _) =
|
||||||
|
isRedexE defs f || isDLamHead f
|
||||||
|
isRedexE defs (t :# a) =
|
||||||
|
isE t || isRedexT defs t || isRedexT defs a
|
||||||
|
isRedexE _ (CloE {}) = True
|
||||||
|
isRedexE _ (DCloE {}) = True
|
||||||
|
|
||||||
|
public export
|
||||||
|
isRedexT : RedexTest Term
|
||||||
|
isRedexT _ (CloT {}) = True
|
||||||
|
isRedexT _ (DCloT {}) = True
|
||||||
|
isRedexT defs (E e) = isAnn e || isRedexE defs e
|
||||||
|
isRedexT _ _ = False
|
||||||
|
|
||||||
parameters (defs : Definitions' q g)
|
mutual
|
||||||
mutual
|
export covering
|
||||||
namespace Elim
|
Whnf Elim Reduce.isRedexE where
|
||||||
export covering
|
whnf defs (F x) with (lookupElim x defs) proof eq
|
||||||
whnf : Elim q d n -> NonRedexElim q d n defs
|
_ | Just y = whnf defs y
|
||||||
whnf (F x) with (lookupElim x defs) proof eq
|
_ | Nothing = Element (F x) $ rewrite eq in Ah
|
||||||
_ | Just y = whnf y
|
|
||||||
_ | Nothing = Element (F x) $ rewrite eq in Ah
|
|
||||||
|
|
||||||
whnf (B i) = Element (B i) Ah
|
whnf _ (B i) = nred $ B i
|
||||||
|
|
||||||
whnf (f :@ s) =
|
whnf defs (f :@ s) =
|
||||||
let Element f fnf = whnf f in
|
let Element f fnf = whnf defs f in
|
||||||
case nchoose $ isLamHead f of
|
case nchoose $ isLamHead f of
|
||||||
Left _ =>
|
Left _ =>
|
||||||
let Lam {body, _} :# Pi {arg, res, _} = f
|
let Lam {body, _} :# Pi {arg, res, _} = f
|
||||||
s = s :# arg
|
s = s :# arg
|
||||||
in
|
in
|
||||||
whnf $ sub1 body s :# sub1 res s
|
whnf defs $ sub1 body s :# sub1 res s
|
||||||
Right nlh => Element (f :@ s) $ fnf `orNo` nlh
|
Right nlh => Element (f :@ s) $ fnf `orNo` nlh
|
||||||
|
|
||||||
whnf (CasePair pi pair r ret x y body) =
|
whnf defs (CasePair pi pair r ret x y body) =
|
||||||
let Element pair pairnf = whnf pair in
|
let Element pair pairnf = whnf defs pair in
|
||||||
case nchoose $ isPairHead pair of
|
case nchoose $ isPairHead pair of
|
||||||
Left _ =>
|
Left _ =>
|
||||||
let Pair {fst, snd} :# Sig {fst = tfst, snd = tsnd, _} = pair
|
let Pair {fst, snd} :# Sig {fst = tfst, snd = tsnd, _} = pair
|
||||||
fst = fst :# tfst
|
fst = fst :# tfst
|
||||||
snd = snd :# sub1 tsnd fst
|
snd = snd :# sub1 tsnd fst
|
||||||
in
|
in
|
||||||
whnf $ subN body [fst, snd] :# sub1 ret pair
|
whnf defs $ subN body [fst, snd] :# sub1 ret pair
|
||||||
Right np =>
|
Right np =>
|
||||||
Element (CasePair pi pair r ret x y body) $ pairnf `orNo` np
|
Element (CasePair pi pair r ret x y body) $ pairnf `orNo` np
|
||||||
|
|
||||||
whnf (f :% p) =
|
whnf defs (f :% p) =
|
||||||
let Element f fnf = whnf f in
|
let Element f fnf = whnf defs f in
|
||||||
case nchoose $ isDLamHead f of
|
case nchoose $ isDLamHead f of
|
||||||
Left _ =>
|
Left _ =>
|
||||||
let DLam {body, _} :# Eq {ty, l, r, _} = f
|
let DLam {body, _} :# Eq {ty, l, r, _} = f
|
||||||
body = endsOr l r (dsub1 body p) p
|
body = endsOr l r (dsub1 body p) p
|
||||||
in
|
in
|
||||||
whnf $ body :# dsub1 ty p
|
whnf defs $ body :# dsub1 ty p
|
||||||
Right ndlh =>
|
Right ndlh =>
|
||||||
Element (f :% p) $ fnf `orNo` ndlh
|
Element (f :% p) $ fnf `orNo` ndlh
|
||||||
|
|
||||||
whnf (s :# a) =
|
whnf defs (s :# a) =
|
||||||
let Element s snf = whnf s in
|
let Element s snf = whnf defs s in
|
||||||
case nchoose $ isE s of
|
case nchoose $ isE s of
|
||||||
Left _ => let E e = s in Element e $ noOr2 snf
|
Left _ => let E e = s in Element e $ noOr2 snf
|
||||||
Right ne =>
|
Right ne =>
|
||||||
let Element a anf = whnf a in
|
let Element a anf = whnf defs a in
|
||||||
Element (s :# a) $ ne `orNo` snf `orNo` anf
|
Element (s :# a) $ ne `orNo` snf `orNo` anf
|
||||||
|
|
||||||
whnf (CloE el th) = whnf $ pushSubstsWith' id th el
|
whnf defs (CloE el th) = whnf defs $ pushSubstsWith' id th el
|
||||||
whnf (DCloE el th) = whnf $ pushSubstsWith' th id el
|
whnf defs (DCloE el th) = whnf defs $ pushSubstsWith' th id el
|
||||||
|
|
||||||
namespace Term
|
export covering
|
||||||
export covering
|
Whnf Term Reduce.isRedexT where
|
||||||
whnf : Term q d n -> NonRedexTerm q d n defs
|
whnf _ t@(TYPE {}) = nred t
|
||||||
whnf t@(TYPE {}) = Element t Ah
|
whnf _ t@(Pi {}) = nred t
|
||||||
whnf t@(Pi {}) = Element t Ah
|
whnf _ t@(Lam {}) = nred t
|
||||||
whnf t@(Lam {}) = Element t Ah
|
whnf _ t@(Sig {}) = nred t
|
||||||
whnf t@(Sig {}) = Element t Ah
|
whnf _ t@(Pair {}) = nred t
|
||||||
whnf t@(Pair {}) = Element t Ah
|
whnf _ t@(Eq {}) = nred t
|
||||||
whnf t@(Eq {}) = Element t Ah
|
whnf _ t@(DLam {}) = nred t
|
||||||
whnf t@(DLam {}) = Element t Ah
|
|
||||||
|
|
||||||
whnf (E e) =
|
whnf defs (E e) =
|
||||||
let Element e enf = whnf e in
|
let Element e enf = whnf defs e in
|
||||||
case nchoose $ isAnn e of
|
case nchoose $ isAnn e of
|
||||||
Left _ => let tm :# _ = e in Element tm $ noOr1 $ noOr2 enf
|
Left _ => let tm :# _ = e in Element tm $ noOr1 $ noOr2 enf
|
||||||
Right na => Element (E e) $ na `orNo` enf
|
Right na => Element (E e) $ na `orNo` enf
|
||||||
|
|
||||||
whnf (CloT tm th) = whnf $ pushSubstsWith' id th tm
|
whnf defs (CloT tm th) = whnf defs $ pushSubstsWith' id th tm
|
||||||
whnf (DCloT tm th) = whnf $ pushSubstsWith' th id tm
|
whnf defs (DCloT tm th) = whnf defs $ pushSubstsWith' th id tm
|
||||||
|
|
Loading…
Reference in a new issue