mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-26 05:01:34 +03:00
Collect constructors on the left hand side of case alternatives (#2919)
* [ fix ] collect constructors on LHS of cases alts * [ tests ] Updated expected These functions do refer to these constructors at runtime, so this is the correct output * Update CHANGELOG.md * [test] update test output again
This commit is contained in:
parent
79116320dd
commit
442c5b529f
@ -71,6 +71,9 @@
|
|||||||
* Additionally some conflicting fixity declarations in the Idris 2 compiler
|
* Additionally some conflicting fixity declarations in the Idris 2 compiler
|
||||||
and libraries have been removed.
|
and libraries have been removed.
|
||||||
|
|
||||||
|
* Constructors mentioned on the left hand side of functions/case alternatives
|
||||||
|
are now included in the `Refers to (runtime)` section of functions debug info.
|
||||||
|
|
||||||
### Library changes
|
### Library changes
|
||||||
|
|
||||||
#### Prelude
|
#### Prelude
|
||||||
|
@ -496,8 +496,8 @@ mutual
|
|||||||
|
|
||||||
addRefsConAlts : NameMap Bool -> List (CConAlt vars) -> NameMap Bool
|
addRefsConAlts : NameMap Bool -> List (CConAlt vars) -> NameMap Bool
|
||||||
addRefsConAlts ds [] = ds
|
addRefsConAlts ds [] = ds
|
||||||
addRefsConAlts ds (MkConAlt _ _ _ _ sc :: rest)
|
addRefsConAlts ds (MkConAlt n _ _ _ sc :: rest)
|
||||||
= addRefsConAlts (addRefs ds sc) rest
|
= addRefsConAlts (addRefs (insert n False ds) sc) rest
|
||||||
|
|
||||||
addRefsConstAlts : NameMap Bool -> List (CConstAlt vars) -> NameMap Bool
|
addRefsConstAlts : NameMap Bool -> List (CConstAlt vars) -> NameMap Bool
|
||||||
addRefsConstAlts ds [] = ds
|
addRefsConstAlts ds [] = ds
|
||||||
|
@ -9,4 +9,6 @@ Prelude.Num.MkIntegral = Constructor tag Just 0 arity 3
|
|||||||
Prelude.EqOrd.MkOrd = Constructor tag Just 0 arity 8
|
Prelude.EqOrd.MkOrd = Constructor tag Just 0 arity 8
|
||||||
Prelude.EqOrd.MkEq = Constructor tag Just 0 arity 2
|
Prelude.EqOrd.MkEq = Constructor tag Just 0 arity 2
|
||||||
Prelude.Interfaces.MkMonoid = Constructor tag Just 0 arity 2
|
Prelude.Interfaces.MkMonoid = Constructor tag Just 0 arity 2
|
||||||
|
Prelude.Interfaces.MkMonad = Constructor tag Just 0 arity 3
|
||||||
Prelude.Interfaces.MkFoldable = Constructor tag Just 0 arity 6
|
Prelude.Interfaces.MkFoldable = Constructor tag Just 0 arity 6
|
||||||
|
Prelude.Interfaces.MkApplicative = Constructor tag Just 0 arity 3
|
||||||
|
@ -12,7 +12,7 @@ Compiled: \ {arg:1}, {arg:2} => case {arg:2} of
|
|||||||
; Main.(::) {tag = 1} [cons] {e:2} {e:3} => Prelude.Basics.(::) {tag = 1} [cons] {e:2} (Main.idL {e:3} (Main.view {e:3}))
|
; Main.(::) {tag = 1} [cons] {e:2} {e:3} => Prelude.Basics.(::) {tag = 1} [cons] {e:2} (Main.idL {e:3} (Main.view {e:3}))
|
||||||
}
|
}
|
||||||
Refers to: Main.view, Main.idL, Prelude.Basics.Nil, Prelude.Basics.(::)
|
Refers to: Main.view, Main.idL, Prelude.Basics.Nil, Prelude.Basics.(::)
|
||||||
Refers to (runtime): Main.view, Main.idL, Prelude.Basics.Nil, Prelude.Basics.(::)
|
Refers to (runtime): Main.view, Main.idL, Main.Nil, Main.(::), Prelude.Basics.Nil, Prelude.Basics.(::)
|
||||||
Flags: covering
|
Flags: covering
|
||||||
Size change: Main.idL: [Just (0, Same), Just (2, Smaller), Nothing]
|
Size change: Main.idL: [Just (0, Same), Just (2, Smaller), Nothing]
|
||||||
Main>
|
Main>
|
||||||
|
@ -9,6 +9,7 @@ Compiled: \ {arg:0} => case {arg:0} of
|
|||||||
; _ => 0
|
; _ => 0
|
||||||
}
|
}
|
||||||
Refers to: Prelude.Basics.True, Prelude.Basics.False
|
Refers to: Prelude.Basics.True, Prelude.Basics.False
|
||||||
|
Refers to (runtime): Prelude.Types.Nat
|
||||||
Flags: allguarded, covering
|
Flags: allguarded, covering
|
||||||
Main> Main.isInt32
|
Main> Main.isInt32
|
||||||
Arguments [{arg:0}]
|
Arguments [{arg:0}]
|
||||||
@ -20,6 +21,7 @@ Compiled: \ {arg:0} => case {arg:0} of
|
|||||||
; _ => 0
|
; _ => 0
|
||||||
}
|
}
|
||||||
Refers to: Prelude.Basics.True, Prelude.Basics.False
|
Refers to: Prelude.Basics.True, Prelude.Basics.False
|
||||||
|
Refers to (runtime): Int32
|
||||||
Flags: allguarded, covering
|
Flags: allguarded, covering
|
||||||
Main> Prelude.Types.plus
|
Main> Prelude.Types.plus
|
||||||
Arguments [{arg:0}, {arg:1}]
|
Arguments [{arg:0}, {arg:1}]
|
||||||
|
Loading…
Reference in New Issue
Block a user