pair stuff
This commit is contained in:
parent
6073ab4705
commit
4b36d8b7c8
16 changed files with 441 additions and 117 deletions
|
@ -52,6 +52,12 @@ mutual
|
|||
||| function term
|
||||
Lam : (x : Name) -> (body : ScopeTerm q d n) -> Term q d n
|
||||
|
||||
||| pair type
|
||||
Sig : (x : Name) -> (fst : Term q d n) -> (snd : ScopeTerm q d n) ->
|
||||
Term q d n
|
||||
||| pair value
|
||||
Pair : (fst, snd : Term q d n) -> Term q d n
|
||||
|
||||
||| equality type
|
||||
Eq : (i : Name) -> (ty : DScopeTerm q d n) -> (l, r : Term q d n) ->
|
||||
Term q d n
|
||||
|
@ -79,6 +85,15 @@ mutual
|
|||
||| term application
|
||||
(:@) : (fun : Elim q d n) -> (arg : Term q d n) -> Elim q d n
|
||||
|
||||
||| pair destruction
|
||||
|||
|
||||
||| `CasePair 𝜋 𝑒 𝑟 𝐴 𝑥 𝑦 𝑡` is
|
||||
||| `𝐜𝐚𝐬𝐞 𝜋 · 𝑒 𝐫𝐞𝐭𝐮𝐫𝐧 𝑟 ⇒ 𝐴 𝐨𝐟 { (𝑥, 𝑦) ⇒ 𝑡 }`
|
||||
CasePair : (qty : q) -> (pair : Elim q d n) ->
|
||||
(r : Name) -> (ret : ScopeTerm q d n) ->
|
||||
(x, y : Name) -> (body : ScopeTermN 2 q d n) ->
|
||||
Elim q d n
|
||||
|
||||
||| dim application
|
||||
(:%) : (fun : Elim q d n) -> (arg : Dim d) -> Elim q d n
|
||||
|
||||
|
|
|
@ -11,18 +11,25 @@ import Data.Vect
|
|||
|
||||
|
||||
private %inline
|
||||
arrowD, lamD, eqndD, dlamD, annD : Pretty.HasEnv m => m (Doc HL)
|
||||
arrowD = hlF Syntax $ ifUnicode "→" "->"
|
||||
lamD = hlF Syntax $ ifUnicode "λ" "fun"
|
||||
eqndD = hlF Syntax $ ifUnicode "≡" "=="
|
||||
dlamD = hlF Syntax $ ifUnicode "λᴰ" "dfun"
|
||||
annD = hlF Syntax $ ifUnicode "∷" "::"
|
||||
arrowD, timesD, darrowD, lamD, eqndD, dlamD, annD :
|
||||
Pretty.HasEnv m => m (Doc HL)
|
||||
arrowD = hlF Syntax $ ifUnicode "→" "->"
|
||||
timesD = hlF Syntax $ ifUnicode "×" "**"
|
||||
darrowD = hlF Syntax $ ifUnicode "⇒" "=>"
|
||||
lamD = hlF Syntax $ ifUnicode "λ" "fun"
|
||||
eqndD = hlF Syntax $ ifUnicode "≡" "=="
|
||||
dlamD = hlF Syntax $ ifUnicode "λᴰ" "dfun"
|
||||
annD = hlF Syntax $ ifUnicode "∷" "::"
|
||||
|
||||
private %inline
|
||||
typeD, eqD, colonD : Doc HL
|
||||
typeD = hl Syntax "Type"
|
||||
eqD = hl Syntax "Eq"
|
||||
colonD = hl Syntax ":"
|
||||
typeD, eqD, colonD, commaD, caseD, returnD, ofD : Doc HL
|
||||
typeD = hl Syntax "Type"
|
||||
eqD = hl Syntax "Eq"
|
||||
colonD = hl Syntax ":"
|
||||
commaD = hl Syntax ","
|
||||
caseD = hl Syntax "case"
|
||||
ofD = hl Syntax "of"
|
||||
returnD = hl Syntax "return"
|
||||
|
||||
mutual
|
||||
export covering
|
||||
|
@ -30,12 +37,15 @@ mutual
|
|||
prettyM (TYPE l) =
|
||||
parensIfM App $ typeD <//> !(withPrec Arg $ prettyM l)
|
||||
prettyM (Pi qty x s t) =
|
||||
parensIfM Outer $ hang 2 $
|
||||
!(prettyBinder [qty] x s) <++> !arrowD
|
||||
<//> !(under T x $ prettyM t)
|
||||
prettyBindType [qty] x s !arrowD t
|
||||
prettyM (Lam x t) =
|
||||
let GotLams {names, body, _} = getLams' [x] t.term Refl in
|
||||
prettyLams T (toList names) body
|
||||
prettyM (Sig x s t) =
|
||||
prettyBindType [] x s !timesD t
|
||||
prettyM (Pair s t) =
|
||||
let GotPairs {init, last, _} = getPairs t in
|
||||
prettyTuple $ s :: init ++ [last]
|
||||
prettyM (Eq _ (DUnused ty) l r) =
|
||||
parensIfM Eq !(withPrec InEq $ pure $
|
||||
sep [!(prettyM l) <++> !eqndD, !(prettyM r) <++> colonD, !(prettyM ty)])
|
||||
|
@ -43,7 +53,7 @@ mutual
|
|||
parensIfM App $
|
||||
eqD <++>
|
||||
sep [bracks !(withPrec Outer $ pure $ hang 2 $
|
||||
sep [hl DVar !(prettyM i) <++> !arrowD,
|
||||
sep [hl DVar !(prettyM i) <++> !darrowD,
|
||||
!(under D i $ prettyM ty)]),
|
||||
!(withPrec Arg $ prettyM l),
|
||||
!(withPrec Arg $ prettyM r)]
|
||||
|
@ -67,6 +77,10 @@ mutual
|
|||
prettyM (e :@ s) =
|
||||
let GotArgs {fun, args, _} = getArgs' e [s] in
|
||||
prettyApps fun args
|
||||
prettyM (CasePair pi p r ret x y body) = do
|
||||
pat <- (parens . separate commaD . map (hl TVar)) <$>
|
||||
traverse prettyM [x, y]
|
||||
prettyCase pi p r ret [([x, y], pat, body)]
|
||||
prettyM (e :% d) =
|
||||
let GotDArgs {fun, args, _} = getDArgs' e [d] in
|
||||
prettyApps fun args
|
||||
|
@ -90,6 +104,15 @@ mutual
|
|||
TSubst q d from to -> m (Doc HL)
|
||||
prettyTSubst s = prettySubstM prettyM (!ask).tnames TVar "[" "]" s
|
||||
|
||||
export covering
|
||||
prettyBindType : Pretty.HasEnv m => PrettyHL q =>
|
||||
List q -> Name -> Term q d n -> Doc HL ->
|
||||
ScopeTerm q d n -> m (Doc HL)
|
||||
prettyBindType qtys x s arr t =
|
||||
parensIfM Outer $ hang 2 $
|
||||
!(prettyBinder qtys x s) <++> arr
|
||||
<//> !(under T x $ prettyM t)
|
||||
|
||||
export covering
|
||||
prettyBinder : Pretty.HasEnv m => PrettyHL q =>
|
||||
List q -> Name -> Term q d n -> m (Doc HL)
|
||||
|
@ -104,7 +127,7 @@ mutual
|
|||
BinderSort -> List Name -> Term q _ _ -> m (Doc HL)
|
||||
prettyLams sort names body = do
|
||||
lam <- case sort of T => lamD; D => dlamD
|
||||
header <- sequence $ [hl TVar <$> prettyM x | x <- names] ++ [arrowD]
|
||||
header <- sequence $ [hl TVar <$> prettyM x | x <- names] ++ [darrowD]
|
||||
body <- unders sort names $ prettyM body
|
||||
parensIfM Outer $ sep (lam :: header) <//> body
|
||||
|
||||
|
@ -114,3 +137,30 @@ mutual
|
|||
prettyApps fun args =
|
||||
parensIfM App =<< withPrec Arg
|
||||
[|prettyM fun <//> (align . sep <$> traverse prettyM args)|]
|
||||
|
||||
export covering
|
||||
prettyTuple : Pretty.HasEnv m => PrettyHL a => List a -> m (Doc HL)
|
||||
prettyTuple = map (parens . align . separate commaD) . traverse prettyM
|
||||
|
||||
export covering
|
||||
prettyArm : Pretty.HasEnv m => PrettyHL a =>
|
||||
(List Name, Doc HL, a) -> m (Doc HL)
|
||||
prettyArm (xs, pat, body) =
|
||||
pure $ hang 2 $ sep
|
||||
[hsep [pat, !darrowD],
|
||||
!(withPrec Outer $ unders T xs $ prettyM body)]
|
||||
|
||||
export covering
|
||||
prettyArms : Pretty.HasEnv m => PrettyHL a =>
|
||||
List (List Name, Doc HL, a) -> m (Doc HL)
|
||||
prettyArms = map (braces . align . sep) . traverse prettyArm
|
||||
|
||||
export covering
|
||||
prettyCase : Pretty.HasEnv m =>
|
||||
PrettyHL q => PrettyHL a => PrettyHL b => PrettyHL c =>
|
||||
q -> a -> Name -> b -> List (List Name, Doc HL, c) -> m (Doc HL)
|
||||
prettyCase pi elim r ret arms =
|
||||
pure $ align $ sep $
|
||||
[hsep [caseD, !(prettyM elim), !(prettyQtyBinds [pi])],
|
||||
hsep [returnD, !(prettyM r), !darrowD, !(under T r $ prettyM ret)],
|
||||
hsep [ofD, !(prettyArms arms)]]
|
||||
|
|
|
@ -3,6 +3,7 @@ module Quox.Syntax.Term.Reduce
|
|||
import Quox.No
|
||||
import Quox.Syntax.Term.Base
|
||||
import Quox.Syntax.Term.Subst
|
||||
import Data.Vect
|
||||
import Data.Maybe
|
||||
|
||||
%default total
|
||||
|
@ -66,6 +67,10 @@ mutual
|
|||
ncloT $ Pi qty x (subs a th ph) (subs body th ph)
|
||||
pushSubstsWith th ph (Lam x body) =
|
||||
ncloT $ Lam x $ subs body th ph
|
||||
pushSubstsWith th ph (Sig x a b) =
|
||||
ncloT $ Sig x (subs a th ph) (subs b th ph)
|
||||
pushSubstsWith th ph (Pair s t) =
|
||||
ncloT $ Pair (subs s th ph) (subs t th ph)
|
||||
pushSubstsWith th ph (Eq i ty l r) =
|
||||
ncloT $ Eq i (subs ty th ph) (subs l th ph) (subs r th ph)
|
||||
pushSubstsWith th ph (DLam i body) =
|
||||
|
@ -96,6 +101,8 @@ mutual
|
|||
Right no => Element res no
|
||||
pushSubstsWith th ph (f :@ s) =
|
||||
ncloE $ subs f th ph :@ subs s th ph
|
||||
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)
|
||||
pushSubstsWith th ph (f :% d) =
|
||||
ncloE $ subs f th ph :% (d // th)
|
||||
pushSubstsWith th ph (s :# a) =
|
||||
|
@ -141,6 +148,11 @@ isDLamHead : Elim {} -> Bool
|
|||
isDLamHead (DLam {} :# Eq {}) = True
|
||||
isDLamHead _ = False
|
||||
|
||||
public export %inline
|
||||
isPairHead : Elim {} -> Bool
|
||||
isPairHead (Pair {} :# Sig {}) = True
|
||||
isPairHead _ = False
|
||||
|
||||
public export %inline
|
||||
isE : Term {} -> Bool
|
||||
isE (E _) = True
|
||||
|
@ -156,13 +168,14 @@ parameters (g : Lookup q d n)
|
|||
namespace Elim
|
||||
public export
|
||||
isRedex : Elim q d n -> Bool
|
||||
isRedex (F x) = isJust $ g x
|
||||
isRedex (B _) = False
|
||||
isRedex (f :@ _) = isRedex f || isLamHead f
|
||||
isRedex (f :% _) = isRedex f || isDLamHead f
|
||||
isRedex (t :# a) = isE t || isRedex t || isRedex a
|
||||
isRedex (CloE {}) = True
|
||||
isRedex (DCloE {}) = True
|
||||
isRedex (F x) = isJust $ g x
|
||||
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
|
||||
public export
|
||||
|
@ -211,24 +224,36 @@ parameters (g : Lookup q d n)
|
|||
whnf $ sub1 body s :# sub1 res s
|
||||
Right nlh => Element (f :@ s) $ fnf `orNo` nlh
|
||||
|
||||
whnf (CasePair pi pair r ret x y body) =
|
||||
let Element pair pairnf = whnf pair in
|
||||
case nchoose $ isPairHead pair of
|
||||
Left _ =>
|
||||
let Pair {fst, snd} :# Sig {fst = tfst, snd = tsnd, _} = pair
|
||||
fst = fst :# tfst
|
||||
snd = snd :# sub1 tsnd fst
|
||||
in
|
||||
whnf $ subN body [fst, snd] :# sub1 ret pair
|
||||
Right np =>
|
||||
Element (CasePair pi pair r ret x y body) $ pairnf `orNo` np
|
||||
|
||||
whnf (f :% p) =
|
||||
let Element f fnf = whnf f in
|
||||
case nchoose $ isDLamHead f of
|
||||
Left _ =>
|
||||
let DLam {body, _} :# Eq {ty, l, r, _} = f
|
||||
body = case p of K e => pick l r e; _ => dsub1 body p
|
||||
body = endsOr l r (dsub1 body p) p
|
||||
in
|
||||
whnf $ body :# dsub1 ty p
|
||||
Right ndlh =>
|
||||
Element (f :% p) $ fnf `orNo` ndlh
|
||||
|
||||
whnf (s :# a) =
|
||||
let Element s snf = whnf s
|
||||
Element a anf = whnf a
|
||||
in
|
||||
let Element s snf = whnf s in
|
||||
case nchoose $ isE s of
|
||||
Left _ => let E e = s in Element e $ noOr2 snf
|
||||
Right ne => Element (s :# a) $ ne `orNo` snf `orNo` anf
|
||||
Right ne =>
|
||||
let Element a anf = whnf a in
|
||||
Element (s :# a) $ ne `orNo` snf `orNo` anf
|
||||
|
||||
whnf (CloE el th) = whnf $ pushSubstsWith' id th el
|
||||
whnf (DCloE el th) = whnf $ pushSubstsWith' th id el
|
||||
|
@ -236,11 +261,13 @@ parameters (g : Lookup q d n)
|
|||
namespace Term
|
||||
export covering
|
||||
whnf : Term q d n -> NonRedexTerm q d n g
|
||||
whnf (TYPE l) = Element (TYPE l) Ah
|
||||
whnf (Pi qty x arg res) = Element (Pi qty x arg res) Ah
|
||||
whnf (Lam x body) = Element (Lam x body) Ah
|
||||
whnf (Eq i ty l r) = Element (Eq i ty l r) Ah
|
||||
whnf (DLam i body) = Element (DLam i body) Ah
|
||||
whnf t@(TYPE {}) = Element t Ah
|
||||
whnf t@(Pi {}) = Element t Ah
|
||||
whnf t@(Lam {}) = Element t Ah
|
||||
whnf t@(Sig {}) = Element t Ah
|
||||
whnf t@(Pair {}) = Element t Ah
|
||||
whnf t@(Eq {}) = Element t Ah
|
||||
whnf t@(DLam {}) = Element t Ah
|
||||
|
||||
whnf (E e) =
|
||||
let Element e enf = whnf e in
|
||||
|
|
|
@ -29,6 +29,16 @@ public export
|
|||
NotDLam = No . isDLam
|
||||
|
||||
|
||||
public export %inline
|
||||
isPair : Term {} -> Bool
|
||||
isPair (Pair {}) = True
|
||||
isPair _ = False
|
||||
|
||||
public export
|
||||
0 NotPair : Pred $ Term {}
|
||||
NotPair = No . isPair
|
||||
|
||||
|
||||
public export %inline
|
||||
isApp : Elim {} -> Bool
|
||||
isApp (_ :@ _) = True
|
||||
|
@ -159,8 +169,8 @@ record GetDLams q d n where
|
|||
|
||||
export
|
||||
getDLams' : forall lams, rest.
|
||||
Vect lams Name -> Term q rest n -> (0 eq : lams + d = rest) ->
|
||||
GetDLams q d n
|
||||
Vect lams Name -> Term q rest n -> (0 eq : lams + d = rest) ->
|
||||
GetDLams q d n
|
||||
getDLams' is s Refl = case nchoose $ isDLam s of
|
||||
Left y => let DLam i body = s in
|
||||
getDLams' (i :: is) (assert_smaller s body.term) Refl
|
||||
|
@ -169,3 +179,19 @@ getDLams' is s Refl = case nchoose $ isDLam s of
|
|||
export %inline
|
||||
getDLams : Term q d n -> GetDLams q d n
|
||||
getDLams s = getDLams' [] s Refl
|
||||
|
||||
|
||||
public export
|
||||
record GetPairs q d n where
|
||||
constructor GotPairs
|
||||
init : List $ Term q d n
|
||||
last : Term q d n
|
||||
notPair : NotPair last
|
||||
|
||||
export
|
||||
getPairs : Term q d n -> GetPairs q d n
|
||||
getPairs t = case nchoose $ isPair t of
|
||||
Left y =>
|
||||
let Pair s t = t; GotPairs ts t np = getPairs t in
|
||||
GotPairs (s :: ts) t np
|
||||
Right n => GotPairs [] t n
|
||||
|
|
|
@ -1,6 +1,7 @@
|
|||
module Quox.Syntax.Term.Subst
|
||||
|
||||
import Quox.Syntax.Term.Base
|
||||
import Data.Vect
|
||||
|
||||
%default total
|
||||
|
||||
|
@ -195,15 +196,25 @@ namespace DScopeTermN
|
|||
(DUsed body).term = body
|
||||
(DUnused body).term = body /// shift s
|
||||
|
||||
|
||||
export %inline
|
||||
subN : ScopeTermN s q d n -> Vect s (Elim q d n) -> Term q d n
|
||||
subN (TUsed body) es = body // fromVect es
|
||||
subN (TUnused body) _ = body
|
||||
|
||||
export %inline
|
||||
sub1 : ScopeTerm q d n -> Elim q d n -> Term q d n
|
||||
sub1 (TUsed body) e = body // one e
|
||||
sub1 (TUnused body) e = body
|
||||
sub1 t e = subN t [e]
|
||||
|
||||
export %inline
|
||||
dsubN : DScopeTermN s q d n -> Vect s (Dim d) -> Term q d n
|
||||
dsubN (DUsed body) ps = body /// fromVect ps
|
||||
dsubN (DUnused body) _ = body
|
||||
|
||||
export %inline
|
||||
dsub1 : DScopeTerm q d n -> Dim d -> Term q d n
|
||||
dsub1 (DUsed body) p = body /// one p
|
||||
dsub1 (DUnused body) p = body
|
||||
dsub1 t p = dsubN t [p]
|
||||
|
||||
|
||||
public export %inline
|
||||
(.zero) : DScopeTerm q d n -> Term q d n
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue