fix(hott): adjust library to new apply tactic semantics
This commit is contained in:
parent
c2a296b1de
commit
0f34f4d4a1
8 changed files with 20 additions and 16 deletions
|
@ -25,6 +25,8 @@ namespace category
|
|||
definition path_of_iso {a b : ob} : a ≅ b → a = b :=
|
||||
iso_of_path⁻¹
|
||||
|
||||
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
|
||||
|
||||
definition ob_1_type : is_trunc nat.zero .+1 ob :=
|
||||
begin
|
||||
apply is_trunc_succ, intros (a, b),
|
||||
|
|
|
@ -28,6 +28,8 @@ namespace precategory
|
|||
|
||||
variables {C : Precategory} {a b c : C}
|
||||
|
||||
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
|
||||
|
||||
theorem compose_op {f : hom a b} {g : hom b c} : f ∘op g = g ∘ f := idp
|
||||
|
||||
-- TODO: Decide whether just to use funext for this theorem or
|
||||
|
|
|
@ -44,6 +44,8 @@ namespace functor
|
|||
apply (prod.rec_on P1), intros (d3, d4), apply idp,
|
||||
end
|
||||
|
||||
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
|
||||
|
||||
protected definition strict_cat_has_functor_hset
|
||||
[HD : is_hset (objects D)] : is_hset (functor C D) :=
|
||||
begin
|
||||
|
|
|
@ -40,6 +40,8 @@ namespace morphism
|
|||
intro S, apply (sigma.rec_on S), intros (f, H), apply idp,
|
||||
end
|
||||
|
||||
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
|
||||
|
||||
-- The statement "f is an isomorphism" is a mere proposition
|
||||
definition is_hprop_of_is_iso : is_hset (is_iso f) :=
|
||||
begin
|
||||
|
|
|
@ -26,11 +26,11 @@ namespace natural_transformation
|
|||
(λ a, η a ∘ θ a)
|
||||
(λ a b f,
|
||||
calc
|
||||
H f ∘ (η a ∘ θ a) = (H f ∘ η a) ∘ θ a : assoc
|
||||
... = (η b ∘ G f) ∘ θ a : naturality η f
|
||||
... = η b ∘ (G f ∘ θ a) : assoc
|
||||
... = η b ∘ (θ b ∘ F f) : naturality θ f
|
||||
... = (η b ∘ θ b) ∘ F f : assoc)
|
||||
H f ∘ (η a ∘ θ a) = (H f ∘ η a) ∘ θ a : assoc
|
||||
... = (η b ∘ G f) ∘ θ a : naturality η f
|
||||
... = η b ∘ (G f ∘ θ a) : assoc
|
||||
... = η b ∘ (θ b ∘ F f) : naturality θ f
|
||||
... = (η b ∘ θ b) ∘ F f : assoc)
|
||||
|
||||
infixr `∘n`:60 := compose
|
||||
|
||||
|
@ -47,6 +47,7 @@ namespace natural_transformation
|
|||
apply (dcongr_arg2 (@natural_transformation.mk C D F G) p₁ p₂),
|
||||
end
|
||||
|
||||
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
|
||||
|
||||
protected definition assoc (η₃ : H ⟹ I) (η₂ : G ⟹ H) (η₁ : F ⟹ G) :
|
||||
η₃ ∘n (η₂ ∘n η₁) = (η₃ ∘n η₂) ∘n η₁ :=
|
||||
|
|
|
@ -28,9 +28,9 @@ namespace truncation
|
|||
(Π (x y : A), is_trunc n (x = y)) ≃ (is_trunc (n .+1) A) :=
|
||||
begin
|
||||
fapply equiv.mk,
|
||||
intro H, apply is_trunc_succ, exact H,
|
||||
intro H, apply is_trunc_succ,
|
||||
fapply is_equiv.adjointify,
|
||||
intros (H, x, y), apply succ_is_trunc, exact H,
|
||||
intros (H, x, y), apply succ_is_trunc,
|
||||
intro H, apply (is_trunc.rec_on H), intro Hint, apply idp,
|
||||
intro P,
|
||||
exact sorry,
|
||||
|
@ -56,9 +56,6 @@ namespace truncation
|
|||
apply trunc_equiv,
|
||||
apply equiv.to_is_equiv,
|
||||
apply is_trunc.pi_char,
|
||||
apply trunc_pi, intro a,
|
||||
apply trunc_pi, intro b,
|
||||
apply (IH (a = b)),
|
||||
end
|
||||
|
||||
end truncation
|
||||
|
|
|
@ -99,7 +99,7 @@ namespace pi
|
|||
apply (trunc_index.rec_on n),
|
||||
intros (B, H),
|
||||
fapply is_contr.mk,
|
||||
intro a, apply center, apply H, --remove "apply H" when term synthesis works correctly
|
||||
intro a, apply center,
|
||||
intro f, apply path_pi,
|
||||
intro x, apply (contr (f x)),
|
||||
intros (n, IH, B, H),
|
||||
|
@ -116,8 +116,7 @@ namespace pi
|
|||
[H : ∀a, is_trunc n (f a = g a)] : is_trunc n (f = g) :=
|
||||
begin
|
||||
apply trunc_equiv', apply equiv.symm,
|
||||
apply path_equiv_homotopy,
|
||||
apply trunc_pi, exact H,
|
||||
apply path_equiv_homotopy
|
||||
end
|
||||
|
||||
end pi
|
||||
|
|
|
@ -451,14 +451,13 @@ namespace sigma
|
|||
intros (A, B, HA, HB),
|
||||
fapply trunc_equiv',
|
||||
apply equiv.symm,
|
||||
apply equiv_center_of_contr, apply HA, --should be solved by term synthesis
|
||||
apply HB,
|
||||
apply equiv_center_of_contr,
|
||||
intros (n, IH, A, B, HA, HB),
|
||||
fapply is_trunc_succ, intros (u, v),
|
||||
fapply trunc_equiv',
|
||||
apply equiv_sigma_path,
|
||||
apply IH,
|
||||
apply succ_is_trunc, assumption,
|
||||
apply succ_is_trunc,
|
||||
intro p,
|
||||
show is_trunc n (p ▹ u .2 = v .2), from
|
||||
succ_is_trunc n (p ▹ u.2) (v.2),
|
||||
|
|
Loading…
Reference in a new issue