Documentation

Mathlib.Tactic.FunProp.Elab

funProp tactic syntax #

fun_prop config elaborator

Equations
  • One or more equations did not get rendered due to their size.
Instances For

    fun_prop solves a goal of the form P f, where P is a predicate and f is a function, by decomposing f into a composition of elementary functions, and proving P on each of those by matching against a set of @[fun_prop] lemmas.

    If fun_prop fails to solve a goal with the error "No theorems found", you can solve this issue by importing or adding new theorems tagged with the @[fun_prop] attribute. See the module documentation for Mathlib/Tactic/FunProp.lean for a detailed explanation.

    • fun_prop (disch := tac) uses tac to solve potential side goals. Setting this option is required to solve ContinuousAt/On/Within goals.
    • fun_prop [c, ...] will unfold the constant(s) c, ... before decomposing f.
    • fun_prop (config := cfg) sets advanced configuration options using cfg : FunProp.Config (see FunProp.Config for details).

    Examples:

    example : Continuous (fun x : ℝ ↦ x * sin x) := by fun_prop
    
    -- Specify a discharger to solve `ContinuousAt`/`Within`/`On` goals:
    example (y : ℝ) (hy : y ≠ 0) : ContinuousAt (fun x : ℝ ↦ 1/x) y := by
      fun_prop (disch := assumption)
    
    example (y : ℝ) (hy : y ≠ 0) : ContinuousAt (fun x => x * (Real.log x) ^ 2 - Real.exp x / x) y := by
      fun_prop (disch := aesop)
    
    Equations
    • One or more equations did not get rendered due to their size.
    Instances For

      Tactic to prove function properties

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        Command that printins all function properties attached to a function.

        For example

        #print_fun_prop_theorems HAdd.hAdd
        

        might print out

        Continuous
          continuous_add, args: [4,5], priority: 1000
          continuous_add_left, args: [5], priority: 1000
          continuous_add_right, args [4], priority: 1000
          ...
        Differentiable
          Differentiable.add, args: [4,5], priority: 1000
          Differentiable.add_const, args: [4], priority: 1000
          Differentiable.const_add, args: [5], priority: 1000
          ...
        

        You can also see only theorems about a concrete function property

        #print_fun_prop_theorems HAdd.hAdd Continuous
        
        Equations
        • One or more equations did not get rendered due to their size.
        Instances For