refactor(library/init/nat,library/data/nat/*): chagne dots to underscores in protected names

This commit is contained in:
Jeremy Avigad 2015-10-22 17:35:27 -04:00 committed by Leonardo de Moura
parent 6dfaf1863c
commit da5bd03656
14 changed files with 66 additions and 67 deletions

View file

@ -121,7 +121,7 @@ theorem pow_gt_one {x : A} {i : } (xgt1 : x > 1) (ipos : i > 0) : x^i > 1 :=
assert xpos : x > 0, from lt.trans zero_lt_one xgt1,
begin
induction i with [i, ih],
{exfalso, exact !nat.lt.irrefl ipos},
{exfalso, exact !lt.irrefl ipos},
have xige1 : x^i ≥ 1, from pow_ge_one _ (le_of_lt xgt1),
rewrite [pow_succ, -mul_one 1],
apply mul_lt_mul xgt1 xige1 zero_lt_one,

View file

@ -218,7 +218,7 @@ private lemma enle.trans (a b c : A) : enle a b → enle b c → enle a c :=
assume h₁ h₂, le.trans h₁ h₂
private lemma enle.total (a b : A) : enle a b enle b a :=
le.total
!le.total
private lemma enle.antisymm (a b : A) : enle a b → enle b a → a = b :=
assume h₁ h₂,

View file

@ -462,7 +462,7 @@ theorem repr_mul : Π (a b : ), repr (a * b) = pmul (repr a) (repr b)
theorem equiv_mul_prep {xa ya xb yb xn yn xm ym : }
(H1 : xa + yb = ya + xb) (H2 : xn + ym = yn + xm)
: xa*xn+ya*yn+(xb*ym+yb*xm) = xa*yn+ya*xn+(xb*xm+yb*ym) :=
nat.add.right_cancel (calc
nat.add_right_cancel (calc
xa*xn+ya*yn + (xb*ym+yb*xm) + (yb*xn+xb*yn + (xb*xn+yb*yn))
= xa*xn+ya*yn + (yb*xn+xb*yn) + (xb*ym+yb*xm + (xb*xn+yb*yn)) : by rewrite add.comm4
... = xa*xn+ya*yn + (yb*xn+xb*yn) + (xb*xn+yb*yn + (xb*ym+yb*xm)) : by rewrite {xb*ym+yb*xm +_}nat.add_comm

View file

@ -396,7 +396,7 @@ theorem exists_least_of_bdd {P : → Prop} [HP : decidable_pred P]
let Hk' := algebra.not_le_of_gt Hk,
rewrite Hzbk,
apply λ p, mt (ge_least_of_lt _ p) Hk',
apply nat.lt.trans Hk,
apply nat.lt_trans Hk,
apply least_lt _ !lt_succ_self H'
end
@ -434,7 +434,7 @@ theorem exists_greatest_of_bdd {P : → Prop} [HP : decidable_pred P]
let Hk' := algebra.not_le_of_gt Hk,
rewrite Hzbk,
apply λ p, mt (ge_least_of_lt _ p) Hk',
apply nat.lt.trans Hk,
apply nat.lt_trans Hk,
apply least_lt _ !lt_succ_self H'
end

View file

@ -473,7 +473,7 @@ theorem nodup_upto : ∀ n, nodup (upto n)
| (n+1) :=
have d : nodup (upto n), from nodup_upto n,
have n : n ∉ upto n, from
assume i : n ∈ upto n, absurd (of_mem_of_all i (upto_less n)) (lt.irrefl n),
assume i : n ∈ upto n, absurd (of_mem_of_all i (upto_less n)) (nat.lt_irrefl n),
nodup_cons n d
theorem lt_of_mem_upto {n i : nat} : i ∈ upto n → i < n :=

View file

@ -145,13 +145,13 @@ nat.induction_on k
... = n + succ (m + l) : add_succ
... = n + (m + succ l) : add_succ)
protected theorem add.left_comm : Π (n m k : ), n + (m + k) = m + (n + k) :=
protected theorem add_left_comm : Π (n m k : ), n + (m + k) = m + (n + k) :=
left_comm nat.add_comm nat.add_assoc
protected theorem add.right_comm : Π (n m k : ), n + m + k = n + k + m :=
protected theorem add_right_comm : Π (n m k : ), n + m + k = n + k + m :=
right_comm nat.add_comm nat.add_assoc
protected theorem add.left_cancel {n m k : } : n + m = n + k → m = k :=
protected theorem add_left_cancel {n m k : } : n + m = n + k → m = k :=
nat.induction_on n
(take H : 0 + m = 0 + k,
!nat.zero_add⁻¹ ⬝ H ⬝ !nat.zero_add)
@ -164,9 +164,9 @@ nat.induction_on n
have n + m = n + k, from succ.inj this,
IH this)
protected theorem add.right_cancel {n m k : } (H : n + m = k + m) : n = k :=
protected theorem add_right_cancel {n m k : } (H : n + m = k + m) : n = k :=
have H2 : m + n = m + k, from !nat.add_comm ⬝ H ⬝ !nat.add_comm,
add.left_cancel H2
nat.add_left_cancel H2
theorem eq_zero_of_add_eq_zero_right {n m : } : n + m = 0 → n = 0 :=
nat.induction_on n
@ -236,7 +236,7 @@ nat.induction_on k
(n + m) * succ l = (n + m) * l + (n + m) : mul_succ
... = n * l + m * l + (n + m) : IH
... = n * l + m * l + n + m : nat.add_assoc
... = n * l + n + m * l + m : nat.add.right_comm
... = n * l + n + m * l + m : nat.add_right_comm
... = n * l + n + (m * l + m) : nat.add_assoc
... = n * succ l + (m * l + m) : mul_succ
... = n * succ l + m * succ l : mul_succ)

View file

@ -26,11 +26,11 @@ or_resolve_right (nat.eq_or_lt_of_le H1)
protected theorem lt_iff_le_and_ne (m n : ) : m < n ↔ m ≤ n ∧ m ≠ n :=
iff.intro
(take H, and.intro (nat.le_of_lt H) (take H1, !lt.irrefl (H1 ▸ H)))
(take H, and.intro (nat.le_of_lt H) (take H1, !nat.lt_irrefl (H1 ▸ H)))
(and.rec nat.lt_of_le_and_ne)
theorem le_add_right (n k : ) : n ≤ n + k :=
nat.rec !le.refl (λ k, le_succ_of_le) k
nat.rec !nat.le_refl (λ k, le_succ_of_le) k
theorem le_add_left (n m : ): n ≤ m + n :=
!add.comm ▸ !le_add_right
@ -42,7 +42,7 @@ theorem le.elim {n m : } : n ≤ m → ∃ k, n + k = m :=
le.rec (exists.intro 0 rfl) (λm h, Exists.rec
(λ k H, exists.intro (succ k) (H ▸ rfl)))
protected theorem le.total {m n : } : m ≤ n n ≤ m :=
protected theorem le_total {m n : } : m ≤ n n ≤ m :=
or.imp_left nat.le_of_lt !nat.lt_or_ge
/- addition -/
@ -54,11 +54,11 @@ protected theorem add_le_add_right {n m : } (H : n ≤ m) (k : ) : n + k
!add.comm ▸ !add.comm ▸ nat.add_le_add_left H k
protected theorem le_of_add_le_add_left {k n m : } (H : k + n ≤ k + m) : n ≤ m :=
obtain l Hl, from le.elim H, le.intro (add.left_cancel (!add.assoc⁻¹ ⬝ Hl))
obtain l Hl, from le.elim H, le.intro (nat.add_left_cancel (!add.assoc⁻¹ ⬝ Hl))
protected theorem lt_of_add_lt_add_left {k n m : } (H : k + n < k + m) : n < m :=
let H' := nat.le_of_lt H in
nat.lt_of_le_and_ne (nat.le_of_add_le_add_left H') (assume Heq, !lt.irrefl (Heq ▸ H))
nat.lt_of_le_and_ne (nat.le_of_add_le_add_left H') (assume Heq, !nat.lt_irrefl (Heq ▸ H))
protected theorem add_lt_add_left {n m : } (H : n < m) (k : ) : k + n < k + m :=
lt_of_succ_le (!add_succ ▸ nat.add_le_add_left (succ_le_of_lt H) k)
@ -80,7 +80,7 @@ theorem mul_le_mul_right {n m : } (k : ) (H : n ≤ m) : n * k ≤ m * k :
!mul.comm ▸ !mul.comm ▸ !mul_le_mul_left H
protected theorem mul_le_mul {n m k l : } (H1 : n ≤ k) (H2 : m ≤ l) : n * m ≤ k * l :=
le.trans (!nat.mul_le_mul_right H1) (!nat.mul_le_mul_left H2)
nat.le_trans (!nat.mul_le_mul_right H1) (!nat.mul_le_mul_left H2)
protected theorem mul_lt_mul_of_pos_left {n m k : } (H : n < m) (Hk : k > 0) : k * n < k * m :=
nat.lt_of_lt_of_le (nat.lt_add_of_pos_right Hk) (!mul_succ ▸ nat.mul_le_mul_left k (succ_le_of_lt H))
@ -95,17 +95,17 @@ open algebra
protected definition decidable_linear_ordered_semiring [reducible] [trans_instance] :
algebra.decidable_linear_ordered_semiring nat :=
⦃ algebra.decidable_linear_ordered_semiring, nat.comm_semiring,
add_left_cancel := @nat.add.left_cancel,
add_right_cancel := @nat.add.right_cancel,
add_left_cancel := @nat.add_left_cancel,
add_right_cancel := @nat.add_right_cancel,
lt := nat.lt,
le := nat.le,
le_refl := nat.le.refl,
le_trans := @nat.le.trans,
le_antisymm := @nat.le.antisymm,
le_total := @nat.le.total,
le_refl := nat.le_refl,
le_trans := @nat.le_trans,
le_antisymm := @nat.le_antisymm,
le_total := @nat.le_total,
le_iff_lt_or_eq := @nat.le_iff_lt_or_eq,
le_of_lt := @nat.le_of_lt,
lt_irrefl := @nat.lt.irrefl,
lt_irrefl := @nat.lt_irrefl,
lt_of_lt_of_le := @nat.lt_of_lt_of_le,
lt_of_le_of_lt := @nat.lt_of_le_of_lt,
lt_of_add_lt_add_left := @nat.lt_of_add_lt_add_left,
@ -119,7 +119,6 @@ algebra.decidable_linear_ordered_semiring nat :=
mul_lt_mul_of_pos_right := @nat.mul_lt_mul_of_pos_right,
decidable_lt := nat.decidable_lt ⦄
definition nat_has_dvd [reducible] [instance] [priority nat.prio] : has_dvd nat :=
has_dvd.mk has_dvd.dvd

View file

@ -61,11 +61,11 @@ theorem pow_cancel_left : ∀ {a b c : nat}, a > 1 → a^b = a^c → b = c
| a (succ b) 0 h₁ h₂ :=
assert a = 1, by rewrite [pow_succ at h₂, pow_zero at h₂]; exact (eq_one_of_mul_eq_one_right h₂),
assert 1 < 1, by rewrite [this at h₁]; exact h₁,
absurd `1 < 1` !lt.irrefl
absurd `1 <[nat] 1` !lt.irrefl
| a 0 (succ c) h₁ h₂ :=
assert a = 1, by rewrite [pow_succ at h₂, pow_zero at h₂]; exact (eq_one_of_mul_eq_one_right (eq.symm h₂)),
assert 1 < 1, by rewrite [this at h₁]; exact h₁,
absurd `1 < 1` !lt.irrefl
absurd `1 <[nat] 1` !lt.irrefl
| a (succ b) (succ c) h₁ h₂ :=
assert a ≠ 0, from assume aeq0, by rewrite [aeq0 at h₁]; exact (absurd h₁ dec_trivial),
assert a^b = a^c, by rewrite [*pow_succ at h₂]; exact (eq_of_mul_eq_mul_left (pos_of_ne_zero this) h₂),

View file

@ -70,7 +70,7 @@ definition pnat_lt_decidable [instance] {p q : +} : decidable (p < q) :=
begin rewrite lt.def, exact nat.decidable_lt p~ q~ end
theorem le.trans {p q r : +} : p ≤ q → q ≤ r → p ≤ r :=
begin rewrite +le.def, apply le.trans end
begin rewrite +le.def, apply nat.le_trans end
definition max (p q : +) : + :=
tag (max p~ q~) (lt_of_lt_of_le (!pnat_pos) (!le_max_right))
@ -189,7 +189,7 @@ theorem mul_lt_mul_left {a b c : +} (H : a < b) : a * c < b * c :=
begin rewrite [lt.def at *], exact mul_lt_mul_of_pos_right H !pnat_pos end
theorem one_lt_two : pone < 2 :=
!nat.le.refl
!nat.le_refl
theorem inv_two_mul_lt_inv (n : +) : (2 * n)⁻¹ < n⁻¹ :=
begin

View file

@ -40,13 +40,13 @@ namespace nat
/- basic definitions on natural numbers -/
inductive le (a : ) : → Prop :=
| refl' : le a a -- use refl' to avoid overloading le.refl
| nat_refl : le a a -- use nat_refl to avoid overloading le.refl
| step : Π {b}, le a b → le a (succ b)
definition nat_has_le [instance] [reducible] [priority nat.prio]: has_le nat := has_le.mk nat.le
protected lemma le.refl [refl] : ∀ a : nat, a ≤ a :=
le.refl'
protected lemma le_refl [refl] : ∀ a : nat, a ≤ a :=
le.nat_refl
protected definition lt [reducible] (n m : ) := succ n ≤ m
definition nat_has_lt [instance] [reducible] [priority nat.prio] : has_lt nat := has_lt.mk nat.lt
@ -85,9 +85,9 @@ namespace nat
/- properties of inequality -/
protected theorem le_of_eq {n m : } (p : n = m) : n ≤ m := p ▸ !le.refl
protected theorem le_of_eq {n m : } (p : n = m) : n ≤ m := p ▸ !nat.le_refl
theorem le_succ (n : ) : n ≤ succ n := le.step !le.refl
theorem le_succ (n : ) : n ≤ succ n := le.step !nat.le_refl
theorem pred_le (n : ) : pred n ≤ n := by cases n;repeat constructor
@ -97,20 +97,20 @@ namespace nat
theorem pred_le_iff_true [simp] (n : ) : pred n ≤ n ↔ true :=
iff_true_intro (pred_le n)
protected theorem le.trans {n m k : } (H1 : n ≤ m) : m ≤ k → n ≤ k :=
protected theorem le_trans {n m k : } (H1 : n ≤ m) : m ≤ k → n ≤ k :=
le.rec H1 (λp H2, le.step)
theorem le_succ_of_le {n m : } (H : n ≤ m) : n ≤ succ m := le.trans H !le_succ
theorem le_succ_of_le {n m : } (H : n ≤ m) : n ≤ succ m := nat.le_trans H !le_succ
theorem le_of_succ_le {n m : } (H : succ n ≤ m) : n ≤ m := le.trans !le_succ H
theorem le_of_succ_le {n m : } (H : succ n ≤ m) : n ≤ m := nat.le_trans !le_succ H
protected theorem le_of_lt {n m : } (H : n < m) : n ≤ m := le_of_succ_le H
theorem succ_le_succ {n m : } : n ≤ m → succ n ≤ succ m :=
le.rec !le.refl (λa b, le.step)
le.rec !nat.le_refl (λa b, le.step)
theorem pred_le_pred {n m : } : n ≤ m → pred n ≤ pred m :=
le.rec !le.refl (nat.rec (λa b, b) (λa b c, le.step))
le.rec !nat.le_refl (nat.rec (λa b, b) (λa b c, le.step))
theorem le_of_succ_le_succ {n m : } : succ n ≤ succ m → n ≤ m :=
pred_le_pred
@ -131,7 +131,7 @@ namespace nat
iff_false_intro not_succ_le_self
theorem zero_le : ∀ (n : ), 0 ≤ n :=
nat.rec !le.refl (λa, le.step)
nat.rec !nat.le_refl (λa, le.step)
theorem zero_le_iff_true [simp] (n : ) : 0 ≤ n ↔ true :=
iff_true_intro !zero_le
@ -144,33 +144,33 @@ namespace nat
theorem zero_lt_succ_iff_true [simp] (n : ) : 0 < succ n ↔ true :=
iff_true_intro (zero_lt_succ n)
protected theorem lt.trans {n m k : } (H1 : n < m) : m < k → n < k :=
le.trans (le.step H1)
protected theorem lt_trans {n m k : } (H1 : n < m) : m < k → n < k :=
nat.le_trans (le.step H1)
protected theorem lt_of_le_of_lt {n m k : } (H1 : n ≤ m) : m < k → n < k :=
le.trans (succ_le_succ H1)
nat.le_trans (succ_le_succ H1)
protected theorem lt_of_lt_of_le {n m k : } : n < m → m ≤ k → n < k := le.trans
protected theorem lt_of_lt_of_le {n m k : } : n < m → m ≤ k → n < k := nat.le_trans
protected theorem lt.irrefl (n : ) : ¬n < n := not_succ_le_self
protected theorem lt_irrefl (n : ) : ¬n < n := not_succ_le_self
theorem self_lt_succ (n : ) : n < succ n := !le.refl
theorem self_lt_succ (n : ) : n < succ n := !nat.le_refl
theorem self_lt_succ_iff_true [simp] (n : ) : n < succ n ↔ true :=
iff_true_intro (self_lt_succ n)
theorem lt.base (n : ) : n < succ n := !le.refl
theorem lt.base (n : ) : n < succ n := !nat.le_refl
theorem le_lt_antisymm {n m : } (H1 : n ≤ m) (H2 : m < n) : false :=
!lt.irrefl (nat.lt_of_le_of_lt H1 H2)
!nat.lt_irrefl (nat.lt_of_le_of_lt H1 H2)
protected theorem le.antisymm {n m : } (H1 : n ≤ m) : m ≤ n → n = m :=
le.cases_on H1 (λa, rfl) (λa b c, absurd (nat.lt_of_le_of_lt b c) !lt.irrefl)
protected theorem le_antisymm {n m : } (H1 : n ≤ m) : m ≤ n → n = m :=
le.cases_on H1 (λa, rfl) (λa b c, absurd (nat.lt_of_le_of_lt b c) !nat.lt_irrefl)
theorem lt_le_antisymm {n m : } (H1 : n < m) (H2 : m ≤ n) : false :=
le_lt_antisymm H2 H1
protected theorem lt.asymm {n m : } (H1 : n < m) : ¬ m < n :=
protected theorem nat.lt_asymm {n m : } (H1 : n < m) : ¬ m < n :=
le_lt_antisymm (nat.le_of_lt H1)
theorem not_lt_zero (a : ) : ¬ a < 0 := !not_succ_le_zero
@ -219,21 +219,21 @@ namespace nat
protected theorem lt_or_ge (a b : ) : a < b a ≥ b :=
nat.rec (inr !zero_le) (λn, or.rec
(λh, inl (le_succ_of_le h))
(λh, elim (nat.eq_or_lt_of_le h) (λe, inl (eq.subst e !le.refl)) inr)) b
(λh, elim (nat.eq_or_lt_of_le h) (λe, inl (eq.subst e !nat.le_refl)) inr)) b
protected definition lt_ge_by_cases {a b : } {P : Type} (H1 : a < b → P) (H2 : a ≥ b → P) : P :=
by_cases H1 (λh, H2 (elim !nat.lt_or_ge (λa, absurd a h) (λa, a)))
protected definition lt.by_cases {a b : } {P : Type} (H1 : a < b → P) (H2 : a = b → P)
protected definition lt_by_cases {a b : } {P : Type} (H1 : a < b → P) (H2 : a = b → P)
(H3 : b < a → P) : P :=
nat.lt_ge_by_cases H1 (λh₁,
nat.lt_ge_by_cases H3 (λh₂, H2 (le.antisymm h₂ h₁)))
nat.lt_ge_by_cases H3 (λh₂, H2 (nat.le_antisymm h₂ h₁)))
protected theorem lt.trichotomy (a b : ) : a < b a = b b < a :=
lt.by_cases (λH, inl H) (λH, inr (inl H)) (λH, inr (inr H))
protected theorem lt_trichotomy (a b : ) : a < b a = b b < a :=
nat.lt_by_cases (λH, inl H) (λH, inr (inl H)) (λH, inr (inr H))
protected theorem eq_or_lt_of_not_lt {a b : } (hnlt : ¬ a < b) : a = b b < a :=
or.rec_on (lt.trichotomy a b)
or.rec_on (nat.lt_trichotomy a b)
(λ hlt, absurd hlt hnlt)
(λ h, h)
@ -257,14 +257,14 @@ namespace nat
eq.symm !zero_sub_eq_zero
theorem sub_le (a b : ) : a - b ≤ a :=
nat.rec_on b !le.refl (λ b₁, le.trans !pred_le)
nat.rec_on b !nat.le_refl (λ b₁, nat.le_trans !pred_le)
theorem sub_le_iff_true [simp] (a b : ) : a - b ≤ a ↔ true :=
iff_true_intro (sub_le a b)
theorem sub_lt {a b : } (H1 : 0 < a) (H2 : 0 < b) : a - b < a :=
!nat.cases_on (λh, absurd h !lt.irrefl)
(λa h, succ_le_succ (!nat.cases_on (λh, absurd h !lt.irrefl)
!nat.cases_on (λh, absurd h !nat.lt_irrefl)
(λa h, succ_le_succ (!nat.cases_on (λh, absurd h !nat.lt_irrefl)
(λb c, eq.substr !succ_sub_succ_eq_sub !sub_le) H2)) H1
theorem sub_lt_succ (a b : ) : a - b < succ a :=

View file

@ -34,7 +34,7 @@ lemma znkω_succ (n : nat) (k : nat) : znkω (n+1) k 0 = 0 :=
rfl
lemma znkω_bound (n : nat) (k : nat) : znkω n k n = k :=
if_neg !lt.irrefl
if_neg !nat.lt_irrefl
lemma zω_eq_znkω (n : nat) (k : nat) : zω =[n] znkω n k :=
λ a altn, begin esimp [zω, znkω], rewrite [if_pos altn] end

View file

@ -50,7 +50,7 @@ private lemma Y_unique : ∀ {P l₁ l₂}, length l₁ = length l₂ → Y P l
| P [] (a₂::l₂) h₁ h₂ h₃ := by contradiction
| P (a₁::l₁) [] h₁ h₂ h₃ := by contradiction
| P (a₁::l₁) (a₂::l₂) h₁ h₂ h₃ :=
have n₁ : length l₁ = length l₂, by rewrite [*length_cons at h₁]; apply add.right_cancel h₁,
have n₁ : length l₁ = length l₂, by rewrite [*length_cons at h₁]; apply nat.add_right_cancel h₁,
have n₂ : Y P l₁, from and.elim_left h₂,
have n₃ : Y P l₂, from and.elim_left h₃,
assert ih : l₁ = l₂, from Y_unique n₁ n₂ n₃,

View file

@ -313,8 +313,8 @@ let N := max N₁ N₂ in
exists.intro N
(take n,
suppose n ≥ N,
have ngtN₁ : n ≥ N₁, from nat.le.trans !le_max_left `n ≥ N`,
have ngtN₂ : n ≥ N₂, from nat.le.trans !le_max_right `n ≥ N`,
have ngtN₁ : n ≥ N₁, from nat.le_trans !le_max_left `n ≥ N`,
have ngtN₂ : n ≥ N₂, from nat.le_trans !le_max_right `n ≥ N`,
show abs ((X n + Y n) - (x + y)) < ε, from calc
abs ((X n + Y n) - (x + y))
= abs ((X n - x) + (Y n - y)) : by rewrite [sub_add_eq_sub_sub, *sub_eq_add_neg,

View file

@ -114,7 +114,7 @@ begin
rewrite [-Pig, -Pjh, -pow_add, pow_mod Pe],
apply mem_sep_of_mem !mem_univ,
existsi ((i + j) mod (succ n)), apply and.intro,
apply nat.lt.trans (mod_lt (i+j) !zero_lt_succ) (succ_lt_succ Plt),
apply nat.lt_trans (mod_lt (i+j) !zero_lt_succ) (succ_lt_succ Plt),
apply rfl
end
@ -129,7 +129,7 @@ begin
rewrite [inv_eq_of_mul_eq_one Pinv],
apply mem_sep_of_mem !mem_univ,
existsi ni, apply and.intro,
apply nat.lt.trans (is_lt ni) (succ_lt_succ Plt),
apply nat.lt_trans (is_lt ni) (succ_lt_succ Plt),
apply rfl
end