feat(library/algebra/ordered_field): move identity about abs to ordered_field

This commit is contained in:
Rob Lewis 2015-06-16 14:06:57 +10:00 committed by Leonardo de Moura
parent 090f00458d
commit ff0ba6687e
2 changed files with 14 additions and 4 deletions

View file

@ -421,5 +421,16 @@ section discrete_linear_ordered_field
exact Hb, exact H
end
theorem abs_one_div : abs (1 / a) = 1 / abs a :=
if H : a > 0 then
by rewrite [abs_of_pos H, abs_of_pos (div_pos_of_pos H)]
else
(if H' : a < 0 then
by rewrite [abs_of_neg H', abs_of_neg (div_neg_of_neg H'),
-(one_div_neg_eq_neg_one_div (ne_of_lt H'))]
else
have Heq [visible] : a = 0, from eq_of_le_of_ge (le_of_not_gt H) (le_of_not_gt H'),
by rewrite [Heq, div_zero, *abs_zero, div_zero])
end discrete_linear_ordered_field
end algebra

View file

@ -25,7 +25,9 @@ namespace s
-- helper lemmas
theorem abs_sub_square (a b : ) : abs (a - b) * abs (a - b) = a * a + b * b - (1 + 1) * a * b :=
sorry --begin rewrite [abs_mul_self, *rat.left_distrib, *rat.right_distrib, *one_mul] end
by rewrite [abs_mul_self, *rat.mul_sub_left_distrib, *rat.mul_sub_right_distrib,
sub_add_eq_sub_sub, sub_neg_eq_add, *rat.right_distrib, sub_add_eq_sub_sub, *one_mul,
*add.assoc, {_ + b * b}add.comm, {_ + (b * b + _)}add.comm, mul.comm b a, *add.assoc]
theorem neg_add_rewrite {a b : } : a + -b = -(b + -a) := sorry
@ -42,9 +44,6 @@ theorem abs_abs_sub_abs_le_abs_sub (a b : ) : abs (abs a - abs b) ≤ abs (a
apply trivial
end
theorem abs_one_div (q : ) : abs (1 / q) = 1 / abs q := sorry
-- does this not exist already??
theorem forall_of_not_exists {A : Type} {P : A → Prop} (H : ¬ ∃ a : A, P a) : ∀ a : A, ¬ P a :=
take a, assume Ha, H (exists.intro a Ha)