Skip to content

Commit 993bc5f

Browse files
committed
fix: simp only [· ∈ ·]
This PR fixes `simp only [· ∈ ·]` after #5020. Fixes #5905
1 parent 811d8fb commit 993bc5f

File tree

4 files changed

+68
-36
lines changed

4 files changed

+68
-36
lines changed

src/Lean/Elab/BuiltinNotation.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -322,7 +322,7 @@ def elabCDotFunctionAlias? (stx : Term) : TermElabM (Option Expr) := do
322322
let stx ← liftMacroM <| expandMacros stx
323323
match stx with
324324
| `(fun $binders* => $f $args*) =>
325-
if binders == args then
325+
if binders.raw.toList.isPerm args.raw.toList then
326326
try Term.resolveId? f catch _ => return none
327327
else
328328
return none
@@ -332,7 +332,7 @@ def elabCDotFunctionAlias? (stx : Term) : TermElabM (Option Expr) := do
332332
| `(fun $binders* => rightact% $f $a $b)
333333
| `(fun $binders* => binrel% $f $a $b)
334334
| `(fun $binders* => binrel_no_prop% $f $a $b) =>
335-
if binders == #[a, b] then
335+
if binders == #[a, b] || binders == #[b, a] then
336336
try Term.resolveId? f catch _ => return none
337337
else
338338
return none

tests/lean/cdotAtSimpArg.lean

Lines changed: 0 additions & 25 deletions
This file was deleted.

tests/lean/cdotAtSimpArg.lean.expected.out

Lines changed: 0 additions & 9 deletions
This file was deleted.

tests/lean/run/cdotAtSimpArg.lean

Lines changed: 66 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,66 @@
1+
/-! Test `·` being able to refer to constants in `simp` -/
2+
3+
example : ¬ true = false := by
4+
simp [(¬ ·)]
5+
6+
/-! Test `binop%` -/
7+
8+
/--
9+
info: y x : Nat
10+
h : y = 0
11+
⊢ Add.add x y = x
12+
-/
13+
#guard_msgs in
14+
example (h : y = 0) : x + y = x := by
15+
simp [(.+.)] -- Expands `HAdd.hAdd
16+
trace_state
17+
simp [Add.add]
18+
simp [h, Nat.add]
19+
done
20+
21+
/--
22+
info: y x : Nat
23+
h : y = 0
24+
⊢ Add.add x y = x
25+
-/
26+
#guard_msgs in
27+
example (h : y = 0) : x + y = x := by
28+
simp [.+.]
29+
trace_state
30+
simp [Add.add]
31+
simp [h, Nat.add]
32+
done
33+
34+
/-! Test `binop%` variant `rightact%` as well -/
35+
36+
/--
37+
error: unsolved goals
38+
x y : Nat
39+
⊢ Pow.pow x y = Pow.pow y x
40+
-/
41+
#guard_msgs in
42+
example (x y : Nat) : x ^ y = y ^ x := by
43+
simp only [.^.]
44+
45+
46+
def f (n m : Nat) : Nat := n + m
47+
macro "ff" t1:term:arg t2:term:arg : term => `(f $t2 $t1)
48+
49+
/--
50+
error: unsolved goals
51+
x y : Nat
52+
⊢ [x + y, y + x].sum > 0
53+
-/
54+
#guard_msgs in
55+
example (x y : Nat) : [f x y, ff x y].sum > 0 := by
56+
simp only [ff · ·] -- NB: ff is syntax, this unfolds also f
57+
58+
59+
/--
60+
error: unsolved goals
61+
x y : Nat
62+
⊢ List.Mem x [x]
63+
-/
64+
#guard_msgs in
65+
example (x y : Nat) : x ∈ [x] := by
66+
simp only [· ∈ ·] -- syntax has arguments swapped, see #5905

0 commit comments

Comments
 (0)