Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

feat: add support for other types in norm_num #8100

Draft
wants to merge 1 commit into
base: master
Choose a base branch
from
Draft
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
10 changes: 5 additions & 5 deletions Mathlib/Tactic/NormNum/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -185,11 +185,11 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
let .app (.app (f : Q($α → $α → $α)) (a : Q($α))) (b : Q($α)) ← whnfR e | failure
let ra ← derive a; let rb ← derive b
match ra, rb with
| .isBool .., _ | _, .isBool .. => failure
| .isNat _ .., .isNat _ .. | .isNat _ .., .isNegNat _ .. | .isNat _ .., .isRat _ ..
| .isNegNat _ .., .isNat _ .. | .isNegNat _ .., .isNegNat _ .. | .isNegNat _ .., .isRat _ ..
| .isRat _ .., .isNat _ .. | .isRat _ .., .isNegNat _ .. | .isRat _ .., .isRat _ .. =>
guard <|← withNewMCtxDepth <| isDefEq f q(HAdd.hAdd (α := $α))
| _, _ => failure
let rec
/-- Main part of `evalAdd`. -/
core : Option (Result e) := do
Expand Down Expand Up @@ -218,7 +218,6 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
let r2 : Q(Nat.mul $da $db = Nat.mul $k $dc) := (q(Eq.refl $t2) : Expr)
return .isRat' dα qc nc dc q(isRat_add (f := $f) (.refl $f) $pa $pb $r1 $r2)
match ra, rb with
| .isBool .., _ | _, .isBool .. => failure
| .isRat dα .., _ | _, .isRat dα .. => ratArm dα
| .isNegNat rα .., _ | _, .isNegNat rα .. => intArm rα
| .isNat _ na pa, .isNat sα nb pb =>
Expand All @@ -228,6 +227,7 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
have c : Q(ℕ) := mkRawNatLit (na.natLit! + nb.natLit!)
haveI' : Nat.add $na $nb =Q $c := ⟨⟩
return .isNat sα c q(isNat_add (f := $f) (.refl $f) $pa $pb (.refl $c))
| _, _ => failure
core

-- see note [norm_num lemma function equality]
Expand Down Expand Up @@ -266,10 +266,10 @@ such that `norm_num` successfully recognises `a`. -/
haveI' : Int.neg $na =Q $nb := ⟨⟩
return .isRat' dα qb nb da q(isRat_neg (f := $f) (.refl $f) $pa (.refl $nb))
match ra with
| .isBool _ .. => failure
| .isNat _ .. => intArm rα
| .isNegNat rα .. => intArm rα
| .isRat dα .. => ratArm dα
| _ => failure
core

-- see note [norm_num lemma function equality]
Expand Down Expand Up @@ -321,10 +321,10 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
let r2 : Q(Nat.mul $da $db = Nat.mul $k $dc) := (q(Eq.refl $t2) : Expr)
return .isRat' dα qc nc dc q(isRat_sub (f := $f) (.refl $f) $pa $pb $r1 $r2)
match ra, rb with
| .isBool .., _ | _, .isBool .. => failure
| .isRat dα .., _ | _, .isRat dα .. => ratArm dα
| .isNegNat rα .., _ | _, .isNegNat rα ..
| .isNat _ .., .isNat _ .. => intArm rα
| _, _ => failure
core

-- see note [norm_num lemma function equality]
Expand Down Expand Up @@ -391,7 +391,6 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
let r2 : Q(Nat.mul $da $db = Nat.mul $k $dc) := (q(Eq.refl $t2) : Expr)
return .isRat' dα qc nc dc q(isRat_mul (f := $f) (.refl $f) $pa $pb $r1 $r2)
match ra, rb with
| .isBool .., _ | _, .isBool .. => failure
| .isRat dα .., _ | _, .isRat dα .. => ratArm dα
| .isNegNat rα .., _ | _, .isNegNat rα .. => intArm rα
| .isNat mα' na pa, .isNat mα nb pb =>
Expand All @@ -400,6 +399,7 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
have c : Q(ℕ) := mkRawNatLit (na.natLit! * nb.natLit!)
haveI' : Nat.mul $na $nb =Q $c := ⟨⟩
return .isNat mα c q(isNat_mul (f := $f) (.refl $f) $pa $pb (.refl $c))
| _, _ => failure
core

