doc-next-gen

Init.Tactics

Module docstring

{}

Lean.Parser.Tactic.as_aux_lemma definition
: Lean.ParserDescr✝
Full source
/--
`as_aux_lemma => tac` does the same as `tac`, except that it wraps the resulting expression
into an auxiliary lemma. In some cases, this significantly reduces the size of expressions
because the proof term is not duplicated.
-/
syntax (name := as_aux_lemma) "as_aux_lemma" " => " tacticSeq : tactic
Auxiliary lemma tactic wrapper
Informal description
The tactic `as_aux_lemma => tac` executes the tactic `tac` and wraps the resulting expression in an auxiliary lemma. This can reduce the size of expressions by avoiding duplication of the proof term.
Lean.Parser.Tactic.withAnnotateState definition
: Lean.ParserDescr✝
Full source
/--
`with_annotate_state stx t` annotates the lexical range of `stx : Syntax` with
the initial and final state of running tactic `t`.
-/
scoped syntax (name := withAnnotateState)
  "with_annotate_state " rawStx ppSpace tactic : tactic
State annotation tactic
Informal description
The tactic `with_annotate_state stx t` annotates the lexical range of the syntax `stx` with the initial and final state of executing the tactic `t`.
Lean.Parser.Tactic.intro definition
: Lean.ParserDescr✝
Full source
/--
Introduces one or more hypotheses, optionally naming and/or pattern-matching them.
For each hypothesis to be introduced, the remaining main goal's target type must
be a `let` or function type.

* `intro` by itself introduces one anonymous hypothesis, which can be accessed
  by e.g. `assumption`.
* `intro x y` introduces two hypotheses and names them. Individual hypotheses
  can be anonymized via `_`, or matched against a pattern:
  ```lean
  -- ... ⊢ α × β → ...
  intro (a, b)
  -- ..., a : α, b : β ⊢ ...
  ```
* Alternatively, `intro` can be combined with pattern matching much like `fun`:
  ```lean
  intro
  | n + 1, 0 => tac
  | ...
  ```
-/
syntax (name := intro) "intro" notFollowedBy("|") (ppSpace colGt term:max)* : tactic
Introduction tactic (`intro`)
Informal description
The tactic `intro` introduces one or more hypotheses in a goal, optionally naming them or pattern-matching against their structure. The goal's target type must be a `let` binding or function type for this to apply. - When used alone (`intro`), it introduces one anonymous hypothesis. - When used with names (`intro x y`), it introduces named hypotheses. - It supports pattern matching, either inline (`intro (a, b)`) or with multiple cases (`intro | n + 1, 0 => ...`).
Lean.Parser.Tactic.intros definition
: Lean.ParserDescr✝
Full source
/--
Introduces zero or more hypotheses, optionally naming them.

- `intros` is equivalent to repeatedly applying `intro`
  until the goal is not an obvious candidate for `intro`, which is to say
  that so long as the goal is a `let` or a pi type (e.g. an implication, function, or universal quantifier),
  the `intros` tactic will introduce an anonymous hypothesis.
  This tactic does not unfold definitions.

- `intros x y ...` is equivalent to `intro x y ...`,
  introducing hypotheses for each supplied argument and unfolding definitions as necessary.
  Each argument can be either an identifier or a `_`.
  An identifier indicates a name to use for the corresponding introduced hypothesis,
  and a `_` indicates that the hypotheses should be introduced anonymously.

## Examples

Basic properties:
```lean
def AllEven (f : Nat → Nat) := ∀ n, f n % 2 = 0

-- Introduces the two obvious hypotheses automatically
example : ∀ (f : Nat → Nat), AllEven f → AllEven (fun k => f (k + 1)) := by
  intros
  /- Tactic state
     f✝ : Nat → Nat
     a✝ : AllEven f✝
     ⊢ AllEven fun k => f✝ (k + 1) -/
  sorry

-- Introduces exactly two hypotheses, naming only the first
example : ∀ (f : Nat → Nat), AllEven f → AllEven (fun k => f (k + 1)) := by
  intros g _
  /- Tactic state
     g : Nat → Nat
     a✝ : AllEven g
     ⊢ AllEven fun k => g (k + 1) -/
  sorry

-- Introduces exactly three hypotheses, which requires unfolding `AllEven`
example : ∀ (f : Nat → Nat), AllEven f → AllEven (fun k => f (k + 1)) := by
  intros f h n
  /- Tactic state
     f : Nat → Nat
     h : AllEven f
     n : Nat
     ⊢ (fun k => f (k + 1)) n % 2 = 0 -/
  apply h
```

Implications:
```lean
example (p q : Prop) : p → q → p := by
  intros
  /- Tactic state
     a✝¹ : p
     a✝ : q
     ⊢ p      -/
  assumption
```

Let bindings:
```lean
example : let n := 1; let k := 2; n + k = 3 := by
  intros
  /- n✝ : Nat := 1
     k✝ : Nat := 2
     ⊢ n✝ + k✝ = 3 -/
  rfl
```
-/
syntax (name := intros) "intros" (ppSpace colGt (ident <|> hole))* : tactic
Introduction tactic for hypotheses (`intros`)
Informal description
The `intros` tactic introduces zero or more hypotheses in a goal, optionally naming them. - When used without arguments (`intros`), it repeatedly applies `intro` until the goal is no longer a `let` binding, function type, implication, or universal quantifier, introducing anonymous hypotheses. - When used with arguments (`intros x y ...`), it introduces hypotheses for each argument, naming them if identifiers are provided or leaving them anonymous if `_` is used. This variant unfolds definitions as needed. **Examples:** 1. For a universally quantified goal `∀ (f : Nat → Nat), AllEven f → AllEven (fun k => f (k + 1))`, `intros` introduces two anonymous hypotheses `f✝` and `a✝`. 2. With `intros g _`, it introduces `g` and an anonymous hypothesis. 3. With `intros f h n`, it introduces `f`, `h`, and `n`, unfolding `AllEven` in the process. 4. For implications like `p → q → p`, `intros` introduces anonymous hypotheses for `p` and `q`. 5. For let bindings like `let n := 1; let k := 2; n + k = 3`, `intros` introduces anonymous bindings for `n✝` and `k✝`.
Lean.Parser.Tactic.rename definition
: Lean.ParserDescr✝
Full source
/--
`rename t => x` renames the most recent hypothesis whose type matches `t`
(which may contain placeholders) to `x`, or fails if no such hypothesis could be found.
-/
syntax (name := rename) "rename " term " => " ident : tactic
Rename Hypothesis Tactic
Informal description
The tactic `rename t => x` renames the most recent hypothesis in the current goal whose type matches the pattern `t` (which may contain placeholders) to the name `x`. If no such hypothesis is found, the tactic fails.
Lean.Parser.Tactic.clear definition
: Lean.ParserDescr✝
Full source
/--
`clear x...` removes the given hypotheses, or fails if there are remaining
references to a hypothesis.
-/
syntax (name := clear) "clear" (ppSpace colGt term:max)+ : tactic
Clear tactic for removing hypotheses
Informal description
The tactic `clear x...` removes the specified hypotheses from the current goal, provided there are no remaining references to them. If any hypothesis is still referenced elsewhere, the tactic fails.
Lean.Parser.Tactic.subst definition
: Lean.ParserDescr✝
Full source
/--
`subst x...` substitutes each `x` with `e` in the goal if there is a hypothesis
of type `x = e` or `e = x`.
If `x` is itself a hypothesis of type `y = e` or `e = y`, `y` is substituted instead.
-/
syntax (name := subst) "subst" (ppSpace colGt term:max)+ : tactic
Substitution tactic for equality hypotheses
Informal description
The tactic `subst` substitutes each variable `x` with `e` in the goal if there exists a hypothesis of the form `x = e` or `e = x`. If `x` is itself a hypothesis of the form `y = e` or `e = y`, then `y` is substituted instead.
Lean.Parser.Tactic.assumption definition
: Lean.ParserDescr✝
Full source
/--
`assumption` tries to solve the main goal using a hypothesis of compatible type, or else fails.
Note also the `‹t›` term notation, which is a shorthand for `show t by assumption`.
-/
syntax (name := assumption) "assumption" : tactic
Assumption tactic
Informal description
The tactic `assumption` attempts to solve the current goal by finding a hypothesis in the context whose type matches the goal. If no such hypothesis exists, the tactic fails. Additionally, the term notation `‹t›` is shorthand for `show t by assumption`, which uses the `assumption` tactic to prove `t`.
Lean.Parser.Tactic.contradiction definition
: Lean.ParserDescr✝
Full source
/--
`contradiction` closes the main goal if its hypotheses are "trivially contradictory".

- Inductive type/family with no applicable constructors
  ```lean
  example (h : False) : p := by contradiction
  ```
- Injectivity of constructors
  ```lean
  example (h : none = some true) : p := by contradiction  --
  ```
- Decidable false proposition
  ```lean
  example (h : 2 + 2 = 3) : p := by contradiction
  ```
- Contradictory hypotheses
  ```lean
  example (h : p) (h' : ¬ p) : q := by contradiction
  ```
- Other simple contradictions such as
  ```lean
  example (x : Nat) (h : x ≠ x) : p := by contradiction
  ```
-/
syntax (name := contradiction) "contradiction" : tactic
Contradiction tactic
Informal description
The tactic `contradiction` closes the current goal when the hypotheses contain a trivial contradiction, such as: - An inductive type with no applicable constructors (e.g., a hypothesis of type `False`). - An equality between distinct constructors of an inductive type (e.g., `none = some true`). - A decidable false proposition (e.g., `2 + 2 = 3`). - Directly contradictory hypotheses (e.g., `h : p` and `h' : ¬p`). - Other simple contradictions like `x ≠ x` for some `x`.
Lean.Parser.Tactic.falseOrByContra definition
: Lean.ParserDescr✝
Full source
/--
Changes the goal to `False`, retaining as much information as possible:

* If the goal is `False`, do nothing.
* If the goal is an implication or a function type, introduce the argument and restart.
  (In particular, if the goal is `x ≠ y`, introduce `x = y`.)
* Otherwise, for a propositional goal `P`, replace it with `¬ ¬ P`
  (attempting to find a `Decidable` instance, but otherwise falling back to working classically)
  and introduce `¬ P`.
* For a non-propositional goal use `False.elim`.
-/
syntax (name := falseOrByContra) "false_or_by_contra" : tactic
False or by contradiction tactic
Informal description
The tactic `false_or_by_contra` transforms the current goal into `False` while preserving as much information as possible. It handles different goal types as follows: - If the goal is already `False`, it does nothing. - If the goal is an implication or function type, it introduces the argument and restarts the process (e.g., for `x ≠ y`, it introduces `x = y`). - For a propositional goal `P`, it replaces it with `¬¬P` (attempting to find a `Decidable` instance, otherwise working classically) and introduces `¬P`. - For a non-propositional goal, it uses `False.elim`.
Lean.Parser.Tactic.apply definition
: Lean.ParserDescr✝
Full source
/--
`apply e` tries to match the current goal against the conclusion of `e`'s type.
If it succeeds, then the tactic returns as many subgoals as the number of premises that
have not been fixed by type inference or type class resolution.
Non-dependent premises are added before dependent ones.

The `apply` tactic uses higher-order pattern matching, type class resolution,
and first-order unification with dependent types.
-/
syntax (name := apply) "apply " term : tactic
Apply tactic
Informal description
The `apply` tactic attempts to match the current goal with the conclusion of the type of the given expression `e`. If successful, it generates subgoals corresponding to the premises of `e` that are not resolved by type inference or type class resolution, prioritizing non-dependent premises over dependent ones. This tactic employs higher-order pattern matching, type class resolution, and first-order unification with dependent types.
Lean.Parser.Tactic.refine definition
: Lean.ParserDescr✝
Full source
/--
`refine e` behaves like `exact e`, except that named (`?x`) or unnamed (`?_`)
holes in `e` that are not solved by unification with the main goal's target type
are converted into new goals, using the hole's name, if any, as the goal case name.
-/
syntax (name := refine) "refine " term : tactic
Refine tactic
Informal description
The `refine` tactic behaves similarly to the `exact` tactic, but with the additional feature that any named (`?x`) or unnamed (`?_`) holes in the expression `e` that are not resolved by unification with the target type of the current goal are converted into new subgoals. The names of these holes, if provided, are used as the case names for the generated subgoals.
Lean.Parser.Tactic.refine' definition
: Lean.ParserDescr✝
Full source
/--
`refine' e` behaves like `refine e`, except that unsolved placeholders (`_`)
and implicit parameters are also converted into new goals.
-/
syntax (name := refine') "refine' " term : tactic
Refine' tactic with implicit parameter handling
Informal description
The tactic `refine' e` behaves similarly to `refine e`, with the additional feature that any unsolved placeholders (`_`) and implicit parameters in the expression `e` are also converted into new goals.
Lean.Parser.Tactic.tacticExfalso definition
: Lean.ParserDescr✝
Full source
macromacro "exfalso" : tactic => `(tactic| refine False.elim ?_)
Exfalso tactic
Informal description
The tactic `exfalso` transforms the current goal `⊢ tgt` into `⊢ False` by applying the elimination principle for `False` (i.e., `False.elim`). This is typically used to replace the current goal with a proof of `False` when a contradiction has been derived.
Lean.Parser.Tactic.constructor definition
: Lean.ParserDescr✝
Full source
/--
If the main goal's target type is an inductive type, `constructor` solves it with
the first matching constructor, or else fails.
-/
syntax (name := constructor) "constructor" : tactic
Constructor tactic for inductive types
Informal description
The tactic `constructor` attempts to solve the current goal by applying the first matching constructor of the goal's target inductive type. If no matching constructor is found, the tactic fails.
Lean.Parser.Tactic.left definition
: Lean.ParserDescr✝
Full source
/--
Applies the first constructor when
the goal is an inductive type with exactly two constructors, or fails otherwise.
```
example : True ∨ False := by
  left
  trivial
```
-/
syntax (name := left) "left" : tactic
Left constructor tactic for binary inductive types
Informal description
The tactic `left` applies the first constructor when the goal is an inductive type with exactly two constructors, and fails otherwise. For example, in the goal `True ∨ False`, applying `left` reduces the goal to `True`.
Lean.Parser.Tactic.right definition
: Lean.ParserDescr✝
Full source
/--
Applies the second constructor when
the goal is an inductive type with exactly two constructors, or fails otherwise.
```
example {p q : Prop} (h : q) : p ∨ q := by
  right
  exact h
```
-/
syntax (name := right) "right" : tactic
Right constructor tactic for binary inductive types
Informal description
The tactic `right` applies the second constructor when the goal is an inductive type with exactly two constructors, and fails otherwise. For example, in the goal `p ∨ q`, applying `right` reduces the goal to `q`.
Lean.Parser.Tactic.case definition
: Lean.ParserDescr✝
Full source
/--
* `case tag => tac` focuses on the goal with case name `tag` and solves it using `tac`,
  or else fails.
* `case tag x₁ ... xₙ => tac` additionally renames the `n` most recent hypotheses
  with inaccessible names to the given names.
* `case tag₁ | tag₂ => tac` is equivalent to `(case tag₁ => tac); (case tag₂ => tac)`.
-/
syntax (name := case) "case " sepBy1(caseArg, " | ") " => " tacticSeq : tactic
Case tactic for goal focusing and solving
Informal description
The `case` tactic focuses on a goal with a specified case name `tag` and attempts to solve it using the provided tactic `tac`. If the tactic fails, the entire operation fails. There are three variants: 1. `case tag => tac`: Solves the goal named `tag` using `tac`. 2. `case tag x₁ ... xₙ => tac`: Additionally renames the `n` most recent hypotheses with inaccessible names to `x₁ ... xₙ` before applying `tac`. 3. `case tag₁ | tag₂ => tac`: Equivalent to applying `tac` sequentially to the goals named `tag₁` and `tag₂`.
Lean.Parser.Tactic.case' definition
: Lean.ParserDescr✝
Full source
/--
`case'` is similar to the `case tag => tac` tactic, but does not ensure the goal
has been solved after applying `tac`, nor admits the goal if `tac` failed.
Recall that `case` closes the goal using `sorry` when `tac` fails, and
the tactic execution is not interrupted.
-/
syntax (name := case') "case' " sepBy1(caseArg, " | ") " => " tacticSeq : tactic
Case tactic (non-admitting variant)
Informal description
The `case'` tactic is a variant of the `case` tactic that focuses on a goal with a specified case name `tag` and applies the tactic `tac` without ensuring the goal is solved afterward. Unlike `case`, it does not admit the goal if `tac` fails. There are three variants: 1. `case' tag => tac`: Applies `tac` to the goal named `tag`. 2. `case' tag x₁ ... xₙ => tac`: Additionally renames the `n` most recent hypotheses with inaccessible names to `x₁ ... xₙ` before applying `tac`. 3. `case' tag₁ | tag₂ => tac`: Applies `tac` sequentially to the goals named `tag₁` and `tag₂`.
Lean.Parser.Tactic.tacticNext_=>_ definition
: Lean.ParserDescr✝
Full source
macromacro nextTk:"next " args:binderIdent* arrowTk:" => " tac:tacticSeq : tactic =>
  -- Limit ref variability for incrementality; see Note [Incremental Macros]
  withRef arrowTk `(tactic| case%$nextTk _ $args* =>%$arrowTk $tac)
Next goal tactic
Informal description
The tactic `next => tac` focuses on the next goal and attempts to solve it using the tactic `tac`. If `tac` fails, the tactic fails. The variant `next x₁ ... xₙ => tac` additionally renames the `n` most recent hypotheses with inaccessible names to `x₁ ... xₙ` before applying `tac`.
Lean.Parser.Tactic.allGoals definition
: Lean.ParserDescr✝
Full source
/--
`all_goals tac` runs `tac` on each goal, concatenating the resulting goals.
If the tactic fails on any goal, the entire `all_goals` tactic fails.

See also `any_goals tac`.
-/
syntax (name := allGoals) "all_goals " tacticSeq : tactic
Apply tactic to all goals
Informal description
The tactic `all_goals tac` applies the tactic sequence `tac` to each goal in the current context, combining the resulting goals. If `tac` fails on any goal, the entire `all_goals` tactic fails.
Lean.Parser.Tactic.anyGoals definition
: Lean.ParserDescr✝
Full source
/--
`any_goals tac` applies the tactic `tac` to every goal,
concating the resulting goals for successful tactic applications.
If the tactic fails on all of the goals, the entire `any_goals` tactic fails.

This tactic is like `all_goals try tac` except that it fails if none of the applications of `tac` succeeds.
-/
syntax (name := anyGoals) "any_goals " tacticSeq : tactic
Apply tactic to any goal (failing if none succeed)
Informal description
The tactic `any_goals tac` applies the tactic sequence `tac` to every goal in the current context, combining the resulting goals from successful applications. If `tac` fails on all goals, the entire `any_goals` tactic fails. This is similar to `all_goals try tac` but ensures failure if no application succeeds.
Lean.Parser.Tactic.focus definition
: Lean.ParserDescr✝
Full source
/--
`focus tac` focuses on the main goal, suppressing all other goals, and runs `tac` on it.
Usually `· tac`, which enforces that the goal is closed by `tac`, should be preferred.
-/
syntax (name := focus) "focus " tacticSeq : tactic
Focus tactic
Informal description
The tactic `focus tac` applies the tactic sequence `tac` to the main goal while temporarily hiding all other goals. It is generally recommended to use `· tac` instead, which ensures that the goal is fully closed by `tac`.
Lean.Parser.Tactic.paren definition
: Lean.ParserDescr✝
Full source
/--
`(tacs)` executes a list of tactics in sequence, without requiring that
the goal be closed at the end like `· tacs`. Like `by` itself, the tactics
can be either separated by newlines or `;`.
-/
syntax (name := paren) "(" withoutPosition(tacticSeq) ")" : tactic
Parenthesized tactic sequence
Informal description
The tactic `(tacs)` executes a sequence of tactics `tacs` enclosed in parentheses, allowing multiple tactics to be run in sequence without requiring that the goal be closed at the end (unlike `· tacs`). The tactics can be separated by newlines or semicolons.
Lean.Parser.Tactic.withReducible definition
: Lean.ParserDescr✝
Full source
/--
`with_reducible tacs` executes `tacs` using the reducible transparency setting.
In this setting only definitions tagged as `[reducible]` are unfolded.
-/
syntax (name := withReducible) "with_reducible " tacticSeq : tactic
Tactic with reducible transparency
Informal description
The tactic `with_reducible tacs` executes the sequence of tactics `tacs` in the reducible transparency setting, where only definitions marked as `[reducible]` are unfolded during execution.
Lean.Parser.Tactic.withReducibleAndInstances definition
: Lean.ParserDescr✝
Full source
/--
`with_reducible_and_instances tacs` executes `tacs` using the `.instances` transparency setting.
In this setting only definitions tagged as `[reducible]` or type class instances are unfolded.
-/
syntax (name := withReducibleAndInstances) "with_reducible_and_instances " tacticSeq : tactic
Tactic with reducible and instances transparency
Informal description
The tactic `with_reducible_and_instances tacs` executes the sequence of tactics `tacs` using the `.instances` transparency setting, where only definitions marked as `[reducible]` or type class instances are unfolded during execution.
Lean.Parser.Tactic.withUnfoldingAll definition
: Lean.ParserDescr✝
Full source
/--
`with_unfolding_all tacs` executes `tacs` using the `.all` transparency setting.
In this setting all definitions that are not opaque are unfolded.
-/
syntax (name := withUnfoldingAll) "with_unfolding_all " tacticSeq : tactic
Tactic with full unfolding transparency
Informal description
The tactic `with_unfolding_all tacs` executes the sequence of tactics `tacs` using the `.all` transparency setting, where all non-opaque definitions are unfolded during execution.
Lean.Parser.Tactic.rotateLeft definition
: Lean.ParserDescr✝
Full source
/--
`rotate_left n` rotates goals to the left by `n`. That is, `rotate_left 1`
takes the main goal and puts it to the back of the subgoal list.
If `n` is omitted, it defaults to `1`.
-/
syntax (name := rotateLeft) "rotate_left" (ppSpace num)? : tactic
Goal rotation tactic (left)
Informal description
The tactic `rotate_left n` shifts the list of goals to the left by `n` positions. Specifically, `rotate_left 1` moves the main goal to the end of the subgoal list. If `n` is not provided, it defaults to `1`.
Lean.Parser.Tactic.rotateRight definition
: Lean.ParserDescr✝
Full source
/--
Rotate the goals to the right by `n`. That is, take the goal at the back
and push it to the front `n` times. If `n` is omitted, it defaults to `1`.
-/
syntax (name := rotateRight) "rotate_right" (ppSpace num)? : tactic
Goal rotation tactic (right)
Informal description
The tactic `rotate_right n` shifts the list of goals to the right by `n` positions. Specifically, it takes the goal at the back and moves it to the front `n` times. If `n` is not provided, it defaults to `1`.
Lean.Parser.Tactic.tacticRfl definition
: Lean.ParserDescr✝
Full source
/--
This tactic applies to a goal whose target has the form `x ~ x`,
where `~` is equality, heterogeneous equality or any relation that
has a reflexivity lemma tagged with the attribute @[refl].
-/
syntax "rfl" : tactic
Reflexivity tactic (`rfl`)
Informal description
The tactic `rfl` applies to a goal of the form `x ~ x`, where `~` is equality, heterogeneous equality, or any reflexive relation (i.e., a relation with a reflexivity lemma tagged with `@[refl]`). It proves such goals by reflexivity.
Lean.Parser.Tactic.applyRfl definition
: Lean.ParserDescr✝
Full source
/--
The same as `rfl`, but without trying `eq_refl` at the end.
-/
syntax (name := applyRfl) "apply_rfl" : tactic
Reflexivity tactic (`apply_rfl`)
Informal description
The tactic `apply_rfl` is similar to `rfl`, but it does not attempt to use `eq_refl` as a fallback. It applies to goals of the form `x ~ x`, where `~` is a reflexive relation (i.e., a relation with a reflexivity lemma tagged with `@[refl]`), and proves such goals by reflexivity.
Lean.Parser.Tactic.acRfl definition
: Lean.ParserDescr✝
Full source
/--
`ac_rfl` proves equalities up to application of an associative and commutative operator.
```
instance : Associative (α := Nat) (.+.) := ⟨Nat.add_assoc⟩
instance : Commutative (α := Nat) (.+.) := ⟨Nat.add_comm⟩

example (a b c d : Nat) : a + b + c + d = d + (b + c) + a := by ac_rfl
```
-/
syntax (name := acRfl) "ac_rfl" : tactic
Reflexivity modulo associativity and commutativity
Informal description
The tactic `ac_rfl` proves equalities up to application of an associative and commutative operator. For example, given instances proving that addition on natural numbers is associative and commutative, `ac_rfl` can prove equalities like `a + b + c + d = d + (b + c) + a`.
Lean.Parser.Tactic.tacticSorry definition
: Lean.ParserDescr✝
Full source
macromacro "sorry" : tactic => `(tactic| exact sorry)
The `sorry` tactic placeholder
Informal description
The `sorry` tactic is a placeholder that temporarily closes a proof goal by using an incomplete proof term. It allows for the creation of syntactically correct proof skeletons while marking parts that still need to be completed. Lean issues a warning whenever `sorry` is used in a proof, and the presence of `sorry` can be detected by checking for `sorryAx` in the theorem's axioms.
Lean.Parser.Tactic.tacticAdmit definition
: Lean.ParserDescr✝
Full source
macromacro "admit" : tactic => `(tactic| sorry)
The `admit` tactic placeholder
Informal description
The `admit` tactic is a synonym for `sorry`, used as a placeholder for incomplete proofs. When used, it temporarily closes a proof goal by using an incomplete proof term, allowing for the creation of syntactically correct proof skeletons while marking parts that still need to be completed. Lean issues a warning whenever `admit` (or `sorry`) is used in a proof.
Lean.Parser.Tactic.valConfigItem definition
: Lean.ParserDescr✝
Full source
/--
`(opt := val)` sets the `opt` configuration option to `val`.

As a special case, `(config := ...)` sets the entire configuration.
-/
syntax valConfigItem := atomic(" (" notFollowedBy(&"discharger" <|> &"disch") (ident <|> &"config")) " := " withoutPosition(term) ")"
Configuration value assignment syntax
Informal description
The syntax `(opt := val)` sets the configuration option `opt` to the value `val`. As a special case, `(config := ...)` sets the entire configuration.
Lean.Parser.Tactic.location definition
: Lean.ParserDescr✝
Full source
/--
Location specifications are used by many tactics that can operate on either the
hypotheses or the goal. It can have one of the forms:
* 'empty' is not actually present in this syntax, but most tactics use
  `(location)?` matchers. It means to target the goal only.
* `at h₁ ... hₙ`: target the hypotheses `h₁`, ..., `hₙ`
* `at h₁ h₂ ⊢`: target the hypotheses `h₁` and `h₂`, and the goal
* `at *`: target all hypotheses and the goal
-/
syntax location := withPosition(ppGroup(" at" (locationWildcard <|> locationHyp)))
Tactic location syntax
Informal description
The syntax for specifying locations where a tactic should operate, which can be: - Empty (default): target only the goal - `at h₁ ... hₙ`: target hypotheses `h₁` through `hₙ` - `at h₁ h₂ ⊢`: target hypotheses `h₁`, `h₂`, and the goal (denoted by `⊢`) - `at *`: target all hypotheses and the goal
Lean.Parser.Tactic.change definition
: Lean.ParserDescr✝
Full source
/--
* `change tgt'` will change the goal from `tgt` to `tgt'`,
  assuming these are definitionally equal.
* `change t' at h` will change hypothesis `h : t` to have type `t'`, assuming
  assuming `t` and `t'` are definitionally equal.
-/
syntax (name := change) "change " term (location)? : tactic
Change tactic for goals and hypotheses
Informal description
The `change` tactic modifies the goal or a hypothesis to a definitionally equal type. Specifically: - `change tgt'` changes the goal from `tgt` to `tgt'`, provided `tgt` and `tgt'` are definitionally equal. - `change t' at h` changes the type of hypothesis `h : t` to `t'`, provided `t` and `t'` are definitionally equal.
Lean.Parser.Tactic.changeWith definition
: Lean.ParserDescr✝
Full source
/--
* `change a with b` will change occurrences of `a` to `b` in the goal,
  assuming `a` and `b` are definitionally equal.
* `change a with b at h` similarly changes `a` to `b` in the type of hypothesis `h`.
-/
syntax (name := changeWith) "change " term " with " term (location)? : tactic
Definitional change tactic
Informal description
The tactic `change a with b` replaces occurrences of `a` with `b` in the goal, provided that `a` and `b` are definitionally equal. The variant `change a with b at h` performs the same replacement in the type of hypothesis `h`.
Lean.Parser.Tactic.rwRule definition
: Lean.ParserDescr✝
Full source
/--
If `thm` is a theorem `a = b`, then as a rewrite rule,
* `thm` means to replace `a` with `b`, and
* `← thm` means to replace `b` with `a`.
-/
syntax rwRule    := patternIgnore("← " <|> "<- ")? term
Rewrite rule syntax
Informal description
A rewrite rule `rwRule` is defined as an optional left arrow `←` or `<-` followed by a term. Given a theorem `thm` stating `a = b`, the rule `thm` replaces `a` with `b`, while `← thm` replaces `b` with `a`.
Lean.Parser.Tactic.rewriteSeq definition
: Lean.ParserDescr✝
Full source
/--
`rewrite [e]` applies identity `e` as a rewrite rule to the target of the main goal.
If `e` is preceded by left arrow (`←` or `<-`), the rewrite is applied in the reverse direction.
If `e` is a defined constant, then the equational theorems associated with `e` are used.
This provides a convenient way to unfold `e`.
- `rewrite [e₁, ..., eₙ]` applies the given rules sequentially.
- `rewrite [e] at l` rewrites `e` at location(s) `l`, where `l` is either `*` or a
  list of hypotheses in the local context. In the latter case, a turnstile `⊢` or `|-`
  can also be used, to signify the target of the goal.

Using `rw (occs := .pos L) [e]`,
where `L : List Nat`, you can control which "occurrences" are rewritten.
(This option applies to each rule, so usually this will only be used with a single rule.)
Occurrences count from `1`.
At each allowed occurrence, arguments of the rewrite rule `e` may be instantiated,
restricting which later rewrites can be found.
(Disallowed occurrences do not result in instantiation.)
`(occs := .neg L)` allows skipping specified occurrences.
-/
syntax (name := rewriteSeq) "rewrite" optConfig rwRuleSeq (location)? : tactic
Rewrite tactic
Informal description
The `rewrite` tactic applies a given identity `e` as a rewrite rule to the target of the main goal. If `e` is preceded by a left arrow (`←` or `<-`), the rewrite is applied in the reverse direction. If `e` is a defined constant, the equational theorems associated with `e` are used, providing a convenient way to unfold `e`. - `rewrite [e₁, ..., eₙ]` applies the given rules sequentially. - `rewrite [e] at l` rewrites `e` at location(s) `l`, where `l` is either `*` (all hypotheses and the goal) or a list of hypotheses in the local context. The turnstile `⊢` or `|-` can also be used to signify the goal. Using `rw (occs := .pos L) [e]`, where `L : List Nat`, allows controlling which occurrences are rewritten (counted from `1`). At each allowed occurrence, arguments of the rewrite rule `e` may be instantiated, restricting later rewrites. `(occs := .neg L)` skips specified occurrences.
Lean.Parser.Tactic.rwSeq definition
: Lean.ParserDescr✝
Full source
macro (name := rwSeq) "rw " c:optConfig s:rwRuleSeq l:(location)? : tactic =>
  match s with
  | `(rwRuleSeq| [$rs,*]%$rbrak) =>
    -- We show the `rfl` state on `]`
    `(tactic| (rewrite $c [$rs,*] $(l)?; with_annotate_state $rbrak (try (with_reducible rfl))))
  | _ => Macro.throwUnsupported
Rewrite with reflexivity tactic (`rw`)
Informal description
The `rw` tactic is similar to the `rewrite` tactic, but additionally attempts to close the goal by applying a "cheap" (reducible) reflexivity check (`rfl`) after performing the rewrites. The tactic accepts an optional configuration and a sequence of rewrite rules, which can be applied to specified locations in the goal or hypotheses.
Lean.Parser.Tactic.injection definition
: Lean.ParserDescr✝
Full source
/--
The `injection` tactic is based on the fact that constructors of inductive data
types are injections.
That means that if `c` is a constructor of an inductive datatype, and if `(c t₁)`
and `(c t₂)` are two terms that are equal then  `t₁` and `t₂` are equal too.
If `q` is a proof of a statement of conclusion `t₁ = t₂`, then injection applies
injectivity to derive the equality of all arguments of `t₁` and `t₂` placed in
the same positions. For example, from `(a::b) = (c::d)` we derive `a=c` and `b=d`.
To use this tactic `t₁` and `t₂` should be constructor applications of the same constructor.
Given `h : a::b = c::d`, the tactic `injection h` adds two new hypothesis with types
`a = c` and `b = d` to the main goal.
The tactic `injection h with h₁ h₂` uses the names `h₁` and `h₂` to name the new hypotheses.
-/
syntax (name := injection) "injection " term (" with" (ppSpace colGt (ident <|> hole))+)? : tactic
Injection tactic for constructor equalities
Informal description
The `injection` tactic leverages the injectivity property of constructors in inductive data types. Specifically, if `c` is a constructor and `c t₁ = c t₂` holds, then `t₁ = t₂` must also hold. Given a hypothesis `h : t₁ = t₂` where `t₁` and `t₂` are applications of the same constructor, the tactic derives equalities between corresponding arguments of `t₁` and `t₂`. For example, from `h : (a::b) = (c::d)`, it derives `a = c` and `b = d`. Optionally, the derived equalities can be named using the `with` clause, e.g., `injection h with h₁ h₂` names the new hypotheses `h₁ : a = c` and `h₂ : b = d`.
Lean.Parser.Tactic.injections definition
: Lean.ParserDescr✝
Full source
/-- `injections` applies `injection` to all hypotheses recursively
(since `injection` can produce new hypotheses). Useful for destructing nested
constructor equalities like `(a::b::c) = (d::e::f)`. -/
-- TODO: add with
syntax (name := injections) "injections" (ppSpace colGt (ident <|> hole))* : tactic
Recursive injection tactic for constructor equalities
Informal description
The tactic `injections` recursively applies the `injection` tactic to all hypotheses, which is useful for deconstructing nested constructor equalities such as `(a::b::c) = (d::e::f)`. Since `injection` can produce new hypotheses, this tactic handles those recursively.
Lean.Parser.Tactic.discharger definition
: Lean.ParserDescr✝
Full source
/--
The discharger clause of `simp` and related tactics.
This is a tactic used to discharge the side conditions on conditional rewrite rules.
-/
syntax discharger := atomic(" (" patternIgnore(&"discharger" <|> &"disch")) " := " withoutPosition(tacticSeq) ")"
Discharger clause syntax for simplification tactics
Informal description
The syntax definition for the `discharger` clause in `simp` and related tactics, which specifies a tactic used to discharge side conditions on conditional rewrite rules.
Lean.Parser.Tactic.simpLemma definition
: Lean.ParserDescr✝
Full source
/--
A simp lemma specification is:
* optional `↑` or `↓` to specify use before or after entering the subterm
* optional `←` to use the lemma backward
* `thm` for the theorem to rewrite with
-/
syntax simpLemma := (simpPre <|> simpPost)? patternIgnore("← " <|> "<- ")? term
Simplification lemma syntax specification
Informal description
The syntax specification for a simplification lemma in the `simp` tactic consists of: 1. An optional `↑` or `↓` to indicate whether the lemma should be applied before or after entering subterms 2. An optional `←` or `<-` to indicate the lemma should be used in the reverse direction 3. A theorem name `thm` to be used for rewriting
Lean.Parser.Tactic.simp definition
: Lean.ParserDescr✝
Full source
/--
The `simp` tactic uses lemmas and hypotheses to simplify the main goal target or
non-dependent hypotheses. It has many variants:
- `simp` simplifies the main goal target using lemmas tagged with the attribute `[simp]`.
- `simp [h₁, h₂, ..., hₙ]` simplifies the main goal target using the lemmas tagged
  with the attribute `[simp]` and the given `hᵢ`'s, where the `hᵢ`'s are expressions.-
- If an `hᵢ` is a defined constant `f`, then `f` is unfolded. If `f` has equational lemmas associated
  with it (and is not a projection or a `reducible` definition), these are used to rewrite with `f`.
- `simp [*]` simplifies the main goal target using the lemmas tagged with the
  attribute `[simp]` and all hypotheses.
- `simp only [h₁, h₂, ..., hₙ]` is like `simp [h₁, h₂, ..., hₙ]` but does not use `[simp]` lemmas.
- `simp [-id₁, ..., -idₙ]` simplifies the main goal target using the lemmas tagged
  with the attribute `[simp]`, but removes the ones named `idᵢ`.
- `simp at h₁ h₂ ... hₙ` simplifies the hypotheses `h₁ : T₁` ... `hₙ : Tₙ`. If
  the target or another hypothesis depends on `hᵢ`, a new simplified hypothesis
  `hᵢ` is introduced, but the old one remains in the local context.
- `simp at *` simplifies all the hypotheses and the target.
- `simp [*] at *` simplifies target and all (propositional) hypotheses using the
  other hypotheses.
-/
syntax (name := simp) "simp" optConfig (discharger)? (&" only")?
  (" [" withoutPosition((simpStar <|> simpErase <|> simpLemma),*,?) "]")? (location)? : tactic
Simplification tactic (`simp`)
Informal description
The `simp` tactic simplifies the main goal target or non-dependent hypotheses using lemmas tagged with the `[simp]` attribute and optionally specified additional lemmas or hypotheses. It supports various configurations including: - Basic simplification with `[simp]` lemmas - Including additional lemmas or hypotheses via `simp [h₁, h₂, ..., hₙ]` - Unfolding definitions when they appear in the arguments - Using all hypotheses with `simp [*]` - Restricting to only specified lemmas with `simp only [h₁, ..., hₙ]` - Excluding specific `[simp]` lemmas with `simp [-id₁, ..., -idₙ]` - Simplifying at specific locations with `simp at h₁ ... hₙ` or `simp at *` - Combining these features with `simp [*] at *`
Lean.Parser.Tactic.simpAll definition
: Lean.ParserDescr✝
Full source
/--
`simp_all` is a stronger version of `simp [*] at *` where the hypotheses and target
are simplified multiple times until no simplification is applicable.
Only non-dependent propositional hypotheses are considered.
-/
syntax (name := simpAll) "simp_all" optConfig (discharger)? (&" only")?
  (" [" withoutPosition((simpErase <|> simpLemma),*,?) "]")? : tactic
Enhanced simplification tactic (`simp_all`)
Informal description
The `simp_all` tactic is an enhanced version of `simp [*] at *` that repeatedly simplifies both the hypotheses and the target until no further simplifications can be applied. It only considers non-dependent propositional hypotheses for simplification.
Lean.Parser.Tactic.dsimp definition
: Lean.ParserDescr✝
Full source
/--
The `dsimp` tactic is the definitional simplifier. It is similar to `simp` but only
applies theorems that hold by reflexivity. Thus, the result is guaranteed to be
definitionally equal to the input.
-/
syntax (name := dsimp) "dsimp" optConfig (discharger)? (&" only")?
  (" [" withoutPosition((simpErase <|> simpLemma),*,?) "]")? (location)? : tactic
Definitional Simplifier Tactic (`dsimp`)
Informal description
The `dsimp` tactic is the definitional simplifier in Lean. Unlike the general `simp` tactic, `dsimp` only applies theorems that hold by reflexivity, ensuring the result is definitionally equal to the input. It supports optional configuration, a discharger clause for handling side conditions, and allows restricting the simplification rules via an `only` modifier or explicit lemma specifications.
Lean.Parser.Tactic.simpArg definition
Full source
/--
A `simpArg` is either a `*`, `-lemma` or a simp lemma specification
(which includes the `↑` `↓` `←` specifications for pre, post, reverse rewriting).
-/
def simpArg := simpStar.binary `orelse (simpErase.binary `orelse simpLemma)
Simplification argument specification
Informal description
A `simpArg` is either: 1. A `*` indicating all available hypotheses should be used as simplification lemmas 2. A `-lemma` indicating a theorem should be removed from the simplification rules 3. A lemma specification with optional direction markers (`↑`, `↓`, `←`) for controlling how the lemma is applied during rewriting
Lean.Parser.Tactic.simpArgs definition
: Lean.ParserDescr✝
Full source
/-- A simp args list is a list of `simpArg`. This is the main argument to `simp`. -/
syntax simpArgs := " [" simpArg,* "]"
Simplification arguments list syntax
Informal description
The syntax for specifying a list of simplification arguments in the `simp` tactic, enclosed in square brackets and separated by commas. Each argument is a `simpArg` which can be: 1. A `*` to use all available hypotheses as simplification lemmas 2. A `-lemma` to remove a theorem from the simplification rules 3. A lemma specification with optional direction markers (`↑`, `↓`, `←`) for controlling the rewriting direction.
Lean.Parser.Tactic.dsimpArg definition
Full source
/--
A `dsimpArg` is similar to `simpArg`, but it does not have the `simpStar` form
because it does not make sense to use hypotheses in `dsimp`.
-/
def dsimpArg := simpErase.binary `orelse simpLemma
`dsimp` tactic argument specification
Informal description
A `dsimpArg` is a component used in the `dsimp` tactic, which can either be a simplification rule to be erased (using the syntax `-thm`) or a simplification lemma (specified by a theorem name `thm`). Unlike `simpArg`, it does not support the `simpStar` form since `dsimp` does not utilize hypotheses.
Lean.Parser.Tactic.dsimpArgs definition
: Lean.ParserDescr✝
Full source
/-- A dsimp args list is a list of `dsimpArg`. This is the main argument to `dsimp`. -/
syntax dsimpArgs := " [" dsimpArg,* "]"
`dsimp` tactic arguments list
Informal description
The syntax `dsimpArgs` defines a list of `dsimpArg` elements enclosed in square brackets, which serves as the primary argument to the `dsimp` tactic. Each `dsimpArg` can be either a simplification rule to be removed (using `-thm`) or a simplification lemma (specified by a theorem name `thm`). Unlike `simpArg`, it does not support the `simpStar` form since `dsimp` does not use hypotheses.
Lean.Parser.Tactic.simpTraceArgsRest definition
: Lean.ParserDescr✝
Full source
/-- The common arguments of `simp?` and `simp?!`. -/
syntax simpTraceArgsRest := optConfig (discharger)? (&" only")? (simpArgs)? (ppSpace location)?
Remaining arguments syntax for `simp?` and `simp?!` tactics
Informal description
The syntax definition for the remaining arguments of `simp?` and `simp?!` tactics, which includes optional configuration options, an optional discharger clause, an optional "only" modifier, optional simplification arguments, and an optional location specification.
Lean.Parser.Tactic.simpTrace definition
: Lean.ParserDescr✝
Full source
/--
`simp?` takes the same arguments as `simp`, but reports an equivalent call to `simp only`
that would be sufficient to close the goal. This is useful for reducing the size of the simp
set in a local invocation to speed up processing.
```
example (x : Nat) : (if True then x + 2 else 3) = x + 2 := by
  simp? -- prints "Try this: simp only [ite_true]"
```

This command can also be used in `simp_all` and `dsimp`.
-/
syntax (name := simpTrace) "simp?" "!"? simpTraceArgsRest : tactic
`simp?` tactic for reporting minimal simplification sets
Informal description
The `simp?` tactic takes the same arguments as `simp`, but instead of applying the simplifications, it reports an equivalent call to `simp only` that would be sufficient to close the goal. This is useful for reducing the size of the simp set in a local invocation to speed up processing. For example, in the case of proving `(if True then x + 2 else 3) = x + 2`, `simp?` would suggest using `simp only [ite_true]`. This command can also be used in `simp_all` and `dsimp`.
Lean.Parser.Tactic.tacticSimp?!_ definition
: Lean.ParserDescr✝
Full source
macromacro tk:"simp?!" rest:simpTraceArgsRest : tactic => `(tactic| simp?%$tk ! $rest)
Interactive aggressive simplification tactic (`simp?!`)
Informal description
The macro `simp?!` is a variant of the `simp` tactic that suggests simplifications interactively, with the `!` modifier indicating that it should use more aggressive simplification rules.
Lean.Parser.Tactic.simpAllTraceArgsRest definition
: Lean.ParserDescr✝
Full source
/-- The common arguments of `simp_all?` and `simp_all?!`. -/
syntax simpAllTraceArgsRest := optConfig (discharger)? (&" only")? (dsimpArgs)?
Remaining arguments syntax for `simp_all?` and `simp_all?!` tactics
Informal description
The syntax definition for the remaining arguments of `simp_all?` and `simp_all?!` tactics, which includes optional configuration, an optional discharger clause, an optional `only` modifier, and an optional list of `dsimp` arguments.
Lean.Parser.Tactic.tacticSimp_all?!_ definition
: Lean.ParserDescr✝
Full source
macromacro tk:"simp_all?!" rest:simpAllTraceArgsRest : tactic => `(tactic| simp_all?%$tk ! $rest)
Aggressive simplification with tracing (`simp_all?!`)
Informal description
The tactic `simp_all?!` is a variant of `simp_all` that performs a more aggressive simplification, indicated by the `!` modifier, and traces the simplification steps, indicated by the `?` modifier. It takes optional configuration arguments for further customization.
Lean.Parser.Tactic.dsimpTraceArgsRest definition
: Lean.ParserDescr✝
Full source
/-- The common arguments of `dsimp?` and `dsimp?!`. -/
syntax dsimpTraceArgsRest := optConfig (&" only")? (dsimpArgs)? (ppSpace location)?
Common arguments for `dsimp?` and `dsimp?!` tactics
Informal description
The syntax for the common arguments of the `dsimp?` and `dsimp?!` tactics, which includes optional configuration options, an optional `only` keyword, an optional list of simplification arguments, and an optional location specification.
Lean.Parser.Tactic.dsimpTrace definition
: Lean.ParserDescr✝
Full source
@[inherit_doc simpTrace]
syntax (name := dsimpTrace) "dsimp?" "!"? dsimpTraceArgsRest : tactic
`dsimp?` tactic syntax
Informal description
The syntax for the `dsimp?` tactic, which optionally takes a `!` and common arguments for `dsimp?` and `dsimp?!` tactics. The common arguments include optional configuration options, an optional `only` keyword, an optional list of simplification arguments, and an optional location specification.
Lean.Parser.Tactic.tacticDsimp?!_ definition
: Lean.ParserDescr✝
Full source
macromacro tk:"dsimp?!" rest:dsimpTraceArgsRest : tactic => `(tactic| dsimp?%$tk ! $rest)
Forced display definitional simplification tactic (`dsimp?!`)
Informal description
The tactic `dsimp?!` performs definitional simplification on the goal, displaying the simplified terms and the list of simplification lemmas used. It is a variant of `dsimp?` that forces the display of the trace information.
Lean.Parser.Tactic.simpaArgsRest definition
: Lean.ParserDescr✝
Full source
/-- The arguments to the `simpa` family tactics. -/
syntax simpaArgsRest := optConfig (discharger)? &" only "? (simpArgs)? (" using " term)?
Remaining arguments syntax for `simpa` tactics
Informal description
The syntax definition for the remaining arguments of the `simpa` family of tactics, which includes optional configuration, an optional discharger clause, an optional `only` modifier, an optional list of simplification arguments, and an optional `using` clause followed by a term.
Lean.Parser.Tactic.simpa definition
: Lean.ParserDescr✝
Full source
/--
This is a "finishing" tactic modification of `simp`. It has two forms.

* `simpa [rules, ⋯] using e` will simplify the goal and the type of
  `e` using `rules`, then try to close the goal using `e`.

  Simplifying the type of `e` makes it more likely to match the goal
  (which has also been simplified). This construction also tends to be
  more robust under changes to the simp lemma set.

* `simpa [rules, ⋯]` will simplify the goal and the type of a
  hypothesis `this` if present in the context, then try to close the goal using
  the `assumption` tactic.
-/
syntax (name := simpa) "simpa" "?"? "!"? simpaArgsRest : tactic
Simplification and Assumption Tactic (`simpa`)
Informal description
The `simpa` tactic is a finishing tactic that simplifies both the goal and the type of a given expression using specified rules, then attempts to close the goal. It has two main forms: 1. `simpa [rules, ...] using e` simplifies the goal and the type of expression `e` using the given rules, then tries to close the goal using `e`. 2. `simpa [rules, ...]` simplifies the goal and the type of a hypothesis `this` (if present in the context), then tries to close the goal using the `assumption` tactic.
Lean.Parser.Tactic.tacticSimpa!_ definition
: Lean.ParserDescr✝
Full source
macromacro "simpa!" rest:simpaArgsRest : tactic =>
  `(tactic| simpa ! $rest:simpaArgsRest)
`simpa!` tactic with aggressive simplification
Informal description
The macro `simpa!` is a variant of the `simpa` tactic that applies the `simplify` tactic with the `!` modifier (which performs more aggressive simplification) followed by the remaining arguments specified in `rest`.
Lean.Parser.Tactic.tacticSimpa?_ definition
: Lean.ParserDescr✝
Full source
macromacro "simpa?" rest:simpaArgsRest : tactic =>
  `(tactic| simpa ? $rest:simpaArgsRest)
Interactive simplification suggestion tactic (`simpa?`)
Informal description
The tactic `simpa?` is a variant of the `simpa` tactic that suggests possible simplification steps to the user. It takes the same optional arguments as `simpa`, including configuration, discharger clause, simplification arguments, and a `using` clause followed by a term.
Lean.Parser.Tactic.delta definition
: Lean.ParserDescr✝
Full source
/--
`delta id1 id2 ...` delta-expands the definitions `id1`, `id2`, ....
This is a low-level tactic, it will expose how recursive definitions have been
compiled by Lean.
-/
syntax (name := delta) "delta" (ppSpace colGt ident)+ (location)? : tactic
Delta-expansion tactic
Informal description
The tactic `delta id₁ id₂ ...` performs delta-expansion on the definitions `id₁`, `id₂`, ..., which unfolds these definitions completely. This is a low-level tactic that reveals how recursive definitions have been compiled internally by Lean.
Lean.Parser.Tactic.unfold definition
: Lean.ParserDescr✝
Full source
/--
* `unfold id` unfolds all occurrences of definition `id` in the target.
* `unfold id1 id2 ...` is equivalent to `unfold id1; unfold id2; ...`.
* `unfold id at h` unfolds at the hypothesis `h`.

Definitions can be either global or local definitions.

For non-recursive global definitions, this tactic is identical to `delta`.
For recursive global definitions, it uses the "unfolding lemma" `id.eq_def`,
which is generated for each recursive definition, to unfold according to the recursive definition given by the user.
Only one level of unfolding is performed, in contrast to `simp only [id]`, which unfolds definition `id` recursively.
-/
syntax (name := unfold) "unfold" (ppSpace colGt ident)+ (location)? : tactic
Definition unfolding tactic
Informal description
The `unfold` tactic expands all occurrences of the specified definitions in the target. It can be used in several ways: - `unfold id` unfolds all instances of definition `id` in the goal. - `unfold id1 id2 ...` sequentially unfolds `id1`, then `id2`, etc. - `unfold id at h` unfolds the definition `id` in hypothesis `h`. For non-recursive global definitions, this tactic behaves identically to `delta`. For recursive global definitions, it uses the unfolding lemma `id.eq_def` to perform one level of unfolding according to the user's definition, unlike `simp only [id]` which unfolds recursively.
Lean.Parser.Tactic.tacticHave_ definition
: Lean.ParserDescr✝
Full source
/--
The `have` tactic is for adding hypotheses to the local context of the main goal.
* `have h : t := e` adds the hypothesis `h : t` if `e` is a term of type `t`.
* `have h := e` uses the type of `e` for `t`.
* `have : t := e` and `have := e` use `this` for the name of the hypothesis.
* `have pat := e` for a pattern `pat` is equivalent to `match e with | pat => _`,
  where `_` stands for the tactics that follow this one.
  It is convenient for types that have only one applicable constructor.
  For example, given `h : p ∧ q ∧ r`, `have ⟨h₁, h₂, h₃⟩ := h` produces the
  hypotheses `h₁ : p`, `h₂ : q`, and `h₃ : r`.
-/
syntax "have " haveDecl : tactic
`have` tactic for introducing hypotheses
Informal description
The `have` tactic introduces a new hypothesis to the local context of the current goal. It has several forms: 1. `have h : t := e` adds `h : t` if `e` has type `t`. 2. `have h := e` uses the type of `e` for `t`. 3. `have : t := e` and `have := e` use `this` as the hypothesis name. 4. `have pat := e` is equivalent to `match e with | pat => _`, where `_` represents subsequent tactics. This is useful for types with a single constructor, e.g., given `h : p ∧ q ∧ r`, `have ⟨h₁, h₂, h₃⟩ := h` produces `h₁ : p`, `h₂ : q`, and `h₃ : r`.
Lean.Parser.Tactic.tacticSuffices_ definition
: Lean.ParserDescr✝
Full source
macromacro "suffices " d:sufficesDecl : tactic => `(tactic| refine_lift suffices $d; ?_)
`suffices` tactic for goal modification
Informal description
The tactic `suffices` allows modifying the current goal by introducing a new subgoal. Given a main goal `⊢ t`, the command `suffices h : t' from e` replaces the main goal with `⊢ t'`, where `e` must be a proof of `t` in the context extended with `h : t'`. A shorthand form `suffices h : t' by tac` is equivalent to `suffices h : t' from by tac`. If the hypothesis name `h` is omitted, the default name `this` is used.
Lean.Parser.Tactic.tacticLet_ definition
: Lean.ParserDescr✝
Full source
macromacro "let " d:letDecl : tactic => `(tactic| refine_lift let $d:letDecl; ?_)
`let` tactic for local definitions
Informal description
The `let` tactic introduces a local definition within the current proof context. It can be used in several forms: 1. `let x : t := e` defines `x` of type `t` with value `e` (where `e` has type `t`) 2. `let x := e` defines `x` with type inferred from `e` 3. `let : t := e` or `let := e` defines an unnamed hypothesis `this` 4. `let pat := e` performs pattern matching, equivalent to `match e with | pat => _` This is particularly useful for destructuring compound types with single constructors, such as extracting components from tuples.
Lean.Parser.Tactic.tacticShow_ definition
: Lean.ParserDescr✝
Full source
macromacro "show " e:term : tactic => `(tactic| refine_lift show $e from ?_)
`show` tactic for goal selection and unification
Informal description
The tactic `show t` identifies the first goal whose target type unifies with `t`, makes it the main goal, performs the unification, and replaces the target with the unified version of `t`.
Lean.Parser.Tactic.inductionAltLHS definition
: Lean.ParserDescr✝
Full source
/--
The left hand side of an induction arm, `| foo a b c` or `| @foo a b c`
where `foo` is a constructor of the inductive type and `a b c` are the arguments
to the constructor.
-/
syntax inductionAltLHS := "| " (("@"? ident) <|> hole) (ident <|> hole)*
Induction case pattern syntax
Informal description
The left-hand side of an induction case, written as `| foo a b c` or `| @foo a b c`, where `foo` is a constructor of the inductive type and `a`, `b`, `c` are the arguments to the constructor.
Lean.Parser.Tactic.inductionAlt definition
: Lean.ParserDescr✝
Full source
/--
In induction alternative, which can have 1 or more cases on the left
and `_`, `?_`, or a tactic sequence after the `=>`.
-/
syntax inductionAlt  := ppDedent(ppLine) inductionAltLHS+ " => " (hole <|> syntheticHole <|> tacticSeq)
Induction alternative syntax
Informal description
An induction alternative consists of one or more cases on the left-hand side (written as `| foo a b c` or similar), followed by `=>` and either `_`, `?_`, or a sequence of tactics. This syntax is used to specify how to handle each case in an induction proof.
Lean.Parser.Tactic.inductionAlts definition
: Lean.ParserDescr✝
Full source
/--
After `with`, there is an optional tactic that runs on all branches, and
then a list of alternatives.
-/
syntax inductionAlts := " with" (ppSpace colGt tactic)? withPosition((colGe inductionAlt)*)
Induction alternatives syntax
Informal description
The syntax for specifying alternatives in an induction proof, consisting of an optional tactic that runs on all branches followed by a list of case alternatives. Each alternative is written as `| pattern => tactics`, where `pattern` matches a constructor of the inductive type and `tactics` are the proof steps for that case.
Lean.Parser.Tactic.elimTarget definition
: Lean.ParserDescr✝
Full source
/--
A target for the `induction` or `cases` tactic, of the form `e` or `h : e`.

The `h : e` syntax introduces a hypotheses of the form `h : e = _` in each goal,
with `_` replaced by the corresponding value of the target.
It is useful when `e` is not a free variable.
-/
syntax elimTarget := atomic(binderIdent " : ")? term
Elimination target for induction or cases tactics
Informal description
A target for the `induction` or `cases` tactic, which can be either an expression `e` or a named expression `h : e`. When using the `h : e` syntax, a hypothesis of the form `h : e = _` is introduced in each goal, where `_` is replaced by the corresponding value of the target. This is particularly useful when `e` is not a free variable.
Lean.Parser.Tactic.induction definition
: Lean.ParserDescr✝
Full source
/--
Assuming `x` is a variable in the local context with an inductive type,
`induction x` applies induction on `x` to the main goal,
producing one goal for each constructor of the inductive type,
in which the target is replaced by a general instance of that constructor
and an inductive hypothesis is added for each recursive argument to the constructor.
If the type of an element in the local context depends on `x`,
that element is reverted and reintroduced afterward,
so that the inductive hypothesis incorporates that hypothesis as well.

For example, given `n : Nat` and a goal with a hypothesis `h : P n` and target `Q n`,
`induction n` produces one goal with hypothesis `h : P 0` and target `Q 0`,
and one goal with hypotheses `h : P (Nat.succ a)` and `ih₁ : P a → Q a` and target `Q (Nat.succ a)`.
Here the names `a` and `ih₁` are chosen automatically and are not accessible.
You can use `with` to provide the variables names for each constructor.
- `induction e`, where `e` is an expression instead of a variable,
  generalizes `e` in the goal, and then performs induction on the resulting variable.
- `induction e using r` allows the user to specify the principle of induction that should be used.
  Here `r` should be a term whose result type must be of the form `C t`,
  where `C` is a bound variable and `t` is a (possibly empty) sequence of bound variables
- `induction e generalizing z₁ ... zₙ`, where `z₁ ... zₙ` are variables in the local context,
  generalizes over `z₁ ... zₙ` before applying the induction but then introduces them in each goal.
  In other words, the net effect is that each inductive hypothesis is generalized.
- Given `x : Nat`, `induction x with | zero => tac₁ | succ x' ih => tac₂`
  uses tactic `tac₁` for the `zero` case, and `tac₂` for the `succ` case.
-/
syntax (name := induction) "induction " elimTarget,+ (" using " term)?
  (" generalizing" (ppSpace colGt term:max)+)? (inductionAlts)? : tactic
