feat(hott/library): various changes and additions.

Most notably:
Give le.refl the attribute [refl]. This simplifies tactic proofs in various places.
Redefine the order of trunc_index, and instantiate it as weak order.
Add more about pointed equivalences.
This commit is contained in:
Floris van Doorn 2016-02-24 19:43:50 -05:00 committed by Leonardo de Moura
parent 65b367ddff
commit e5d5ef9d55
18 changed files with 414 additions and 170 deletions

View file

@ -40,13 +40,13 @@ namespace eq
definition fundamental_group [constructor] (A : Type*) : Group := definition fundamental_group [constructor] (A : Type*) : Group :=
ghomotopy_group zero A ghomotopy_group zero A
notation `πG[`:95 n:0 ` +1] `:0 A:95 := ghomotopy_group n A notation `πg[`:95 n:0 ` +1] `:0 A:95 := ghomotopy_group n A
notation `πaG[`:95 n:0 ` +2] `:0 A:95 := cghomotopy_group n A notation `πag[`:95 n:0 ` +2] `:0 A:95 := cghomotopy_group n A
prefix `π₁`:95 := fundamental_group prefix `π₁`:95 := fundamental_group
open equiv unit open equiv unit
theorem trivial_homotopy_of_is_set (A : Type*) [H : is_set A] (n : ) : πG[n+1] A = G0 := theorem trivial_homotopy_of_is_set (A : Type*) [H : is_set A] (n : ) : πg[n+1] A = G0 :=
begin begin
apply trivial_group_of_is_contr, apply trivial_group_of_is_contr,
apply is_trunc_trunc_of_is_trunc, apply is_trunc_trunc_of_is_trunc,
@ -54,9 +54,9 @@ namespace eq
apply is_trunc_succ_succ_of_is_set apply is_trunc_succ_succ_of_is_set
end end
definition homotopy_group_succ_out (A : Type*) (n : ) : πG[ n +1] A = π₁ Ω[n] A := idp definition homotopy_group_succ_out (A : Type*) (n : ) : πg[ n +1] A = π₁ Ω[n] A := idp
definition homotopy_group_succ_in (A : Type*) (n : ) : πG[succ n +1] A = πG[n +1] Ω A := definition homotopy_group_succ_in (A : Type*) (n : ) : πg[succ n +1] A = πg[n +1] Ω A :=
begin begin
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))},
@ -65,7 +65,7 @@ namespace eq
apply loop_space_succ_eq_in_concat end end}, apply loop_space_succ_eq_in_concat end end},
end end
definition homotopy_group_add (A : Type*) (n m : ) : πG[n+m +1] A = πG[n +1] Ω[m] A := definition homotopy_group_add (A : Type*) (n m : ) : πg[n+m +1] A = πg[n +1] Ω[m] A :=
begin begin
revert A, induction m with m IH: intro A, revert A, induction m with m IH: intro A,
{ reflexivity}, { reflexivity},
@ -74,7 +74,7 @@ namespace eq
end end
theorem trivial_homotopy_of_is_set_loop_space {A : Type*} {n : } (m : ) (H : is_set (Ω[n] A)) theorem trivial_homotopy_of_is_set_loop_space {A : Type*} {n : } (m : ) (H : is_set (Ω[n] A))
: πG[m+n+1] A = G0 := : πg[m+n+1] A = G0 :=
!homotopy_group_add ⬝ !trivial_homotopy_of_is_set !homotopy_group_add ⬝ !trivial_homotopy_of_is_set
definition phomotopy_group_functor [constructor] (n : ) {A B : Type*} (f : A →* B) definition phomotopy_group_functor [constructor] (n : ) {A B : Type*} (f : A →* B)

View file

@ -23,7 +23,7 @@ section
variable [s : weak_order A] variable [s : weak_order A]
include s include s
definition le.refl (a : A) : a ≤ a := !weak_order.le_refl definition le.refl [refl] (a : A) : a ≤ a := !weak_order.le_refl
definition le_of_eq {a b : A} (H : a = b) : a ≤ b := H ▸ le.refl a definition le_of_eq {a b : A} (H : a = b) : a ≤ b := H ▸ le.refl a
@ -307,13 +307,13 @@ section
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
(assume H : a ≤ b, by rewrite [↑min, if_pos H]; apply le.refl) (assume H : a ≤ b, by rewrite [↑min, if_pos H])
(assume H : ¬ a ≤ b, by rewrite [↑min, if_neg H]; apply le_of_lt (lt_of_not_ge H)) (assume H : ¬ a ≤ b, by rewrite [↑min, if_neg H]; apply le_of_lt (lt_of_not_ge H))
theorem min_le_right (a b : A) : min a b ≤ b := theorem min_le_right (a b : A) : min a b ≤ b :=
by_cases by_cases
(assume H : a ≤ b, by rewrite [↑min, if_pos H]; apply H) (assume H : a ≤ b, by rewrite [↑min, if_pos H]; apply H)
(assume H : ¬ a ≤ b, by rewrite [↑min, if_neg H]; apply le.refl) (assume H : ¬ a ≤ b, by rewrite [↑min, if_neg H])
theorem le_min {a b c : A} (H₁ : c ≤ a) (H₂ : c ≤ b) : c ≤ min a b := theorem le_min {a b c : A} (H₁ : c ≤ a) (H₂ : c ≤ b) : c ≤ min a b :=
by_cases by_cases
@ -323,11 +323,11 @@ section
theorem le_max_left (a b : A) : a ≤ max a b := theorem le_max_left (a b : A) : a ≤ max a b :=
by_cases by_cases
(assume H : a ≤ b, by rewrite [↑max, if_pos H]; apply H) (assume H : a ≤ b, by rewrite [↑max, if_pos H]; apply H)
(assume H : ¬ a ≤ b, by rewrite [↑max, if_neg H]; apply le.refl) (assume H : ¬ a ≤ b, by rewrite [↑max, if_neg H])
theorem le_max_right (a b : A) : b ≤ max a b := theorem le_max_right (a b : A) : b ≤ max a b :=
by_cases by_cases
(assume H : a ≤ b, by rewrite [↑max, if_pos H]; apply le.refl) (assume H : a ≤ b, by rewrite [↑max, if_pos H])
(assume H : ¬ a ≤ b, by rewrite [↑max, if_neg H]; apply le_of_lt (lt_of_not_ge H)) (assume H : ¬ a ≤ b, by rewrite [↑max, if_neg H]; apply le_of_lt (lt_of_not_ge H))
theorem max_le {a b c : A} (H₁ : a ≤ c) (H₂ : b ≤ c) : max a b ≤ c := theorem max_le {a b c : A} (H₁ : a ≤ c) (H₂ : b ≤ c) : max a b ≤ c :=

View file

@ -16,7 +16,7 @@ variable {A : Type}
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 :=
(add_le_add_left : Πa b, le a b → Πc, le (add c a) (add c b)) (add_le_add_left : Πa b, le a b → Πc, le (add c a) (add c b))
(le_of_add_le_add_left : Πa b c, le (add a b) (add a c) → le b c) (le_of_add_le_add_left : Πa b c, le (add a b) (add a c) → le b c)
(add_lt_add_left : Πa b, lt a b → Πc, lt (add c a) (add c b)) (add_lt_add_left : Πa b, lt a b → Πc, lt (add c a) (add c b))

View file

@ -252,7 +252,7 @@ namespace circle
end end
open nat open nat
definition homotopy_group_of_circle (n : ) : πG[n+1 +1] S¹. = G0 := definition homotopy_group_of_circle (n : ) : πg[n+1 +1] S¹. = G0 :=
begin begin
refine @trivial_homotopy_of_is_set_loop_space S¹. 1 n _, refine @trivial_homotopy_of_is_set_loop_space S¹. 1 n _,
apply is_trunc_equiv_closed_rev, apply base_eq_base_equiv apply is_trunc_equiv_closed_rev, apply base_eq_base_equiv
@ -272,7 +272,7 @@ namespace circle
definition is_trunc_circle [instance] : is_trunc 1 S¹ := definition is_trunc_circle [instance] : is_trunc 1 S¹ :=
begin begin
apply is_trunc_succ_of_is_trunc_loop, apply is_trunc_succ_of_is_trunc_loop,
{ exact unit.star}, { apply trunc_index.minus_one_le_succ},
{ intro x, apply is_trunc_equiv_closed_rev, apply eq_equiv_Z} { intro x, apply is_trunc_equiv_closed_rev, apply eq_equiv_Z}
end end

View file

