natural numbers
This commit is contained in:
parent
fae534dae0
commit
9250789219
15 changed files with 305 additions and 10 deletions
|
@ -75,6 +75,12 @@ mutual
|
|||
||| equality term
|
||||
DLam : (body : DScopeTerm q d n) -> Term q d n
|
||||
|
||||
||| natural numbers (temporary until 𝐖 gets added)
|
||||
Nat : Term q d n
|
||||
-- [todo] can these be elims?
|
||||
Zero : Term q d n
|
||||
Succ : (p : Term q d n) -> Term q d n
|
||||
|
||||
||| elimination
|
||||
E : (e : Elim q d n) -> Term q d n
|
||||
|
||||
|
@ -111,6 +117,13 @@ mutual
|
|||
(arms : CaseEnumArms q d n) ->
|
||||
Elim q d n
|
||||
|
||||
||| nat matching
|
||||
CaseNat : (qty, qtyIH : q) -> (nat : Elim q d n) ->
|
||||
(ret : ScopeTerm q d n) ->
|
||||
(zero : Term q d n) ->
|
||||
(succ : ScopeTermN 2 q d n) ->
|
||||
Elim q d n
|
||||
|
||||
||| dim application
|
||||
(:%) : (fun : Elim q d n) -> (arg : Dim d) -> Elim q d n
|
||||
|
||||
|
|
|
@ -11,8 +11,8 @@ import Data.Vect
|
|||
%default total
|
||||
|
||||
|
||||
private %inline
|
||||
typeD, arrowD, darrowD, timesD, lamD, eqndD, dlamD, annD :
|
||||
export %inline
|
||||
typeD, arrowD, darrowD, timesD, lamD, eqndD, dlamD, annD, natD :
|
||||
Pretty.HasEnv m => m (Doc HL)
|
||||
typeD = hlF Syntax $ ifUnicode "★" "Type"
|
||||
arrowD = hlF Syntax $ ifUnicode "→" "->"
|
||||
|
@ -22,9 +22,10 @@ lamD = hlF Syntax $ ifUnicode "λ" "fun"
|
|||
eqndD = hlF Syntax $ ifUnicode "≡" "=="
|
||||
dlamD = hlF Syntax $ ifUnicode "δ" "dfun"
|
||||
annD = hlF Syntax $ ifUnicode "∷" "::"
|
||||
natD = hlF Syntax $ ifUnicode "ℕ" "Nat"
|
||||
|
||||
private %inline
|
||||
eqD, colonD, commaD, caseD, returnD, ofD, dotD : Doc HL
|
||||
export %inline
|
||||
eqD, colonD, commaD, caseD, returnD, ofD, dotD, zeroD, succD : Doc HL
|
||||
eqD = hl Syntax "Eq"
|
||||
colonD = hl Syntax ":"
|
||||
commaD = hl Syntax ","
|
||||
|
@ -32,6 +33,8 @@ caseD = hl Syntax "case"
|
|||
ofD = hl Syntax "of"
|
||||
returnD = hl Syntax "return"
|
||||
dotD = hl Delim "."
|
||||
zeroD = hl Syntax "zero"
|
||||
succD = hl Syntax "succ"
|
||||
|
||||
|
||||
export
|
||||
|
@ -157,6 +160,16 @@ export
|
|||
prettyTagBare : TagVal -> Doc HL
|
||||
prettyTagBare t = hl Tag $ quoteTag t
|
||||
|
||||
export
|
||||
toNatLit : Term q d n -> Maybe Nat
|
||||
toNatLit Zero = Just 0
|
||||
toNatLit (Succ n) = [|S $ toNatLit n|]
|
||||
toNatLit _ = Nothing
|
||||
|
||||
private
|
||||
eterm : Term q d n -> Exists (Term q d)
|
||||
eterm = Evidence n
|
||||
|
||||
|
||||
parameters (showSubsts : Bool)
|
||||
mutual
|
||||
|
@ -201,6 +214,14 @@ parameters (showSubsts : Bool)
|
|||
prettyM (DLam (S i t)) =
|
||||
let GotDLams {names, body, _} = getDLams' i t.term Refl in
|
||||
prettyLams (Just !dlamD) D (toSnocList' names) body
|
||||
prettyM Nat = natD
|
||||
prettyM Zero = pure $ hl Syntax "0"
|
||||
prettyM (Succ n) =
|
||||
case toNatLit n of
|
||||
Just n => pure $ hl Syntax $ pretty $ S n
|
||||
Nothing => do
|
||||
n <- withPrec Arg $ prettyM n
|
||||
parensIfM App $ succD <++> n
|
||||
prettyM (E e) = prettyM e
|
||||
prettyM (CloT s th) =
|
||||
if showSubsts then
|
||||
|
@ -231,6 +252,14 @@ parameters (showSubsts : Bool)
|
|||
prettyM (CaseEnum pi t (S [< r] ret) arms) =
|
||||
prettyCase pi t r ret.term
|
||||
[([<], prettyTag t, b) | (t, b) <- SortedMap.toList arms]
|
||||
prettyM (CaseNat pi pi' nat (S [< r] ret) zer (S [< s, ih] suc)) =
|
||||
prettyCase pi nat r ret.term
|
||||
[([<], zeroD, eterm zer),
|
||||
([< s, ih], !succPat, eterm suc.term)]
|
||||
where
|
||||
succPat : m (Doc HL)
|
||||
succPat = pure $
|
||||
hsep [succD, !(pretty0M s), bracks !(pretty0M $ MkWithQty pi' ih)]
|
||||
prettyM (e :% d) =
|
||||
let GotDArgs {fun, args, _} = getDArgs' e [d] in
|
||||
prettyApps (Just "@") fun args
|
||||
|
|
|
@ -268,6 +268,9 @@ mutual
|
|||
nclo $ Eq (ty // th // ph) (l // th // ph) (r // th // ph)
|
||||
pushSubstsWith th ph (DLam body) =
|
||||
nclo $ DLam (body // th // ph)
|
||||
pushSubstsWith _ _ Nat = nclo Nat
|
||||
pushSubstsWith _ _ Zero = nclo Zero
|
||||
pushSubstsWith th ph (Succ n) = nclo $ Succ $ n // th // ph
|
||||
pushSubstsWith th ph (E e) =
|
||||
let Element e nc = pushSubstsWith th ph e in nclo $ E e
|
||||
pushSubstsWith th ph (CloT s ps) =
|
||||
|
@ -291,6 +294,9 @@ mutual
|
|||
pushSubstsWith th ph (CaseEnum pi t r arms) =
|
||||
nclo $ CaseEnum pi (t // th // ph) (r // th // ph)
|
||||
(map (\b => b // th // ph) arms)
|
||||
pushSubstsWith th ph (CaseNat pi pi' n r z s) =
|
||||
nclo $ CaseNat pi pi' (n // th // ph) (r // th // ph)
|
||||
(z // th // ph) (s // th // ph)
|
||||
pushSubstsWith th ph (f :% d) =
|
||||
nclo $ (f // th // ph) :% (d // th)
|
||||
pushSubstsWith th ph (s :# a) =
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue