You signed in with another tab or window. Reload to refresh your session.You signed out in another tab or window. Reload to refresh your session.You switched accounts on another tab or window. Reload to refresh your session.Dismiss alert
I expect aesop? to only suggest things that actually close the proof. However, here is an example where it returns something that fails:
import Mathlib.Data.Real.Basic
import Mathlib.Algebra.Order.Positive.Field
abbrev PosReal : Type := { x : ℝ // 0 < x }
notation"ℝ+" => PosReal
abbrev solution_set : Set (ℝ+ → ℝ+) := { fun x ↦ x + 1 }
defmyTacGen : Aesop.TacGen := fun _ => do
return #[("exact ⟨val - f { val := val, property := property }, fun a ha => by simpa⟩",
0.9)]
macro "foo" : tactic => `(tactic| aesop? (add 90% myTacGen))
theoremusa2023_p2b (f : ℝ+ → ℝ+)
(P : ∀ x y, f (x * y + (f x)) = x * (f y) + ⟨2, two_pos⟩) :
f ∈ solution_set := bysuffices h : ∃ a b : ℝ, 0 < a ∧ ∀ x, (f x).val = a * x.val + b bysorrylet c := f 1
foo
/- Returns the following, which does not actually work: ``` Try this: simp_all only [Subtype.forall, exists_and_left] unhygienic with_reducible aesop_destruct_products apply Exists.intro apply And.intro on_goal 2 => exact ⟨val - f { val := val, property := property }, fun a ha => by simpa⟩ exact property ``` -/
import Aesop
@[aesop 90%]defmyTacGen : Aesop.TacGen := fun _ => do
return #[("exact ⟨val - f { val := val, property := property }, fun a ha => by simpa⟩",
0.9)]
theoremfoo2 (f : { x // 0 < x } → { x // 0 < x })
(val : Nat)
(property : 0 < val) :
∃ w x, ∀ (a : Nat) (b : 0 < a), ↑(f { val := a, property := b }) = w * a + x := by
constructor
aesop?
-- Try this:-- exact ⟨val - f { val := val, property := property }, fun a ha => by simpa⟩
I expect
aesop?
to only suggest things that actually close the proof. However, here is an example where it returns something that fails:This is with mathlib commit
8d0c6e6f821392660468fdce7c5463740d7988be
and aesop commit5fefb40a7c9038a7150e7edd92e43b1b94c49e79
The text was updated successfully, but these errors were encountered: