Skip to content

Commit

Permalink
fix: match discriminant reduction should not unfold irreducible defs
Browse files Browse the repository at this point in the history
  • Loading branch information
Kha committed Mar 23, 2023
1 parent 8650804 commit a9eedb7
Show file tree
Hide file tree
Showing 3 changed files with 31 additions and 3 deletions.
6 changes: 3 additions & 3 deletions src/Lean/Meta/WHNF.lean
Original file line number Diff line number Diff line change
Expand Up @@ -368,7 +368,7 @@ inductive ReduceMatcherResult where
should reduce in any transparency mode.
Thus, we define a custom `canUnfoldAtMatcher` predicate for `whnfMatcher`.
This solution is not very modular because modications at the `match` compiler require changes here.
This solution is not very modular because modifications at the `match` compiler require changes here.
We claim this is defensible because it is reducing the auxiliary declaration defined by the `match` compiler.
Alternative solution: tactics that use `TransparencyMode.reducible` should rely on the equations we generated for match-expressions.
Expand All @@ -378,7 +378,7 @@ inductive ReduceMatcherResult where
def canUnfoldAtMatcher (cfg : Config) (info : ConstantInfo) : CoreM Bool := do
match cfg.transparency with
| TransparencyMode.all => return true
| TransparencyMode.default => return true
| TransparencyMode.default => return !(← isIrreducible info.name)
| _ =>
if (← isReducible info.name) || isGlobalInstance (← getEnv) info.name then
return true
Expand All @@ -402,7 +402,7 @@ def canUnfoldAtMatcher (cfg : Config) (info : ConstantInfo) : CoreM Bool := do

private def whnfMatcher (e : Expr) : MetaM Expr := do
/- When reducing `match` expressions, if the reducibility setting is at `TransparencyMode.reducible`,
we increase it to `TransparencyMode.instance`. We use the `TransparencyMode.reducible` in many places (e.g., `simp`),
we increase it to `TransparencyMode.instances`. We use the `TransparencyMode.reducible` in many places (e.g., `simp`),
and this setting prevents us from reducing `match` expressions where the discriminants are terms such as `OfNat.ofNat α n inst`.
For example, `simp [Int.div]` will not unfold the application `Int.div 2 1` occuring in the target.
Expand Down
22 changes: 22 additions & 0 deletions tests/lean/2161.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,22 @@
structure Foo where num : Nat deriving DecidableEq

namespace Foo

instance : OfNat Foo n := ⟨⟨n⟩⟩

/-! # Example 1 -/

@[irreducible] def mul (a b : Foo) : Foo :=
let d := Nat.gcd a.num 1
⟨(a.num.div d) * (b.num.div d)⟩

-- should fail fast; exact heartbeat count at time of writing is 31
set_option maxHeartbeats 310
example : ((Foo.mul 4 1).mul 1).mul 1 = 4 := by decide

/-! # Example 2 -/

@[irreducible] def add (a b : Foo) : Foo := ⟨a.num * b.num⟩

-- should not succeed (and fail fast); exact heartbeat count at time of writing is 21
example : ((Foo.add 4 1).add 1).add 1 = 4 := by decide
6 changes: 6 additions & 0 deletions tests/lean/2161.lean.expected.out
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
2161.lean:14:48-14:54: error: failed to reduce to 'true'
Decidable.rec (fun h => (fun x => false) h) (fun h => (fun x => true) h)
(instDecidableEqFoo (mul (mul (mul 4 1) 1) 1) 4)
2161.lean:22:48-22:54: error: failed to reduce to 'true'
Decidable.rec (fun h => (fun x => false) h) (fun h => (fun x => true) h)
(instDecidableEqFoo (add (add (add 4 1) 1) 1) 4)

0 comments on commit a9eedb7

Please sign in to comment.