A tactic for canceling numeric denominators
This file defines tactics that cancel numeric denominators from field expressions.
As an example, we want to transform a comparison 5*(a/3 + b/4) < c/3
into the equivalent
5*(4*a + 3*b) < 4*c
.
Implementation notes
The tooling here was originally written for linarith
, not intended as an interactive tactic.
The interactive version has been split off because it is sometimes convenient to use on its own.
There are likely some rough edges to it.
Improving this tactic would be a good project for someone interested in learning tactic programming.
Lemmas used in the procedure
Computing cancelation factors
cancel_denominators_in_type h
assumes that h
is of the form lhs R rhs
,
where R ∈ {<, ≤, =, ≥, >}
.
It produces an expression h'
of the form lhs' R rhs'
and a proof that h = h'
.
Numeric denominators have been canceled in lhs'
and rhs'
.
Interactive version
cancel_denoms
attempts to remove numerals from the denominators of fractions.
It works on propositions that are field-valued inequalities.
variables {α : Type} [linear_ordered_field α] (a b c : α)
example (h : a / 5 + b / 4 < c) : 4*a + 5*b < 20*c :=
begin
cancel_denoms at h,
exact h
end
example (h : a > 0) : a / 5 > 0 :=
begin
cancel_denoms,
exact h
end