Use full names for constructors in case split

Fixes #184
This commit is contained in:
Edwin Brady 2020-06-27 15:47:38 +01:00
parent 574524d8de
commit 6c007fc046
6 changed files with 32 additions and 2 deletions

View File

@ -84,6 +84,8 @@ getDefining tm
Ref _ _ fn => Just fn
_ => Nothing
-- For the name on the lhs, return the function name being defined, the
-- type name, and the possible constructors.
findCons : {auto c : Ref Ctxt Defs} ->
Name -> Term [] -> Core (SplitResult (Name, Name, List Name))
findCons n lhs
@ -102,7 +104,8 @@ findCons n lhs
(CantSplitThis n
("Not a type constructor " ++
show res)))
pure (OK (fn, tyn, cons))
pure (OK (fn, !(toFullNames tyn),
!(traverse toFullNames cons)))
findAllVars : Term vars -> List Name
findAllVars (Bind _ x (PVar c p ty) sc)

View File

@ -97,7 +97,7 @@ idrisTests
"reg001", "reg002", "reg003", "reg004", "reg005", "reg006", "reg007",
"reg008", "reg009", "reg010", "reg011", "reg012", "reg013", "reg014",
"reg015", "reg016", "reg017", "reg018", "reg019", "reg020", "reg021",
"reg022", "reg023",
"reg022", "reg023", "reg024",
-- Totality checking
"total001", "total002", "total003", "total004", "total005",
"total006", "total007", "total008",

View File

@ -0,0 +1,10 @@
1/1: Building split (split.idr)
Main> notLtz BLT_ZO impossible
notLtz BLT_ZE impossible
notLtz (BLT_OO x) impossible
notLtz (BLT_OE x) impossible
notLtz BLT_OE_Eq impossible
notLtz (BLT_EO x) impossible
notLtz (BLT_EE x) impossible
Main>
Bye for now!

View File

@ -0,0 +1 @@
:cs 13 8 x

3
tests/idris2/reg024/run Executable file
View File

@ -0,0 +1,3 @@
$1 --no-banner split.idr < input
rm -rf build

View File

@ -0,0 +1,13 @@
data BNat = BZ | O BNat | E BNat
data BLT : BNat -> BNat -> Type where
BLT_ZO : BLT BZ (O bn)
BLT_ZE : BLT BZ (E bn)
BLT_OO : BLT bn bm -> BLT (O bn) (O bm)
BLT_OE : BLT bn bm -> BLT (O bn) (E bm)
BLT_OE_Eq : BLT (O bn) (E bn)
BLT_EO : BLT bn bm -> BLT (E bn) (O bm)
BLT_EE : BLT bn bm -> BLT (E bn) (E bm)
notLtz : BLT bn BZ -> Void
notLtz x = ?notLtz_rhs