@ -48,16 +48,25 @@ namespace sphere_index
notation `-1` := minus_one notation `-1` := minus_one
notation `ℕ₋₁` := sphere_index notation `ℕ₋₁` := sphere_index
definition add (n m : sphere_index) : sphere_index := definition add_plus_one (n m : sphere_index) : sphere_index :=
sphere_index.rec_on m n (λ k l, l .+1) sphere_index.rec_on m n (λ k l, l .+1)
definition leq (n m : sphere_index) : Type₀ := -- addition of sphere_indices, where (-1 + -1) is defined to be -1.
protected definition add (n m : sphere_index) : sphere_index :=
sphere_index.cases_on m
(sphere_index.cases_on n -1 id)
(sphere_index.rec n (λn' r, succ r))
protected definition le (n m : sphere_index) : Type₀ :=
sphere_index.rec_on n (λm, unit) (λ n p m, sphere_index.rec_on m (λ p, empty) (λ m q p, p m) p) m sphere_index.rec_on n (λm, unit) (λ n p m, sphere_index.rec_on m (λ p, empty) (λ m q p, p m) p) m
infix `+1+`:65 := sphere_index.add infix `+1+`:65 := sphere_index.add_plus_one
definition has_add_sphere_index [instance] [priority 2000] [reducible] : has_add ℕ₋₁ :=
has_add.mk sphere_index.add
definition has_le_sphere_index [instance] : has_le sphere_index := definition has_le_sphere_index [instance] : has_le sphere_index :=
has_le.mk leq has_le.mk sphere_index.le
definition succ_le_succ {n m : sphere_index} (H : n ≤ m) : n.+1 ≤ m.+1 := proof H qed 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 := proof H qed definition le_of_succ_le_succ {n m : sphere_index} (H : n.+1 ≤ m.+1) : n ≤ m := proof H qed
@ -125,16 +134,17 @@ namespace sphere
definition sphere_eq_bool : S 0 = bool := definition sphere_eq_bool : S 0 = bool :=
ua sphere_equiv_bool ua sphere_equiv_bool
definition sphere_eq_bool_pointed : S. 0 = pbool := definition sphere_eq_pbool : S. 0 = pbool :=
pType_eq sphere_equiv_bool idp pType_eq sphere_equiv_bool idp
-- TODO: the commented-out part makes the forward function below "apn _ surf" -- TODO1: the commented-out part makes the forward function below "apn _ surf"
-- TODO2: we could make this a pointed equivalence
definition pmap_sphere (A : Type*) (n : ) : map₊ (S. n) A ≃ Ω[n] A := definition pmap_sphere (A : Type*) (n : ) : map₊ (S. n) A ≃ Ω[n] A :=
begin begin
-- fapply equiv_change_fun, -- fapply equiv_change_fun,
-- { -- {
revert A, induction n with n IH: intro A, revert A, induction n with n IH: intro A,
{ krewrite [sphere_eq_bool_pointed], apply pmap_bool_equiv}, { apply tr_rev (λx, x →* _ ≃ _) sphere_eq_pbool, apply pmap_bool_equiv},
{ refine susp_adjoint_loop (S. n) A ⬝e !IH ⬝e _, rewrite [loop_space_succ_eq_in]} { refine susp_adjoint_loop (S. n) A ⬝e !IH ⬝e _, rewrite [loop_space_succ_eq_in]}
-- }, -- },
-- { intro f, exact apn n f surf}, -- { intro f, exact apn n f surf},

View file

@ -56,11 +56,11 @@ namespace eq
-- Concatenation is associative. -- Concatenation is associative.
definition con.assoc' (p : x = y) (q : y = z) (r : z = t) : definition con.assoc' (p : x = y) (q : y = z) (r : z = t) :
p ⬝ (q ⬝ r) = (p ⬝ q) ⬝ r := p ⬝ (q ⬝ r) = (p ⬝ q) ⬝ r :=
by induction r; induction q; reflexivity by induction r; reflexivity
definition con.assoc (p : x = y) (q : y = z) (r : z = t) : definition con.assoc (p : x = y) (q : y = z) (r : z = t) :
(p ⬝ q) ⬝ r = p ⬝ (q ⬝ r) := (p ⬝ q) ⬝ r = p ⬝ (q ⬝ r) :=
by induction r; induction q; reflexivity by induction r; reflexivity
-- The left inverse law. -- The left inverse law.
definition con.right_inv [unfold 4] (p : x = y) : p ⬝ p⁻¹ = idp := definition con.right_inv [unfold 4] (p : x = y) : p ⬝ p⁻¹ = idp :=
@ -80,10 +80,10 @@ namespace eq
by induction q; induction p; reflexivity by induction q; induction p; reflexivity
definition con_inv_cancel_right (p : x = y) (q : y = z) : (p ⬝ q) ⬝ q⁻¹ = p := definition con_inv_cancel_right (p : x = y) (q : y = z) : (p ⬝ q) ⬝ q⁻¹ = p :=
by induction q; induction p; reflexivity by induction q; reflexivity
definition inv_con_cancel_right (p : x = z) (q : y = z) : (p ⬝ q⁻¹) ⬝ q = p := definition inv_con_cancel_right (p : x = z) (q : y = z) : (p ⬝ q⁻¹) ⬝ q = p :=
by induction q; induction p; reflexivity by induction q; reflexivity
-- Inverse distributes over concatenation -- Inverse distributes over concatenation
definition con_inv (p : x = y) (q : y = z) : (p ⬝ q)⁻¹ = q⁻¹ ⬝ p⁻¹ := definition con_inv (p : x = y) (q : y = z) : (p ⬝ q)⁻¹ = q⁻¹ ⬝ p⁻¹ :=
@ -168,22 +168,22 @@ namespace eq
by induction p; intro h; exact h ⬝ !idp_con by induction p; intro h; exact h ⬝ !idp_con
definition con_inv_eq_idp [unfold 6] {p q : x = y} (r : p = q) : p ⬝ q⁻¹ = idp := definition con_inv_eq_idp [unfold 6] {p q : x = y} (r : p = q) : p ⬝ q⁻¹ = idp :=
by cases r;apply con.right_inv by cases r; apply con.right_inv
definition inv_con_eq_idp [unfold 6] {p q : x = y} (r : p = q) : q⁻¹ ⬝ p = idp := definition inv_con_eq_idp [unfold 6] {p q : x = y} (r : p = q) : q⁻¹ ⬝ p = idp :=
by cases r;apply con.left_inv by cases r; apply con.left_inv
definition con_eq_idp {p : x = y} {q : y = x} (r : p = q⁻¹) : p ⬝ q = idp := definition con_eq_idp {p : x = y} {q : y = x} (r : p = q⁻¹) : p ⬝ q = idp :=
by cases q;exact r by cases q; exact r
definition idp_eq_inv_con {p q : x = y} (r : p = q) : idp = p⁻¹ ⬝ q := definition idp_eq_inv_con {p q : x = y} (r : p = q) : idp = p⁻¹ ⬝ q :=
by cases r;exact !con.left_inv⁻¹ by cases r; exact !con.left_inv⁻¹
definition idp_eq_con_inv {p q : x = y} (r : p = q) : idp = q ⬝ p⁻¹ := definition idp_eq_con_inv {p q : x = y} (r : p = q) : idp = q ⬝ p⁻¹ :=
by cases r;exact !con.right_inv⁻¹ by cases r; exact !con.right_inv⁻¹
definition idp_eq_con {p : x = y} {q : y = x} (r : p⁻¹ = q) : idp = q ⬝ p := definition idp_eq_con {p : x = y} {q : y = x} (r : p⁻¹ = q) : idp = q ⬝ p :=
by cases p;exact r by cases p; exact r
/- Transport -/ /- Transport -/
@ -248,7 +248,7 @@ namespace eq
definition apd [unfold 6] (f : Πa, P a) {x y : A} (p : x = y) : p ▸ f x = f y := definition apd [unfold 6] (f : Πa, P a) {x y : A} (p : x = y) : p ▸ f x = f y :=
by induction p; reflexivity by induction p; reflexivity
definition ap011 (f : A → B → C) (Ha : a = a') (Hb : b = b') : f a b = f a' b' := definition ap011 [unfold 9] (f : A → B → C) (Ha : a = a') (Hb : b = b') : f a b = f a' b' :=
by cases Ha; exact ap (f a) Hb by cases Ha; exact ap (f a) Hb
/- More theorems for moving things around in equations -/ /- More theorems for moving things around in equations -/

View file

@ -23,15 +23,9 @@ namespace is_trunc
open trunc_index open trunc_index
definition has_zero_trunc_index [instance] [priority 2000] : has_zero trunc_index :=
has_zero.mk (succ (succ minus_two))
definition has_one_trunc_index [instance] [priority 2000] : 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 )
-/ -/
notation `-1` := trunc_index.succ trunc_index.minus_two -- ISSUE: -1 gets printed as -2.+1? notation `-1` := trunc_index.succ trunc_index.minus_two -- ISSUE: -1 gets printed as -2.+1?
notation `-2` := trunc_index.minus_two notation `-2` := trunc_index.minus_two
@ -39,47 +33,63 @@ namespace is_trunc
postfix ` .+2`:(max+1) := λn, (n .+1 .+1) postfix ` .+2`:(max+1) := λn, (n .+1 .+1)
notation `ℕ₋₂` := trunc_index -- input using \N-2 notation `ℕ₋₂` := trunc_index -- input using \N-2
definition has_zero_trunc_index [instance] [priority 2000] : has_zero ℕ₋₂ :=
has_zero.mk (succ (succ minus_two))
definition has_one_trunc_index [instance] [priority 2000] : has_one ℕ₋₂ :=
has_one.mk (succ (succ (succ minus_two)))
namespace trunc_index namespace trunc_index
--addition, where we add two to the result --addition, where we add two to the result
definition add_plus_two [reducible] (n m : trunc_index) : trunc_index := definition add_plus_two [reducible] (n m : ℕ₋₂) : ℕ₋₂ :=
trunc_index.rec_on m n (λ k l, l .+1) trunc_index.rec_on m n (λ k l, l .+1)
-- addition of trunc_indices, where results smaller than -2 are changed to -2 -- addition of trunc_indices, where results smaller than -2 are changed to -2
definition tr_add (n m : trunc_index) : trunc_index := protected definition add (n m : ℕ₋₂) : ℕ₋₂ :=
trunc_index.cases_on m trunc_index.cases_on m
(trunc_index.cases_on n -2 (λn', (trunc_index.cases_on n' -2 id))) (trunc_index.cases_on n -2 (λn', (trunc_index.cases_on n' -2 id)))
(λm', trunc_index.cases_on m' (λm', trunc_index.cases_on m'
(trunc_index.cases_on n -2 id) (trunc_index.cases_on n -2 id)
(trunc_index.rec n (λn' r, succ r))) (trunc_index.rec n (λn' r, succ r)))
definition leq [reducible] (n m : trunc_index) : Type₀ := /- we give a weird name to the reflexivity step to avoid overloading le.refl
trunc_index.rec_on n (λm, unit) (λ n p m, trunc_index.rec_on m (λ p, empty) (λ m q p, p m) p) m (which can be used if types.trunc is imported) -/
inductive le (a : ℕ₋₂) : ℕ₋₂ → Type :=
definition has_le_trunc_index [instance] [priority 2000] : has_le trunc_index := | tr_refl : le a a
has_le.mk leq | step : Π {b}, le a b → le a (b.+1)
end trunc_index end trunc_index
attribute trunc_index.tr_add [reducible] definition has_le_trunc_index [instance] [priority 2000] : has_le ℕ₋₂ :=
has_le.mk trunc_index.le
attribute trunc_index.add [reducible]
infix `+2+`:65 := trunc_index.add_plus_two infix `+2+`:65 := trunc_index.add_plus_two
definition has_add_trunc_index [instance] [priority 2000] : has_add ℕ₋₂ := definition has_add_trunc_index [instance] [priority 2000] : has_add ℕ₋₂ :=
has_add.mk trunc_index.tr_add has_add.mk trunc_index.add
namespace trunc_index definition sub_two [reducible] (n : ) : ℕ₋₂ :=
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 := proof H qed
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 empty_of_succ_le_minus_two {n : trunc_index} (H : n .+1 ≤ -2) : empty := H
end trunc_index
definition trunc_index.of_nat [coercion] [reducible] (n : nat) : trunc_index :=
(nat.rec_on n -2 (λ n k, k.+1)).+2
definition sub_two [reducible] (n : nat) : trunc_index :=
nat.rec_on n -2 (λ n k, k.+1) nat.rec_on n -2 (λ n k, k.+1)
definition add_two [reducible] (n : ℕ₋₂) : :=
trunc_index.rec_on n nat.zero (λ n k, nat.succ k)
postfix ` .-2`:(max+1) := sub_two postfix ` .-2`:(max+1) := sub_two
definition trunc_index.of_nat [coercion] [reducible] (n : ) : ℕ₋₂ :=
n.-2.+2
namespace trunc_index
definition succ_le_succ {n m : ℕ₋₂} (H : n ≤ m) : n.+1 ≤ m.+1 :=
by induction H with m H IH; apply le.tr_refl; exact le.step IH
definition minus_two_le (n : ℕ₋₂) : -2 ≤ n :=
by induction n with n IH; apply le.tr_refl; exact le.step IH
protected definition le_refl (n : ℕ₋₂) : n ≤ n :=
le.tr_refl n
end trunc_index
/- truncated types -/ /- truncated types -/
/- /-
@ -91,14 +101,14 @@ namespace is_trunc
(center : A) (center : A)
(center_eq : Π(a : A), center = a) (center_eq : Π(a : A), center = a)
definition is_trunc_internal (n : trunc_index) : Type → Type := definition is_trunc_internal (n : ℕ₋₂) : Type → Type :=
trunc_index.rec_on n trunc_index.rec_on n
(λA, contr_internal A) (λA, contr_internal A)
(λn trunc_n A, (Π(x y : A), trunc_n (x = y))) (λn trunc_n A, (Π(x y : A), trunc_n (x = y)))
end is_trunc open is_trunc end is_trunc open is_trunc
structure is_trunc [class] (n : trunc_index) (A : Type) := structure is_trunc [class] (n : ℕ₋₂) (A : Type) :=
(to_internal : is_trunc_internal n A) (to_internal : is_trunc_internal n A)
open nat num is_trunc.trunc_index open nat num is_trunc.trunc_index
@ -111,12 +121,12 @@ namespace is_trunc
variables {A B : Type} variables {A B : Type}
definition is_trunc_succ_intro (A : Type) (n : trunc_index) [H : ∀x y : A, is_trunc n (x = y)] definition is_trunc_succ_intro (A : Type) (n : ℕ₋₂) [H : ∀x y : A, is_trunc n (x = y)]
: is_trunc n.+1 A := : is_trunc n.+1 A :=
is_trunc.mk (λ x y, !is_trunc.to_internal) is_trunc.mk (λ x y, !is_trunc.to_internal)
definition is_trunc_eq [instance] [priority 1200] definition is_trunc_eq [instance] [priority 1200]
(n : trunc_index) [H : is_trunc (n.+1) A] (x y : A) : is_trunc n (x = y) := (n : ℕ₋₂) [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)
/- contractibility -/ /- contractibility -/
@ -144,7 +154,7 @@ namespace is_trunc
/- truncation is upward close -/ /- truncation is upward close -/
-- n-types are also (n+1)-types -- n-types are also (n+1)-types
theorem is_trunc_succ [instance] [priority 900] (A : Type) (n : trunc_index) theorem is_trunc_succ [instance] [priority 900] (A : Type) (n : ℕ₋₂)
[H : is_trunc n A] : is_trunc (n.+1) A := [H : is_trunc n A] : is_trunc (n.+1) A :=
trunc_index.rec_on n trunc_index.rec_on n
(λ A (H : is_contr A), !is_trunc_succ_intro) (λ A (H : is_contr A), !is_trunc_succ_intro)
@ -152,43 +162,35 @@ 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_of_leq.{l} (A : Type.{l}) {n m : trunc_index} (Hnm : n ≤ m) theorem is_trunc_of_le.{l} (A : Type.{l}) {n m : ℕ₋₂} (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 begin
λ k A, trunc_index.cases_on k induction Hnm with m Hnm IH,
(λh1 h2, h2) { exact Hn},
(λk h1 h2, empty.elim (trunc_index.empty_of_succ_le_minus_two h1)), { exact _}
have step : Π (m : trunc_index) end
(IHm : Π (n : trunc_index) (A : Type), n ≤ m → is_trunc n A → is_trunc m A)
(n : trunc_index) (A : Type)
(Hnm : n ≤ m .+1) (Hn : is_trunc n A), is_trunc m .+1 A, from
λm IHm n, trunc_index.rec_on n
(λA Hnm Hn, @is_trunc_succ A m (IHm -2 A star Hn))
(λn IHn A Hnm (Hn : is_trunc n.+1 A),
@is_trunc_succ_intro A m (λx y, IHm n (x = y) (trunc_index.le_of_succ_le_succ Hnm) _)),
trunc_index.rec_on m base step n A Hnm Hn
definition is_trunc_of_imp_is_trunc {n : trunc_index} (H : A → is_trunc (n.+1) A) definition is_trunc_of_imp_is_trunc {n : ℕ₋₂} (H : A → is_trunc (n.+1) A)
: is_trunc (n.+1) A := : is_trunc (n.+1) A :=
@is_trunc_succ_intro _ _ (λx y, @is_trunc_eq _ _ (H x) x y) @is_trunc_succ_intro _ _ (λx y, @is_trunc_eq _ _ (H x) x y)
definition is_trunc_of_imp_is_trunc_of_leq {n : trunc_index} (Hn : -1 ≤ n) (H : A → is_trunc n A) definition is_trunc_of_imp_is_trunc_of_le {n : ℕ₋₂} (Hn : -1 ≤ n) (H : A → is_trunc n A)
: is_trunc n A := : is_trunc n A :=
trunc_index.rec_on n (λHn H, empty.rec _ Hn) begin
(λn IH Hn, is_trunc_of_imp_is_trunc) cases Hn with n' Hn': apply is_trunc_of_imp_is_trunc H
Hn H end
-- these must be definitions, because we need them to compute sometimes -- these must be definitions, because we need them to compute sometimes
definition is_trunc_of_is_contr (A : Type) (n : trunc_index) [H : is_contr A] : is_trunc n A := definition is_trunc_of_is_contr (A : Type) (n : ℕ₋₂) [H : is_contr A] : is_trunc n A :=
trunc_index.rec_on n H (λn H, _) trunc_index.rec_on n H (λn H, _)
definition is_trunc_succ_of_is_prop (A : Type) (n : trunc_index) [H : is_prop A] definition is_trunc_succ_of_is_prop (A : Type) (n : ℕ₋₂) [H : is_prop A]
: is_trunc (n.+1) A := : is_trunc (n.+1) A :=
is_trunc_of_leq A (show -1 ≤ n.+1, from star) is_trunc_of_le A (show -1 ≤ n.+1, from succ_le_succ (minus_two_le n))
definition is_trunc_succ_succ_of_is_set (A : Type) (n : trunc_index) [H : is_set A] definition is_trunc_succ_succ_of_is_set (A : Type) (n : ℕ₋₂) [H : is_set A]
: is_trunc (n.+2) A := : is_trunc (n.+2) A :=
@(is_trunc_of_leq A (show 0 ≤ n.+2, from proof star qed)) H is_trunc_of_le A (show 0 ≤ n.+2, from succ_le_succ (succ_le_succ (minus_two_le n)))
/- props -/ /- props -/
@ -244,10 +246,10 @@ namespace is_trunc
local attribute is_contr_unit is_prop_empty [instance] local attribute is_contr_unit is_prop_empty [instance]
definition is_trunc_unit [instance] (n : trunc_index) : is_trunc n unit := definition is_trunc_unit [instance] (n : ℕ₋₂) : is_trunc n unit :=
!is_trunc_of_is_contr !is_trunc_of_is_contr
definition is_trunc_empty [instance] (n : trunc_index) : is_trunc (n.+1) empty := definition is_trunc_empty [instance] (n : ℕ₋₂) : is_trunc (n.+1) empty :=
!is_trunc_succ_of_is_prop !is_trunc_succ_of_is_prop
/- interaction with equivalences -/ /- interaction with equivalences -/
@ -267,7 +269,7 @@ namespace is_trunc
(λa, center B) (λa, center B)
(is_equiv.adjointify (λa, center B) (λb, center A) center_eq center_eq) (is_equiv.adjointify (λa, center B) (λb, center A) center_eq center_eq)
theorem is_trunc_is_equiv_closed (n : trunc_index) (f : A → B) [H : is_equiv f] theorem is_trunc_is_equiv_closed (n : ℕ₋₂) (f : A → B) [H : is_equiv f]
[HA : is_trunc n A] : is_trunc n B := [HA : is_trunc n A] : is_trunc n B :=
trunc_index.rec_on n trunc_index.rec_on n
(λA (HA : is_contr A) B f (H : is_equiv f), is_contr_is_equiv_closed f) (λA (HA : is_contr A) B f (H : is_equiv f), is_contr_is_equiv_closed f)
@ -275,15 +277,15 @@ namespace is_trunc
IH (f⁻¹ x = f⁻¹ y) _ (x = y) (ap f⁻¹)⁻¹ !is_equiv_inv)) IH (f⁻¹ x = f⁻¹ y) _ (x = y) (ap f⁻¹)⁻¹ !is_equiv_inv))
A HA B f H A HA B f H
definition is_trunc_is_equiv_closed_rev (n : trunc_index) (f : A → B) [H : is_equiv f] definition is_trunc_is_equiv_closed_rev (n : ℕ₋₂) (f : A → B) [H : is_equiv f]
[HA : is_trunc n B] : is_trunc n A := [HA : is_trunc n B] : is_trunc n A :=
is_trunc_is_equiv_closed n f⁻¹ is_trunc_is_equiv_closed n f⁻¹
definition is_trunc_equiv_closed (n : trunc_index) (f : A ≃ B) [HA : is_trunc n A] definition is_trunc_equiv_closed (n : ℕ₋₂) (f : A ≃ B) [HA : is_trunc n A]
: is_trunc n B := : is_trunc n B :=
is_trunc_is_equiv_closed n (to_fun f) is_trunc_is_equiv_closed n (to_fun f)
definition is_trunc_equiv_closed_rev (n : trunc_index) (f : A ≃ B) [HA : is_trunc n B] definition is_trunc_equiv_closed_rev (n : ℕ₋₂) (f : A ≃ B) [HA : is_trunc n B]
: is_trunc n A := : is_trunc n A :=
is_trunc_is_equiv_closed n (to_inv f) is_trunc_is_equiv_closed n (to_inv f)
@ -299,7 +301,7 @@ namespace is_trunc
equiv_of_is_prop (iff.elim_left H) (iff.elim_right H) equiv_of_is_prop (iff.elim_left H) (iff.elim_right H)
/- truncatedness of lift -/ /- truncatedness of lift -/
definition is_trunc_lift [instance] [priority 1450] (A : Type) (n : trunc_index) definition is_trunc_lift [instance] [priority 1450] (A : Type) (n : ℕ₋₂)
[H : is_trunc n A] : is_trunc n (lift A) := [H : is_trunc n A] : is_trunc n (lift A) :=
is_trunc_equiv_closed _ !equiv_lift is_trunc_equiv_closed _ !equiv_lift
@ -326,7 +328,7 @@ namespace is_trunc
pathover_of_eq_tr !is_prop.elim pathover_of_eq_tr !is_prop.elim
definition is_trunc_pathover [instance] definition is_trunc_pathover [instance]
(n : trunc_index) [H : is_trunc (n.+1) (C a)] : is_trunc n (c =[p] c₂) := (n : ℕ₋₂) [H : is_trunc (n.+1) (C a)] : is_trunc n (c =[p] c₂) :=
is_trunc_equiv_closed_rev n !pathover_equiv_eq_tr is_trunc_equiv_closed_rev n !pathover_equiv_eq_tr
variables {p c c₂} variables {p c c₂}
@ -339,7 +341,7 @@ namespace is_trunc
end is_trunc end is_trunc
structure trunctype (n : trunc_index) := structure trunctype (n : ℕ₋₂) :=
(carrier : Type) (carrier : Type)
(struct : is_trunc n carrier) (struct : is_trunc n carrier)
@ -353,13 +355,13 @@ attribute trunctype.struct [instance] [priority 1400]
protected abbreviation Prop.mk := @trunctype.mk -1 protected abbreviation Prop.mk := @trunctype.mk -1
protected abbreviation Set.mk := @trunctype.mk (-1.+1) protected abbreviation Set.mk := @trunctype.mk (-1.+1)
protected definition trunctype.mk' [constructor] (n : trunc_index) (A : Type) [H : is_trunc n A] protected definition trunctype.mk' [constructor] (n : ℕ₋₂) (A : Type) [H : is_trunc n A]
: n-Type := : n-Type :=
trunctype.mk A H trunctype.mk A H
namespace is_trunc namespace is_trunc
definition tlift.{u v} [constructor] {n : trunc_index} (A : trunctype.{u} n) definition tlift.{u v} [constructor] {n : ℕ₋₂} (A : trunctype.{u} n)
: trunctype.{max u v} n := : trunctype.{max u v} n :=
trunctype.mk (lift A) !is_trunc_lift trunctype.mk (lift A) !is_trunc_lift

