sop → elab-util
This commit is contained in:
parent
04d3c9368a
commit
fc3c2dc8ab
9 changed files with 40 additions and 31 deletions
|
@ -1,6 +1,6 @@
|
||||||
module Quox.CharExtra
|
module Quox.CharExtra
|
||||||
|
|
||||||
import Generics.Derive
|
import Derive.Prelude
|
||||||
|
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
|
@ -10,38 +10,38 @@ import Generics.Derive
|
||||||
namespace Letter
|
namespace Letter
|
||||||
public export
|
public export
|
||||||
data Letter = Uppercase | Lowercase | Titlecase | Modifier | Other
|
data Letter = Uppercase | Lowercase | Titlecase | Modifier | Other
|
||||||
%runElab derive "Letter" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Letter" [Eq, Ord, Show]
|
||||||
|
|
||||||
namespace Mark
|
namespace Mark
|
||||||
public export
|
public export
|
||||||
data Mark = NonSpacing | SpacingCombining | Enclosing
|
data Mark = NonSpacing | SpacingCombining | Enclosing
|
||||||
%runElab derive "Mark" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Mark" [Eq, Ord, Show]
|
||||||
|
|
||||||
namespace Number
|
namespace Number
|
||||||
public export
|
public export
|
||||||
data Number = Decimal | Letter | Other
|
data Number = Decimal | Letter | Other
|
||||||
%runElab derive "Number" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Number" [Eq, Ord, Show]
|
||||||
|
|
||||||
namespace Punctuation
|
namespace Punctuation
|
||||||
public export
|
public export
|
||||||
data Punctuation = Connector | Dash | Open | Close
|
data Punctuation = Connector | Dash | Open | Close
|
||||||
| InitialQuote | FinalQuote | Other
|
| InitialQuote | FinalQuote | Other
|
||||||
%runElab derive "Punctuation" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Punctuation" [Eq, Ord, Show]
|
||||||
|
|
||||||
namespace Symbol
|
namespace Symbol
|
||||||
public export
|
public export
|
||||||
data Symbol = Math | Currency | Modifier | Other
|
data Symbol = Math | Currency | Modifier | Other
|
||||||
%runElab derive "Symbol" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Symbol" [Eq, Ord, Show]
|
||||||
|
|
||||||
namespace Separator
|
namespace Separator
|
||||||
public export
|
public export
|
||||||
data Separator = Space | Line | Paragraph
|
data Separator = Space | Line | Paragraph
|
||||||
%runElab derive "Separator" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Separator" [Eq, Ord, Show]
|
||||||
|
|
||||||
namespace Other
|
namespace Other
|
||||||
public export
|
public export
|
||||||
data Other = Control | Format | Surrogate | PrivateUse | NotAssigned
|
data Other = Control | Format | Surrogate | PrivateUse | NotAssigned
|
||||||
%runElab derive "Other" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Other" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
|
@ -53,7 +53,7 @@ data GeneralCategory
|
||||||
| Symbol Symbol
|
| Symbol Symbol
|
||||||
| Separator Separator
|
| Separator Separator
|
||||||
| Other Other
|
| Other Other
|
||||||
%runElab derive "GeneralCategory" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "GeneralCategory" [Eq, Ord, Show]
|
||||||
|
|
||||||
private
|
private
|
||||||
%foreign "scheme:(lambda (c) (symbol->string (char-general-category c)))"
|
%foreign "scheme:(lambda (c) (symbol->string (char-general-category c)))"
|
||||||
|
|
|
@ -6,7 +6,7 @@ import Data.String.Extra
|
||||||
import Data.SortedMap
|
import Data.SortedMap
|
||||||
import Text.Lexer
|
import Text.Lexer
|
||||||
import Text.Lexer.Tokenizer
|
import Text.Lexer.Tokenizer
|
||||||
import Generics.Derive
|
import Derive.Prelude
|
||||||
%hide TT.Name
|
%hide TT.Name
|
||||||
|
|
||||||
|
|
||||||
|
@ -22,7 +22,7 @@ import Generics.Derive
|
||||||
||| @ TYPE "Type" or "★" with subscript
|
||| @ TYPE "Type" or "★" with subscript
|
||||||
public export
|
public export
|
||||||
data Token = R String | I Name | N Nat | S String | T String | TYPE Nat
|
data Token = R String | I Name | N Nat | S String | T String | TYPE Nat
|
||||||
%runElab derive "Token" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Token" [Eq, Ord, Show]
|
||||||
|
|
||||||
-- token or whitespace
|
-- token or whitespace
|
||||||
private
|
private
|
||||||
|
@ -37,8 +37,8 @@ record Error where
|
||||||
line, col : Int
|
line, col : Int
|
||||||
||| `Nothing` if the error is at the end of the input
|
||| `Nothing` if the error is at the end of the input
|
||||||
char : Maybe Char
|
char : Maybe Char
|
||||||
%runElab derive "StopReason" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "StopReason" [Eq, Ord, Show]
|
||||||
%runElab derive "Error" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Error" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
|
||||||
skip : Lexer -> Tokenizer TokenW
|
skip : Lexer -> Tokenizer TokenW
|
||||||
|
|
|
@ -2,7 +2,7 @@ module Quox.Name
|
||||||
|
|
||||||
import public Data.SnocList
|
import public Data.SnocList
|
||||||
import Data.List
|
import Data.List
|
||||||
import Generics.Derive
|
import Derive.Prelude
|
||||||
|
|
||||||
%hide TT.Name
|
%hide TT.Name
|
||||||
|
|
||||||
|
@ -13,7 +13,7 @@ import Generics.Derive
|
||||||
public export
|
public export
|
||||||
data BaseName
|
data BaseName
|
||||||
= UN String -- user-given name
|
= UN String -- user-given name
|
||||||
%runElab derive "BaseName" [Generic, Meta, Eq, Ord, DecEq]
|
%runElab derive "BaseName" [Eq, Ord]
|
||||||
|
|
||||||
export
|
export
|
||||||
Show BaseName where
|
Show BaseName where
|
||||||
|
@ -34,7 +34,7 @@ record Name where
|
||||||
constructor MakeName
|
constructor MakeName
|
||||||
mods : SnocList String
|
mods : SnocList String
|
||||||
base : BaseName
|
base : BaseName
|
||||||
%runElab derive "Name" [Generic, Meta, Eq, Ord, DecEq]
|
%runElab derive "Name" [Eq, Ord]
|
||||||
|
|
||||||
export
|
export
|
||||||
Show Name where
|
Show Name where
|
||||||
|
|
|
@ -10,7 +10,7 @@ import Data.DPair
|
||||||
|
|
||||||
import public Control.Monad.Identity
|
import public Control.Monad.Identity
|
||||||
import public Control.Monad.Reader
|
import public Control.Monad.Reader
|
||||||
import Generics.Derive
|
import Derive.Prelude
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
%language ElabReflection
|
%language ElabReflection
|
||||||
|
@ -36,7 +36,7 @@ data HL
|
||||||
| Qty
|
| Qty
|
||||||
| Syntax
|
| Syntax
|
||||||
| Tag
|
| Tag
|
||||||
%runElab derive "HL" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "HL" [Eq, Ord, Show]
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data PPrec
|
data PPrec
|
||||||
|
@ -49,7 +49,7 @@ data PPrec
|
||||||
| App -- term/dimension application
|
| App -- term/dimension application
|
||||||
| SApp -- substitution application
|
| SApp -- substitution application
|
||||||
| Arg -- argument to nonfix function
|
| Arg -- argument to nonfix function
|
||||||
%runElab derive "PPrec" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "PPrec" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
|
|
|
@ -6,18 +6,17 @@ import Quox.Pretty
|
||||||
|
|
||||||
import Decidable.Equality
|
import Decidable.Equality
|
||||||
import Control.Function
|
import Control.Function
|
||||||
import Generics.Derive
|
import Derive.Prelude
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
%language ElabReflection
|
%language ElabReflection
|
||||||
%hide SOP.from; %hide SOP.to
|
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data DimConst = Zero | One
|
data DimConst = Zero | One
|
||||||
%name DimConst e
|
%name DimConst e
|
||||||
|
|
||||||
%runElab derive "DimConst" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "DimConst" [Eq, Ord, Show]
|
||||||
|
|
||||||
||| `ends l r e` returns `l` if `e` is `Zero`, or `r` if it is `One`.
|
||| `ends l r e` returns `l` if `e` is `Zero`, or `r` if it is `One`.
|
||||||
public export
|
public export
|
||||||
|
|
|
@ -2,7 +2,7 @@ module Quox.Syntax.Qty.Three
|
||||||
|
|
||||||
import Quox.Pretty
|
import Quox.Pretty
|
||||||
import public Quox.Syntax.Qty
|
import public Quox.Syntax.Qty
|
||||||
import Generics.Derive
|
import Derive.Prelude
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
%language ElabReflection
|
%language ElabReflection
|
||||||
|
@ -12,7 +12,7 @@ public export
|
||||||
data Three = Zero | One | Any
|
data Three = Zero | One | Any
|
||||||
%name Three pi, rh
|
%name Three pi, rh
|
||||||
|
|
||||||
%runElab derive "Three" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Three" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
|
||||||
export
|
export
|
||||||
|
@ -23,6 +23,18 @@ PrettyHL Three where
|
||||||
One => ifUnicode "𝟭" "1"
|
One => ifUnicode "𝟭" "1"
|
||||||
Any => ifUnicode "𝛚" "*"
|
Any => ifUnicode "𝛚" "*"
|
||||||
|
|
||||||
|
public export
|
||||||
|
DecEq Three where
|
||||||
|
decEq Zero Zero = Yes Refl
|
||||||
|
decEq Zero One = No $ \case _ impossible
|
||||||
|
decEq Zero Any = No $ \case _ impossible
|
||||||
|
decEq One Zero = No $ \case _ impossible
|
||||||
|
decEq One One = Yes Refl
|
||||||
|
decEq One Any = No $ \case _ impossible
|
||||||
|
decEq Any Zero = No $ \case _ impossible
|
||||||
|
decEq Any One = No $ \case _ impossible
|
||||||
|
decEq Any Any = Yes Refl
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
plus : Three -> Three -> Three
|
plus : Three -> Three -> Three
|
||||||
|
|
|
@ -3,7 +3,7 @@ module Quox.Syntax.Universe
|
||||||
import Quox.Pretty
|
import Quox.Pretty
|
||||||
|
|
||||||
import Data.Fin
|
import Data.Fin
|
||||||
import Generics.Derive
|
import Derive.Prelude
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
%language ElabReflection
|
%language ElabReflection
|
||||||
|
@ -15,7 +15,7 @@ public export
|
||||||
data Universe = U Nat | UAny
|
data Universe = U Nat | UAny
|
||||||
%name Universe l
|
%name Universe l
|
||||||
|
|
||||||
%runElab derive "Universe" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "Universe" [Eq, Ord, Show]
|
||||||
|
|
||||||
export
|
export
|
||||||
PrettyHL Universe where
|
PrettyHL Universe where
|
||||||
|
|
|
@ -6,11 +6,9 @@ import public Quox.Reduce
|
||||||
|
|
||||||
import public Data.SortedMap
|
import public Data.SortedMap
|
||||||
import public Control.Monad.Either
|
import public Control.Monad.Either
|
||||||
import Generics.Derive
|
import Derive.Prelude
|
||||||
|
|
||||||
%hide TT.Name
|
%hide TT.Name
|
||||||
%hide SOP.from
|
|
||||||
%hide SOP.to
|
|
||||||
%default total
|
%default total
|
||||||
%language ElabReflection
|
%language ElabReflection
|
||||||
|
|
||||||
|
@ -103,7 +101,7 @@ InferResult eqs q d n = IfConsistent eqs $ InferResult' q d n
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data EqMode = Equal | Sub | Super
|
data EqMode = Equal | Sub | Super
|
||||||
%runElab derive "EqMode" [Generic, Meta, Eq, Ord, DecEq, Show]
|
%runElab derive "EqMode" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
|
|
|
@ -5,7 +5,7 @@ authors = "rhiannon morris"
|
||||||
sourceloc = "https://git.rhiannon.website/rhi/quox"
|
sourceloc = "https://git.rhiannon.website/rhi/quox"
|
||||||
license = "acsl"
|
license = "acsl"
|
||||||
|
|
||||||
depends = base, contrib, elab-util, sop, snocvect
|
depends = base, contrib, elab-util, snocvect
|
||||||
|
|
||||||
modules =
|
modules =
|
||||||
Quox.NatExtra,
|
Quox.NatExtra,
|
||||||
|
|
Loading…
Reference in a new issue