Case bash on variables in finite intervals #
This file provides the tactic interval_cases. interval_cases n will:
- inspect hypotheses looking for lower and upper bounds of the form a ≤ nora < nandn < born ≤ b, including the bound0 ≤ nforn : ℕautomatically.
- call fin_caseson the synthesised hypothesisn ∈ Set.Ico a b, assuming an appropriateFintypeinstance can be found for the type ofn.
Currently, n must be of type ℕ or ℤ (TODO: generalize).
You can also explicitly specify a lower and upper bound to use, as interval_cases using hl hu,
where the hypotheses should be of the form hl : a ≤ n and hu : n < b. In that case,
interval_cases calls fin_cases on the resulting hypothesis h : n ∈ Set.Ico a b.
The result of interval_cases is a list of goals,
one for each integer value between the bounds.
- rhs : Lean.ExprThe target expression, a numeral in the input type 
- value : ℤThe numeric value of the target expression 
- goal : Lean.MVarIdThe new subgoal, of the form ⊢ x = rhs → tgt
Instances For
A Bound represents the result of analyzing a lower or upper bound expression.
If e is the scrutinee expression, then a lower bound expression like 3 < e
is normalized to ¬e ≤ 3 and represented as .lt 3, and an upper bound expression
like e ≤ 5 is represented as .le 5.
- lt: ℤ → Mathlib.Tactic.IntervalCases.BoundA strictly less-than lower bound n ≱ eor upper bounde ≱ n. (interval_casesuses less-equal exclusively, so less-than bounds are actually written as not-less-equal with flipped arguments.)
- le: ℤ → Mathlib.Tactic.IntervalCases.BoundA less-than-or-equal lower bound n ≤ eor upper bounde ≤ n.
Instances For
Assuming Bound represents a lower bound, this returns the (inclusive)
least integer value which is allowed. So 3 ≤ e means the lower bound is 3 and
3 < e means the lower bound is 4.
Equations
- x.asLower = match x with | Mathlib.Tactic.IntervalCases.Bound.lt n => n + 1 | Mathlib.Tactic.IntervalCases.Bound.le n => n
Instances For
Assuming Bound represents an upper bound, this returns the (inclusive)
greatest integer value which is allowed. So e ≤ 3 means the lower bound is 3 and
e < 3 means the upper bound is 2. Note that in the case of e < 0 on Nat
the upper bound is -1, which is not representable as a Nat;
this is why we have to treat the .lt and .le cases separately instead of normalizing
everything to .le bounds.
Equations
- x.asUpper = match x with | Mathlib.Tactic.IntervalCases.Bound.lt n => n - 1 | Mathlib.Tactic.IntervalCases.Bound.le n => n
Instances For
Given a type ty (the type of a hypothesis in the context or a provided expression),
attempt to parse it as an inequality, and return (a, b, strict, positive), where
positive means it is a negated inequality and strict means it is a strict inequality
(a < b or a ≱ b). a is always the lesser argument and b the greater one.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A "typeclass" (not actually a class) of methods for the type-specific handling of
interval_cases. To add support for a new type, you have to implement this interface and add
a dispatch case for it in intervalCases.
- initLB : Lean.Expr → Lean.MetaM (Mathlib.Tactic.IntervalCases.Bound × Lean.Expr × Lean.Expr)Given e, construct(bound, n, p)wherepis a proof ofn ≤ eorn < e(characterized bybound), orfailureif the type is not lower-bounded.
- initUB : Lean.Expr → Lean.MetaM (Mathlib.Tactic.IntervalCases.Bound × Lean.Expr × Lean.Expr)Given e, construct(bound, n, p)wherepis a proof ofe ≤ nore < n(characterized bybound), orfailureif the type is not upper-bounded.
- proveLE : Lean.Expr → Lean.Expr → Lean.MetaM Lean.ExprGiven a, b, provea ≤ bor fail.
- proveLT : Lean.Expr → Lean.Expr → Lean.MetaM Lean.ExprGiven a, b, provea ≱ bor fail.
- Given - a, b, a', pwhere- pproves- a ≱ band- a' := a+1, prove- a' ≤ b.
- Given - a, b, b', pwhere- pproves- a ≱ band- b' := b-1, prove- a ≤ b'.
- Given - e, return- (z, n, p)where- p : e = nand- nis a numeral appropriate for the type denoting the integer- z.
- mkNumeral : ℤ → Lean.MetaM Lean.ExprConstruct the canonical numeral for integer z, or fail ifzis out of range.
Instances For
Given a proof pf, attempts to parse it as an upper (lb = false) or lower (lb = true)
bound on n. If successful, it returns (bound, n, pf') where n is a numeral and
pf' proves n ≤ e or n ≱ e (as described by bound).
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given (z1, e1, p1) a lower bound on e and (z2, e2, p2) an upper bound on e,
such that the distance between the bounds is negative, returns a proof of False.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given (z1, e1, p1) a lower bound on e and (z2, e2, p2) an upper bound on e, such that the
distance between the bounds matches the number of cases in the subarray (which must be positive),
proves the goal g using the metavariables in the array by recursive bisection.
This is the core of the tactic, producing a case tree of if statements which bottoms out
at the cases.
A Methods implementation for ℕ.
This tells interval_cases how to work on natural numbers.
Instances For
A Methods implementation for ℤ.
This tells interval_cases how to work on integers.
Instances For
intervalCases proves goal g by splitting into cases for each integer between the given bounds.
Parameters:
- g: the goal, which can have any type- ⊢ tgt(it works in both proofs and programs)
- e: the scrutinee, the expression we are proving is bounded between integers
- e': a version of- eused for error messages. (This is used by the- interval_casesfrontend tactic because it uses a fresh variable for- e, so it is more helpful to show the pre-generalized expression in error messages.)
- lbs: A list of candidate lower bound expressions. The tactic will automatically pick the best lower bound it can find from the list.
- ubs: A list of candidate upper bound expressions. The tactic will automatically pick the best upper bound it can find from the list.
- mustUseBounds: If true, the tactic will fail if it is unable to parse any of the given- ubsor- lbsinto bounds. If false (the default), these will be silently skipped and an error message is only produced if we could not find any bounds (including those supplied by the type itself, e.g. if we are working over- Nator- Fin n).
Returns an array of IntervalCasesSubgoal, one per subgoal. A subgoal has the following fields:
- rhs: the numeral expression for this case
- value: the integral value of- rhs
- goal: the subgoal of type- ⊢ e = rhs → tgt
Note that this tactic does not perform any substitution or introduction steps -
all subgoals are in the same context as goal itself.
Equations
- One or more equations did not get rendered due to their size.
Instances For
interval_cases n searches for upper and lower bounds on a variable n,
and if bounds are found,
splits into separate cases for each possible value of n.
As an example, in
example (n : ℕ) (w₁ : n ≥ 3) (w₂ : n < 5) : n = 3 ∨ n = 4 := by
  interval_cases n
  all_goals simp
after interval_cases n, the goals are 3 = 3 ∨ 3 = 4 and 4 = 3 ∨ 4 = 4.
You can also explicitly specify a lower and upper bound to use,
as interval_cases using hl, hu.
The hypotheses should be in the form hl : a ≤ n and hu : n < b,
in which case interval_cases calls fin_cases on the resulting fact n ∈ Set.Ico a b.
You can specify a name h for the new hypothesis,
as interval_cases h : n or interval_cases h : n using hl, hu.
Equations
- One or more equations did not get rendered due to their size.