theorem isRat_div [DivisionRing α] : {a b : α} → {cn : ℤ} → {cd : ℕ} → IsRat (a * b⁻¹) cn cd →
Expand Down
11 changes: 11 additions & 0 deletions Mathlib/Tactic/NormNum/Eq.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro
-/
import Mathlib.Tactic.NormNum.Inv
import Mathlib.Tactic.NormNum.Structural
import Mathlib.Algebra.Order.Invertible

/-!
Expand Down Expand Up @@ -35,6 +36,10 @@ theorem isRat_eq_false [Ring α] [CharZero α] : {a b : α} → {na nb : ℤ}
| _, _, _, _, _, _, ⟨_, rfl⟩, ⟨_, rfl⟩, h => by
rw [Rat.invOf_denom_swap]; exact_mod_cast of_decide_eq_false h

theorem eq_of_eq {a b a' b' : α} (ha : a = a') (hb : b = b') (h : a' = b') : a = b := ha ▸ hb ▸ h

theorem ne_of_ne {a b a' b' : α} (ha : a = a') (hb : b = b') (h : a' ≠ b') : a ≠ b := ha ▸ hb ▸ h

/-- The `norm_num` extension which identifies expressions of the form `a = b`,
such that `norm_num` successfully recognises both `a` and `b`. -/
@[norm_num _ = _, Eq _ _] def evalEq : NormNumExt where eval {v β} e := do
Expand Down Expand Up @@ -92,5 +97,11 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
return .isFalse q(isNat_eq_false $pa $pb $r)
else
failure --TODO: nonzero characteristic ≠
| .other .., _ | _, .other .. =>
let ⟨a', pa⟩ := ra.toRawEq
let ⟨b', pb⟩ := rb.toRawEq
match ← evalStructuralEq a' b' with
| .inl p => return .isTrue q(eq_of_eq $pa $pb $p)
| .inr p => return .isFalse q(ne_of_ne $pa $pb $p)

end Mathlib.Meta.NormNum
4 changes: 2 additions & 2 deletions Mathlib/Tactic/NormNum/Ineq.lean
Original file line number Diff line number Diff line change
Expand Up @@ -133,7 +133,6 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
let r : Q(decide ($nb * $da < $na * $db) = true) := (q(Eq.refl true) : Expr)
return .isFalse q(isRat_le_false $pa $pb $r)
match ra, rb with
| .isBool .., _ | _, .isBool .. => failure
| .isRat _ .., _ | _, .isRat _ .. => ratArm
| .isNegNat _ .., _ | _, .isNegNat _ .. => intArm
| .isNat ra na pa, .isNat rb nb pb =>
Expand All @@ -150,6 +149,7 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
return .isFalse q(isNat_le_false $pa $pb $r)
else -- Nats can appear in an `OrderedRing` without `CharZero`.
intArm
| _, _ => failure

/-- The `norm_num` extension which identifies expressions of the form `a < b`,
such that `norm_num` successfully recognises both `a` and `b`. -/
Expand Down Expand Up @@ -190,7 +190,6 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
let r : Q(decide ($nb * $da ≤ $na * $db) = true) := (q(Eq.refl true) : Expr)
return .isFalse q(isRat_lt_false $pa $pb $r)
match ra, rb with
| .isBool .., _ | _, .isBool .. => failure
| .isRat _ .., _ | _, .isRat _ .. => ratArm
| .isNegNat _ .., _ | _, .isNegNat _ .. => intArm
| .isNat ra na pa, .isNat rb nb pb =>
Expand All @@ -208,5 +207,6 @@ such that `norm_num` successfully recognises both `a` and `b`. -/
else
let r : Q(Nat.ble $nb $na = true) := (q(Eq.refl true) : Expr)
return .isFalse q(isNat_lt_false $pa $pb $r)
| _, _ => failure

