Skip to content

Commit

Permalink
[ fix ] respect visibility in forward data declarations
Browse files Browse the repository at this point in the history
  • Loading branch information
gallais committed Feb 8, 2023
1 parent b173267 commit f76c4c4
Show file tree
Hide file tree
Showing 5 changed files with 21 additions and 5 deletions.
9 changes: 5 additions & 4 deletions src/TTImp/ProcessData.idr
Original file line number Diff line number Diff line change
Expand Up @@ -467,17 +467,18 @@ processData {vars} eopts nest env fc vis mbtot (MkImpData dfc n_in mty_raw opts
-- When looking up, note the data types which were undefined at the
-- point of declaration.
ndefm <- lookupCtxtExact n (gamma defs)
(mw, fullty) <- the (Core (List Name, ClosedTerm)) $ case ndefm of
(mw, vis, fullty) <- the (Core (List Name, Visibility, ClosedTerm)) $ case ndefm of
Nothing => case mfullty of
Nothing => throw (GenericMsg fc "Missing telescope for data definition \{show n_in}")
Just fullty => pure ([], fullty)
Just fullty => pure ([], vis, fullty)
Just ndef =>
let vis = max (visibility ndef) vis in
case definition ndef of
TCon _ _ _ _ _ mw [] _ => case mfullty of
Nothing => pure (mw, type ndef)
Nothing => pure (mw, vis, type ndef)
Just fullty =>
do ok <- convert defs [] fullty (type ndef)
if ok then pure (mw, fullty)
if ok then pure (mw, vis, fullty)
else do logTermNF "declare.data" 1 "Previous" [] (type ndef)
logTermNF "declare.data" 1 "Now" [] fullty
throw (AlreadyDefined fc n)
Expand Down
2 changes: 1 addition & 1 deletion tests/Main.idr
Original file line number Diff line number Diff line change
Expand Up @@ -269,7 +269,7 @@ idrisTests = MkTestPool "Misc" [] Nothing
-- Implicit laziness, lazy evaluation
"lazy001", "lazy002",
-- Namespace blocks
"namespace001", "namespace002", "namespace003",
"namespace001", "namespace002", "namespace003", "namespace004",
-- Parameters blocks
"params001", "params002", "params003",
-- Larger programs arising from real usage. Typically things with
Expand Down
11 changes: 11 additions & 0 deletions tests/idris2/namespace004/Export.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,11 @@
namespace Core

public export
data D : Type


data D : Type where
MkD : D

d : D
d = MkD
1 change: 1 addition & 0 deletions tests/idris2/namespace004/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
1/1: Building Export (Export.idr)
3 changes: 3 additions & 0 deletions tests/idris2/namespace004/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
rm -rf build

$1 --no-color --console-width 0 --check Export.idr

0 comments on commit f76c4c4

Please sign in to comment.