Idris2/tests/idris2/perf004/bigdpair.idr
Edwin Brady 3f914889b8 Add visibility rules on types
Can't export a type which refers to a private name. This has caught a
couple of visibility errors in the libraries, code and tests, so they've
been updated too.
2020-05-30 17:03:15 +01:00

32 lines
576 B
Idris

import Data.Vect
MkBig : Nat -> Type
MkBig Z = Int
MkBig (S k) = ((n ** Vect n Int), MkBig k)
bigEx : (n : Nat) -> MkBig n
bigEx Z = 94
bigEx (S k) = ((2 ** [0,0]), bigEx k)
data VWrap : Type -> Type where
MkVWrap : (0 n : Nat) -> Vect n a -> VWrap a
export
MkBig' : Nat -> Type
MkBig' Z = Int
MkBig' (S k) = (VWrap Int, MkBig' k)
namespace Foo
public
export
bigEx' : (n : Nat) -> MkBig' n
bigEx' Z = 94
bigEx' (S k) = (MkVWrap 1 [0], bigEx' k)
eqBigs : bigEx 1000000 = bigEx 1000000
eqBigs = Refl
eqBigs' : bigEx' 800000 = bigEx' 800000
eqBigs' = Refl