feat(library): enforce name conventions on old nat declarations

This commit is contained in:
Leonardo de Moura 2015-04-18 10:50:30 -07:00
parent 3e5796acb2
commit cc63a40a01
9 changed files with 37 additions and 36 deletions

View file

@ -42,4 +42,4 @@ end
end end
theorem not_countable_nat_arrow_nat : (countable (nat → nat)) → false := theorem not_countable_nat_arrow_nat : (countable (nat → nat)) → false :=
assume h, absurd (f_eq h) succ.ne_self assume h, absurd (f_eq h) succ_ne_self

View file

@ -60,7 +60,7 @@ theorem length_append : ∀ (s t : list T), length (s ++ t) = length s + length
| (a :: s) t := calc | (a :: s) t := calc
length (a :: s ++ t) = length (s ++ t) + 1 : rfl length (a :: s ++ t) = length (s ++ t) + 1 : rfl
... = length s + length t + 1 : length_append ... = length s + length t + 1 : length_append
... = (length s + 1) + length t : add.succ_left ... = (length s + 1) + length t : succ_add
... = length (a :: s) + length t : rfl ... = length (a :: s) + length t : rfl
theorem eq_nil_of_length_eq_zero : ∀ {l : list T}, length l = 0 → l = [] theorem eq_nil_of_length_eq_zero : ∀ {l : list T}, length l = 0 → l = []

View file

@ -18,7 +18,7 @@ definition addl (x y : ) : :=
nat.rec y (λ n r, succ r) x nat.rec y (λ n r, succ r) x
infix `⊕`:65 := addl infix `⊕`:65 := addl
theorem addl.succ_right (n m : ) : n ⊕ succ m = succ (n ⊕ m) := theorem addl_succ_right (n m : ) : n ⊕ succ m = succ (n ⊕ m) :=
nat.induction_on n nat.induction_on n
rfl rfl
(λ n₁ ih, calc (λ n₁ ih, calc
@ -42,7 +42,7 @@ nat.induction_on x
(λ y₁ ih₂, calc (λ y₁ ih₂, calc
succ x₁ + succ y₁ = succ (succ x₁ + y₁) : rfl succ x₁ + succ y₁ = succ (succ x₁ + y₁) : rfl
... = succ (succ x₁ ⊕ y₁) : {ih₂} ... = succ (succ x₁ ⊕ y₁) : {ih₂}
... = succ x₁ ⊕ succ y₁ : addl.succ_right)) ... = succ x₁ ⊕ succ y₁ : addl_succ_right))
/- successor and predecessor -/ /- successor and predecessor -/
@ -66,15 +66,15 @@ nat.induction_on n
theorem exists_eq_succ_of_ne_zero {n : } (H : n ≠ 0) : ∃k : , n = succ k := theorem exists_eq_succ_of_ne_zero {n : } (H : n ≠ 0) : ∃k : , n = succ k :=
exists.intro _ (or_resolve_right !eq_zero_or_eq_succ_pred H) exists.intro _ (or_resolve_right !eq_zero_or_eq_succ_pred H)
theorem succ.inj {n m : } (H : succ n = succ m) : n = m := theorem succ_inj {n m : } (H : succ n = succ m) : n = m :=
nat.no_confusion H (λe, e) nat.no_confusion H (λe, e)
theorem succ.ne_self {n : } : succ n ≠ n := theorem succ_ne_self {n : } : succ n ≠ n :=
nat.induction_on n nat.induction_on n
(take H : 1 = 0, (take H : 1 = 0,
have ne : 1 ≠ 0, from !succ_ne_zero, have ne : 1 ≠ 0, from !succ_ne_zero,
absurd H ne) absurd H ne)
(take k IH H, IH (succ.inj H)) (take k IH H, IH (succ_inj H))
theorem discriminate {B : Prop} {n : } (H1: n = 0 → B) (H2 : ∀m, n = succ m → B) : B := theorem discriminate {B : Prop} {n : } (H1: n = 0 → B) (H2 : ∀m, n = succ m → B) : B :=
have H : n = n → B, from nat.cases_on n H1 H2, have H : n = n → B, from nat.cases_on n H1 H2,
@ -117,7 +117,7 @@ nat.induction_on n
0 + succ m = succ (0 + m) : add_succ 0 + succ m = succ (0 + m) : add_succ
... = succ m : IH) ... = succ m : IH)
theorem add.succ_left (n m : ) : (succ n) + m = succ (n + m) := theorem succ_add (n m : ) : (succ n) + m = succ (n + m) :=
nat.induction_on m nat.induction_on m
(!add_zero ▸ !add_zero) (!add_zero ▸ !add_zero)
(take k IH, calc (take k IH, calc
@ -131,10 +131,10 @@ nat.induction_on m
(take k IH, calc (take k IH, calc
n + succ k = succ (n+k) : add_succ n + succ k = succ (n+k) : add_succ
... = succ (k + n) : IH ... = succ (k + n) : IH
... = succ k + n : add.succ_left) ... = succ k + n : succ_add)
theorem succ_add_eq_succ_add (n m : ) : succ n + m = n + succ m := theorem succ_add_eq_succ_add (n m : ) : succ n + m = n + succ m :=
!add.succ_left ⬝ !add_succ⁻¹ !succ_add ⬝ !add_succ⁻¹
theorem add.assoc (n m k : ) : (n + m) + k = n + (m + k) := theorem add.assoc (n m k : ) : (n + m) + k = n + (m + k) :=
nat.induction_on k nat.induction_on k
@ -159,10 +159,10 @@ nat.induction_on n
(take (n : ) (IH : n + m = n + k → m = k) (H : succ n + m = succ n + k), (take (n : ) (IH : n + m = n + k → m = k) (H : succ n + m = succ n + k),
have H2 : succ (n + m) = succ (n + k), have H2 : succ (n + m) = succ (n + k),
from calc from calc
succ (n + m) = succ n + m : add.succ_left succ (n + m) = succ n + m : succ_add
... = succ n + k : H ... = succ n + k : H
... = succ (n + k) : add.succ_left, ... = succ (n + k) : succ_add,
have H3 : n + m = n + k, from succ.inj H2, have H3 : n + m = n + k, from succ_inj H2,
IH H3) IH H3)
theorem add.cancel_right {n m k : } (H : n + m = k + m) : n = k := theorem add.cancel_right {n m k : } (H : n + m = k + m) : n = k :=
@ -176,21 +176,21 @@ nat.induction_on n
assume H : succ k + m = 0, assume H : succ k + m = 0,
absurd absurd
(show succ (k + m) = 0, from calc (show succ (k + m) = 0, from calc
succ (k + m) = succ k + m : add.succ_left succ (k + m) = succ k + m : succ_add
... = 0 : H) ... = 0 : H)
!succ_ne_zero) !succ_ne_zero)
theorem eq_zero_of_add_eq_zero_left {n m : } (H : n + m = 0) : m = 0 := theorem eq_zero_of_add_eq_zero_left {n m : } (H : n + m = 0) : m = 0 :=
eq_zero_of_add_eq_zero_right (!add.comm ⬝ H) eq_zero_of_add_eq_zero_right (!add.comm ⬝ H)
theorem add.eq_zero {n m : } (H : n + m = 0) : n = 0 ∧ m = 0 := theorem eq_zero_and_eq_zero_of_add_eq_zero {n m : } (H : n + m = 0) : n = 0 ∧ m = 0 :=
and.intro (eq_zero_of_add_eq_zero_right H) (eq_zero_of_add_eq_zero_left H) and.intro (eq_zero_of_add_eq_zero_right H) (eq_zero_of_add_eq_zero_left H)
theorem add_one (n : ) : n + 1 = succ n := theorem add_one (n : ) : n + 1 = succ n :=
!add_zero ▸ !add_succ !add_zero ▸ !add_succ
theorem one_add (n : ) : 1 + n = succ n := theorem one_add (n : ) : 1 + n = succ n :=
!zero_add ▸ !add.succ_left !zero_add ▸ !succ_add
/- multiplication -/ /- multiplication -/

