2014-12-20 19:05:28 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: Floris van Doorn, Leonardo de Moura
|
|
|
|
|
-/
|
|
|
|
|
prelude
|
2015-04-24 23:58:50 +00:00
|
|
|
|
import init.wf init.tactic init.hedberg init.util init.types
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-05 21:25:35 +00:00
|
|
|
|
open eq decidable sum lift
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
|
|
|
|
namespace nat
|
|
|
|
|
notation `ℕ` := nat
|
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
/- basic definitions on natural numbers -/
|
|
|
|
|
inductive le (a : ℕ) : ℕ → Type₀ :=
|
|
|
|
|
| refl : le a a
|
|
|
|
|
| step : Π {b}, le a b → le a (succ b)
|
|
|
|
|
|
|
|
|
|
infix `≤` := le
|
|
|
|
|
attribute le.refl [refl]
|
|
|
|
|
|
|
|
|
|
definition lt [reducible] (n m : ℕ) := succ n ≤ m
|
|
|
|
|
definition ge [reducible] (n m : ℕ) := m ≤ n
|
|
|
|
|
definition gt [reducible] (n m : ℕ) := succ m ≤ n
|
|
|
|
|
infix `<` := lt
|
|
|
|
|
infix `≥` := ge
|
|
|
|
|
infix `>` := gt
|
|
|
|
|
|
|
|
|
|
definition pred [unfold-c 1] (a : nat) : nat :=
|
|
|
|
|
nat.cases_on a zero (λ a₁, a₁)
|
|
|
|
|
|
|
|
|
|
-- add is defined in init.num
|
2015-06-04 23:16:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
definition sub (a b : nat) : nat :=
|
|
|
|
|
nat.rec_on b a (λ b₁ r, pred r)
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
definition mul (a b : nat) : nat :=
|
|
|
|
|
nat.rec_on b zero (λ b₁ r, r + a)
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
notation a - b := sub a b
|
|
|
|
|
notation a * b := mul a b
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
/- properties of ℕ -/
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
|
|
|
|
protected definition is_inhabited [instance] : inhabited nat :=
|
|
|
|
|
inhabited.mk zero
|
|
|
|
|
|
2015-05-05 21:25:35 +00:00
|
|
|
|
protected definition has_decidable_eq [instance] : ∀ x y : nat, decidable (x = y)
|
|
|
|
|
| has_decidable_eq zero zero := inl rfl
|
|
|
|
|
| has_decidable_eq (succ x) zero := inr (by contradiction)
|
|
|
|
|
| has_decidable_eq zero (succ y) := inr (by contradiction)
|
|
|
|
|
| has_decidable_eq (succ x) (succ y) :=
|
|
|
|
|
match has_decidable_eq x y with
|
|
|
|
|
| inl xeqy := inl (by rewrite xeqy)
|
|
|
|
|
| inr xney := inr (λ h : succ x = succ y, by injection h with xeqy; exact absurd xeqy xney)
|
|
|
|
|
end
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
/- properties of inequality -/
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_of_eq {n m : ℕ} (p : n = m) : n ≤ m := p ▸ le.refl n
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_succ (n : ℕ) : n ≤ succ n := by repeat constructor
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem pred_le (n : ℕ) : pred n ≤ n := by cases n;all_goals (repeat constructor)
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le.trans [trans] {n m k : ℕ} (H1 : n ≤ m) (H2 : m ≤ k) : n ≤ k :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by induction H2 with n H2 IH;exact H1;exact le.step IH
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_succ_of_le {n m : ℕ} (H : n ≤ m) : n ≤ succ m := le.trans H !le_succ
|
2015-05-29 18:50:19 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_of_succ_le {n m : ℕ} (H : succ n ≤ m) : n ≤ m := le.trans !le_succ H
|
2015-05-29 18:50:19 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_of_lt {n m : ℕ} (H : n < m) : n ≤ m := le_of_succ_le H
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem succ_le_succ [unfold-c 3] {n m : ℕ} (H : n ≤ m) : succ n ≤ succ m :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by induction H;reflexivity;exact le.step v_0
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem pred_le_pred [unfold-c 3] {n m : ℕ} (H : n ≤ m) : pred n ≤ pred m :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by induction H;reflexivity;cases b;exact v_0;exact le.step v_0
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_of_succ_le_succ [unfold-c 3] {n m : ℕ} (H : succ n ≤ succ m) : n ≤ m :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
pred_le_pred H
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_succ_of_pred_le [unfold-c 1] {n m : ℕ} (H : pred n ≤ m) : n ≤ succ m :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by cases n;exact le.step H;exact succ_le_succ H
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem not_succ_le_self {n : ℕ} : ¬succ n ≤ n :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by induction n with n IH;all_goals intros;cases a;apply IH;exact le_of_succ_le_succ a
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem zero_le (n : ℕ) : 0 ≤ n :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by induction n with n IH;apply le.refl;exact le.step IH
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem lt.step {n m : ℕ} (H : n < m) : n < succ m :=
|
|
|
|
|
le.step H
|
|
|
|
|
|
|
|
|
|
theorem zero_lt_succ (n : ℕ) : 0 < succ n :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by induction n with n IH;apply le.refl;exact le.step IH
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem lt.trans [trans] {n m k : ℕ} (H1 : n < m) (H2 : m < k) : n < k :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
le.trans (le.step H1) H2
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem lt_of_le_of_lt [trans] {n m k : ℕ} (H1 : n ≤ m) (H2 : m < k) : n < k :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
le.trans (succ_le_succ H1) H2
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem lt_of_lt_of_le [trans] {n m k : ℕ} (H1 : n < m) (H2 : m ≤ k) : n < k :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
le.trans H1 H2
|
|
|
|
|
|
|
|
|
|
theorem le.antisymm {n m : ℕ} (H1 : n ≤ m) (H2 : m ≤ n) : n = m :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
begin
|
2015-05-29 18:50:19 +00:00
|
|
|
|
cases H1 with m' H1',
|
|
|
|
|
{ reflexivity},
|
|
|
|
|
{ cases H2 with n' H2',
|
|
|
|
|
{ reflexivity},
|
|
|
|
|
{ exfalso, apply not_succ_le_self, exact lt.trans H1' H2'}},
|
2014-12-20 19:05:28 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem not_succ_le_zero (n : ℕ) : ¬succ n ≤ zero :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by intro H; cases H
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
theorem lt.irrefl (n : ℕ) : ¬n < n := not_succ_le_self
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
theorem self_lt_succ (n : ℕ) : n < succ n := !le.refl
|
|
|
|
|
theorem lt.base (n : ℕ) : n < succ n := !le.refl
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
theorem le_lt_antisymm {n m : ℕ} (H1 : n ≤ m) (H2 : m < n) : empty :=
|
|
|
|
|
!lt.irrefl (lt_of_le_of_lt H1 H2)
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
theorem lt_le_antisymm {n m : ℕ} (H1 : n < m) (H2 : m ≤ n) : empty :=
|
|
|
|
|
le_lt_antisymm H2 H1
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
theorem lt.asymm {n m : ℕ} (H1 : n < m) (H2 : m < n) : empty :=
|
|
|
|
|
le_lt_antisymm (le_of_lt H1) H2
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
definition lt.by_cases {a b : ℕ} {P : Type} (H1 : a < b → P) (H2 : a = b → P) (H3 : b < a → P) : P :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
begin
|
2015-05-29 18:50:19 +00:00
|
|
|
|
revert b H1 H2 H3, induction a with a IH,
|
|
|
|
|
{ intros, cases b,
|
|
|
|
|
exact H2 idp,
|
|
|
|
|
exact H1 !zero_lt_succ},
|
|
|
|
|
{ intros, cases b with b,
|
|
|
|
|
exact H3 !zero_lt_succ,
|
|
|
|
|
{ apply IH,
|
|
|
|
|
intro H, exact H1 (succ_le_succ H),
|
|
|
|
|
intro H, exact H2 (ap succ H),
|
|
|
|
|
intro H, exact H3 (succ_le_succ H)}}
|
2014-12-20 19:05:28 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem lt.trichotomy (a b : ℕ) : a < b ⊎ a = b ⊎ b < a :=
|
|
|
|
|
lt.by_cases inl (λH, inr (inl H)) (λH, inr (inr H))
|
2015-05-29 18:50:19 +00:00
|
|
|
|
|
|
|
|
|
definition lt_ge_by_cases {a b : ℕ} {P : Type} (H1 : a < b → P) (H2 : a ≥ b → P) : P :=
|
2015-06-04 23:16:28 +00:00
|
|
|
|
lt.by_cases H1 (λH, H2 (le_of_eq H⁻¹)) (λH, H2 (le_of_lt H))
|
|
|
|
|
|
|
|
|
|
theorem lt_or_ge (a b : ℕ) : (a < b) ⊎ (a ≥ b) :=
|
|
|
|
|
lt_ge_by_cases inl inr
|
2015-05-29 18:50:19 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem not_lt_zero (a : ℕ) : ¬ a < zero :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by intro H; cases H
|
|
|
|
|
|
|
|
|
|
-- less-than is well-founded
|
|
|
|
|
definition lt.wf [instance] : well_founded lt :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
begin
|
2015-05-29 18:50:19 +00:00
|
|
|
|
constructor, intro n, induction n with n IH,
|
|
|
|
|
{ constructor, intros n H, exfalso, exact !not_lt_zero H},
|
|
|
|
|
{ constructor, intros m H,
|
|
|
|
|
assert aux : ∀ {n₁} (hlt : m < n₁), succ n = n₁ → acc lt m,
|
|
|
|
|
{ intros n₁ hlt, induction hlt,
|
|
|
|
|
{ intro p, injection p with q, exact q ▸ IH},
|
|
|
|
|
{ intro p, injection p with q, exact (acc.inv (q ▸ IH) a)}},
|
|
|
|
|
apply aux H idp},
|
2014-12-20 19:05:28 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
definition measure {A : Type} (f : A → ℕ) : A → A → Type₀ :=
|
|
|
|
|
inv_image lt f
|
|
|
|
|
|
|
|
|
|
definition measure.wf {A : Type} (f : A → ℕ) : well_founded (measure f) :=
|
|
|
|
|
inv_image.wf f lt.wf
|
|
|
|
|
|
|
|
|
|
theorem succ_lt_succ {a b : ℕ} (H : a < b) : succ a < succ b :=
|
|
|
|
|
succ_le_succ H
|
|
|
|
|
|
|
|
|
|
theorem lt_of_succ_lt {a b : ℕ} (H : succ a < b) : a < b :=
|
|
|
|
|
le_of_succ_le H
|
|
|
|
|
|
|
|
|
|
theorem lt_of_succ_lt_succ {a b : ℕ} (H : succ a < succ b) : a < b :=
|
|
|
|
|
le_of_succ_le_succ H
|
|
|
|
|
|
|
|
|
|
definition decidable_le [instance] : decidable_rel le :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
begin
|
2015-05-29 18:50:19 +00:00
|
|
|
|
intros n, induction n with n IH,
|
|
|
|
|
{ intro m, left, apply zero_le},
|
|
|
|
|
{ intro m, cases m with m,
|
|
|
|
|
{ right, apply not_succ_le_zero},
|
|
|
|
|
{ let H := IH m, clear IH,
|
|
|
|
|
cases H with H H,
|
|
|
|
|
left, exact succ_le_succ H,
|
|
|
|
|
right, intro H2, exact H (le_of_succ_le_succ H2)}}
|
2014-12-20 19:05:28 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
definition decidable_lt [instance] : decidable_rel lt := _
|
|
|
|
|
definition decidable_gt [instance] : decidable_rel gt := _
|
|
|
|
|
definition decidable_ge [instance] : decidable_rel ge := _
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem eq_or_lt_of_le {a b : ℕ} (H : a ≤ b) : a = b ⊎ a < b :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by cases H with b' H; exact sum.inl rfl; exact sum.inr (succ_le_succ H)
|
|
|
|
|
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_of_eq_or_lt {a b : ℕ} (H : a = b ⊎ a < b) : a ≤ b :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by cases H with H H; exact le_of_eq H; exact le_of_lt H
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem eq_or_lt_of_not_lt {a b : ℕ} (hnlt : ¬ a < b) : a = b ⊎ b < a :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
sum.rec_on (lt.trichotomy a b)
|
|
|
|
|
(λ hlt, absurd hlt hnlt)
|
|
|
|
|
(λ h, h)
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem lt_succ_of_le {a b : ℕ} (h : a ≤ b) : a < succ b :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
succ_le_succ h
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem lt_of_succ_le {a b : ℕ} (h : succ a ≤ b) : a < b := h
|
2015-05-29 18:50:19 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem succ_le_of_lt {a b : ℕ} (h : a < b) : succ a ≤ b := h
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-05-29 18:50:19 +00:00
|
|
|
|
definition max (a b : ℕ) : ℕ := if a < b then b else a
|
|
|
|
|
definition min (a b : ℕ) : ℕ := if a < b then a else b
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem max_self (a : ℕ) : max a a = a :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
eq.rec_on !if_t_t rfl
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem max_eq_right {a b : ℕ} (H : a < b) : max a b = b :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
if_pos H
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem max_eq_left {a b : ℕ} (H : ¬ a < b) : max a b = a :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
if_neg H
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem eq_max_right {a b : ℕ} (H : a < b) : b = max a b :=
|
2015-05-05 21:25:35 +00:00
|
|
|
|
eq.rec_on (max_eq_right H) rfl
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem eq_max_left {a b : ℕ} (H : ¬ a < b) : a = max a b :=
|
2015-05-05 21:25:35 +00:00
|
|
|
|
eq.rec_on (max_eq_left H) rfl
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_max_left (a b : ℕ) : a ≤ max a b :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
by_cases
|
2015-05-05 21:25:35 +00:00
|
|
|
|
(λ h : a < b, le_of_lt (eq.rec_on (eq_max_right h) h))
|
|
|
|
|
(λ h : ¬ a < b, eq.rec_on (eq_max_left h) !le.refl)
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem le_max_right (a b : ℕ) : b ≤ max a b :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
by_cases
|
2015-05-05 21:25:35 +00:00
|
|
|
|
(λ h : a < b, eq.rec_on (eq_max_right h) !le.refl)
|
2014-12-20 19:05:28 +00:00
|
|
|
|
(λ h : ¬ a < b, sum.rec_on (eq_or_lt_of_not_lt h)
|
2015-05-05 21:25:35 +00:00
|
|
|
|
(λ heq, eq.rec_on heq (eq.rec_on (inverse (max_self a)) !le.refl))
|
2014-12-20 19:05:28 +00:00
|
|
|
|
(λ h : b < a,
|
2015-05-05 21:25:35 +00:00
|
|
|
|
have aux : a = max a b, from eq_max_left (lt.asymm h),
|
|
|
|
|
eq.rec_on aux (le_of_lt h)))
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem succ_sub_succ_eq_sub (a b : ℕ) : succ a - succ b = a - b :=
|
2015-05-29 18:50:19 +00:00
|
|
|
|
by induction b with b IH; reflexivity; apply ap pred IH
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem sub_eq_succ_sub_succ (a b : ℕ) : a - b = succ a - succ b :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
eq.rec_on (succ_sub_succ_eq_sub a b) rfl
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem zero_sub_eq_zero (a : ℕ) : zero - a = zero :=
|
2015-02-11 20:49:27 +00:00
|
|
|
|
nat.rec_on a
|
2014-12-20 19:05:28 +00:00
|
|
|
|
rfl
|
2015-05-05 21:25:35 +00:00
|
|
|
|
(λ a₁ (ih : zero - a₁ = zero), ap pred ih)
|
2014-12-20 19:05:28 +00:00
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem zero_eq_zero_sub (a : ℕ) : zero = zero - a :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
eq.rec_on (zero_sub_eq_zero a) rfl
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem sub_lt {a b : ℕ} : zero < a → zero < b → a - b < a :=
|
2014-12-20 19:05:28 +00:00
|
|
|
|
have aux : Π {a}, zero < a → Π {b}, zero < b → a - b < a, from
|
2015-05-29 18:50:19 +00:00
|
|
|
|
λa h₁, le.rec_on h₁
|
|
|
|
|
(λb h₂, le.cases_on h₂
|
2014-12-20 19:05:28 +00:00
|
|
|
|
(lt.base zero)
|
|
|
|
|
(λ b₁ bpos,
|
|
|
|
|
eq.rec_on (sub_eq_succ_sub_succ zero b₁)
|
|
|
|
|
(eq.rec_on (zero_eq_zero_sub b₁) (lt.base zero))))
|
2015-05-29 18:50:19 +00:00
|
|
|
|
(λa₁ apos ih b h₂, le.cases_on h₂
|
2014-12-20 19:05:28 +00:00
|
|
|
|
(lt.base a₁)
|
|
|
|
|
(λ b₁ bpos,
|
|
|
|
|
eq.rec_on (sub_eq_succ_sub_succ a₁ b₁)
|
|
|
|
|
(lt.trans (@ih b₁ bpos) (lt.base a₁)))),
|
|
|
|
|
λ h₁ h₂, aux h₁ h₂
|
|
|
|
|
|
2015-06-04 23:16:28 +00:00
|
|
|
|
theorem sub_le (a b : ℕ) : a - b ≤ a :=
|
2015-02-11 20:49:27 +00:00
|
|
|
|
nat.rec_on b
|
2014-12-20 19:05:28 +00:00
|
|
|
|
(le.refl a)
|
|
|
|
|
(λ b₁ ih, le.trans !pred_le ih)
|
|
|
|
|
|
2015-06-15 11:38:41 +00:00
|
|
|
|
lemma sub_lt_succ (a b : ℕ) : a - b < succ a := lt_succ_of_le (sub_le a b)
|
2014-12-20 19:05:28 +00:00
|
|
|
|
end nat
|