View file

@ -84,6 +84,12 @@ namespace pointed
prefix `Ω`:(max+5) := ploop_space prefix `Ω`:(max+5) := ploop_space
notation `Ω[`:95 n:0 `] `:0 A:95 := iterated_ploop_space n A notation `Ω[`:95 n:0 `] `:0 A:95 := iterated_ploop_space n A
definition iterated_ploop_space_zero [unfold_full] (A : Type*)
: Ω[0] A = A := rfl
definition iterated_ploop_space_succ [unfold_full] (k : ) (A : Type*)
: Ω[succ k] A = Ω Ω[k] A := rfl
definition rfln [constructor] [reducible] {A : Type*} {n : } : Ω[n] A := pt definition rfln [constructor] [reducible] {A : Type*} {n : } : Ω[n] A := pt
definition refln [constructor] [reducible] (A : Type*) (n : ) : Ω[n] A := pt definition refln [constructor] [reducible] (A : Type*) (n : ) : Ω[n] A := pt
definition refln_eq_refl (A : Type*) (n : ) : rfln = rfl :> Ω[succ n] A := rfl definition refln_eq_refl (A : Type*) (n : ) : rfln = rfl :> Ω[succ n] A := rfl
@ -246,10 +252,10 @@ namespace pointed
/- categorical properties of pointed maps -/ /- categorical properties of pointed maps -/
definition pid [constructor] (A : Type*) : A →* A := definition pid [constructor] [refl] (A : Type*) : A →* A :=
pmap.mk id idp pmap.mk id idp
definition pcompose [constructor] (g : B →* C) (f : A →* B) : A →* C := definition pcompose [constructor] [trans] (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)
infixr ` ∘* `:60 := pcompose infixr ` ∘* `:60 := pcompose
@ -327,12 +333,16 @@ namespace pointed
{ esimp, exact !con.left_inv⁻¹}}, { esimp, exact !con.left_inv⁻¹}},
end end
/- instances of pointed maps -/
-- The constant pointed map between any two types -- The constant pointed map between any two types
definition pconst [constructor] (A B : Type*) : A →* B := definition pconst [constructor] (A B : Type*) : A →* B :=
pmap.mk (λ a, Point B) idp pmap.mk (λ a, Point B) idp
-- the pointed type of pointed maps
definition ppmap [constructor] (A B : Type*) : Type* :=
pType.mk (A →* B) (pconst A B)
/- instances of pointed maps -/
definition ap1 [constructor] (f : A →* B) : Ω A →* Ω B := definition ap1 [constructor] (f : A →* B) : Ω A →* Ω B :=
begin begin
fconstructor, fconstructor,
@ -340,7 +350,7 @@ namespace pointed
{ esimp, apply con.left_inv} { esimp, apply con.left_inv}
end end
definition apn [unfold 3] (n : ) (f : map₊ A B) : Ω[n] A →* Ω[n] B := definition apn (n : ) (f : map₊ A B) : Ω[n] A →* Ω[n] B :=
begin begin
induction n with n IH, induction n with n IH,
{ exact f}, { exact f},
@ -362,10 +372,6 @@ namespace pointed
intro p, exact !idp_con⁻¹ intro p, exact !idp_con⁻¹
end end
-- TODO:
-- definition apn_compose (n : ) (g : B →* C) (f : A →* B) : apn n (g ∘* f) ~* apn n g ∘* apn n f :=
-- _
definition ap1_compose (g : B →* C) (f : A →* B) : ap1 (g ∘* f) ~* ap1 g ∘* ap1 f := definition ap1_compose (g : B →* C) (f : A →* B) : ap1 (g ∘* f) ~* ap1 g ∘* ap1 f :=
begin begin
induction B, induction C, induction g with g pg, induction f with f pf, esimp at *, induction B, induction C, induction g with g pg, induction f with f pf, esimp at *,
@ -484,6 +490,13 @@ namespace pointed
{ esimp, } { esimp, }
end -/ end -/
definition apn_compose (n : ) (g : B →* C) (f : A →* B) : apn n (g ∘* f) ~* apn n g ∘* apn n f :=
begin
induction n with n IH,
{ reflexivity},
{ refine ap1_phomotopy IH ⬝* _, apply ap1_compose}
end
infix ` ⬝*p `:75 := pconcat_eq infix ` ⬝*p `:75 := pconcat_eq
infix ` ⬝p* `:75 := eq_pconcat infix ` ⬝p* `:75 := eq_pconcat

