feat(library/algebra): add missing theorems to algebra library
This commit is contained in:
parent
4312f1e54b
commit
f5dcb1e0a9
5 changed files with 47 additions and 3 deletions
|
@ -541,6 +541,7 @@ section add_group
|
|||
|
||||
theorem add_eq_of_eq_sub {a b c : A} (H : a = c - b) : a + b = c :=
|
||||
add_eq_of_eq_add_neg H
|
||||
|
||||
end add_group
|
||||
|
||||
structure add_comm_group [class] (A : Type) extends add_group A, add_comm_monoid A
|
||||
|
@ -585,6 +586,8 @@ section add_comm_group
|
|||
theorem sub_eq_sub_add_sub (a b c : A) : a - b = c - b + (a - c) :=
|
||||
by rewrite [add_sub, sub_add_cancel] ⬝ !add.comm
|
||||
|
||||
theorem neg_neg_sub_neg (a b : A) : - (-a - -b) = a - b :=
|
||||
by rewrite [neg_sub, sub_neg_eq_add, neg_add_eq_sub]
|
||||
end add_comm_group
|
||||
|
||||
definition group_of_add_group (A : Type) [G : add_group A] : group A :=
|
||||
|
|
|
@ -290,6 +290,17 @@ section
|
|||
(assume H : ¬ a ≤ b,
|
||||
(inr (assume H1 : a = b, H (H1 ▸ !le.refl))))
|
||||
|
||||
theorem eq_or_lt_of_not_lt {a b : A} (H : ¬ a < b) : a = b ∨ b < a :=
|
||||
if Heq : a = b then or.inl Heq else or.inr (lt_of_not_ge (λ Hge, H (lt_of_le_of_ne Hge Heq)))
|
||||
|
||||
theorem eq_or_lt_of_le {a b : A} (H : a ≤ b) : a = b ∨ a < b :=
|
||||
begin
|
||||
cases eq_or_lt_of_not_lt (not_lt_of_ge H),
|
||||
exact or.inl a_1⁻¹,
|
||||
exact or.inr a_1
|
||||
end
|
||||
|
||||
|
||||
-- testing equality first may result in more definitional equalities
|
||||
definition lt.cases {B : Type} (a b : A) (t_lt t_eq t_gt : B) : B :=
|
||||
if a = b then t_eq else (if a < b then t_lt else t_gt)
|
||||
|
|
|
@ -493,6 +493,20 @@ section
|
|||
theorem sub_lt_right_of_lt_add : a < b + c → a - c < b :=
|
||||
iff.mp !lt_add_iff_sub_lt_right
|
||||
|
||||
theorem sub_lt_of_sub_lt : a - b < c → a - c < b :=
|
||||
begin
|
||||
intro H,
|
||||
apply sub_lt_left_of_lt_add,
|
||||
apply lt_add_of_sub_lt_right _ _ _ H
|
||||
end
|
||||
|
||||
theorem sub_le_of_sub_le : a - b ≤ c → a - c ≤ b :=
|
||||
begin
|
||||
intro H,
|
||||
apply sub_left_le_of_le_add,
|
||||
apply le_add_of_sub_right_le _ _ _ H
|
||||
end
|
||||
|
||||
-- TODO: the Isabelle library has varations on a + b ≤ b ↔ a ≤ 0
|
||||
theorem le_iff_le_of_sub_eq_sub {a b c d : A} (H : a - b = c - d) : a ≤ b ↔ c ≤ d :=
|
||||
calc
|
||||
|
@ -553,6 +567,9 @@ section
|
|||
theorem sub_le_of_nonneg (H : b ≥ 0) : a - b ≤ a :=
|
||||
add_le_of_le_of_nonpos (le.refl a) (neg_nonpos_of_nonneg H)
|
||||
|
||||
theorem sub_lt_of_pos {b : A} (H : b > 0) : a - b < a :=
|
||||
add_lt_of_le_of_neg (le.refl a) (neg_neg_of_pos H)
|
||||
|
||||
theorem neg_add_neg_le_neg_of_pos {a : A} (H : a > 0) : -a + -a ≤ -a :=
|
||||
!neg_add ▸ neg_le_neg (le_add_of_nonneg_left (le_of_lt H))
|
||||
end
|
||||
|
|
|
@ -672,6 +672,19 @@ section
|
|||
theorem sub_le_of_abs_sub_le_right (H : abs (a - b) ≤ c) : a - c ≤ b :=
|
||||
sub_le_of_abs_sub_le_left (!abs_sub ▸ H)
|
||||
|
||||
theorem sub_lt_of_abs_sub_lt_left (H : abs (a - b) < c) : b - c < a :=
|
||||
if Hz : 0 ≤ a - b then
|
||||
(calc
|
||||
a ≥ b : (iff.mp !sub_nonneg_iff_le) Hz
|
||||
... > b - c : sub_lt_of_pos _ (lt_of_le_of_lt !abs_nonneg H))
|
||||
else
|
||||
(have Habs : b - a < c, by rewrite [abs_of_neg (lt_of_not_ge Hz) at H, neg_sub at H]; apply H,
|
||||
have Habs' : b < c + a, from lt_add_of_sub_lt_right _ _ _ Habs,
|
||||
sub_lt_left_of_lt_add _ _ _ Habs')
|
||||
|
||||
theorem sub_lt_of_abs_sub_lt_right (H : abs (a - b) < c) : a - c < b :=
|
||||
sub_lt_of_abs_sub_lt_left (!abs_sub ▸ H)
|
||||
|
||||
theorem abs_sub_square (a b : A) : abs (a - b) * abs (a - b) = a * a + b * b - (1 + 1) * a * b :=
|
||||
by rewrite [abs_mul_abs_self, *mul_sub_left_distrib, *mul_sub_right_distrib,
|
||||
sub_add_eq_sub_sub, sub_neg_eq_add, *right_distrib, sub_add_eq_sub_sub, *one_mul,
|
||||
|
|
|
@ -7,9 +7,9 @@ This construction follows Bishop and Bridges (1985).
|
|||
|
||||
At this point, we no longer proceed constructively: this file makes heavy use of decidability,
|
||||
excluded middle, and Hilbert choice. Two sets of definitions of Cauchy sequences, convergence,
|
||||
etc are available, one with rates and one without. The definitions with rates are amenable
|
||||
to be used constructively, if and when that development takes place. The second set of
|
||||
definitions are the usual classical ones.
|
||||
etc are available in the libray, one with rates and one without. The definitions here, with rates,
|
||||
are amenable to be used constructively if and when that development takes place. The second set of
|
||||
definitions available in /library/theories/analysis/metric_space.lean are the usual classical ones.
|
||||
|
||||
Here, we show that ℝ is complete. The proofs of Cauchy completeness and the supremum property
|
||||
are independent of each other.
|
||||
|
|
Loading…
Reference in a new issue