feat(library/algebra/order.lean, data/int/{basic,order}.lean): add theorem, correct gt_trans

This commit is contained in:
Jeremy Avigad 2015-04-15 22:25:02 -04:00 committed by Leonardo de Moura
parent 53919699bc
commit 5eb7fef564
3 changed files with 14 additions and 10 deletions

View file

@ -87,10 +87,13 @@ section
theorem gt.trans {a b c : A} (H1 : a > b) (H2: b > c) : a > c := lt.trans H2 H1
calc_trans gt.trans
theorem ne_of_lt {a b : A} : a < b → a ≠ b :=
assume lt_ab : a < b, assume eq_ab : a = b,
theorem ne_of_lt {a b : A} (lt_ab : a < b) : a ≠ b :=
assume eq_ab : a = b,
show false, from lt.irrefl b (eq_ab ▸ lt_ab)
theorem ne_of_gt {a b : A} (gt_ab : a > b) : a ≠ b :=
ne.symm (ne_of_lt gt_ab)
theorem lt.asymm {a b : A} (H : a < b) : ¬ b < a :=
assume H1 : b < a, lt.irrefl _ (lt.trans H H1)
end

View file

@ -150,19 +150,19 @@ int.cases_on a
/- int is a quotient of ordered pairs of natural numbers -/
definition equiv (p q : × ) : Prop := pr1 p + pr2 q = pr2 p + pr1 q
protected definition equiv (p q : × ) : Prop := pr1 p + pr2 q = pr2 p + pr1 q
local notation p `≡` q := equiv p q
theorem equiv.refl {p : × } : p ≡ p := !add.comm
protected theorem equiv.refl {p : × } : p ≡ p := !add.comm
theorem equiv.symm {p q : × } (H : p ≡ q) : q ≡ p :=
protected theorem equiv.symm {p q : × } (H : p ≡ q) : q ≡ p :=
calc
pr1 q + pr2 p = pr2 p + pr1 q : !add.comm
... = pr1 p + pr2 q : H⁻¹
... = pr2 q + pr1 p : !add.comm
theorem equiv.trans {p q r : × } (H1 : p ≡ q) (H2 : q ≡ r) : p ≡ r :=
protected theorem equiv.trans {p q r : × } (H1 : p ≡ q) (H2 : q ≡ r) : p ≡ r :=
have H3 : pr1 p + pr2 r + pr2 q = pr2 p + pr1 r + pr2 q, from
calc
pr1 p + pr2 r + pr2 q = pr1 p + pr2 q + pr2 r : by simp
@ -172,10 +172,10 @@ have H3 : pr1 p + pr2 r + pr2 q = pr2 p + pr1 r + pr2 q, from
... = pr2 p + pr1 r + pr2 q : by simp,
show pr1 p + pr2 r = pr2 p + pr1 r, from add.cancel_right H3
theorem equiv_equiv : is_equivalence equiv :=
protected theorem equiv_equiv : is_equivalence equiv :=
is_equivalence.mk @equiv.refl @equiv.symm @equiv.trans
theorem equiv_cases {p q : × } (H : equiv p q) :
protected theorem equiv_cases {p q : × } (H : equiv p q) :
(pr1 p ≥ pr2 p ∧ pr1 q ≥ pr2 q) (pr1 p < pr2 p ∧ pr1 q < pr2 q) :=
or.elim (@le_or_gt (pr2 p) (pr1 p))
(assume H1: pr1 p ≥ pr2 p,
@ -185,7 +185,7 @@ or.elim (@le_or_gt (pr2 p) (pr1 p))
have H2 : pr2 p + pr1 q < pr2 p + pr2 q, from H ▸ add_lt_add_right H1 (pr2 q),
or.inr (and.intro H1 (lt_of_add_lt_add_left H2)))
theorem equiv_of_eq {p q : × } (H : p = q) : p ≡ q := H ▸ equiv.refl
protected theorem equiv_of_eq {p q : × } (H : p = q) : p ≡ q := H ▸ equiv.refl
calc_trans equiv.trans
calc_refl equiv.refl

View file

@ -259,7 +259,8 @@ section port_algebra
show decidable (b < a), from _
theorem ge.trans: ∀{a b c : }, a ≥ b → b ≥ c → a ≥ c := @algebra.ge.trans _ _
theorem gt.trans: ∀{a b c : }, a ≥ b → b ≥ c → a ≥ c := @algebra.ge.trans _ _
theorem gt.trans: ∀{a b c : }, a > b → b > c → a > c := @algebra.gt.trans _ _
theorem ne_of_gt: ∀{a b : }, a > b → a ≠ b := @algebra.ne_of_gt _ _
theorem gt_of_gt_of_ge : ∀{a b c : }, a > b → b ≥ c → a > c := @algebra.gt_of_gt_of_ge _ _
theorem gt_of_ge_of_gt : ∀{a b c : }, a ≥ b → b > c → a > c := @algebra.gt_of_ge_of_gt _ _
theorem lt.asymm : ∀{a b : }, a < b → ¬ b < a := @algebra.lt.asymm _ _