View file

@ -187,7 +187,7 @@ sub_induction n m
take H : succ k ≤ succ l, take H : succ k ≤ succ l,
calc calc
succ k + (succ l - succ k) = succ k + (l - k) : succ_sub_succ succ k + (succ l - succ k) = succ k + (l - k) : succ_sub_succ
... = succ (k + (l - k)) : add.succ_left ... = succ (k + (l - k)) : succ_add
... = succ l : IH (le_of_succ_le_succ H)) ... = succ l : IH (le_of_succ_le_succ H))
theorem add_sub_of_ge {n m : } (H : n ≥ m) : n + (m - n) = n := theorem add_sub_of_ge {n m : } (H : n ≥ m) : n + (m - n) = n :=

View file

@ -220,33 +220,33 @@ namespace nat
definition min (a b : nat) : nat := definition min (a b : nat) : nat :=
if a < b then a else b if a < b then a else b
definition max_a_a (a : nat) : a = max a a := definition max_self (a : nat) : max a a = a :=
eq.rec_on !if_t_t rfl eq.rec_on !if_t_t rfl
definition max.eq_right {a b : nat} (H : a < b) : max a b = b := definition max_eq_right {a b : nat} (H : a < b) : max a b = b :=
if_pos H if_pos H
definition max.eq_left {a b : nat} (H : ¬ a < b) : max a b = a := definition max_eq_left {a b : nat} (H : ¬ a < b) : max a b = a :=
if_neg H if_neg H
definition max.right_eq {a b : nat} (H : a < b) : b = max a b := definition eq_max_right {a b : nat} (H : a < b) : b = max a b :=
eq.rec_on (max.eq_right H) rfl eq.rec_on (max_eq_right H) rfl
definition max.left_eq {a b : nat} (H : ¬ a < b) : a = max a b := definition eq_max_left {a b : nat} (H : ¬ a < b) : a = max a b :=
eq.rec_on (max.eq_left H) rfl eq.rec_on (max_eq_left H) rfl
definition max.left (a b : nat) : a ≤ max a b := definition le_max_left (a b : nat) : a ≤ max a b :=
by_cases by_cases
(λ h : a < b, le_of_lt (eq.rec_on (max.right_eq h) h)) (λ h : a < b, le_of_lt (eq.rec_on (eq_max_right h) h))
(λ h : ¬ a < b, eq.rec_on (max.eq_left h) !le.refl) (λ h : ¬ a < b, eq.rec_on (eq_max_left h) !le.refl)
definition max.right (a b : nat) : b ≤ max a b := definition le_max_right (a b : nat) : b ≤ max a b :=
by_cases by_cases
(λ h : a < b, eq.rec_on (max.eq_right h) !le.refl) (λ h : a < b, eq.rec_on (eq_max_right h) !le.refl)
(λ h : ¬ a < b, or.rec_on (eq_or_lt_of_not_lt h) (λ h : ¬ a < b, or.rec_on (eq_or_lt_of_not_lt h)
(λ heq, eq.rec_on heq (eq.rec_on (max_a_a a) !le.refl)) (λ heq, eq.rec_on heq (eq.rec_on (eq.symm (max_self a)) !le.refl))
(λ h : b < a, (λ h : b < a,
have aux : a = max a b, from max.left_eq (lt.asymm h), have aux : a = max a b, from eq_max_left (lt.asymm h),
eq.rec_on aux (le_of_lt h))) eq.rec_on aux (le_of_lt h)))
definition gt [reducible] a b := lt b a definition gt [reducible] a b := lt b a

