Skip to content

Commit

Permalink
fix: ignore unused alternatives in Ord derive handler (#3725)
Browse files Browse the repository at this point in the history
Closes #3706

This derive handler's implementation is very similar to `BEq`'s, which
already ignores unused alternative so as to work correctly on indexed
inductive types. This PR simply implements the same solution as the one
present in
[`BEq.lean`](https://github.com/leanprover/lean4/blob/2c15cdda044e77bb8c3937c63501850790e60dc6/src/Lean/Elab/Deriving/BEq.lean#L94).

After some tests, it doesn't seem like any other derive handler present
in Core suffers from the same issue (though some handlers don't work on
indexed inductives for other reasons).
  • Loading branch information
arthur-adjedj authored Mar 21, 2024
1 parent 4d4e467 commit bf8b66c
Show file tree
Hide file tree
Showing 2 changed files with 14 additions and 0 deletions.
1 change: 1 addition & 0 deletions src/Lean/Elab/Deriving/Ord.lean
Original file line number Diff line number Diff line change
Expand Up @@ -84,6 +84,7 @@ def mkMutualBlock (ctx : Context) : TermElabM Syntax := do
for i in [:ctx.typeInfos.size] do
auxDefs := auxDefs.push (← mkAuxFunction ctx i)
`(mutual
set_option match.ignoreUnusedAlts true
$auxDefs:command*
end)

Expand Down
13 changes: 13 additions & 0 deletions tests/lean/run/3706.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
/--Ensure that `deriving Ord` works on indexed inductive types.-/
inductive Ty where
| int
| bool
inductive Expr : Ty → Type where
| a : Expr .int
| b : Expr .bool
deriving Ord

inductive Vec (α : Type u) : Nat → Type u
| nil : Vec α 0
| cons : α → {n : Nat} → Vec α n → Vec α (n+1)
deriving Ord

0 comments on commit bf8b66c

Please sign in to comment.