remove square brackets around type lines
(parens are needed if they are anything other than a `term arg`)
This commit is contained in:
parent
d631b86be3
commit
64de93a13c
7 changed files with 42 additions and 35 deletions
|
@ -17,7 +17,7 @@ def0 If : 1.Bool → 0.★₀ → 0.★₀ → ★₀ =
|
|||
def0 T : ω.Bool → ★₀ = λ b ⇒ If b True False;
|
||||
|
||||
def true-not-false : Not ('true ≡ 'false : Bool) =
|
||||
λ eq ⇒ coe [i ⇒ T (eq @i)] 'true;
|
||||
λ eq ⇒ coe (i ⇒ T (eq @i)) 'true;
|
||||
|
||||
|
||||
-- [todo] infix
|
||||
|
|
|
@ -13,7 +13,7 @@ def0 All : 0.(A : ★₀) → 0.(Pred A) → ★₁ =
|
|||
|
||||
def cong :
|
||||
0.(A : ★₀) → 0.(P : Pred A) → 1.(p : All A P) →
|
||||
0.(x y : A) → 1.(xy : x ≡ y : A) → Eq [𝑖 ⇒ P (xy @𝑖)] (p x) (p y) =
|
||||
0.(x y : A) → 1.(xy : x ≡ y : A) → Eq (𝑖 ⇒ P (xy @𝑖)) (p x) (p y) =
|
||||
λ A P p x y xy ⇒ δ 𝑖 ⇒ p (xy @𝑖);
|
||||
|
||||
def0 eq-f :
|
||||
|
@ -28,9 +28,9 @@ def funext :
|
|||
λ A P p q eq ⇒ δ 𝑖 ⇒ λ x ⇒ eq x @𝑖;
|
||||
|
||||
def sym : 0.(A : ★₀) → 0.(x y : A) → 1.(x ≡ y : A) → y ≡ x : A =
|
||||
λ A x y eq ⇒ δ 𝑖 ⇒ comp [A] (eq @0) @𝑖 { 0 𝑗 ⇒ eq @𝑗; 1 _ ⇒ eq @0 };
|
||||
λ A x y eq ⇒ δ 𝑖 ⇒ comp A (eq @0) @𝑖 { 0 𝑗 ⇒ eq @𝑗; 1 _ ⇒ eq @0 };
|
||||
|
||||
def trans : 0.(A : ★₀) → 0.(x y z : A) →
|
||||
ω.(x ≡ y : A) → ω.(y ≡ z : A) → x ≡ z : A =
|
||||
λ A x y z eq1 eq2 ⇒ δ 𝑖 ⇒
|
||||
comp [A] (eq1 @𝑖) @𝑖 { 0 _ ⇒ eq1 @0; 1 𝑗 ⇒ eq2 @𝑗 };
|
||||
comp A (eq1 @𝑖) @𝑖 { 0 _ ⇒ eq1 @0; 1 𝑗 ⇒ eq2 @𝑗 };
|
||||
|
|
|
@ -48,10 +48,10 @@ def isSucc? : ω.(n : ℕ) → Dec (IsSucc n) =
|
|||
};
|
||||
|
||||
def zero-not-succ : 0.(m : ℕ) → Not (zero ≡ succ m : ℕ) =
|
||||
λ m eq ⇒ coe [𝑖 ⇒ IsSucc (eq @𝑖)] @1 @0 'true;
|
||||
λ m eq ⇒ coe (𝑖 ⇒ IsSucc (eq @𝑖)) @1 @0 'true;
|
||||
|
||||
def succ-not-zero : 0.(m : ℕ) → Not (succ m ≡ zero : ℕ) =
|
||||
λ m eq ⇒ coe [𝑖 ⇒ IsSucc (eq @𝑖)] 'true;
|
||||
λ m eq ⇒ coe (𝑖 ⇒ IsSucc (eq @𝑖)) 'true;
|
||||
|
||||
|
||||
def0 not-succ-self : 0.(m : ℕ) → Not (m ≡ succ m : ℕ) =
|
||||
|
|
|
@ -245,15 +245,6 @@ export
|
|||
term : FileName -> Grammar True PTerm
|
||||
-- defined after all the subterm parsers
|
||||
|
||||
export
|
||||
typeLine : FileName -> Grammar True (PatVar, PTerm)
|
||||
typeLine fname = do
|
||||
resC "["
|
||||
mustWork $ do
|
||||
i <- patVar fname <* resC "⇒" <|> unused fname
|
||||
t <- assert_total term fname <* needRes "]"
|
||||
pure (i, t)
|
||||
|
||||
||| box term `[t]` or type `[π.A]`
|
||||
export
|
||||
boxTerm : FileName -> Grammar True PTerm
|
||||
|
@ -287,6 +278,19 @@ termArg fname = withLoc fname $
|
|||
<|> [|V qname|]
|
||||
<|> const <$> tupleTerm fname
|
||||
|
||||
export
|
||||
properTypeLine : FileName -> Grammar True (PatVar, PTerm)
|
||||
properTypeLine fname = do
|
||||
resC "("
|
||||
i <- patVar fname <* resC "⇒" <|> unused fname
|
||||
t <- assert_total term fname <* needRes ")"
|
||||
pure (i, t)
|
||||
|
||||
export
|
||||
typeLine : FileName -> Grammar True (PatVar, PTerm)
|
||||
typeLine fname =
|
||||
properTypeLine fname <|> [|(,) (unused fname) (termArg fname)|]
|
||||
|
||||
||| optionally, two dimension arguments. if absent default to `@0 @1`
|
||||
private
|
||||
optDirection : FileName -> Grammar False (PDim, PDim)
|
||||
|
|
|
@ -148,6 +148,10 @@ export %inline
|
|||
braces : {opts : _} -> Doc opts -> Eff Pretty (Doc opts)
|
||||
braces = looseDelims "{" "}"
|
||||
|
||||
export %inline
|
||||
tightBraces : {opts : _} -> Doc opts -> Eff Pretty (Doc opts)
|
||||
tightBraces = tightDelims "{" "}"
|
||||
|
||||
export %inline
|
||||
parensIf : {opts : _} -> Bool -> Doc opts -> Eff Pretty (Doc opts)
|
||||
parensIf True = parens
|
||||
|
|
|
@ -136,9 +136,9 @@ prettyTypeLine : {opts : _} ->
|
|||
DScopeTerm d n ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyTypeLine dnames tnames (S _ (N body)) =
|
||||
bracks =<< withPrec Outer (prettyTerm dnames tnames body)
|
||||
withPrec Arg (prettyTerm dnames tnames body)
|
||||
prettyTypeLine dnames tnames (S [< i] (Y body)) =
|
||||
bracks =<< do
|
||||
parens =<< do
|
||||
i' <- prettyDBind i
|
||||
ty' <- withPrec Outer $ prettyTerm (dnames :< i) tnames body
|
||||
pure $ sep [hsep [i', !darrowD], ty']
|
||||
|
@ -281,9 +281,9 @@ prettyTag tag = hl Tag $ text $ "'" ++ quoteTag tag
|
|||
export
|
||||
prettyEnum : {opts : _} -> List String -> Eff Pretty (Doc opts)
|
||||
prettyEnum cases =
|
||||
tightDelims "{" "}" =<<
|
||||
tightBraces =<<
|
||||
fillSeparateTight !commaD <$>
|
||||
traverse (hl Tag . text . quoteTag) cases
|
||||
traverse (hl Tag . Doc.text . quoteTag) cases
|
||||
|
||||
private
|
||||
prettyCaseRet : {opts : _} ->
|
||||
|
|
|
@ -228,13 +228,13 @@ tests = "parser" :- [
|
|||
],
|
||||
|
||||
"equality type" :- [
|
||||
parseMatch term "Eq [i ⇒ A] s t"
|
||||
parseMatch term "Eq (i ⇒ A) s t"
|
||||
`(Eq (PV "i" _, V "A" _) (V "s" _) (V "t" _) _),
|
||||
parseMatch term "Eq [i ⇒ A (B @i)] (f x) (g y)"
|
||||
parseMatch term "Eq (i ⇒ A (B @i)) (f x) (g y)"
|
||||
`(Eq (PV "i" _, App (V "A" _) (DApp (V "B" _) (V "i" _) _) _)
|
||||
(App (V "f" _) (V "x" _) _)
|
||||
(App (V "g" _) (V "y" _) _) _),
|
||||
parseMatch term "Eq [A] s t"
|
||||
parseMatch term "Eq A s t"
|
||||
`(Eq (Unused _, V "A" _) (V "s" _) (V "t" _) _),
|
||||
parseMatch term "s ≡ t : A"
|
||||
`(Eq (Unused _, V "A" _) (V "s" _) (V "t" _) _),
|
||||
|
@ -286,33 +286,32 @@ tests = "parser" :- [
|
|||
],
|
||||
|
||||
"coe" :- [
|
||||
parseMatch term "coe [A] @p @q x"
|
||||
parseMatch term "coe A @p @q x"
|
||||
`(Coe (Unused _, V "A" _) (V "p" _) (V "q" _) (V "x" _) _),
|
||||
parseMatch term "coe [i ⇒ A] @p @q x"
|
||||
parseMatch term "coe (i ⇒ A) @p @q x"
|
||||
`(Coe (PV "i" _, V "A" _) (V "p" _) (V "q" _) (V "x" _) _),
|
||||
parseMatch term "coe [A] x"
|
||||
parseMatch term "coe A x"
|
||||
`(Coe (Unused _, V "A" _) (K Zero _) (K One _) (V "x" _) _),
|
||||
parseFails term "coe [A] @p @q",
|
||||
parseFails term "coe A @p @q x",
|
||||
parseFails term "coe [i ⇒ A] @p q x"
|
||||
parseFails term "coe A @p @q",
|
||||
parseFails term "coe (i ⇒ A) @p q x"
|
||||
],
|
||||
|
||||
"comp" :- [
|
||||
parseMatch term "comp [A] @p @q s @r { 0 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁ }"
|
||||
parseMatch term "comp A @p @q s @r { 0 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁ }"
|
||||
`(Comp (Unused _, V "A" _) (V "p" _) (V "q" _) (V "s" _) (V "r" _)
|
||||
(PV "𝑗" _, V "s₀" _) (PV "𝑘" _, V "s₁" _) _),
|
||||
parseMatch term "comp [𝑖 ⇒ A] @p @q s @r { 0 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁ }"
|
||||
parseMatch term "comp (𝑖 ⇒ A) @p @q s @r { 0 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁ }"
|
||||
`(Comp (PV "𝑖" _, V "A" _) (V "p" _) (V "q" _) (V "s" _) (V "r" _)
|
||||
(PV "𝑗" _, V "s₀" _) (PV "𝑘" _, V "s₁" _) _),
|
||||
parseMatch term "comp [A] @p @q s @r { 1 𝑗 ⇒ s₀; 0 𝑘 ⇒ s₁; }"
|
||||
parseMatch term "comp A @p @q s @r { 1 𝑗 ⇒ s₀; 0 𝑘 ⇒ s₁; }"
|
||||
`(Comp (Unused _, V "A" _) (V "p" _) (V "q" _) (V "s" _) (V "r" _)
|
||||
(PV "𝑘" _, V "s₁" _) (PV "𝑗" _, V "s₀" _) _),
|
||||
parseMatch term "comp [A] s @r { 0 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁ }"
|
||||
parseMatch term "comp A s @r { 0 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁ }"
|
||||
`(Comp (Unused _, V "A" _) (K Zero _) (K One _) (V "s" _) (V "r" _)
|
||||
(PV "𝑗" _, V "s₀" _) (PV "𝑘" _, V "s₁" _) _),
|
||||
parseFails term "comp [A] @p @q s @r { 1 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁; }",
|
||||
parseFails term "comp [A] @p @q s @r { 0 𝑗 ⇒ s₀ }",
|
||||
parseFails term "comp [A] @p @q s @r { }"
|
||||
parseFails term "comp A @p @q s @r { 1 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁; }",
|
||||
parseFails term "comp A @p @q s @r { 0 𝑗 ⇒ s₀ }",
|
||||
parseFails term "comp A @p @q s @r { }"
|
||||
],
|
||||
|
||||
"case" :- [
|
||||
|
|
Loading…
Reference in a new issue