2017-07-08 12:39:23 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2017 Floris van Doorn and Ulrik Buchholtz. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: Floris van Doorn, Ulrik Buchholtz
|
|
|
|
|
|
|
|
|
|
Truncatedness and truncation of spectra
|
|
|
|
|
-/
|
|
|
|
|
|
2017-06-28 12:14:48 +00:00
|
|
|
|
import .spectrum .EM
|
|
|
|
|
|
2017-06-28 16:19:23 +00:00
|
|
|
|
namespace int
|
|
|
|
|
|
2017-07-01 19:02:31 +00:00
|
|
|
|
-- TODO move this
|
2017-07-08 12:26:34 +00:00
|
|
|
|
open nat algebra
|
2017-07-01 12:02:23 +00:00
|
|
|
|
section
|
|
|
|
|
private definition maxm2_le.lemma₁ {n k : ℕ} : n+(1:int) + -[1+ k] ≤ n :=
|
|
|
|
|
le.intro (
|
2017-07-08 12:26:34 +00:00
|
|
|
|
calc n + 1 + -[1+ k] + k
|
|
|
|
|
= n + 1 + (-(k + 1)) + k : by reflexivity
|
|
|
|
|
... = n + 1 + (-1 - k) + k : by krewrite (neg_add_rev k 1)
|
|
|
|
|
... = n + 1 + (-1 - k + k) : add.assoc
|
|
|
|
|
... = n + 1 + (-1 + -k + k) : by reflexivity
|
|
|
|
|
... = n + 1 + (-1 + (-k + k)) : add.assoc
|
|
|
|
|
... = n + 1 + (-1 + 0) : add.left_inv
|
|
|
|
|
... = n + (1 + (-1 + 0)) : add.assoc
|
|
|
|
|
... = n : int.add_zero)
|
2017-07-01 12:02:23 +00:00
|
|
|
|
|
|
|
|
|
private definition maxm2_le.lemma₂ {n : ℕ} {k : ℤ} : -[1+ n] + 1 + k ≤ k :=
|
|
|
|
|
le.intro (
|
2017-07-08 12:26:34 +00:00
|
|
|
|
calc -[1+ n] + 1 + k + n
|
|
|
|
|
= - (n + 1) + 1 + k + n : by reflexivity
|
|
|
|
|
... = -n - 1 + 1 + k + n : by rewrite (neg_add n 1)
|
|
|
|
|
... = -n + (-1 + 1) + k + n : by krewrite (int.add_assoc (-n) (-1) 1)
|
|
|
|
|
... = -n + 0 + k + n : add.left_inv 1
|
|
|
|
|
... = -n + k + n : int.add_zero
|
|
|
|
|
... = k + -n + n : int.add_comm
|
|
|
|
|
... = k + (-n + n) : int.add_assoc
|
|
|
|
|
... = k + 0 : add.left_inv n
|
|
|
|
|
... = k : int.add_zero)
|
|
|
|
|
|
|
|
|
|
open trunc_index
|
2017-07-01 12:02:23 +00:00
|
|
|
|
|
|
|
|
|
definition maxm2_le (n k : ℤ) : maxm2 (n+1+k) ≤ (maxm1m1 n).+1+2+(maxm1m1 k) :=
|
2017-06-29 19:06:47 +00:00
|
|
|
|
begin
|
2017-07-01 12:02:23 +00:00
|
|
|
|
rewrite [-(maxm1_eq_succ n)],
|
2017-06-29 19:06:47 +00:00
|
|
|
|
induction n with n n,
|
2017-07-01 12:02:23 +00:00
|
|
|
|
{ induction k with k k,
|
|
|
|
|
{ induction k with k IH,
|
|
|
|
|
{ apply le.tr_refl },
|
|
|
|
|
{ exact succ_le_succ IH } },
|
|
|
|
|
{ exact trunc_index.le_trans (maxm2_monotone maxm2_le.lemma₁)
|
|
|
|
|
(maxm2_le_maxm1 n) } },
|
|
|
|
|
{ krewrite (add_plus_two_comm -1 (maxm1m1 k)),
|
|
|
|
|
rewrite [-(maxm1_eq_succ k)],
|
|
|
|
|
exact trunc_index.le_trans (maxm2_monotone maxm2_le.lemma₂)
|
|
|
|
|
(maxm2_le_maxm1 k) }
|
2017-06-29 19:06:47 +00:00
|
|
|
|
end
|
2017-06-28 16:19:23 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end int
|
|
|
|
|
|
2017-06-28 12:14:48 +00:00
|
|
|
|
open int trunc eq is_trunc lift unit pointed equiv is_equiv algebra EM
|
2017-06-28 16:19:23 +00:00
|
|
|
|
|
2017-06-28 12:14:48 +00:00
|
|
|
|
namespace spectrum
|
|
|
|
|
|
2017-07-01 19:02:31 +00:00
|
|
|
|
definition ptrunc_maxm2_change_int {k l : ℤ} (p : k = l) (X : Type*)
|
2017-06-29 19:06:47 +00:00
|
|
|
|
: ptrunc (maxm2 k) X ≃* ptrunc (maxm2 l) X :=
|
2017-07-05 19:40:15 +00:00
|
|
|
|
ptrunc_change_index (ap maxm2 p) X
|
2017-06-28 22:02:09 +00:00
|
|
|
|
|
2017-07-01 19:02:31 +00:00
|
|
|
|
definition is_trunc_maxm2_change_int {k l : ℤ} (X : pType) (p : k = l)
|
|
|
|
|
: is_trunc (maxm2 k) X → is_trunc (maxm2 l) X :=
|
|
|
|
|
by induction p; exact id
|
|
|
|
|
|
|
|
|
|
definition is_trunc_maxm2_loop {k : ℤ} {A : Type*} (H : is_trunc (maxm2 (k+1)) A) :
|
|
|
|
|
is_trunc (maxm2 k) (Ω A) :=
|
2017-06-28 12:14:48 +00:00
|
|
|
|
begin
|
2017-07-01 19:02:31 +00:00
|
|
|
|
induction k with k k,
|
|
|
|
|
apply is_trunc_loop, exact H,
|
|
|
|
|
apply is_contr_loop,
|
|
|
|
|
cases k with k,
|
|
|
|
|
{ exact H },
|
|
|
|
|
{ apply is_trunc_succ, apply is_trunc_succ, exact H }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition loop_ptrunc_maxm2_pequiv {k : ℤ} {l : ℕ₋₂} (p : maxm2 (k+1) = l) (X : Type*) :
|
|
|
|
|
Ω (ptrunc l X) ≃* ptrunc (maxm2 k) (Ω X) :=
|
|
|
|
|
begin
|
|
|
|
|
induction p,
|
2017-06-28 12:14:48 +00:00
|
|
|
|
induction k with k k,
|
2017-06-29 19:06:47 +00:00
|
|
|
|
{ exact loop_ptrunc_pequiv k X },
|
2017-06-30 14:29:52 +00:00
|
|
|
|
{ refine pequiv_of_is_contr _ _ _ !is_trunc_trunc,
|
|
|
|
|
apply is_contr_loop,
|
2017-06-30 14:14:55 +00:00
|
|
|
|
cases k with k,
|
|
|
|
|
{ change is_set (trunc 0 X), apply _ },
|
|
|
|
|
{ change is_set (trunc -2 X), apply _ }}
|
2017-06-28 12:14:48 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-01 19:02:31 +00:00
|
|
|
|
definition ptrunc_elim_phomotopy2 [constructor] (k : ℕ₋₂) {A B : Type*} {f g : A →* B} (H₁ : is_trunc k B)
|
|
|
|
|
(H₂ : is_trunc k B) (p : f ~* g) : @ptrunc.elim k A B H₁ f ~* @ptrunc.elim k A B H₂ g :=
|
|
|
|
|
begin
|
|
|
|
|
fapply phomotopy.mk,
|
|
|
|
|
{ intro x, induction x with a, exact p a },
|
|
|
|
|
{ exact to_homotopy_pt p }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition loop_ptrunc_maxm2_pequiv_ptrunc_elim' {k : ℤ} {l : ℕ₋₂} (p : maxm2 (k+1) = l)
|
|
|
|
|
{A B : Type*} (f : A →* B) {H : is_trunc l B} :
|
|
|
|
|
Ω→ (ptrunc.elim l f) ∘* (loop_ptrunc_maxm2_pequiv p A)⁻¹ᵉ* ~*
|
|
|
|
|
@ptrunc.elim (maxm2 k) _ _ (is_trunc_maxm2_loop (is_trunc_of_eq p⁻¹ H)) (Ω→ f) :=
|
|
|
|
|
begin
|
|
|
|
|
induction p, induction k with k k,
|
|
|
|
|
{ refine pwhisker_right _ (ap1_phomotopy _) ⬝* @(ap1_ptrunc_elim k f) H,
|
|
|
|
|
apply ptrunc_elim_phomotopy2, reflexivity },
|
|
|
|
|
{ apply phomotopy_of_is_contr_cod, exact is_trunc_maxm2_loop H }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition loop_ptrunc_maxm2_pequiv_ptrunc_elim {k : ℤ} {l : ℕ₋₂} (p : maxm2 (k+1) = l)
|
|
|
|
|
{A B : Type*} (f : A →* B) {H1 : is_trunc ((maxm2 k).+1) B } {H2 : is_trunc l B} :
|
|
|
|
|
Ω→ (ptrunc.elim l f) ∘* (loop_ptrunc_maxm2_pequiv p A)⁻¹ᵉ* ~* ptrunc.elim (maxm2 k) (Ω→ f) :=
|
|
|
|
|
begin
|
|
|
|
|
induction p, induction k with k k: esimp at H1,
|
|
|
|
|
{ refine pwhisker_right _ (ap1_phomotopy _) ⬝* ap1_ptrunc_elim k f,
|
|
|
|
|
apply ptrunc_elim_phomotopy2, reflexivity },
|
|
|
|
|
{ apply phomotopy_of_is_contr_cod }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition loop_ptrunc_maxm2_pequiv_ptr {k : ℤ} {l : ℕ₋₂} (p : maxm2 (k+1) = l) (A : Type*) :
|
|
|
|
|
Ω→ (ptr l A) ~* (loop_ptrunc_maxm2_pequiv p A)⁻¹ᵉ* ∘* ptr (maxm2 k) (Ω A) :=
|
|
|
|
|
begin
|
|
|
|
|
induction p, induction k with k k,
|
|
|
|
|
{ exact ap1_ptr k A },
|
|
|
|
|
{ apply phomotopy_pinv_left_of_phomotopy, apply phomotopy_of_is_contr_cod, apply is_trunc_trunc }
|
|
|
|
|
end
|
|
|
|
|
|
2017-06-29 19:06:47 +00:00
|
|
|
|
definition is_trunc_of_is_trunc_maxm2 (k : ℤ) (X : Type)
|
|
|
|
|
: is_trunc (maxm2 k) X → is_trunc (max0 k) X :=
|
|
|
|
|
λ H, @is_trunc_of_le X _ _ (maxm2_le_maxm0 k) H
|
2017-06-28 12:14:48 +00:00
|
|
|
|
|
2017-06-28 14:21:11 +00:00
|
|
|
|
definition strunc [constructor] (k : ℤ) (E : spectrum) : spectrum :=
|
2017-06-29 19:06:47 +00:00
|
|
|
|
spectrum.MK (λ(n : ℤ), ptrunc (maxm2 (k + n)) (E n))
|
|
|
|
|
(λ(n : ℤ), ptrunc_pequiv_ptrunc (maxm2 (k + n)) (equiv_glue E n)
|
2017-07-01 19:02:31 +00:00
|
|
|
|
⬝e* (loop_ptrunc_maxm2_pequiv (ap maxm2 (add.assoc k n 1)) (E (n+1)))⁻¹ᵉ*)
|
2017-06-28 12:14:48 +00:00
|
|
|
|
|
2017-06-28 14:21:11 +00:00
|
|
|
|
definition strunc_change_int [constructor] {k l : ℤ} (E : spectrum) (p : k = l) :
|
|
|
|
|
strunc k E →ₛ strunc l E :=
|
2017-06-28 12:14:48 +00:00
|
|
|
|
begin induction p, reflexivity end
|
|
|
|
|
|
2017-06-30 14:14:55 +00:00
|
|
|
|
definition is_strunc [reducible] (k : ℤ) (E : spectrum) : Type :=
|
2017-06-29 19:06:47 +00:00
|
|
|
|
Π (n : ℤ), is_trunc (maxm2 (k + n)) (E n)
|
2017-06-28 14:21:11 +00:00
|
|
|
|
|
|
|
|
|
definition is_strunc_change_int {k l : ℤ} (E : spectrum) (p : k = l)
|
|
|
|
|
: is_strunc k E → is_strunc l E :=
|
|
|
|
|
begin induction p, exact id end
|
|
|
|
|
|
2017-07-01 13:26:38 +00:00
|
|
|
|
definition is_strunc_of_le {k l : ℤ} (E : spectrum) (H : k ≤ l)
|
|
|
|
|
: is_strunc k E → is_strunc l E :=
|
|
|
|
|
begin
|
|
|
|
|
intro T, intro n, exact is_trunc_of_le (E n)
|
|
|
|
|
(maxm2_monotone (algebra.add_le_add_right H n))
|
|
|
|
|
end
|
|
|
|
|
|
2017-07-02 00:14:18 +00:00
|
|
|
|
definition is_strunc_pequiv_closed {k : ℤ} {E F : spectrum} (H : Πn, E n ≃* F n)
|
|
|
|
|
(H2 : is_strunc k E) : is_strunc k F :=
|
|
|
|
|
λn, is_trunc_equiv_closed (maxm2 (k + n)) (H n)
|
|
|
|
|
|
2017-06-28 14:21:11 +00:00
|
|
|
|
definition is_strunc_strunc (k : ℤ) (E : spectrum)
|
|
|
|
|
: is_strunc k (strunc k E) :=
|
2017-06-29 19:06:47 +00:00
|
|
|
|
λ n, is_trunc_trunc (maxm2 (k + n)) (E n)
|
|
|
|
|
|
2017-07-07 21:32:57 +00:00
|
|
|
|
definition is_strunc_strunc_of_is_strunc (k : ℤ) {l : ℤ} {E : spectrum} (H : is_strunc l E)
|
|
|
|
|
: is_strunc l (strunc k E) :=
|
|
|
|
|
λ n, !is_trunc_trunc_of_is_trunc
|
|
|
|
|
|
2017-07-01 19:02:31 +00:00
|
|
|
|
definition str [constructor] (k : ℤ) (E : spectrum) : E →ₛ strunc k E :=
|
|
|
|
|
smap.mk (λ n, ptr (maxm2 (k + n)) (E n))
|
|
|
|
|
abstract begin
|
|
|
|
|
intro n,
|
|
|
|
|
apply psquare_of_phomotopy,
|
|
|
|
|
refine !passoc ⬝* pwhisker_left _ !ptr_natural ⬝* _,
|
|
|
|
|
refine !passoc⁻¹* ⬝* pwhisker_right _ !loop_ptrunc_maxm2_pequiv_ptr⁻¹*,
|
|
|
|
|
end end
|
|
|
|
|
|
|
|
|
|
definition strunc_elim [constructor] {k : ℤ} {E F : spectrum} (f : E →ₛ F)
|
|
|
|
|
(H : is_strunc k F) : strunc k E →ₛ F :=
|
|
|
|
|
smap.mk (λn, ptrunc.elim (maxm2 (k + n)) (f n))
|
|
|
|
|
abstract begin
|
|
|
|
|
intro n,
|
|
|
|
|
apply psquare_of_phomotopy,
|
|
|
|
|
symmetry,
|
|
|
|
|
refine !passoc⁻¹* ⬝* pwhisker_right _ !loop_ptrunc_maxm2_pequiv_ptrunc_elim' ⬝* _,
|
|
|
|
|
refine @(ptrunc_elim_ptrunc_functor _ _ _) _ ⬝* _,
|
|
|
|
|
refine _ ⬝* @(ptrunc_elim_pcompose _ _ _) _ _,
|
|
|
|
|
apply is_trunc_maxm2_loop,
|
|
|
|
|
refine is_trunc_of_eq _ (H (n+1)), exact ap maxm2 (add.assoc k n 1)⁻¹,
|
|
|
|
|
apply ptrunc_elim_phomotopy2,
|
|
|
|
|
apply phomotopy_of_psquare,
|
|
|
|
|
apply ptranspose,
|
|
|
|
|
apply smap.glue_square
|
|
|
|
|
end end
|
2017-06-28 14:21:11 +00:00
|
|
|
|
|
2017-06-30 14:14:55 +00:00
|
|
|
|
definition strunc_functor [constructor] (k : ℤ) {E F : spectrum} (f : E →ₛ F) :
|
|
|
|
|
strunc k E →ₛ strunc k F :=
|
2017-07-01 19:02:31 +00:00
|
|
|
|
strunc_elim (str k F ∘ₛ f) (is_strunc_strunc k F)
|
2017-06-30 14:14:55 +00:00
|
|
|
|
|
2017-07-03 12:37:02 +00:00
|
|
|
|
definition is_strunc_sunit (n : ℤ) : is_strunc n sunit :=
|
|
|
|
|
begin
|
|
|
|
|
intro k, apply is_trunc_lift, apply is_trunc_unit
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
open option
|
|
|
|
|
definition is_strunc_add_point_spectrum {X : Type} {Y : X → spectrum} {s₀ : ℤ}
|
|
|
|
|
(H : Πx, is_strunc s₀ (Y x)) : Π(x : X₊), is_strunc s₀ (add_point_spectrum Y x)
|
2017-07-08 14:49:30 +00:00
|
|
|
|
| (some x) := proof H x qed
|
|
|
|
|
| none := begin intro k, apply is_trunc_lift, apply is_trunc_unit end
|
2017-07-03 12:37:02 +00:00
|
|
|
|
|
2017-06-28 14:21:11 +00:00
|
|
|
|
definition is_strunc_EM_spectrum (G : AbGroup)
|
|
|
|
|
: is_strunc 0 (EM_spectrum G) :=
|
|
|
|
|
begin
|
|
|
|
|
intro n, induction n with n n,
|
|
|
|
|
{ -- case ≥ 0
|
2017-06-29 19:06:47 +00:00
|
|
|
|
apply is_trunc_maxm2_change_int (EM G n) (zero_add n)⁻¹,
|
2017-06-28 14:21:11 +00:00
|
|
|
|
apply is_trunc_EM },
|
2017-06-30 14:14:55 +00:00
|
|
|
|
{ change is_contr (EM_spectrum G (-[1+n])),
|
|
|
|
|
induction n with n IH,
|
2017-06-28 14:21:11 +00:00
|
|
|
|
{ -- case = -1
|
2017-06-30 14:14:55 +00:00
|
|
|
|
apply is_contr_loop, exact is_trunc_EM G 0 },
|
2017-06-28 14:21:11 +00:00
|
|
|
|
{ -- case < -1
|
2017-06-30 14:14:55 +00:00
|
|
|
|
apply is_trunc_loop, apply is_trunc_succ, exact IH }}
|
2017-06-28 14:21:11 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition trivial_shomotopy_group_of_is_strunc (E : spectrum)
|
|
|
|
|
{n k : ℤ} (K : is_strunc n E) (H : n < k)
|
|
|
|
|
: is_contr (πₛ[k] E) :=
|
2017-06-28 16:19:23 +00:00
|
|
|
|
let m := n + (2 - k) in
|
|
|
|
|
have I : m < 2, from
|
|
|
|
|
calc
|
|
|
|
|
m = (2 - k) + n : int.add_comm n (2 - k)
|
|
|
|
|
... < (2 - k) + k : add_lt_add_left H (2 - k)
|
|
|
|
|
... = 2 : sub_add_cancel 2 k,
|
|
|
|
|
@trivial_homotopy_group_of_is_trunc (E (2 - k)) (max0 m) 2
|
2017-06-29 19:06:47 +00:00
|
|
|
|
(is_trunc_of_is_trunc_maxm2 m (E (2 - k)) (K (2 - k)))
|
2017-06-28 16:19:23 +00:00
|
|
|
|
(nat.succ_le_succ (max0_le_of_le (le_sub_one_of_lt I)))
|
2017-06-28 14:21:11 +00:00
|
|
|
|
|
2017-07-01 12:02:23 +00:00
|
|
|
|
structure truncspectrum (n : ℤ) :=
|
|
|
|
|
(carrier : spectrum)
|
|
|
|
|
(struct : is_strunc n carrier)
|
|
|
|
|
|
|
|
|
|
notation n `-spectrum` := truncspectrum n
|
|
|
|
|
|
|
|
|
|
attribute truncspectrum.carrier [coercion]
|
|
|
|
|
|
|
|
|
|
definition genspectrum_of_truncspectrum (n : ℤ)
|
|
|
|
|
: n-spectrum → gen_spectrum +ℤ :=
|
|
|
|
|
λ E, truncspectrum.carrier E
|
|
|
|
|
|
|
|
|
|
attribute genspectrum_of_truncspectrum [coercion]
|
|
|
|
|
|
|
|
|
|
section
|
|
|
|
|
|
|
|
|
|
open is_conn
|
|
|
|
|
|
|
|
|
|
definition is_conn_maxm1_of_maxm2 (A : Type*) (n : ℤ)
|
|
|
|
|
: is_conn (maxm2 n) A → is_conn (maxm1m1 n).+1 A :=
|
|
|
|
|
begin
|
|
|
|
|
intro H, induction n with n n,
|
|
|
|
|
{ exact H },
|
|
|
|
|
{ exact is_conn_minus_one A (tr pt) }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition is_trunc_maxm2_of_maxm1 (A : Type*) (n : ℤ)
|
|
|
|
|
: is_trunc (maxm1m1 n).+1 A → is_trunc (maxm2 n) A :=
|
|
|
|
|
begin
|
|
|
|
|
intro H, induction n with n n,
|
|
|
|
|
{ exact H},
|
|
|
|
|
{ apply is_contr_of_merely_prop,
|
|
|
|
|
{ exact H },
|
|
|
|
|
{ exact tr pt } }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
variables (A : Type*) (n : ℤ) [H : is_conn (maxm2 n) A]
|
|
|
|
|
include H
|
|
|
|
|
|
2017-07-05 14:42:27 +00:00
|
|
|
|
definition is_trunc_maxm2_ppi (k l : ℤ) (H3 : l ≤ n+1+k) (P : A → Type*)
|
|
|
|
|
(H2 : Πa, is_trunc (maxm2 l) (P a)) : is_trunc (maxm2 k) (Π*(a : A), P a) :=
|
2017-07-01 12:02:23 +00:00
|
|
|
|
is_trunc_maxm2_of_maxm1 (Π*(a : A), P a) k
|
2017-07-05 14:42:27 +00:00
|
|
|
|
(@is_trunc_ppi_of_is_conn A (maxm1m1 n) (is_conn_maxm1_of_maxm2 A n H) (maxm1m1 k) _
|
|
|
|
|
(le.trans (maxm2_monotone H3) (maxm2_le n k)) P H2)
|
2017-07-01 12:02:23 +00:00
|
|
|
|
|
2017-07-05 14:42:27 +00:00
|
|
|
|
definition is_strunc_spi_of_is_conn (k l : ℤ) (H3 : l ≤ n+1+k) (P : A → spectrum)
|
|
|
|
|
(H2 : Πa, is_strunc l (P a)) : is_strunc k (spi A P) :=
|
2017-07-01 12:02:23 +00:00
|
|
|
|
begin
|
|
|
|
|
intro m, unfold spi,
|
2017-07-05 14:42:27 +00:00
|
|
|
|
exact is_trunc_maxm2_ppi A n (k+m) _ (le.trans (add_le_add_right H3 _)
|
|
|
|
|
(le_of_eq (add.assoc (n+1) k m))) (λ a, P a m) (λa, H2 a m)
|
2017-07-01 12:02:23 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end
|
|
|
|
|
|
2017-07-05 14:42:27 +00:00
|
|
|
|
definition is_strunc_spi_of_le {A : Type*} (k n : ℤ) (H : n ≤ k) (P : A → spectrum)
|
|
|
|
|
(H2 : Πa, is_strunc n (P a)) : is_strunc k (spi A P) :=
|
2017-07-01 13:26:38 +00:00
|
|
|
|
begin
|
|
|
|
|
assert K : n ≤ -[1+ 0] + 1 + k,
|
|
|
|
|
{ krewrite (int.zero_add k), exact H },
|
2017-07-05 14:42:27 +00:00
|
|
|
|
{ exact @is_strunc_spi_of_is_conn A (-[1+ 0]) (is_conn.is_conn_minus_two A) k _ K P H2 }
|
2017-07-01 13:26:38 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-02 00:14:18 +00:00
|
|
|
|
definition is_strunc_spi {A : Type*} (n : ℤ) (P : A → spectrum) (H : Πa, is_strunc n (P a))
|
|
|
|
|
: is_strunc n (spi A P) :=
|
2017-07-05 14:42:27 +00:00
|
|
|
|
is_strunc_spi_of_le n n !le.refl P H
|
2017-07-02 00:14:18 +00:00
|
|
|
|
|
|
|
|
|
definition is_strunc_sp_cotensor (n : ℤ) (A : Type*) {Y : spectrum} (H : is_strunc n Y)
|
|
|
|
|
: is_strunc n (sp_cotensor A Y) :=
|
|
|
|
|
is_strunc_pequiv_closed (λn, !pppi_pequiv_ppmap) (is_strunc_spi n (λa, Y) (λa, H))
|
|
|
|
|
|
2017-07-07 21:32:57 +00:00
|
|
|
|
definition is_strunc_sp_ucotensor (n : ℤ) (A : Type) {Y : spectrum} (H : is_strunc n Y)
|
|
|
|
|
: is_strunc n (sp_ucotensor A Y) :=
|
|
|
|
|
λk, !pi.is_trunc_arrow
|
|
|
|
|
|
2017-06-28 12:14:48 +00:00
|
|
|
|
end spectrum
|