remove universe type
This commit is contained in:
parent
0cae84c75b
commit
b7acf39c39
8 changed files with 21 additions and 44 deletions
|
@ -84,7 +84,7 @@ mutual
|
|||
(t : Term Three d n) -> (0 _ : NotClo t) =>
|
||||
PTerm
|
||||
toPTermWith' ds ns s = case s of
|
||||
TYPE (U l) =>
|
||||
TYPE l =>
|
||||
TYPE l
|
||||
Pi qty arg (S [x] res) =>
|
||||
Pi qty (Just x) (toPTermWith ds ns arg)
|
||||
|
|
|
@ -6,5 +6,4 @@ import public Quox.Syntax.Qty
|
|||
import public Quox.Syntax.Shift
|
||||
import public Quox.Syntax.Subst
|
||||
import public Quox.Syntax.Term
|
||||
import public Quox.Syntax.Universe
|
||||
import public Quox.Syntax.Var
|
||||
|
|
|
@ -3,7 +3,6 @@ module Quox.Syntax.Term.Base
|
|||
import public Quox.Syntax.Var
|
||||
import public Quox.Syntax.Shift
|
||||
import public Quox.Syntax.Subst
|
||||
import public Quox.Syntax.Universe
|
||||
import public Quox.Syntax.Qty
|
||||
import public Quox.Syntax.Dim
|
||||
import public Quox.Name
|
||||
|
@ -32,6 +31,10 @@ public export
|
|||
0 TSubstLike : Type
|
||||
TSubstLike = Type -> Nat -> Nat -> Nat -> Type
|
||||
|
||||
public export
|
||||
0 Universe : Type
|
||||
Universe = Nat
|
||||
|
||||
public export
|
||||
0 TagVal : Type
|
||||
TagVal = String
|
||||
|
|
|
@ -32,6 +32,16 @@ ofD = hl Syntax "of"
|
|||
returnD = hl Syntax "return"
|
||||
|
||||
|
||||
export
|
||||
prettyUnivSuffix : Pretty.HasEnv m => Universe -> m (Doc HL)
|
||||
prettyUnivSuffix l =
|
||||
ifUnicode (pretty $ pack $ map sub $ unpack $ show l) (pretty l)
|
||||
where
|
||||
sub : Char -> Char
|
||||
sub c = case c of
|
||||
'0' => '₀'; '1' => '₁'; '2' => '₂'; '3' => '₃'; '4' => '₄'
|
||||
'5' => '₅'; '6' => '₆'; '7' => '₇'; '8' => '₈'; '9' => '₉'; _ => c
|
||||
|
||||
export
|
||||
prettyBindType : PrettyHL a => PrettyHL b => PrettyHL q =>
|
||||
Pretty.HasEnv m =>
|
||||
|
|
|
@ -1,34 +0,0 @@
|
|||
module Quox.Syntax.Universe
|
||||
|
||||
import Quox.Pretty
|
||||
|
||||
import Data.Fin
|
||||
import Derive.Prelude
|
||||
|
||||
%default total
|
||||
%language ElabReflection
|
||||
|
||||
|
||||
public export
|
||||
data Universe = U Nat
|
||||
%name Universe l
|
||||
|
||||
%runElab derive "Universe" [Eq, Ord, Show]
|
||||
|
||||
export
|
||||
PrettyHL Universe where
|
||||
prettyM (U l) = pure $ hl Free $ pretty l
|
||||
|
||||
export
|
||||
prettyUnivSuffix : Pretty.HasEnv m => Universe -> m (Doc HL)
|
||||
prettyUnivSuffix (U l) =
|
||||
ifUnicode (pretty $ pack $ map sub $ unpack $ show l) (pretty l)
|
||||
where
|
||||
sub : Char -> Char
|
||||
sub c = case c of
|
||||
'0' => '₀'; '1' => '₁'; '2' => '₂'; '3' => '₃'; '4' => '₄'
|
||||
'5' => '₅'; '6' => '₆'; '7' => '₇'; '8' => '₈'; '9' => '₉'; _ => c
|
||||
|
||||
export %inline
|
||||
fromInteger : (x : Integer) -> (0 _ : So (x >= 0)) => Universe
|
||||
fromInteger x = U $ fromInteger x
|
|
@ -26,7 +26,6 @@ modules =
|
|||
Quox.Syntax.Term.Pretty,
|
||||
Quox.Syntax.Term.Split,
|
||||
Quox.Syntax.Term.Subst,
|
||||
Quox.Syntax.Universe,
|
||||
Quox.Syntax.Var,
|
||||
Quox.Definition,
|
||||
Quox.Reduce,
|
||||
|
|
|
@ -260,11 +260,11 @@ tests = "equality & subtyping" :- [
|
|||
|
||||
"free var" :-
|
||||
let au_bu = fromList
|
||||
[("A", mkDef Any (TYPE (U 1)) (TYPE (U 0))),
|
||||
("B", mkDef Any (TYPE (U 1)) (TYPE (U 0)))]
|
||||
[("A", mkDef Any (TYPE 1) (TYPE 0)),
|
||||
("B", mkDef Any (TYPE 1) (TYPE 0))]
|
||||
au_ba = fromList
|
||||
[("A", mkDef Any (TYPE (U 1)) (TYPE (U 0))),
|
||||
("B", mkDef Any (TYPE (U 1)) (FT "A"))]
|
||||
[("A", mkDef Any (TYPE 1) (TYPE 0)),
|
||||
("B", mkDef Any (TYPE 1) (FT "A"))]
|
||||
in [
|
||||
testEq "A = A" $
|
||||
equalE empty (F "A") (F "A"),
|
||||
|
|
|
@ -63,8 +63,8 @@ PrettyHL q => ToInfo (Error q) where
|
|||
toInfo (ClashU mode k l) =
|
||||
[("type", "ClashU"),
|
||||
("mode", show mode),
|
||||
("left", prettyStr True k),
|
||||
("right", prettyStr True l)]
|
||||
("left", show k),
|
||||
("right", show l)]
|
||||
toInfo (ClashQ pi rh) =
|
||||
[("type", "ClashQ"),
|
||||
("left", prettyStr True pi),
|
||||
|
|
Loading…
Reference in a new issue