Mathematical Induction Tactic
Informal description
The `induction` tactic applies mathematical induction on a variable `x` of an inductive type, generating one subgoal for each constructor of the type. Each subgoal replaces the target with a general instance of the constructor and includes inductive hypotheses for recursive arguments. If the type of a local hypothesis depends on `x`, it is temporarily removed and reintroduced to ensure the inductive hypothesis incorporates it. For example, given `n : ℕ` and a goal with hypothesis `h : P n` and target `Q n`, applying `induction n` produces: 1. A subgoal with hypothesis `h : P 0` and target `Q 0`. 2. A subgoal with hypotheses `h : P (Nat.succ a)` and `ih₁ : P a → Q a`, and target `Q (Nat.succ a)`. Additional features include: - Generalizing expressions before induction via `induction e`. - Specifying a custom induction principle via `induction e using r`. - Generalizing additional variables via `induction e generalizing z₁ ... zₙ`. - Explicit case handling via `induction x with | zero => tac₁ | succ x' ih => tac₂`.
Lean.Parser.Tactic.generalize definition
: Lean.ParserDescr✝
Full source
/--
* `generalize ([h :] e = x),+` replaces all occurrences `e`s in the main goal
  with a fresh hypothesis `x`s. If `h` is given, `h : e = x` is introduced as well.
* `generalize e = x at h₁ ... hₙ` also generalizes occurrences of `e`
  inside `h₁`, ..., `hₙ`.
* `generalize e = x at *` will generalize occurrences of `e` everywhere.
-/
syntax (name := generalize) "generalize " generalizeArg,+ (location)? : tactic
Generalize tactic
Informal description
The `generalize` tactic replaces all occurrences of an expression `e` in the main goal with a fresh variable `x`. Optionally, it can introduce a hypothesis `h : e = x` if a name `h` is provided. The tactic can also generalize occurrences of `e` in specified hypotheses or in all hypotheses and the goal when using `at *`.
Lean.Parser.Tactic.cases definition
: Lean.ParserDescr✝
Full source
/--
Assuming `x` is a variable in the local context with an inductive type,
`cases x` splits the main goal, producing one goal for each constructor of the
inductive type, in which the target is replaced by a general instance of that constructor.
If the type of an element in the local context depends on `x`,
that element is reverted and reintroduced afterward,
so that the case split affects that hypothesis as well.
`cases` detects unreachable cases and closes them automatically.

For example, given `n : Nat` and a goal with a hypothesis `h : P n` and target `Q n`,
`cases n` produces one goal with hypothesis `h : P 0` and target `Q 0`,
and one goal with hypothesis `h : P (Nat.succ a)` and target `Q (Nat.succ a)`.
Here the name `a` is chosen automatically and is not accessible.
You can use `with` to provide the variables names for each constructor.
- `cases e`, where `e` is an expression instead of a variable, generalizes `e` in the goal,
  and then cases on the resulting variable.
- Given `as : List α`, `cases as with | nil => tac₁ | cons a as' => tac₂`,
  uses tactic `tac₁` for the `nil` case, and `tac₂` for the `cons` case,
  and `a` and `as'` are used as names for the new variables introduced.
- `cases h : e`, where `e` is a variable or an expression,
  performs cases on `e` as above, but also adds a hypothesis `h : e = ...` to each hypothesis,
  where `...` is the constructor instance for that particular case.
-/
syntax (name := cases) "cases " elimTarget,+ (" using " term)? (inductionAlts)? : tactic
Case-splitting tactic for inductive types
Informal description
The `cases` tactic splits a goal into multiple subgoals, one for each constructor of an inductive type. Given a variable `x` of an inductive type, `cases x` replaces the goal with instances for each constructor of `x`. If other hypotheses depend on `x`, they are temporarily removed and reintroduced after the case split. Unreachable cases are automatically closed. For example, if `n : Nat` and the goal has a hypothesis `h : P n` and target `Q n`, applying `cases n` produces two subgoals: one with `h : P 0` and target `Q 0`, and another with `h : P (Nat.succ a)` and target `Q (Nat.succ a)`, where `a` is an automatically generated name. The tactic can also: - Handle expressions (not just variables) by generalizing them first. - Use explicit constructor names via `with` syntax, e.g., `cases as with | nil => tac₁ | cons a as' => tac₂`. - Add equality hypotheses via `cases h : e`, where `h : e = ...` is added to each subgoal, with `...` replaced by the constructor instance.
Lean.Parser.Tactic.funInduction definition
: Lean.ParserDescr✝
Full source
/--
The `fun_induction` tactic is a convenience wrapper of the `induction` tactic when using a functional
induction principle.

The tactic invocation
```
fun_induction f x₁ ... xₙ y₁ ... yₘ
```
where `f` is a function defined by non-mutual structural or well-founded recursion, is equivalent to
```
induction y₁, ... yₘ using f.induct x₁ ... xₙ
```
where the arguments of `f` are used as arguments to `f.induct` or targets of the induction, as
appropriate.

The form
```
fun_induction f
```
(with no arguments to `f`) searches the goal for an unique eligible application of `f`, and uses
these arguments. An application of `f` is eligible if it is saturated and the arguments that will
become targets are free variables.

The forms `fun_induction f x y generalizing z₁ ... zₙ` and
`fun_induction f x y with | case1 => tac₁ | case2 x' ih => tac₂` work like with `induction.`
-/
syntax (name := funInduction) "fun_induction " term
  (" generalizing" (ppSpace colGt term:max)+)? (inductionAlts)? : tactic
Functional induction tactic
Informal description
The `fun_induction` tactic is a convenience wrapper for performing induction on a function `f` defined by structural or well-founded recursion. Given a function `f` with arguments `x₁ ... xₙ` (fixed parameters) and `y₁ ... yₘ` (targets for induction), the tactic invocation `fun_induction f x₁ ... xₙ y₁ ... yₘ` is equivalent to `induction y₁, ... yₘ using f.induct x₁ ... xₙ`. If no arguments are provided (`fun_induction f`), the tactic searches the goal for a unique eligible application of `f` (where the target arguments are free variables) and uses those arguments automatically. The tactic also supports additional forms like `generalizing` clauses and case alternatives, mirroring the behavior of the standard `induction` tactic.
Lean.Parser.Tactic.funCases definition
: Lean.ParserDescr✝
Full source
/--
The `fun_cass` tactic is a convenience wrapper of the `cases` tactic when using a functional
cases principle.

The tactic invocation
```
fun_cases f x ... y ...`
```
is equivalent to
```
cases y, ... using f.fun_cases x ...
```
where the arguments of `f` are used as arguments to `f.fun_cases` or targets of the case analysis, as
appropriate.

The form
```
fun_cases f
```
(with no arguments to `f`) searches the goal for an unique eligible application of `f`, and uses
these arguments. An application of `f` is eligible if it is saturated and the arguments that will
become targets are free variables.

The form `fun_cases f x y with | case1 => tac₁ | case2 x' ih => tac₂` works like with `cases`.
-/
syntax (name := funCases) "fun_cases " term (inductionAlts)? : tactic
Functional case-splitting tactic
Informal description
The `fun_cases` tactic is a convenience wrapper for the `cases` tactic when using a functional cases principle. Given a function `f` with arguments `x₁ ... xₙ` (fixed parameters) and `y₁ ... yₖ` (targets for case analysis), the invocation `fun_cases f x₁ ... xₙ y₁ ... yₖ` is equivalent to `cases y₁, ... yₖ using f.fun_cases x₁ ... xₙ`. If no arguments are provided (`fun_cases f`), the tactic searches the goal for a unique eligible application of `f` (where the target arguments are free variables) and uses those arguments automatically. The tactic also supports case alternatives via `with` syntax, mirroring the behavior of the standard `cases` tactic.
Lean.Parser.Tactic.tacticRepeat_ definition
: Lean.ParserDescr✝
Full source
/--
`repeat tac` repeatedly applies `tac` so long as it succeeds.
The tactic `tac` may be a tactic sequence, and if `tac` fails at any point in its execution,
`repeat` will revert any partial changes that `tac` made to the tactic state.

The tactic `tac` should eventually fail, otherwise `repeat tac` will run indefinitely.

See also:
* `try tac` is like `repeat tac` but will apply `tac` at most once.
* `repeat' tac` recursively applies `tac` to each goal.
* `first | tac1 | tac2` implements the backtracking used by `repeat`
-/
syntax "repeat " tacticSeq : tactic
Repeat tactic
Informal description
The tactic `repeat tac` applies the tactic `tac` repeatedly as long as it succeeds. If `tac` fails at any point during execution, `repeat` reverts any partial changes made to the tactic state. The tactic `tac` should eventually fail to prevent `repeat tac` from running indefinitely.
Lean.Parser.Tactic.repeat' definition
: Lean.ParserDescr✝
Full source
/--
`repeat' tac` recursively applies `tac` on all of the goals so long as it succeeds.
That is to say, if `tac` produces multiple subgoals, then `repeat' tac` is applied to each of them.

