mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 01:01:59 +03:00
[ elab, minor ] Implement Functor
for PiInfo
This commit is contained in:
parent
f561c78812
commit
715a304137
@ -191,6 +191,8 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
|
|||||||
|
|
||||||
* `Decidable.Decidable.decison` is now `public export`.
|
* `Decidable.Decidable.decison` is now `public export`.
|
||||||
|
|
||||||
|
* `Functor` is implemented for `PiInfo` from `Language.Reflection.TT`.
|
||||||
|
|
||||||
#### Contrib
|
#### Contrib
|
||||||
|
|
||||||
* `Data.List.Lazy` was moved from `contrib` to `base`.
|
* `Data.List.Lazy` was moved from `contrib` to `base`.
|
||||||
|
@ -229,6 +229,13 @@ public export
|
|||||||
data PiInfo t = ImplicitArg | ExplicitArg | AutoImplicit | DefImplicit t
|
data PiInfo t = ImplicitArg | ExplicitArg | AutoImplicit | DefImplicit t
|
||||||
%name PiInfo pinfo
|
%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
|
export
|
||||||
showPiInfo : Show a => {default True wrapExplicit : Bool} -> PiInfo a -> String -> String
|
showPiInfo : Show a => {default True wrapExplicit : Bool} -> PiInfo a -> String -> String
|
||||||
showPiInfo ImplicitArg s = "{\{s}}"
|
showPiInfo ImplicitArg s = "{\{s}}"
|
||||||
|
Loading…
Reference in New Issue
Block a user