Skip to content
This repository has been archived by the owner on Jul 24, 2024. It is now read-only.

[Merged by Bors] - feat(tactic/linear_combination): allow linear_combination with { exponent := n } #15428

Closed
wants to merge 7 commits into from

Conversation

robertylewis
Copy link
Member

@robertylewis robertylewis commented Jul 16, 2022

#15425 will solve problems where the goal is not a linear combination of hypotheses, but a power of the goal is. This PR provides a more natural certificate for these proofs. Writing linear_combination ... with { exponent := 2 } will square the goal before subtracting the linear combination. In principle this could be useful even outside of polyrith.


Open in Gitpod

@robertylewis robertylewis added awaiting-review The author would like community review of the PR t-meta Tactics, attributes or user commands modifies-tactic-syntax This PR adds a new interactive tactic or modifies the syntax of an existing tactic. labels Jul 16, 2022
@@ -415,10 +435,16 @@ by linear_combination 3 * h a b + hqc
meta def _root_.tactic.interactive.linear_combination
(input : parse (as_linear_combo ff [] <$> texpr)?)
(_ : parse (tk "with")?)
(exponent : parse (prod.mk <$> ident <*> small_nat)?)
Copy link
Member

Choose a reason for hiding this comment

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

This results in the slightly unhelpful help text,

linear_combination expr? with? (id n)? linear_combo.linear_combination_config?

I think either:

  • exponent should appear in this help text
  • we should make exponent a field of linear_combination_config

Copy link
Member

@eric-wieser eric-wieser Aug 2, 2022

Choose a reason for hiding this comment

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

I'd lean towards the second; linear_combination _ with { exponent := 2 } is much easier to extend later, doesn't involve any surprising new syntax, and has default values built in.

Copy link
Member Author

Choose a reason for hiding this comment

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

I'd lean towards the second; linear_combination _ with { exponent := 2 } is much easier to extend later, doesn't involve any surprising new syntax, and has default values built in.

I'm not opposed, but there's another linear_combination feature under discussion that wouldn't fit this pattern. We're scheming up syntax for a "contraposed" version, something like linear_combination 3*⊢ with target zero_ne_one that would work on disequality goals and a single disequality hypothesis/proof term. I don't think we can take an arbitrary proof of an arbitrary proposition in the config structure with convenient syntax and a default value.

So that would leave us at linear_combination _ with target zero_ne_one {exponent := 2} which is a little funny.

Copy link
Member

Choose a reason for hiding this comment

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

What would your proposed spelling me for both options simultaneously?

Copy link
Member

Choose a reason for hiding this comment

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

At a glance, "linear_combination _ with target x" feels a bit like simpa _ using x

Copy link
Member Author

Choose a reason for hiding this comment

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

@eric-wieser I've taken your suggestion and reimplemented this as a field of the config object!

@robertylewis
Copy link
Member Author

Sorry for the delay in updating this. I think it's ready for another review.

Copy link
Member

@eric-wieser eric-wieser left a comment

Choose a reason for hiding this comment

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

LGTM, documentation just needs some tweaks

@YaelDillies YaelDillies requested a review from a team as a code owner March 24, 2023 11:36
@YaelDillies YaelDillies changed the title feat(tactic/linear_combination): allow linear_combination with exponent n feat(tactic/linear_combination): allow linear_combination with { exponent := n } Mar 24, 2023
@YaelDillies
Copy link
Collaborator

I took the freedom to apply Eric's last suggestions. I assume this will be quite easy to port to mathlib4.

maintainer merge

@github-actions
Copy link

🚀 Pull request has been placed on the maintainer queue by YaelDillies.

@Vierkantor
Copy link
Collaborator

Implementing the exponent as a config option with a default value is a nice idea, I should remember that!

bors r+

@github-actions github-actions bot added ready-to-merge All that is left is for bors to build and merge this PR. (Remember you need to say `bors r+`.) and removed awaiting-review The author would like community review of the PR labels Mar 24, 2023
bors bot pushed a commit that referenced this pull request Mar 24, 2023
…onent := n }` (#15428)

#15425 will solve problems where the goal is not a linear combination of hypotheses, but a *power* of the goal is. This PR provides a more natural certificate for these proofs. Writing `linear_combination ... with { exponent  := 2 }` will square the goal before subtracting the linear combination. In principle this could be useful even outside of `polyrith`. 



Co-authored-by: Rob Lewis <rob.y.lewis@gmail.com>
Co-authored-by: Yaël Dillies <yael.dillies@gmail.com>
@bors
Copy link

bors bot commented Mar 24, 2023

Pull request successfully merged into master.

Build succeeded:

@bors bors bot changed the title feat(tactic/linear_combination): allow linear_combination with { exponent := n } [Merged by Bors] - feat(tactic/linear_combination): allow linear_combination with { exponent := n } Mar 24, 2023
@bors bors bot closed this Mar 24, 2023
@bors bors bot deleted the lc-exponent branch March 24, 2023 14:05
bors bot pushed a commit to leanprover-community/mathlib4 that referenced this pull request Oct 20, 2023
[mathlib3#15428](leanprover-community/mathlib3#15428) added an exponent option to `linear_combination`. The idea is that `linear_combination (exp := 2) ...` will take a linear combination of hypotheses adding up to the *square* of the goal. This is only mildly useful on its own, but it's a very useful certificate syntax for [mathlib3#15425](leanprover-community/mathlib3#15425).



Co-authored-by: Rob Lewis <rob.y.lewis@gmail.com>
Sign up for free to subscribe to this conversation on GitHub. Already have an account? Sign in.
Labels
modifies-tactic-syntax This PR adds a new interactive tactic or modifies the syntax of an existing tactic. ready-to-merge All that is left is for bors to build and merge this PR. (Remember you need to say `bors r+`.) t-meta Tactics, attributes or user commands
Projects
None yet
Development

Successfully merging this pull request may close these issues.

4 participants