end Mathlib.Meta.NormNum
2 changes: 1 addition & 1 deletion Mathlib/Tactic/NormNum/Pow.lean
Original file line number Diff line number Diff line change
Expand Up @@ -174,7 +174,6 @@ def evalPow : NormNumExt where eval {u α} e := do
/-- Main part of `evalPow`. -/
core : Option (Result e) := do
match ra with
| .isBool .. => failure
| .isNat sα na pa =>
assumeInstancesCommute
have ⟨c, r⟩ := evalNatPow na nb
Expand All @@ -190,4 +189,5 @@ def evalPow : NormNumExt where eval {u α} e := do
have ⟨dc, r2⟩ := evalNatPow da nb
let qc := mkRat zc dc.natLit!
return .isRat' dα qc nc dc q(isRat_pow (f := $f) (.refl $f) $pa $pb $r1 $r2)
| _ => failure
core
19 changes: 15 additions & 4 deletions Mathlib/Tactic/NormNum/Result.lean
Original file line number Diff line number Diff line change
Expand Up @@ -243,6 +243,8 @@ inductive Result' where
| isNegNat (inst lit proof : Expr)
/-- Untyped version of `Result.isRat`. -/
| isRat (inst : Expr) (q : Rat) (n d proof : Expr)
/-- Untyped version of `Result.other`. -/
| other (e proof : Expr)
deriving Inhabited

section
Expand Down Expand Up @@ -279,6 +281,12 @@ and `q` is the value of `n / d`. -/
∀ (inst : Q(DivisionRing $α) := by assumption) (q : Rat) (n : Q(ℤ)) (d : Q(ℕ))
(proof : Q(IsRat $x $n $d)), Result x := Result'.isRat

/-- The result is `proof : isRat x n d`, where `n` is either `.ofNat lit` or `.negOfNat lit`
with `lit` a raw nat literal and `d` is a raw nat literal (not 0 or 1),
and `q` is the value of `n / d`. -/
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

This doc comment is a verbatim copy of the Result.isRat comment. Did you mean to have something different here?

@[match_pattern, inline] def Result.other {α : Q(Type u)} {x : Q($α)} :
∀ (e : Q($α)) (proof : Q($x = $e)), Result x := Result'.other

end

/-- The result is `z : ℤ` and `proof : isNat x z`. -/
Expand Down Expand Up @@ -310,21 +318,22 @@ instance : ToMessageData (Result x) where
| .isNat _ lit proof => m!"isNat {lit} ({proof})"
| .isNegNat _ lit proof => m!"isNegNat {lit} ({proof})"
| .isRat _ q _ _ proof => m!"isRat {q} ({proof})"
| .other e proof => m!"other {e} ({proof})"

/-- Returns the rational number that is the result of `norm_num` evaluation. -/
def Result.toRat : Result e → Option Rat
| .isBool .. => none
| .isNat _ lit _ => some lit.natLit!
| .isNegNat _ lit _ => some (-lit.natLit!)
| .isRat _ q .. => some q
| _ => none

/-- Returns the rational number that is the result of `norm_num` evaluation, along with a proof
that the denominator is nonzero in the `isRat` case. -/
def Result.toRatNZ : Result e → Option (Rat × Option Expr)
| .isBool .. => none
| .isNat _ lit _ => some (lit.natLit!, none)
| .isNegNat _ lit _ => some (-lit.natLit!, none)
| .isRat _ q _ _ p => some (q, q(IsRat.den_nz $p))
| _ => none

/--
Extract from a `Result` the integer value (as both a term and an expression),
Expand All @@ -345,7 +354,6 @@ and the proof that the original expression is equal to this rational number.
def Result.toRat' {α : Q(Type u)} {e : Q($α)}
(_i : Q(DivisionRing $α) := by with_reducible assumption) :
Result e → Option (ℚ × (n : Q(ℤ)) × (d : Q(ℕ)) × Q(IsRat $e $n $d))
| .isBool .. => none
| .isNat _ lit proof =>
have proof : Q(@IsNat _ instAddMonoidWithOne $e $lit) := proof
some ⟨lit.natLit!, q(.ofNat $lit), q(nat_lit 1), q(($proof).to_isRat)⟩
Expand All @@ -354,6 +362,7 @@ def Result.toRat' {α : Q(Type u)} {e : Q($α)}
some ⟨-lit.natLit!, q(.negOfNat $lit), q(nat_lit 1),
(q(@IsInt.to_isRat _ DivisionRing.toRing _ _ $proof) : Expr)⟩
| .isRat _ q n d proof => some ⟨q, n, d, proof⟩
| _ => none

