[ elab, minor ] Implement Functor for PiInfo

This commit is contained in:
Denis Buzdalov 2024-06-27 16:27:56 +03:00 committed by G. Allais
parent f561c78812
commit 715a304137
2 changed files with 9 additions and 0 deletions

View File

@ -191,6 +191,8 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
* `Decidable.Decidable.decison` is now `public export`.
* `Functor` is implemented for `PiInfo` from `Language.Reflection.TT`.
#### Contrib
* `Data.List.Lazy` was moved from `contrib` to `base`.

View File

@ -229,6 +229,13 @@ public export
data PiInfo t = ImplicitArg | ExplicitArg | AutoImplicit | DefImplicit t
%name PiInfo pinfo
public export
Functor PiInfo where
map f ImplicitArg = ImplicitArg
map f ExplicitArg = ExplicitArg
map f AutoImplicit = AutoImplicit
map f $ DefImplicit x = DefImplicit $ f x
export
showPiInfo : Show a => {default True wrapExplicit : Bool} -> PiInfo a -> String -> String
showPiInfo ImplicitArg s = "{\{s}}"