View file

@ -36,11 +36,15 @@ namespace pointed
definition pequiv_of_equiv [constructor] (f : A ≃ B) (H : f pt = pt) : A ≃* B := definition pequiv_of_equiv [constructor] (f : A ≃ B) (H : f pt = pt) : A ≃* B :=
pequiv.mk f _ H pequiv.mk f _ H
protected definition pequiv.MK [constructor] (f : A →* B) (g : B →* A)
(gf : Πa, g (f a) = a) (fg : Πb, f (g b) = b) : A ≃* B :=
pequiv.mk f (adjointify f g fg gf) (respect_pt f)
definition equiv_of_pequiv [constructor] (f : A ≃* B) : A ≃ B := definition equiv_of_pequiv [constructor] (f : A ≃* B) : A ≃ B :=
equiv.mk f _ equiv.mk f _
definition to_pinv [constructor] (f : A ≃* B) : B →* A := definition to_pinv [constructor] (f : A ≃* B) : B →* A :=
pmap.mk f⁻¹ (ap f⁻¹ (respect_pt f)⁻¹ ⬝ !left_inv) pmap.mk f⁻¹ ((ap f⁻¹ (respect_pt f))⁻¹ ⬝ !left_inv)
definition pua {A B : Type*} (f : A ≃* B) : A = B := definition pua {A B : Type*} (f : A ≃* B) : A = B :=
pType_eq (equiv_of_pequiv f) !respect_pt pType_eq (equiv_of_pequiv f) !respect_pt
@ -103,7 +107,7 @@ namespace pointed
definition pleft_inv (f : A ≃* B) : f⁻¹ᵉ* ∘* f ~* pid A := definition pleft_inv (f : A ≃* B) : f⁻¹ᵉ* ∘* f ~* pid A :=
phomotopy.mk (left_inv f) phomotopy.mk (left_inv f)
abstract begin abstract begin
esimp, rewrite ap_inv, symmetry, apply con_inv_cancel_left esimp, symmetry, apply con_inv_cancel_left
end end end end
definition pright_inv (f : A ≃* B) : f ∘* f⁻¹ᵉ* ~* pid B := definition pright_inv (f : A ≃* B) : f ∘* f⁻¹ᵉ* ~* pid B :=
@ -182,9 +186,55 @@ namespace pointed
(p : h ~* f ∘* g) : f⁻¹ᵉ* ∘* h ~* g := (p : h ~* f ∘* g) : f⁻¹ᵉ* ∘* h ~* g :=
(phomotopy_pinv_left_of_phomotopy p⁻¹*)⁻¹* (phomotopy_pinv_left_of_phomotopy p⁻¹*)⁻¹*
definition phomotopy_of_phomotopy_pinv_left {f : C ≃* B} {g : A →* B} {h : A →* C} definition phomotopy_of_phomotopy_pinv_left {f : C ≃* B} {g : A →* B} {h : A →* C}
(p : h ~* f⁻¹ᵉ* ∘* g) : f ∘* h ~* g := (p : h ~* f⁻¹ᵉ* ∘* g) : f ∘* h ~* g :=
(phomotopy_of_pinv_left_phomotopy p⁻¹*)⁻¹* (phomotopy_of_pinv_left_phomotopy p⁻¹*)⁻¹*
/- pointed equivalences between particular pointed types -/
definition loop_pequiv_loop [constructor] (f : A ≃* B) : Ω A ≃* Ω B :=
pequiv.MK (ap1 f) (ap1 f⁻¹ᵉ*)
abstract begin
intro p,
refine ((ap1_compose f⁻¹ᵉ* f) p)⁻¹ ⬝ _,
refine ap1_phomotopy (pleft_inv f) p ⬝ _,
exact ap1_id p
end end
abstract begin
intro p,
refine ((ap1_compose f f⁻¹ᵉ*) p)⁻¹ ⬝ _,
refine ap1_phomotopy (pright_inv f) p ⬝ _,
exact ap1_id p
end end
definition loopn_pequiv_loopn (n : ) (f : A ≃* B) : Ω[n] A ≃* Ω[n] B :=
begin
induction n with n IH,
{ exact f},
{ exact loop_pequiv_loop IH}
end
definition pmap_functor [constructor] {A A' B B' : Type*} (f : A' →* A) (g : B →* B') :
ppmap A B →* ppmap A' B' :=
pmap.mk (λh, g ∘* h ∘* f)
abstract begin
fapply pmap_eq,
{ esimp, intro a, exact respect_pt g},
{ rewrite [▸*, ap_constant], apply idp_con}
end end
/-
definition pmap_pequiv_pmap {A A' B B' : Type*} (f : A ≃* A') (g : B ≃* B') :
ppmap A B ≃* ppmap A' B' :=
pequiv.MK (pmap_functor f⁻¹ᵉ* g) (pmap_functor f g⁻¹ᵉ*)
abstract begin
intro a, esimp, apply pmap_eq,
{ esimp, },
{ }
end end
abstract begin
end end
-/
end pointed end pointed

