[ fix, re #3063 ] Fix forgotten reflection values

This commit is contained in:
Denis Buzdalov 2023-12-18 19:13:40 +03:00 committed by CodingCellist
parent 58e5d15662
commit 36132f6539
5 changed files with 22 additions and 2 deletions

View File

@ -330,9 +330,9 @@ export
Reflect a => Reflect (WithDefault a def) where
reflect fc defs lhs env def
= onWithDefault
(appCon fc defs (reflectionttimp "Default") [Erased fc Placeholder, Erased fc Placeholder])
(appCon fc defs (reflectionttimp "DefaultedValue") [Erased fc Placeholder, Erased fc Placeholder])
(\x => do x' <- reflect fc defs lhs env x
appCon fc defs (reflectionttimp "Value") [Erased fc Placeholder, Erased fc Placeholder, x'])
appCon fc defs (reflectionttimp "SpecifiedValue") [Erased fc Placeholder, Erased fc Placeholder, x'])
def
export

View File

@ -0,0 +1,15 @@
import Language.Reflection
%language ElabReflection
data_decl : List Decl
data_decl = `[data T = A | B]
record_decl : List Decl
record_decl = `[record R where]
data_decl' : List Decl
data_decl' = `[public export data T = A | B]
record_decl' : List Decl
record_decl' = `[public export record R where]

View File

@ -0,0 +1 @@
1/1: Building Issue3168 (Issue3168.idr)

View File

@ -0,0 +1,3 @@
. ../../../testutils.sh
check Issue3168.idr

View File

@ -0,0 +1 @@
package a-test