[ re #2423 ] Make Eq implementations in Language.Reflection.TTImp total (#2430)

* [ re #2423 ] Turn newly added `Eq` implementations to be total
* [ fix ] Add lacking `Eq TTImp` clauses and lacking `Eq` implementations
This commit is contained in:
Denis Buzdalov 2022-04-22 20:01:42 +03:00 committed by GitHub
parent 68bcacf3ec
commit 77cf47f962
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
2 changed files with 114 additions and 30 deletions

View File

@ -182,3 +182,24 @@ data Visibility = Private | Export | Public
public export
data BuiltinType = BuiltinNatural | NaturalToInteger | IntegerToNatural
public export
Eq TotalReq where
Total == Total = True
CoveringOnly == CoveringOnly = True
PartialOK == PartialOK = True
_ == _ = False
public export
Eq Visibility where
Private == Private = True
Export == Export = True
Public == Public = True
_ == _ = False
public export
Eq BuiltinType where
BuiltinNatural == BuiltinNatural = True
NaturalToInteger == NaturalToInteger = True
IntegerToNatural == IntegerToNatural = True
_ == _ = False

View File

@ -214,7 +214,6 @@ getFC (IWithUnambigNames fc _ _) = fc
public export
partial
Eq LazyReason where
LInf == LInf = True
LLazy == LLazy = True
@ -222,12 +221,10 @@ Eq LazyReason where
_ == _ = False
public export
partial
Eq Namespace where
MkNS ns == MkNS ns' = ns == ns'
public export
partial
Eq Count where
M0 == M0 = True
M1 == M1 = True
@ -235,7 +232,6 @@ Eq Count where
_ == _ = False
public export
partial
Eq BindMode where
PI c == PI c' = c == c'
PATTERN == PATTERN = True
@ -243,14 +239,12 @@ Eq BindMode where
_ == _ = False
public export
partial
Eq UseSide where
UseLeft == UseLeft = True
UseRight == UseRight = True
_ == _ = False
public export
partial
Eq DotReason where
NonLinearVar == NonLinearVar = True
VarApplied == VarApplied = True
@ -262,12 +256,10 @@ Eq DotReason where
_ == _ = False
public export
partial
Eq WithFlag where
Syntactic == Syntactic = True
public export
partial
Eq UserName where
Basic n == Basic n' = n == n'
Field n == Field n' = n == n'
@ -275,7 +267,6 @@ Eq UserName where
_ == _ = False
public export
partial
Eq Name where
NS ns n == NS ns' n' = ns == ns' && n == n'
UN n == UN n' = n == n'
@ -287,7 +278,6 @@ Eq Name where
_ == _ = False
public export
partial
Eq PrimType where
IntType == IntType = True
IntegerType == IntegerType = True
@ -306,7 +296,6 @@ Eq PrimType where
_ == _ = False
public export
partial
Eq Constant where
I c == I c' = c == c'
BI c == BI c' = c == c'
@ -325,28 +314,25 @@ Eq Constant where
WorldVal == WorldVal = True
_ == _ = False
public export
Eq a => Eq (PiInfo a) where
ImplicitArg == ImplicitArg = True
ExplicitArg == ExplicitArg = True
AutoImplicit == AutoImplicit = True
DefImplicit t == DefImplicit t' = t == t'
_ == _ = False
mutual
public export
partial
Eq (PiInfo TTImp) where
ImplicitArg == ImplicitArg = True
ExplicitArg == ExplicitArg = True
AutoImplicit == AutoImplicit = True
DefImplicit t == DefImplicit t' = t == t'
_ == _ = False
public export
partial
Eq Clause where
PatClause _ lhs rhs == PatClause _ lhs' rhs' =
lhs == lhs' && rhs == rhs'
WithClause _ l r w p f cs == WithClause _ l' r' w' p' f' cs' =
l == l' && r == r' && w == w' && p == p' && f == f' && cs == cs'
l == l' && r == r' && w == w' && p == p' && f == f' && (assert_total $ cs == cs')
ImpossibleClause _ l == ImpossibleClause _ l' = l == l'
_ == _ = False
public export
partial
Eq IFieldUpdate where
ISetField p t == ISetField p' t' =
p == p' && t == t'
@ -355,7 +341,6 @@ mutual
_ == _ = False
public export
partial
Eq AltType where
FirstSuccess == FirstSuccess = True
Unique == Unique = True
@ -363,19 +348,96 @@ mutual
_ == _ = False
public export
partial
Eq NoMangleDirective where
CommonName s == CommonName s' = s == s'
BackendNames ns == BackendNames ns' = ns == ns'
_ == _ = False
public export
Eq FnOpt where
Inline == Inline = True
NoInline == NoInline = True
Deprecate == Deprecate = True
TCInline == TCInline = True
Hint b == Hint b' = b == b'
GlobalHint b == GlobalHint b' = b == b'
ExternFn == ExternFn = True
ForeignFn es == ForeignFn es' = es == es'
Invertible == Invertible = True
Totality tr == Totality tr' = tr == tr'
Macro == Macro = True
SpecArgs ns == SpecArgs ns' = ns == ns'
NoMangle nm == NoMangle nm' = nm == nm'
_ == _ = False
public export
Eq DataOpt where
SearchBy ns == SearchBy ns' = ns == ns'
NoHints == NoHints = True
UniqueSearch == UniqueSearch = True
External == External = True
NoNewtype == NoNewtype = True
_ == _ = False
public export
Eq Data where
MkData _ n tc os dc == MkData _ n' tc' os' dc' =
n == n' && tc == tc' && os == os' && dc == dc'
MkLater _ n tc == MkLater _ n' tc' =
n == n' && tc == tc'
_ == _ = False
public export
Eq ITy where
MkTy _ _ n ty == MkTy _ _ n' ty' = n == n' && ty == ty'
public export
Eq IField where
MkIField _ c pi n e == MkIField _ c' pi' n' e' =
c == c' && pi == pi' && n == n' && e == e'
public export
Eq Record where
MkRecord _ n ps cn fs == MkRecord _ n' ps' cn' fs' =
n == n' && ps == ps' && cn == cn' && fs == fs'
public export
Eq Decl where
IClaim _ c v fos t == IClaim _ c' v' fos' t' =
c == c' && v == v' && fos == fos' && t == t'
IData _ v t d == IData _ v' t' d' =
v == v' && t == t' && d == d'
IDef _ n cs == IDef _ n' cs' =
n == n' && cs == cs'
IParameters _ ps ds == IParameters _ ps' ds' =
ps == ps' && (assert_total $ ds == ds')
IRecord _ ns v tr r == IRecord _ ns' v' tr' r' =
ns == ns' && v == v' && tr == tr' && r == r'
INamespace _ ns ds == INamespace _ ns' ds' =
ns == ns' && (assert_total $ ds == ds')
ITransform _ n f t == ITransform _ n' f' t' =
n == n' && f == f' && t == t'
IRunElabDecl _ e == IRunElabDecl _ e' = e == e'
ILog p == ILog p' = p == p'
IBuiltin _ t n == IBuiltin _ t' n' =
t == t' && n == n'
_ == _ = False
public export
Eq TTImp where
IVar _ v == IVar _ v' = v == v'
IPi _ c i n a r == IPi _ c' i' n' a' r' =
c == c' && i == i' && n == n' && a == a' && r == r'
c == c' && (assert_total $ i == i') && n == n' && a == a' && r == r'
ILam _ c i n a r == ILam _ c' i' n' a' r' =
c == c' && i == i' && n == n' && a == a' && r == r'
c == c' && (assert_total $ i == i') && n == n' && a == a' && r == r'
ILet _ _ c n ty val s == ILet _ _ c' n' ty' val' s' =
c == c' && n == n' && ty == ty' && val == val' && s == s'
ICase _ t ty cs == ICase _ t' ty' cs'
= t == t' && ty == ty' && cs == cs'
= t == t' && ty == ty' && (assert_total $ cs == cs')
ILocal _ ds e == ILocal _ ds' e' =
(assert_total $ ds == ds') && e == e'
IUpdate _ fs t == IUpdate _ fs' t' =
fs == fs' && t == t'
(assert_total $ fs == fs') && t == t'
IApp _ f x == IApp _ f' x' = f == f' && x == x'
INamedApp _ f n x == INamedApp _ f' n' x' =
@ -385,7 +447,7 @@ mutual
ISearch _ n == ISearch _ n' = n == n'
IAlternative _ t as == IAlternative _ t' as' =
t == t' && as == as'
(assert_total $ t == t') && (assert_total $ as == as')
IRewrite _ p q == IRewrite _ p' q' =
p == p' && q == q'
@ -403,6 +465,7 @@ mutual
IQuote _ tm == IQuote _ tm' = tm == tm'
IQuoteName _ n == IQuoteName _ n' = n == n'
IQuoteDecl _ ds == IQuoteDecl _ ds' = assert_total $ ds == ds'
IUnquote _ tm == IUnquote _ tm' = tm == tm'
IPrimVal _ c == IPrimVal _ c' = c == c'