View file

@ -8,16 +8,152 @@ Properties of is_trunc and trunctype
-- NOTE: the fact that (is_trunc n A) is a mere proposition is proved in .prop_trunc -- NOTE: the fact that (is_trunc n A) is a mere proposition is proved in .prop_trunc
import types.pi types.eq types.equiv ..function import .pointed2 ..function algebra.order types.nat.order
open eq sigma sigma.ops pi function equiv trunctype open eq sigma sigma.ops pi function equiv trunctype
is_equiv prod is_trunc.trunc_index pointed nat is_equiv prod is_trunc.trunc_index pointed nat is_trunc algebra
namespace is_trunc namespace is_trunc
variables {A B : Type} {n : trunc_index}
namespace trunc_index
definition minus_one_le_succ (n : trunc_index) : -1 ≤ n.+1 :=
succ_le_succ (minus_two_le n)
definition zero_le_of_nat (n : ) : 0 ≤ of_nat n :=
succ_le_succ !minus_one_le_succ
open decidable
protected definition has_decidable_eq [instance] : Π(n m : ℕ₋₂), decidable (n = m)
| has_decidable_eq -2 -2 := inl rfl
| has_decidable_eq (n.+1) -2 := inr (by contradiction)
| has_decidable_eq -2 (m.+1) := inr (by contradiction)
| has_decidable_eq (n.+1) (m.+1) :=
match has_decidable_eq n m with
| inl xeqy := inl (by rewrite xeqy)
| inr xney := inr (λ h : succ n = succ m, by injection h with xeqy; exact absurd xeqy xney)
end
definition not_succ_le_minus_two {n : trunc_index} (H : n .+1 ≤ -2) : empty :=
by cases H
protected definition le_trans {n m k : ℕ₋₂} (H1 : n ≤ m) (H2 : m ≤ k) : n ≤ k :=
begin
induction H2 with k H2 IH,
{ exact H1},
{ exact le.step IH}
end
definition le_of_succ_le_succ {n m : trunc_index} (H : n.+1 ≤ m.+1) : n ≤ m :=
begin
cases H with m H',
{ apply le.tr_refl},
{ exact trunc_index.le_trans (le.step !le.tr_refl) H'}
end
theorem not_succ_le_self {n : ℕ₋₂} : ¬n.+1 ≤ n :=
begin
induction n with n IH: intro H,
{ exact not_succ_le_minus_two H},
{ exact IH (le_of_succ_le_succ H)}
end
protected definition le_antisymm {n m : ℕ₋₂} (H1 : n ≤ m) (H2 : m ≤ n) : n = m :=
begin
induction H2 with n H2 IH,
{ reflexivity},
{ exfalso, apply @not_succ_le_self n, exact trunc_index.le_trans H1 H2}
end
protected definition le_succ {n m : ℕ₋₂} (H1 : n ≤ m): n ≤ m.+1 :=
le.step H1
end trunc_index open trunc_index
definition weak_order_trunc_index [trans_instance] [reducible] : weak_order trunc_index :=
weak_order.mk le trunc_index.le_refl @trunc_index.le_trans @trunc_index.le_antisymm
namespace trunc_index
/- more theorems about truncation indices -/
definition succ_add_nat (n : ℕ₋₂) (m : ) : n.+1 + m = (n + m).+1 :=
by induction m with m IH; reflexivity; exact ap succ IH
definition nat_add_succ (n : ) (m : ℕ₋₂) : n + m.+1 = (n + m).+1 :=
begin
cases m with m, reflexivity,
cases m with m, reflexivity,
induction m with m IH, reflexivity, exact ap succ IH
end
definition add_nat_succ (n : ℕ₋₂) (m : ) : n + (nat.succ m) = (n + m).+1 :=
by reflexivity
definition nat_succ_add (n : ) (m : ℕ₋₂) : (nat.succ n) + m = (n + m).+1 :=
begin
cases m with m, reflexivity,
cases m with m, reflexivity,
induction m with m IH, reflexivity, exact ap succ IH
end
definition sub_two_add_two (n : ℕ₋₂) : sub_two (add_two n) = n :=
begin
induction n with n IH,
{ reflexivity},
{ apply ap succ IH}
end
definition add_two_sub_two (n : ) : add_two (sub_two n) = n :=
begin
induction n with n IH,
{ reflexivity},
{ apply ap nat.succ IH}
end
definition succ_sub_two (n : ) : (nat.succ n).-2 = n.-2 .+1 := rfl
definition sub_two_succ_succ (n : ) : n.-2.+1.+1 = n := rfl
definition succ_sub_two_succ (n : ) : (nat.succ n).-2.+1 = n := rfl
definition of_nat_le_of_nat {n m : } (H : n ≤ m) : (of_nat n ≤ of_nat m) :=
begin
induction H with m H IH,
{ apply le.refl},
{ exact trunc_index.le_succ IH}
end
definition sub_two_le_sub_two {n m : } (H : n ≤ m) : n.-2 ≤ m.-2 :=
begin
induction H with m H IH,
{ apply le.refl},
{ exact trunc_index.le_succ IH}
end
definition add_two_le_add_two {n m : ℕ₋₂} (H : n ≤ m) : add_two n ≤ add_two m :=
begin
induction H with m H IH,
{ reflexivity},
{ constructor, exact IH},
end
definition le_of_sub_two_le_sub_two {n m : } (H : n.-2 ≤ m.-2) : n ≤ m :=
begin
rewrite [-add_two_sub_two n, -add_two_sub_two m],
exact add_two_le_add_two H,
end
definition le_of_of_nat_le_of_nat {n m : } (H : of_nat n ≤ of_nat m) : n ≤ m :=
begin
apply le_of_sub_two_le_sub_two,
exact le_of_succ_le_succ (le_of_succ_le_succ H)
end
end trunc_index open trunc_index
variables {A B : Type} {n : ℕ₋₂}
/- theorems about trunctype -/ /- theorems about trunctype -/
protected definition trunctype.sigma_char.{l} (n : trunc_index) : protected definition trunctype.sigma_char.{l} [constructor] (n : ℕ₋₂) :
(trunctype.{l} n) ≃ (Σ (A : Type.{l}), is_trunc n A) := (trunctype.{l} n) ≃ (Σ (A : Type.{l}), is_trunc n A) :=
begin begin
fapply equiv.MK, fapply equiv.MK,
@ -27,7 +163,7 @@ namespace is_trunc
{ intro A, induction A with A1 A2, reflexivity}, { intro A, induction A with A1 A2, reflexivity},
end end
definition trunctype_eq_equiv (n : trunc_index) (A B : n-Type) : definition trunctype_eq_equiv [constructor] (n : ℕ₋₂) (A B : n-Type) :
(A = B) ≃ (carrier A = carrier B) := (A = B) ≃ (carrier A = carrier B) :=
calc calc
(A = B) ≃ (to_fun (trunctype.sigma_char n) A = to_fun (trunctype.sigma_char n) B) (A = B) ≃ (to_fun (trunctype.sigma_char n) A = to_fun (trunctype.sigma_char n) B)
@ -40,13 +176,13 @@ namespace is_trunc
(Hn : -1 ≤ n) : is_trunc n A := (Hn : -1 ≤ n) : is_trunc n A :=
begin begin
induction n with n, induction n with n,
{exact !empty.elim Hn}, {exfalso, exact not_succ_le_minus_two Hn},
{apply is_trunc_succ_intro, intro a a', {apply is_trunc_succ_intro, intro a a',
fapply @is_trunc_is_equiv_closed_rev _ _ n (ap f)} fapply @is_trunc_is_equiv_closed_rev _ _ n (ap f)}
end end
theorem is_trunc_is_retraction_closed (f : A → B) [Hf : is_retraction f] theorem is_trunc_is_retraction_closed (f : A → B) [Hf : is_retraction f]
(n : trunc_index) [HA : is_trunc n A] : is_trunc n B := (n : ℕ₋₂) [HA : is_trunc n A] : is_trunc n B :=
begin begin
revert A B f Hf HA, revert A B f Hf HA,
induction n with n IH, induction n with n IH,
@ -67,22 +203,22 @@ namespace is_trunc
definition is_embedding_to_fun (A B : Type) : is_embedding (@to_fun A B) := definition is_embedding_to_fun (A B : Type) : is_embedding (@to_fun A B) :=
λf f', !is_equiv_ap_to_fun λf f', !is_equiv_ap_to_fun
theorem is_trunc_trunctype [instance] (n : trunc_index) : is_trunc n.+1 (n-Type) := theorem is_trunc_trunctype [instance] (n : ℕ₋₂) : is_trunc n.+1 (n-Type) :=
begin begin
apply is_trunc_succ_intro, intro X Y, apply is_trunc_succ_intro, intro X Y,
fapply is_trunc_equiv_closed, fapply is_trunc_equiv_closed,
{apply equiv.symm, apply trunctype_eq_equiv}, { apply equiv.symm, apply trunctype_eq_equiv},
fapply is_trunc_equiv_closed, fapply is_trunc_equiv_closed,
{apply equiv.symm, apply eq_equiv_equiv}, { apply equiv.symm, apply eq_equiv_equiv},
induction n, induction n,
{apply @is_contr_of_inhabited_prop, { apply @is_contr_of_inhabited_prop,
{apply is_trunc_is_embedding_closed, { apply is_trunc_is_embedding_closed,
{apply is_embedding_to_fun} , { apply is_embedding_to_fun} ,
{exact unit.star}}, { reflexivity}},
{apply equiv_of_is_contr_of_is_contr}}, { apply equiv_of_is_contr_of_is_contr}},
{apply is_trunc_is_embedding_closed, { apply is_trunc_is_embedding_closed,
{apply is_embedding_to_fun}, { apply is_embedding_to_fun},
{exact unit.star}} { apply minus_one_le_succ}}
end end
@ -126,15 +262,15 @@ namespace is_trunc
is_set_of_double_neg_elim (λa b, by_contradiction) is_set_of_double_neg_elim (λa b, by_contradiction)
end end
theorem is_trunc_of_axiom_K_of_leq {A : Type} (n : trunc_index) (H : -1 ≤ n) theorem is_trunc_of_axiom_K_of_le {A : Type} (n : ℕ₋₂) (H : -1 ≤ n)
(K : Π(a : A), is_trunc n (a = a)) : is_trunc (n.+1) A := (K : Π(a : A), is_trunc n (a = a)) : is_trunc (n.+1) A :=
@is_trunc_succ_intro _ _ (λa b, is_trunc_of_imp_is_trunc_of_leq H (λp, eq.rec_on p !K)) @is_trunc_succ_intro _ _ (λa b, is_trunc_of_imp_is_trunc_of_le H (λp, eq.rec_on p !K))
theorem is_trunc_succ_of_is_trunc_loop (Hn : -1 ≤ n) (Hp : Π(a : A), is_trunc n (a = a)) theorem is_trunc_succ_of_is_trunc_loop (Hn : -1 ≤ n) (Hp : Π(a : A), is_trunc n (a = a))
: is_trunc (n.+1) A := : is_trunc (n.+1) A :=
begin begin
apply is_trunc_succ_intro, intros a a', apply is_trunc_succ_intro, intros a a',
apply is_trunc_of_imp_is_trunc_of_leq Hn, intro p, apply is_trunc_of_imp_is_trunc_of_le Hn, intro p,
induction p, apply Hp induction p, apply Hp
end end
@ -154,7 +290,8 @@ namespace is_trunc
{ apply is_trunc_succ_iff_is_trunc_loop, apply le.refl}, { apply is_trunc_succ_iff_is_trunc_loop, apply le.refl},
{ apply pi_iff_pi, intro a, esimp, apply is_prop_iff_is_contr, reflexivity}}, { apply pi_iff_pi, intro a, esimp, apply is_prop_iff_is_contr, reflexivity}},
{ intro A, esimp [iterated_ploop_space], { intro A, esimp [iterated_ploop_space],
transitivity _, apply @is_trunc_succ_iff_is_trunc_loop @n, esimp, constructor, transitivity _,
{ apply @is_trunc_succ_iff_is_trunc_loop @n, esimp, apply minus_one_le_succ},
apply pi_iff_pi, intro a, transitivity _, apply IH, apply pi_iff_pi, intro a, transitivity _, apply IH,
transitivity _, apply 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,
@ -178,27 +315,27 @@ namespace is_trunc
apply iff.mp !is_trunc_iff_is_contr_loop H apply iff.mp !is_trunc_iff_is_contr_loop H
end end
end is_trunc open is_trunc end is_trunc open is_trunc is_trunc.trunc_index
namespace trunc namespace trunc
variable {A : Type} variable {A : Type}
protected definition code (n : trunc_index) (aa aa' : trunc n.+1 A) : n-Type := protected definition code (n : ℕ₋₂) (aa aa' : trunc n.+1 A) : n-Type :=
trunc.rec_on aa (λa, trunc.rec_on aa' (λa', trunctype.mk' n (trunc n (a = a')))) trunc.rec_on aa (λa, trunc.rec_on aa' (λa', trunctype.mk' n (trunc n (a = a'))))
protected definition encode (n : trunc_index) (aa aa' : trunc n.+1 A) : aa = aa' → trunc.code n aa aa' := protected definition encode (n : ℕ₋₂) (aa aa' : trunc n.+1 A) : aa = aa' → trunc.code n aa aa' :=
begin begin
intro p, induction p, induction aa with a, esimp [trunc.code,trunc.rec_on], exact (tr idp) intro p, induction p, induction aa with a, esimp [trunc.code,trunc.rec_on], exact (tr idp)
end end
protected definition decode (n : trunc_index) (aa aa' : trunc n.+1 A) : trunc.code n aa aa' → aa = aa' := protected definition decode (n : ℕ₋₂) (aa aa' : trunc n.+1 A) : trunc.code n aa aa' → aa = aa' :=
begin begin
induction aa' with a', induction aa with a, induction aa' with a', induction aa with a,
esimp [trunc.code, trunc.rec_on], intro x, esimp [trunc.code, trunc.rec_on], intro x,
induction x with p, exact ap tr p, induction x with p, exact ap tr p,
end end
definition trunc_eq_equiv [constructor] (n : trunc_index) (aa aa' : trunc n.+1 A) definition trunc_eq_equiv [constructor] (n : ℕ₋₂) (aa aa' : trunc n.+1 A)
: aa = aa' ≃ trunc.code n aa aa' := : aa = aa' ≃ trunc.code n aa aa' :=
begin begin
fapply equiv.MK, fapply equiv.MK,
@ -212,18 +349,18 @@ namespace trunc
{ intro p, induction p, apply (trunc.rec_on aa), intro a, exact idp}, { intro p, induction p, apply (trunc.rec_on aa), intro a, exact idp},
end end
definition tr_eq_tr_equiv [constructor] (n : trunc_index) (a a' : A) definition tr_eq_tr_equiv [constructor] (n : ℕ₋₂) (a a' : A)
: (tr a = tr a' :> trunc n.+1 A) ≃ trunc n (a = a') := : (tr a = tr a' :> trunc n.+1 A) ≃ trunc n (a = a') :=
!trunc_eq_equiv !trunc_eq_equiv
definition is_trunc_trunc_of_is_trunc [instance] [priority 500] (A : Type) definition is_trunc_trunc_of_is_trunc [instance] [priority 500] (A : Type)
(n m : trunc_index) [H : is_trunc n A] : is_trunc n (trunc m A) := (n m : ℕ₋₂) [H : is_trunc n A] : is_trunc n (trunc m A) :=
begin begin
revert A m H, eapply (trunc_index.rec_on n), revert A m H, eapply (trunc_index.rec_on n),
{ clear n, intro A m H, apply is_contr_equiv_closed, { clear n, intro A m H, apply is_contr_equiv_closed,
{ apply equiv.symm, apply trunc_equiv, apply (@is_trunc_of_leq _ -2), exact unit.star} }, { apply equiv.symm, apply trunc_equiv, apply (@is_trunc_of_le _ -2), apply minus_two_le} },
{ clear n, intro n IH A m H, induction m with m, { clear n, intro n IH A m H, induction m with m,
{ apply (@is_trunc_of_leq _ -2), exact unit.star}, { apply (@is_trunc_of_le _ -2), apply minus_two_le},
{ apply is_trunc_succ_intro, intro aa aa', { apply is_trunc_succ_intro, intro aa aa',
apply (@trunc.rec_on _ _ _ aa (λy, !is_trunc_succ_of_is_prop)), apply (@trunc.rec_on _ _ _ aa (λy, !is_trunc_succ_of_is_prop)),
eapply (@trunc.rec_on _ _ _ aa' (λy, !is_trunc_succ_of_is_prop)), eapply (@trunc.rec_on _ _ _ aa' (λy, !is_trunc_succ_of_is_prop)),
@ -238,10 +375,28 @@ namespace trunc
!trunc_equiv (f a) !trunc_equiv (f a)
/- transport over a truncated family -/ /- transport over a truncated family -/
definition trunc_transport {a a' : A} {P : A → Type} (p : a = a') (n : trunc_index) (x : P a) definition trunc_transport {a a' : A} {P : A → Type} (p : a = a') (n : ℕ₋₂) (x : P a)
: transport (λa, trunc n (P a)) p (tr x) = tr (p ▸ x) := : transport (λa, trunc n (P a)) p (tr x) = tr (p ▸ x) :=
by induction p; reflexivity by induction p; reflexivity
definition trunc_trunc_equiv_left [constructor] (A : Type) (n m : ℕ₋₂) (H : n ≤ m)
: trunc n (trunc m A) ≃ trunc n A :=
begin
note H2 := is_trunc_of_le (trunc n A) H,
fapply equiv.MK,
{ intro x, induction x with x, induction x with x, exact tr x},
{ intro x, induction x with x, exact tr (tr x)},
{ intro x, induction x with x, reflexivity},
{ intro x, induction x with x, induction x with x, reflexivity}
end
definition trunc_trunc_equiv_right [constructor] (A : Type) (n m : ℕ₋₂) (H : n ≤ m)
: trunc m (trunc n A) ≃ trunc n A :=
begin
apply trunc_equiv,
exact is_trunc_of_le _ H,
end
definition image [constructor] {A B : Type} (f : A → B) (b : B) : Prop := ∥ fiber f b ∥ definition image [constructor] {A B : Type} (f : A → B) (b : B) : Prop := ∥ fiber f b ∥
definition image.mk [constructor] {A B : Type} {f : A → B} {b : B} (a : A) (p : f a = b) definition image.mk [constructor] {A B : Type} {f : A → B} {b : B} (a : A) (p : f a = b)
@ -249,13 +404,27 @@ namespace trunc
tr (fiber.mk a p) tr (fiber.mk a p)
-- truncation of pointed types -- truncation of pointed types
definition ptrunc [constructor] (n : trunc_index) (X : Type*) : n-Type* := definition ptrunc [constructor] (n : ℕ₋₂) (X : Type*) : n-Type* :=
ptrunctype.mk (trunc n X) _ (tr pt) ptrunctype.mk (trunc n X) _ (tr pt)
definition ptrunc_functor [constructor] {X Y : Type*} (n : ℕ₋₂) (f : X →* Y) definition ptrunc_functor [constructor] {X Y : Type*} (n : ℕ₋₂) (f : X →* Y)
: ptrunc n X →* ptrunc n Y := : ptrunc n X →* ptrunc n Y :=
pmap.mk (trunc_functor n f) (ap tr (respect_pt f)) pmap.mk (trunc_functor n f) (ap tr (respect_pt f))
definition loop_ptrunc_pequiv [constructor] (n : ℕ₋₂) (A : Type*) :
Ω (ptrunc (n+1) A) ≃* ptrunc n (Ω A) :=
pequiv_of_equiv !tr_eq_tr_equiv idp
definition iterated_loop_ptrunc_pequiv [constructor] (n : ℕ₋₂) (k : ) (A : Type*) :
Ω[k] (ptrunc (n+k) A) ≃* ptrunc n (Ω[k] A) :=
begin
revert n, induction k with k IH: intro n,
{ reflexivity},
{ refine _ ⬝e* loop_ptrunc_pequiv n (Ω[k] A),
rewrite [iterated_ploop_space_succ], apply loop_pequiv_loop,
refine _ ⬝e* IH (n.+1),
rewrite succ_add_nat}
end
end trunc open trunc end trunc open trunc

View file

@ -45,7 +45,7 @@ namespace univ
definition not_is_set_type : ¬is_set Type.{u} := definition not_is_set_type : ¬is_set Type.{u} :=
assume H : is_set Type, assume H : is_set Type,
absurd (is_trunc_is_embedding_closed lift star) not_is_set_type0 absurd (is_trunc_is_embedding_closed lift !trunc_index.minus_one_le_succ) not_is_set_type0
definition not_double_negation_elimination0 : ¬Π(A : Type₀), ¬¬A → A := definition not_double_negation_elimination0 : ¬Π(A : Type₀), ¬¬A → A :=
begin begin

View file

@ -211,7 +211,7 @@ have h₂ : ∀ b, f b = b → a ≤ b, from
take b, take b,
suppose f b = b, suppose f b = b,
have b ∈ {u | f u ≤ u}, from have b ∈ {u | f u ≤ u}, from
show f b ≤ b, by rewrite this; apply le.refl, show f b ≤ b, by rewrite this,
Inf_le this, Inf_le this,
exists.intro a (and.intro h₁ h₂) exists.intro a (and.intro h₁ h₂)
@ -233,7 +233,7 @@ have h₁ : f a = a, from
have h₂ : ∀ b, f b = b → b ≤ a, from have h₂ : ∀ b, f b = b → b ≤ a, from
take b, take b,
suppose f b = b, suppose f b = b,
have b ≤ f b, by rewrite this; apply le.refl, have b ≤ f b, by rewrite this,
le_Sup this, le_Sup this,
exists.intro a (and.intro h₁ h₂) exists.intro a (and.intro h₁ h₂)
@ -266,12 +266,12 @@ lemma Inf_singleton {a : A} : ⨅'{a} = a :=
have ⨅'{a} ≤ a, from have ⨅'{a} ≤ a, from
Inf_le !mem_insert, Inf_le !mem_insert,
have a ≤ ⨅'{a}, from have a ≤ ⨅'{a}, from
le_Inf (take b, suppose b ∈ '{a}, have b = a, from eq_of_mem_singleton this, by rewrite this; apply le.refl), le_Inf (take b, suppose b ∈ '{a}, have b = a, from eq_of_mem_singleton this, by rewrite this),
le.antisymm `⨅'{a} ≤ a` `a ≤ ⨅'{a}` le.antisymm `⨅'{a} ≤ a` `a ≤ ⨅'{a}`
lemma Sup_singleton {a : A} : ⨆'{a} = a := lemma Sup_singleton {a : A} : ⨆'{a} = a :=
have ⨆'{a} ≤ a, from have ⨆'{a} ≤ a, from
Sup_le (take b, suppose b ∈ '{a}, have b = a, from eq_of_mem_singleton this, by rewrite this; apply le.refl), Sup_le (take b, suppose b ∈ '{a}, have b = a, from eq_of_mem_singleton this, by rewrite this),
have a ≤ ⨆'{a}, from have a ≤ ⨆'{a}, from
le_Sup !mem_insert, le_Sup !mem_insert,
le.antisymm `⨆'{a} ≤ a` `a ≤ ⨆'{a}` le.antisymm `⨆'{a} ≤ a` `a ≤ ⨆'{a}`

View file

@ -20,7 +20,7 @@ structure weak_order [class] (A : Type) extends has_le A :=
section section
variables [weak_order A] variables [weak_order A]
theorem le.refl (a : A) : a ≤ a := !weak_order.le_refl theorem le.refl [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_of_eq {a b : A} (H : a = b) : a ≤ b := H ▸ le.refl a
@ -315,13 +315,13 @@ section
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
(assume H : a ≤ b, by rewrite [↑min, if_pos H]; apply le.refl) (assume H : a ≤ b, by rewrite [↑min, if_pos H])
(assume H : ¬ a ≤ b, by rewrite [↑min, if_neg H]; apply le_of_lt (lt_of_not_ge H)) (assume H : ¬ a ≤ b, by rewrite [↑min, if_neg H]; apply le_of_lt (lt_of_not_ge H))
theorem min_le_right (a b : A) : min a b ≤ b := theorem min_le_right (a b : A) : min a b ≤ b :=
by_cases by_cases
(assume H : a ≤ b, by rewrite [↑min, if_pos H]; apply H) (assume H : a ≤ b, by rewrite [↑min, if_pos H]; apply H)
(assume H : ¬ a ≤ b, by rewrite [↑min, if_neg H]; apply le.refl) (assume H : ¬ a ≤ b, by rewrite [↑min, if_neg H])
theorem le_min {a b c : A} (H₁ : c ≤ a) (H₂ : c ≤ b) : c ≤ min a b := theorem le_min {a b c : A} (H₁ : c ≤ a) (H₂ : c ≤ b) : c ≤ min a b :=
by_cases by_cases
@ -331,11 +331,11 @@ section
theorem le_max_left (a b : A) : a ≤ max a b := theorem le_max_left (a b : A) : a ≤ max a b :=
by_cases by_cases
(assume H : a ≤ b, by rewrite [↑max, if_pos H]; apply H) (assume H : a ≤ b, by rewrite [↑max, if_pos H]; apply H)
(assume H : ¬ a ≤ b, by rewrite [↑max, if_neg H]; apply le.refl) (assume H : ¬ a ≤ b, by rewrite [↑max, if_neg H])
theorem le_max_right (a b : A) : b ≤ max a b := theorem le_max_right (a b : A) : b ≤ max a b :=
by_cases by_cases
(assume H : a ≤ b, by rewrite [↑max, if_pos H]; apply le.refl) (assume H : a ≤ b, by rewrite [↑max, if_pos H])
(assume H : ¬ a ≤ b, by rewrite [↑max, if_neg H]; apply le_of_lt (lt_of_not_ge H)) (assume H : ¬ a ≤ b, by rewrite [↑max, if_neg H]; apply le_of_lt (lt_of_not_ge H))
theorem max_le {a b c : A} (H₁ : a ≤ c) (H₂ : b ≤ c) : max a b ≤ c := theorem max_le {a b c : A} (H₁ : a ≤ c) (H₂ : b ≤ c) : max a b ≤ c :=

View file

@ -79,7 +79,7 @@ section decidable_linear_ordered_cancel_comm_monoid_B
induction s with a' s' a'nins' ih, induction s with a' s' a'nins' ih,
{exact false.elim (not_mem_empty a H)}, {exact false.elim (not_mem_empty a H)},
cases (decidable.em (s' = ∅)) with s'empty s'nempty, cases (decidable.em (s' = ∅)) with s'empty s'nempty,
{rewrite [s'empty at *, Max_singleton, eq_of_mem_singleton H], apply le.refl}, {rewrite [s'empty at *, Max_singleton, eq_of_mem_singleton H]},
rewrite [Max_insert f a'nins' s'nempty], rewrite [Max_insert f a'nins' s'nempty],
cases (eq_or_mem_of_mem_insert H) with aeqa' ains', cases (eq_or_mem_of_mem_insert H) with aeqa' ains',
{rewrite aeqa', apply le_max_left}, {rewrite aeqa', apply le_max_left},
@ -154,7 +154,7 @@ section decidable_linear_ordered_cancel_comm_monoid_B
induction s with a' s' a'nins' ih, induction s with a' s' a'nins' ih,
{exact false.elim (not_mem_empty a H)}, {exact false.elim (not_mem_empty a H)},
cases (decidable.em (s' = ∅)) with s'empty s'nempty, cases (decidable.em (s' = ∅)) with s'empty s'nempty,
{rewrite [s'empty at *, Min_singleton, eq_of_mem_singleton H], apply le.refl}, {rewrite [s'empty at *, Min_singleton, eq_of_mem_singleton H]},
rewrite [Min_insert f a'nins' s'nempty], rewrite [Min_insert f a'nins' s'nempty],
cases (eq_or_mem_of_mem_insert H) with aeqa' ains', cases (eq_or_mem_of_mem_insert H) with aeqa' ains',
{rewrite aeqa', apply min_le_left}, {rewrite aeqa', apply min_le_left},

View file

@ -82,7 +82,7 @@ section decidable_linear_ordered_comm_group
proposition abs_Sum_le (f : A → B) (s : finset A) : abs (∑ x ∈ s, f x) ≤ (∑ x ∈ s, abs (f x)) := proposition abs_Sum_le (f : A → B) (s : finset A) : abs (∑ x ∈ s, f x) ≤ (∑ x ∈ s, abs (f x)) :=
begin begin
induction s with a s ans ih, induction s with a s ans ih,
{rewrite [+Sum_empty, abs_zero], apply le.refl}, {rewrite [+Sum_empty, abs_zero]},
rewrite [Sum_insert_of_not_mem f ans, Sum_insert_of_not_mem _ ans], rewrite [Sum_insert_of_not_mem f ans, Sum_insert_of_not_mem _ ans],
apply le.trans, apply le.trans,
apply abs_add_le_abs_add_abs, apply abs_add_le_abs_add_abs,
@ -141,13 +141,12 @@ section ordered_comm_group
begin begin
cases (em (finite s)) with fins nfins, cases (em (finite s)) with fins nfins,
{induction fins with a s fins ans ih, {induction fins with a s fins ans ih,
{rewrite +Sum_empty; apply le.refl}, {rewrite +Sum_empty},
{rewrite [Sum_insert_of_not_mem f ans, Sum_insert_of_not_mem g ans], {rewrite [Sum_insert_of_not_mem f ans, Sum_insert_of_not_mem g ans],
have H1 : f a ≤ g a, from H !mem_insert, have H1 : f a ≤ g a, from H !mem_insert,
have H2 : (∑ x ∈ s, f x) ≤ (∑ x ∈ s, g x), from ih (forall_of_forall_insert H), have H2 : (∑ x ∈ s, f x) ≤ (∑ x ∈ s, g x), from ih (forall_of_forall_insert H),
apply add_le_add H1 H2}}, apply add_le_add H1 H2}},
rewrite [+Sum_of_not_finite nfins], rewrite [+Sum_of_not_finite nfins]
apply le.refl
end end
proposition Sum_nonneg (f : A → B) {s : set A} (H : ∀₀ x ∈ s, f x ≥ 0) : proposition Sum_nonneg (f : A → B) {s : set A} (H : ∀₀ x ∈ s, f x ≥ 0) :
@ -171,9 +170,9 @@ section decidable_linear_ordered_comm_group
begin begin
cases (em (finite s)) with fins nfins, cases (em (finite s)) with fins nfins,
rotate 1, rotate 1,
{rewrite [+Sum_of_not_finite nfins, abs_zero], apply le.refl}, {rewrite [+Sum_of_not_finite nfins, abs_zero]},
induction fins with a s fins ans ih, induction fins with a s fins ans ih,
{rewrite [+Sum_empty, abs_zero], apply le.refl}, {rewrite [+Sum_empty, abs_zero]},
rewrite [Sum_insert_of_not_mem f ans, Sum_insert_of_not_mem _ ans], rewrite [Sum_insert_of_not_mem f ans, Sum_insert_of_not_mem _ ans],
apply le.trans, apply le.trans,
apply abs_add_le_abs_add_abs, apply abs_add_le_abs_add_abs,

View file

@ -115,7 +115,6 @@ theorem equiv_abs_of_ge_zero {s : seq} (Hs : regular s) (Hz : s_le zero s) : s_a
apply add_le_add, apply add_le_add,
repeat (apply inv_ge_of_le; apply Hn), repeat (apply inv_ge_of_le; apply Hn),
krewrite pnat.add_halves, krewrite pnat.add_halves,
apply le.refl
end end
theorem equiv_neg_abs_of_le_zero {s : seq} (Hs : regular s) (Hz : s_le s zero) : s_abs s ≡ sneg s := theorem equiv_neg_abs_of_le_zero {s : seq} (Hs : regular s) (Hz : s_le s zero) : s_abs s ≡ sneg s :=
@ -146,7 +145,6 @@ theorem equiv_neg_abs_of_le_zero {s : seq} (Hs : regular s) (Hz : s_le s zero) :
apply add_le_add, apply add_le_add,
repeat (apply inv_ge_of_le; apply Hn), repeat (apply inv_ge_of_le; apply Hn),
krewrite pnat.add_halves, krewrite pnat.add_halves,
apply le.refl,
let Hneg' := lt_of_not_ge Hneg, let Hneg' := lt_of_not_ge Hneg,
rewrite [abs_of_neg Hneg', ↑sneg, sub_neg_eq_add, neg_add_eq_sub, sub_self, rewrite [abs_of_neg Hneg', ↑sneg, sub_neg_eq_add, neg_add_eq_sub, sub_self,
abs_zero], abs_zero],
@ -289,7 +287,6 @@ theorem cauchy_with_rate_of_converges_to_with_rate {X : r_seq} {a : } {N :
xrewrite -of_rat_add, xrewrite -of_rat_add,
apply of_rat_le_of_rat_of_le, apply of_rat_le_of_rat_of_le,
krewrite pnat.add_halves, krewrite pnat.add_halves,
apply rat.le_refl
end end
private definition Nb (M : + → +) := λ k, pnat.max (3 * k) (M (2 * k)) private definition Nb (M : + → +) := λ k, pnat.max (3 * k) (M (2 * k))
@ -425,7 +422,6 @@ theorem converges_to_with_rate_of_cauchy_with_rate : converges_to_with_rate X li
apply Nb_spec_left, apply Nb_spec_left,
rewrite -*pnat.mul_assoc, rewrite -*pnat.mul_assoc,
krewrite pnat.p_add_fractions, krewrite pnat.p_add_fractions,
apply rat.le_refl
end end
end lim_seq end lim_seq

View file

@ -330,7 +330,7 @@ theorem neg_infty_le : ∀ v, -∞ ≤ v
protected theorem le_refl : ∀ u : ereal, u ≤ u protected theorem le_refl : ∀ u : ereal, u ≤ u
| ∞ := trivial | ∞ := trivial
| -∞ := trivial | -∞ := trivial
| (of_real x) := by rewrite [of_real_le_of_real]; apply le.refl | (of_real x) := by rewrite [of_real_le_of_real]
protected theorem le_trans : ∀ u v w : ereal, u ≤ v → v ≤ w → u ≤ w protected theorem le_trans : ∀ u v w : ereal, u ≤ v → v ≤ w → u ≤ w
| u v ∞ H1 H2 := !le_infty | u v ∞ H1 H2 := !le_infty

View file

@ -229,6 +229,10 @@ order for the change to take effect."
("m=" . ("")) ("m=" . (""))
("?=" . ("")) ("?=" . (""))
;; Pointed types in HoTT
("pequiv" . ("≃*"))
("phomotopy" . ("~*"))
("1" . ("")) ("1" . (""))
("2" . ("")) ("2" . (""))
("3" . ("")) ("3" . (""))
@ -338,7 +342,7 @@ order for the change to take effect."
("dot" . ("")) ("dot" . (""))
("sy" . ("⁻¹")) ("sy" . ("⁻¹"))
("inv" . ("⁻¹")) ("inv" . ("⁻¹"))
("-1" . ("⁻¹")) ("-1" . ("⁻¹" "₋₁"))
("-1e" . ("⁻¹ᵉ")) ("-1e" . ("⁻¹ᵉ"))
("-1f" . ("⁻¹ᶠ")) ("-1f" . ("⁻¹ᶠ"))
("-1g" . ("⁻¹ᵍ")) ("-1g" . ("⁻¹ᵍ"))
@ -351,7 +355,7 @@ order for the change to take effect."
("-1s" . ("⁻¹ˢ")) ("-1s" . ("⁻¹ˢ"))
("-1v" . ("⁻¹ᵛ")) ("-1v" . ("⁻¹ᵛ"))
("-1E" . ("⁻¹ᴱ")) ("-1E" . ("⁻¹ᴱ"))
("-2" . ("⁻²")) ("-2" . ("⁻²" "₋₂"))
("-2o" . ("⁻²ᵒ")) ("-2o" . ("⁻²ᵒ"))
("-3" . ("⁻³")) ("-3" . ("⁻³"))
("qed" . ("")) ("qed" . (""))
@ -598,6 +602,7 @@ order for the change to take effect."
("Nat" . ("")) ("Nat" . (""))
("N" . ("")) ("N" . (""))
("N-2" . ("ℕ₋₂")) ("N-2" . ("ℕ₋₂"))
("N-1" . ("ℕ₋₁"))
("int" . ("")) ("int" . (""))
("Int" . ("")) ("Int" . (""))
("Z" . ("")) ("Z" . (""))