See also:
* `repeat tac` simply repeatedly applies `tac`.
* `repeat1' tac` is `repeat' tac` but requires that `tac` succeed for some goal at least once.
-/
syntax (name := repeat') "repeat' " tacticSeq : tactic
Recursive tactic application (`repeat'`)
Informal description
The tactic `repeat' tac` recursively applies the tactic `tac` to all current goals as long as it succeeds. If `tac` produces multiple subgoals, `repeat' tac` is then applied to each of these subgoals. This differs from `repeat tac` which simply applies `tac` repeatedly without recursion, and from `repeat1' tac` which additionally requires that `tac` succeeds on at least one goal initially.
Lean.Parser.Tactic.repeat1' definition
: Lean.ParserDescr✝
Full source
/--
`repeat1' tac` recursively applies to `tac` on all of the goals so long as it succeeds,
but `repeat1' tac` fails if `tac` succeeds on none of the initial goals.

See also:
* `repeat tac` simply applies `tac` repeatedly.
* `repeat' tac` is like `repeat1' tac` but it does not require that `tac` succeed at least once.
-/
syntax (name := repeat1') "repeat1' " tacticSeq : tactic
Recursive tactic application requiring initial success (`repeat1'`)
Informal description
The tactic `repeat1' tac` applies the tactic `tac` recursively to all goals as long as it succeeds, but fails if `tac` does not succeed on any of the initial goals. This differs from `repeat tac` (which applies `tac` repeatedly without requiring success) and `repeat' tac` (which behaves like `repeat1' tac` but does not require `tac` to succeed at least once).
Lean.Parser.Tactic.tacticTrivial definition
: Lean.ParserDescr✝
Full source
/--
`trivial` tries different simple tactics (e.g., `rfl`, `contradiction`, ...)
to close the current goal.
You can use the command `macro_rules` to extend the set of tactics used. Example:
```
macro_rules | `(tactic| trivial) => `(tactic| simp)
```
-/
syntax "trivial" : tactic
Trivial tactic
Informal description
The tactic `trivial` attempts to close the current goal by applying various simple tactics such as reflexivity (`rfl`), contradiction, and others. The set of tactics used can be extended via the `macro_rules` command.
Lean.Parser.Tactic.classical definition
: Lean.ParserDescr✝
Full source
/--
`classical tacs` runs `tacs` in a scope where `Classical.propDecidable` is a low priority
local instance.

Note that `classical` is a scoping tactic: it adds the instance only within the
scope of the tactic.
-/
syntax (name := classical) "classical" ppDedent(tacticSeq) : tactic
Classical reasoning tactic
Informal description
The tactic `classical tacs` executes the sequence of tactics `tacs` in a scope where the law of excluded middle (`Classical.propDecidable`) is available as a low-priority local instance. This allows classical reasoning within the scope of the tactic sequence.
Lean.Parser.Tactic.split definition
: Lean.ParserDescr✝
Full source
/--
The `split` tactic is useful for breaking nested if-then-else and `match` expressions into separate cases.
For a `match` expression with `n` cases, the `split` tactic generates at most `n` subgoals.

For example, given `n : Nat`, and a target `if n = 0 then Q else R`, `split` will generate
one goal with hypothesis `n = 0` and target `Q`, and a second goal with hypothesis
`¬n = 0` and target `R`.  Note that the introduced hypothesis is unnamed, and is commonly
renamed used the `case` or `next` tactics.

- `split` will split the goal (target).
- `split at h` will split the hypothesis `h`.
-/
syntax (name := split) "split" (ppSpace colGt term)? (location)? : tactic
Split tactic for case analysis
Informal description
The `split` tactic is used to break down nested if-then-else and `match` expressions into separate cases. For a `match` expression with `n` cases, `split` generates at most `n` subgoals. For example, given a goal of the form `if n = 0 then Q else R`, `split` produces two subgoals: one with the hypothesis `n = 0` and target `Q`, and another with the hypothesis `¬n = 0` and target `R`. The tactic can also be applied to hypotheses using the syntax `split at h`.
Lean.Parser.Tactic.dbgTrace definition
: Lean.ParserDescr✝
Full source
/-- `dbg_trace "foo"` prints `foo` when elaborated.
Useful for debugging tactic control flow:
```
example : False ∨ True := by
  first
  | apply Or.inl; trivial; dbg_trace "left"
  | apply Or.inr; trivial; dbg_trace "right"
```
-/
syntax (name := dbgTrace) "dbg_trace " str : tactic
Debug trace tactic
Informal description
The tactic `dbg_trace "foo"` prints the string `"foo"` during elaboration. It is primarily used for debugging tactic control flow, allowing users to track the execution path of their proofs by inserting trace statements at strategic points.
Lean.Parser.Tactic.tacticStop_ definition
: Lean.ParserDescr✝
Full source
macromacro "stop" tacticSeq : tactic => `(tactic| repeat sorry)
Stop tactic (discard remainder of proof)
Informal description
The tactic `stop` is a helper that discards the remainder of a proof by repeatedly applying the `sorry` tactic. It is useful during proof development to temporarily ignore unfinished parts of a proof without commenting them out.
Lean.Parser.Tactic.specialize definition
: Lean.ParserDescr✝
Full source
/--
The tactic `specialize h a₁ ... aₙ` works on local hypothesis `h`.
The premises of this hypothesis, either universal quantifications or
non-dependent implications, are instantiated by concrete terms coming
from arguments `a₁` ... `aₙ`.
The tactic adds a new hypothesis with the same name `h := h a₁ ... aₙ`
and tries to clear the previous one.
-/
syntax (name := specialize) "specialize " term : tactic
Specialization tactic
Informal description
The tactic `specialize h a₁ ... aₙ` instantiates the premises of a local hypothesis `h` (which may be universal quantifications or non-dependent implications) with concrete terms `a₁` through `aₙ`. It then replaces `h` with the new specialized hypothesis `h := h a₁ ... aₙ` and attempts to clear the original hypothesis.
Lean.Parser.Tactic.tacticUnhygienic_ definition
: Lean.ParserDescr✝
Full source
macromacro "unhygienic " t:tacticSeq : tactic => `(tactic| set_option tactic.hygienic false in $t)
Unhygienic tactic execution
Informal description
The tactic `unhygienic` executes a sequence of tactics `tacs` with name hygiene disabled, allowing tactics that normally create inaccessible names to instead create regular variables. This is useful for referring to autogenerated names directly, though it is brittle as naming strategies may change.
Lean.Parser.Tactic.congr definition
: Lean.ParserDescr✝
Full source
/--
Apply congruence (recursively) to goals of the form `⊢ f as = f bs` and `⊢ HEq (f as) (f bs)`.
The optional parameter is the depth of the recursive applications.
This is useful when `congr` is too aggressive in breaking down the goal.
For example, given `⊢ f (g (x + y)) = f (g (y + x))`,
`congr` produces the goals `⊢ x = y` and `⊢ y = x`,
while `congr 2` produces the intended `⊢ x + y = y + x`.
-/
syntax (name := congr) "congr" (ppSpace num)? : tactic
Congruence tactic
Informal description
The tactic `congr` applies congruence recursively to goals of the form `f as = f bs` or `HEq (f as) (f bs)`. An optional parameter specifies the depth of recursive application, allowing control over how aggressively the goal is broken down. For example, given the goal `f (g (x + y)) = f (g (y + x))`, `congr` would produce subgoals `x = y` and `y = x`, while `congr 2` would produce the desired `x + y = y + x`.
Lean.Parser.Tactic.tacDepIfThenElse definition
: Lean.ParserDescr✝
Full source
/--
In tactic mode, `if h : t then tac1 else tac2` can be used as alternative syntax for:
```
by_cases h : t
· tac1
· tac2
```
It performs case distinction on `h : t` or `h : ¬t` and `tac1` and `tac2` are the subproofs.

You can use `?_` or `_` for either subproof to delay the goal to after the tactic, but
if a tactic sequence is provided for `tac1` or `tac2` then it will require the goal to be closed
by the end of the block.
-/
syntax (name := tacDepIfThenElse)
  ppRealGroup(ppRealFill(ppIndent("if " binderIdent " : " term " then") ppSpace matchRhsTacticSeq)
    ppDedent(ppSpace) ppRealFill("else " matchRhsTacticSeq)) : tactic
Dependent if-then-else tactic
Informal description
The tactic `if h : t then tac1 else tac2` provides an alternative syntax for performing case distinction on a term `t` with a named hypothesis `h`. It splits the proof into two cases: one where `h : t` holds (executing `tac1`) and one where `h : ¬t` holds (executing `tac2`). This is equivalent to using `by_cases h : t` followed by the respective subproofs. Placeholders `?_` or `_` can be used for either subproof to delay the goal, but if a tactic sequence is provided, it must close the goal by the end of the block.
Lean.Parser.Tactic.tacIfThenElse definition
: Lean.ParserDescr✝
Full source
/--
In tactic mode, `if t then tac1 else tac2` is alternative syntax for:
```
by_cases t
· tac1
· tac2
```
It performs case distinction on `h† : t` or `h† : ¬t`, where `h†` is an anonymous
hypothesis, and `tac1` and `tac2` are the subproofs. (It doesn't actually use
nondependent `if`, since this wouldn't add anything to the context and hence would be
useless for proving theorems. To actually insert an `ite` application use
`refine if t then ?_ else ?_`.)
-/
syntax (name := tacIfThenElse)
  ppRealGroup(ppRealFill(ppIndent("if " term " then") ppSpace matchRhsTacticSeq)
    ppDedent(ppSpace) ppRealFill("else " matchRhsTacticSeq)) : tactic
Conditional case distinction tactic (`if-then-else`)
Informal description
The tactic `if t then tac1 else tac2` provides an alternative syntax for performing case distinction on a term `t`. It splits the proof into two cases: one where `t` holds (executing `tac1`) and one where `¬t` holds (executing `tac2`). This is equivalent to using `by_cases t` followed by the respective subproofs. The tactic does not use the non-dependent `if` construct, as that would not add any useful information to the context. To explicitly insert an `ite` application, one should use `refine if t then ?_ else ?_` instead.
Lean.Parser.Tactic.tacticNofun definition
: Lean.ParserDescr✝
Full source
macromacro "nofun" : tactic => `(tactic| exact nofun)
`nofun` tactic
Informal description
The tactic `nofun` is a shorthand for `exact nofun`: it introduces the given assumptions and then performs an empty pattern match, which closes the goal if the introduced pattern is impossible to match.
Lean.Parser.Tactic.tacticNomatch_,, definition
: Lean.ParserDescr✝
Full source
macromacro "nomatch " es:term,+ : tactic =>
  `(tactic| exact nomatch $es:term,*)
`nomatch` tactic
Informal description
The tactic `nomatch` is a shorthand for `exact nomatch`, which is used to perform an empty pattern match on the given expressions. This tactic closes the goal if the introduced pattern is impossible to match.
Lean.Parser.Tactic.replace definition
: Lean.ParserDescr✝
Full source
/--
Acts like `have`, but removes a hypothesis with the same name as
this one if possible. For example, if the state is:

```lean
f : α → β
h : α
⊢ goal
```

Then after `replace h := f h` the state will be:

```lean
f : α → β
h : β
⊢ goal
```

whereas `have h := f h` would result in:

```lean
f : α → β
h† : α
h : β
⊢ goal
```

This can be used to simulate the `specialize` and `apply at` tactics of Coq.
-/
syntax (name := replace) "replace" haveDecl : tactic
`replace` tactic (hypothesis replacement variant of `have`)
Informal description
The tactic `replace` acts similarly to `have`, but additionally removes an existing hypothesis with the same name if possible. For example, given the context: ``` f : α → β h : α ⊢ goal ``` executing `replace h := f h` will result in: ``` f : α → β h : β ⊢ goal ``` whereas `have h := f h` would keep the original hypothesis `h : α` (marked as `h†`). This tactic can be used to simulate Coq's `specialize` and `apply at` tactics.
Lean.Parser.Tactic.substEqs definition
: Lean.ParserDescr✝
Full source
/--
`subst_eq` repeatedly substitutes according to the equality proof hypotheses in the context,
replacing the left side of the equality with the right, until no more progress can be made.
-/
syntax (name := substEqs) "subst_eqs" : tactic
Repeated equality substitution tactic (`subst_eqs`)
Informal description
The tactic `subst_eqs` repeatedly substitutes terms in the proof context according to available equality hypotheses, replacing the left-hand side of each equality with the right-hand side, until no further substitutions can be made.
Lean.Parser.Tactic.tacticHaveI_ definition
: Lean.ParserDescr✝
Full source
macromacro "haveI" d:haveDecl : tactic => `(tactic| refine_lift haveI $d:haveDecl; ?_)
Inline hypothesis introduction tactic (`haveI`)
Informal description
The tactic `haveI` introduces a new hypothesis in the current proof context, similar to `have`, but it inlines the value directly rather than creating a `let_fun` term. This can be useful for avoiding unnecessary let bindings in the proof term.
Lean.Parser.Tactic.DecideConfig structure
Full source
/--
Configuration for the `decide` tactic family.
-/
structure DecideConfig where
  /-- If true (default: false), then use only kernel reduction when reducing the `Decidable` instance.
  This is more efficient, since the default mode reduces twice (once in the elaborator and again in the kernel),
  however kernel reduction ignores transparency settings. -/
  kernel : Bool := false
  /-- If true (default: false), then uses the native code compiler to evaluate the `Decidable` instance,
  admitting the result via the axiom `Lean.ofReduceBool`.  This can be significantly more efficient,
  but it is at the cost of increasing the trusted code base, namely the Lean compiler
  and all definitions with an `@[implemented_by]` attribute.
  The instance is only evaluated once. The `native_decide` tactic is a synonym for `decide +native`. -/
  native : Bool := false
  /-- If true (default: true), then when preprocessing the goal, do zeta reduction to attempt to eliminate free variables. -/
  zetaReduce : Bool := true
  /-- If true (default: false), then when preprocessing, removes irrelevant variables and reverts the local context.
  A variable is *relevant* if it appears in the target, if it appears in a relevant variable,
  or if it is a proposition that refers to a relevant variable. -/
  revert : Bool := false
Configuration for the `decide` tactics
Informal description
The structure `DecideConfig` represents the configuration options for the `decide` family of tactics in Lean, which are used to automatically prove decidable propositions.
Lean.Parser.Tactic.decide definition
: Lean.ParserDescr✝
Full source
/--
`decide` attempts to prove the main goal (with target type `p`) by synthesizing an instance of `Decidable p`
and then reducing that instance to evaluate the truth value of `p`.
If it reduces to `isTrue h`, then `h` is a proof of `p` that closes the goal.

The target is not allowed to contain local variables or metavariables.
If there are local variables, you can first try using the `revert` tactic with these local variables to move them into the target,
or you can use the `+revert` option, described below.

Options:
- `decide +revert` begins by reverting local variables that the target depends on,
  after cleaning up the local context of irrelevant variables.
  A variable is *relevant* if it appears in the target, if it appears in a relevant variable,
  or if it is a proposition that refers to a relevant variable.
- `decide +kernel` uses kernel for reduction instead of the elaborator.
  It has two key properties: (1) since it uses the kernel, it ignores transparency and can unfold everything,
  and (2) it reduces the `Decidable` instance only once instead of twice.
- `decide +native` uses the native code compiler (`#eval`) to evaluate the `Decidable` instance,
  admitting the result via the `Lean.ofReduceBool` axiom.
  This can be significantly more efficient than using reduction, but it is at the cost of increasing the size
  of the trusted code base.
  Namely, it depends on the correctness of the Lean compiler and all definitions with an `@[implemented_by]` attribute.
  Like with `+kernel`, the `Decidable` instance is evaluated only once.

Limitation: In the default mode or `+kernel` mode, since `decide` uses reduction to evaluate the term,
`Decidable` instances defined by well-founded recursion might not work because evaluating them requires reducing proofs.
Reduction can also get stuck on `Decidable` instances with `Eq.rec` terms.
These can appear in instances defined using tactics (such as `rw` and `simp`).
To avoid this, create such instances using definitions such as `decidable_of_iff` instead.

## Examples

Proving inequalities:
```lean
example : 2 + 2 ≠ 5 := by decide
```

Trying to prove a false proposition:
```lean
example : 1 ≠ 1 := by decide
/-
tactic 'decide' proved that the proposition
  1 ≠ 1
is false
-/
```

Trying to prove a proposition whose `Decidable` instance fails to reduce
```lean
opaque unknownProp : Prop

open scoped Classical in
example : unknownProp := by decide
/-
tactic 'decide' failed for proposition
  unknownProp
since its 'Decidable' instance reduced to
  Classical.choice ⋯
rather than to the 'isTrue' constructor.
-/
```

## Properties and relations

For equality goals for types with decidable equality, usually `rfl` can be used in place of `decide`.
```lean
example : 1 + 1 = 2 := by decide
example : 1 + 1 = 2 := by rfl
```
-/
syntax (name := decide) "decide" optConfig : tactic
Decide tactic for proving decidable propositions
Informal description
The `decide` tactic attempts to prove a goal of type `p` by synthesizing an instance of `Decidable p` and evaluating it to determine the truth value of `p`. If the evaluation results in `isTrue h`, then `h` is used as a proof to close the goal. The goal must not contain local variables or metavariables. The tactic supports several options: - `+revert`: Reverts relevant local variables before attempting to decide. - `+kernel`: Uses kernel reduction for evaluation, ignoring transparency and unfolding all definitions. - `+native`: Uses the native code compiler for efficient evaluation, at the cost of trusting more code.
Lean.Parser.Tactic.nativeDecide definition
: Lean.ParserDescr✝
Full source
/--
`native_decide` is a synonym for `decide +native`.
It will attempt to prove a goal of type `p` by synthesizing an instance
of `Decidable p` and then evaluating it to `isTrue ..`. Unlike `decide`, this
uses `#eval` to evaluate the decidability instance.

This should be used with care because it adds the entire lean compiler to the trusted
part, and the axiom `Lean.ofReduceBool` will show up in `#print axioms` for theorems using
this method or anything that transitively depends on them. Nevertheless, because it is
compiled, this can be significantly more efficient than using `decide`, and for very
large computations this is one way to run external programs and trust the result.
```lean
example : (List.range 1000).length = 1000 := by native_decide
```
-/
syntax (name := nativeDecide) "native_decide" optConfig : tactic
Native decision procedure tactic
Informal description
The tactic `native_decide` attempts to prove a goal of type `p` by synthesizing a `Decidable p` instance and evaluating it to `isTrue ..`. Unlike `decide`, it uses `#eval` for evaluation, which is more efficient but adds the Lean compiler to the trusted codebase. This should be used with caution as it introduces the axiom `Lean.ofReduceBool` in the proof.
Lean.Parser.Tactic.omega definition
: Lean.ParserDescr✝
Full source
/--
The `omega` tactic, for resolving integer and natural linear arithmetic problems.

It is not yet a full decision procedure (no "dark" or "grey" shadows),
but should be effective on many problems.

We handle hypotheses of the form `x = y`, `x < y`, `x ≤ y`, and `k ∣ x` for `x y` in `Nat` or `Int`
(and `k` a literal), along with negations of these statements.

We decompose the sides of the inequalities as linear combinations of atoms.

If we encounter `x / k` or `x % k` for literal integers `k` we introduce new auxiliary variables
and the relevant inequalities.

On the first pass, we do not perform case splits on natural subtraction.
If `omega` fails, we recursively perform a case split on
a natural subtraction appearing in a hypothesis, and try again.

The options
```
omega +splitDisjunctions +splitNatSub +splitNatAbs +splitMinMax
```
can be used to:
* `splitDisjunctions`: split any disjunctions found in the context,
  if the problem is not otherwise solvable.
* `splitNatSub`: for each appearance of `((a - b : Nat) : Int)`, split on `a ≤ b` if necessary.
* `splitNatAbs`: for each appearance of `Int.natAbs a`, split on `0 ≤ a` if necessary.
* `splitMinMax`: for each occurrence of `min a b`, split on `min a b = a ∨ min a b = b`
Currently, all of these are on by default.
-/
syntax (name := omega) "omega" optConfig : tactic
Omega tactic for linear arithmetic
Informal description
The `omega` tactic is a partial decision procedure for solving linear arithmetic problems over the integers and natural numbers. It handles hypotheses of the form $x = y$, $x < y$, $x \leq y$, and $k \mid x$ (for $k$ a literal constant), along with their negations. The tactic decomposes inequalities into linear combinations of atoms and introduces auxiliary variables for division and modulus operations with literal denominators. By default, it performs case splits on natural subtraction, absolute value, and min/max operations when necessary, though these behaviors can be configured via options.
Lean.Parser.Tactic.tacticBv_omega definition
: Lean.ParserDescr✝
Full source
macromacro "bv_omega" : tactic => `(tactic| (try simp -implicitDefEqProofs only [bitvec_to_nat] at *) <;> omega)
Bit vector omega tactic
Informal description
The tactic `bv_omega` is an extension of the `omega` tactic that includes a preprocessor to convert statements about bit vectors (`BitVec`) into equivalent statements about natural numbers (`Nat`). The preprocessor applies the simplification rule `bitvec_to_nat` to all relevant hypotheses and goals before invoking the `omega` tactic. The `bitvec_to_nat` rule can be extended with additional theorems marked with the `@[simp]` attribute.
Lean.Parser.Tactic.acNf0 definition
: Lean.ParserDescr✝
Full source
/-- Implementation of `ac_nf` (the full `ac_nf` calls `trivial` afterwards). -/
syntax (name := acNf0) "ac_nf0" (location)? : tactic
AC normalization tactic (partial)
Informal description
The tactic `ac_nf0` is a partial implementation of the `ac_nf` tactic, which performs normalization modulo associativity and commutativity (AC) without the subsequent call to the `trivial` tactic that `ac_nf` includes. The optional `location` parameter specifies where the tactic should be applied, following the same syntax as other location-aware tactics.
Lean.Parser.Tactic.normCast0 definition
: Lean.ParserDescr✝
Full source
/-- Implementation of `norm_cast` (the full `norm_cast` calls `trivial` afterwards). -/
syntax (name := normCast0) "norm_cast0" optConfig (location)? : tactic
Partial coercion normalization tactic
Informal description
The tactic `norm_cast0` is a partial implementation of the `norm_cast` tactic, which normalizes expressions involving coercions but does not subsequently call the `trivial` tactic. The optional configuration and location parameters allow for customization of where and how the tactic is applied.
Lean.Parser.Tactic.tacticAssumption_mod_cast_ definition
: Lean.ParserDescr✝
Full source
macromacro "assumption_mod_cast" cfg:optConfig : tactic => `(tactic| norm_cast0 $cfg at * <;> assumption)
`assumption_mod_cast` tactic (cast-normalizing variant)
Informal description
The tactic `assumption_mod_cast` is a variant of the `assumption` tactic that solves the goal using a hypothesis. It first normalizes the goal and all local hypotheses by moving casts as far outward as possible using `norm_cast`, then attempts to close the goal with the modified hypotheses.
Lean.Parser.Tactic.tacticNorm_cast__ definition
: Lean.ParserDescr✝
Full source
macromacro "norm_cast" cfg:optConfig loc:(location)? : tactic =>
  `(tactic| norm_cast0 $cfg $[$loc]? <;> try trivial)
Normalization of coercions tactic (`norm_cast`)
Informal description
The `norm_cast` tactic normalizes coercions (casts) in expressions by moving them upwards. It can be applied to the goal or to specific hypotheses using the `at` syntax. The tactic is designed to be safe and efficient, with special handling for numerals. It is particularly useful for simplifying expressions involving casts between different numeric types.
Lean.Parser.Tactic.pushCast definition
: Lean.ParserDescr✝
Full source
/--
`push_cast` rewrites the goal to move certain coercions (*casts*) inward, toward the leaf nodes.
This uses `norm_cast` lemmas in the forward direction.
For example, `↑(a + b)` will be written to `↑a + ↑b`.
- `push_cast` moves casts inward in the goal.
- `push_cast at h` moves casts inward in the hypothesis `h`.
It can be used with extra simp lemmas with, for example, `push_cast [Int.add_zero]`.

Example:
```lean
example (a b : Nat)
    (h1 : ((a + b : Nat) : Int) = 10)
    (h2 : ((a + b + 0 : Nat) : Int) = 10) :
    ((a + b : Nat) : Int) = 10 := by
  /-
  h1 : ↑(a + b) = 10
  h2 : ↑(a + b + 0) = 10
  ⊢ ↑(a + b) = 10
  -/
  push_cast
  /- Now
  ⊢ ↑a + ↑b = 10
  -/
  push_cast at h1
  push_cast [Int.add_zero] at h2
  /- Now
  h1 h2 : ↑a + ↑b = 10
  -/
  exact h1
```

See also `norm_cast`.
-/
syntax (name := pushCast) "push_cast" optConfig (discharger)? (&" only")?
  (" [" (simpStar <|> simpErase <|> simpLemma),* "]")? (location)? : tactic
Push coercions inward tactic (`push_cast`)
Informal description
The `push_cast` tactic rewrites expressions to move coercions (casts) inward toward the leaf nodes of the expression, using `norm_cast` lemmas in the forward direction. For example, it rewrites $\uparrow(a + b)$ to $\uparrow a + \uparrow b$. The tactic can be applied to the goal or to specific hypotheses using the `at` syntax, and additional simplification lemmas can be provided in square brackets.
Lean.Parser.Tactic.tacticAc_nf_ definition
: Lean.ParserDescr✝
Full source
macromacro "ac_nf" loc:(location)? : tactic =>
  `(tactic| ac_nf0 $[$loc]? <;> try trivial)
Associative-Commutative Normal Form Tactic
Informal description
The tactic `ac_nf` normalizes equalities involving an associative and commutative operator by rewriting both sides to a common normal form. It can be applied to the goal target or specified hypotheses. For example, given an equality `a + b + c + d = d + (b + c) + a` where `+` is associative and commutative, `ac_nf` rewrites both sides to `a + (b + (c + d))`.
Lean.Parser.Tactic.symm definition
: Lean.ParserDescr✝
Full source
/--
* `symm` applies to a goal whose target has the form `t ~ u` where `~` is a symmetric relation,
  that is, a relation which has a symmetry lemma tagged with the attribute [symm].
  It replaces the target with `u ~ t`.
* `symm at h` will rewrite a hypothesis `h : t ~ u` to `h : u ~ t`.
-/
syntax (name := symm) "symm" (location)? : tactic
Symmetry tactic for symmetric relations
Informal description
The tactic `symm` applies to a goal whose target is of the form `t ~ u`, where `~` is a symmetric relation (i.e., a relation with a symmetry lemma tagged with the attribute `[symm]`). It replaces the target with `u ~ t`. When used as `symm at h`, it rewrites a hypothesis `h : t ~ u` to `h : u ~ t`.
Lean.Parser.Tactic.solveByElim definition
: Lean.ParserDescr✝
Full source
/--
`solve_by_elim` calls `apply` on the main goal to find an assumption whose head matches
and then repeatedly calls `apply` on the generated subgoals until no subgoals remain,
performing at most `maxDepth` (defaults to 6) recursive steps.

`solve_by_elim` discharges the current goal or fails.

`solve_by_elim` performs backtracking if subgoals can not be solved.

By default, the assumptions passed to `apply` are the local context, `rfl`, `trivial`,
`congrFun` and `congrArg`.

The assumptions can be modified with similar syntax as for `simp`:
* `solve_by_elim [h₁, h₂, ..., hᵣ]` also applies the given expressions.
* `solve_by_elim only [h₁, h₂, ..., hᵣ]` does not include the local context,
  `rfl`, `trivial`, `congrFun`, or `congrArg` unless they are explicitly included.
* `solve_by_elim [-h₁, ... -hₙ]` removes the given local hypotheses.
* `solve_by_elim using [a₁, ...]` uses all lemmas which have been labelled
  with the attributes `aᵢ` (these attributes must be created using `register_label_attr`).

`solve_by_elim*` tries to solve all goals together, using backtracking if a solution for one goal
makes other goals impossible.
(Adding or removing local hypotheses may not be well-behaved when starting with multiple goals.)

Optional arguments passed via a configuration argument as `solve_by_elim (config := { ... })`
- `maxDepth`: number of attempts at discharging generated subgoals
- `symm`: adds all hypotheses derived by `symm` (defaults to `true`).
- `exfalso`: allow calling `exfalso` and trying again if `solve_by_elim` fails
  (defaults to `true`).
- `transparency`: change the transparency mode when calling `apply`. Defaults to `.default`,
  but it is often useful to change to `.reducible`,
  so semireducible definitions will not be unfolded when trying to apply a lemma.

See also the doc-comment for `Lean.Meta.Tactic.Backtrack.BacktrackConfig` for the options
`proc`, `suspend`, and `discharge` which allow further customization of `solve_by_elim`.
Both `apply_assumption` and `apply_rules` are implemented via these hooks.
-/
syntax (name := solveByElim)
  "solve_by_elim" "*"? optConfig (&" only")? (args)? (using_)? : tactic
`solve_by_elim` tactic for goal solving via assumption application
Informal description
The `solve_by_elim` tactic attempts to solve the current goal by repeatedly applying assumptions from the local context, `rfl`, `trivial`, `congrFun`, and `congrArg`. It performs backtracking if subgoals cannot be solved and allows customization of the assumptions used, including adding or removing specific hypotheses, restricting to a given list of lemmas, or using lemmas labeled with specific attributes. The tactic can be configured with options such as maximum recursion depth, symmetry handling, and transparency mode for applying lemmas.
Lean.Parser.Tactic.applyAssumption definition
: Lean.ParserDescr✝
Full source
/--
`apply_assumption` looks for an assumption of the form `... → ∀ _, ... → head`
where `head` matches the current goal.

You can specify additional rules to apply using `apply_assumption [...]`.
By default `apply_assumption` will also try `rfl`, `trivial`, `congrFun`, and `congrArg`.
If you don't want these, or don't want to use all hypotheses, use `apply_assumption only [...]`.
You can use `apply_assumption [-h]` to omit a local hypothesis.
You can use `apply_assumption using [a₁, ...]` to use all lemmas which have been labelled
with the attributes `aᵢ` (these attributes must be created using `register_label_attr`).

`apply_assumption` will use consequences of local hypotheses obtained via `symm`.

If `apply_assumption` fails, it will call `exfalso` and try again.
Thus if there is an assumption of the form `P → ¬ Q`, the new tactic state
will have two goals, `P` and `Q`.

You can pass a further configuration via the syntax `apply_rules (config := {...}) lemmas`.
The options supported are the same as for `solve_by_elim` (and include all the options for `apply`).
-/
syntax (name := applyAssumption)
  "apply_assumption" optConfig (&" only")? (args)? (using_)? : tactic
`apply_assumption` tactic
Informal description
The tactic `apply_assumption` searches for an assumption of the form `... → ∀ _, ... → head` where `head` matches the current goal. It can be configured with additional rules to apply, and by default also tries `rfl`, `trivial`, `congrFun`, and `congrArg`. The tactic can be restricted to use only specified hypotheses or lemmas, and will use consequences of local hypotheses obtained via `symm`. If it fails, it calls `exfalso` and tries again, potentially splitting the goal into two if an assumption of the form `P → ¬ Q` is found.
Lean.Parser.Tactic.applyRules definition
: Lean.ParserDescr✝
Full source
/--
`apply_rules [l₁, l₂, ...]` tries to solve the main goal by iteratively
applying the list of lemmas `[l₁, l₂, ...]` or by applying a local hypothesis.
If `apply` generates new goals, `apply_rules` iteratively tries to solve those goals.
You can use `apply_rules [-h]` to omit a local hypothesis.

`apply_rules` will also use `rfl`, `trivial`, `congrFun` and `congrArg`.
These can be disabled, as can local hypotheses, by using `apply_rules only [...]`.

You can use `apply_rules using [a₁, ...]` to use all lemmas which have been labelled
with the attributes `aᵢ` (these attributes must be created using `register_label_attr`).

You can pass a further configuration via the syntax `apply_rules (config := {...})`.
The options supported are the same as for `solve_by_elim` (and include all the options for `apply`).

`apply_rules` will try calling `symm` on hypotheses and `exfalso` on the goal as needed.
This can be disabled with `apply_rules (config := {symm := false, exfalso := false})`.

You can bound the iteration depth using the syntax `apply_rules (config := {maxDepth := n})`.

Unlike `solve_by_elim`, `apply_rules` does not perform backtracking, and greedily applies
a lemma from the list until it gets stuck.
-/
syntax (name := applyRules) "apply_rules" optConfig (&" only")? (args)? (using_)? : tactic
Iterative Lemma Application Tactic (`apply_rules`)
Informal description
The tactic `apply_rules` attempts to solve the main goal by iteratively applying a list of lemmas `[l₁, l₂, ...]` or local hypotheses. If applying a lemma generates new subgoals, `apply_rules` continues to attempt solving those subgoals. The tactic can be configured to omit certain local hypotheses (`apply_rules [-h]`), disable built-in lemmas (`apply_rules only [...]`), or use lemmas labeled with specific attributes (`apply_rules using [a₁, ...]`). Additional configuration options, such as disabling symmetry application (`symm := false`) or bounding the iteration depth (`maxDepth := n`), can be passed via the syntax `apply_rules (config := {...})`. Unlike `solve_by_elim`, `apply_rules` does not perform backtracking and applies lemmas greedily until it gets stuck.
Lean.Parser.Tactic.exact? definition
: Lean.ParserDescr✝
Full source
/--
Searches environment for definitions or theorems that can solve the goal using `exact`
with conditions resolved by `solve_by_elim`.

The optional `using` clause provides identifiers in the local context that must be
used by `exact?` when closing the goal.  This is most useful if there are multiple
ways to resolve the goal, and one wants to guide which lemma is used.
-/
syntax (name := exact?) "exact?" (" using " (colGt ident),+)? : tactic
Interactive exact tactic with search
Informal description
The tactic `exact?` searches the environment for definitions or theorems that can solve the current goal using `exact`, with any necessary conditions resolved by `solve_by_elim`. An optional `using` clause can specify identifiers from the local context that must be used to close the goal, which is useful when there are multiple ways to resolve the goal and guidance is needed on which lemma to use.
Lean.Parser.Tactic.apply? definition
: Lean.ParserDescr✝
Full source
/--
Searches environment for definitions or theorems that can refine the goal using `apply`
with conditions resolved when possible with `solve_by_elim`.

The optional `using` clause provides identifiers in the local context that must be
used when closing the goal.
-/
syntax (name := apply?) "apply?" (" using " (colGt term),+)? : tactic
Interactive apply tactic with search
Informal description
The tactic `apply?` searches the environment for definitions or theorems that can refine the current goal using the `apply` tactic, attempting to resolve any conditions with `solve_by_elim` when possible. An optional `using` clause can specify identifiers from the local context that must be used to close the goal.
Lean.Parser.Tactic.rewrites? definition
: Lean.ParserDescr✝
Full source
/--
`rw?` tries to find a lemma which can rewrite the goal.

`rw?` should not be left in proofs; it is a search tool, like `apply?`.

Suggestions are printed as `rw [h]` or `rw [← h]`.

You can use `rw? [-my_lemma, -my_theorem]` to prevent `rw?` using the named lemmas.
-/
syntax (name := rewrites?) "rw?" (ppSpace location)? (rewrites_forbidden)? : tactic
Interactive rewrite search tactic
Informal description
The tactic `rw?` searches for lemmas that can rewrite the current goal, suggesting them in the form `rw [h]` or `rw [← h]`. It is intended as a search tool (like `apply?`) and should not remain in final proofs. Optionally, specific lemmas can be excluded from the search using the syntax `rw? [-my_lemma, -my_theorem]`.
Lean.Parser.Tactic.showTerm definition
: Lean.ParserDescr✝
Full source
/--
`show_term tac` runs `tac`, then prints the generated term in the form
"exact X Y Z" or "refine X ?_ Z" if there are remaining subgoals.

(For some tactics, the printed term will not be human readable.)
-/
syntax (name := showTerm) "show_term " tacticSeq : tactic
Show Term Tactic
Informal description
The tactic `show_term tac` executes the tactic sequence `tac` and then displays the generated term in the form `exact X Y Z` or `refine X ?_ Z` if there are remaining subgoals. Note that for some tactics, the printed term may not be human-readable.
Lean.Parser.Tactic.exposeNames definition
: Lean.ParserDescr✝
Full source
/--
`expose_names` renames all inaccessible variables with accessible names, making them available
for reference in generated tactics. However, this renaming introduces machine-generated names
that are not fully under user control. `expose_names` is primarily intended as a preamble for
auto-generated end-game tactic scripts. It is also useful as an alternative to
`set_option tactic.hygienic false`. If explicit control over renaming is needed in the
middle of a tactic script, consider using structured tactic scripts with
`match .. with`, `induction .. with`, or `intro` with explicit user-defined names,
as well as tactics such as `next`, `case`, and `rename_i`.
-/
syntax (name := exposeNames) "expose_names" : tactic
Variable exposure tactic (`expose_names`)
Informal description
The tactic `expose_names` renames all inaccessible variables with accessible names, making them available for reference in generated tactics. This is primarily intended as a preamble for auto-generated end-game tactic scripts and serves as an alternative to disabling tactic hygiene via `set_option tactic.hygienic false`. For more control over renaming in the middle of a tactic script, structured tactics like `match`, `induction`, or `intro` with explicit names should be used, along with tactics such as `next`, `case`, and `rename_i`.
Lean.Parser.Tactic.suggestPremises definition
: Lean.ParserDescr✝
Full source
/--
`#suggest_premises` will suggest premises for the current goal, using the currently registered premise selector.

The suggestions are printed in the order of their confidence, from highest to lowest.
-/
syntax (name := suggestPremises) "suggest_premises" : tactic
Premise suggestion tactic
Informal description
The `suggest_premises` tactic suggests relevant premises for the current goal, ordered by confidence from highest to lowest, using the currently registered premise selector.
Lean.Parser.Tactic.bvDecideMacro definition
: Lean.ParserDescr✝
Full source
macro (name := bvDecideMacro) (priority:=low) "bv_decide" optConfig : tactic =>
  Macro.throwError "to use `bv_decide`, please include `import Std.Tactic.BVDecide`"
Bit Vector and Boolean Decision Procedure (`bv_decide`)
Informal description
The macro `bv_decide` is a tactic that automatically solves goals involving fixed-width bit vectors (`BitVec`) and Boolean values (`Bool`) by querying an external SAT solver and verifying the proof within Lean. It handles goals equivalent to the `QF_BV` logic in SMT-LIB, and can decompose structures containing `BitVec` or `Bool` information. If the tactic encounters an unknown definition, it treats it as an unconstrained `BitVec` variable. If the goal cannot be solved, `bv_decide` provides a counterexample with assignments for all considered variables. The tactic supports proof caching via `bv_decide?` and can be configured with the `bv.ac_nf` option for problems requiring associativity or commutativity. Note that `bv_decide` relies on the correctness of the code generator through `ofReduceBool`.
Lean.Parser.Tactic.bvTraceMacro definition
: Lean.ParserDescr✝
Full source
macro (name := bvTraceMacro) (priority:=low) "bv_decide?" optConfig : tactic =>
  Macro.throwError "to use `bv_decide?`, please include `import Std.Tactic.BVDecide`"
Proof script suggestion for `bv_decide` tactic
Informal description
A macro that suggests a proof script for a `bv_decide` tactic call, useful for caching LRAT proofs. When invoked with `bv_decide?`, it checks if the `Std.Tactic.BVDecide` module is imported and provides guidance accordingly.
Lean.Parser.Tactic.bvNormalizeMacro definition
: Lean.ParserDescr✝
Full source
macro (name := bvNormalizeMacro) (priority:=low) "bv_normalize" optConfig : tactic =>
  Macro.throwError "to use `bv_normalize`, please include `import Std.Tactic.BVDecide`"
Bit-vector normalization macro
Informal description
The macro `bv_normalize` runs the normalization procedure of the `bv_decide` tactic, which can sometimes solve basic goals involving bit vectors (`BitVec`). **Note:** To use this macro, you must include `import Std.Tactic.BVDecide` in your file.
Lean.Parser.Attr.simp definition
: Lean.ParserDescr✝
Full source
/--
Theorems tagged with the `simp` attribute are used by the simplifier
(i.e., the `simp` tactic, and its variants) to simplify expressions occurring in your goals.
We call theorems tagged with the `simp` attribute "simp theorems" or "simp lemmas".
Lean maintains a database/index containing all active simp theorems.
Here is an example of a simp theorem.
```lean
@[simp] theorem ne_eq (a b : α) : (a ≠ b) = Not (a = b) := rfl
```
This simp theorem instructs the simplifier to replace instances of the term
`a ≠ b` (e.g. `x + 0 ≠ y`) with `Not (a = b)` (e.g., `Not (x + 0 = y)`).
The simplifier applies simp theorems in one direction only:
if `A = B` is a simp theorem, then `simp` replaces `A`s with `B`s,
but it doesn't replace `B`s with `A`s. Hence a simp theorem should have the
property that its right-hand side is "simpler" than its left-hand side.
In particular, `=` and `↔` should not be viewed as symmetric operators in this situation.
The following would be a terrible simp theorem (if it were even allowed):
```lean
@[simp] lemma mul_right_inv_bad (a : G) : 1 = a * a⁻¹ := ...
```
Replacing 1 with a * a⁻¹ is not a sensible default direction to travel.
Even worse would be a theorem that causes expressions to grow without bound,
causing simp to loop forever.

By default the simplifier applies `simp` theorems to an expression `e`
after its sub-expressions have been simplified.
We say it performs a bottom-up simplification.
You can instruct the simplifier to apply a theorem before its sub-expressions
have been simplified by using the modifier `↓`. Here is an example
```lean
@[simp↓] theorem not_and_eq (p q : Prop) : (¬ (p ∧ q)) = (¬p ∨ ¬q) :=
```

You can instruct the simplifier to rewrite the lemma from right-to-left:
```lean
attribute @[simp ←] and_assoc
```

When multiple simp theorems are applicable, the simplifier uses the one with highest priority.
The equational theorems of function are applied at very low priority (100 and below).
If there are several with the same priority, it is uses the "most recent one". Example:
```lean
@[simp high] theorem cond_true (a b : α) : cond true a b = a := rfl
@[simp low+1] theorem or_true (p : Prop) : (p ∨ True) = True :=
  propext <| Iff.intro (fun _ => trivial) (fun _ => Or.inr trivial)
@[simp 100] theorem ite_self {d : Decidable c} (a : α) : ite c a a = a := by
  cases d <;> rfl
```
-/
syntax (name := simp) "simp" (Tactic.simpPre <|> Tactic.simpPost)? patternIgnore("← " <|> "<- ")? (ppSpace prio)? : attr
Simplifier attribute for theorems
Informal description
The `simp` attribute is used to tag theorems that the simplifier (i.e., the `simp` tactic and its variants) can use to simplify expressions in goals. A theorem tagged with `@[simp]` is called a "simp theorem" or "simp lemma". The simplifier applies these theorems in one direction only: from left to right, replacing instances of the left-hand side with the right-hand side, which should be simpler. The attribute supports modifiers for controlling the application order (`↓` for pre-simplification, `←` for right-to-left rewriting) and priority levels to resolve conflicts when multiple simp theorems are applicable.
Lean.Parser.Attr.wf_preprocess definition
: Lean.ParserDescr✝
Full source
/--
Theorems tagged with the `wf_preprocess` attribute are used during the processing of functions defined
by well-founded recursion. They are applied to the function's body to add additional hypotheses,
such as replacing `if c then _ else _` with `if h : c then _ else _` or `xs.map` with
`xs.attach.map`. Also see `wfParam`.
-/
syntax (name := wf_preprocess) "wf_preprocess" (Tactic.simpPre <|> Tactic.simpPost)? patternIgnore("← " <|> "<- ")? (ppSpace prio)? : attr
Well-founded recursion preprocessing attribute
Informal description
The `wf_preprocess` attribute is used to tag theorems that assist in processing functions defined by well-founded recursion. These theorems are applied to the function's body to introduce additional hypotheses, such as converting `if c then _ else _` to `if h : c then _ else _` or transforming `xs.map` to `xs.attach.map`. The attribute can optionally include a simplification direction (pre or post) and a priority.
Lean.Parser.Attr.norm_cast definition
: Lean.ParserDescr✝
Full source
/--
The `norm_cast` attribute should be given to lemmas that describe the
behaviour of a coercion with respect to an operator, a relation, or a particular
function.

It only concerns equality or iff lemmas involving `↑`, `⇑` and `↥`, describing the behavior of
the coercion functions.
It does not apply to the explicit functions that define the coercions.

Examples:
```lean
@[norm_cast] theorem coe_nat_inj' {m n : ℕ} : (↑m : ℤ) = ↑n ↔ m = n

@[norm_cast] theorem coe_int_denom (n : ℤ) : (n : ℚ).denom = 1

@[norm_cast] theorem cast_id : ∀ n : ℚ, ↑n = n

@[norm_cast] theorem coe_nat_add (m n : ℕ) : (↑(m + n) : ℤ) = ↑m + ↑n

@[norm_cast] theorem cast_coe_nat (n : ℕ) : ((n : ℤ) : α) = n

@[norm_cast] theorem cast_one : ((1 : ℚ) : α) = 1
```

Lemmas tagged with `@[norm_cast]` are classified into three categories: `move`, `elim`, and
`squash`. They are classified roughly as follows:

* elim lemma:   LHS has 0 head coes and ≥ 1 internal coe
* move lemma:   LHS has 1 head coe and 0 internal coes,    RHS has 0 head coes and ≥ 1 internal coes
* squash lemma: LHS has ≥ 1 head coes and 0 internal coes, RHS has fewer head coes

`norm_cast` uses `move` and `elim` lemmas to factor coercions toward the root of an expression
and to cancel them from both sides of an equation or relation. It uses `squash` lemmas to clean
up the result.

It is typically not necessary to specify these categories, as `norm_cast` lemmas are
automatically classified by default. The automatic classification can be overridden by
giving an optional `elim`, `move`, or `squash` parameter to the attribute.

```lean
@[simp, norm_cast elim] lemma nat_cast_re (n : ℕ) : (n : ℂ).re = n := by
  rw [← of_real_nat_cast, of_real_re]
```

Don't do this unless you understand what you are doing.
-/
syntax (name := norm_cast) "norm_cast" (ppSpace normCastLabel)? (ppSpace num)? : attr
`norm_cast` attribute for coercion lemmas
Informal description
The `norm_cast` attribute is used to tag lemmas that describe the behavior of coercions with respect to operators, relations, or specific functions. These lemmas must involve equality or equivalence statements with coercion symbols (`↑`, `⇑`, `↥`) and describe how coercions interact with other operations. The attribute classifies such lemmas into three categories: `elim`, `move`, and `squash`, which are used by the `norm_cast` tactic to simplify expressions involving coercions by moving them toward the root of an expression, eliminating them, or cleaning up the result.
term‹_› definition
: Lean.ParserDescr✝
Full source
/--
`‹t›` resolves to an (arbitrary) hypothesis of type `t`.
It is useful for referring to hypotheses without accessible names.
`t` may contain holes that are solved by unification with the expected type;
in particular, `‹_›` is a shortcut for `by assumption`.
-/
syntax "‹" withoutPosition(term) "›" : term
Anonymous hypothesis reference
Informal description
The syntax `‹t›` resolves to an arbitrary hypothesis of type `t`, where `t` can be a term with holes filled by unification. In particular, `‹_›` is shorthand for `by assumption`, used to refer to unnamed hypotheses.
tacticGet_elem_tactic_trivial definition
: Lean.ParserDescr✝
Full source
/--
`get_elem_tactic_trivial` is an extensible tactic automatically called
by the notation `arr[i]` to prove any side conditions that arise when
constructing the term (e.g. the index is in bounds of the array).
The default behavior is to just try `trivial` (which handles the case
where `i < arr.size` is in the context) and `simp +arith` and `omega`
(for doing linear arithmetic in the index).
-/
syntax "get_elem_tactic_trivial" : tactic
Tactic for array access side conditions
Informal description
The tactic `get_elem_tactic_trivial` is automatically invoked by the array access notation `arr[i]` to handle side conditions such as proving that the index `i` is within the bounds of the array `arr`. By default, it attempts to solve these conditions using `trivial` (for cases where `i < arr.size` is already in the context), `simp` combined with arithmetic reasoning (`arith`), and the `omega` tactic for linear arithmetic on the index.
tacticGet_elem_tactic definition
: Lean.ParserDescr✝
Full source
macromacro "get_elem_tactic" : tactic =>
  `(tactic| first
      /-
      Recall that `macro_rules` (namely, for `get_elem_tactic_trivial`) are tried in reverse order.
      We first, however, try `done`, since the necessary proof may already have been
      found during unification, in which case there is no goal to solve (see #6999).
      If a goal is present, we want `assumption` to be tried first.
      This is important for theorems such as
      ```
      [simp] theorem getElem_pop (a : Array α) (i : Nat) (hi : i < a.pop.size) :
      a.pop[i] = a[i]'(Nat.lt_of_lt_of_le (a.size_pop ▸ hi) (Nat.sub_le _ _)) :=
      ```
      There is a proof embedded in the right-hand-side, and we want it to be just `hi`.
      If `omega` is used to "fill" this proof, we will have a more complex proof term that
      cannot be inferred by unification.
      We hardcoded `assumption` here to ensure users cannot accidentally break this IF
      they add new `macro_rules` for `get_elem_tactic_trivial`.

      TODO: Implement priorities for `macro_rules`.
      TODO: Ensure we have **high-priority** macro_rules for `get_elem_tactic_trivial` which are
            just `done` and `assumption`.
      -/
    | done
    | assumption
    | get_elem_tactic_trivial
    | fail "failed to prove index is valid, possible solutions:
  - Use `have`-expressions to prove the index is valid
  - Use `a[i]!` notation instead, runtime check is performed, and 'Panic' error message is produced if index is not valid
  - Use `a[i]?` notation instead, result is an `Option` type
  - Use `a[i]'h` notation instead, where `h` is a proof that index is valid"
   )
Automatic index validation tactic for array access
Informal description
The tactic `get_elem_tactic` is automatically invoked by the notation `arr[i]` to handle side conditions that arise when constructing the term (such as ensuring the index is within the bounds of the array). It first attempts to use the `done` tactic if the proof is already resolved during unification, then tries `assumption` for existing hypotheses, and finally delegates to `get_elem_tactic_trivial`. If all attempts fail, it provides a diagnostic error message suggesting alternative notations or explicit proofs.
autoParam abbrev
(α : Sort u) (tactic : Lean.Syntax) : Sort u
Full source
/--
  Gadget for automatic parameter support. This is similar to the `optParam` gadget, but it uses
  the given tactic.
  Like `optParam`, this gadget only affects elaboration.
  For example, the tactic will *not* be invoked during type class resolution. -/
abbrev autoParam.{u} (α : Sort u) (tactic : Lean.Syntax) : Sort u := α
Automatic Parameter Gadget with Tactic Support
Informal description
The `autoParam` gadget is used for automatic parameter support in Lean, similar to `optParam` but utilizing a specified tactic for elaboration. It affects only the elaboration process and does not invoke the tactic during type class resolution.