feat(hott): port nat and int from the standard library

This commit is contained in:
Floris van Doorn 2015-12-09 00:02:05 -05:00 committed by Leonardo de Moura
parent 46739c8b70
commit 2325d23f68
44 changed files with 2626 additions and 1766 deletions

View file

@ -84,3 +84,34 @@ namespace binary
{A B : Type} (f : A → A → A) (g : B → A) (lcomm : left_commutative f) : left_commutative (compose_left f g) := {A B : Type} (f : A → A → A) (g : B → A) (lcomm : left_commutative f) : left_commutative (compose_left f g) :=
λ a b₁ b₂, !lcomm λ a b₁ b₂, !lcomm
end binary end binary
open eq
namespace is_equiv
definition inv_preserve_binary {A B : Type} (f : A → B) [H : is_equiv f]
(mA : A → A → A) (mB : B → B → B) (H : Π(a a' : A), mB (f a) (f a') = f (mA a a'))
(b b' : B) : f⁻¹ (mB b b') = mA (f⁻¹ b) (f⁻¹ b') :=
begin
have H2 : f⁻¹ (mB (f (f⁻¹ b)) (f (f⁻¹ b'))) = f⁻¹ (f (mA (f⁻¹ b) (f⁻¹ b'))), from ap f⁻¹ !H,
rewrite [+right_inv f at H2,left_inv f at H2,▸* at H2,H2]
end
definition preserve_binary_of_inv_preserve {A B : Type} (f : A → B) [H : is_equiv f]
(mA : A → A → A) (mB : B → B → B) (H : Π(b b' : B), mA (f⁻¹ b) (f⁻¹ b') = f⁻¹ (mB b b'))
(a a' : A) : f (mA a a') = mB (f a) (f a') :=
begin
have H2 : f (mA (f⁻¹ (f a)) (f⁻¹ (f a'))) = f (f⁻¹ (mB (f a) (f a'))), from ap f !H,
rewrite [right_inv f at H2,+left_inv f at H2,▸* at H2,H2]
end
end is_equiv
namespace equiv
open is_equiv equiv.ops
definition inv_preserve_binary {A B : Type} (f : A ≃ B)
(mA : A → A → A) (mB : B → B → B) (H : Π(a a' : A), mB (f a) (f a') = f (mA a a'))
(b b' : B) : f⁻¹ (mB b b') = mA (f⁻¹ b) (f⁻¹ b') :=
inv_preserve_binary f mA mB H b b'
definition preserve_binary_of_inv_preserve {A B : Type} (f : A ≃ B)
(mA : A → A → A) (mB : B → B → B) (H : Π(b b' : B), mA (f⁻¹ b) (f⁻¹ b') = f⁻¹ (mB b b'))
(a a' : A) : f (mA a a') = mB (f a) (f a') :=
preserve_binary_of_inv_preserve f mA mB H a a'
end equiv

View file

@ -35,7 +35,7 @@ namespace functor
apply nat_trans_eq, apply nat_trans_eq,
intro d, calc intro d, calc
natural_map (Fhom F (f' ∘ f)) d = F (f' ∘ f, id) : by esimp natural_map (Fhom F (f' ∘ f)) d = F (f' ∘ f, id) : by esimp
... = F (f' ∘ f, id ∘ id) : by rewrite id_id ... = F (f' ∘ f, category.id ∘ category.id) : by rewrite id_id
... = F ((f',id) ∘ (f, id)) : by esimp ... = F ((f',id) ∘ (f, id)) : by esimp
... = F (f',id) ∘ F (f, id) : by rewrite [respect_comp F] ... = F (f',id) ∘ F (f, id) : by rewrite [respect_comp F]
... = natural_map ((Fhom F f') ∘ (Fhom F f)) d : by esimp ... = natural_map ((Fhom F f') ∘ (Fhom F f)) d : by esimp
@ -119,10 +119,11 @@ namespace functor
apply id_leftright, apply id_leftright,
show (functor_uncurry (functor_curry F)) (f, g) = F (f,g), show (functor_uncurry (functor_curry F)) (f, g) = F (f,g),
from calc from calc
(functor_uncurry (functor_curry F)) (f, g) = to_fun_hom F (id, g) ∘ to_fun_hom F (f, id) : by esimp (functor_uncurry (functor_curry F)) (f, g)
... = F (id ∘ f, g ∘ id) : by krewrite [-respect_comp F (id,g) (f,id)] = to_fun_hom F (id, g) ∘ to_fun_hom F (f, id) : by esimp
... = F (f, g ∘ id) : by rewrite id_left ... = F (category.id ∘ f, g ∘ category.id) : (respect_comp F (id,g) (f,id))⁻¹
... = F (f,g) : by rewrite id_right, ... = F (f, g ∘ category.id) : by rewrite id_left
... = F (f,g) : by rewrite id_right,
end end
definition functor_curry_functor_uncurry_ob (c : C) definition functor_curry_functor_uncurry_ob (c : C)

View file

@ -3,7 +3,7 @@ Copyright (c) 2014 Robert Lewis. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Authors: Robert Lewis Authors: Robert Lewis
Structures with multiplicative and additive components, including division rings and fields. Structures with multiplicative prod additive components, including division rings prod fields.
The development is modeled after Isabelle's library. The development is modeled after Isabelle's library.
-/ -/
import algebra.binary algebra.group algebra.ring import algebra.binary algebra.group algebra.ring
@ -73,7 +73,7 @@ section division_ring
absurd C1 Ha absurd C1 Ha
theorem mul_ne_zero_comm (H : a * b ≠ 0) : b * a ≠ 0 := theorem mul_ne_zero_comm (H : a * b ≠ 0) : b * a ≠ 0 :=
have H2 : a ≠ 0 × b ≠ 0, from ne_zero_and_ne_zero_of_mul_ne_zero H, have H2 : a ≠ 0 × b ≠ 0, from ne_zero_prod_ne_zero_of_mul_ne_zero H,
division_ring.mul_ne_zero (prod.pr2 H2) (prod.pr1 H2) division_ring.mul_ne_zero (prod.pr2 H2) (prod.pr1 H2)
theorem eq_one_div_of_mul_eq_one (H : a * b = 1) : b = 1 / a := theorem eq_one_div_of_mul_eq_one (H : a * b = 1) : b = 1 / a :=
@ -222,7 +222,7 @@ section field
by rewrite [(division_ring.one_div_mul_one_div Ha Hb), mul.comm b] by rewrite [(division_ring.one_div_mul_one_div Ha Hb), mul.comm b]
theorem field.div_mul_right (Hb : b ≠ 0) (H : a * b ≠ 0) : a / (a * b) = 1 / b := theorem field.div_mul_right (Hb : b ≠ 0) (H : a * b ≠ 0) : a / (a * b) = 1 / b :=
have a ≠ 0, from prod.pr1 (ne_zero_and_ne_zero_of_mul_ne_zero H), have a ≠ 0, from prod.pr1 (ne_zero_prod_ne_zero_of_mul_ne_zero H),
symm (calc symm (calc
1 / b = 1 * (1 / b) : one_mul 1 / b = 1 * (1 / b) : one_mul
... = (a * a⁻¹) * (1 / b) : mul_inv_cancel this ... = (a * a⁻¹) * (1 / b) : mul_inv_cancel this
@ -324,10 +324,10 @@ section discrete_field
include s include s
variables {a b c d : A} variables {a b c d : A}
-- many of the theorems in discrete_field are the same as theorems in field or division ring, -- many of the theorems in discrete_field are the same as theorems in field sum division ring,
-- but with fewer hypotheses since 0⁻¹ = 0 and equality is decidable. -- but with fewer hypotheses since 0⁻¹ = 0 prod equality is decidable.
theorem discrete_field.eq_zero_or_eq_zero_of_mul_eq_zero theorem discrete_field.eq_zero_sum_eq_zero_of_mul_eq_zero
(x y : A) (H : x * y = 0) : x = 0 ⊎ y = 0 := (x y : A) (H : x * y = 0) : x = 0 ⊎ y = 0 :=
decidable.by_cases decidable.by_cases
(suppose x = 0, sum.inl this) (suppose x = 0, sum.inl this)
@ -337,7 +337,7 @@ section discrete_field
definition discrete_field.to_integral_domain [trans_instance] [reducible] : definition discrete_field.to_integral_domain [trans_instance] [reducible] :
integral_domain A := integral_domain A :=
⦃ integral_domain, s, ⦃ integral_domain, s,
eq_zero_or_eq_zero_of_mul_eq_zero := discrete_field.eq_zero_or_eq_zero_of_mul_eq_zero⦄ eq_zero_sum_eq_zero_of_mul_eq_zero := discrete_field.eq_zero_sum_eq_zero_of_mul_eq_zero⦄
theorem inv_zero : 0⁻¹ = (0:A) := !discrete_field.inv_zero theorem inv_zero : 0⁻¹ = (0:A) := !discrete_field.inv_zero
@ -524,5 +524,4 @@ theorem subst_into_div [s : has_div A] (a₁ b₁ a₂ b₂ v : A) (H : a₁ / b
by rewrite [H1, H2, H] by rewrite [H1, H2, H]
end norm_num end norm_num
end algebra end algebra

View file

@ -9,7 +9,7 @@ Various multiplicative and additive structures. Partially modeled on Isabelle's
import algebra.binary algebra.priority import algebra.binary algebra.priority
open eq eq.ops -- note: ⁻¹ will be overloaded open eq eq.ops -- note: ⁻¹ will be overloaded
open binary algebra open binary algebra is_trunc
set_option class.force_new true set_option class.force_new true
variable {A : Type} variable {A : Type}
@ -19,8 +19,11 @@ variable {A : Type}
namespace algebra namespace algebra
structure semigroup [class] (A : Type) extends has_mul A := structure semigroup [class] (A : Type) extends has_mul A :=
(is_hset_carrier : is_hset A)
(mul_assoc : Πa b c, mul (mul a b) c = mul a (mul b c)) (mul_assoc : Πa b c, mul (mul a b) c = mul a (mul b c))
attribute semigroup.is_hset_carrier [instance]
theorem mul.assoc [s : semigroup A] (a b c : A) : a * b * c = a * (b * c) := theorem mul.assoc [s : semigroup A] (a b c : A) : a * b * c = a * (b * c) :=
!semigroup.mul_assoc !semigroup.mul_assoc
@ -57,8 +60,11 @@ abbreviation eq_of_mul_eq_mul_right' := @mul.right_cancel
/- additive semigroup -/ /- additive semigroup -/
structure add_semigroup [class] (A : Type) extends has_add A := structure add_semigroup [class] (A : Type) extends has_add A :=
(is_hset_carrier : is_hset A)
(add_assoc : Πa b c, add (add a b) c = add a (add b c)) (add_assoc : Πa b c, add (add a b) c = add a (add b c))
attribute add_semigroup.is_hset_carrier [instance]
theorem add.assoc [s : add_semigroup A] (a b c : A) : a + b + c = a + (b + c) := theorem add.assoc [s : add_semigroup A] (a b c : A) : a + b + c = a + (b + c) :=
!add_semigroup.add_assoc !add_semigroup.add_assoc
@ -121,7 +127,8 @@ definition add_monoid.to_monoid {A : Type} [s : add_monoid A] : monoid A :=
mul_assoc := add_monoid.add_assoc, mul_assoc := add_monoid.add_assoc,
one := add_monoid.zero A, one := add_monoid.zero A,
mul_one := add_monoid.add_zero, mul_one := add_monoid.add_zero,
one_mul := add_monoid.zero_add one_mul := add_monoid.zero_add,
is_hset_carrier := _
definition add_comm_monoid.to_comm_monoid {A : Type} [s : add_comm_monoid A] : comm_monoid A := definition add_comm_monoid.to_comm_monoid {A : Type} [s : add_comm_monoid A] : comm_monoid A :=
@ -577,7 +584,8 @@ definition group_of_add_group (A : Type) [G : add_group A] : group A :=
one_mul := zero_add, one_mul := zero_add,
mul_one := add_zero, mul_one := add_zero,
inv := has_neg.neg, inv := has_neg.neg,
mul_left_inv := add.left_inv⦄ mul_left_inv := add.left_inv,
is_hset_carrier := _⦄
namespace norm_num namespace norm_num
reveal add.assoc reveal add.assoc

View file

@ -65,8 +65,7 @@ namespace eq
fapply Group_eq, fapply Group_eq,
{ apply equiv_of_eq, exact ap (λ(X : Type*), trunc 0 X) (loop_space_succ_eq_in A (succ n))}, { apply equiv_of_eq, exact ap (λ(X : Type*), trunc 0 X) (loop_space_succ_eq_in A (succ n))},
{ exact abstract [irreducible] begin refine trunc.rec _, intro p, refine trunc.rec _, intro q, { exact abstract [irreducible] begin refine trunc.rec _, intro p, refine trunc.rec _, intro q,
rewrite [▸*,-+tr_eq_cast_ap, +trunc_transport, ↑[group_homotopy_group, group.to_monoid, rewrite [▸*,-+tr_eq_cast_ap, +trunc_transport], refine !trunc_transport ⬝ _, apply ap tr,
monoid.to_semigroup, semigroup.to_has_mul, trunc_mul], trunc_transport], apply ap tr,
apply loop_space_succ_eq_in_concat end end}, apply loop_space_succ_eq_in_concat end end},
end end

View file

@ -6,11 +6,20 @@ Author: Floris van Doorn
Theorems about algebra specific to HoTT Theorems about algebra specific to HoTT
-/ -/
import .group arity types.pi hprop_trunc types.unit import .group arity types.pi hprop_trunc types.unit .bundled
open equiv eq equiv.ops is_trunc open equiv eq equiv.ops is_trunc unit
namespace algebra namespace algebra
definition trivial_group [constructor] : group unit :=
group.mk (λx y, star) _ (λx y z, idp) star (unit.rec idp) (unit.rec idp) (λx, star) (λx, idp)
definition Trivial_group [constructor] : Group :=
Group.mk _ trivial_group
notation `G0` := Trivial_group
open Group has_mul has_inv open Group has_mul has_inv
-- we prove under which conditions two groups are equal -- we prove under which conditions two groups are equal

View file

@ -3,11 +3,11 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Author: Jeremy Avigad Author: Jeremy Avigad
Weak orders "≤", strict orders "<", and structures that include both. Weak orders "≤", strict orders "<", prod structures that include both.
-/ -/
import algebra.binary algebra.priority import algebra.binary algebra.priority
open eq eq.ops algebra open eq eq.ops algebra
--set_option class.force_new true -- set_option class.force_new true
variable {A : Type} variable {A : Type}
@ -25,6 +25,8 @@ section
theorem le.refl (a : A) : a ≤ a := !weak_order.le_refl theorem le.refl (a : A) : a ≤ a := !weak_order.le_refl
theorem le_of_eq {a b : A} (H : a = b) : a ≤ b := H ▸ le.refl a
theorem le.trans [trans] {a b c : A} : a ≤ b → b ≤ c → a ≤ c := !weak_order.le_trans theorem le.trans [trans] {a b c : A} : a ≤ b → b ≤ c → a ≤ c := !weak_order.le_trans
theorem ge.trans [trans] {a b c : A} (H1 : a ≥ b) (H2: b ≥ c) : a ≥ c := le.trans H2 H1 theorem ge.trans [trans] {a b c : A} (H1 : a ≥ b) (H2: b ≥ c) : a ≥ c := le.trans H2 H1
@ -83,9 +85,7 @@ definition wf.rec_on {A : Type} [s : wf_strict_order A] {P : A → Type}
(x : A) (H : Πx, (Πy, wf_strict_order.lt y x → P y) → P x) : P x := (x : A) (H : Πx, (Πy, wf_strict_order.lt y x → P y) → P x) : P x :=
wf_strict_order.wf_rec P H x wf_strict_order.wf_rec P H x
definition wf.ind_on := @wf.rec_on /- structures with a weak prod a strict order -/
/- structures with a weak and a strict order -/
structure order_pair [class] (A : Type) extends weak_order A, has_lt A := structure order_pair [class] (A : Type) extends weak_order A, has_lt A :=
(le_of_lt : Π a b, lt a b → le a b) (le_of_lt : Π a b, lt a b → le a b)
@ -126,36 +126,36 @@ section
end end
structure strong_order_pair [class] (A : Type) extends weak_order A, has_lt A := structure strong_order_pair [class] (A : Type) extends weak_order A, has_lt A :=
(le_iff_lt_or_eq : Πa b, le a b ↔ lt a b ⊎ a = b) (le_iff_lt_sum_eq : Πa b, le a b ↔ lt a b ⊎ a = b)
(lt_irrefl : Π a, ¬ lt a a) (lt_irrefl : Π a, ¬ lt a a)
theorem le_iff_lt_or_eq [s : strong_order_pair A] {a b : A} : a ≤ b ↔ a < b ⊎ a = b := theorem le_iff_lt_sum_eq [s : strong_order_pair A] {a b : A} : a ≤ b ↔ a < b ⊎ a = b :=
!strong_order_pair.le_iff_lt_or_eq !strong_order_pair.le_iff_lt_sum_eq
theorem lt_or_eq_of_le [s : strong_order_pair A] {a b : A} (le_ab : a ≤ b) : a < b ⊎ a = b := theorem lt_sum_eq_of_le [s : strong_order_pair A] {a b : A} (le_ab : a ≤ b) : a < b ⊎ a = b :=
iff.mp le_iff_lt_or_eq le_ab iff.mp le_iff_lt_sum_eq le_ab
theorem le_of_lt_or_eq [s : strong_order_pair A] {a b : A} (lt_or_eq : a < b ⊎ a = b) : a ≤ b := theorem le_of_lt_sum_eq [s : strong_order_pair A] {a b : A} (lt_sum_eq : a < b ⊎ a = b) : a ≤ b :=
iff.mpr le_iff_lt_or_eq lt_or_eq iff.mpr le_iff_lt_sum_eq lt_sum_eq
private theorem lt_irrefl' [s : strong_order_pair A] (a : A) : ¬ a < a := private theorem lt_irrefl' [s : strong_order_pair A] (a : A) : ¬ a < a :=
!strong_order_pair.lt_irrefl !strong_order_pair.lt_irrefl
private theorem le_of_lt' [s : strong_order_pair A] (a b : A) : a < b → a ≤ b := private theorem le_of_lt' [s : strong_order_pair A] (a b : A) : a < b → a ≤ b :=
take Hlt, le_of_lt_or_eq (sum.inl Hlt) take Hlt, le_of_lt_sum_eq (sum.inl Hlt)
private theorem lt_iff_le_and_ne [s : strong_order_pair A] {a b : A} : a < b ↔ (a ≤ b × a ≠ b) := private theorem lt_iff_le_prod_ne [s : strong_order_pair A] {a b : A} : a < b ↔ (a ≤ b × a ≠ b) :=
iff.intro iff.intro
(take Hlt, pair (le_of_lt_or_eq (sum.inl Hlt)) (take Hab, absurd (Hab ▸ Hlt) !lt_irrefl')) (take Hlt, pair (le_of_lt_sum_eq (sum.inl Hlt)) (take Hab, absurd (Hab ▸ Hlt) !lt_irrefl'))
(take Hand, (take Hand,
have Hor : a < b ⊎ a = b, from lt_or_eq_of_le (prod.pr1 Hand), have Hor : a < b ⊎ a = b, from lt_sum_eq_of_le (prod.pr1 Hand),
sum_resolve_left Hor (prod.pr2 Hand)) sum_resolve_left Hor (prod.pr2 Hand))
theorem lt_of_le_of_ne [s : strong_order_pair A] {a b : A} : a ≤ b → a ≠ b → a < b := theorem lt_of_le_of_ne [s : strong_order_pair A] {a b : A} : a ≤ b → a ≠ b → a < b :=
take H1 H2, iff.mpr lt_iff_le_and_ne (pair H1 H2) take H1 H2, iff.mpr lt_iff_le_prod_ne (pair H1 H2)
private theorem ne_of_lt' [s : strong_order_pair A] {a b : A} (H : a < b) : a ≠ b := private theorem ne_of_lt' [s : strong_order_pair A] {a b : A} (H : a < b) : a ≠ b :=
prod.pr2 (iff.mp (@lt_iff_le_and_ne _ _ _ _) H) prod.pr2 ((iff.mp (@lt_iff_le_prod_ne _ _ _ _)) H)
private theorem lt_of_lt_of_le' [s : strong_order_pair A] (a b c : A) : a < b → b ≤ c → a < c := private theorem lt_of_lt_of_le' [s : strong_order_pair A] (a b c : A) : a < b → b ≤ c → a < c :=
assume lt_ab : a < b, assume lt_ab : a < b,
@ -166,7 +166,7 @@ have ne_ac : a ≠ c, from
have le_ba : b ≤ a, from eq_ac⁻¹ ▸ le_bc, have le_ba : b ≤ a, from eq_ac⁻¹ ▸ le_bc,
have eq_ab : a = b, from le.antisymm (le_of_lt' _ _ lt_ab) le_ba, have eq_ab : a = b, from le.antisymm (le_of_lt' _ _ lt_ab) le_ba,
show empty, from ne_of_lt' lt_ab eq_ab, show empty, from ne_of_lt' lt_ab eq_ab,
show a < c, from iff.mpr (lt_iff_le_and_ne) (pair le_ac ne_ac) show a < c, from iff.mpr (lt_iff_le_prod_ne) (pair le_ac ne_ac)
theorem lt_of_le_of_lt' [s : strong_order_pair A] (a b c : A) : a ≤ b → b < c → a < c := theorem lt_of_le_of_lt' [s : strong_order_pair A] (a b c : A) : a ≤ b → b < c → a < c :=
assume le_ab : a ≤ b, assume le_ab : a ≤ b,
@ -177,7 +177,7 @@ have ne_ac : a ≠ c, from
have le_cb : c ≤ b, from eq_ac ▸ le_ab, have le_cb : c ≤ b, from eq_ac ▸ le_ab,
have eq_bc : b = c, from le.antisymm (le_of_lt' _ _ lt_bc) le_cb, have eq_bc : b = c, from le.antisymm (le_of_lt' _ _ lt_bc) le_cb,
show empty, from ne_of_lt' lt_bc eq_bc, show empty, from ne_of_lt' lt_bc eq_bc,
show a < c, from iff.mpr (lt_iff_le_and_ne) (pair le_ac ne_ac) show a < c, from iff.mpr (lt_iff_le_prod_ne) (pair le_ac ne_ac)
definition strong_order_pair.to_order_pair [trans_instance] [reducible] definition strong_order_pair.to_order_pair [trans_instance] [reducible]
[s : strong_order_pair A] : order_pair A := [s : strong_order_pair A] : order_pair A :=
@ -206,18 +206,21 @@ section
theorem lt.trichotomy : a < b ⊎ a = b ⊎ b < a := theorem lt.trichotomy : a < b ⊎ a = b ⊎ b < a :=
sum.elim (le.total a b) sum.elim (le.total a b)
(assume H : a ≤ b, (assume H : a ≤ b,
sum.elim (iff.mp (@le_iff_lt_or_eq _ _ _ _) H) (assume H1, sum.inl H1) (assume H1, sum.inr (sum.inl H1))) sum.elim (iff.mp !le_iff_lt_sum_eq H) (assume H1, sum.inl H1) (assume H1, sum.inr (sum.inl H1)))
(assume H : b ≤ a, (assume H : b ≤ a,
sum.elim (iff.mp (@le_iff_lt_or_eq _ _ _ _) H) sum.elim (iff.mp !le_iff_lt_sum_eq H)
(assume H1, sum.inr (sum.inr H1)) (assume H1, sum.inr (sum.inr H1))
(assume H1, sum.inr (sum.inl (H1⁻¹)))) (assume H1, sum.inr (sum.inl (H1⁻¹))))
theorem lt.by_cases {a b : A} {P : Type} definition lt.by_cases {a b : A} {P : Type}
(H1 : a < b → P) (H2 : a = b → P) (H3 : b < a → P) : P := (H1 : a < b → P) (H2 : a = b → P) (H3 : b < a → P) : P :=
sum.elim !lt.trichotomy sum.elim !lt.trichotomy
(assume H, H1 H) (assume H, H1 H)
(assume H, sum.elim H (assume H', H2 H') (assume H', H3 H')) (assume H, sum.elim H (assume H', H2 H') (assume H', H3 H'))
definition lt_ge_by_cases {a b : A} {P : Type} (H1 : a < b → P) (H2 : a ≥ b → P) : P :=
lt.by_cases H1 (λH, H2 (H ▸ le.refl a)) (λH, H2 (le_of_lt H))
theorem le_of_not_gt {a b : A} (H : ¬ a > b) : a ≤ b := theorem le_of_not_gt {a b : A} (H : ¬ a > b) : a ≤ b :=
lt.by_cases (assume H', absurd H' H) (assume H', H' ▸ !le.refl) (assume H', le_of_lt H') lt.by_cases (assume H', absurd H' H) (assume H', H' ▸ !le.refl) (assume H', le_of_lt H')
@ -227,16 +230,16 @@ section
(assume H', absurd (H' ▸ !le.refl) H) (assume H', absurd (H' ▸ !le.refl) H)
(assume H', H') (assume H', H')
theorem lt_or_ge : a < b ⊎ a ≥ b := theorem lt_sum_ge : a < b ⊎ a ≥ b :=
lt.by_cases lt.by_cases
(assume H1 : a < b, sum.inl H1) (assume H1 : a < b, sum.inl H1)
(assume H1 : a = b, sum.inr (H1 ▸ le.refl a)) (assume H1 : a = b, sum.inr (H1 ▸ le.refl a))
(assume H1 : a > b, sum.inr (le_of_lt H1)) (assume H1 : a > b, sum.inr (le_of_lt H1))
theorem le_or_gt : a ≤ b ⊎ a > b := theorem le_sum_gt : a ≤ b ⊎ a > b :=
!sum.swap (lt_or_ge b a) !sum.swap (lt_sum_ge b a)
theorem lt_or_gt_of_ne {a b : A} (H : a ≠ b) : a < b ⊎ a > b := theorem lt_sum_gt_of_ne {a b : A} (H : a ≠ b) : a < b ⊎ a > b :=
lt.by_cases (assume H1, sum.inl H1) (assume H1, absurd H1 H) (assume H1, sum.inr H1) lt.by_cases (assume H1, sum.inl H1) (assume H1, absurd H1 H) (assume H1, sum.inr H1)
end end
@ -272,12 +275,12 @@ section
(assume H : ¬ a ≤ b, (assume H : ¬ a ≤ b,
(inr (assume H1 : a = b, H (H1 ▸ !le.refl)))) (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 := theorem eq_sum_lt_of_not_lt {a b : A} (H : ¬ a < b) : a = b ⊎ b < a :=
if Heq : a = b then sum.inl Heq else sum.inr (lt_of_not_ge (λ Hge, H (lt_of_le_of_ne Hge Heq))) if Heq : a = b then sum.inl Heq else sum.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 := theorem eq_sum_lt_of_le {a b : A} (H : a ≤ b) : a = b ⊎ a < b :=
begin begin
cases eq_or_lt_of_not_lt (not_lt_of_ge H), cases eq_sum_lt_of_not_lt (not_lt_of_ge H),
exact sum.inl a_1⁻¹, exact sum.inl a_1⁻¹,
exact sum.inr a_1 exact sum.inr a_1
end end
@ -301,7 +304,7 @@ section
definition min (a b : A) : A := if a ≤ b then a else b definition min (a b : A) : A := if a ≤ b then a else b
definition max (a b : A) : A := if a ≤ b then b else a definition max (a b : A) : A := if a ≤ b then b else a
/- these show min and max form a lattice -/ /- these show min prod max form a lattice -/
theorem min_le_left (a b : A) : min a b ≤ a := theorem min_le_left (a b : A) : min a b ≤ a :=
by_cases by_cases
@ -339,7 +342,7 @@ section
theorem le_max_right_iff_unit (a b : A) : b ≤ max a b ↔ unit := theorem le_max_right_iff_unit (a b : A) : b ≤ max a b ↔ unit :=
iff_unit_intro (le_max_right a b) iff_unit_intro (le_max_right a b)
/- these are also proved for lattices, but with inf and sup in place of min and max -/ /- these are also proved for lattices, but with inf prod sup in place of min prod max -/
theorem eq_min {a b c : A} (H₁ : c ≤ a) (H₂ : c ≤ b) (H₃ : Π{d}, d ≤ a → d ≤ b → d ≤ c) : theorem eq_min {a b c : A} (H₁ : c ≤ a) (H₂ : c ≤ b) (H₃ : Π{d}, d ≤ a → d ≤ b → d ≤ c) :
c = min a b := c = min a b :=
@ -420,12 +423,12 @@ section
/- these use the fact that it is a linear ordering -/ /- these use the fact that it is a linear ordering -/
theorem lt_min {a b c : A} (H₁ : a < b) (H₂ : a < c) : a < min b c := theorem lt_min {a b c : A} (H₁ : a < b) (H₂ : a < c) : a < min b c :=
sum.elim !le_or_gt sum.elim !le_sum_gt
(assume H : b ≤ c, by rewrite (min_eq_left H); apply H₁) (assume H : b ≤ c, by rewrite (min_eq_left H); apply H₁)
(assume H : b > c, by rewrite (min_eq_right_of_lt H); apply H₂) (assume H : b > c, by rewrite (min_eq_right_of_lt H); apply H₂)
theorem max_lt {a b c : A} (H₁ : a < c) (H₂ : b < c) : max a b < c := theorem max_lt {a b c : A} (H₁ : a < c) (H₂ : b < c) : max a b < c :=
sum.elim !le_or_gt sum.elim !le_sum_gt
(assume H : a ≤ b, by rewrite (max_eq_right H); apply H₂) (assume H : a ≤ b, by rewrite (max_eq_right H); apply H₂)
(assume H : a > b, by rewrite (max_eq_left_of_lt H); apply H₁) (assume H : a > b, by rewrite (max_eq_left_of_lt H); apply H₁)
end end

View file

@ -5,6 +5,7 @@ Authors: Robert Lewis
-/ -/
import algebra.ordered_ring algebra.field import algebra.ordered_ring algebra.field
open eq eq.ops algebra open eq eq.ops algebra
set_option class.force_new true
namespace algebra namespace algebra
structure linear_ordered_field [class] (A : Type) extends linear_ordered_ring A, field A structure linear_ordered_field [class] (A : Type) extends linear_ordered_ring A, field A
@ -339,7 +340,7 @@ section linear_ordered_field
apply one_div_pos_of_pos He apply one_div_pos_of_pos He
end end
theorem exists_add_lt_and_pos_of_lt (H : b < a) : Σ c : A, b + c < a × c > 0 := theorem exists_add_lt_prod_pos_of_lt (H : b < a) : Σ c : A, b + c < a × c > 0 :=
sigma.mk ((a - b) / (1 + 1)) sigma.mk ((a - b) / (1 + 1))
(pair (assert H2 : a + a > (b + b) + (a - b), from calc (pair (assert H2 : a + a > (b + b) + (a - b), from calc
a + a > b + a : add_lt_add_right H a + a > b + a : add_lt_add_right H
@ -356,7 +357,7 @@ section linear_ordered_field
begin begin
apply le_of_not_gt, apply le_of_not_gt,
intro Hb, intro Hb,
cases exists_add_lt_and_pos_of_lt Hb with [c, Hc], cases exists_add_lt_prod_pos_of_lt Hb with [c, Hc],
let Hc' := H c (prod.pr2 Hc), let Hc' := H c (prod.pr2 Hc),
apply (not_le_of_gt (prod.pr1 Hc)) (iff.mpr !le_add_iff_sub_right_le Hc') apply (not_le_of_gt (prod.pr1 Hc)) (iff.mpr !le_add_iff_sub_right_le Hc')
end end

View file

@ -13,6 +13,7 @@ set_option class.force_new true
variable {A : Type} variable {A : Type}
/- partially ordered monoids, such as the natural numbers -/ /- partially ordered monoids, such as the natural numbers -/
namespace algebra namespace algebra
structure ordered_cancel_comm_monoid [class] (A : Type) extends add_comm_monoid A, structure ordered_cancel_comm_monoid [class] (A : Type) extends add_comm_monoid A,
add_left_cancel_semigroup A, add_right_cancel_semigroup A, order_pair A := add_left_cancel_semigroup A, add_right_cancel_semigroup A, order_pair A :=
@ -122,7 +123,7 @@ section
!zero_add ▸ (add_lt_add_of_le_of_lt Ha Hb) !zero_add ▸ (add_lt_add_of_le_of_lt Ha Hb)
-- TODO: add nonpos version (will be easier with simplifier) -- TODO: add nonpos version (will be easier with simplifier)
theorem add_eq_zero_iff_eq_zero_and_eq_zero_of_nonneg_of_nonneg theorem add_eq_zero_iff_eq_zero_prod_eq_zero_of_nonneg_of_nonneg
(Ha : 0 ≤ a) (Hb : 0 ≤ b) : a + b = 0 ↔ a = 0 × b = 0 := (Ha : 0 ≤ a) (Hb : 0 ≤ b) : a + b = 0 ↔ a = 0 × b = 0 :=
iff.intro iff.intro
(assume Hab : a + b = 0, (assume Hab : a + b = 0,
@ -336,7 +337,7 @@ section
iff.mp !add_le_iff_le_sub_left iff.mp !add_le_iff_le_sub_left
theorem add_le_iff_le_sub_right : a + b ≤ c ↔ a ≤ c - b := theorem add_le_iff_le_sub_right : a + b ≤ c ↔ a ≤ c - b :=
have H: a + b ≤ c ↔ a + b - b ≤ c - b, from proof iff.symm (!add_le_add_right_iff) qed, have H: a + b ≤ c ↔ a + b - b ≤ c - b, from iff.symm (!add_le_add_right_iff),
!add_neg_cancel_right ▸ H !add_neg_cancel_right ▸ H
theorem add_le_of_le_sub_right {a b c : A} : a ≤ c - b → a + b ≤ c := theorem add_le_of_le_sub_right {a b c : A} : a ≤ c - b → a + b ≤ c :=
@ -718,7 +719,7 @@ section
show a = b, from eq_of_sub_eq_zero this show a = b, from eq_of_sub_eq_zero this
theorem abs_pos_of_ne_zero (H : a ≠ 0) : abs a > 0 := theorem abs_pos_of_ne_zero (H : a ≠ 0) : abs a > 0 :=
sum.elim (lt_or_gt_of_ne H) abs_pos_of_neg abs_pos_of_pos sum.elim (lt_sum_gt_of_ne H) abs_pos_of_neg abs_pos_of_pos
theorem abs.by_cases {P : A → Type} {a : A} (H1 : P a) (H2 : P (-a)) : P (abs a) := theorem abs.by_cases {P : A → Type} {a : A} (H1 : P a) (H2 : P (-a)) : P (abs a) :=
sum.elim (le.total 0 a) sum.elim (le.total 0 a)
@ -820,5 +821,4 @@ section
end end
end end
end algebra end algebra

View file

@ -4,16 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad Authors: Jeremy Avigad
Here an "ordered_ring" is partially ordered ring, which is ordered with respect to both a weak Here an "ordered_ring" is partially ordered ring, which is ordered with respect to both a weak
order and an associated strict order. Our numeric structures (int, rat, and real) will be instances order prod an associated strict order. Our numeric structures (int, rat, prod real) will be instances
of "linear_ordered_comm_ring". This development is modeled after Isabelle's library. of "linear_ordered_comm_ring". This development is modeled after Isabelle's library.
-/ -/
import algebra.ordered_group algebra.ring import algebra.ordered_group algebra.ring
open eq eq.ops open eq eq.ops algebra
set_option class.force_new true set_option class.force_new true
variable {A : Type} variable {A : Type}
namespace algebra namespace algebra
private definition absurd_a_lt_a {B : Type} {a : A} [s : strict_order A] (H : a < a) : B := private definition absurd_a_lt_a {B : Type} {a : A} [s : strict_order A] (H : a < a) : B :=
absurd H (lt.irrefl a) absurd H (lt.irrefl a)
@ -335,7 +334,7 @@ definition linear_ordered_ring.to_linear_ordered_semiring [trans_instance] [redu
structure linear_ordered_comm_ring [class] (A : Type) extends linear_ordered_ring A, comm_monoid A structure linear_ordered_comm_ring [class] (A : Type) extends linear_ordered_ring A, comm_monoid A
theorem linear_ordered_comm_ring.eq_zero_or_eq_zero_of_mul_eq_zero [s : linear_ordered_comm_ring A] theorem linear_ordered_comm_ring.eq_zero_sum_eq_zero_of_mul_eq_zero [s : linear_ordered_comm_ring A]
{a b : A} (H : a * b = 0) : a = 0 ⊎ b = 0 := {a b : A} (H : a * b = 0) : a = 0 ⊎ b = 0 :=
lt.by_cases lt.by_cases
(assume Ha : 0 < a, (assume Ha : 0 < a,
@ -374,8 +373,8 @@ lt.by_cases
definition linear_ordered_comm_ring.to_integral_domain [trans_instance] [reducible] definition linear_ordered_comm_ring.to_integral_domain [trans_instance] [reducible]
[s: linear_ordered_comm_ring A] : integral_domain A := [s: linear_ordered_comm_ring A] : integral_domain A :=
⦃ integral_domain, s, ⦃ integral_domain, s,
eq_zero_or_eq_zero_of_mul_eq_zero := eq_zero_sum_eq_zero_of_mul_eq_zero :=
@linear_ordered_comm_ring.eq_zero_or_eq_zero_of_mul_eq_zero A s ⦄ @linear_ordered_comm_ring.eq_zero_sum_eq_zero_of_mul_eq_zero A s ⦄
section section
variable [s : linear_ordered_ring A] variable [s : linear_ordered_ring A]
@ -389,7 +388,7 @@ section
theorem zero_le_one : 0 ≤ (1:A) := one_mul 1 ▸ mul_self_nonneg 1 theorem zero_le_one : 0 ≤ (1:A) := one_mul 1 ▸ mul_self_nonneg 1
theorem pos_and_pos_or_neg_and_neg_of_mul_pos {a b : A} (Hab : a * b > 0) : theorem pos_prod_pos_sum_neg_prod_neg_of_mul_pos {a b : A} (Hab : a * b > 0) :
(a > 0 × b > 0) ⊎ (a < 0 × b < 0) := (a > 0 × b > 0) ⊎ (a < 0 × b < 0) :=
lt.by_cases lt.by_cases
(assume Ha : 0 < a, (assume Ha : 0 < a,
@ -712,7 +711,7 @@ section
end end
/- TODO: Multiplication and one, starting with mult_right_le_one_le. -/ /- TODO: Multiplication prod one, starting with mult_right_le_one_le. -/
namespace norm_num namespace norm_num
@ -740,5 +739,4 @@ theorem nonzero_of_neg_helper [s : linear_ordered_ring A] (a : A) (H : a ≠ 0)
begin intro Ha, apply H, apply eq_of_neg_eq_neg, rewrite neg_zero, exact Ha end begin intro Ha, apply H, apply eq_of_neg_eq_neg, rewrite neg_zero, exact Ha end
end norm_num end norm_num
end algebra end algebra

View file

@ -1,9 +1,9 @@
We have ported a lot of algebra files from the standard library to the HoTT library. We port a lot of algebra files from the standard library to the HoTT library.
Port instructions for the abstract structures: Port instructions:
- use the script port.pl in scripts/ to port the file. e.g. execute in the scripts file: - use the script port.pl in scripts/ to port the file. e.g. execute the following in the `scripts` folder: `./port.pl ../library/algebra/lattice.lean ../hott/algebra/lattice.hlean`
`./port.pl ../library/algebra/lattice.lean ../hott/algebra/lattice.hlean` - remove imports starting with `data.` or `logic.` (sometimes you need to replace a `data.` import by the corresponding `types.` import)
- remove imports starting with `data.` or `logic.` - All of the algebraic hierarchy is in the algebra namespace in the HoTT library.
- open namespace algebra, and put every identifier in namespace algebra - Open namespaces `eq` and `algebra` if needed
- add option `set_option class.force_new true` - (optional) add option `set_option class.force_new true`
- fix all remaining errors (open namespace `eq` if needed) - fix all remaining errors

View file

@ -3,19 +3,18 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura Authors: Jeremy Avigad, Leonardo de Moura
Structures with multiplicative and additive components, including semirings, rings, and fields. Structures with multiplicative prod additive components, including semirings, rings, prod fields.
The development is modeled after Isabelle's library. The development is modeled after Isabelle's library.
-/ -/
import algebra.group import algebra.binary algebra.group
open algebra eq open eq eq.ops algebra
variable {A : Type}
set_option class.force_new true set_option class.force_new true
variable {A : Type}
namespace algebra
/- auxiliary classes -/ /- auxiliary classes -/
namespace algebra
structure distrib [class] (A : Type) extends has_mul A, has_add A := structure distrib [class] (A : Type) extends has_mul A, has_add A :=
(left_distrib : Πa b c, mul a (add b c) = add (mul a b) (mul a c)) (left_distrib : Πa b c, mul a (add b c) = add (mul a b) (mul a c))
(right_distrib : Πa b c, mul (add a b) c = add (mul a c) (mul b c)) (right_distrib : Πa b c, mul (add a b) c = add (mul a c) (mul b c))
@ -247,7 +246,7 @@ section
... = 0 : mul_zero, ... = 0 : mul_zero,
symm (neg_eq_of_add_eq_zero this) symm (neg_eq_of_add_eq_zero this)
theorem ne_zero_and_ne_zero_of_mul_ne_zero {a b : A} (H : a * b ≠ 0) : a ≠ 0 × b ≠ 0 := theorem ne_zero_prod_ne_zero_of_mul_ne_zero {a b : A} (H : a * b ≠ 0) : a ≠ 0 × b ≠ 0 :=
have a ≠ 0, from have a ≠ 0, from
(suppose a = 0, (suppose a = 0,
have a * b = 0, by rewrite [this, zero_mul], have a * b = 0, by rewrite [this, zero_mul],
@ -256,7 +255,7 @@ section
(suppose b = 0, (suppose b = 0,
have a * b = 0, by rewrite [this, mul_zero], have a * b = 0, by rewrite [this, mul_zero],
absurd this H), absurd this H),
pair `a ≠ 0` `b ≠ 0` prod.mk `a ≠ 0` `b ≠ 0`
end end
structure comm_ring [class] (A : Type) extends ring A, comm_semigroup A structure comm_ring [class] (A : Type) extends ring A, comm_semigroup A
@ -327,11 +326,11 @@ end
/- integral domains -/ /- integral domains -/
structure no_zero_divisors [class] (A : Type) extends has_mul A, has_zero A := structure no_zero_divisors [class] (A : Type) extends has_mul A, has_zero A :=
(eq_zero_or_eq_zero_of_mul_eq_zero : Πa b, mul a b = zero → a = zero ⊎ b = zero) (eq_zero_sum_eq_zero_of_mul_eq_zero : Πa b, mul a b = zero → a = zero ⊎ b = zero)
theorem eq_zero_or_eq_zero_of_mul_eq_zero {A : Type} [s : no_zero_divisors A] {a b : A} theorem eq_zero_sum_eq_zero_of_mul_eq_zero {A : Type} [s : no_zero_divisors A] {a b : A}
(H : a * b = 0) : (H : a * b = 0) :
a = 0 ⊎ b = 0 := !no_zero_divisors.eq_zero_or_eq_zero_of_mul_eq_zero H a = 0 ⊎ b = 0 := !no_zero_divisors.eq_zero_sum_eq_zero_of_mul_eq_zero H
structure integral_domain [class] (A : Type) extends comm_ring A, no_zero_divisors A, structure integral_domain [class] (A : Type) extends comm_ring A, no_zero_divisors A,
zero_ne_one_class A zero_ne_one_class A
@ -342,18 +341,18 @@ section
theorem mul_ne_zero {a b : A} (H1 : a ≠ 0) (H2 : b ≠ 0) : a * b ≠ 0 := theorem mul_ne_zero {a b : A} (H1 : a ≠ 0) (H2 : b ≠ 0) : a * b ≠ 0 :=
suppose a * b = 0, suppose a * b = 0,
sum.elim (eq_zero_or_eq_zero_of_mul_eq_zero this) (assume H3, H1 H3) (assume H4, H2 H4) sum.elim (eq_zero_sum_eq_zero_of_mul_eq_zero this) (assume H3, H1 H3) (assume H4, H2 H4)
theorem eq_of_mul_eq_mul_right {a b c : A} (Ha : a ≠ 0) (H : b * a = c * a) : b = c := theorem eq_of_mul_eq_mul_right {a b c : A} (Ha : a ≠ 0) (H : b * a = c * a) : b = c :=
have b * a - c * a = 0, from iff.mp !eq_iff_sub_eq_zero H, have b * a - c * a = 0, from iff.mp !eq_iff_sub_eq_zero H,
have (b - c) * a = 0, using this, by rewrite [mul_sub_right_distrib, this], have (b - c) * a = 0, using this, by rewrite [mul_sub_right_distrib, this],
have b - c = 0, from sum_resolve_left (eq_zero_or_eq_zero_of_mul_eq_zero this) Ha, have b - c = 0, from sum_resolve_left (eq_zero_sum_eq_zero_of_mul_eq_zero this) Ha,
iff.elim_right !eq_iff_sub_eq_zero this iff.elim_right !eq_iff_sub_eq_zero this
theorem eq_of_mul_eq_mul_left {a b c : A} (Ha : a ≠ 0) (H : a * b = a * c) : b = c := theorem eq_of_mul_eq_mul_left {a b c : A} (Ha : a ≠ 0) (H : a * b = a * c) : b = c :=
have a * b - a * c = 0, from iff.mp !eq_iff_sub_eq_zero H, have a * b - a * c = 0, from iff.mp !eq_iff_sub_eq_zero H,
have a * (b - c) = 0, using this, by rewrite [mul_sub_left_distrib, this], have a * (b - c) = 0, using this, by rewrite [mul_sub_left_distrib, this],
have b - c = 0, from sum_resolve_right (eq_zero_or_eq_zero_of_mul_eq_zero this) Ha, have b - c = 0, from sum_resolve_right (eq_zero_sum_eq_zero_of_mul_eq_zero this) Ha,
iff.elim_right !eq_iff_sub_eq_zero this iff.elim_right !eq_iff_sub_eq_zero this
-- TODO: do we want the iff versions? -- TODO: do we want the iff versions?
@ -363,7 +362,7 @@ section
suppose b - 1 = 0, H₁ (!zero_add ▸ eq_add_of_sub_eq this), suppose b - 1 = 0, H₁ (!zero_add ▸ eq_add_of_sub_eq this),
have a * b - a = 0, by rewrite H₂; apply sub_self, have a * b - a = 0, by rewrite H₂; apply sub_self,
have a * (b - 1) = 0, by+ rewrite [mul_sub_left_distrib, mul_one]; apply this, have a * (b - 1) = 0, by+ rewrite [mul_sub_left_distrib, mul_one]; apply this,
show a = 0, from sum_resolve_left (eq_zero_or_eq_zero_of_mul_eq_zero this) `b - 1 ≠ 0` show a = 0, from sum_resolve_left (eq_zero_sum_eq_zero_of_mul_eq_zero this) `b - 1 ≠ 0`
theorem eq_zero_of_mul_eq_self_left {a b : A} (H₁ : b ≠ 1) (H₂ : b * a = a) : a = 0 := theorem eq_zero_of_mul_eq_self_left {a b : A} (H₁ : b ≠ 1) (H₂ : b * a = a) : a = 0 :=
eq_zero_of_mul_eq_self_right H₁ (!mul.comm ▸ H₂) eq_zero_of_mul_eq_self_right H₁ (!mul.comm ▸ H₂)
@ -373,7 +372,7 @@ section
(suppose a * a = b * b, (suppose a * a = b * b,
have (a - b) * (a + b) = 0, have (a - b) * (a + b) = 0,
by rewrite [mul.comm, -mul_self_sub_mul_self_eq, this, sub_self], by rewrite [mul.comm, -mul_self_sub_mul_self_eq, this, sub_self],
assert a - b = 0 ⊎ a + b = 0, from !eq_zero_or_eq_zero_of_mul_eq_zero this, assert a - b = 0 ⊎ a + b = 0, from !eq_zero_sum_eq_zero_of_mul_eq_zero this,
sum.elim this sum.elim this
(suppose a - b = 0, sum.inl (eq_of_sub_eq_zero this)) (suppose a - b = 0, sum.inl (eq_of_sub_eq_zero this))
(suppose a + b = 0, sum.inr (eq_neg_of_add_eq_zero this))) (suppose a + b = 0, sum.inr (eq_neg_of_add_eq_zero this)))
@ -385,7 +384,7 @@ section
assert a * a = 1 * 1 ↔ a = 1 ⊎ a = -1, from mul_self_eq_mul_self_iff a 1, assert a * a = 1 * 1 ↔ a = 1 ⊎ a = -1, from mul_self_eq_mul_self_iff a 1,
by rewrite mul_one at this; exact this by rewrite mul_one at this; exact this
-- TODO: c - b * c → c = 0 ⊎ b = 1 and variants -- TODO: c - b * c → c = 0 ⊎ b = 1 prod variants
theorem dvd_of_mul_dvd_mul_left {a b c : A} (Ha : a ≠ 0) (Hdvd : (a * b a * c)) : (b c) := theorem dvd_of_mul_dvd_mul_left {a b c : A} (Ha : a ≠ 0) (Hdvd : (a * b a * c)) : (b c) :=
dvd.elim Hdvd dvd.elim Hdvd

View file

@ -10,7 +10,7 @@ import .sphere
import types.bool types.int.hott types.equiv import types.bool types.int.hott types.equiv
import algebra.homotopy_group algebra.hott import algebra.homotopy_group algebra.hott
open eq susp bool sphere_index is_equiv equiv equiv.ops is_trunc pi open eq susp bool sphere_index is_equiv equiv equiv.ops is_trunc pi algebra
definition circle : Type₀ := sphere 1 definition circle : Type₀ := sphere 1
@ -227,16 +227,18 @@ namespace circle
definition base_eq_base_equiv [constructor] : base = base ≃ := definition base_eq_base_equiv [constructor] : base = base ≃ :=
circle_eq_equiv base circle_eq_equiv base
definition decode_add (a b : ) : circle.decode a ⬝ circle.decode b = circle.decode (a + b) := definition decode_add (a b : ) : circle.decode a ⬝ circle.decode b = circle.decode (a +[] b) :=
!power_con_power !power_con_power
definition encode_con (p q : base = base) : circle.encode (p ⬝ q) = circle.encode p + circle.encode q := definition encode_con (p q : base = base)
preserve_binary_of_inv_preserve base_eq_base_equiv concat add decode_add p q : circle.encode (p ⬝ q) = circle.encode p +[] circle.encode q :=
preserve_binary_of_inv_preserve base_eq_base_equiv concat (@add _) decode_add p q
--the carrier of π₁(S¹) is the set-truncation of base = base. --the carrier of π₁(S¹) is the set-truncation of base = base.
open algebra trunc equiv.ops open algebra trunc equiv.ops
definition fg_carrier_equiv_int : π[1](S¹.) ≃ := definition fg_carrier_equiv_int : π[1](S¹.) ≃ :=
trunc_equiv_trunc 0 base_eq_base_equiv ⬝e !trunc_equiv trunc_equiv_trunc 0 base_eq_base_equiv ⬝e @(trunc_equiv _) proof _ qed
definition con_comm_base (p q : base = base) : p ⬝ q = q ⬝ p := definition con_comm_base (p q : base = base) : p ⬝ q = q ⬝ p :=
eq_of_fn_eq_fn base_eq_base_equiv (by esimp;rewrite [+encode_con,add.comm]) eq_of_fn_eq_fn base_eq_base_equiv (by esimp;rewrite [+encode_con,add.comm])

View file

@ -36,10 +36,16 @@ namespace sphere_index
notation for sphere_index is -1, 0, 1, ... notation for sphere_index is -1, 0, 1, ...
from 0 and up this comes from a coercion from num to sphere_index (via nat) from 0 and up this comes from a coercion from num to sphere_index (via nat)
-/ -/
definition has_zero_sphere_index [instance] [reducible] : has_zero sphere_index :=
has_zero.mk (succ minus_one)
definition has_one_sphere_index [instance] [reducible] : has_one sphere_index :=
has_one.mk (succ (succ minus_one))
postfix `.+1`:(max+1) := sphere_index.succ postfix `.+1`:(max+1) := sphere_index.succ
postfix `.+2`:(max+1) := λ(n : sphere_index), (n .+1 .+1) postfix `.+2`:(max+1) := λ(n : sphere_index), (n .+1 .+1)
notation `-1` := minus_one notation `-1` := minus_one
export [coercions] nat
notation `ℕ₋₁` := sphere_index notation `ℕ₋₁` := sphere_index
definition add (n m : sphere_index) : sphere_index := definition add (n m : sphere_index) : sphere_index :=
@ -50,11 +56,11 @@ namespace sphere_index
infix `+1+`:65 := sphere_index.add infix `+1+`:65 := sphere_index.add
notation x <= y := sphere_index.leq x y definition has_le_sphere_index [instance] [reducible] : has_le sphere_index :=
notation x ≤ y := sphere_index.leq x y has_le.mk leq
definition succ_le_succ {n m : sphere_index} (H : n ≤ m) : n.+1 ≤ m.+1 := H definition succ_le_succ {n m : sphere_index} (H : n ≤ m) : n.+1 ≤ m.+1 := proof H qed
definition le_of_succ_le_succ {n m : sphere_index} (H : n.+1 ≤ m.+1) : n ≤ m := H definition le_of_succ_le_succ {n m : sphere_index} (H : n.+1 ≤ m.+1) : n ≤ m := proof H qed
definition minus_two_le (n : sphere_index) : -1 ≤ n := star definition minus_two_le (n : sphere_index) : -1 ≤ n := star
definition empty_of_succ_le_minus_two {n : sphere_index} (H : n .+1 ≤ -1) : empty := H definition empty_of_succ_le_minus_two {n : sphere_index} (H : n .+1 ≤ -1) : empty := H
@ -104,17 +110,17 @@ namespace sphere
definition bool_of_sphere : S 0 → bool := definition bool_of_sphere : S 0 → bool :=
susp.rec ff tt (λx, empty.elim x) proof susp.rec ff tt (λx, empty.elim x) qed
definition sphere_of_bool : bool → S 0 definition sphere_of_bool : bool → S 0
| ff := north | ff := proof north qed
| tt := south | tt := proof south qed
definition sphere_equiv_bool : S 0 ≃ bool := definition sphere_equiv_bool : S 0 ≃ bool :=
equiv.MK bool_of_sphere equiv.MK bool_of_sphere
sphere_of_bool sphere_of_bool
(λb, match b with | tt := idp | ff := idp end) (λb, match b with | tt := idp | ff := idp end)
(λx, susp.rec_on x idp idp (empty.rec _)) (λx, proof susp.rec_on x idp idp (empty.rec _) qed)
definition sphere_eq_bool : S 0 = bool := definition sphere_eq_bool : S 0 = bool :=
ua sphere_equiv_bool ua sphere_equiv_bool

View file

@ -169,7 +169,7 @@ namespace susp
(!ap_con ⬝ (!ap_con ⬝
whisker_left _ !ap_inv) ⬝ whisker_left _ !ap_inv) ⬝
(!elim_merid ◾ inverse2 !elim_merid)}, (!elim_merid ◾ inverse2 !elim_merid)},
{ rewrite [▸*,inverse2_right_inv (elim_merid function.id idp)], { rewrite [▸*,inverse2_right_inv (elim_merid id idp)],
refine !con.assoc ⬝ _, refine !con.assoc ⬝ _,
xrewrite [ap_con_right_inv (susp.elim x x (λa, a)) (merid idp),idp_con_idp,-ap_compose]} xrewrite [ap_con_right_inv (susp.elim x x (λa, a)) (merid idp),idp_con_idp,-ap_compose]}
end end

155
hott/init/connectives.hlean Normal file
View file

@ -0,0 +1,155 @@
/-
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura, Haitao Zhang
The propositional connectives.
-/
prelude
import .types
open unit
variables {a b c d : Type}
/- implies -/
definition imp (a b : Type) : Type := a → b
definition imp.id (H : a) : a := H
definition imp.intro (H : a) (H₂ : b) : a := H
definition imp.mp (H : a) (H₂ : a → b) : b :=
H₂ H
definition imp.syl (H : a → b) (H₂ : c → a) (Hc : c) : b :=
H (H₂ Hc)
definition imp.left (H : a → b) (H₂ : b → c) (Ha : a) : c :=
H₂ (H Ha)
definition imp_unit (a : Type) : (a → unit) ↔ unit :=
iff_unit_intro (imp.intro star)
definition unit_imp (a : Type) : (unit → a) ↔ a :=
iff.intro (assume H, H star) imp.intro
definition imp_empty (a : Type) : (a → empty) ↔ ¬ a := iff.rfl
definition empty_imp (a : Type) : (empty → a) ↔ unit :=
iff_unit_intro empty.elim
/- not -/
definition not.elim {A : Type} (H1 : ¬a) (H2 : a) : A := absurd H2 H1
definition not.mto {a b : Type} : (a → b) → ¬b → ¬a := imp.left
definition not_imp_not_of_imp {a b : Type} : (a → b) → ¬b → ¬a := not.mto
definition not_not_of_not_implies : ¬(a → b) → ¬¬a :=
not.mto not.elim
definition not_of_not_implies : ¬(a → b) → ¬b :=
not.mto imp.intro
definition not_not_em : ¬¬(a ⊎ ¬a) :=
assume not_em : ¬(a ⊎ ¬a),
not_em (sum.inr (not.mto sum.inl not_em))
definition not_iff_not (H : a ↔ b) : ¬a ↔ ¬b :=
iff.intro (not.mto (iff.mpr H)) (not.mto (iff.mp H))
/- prod -/
definition not_prod_of_not_left (b : Type) : ¬a → ¬(a × b) :=
not.mto prod.pr1
definition not_prod_of_not_right (a : Type) {b : Type} : ¬b → ¬(a × b) :=
not.mto prod.pr2
definition prod.imp_left (H : a → b) : a × c → b × c :=
prod.imp H imp.id
definition prod.imp_right (H : a → b) : c × a → c × b :=
prod.imp imp.id H
definition prod_of_prod_of_imp_of_imp (H₁ : a × b) (H₂ : a → c) (H₃ : b → d) : c × d :=
prod.imp H₂ H₃ H₁
definition prod_of_prod_of_imp_left (H₁ : a × c) (H : a → b) : b × c :=
prod.imp_left H H₁
definition prod_of_prod_of_imp_right (H₁ : c × a) (H : a → b) : c × b :=
prod.imp_right H H₁
definition prod_imp_iff (a b c : Type) : (a × b → c) ↔ (a → b → c) :=
iff.intro (λH a b, H (pair a b)) prod.rec
/- sum -/
definition not_sum : ¬a → ¬b → ¬(a ⊎ b) := sum.rec
definition sum_of_sum_of_imp_of_imp (H₁ : a ⊎ b) (H₂ : a → c) (H₃ : b → d) : c ⊎ d :=
sum.imp H₂ H₃ H₁
definition sum_of_sum_of_imp_left (H₁ : a ⊎ c) (H : a → b) : b ⊎ c :=
sum.imp_left H H₁
definition sum_of_sum_of_imp_right (H₁ : c ⊎ a) (H : a → b) : c ⊎ b :=
sum.imp_right H H₁
definition sum.elim3 (H : a ⊎ b ⊎ c) (Ha : a → d) (Hb : b → d) (Hc : c → d) : d :=
sum.elim H Ha (assume H₂, sum.elim H₂ Hb Hc)
definition sum_resolve_right (H₁ : a ⊎ b) (H₂ : ¬a) : b :=
sum.elim H₁ (not.elim H₂) imp.id
definition sum_resolve_left (H₁ : a ⊎ b) : ¬b → a :=
sum_resolve_right (sum.swap H₁)
definition sum.imp_distrib : ((a ⊎ b) → c) ↔ ((a → c) × (b → c)) :=
iff.intro
(λH, pair (imp.syl H sum.inl) (imp.syl H sum.inr))
(prod.rec sum.rec)
definition sum_iff_right_of_imp {a b : Type} (Ha : a → b) : (a ⊎ b) ↔ b :=
iff.intro (sum.rec Ha imp.id) sum.inr
definition sum_iff_left_of_imp {a b : Type} (Hb : b → a) : (a ⊎ b) ↔ a :=
iff.intro (sum.rec imp.id Hb) sum.inl
definition sum_iff_sum (H1 : a ↔ c) (H2 : b ↔ d) : (a ⊎ b) ↔ (c ⊎ d) :=
iff.intro (sum.imp (iff.mp H1) (iff.mp H2)) (sum.imp (iff.mpr H1) (iff.mpr H2))
/- distributivity -/
definition prod.pr1_distrib (a b c : Type) : a × (b ⊎ c) ↔ (a × b) ⊎ (a × c) :=
iff.intro
(prod.rec (λH, sum.imp (pair H) (pair H)))
(sum.rec (prod.imp_right sum.inl) (prod.imp_right sum.inr))
definition prod.pr2_distrib (a b c : Type) : (a ⊎ b) × c ↔ (a × c) ⊎ (b × c) :=
iff.trans (iff.trans !prod.comm !prod.pr1_distrib) (sum_iff_sum !prod.comm !prod.comm)
definition sum.left_distrib (a b c : Type) : a ⊎ (b × c) ↔ (a ⊎ b) × (a ⊎ c) :=
iff.intro
(sum.rec (λH, pair (sum.inl H) (sum.inl H)) (prod.imp sum.inr sum.inr))
(prod.rec (sum.rec (imp.syl imp.intro sum.inl) (imp.syl sum.imp_right pair)))
definition sum.right_distrib (a b c : Type) : (a × b) ⊎ c ↔ (a ⊎ c) × (b ⊎ c) :=
iff.trans (iff.trans !sum.comm !sum.left_distrib) (prod_congr !sum.comm !sum.comm)
/- iff -/
definition iff.def : (a ↔ b) = ((a → b) × (b → a)) := rfl
definition pi_imp_pi {A : Type} {P Q : A → Type} (H : Πa, (P a → Q a)) (p : Πa, P a) (a : A) : Q a :=
(H a) (p a)
definition pi_iff_pi {A : Type} {P Q : A → Type} (H : Πa, (P a ↔ Q a)) : (Πa, P a) ↔ (Πa, Q a) :=
iff.intro (λp a, iff.elim_left (H a) (p a)) (λq a, iff.elim_right (H a) (q a))
definition imp_iff {P : Type} (Q : Type) (p : P) : (P → Q) ↔ Q :=
iff.intro (λf, f p) imp.intro

View file

@ -7,13 +7,15 @@ Authors: Leonardo de Moura, Jakob von Raumer, Floris van Doorn
prelude prelude
import init.datatypes init.reserved_notation init.tactic init.logic import init.datatypes init.reserved_notation init.tactic init.logic
import init.bool init.num init.relation init.wf import init.bool init.num init.relation init.wf
import init.types import init.types init.connectives
import init.trunc init.path init.equiv init.util import init.trunc init.path init.equiv init.util
import init.ua init.funext import init.ua init.funext
import init.hedberg init.nat init.hit init.pathover import init.hedberg init.nat init.hit init.pathover
namespace core namespace core
export bool empty unit sum export bool unit
export empty (hiding elim)
export sum (hiding elim)
export sigma (hiding pr1 pr2) export sigma (hiding pr1 pr2)
export [notations] prod export [notations] prod
export [notations] nat export [notations] nat

View file

@ -24,9 +24,6 @@ definition compose_right [reducible] [unfold_full] (f : B → B → B) (g : A
definition compose_left [reducible] [unfold_full] (f : B → B → B) (g : A → B) : A → B → B := definition compose_left [reducible] [unfold_full] (f : B → B → B) (g : A → B) : A → B → B :=
λ a b, f (g a) b λ a b, f (g a) b
definition id [reducible] [unfold_full] (a : A) : A :=
a
definition on_fun [reducible] [unfold_full] (f : B → B → C) (g : A → B) : A → A → C := definition on_fun [reducible] [unfold_full] (f : B → B → C) (g : A → B) : A → A → C :=
λx y, f (g x) (g y) λx y, f (g x) (g y)

View file

@ -8,6 +8,9 @@ prelude
import init.reserved_notation import init.reserved_notation
open unit open unit
definition id [reducible] [unfold_full] {A : Type} (a : A) : A :=
a
/- not -/ /- not -/
definition not [reducible] (a : Type) := a → empty definition not [reducible] (a : Type) := a → empty
@ -19,24 +22,20 @@ empty.rec (λ e, b) (H₂ H₁)
definition mt {a b : Type} (H₁ : a → b) (H₂ : ¬b) : ¬a := definition mt {a b : Type} (H₁ : a → b) (H₂ : ¬b) : ¬a :=
assume Ha : a, absurd (H₁ Ha) H₂ assume Ha : a, absurd (H₁ Ha) H₂
protected definition not_empty : ¬ empty := definition not_empty : ¬empty :=
assume H : empty, H assume H : empty, H
definition not_not_intro {a : Type} (Ha : a) : ¬¬a := definition non_contradictory (a : Type) : Type := ¬¬a
definition non_contradictory_intro {a : Type} (Ha : a) : ¬¬a :=
assume Hna : ¬a, absurd Ha Hna assume Hna : ¬a, absurd Ha Hna
theorem not_of_not_not_not {a : Type} (H : ¬¬¬a) : ¬a :=
λ Ha, absurd (not_not_intro Ha) H
definition not.elim {a : Type} (H₁ : ¬a) (H₂ : a) : empty := H₁ H₂
definition not.intro {a : Type} (H : a → empty) : ¬a := H definition not.intro {a : Type} (H : a → empty) : ¬a := H
definition not_not_of_not_implies {a b : Type} (H : ¬(a → b)) : ¬¬a := /- empty -/
assume Hna : ¬a, absurd (assume Ha : a, absurd Ha Hna) H
definition not_of_not_implies {a b : Type} (H : ¬(a → b)) : ¬b := definition empty.elim {c : Type} (H : empty) : c :=
assume Hb : b, absurd (assume Ha : a, Hb) H empty.rec _ H
/- eq -/ /- eq -/
@ -55,10 +54,10 @@ namespace eq
definition symm [unfold 4] (H : a = b) : b = a := definition symm [unfold 4] (H : a = b) : b = a :=
subst H (refl a) subst H (refl a)
theorem mp {a b : Type} : (a = b) → a → b := definition mp {a b : Type} : (a = b) → a → b :=
eq.rec_on eq.rec_on
theorem mpr {a b : Type} : (a = b) → b → a := definition mpr {a b : Type} : (a = b) → b → a :=
assume H₁ H₂, eq.rec_on (eq.symm H₁) H₂ assume H₁ H₂, eq.rec_on (eq.symm H₁) H₂
namespace ops end ops -- this is just to ensure that this namespace exists. There is nothing in it namespace ops end ops -- this is just to ensure that this namespace exists. There is nothing in it
@ -75,13 +74,13 @@ eq.rec H₁ H₂
definition congr {A B : Type} {f₁ f₂ : A → B} {a₁ a₂ : A} (H₁ : f₁ = f₂) (H₂ : a₁ = a₂) : f₁ a₁ = f₂ a₂ := definition congr {A B : Type} {f₁ f₂ : A → B} {a₁ a₂ : A} (H₁ : f₁ = f₂) (H₂ : a₁ = a₂) : f₁ a₁ = f₂ a₂ :=
eq.subst H₁ (eq.subst H₂ rfl) eq.subst H₁ (eq.subst H₂ rfl)
theorem congr_fun {A : Type} {B : A → Type} {f g : Π x, B x} (H : f = g) (a : A) : f a = g a := definition congr_fun {A : Type} {B : A → Type} {f g : Π x, B x} (H : f = g) (a : A) : f a = g a :=
eq.subst H (eq.refl (f a)) eq.subst H (eq.refl (f a))
theorem congr_arg {A B : Type} (a a' : A) (f : A → B) (Ha : a = a') : f a = f a' := definition congr_arg {A B : Type} (a a' : A) (f : A → B) (Ha : a = a') : f a = f a' :=
eq.subst Ha rfl eq.subst Ha rfl
theorem congr_arg2 {A B C : Type} (a a' : A) (b b' : B) (f : A → B → C) (Ha : a = a') (Hb : b = b') : f a b = f a' b' := definition congr_arg2 {A B C : Type} (a a' : A) (b b' : B) (f : A → B → C) (Ha : a = a') (Hb : b = b') : f a b = f a' b' :=
eq.subst Ha (eq.subst Hb rfl) eq.subst Ha (eq.subst Hb rfl)
section section
@ -110,318 +109,563 @@ end lift
/- ne -/ /- ne -/
definition ne {A : Type} (a b : A) := ¬(a = b) definition ne [reducible] {A : Type} (a b : A) := ¬(a = b)
infix ≠ := ne notation a ≠ b := ne a b
namespace ne namespace ne
open eq.ops open eq.ops
variable {A : Type} variable {A : Type}
variables {a b : A} variables {a b : A}
definition intro : (a = b → empty) → a ≠ b := definition intro (H : a = b → empty) : a ≠ b := H
assume H, H
definition elim : a ≠ b → a = b → empty := definition elim (H : a ≠ b) : a = b → empty := H
assume H₁ H₂, H₁ H₂
definition irrefl : a ≠ a → empty := definition irrefl (H : a ≠ a) : empty := H rfl
assume H, H rfl
definition symm : a ≠ b → b ≠ a := definition symm (H : a ≠ b) : b ≠ a :=
assume (H : a ≠ b) (H₁ : b = a), H H₁⁻¹ assume (H₁ : b = a), H (H₁⁻¹)
end ne end ne
definition empty_of_ne {A : Type} {a : A} : a ≠ a → empty := ne.irrefl
section section
open eq.ops open eq.ops
variables {A : Type} {a b c : A} variables {p : Type₀}
definition empty.of_ne : a ≠ a → empty := definition ne_empty_of_self : p → p ≠ empty :=
assume H, H rfl assume (Hp : p) (Heq : p = empty), Heq ▸ Hp
definition ne.of_eq_of_ne : a = b → b ≠ c → a ≠ c := definition ne_unit_of_not : ¬p → p ≠ unit :=
assume H₁ H₂, H₁⁻¹ ▸ H₂ assume (Hnp : ¬p) (Heq : p = unit), (Heq ▸ Hnp) star
definition ne.of_ne_of_eq : a ≠ b → b = c → a ≠ c := definition unit_ne_empty : ¬unit = empty :=
assume H₁ H₂, H₂ ▸ H₁ ne_empty_of_self star
end end
/- prod -/
abbreviation pair [constructor] := @prod.mk
infixr × := prod
variables {a b c d : Type}
attribute prod.rec [elim]
attribute prod.mk [intro!]
protected definition prod.elim [unfold 4] (H₁ : a × b) (H₂ : a → b → c) : c :=
prod.rec H₂ H₁
definition prod.swap [unfold 3] : a × b → b × a :=
prod.rec (λHa Hb, prod.mk Hb Ha)
/- sum -/
infixr ⊎ := sum
infixr + := sum
attribute sum.rec [elim]
protected definition sum.elim [unfold 4] (H₁ : a ⊎ b) (H₂ : a → c) (H₃ : b → c) : c :=
sum.rec H₂ H₃ H₁
definition non_contradictory_em (a : Type) : ¬¬(a ⊎ ¬a) :=
assume not_em : ¬(a ⊎ ¬a),
have neg_a : ¬a, from
assume pos_a : a, absurd (sum.inl pos_a) not_em,
absurd (sum.inr neg_a) not_em
definition sum.swap : a ⊎ b → b ⊎ a := sum.rec sum.inr sum.inl
/- iff -/ /- iff -/
definition iff (a b : Type) := prod (a → b) (b → a) definition iff (a b : Type) := (a → b) × (b → a)
infix <-> := iff notation a <-> b := iff a b
infix ↔ := iff notation a ↔ b := iff a b
variables {a b c : Type}
namespace iff definition iff.intro : (a → b) → (b → a) → (a ↔ b) := prod.mk
definition def : (a ↔ b) = (prod (a → b) (b → a)) := attribute iff.intro [intro!]
rfl
definition intro (H₁ : a → b) (H₂ : b → a) : a ↔ b := definition iff.elim : ((a → b) → (b → a) → c) → (a ↔ b) → c := prod.rec
prod.mk H₁ H₂
definition elim (H₁ : (a → b) → (b → a) → c) (H₂ : a ↔ b) : c := attribute iff.elim [recursor 5] [elim]
prod.rec H₁ H₂
definition elim_left (H : a ↔ b) : a → b := definition iff.elim_left : (a ↔ b) → a → b := prod.pr1
elim (assume H₁ H₂, H₁) H
definition mp := @elim_left definition iff.mp := @iff.elim_left
definition elim_right (H : a ↔ b) : b → a := definition iff.elim_right : (a ↔ b) → b → a := prod.pr2
elim (assume H₁ H₂, H₂) H
definition mpr := @elim_right definition iff.mpr := @iff.elim_right
definition flip_sign (H₁ : a ↔ b) : ¬a ↔ ¬b := definition iff.refl [refl] (a : Type) : a ↔ a :=
intro iff.intro (assume H, H) (assume H, H)
(assume Hna, mt (elim_right H₁) Hna)
(assume Hnb, mt (elim_left H₁) Hnb)
definition refl (a : Type) : a ↔ a := definition iff.rfl {a : Type} : a ↔ a :=
intro (assume H, H) (assume H, H) iff.refl a
definition rfl {a : Type} : a ↔ a := definition iff.trans [trans] (H₁ : a ↔ b) (H₂ : b ↔ c) : a ↔ c :=
refl a iff.intro
(assume Ha, iff.mp H₂ (iff.mp H₁ Ha))
(assume Hc, iff.mpr H₁ (iff.mpr H₂ Hc))
definition iff_of_eq (a b : Type) (p : a = b) : a ↔ b := definition iff.symm [symm] (H : a ↔ b) : b ↔ a :=
eq.rec rfl p iff.intro (iff.elim_right H) (iff.elim_left H)
definition trans (H₁ : a ↔ b) (H₂ : b ↔ c) : a ↔ c := definition iff.comm : (a ↔ b) ↔ (b ↔ a) :=
intro iff.intro iff.symm iff.symm
(assume Ha, elim_left H₂ (elim_left H₁ Ha))
(assume Hc, elim_right H₁ (elim_right H₂ Hc))
definition symm (H : a ↔ b) : b ↔ a := definition iff.of_eq {a b : Type} (H : a = b) : a ↔ b :=
intro eq.rec_on H iff.rfl
(assume Hb, elim_right H Hb)
(assume Ha, elim_left H Ha)
definition unit_elim (H : a ↔ unit) : a := definition not_iff_not_of_iff (H₁ : a ↔ b) : ¬a ↔ ¬b :=
mp (symm H) unit.star
definition empty_elim (H : a ↔ empty) : ¬a :=
assume Ha : a, mp H Ha
open eq.ops
definition of_eq {a b : Type} (H : a = b) : a ↔ b :=
iff.intro (λ Ha, H ▸ Ha) (λ Hb, H⁻¹ ▸ Hb)
definition pi_iff_pi {A : Type} {P Q : A → Type} (H : Πa, (P a ↔ Q a)) : (Πa, P a) ↔ Πa, Q a :=
iff.intro (λp a, iff.elim_left (H a) (p a)) (λq a, iff.elim_right (H a) (q a))
theorem imp_iff {P : Type} (Q : Type) (p : P) : (P → Q) ↔ Q :=
iff.intro (λf, f p) (λq p, q)
end iff
theorem not_iff_not_of_iff (H₁ : a ↔ b) : ¬a ↔ ¬b :=
iff.intro iff.intro
(assume (Hna : ¬ a) (Hb : b), Hna (iff.elim_right H₁ Hb)) (assume (Hna : ¬ a) (Hb : b), Hna (iff.elim_right H₁ Hb))
(assume (Hnb : ¬ b) (Ha : a), Hnb (iff.elim_left H₁ Ha)) (assume (Hnb : ¬ b) (Ha : a), Hnb (iff.elim_left H₁ Ha))
theorem of_iff_unit (H : a ↔ unit) : a := definition of_iff_unit (H : a ↔ unit) : a :=
iff.mp (iff.symm H) star iff.mp (iff.symm H) star
theorem not_of_iff_empty : (a ↔ empty) → ¬a := iff.mp definition not_of_iff_empty : (a ↔ empty) → ¬a := iff.mp
theorem iff_unit_intro (H : a) : a ↔ unit := definition iff_unit_intro (H : a) : a ↔ unit :=
iff.intro iff.intro
(λ Hl, star) (λ Hl, star)
(λ Hr, H) (λ Hr, H)
theorem iff_empty_intro (H : ¬a) : a ↔ empty := definition iff_empty_intro (H : ¬a) : a ↔ empty :=
iff.intro H (empty.rec _) iff.intro H (empty.rec _)
theorem not_non_contradictory_iff_absurd (a : Type) : ¬¬¬a ↔ ¬a := definition not_non_contradictory_iff_absurd (a : Type) : ¬¬¬a ↔ ¬a :=
iff.intro iff.intro
(λ (Hl : ¬¬¬a) (Ha : a), Hl (λf, f Ha)) (λ (Hl : ¬¬¬a) (Ha : a), Hl (non_contradictory_intro Ha))
absurd absurd
attribute iff.refl [refl] definition imp_congr [congr] (H1 : a ↔ c) (H2 : b ↔ d) : (a → b) ↔ (c → d) :=
attribute iff.trans [trans] iff.intro
attribute iff.symm [symm] (λHab Hc, iff.mp H2 (Hab (iff.mpr H1 Hc)))
(λHcd Ha, iff.mpr H2 (Hcd (iff.mp H1 Ha)))
definition not_not_intro (Ha : a) : ¬¬a :=
assume Hna : ¬a, Hna Ha
definition not_of_not_not_not (H : ¬¬¬a) : ¬a :=
λ Ha, absurd (not_not_intro Ha) H
definition not_unit [simp] : (¬ unit) ↔ empty :=
iff_empty_intro (not_not_intro star)
definition not_empty_iff [simp] : (¬ empty) ↔ unit :=
iff_unit_intro not_empty
definition not_congr [congr] (H : a ↔ b) : ¬a ↔ ¬b :=
iff.intro (λ H₁ H₂, H₁ (iff.mpr H H₂)) (λ H₁ H₂, H₁ (iff.mp H H₂))
definition ne_self_iff_empty [simp] {A : Type} (a : A) : (not (a = a)) ↔ empty :=
iff.intro empty_of_ne empty.elim
definition eq_self_iff_unit [simp] {A : Type} (a : A) : (a = a) ↔ unit :=
iff_unit_intro rfl
definition iff_not_self [simp] (a : Type) : (a ↔ ¬a) ↔ empty :=
iff_empty_intro (λ H,
have H' : ¬a, from (λ Ha, (iff.mp H Ha) Ha),
H' (iff.mpr H H'))
definition not_iff_self [simp] (a : Type) : (¬a ↔ a) ↔ empty :=
iff_empty_intro (λ H,
have H' : ¬a, from (λ Ha, (iff.mpr H Ha) Ha),
H' (iff.mp H H'))
definition unit_iff_empty [simp] : (unit ↔ empty) ↔ empty :=
iff_empty_intro (λ H, iff.mp H star)
definition empty_iff_unit [simp] : (empty ↔ unit) ↔ empty :=
iff_empty_intro (λ H, iff.mpr H star)
definition empty_of_unit_iff_empty : (unit ↔ empty) → empty :=
assume H, iff.mp H star
/- prod simp rules -/
definition prod.imp (H₂ : a → c) (H₃ : b → d) : a × b → c × d :=
prod.rec (λHa Hb, prod.mk (H₂ Ha) (H₃ Hb))
definition prod_congr [congr] (H1 : a ↔ c) (H2 : b ↔ d) : (a × b) ↔ (c × d) :=
iff.intro (prod.imp (iff.mp H1) (iff.mp H2)) (prod.imp (iff.mpr H1) (iff.mpr H2))
definition prod.comm [simp] : a × b ↔ b × a :=
iff.intro prod.swap prod.swap
definition prod.assoc [simp] : (a × b) × c ↔ a × (b × c) :=
iff.intro
(prod.rec (λ H' Hc, prod.rec (λ Ha Hb, prod.mk Ha (prod.mk Hb Hc)) H'))
(prod.rec (λ Ha, prod.rec (λ Hb Hc, prod.mk (prod.mk Ha Hb) Hc)))
definition prod.pr1_comm [simp] : a × (b × c) ↔ b × (a × c) :=
iff.trans (iff.symm !prod.assoc) (iff.trans (prod_congr !prod.comm !iff.refl) !prod.assoc)
definition prod_iff_left {a b : Type} (Hb : b) : (a × b) ↔ a :=
iff.intro prod.pr1 (λHa, prod.mk Ha Hb)
definition prod_iff_right {a b : Type} (Ha : a) : (a × b) ↔ b :=
iff.intro prod.pr2 (prod.mk Ha)
definition prod_unit [simp] (a : Type) : a × unit ↔ a :=
prod_iff_left star
definition unit_prod [simp] (a : Type) : unit × a ↔ a :=
prod_iff_right star
definition prod_empty [simp] (a : Type) : a × empty ↔ empty :=
iff_empty_intro prod.pr2
definition empty_prod [simp] (a : Type) : empty × a ↔ empty :=
iff_empty_intro prod.pr1
definition not_prod_self [simp] (a : Type) : (¬a × a) ↔ empty :=
iff_empty_intro (λ H, prod.elim H (λ H₁ H₂, absurd H₂ H₁))
definition prod_not_self [simp] (a : Type) : (a × ¬a) ↔ empty :=
iff_empty_intro (λ H, prod.elim H (λ H₁ H₂, absurd H₁ H₂))
definition prod_self [simp] (a : Type) : a × a ↔ a :=
iff.intro prod.pr1 (assume H, prod.mk H H)
/- sum simp rules -/
definition sum.imp (H₂ : a → c) (H₃ : b → d) : a ⊎ b → c ⊎ d :=
sum.rec (λ H, sum.inl (H₂ H)) (λ H, sum.inr (H₃ H))
definition sum.imp_left (H : a → b) : a ⊎ c → b ⊎ c :=
sum.imp H id
definition sum.imp_right (H : a → b) : c ⊎ a → c ⊎ b :=
sum.imp id H
definition sum_congr [congr] (H1 : a ↔ c) (H2 : b ↔ d) : (a ⊎ b) ↔ (c ⊎ d) :=
iff.intro (sum.imp (iff.mp H1) (iff.mp H2)) (sum.imp (iff.mpr H1) (iff.mpr H2))
definition sum.comm [simp] : a ⊎ b ↔ b ⊎ a := iff.intro sum.swap sum.swap
definition sum.assoc [simp] : (a ⊎ b) ⊎ c ↔ a ⊎ (b ⊎ c) :=
iff.intro
(sum.rec (sum.imp_right sum.inl) (λ H, sum.inr (sum.inr H)))
(sum.rec (λ H, sum.inl (sum.inl H)) (sum.imp_left sum.inr))
definition sum.left_comm [simp] : a ⊎ (b ⊎ c) ↔ b ⊎ (a ⊎ c) :=
iff.trans (iff.symm !sum.assoc) (iff.trans (sum_congr !sum.comm !iff.refl) !sum.assoc)
definition sum_unit [simp] (a : Type) : a ⊎ unit ↔ unit :=
iff_unit_intro (sum.inr star)
definition unit_sum [simp] (a : Type) : unit ⊎ a ↔ unit :=
iff_unit_intro (sum.inl star)
definition sum_empty [simp] (a : Type) : a ⊎ empty ↔ a :=
iff.intro (sum.rec id empty.elim) sum.inl
definition empty_sum [simp] (a : Type) : empty ⊎ a ↔ a :=
iff.trans sum.comm !sum_empty
definition sum_self [simp] (a : Type) : a ⊎ a ↔ a :=
iff.intro (sum.rec id id) sum.inl
/- sum resolution rulse -/
definition sum.resolve_left {a b : Type} (H : a ⊎ b) (na : ¬ a) : b :=
sum.elim H (λ Ha, absurd Ha na) id
definition sum.neg_resolve_left {a b : Type} (H : ¬ a ⊎ b) (Ha : a) : b :=
sum.elim H (λ na, absurd Ha na) id
definition sum.resolve_right {a b : Type} (H : a ⊎ b) (nb : ¬ b) : a :=
sum.elim H id (λ Hb, absurd Hb nb)
definition sum.neg_resolve_right {a b : Type} (H : a ⊎ ¬ b) (Hb : b) : a :=
sum.elim H id (λ nb, absurd Hb nb)
/- iff simp rules -/
definition iff_unit [simp] (a : Type) : (a ↔ unit) ↔ a :=
iff.intro (assume H, iff.mpr H star) iff_unit_intro
definition unit_iff [simp] (a : Type) : (unit ↔ a) ↔ a :=
iff.trans iff.comm !iff_unit
definition iff_empty [simp] (a : Type) : (a ↔ empty) ↔ ¬ a :=
iff.intro prod.pr1 iff_empty_intro
definition empty_iff [simp] (a : Type) : (empty ↔ a) ↔ ¬ a :=
iff.trans iff.comm !iff_empty
definition iff_self [simp] (a : Type) : (a ↔ a) ↔ unit :=
iff_unit_intro iff.rfl
definition iff_congr [congr] (H1 : a ↔ c) (H2 : b ↔ d) : (a ↔ b) ↔ (c ↔ d) :=
prod_congr (imp_congr H1 H2) (imp_congr H2 H1)
/- decidable -/
inductive decidable [class] (p : Type) : Type :=
| inl : p → decidable p
| inr : ¬p → decidable p
definition decidable_unit [instance] : decidable unit :=
decidable.inl star
definition decidable_empty [instance] : decidable empty :=
decidable.inr not_empty
-- We use "dependent" if-then-else to be able to communicate the if-then-else condition
-- to the branches
definition dite (c : Type) [H : decidable c] {A : Type} : (c → A) → (¬ c → A) → A :=
decidable.rec_on H
/- if-then-else -/
definition ite (c : Type) [H : decidable c] {A : Type} (t e : A) : A :=
decidable.rec_on H (λ Hc, t) (λ Hnc, e)
namespace decidable
variables {p q : Type}
definition by_cases {q : Type} [C : decidable p] : (p → q) → (¬p → q) → q := !dite
theorem em (p : Type) [H : decidable p] : p ⊎ ¬p := by_cases sum.inl sum.inr
theorem by_contradiction [Hp : decidable p] (H : ¬p → empty) : p :=
if H1 : p then H1 else empty.rec _ (H H1)
end decidable
section
variables {p q : Type}
open decidable
definition decidable_of_decidable_of_iff (Hp : decidable p) (H : p ↔ q) : decidable q :=
if Hp : p then inl (iff.mp H Hp)
else inr (iff.mp (not_iff_not_of_iff H) Hp)
definition decidable_of_decidable_of_eq {p q : Type} (Hp : decidable p) (H : p = q)
: decidable q :=
decidable_of_decidable_of_iff Hp (iff.of_eq H)
protected definition sum.by_cases [Hp : decidable p] [Hq : decidable q] {A : Type}
(h : p ⊎ q) (h₁ : p → A) (h₂ : q → A) : A :=
if hp : p then h₁ hp else
if hq : q then h₂ hq else
empty.rec _ (sum.elim h hp hq)
end
section
variables {p q : Type}
open decidable (rec_on inl inr)
definition decidable_prod [instance] [Hp : decidable p] [Hq : decidable q] : decidable (p × q) :=
if hp : p then
if hq : q then inl (prod.mk hp hq)
else inr (assume H : p × q, hq (prod.pr2 H))
else inr (assume H : p × q, hp (prod.pr1 H))
definition decidable_sum [instance] [Hp : decidable p] [Hq : decidable q] : decidable (p ⊎ q) :=
if hp : p then inl (sum.inl hp) else
if hq : q then inl (sum.inr hq) else
inr (sum.rec hp hq)
definition decidable_not [instance] [Hp : decidable p] : decidable (¬p) :=
if hp : p then inr (absurd hp) else inl hp
definition decidable_implies [instance] [Hp : decidable p] [Hq : decidable q] : decidable (p → q) :=
if hp : p then
if hq : q then inl (assume H, hq)
else inr (assume H : p → q, absurd (H hp) hq)
else inl (assume Hp, absurd Hp hp)
definition decidable_iff [instance] [Hp : decidable p] [Hq : decidable q] : decidable (p ↔ q) :=
decidable_prod
end
definition decidable_pred [reducible] {A : Type} (R : A → Type) := Π (a : A), decidable (R a)
definition decidable_rel [reducible] {A : Type} (R : A → A → Type) := Π (a b : A), decidable (R a b)
definition decidable_eq [reducible] (A : Type) := decidable_rel (@eq A)
definition decidable_ne [instance] {A : Type} [H : decidable_eq A] (a b : A) : decidable (a ≠ b) :=
decidable_implies
namespace bool
theorem ff_ne_tt : ff = tt → empty
| [none]
end bool
open bool
definition is_dec_eq {A : Type} (p : A → A → bool) : Type := Π ⦃x y : A⦄, p x y = tt → x = y
definition is_dec_refl {A : Type} (p : A → A → bool) : Type := Πx, p x x = tt
open decidable
protected definition bool.has_decidable_eq [instance] : Πa b : bool, decidable (a = b)
| ff ff := inl rfl
| ff tt := inr ff_ne_tt
| tt ff := inr (ne.symm ff_ne_tt)
| tt tt := inl rfl
definition decidable_eq_of_bool_pred {A : Type} {p : A → A → bool} (H₁ : is_dec_eq p) (H₂ : is_dec_refl p) : decidable_eq A :=
take x y : A, if Hp : p x y = tt then inl (H₁ Hp)
else inr (assume Hxy : x = y, (eq.subst Hxy Hp) (H₂ y))
/- inhabited -/ /- inhabited -/
inductive inhabited [class] (A : Type) : Type := inductive inhabited [class] (A : Type) : Type :=
mk : A → inhabited A mk : A → inhabited A
namespace inhabited protected definition inhabited.value {A : Type} : inhabited A → A :=
inhabited.rec (λa, a)
protected definition destruct {A : Type} {B : Type} (H1 : inhabited A) (H2 : A → B) : B := protected definition inhabited.destruct {A : Type} {B : Type} (H1 : inhabited A) (H2 : A → B) : B :=
inhabited.rec H2 H1 inhabited.rec H2 H1
definition default (A : Type) [H : inhabited A] : A :=
inhabited.value H
definition arbitrary [irreducible] (A : Type) [H : inhabited A] : A :=
inhabited.value H
definition Type.is_inhabited [instance] : inhabited Type :=
inhabited.mk (lift unit)
definition inhabited_fun [instance] (A : Type) {B : Type} [H : inhabited B] : inhabited (A → B) := definition inhabited_fun [instance] (A : Type) {B : Type} [H : inhabited B] : inhabited (A → B) :=
inhabited.destruct H (λb, mk (λa, b)) inhabited.rec_on H (λb, inhabited.mk (λa, b))
definition inhabited_Pi [instance] (A : Type) {B : A → Type} [H : Πx, inhabited (B x)] : definition inhabited_Pi [instance] (A : Type) {B : A → Type} [H : Πx, inhabited (B x)] :
inhabited (Πx, B x) := inhabited (Πx, B x) :=
mk (λa, inhabited.destruct (H a) (λb, b)) inhabited.mk (λa, !default)
definition default (A : Type) [H : inhabited A] : A := inhabited.destruct H (take a, a) protected definition bool.is_inhabited [instance] : inhabited bool :=
inhabited.mk ff
end inhabited protected definition pos_num.is_inhabited [instance] : inhabited pos_num :=
inhabited.mk pos_num.one
/- decidable -/ protected definition num.is_inhabited [instance] : inhabited num :=
inhabited.mk num.zero
inductive decidable.{l} [class] (p : Type.{l}) : Type.{l} := inductive nonempty [class] (A : Type) : Type :=
| inl : p → decidable p intro : A → nonempty A
| inr : ¬p → decidable p
namespace decidable protected definition nonempty.elim {A : Type} {B : Type} (H1 : nonempty A) (H2 : A → B) : B :=
variables {p q : Type} nonempty.rec H2 H1
definition pos_witness [C : decidable p] (H : p) : p := theorem nonempty_of_inhabited [instance] {A : Type} [H : inhabited A] : nonempty A :=
decidable.rec_on C (λ Hp, Hp) (λ Hnp, absurd H Hnp) nonempty.intro !default
definition neg_witness [C : decidable p] (H : ¬ p) : ¬ p := theorem nonempty_of_exists {A : Type} {P : A → Type} : (sigma P) → nonempty A :=
decidable.rec_on C (λ Hp, absurd Hp H) (λ Hnp, Hnp) sigma.rec (λw H, nonempty.intro w)
definition by_cases {q : Type} [C : decidable p] (Hpq : p → q) (Hnpq : ¬p → q) : q := /- subsingleton -/
decidable.rec_on C (assume Hp, Hpq Hp) (assume Hnp, Hnpq Hnp)
definition em (p : Type) [H : decidable p] : sum p ¬p := inductive subsingleton [class] (A : Type) : Type :=
by_cases (λ Hp, sum.inl Hp) (λ Hnp, sum.inr Hnp) intro : (Π a b : A, a = b) → subsingleton A
definition by_contradiction [Hp : decidable p] (H : ¬p → empty) : p := protected definition subsingleton.elim {A : Type} [H : subsingleton A] : Π(a b : A), a = b :=
by_cases subsingleton.rec (λp, p) H
(assume H₁ : p, H₁)
(assume H₁ : ¬p, empty.rec (λ e, p) (H H₁))
definition decidable_iff_equiv (Hp : decidable p) (H : p ↔ q) : decidable q := protected theorem rec_subsingleton {p : Type} [H : decidable p]
decidable.rec_on Hp {H1 : p → Type} {H2 : ¬p → Type}
(assume Hp : p, inl (iff.elim_left H Hp)) [H3 : Π(h : p), subsingleton (H1 h)] [H4 : Π(h : ¬p), subsingleton (H2 h)]
(assume Hnp : ¬p, inr (iff.elim_left (iff.flip_sign H) Hnp)) : subsingleton (decidable.rec_on H H1 H2) :=
decidable.rec_on H (λh, H3 h) (λh, H4 h) --this can be proven using dependent version of "by_cases"
definition decidable_eq_equiv.{l} {p q : Type.{l}} (Hp : decidable p) (H : p = q) : decidable q := theorem if_pos {c : Type} [H : decidable c] (Hc : c) {A : Type} {t e : A} : (ite c t e) = t :=
decidable_iff_equiv Hp (iff.of_eq H)
end decidable
section
variables {p q : Type}
open decidable (rec_on inl inr)
definition decidable_unit [instance] : decidable unit :=
inl unit.star
definition decidable_empty [instance] : decidable empty :=
inr not_empty
definition decidable_prod [instance] [Hp : decidable p] [Hq : decidable q] : decidable (prod p q) :=
rec_on Hp
(assume Hp : p, rec_on Hq
(assume Hq : q, inl (prod.mk Hp Hq))
(assume Hnq : ¬q, inr (λ H : prod p q, prod.rec_on H (λ Hp Hq, absurd Hq Hnq))))
(assume Hnp : ¬p, inr (λ H : prod p q, prod.rec_on H (λ Hp Hq, absurd Hp Hnp)))
definition decidable_sum [instance] [Hp : decidable p] [Hq : decidable q] : decidable (sum p q) :=
rec_on Hp
(assume Hp : p, inl (sum.inl Hp))
(assume Hnp : ¬p, rec_on Hq
(assume Hq : q, inl (sum.inr Hq))
(assume Hnq : ¬q, inr (λ H : sum p q, sum.rec_on H (λ Hp, absurd Hp Hnp) (λ Hq, absurd Hq Hnq))))
definition decidable_not [instance] [Hp : decidable p] : decidable (¬p) :=
rec_on Hp
(assume Hp, inr (not_not_intro Hp))
(assume Hnp, inl Hnp)
definition decidable_implies [instance] [Hp : decidable p] [Hq : decidable q] : decidable (p → q) :=
rec_on Hp
(assume Hp : p, rec_on Hq
(assume Hq : q, inl (assume H, Hq))
(assume Hnq : ¬q, inr (assume H : p → q, absurd (H Hp) Hnq)))
(assume Hnp : ¬p, inl (assume Hp, absurd Hp Hnp))
definition decidable_if [instance] [Hp : decidable p] [Hq : decidable q] : decidable (p ↔ q) :=
show decidable (prod (p → q) (q → p)), from _
end
definition decidable_pred [reducible] {A : Type} (R : A → Type) := Π (a : A), decidable (R a)
definition decidable_rel [reducible] {A : Type} (R : A → A → Type) := Π (a b : A), decidable (R a b)
definition decidable_eq [reducible] (A : Type) := decidable_rel (@eq A)
definition decidable_ne [instance] {A : Type} [H : decidable_eq A] : decidable_rel (@ne A) :=
show Π x y : A, decidable (x = y → empty), from _
definition ite (c : Type) [H : decidable c] {A : Type} (t e : A) : A :=
decidable.rec_on H (λ Hc, t) (λ Hnc, e)
definition if_pos {c : Type} [H : decidable c] (Hc : c) {A : Type} {t e : A} : (if c then t else e) = t :=
decidable.rec decidable.rec
(λ Hc : c, eq.refl (@ite c (decidable.inl Hc) A t e)) (λ Hc : c, eq.refl (@ite c (decidable.inl Hc) A t e))
(λ Hnc : ¬c, absurd Hc Hnc) (λ Hnc : ¬c, absurd Hc Hnc)
H H
definition if_neg {c : Type} [H : decidable c] (Hnc : ¬c) {A : Type} {t e : A} : (if c then t else e) = e := theorem if_neg {c : Type} [H : decidable c] (Hnc : ¬c) {A : Type} {t e : A} : (ite c t e) = e :=
decidable.rec decidable.rec
(λ Hc : c, absurd Hc Hnc) (λ Hc : c, absurd Hc Hnc)
(λ Hnc : ¬c, eq.refl (@ite c (decidable.inr Hnc) A t e)) (λ Hnc : ¬c, eq.refl (@ite c (decidable.inr Hnc) A t e))
H H
definition if_t_t (c : Type) [H : decidable c] {A : Type} (t : A) : (if c then t else t) = t := theorem if_t_t [simp] (c : Type) [H : decidable c] {A : Type} (t : A) : (ite c t t) = t :=
decidable.rec decidable.rec
(λ Hc : c, eq.refl (@ite c (decidable.inl Hc) A t t)) (λ Hc : c, eq.refl (@ite c (decidable.inl Hc) A t t))
(λ Hnc : ¬c, eq.refl (@ite c (decidable.inr Hnc) A t t)) (λ Hnc : ¬c, eq.refl (@ite c (decidable.inr Hnc) A t t))
H H
definition if_unit {A : Type} (t e : A) : (if unit then t else e) = t := theorem implies_of_if_pos {c t e : Type} [H : decidable c] (h : ite c t e) : c → t :=
if_pos unit.star
definition if_empty {A : Type} (t e : A) : (if empty then t else e) = e :=
if_neg not_empty
section
open eq.ops
definition if_cond_congr {c₁ c₂ : Type} [H₁ : decidable c₁] [H₂ : decidable c₂] (Heq : c₁ ↔ c₂) {A : Type} (t e : A)
: (if c₁ then t else e) = (if c₂ then t else e) :=
decidable.rec_on H₁
(λ Hc₁ : c₁, decidable.rec_on H₂
(λ Hc₂ : c₂, if_pos Hc₁ ⬝ (if_pos Hc₂)⁻¹)
(λ Hnc₂ : ¬c₂, absurd (iff.elim_left Heq Hc₁) Hnc₂))
(λ Hnc₁ : ¬c₁, decidable.rec_on H₂
(λ Hc₂ : c₂, absurd (iff.elim_right Heq Hc₂) Hnc₁)
(λ Hnc₂ : ¬c₂, if_neg Hnc₁ ⬝ (if_neg Hnc₂)⁻¹))
definition if_congr_aux {c₁ c₂ : Type} [H₁ : decidable c₁] [H₂ : decidable c₂] {A : Type} {t₁ t₂ e₁ e₂ : A}
(Hc : c₁ ↔ c₂) (Ht : t₁ = t₂) (He : e₁ = e₂) :
(if c₁ then t₁ else e₁) = (if c₂ then t₂ else e₂) :=
Ht ▸ He ▸ (if_cond_congr Hc t₁ e₁)
definition if_congr {c₁ c₂ : Type} [H₁ : decidable c₁] {A : Type} {t₁ t₂ e₁ e₂ : A} (Hc : c₁ ↔ c₂) (Ht : t₁ = t₂) (He : e₁ = e₂) :
(if c₁ then t₁ else e₁) = (@ite c₂ (decidable.decidable_iff_equiv H₁ Hc) A t₂ e₂) :=
have H2 [visible] : decidable c₂, from (decidable.decidable_iff_equiv H₁ Hc),
if_congr_aux Hc Ht He
theorem implies_of_if_pos {c t e : Type} [H : decidable c] (h : if c then t else e) : c → t :=
assume Hc, eq.rec_on (if_pos Hc) h assume Hc, eq.rec_on (if_pos Hc) h
theorem implies_of_if_neg {c t e : Type} [H : decidable c] (h : if c then t else e) : ¬c → e := theorem implies_of_if_neg {c t e : Type} [H : decidable c] (h : ite c t e) : ¬c → e :=
assume Hnc, eq.rec_on (if_neg Hnc) h assume Hnc, eq.rec_on (if_neg Hnc) h
-- We use "dependent" if-then-else to be able to communicate the if-then-else condition theorem if_ctx_congr {A : Type} {b c : Type} [dec_b : decidable b] [dec_c : decidable c]
-- to the branches {x y u v : A}
definition dite (c : Type) [H : decidable c] {A : Type} (t : c → A) (e : ¬ c → A) : A := (h_c : b ↔ c) (h_t : c → x = u) (h_e : ¬c → y = v) :
decidable.rec_on H (λ Hc, t Hc) (λ Hnc, e Hnc) ite b x y = ite c u v :=
decidable.rec_on dec_b
(λ hp : b, calc
ite b x y = x : if_pos hp
... = u : h_t (iff.mp h_c hp)
... = ite c u v : if_pos (iff.mp h_c hp))
(λ hn : ¬b, calc
ite b x y = y : if_neg hn
... = v : h_e (iff.mp (not_iff_not_of_iff h_c) hn)
... = ite c u v : if_neg (iff.mp (not_iff_not_of_iff h_c) hn))
definition dif_pos {c : Type} [H : decidable c] (Hc : c) {A : Type} {t : c → A} {e : ¬ c → A} : (if H : c then t H else e H) = t (decidable.pos_witness Hc) := theorem if_congr [congr] {A : Type} {b c : Type} [dec_b : decidable b] [dec_c : decidable c]
decidable.rec {x y u v : A}
(λ Hc : c, eq.refl (@dite c (decidable.inl Hc) A t e)) (h_c : b ↔ c) (h_t : x = u) (h_e : y = v) :
(λ Hnc : ¬c, absurd Hc Hnc) ite b x y = ite c u v :=
H @if_ctx_congr A b c dec_b dec_c x y u v h_c (λ h, h_t) (λ h, h_e)
definition dif_neg {c : Type} [H : decidable c] (Hnc : ¬c) {A : Type} {t : c → A} {e : ¬ c → A} : (if H : c then t H else e H) = e (decidable.neg_witness Hnc) := theorem if_ctx_simp_congr {A : Type} {b c : Type} [dec_b : decidable b] {x y u v : A}
decidable.rec (h_c : b ↔ c) (h_t : c → x = u) (h_e : ¬c → y = v) :
(λ Hc : c, absurd Hc Hnc) ite b x y = (@ite c (decidable_of_decidable_of_iff dec_b h_c) A u v) :=
(λ Hnc : ¬c, eq.refl (@dite c (decidable.inr Hnc) A t e)) @if_ctx_congr A b c dec_b (decidable_of_decidable_of_iff dec_b h_c) x y u v h_c h_t h_e
H
-- Remark: dite and ite are "definitionally equal" when we ignore the proofs. theorem if_simp_congr [congr] {A : Type} {b c : Type} [dec_b : decidable b] {x y u v : A}
definition dite_ite_eq (c : Type) [H : decidable c] {A : Type} (t : A) (e : A) : dite c (λh, t) (λh, e) = ite c t e := (h_c : b ↔ c) (h_t : x = u) (h_e : y = v) :
ite b x y = (@ite c (decidable_of_decidable_of_iff dec_b h_c) A u v) :=
@if_ctx_simp_congr A b c dec_b x y u v h_c (λ h, h_t) (λ h, h_e)
definition if_unit [simp] {A : Type} (t e : A) : (if unit then t else e) = t :=
if_pos star
definition if_empty [simp] {A : Type} (t e : A) : (if empty then t else e) = e :=
if_neg not_empty
theorem if_ctx_congr_prop {b c x y u v : Type} [dec_b : decidable b] [dec_c : decidable c]
(h_c : b ↔ c) (h_t : c → (x ↔ u)) (h_e : ¬c → (y ↔ v)) :
ite b x y ↔ ite c u v :=
decidable.rec_on dec_b
(λ hp : b, calc
ite b x y ↔ x : iff.of_eq (if_pos hp)
... ↔ u : h_t (iff.mp h_c hp)
... ↔ ite c u v : iff.of_eq (if_pos (iff.mp h_c hp)))
(λ hn : ¬b, calc
ite b x y ↔ y : iff.of_eq (if_neg hn)
... ↔ v : h_e (iff.mp (not_iff_not_of_iff h_c) hn)
... ↔ ite c u v : iff.of_eq (if_neg (iff.mp (not_iff_not_of_iff h_c) hn)))
theorem if_congr_prop [congr] {b c x y u v : Type} [dec_b : decidable b] [dec_c : decidable c]
(h_c : b ↔ c) (h_t : x ↔ u) (h_e : y ↔ v) :
ite b x y ↔ ite c u v :=
if_ctx_congr_prop h_c (λ h, h_t) (λ h, h_e)
theorem if_ctx_simp_congr_prop {b c x y u v : Type} [dec_b : decidable b]
(h_c : b ↔ c) (h_t : c → (x ↔ u)) (h_e : ¬c → (y ↔ v)) :
ite b x y ↔ (@ite c (decidable_of_decidable_of_iff dec_b h_c) Type u v) :=
@if_ctx_congr_prop b c x y u v dec_b (decidable_of_decidable_of_iff dec_b h_c) h_c h_t h_e
theorem if_simp_congr_prop [congr] {b c x y u v : Type} [dec_b : decidable b]
(h_c : b ↔ c) (h_t : x ↔ u) (h_e : y ↔ v) :
ite b x y ↔ (@ite c (decidable_of_decidable_of_iff dec_b h_c) Type u v) :=
@if_ctx_simp_congr_prop b c x y u v dec_b h_c (λ h, h_t) (λ h, h_e)
-- Remark: dite prod ite are "definitionally equal" when we ignore the proofs.
theorem dite_ite_eq (c : Type) [H : decidable c] {A : Type} (t : A) (e : A) : dite c (λh, t) (λh, e) = ite c t e :=
rfl rfl
end
open eq.ops unit
definition is_unit (c : Type) [H : decidable c] : Type₀ := definition is_unit (c : Type) [H : decidable c] : Type₀ :=
if c then unit else empty if c then unit else empty
@ -429,16 +673,26 @@ if c then unit else empty
definition is_empty (c : Type) [H : decidable c] : Type₀ := definition is_empty (c : Type) [H : decidable c] : Type₀ :=
if c then empty else unit if c then empty else unit
theorem of_is_unit {c : Type} [H₁ : decidable c] (H₂ : is_unit c) : c := definition of_is_unit {c : Type} [H₁ : decidable c] (H₂ : is_unit c) : c :=
decidable.rec_on H₁ (λ Hc, Hc) (λ Hnc, empty.rec _ (if_neg Hnc ▸ H₂)) decidable.rec_on H₁ (λ Hc, Hc) (λ Hnc, empty.rec _ (if_neg Hnc ▸ H₂))
notation `dec_trivial` := of_is_unit star notation `dec_star` := of_is_unit star
theorem not_of_not_is_unit {c : Type} [H₁ : decidable c] (H₂ : ¬ is_unit c) : ¬ c := theorem not_of_not_is_unit {c : Type} [H₁ : decidable c] (H₂ : ¬ is_unit c) : ¬ c :=
decidable.rec_on H₁ (λ Hc, absurd star (if_pos Hc ▸ H₂)) (λ Hnc, Hnc) if Hc : c then absurd star (if_pos Hc ▸ H₂) else Hc
theorem not_of_is_empty {c : Type} [H₁ : decidable c] (H₂ : is_empty c) : ¬ c := theorem not_of_is_empty {c : Type} [H₁ : decidable c] (H₂ : is_empty c) : ¬ c :=
decidable.rec_on H₁ (λ Hc, empty.rec _ (if_pos Hc ▸ H₂)) (λ Hnc, Hnc) if Hc : c then empty.rec _ (if_pos Hc ▸ H₂) else Hc
theorem of_not_is_empty {c : Type} [H₁ : decidable c] (H₂ : ¬ is_empty c) : c := theorem of_not_is_empty {c : Type} [H₁ : decidable c] (H₂ : ¬ is_empty c) : c :=
decidable.rec_on H₁ (λ Hc, Hc) (λ Hnc, absurd star (if_neg Hnc ▸ H₂)) if Hc : c then Hc else absurd star (if_neg Hc ▸ H₂)
-- The following symbols should not be considered in the pattern inference procedure used by
-- heuristic instantiation.
attribute prod sum not iff ite dite eq ne [no_pattern]
-- namespace used to collect congruence rules for "contextual simplification"
namespace contextual
attribute if_ctx_simp_congr [congr]
attribute if_ctx_simp_congr_prop [congr]
end contextual

View file

@ -179,10 +179,10 @@ namespace nat
theorem lt_zero_iff_empty [simp] (a : ) : a < 0 ↔ empty := theorem lt_zero_iff_empty [simp] (a : ) : a < 0 ↔ empty :=
iff_empty_intro (not_lt_zero a) iff_empty_intro (not_lt_zero a)
protected theorem eq_or_lt_of_le {a b : } (H : a ≤ b) : a = b ⊎ a < b := protected theorem eq_sum_lt_of_le {a b : } (H : a ≤ b) : a = b ⊎ a < b :=
le.cases_on H (inl rfl) (λn h, inr (succ_le_succ h)) le.cases_on H (inl rfl) (λn h, inr (succ_le_succ h))
protected theorem le_of_eq_or_lt {a b : } (H : a = b ⊎ a < b) : a ≤ b := protected theorem le_of_eq_sum_lt {a b : } (H : a = b ⊎ a < b) : a ≤ b :=
sum.rec_on H !nat.le_of_eq !nat.le_of_lt sum.rec_on H !nat.le_of_eq !nat.le_of_lt
-- less-than is well-founded -- less-than is well-founded
@ -222,13 +222,13 @@ namespace nat
definition decidable_lt [instance] [priority nat.prio] : Π a b : nat, decidable (a < b) := definition decidable_lt [instance] [priority nat.prio] : Π a b : nat, decidable (a < b) :=
λ a b, decidable_le (succ a) b λ a b, decidable_le (succ a) b
protected theorem lt_or_ge (a b : ) : a < b ⊎ a ≥ b := protected theorem lt_sum_ge (a b : ) : a < b ⊎ a ≥ b :=
nat.rec (inr !zero_le) (λn, sum.rec nat.rec (inr !zero_le) (λn, sum.rec
(λh, inl (le_succ_of_le h)) (λh, inl (le_succ_of_le h))
(λh, sum.rec_on (nat.eq_or_lt_of_le h) (λe, inl (eq.subst e !nat.le_refl)) inr)) b (λh, sum.rec_on (nat.eq_sum_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 := protected definition lt_ge_by_cases {a b : } {P : Type} (H1 : a < b → P) (H2 : a ≥ b → P) : P :=
by_cases H1 (λh, H2 (sum.rec_on !nat.lt_or_ge (λa, absurd a h) (λa, a))) by_cases H1 (λh, H2 (sum.rec_on !nat.lt_sum_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 := (H3 : b < a → P) : P :=
@ -238,7 +238,7 @@ namespace nat
protected theorem lt_trichotomy (a b : ) : a < b ⊎ a = b ⊎ b < a := 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)) 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 := protected theorem eq_sum_lt_of_not_lt {a b : } (hnlt : ¬ a < b) : a = b ⊎ b < a :=
sum.rec_on (nat.lt_trichotomy a b) sum.rec_on (nat.lt_trichotomy a b)
(λ hlt, absurd hlt hnlt) (λ hlt, absurd hlt hnlt)
(λ h, h) (λ h, h)

View file

@ -155,8 +155,8 @@ reserve infixr ` ▹ `:75
/- types and type constructors -/ /- types and type constructors -/
reserve infixr ` ⊎ `:25 reserve infixr ` ⊎ `:30
reserve infixr ` × `:30 reserve infixr ` × `:35
/- arithmetic operations -/ /- arithmetic operations -/

View file

@ -28,6 +28,9 @@ namespace is_trunc
definition has_zero_trunc_index [instance] [reducible] : has_zero trunc_index := definition has_zero_trunc_index [instance] [reducible] : has_zero trunc_index :=
has_zero.mk (succ (succ minus_two)) has_zero.mk (succ (succ minus_two))
definition has_one_trunc_index [instance] [reducible] : has_one trunc_index :=
has_one.mk (succ (succ (succ minus_two)))
/- /-
notation for trunc_index is -2, -1, 0, 1, ... notation for trunc_index is -2, -1, 0, 1, ...
from 0 and up this comes from a coercion from num to trunc_index (via nat) from 0 and up this comes from a coercion from num to trunc_index (via nat)
@ -44,15 +47,17 @@ namespace is_trunc
definition leq (n m : trunc_index) : Type₀ := definition leq (n m : trunc_index) : Type₀ :=
trunc_index.rec_on n (λm, unit) (λ n p m, trunc_index.rec_on m (λ p, empty) (λ m q p, p m) p) m trunc_index.rec_on n (λm, unit) (λ n p m, trunc_index.rec_on m (λ p, empty) (λ m q p, p m) p) m
infix <= := trunc_index.leq
infix ≤ := trunc_index.leq definition has_le_trunc_index [instance] [reducible] : has_le trunc_index :=
has_le.mk leq
end trunc_index end trunc_index
infix `+2+`:65 := trunc_index.add infix `+2+`:65 := trunc_index.add
namespace trunc_index namespace trunc_index
definition succ_le_succ {n m : trunc_index} (H : n ≤ m) : n.+1 ≤ m.+1 := H definition succ_le_succ {n m : trunc_index} (H : n ≤ m) : n.+1 ≤ m.+1 := proof H qed
definition le_of_succ_le_succ {n m : trunc_index} (H : n.+1 ≤ m.+1) : n ≤ m := H definition le_of_succ_le_succ {n m : trunc_index} (H : n.+1 ≤ m.+1) : n ≤ m := proof H qed
definition minus_two_le (n : trunc_index) : -2 ≤ n := star definition minus_two_le (n : trunc_index) : -2 ≤ n := star
definition le.refl (n : trunc_index) : n ≤ n := by induction n with n IH; exact star; exact IH definition le.refl (n : trunc_index) : n ≤ n := by induction n with n IH; exact star; exact IH
definition empty_of_succ_le_minus_two {n : trunc_index} (H : n .+1 ≤ -2) : empty := H definition empty_of_succ_le_minus_two {n : trunc_index} (H : n .+1 ≤ -2) : empty := H
@ -101,6 +106,10 @@ namespace is_trunc
(n : trunc_index) [H : is_trunc (n.+1) A] (x y : A) : is_trunc n (x = y) := (n : trunc_index) [H : is_trunc (n.+1) A] (x y : A) : is_trunc n (x = y) :=
is_trunc.mk (is_trunc.to_internal (n.+1) A x y) is_trunc.mk (is_trunc.to_internal (n.+1) A x y)
definition is_trunc_eq_zero [instance] [priority 1250] [H : is_trunc 1 A] (x y : A)
: is_hset (x = y) :=
@is_trunc_eq A 0 H x y
/- contractibility -/ /- contractibility -/
definition is_contr.mk (center : A) (center_eq : Π(a : A), center = a) : is_contr A := definition is_contr.mk (center : A) (center_eq : Π(a : A), center = a) : is_contr A :=
@ -134,6 +143,9 @@ namespace is_trunc
A H A H
--in the proof the type of H is given explicitly to make it available for class inference --in the proof the type of H is given explicitly to make it available for class inference
theorem is_trunc_succ_zero [instance] [priority 950] (A : Type) [H : is_hset A] : is_trunc 1 A :=
!is_trunc_succ
theorem is_trunc_of_leq.{l} (A : Type.{l}) {n m : trunc_index} (Hnm : n ≤ m) theorem is_trunc_of_leq.{l} (A : Type.{l}) {n m : trunc_index} (Hnm : n ≤ m)
[Hn : is_trunc n A] : is_trunc m A := [Hn : is_trunc n A] : is_trunc m A :=
have base : ∀k A, k ≤ -2 → is_trunc k A → (is_trunc -2 A), from have base : ∀k A, k ≤ -2 → is_trunc k A → (is_trunc -2 A), from

View file

@ -11,13 +11,6 @@ open iff
-- Empty type -- Empty type
-- ---------- -- ----------
namespace empty
protected theorem elim {A : Type} (H : empty) : A :=
empty.rec (λe, A) H
end empty
protected definition empty.has_decidable_eq [instance] : decidable_eq empty := protected definition empty.has_decidable_eq [instance] : decidable_eq empty :=
take (a b : empty), decidable.inl (!empty.elim a) take (a b : empty), decidable.inl (!empty.elim a)
@ -48,8 +41,6 @@ end sigma
-- Sum type -- Sum type
-- -------- -- --------
infixr ⊎ := sum
infixr + := sum
namespace sum namespace sum
infixr [parsing_only] `+t`:25 := sum -- notation which is never overloaded infixr [parsing_only] `+t`:25 := sum -- notation which is never overloaded
@ -60,8 +51,6 @@ namespace sum
variables {a b c d : Type} variables {a b c d : Type}
protected definition elim (H : a ⊎ b) (f : a → c) (g : b → c) := sum.rec_on H f g
definition sum_of_sum_of_imp_of_imp (H₁ : a ⊎ b) (H₂ : a → c) (H₃ : b → d) : c ⊎ d := definition sum_of_sum_of_imp_of_imp (H₁ : a ⊎ b) (H₂ : a → c) (H₃ : b → d) : c ⊎ d :=
sum.rec_on H₁ sum.rec_on H₁
(assume Ha : a, sum.inl (H₂ Ha)) (assume Ha : a, sum.inl (H₂ Ha))
@ -81,8 +70,6 @@ end sum
-- Product type -- Product type
-- ------------ -- ------------
abbreviation pair [constructor] := @prod.mk
infixr × := prod
namespace prod namespace prod
@ -168,168 +155,9 @@ namespace prod
end prod end prod
/- logic using prod and sum -/ /- logic (ported from standard library as second half of logic file) -/
/- iff -/
variables {a b c d : Type} variables {a b c d : Type}
open prod sum unit open prod sum unit
/- prod -/
definition not_prod_of_not_left (b : Type) (Hna : ¬a) : ¬(a × b) :=
assume H : a × b, absurd (pr1 H) Hna
definition not_prod_of_not_right (a : Type) {b : Type} (Hnb : ¬b) : ¬(a × b) :=
assume H : a × b, absurd (pr2 H) Hnb
definition prod.swap (H : a × b) : b × a :=
pair (pr2 H) (pr1 H)
definition prod_of_prod_of_imp_of_imp (H₁ : a × b) (H₂ : a → c) (H₃ : b → d) : c × d :=
by cases H₁ with aa bb; exact (H₂ aa, H₃ bb)
definition prod_of_prod_of_imp_left (H₁ : a × c) (H : a → b) : b × c :=
by cases H₁ with aa cc; exact (H aa, cc)
definition prod_of_prod_of_imp_right (H₁ : c × a) (H : a → b) : c × b :=
by cases H₁ with cc aa; exact (cc, H aa)
definition prod.comm : a × b ↔ b × a :=
iff.intro (λH, prod.swap H) (λH, prod.swap H)
definition prod.assoc : (a × b) × c ↔ a × (b × c) :=
iff.intro
(assume H, pair
(pr1 (pr1 H))
(pair (pr2 (pr1 H)) (pr2 H)))
(assume H, pair
(pair (pr1 H) (pr1 (pr2 H)))
(pr2 (pr2 H)))
definition prod_unit (a : Type) : a × unit ↔ a :=
iff.intro (assume H, pr1 H) (assume H, pair H star)
definition unit_prod (a : Type) : unit × a ↔ a :=
iff.intro (assume H, pr2 H) (assume H, pair star H)
definition prod_empty (a : Type) : a × empty ↔ empty :=
iff.intro (assume H, pr2 H) (assume H, !empty.elim H)
definition empty_prod (a : Type) : empty × a ↔ empty :=
iff.intro (assume H, pr1 H) (assume H, !empty.elim H)
definition prod_self (a : Type) : a × a ↔ a :=
iff.intro (assume H, pr1 H) (assume H, pair H H)
/- sum -/
definition not_sum (Hna : ¬a) (Hnb : ¬b) : ¬(a ⊎ b) :=
assume H : a ⊎ b, sum.rec_on H
(assume Ha, absurd Ha Hna)
(assume Hb, absurd Hb Hnb)
definition sum_of_sum_of_imp_of_imp (H₁ : a ⊎ b) (H₂ : a → c) (H₃ : b → d) : c ⊎ d :=
sum.rec_on H₁
(assume Ha : a, sum.inl (H₂ Ha))
(assume Hb : b, sum.inr (H₃ Hb))
definition sum_of_sum_of_imp_left (H₁ : a ⊎ c) (H : a → b) : b ⊎ c :=
sum.rec_on H₁
(assume H₂ : a, sum.inl (H H₂))
(assume H₂ : c, sum.inr H₂)
definition sum_of_sum_of_imp_right (H₁ : c ⊎ a) (H : a → b) : c ⊎ b :=
sum.rec_on H₁
(assume H₂ : c, sum.inl H₂)
(assume H₂ : a, sum.inr (H H₂))
definition sum.elim3 (H : a ⊎ b ⊎ c) (Ha : a → d) (Hb : b → d) (Hc : c → d) : d :=
sum.rec_on H Ha (assume H₂, sum.rec_on H₂ Hb Hc)
definition sum_resolve_right (H₁ : a ⊎ b) (H₂ : ¬a) : b :=
sum.rec_on H₁ (assume Ha, absurd Ha H₂) (assume Hb, Hb)
definition sum_resolve_left (H₁ : a ⊎ b) (H₂ : ¬b) : a :=
sum.rec_on H₁ (assume Ha, Ha) (assume Hb, absurd Hb H₂)
definition sum.swap (H : a ⊎ b) : b ⊎ a :=
sum.rec_on H (assume Ha, sum.inr Ha) (assume Hb, sum.inl Hb)
definition sum.comm : a ⊎ b ↔ b ⊎ a :=
iff.intro (λH, sum.swap H) (λH, sum.swap H)
definition sum.assoc : (a ⊎ b) ⊎ c ↔ a ⊎ (b ⊎ c) :=
iff.intro
(assume H, sum.rec_on H
(assume H₁, sum.rec_on H₁
(assume Ha, sum.inl Ha)
(assume Hb, sum.inr (sum.inl Hb)))
(assume Hc, sum.inr (sum.inr Hc)))
(assume H, sum.rec_on H
(assume Ha, (sum.inl (sum.inl Ha)))
(assume H₁, sum.rec_on H₁
(assume Hb, sum.inl (sum.inr Hb))
(assume Hc, sum.inr Hc)))
definition sum_unit (a : Type) : a ⊎ unit ↔ unit :=
iff.intro (assume H, star) (assume H, sum.inr H)
definition unit_sum (a : Type) : unit ⊎ a ↔ unit :=
iff.intro (assume H, star) (assume H, sum.inl H)
definition sum_empty (a : Type) : a ⊎ empty ↔ a :=
iff.intro
(assume H, sum.rec_on H (assume H1 : a, H1) (assume H1 : empty, !empty.elim H1))
(assume H, sum.inl H)
definition empty_sum (a : Type) : empty ⊎ a ↔ a :=
iff.intro
(assume H, sum.rec_on H (assume H1 : empty, !empty.elim H1) (assume H1 : a, H1))
(assume H, sum.inr H)
definition sum_self (a : Type) : a ⊎ a ↔ a :=
iff.intro
(assume H, sum.rec_on H (assume H1, H1) (assume H1, H1))
(assume H, sum.inl H)
/- TODO
theorem sum.right_comm (a b c : Type) : (a + b) + c ↔ (a + c) + b :=
calc
(a + b) + c ↔ a + (b + c) : sum.assoc
... ↔ a + (c + b) : {sum.comm}
... ↔ (a + c) + b : iff.symm sum.assoc
theorem sum.left_comm (a b c : Type) : a + (b + c) ↔ b + (a + c) :=
calc
a + (b + c) ↔ (a + b) + c : iff.symm sum.assoc
... ↔ (b + a) + c : {sum.comm}
... ↔ b + (a + c) : sum.assoc
theorem prod.right_comm (a b c : Type) : (a × b) × c ↔ (a × c) × b :=
calc
(a × b) × c ↔ a × (b × c) : prod.assoc
... ↔ a × (c × b) : _
... ↔ (a × c) × b : iff.symm prod.assoc
theorem prod_not_self_iff {a : Type} : a × ¬ a ↔ false :=
iff.intro (assume H, (prod.right H) (prod.left H)) (assume H, false.elim H)
theorem not_prod_self_iff {a : Type} : ¬ a × a ↔ false :=
!prod.comm ▸ !prod_not_self_iff
theorem prod.left_comm [simp] (a b c : Type) : a × (b × c) ↔ b × (a × c) :=
calc
a × (b × c) ↔ (a × b) × c : iff.symm prod.assoc
... ↔ (b × a) × c : {prod.comm}
... ↔ b × (a × c) : prod.assoc
-/
theorem imp.syl (H : a → b) (H₂ : c → a) (Hc : c) : b :=
H (H₂ Hc)
theorem sum.imp_distrib : ((a + b) → c) ↔ ((a → c) × (b → c)) :=
iff.intro
(λH, prod.mk (imp.syl H sum.inl) (imp.syl H sum.inr))
(prod.rec sum.rec)
theorem not_sum_iff_not_prod_not {a b : Type} : ¬(a + b) ↔ ¬a × ¬b :=
sum.imp_distrib

View file

@ -0,0 +1,15 @@
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
-- Released under Apache 2.0 license as described in the file LICENSE.
-- Author: Leonardo de Moura, Jeremy Avigad
-- tools.helper_tactics
-- ====================
-- Useful tactics.
open tactic
namespace helper_tactics
definition apply_refl := apply eq.refl
tactic_hint apply_refl
end helper_tactics

6
hott/tools/tools.md Normal file
View file

@ -0,0 +1,6 @@
tools
=====
Various additional tools.
* [helper_tactics](helper_tactics.lean) : useful tactics

View file

@ -1,24 +1,146 @@
/- /-
Copyright (c) 2015 Floris van Doorn. All rights reserved. Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn Author: Leonardo de Moura, Floris van Doorn
Theorems about the booleans Partially ported from the standard library
-/ -/
open is_equiv eq equiv function is_trunc option unit decidable open eq eq.ops decidable
namespace bool namespace bool
local attribute bor [reducible]
local attribute band [reducible]
theorem dichotomy (b : bool) : b = ff ⊎ b = tt :=
bool.cases_on b (sum.inl rfl) (sum.inr rfl)
theorem cond_ff {A : Type} (t e : A) : cond ff t e = e :=
rfl
theorem cond_tt {A : Type} (t e : A) : cond tt t e = t :=
rfl
theorem eq_tt_of_ne_ff : Π {a : bool}, a ≠ ff → a = tt
| @eq_tt_of_ne_ff tt H := rfl
| @eq_tt_of_ne_ff ff H := absurd rfl H
theorem eq_ff_of_ne_tt : Π {a : bool}, a ≠ tt → a = ff
| @eq_ff_of_ne_tt tt H := absurd rfl H
| @eq_ff_of_ne_tt ff H := rfl
theorem absurd_of_eq_ff_of_eq_tt {B : Type} {a : bool} (H₁ : a = ff) (H₂ : a = tt) : B :=
absurd (H₁⁻¹ ⬝ H₂) ff_ne_tt
theorem tt_bor (a : bool) : bor tt a = tt :=
rfl
notation a || b := bor a b
theorem bor_tt (a : bool) : a || tt = tt :=
bool.cases_on a rfl rfl
theorem ff_bor (a : bool) : ff || a = a :=
bool.cases_on a rfl rfl
theorem bor_ff (a : bool) : a || ff = a :=
bool.cases_on a rfl rfl
theorem bor_self (a : bool) : a || a = a :=
bool.cases_on a rfl rfl
theorem bor.comm (a b : bool) : a || b = b || a :=
by cases a; repeat (cases b | reflexivity)
theorem bor.assoc (a b c : bool) : (a || b) || c = a || (b || c) :=
match a with
| ff := by rewrite *ff_bor
| tt := by rewrite *tt_bor
end
theorem or_of_bor_eq {a b : bool} : a || b = tt → a = tt ⊎ b = tt :=
bool.rec_on a
(suppose ff || b = tt,
have b = tt, from !ff_bor ▸ this,
sum.inr this)
(suppose tt || b = tt,
sum.inl rfl)
theorem bor_inl {a b : bool} (H : a = tt) : a || b = tt :=
by rewrite H
theorem bor_inr {a b : bool} (H : b = tt) : a || b = tt :=
bool.rec_on a (by rewrite H) (by rewrite H)
theorem ff_band (a : bool) : ff && a = ff :=
rfl
theorem tt_band (a : bool) : tt && a = a :=
bool.cases_on a rfl rfl
theorem band_ff (a : bool) : a && ff = ff :=
bool.cases_on a rfl rfl
theorem band_tt (a : bool) : a && tt = a :=
bool.cases_on a rfl rfl
theorem band_self (a : bool) : a && a = a :=
bool.cases_on a rfl rfl
theorem band.comm (a b : bool) : a && b = b && a :=
bool.cases_on a
(bool.cases_on b rfl rfl)
(bool.cases_on b rfl rfl)
theorem band.assoc (a b c : bool) : (a && b) && c = a && (b && c) :=
match a with
| ff := by rewrite *ff_band
| tt := by rewrite *tt_band
end
theorem band_elim_left {a b : bool} (H : a && b = tt) : a = tt :=
sum.elim (dichotomy a)
(suppose a = ff,
absurd
(calc ff = ff && b : ff_band
... = a && b : this
... = tt : H)
ff_ne_tt)
(suppose a = tt, this)
theorem band_intro {a b : bool} (H₁ : a = tt) (H₂ : b = tt) : a && b = tt :=
by rewrite [H₁, H₂]
theorem band_elim_right {a b : bool} (H : a && b = tt) : b = tt :=
band_elim_left (!band.comm ⬝ H)
theorem bnot_bnot (a : bool) : bnot (bnot a) = a :=
bool.cases_on a rfl rfl
theorem bnot_empty : bnot ff = tt :=
rfl
theorem bnot_unit : bnot tt = ff :=
rfl
theorem eq_tt_of_bnot_eq_ff {a : bool} : bnot a = ff → a = tt :=
bool.cases_on a (by contradiction) (λ h, rfl)
theorem eq_ff_of_bnot_eq_tt {a : bool} : bnot a = tt → a = ff :=
bool.cases_on a (λ h, rfl) (by contradiction)
definition bxor (x:bool) (y:bool) := cond x (bnot y) y
/- HoTT-related stuff -/
open is_equiv equiv function is_trunc option unit decidable
definition ff_ne_tt : ff = tt → empty
| [none]
definition is_equiv_bnot [constructor] [instance] [priority 500] : is_equiv bnot := definition is_equiv_bnot [constructor] [instance] [priority 500] : is_equiv bnot :=
begin begin
fapply is_equiv.mk, fapply is_equiv.mk,
exact bnot, exact bnot,
all_goals (intro b;cases b), do 6 reflexivity all_goals (intro b;cases b), do 6 reflexivity
-- all_goals (focus (intro b;cases b;all_goals reflexivity)), -- all_goals (focus (intro b;cases b;all_goals reflexivity)),
end end
definition bnot_ne : Π(b : bool), bnot b ≠ b definition bnot_ne : Π(b : bool), bnot b ≠ b
@ -43,10 +165,4 @@ namespace bool
{ intro b, cases b, reflexivity, reflexivity}, { intro b, cases b, reflexivity, reflexivity},
end end
protected definition has_decidable_eq [instance] : ∀ x y : bool, decidable (x = y)
| has_decidable_eq ff ff := inl rfl
| has_decidable_eq ff tt := inr (by contradiction)
| has_decidable_eq tt ff := inr (by contradiction)
| has_decidable_eq tt tt := inl rfl
end bool end bool

File diff suppressed because it is too large Load diff

View file

@ -6,11 +6,12 @@ Author: Floris van Doorn
Theorems about the integers specific to HoTT Theorems about the integers specific to HoTT
-/ -/
import .basic types.eq arity import .basic types.eq arity algebra.bundled
open core eq is_equiv equiv equiv.ops open core eq is_equiv equiv equiv.ops algebra is_trunc
open nat (hiding pred) open nat (hiding pred)
namespace int namespace int
section section
open algebra open algebra
definition group_integers : Group := definition group_integers : Group :=
@ -21,7 +22,7 @@ namespace int
adjointify succ pred (λa, !add_sub_cancel) (λa, !sub_add_cancel) adjointify succ pred (λa, !add_sub_cancel) (λa, !sub_add_cancel)
definition equiv_succ : := equiv.mk succ _ definition equiv_succ : := equiv.mk succ _
definition is_equiv_neg [instance] : is_equiv neg := definition is_equiv_neg [instance] : is_equiv (neg : ) :=
adjointify neg neg (λx, !neg_neg) (λa, !neg_neg) adjointify neg neg (λx, !neg_neg) (λa, !neg_neg)
definition equiv_neg : := equiv.mk neg _ definition equiv_neg : := equiv.mk neg _
@ -90,8 +91,9 @@ namespace eq
idp idp
(λn IH, idp) (λn IH, idp)
(λn IH, calc (λn IH, calc
power p (-succ n) ⬝ p = (power p (-n) ⬝ p⁻¹) ⬝ p : by rewrite [↑power,-rec_nat_on_neg] power p (-succ n) ⬝ p
... = power p (-n) : inv_con_cancel_right = (power p (-int.of_nat n) ⬝ p⁻¹) ⬝ p : by rewrite [↑power,-rec_nat_on_neg]
... = power p (-int.of_nat n) : inv_con_cancel_right
... = power p (succ (-succ n)) : by rewrite -succ_neg_succ) ... = power p (succ (-succ n)) : by rewrite -succ_neg_succ)
definition power_con_inv : power p b ⬝ p⁻¹ = power p (pred b) := definition power_con_inv : power p b ⬝ p⁻¹ = power p (pred b) :=
@ -101,7 +103,8 @@ namespace eq
power p (succ n) ⬝ p⁻¹ = power p n : by apply con_inv_cancel_right power p (succ n) ⬝ p⁻¹ = power p n : by apply con_inv_cancel_right
... = power p (pred (succ n)) : by rewrite pred_nat_succ) ... = power p (pred (succ n)) : by rewrite pred_nat_succ)
(λn IH, calc (λn IH, calc
power p (-succ n) ⬝ p⁻¹ = power p (-succ (succ n)) : by rewrite [↑power,-rec_nat_on_neg] power p (-int.of_nat (succ n)) ⬝ p⁻¹
= power p (-int.of_nat (succ (succ n))) : by rewrite [↑power,-rec_nat_on_neg]
... = power p (pred (-succ n)) : by rewrite -neg_succ) ... = power p (pred (-succ n)) : by rewrite -neg_succ)
definition con_power : p ⬝ power p b = power p (succ b) := definition con_power : p ⬝ power p b = power p (succ b) :=
@ -111,12 +114,12 @@ namespace eq
p ⬝ power p (succ n) = (p ⬝ power p n) ⬝ p : con.assoc p _ p p ⬝ power p (succ n) = (p ⬝ power p n) ⬝ p : con.assoc p _ p
... = power p (succ (succ n)) : by rewrite IH qed) ... = power p (succ (succ n)) : by rewrite IH qed)
( λn IH, calc ( λn IH, calc
p ⬝ power p (-succ n) p ⬝ power p (-int.of_nat (succ n))
= p ⬝ (power p (-n) ⬝ p⁻¹) : by rewrite [↑power,rec_nat_on_neg] = p ⬝ (power p (-int.of_nat n) ⬝ p⁻¹) : by rewrite [↑power, rec_nat_on_neg]
... = (p ⬝ power p (-n)) ⬝ p⁻¹ : con.assoc ... = (p ⬝ power p (-int.of_nat n)) ⬝ p⁻¹ : con.assoc
... = power p (succ (-n)) ⬝ p⁻¹ : by rewrite IH ... = power p (succ (-int.of_nat n)) ⬝ p⁻¹ : by rewrite IH
... = power p (pred (succ (-n))) : power_con_inv ... = power p (pred (succ (-int.of_nat n))) : power_con_inv
... = power p (succ (-succ n)) : by rewrite [succ_neg_nat_succ,int.pred_succ]) ... = power p (succ (-int.of_nat (succ n))) : by rewrite [succ_neg_nat_succ,int.pred_succ])
definition inv_con_power : p⁻¹ ⬝ power p b = power p (pred b) := definition inv_con_power : p⁻¹ ⬝ power p b = power p (pred b) :=
rec_nat_on b rec_nat_on b
@ -127,18 +130,20 @@ namespace eq
... = power p (succ (pred n)) : power_con ... = power p (succ (pred n)) : power_con
... = power p (pred (succ n)) : by rewrite [succ_pred,-int.pred_succ n]) ... = power p (pred (succ n)) : by rewrite [succ_pred,-int.pred_succ n])
( λn IH, calc ( λn IH, calc
p⁻¹ ⬝ power p (-succ n) = p⁻¹ ⬝ (power p (-n) ⬝ p⁻¹) : by rewrite [↑power,rec_nat_on_neg] p⁻¹ ⬝ power p (-int.of_nat (succ n))
... = (p⁻¹ ⬝ power p (-n)) ⬝ p⁻¹ : con.assoc = p⁻¹ ⬝ (power p (-int.of_nat n) ⬝ p⁻¹) : by rewrite [↑power,rec_nat_on_neg]
... = power p (pred (-n)) ⬝ p⁻¹ : by rewrite IH ... = (p⁻¹ ⬝ power p (-int.of_nat n)) ⬝ p⁻¹ : con.assoc
... = power p (-succ n) ⬝ p⁻¹ : by rewrite -neg_succ ... = power p (pred (-int.of_nat n)) ⬝ p⁻¹ : by rewrite IH
... = power p (-int.of_nat (succ n)) ⬝ p⁻¹ : by rewrite -neg_succ
... = power p (-succ (succ n)) : by rewrite [↑power,-rec_nat_on_neg] ... = power p (-succ (succ n)) : by rewrite [↑power,-rec_nat_on_neg]
... = power p (pred (-succ n)) : by rewrite -neg_succ) ... = power p (pred (-succ n)) : by rewrite -neg_succ)
definition power_con_power : Π(b : ), power p b ⬝ power p c = power p (b + c) := definition power_con_power : Π(b : ), power p b ⬝ power p c = power p (b + c) :=
rec_nat_on c rec_nat_on c
(λb, by rewrite int.add_zero) (λb, by rewrite int.add_zero)
(λn IH b, by rewrite [-con_power,-con.assoc,power_con,IH,↑succ,int.add.assoc,int.add.comm 1 n]) (λn IH b, by rewrite [-con_power,-con.assoc,power_con,IH,↑succ,add.assoc,
add.comm (int.of_nat n)])
(λn IH b, by rewrite [neg_nat_succ,-inv_con_power,-con.assoc,power_con_inv,IH,↑pred, (λn IH b, by rewrite [neg_nat_succ,-inv_con_power,-con.assoc,power_con_inv,IH,↑pred,
+sub_eq_add_neg,int.add.assoc,int.add.comm (-1) (-n)]) +sub_eq_add_neg,add.assoc,add.comm (-n)])
end eq end eq

View file

@ -10,7 +10,7 @@ Some lemmas are commented out, their proofs need to be repaired when needed
import .pointed .nat .pi import .pointed .nat .pi
open eq lift nat is_trunc pi pointed sum function prod option sigma open eq lift nat is_trunc pi pointed sum function prod option sigma algebra
inductive list (T : Type) : Type := inductive list (T : Type) : Type :=
| nil {} : list T | nil {} : list T

View file

@ -1,13 +1,13 @@
/- /-
Copyright (c) 2014 Floris van Doorn. All rights reserved. Copyright (c) 2014 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
(Ported from standard library file data.nat.basic on May 02, 2015) (Ported from standard library)
Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad
Basic operations on the natural numbers. Basic operations on the natural numbers.
-/ -/
import algebra.ring import ..num algebra.ring
open core prod binary open prod binary eq algebra lift is_trunc
namespace nat namespace nat
@ -17,7 +17,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
definition addl_succ_right (n m : ) : n ⊕ succ m = succ (n ⊕ m) := theorem addl_succ_right (n m : ) : n ⊕ succ m = succ (n ⊕ m) :=
nat.rec_on n nat.rec_on n
rfl rfl
(λ n₁ ih, calc (λ n₁ ih, calc
@ -25,115 +25,120 @@ nat.rec_on n
... = succ (succ (n₁ ⊕ m)) : ih ... = succ (succ (n₁ ⊕ m)) : ih
... = succ (succ n₁ ⊕ m) : rfl) ... = succ (succ n₁ ⊕ m) : rfl)
definition add_eq_addl (x : ) : ∀y, x + y = x ⊕ y := theorem add_eq_addl (x : ) : Πy, x + y = x ⊕ y :=
nat.rec_on x nat.rec_on x
(λ y, nat.rec_on y (λ y, nat.rec_on y
rfl rfl
(λ y₁ ih, calc (λ y₁ ih, calc
zero + succ y₁ = succ (zero + y₁) : rfl 0 + succ y₁ = succ (0 + y₁) : rfl
... = succ (zero ⊕ y₁) : {ih} ... = succ (0 ⊕ y₁) : {ih}
... = zero ⊕ (succ y₁) : rfl)) ... = 0 ⊕ (succ y₁) : rfl))
(λ x₁ ih₁ y, nat.rec_on y (λ x₁ ih₁ y, nat.rec_on y
(calc (calc
succ x₁ + zero = succ (x₁ + zero) : rfl succ x₁ + 0 = succ (x₁ + 0) : rfl
... = succ (x₁ ⊕ zero) : {ih₁ zero} ... = succ (x₁ ⊕ 0) : {ih₁ 0}
... = succ x₁ ⊕ zero : rfl) ... = succ x₁ ⊕ 0 : rfl)
(λ 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 prod predecessor -/
definition succ_ne_zero (n : ) : succ n ≠ 0 := theorem succ_ne_zero (n : ) : succ n ≠ 0 :=
by contradiction by contradiction
-- add_rewrite succ_ne_zero -- add_rewrite succ_ne_zero
definition pred_zero : pred 0 = 0 := theorem pred_zero [simp] : pred 0 = 0 :=
rfl rfl
definition pred_succ (n : ) : pred (succ n) = n := theorem pred_succ [simp] (n : ) : pred (succ n) = n :=
rfl rfl
definition eq_zero_or_eq_succ_pred (n : ) : n = 0 ⊎ n = succ (pred n) := theorem eq_zero_sum_eq_succ_pred (n : ) : n = 0 ⊎ n = succ (pred n) :=
nat.rec_on n nat.rec_on n
(sum.inl rfl) (sum.inl rfl)
(take m IH, sum.inr rfl) (take m IH, sum.inr
(show succ m = succ (pred (succ m)), from ap succ !pred_succ⁻¹))
definition 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 :=
sigma.mk _ (sum_resolve_right !eq_zero_or_eq_succ_pred H) sigma.mk _ (sum_resolve_right !eq_zero_sum_eq_succ_pred H)
definition succ.inj {n m : } (H : succ n = succ m) : n = m := theorem succ.inj {n m : } (H : succ n = succ m) : n = m :=
lift.down (nat.no_confusion H (λe, e)) down (nat.no_confusion H imp.id)
definition succ_ne_self {n : } : succ n ≠ n := abbreviation eq_of_succ_eq_succ := @succ.inj
theorem succ_ne_self {n : } : succ n ≠ n :=
nat.rec_on n nat.rec_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))
definition discriminate {B : Type} {n : } (H1: n = 0 → B) (H2 : ∀m, n = succ m → B) : B := theorem discriminate {B : Type} {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,
H rfl H rfl
definition two_step_induction_on {P : → Type} (a : ) (H1 : P 0) (H2 : P 1) theorem two_step_rec_on {P : → Type} (a : ) (H1 : P 0) (H2 : P 1)
(H3 : (n : ) (IH1 : P n) (IH2 : P (succ n)), P (succ (succ n))) : P a := (H3 : Π (n : ) (IH1 : P n) (IH2 : P (succ n)), P (succ (succ n))) : P a :=
have stronger : P a × P (succ a), from have stronger : P a × P (succ a), from
nat.rec_on a nat.rec_on a
(pair H1 H2) (pair H1 H2)
(take k IH, (take k IH,
have IH1 : P k, from pr1 IH, have IH1 : P k, from prod.pr1 IH,
have IH2 : P (succ k), from pr2 IH, have IH2 : P (succ k), from prod.pr2 IH,
pair IH2 (H3 k IH1 IH2)), pair IH2 (H3 k IH1 IH2)),
pr1 stronger prod.pr1 stronger
definition sub_induction {P : → Type} (n m : ) (H1 : ∀m, P 0 m) theorem sub_induction {P : → Type} (n m : ) (H1 : Πm, P 0 m)
(H2 : ∀n, P (succ n) 0) (H3 : ∀n m, P n m → P (succ n) (succ m)) : P n m := (H2 : Πn, P (succ n) 0) (H3 : Πn m, P n m → P (succ n) (succ m)) : P n m :=
have general : ∀m, P n m, from nat.rec_on n have general : Πm, P n m, from nat.rec_on n H1
(take m : , H1 m)
(take k : , (take k : ,
assume IH : m, P k m, assume IH : Πm, P k m,
take m : , take m : ,
nat.cases_on m (H2 k) (take l, (H3 k l (IH l)))), nat.cases_on m (H2 k) (take l, (H3 k l (IH l)))),
general m general m
/- addition -/ /- addition -/
definition add_zero (n : ) : n + 0 = n := protected theorem add_zero [simp] (n : ) : n + 0 = n :=
rfl rfl
definition add_succ (n m : ) : n + succ m = succ (n + m) := theorem add_succ [simp] (n m : ) : n + succ m = succ (n + m) :=
rfl rfl
definition zero_add (n : ) : 0 + n = n := protected theorem zero_add [simp] (n : ) : 0 + n = n :=
nat.rec_on n nat.rec_on n
!add_zero !nat.add_zero
(take m IH, show 0 + succ m = succ m, from (take m IH, show 0 + succ m = succ m, from
calc calc
0 + succ m = succ (0 + m) : add_succ 0 + succ m = succ (0 + m) : add_succ
... = succ m : IH) ... = succ m : IH)
definition succ_add (n m : ) : (succ n) + m = succ (n + m) := theorem succ_add [simp] (n m : ) : (succ n) + m = succ (n + m) :=
nat.rec_on m nat.rec_on m
(rfl) (!nat.add_zero ▸ !nat.add_zero)
(take k IH, eq.ap succ IH) (take k IH, calc
succ n + succ k = succ (succ n + k) : add_succ
... = succ (succ (n + k)) : IH
... = succ (n + succ k) : add_succ)
definition add.comm (n m : ) : n + m = m + n := protected theorem add_comm [simp] (n m : ) : n + m = m + n :=
nat.rec_on m nat.rec_on m
(!add_zero ⬝ !zero_add⁻¹) (by rewrite [nat.add_zero, nat.zero_add])
(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 : succ_add) ... = succ k + n : succ_add)
definition 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 :=
!succ_add ⬝ !add_succ⁻¹ !succ_add ⬝ !add_succ⁻¹
definition add.assoc (n m k : ) : (n + m) + k = n + (m + k) := protected theorem add_assoc [simp] (n m k : ) : (n + m) + k = n + (m + k) :=
nat.rec_on k nat.rec_on k
(!add_zero ▸ !add_zero) (by rewrite +nat.add_zero)
(take l IH, (take l IH,
calc calc
(n + m) + succ l = succ ((n + m) + l) : add_succ (n + m) + succ l = succ ((n + m) + l) : add_succ
@ -141,33 +146,30 @@ nat.rec_on k
... = n + succ (m + l) : add_succ ... = n + succ (m + l) : add_succ
... = n + (m + succ l) : add_succ) ... = n + (m + succ l) : add_succ)
definition 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 add.comm add.assoc n m k left_comm nat.add_comm nat.add_assoc
definition 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 add.comm add.assoc n m k right_comm nat.add_comm nat.add_assoc
theorem add.comm4 : Π {n m k l : }, n + m + (k + l) = n + k + (m + l) := protected theorem add_left_cancel {n m k : } : n + m = n + k → m = k :=
comm4 add.comm add.assoc
definition add.cancel_left {n m k : } : n + m = n + k → m = k :=
nat.rec_on n nat.rec_on n
(take H : 0 + m = 0 + k, (take H : 0 + m = 0 + k,
!zero_add⁻¹ ⬝ H ⬝ !zero_add) !nat.zero_add⁻¹ ⬝ H ⬝ !nat.zero_add)
(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 succ (n + m) = succ (n + k),
from calc from calc
succ (n + m) = succ n + m : succ_add succ (n + m) = succ n + m : succ_add
... = succ n + k : H ... = succ n + k : H
... = succ (n + k) : succ_add, ... = succ (n + k) : succ_add,
have H3 : n + m = n + k, from succ.inj H2, have n + m = n + k, from succ.inj this,
IH H3) IH this)
definition add.cancel_right {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 !add.comm ⬝ H ⬝ !add.comm, have H2 : m + n = m + k, from !nat.add_comm ⬝ H ⬝ !nat.add_comm,
add.cancel_left H2 nat.add_left_cancel H2
definition eq_zero_of_add_eq_zero_right {n m : } : n + m = 0 → n = 0 := theorem eq_zero_of_add_eq_zero_right {n m : } : n + m = 0 → n = 0 :=
nat.rec_on n nat.rec_on n
(take (H : 0 + m = 0), rfl) (take (H : 0 + m = 0), rfl)
(take k IH, (take k IH,
@ -178,99 +180,98 @@ nat.rec_on n
... = 0 : H) ... = 0 : H)
!succ_ne_zero) !succ_ne_zero)
definition 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 (!nat.add_comm ⬝ H)
definition eq_zero_and_eq_zero_of_add_eq_zero {n m : } (H : n + m = 0) : n = 0 × m = 0 := theorem eq_zero_prod_eq_zero_of_add_eq_zero {n m : } (H : n + m = 0) : n = 0 × m = 0 :=
pair (eq_zero_of_add_eq_zero_right H) (eq_zero_of_add_eq_zero_left H) pair (eq_zero_of_add_eq_zero_right H) (eq_zero_of_add_eq_zero_left H)
definition add_one (n : ) : n + 1 = succ n := theorem add_one [simp] (n : ) : n + 1 = succ n := rfl
!add_zero ▸ !add_succ
definition one_add (n : ) : 1 + n = succ n := theorem one_add (n : ) : 1 + n = succ n :=
!zero_add ▸ !succ_add !nat.zero_add ▸ !succ_add
/- multiplication -/ /- multiplication -/
definition mul_zero (n : ) : n * 0 = 0 := protected theorem mul_zero [simp] (n : ) : n * 0 = 0 :=
rfl rfl
definition mul_succ (n m : ) : n * succ m = n * m + n := theorem mul_succ [simp] (n m : ) : n * succ m = n * m + n :=
rfl rfl
-- commutativity, distributivity, associativity, identity -- commutativity, distributivity, associativity, identity
definition zero_mul (n : ) : 0 * n = 0 := protected theorem zero_mul [simp] (n : ) : 0 * n = 0 :=
nat.rec_on n nat.rec_on n
!mul_zero !nat.mul_zero
(take m IH, !mul_succ ⬝ !add_zero ⬝ IH) (take m IH, !mul_succ ⬝ !nat.add_zero ⬝ IH)
definition succ_mul (n m : ) : (succ n) * m = (n * m) + m := theorem succ_mul [simp] (n m : ) : (succ n) * m = (n * m) + m :=
nat.rec_on m nat.rec_on m
(!mul_zero ⬝ !mul_zero⁻¹ ⬝ !add_zero⁻¹) (by rewrite nat.mul_zero)
(take k IH, calc (take k IH, calc
succ n * succ k = succ n * k + succ n : mul_succ succ n * succ k = succ n * k + succ n : mul_succ
... = n * k + k + succ n : IH ... = n * k + k + succ n : IH
... = n * k + (k + succ n) : add.assoc ... = n * k + (k + succ n) : nat.add_assoc
... = n * k + (succ n + k) : add.comm ... = n * k + (succ n + k) : nat.add_comm
... = n * k + (n + succ k) : succ_add_eq_succ_add ... = n * k + (n + succ k) : succ_add_eq_succ_add
... = n * k + n + succ k : add.assoc ... = n * k + n + succ k : nat.add_assoc
... = n * succ k + succ k : mul_succ) ... = n * succ k + succ k : mul_succ)
definition mul.comm (n m : ) : n * m = m * n := protected theorem mul_comm [simp] (n m : ) : n * m = m * n :=
nat.rec_on m nat.rec_on m
(!mul_zero ⬝ !zero_mul⁻¹) (!nat.mul_zero ⬝ !nat.zero_mul⁻¹)
(take k IH, calc (take k IH, calc
n * succ k = n * k + n : mul_succ n * succ k = n * k + n : mul_succ
... = k * n + n : IH ... = k * n + n : IH
... = (succ k) * n : succ_mul) ... = (succ k) * n : succ_mul)
definition mul.right_distrib (n m k : ) : (n + m) * k = n * k + m * k := protected theorem right_distrib (n m k : ) : (n + m) * k = n * k + m * k :=
nat.rec_on k nat.rec_on k
(calc (calc
(n + m) * 0 = 0 : mul_zero (n + m) * 0 = 0 : nat.mul_zero
... = 0 + 0 : add_zero ... = 0 + 0 : nat.add_zero
... = n * 0 + 0 : mul_zero ... = n * 0 + 0 : nat.mul_zero
... = n * 0 + m * 0 : mul_zero) ... = n * 0 + m * 0 : nat.mul_zero)
(take l IH, calc (take l IH, calc
(n + m) * succ l = (n + m) * l + (n + m) : mul_succ (n + m) * succ l = (n + m) * l + (n + m) : mul_succ
... = n * l + m * l + (n + m) : IH ... = n * l + m * l + (n + m) : IH
... = n * l + m * l + n + m : add.assoc ... = n * l + m * l + n + m : nat.add_assoc
... = n * l + n + m * l + m : add.right_comm ... = n * l + n + m * l + m : nat.add_right_comm
... = n * l + n + (m * l + m) : add.assoc ... = n * l + n + (m * l + m) : nat.add_assoc
... = n * succ l + (m * l + m) : mul_succ ... = n * succ l + (m * l + m) : mul_succ
... = n * succ l + m * succ l : mul_succ) ... = n * succ l + m * succ l : mul_succ)
definition mul.left_distrib (n m k : ) : n * (m + k) = n * m + n * k := protected theorem left_distrib (n m k : ) : n * (m + k) = n * m + n * k :=
calc calc
n * (m + k) = (m + k) * n : mul.comm n * (m + k) = (m + k) * n : nat.mul_comm
... = m * n + k * n : mul.right_distrib ... = m * n + k * n : nat.right_distrib
... = n * m + k * n : mul.comm ... = n * m + k * n : nat.mul_comm
... = n * m + n * k : mul.comm ... = n * m + n * k : nat.mul_comm
definition mul.assoc (n m k : ) : (n * m) * k = n * (m * k) := protected theorem mul_assoc [simp] (n m k : ) : (n * m) * k = n * (m * k) :=
nat.rec_on k nat.rec_on k
(calc (calc
(n * m) * 0 = n * (m * 0) : mul_zero) (n * m) * 0 = n * (m * 0) : nat.mul_zero)
(take l IH, (take l IH,
calc calc
(n * m) * succ l = (n * m) * l + n * m : mul_succ (n * m) * succ l = (n * m) * l + n * m : mul_succ
... = n * (m * l) + n * m : IH ... = n * (m * l) + n * m : IH
... = n * (m * l + m) : mul.left_distrib ... = n * (m * l + m) : nat.left_distrib
... = n * (m * succ l) : mul_succ) ... = n * (m * succ l) : mul_succ)
definition mul_one (n : ) : n * 1 = n := protected theorem mul_one [simp] (n : ) : n * 1 = n :=
calc calc
n * 1 = n * 0 + n : mul_succ n * 1 = n * 0 + n : mul_succ
... = 0 + n : mul_zero ... = 0 + n : nat.mul_zero
... = n : zero_add ... = n : nat.zero_add
definition one_mul (n : ) : 1 * n = n := protected theorem one_mul [simp] (n : ) : 1 * n = n :=
calc calc
1 * n = n * 1 : mul.comm 1 * n = n * 1 : nat.mul_comm
... = n : mul_one ... = n : nat.mul_one
definition eq_zero_or_eq_zero_of_mul_eq_zero {n m : } : n * m = 0 → n = 0 ⊎ m = 0 := theorem eq_zero_sum_eq_zero_of_mul_eq_zero {n m : } : n * m = 0 → n = 0 ⊎ m = 0 :=
nat.cases_on n nat.cases_on n
(assume H, sum.inl rfl) (assume H, sum.inl rfl)
(take n', (take n',
@ -279,72 +280,38 @@ nat.cases_on n
(take m', (take m',
assume H : succ n' * succ m' = 0, assume H : succ n' * succ m' = 0,
absurd absurd
((calc (calc
0 = succ n' * succ m' : H 0 = succ n' * succ m' : H
... = succ n' * m' + succ n' : mul_succ ... = succ n' * m' + succ n' : mul_succ
... = succ (succ n' * m' + n') : add_succ)⁻¹) ... = succ (succ n' * m' + n') : add_succ)⁻¹
!succ_ne_zero)) !succ_ne_zero))
section protected definition comm_semiring [reducible] [trans_instance] : comm_semiring nat :=
open [classes] algebra ⦃comm_semiring,
add := nat.add,
protected definition comm_semiring [instance] [reducible] : algebra.comm_semiring nat := add_assoc := nat.add_assoc,
⦃algebra.comm_semiring, zero := nat.zero,
add := add, zero_add := nat.zero_add,
add_assoc := add.assoc, add_zero := nat.add_zero,
zero := zero, add_comm := nat.add_comm,
zero_add := zero_add, mul := nat.mul,
add_zero := add_zero, mul_assoc := nat.mul_assoc,
add_comm := add.comm, one := nat.succ nat.zero,
mul := mul, one_mul := nat.one_mul,
mul_assoc := mul.assoc, mul_one := nat.mul_one,
one := succ zero, left_distrib := nat.left_distrib,
one_mul := one_mul, right_distrib := nat.right_distrib,
mul_one := mul_one, zero_mul := nat.zero_mul,
left_distrib := mul.left_distrib, mul_zero := nat.mul_zero,
right_distrib := mul.right_distrib, mul_comm := nat.mul_comm,
zero_mul := zero_mul, is_hset_carrier:= _⦄
mul_zero := mul_zero,
mul_comm := mul.comm,
is_hset_carrier := is_hset_of_decidable_eq⦄
end
section port_algebra
open [classes] algebra
definition mul.left_comm : ∀a b c : , a * (b * c) = b * (a * c) := algebra.mul.left_comm
definition mul.right_comm : ∀a b c : , (a * b) * c = (a * c) * b := algebra.mul.right_comm
definition dvd (a b : ) : Type₀ := algebra.dvd a b
notation a b := dvd a b
definition dvd.intro : ∀{a b c : } (H : a * c = b), a b := @algebra.dvd.intro _ _
definition dvd.intro_left : ∀{a b c : } (H : c * a = b), a b := @algebra.dvd.intro_left _ _
definition exists_eq_mul_right_of_dvd : ∀{a b : } (H : a b), Σc, b = a * c :=
@algebra.exists_eq_mul_right_of_dvd _ _
definition dvd.elim : ∀{P : Type} {a b : } (H₁ : a b) (H₂ : ∀c, b = a * c → P), P :=
@algebra.dvd.elim _ _
definition exists_eq_mul_left_of_dvd : ∀{a b : } (H : a b), Σc, b = c * a :=
@algebra.exists_eq_mul_left_of_dvd _ _
definition dvd.elim_left : ∀{P : Type} {a b : } (H₁ : a b) (H₂ : ∀c, b = c * a → P), P :=
@algebra.dvd.elim_left _ _
definition dvd.refl : ∀a : , a a := algebra.dvd.refl
definition dvd.trans : ∀{a b c : }, a b → b c → a c := @algebra.dvd.trans _ _
definition eq_zero_of_zero_dvd : ∀{a : }, 0 a → a = 0 := @algebra.eq_zero_of_zero_dvd _ _
definition dvd_zero : ∀a : , a 0 := algebra.dvd_zero
definition one_dvd : ∀a : , 1 a := algebra.one_dvd
definition dvd_mul_right : ∀a b : , a a * b := algebra.dvd_mul_right
definition dvd_mul_left : ∀a b : , a b * a := algebra.dvd_mul_left
definition dvd_mul_of_dvd_left : ∀{a b : } (H : a b) (c : ), a b * c :=
@algebra.dvd_mul_of_dvd_left _ _
definition dvd_mul_of_dvd_right : ∀{a b : } (H : a b) (c : ), a c * b :=
@algebra.dvd_mul_of_dvd_right _ _
definition mul_dvd_mul : ∀{a b c d : }, a b → c d → a * c b * d :=
@algebra.mul_dvd_mul _ _
definition dvd_of_mul_right_dvd : ∀{a b c : }, a * b c → a c :=
@algebra.dvd_of_mul_right_dvd _ _
definition dvd_of_mul_left_dvd : ∀{a b c : }, a * b c → b c :=
@algebra.dvd_of_mul_left_dvd _ _
definition dvd_add : ∀{a b c : }, a b → a c → a b + c := @algebra.dvd_add _ _
end port_algebra
end nat end nat
section
open nat
definition iterate {A : Type} (op : A → A) : → A → A
| 0 := λ a, a
| (succ k) := λ a, op (iterate k a)
notation f`^[`n`]` := iterate f n
end

View file

@ -6,9 +6,9 @@ Author: Floris van Doorn
Theorems about the natural numbers specific to HoTT Theorems about the natural numbers specific to HoTT
-/ -/
import .basic import .order
open is_trunc unit empty eq equiv open is_trunc unit empty eq equiv algebra
namespace nat namespace nat
definition is_hprop_le [instance] (n m : ) : is_hprop (n ≤ m) := definition is_hprop_le [instance] (n m : ) : is_hprop (n ≤ m) :=
@ -25,6 +25,8 @@ namespace nat
{ exact ap le.step !v_0}}, { exact ap le.step !v_0}},
end end
definition is_hprop_lt [instance] (n m : ) : is_hprop (n < m) := !is_hprop_le
definition le_equiv_succ_le_succ (n m : ) : (n ≤ m) ≃ (succ n ≤ succ m) := definition le_equiv_succ_le_succ (n m : ) : (n ≤ m) ≃ (succ n ≤ succ m) :=
equiv_of_is_hprop succ_le_succ le_of_succ_le_succ equiv_of_is_hprop succ_le_succ le_of_succ_le_succ
definition le_succ_equiv_pred_le (n m : ) : (n ≤ succ m) ≃ (pred n ≤ m) := definition le_succ_equiv_pred_le (n m : ) : (n ≤ succ m) ≃ (pred n ≤ m) :=
@ -73,8 +75,9 @@ namespace nat
unfold [lt_ge_by_cases,lt.by_cases], induction (lt.trichotomy n m) with H' H', unfold [lt_ge_by_cases,lt.by_cases], induction (lt.trichotomy n m) with H' H',
{ esimp, apply ap H1 !is_hprop.elim}, { esimp, apply ap H1 !is_hprop.elim},
{ cases H' with H' H', { cases H' with H' H',
esimp, exact !Heq⁻¹ ⬝ ap H1 !is_hprop.elim, { esimp, induction H', esimp, symmetry,
exfalso, apply lt.irrefl, apply lt_of_le_of_lt H H'} exact ap H1 !is_hprop.elim ⬝ Heq idp ⬝ ap H2 !is_hprop.elim},
{ exfalso, apply lt.irrefl, apply lt_of_le_of_lt H H'}}
end end
protected definition code [reducible] [unfold 1 2] : → Type₀ protected definition code [reducible] [unfold 1 2] : → Type₀

View file

@ -4,180 +4,158 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad
The order relation on the natural numbers. The order relation on the natural numbers.
Note: this file has significant differences than the standard library version
-/ -/
import .basic algebra.ordered_ring import .basic algebra.ordered_ring
open prod decidable sum eq sigma sigma.ops open eq eq.ops algebra algebra
namespace nat namespace nat
/- lt and le -/ /- lt prod le -/
theorem le_of_lt_or_eq {m n : } (H : m < n ⊎ m = n) : m ≤ n := protected theorem le_of_lt_sum_eq {m n : } (H : m < n ⊎ m = n) : m ≤ n :=
sum.rec_on H (take H1, le_of_lt H1) (take H1, H1 ▸ !le.refl) nat.le_of_eq_sum_lt (sum.swap H)
theorem lt_or_eq_of_le {m n : } (H : m ≤ n) : m < n ⊎ m = n := protected theorem lt_sum_eq_of_le {m n : } (H : m ≤ n) : m < n ⊎ m = n :=
lt.by_cases sum.swap (nat.eq_sum_lt_of_le H)
(assume H1 : m < n, sum.inl H1)
(assume H1 : m = n, sum.inr H1)
(assume H1 : m > n, absurd (lt_of_le_of_lt H H1) !lt.irrefl)
theorem le_iff_lt_or_eq (m n : ) : m ≤ n ↔ m < n ⊎ m = n := protected theorem le_iff_lt_sum_eq (m n : ) : m ≤ n ↔ m < n ⊎ m = n :=
iff.intro lt_or_eq_of_le le_of_lt_or_eq iff.intro nat.lt_sum_eq_of_le nat.le_of_lt_sum_eq
theorem lt_of_le_and_ne {m n : } (H1 : m ≤ n) (H2 : m ≠ n) : m < n := protected theorem lt_of_le_prod_ne {m n : } (H1 : m ≤ n) : m ≠ n → m < n :=
sum.rec_on (lt_or_eq_of_le H1) sum_resolve_right (nat.eq_sum_lt_of_le H1)
(take H3 : m < n, H3)
(take H3 : m = n, absurd H3 H2)
theorem lt_iff_le_and_ne (m n : ) : m < n ↔ m ≤ n × m ≠ n := protected theorem lt_iff_le_prod_ne (m n : ) : m < n ↔ m ≤ n × m ≠ n :=
iff.intro iff.intro
(take H, pair (le_of_lt H) (take H1, lt.irrefl _ (H1 ▸ H))) (take H, pair (nat.le_of_lt H) (take H1, !nat.lt_irrefl (H1 ▸ H)))
(take H, lt_of_le_and_ne (pr1 H) (pr2 H)) (prod.rec nat.lt_of_le_prod_ne)
theorem le_add_right (n k : ) : n ≤ n + k := theorem le_add_right (n k : ) : n ≤ n + k :=
nat.rec_on k nat.rec !nat.le_refl (λ k, le_succ_of_le) k
(calc n ≤ n : le.refl n
... = n + zero : add_zero)
(λ k (ih : n ≤ n + k), calc
n ≤ succ (n + k) : le_succ_of_le ih
... = n + succ k : add_succ)
theorem le_add_left (n m : ): n ≤ m + n := theorem le_add_left (n m : ): n ≤ m + n :=
!add.comm ▸ !le_add_right !add.comm ▸ !le_add_right
theorem le.intro {n m k : } (h : n + k = m) : n ≤ m := theorem le.intro {n m k : } (h : n + k = m) : n ≤ m :=
h ▸ le_add_right n k h ▸ !le_add_right
theorem le.elim {n m : } (h : n ≤ m) : Σk, n + k = m := theorem le.elim {n m : } : n ≤ m → Σ k, n + k = m :=
by induction h with m h ih;exact ⟨0, idp⟩;exact ⟨succ ih.1, ap succ ih.2⟩ le.rec (sigma.mk 0 rfl) (λm h, sigma.rec
(λ k H, sigma.mk (succ k) (H ▸ rfl)))
theorem le.total {m n : } : m ≤ n ⊎ n ≤ m := protected theorem le_total {m n : } : m ≤ n ⊎ n ≤ m :=
lt.by_cases sum.imp_left nat.le_of_lt !nat.lt_sum_ge
(assume H : m < n, sum.inl (le_of_lt H))
(assume H : m = n, sum.inl (H ▸ !le.refl))
(assume H : m > n, sum.inr (le_of_lt H))
/- addition -/ /- addition -/
theorem add_le_add_left {n m : } (H : n ≤ m) (k : ) : k + n ≤ k + m := protected theorem add_le_add_left {n m : } (H : n ≤ m) (k : ) : k + n ≤ k + m :=
sigma.rec_on (le.elim H) (λ(l : ) (Hl : n + l = m), obtain l Hl, from le.elim H, le.intro (Hl ▸ !add.assoc)
le.intro
(calc
k + n + l = k + (n + l) : !add.assoc
... = k + m : {Hl}))
theorem add_le_add_right {n m : } (H : n ≤ m) (k : ) : n + k ≤ m + k := protected theorem add_le_add_right {n m : } (H : n ≤ m) (k : ) : n + k ≤ m + k :=
!add.comm ▸ !add.comm ▸ add_le_add_left H k !add.comm ▸ !add.comm ▸ nat.add_le_add_left H k
theorem le_of_add_le_add_left {k n m : } (H : k + n ≤ k + m) : n ≤ m := protected theorem le_of_add_le_add_left {k n m : } (H : k + n ≤ k + m) : n ≤ m :=
sigma.rec_on (le.elim H) (λ(l : ) (Hl : k + n + l = k + m), obtain l Hl, from le.elim H, le.intro (nat.add_left_cancel (!add.assoc⁻¹ ⬝ Hl))
le.intro (add.cancel_left
(calc
k + (n + l) = k + n + l : (!add.assoc)⁻¹
... = k + m : Hl)))
theorem add_lt_add_left {n m : } (H : n < m) (k : ) : k + n < k + m := protected theorem lt_of_add_lt_add_left {k n m : } (H : k + n < k + m) : n < m :=
lt_of_succ_le (!add_succ ▸ add_le_add_left (succ_le_of_lt H) k) let H' := nat.le_of_lt H in
nat.lt_of_le_prod_ne (nat.le_of_add_le_add_left H') (assume Heq, !nat.lt_irrefl (Heq ▸ H))
theorem add_lt_add_right {n m : } (H : n < m) (k : ) : n + k < m + k := protected theorem add_lt_add_left {n m : } (H : n < m) (k : ) : k + n < k + m :=
!add.comm ▸ !add.comm ▸ add_lt_add_left H k lt_of_succ_le (!add_succ ▸ nat.add_le_add_left (succ_le_of_lt H) k)
theorem lt_add_of_pos_right {n k : } (H : k > 0) : n < n + k := protected theorem add_lt_add_right {n m : } (H : n < m) (k : ) : n + k < m + k :=
!add_zero ▸ add_lt_add_left H n !add.comm ▸ !add.comm ▸ nat.add_lt_add_left H k
protected theorem lt_add_of_pos_right {n k : } (H : k > 0) : n < n + k :=
!add_zero ▸ nat.add_lt_add_left H n
/- multiplication -/ /- multiplication -/
theorem mul_le_mul_left {n m : } (H : n ≤ m) (k : ) : k * n ≤ k * m := theorem mul_le_mul_left {n m : } (k : ) (H : n ≤ m) : k * n ≤ k * m :=
sigma.rec_on (le.elim H) (λ(l : ) (Hl : n + l = m), obtain (l : ) (Hl : n + l = m), from le.elim H,
have H2 : k * n + k * l = k * m, by rewrite [-mul.left_distrib, Hl], have k * n + k * l = k * m, by rewrite [-left_distrib, Hl],
le.intro H2) le.intro this
theorem mul_le_mul_right {n m : } (H : n ≤ m) (k : ) : n * k ≤ m * k := theorem mul_le_mul_right {n m : } (k : ) (H : n ≤ m) : n * k ≤ m * k :=
!mul.comm ▸ !mul.comm ▸ (mul_le_mul_left H k) !mul.comm ▸ !mul.comm ▸ !mul_le_mul_left H
theorem mul_le_mul {n m k l : } (H1 : n ≤ k) (H2 : m ≤ l) : n * m ≤ k * l := protected theorem mul_le_mul {n m k l : } (H1 : n ≤ k) (H2 : m ≤ l) : n * m ≤ k * l :=
le.trans (mul_le_mul_right H1 m) (mul_le_mul_left H2 k) nat.le_trans (!nat.mul_le_mul_right H1) (!nat.mul_le_mul_left H2)
theorem mul_lt_mul_of_pos_left {n m k : } (H : n < m) (Hk : k > 0) : k * n < k * m := protected theorem mul_lt_mul_of_pos_left {n m k : } (H : n < m) (Hk : k > 0) : k * n < k * m :=
have H2 : k * n < k * n + k, from lt_add_of_pos_right Hk, 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))
have H3 : k * n + k ≤ k * m, from !mul_succ ▸ mul_le_mul_left (succ_le_of_lt H) k,
lt_of_lt_of_le H2 H3
theorem mul_lt_mul_of_pos_right {n m k : } (H : n < m) (Hk : k > 0) : n * k < m * k := protected theorem mul_lt_mul_of_pos_right {n m k : } (H : n < m) (Hk : k > 0) : n * k < m * k :=
!mul.comm ▸ !mul.comm ▸ mul_lt_mul_of_pos_left H Hk !mul.comm ▸ !mul.comm ▸ nat.mul_lt_mul_of_pos_left H Hk
/- nat is an instance of a linearly ordered semiring -/ /- nat is an instance of a linearly ordered semiring prod a lattice -/
section protected definition decidable_linear_ordered_semiring [reducible] [trans_instance] :
open [classes] algebra decidable_linear_ordered_semiring nat :=
⦃ decidable_linear_ordered_semiring, nat.comm_semiring,
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_iff_lt_sum_eq := @nat.le_iff_lt_sum_eq,
le_of_lt := @nat.le_of_lt,
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,
add_lt_add_left := @nat.add_lt_add_left,
add_le_add_left := @nat.add_le_add_left,
le_of_add_le_add_left := @nat.le_of_add_le_add_left,
zero_lt_one := zero_lt_succ 0,
mul_le_mul_of_nonneg_left := (take a b c H1 H2, nat.mul_le_mul_left c H1),
mul_le_mul_of_nonneg_right := (take a b c H1 H2, nat.mul_le_mul_right c H1),
mul_lt_mul_of_pos_left := @nat.mul_lt_mul_of_pos_left,
mul_lt_mul_of_pos_right := @nat.mul_lt_mul_of_pos_right,
decidable_lt := nat.decidable_lt ⦄
protected definition linear_ordered_semiring [instance] [reducible] : definition nat_has_dvd [reducible] [instance] [priority nat.prio] : has_dvd nat :=
algebra.linear_ordered_semiring nat := has_dvd.mk has_dvd.dvd
⦃ algebra.linear_ordered_semiring, nat.comm_semiring,
add_left_cancel := @add.cancel_left,
add_right_cancel := @add.cancel_right,
lt := lt,
le := le,
le_refl := le.refl,
le_trans := @le.trans,
le_antisymm := @le.antisymm,
le_total := @le.total,
le_iff_lt_or_eq := @le_iff_lt_or_eq,
lt_iff_le_and_ne := lt_iff_le_and_ne,
add_le_add_left := @add_le_add_left,
le_of_add_le_add_left := @le_of_add_le_add_left,
zero_ne_one := ne.symm (succ_ne_zero zero),
mul_le_mul_of_nonneg_left := (take a b c H1 H2, mul_le_mul_left H1 c),
mul_le_mul_of_nonneg_right := (take a b c H1 H2, mul_le_mul_right H1 c),
mul_lt_mul_of_pos_left := @mul_lt_mul_of_pos_left,
mul_lt_mul_of_pos_right := @mul_lt_mul_of_pos_right ⦄
variables {a b c d : nat} theorem add_pos_left {a : } (H : 0 < a) (b : ) : 0 < a + b :=
theorem ne_of_lt (lt_ab : a < b) : a ≠ b := algebra.ne_of_lt lt_ab @add_pos_of_pos_of_nonneg _ _ a b H !zero_le
theorem ne_of_gt (gt_ab : a > b) : a ≠ b := algebra.ne_of_gt gt_ab
theorem lt_of_not_le (H : ¬ a ≥ b) : a < b := algebra.lt_of_not_le H
theorem le_or_gt (a b : nat) : sum (a ≤ b) (a > b) := algebra.le_or_gt a b
theorem le_of_mul_le_mul_left (H : c * a ≤ c * b) (Hc : c > 0) : a ≤ b := algebra.le_of_mul_le_mul_left H Hc
theorem not_lt_of_le (H : a ≤ b) : ¬ b < a := algebra.not_lt_of_le H
theorem not_le_of_lt (H : a < b) : ¬ b ≤ a := algebra.not_le_of_lt H
theorem add_le_add (Hab : a ≤ b) (Hcd : c ≤ d) : a + c ≤ b + d := algebra.add_le_add Hab Hcd
theorem lt_of_add_lt_add_right (H : a + b < c + b) : a < c := algebra.lt_of_add_lt_add_right H
theorem lt_of_add_lt_add_left (H : a + b < a + c) : b < c := algebra.lt_of_add_lt_add_left H
end
section port_algebra theorem add_pos_right {a : } (H : 0 < a) (b : ) : 0 < b + a :=
open [classes] algebra by rewrite add.comm; apply add_pos_left H b
theorem add_pos_left : Π{a : }, 0 < a → Πb : , 0 < a + b :=
take a H b, @algebra.add_pos_of_pos_of_nonneg _ _ a b H !zero_le
theorem add_pos_right : Π{a : }, 0 < a → Πb : , 0 < b + a :=
take a H b, !add.comm ▸ add_pos_left H b
theorem add_eq_zero_iff_eq_zero_and_eq_zero : Π{a b : },
a + b = 0 ↔ a = 0 × b = 0 :=
take a b : ,
@algebra.add_eq_zero_iff_eq_zero_and_eq_zero_of_nonneg_of_nonneg _ _ a b !zero_le !zero_le
theorem le_add_of_le_left : Π{a b c : }, b ≤ c → b ≤ a + c :=
take a b c H, @algebra.le_add_of_nonneg_of_le _ _ a b c !zero_le H
theorem le_add_of_le_right : Π{a b c : }, b ≤ c → b ≤ c + a :=
take a b c H, @algebra.le_add_of_le_of_nonneg _ _ a b c H !zero_le
theorem lt_add_of_lt_left : Π{b c : }, b < c → Πa, b < a + c :=
take b c H a, @algebra.lt_add_of_nonneg_of_lt _ _ a b c !zero_le H
theorem lt_add_of_lt_right : Π{b c : }, b < c → Πa, b < c + a :=
take b c H a, @algebra.lt_add_of_lt_of_nonneg _ _ a b c H !zero_le
theorem lt_of_mul_lt_mul_left : Π{a b c : }, c * a < c * b → a < b :=
take a b c H, @algebra.lt_of_mul_lt_mul_left _ _ a b c H !zero_le
theorem lt_of_mul_lt_mul_right : Π{a b c : }, a * c < b * c → a < b :=
take a b c H, @algebra.lt_of_mul_lt_mul_right _ _ a b c H !zero_le
theorem pos_of_mul_pos_left : Π{a b : }, 0 < a * b → 0 < b :=
take a b H, @algebra.pos_of_mul_pos_left _ _ a b H !zero_le
theorem pos_of_mul_pos_right : Π{a b : }, 0 < a * b → 0 < a :=
take a b H, @algebra.pos_of_mul_pos_right _ _ a b H !zero_le
end port_algebra
theorem zero_le_one : 0 ≤ 1 := dec_trivial theorem add_eq_zero_iff_eq_zero_prod_eq_zero {a b : } :
theorem zero_lt_one : 0 < 1 := dec_trivial a + b = 0 ↔ a = 0 × b = 0 :=
@add_eq_zero_iff_eq_zero_prod_eq_zero_of_nonneg_of_nonneg _ _ a b !zero_le !zero_le
theorem le_add_of_le_left {a b c : } (H : b ≤ c) : b ≤ a + c :=
@le_add_of_nonneg_of_le _ _ a b c !zero_le H
theorem le_add_of_le_right {a b c : } (H : b ≤ c) : b ≤ c + a :=
@le_add_of_le_of_nonneg _ _ a b c H !zero_le
theorem lt_add_of_lt_left {b c : } (H : b < c) (a : ) : b < a + c :=
@lt_add_of_nonneg_of_lt _ _ a b c !zero_le H
theorem lt_add_of_lt_right {b c : } (H : b < c) (a : ) : b < c + a :=
@lt_add_of_lt_of_nonneg _ _ a b c H !zero_le
theorem lt_of_mul_lt_mul_left {a b c : } (H : c * a < c * b) : a < b :=
@lt_of_mul_lt_mul_left _ _ a b c H !zero_le
theorem lt_of_mul_lt_mul_right {a b c : } (H : a * c < b * c) : a < b :=
@lt_of_mul_lt_mul_right _ _ a b c H !zero_le
theorem pos_of_mul_pos_left {a b : } (H : 0 < a * b) : 0 < b :=
@pos_of_mul_pos_left _ _ a b H !zero_le
theorem pos_of_mul_pos_right {a b : } (H : 0 < a * b) : 0 < a :=
@pos_of_mul_pos_right _ _ a b H !zero_le
theorem zero_le_one : (0:nat) ≤ 1 :=
dec_star
/- properties specific to nat -/ /- properties specific to nat -/
@ -194,116 +172,95 @@ theorem eq_zero_of_le_zero {n : } (H : n ≤ 0) : n = 0 :=
obtain (k : ) (Hk : n + k = 0), from le.elim H, obtain (k : ) (Hk : n + k = 0), from le.elim H,
eq_zero_of_add_eq_zero_right Hk eq_zero_of_add_eq_zero_right Hk
/- succ and pred -/ /- succ prod pred -/
theorem le_of_lt_succ {m n : nat} : m < succ n → m ≤ n :=
le_of_succ_le_succ
theorem lt_iff_succ_le (m n : nat) : m < n ↔ succ m ≤ n := theorem lt_iff_succ_le (m n : nat) : m < n ↔ succ m ≤ n :=
iff.rfl iff.rfl
theorem lt_succ_iff_le (m n : nat) : m < succ n ↔ m ≤ n :=
iff.intro le_of_lt_succ lt_succ_of_le
theorem self_le_succ (n : ) : n ≤ succ n := theorem self_le_succ (n : ) : n ≤ succ n :=
le.intro !add_one le.intro !add_one
theorem succ_le_or_eq_of_le {n m : } (H : n ≤ m) : succ n ≤ m ⊎ n = m := theorem succ_le_sum_eq_of_le {n m : } : n ≤ m → succ n ≤ m ⊎ n = m :=
sum.rec_on (lt_or_eq_of_le H) lt_sum_eq_of_le
(assume H1 : n < m, sum.inl (succ_le_of_lt H1))
(assume H1 : n = m, sum.inr H1)
theorem pred_le_of_le_succ {n m : } : n ≤ succ m → pred n ≤ m := theorem pred_le_of_le_succ {n m : } : n ≤ succ m → pred n ≤ m :=
nat.cases_on n pred_le_pred
(assume H, !pred_zero⁻¹ ▸ zero_le m)
(take n',
assume H : succ n' ≤ succ m,
have H1 : n' ≤ m, from le_of_succ_le_succ H,
!pred_succ⁻¹ ▸ H1)
theorem succ_le_of_le_pred {n m : } : succ n ≤ m → n ≤ pred m := theorem succ_le_of_le_pred {n m : } : succ n ≤ m → n ≤ pred m :=
nat.cases_on m pred_le_pred
(assume H, absurd H !not_succ_le_zero)
(take m',
assume H : succ n ≤ succ m',
have H1 : n ≤ m', from le_of_succ_le_succ H,
!pred_succ⁻¹ ▸ H1)
theorem pred_le_pred_of_le {n m : } : n ≤ m → pred n ≤ pred m := theorem pred_le_pred_of_le {n m : } : n ≤ m → pred n ≤ pred m :=
nat.cases_on n pred_le_pred
(assume H, pred_zero⁻¹ ▸ zero_le (pred m))
(take n', theorem pre_lt_of_lt {n m : } : n < m → pred n < m :=
assume H : succ n' ≤ m, lt_of_le_of_lt !pred_le
!pred_succ⁻¹ ▸ succ_le_of_le_pred H)
theorem lt_of_pred_lt_pred {n m : } (H : pred n < pred m) : n < m := theorem lt_of_pred_lt_pred {n m : } (H : pred n < pred m) : n < m :=
lt_of_not_le lt_of_not_ge
(take H1 : m ≤ n, (suppose m ≤ n,
not_lt_of_le (pred_le_pred_of_le H1) H) not_lt_of_ge (pred_le_pred_of_le this) H)
theorem le_or_eq_succ_of_le_succ {n m : } (H : n ≤ succ m) : n ≤ m ⊎ n = succ m := theorem le_sum_eq_succ_of_le_succ {n m : } (H : n ≤ succ m) : n ≤ m ⊎ n = succ m :=
sum_of_sum_of_imp_left (succ_le_or_eq_of_le H) sum.imp_left le_of_succ_le_succ (succ_le_sum_eq_of_le H)
(take H2 : succ n ≤ succ m, show n ≤ m, from le_of_succ_le_succ H2)
theorem le_pred_self (n : ) : pred n ≤ n := theorem le_pred_self (n : ) : pred n ≤ n :=
nat.cases_on n !pred_le
(pred_zero⁻¹ ▸ !le.refl)
(take k : , (!pred_succ)⁻¹ ▸ !self_le_succ)
theorem succ_pos (n : ) : 0 < succ n := theorem succ_pos (n : ) : 0 < succ n :=
!zero_lt_succ !zero_lt_succ
theorem succ_pred_of_pos {n : } (H : n > 0) : succ (pred n) = n := theorem succ_pred_of_pos {n : } (H : n > 0) : succ (pred n) = n :=
(sum_resolve_right (eq_zero_or_eq_succ_pred n) (ne.symm (ne_of_lt H)))⁻¹ (sum_resolve_right (eq_zero_sum_eq_succ_pred n) (ne.symm (ne_of_lt H)))⁻¹
theorem exists_eq_succ_of_lt {n m : } (H : n < m) : Σk, m = succ k := theorem exists_eq_succ_of_lt {n : } : Π {m : }, n < m → Σk, m = succ k
discriminate | 0 H := absurd H !not_lt_zero
(take (Hm : m = 0), absurd (Hm ▸ H) !not_lt_zero) | (succ k) H := sigma.mk k rfl
(take (l : ) (Hm : m = succ l), sigma.mk l Hm)
theorem lt_succ_self (n : ) : n < succ n := theorem lt_succ_self (n : ) : n < succ n :=
lt.base n lt.base n
theorem le_of_lt_succ {n m : } (H : n < succ m) : n ≤ m := lemma lt_succ_of_lt {i j : nat} : i < j → i < succ j :=
le_of_succ_le_succ (succ_le_of_lt H) assume Plt, lt.trans Plt (self_lt_succ j)
/- other forms of rec -/ /- other forms of induction -/
protected theorem strong_induction_on {P : nat → Type} (n : ) (H : Πn, (Πm, m < n → P m) → P n) : protected definition strong_rec_on {P : nat → Type} (n : ) (H : Πn, (Πm, m < n → P m) → P n) : P n :=
P n := nat.rec (λm h, absurd h !not_lt_zero)
have H1 : Π {n m : nat}, m < n → P m, from (λn' (IH : Π {m : }, m < n' → P m) m l,
take n, sum.elim (lt_sum_eq_of_le (le_of_lt_succ l))
nat.rec_on n IH (λ e, eq.rec (H n' @IH) e⁻¹)) (succ n) n !lt_succ_self
(show Πm, m < 0 → P m, from take m H, absurd H !not_lt_zero)
(take n',
assume IH : Π {m : nat}, m < n' → P m,
have H2: P n', from H n' @IH,
show Πm, m < succ n' → P m, from
take m,
assume H3 : m < succ n',
sum.rec_on (lt_or_eq_of_le (le_of_lt_succ H3))
(assume H4: m < n', IH H4)
(assume H4: m = n', H4⁻¹ ▸ H2)),
H1 !lt_succ_self
protected theorem case_strong_induction_on {P : nat → Type} (a : nat) (H0 : P 0) protected theorem case_strong_rec_on {P : nat → Type} (a : nat) (H0 : P 0)
(Hind : Π(n : nat), (Πm, m ≤ n → P m) → P (succ n)) : P a := (Hind : Π(n : nat), (Πm, m ≤ n → P m) → P (succ n)) : P a :=
nat.strong_induction_on a nat.strong_rec_on a
(take n, (take n,
show (Π m, m < n → P m) → P n, from show (Π m, m < n → P m) → P n, from
nat.cases_on n nat.cases_on n
(assume H : (Πm, m < 0 → P m), show P 0, from H0) (suppose (Π m, m < 0 → P m), show P 0, from H0)
(take n, (take n,
assume H : (Πm, m < succ n → P m), suppose (Π m, m < succ n → P m),
show P (succ n), from show P (succ n), from
Hind n (take m, assume H1 : m ≤ n, H _ (lt_succ_of_le H1)))) Hind n (take m, assume H1 : m ≤ n, this _ (lt_succ_of_le H1))))
/- pos -/ /- pos -/
theorem by_cases_zero_pos {P : → Type} (y : ) (H0 : P 0) (H1 : Π {y : nat}, y > 0 → P y) : P y := theorem by_cases_zero_pos {P : → Type} (y : ) (H0 : P 0) (H1 : Π {y : nat}, y > 0 → P y) :
P y :=
nat.cases_on y H0 (take y, H1 !succ_pos) nat.cases_on y H0 (take y, H1 !succ_pos)
theorem eq_zero_or_pos (n : ) : n = 0 ⊎ n > 0 := theorem eq_zero_sum_pos (n : ) : n = 0 ⊎ n > 0 :=
sum_of_sum_of_imp_left sum_of_sum_of_imp_left
(sum.swap (lt_or_eq_of_le !zero_le)) (sum.swap (lt_sum_eq_of_le !zero_le))
(take H : 0 = n, H⁻¹) (suppose 0 = n, by subst n)
theorem pos_of_ne_zero {n : } (H : n ≠ 0) : n > 0 := theorem pos_of_ne_zero {n : } (H : n ≠ 0) : n > 0 :=
sum.rec_on !eq_zero_or_pos (take H2 : n = 0, absurd H2 H) (take H2 : n > 0, H2) sum.elim !eq_zero_sum_pos (take H2 : n = 0, by contradiction) (take H2 : n > 0, H2)
theorem ne_zero_of_pos {n : } (H : n > 0) : n ≠ 0 := theorem ne_zero_of_pos {n : } (H : n > 0) : n ≠ 0 :=
ne.symm (ne_of_lt H) ne.symm (ne_of_lt H)
@ -313,53 +270,53 @@ exists_eq_succ_of_lt H
theorem pos_of_dvd_of_pos {m n : } (H1 : m n) (H2 : n > 0) : m > 0 := theorem pos_of_dvd_of_pos {m n : } (H1 : m n) (H2 : n > 0) : m > 0 :=
pos_of_ne_zero pos_of_ne_zero
(assume H3 : m = 0, (suppose m = 0,
have H4 : n = 0, from eq_zero_of_zero_dvd (H3 ▸ H1), assert n = 0, from eq_zero_of_zero_dvd (this ▸ H1),
ne_of_lt H2 H4⁻¹) ne_of_lt H2 (by subst n))
/- multiplication -/ /- multiplication -/
theorem mul_lt_mul_of_le_of_lt {n m k l : } (Hk : k > 0) (H1 : n ≤ k) (H2 : m < l) : theorem mul_lt_mul_of_le_of_lt {n m k l : } (Hk : k > 0) (H1 : n ≤ k) (H2 : m < l) :
n * m < k * l := n * m < k * l :=
lt_of_le_of_lt (mul_le_mul_right H1 m) (mul_lt_mul_of_pos_left H2 Hk) lt_of_le_of_lt (mul_le_mul_right m H1) (mul_lt_mul_of_pos_left H2 Hk)
theorem mul_lt_mul_of_lt_of_le {n m k l : } (Hl : l > 0) (H1 : n < k) (H2 : m ≤ l) : theorem mul_lt_mul_of_lt_of_le {n m k l : } (Hl : l > 0) (H1 : n < k) (H2 : m ≤ l) :
n * m < k * l := n * m < k * l :=
lt_of_le_of_lt (mul_le_mul_left H2 n) (mul_lt_mul_of_pos_right H1 Hl) lt_of_le_of_lt (mul_le_mul_left n H2) (mul_lt_mul_of_pos_right H1 Hl)
theorem mul_lt_mul_of_le_of_le {n m k l : } (H1 : n < k) (H2 : m < l) : n * m < k * l := theorem mul_lt_mul_of_le_of_le {n m k l : } (H1 : n < k) (H2 : m < l) : n * m < k * l :=
have H3 : n * m ≤ k * m, from mul_le_mul_right (le_of_lt H1) m, have H3 : n * m ≤ k * m, from mul_le_mul_right m (le_of_lt H1),
have H4 : k * m < k * l, from mul_lt_mul_of_pos_left H2 (lt_of_le_of_lt !zero_le H1), have H4 : k * m < k * l, from mul_lt_mul_of_pos_left H2 (lt_of_le_of_lt !zero_le H1),
lt_of_le_of_lt H3 H4 lt_of_le_of_lt H3 H4
theorem eq_of_mul_eq_mul_left {m k n : } (Hn : n > 0) (H : n * m = n * k) : m = k := theorem eq_of_mul_eq_mul_left {m k n : } (Hn : n > 0) (H : n * m = n * k) : m = k :=
have H2 : n * m ≤ n * k, from H ▸ !le.refl, have n * m ≤ n * k, by rewrite H,
have H3 : n * k ≤ n * m, from H ▸ !le.refl, have m ≤ k, from le_of_mul_le_mul_left this Hn,
have H4 : m ≤ k, from le_of_mul_le_mul_left H2 Hn, have n * k ≤ n * m, by rewrite H,
have H5 : k ≤ m, from le_of_mul_le_mul_left H3 Hn, have k ≤ m, from le_of_mul_le_mul_left this Hn,
le.antisymm H4 H5 le.antisymm `m ≤ k` this
theorem eq_of_mul_eq_mul_right {n m k : } (Hm : m > 0) (H : n * m = k * m) : n = k := theorem eq_of_mul_eq_mul_right {n m k : } (Hm : m > 0) (H : n * m = k * m) : n = k :=
eq_of_mul_eq_mul_left Hm (!mul.comm ▸ !mul.comm ▸ H) eq_of_mul_eq_mul_left Hm (!mul.comm ▸ !mul.comm ▸ H)
theorem eq_zero_or_eq_of_mul_eq_mul_left {n m k : } (H : n * m = n * k) : n = 0 ⊎ m = k := theorem eq_zero_sum_eq_of_mul_eq_mul_left {n m k : } (H : n * m = n * k) : n = 0 ⊎ m = k :=
sum_of_sum_of_imp_right !eq_zero_or_pos sum_of_sum_of_imp_right !eq_zero_sum_pos
(assume Hn : n > 0, eq_of_mul_eq_mul_left Hn H) (assume Hn : n > 0, eq_of_mul_eq_mul_left Hn H)
theorem eq_zero_or_eq_of_mul_eq_mul_right {n m k : } (H : n * m = k * m) : m = 0 ⊎ n = k := theorem eq_zero_sum_eq_of_mul_eq_mul_right {n m k : } (H : n * m = k * m) : m = 0 ⊎ n = k :=
eq_zero_or_eq_of_mul_eq_mul_left (!mul.comm ▸ !mul.comm ▸ H) eq_zero_sum_eq_of_mul_eq_mul_left (!mul.comm ▸ !mul.comm ▸ H)
theorem eq_one_of_mul_eq_one_right {n m : } (H : n * m = 1) : n = 1 := theorem eq_one_of_mul_eq_one_right {n m : } (H : n * m = 1) : n = 1 :=
have H2 : n * m > 0, from H⁻¹ ▸ !succ_pos, have H2 : n * m > 0, by rewrite H; apply succ_pos,
have H3 : n > 0, from pos_of_mul_pos_right H2, sum.elim (le_sum_gt n 1)
have H4 : m > 0, from pos_of_mul_pos_left H2, (suppose n ≤ 1,
sum.rec_on (le_or_gt n 1) have n > 0, from pos_of_mul_pos_right H2,
(assume H5 : n ≤ 1, show n = 1, from le.antisymm `n ≤ 1` (succ_le_of_lt this))
show n = 1, from le.antisymm H5 (succ_le_of_lt H3)) (suppose n > 1,
(assume H5 : n > 1, have m > 0, from pos_of_mul_pos_left H2,
have H6 : n * m ≥ 2 * 1, from mul_le_mul (succ_le_of_lt H5) (succ_le_of_lt H4), have n * m ≥ 2 * 1, from nat.mul_le_mul (succ_le_of_lt `n > 1`) (succ_le_of_lt this),
have H7 : 1 ≥ 2, from !mul_one ▸ H ▸ H6, have 1 ≥ 2, from !mul_one ▸ H ▸ this,
absurd !lt_succ_self (not_lt_of_le H7)) absurd !lt_succ_self (not_lt_of_ge this))
theorem eq_one_of_mul_eq_one_left {n m : } (H : n * m = 1) : m = 1 := theorem eq_one_of_mul_eq_one_left {n m : } (H : n * m = 1) : m = 1 :=
eq_one_of_mul_eq_one_right (!mul.comm ▸ H) eq_one_of_mul_eq_one_right (!mul.comm ▸ H)
@ -372,8 +329,164 @@ eq_one_of_mul_eq_self_left Hpos (!mul.comm ▸ H)
theorem eq_one_of_dvd_one {n : } (H : n 1) : n = 1 := theorem eq_one_of_dvd_one {n : } (H : n 1) : n = 1 :=
dvd.elim H dvd.elim H
(take m, (take m, suppose 1 = n * m,
assume H1 : 1 = n * m, eq_one_of_mul_eq_one_right this⁻¹)
eq_one_of_mul_eq_one_right H1⁻¹)
/- min prod max -/
open decidable
theorem min_zero [simp] (a : ) : min a 0 = 0 :=
by rewrite [min_eq_right !zero_le]
theorem zero_min [simp] (a : ) : min 0 a = 0 :=
by rewrite [min_eq_left !zero_le]
theorem max_zero [simp] (a : ) : max a 0 = a :=
by rewrite [max_eq_left !zero_le]
theorem zero_max [simp] (a : ) : max 0 a = a :=
by rewrite [max_eq_right !zero_le]
theorem min_succ_succ [simp] (a b : ) : min (succ a) (succ b) = succ (min a b) :=
sum.elim !lt_sum_ge
(suppose a < b, by rewrite [min_eq_left_of_lt this, min_eq_left_of_lt (succ_lt_succ this)])
(suppose a ≥ b, by rewrite [min_eq_right this, min_eq_right (succ_le_succ this)])
theorem max_succ_succ [simp] (a b : ) : max (succ a) (succ b) = succ (max a b) :=
sum.elim !lt_sum_ge
(suppose a < b, by rewrite [max_eq_right_of_lt this, max_eq_right_of_lt (succ_lt_succ this)])
(suppose a ≥ b, by rewrite [max_eq_left this, max_eq_left (succ_le_succ this)])
/- In algebra.ordered_group, these next four are only proved for additive groups, not additive
semigroups. -/
protected theorem min_add_add_left (a b c : ) : min (a + b) (a + c) = a + min b c :=
decidable.by_cases
(suppose b ≤ c,
assert a + b ≤ a + c, from add_le_add_left this _,
by rewrite [min_eq_left `b ≤ c`, min_eq_left this])
(suppose ¬ b ≤ c,
assert c ≤ b, from le_of_lt (lt_of_not_ge this),
assert a + c ≤ a + b, from add_le_add_left this _,
by rewrite [min_eq_right `c ≤ b`, min_eq_right this])
protected theorem min_add_add_right (a b c : ) : min (a + c) (b + c) = min a b + c :=
by rewrite [add.comm a c, add.comm b c, add.comm _ c]; apply nat.min_add_add_left
protected theorem max_add_add_left (a b c : ) : max (a + b) (a + c) = a + max b c :=
decidable.by_cases
(suppose b ≤ c,
assert a + b ≤ a + c, from add_le_add_left this _,
by rewrite [max_eq_right `b ≤ c`, max_eq_right this])
(suppose ¬ b ≤ c,
assert c ≤ b, from le_of_lt (lt_of_not_ge this),
assert a + c ≤ a + b, from add_le_add_left this _,
by rewrite [max_eq_left `c ≤ b`, max_eq_left this])
protected theorem max_add_add_right (a b c : ) : max (a + c) (b + c) = max a b + c :=
by rewrite [add.comm a c, add.comm b c, add.comm _ c]; apply nat.max_add_add_left
/- least prod greatest -/
section least_prod_greatest
variable (P : → Type)
variable [decP : Π n, decidable (P n)]
include decP
-- returns the least i < n satisfying P, sum n if there is none
definition least :
| 0 := 0
| (succ n) := if P (least n) then least n else succ n
theorem least_of_bound {n : } (H : P n) : P (least P n) :=
begin
induction n with [m, ih],
rewrite ↑least,
apply H,
rewrite ↑least,
cases decidable.em (P (least P m)) with [Hlp, Hlp],
rewrite [if_pos Hlp],
apply Hlp,
rewrite [if_neg Hlp],
apply H
end
theorem least_le (n : ) : least P n ≤ n:=
begin
induction n with [m, ih],
{rewrite ↑least},
rewrite ↑least,
cases decidable.em (P (least P m)) with [Psm, Pnsm],
rewrite [if_pos Psm],
apply le.trans ih !le_succ,
rewrite [if_neg Pnsm]
end
theorem least_of_lt {i n : } (ltin : i < n) (H : P i) : P (least P n) :=
begin
induction n with [m, ih],
exact absurd ltin !not_lt_zero,
rewrite ↑least,
cases decidable.em (P (least P m)) with [Psm, Pnsm],
rewrite [if_pos Psm],
apply Psm,
rewrite [if_neg Pnsm],
cases (lt_sum_eq_of_le (le_of_lt_succ ltin)) with [Hlt, Heq],
exact absurd (ih Hlt) Pnsm,
rewrite Heq at H,
exact absurd (least_of_bound P H) Pnsm
end
theorem ge_least_of_lt {i n : } (ltin : i < n) (Hi : P i) : i ≥ least P n :=
begin
induction n with [m, ih],
exact absurd ltin !not_lt_zero,
rewrite ↑least,
cases decidable.em (P (least P m)) with [Psm, Pnsm],
rewrite [if_pos Psm],
cases (lt_sum_eq_of_le (le_of_lt_succ ltin)) with [Hlt, Heq],
apply ih Hlt,
rewrite Heq,
apply least_le,
rewrite [if_neg Pnsm],
cases (lt_sum_eq_of_le (le_of_lt_succ ltin)) with [Hlt, Heq],
apply absurd (least_of_lt P Hlt Hi) Pnsm,
rewrite Heq at Hi,
apply absurd (least_of_bound P Hi) Pnsm
end
theorem least_lt {n i : } (ltin : i < n) (Hi : P i) : least P n < n :=
lt_of_le_of_lt (ge_least_of_lt P ltin Hi) ltin
-- returns the largest i < n satisfying P, sum n if there is none.
definition greatest :
| 0 := 0
| (succ n) := if P n then n else greatest n
theorem greatest_of_lt {i n : } (ltin : i < n) (Hi : P i) : P (greatest P n) :=
begin
induction n with [m, ih],
{exact absurd ltin !not_lt_zero},
{cases (decidable.em (P m)) with [Psm, Pnsm],
{rewrite [↑greatest, if_pos Psm]; exact Psm},
{rewrite [↑greatest, if_neg Pnsm],
have neim : i ≠ m, from assume H : i = m, absurd (H ▸ Hi) Pnsm,
have ltim : i < m, from lt_of_le_of_ne (le_of_lt_succ ltin) neim,
apply ih ltim}}
end
theorem le_greatest_of_lt {i n : } (ltin : i < n) (Hi : P i) : i ≤ greatest P n :=
begin
induction n with [m, ih],
{exact absurd ltin !not_lt_zero},
{cases (decidable.em (P m)) with [Psm, Pnsm],
{rewrite [↑greatest, if_pos Psm], apply le_of_lt_succ ltin},
{rewrite [↑greatest, if_neg Pnsm],
have neim : i ≠ m, from assume H : i = m, absurd (H ▸ Hi) Pnsm,
have ltim : i < m, from lt_of_le_of_ne (le_of_lt_succ ltin) neim,
apply ih ltim}}
end
end least_prod_greatest
end nat end nat

View file

@ -3,27 +3,23 @@ Copyright (c) 2014 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn, Jeremy Avigad Authors: Floris van Doorn, Jeremy Avigad
Subtraction on the natural numbers, as well as min, max, and distance. Subtraction on the natural numbers, as well as min, max, prod distance.
Ported from standard library
-/ -/
import .order import .order
open eq.ops algebra eq
open core
namespace nat namespace nat
/- subtraction -/ /- subtraction -/
definition sub_zero (n : ) : n - 0 = n := protected theorem sub_zero (n : ) : n - 0 = n :=
rfl rfl
definition sub_succ (n m : ) : n - succ m = pred (n - m) := theorem sub_succ (n m : ) : n - succ m = pred (n - m) :=
rfl rfl
definition zero_sub (n : ) : 0 - n = 0 := protected theorem zero_sub (n : ) : 0 - n = 0 :=
nat.rec_on n !sub_zero nat.rec_on n !nat.sub_zero
(take k : nat, (take k : nat,
assume IH : 0 - k = 0, assume IH : 0 - k = 0,
calc calc
@ -31,13 +27,13 @@ nat.rec_on n !sub_zero
... = pred 0 : IH ... = pred 0 : IH
... = 0 : pred_zero) ... = 0 : pred_zero)
definition succ_sub_succ (n m : ) : succ n - succ m = n - m := theorem succ_sub_succ (n m : ) : succ n - succ m = n - m :=
succ_sub_succ_eq_sub n m succ_sub_succ_eq_sub n m
definition sub_self (n : ) : n - n = 0 := protected theorem sub_self (n : ) : n - n = 0 :=
nat.rec_on n !sub_zero (take k IH, !succ_sub_succ ⬝ IH) nat.rec_on n !nat.sub_zero (take k IH, !succ_sub_succ ⬝ IH)
definition add_sub_add_right (n k m : ) : (n + k) - (m + k) = n - m := protected theorem add_sub_add_right (n k m : ) : (n + k) - (m + k) = n - m :=
nat.rec_on k nat.rec_on k
(calc (calc
(n + 0) - (m + 0) = n - (m + 0) : {!add_zero} (n + 0) - (m + 0) = n - (m + 0) : {!add_zero}
@ -49,13 +45,12 @@ nat.rec_on k
... = succ (n + l) - succ (m + l) : {!add_succ} ... = succ (n + l) - succ (m + l) : {!add_succ}
... = (n + l) - (m + l) : !succ_sub_succ ... = (n + l) - (m + l) : !succ_sub_succ
... = n - m : IH) ... = n - m : IH)
protected theorem add_sub_add_left (k n m : ) : (k + n) - (k + m) = n - m :=
!add.comm ▸ !add.comm ▸ !nat.add_sub_add_right
definition add_sub_add_left (k n m : ) : (k + n) - (k + m) = n - m := protected theorem add_sub_cancel (n m : ) : n + m - m = n :=
!add.comm ▸ !add.comm ▸ !add_sub_add_right
definition add_sub_cancel (n m : ) : n + m - m = n :=
nat.rec_on m nat.rec_on m
(!add_zero⁻¹ ▸ !sub_zero) (begin rewrite add_zero end)
(take k : , (take k : ,
assume IH : n + k - k = n, assume IH : n + k - k = n,
calc calc
@ -63,13 +58,13 @@ nat.rec_on m
... = n + k - k : succ_sub_succ ... = n + k - k : succ_sub_succ
... = n : IH) ... = n : IH)
definition add_sub_cancel_left (n m : ) : n + m - n = m := protected theorem add_sub_cancel_left (n m : ) : n + m - n = m :=
!add.comm ▸ !add_sub_cancel !add.comm ▸ !nat.add_sub_cancel
definition sub_sub (n m k : ) : n - m - k = n - (m + k) := protected theorem sub_sub (n m k : ) : n - m - k = n - (m + k) :=
nat.rec_on k nat.rec_on k
(calc (calc
n - m - 0 = n - m : sub_zero n - m - 0 = n - m : nat.sub_zero
... = n - (m + 0) : add_zero) ... = n - (m + 0) : add_zero)
(take l : nat, (take l : nat,
assume IH : n - m - l = n - (m + l), assume IH : n - m - l = n - (m + l),
@ -77,60 +72,60 @@ nat.rec_on k
n - m - succ l = pred (n - m - l) : !sub_succ n - m - succ l = pred (n - m - l) : !sub_succ
... = pred (n - (m + l)) : IH ... = pred (n - (m + l)) : IH
... = n - succ (m + l) : sub_succ ... = n - succ (m + l) : sub_succ
... = n - (m + succ l) : {!add_succ⁻¹}) ... = n - (m + succ l) : by rewrite add_succ)
definition succ_sub_sub_succ (n m k : ) : succ n - m - succ k = n - m - k := theorem succ_sub_sub_succ (n m k : ) : succ n - m - succ k = n - m - k :=
calc calc
succ n - m - succ k = succ n - (m + succ k) : sub_sub succ n - m - succ k = succ n - (m + succ k) : nat.sub_sub
... = succ n - succ (m + k) : add_succ ... = succ n - succ (m + k) : add_succ
... = n - (m + k) : succ_sub_succ ... = n - (m + k) : succ_sub_succ
... = n - m - k : sub_sub ... = n - m - k : nat.sub_sub
definition sub_self_add (n m : ) : n - (n + m) = 0 := theorem sub_self_add (n m : ) : n - (n + m) = 0 :=
calc calc
n - (n + m) = n - n - m : sub_sub n - (n + m) = n - n - m : nat.sub_sub
... = 0 - m : sub_self ... = 0 - m : nat.sub_self
... = 0 : zero_sub ... = 0 : nat.zero_sub
definition sub.right_comm (m n k : ) : m - n - k = m - k - n := protected theorem sub.right_comm (m n k : ) : m - n - k = m - k - n :=
calc calc
m - n - k = m - (n + k) : !sub_sub m - n - k = m - (n + k) : !nat.sub_sub
... = m - (k + n) : {!add.comm} ... = m - (k + n) : {!add.comm}
... = m - k - n : !sub_sub⁻¹ ... = m - k - n : !nat.sub_sub⁻¹
definition sub_one (n : ) : n - 1 = pred n := theorem sub_one (n : ) : n - 1 = pred n :=
rfl rfl
definition succ_sub_one (n : ) : succ n - 1 = n := theorem succ_sub_one (n : ) : succ n - 1 = n :=
rfl rfl
/- interaction with multiplication -/ /- interaction with multiplication -/
definition mul_pred_left (n m : ) : pred n * m = n * m - m := theorem mul_pred_left (n m : ) : pred n * m = n * m - m :=
nat.rec_on n nat.rec_on n
(calc (calc
pred 0 * m = 0 * m : pred_zero pred 0 * m = 0 * m : pred_zero
... = 0 : zero_mul ... = 0 : zero_mul
... = 0 - m : zero_sub ... = 0 - m : nat.zero_sub
... = 0 * m - m : zero_mul) ... = 0 * m - m : zero_mul)
(take k : nat, (take k : nat,
assume IH : pred k * m = k * m - m, assume IH : pred k * m = k * m - m,
calc calc
pred (succ k) * m = k * m : pred_succ pred (succ k) * m = k * m : pred_succ
... = k * m + m - m : add_sub_cancel ... = k * m + m - m : nat.add_sub_cancel
... = succ k * m - m : succ_mul) ... = succ k * m - m : succ_mul)
definition mul_pred_right (n m : ) : n * pred m = n * m - n := theorem mul_pred_right (n m : ) : n * pred m = n * m - n :=
calc calc
n * pred m = pred m * n : mul.comm n * pred m = pred m * n : mul.comm
... = m * n - n : mul_pred_left ... = m * n - n : mul_pred_left
... = n * m - n : mul.comm ... = n * m - n : mul.comm
definition mul_sub_right_distrib (n m k : ) : (n - m) * k = n * k - m * k := protected theorem mul_sub_right_distrib (n m k : ) : (n - m) * k = n * k - m * k :=
nat.rec_on m nat.rec_on m
(calc (calc
(n - 0) * k = n * k : sub_zero (n - 0) * k = n * k : nat.sub_zero
... = n * k - 0 : sub_zero ... = n * k - 0 : nat.sub_zero
... = n * k - 0 * k : zero_mul) ... = n * k - 0 * k : zero_mul)
(take l : nat, (take l : nat,
assume IH : (n - l) * k = n * k - l * k, assume IH : (n - l) * k = n * k - l * k,
@ -138,26 +133,27 @@ nat.rec_on m
(n - succ l) * k = pred (n - l) * k : sub_succ (n - succ l) * k = pred (n - l) * k : sub_succ
... = (n - l) * k - k : mul_pred_left ... = (n - l) * k - k : mul_pred_left
... = n * k - l * k - k : IH ... = n * k - l * k - k : IH
... = n * k - (l * k + k) : sub_sub ... = n * k - (l * k + k) : nat.sub_sub
... = n * k - (succ l * k) : succ_mul) ... = n * k - (succ l * k) : succ_mul)
definition mul_sub_left_distrib (n m k : ) : n * (m - k) = n * m - n * k := protected theorem mul_sub_left_distrib (n m k : ) : n * (m - k) = n * m - n * k :=
calc calc
n * (m - k) = (m - k) * n : !mul.comm n * (m - k) = (m - k) * n : !mul.comm
... = m * n - k * n : !mul_sub_right_distrib ... = m * n - k * n : !nat.mul_sub_right_distrib
... = n * m - k * n : {!mul.comm} ... = n * m - k * n : {!mul.comm}
... = n * m - n * k : {!mul.comm} ... = n * m - n * k : {!mul.comm}
definition mul_self_sub_mul_self_eq (a b : nat) : a * a - b * b = (a + b) * (a - b) := protected theorem mul_self_sub_mul_self_eq (a b : nat) : a * a - b * b = (a + b) * (a - b) :=
by rewrite [mul_sub_left_distrib, *mul.right_distrib, mul.comm b a, add.comm (a*a) (a*b), add_sub_add_left] by rewrite [nat.mul_sub_left_distrib, *right_distrib, mul.comm b a, add.comm (a*a) (a*b),
nat.add_sub_add_left]
definition succ_mul_succ_eq (a : nat) : succ a * succ a = a*a + a + a + 1 := theorem succ_mul_succ_eq (a : nat) : succ a * succ a = a*a + a + a + 1 :=
calc succ a * succ a = (a+1)*(a+1) : by rewrite [add_one] calc succ a * succ a = (a+1)*(a+1) : by rewrite [add_one]
... = a*a + a + a + 1 : by rewrite [mul.right_distrib, mul.left_distrib, one_mul, mul_one] ... = a*a + a + a + 1 : by rewrite [right_distrib, left_distrib, one_mul, mul_one]
/- interaction with inequalities -/ /- interaction with inequalities -/
definition succ_sub {m n : } : m ≥ n → succ m - n = succ (m - n) := theorem succ_sub {m n : } : m ≥ n → succ m - n = succ (m - n) :=
sub_induction n m sub_induction n m
(take k, assume H : 0 ≤ k, rfl) (take k, assume H : 0 ≤ k, rfl)
(take k, (take k,
@ -171,16 +167,16 @@ sub_induction n m
... = succ (l - k) : IH (le_of_succ_le_succ H) ... = succ (l - k) : IH (le_of_succ_le_succ H)
... = succ (succ l - succ k) : succ_sub_succ) ... = succ (succ l - succ k) : succ_sub_succ)
definition sub_eq_zero_of_le {n m : } (H : n ≤ m) : n - m = 0 := theorem sub_eq_zero_of_le {n m : } (H : n ≤ m) : n - m = 0 :=
obtain (k : ) (Hk : n + k = m), from le.elim H, Hk ▸ !sub_self_add obtain (k : ) (Hk : n + k = m), from le.elim H, Hk ▸ !sub_self_add
definition add_sub_of_le {n m : } : n ≤ m → n + (m - n) = m := theorem add_sub_of_le {n m : } : n ≤ m → n + (m - n) = m :=
sub_induction n m sub_induction n m
(take k, (take k,
assume H : 0 ≤ k, assume H : 0 ≤ k,
calc calc
0 + (k - 0) = k - 0 : zero_add 0 + (k - 0) = k - 0 : zero_add
... = k : sub_zero) ... = k : nat.sub_zero)
(take k, assume H : succ k ≤ 0, absurd H !not_succ_le_zero) (take k, assume H : succ k ≤ 0, absurd H !not_succ_le_zero)
(take k l, (take k l,
assume IH : k ≤ l → k + (l - k) = l, assume IH : k ≤ l → k + (l - k) = l,
@ -190,38 +186,38 @@ sub_induction n m
... = succ (k + (l - k)) : succ_add ... = succ (k + (l - k)) : succ_add
... = succ l : IH (le_of_succ_le_succ H)) ... = succ l : IH (le_of_succ_le_succ H))
definition 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 :=
calc calc
n + (m - n) = n + 0 : sub_eq_zero_of_le H n + (m - n) = n + 0 : sub_eq_zero_of_le H
... = n : add_zero ... = n : add_zero
definition sub_add_cancel {n m : } : n ≥ m → n - m + m = n := protected theorem sub_add_cancel {n m : } : n ≥ m → n - m + m = n :=
!add.comm ▸ !add_sub_of_le !add.comm ▸ !add_sub_of_le
definition sub_add_of_le {n m : } : n ≤ m → n - m + m = m := theorem sub_add_of_le {n m : } : n ≤ m → n - m + m = m :=
!add.comm ▸ add_sub_of_ge !add.comm ▸ add_sub_of_ge
definition sub.cases {P : → Type} {n m : } (H1 : n ≤ m → P 0) (H2 : Πk, m + k = n -> P k) theorem sub.cases {P : → Type} {n m : } (H1 : n ≤ m → P 0) (H2 : Πk, m + k = n -> P k)
: P (n - m) := : P (n - m) :=
sum.rec_on !le.total sum.elim !le.total
(assume H3 : n ≤ m, (sub_eq_zero_of_le H3)⁻¹ ▸ (H1 H3)) (assume H3 : n ≤ m, (sub_eq_zero_of_le H3)⁻¹ ▸ (H1 H3))
(assume H3 : m ≤ n, H2 (n - m) (add_sub_of_le H3)) (assume H3 : m ≤ n, H2 (n - m) (add_sub_of_le H3))
definition exists_sub_eq_of_le {n m : } (H : n ≤ m) : Σk, m - k = n := theorem exists_sub_eq_of_le {n m : } (H : n ≤ m) : Σk, m - k = n :=
obtain (k : ) (Hk : n + k = m), from le.elim H, obtain (k : ) (Hk : n + k = m), from le.elim H,
sigma.mk k sigma.mk k
(calc (calc
m - k = n + k - k : Hk m - k = n + k - k : by rewrite Hk
... = n : add_sub_cancel) ... = n : nat.add_sub_cancel)
definition add_sub_assoc {m k : } (H : k ≤ m) (n : ) : n + m - k = n + (m - k) := protected theorem add_sub_assoc {m k : } (H : k ≤ m) (n : ) : n + m - k = n + (m - k) :=
have l1 : k ≤ m → n + m - k = n + (m - k), from have l1 : k ≤ m → n + m - k = n + (m - k), from
sub_induction k m sub_induction k m
(take m : , (take m : ,
assume H : 0 ≤ m, assume H : 0 ≤ m,
calc calc
n + m - 0 = n + m : sub_zero n + m - 0 = n + m : nat.sub_zero
... = n + (m - 0) : sub_zero) ... = n + (m - 0) : nat.sub_zero)
(take k : , assume H : succ k ≤ 0, absurd H !not_succ_le_zero) (take k : , assume H : succ k ≤ 0, absurd H !not_succ_le_zero)
(take k m, (take k m,
assume IH : k ≤ m → n + m - k = n + (m - k), assume IH : k ≤ m → n + m - k = n + (m - k),
@ -233,7 +229,7 @@ have l1 : k ≤ m → n + m - k = n + (m - k), from
... = n + (succ m - succ k) : succ_sub_succ), ... = n + (succ m - succ k) : succ_sub_succ),
l1 H l1 H
definition le_of_sub_eq_zero {n m : } : n - m = 0 → n ≤ m := theorem le_of_sub_eq_zero {n m : } : n - m = 0 → n ≤ m :=
sub.cases sub.cases
(assume H1 : n ≤ m, assume H2 : 0 = 0, H1) (assume H1 : n ≤ m, assume H2 : 0 = 0, H1)
(take k : , (take k : ,
@ -242,36 +238,42 @@ sub.cases
have H3 : n = m, from !add_zero ▸ H2 ▸ H1⁻¹, have H3 : n = m, from !add_zero ▸ H2 ▸ H1⁻¹,
H3 ▸ !le.refl) H3 ▸ !le.refl)
definition sub_sub.cases {P : → Type} {n m : } (H1 : Πk, n = m + k -> P k 0) theorem sub_sub.cases {P : → Type} {n m : } (H1 : Πk, n = m + k -> P k 0)
(H2 : Πk, m = n + k → P 0 k) : P (n - m) (m - n) := (H2 : Πk, m = n + k → P 0 k) : P (n - m) (m - n) :=
sum.rec_on !le.total sum.elim !le.total
(assume H3 : n ≤ m, (assume H3 : n ≤ m,
(sub_eq_zero_of_le H3)⁻¹ ▸ (H2 (m - n) (add_sub_of_le H3)⁻¹)) (sub_eq_zero_of_le H3)⁻¹ ▸ (H2 (m - n) (add_sub_of_le H3)⁻¹))
(assume H3 : m ≤ n, (assume H3 : m ≤ n,
(sub_eq_zero_of_le H3)⁻¹ ▸ (H1 (n - m) (add_sub_of_le H3)⁻¹)) (sub_eq_zero_of_le H3)⁻¹ ▸ (H1 (n - m) (add_sub_of_le H3)⁻¹))
definition sub_eq_of_add_eq {n m k : } (H : n + m = k) : k - n = m := protected theorem sub_eq_of_add_eq {n m k : } (H : n + m = k) : k - n = m :=
have H2 : k - n + n = m + n, from have H2 : k - n + n = m + n, from
calc calc
k - n + n = k : sub_add_cancel (le.intro H) k - n + n = k : nat.sub_add_cancel (le.intro H)
... = n + m : H⁻¹ ... = n + m : H⁻¹
... = m + n : !add.comm, ... = m + n : !add.comm,
add.cancel_right H2 add.right_cancel H2
definition sub_le_sub_right {n m : } (H : n ≤ m) (k : ) : n - k ≤ m - k := protected theorem eq_sub_of_add_eq {a b c : } (H : a + c = b) : a = b - c :=
(nat.sub_eq_of_add_eq (!add.comm ▸ H))⁻¹
protected theorem sub_eq_of_eq_add {a b c : } (H : a = c + b) : a - b = c :=
nat.sub_eq_of_add_eq (!add.comm ▸ H⁻¹)
protected theorem sub_le_sub_right {n m : } (H : n ≤ m) (k : ) : n - k ≤ m - k :=
obtain (l : ) (Hl : n + l = m), from le.elim H, obtain (l : ) (Hl : n + l = m), from le.elim H,
sum.rec_on !le.total sum.elim !le.total
(assume H2 : n ≤ k, (sub_eq_zero_of_le H2)⁻¹ ▸ !zero_le) (assume H2 : n ≤ k, (sub_eq_zero_of_le H2)⁻¹ ▸ !zero_le)
(assume H2 : k ≤ n, (assume H2 : k ≤ n,
have H3 : n - k + l = m - k, from have H3 : n - k + l = m - k, from
calc calc
n - k + l = l + (n - k) : add.comm n - k + l = l + (n - k) : add.comm
... = l + n - k : add_sub_assoc H2 l ... = l + n - k : nat.add_sub_assoc H2 l
... = n + l - k : add.comm ... = n + l - k : add.comm
... = m - k : Hl, ... = m - k : Hl,
le.intro H3) le.intro H3)
definition sub_le_sub_left {n m : } (H : n ≤ m) (k : ) : k - m ≤ k - n := protected theorem sub_le_sub_left {n m : } (H : n ≤ m) (k : ) : k - m ≤ k - n :=
obtain (l : ) (Hl : n + l = m), from le.elim H, obtain (l : ) (Hl : n + l = m), from le.elim H,
sub.cases sub.cases
(assume H2 : k ≤ m, !zero_le) (assume H2 : k ≤ m, !zero_le)
@ -285,42 +287,42 @@ sub.cases
... = n + l + m' : add.assoc ... = n + l + m' : add.assoc
... = m + m' : Hl ... = m + m' : Hl
... = k : Hm ... = k : Hm
... = k - n + n : sub_add_cancel H3, ... = k - n + n : nat.sub_add_cancel H3,
le.intro (add.cancel_right H4)) le.intro (add.right_cancel H4))
definition sub_pos_of_lt {m n : } (H : m < n) : n - m > 0 := protected theorem sub_pos_of_lt {m n : } (H : m < n) : n - m > 0 :=
have H1 : n = n - m + m, from (sub_add_cancel (le_of_lt H))⁻¹, assert H1 : n = n - m + m, from (nat.sub_add_cancel (le_of_lt H))⁻¹,
have H2 : 0 + m < n - m + m, from (zero_add m)⁻¹ ▸ H1 ▸ H, have H2 : 0 + m < n - m + m, begin rewrite [zero_add, -H1], exact H end,
!lt_of_add_lt_add_right H2 !lt_of_add_lt_add_right H2
definition lt_of_sub_pos {m n : } (H : n - m > 0) : m < n := protected theorem lt_of_sub_pos {m n : } (H : n - m > 0) : m < n :=
lt_of_not_le lt_of_not_ge
(take H1 : m ≥ n, (take H1 : m ≥ n,
have H2 : n - m = 0, from sub_eq_zero_of_le H1, have H2 : n - m = 0, from sub_eq_zero_of_le H1,
!lt.irrefl (H2 ▸ H)) !lt.irrefl (H2 ▸ H))
definition lt_of_sub_lt_sub_right {n m k : } (H : n - k < m - k) : n < m := protected theorem lt_of_sub_lt_sub_right {n m k : } (H : n - k < m - k) : n < m :=
lt_of_not_le lt_of_not_ge
(assume H1 : m ≤ n, (assume H1 : m ≤ n,
have H2 : m - k ≤ n - k, from sub_le_sub_right H1 _, have H2 : m - k ≤ n - k, from nat.sub_le_sub_right H1 _,
not_le_of_lt H H2) not_le_of_gt H H2)
definition lt_of_sub_lt_sub_left {n m k : } (H : n - m < n - k) : k < m := protected theorem lt_of_sub_lt_sub_left {n m k : } (H : n - m < n - k) : k < m :=
lt_of_not_le lt_of_not_ge
(assume H1 : m ≤ k, (assume H1 : m ≤ k,
have H2 : n - k ≤ n - m, from sub_le_sub_left H1 _, have H2 : n - k ≤ n - m, from nat.sub_le_sub_left H1 _,
not_le_of_lt H H2) not_le_of_gt H H2)
definition sub_lt_sub_add_sub (n m k : ) : n - k ≤ (n - m) + (m - k) := protected theorem sub_lt_sub_add_sub (n m k : ) : n - k ≤ (n - m) + (m - k) :=
sub.cases sub.cases
(assume H : n ≤ m, !zero_add⁻¹ ▸ sub_le_sub_right H k) (assume H : n ≤ m, !zero_add⁻¹ ▸ nat.sub_le_sub_right H k)
(take mn : , (take mn : ,
assume Hmn : m + mn = n, assume Hmn : m + mn = n,
sub.cases sub.cases
(assume H : m ≤ k, (assume H : m ≤ k,
have H2 : n - k ≤ n - m, from sub_le_sub_left H n, have H2 : n - k ≤ n - m, from nat.sub_le_sub_left H n,
have H3 : n - k ≤ mn, from sub_eq_of_add_eq Hmn ▸ H2, assert H3 : n - k ≤ mn, from nat.sub_eq_of_add_eq Hmn ▸ H2,
show n - k ≤ mn + 0, from !add_zero⁻¹ ▸ H3) show n - k ≤ mn + 0, begin rewrite add_zero, assumption end)
(take km : , (take km : ,
assume Hkm : k + km = m, assume Hkm : k + km = m,
have H : k + (mn + km) = n, from have H : k + (mn + km) = n, from
@ -329,10 +331,10 @@ sub.cases
... = k + km + mn : add.assoc ... = k + km + mn : add.assoc
... = m + mn : Hkm ... = m + mn : Hkm
... = n : Hmn, ... = n : Hmn,
have H2 : n - k = mn + km, from sub_eq_of_add_eq H, have H2 : n - k = mn + km, from nat.sub_eq_of_add_eq H,
H2 ▸ !le.refl)) H2 ▸ !le.refl))
definition sub_lt_self {m n : } (H1 : m > 0) (H2 : n > 0) : m - n < m := protected theorem sub_lt_self {m n : } (H1 : m > 0) (H2 : n > 0) : m - n < m :=
calc calc
m - n = succ (pred m) - n : succ_pred_of_pos H1 m - n = succ (pred m) - n : succ_pred_of_pos H1
... = succ (pred m) - succ (pred n) : succ_pred_of_pos H2 ... = succ (pred m) - succ (pred n) : succ_pred_of_pos H2
@ -341,127 +343,160 @@ calc
... < succ (pred m) : lt_succ_self ... < succ (pred m) : lt_succ_self
... = m : succ_pred_of_pos H1 ... = m : succ_pred_of_pos H1
definition le_sub_of_add_le {m n k : } (H : m + k ≤ n) : m ≤ n - k := protected theorem le_sub_of_add_le {m n k : } (H : m + k ≤ n) : m ≤ n - k :=
calc calc
m = m + k - k : add_sub_cancel m = m + k - k : nat.add_sub_cancel
... ≤ n - k : sub_le_sub_right H k ... ≤ n - k : nat.sub_le_sub_right H k
definition lt_sub_of_add_lt {m n k : } (H : m + k < n) (H2 : k ≤ n) : m < n - k := protected theorem lt_sub_of_add_lt {m n k : } (H : m + k < n) (H2 : k ≤ n) : m < n - k :=
lt_of_succ_le (le_sub_of_add_le (calc lt_of_succ_le (nat.le_sub_of_add_le (calc
succ m + k = succ (m + k) : succ_add_eq_succ_add succ m + k = succ (m + k) : succ_add_eq_succ_add
... ≤ n : succ_le_of_lt H)) ... ≤ n : succ_le_of_lt H))
protected theorem sub_lt_of_lt_add {v n m : nat} (h₁ : v < n + m) (h₂ : n ≤ v) : v - n < m :=
have succ v ≤ n + m, from succ_le_of_lt h₁,
have succ (v - n) ≤ m, from
calc succ (v - n) = succ v - n : succ_sub h₂
... ≤ n + m - n : nat.sub_le_sub_right this n
... = m : nat.add_sub_cancel_left,
lt_of_succ_le this
/- distance -/ /- distance -/
definition dist [reducible] (n m : ) := (n - m) + (m - n) definition dist [reducible] (n m : ) := (n - m) + (m - n)
definition dist.comm (n m : ) : dist n m = dist m n := theorem dist.comm (n m : ) : dist n m = dist m n :=
!add.comm !add.comm
definition dist_self (n : ) : dist n n = 0 := theorem dist_self (n : ) : dist n n = 0 :=
calc calc
(n - n) + (n - n) = 0 + (n - n) : sub_self (n - n) + (n - n) = 0 + (n - n) : nat.sub_self
... = 0 + 0 : sub_self ... = 0 + 0 : nat.sub_self
... = 0 : rfl ... = 0 : rfl
definition eq_of_dist_eq_zero {n m : } (H : dist n m = 0) : n = m := theorem eq_of_dist_eq_zero {n m : } (H : dist n m = 0) : n = m :=
have H2 : n - m = 0, from eq_zero_of_add_eq_zero_right H, have H2 : n - m = 0, from eq_zero_of_add_eq_zero_right H,
have H3 : n ≤ m, from le_of_sub_eq_zero H2, have H3 : n ≤ m, from le_of_sub_eq_zero H2,
have H4 : m - n = 0, from eq_zero_of_add_eq_zero_left H, have H4 : m - n = 0, from eq_zero_of_add_eq_zero_left H,
have H5 : m ≤ n, from le_of_sub_eq_zero H4, have H5 : m ≤ n, from le_of_sub_eq_zero H4,
le.antisymm H3 H5 le.antisymm H3 H5
definition dist_eq_sub_of_le {n m : } (H : n ≤ m) : dist n m = m - n := theorem dist_eq_zero {n m : } (H : n = m) : dist n m = 0 :=
by substvars; rewrite [↑dist, *nat.sub_self, add_zero]
theorem dist_eq_sub_of_le {n m : } (H : n ≤ m) : dist n m = m - n :=
calc calc
dist n m = 0 + (m - n) : {sub_eq_zero_of_le H} dist n m = 0 + (m - n) : {sub_eq_zero_of_le H}
... = m - n : zero_add ... = m - n : zero_add
definition dist_eq_sub_of_ge {n m : } (H : n ≥ m) : dist n m = n - m := theorem dist_eq_sub_of_lt {n m : } (H : n < m) : dist n m = m - n :=
dist_eq_sub_of_le (le_of_lt H)
theorem dist_eq_sub_of_ge {n m : } (H : n ≥ m) : dist n m = n - m :=
!dist.comm ▸ dist_eq_sub_of_le H !dist.comm ▸ dist_eq_sub_of_le H
definition dist_zero_right (n : ) : dist n 0 = n := theorem dist_eq_sub_of_gt {n m : } (H : n > m) : dist n m = n - m :=
dist_eq_sub_of_ge !zero_le ⬝ !sub_zero dist_eq_sub_of_ge (le_of_lt H)
definition dist_zero_left (n : ) : dist 0 n = n := theorem dist_zero_right (n : ) : dist n 0 = n :=
dist_eq_sub_of_le !zero_le ⬝ !sub_zero dist_eq_sub_of_ge !zero_le ⬝ !nat.sub_zero
definition dist.intro {n m k : } (H : n + m = k) : dist k n = m := theorem dist_zero_left (n : ) : dist 0 n = n :=
dist_eq_sub_of_le !zero_le ⬝ !nat.sub_zero
theorem dist.intro {n m k : } (H : n + m = k) : dist k n = m :=
calc calc
dist k n = k - n : dist_eq_sub_of_ge (le.intro H) dist k n = k - n : dist_eq_sub_of_ge (le.intro H)
... = m : sub_eq_of_add_eq H ... = m : nat.sub_eq_of_add_eq H
definition dist_add_add_right (n k m : ) : dist (n + k) (m + k) = dist n m := theorem dist_add_add_right (n k m : ) : dist (n + k) (m + k) = dist n m :=
calc calc
dist (n + k) (m + k) = ((n+k) - (m+k)) + ((m+k)-(n+k)) : rfl dist (n + k) (m + k) = ((n+k) - (m+k)) + ((m+k)-(n+k)) : rfl
... = (n - m) + ((m + k) - (n + k)) : add_sub_add_right ... = (n - m) + ((m + k) - (n + k)) : nat.add_sub_add_right
... = (n - m) + (m - n) : add_sub_add_right ... = (n - m) + (m - n) : nat.add_sub_add_right
definition dist_add_add_left (k n m : ) : dist (k + n) (k + m) = dist n m := theorem dist_add_add_left (k n m : ) : dist (k + n) (k + m) = dist n m :=
!add.comm ▸ !add.comm ▸ !dist_add_add_right begin rewrite [add.comm k n, add.comm k m]; apply dist_add_add_right end
definition dist_add_eq_of_ge {n m : } (H : n ≥ m) : dist n m + m = n := theorem dist_add_eq_of_ge {n m : } (H : n ≥ m) : dist n m + m = n :=
calc calc
dist n m + m = n - m + m : {dist_eq_sub_of_ge H} dist n m + m = n - m + m : {dist_eq_sub_of_ge H}
... = n : sub_add_cancel H ... = n : nat.sub_add_cancel H
definition dist_eq_intro {n m k l : } (H : n + m = k + l) : dist n k = dist l m := theorem dist_eq_intro {n m k l : } (H : n + m = k + l) : dist n k = dist l m :=
calc calc
dist n k = dist (n + m) (k + m) : dist_add_add_right dist n k = dist (n + m) (k + m) : dist_add_add_right
... = dist (k + l) (k + m) : H ... = dist (k + l) (k + m) : H
... = dist l m : dist_add_add_left ... = dist l m : dist_add_add_left
definition dist_sub_eq_dist_add_left {n m : } (H : n ≥ m) (k : ) : theorem dist_sub_eq_dist_add_left {n m : } (H : n ≥ m) (k : ) :
dist (n - m) k = dist n (k + m) := dist (n - m) k = dist n (k + m) :=
have H2 : n - m + (k + m) = k + n, from have H2 : n - m + (k + m) = k + n, from
calc calc
n - m + (k + m) = n - m + (m + k) : add.comm n - m + (k + m) = n - m + (m + k) : add.comm
... = n - m + m + k : add.assoc ... = n - m + m + k : add.assoc
... = n + k : sub_add_cancel H ... = n + k : nat.sub_add_cancel H
... = k + n : add.comm, ... = k + n : add.comm,
dist_eq_intro H2 dist_eq_intro H2
definition dist_sub_eq_dist_add_right {k m : } (H : k ≥ m) (n : ) : theorem dist_sub_eq_dist_add_right {k m : } (H : k ≥ m) (n : ) :
dist n (k - m) = dist (n + m) k := dist n (k - m) = dist (n + m) k :=
(dist_sub_eq_dist_add_left H n ▸ !dist.comm) ▸ !dist.comm (dist_sub_eq_dist_add_left H n ▸ !dist.comm) ▸ !dist.comm
definition dist.triangle_inequality (n m k : ) : dist n k ≤ dist n m + dist m k := theorem dist.triangle_inequality (n m k : ) : dist n k ≤ dist n m + dist m k :=
assert (m - k) + ((k - m) + (m - n)) = (m - n) + ((m - k) + (k - m)),
begin
generalize m - k, generalize k - m, generalize m - n, intro x y z,
rewrite [add.comm y x, add.left_comm]
end,
have (n - m) + (m - k) + ((k - m) + (m - n)) = (n - m) + (m - n) + ((m - k) + (k - m)), have (n - m) + (m - k) + ((k - m) + (m - n)) = (n - m) + (m - n) + ((m - k) + (k - m)),
by rewrite [add.assoc, this, -add.assoc], begin rewrite [add.comm (k - m) (m - n),
this ▸ add_le_add !sub_lt_sub_add_sub !sub_lt_sub_add_sub {n - m + _ + _}add.assoc,
{m - k + _}add.left_comm, -add.assoc] end,
this ▸ add_le_add !nat.sub_lt_sub_add_sub !nat.sub_lt_sub_add_sub
definition dist_add_add_le_add_dist_dist (n m k l : ) : dist (n + m) (k + l) ≤ dist n k + dist m l := theorem dist_add_add_le_add_dist_dist (n m k l : ) : dist (n + m) (k + l) ≤ dist n k + dist m l :=
have H : dist (n + m) (k + m) + dist (k + m) (k + l) = dist n k + dist m l, from assert H : dist (n + m) (k + m) + dist (k + m) (k + l) = dist n k + dist m l,
!dist_add_add_left ▸ !dist_add_add_right ▸ rfl, by rewrite [dist_add_add_left, dist_add_add_right],
H ▸ !dist.triangle_inequality by rewrite -H; apply dist.triangle_inequality
theorem dist_mul_right (n k m : ) : dist (n * k) (m * k) = dist n m * k := theorem dist_mul_right (n k m : ) : dist (n * k) (m * k) = dist n m * k :=
assert n m, dist n m = n - m + (m - n), from take n m, rfl, assert Π n m, dist n m = n - m + (m - n), from take n m, rfl,
by rewrite [this, this n m, mul.right_distrib, *mul_sub_right_distrib] by rewrite [this, this n m, right_distrib, *nat.mul_sub_right_distrib]
theorem dist_mul_left (k n m : ) : dist (k * n) (k * m) = k * dist n m := theorem dist_mul_left (k n m : ) : dist (k * n) (k * m) = k * dist n m :=
by rewrite [mul.comm k n, mul.comm k m, dist_mul_right, mul.comm] begin rewrite [mul.comm k n, mul.comm k m, dist_mul_right, mul.comm] end
definition dist_mul_dist (n m k l : ) : dist n m * dist k l = dist (n * k + m * l) (n * l + m * k) := theorem dist_mul_dist (n m k l : ) : dist n m * dist k l = dist (n * k + m * l) (n * l + m * k) :=
have aux : Πk l, k ≥ l → dist n m * dist k l = dist (n * k + m * l) (n * l + m * k), from have aux : Πk l, k ≥ l → dist n m * dist k l = dist (n * k + m * l) (n * l + m * k), from
take k l : , take k l : ,
assume H : k ≥ l, assume H : k ≥ l,
have H2 : m * k ≥ m * l, from mul_le_mul_left H m, have H2 : m * k ≥ m * l, from !mul_le_mul_left H,
have H3 : n * l + m * k ≥ m * l, from le.trans H2 !le_add_left, have H3 : n * l + m * k ≥ m * l, from le.trans H2 !le_add_left,
calc calc
dist n m * dist k l = dist n m * (k - l) : dist_eq_sub_of_ge H dist n m * dist k l = dist n m * (k - l) : dist_eq_sub_of_ge H
... = dist (n * (k - l)) (m * (k - l)) : dist_mul_right ... = dist (n * (k - l)) (m * (k - l)) : dist_mul_right
... = dist (n * k - n * l) (m * k - m * l) : by rewrite [*mul_sub_left_distrib] ... = dist (n * k - n * l) (m * k - m * l) : by rewrite [*nat.mul_sub_left_distrib]
... = dist (n * k) (m * k - m * l + n * l) : dist_sub_eq_dist_add_left (mul_le_mul_left H n) ... = dist (n * k) (m * k - m * l + n * l) : dist_sub_eq_dist_add_left (!mul_le_mul_left H)
... = dist (n * k) (n * l + (m * k - m * l)) : add.comm ... = dist (n * k) (n * l + (m * k - m * l)) : add.comm
... = dist (n * k) (n * l + m * k - m * l) : add_sub_assoc H2 (n * l) ... = dist (n * k) (n * l + m * k - m * l) : nat.add_sub_assoc H2 (n * l)
... = dist (n * k + m * l) (n * l + m * k) : dist_sub_eq_dist_add_right H3 _, ... = dist (n * k + m * l) (n * l + m * k) : dist_sub_eq_dist_add_right H3 _,
sum.rec_on !le.total sum.elim !le.total
(assume H : k ≤ l, !dist.comm ▸ !dist.comm ▸ aux l k H) (assume H : k ≤ l, !dist.comm ▸ !dist.comm ▸ aux l k H)
(assume H : l ≤ k, aux k l H) (assume H : l ≤ k, aux k l H)
lemma dist_eq_max_sub_min {i j : nat} : dist i j = (max i j) - min i j :=
sum.elim (lt_sum_ge i j)
(suppose i < j,
by rewrite [max_eq_right_of_lt this, min_eq_left_of_lt this, dist_eq_sub_of_lt this])
(suppose i ≥ j,
by rewrite [max_eq_left this , min_eq_right this, dist_eq_sub_of_ge this])
lemma dist_succ {i j : nat} : dist (succ i) (succ j) = dist i j :=
by rewrite [↑dist, *succ_sub_succ]
lemma dist_le_max {i j : nat} : dist i j ≤ max i j :=
begin rewrite dist_eq_max_sub_min, apply sub_le end
lemma dist_pos_of_ne {i j : nat} : i ≠ j → dist i j > 0 :=
assume Pne, lt.by_cases
(suppose i < j, begin rewrite [dist_eq_sub_of_lt this], apply nat.sub_pos_of_lt this end)
(suppose i = j, by contradiction)
(suppose i > j, begin rewrite [dist_eq_sub_of_gt this], apply nat.sub_pos_of_lt this end)
end nat end nat

523
hott/types/num.hlean Normal file
View file

@ -0,0 +1,523 @@
/-
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Author: Leonardo de Moura
-/
import types.bool tools.helper_tactics
open bool eq eq.ops decidable helper_tactics
namespace pos_num
theorem succ_not_is_one (a : pos_num) : is_one (succ a) = ff :=
pos_num.rec_on a rfl (take n iH, rfl) (take n iH, rfl)
theorem succ_one : succ one = bit0 one
theorem succ_bit1 (a : pos_num) : succ (bit1 a) = bit0 (succ a)
theorem succ_bit0 (a : pos_num) : succ (bit0 a) = bit1 a
theorem ne_of_bit0_ne_bit0 {a b : pos_num} (H₁ : bit0 a ≠ bit0 b) : a ≠ b :=
suppose a = b,
absurd rfl (this ▸ H₁)
theorem ne_of_bit1_ne_bit1 {a b : pos_num} (H₁ : bit1 a ≠ bit1 b) : a ≠ b :=
suppose a = b,
absurd rfl (this ▸ H₁)
theorem pred_succ : Π (a : pos_num), pred (succ a) = a
| one := rfl
| (bit0 a) := by rewrite succ_bit0
| (bit1 a) :=
calc
pred (succ (bit1 a)) = cond (is_one (succ a)) one (bit1 (pred (succ a))) : rfl
... = cond ff one (bit1 (pred (succ a))) : succ_not_is_one
... = bit1 (pred (succ a)) : rfl
... = bit1 a : pred_succ a
section
variables (a b : pos_num)
theorem one_add_one : one + one = bit0 one
theorem one_add_bit0 : one + (bit0 a) = bit1 a
theorem one_add_bit1 : one + (bit1 a) = succ (bit1 a)
theorem bit0_add_one : (bit0 a) + one = bit1 a
theorem bit1_add_one : (bit1 a) + one = succ (bit1 a)
theorem bit0_add_bit0 : (bit0 a) + (bit0 b) = bit0 (a + b)
theorem bit0_add_bit1 : (bit0 a) + (bit1 b) = bit1 (a + b)
theorem bit1_add_bit0 : (bit1 a) + (bit0 b) = bit1 (a + b)
theorem bit1_add_bit1 : (bit1 a) + (bit1 b) = succ (bit1 (a + b))
theorem one_mul : one * a = a
end
theorem mul_one : Π a, a * one = a
| one := rfl
| (bit1 n) :=
calc bit1 n * one = bit0 (n * one) + one : rfl
... = bit0 n + one : mul_one n
... = bit1 n : bit0_add_one
| (bit0 n) :=
calc bit0 n * one = bit0 (n * one) : rfl
... = bit0 n : mul_one n
theorem decidable_eq [instance] : Π (a b : pos_num), decidable (a = b)
| one one := inl rfl
| one (bit0 b) := inr (by contradiction)
| one (bit1 b) := inr (by contradiction)
| (bit0 a) one := inr (by contradiction)
| (bit0 a) (bit0 b) :=
match decidable_eq a b with
| inl H₁ := inl (by rewrite H₁)
| inr H₁ := inr (by intro H; injection H; contradiction)
end
| (bit0 a) (bit1 b) := inr (by contradiction)
| (bit1 a) one := inr (by contradiction)
| (bit1 a) (bit0 b) := inr (by contradiction)
| (bit1 a) (bit1 b) :=
match decidable_eq a b with
| inl H₁ := inl (by rewrite H₁)
| inr H₁ := inr (by intro H; injection H; contradiction)
end
local notation a < b := (lt a b = tt)
local notation a ` ≮ `:50 b:50 := (lt a b = ff)
theorem lt_one_right_eq_ff : Π a : pos_num, a ≮ one
| one := rfl
| (bit0 a) := rfl
| (bit1 a) := rfl
theorem lt_one_succ_eq_tt : Π a : pos_num, one < succ a
| one := rfl
| (bit0 a) := rfl
| (bit1 a) := rfl
theorem lt_of_lt_bit0_bit0 {a b : pos_num} (H : bit0 a < bit0 b) : a < b := H
theorem lt_of_lt_bit0_bit1 {a b : pos_num} (H : bit1 a < bit0 b) : a < b := H
theorem lt_of_lt_bit1_bit1 {a b : pos_num} (H : bit1 a < bit1 b) : a < b := H
theorem lt_of_lt_bit1_bit0 {a b : pos_num} (H : bit0 a < bit1 b) : a < succ b := H
theorem lt_bit0_bit0_eq_lt (a b : pos_num) : lt (bit0 a) (bit0 b) = lt a b :=
rfl
theorem lt_bit1_bit1_eq_lt (a b : pos_num) : lt (bit1 a) (bit1 b) = lt a b :=
rfl
theorem lt_bit1_bit0_eq_lt (a b : pos_num) : lt (bit1 a) (bit0 b) = lt a b :=
rfl
theorem lt_bit0_bit1_eq_lt_succ (a b : pos_num) : lt (bit0 a) (bit1 b) = lt a (succ b) :=
rfl
theorem lt_irrefl : Π (a : pos_num), a ≮ a
| one := rfl
| (bit0 a) :=
begin
rewrite lt_bit0_bit0_eq_lt, apply lt_irrefl
end
| (bit1 a) :=
begin
rewrite lt_bit1_bit1_eq_lt, apply lt_irrefl
end
theorem ne_of_lt_eq_tt : Π {a b : pos_num}, a < b → a = b → empty
| one ⌞one⌟ H₁ (eq.refl one) := absurd H₁ ff_ne_tt
| (bit0 a) ⌞(bit0 a)⌟ H₁ (eq.refl (bit0 a)) :=
begin
rewrite lt_bit0_bit0_eq_lt at H₁,
apply ne_of_lt_eq_tt H₁ (eq.refl a)
end
| (bit1 a) ⌞(bit1 a)⌟ H₁ (eq.refl (bit1 a)) :=
begin
rewrite lt_bit1_bit1_eq_lt at H₁,
apply ne_of_lt_eq_tt H₁ (eq.refl a)
end
theorem lt_base : Π a : pos_num, a < succ a
| one := rfl
| (bit0 a) :=
begin
rewrite [succ_bit0, lt_bit0_bit1_eq_lt_succ],
apply lt_base
end
| (bit1 a) :=
begin
rewrite [succ_bit1, lt_bit1_bit0_eq_lt],
apply lt_base
end
theorem lt_step : Π {a b : pos_num}, a < b → a < succ b
| one one H := rfl
| one (bit0 b) H := rfl
| one (bit1 b) H := rfl
| (bit0 a) one H := absurd H ff_ne_tt
| (bit0 a) (bit0 b) H :=
begin
rewrite [succ_bit0, lt_bit0_bit1_eq_lt_succ, lt_bit0_bit0_eq_lt at H],
apply lt_step H
end
| (bit0 a) (bit1 b) H :=
begin
rewrite [succ_bit1, lt_bit0_bit0_eq_lt, lt_bit0_bit1_eq_lt_succ at H],
exact H
end
| (bit1 a) one H := absurd H ff_ne_tt
| (bit1 a) (bit0 b) H :=
begin
rewrite [succ_bit0, lt_bit1_bit1_eq_lt, lt_bit1_bit0_eq_lt at H],
exact H
end
| (bit1 a) (bit1 b) H :=
begin
rewrite [succ_bit1, lt_bit1_bit0_eq_lt, lt_bit1_bit1_eq_lt at H],
apply lt_step H
end
theorem lt_of_lt_succ_succ : Π {a b : pos_num}, succ a < succ b → a < b
| one one H := absurd H ff_ne_tt
| one (bit0 b) H := rfl
| one (bit1 b) H := rfl
| (bit0 a) one H :=
begin
rewrite [succ_bit0 at H, succ_one at H, lt_bit1_bit0_eq_lt at H],
apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff a) H
end
| (bit0 a) (bit0 b) H := by exact H
| (bit0 a) (bit1 b) H := by exact H
| (bit1 a) one H :=
begin
rewrite [succ_bit1 at H, succ_one at H, lt_bit0_bit0_eq_lt at H],
apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff (succ a)) H
end
| (bit1 a) (bit0 b) H :=
begin
rewrite [succ_bit1 at H, succ_bit0 at H, lt_bit0_bit1_eq_lt_succ at H],
rewrite lt_bit1_bit0_eq_lt,
apply lt_of_lt_succ_succ H
end
| (bit1 a) (bit1 b) H :=
begin
rewrite [lt_bit1_bit1_eq_lt, *succ_bit1 at H, lt_bit0_bit0_eq_lt at H],
apply lt_of_lt_succ_succ H
end
theorem lt_succ_succ : Π {a b : pos_num}, a < b → succ a < succ b
| one one H := absurd H ff_ne_tt
| one (bit0 b) H :=
begin
rewrite [succ_bit0, succ_one, lt_bit0_bit1_eq_lt_succ],
apply lt_one_succ_eq_tt
end
| one (bit1 b) H :=
begin
rewrite [succ_one, succ_bit1, lt_bit0_bit0_eq_lt],
apply lt_one_succ_eq_tt
end
| (bit0 a) one H := absurd H ff_ne_tt
| (bit0 a) (bit0 b) H := by exact H
| (bit0 a) (bit1 b) H := by exact H
| (bit1 a) one H := absurd H ff_ne_tt
| (bit1 a) (bit0 b) H :=
begin
rewrite [succ_bit1, succ_bit0, lt_bit0_bit1_eq_lt_succ, lt_bit1_bit0_eq_lt at H],
apply lt_succ_succ H
end
| (bit1 a) (bit1 b) H :=
begin
rewrite [lt_bit1_bit1_eq_lt at H, *succ_bit1, lt_bit0_bit0_eq_lt],
apply lt_succ_succ H
end
theorem lt_of_lt_succ : Π {a b : pos_num}, succ a < b → a < b
| one one H := absurd_of_eq_ff_of_eq_tt !lt_one_right_eq_ff H
| one (bit0 b) H := rfl
| one (bit1 b) H := rfl
| (bit0 a) one H := absurd_of_eq_ff_of_eq_tt !lt_one_right_eq_ff H
| (bit0 a) (bit0 b) H := by exact H
| (bit0 a) (bit1 b) H :=
begin
rewrite [succ_bit0 at H, lt_bit1_bit1_eq_lt at H, lt_bit0_bit1_eq_lt_succ],
apply lt_step H
end
| (bit1 a) one H := absurd_of_eq_ff_of_eq_tt !lt_one_right_eq_ff H
| (bit1 a) (bit0 b) H :=
begin
rewrite [lt_bit1_bit0_eq_lt, succ_bit1 at H, lt_bit0_bit0_eq_lt at H],
apply lt_of_lt_succ H
end
| (bit1 a) (bit1 b) H :=
begin
rewrite [succ_bit1 at H, lt_bit0_bit1_eq_lt_succ at H, lt_bit1_bit1_eq_lt],
apply lt_of_lt_succ_succ H
end
theorem lt_of_lt_succ_of_ne : Π {a b : pos_num}, a < succ b → a ≠ b → a < b
| one one H₁ H₂ := absurd rfl H₂
| one (bit0 b) H₁ H₂ := rfl
| one (bit1 b) H₁ H₂ := rfl
| (bit0 a) one H₁ H₂ :=
begin
rewrite [succ_one at H₁, lt_bit0_bit0_eq_lt at H₁],
apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff _) H₁
end
| (bit0 a) (bit0 b) H₁ H₂ :=
begin
rewrite [lt_bit0_bit0_eq_lt, succ_bit0 at H₁, lt_bit0_bit1_eq_lt_succ at H₁],
apply lt_of_lt_succ_of_ne H₁ (ne_of_bit0_ne_bit0 H₂)
end
| (bit0 a) (bit1 b) H₁ H₂ :=
begin
rewrite [succ_bit1 at H₁, lt_bit0_bit0_eq_lt at H₁, lt_bit0_bit1_eq_lt_succ],
exact H₁
end
| (bit1 a) one H₁ H₂ :=
begin
rewrite [succ_one at H₁, lt_bit1_bit0_eq_lt at H₁],
apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff _) H₁
end
| (bit1 a) (bit0 b) H₁ H₂ :=
begin
rewrite [succ_bit0 at H₁, lt_bit1_bit1_eq_lt at H₁, lt_bit1_bit0_eq_lt],
exact H₁
end
| (bit1 a) (bit1 b) H₁ H₂ :=
begin
rewrite [succ_bit1 at H₁, lt_bit1_bit0_eq_lt at H₁, lt_bit1_bit1_eq_lt],
apply lt_of_lt_succ_of_ne H₁ (ne_of_bit1_ne_bit1 H₂)
end
theorem lt_trans : Π {a b c : pos_num}, a < b → b < c → a < c
| one b (bit0 c) H₁ H₂ := rfl
| one b (bit1 c) H₁ H₂ := rfl
| a (bit0 b) one H₁ H₂ := absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff _) H₂
| a (bit1 b) one H₁ H₂ := absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff _) H₂
| (bit0 a) (bit0 b) (bit0 c) H₁ H₂ :=
begin
rewrite lt_bit0_bit0_eq_lt at *, apply lt_trans H₁ H₂
end
| (bit0 a) (bit0 b) (bit1 c) H₁ H₂ :=
begin
rewrite [lt_bit0_bit1_eq_lt_succ at *, lt_bit0_bit0_eq_lt at H₁],
apply lt_trans H₁ H₂
end
| (bit0 a) (bit1 b) (bit0 c) H₁ H₂ :=
begin
rewrite [lt_bit0_bit1_eq_lt_succ at H₁, lt_bit1_bit0_eq_lt at H₂, lt_bit0_bit0_eq_lt],
apply @by_cases (a = b),
begin
intro H, rewrite -H at H₂, exact H₂
end,
begin
intro H,
apply lt_trans (lt_of_lt_succ_of_ne H₁ H) H₂
end
end
| (bit0 a) (bit1 b) (bit1 c) H₁ H₂ :=
begin
rewrite [lt_bit0_bit1_eq_lt_succ at *, lt_bit1_bit1_eq_lt at H₂],
apply lt_trans H₁ (lt_succ_succ H₂)
end
| (bit1 a) (bit0 b) (bit0 c) H₁ H₂ :=
begin
rewrite [lt_bit0_bit0_eq_lt at H₂, lt_bit1_bit0_eq_lt at *],
apply lt_trans H₁ H₂
end
| (bit1 a) (bit0 b) (bit1 c) H₁ H₂ :=
begin
rewrite [lt_bit1_bit0_eq_lt at H₁, lt_bit0_bit1_eq_lt_succ at H₂, lt_bit1_bit1_eq_lt],
apply @by_cases (b = c),
begin
intro H, rewrite H at H₁, exact H₁
end,
begin
intro H,
apply lt_trans H₁ (lt_of_lt_succ_of_ne H₂ H)
end
end
| (bit1 a) (bit1 b) (bit0 c) H₁ H₂ :=
begin
rewrite [lt_bit1_bit1_eq_lt at H₁, lt_bit1_bit0_eq_lt at H₂, lt_bit1_bit0_eq_lt],
apply lt_trans H₁ H₂
end
| (bit1 a) (bit1 b) (bit1 c) H₁ H₂ :=
begin
rewrite lt_bit1_bit1_eq_lt at *,
apply lt_trans H₁ H₂
end
theorem lt_antisymm : Π {a b : pos_num}, a < b → b ≮ a
| one one H := rfl
| one (bit0 b) H := rfl
| one (bit1 b) H := rfl
| (bit0 a) one H := absurd H ff_ne_tt
| (bit0 a) (bit0 b) H :=
begin
rewrite lt_bit0_bit0_eq_lt at *,
apply lt_antisymm H
end
| (bit0 a) (bit1 b) H :=
begin
rewrite lt_bit1_bit0_eq_lt,
rewrite lt_bit0_bit1_eq_lt_succ at H,
have H₁ : succ b ≮ a, from lt_antisymm H,
apply eq_ff_of_ne_tt,
intro H₂,
apply @by_cases (succ b = a),
show succ b = a → empty,
begin
intro Hp,
rewrite -Hp at H,
apply absurd_of_eq_ff_of_eq_tt (lt_irrefl (succ b)) H
end,
show succ b ≠ a → empty,
begin
intro Hn,
have H₃ : succ b < succ a, from lt_succ_succ H₂,
have H₄ : succ b < a, from lt_of_lt_succ_of_ne H₃ Hn,
apply absurd_of_eq_ff_of_eq_tt H₁ H₄
end,
end
| (bit1 a) one H := absurd H ff_ne_tt
| (bit1 a) (bit0 b) H :=
begin
rewrite lt_bit0_bit1_eq_lt_succ,
rewrite lt_bit1_bit0_eq_lt at H,
have H₁ : lt b a = ff, from lt_antisymm H,
apply eq_ff_of_ne_tt,
intro H₂,
apply @by_cases (b = a),
show b = a → empty,
begin
intro Hp,
rewrite -Hp at H,
apply absurd_of_eq_ff_of_eq_tt (lt_irrefl b) H
end,
show b ≠ a → empty,
begin
intro Hn,
have H₃ : b < a, from lt_of_lt_succ_of_ne H₂ Hn,
apply absurd_of_eq_ff_of_eq_tt H₁ H₃
end,
end
| (bit1 a) (bit1 b) H :=
begin
rewrite lt_bit1_bit1_eq_lt at *,
apply lt_antisymm H
end
local notation a ≤ b := (le a b = tt)
theorem le_refl : Π a : pos_num, a ≤ a :=
lt_base
theorem le_eq_lt_succ {a b : pos_num} : le a b = lt a (succ b) :=
rfl
theorem not_lt_of_le : Π {a b : pos_num}, a ≤ b → b < a → empty
| one one H₁ H₂ := absurd H₂ ff_ne_tt
| one (bit0 b) H₁ H₂ := absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff _) H₂
| one (bit1 b) H₁ H₂ := absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff _) H₂
| (bit0 a) one H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at H₁, succ_one at H₁, lt_bit0_bit0_eq_lt at H₁],
apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff _) H₁
end
| (bit0 a) (bit0 b) H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at H₁, succ_bit0 at H₁, lt_bit0_bit1_eq_lt_succ at H₁],
rewrite [lt_bit0_bit0_eq_lt at H₂],
apply not_lt_of_le H₁ H₂
end
| (bit0 a) (bit1 b) H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at H₁, succ_bit1 at H₁, lt_bit0_bit0_eq_lt at H₁],
rewrite [lt_bit1_bit0_eq_lt at H₂],
apply not_lt_of_le H₁ H₂
end
| (bit1 a) one H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at H₁, succ_one at H₁, lt_bit1_bit0_eq_lt at H₁],
apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff _) H₁
end
| (bit1 a) (bit0 b) H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at H₁, succ_bit0 at H₁, lt_bit1_bit1_eq_lt at H₁],
rewrite lt_bit0_bit1_eq_lt_succ at H₂,
have H₃ : a < succ b, from lt_step H₁,
apply @by_cases (b = a),
begin
intro Hba, rewrite -Hba at H₁,
apply absurd_of_eq_ff_of_eq_tt (lt_irrefl b) H₁
end,
begin
intro Hnba,
have H₄ : b < a, from lt_of_lt_succ_of_ne H₂ Hnba,
apply not_lt_of_le H₃ H₄
end
end
| (bit1 a) (bit1 b) H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at H₁, succ_bit1 at H₁, lt_bit1_bit0_eq_lt at H₁],
rewrite [lt_bit1_bit1_eq_lt at H₂],
apply not_lt_of_le H₁ H₂
end
theorem le_antisymm : Π {a b : pos_num}, a ≤ b → b ≤ a → a = b
| one one H₁ H₂ := rfl
| one (bit0 b) H₁ H₂ :=
by apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff b) H₂
| one (bit1 b) H₁ H₂ :=
by apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff b) H₂
| (bit0 a) one H₁ H₂ :=
by apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff a) H₁
| (bit0 a) (bit0 b) H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at *, succ_bit0 at *, lt_bit0_bit1_eq_lt_succ at *],
have H : a = b, from le_antisymm H₁ H₂,
rewrite H
end
| (bit0 a) (bit1 b) H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at *, succ_bit1 at H₁, succ_bit0 at H₂],
rewrite [lt_bit0_bit0_eq_lt at H₁, lt_bit1_bit1_eq_lt at H₂],
apply empty.rec _ (not_lt_of_le H₁ H₂)
end
| (bit1 a) one H₁ H₂ :=
by apply absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff a) H₁
| (bit1 a) (bit0 b) H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at *, succ_bit0 at H₁, succ_bit1 at H₂],
rewrite [lt_bit1_bit1_eq_lt at H₁, lt_bit0_bit0_eq_lt at H₂],
apply empty.rec _ (not_lt_of_le H₂ H₁)
end
| (bit1 a) (bit1 b) H₁ H₂ :=
begin
rewrite [le_eq_lt_succ at *, succ_bit1 at *, lt_bit1_bit0_eq_lt at *],
have H : a = b, from le_antisymm H₁ H₂,
rewrite H
end
theorem le_trans {a b c : pos_num} : a ≤ b → b ≤ c → a ≤ c :=
begin
intro H₁ H₂,
rewrite [le_eq_lt_succ at *],
apply @by_cases (a = b),
begin
intro Hab, rewrite Hab, exact H₂
end,
begin
intro Hnab,
have Haltb : a < b, from lt_of_lt_succ_of_ne H₁ Hnab,
apply lt_trans Haltb H₂
end,
end
end pos_num
namespace num
open pos_num
theorem decidable_eq [instance] : Π (a b : num), decidable (a = b)
| zero zero := inl rfl
| zero (pos b) := inr (by contradiction)
| (pos a) zero := inr (by contradiction)
| (pos a) (pos b) :=
if H : a = b then inl (by rewrite H) else inr (suppose pos a = pos b, begin injection this, contradiction end)
end num

View file

@ -7,7 +7,7 @@ Ported from Coq HoTT
-/ -/
import arity .eq .bool .unit .sigma .nat.basic import arity .eq .bool .unit .sigma .nat.basic
open is_trunc eq prod sigma nat equiv option is_equiv bool unit open is_trunc eq prod sigma nat equiv option is_equiv bool unit algebra
structure pointed [class] (A : Type) := structure pointed [class] (A : Type) :=
(point : A) (point : A)
@ -134,7 +134,7 @@ namespace pointed
end end
definition pid [constructor] (A : Type*) : A →* A := definition pid [constructor] (A : Type*) : A →* A :=
pmap.mk function.id idp pmap.mk id idp
definition pcompose [constructor] (g : B →* C) (f : A →* B) : A →* C := definition pcompose [constructor] (g : B →* C) (f : A →* B) : A →* C :=
pmap.mk (λa, g (f a)) (ap g (respect_pt f) ⬝ respect_pt g) pmap.mk (λa, g (f a)) (ap g (respect_pt f) ⬝ respect_pt g)
@ -273,7 +273,7 @@ namespace pointed
Ω[succ n](Pointed.mk p) = Ω[n](Ω (Pointed.mk p)) : loop_space_succ_eq_in Ω[succ n](Pointed.mk p) = Ω[n](Ω (Pointed.mk p)) : loop_space_succ_eq_in
... = Ω[n] (Ω[2] A) : loop_space_loop_irrel ... = Ω[n] (Ω[2] A) : loop_space_loop_irrel
... = Ω[2+n] A : loop_space_add ... = Ω[2+n] A : loop_space_add
... = Ω[n+2] A : add.comm ... = Ω[n+2] A : by rewrite [algebra.add.comm]
-- TODO: -- TODO:
-- definition apn_compose (n : ) (g : B →* C) (f : A →* B) : apn n (g ∘* f) ~* apn n g ∘* apn n f := -- definition apn_compose (n : ) (g : B →* C) (f : A →* B) : apn n (g ∘* f) ~* apn n g ∘* apn n f :=

View file

@ -152,13 +152,13 @@ namespace is_trunc
revert A, induction n with n IH, revert A, induction n with n IH,
{ intro A, esimp [Iterated_loop_space], transitivity _, { intro A, esimp [Iterated_loop_space], transitivity _,
{ apply is_trunc_succ_iff_is_trunc_loop, apply le.refl}, { apply is_trunc_succ_iff_is_trunc_loop, apply le.refl},
{ apply iff.pi_iff_pi, intro a, esimp, apply is_hprop_iff_is_contr, reflexivity}}, { apply pi_iff_pi, intro a, esimp, apply is_hprop_iff_is_contr, reflexivity}},
{ intro A, esimp [Iterated_loop_space], { intro A, esimp [Iterated_loop_space],
transitivity _, apply @is_trunc_succ_iff_is_trunc_loop @n, esimp, constructor, transitivity _, apply @is_trunc_succ_iff_is_trunc_loop @n, esimp, constructor,
apply iff.pi_iff_pi, intro a, transitivity _, apply IH, apply pi_iff_pi, intro a, transitivity _, apply IH,
transitivity _, apply iff.pi_iff_pi, intro p, transitivity _, apply pi_iff_pi, intro p,
rewrite [iterated_loop_space_loop_irrel n p], apply iff.refl, esimp, rewrite [iterated_loop_space_loop_irrel n p], apply iff.refl, esimp,
apply iff.imp_iff, reflexivity} apply imp_iff, reflexivity}
end end
theorem is_trunc_iff_is_contr_loop (n : ) (A : Type) theorem is_trunc_iff_is_contr_loop (n : ) (A : Type)

View file

@ -6,8 +6,6 @@ Authors: Floris van Doorn
Theorems about the unit type Theorems about the unit type
-/ -/
import algebra.group
open equiv option eq open equiv option eq
namespace unit namespace unit
@ -36,15 +34,3 @@ namespace unit
end unit end unit
open unit is_trunc open unit is_trunc
namespace algebra
definition trivial_group [constructor] : group unit :=
group.mk (λx y, star) _ (λx y z, idp) star (unit.rec idp) (unit.rec idp) (λx, star) (λx, idp)
definition Trivial_group [constructor] : Group :=
Group.mk _ trivial_group
notation `G0` := Trivial_group
end algebra

View file

@ -23,6 +23,8 @@ section
theorem le.refl (a : A) : a ≤ a := !weak_order.le_refl theorem le.refl (a : A) : a ≤ a := !weak_order.le_refl
theorem le_of_eq {a b : A} (H : a = b) : a ≤ b := H ▸ le.refl a
theorem le.trans [trans] {a b c : A} : a ≤ b → b ≤ c → a ≤ c := !weak_order.le_trans theorem le.trans [trans] {a b c : A} : a ≤ b → b ≤ c → a ≤ c := !weak_order.le_trans
theorem ge.trans [trans] {a b c : A} (H1 : a ≥ b) (H2: b ≥ c) : a ≥ c := le.trans H2 H1 theorem ge.trans [trans] {a b c : A} (H1 : a ≥ b) (H2: b ≥ c) : a ≥ c := le.trans H2 H1
@ -218,6 +220,9 @@ section
(assume H, H1 H) (assume H, H1 H)
(assume H, or.elim H (assume H', H2 H') (assume H', H3 H')) (assume H, or.elim H (assume H', H2 H') (assume H', H3 H'))
definition lt_ge_by_cases {a b : A} {P : Prop} (H1 : a < b → P) (H2 : a ≥ b → P) : P :=
lt.by_cases H1 (λH, H2 (H ▸ le.refl a)) (λH, H2 (le_of_lt H))
theorem le_of_not_gt {a b : A} (H : ¬ a > b) : a ≤ b := theorem le_of_not_gt {a b : A} (H : ¬ a > b) : a ≤ b :=
lt.by_cases (assume H', absurd H' H) (assume H', H' ▸ !le.refl) (assume H', le_of_lt H') lt.by_cases (assume H', absurd H' H) (assume H', H' ▸ !le.refl) (assume H', le_of_lt H')

View file

@ -25,8 +25,7 @@ following:
padd_congr (p p' q q' : × ) (H1 : p ≡ p') (H2 : q ≡ q') : padd p q ≡ p' q' padd_congr (p p' q q' : × ) (H1 : p ≡ p') (H2 : q ≡ q') : padd p q ≡ p' q'
-/ -/
import data.nat.basic data.nat.order data.nat.sub data.prod import data.nat.sub algebra.relation data.prod
import algebra.relation algebra.binary algebra.ordered_ring
open eq.ops open eq.ops
open prod relation nat open prod relation nat
open decidable binary open decidable binary
@ -495,9 +494,11 @@ private theorem pmul_assoc_prep {p1 p2 q1 q2 r1 r2 : } :
((p1*q1+p2*q2)*r1+(p1*q2+p2*q1)*r2, (p1*q1+p2*q2)*r2+(p1*q2+p2*q1)*r1) = ((p1*q1+p2*q2)*r1+(p1*q2+p2*q1)*r2, (p1*q1+p2*q2)*r2+(p1*q2+p2*q1)*r1) =
(p1*(q1*r1+q2*r2)+p2*(q1*r2+q2*r1), p1*(q1*r2+q2*r1)+p2*(q1*r1+q2*r2)) := (p1*(q1*r1+q2*r2)+p2*(q1*r2+q2*r1), p1*(q1*r2+q2*r1)+p2*(q1*r1+q2*r2)) :=
begin begin
rewrite[+left_distrib,+right_distrib,*mul.assoc], rewrite [+left_distrib, +right_distrib, *mul.assoc],
exact (congr_arg2 pair (!add.comm4 ⬝ (!congr_arg !nat.add_comm)) rewrite (add.comm4 (p1 * (q1 * r1)) (p2 * (q2 * r1)) (p1 * (q2 * r2)) (p2 * (q1 * r2))),
(!add.comm4 ⬝ (!congr_arg !nat.add_comm))) rewrite (add.comm (p2 * (q2 * r1)) (p2 * (q1 * r2))),
rewrite (add.comm4 (p1 * (q1 * r2)) (p2 * (q2 * r2)) (p1 * (q2 * r1)) (p2 * (q1 * r1))),
rewrite (add.comm (p2 * (q2 * r2)) (p2 * (q1 * r1)))
end end
theorem pmul_assoc (p q r: × ) : pmul (pmul p q) r = pmul p (pmul q r) := pmul_assoc_prep theorem pmul_assoc (p q r: × ) : pmul (pmul p q) r = pmul p (pmul q r) := pmul_assoc_prep
@ -592,6 +593,7 @@ by rewrite [neg_succ_of_nat_eq, neg_add]
definition succ (a : ) := a + (succ zero) definition succ (a : ) := a + (succ zero)
definition pred (a : ) := a - (succ zero) definition pred (a : ) := a - (succ zero)
definition nat_succ_eq_int_succ (n : ) : nat.succ n = int.succ n := rfl
theorem pred_succ (a : ) : pred (succ a) = a := !sub_add_cancel theorem pred_succ (a : ) : pred (succ a) = a := !sub_add_cancel
theorem succ_pred (a : ) : succ (pred a) = a := !add_sub_cancel theorem succ_pred (a : ) : succ (pred a) = a := !add_sub_cancel

View file

@ -5,7 +5,7 @@ Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad
Basic operations on the natural numbers. Basic operations on the natural numbers.
-/ -/
import logic.connectives data.num algebra.binary algebra.ring import ..num algebra.ring
open binary eq.ops open binary eq.ops
namespace nat namespace nat

View file

@ -5,7 +5,7 @@ Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad
The order relation on the natural numbers. The order relation on the natural numbers.
-/ -/
import data.nat.basic algebra.ordered_ring import .basic algebra.ordered_ring
open eq.ops open eq.ops
namespace nat namespace nat
@ -269,7 +269,7 @@ or.elim !eq_zero_or_pos (take H2 : n = 0, by contradiction) (take H2 : n > 0, H2
theorem ne_zero_of_pos {n : } (H : n > 0) : n ≠ 0 := theorem ne_zero_of_pos {n : } (H : n > 0) : n ≠ 0 :=
ne.symm (ne_of_lt H) ne.symm (ne_of_lt H)
theorem exists_eq_succ_of_pos {n : } (H : n > 0) : exists l, n = succ l := theorem exists_eq_succ_of_pos {n : } (H : n > 0) : l, n = succ l :=
exists_eq_succ_of_lt H exists_eq_succ_of_lt H
theorem pos_of_dvd_of_pos {m n : } (H1 : m n) (H2 : n > 0) : m > 0 := theorem pos_of_dvd_of_pos {m n : } (H1 : m n) (H2 : n > 0) : m > 0 :=

View file

@ -145,8 +145,13 @@ iff.trans (iff.trans !or.comm !or.left_distrib) (and_congr !or.comm !or.comm)
definition iff.def : (a ↔ b) = ((a → b) ∧ (b → a)) := rfl definition iff.def : (a ↔ b) = ((a → b) ∧ (b → a)) := rfl
theorem forall_imp_forall {A : Type} {P Q : A → Prop} (H : ∀a, (P a → Q a)) (p : ∀a, P a) (a : A) : Q a := theorem forall_imp_forall {A : Type} {P Q : A → Prop} (H : ∀a, (P a → Q a)) (p : ∀a, P a) (a : A)
: Q a :=
(H a) (p a) (H a) (p a)
theorem forall_iff_forall {A : Type} {P Q : A → Prop} (H : ∀a, (P a ↔ Q a))
: (∀a, P a) ↔ (∀a, Q a) :=
iff.intro (λp a, iff.elim_left (H a) (p a)) (λq a, iff.elim_right (H a) (q a))
theorem imp_iff {P : Prop} (Q : Prop) (p : P) : (P → Q) ↔ Q := theorem imp_iff {P : Prop} (Q : Prop) (p : P) : (P → Q) ↔ Q :=
iff.intro (λf, f p) imp.intro iff.intro (λf, f p) imp.intro

View file

@ -1,16 +0,0 @@
# usage:
# Make sure port.sh and port.pl are executable (chmod u+x port.pl port.sh)
# in the scripts directory, type ./port.sh to port the files specified below
# from the standard library to the HoTT library
# This file requires both port.pl and port.txt to be in the scripts folder
#
# WARNING: This will overwrite all destination files without warning!
#
# See port.pl for the syntax, if you want to add new files to port.
now=$(date +"%B %d, %Y")
./port.pl ../library/data/nat/basic.lean ../hott/types/nat/basic2.hlean "Module: data.nat.basic" "Module: types.nat.basic
(Ported from standard library file data.nat.basic on $now)" "import logic.connectives data.num algebra.binary algebra.ring" "import algebra.ring" "open binary eq.ops" "open core prod binary" "nat.no_confusion H \(λe, e\)" "lift.down (nat.no_confusion H (λe, e))"
# ./port.pl ../library/logic/connectives.lean ../hott/logic.hlean
/port.pl ../library/algebra/ring.lean ../hott/algebra/ring.hlean "import logic.eq logic.connectives data.unit data.sigma data.prod" "import algebra.group" "import algebra.function algebra.binary algebra.group" "" "open eq eq.ops" "open core"

View file

@ -8,32 +8,35 @@ false:empty
induction_on:rec_on induction_on:rec_on
;⊎ ;⊎
or.elim:sum.elim or:sum
or.inl:sum.inl sum.intro_left _;sum.inl
or.inr:sum.inr sum.intro_right _;sum.inr
or.intro_left _;sum.inl or.intro_left _;sum.inl
or.intro_right _;sum.inr or.intro_right _;sum.inr
or_resolve_right:sum_resolve_right
or_resolve_left:sum_resolve_left
or.swap:sum.swap
or.rec_on:sum.rec_on
or_of_or_of_imp_of_imp:sum_of_sum_of_imp_of_imp
or_of_or_of_imp_left:sum_of_sum_of_imp_left
or_of_or_of_imp_right:sum_of_sum_of_imp_right
∧;× ∧;×
and:prod
and.intro:pair and.intro:pair
and.left:
and.elim_left:prod.pr1 and.elim_left:prod.pr1
and.left:prod.pr1 and.left:prod.pr1
and.elim_right:prod.pr2 and.elim_right:prod.pr2
and.right:prod.pr2 and.right:prod.pr2
prod.intro:pair
prod.elim_left:prod.pr1
prod.left:prod.pr1
prod.elim_right:prod.pr2
prod.right:prod.pr2
∀;Π ∀;Π
∃;Σ ∃;Σ
exists.intro:sigma.mk exists.intro:sigma.mk
exists.elim:sigma.rec_on exists.elim:sigma.rec_on
Exists.rec:sigma.rec
eq.symm:inverse eq.symm:inverse
congr_arg:ap congr_arg:ap