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

aesop? with a custom TacGen suggests a proof that does not work #125

Closed
dwrensha opened this issue Apr 29, 2024 · 1 comment
Closed

aesop? with a custom TacGen suggests a proof that does not work #125

dwrensha opened this issue Apr 29, 2024 · 1 comment

Comments

@dwrensha
Copy link
Member

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 }

def myTacGen : 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))

theorem usa2023_p2b (f : ℝ+ → ℝ+)
    (P : ∀ x y, f (x * y + (f x)) = x * (f y) + ⟨2, two_pos⟩) :
    f ∈ solution_set := by
  suffices h : ∃ a b : ℝ, 0 < a ∧ ∀ x, (f x).val = a * x.val + b by sorry
  let 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
   ```
  -/

This is with mathlib commit 8d0c6e6f821392660468fdce7c5463740d7988be and aesop commit 5fefb40a7c9038a7150e7edd92e43b1b94c49e79

@JLimperg
Copy link
Collaborator

Minimised:

import Aesop

@[aesop 90%]
def myTacGen : Aesop.TacGen := fun _ => do
  return #[("exact ⟨val - f { val := val, property := property }, fun a ha => by simpa⟩",
            0.9)]

theorem foo2 (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⟩

JLimperg added a commit that referenced this issue May 27, 2024
... in the presence of delayed-assigned mvars.

Fixes #125
Sign up for free to join this conversation on GitHub. Already have an account? Sign in to comment
Labels
None yet
Projects
None yet
Development

No branches or pull requests

2 participants