View file

@ -1,3 +1,4 @@
exit
import data.finset algebra.function algebra.binary import data.finset algebra.function algebra.binary
open nat list finset function binary open nat list finset function binary

View file

@ -42,12 +42,12 @@ aux
definition tree_forest.height_lt.cons₁ {A : Type} (t : tree A) (f : forest A) : sum.inl t ≺ sum.inr (forest.cons t f) := definition tree_forest.height_lt.cons₁ {A : Type} (t : tree A) (f : forest A) : sum.inl t ≺ sum.inr (forest.cons t f) :=
have aux : tree.height t < forest.height (forest.cons t f), from have aux : tree.height t < forest.height (forest.cons t f), from
lt_succ_of_le (max.left _ _), lt_succ_of_le (le_max_left _ _),
aux aux
definition tree_forest.height_lt.cons₂ {A : Type} (t : tree A) (f : forest A) : sum.inr f ≺ sum.inr (forest.cons t f) := definition tree_forest.height_lt.cons₂ {A : Type} (t : tree A) (f : forest A) : sum.inr f ≺ sum.inr (forest.cons t f) :=
have aux : forest.height f < forest.height (forest.cons t f), from have aux : forest.height f < forest.height (forest.cons t f), from
lt_succ_of_le (max.right _ _), lt_succ_of_le (le_max_right _ _),
aux aux
definition kind {A : Type} (t : tree_forest A) : bool := definition kind {A : Type} (t : tree_forest A) : bool :=

View file

@ -19,10 +19,10 @@ definition height_lt.wf (A : Type) : well_founded (@height_lt A) :=
inv_image.wf height lt.wf inv_image.wf height lt.wf
theorem height_lt.node_left {A : Type} (t₁ t₂ : tree A) : height_lt t₁ (node t₁ t₂) := theorem height_lt.node_left {A : Type} (t₁ t₂ : tree A) : height_lt t₁ (node t₁ t₂) :=
lt_succ_of_le (max.left (height t₁) (height t₂)) lt_succ_of_le (le_max_left (height t₁) (height t₂))
theorem height_lt.node_right {A : Type} (t₁ t₂ : tree A) : height_lt t₂ (node t₁ t₂) := theorem height_lt.node_right {A : Type} (t₁ t₂ : tree A) : height_lt t₂ (node t₁ t₂) :=
lt_succ_of_le (max.right (height t₁) (height t₂)) lt_succ_of_le (le_max_right (height t₁) (height t₂))
theorem height_lt.trans {A : Type} : transitive (@height_lt A) := theorem height_lt.trans {A : Type} : transitive (@height_lt A) :=
inv_image.trans lt height @lt.trans inv_image.trans lt height @lt.trans

View file

@ -118,7 +118,7 @@ list_induction_on s
calc calc
length (concat (cons x s) t ) = succ (length (concat s t)) : refl _ length (concat (cons x s) t ) = succ (length (concat s t)) : refl _
... = succ (length s + length t) : { H } ... = succ (length s + length t) : { H }
... = succ (length s) + length t : {symm !add.succ_left} ... = succ (length s) + length t : {symm !succ_add}
... = length (cons x s) + length t : refl _) ... = length (cons x s) + length t : refl _)
-- Reverse -- Reverse