remove src directories
This commit is contained in:
parent
79211cff84
commit
804f1e3638
36 changed files with 0 additions and 3 deletions
164
lib/Quox/Syntax/Term/Reduce.idr
Normal file
164
lib/Quox/Syntax/Term/Reduce.idr
Normal file
|
@ -0,0 +1,164 @@
|
|||
module Quox.Syntax.Term.Reduce
|
||||
|
||||
import Quox.Syntax.Term.Base
|
||||
import Quox.Syntax.Term.Subst
|
||||
|
||||
%default total
|
||||
|
||||
|
||||
mutual
|
||||
||| true if a term has a closure or dimension closure at the top level,
|
||||
||| or is `E` applied to such an elimination
|
||||
public export %inline
|
||||
topCloT : Term d n -> Bool
|
||||
topCloT (CloT _ _) = True
|
||||
topCloT (DCloT _ _) = True
|
||||
topCloT (E e) = topCloE e
|
||||
topCloT _ = False
|
||||
|
||||
||| true if an elimination has a closure or dimension closure at the top level
|
||||
public export %inline
|
||||
topCloE : Elim d n -> Bool
|
||||
topCloE (CloE _ _) = True
|
||||
topCloE (DCloE _ _) = True
|
||||
topCloE _ = False
|
||||
|
||||
|
||||
public export IsNotCloT : Term d n -> Type
|
||||
IsNotCloT = So . not . topCloT
|
||||
|
||||
||| a term which is not a top level closure
|
||||
public export NotCloTerm : Nat -> Nat -> Type
|
||||
NotCloTerm d n = Subset (Term d n) IsNotCloT
|
||||
|
||||
public export IsNotCloE : Elim d n -> Type
|
||||
IsNotCloE = So . not . topCloE
|
||||
|
||||
||| an elimination which is not a top level closure
|
||||
public export NotCloElim : Nat -> Nat -> Type
|
||||
NotCloElim d n = Subset (Elim d n) IsNotCloE
|
||||
|
||||
public export %inline
|
||||
ncloT : (t : Term d n) -> (0 _ : IsNotCloT t) => NotCloTerm d n
|
||||
ncloT t @{p} = Element t p
|
||||
|
||||
public export %inline
|
||||
ncloE : (t : Elim d n) -> (0 _ : IsNotCloE t) => NotCloElim d n
|
||||
ncloE e @{p} = Element e p
|
||||
|
||||
|
||||
|
||||
mutual
|
||||
||| if the input term has any top-level closures, push them under one layer of
|
||||
||| syntax
|
||||
export %inline
|
||||
pushSubstsT : Term d n -> NotCloTerm d n
|
||||
pushSubstsT s = pushSubstsTWith id id s
|
||||
|
||||
||| if the input elimination has any top-level closures, push them under one
|
||||
||| layer of syntax
|
||||
export %inline
|
||||
pushSubstsE : Elim d n -> NotCloElim d n
|
||||
pushSubstsE e = pushSubstsEWith id id e
|
||||
|
||||
export
|
||||
pushSubstsTWith : DSubst dfrom dto -> TSubst dto from to ->
|
||||
Term dfrom from -> NotCloTerm dto to
|
||||
pushSubstsTWith th ph (TYPE l) =
|
||||
ncloT $ TYPE l
|
||||
pushSubstsTWith th ph (Pi qty x a body) =
|
||||
ncloT $ Pi qty x (subs a th ph) (subs body th ph)
|
||||
pushSubstsTWith th ph (Lam x body) =
|
||||
ncloT $ Lam x $ subs body th ph
|
||||
pushSubstsTWith th ph (E e) =
|
||||
let Element e _ = pushSubstsEWith th ph e in ncloT $ E e
|
||||
pushSubstsTWith th ph (CloT s ps) =
|
||||
pushSubstsTWith th (comp' th ps ph) s
|
||||
pushSubstsTWith th ph (DCloT s ps) =
|
||||
pushSubstsTWith (ps . th) ph s
|
||||
|
||||
export
|
||||
pushSubstsEWith : DSubst dfrom dto -> TSubst dto from to ->
|
||||
Elim dfrom from -> NotCloElim dto to
|
||||
pushSubstsEWith th ph (F x) =
|
||||
ncloE $ F x
|
||||
pushSubstsEWith th ph (B i) =
|
||||
assert_total pushSubstsE $ ph !! i
|
||||
pushSubstsEWith th ph (f :@ s) =
|
||||
ncloE $ subs f th ph :@ subs s th ph
|
||||
pushSubstsEWith th ph (s :# a) =
|
||||
ncloE $ subs s th ph :# subs a th ph
|
||||
pushSubstsEWith th ph (CloE e ps) =
|
||||
pushSubstsEWith th (comp' th ps ph) e
|
||||
pushSubstsEWith th ph (DCloE e ps) =
|
||||
pushSubstsEWith (ps . th) ph e
|
||||
|
||||
|
||||
parameters (th : DSubst dfrom dto) (ph : TSubst dto from to)
|
||||
public export %inline
|
||||
pushSubstsTWith' : Term dfrom from -> Term dto to
|
||||
pushSubstsTWith' s = (pushSubstsTWith th ph s).fst
|
||||
|
||||
public export %inline
|
||||
pushSubstsEWith' : Elim dfrom from -> Elim dto to
|
||||
pushSubstsEWith' e = (pushSubstsEWith th ph e).fst
|
||||
|
||||
|
||||
public export %inline
|
||||
pushSubstsT' : Term d n -> Term d n
|
||||
pushSubstsT' s = (pushSubstsT s).fst
|
||||
|
||||
public export %inline
|
||||
pushSubstsE' : Elim d n -> Elim d n
|
||||
pushSubstsE' e = (pushSubstsE e).fst
|
||||
|
||||
|
||||
mutual
|
||||
-- tightening a term/elim also causes substitutions to be pushed through.
|
||||
-- this is because otherwise a variable in an unused part of the subst
|
||||
-- would cause it to incorrectly fail
|
||||
|
||||
export covering
|
||||
Tighten (Term d) where
|
||||
tighten p (TYPE l) =
|
||||
pure $ TYPE l
|
||||
tighten p (Pi qty x arg res) =
|
||||
Pi qty x <$> tighten p arg
|
||||
<*> tighten p res
|
||||
tighten p (Lam x body) =
|
||||
Lam x <$> tighten p body
|
||||
tighten p (E e) =
|
||||
E <$> tighten p e
|
||||
tighten p (CloT tm th) =
|
||||
tighten p $ pushSubstsTWith' id th tm
|
||||
tighten p (DCloT tm th) =
|
||||
tighten p $ pushSubstsTWith' th id tm
|
||||
|
||||
export covering
|
||||
Tighten (Elim d) where
|
||||
tighten p (F x) =
|
||||
pure $ F x
|
||||
tighten p (B i) =
|
||||
B <$> tighten p i
|
||||
tighten p (fun :@ arg) =
|
||||
[|tighten p fun :@ tighten p arg|]
|
||||
tighten p (tm :# ty) =
|
||||
[|tighten p tm :# tighten p ty|]
|
||||
tighten p (CloE el th) =
|
||||
tighten p $ pushSubstsEWith' id th el
|
||||
tighten p (DCloE el th) =
|
||||
tighten p $ pushSubstsEWith' th id el
|
||||
|
||||
export covering
|
||||
Tighten (ScopeTerm d) where
|
||||
tighten p (TUsed body) = TUsed <$> tighten (Keep p) body
|
||||
tighten p (TUnused body) = TUnused <$> tighten p body
|
||||
|
||||
|
||||
public export %inline
|
||||
weakT : Term d n -> Term d (S n)
|
||||
weakT t = t //. shift 1
|
||||
|
||||
public export %inline
|
||||
weakE : Elim d n -> Elim d (S n)
|
||||
weakE t = t //. shift 1
|
Loading…
Add table
Add a link
Reference in a new issue