split up whnf module
This commit is contained in:
parent
a24ebe0702
commit
8264a1bb81
9 changed files with 854 additions and 788 deletions
216
lib/Quox/Whnf/Interface.idr
Normal file
216
lib/Quox/Whnf/Interface.idr
Normal file
|
@ -0,0 +1,216 @@
|
|||
module Quox.Whnf.Interface
|
||||
|
||||
import public Quox.No
|
||||
import public Quox.Syntax
|
||||
import public Quox.Definition
|
||||
import public Quox.Typing.Context
|
||||
import public Quox.Typing.Error
|
||||
import public Data.Maybe
|
||||
import public Control.Eff
|
||||
|
||||
%default total
|
||||
|
||||
|
||||
public export
|
||||
Whnf : List (Type -> Type)
|
||||
Whnf = [NameGen, Except Error]
|
||||
|
||||
export
|
||||
runWhnfWith : NameSuf -> Eff Whnf a -> (Either Error a, NameSuf)
|
||||
runWhnfWith suf act = extract $ runStateAt GEN suf $ runExcept act
|
||||
|
||||
export
|
||||
runWhnf : Eff Whnf a -> Either Error a
|
||||
runWhnf = fst . runWhnfWith 0
|
||||
|
||||
|
||||
public export
|
||||
0 RedexTest : TermLike -> Type
|
||||
RedexTest tm = {d, n : Nat} -> Definitions -> tm d n -> Bool
|
||||
|
||||
public export
|
||||
interface CanWhnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm
|
||||
where
|
||||
whnf : {d, n : Nat} -> (defs : Definitions) ->
|
||||
(ctx : WhnfContext d n) ->
|
||||
tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs))
|
||||
-- having isRedex be part of the class header, and needing to be explicitly
|
||||
-- quantified on every use since idris can't infer its type, is a little ugly.
|
||||
-- but none of the alternatives i've thought of so far work. e.g. in some
|
||||
-- cases idris can't tell that `isRedex` and `isRedexT` are the same thing
|
||||
|
||||
public export %inline
|
||||
whnf0 : {d, n : Nat} -> {0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
|
||||
(defs : Definitions) -> WhnfContext d n -> tm d n -> Eff Whnf (tm d n)
|
||||
whnf0 defs ctx t = fst <$> whnf defs ctx t
|
||||
|
||||
public export
|
||||
0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
|
||||
Definitions -> Pred (tm d n)
|
||||
IsRedex defs = So . isRedex defs
|
||||
NotRedex defs = No . isRedex defs
|
||||
|
||||
public export
|
||||
0 NonRedex : (tm : TermLike) -> {isRedex : RedexTest tm} ->
|
||||
CanWhnf tm isRedex => (d, n : Nat) -> (defs : Definitions) -> Type
|
||||
NonRedex tm d n defs = Subset (tm d n) (NotRedex defs)
|
||||
|
||||
public export %inline
|
||||
nred : {0 isRedex : RedexTest tm} -> (0 _ : CanWhnf tm isRedex) =>
|
||||
(t : tm d n) -> (0 nr : NotRedex defs t) => NonRedex tm d n defs
|
||||
nred t = Element t nr
|
||||
|
||||
|
||||
||| an expression like `(λ x ⇒ s) ∷ π.(x : A) → B`
|
||||
public export %inline
|
||||
isLamHead : Elim {} -> Bool
|
||||
isLamHead (Ann (Lam {}) (Pi {}) _) = True
|
||||
isLamHead (Coe {}) = True
|
||||
isLamHead _ = False
|
||||
|
||||
||| an expression like `(δ 𝑖 ⇒ s) ∷ Eq (𝑖 ⇒ A) s t`
|
||||
public export %inline
|
||||
isDLamHead : Elim {} -> Bool
|
||||
isDLamHead (Ann (DLam {}) (Eq {}) _) = True
|
||||
isDLamHead (Coe {}) = True
|
||||
isDLamHead _ = False
|
||||
|
||||
||| an expression like `(s, t) ∷ (x : A) × B`
|
||||
public export %inline
|
||||
isPairHead : Elim {} -> Bool
|
||||
isPairHead (Ann (Pair {}) (Sig {}) _) = True
|
||||
isPairHead (Coe {}) = True
|
||||
isPairHead _ = False
|
||||
|
||||
||| an expression like `'a ∷ {a, b, c}`
|
||||
public export %inline
|
||||
isTagHead : Elim {} -> Bool
|
||||
isTagHead (Ann (Tag {}) (Enum {}) _) = True
|
||||
isTagHead (Coe {}) = True
|
||||
isTagHead _ = False
|
||||
|
||||
||| an expression like `0 ∷ ℕ` or `suc n ∷ ℕ`
|
||||
public export %inline
|
||||
isNatHead : Elim {} -> Bool
|
||||
isNatHead (Ann (Zero {}) (Nat {}) _) = True
|
||||
isNatHead (Ann (Succ {}) (Nat {}) _) = True
|
||||
isNatHead (Coe {}) = True
|
||||
isNatHead _ = False
|
||||
|
||||
||| an expression like `[s] ∷ [π. A]`
|
||||
public export %inline
|
||||
isBoxHead : Elim {} -> Bool
|
||||
isBoxHead (Ann (Box {}) (BOX {}) _) = True
|
||||
isBoxHead (Coe {}) = True
|
||||
isBoxHead _ = False
|
||||
|
||||
||| an elimination in a term context
|
||||
public export %inline
|
||||
isE : Term {} -> Bool
|
||||
isE (E {}) = True
|
||||
isE _ = False
|
||||
|
||||
||| an expression like `s ∷ A`
|
||||
public export %inline
|
||||
isAnn : Elim {} -> Bool
|
||||
isAnn (Ann {}) = True
|
||||
isAnn _ = False
|
||||
|
||||
||| a syntactic type
|
||||
public export %inline
|
||||
isTyCon : Term {} -> Bool
|
||||
isTyCon (TYPE {}) = True
|
||||
isTyCon (Pi {}) = True
|
||||
isTyCon (Lam {}) = False
|
||||
isTyCon (Sig {}) = True
|
||||
isTyCon (Pair {}) = False
|
||||
isTyCon (Enum {}) = True
|
||||
isTyCon (Tag {}) = False
|
||||
isTyCon (Eq {}) = True
|
||||
isTyCon (DLam {}) = False
|
||||
isTyCon (Nat {}) = True
|
||||
isTyCon (Zero {}) = False
|
||||
isTyCon (Succ {}) = False
|
||||
isTyCon (BOX {}) = True
|
||||
isTyCon (Box {}) = False
|
||||
isTyCon (E {}) = False
|
||||
isTyCon (CloT {}) = False
|
||||
isTyCon (DCloT {}) = False
|
||||
|
||||
||| a syntactic type, or a neutral
|
||||
public export %inline
|
||||
isTyConE : Term {} -> Bool
|
||||
isTyConE s = isTyCon s || isE s
|
||||
|
||||
||| a syntactic type with an annotation `★ᵢ`
|
||||
public export %inline
|
||||
isAnnTyCon : Elim {} -> Bool
|
||||
isAnnTyCon (Ann ty (TYPE {}) _) = isTyCon ty
|
||||
isAnnTyCon _ = False
|
||||
|
||||
||| 0 or 1
|
||||
public export %inline
|
||||
isK : Dim d -> Bool
|
||||
isK (K {}) = True
|
||||
isK _ = False
|
||||
|
||||
|
||||
mutual
|
||||
||| a reducible elimination
|
||||
|||
|
||||
||| - a free variable, if its definition is known
|
||||
||| - an application whose head is an annotated lambda,
|
||||
||| a case expression whose head is an annotated constructor form, etc
|
||||
||| - a redundant annotation, or one whose term or type is reducible
|
||||
||| - coercions `coe (𝑖 ⇒ A) @p @q s` where:
|
||||
||| - `A` is in whnf, or
|
||||
||| - `𝑖` is not mentioned in `A`, or
|
||||
||| - `p = q`
|
||||
||| - [fixme] this is not true right now but that's because i wrote it
|
||||
||| wrong
|
||||
||| - compositions `comp A @p @q s @r {⋯}` where:
|
||||
||| - `A` is in whnf, or
|
||||
||| - `p = q`, or
|
||||
||| - `r = 0` or `r = 1`
|
||||
||| - [fixme] the p=q condition is also missing here
|
||||
||| - closures
|
||||
public export
|
||||
isRedexE : RedexTest Elim
|
||||
isRedexE defs (F {x, _}) {d, n} =
|
||||
isJust $ lookupElim x defs {d, n}
|
||||
isRedexE _ (B {}) = False
|
||||
isRedexE defs (App {fun, _}) =
|
||||
isRedexE defs fun || isLamHead fun
|
||||
isRedexE defs (CasePair {pair, _}) =
|
||||
isRedexE defs pair || isPairHead pair
|
||||
isRedexE defs (CaseEnum {tag, _}) =
|
||||
isRedexE defs tag || isTagHead tag
|
||||
isRedexE defs (CaseNat {nat, _}) =
|
||||
isRedexE defs nat || isNatHead nat
|
||||
isRedexE defs (CaseBox {box, _}) =
|
||||
isRedexE defs box || isBoxHead box
|
||||
isRedexE defs (DApp {fun, arg, _}) =
|
||||
isRedexE defs fun || isDLamHead fun || isK arg
|
||||
isRedexE defs (Ann {tm, ty, _}) =
|
||||
isE tm || isRedexT defs tm || isRedexT defs ty
|
||||
isRedexE defs (Coe {val, _}) =
|
||||
isRedexT defs val || not (isE val)
|
||||
isRedexE defs (Comp {ty, r, _}) =
|
||||
isRedexT defs ty || isK r
|
||||
isRedexE defs (TypeCase {ty, ret, _}) =
|
||||
isRedexE defs ty || isRedexT defs ret || isAnnTyCon ty
|
||||
isRedexE _ (CloE {}) = True
|
||||
isRedexE _ (DCloE {}) = True
|
||||
|
||||
||| a reducible term
|
||||
|||
|
||||
||| - a reducible elimination, as `isRedexE`
|
||||
||| - an annotated elimination
|
||||
||| (the annotation is redundant in a checkable context)
|
||||
||| - a closure
|
||||
public export
|
||||
isRedexT : RedexTest Term
|
||||
isRedexT _ (CloT {}) = True
|
||||
isRedexT _ (DCloT {}) = True
|
||||
isRedexT defs (E {e, _}) = isAnn e || isRedexE defs e
|
||||
isRedexT _ _ = False
|
Loading…
Add table
Add a link
Reference in a new issue