/--
Given a `NormNum.Result e` (which uses `IsNat`, `IsInt`, `IsRat` to express equality to a rational
Expand All @@ -370,6 +379,7 @@ def Result.toRawEq {α : Q(Type u)} {e : Q($α)} : Result e → (e' : Q($α)) ×
| .isNat _ lit p => ⟨q(Nat.rawCast $lit), q(IsNat.to_raw_eq $p)⟩
| .isNegNat _ lit p => ⟨q(Int.rawCast (.negOfNat $lit)), q(IsInt.to_raw_eq $p)⟩
| .isRat _ _ n d p => ⟨q(Rat.rawCast $n $d), q(IsRat.to_raw_eq $p)⟩
| .other e p => ⟨e, p⟩

/--
`Result.toRawEq` but providing an integer. Given a `NormNum.Result e` for something known to be an
Expand All @@ -381,7 +391,7 @@ def Result.toRawIntEq {α : Q(Type u)} {e : Q($α)} : Result e →
Option (ℤ × (e' : Q($α)) × Q($e = $e'))
| .isNat _ lit p => some ⟨lit.natLit!, q(Nat.rawCast $lit), q(IsNat.to_raw_eq $p)⟩
| .isNegNat _ lit p => some ⟨-lit.natLit!, q(Int.rawCast (.negOfNat $lit)), q(IsInt.to_raw_eq $p)⟩
| .isRat _ .. | .isBool .. => none
| _ => none

/-- Constructs a `Result` out of a raw nat cast. Assumes `e` is a raw nat cast expression. -/
def Result.ofRawNat {α : Q(Type u)} (e : Q($α)) : Result e := Id.run do
Expand Down Expand Up @@ -429,6 +439,7 @@ def Result.toSimpResult {α : Q(Type u)} {e : Q($α)} : Result e → MetaM Simp.
let ⟨n', pn'⟩ ← mkOfNat α q(AddCommMonoidWithOne.toAddMonoidWithOne) lit
let ⟨d', pd'⟩ ← mkOfNat α q(AddCommMonoidWithOne.toAddMonoidWithOne) d
return { expr := q($n' / $d'), proof? := q(IsRat.nonneg_to_eq $p $pn' $pd') }
| .other expr proof? => pure { expr, proof? }

/-- Given `Mathlib.Meta.NormNum.Result.isBool p b`, this is the type of `p`.
Note that `BoolResult p b` is definitionally equal to `Expr`, and if you write `match b with ...`,
Expand Down
141 changes: 141 additions & 0 deletions Mathlib/Tactic/NormNum/Structural.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,141 @@
/-
Copyright (c) 2022 Mario Carneiro. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro
-/
import Mathlib.Tactic.NormNum.Core

/-!
# `norm_num` extension for inductive constructors
-/

set_option autoImplicit true

open Lean Meta Qq

namespace Mathlib.Meta.NormNum

theorem prod_ne_1 {a₁ b₁ : α} {a₂ b₂ : β} (h : a₁ ≠ b₁) : (a₁, a₂) ≠ (b₁, b₂) | rfl => h rfl
theorem prod_ne_2 {a₁ b₁ : α} {a₂ b₂ : β} (h : a₂ ≠ b₂) : (a₁, a₂) ≠ (b₁, b₂) | rfl => h rfl

Check failure on line 19 in Mathlib/Tactic/NormNum/Structural.lean

View workflow job for this annotation

GitHub Actions / Lint style

Mathlib/Tactic/NormNum/Structural.lean#L19: ERR_IND: If the theorem/def statement requires multiple lines, indent it correctly (4 spaces or 2 for `|`)

theorem cons_ne_1 {a₁ b₁ : α} {a₂ b₂ : List α} (h : a₁ ≠ b₁) :
(a₁ :: a₂) ≠ (b₁ :: b₂) | rfl => h rfl
theorem cons_ne_2 {a₁ b₁ : α} {a₂ b₂ : List α} (h : a₂ ≠ b₂) :
(a₁ :: a₂) ≠ (b₁ :: b₂) | rfl => h rfl

inductive MatchList {α : Q(Type u)} (a : Q(List $α)) where
| nil (_ : $a =Q [])
| cons (a₁ : Q($α)) (a₂ : Q(List $α)) (_ : $a =Q $a₁ :: $a₂)

def matchList {α : Q(Type u)} (a : Q(List $α)) : MetaM (MatchList a) := do
match ← whnfR a with
| .app (.const ``List.nil _) _ => return .nil ⟨⟩
| .app (.app (.app (.const ``List.cons _) _) a₁) a₂ => return .cons a₁ a₂ ⟨⟩
| _ => failure

/-- The `norm_num` extension which identifies expressions of the form `(a, b)`,
such that `norm_num` successfully recognises both `a` and `b`. -/
@[norm_num (_, _)] def evalProdMk : NormNumExt where eval {u α} e := do
trace[Meta.debug] "norm_num prod.mk {e}"
let .app (.app (.app (.app (.const ``Prod.mk [u₁, u₂])
(α₁ : Q(Type u₁))) (α₂ : Q(Type u₂))) (a₁ : Q($α₁))) (a₂ : Q($α₂)) ← whnf e | failure
have : u =QL max u₁ u₂ := ⟨⟩
have : $α =Q ($α₁ × $α₂) := ⟨⟩
have : $e =Q ($a₁, $a₂) := ⟨⟩
let ⟨b₁, p₁⟩ := (← derive a₁).toRawEq
let ⟨b₂, p₂⟩ := (← derive a₂).toRawEq
return .other q(($b₁, $b₂)) q(congr (congrArg _ $p₁) $p₂)

/-- The `norm_num` extension which identifies expressions of the form `[]`. -/
@[norm_num []] def evalListNil : NormNumExt where eval {u α} e := do
let e' : Q($α) ← whnf e
have : $e =Q $e' := ⟨⟩
return .other e' q(@rfl _ $e')

/-- The `norm_num` extension which identifies expressions of the form `(a, b)`,
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

(a, b) -> a::b

such that `norm_num` successfully recognises both `a` and `b`. -/
@[norm_num (_ :: _)] def evalListCons : NormNumExt where eval {u α} e := do
let .app (.app (.app (.const ``List.cons _)
(β : Q(Type u))) (a₁ : Q($β))) (a₂ : Q(List $β)) ← whnf e | failure
have : $α =Q List $β := ⟨⟩
have : $e =Q $a₁ :: $a₂ := ⟨⟩
let ⟨b₁, p₁⟩ := (← derive a₁).toRawEq
let ⟨b₂, p₂⟩ := (← derive (u := u) a₂).toRawEq
return .other q(($b₁ :: $b₂)) q(congr (congrArg _ $p₁) $p₂)

/-- Evaluates an equality up to inductive constructors, using `norm_num` on the leaves. -/
partial def evalStructuralEq {α : Q(Type u)} (a b : Q($α)) : MetaM <| Q($a = $b) ⊕ Q($a ≠ $b) := do
let α ← whnf α
let .const f ls := α.getAppFn | failure
match f with
| ``Nat | ``Int | ``_root_.Rat =>
match ← derive q($a = $b) with
| .isTrue p => return .inl p
| .isFalse p => return .inr p
| _ => failure
| ``Prod =>
let [u₁, u₂] := ls | failure
have : u =QL max u₁ u₂ := ⟨⟩
have α₁ : Q(Type u₁) := α.getRevArg! 1
have α₂ : Q(Type u₂) := α.getRevArg! 0
have : $α =Q ($α₁ × $α₂) := ⟨⟩
let .app (.app fa (a₁ : Q($α₁))) (a₂ : Q($α₂)) ← whnfR a | failure
let .app (.app fb (b₁ : Q($α₁))) (b₂ : Q($α₂)) ← whnfR b | failure
have : $a =Q ($a₁, $a₂) := ⟨⟩
have : $b =Q ($b₁, $b₂) := ⟨⟩
unless fa.isAppOfArity ``Prod.mk 2 && fb.isAppOfArity ``Prod.mk 2 do failure
match ← evalStructuralEq a₁ b₁ with
| .inr p₁ => return .inr q(prod_ne_1 $p₁)
| .inl p₁ => match ← evalStructuralEq a₂ b₂ with
| .inr p₂ => return .inr q(prod_ne_2 $p₂)
| .inl p₂ => return .inl q(congr (congrArg _ $p₁) $p₂)
| ``List =>
have β : Q(Type u) := α.getRevArg! 0
let rec evalList (a b : Q(List $β)) :
MetaM <| Q($a ≠ $b) ⊕ MetaM (Q($a = $b) ⊕ Q($a ≠ $b)) := do
match ← matchList a, ← matchList b with
| .nil _, .nil _ => pure <| .inr <| pure <| .inl q(rfl)
| .cons .., .nil _ => pure <| .inl q(List.cons_ne_nil _ _)
| .nil _, .cons .. => pure <| .inl q((List.cons_ne_nil _ _).symm)
| .cons a₁ a₂ _, .cons b₁ b₂ _ =>
match ← evalList a₂ b₂ with
| .inl no => return .inl q(cons_ne_2 $no)
| .inr p₂ => return .inr do
match ← evalStructuralEq a₁ b₁ with
| .inr p₁ => return .inr q(cons_ne_1 $p₁)
| .inl p₁ => match ← p₂ with
| .inr p₂ => return .inr q(cons_ne_2 $p₂)
| .inl p₂ => return .inl q(congr (congrArg _ $p₁) $p₂)
have : $α =Q List $β := ⟨⟩
match ← evalList a b with
| .inl no => return .inr no
| .inr k => k
| _ =>
matchConstInduct α.getAppFn (fun _ => failure) fun v us => do
let a ← whnfR a; let b ← whnfR b
let .const fa _ := a.getAppFn | failure
let .const fb _ := b.getAppFn | failure
let some (.ctorInfo info) := (← getEnv).find? fa | failure
unless fa == fb do
let some (.ctorInfo _) := (← getEnv).find? fb | failure
return .inr <| mkAppN (.const (.str v.name "noConfusion") (.zero :: us))
(α.getAppArgs ++ #[(q(False) : Expr), a, b])
let aArgs := a.getAppArgs; let bArgs := b.getAppArgs
let mut args : Array ((u' : Level) × (α' : Q(Type u')) × Q($α') × Q($α')) := #[]
for ai in aArgs[info.numParams:], bi in bArgs[info.numParams:] do
let ⟨_, α', ai⟩ ← inferTypeQ' ai
unless ← isDefEq α' (← inferType bi) do failure
args := args.push ⟨_, _, ai, bi⟩
let mut eq : Expr ← mkEqRefl <| mkAppN a.getAppFn a.getAppArgs[:info.numParams]
for ⟨_, _, ai, bi⟩ in args, i in [0:args.size] do
match ← evalStructuralEq ai bi with
| .inl eq' => eq ← mkCongr eq eq'
| .inr ne =>
return .inr <| ← withLocalDeclD `h q($a = $b) fun h => do
let e := mkAppN (.const (.str v.name "noConfusion") (.zero :: us))
(α.getAppArgs ++ (#[q(False), a, b, h] : Array Expr))
let .forallE _ dom _ _ ← whnf (← inferType e) | failure
forallTelescopeReducing dom fun xs _ => do
unless xs.size == args.size do failure
mkLambdaFVars #[h] <| mkApp e <| ← mkLambdaFVars xs (.app ne xs[i]!)
return .inl eq
2 changes: 2 additions & 0 deletions test/norm_num.lean
Original file line number Diff line number Diff line change
Expand Up @@ -679,3 +679,5 @@ example : (1 : R PUnit.{u+1} PUnit.{v+1}) <= 2 := by

-- Check that we avoid deep recursion in evaluating large powers.
example : 10^40000000 = 10^40000000 := by norm_num

example : (1 + 1, 2 * 2) = (2, 4) := by norm_num1
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

It would be good to add a similar test for lists.

Loading