fix(hott): adjust library to new apply tactic semantics

This commit is contained in:
Leonardo de Moura 2015-02-06 17:27:56 -08:00
parent c2a296b1de
commit 0f34f4d4a1
8 changed files with 20 additions and 16 deletions

View file

@ -25,6 +25,8 @@ namespace category
definition path_of_iso {a b : ob} : a ≅ b → a = b := definition path_of_iso {a b : ob} : a ≅ b → a = b :=
iso_of_path⁻¹ 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 := definition ob_1_type : is_trunc nat.zero .+1 ob :=
begin begin
apply is_trunc_succ, intros (a, b), apply is_trunc_succ, intros (a, b),

View file

@ -28,6 +28,8 @@ namespace precategory
variables {C : Precategory} {a b c : C} 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 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 -- TODO: Decide whether just to use funext for this theorem or

View file

@ -44,6 +44,8 @@ namespace functor
apply (prod.rec_on P1), intros (d3, d4), apply idp, apply (prod.rec_on P1), intros (d3, d4), apply idp,
end end
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
protected definition strict_cat_has_functor_hset protected definition strict_cat_has_functor_hset
[HD : is_hset (objects D)] : is_hset (functor C D) := [HD : is_hset (objects D)] : is_hset (functor C D) :=
begin begin

View file

@ -40,6 +40,8 @@ namespace morphism
intro S, apply (sigma.rec_on S), intros (f, H), apply idp, intro S, apply (sigma.rec_on S), intros (f, H), apply idp,
end 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 -- The statement "f is an isomorphism" is a mere proposition
definition is_hprop_of_is_iso : is_hset (is_iso f) := definition is_hprop_of_is_iso : is_hset (is_iso f) :=
begin begin

View file

@ -47,6 +47,7 @@ namespace natural_transformation
apply (dcongr_arg2 (@natural_transformation.mk C D F G) p₁ p₂), apply (dcongr_arg2 (@natural_transformation.mk C D F G) p₁ p₂),
end end
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
protected definition assoc (η₃ : H ⟹ I) (η₂ : G ⟹ H) (η₁ : F ⟹ G) : protected definition assoc (η₃ : H ⟹ I) (η₂ : G ⟹ H) (η₁ : F ⟹ G) :
η₃ ∘n (η₂ ∘n η₁) = (η₃ ∘n η₂) ∘n η₁ := η₃ ∘n (η₂ ∘n η₁) = (η₃ ∘n η₂) ∘n η₁ :=

View file

@ -28,9 +28,9 @@ namespace truncation
(Π (x y : A), is_trunc n (x = y)) ≃ (is_trunc (n .+1) A) := (Π (x y : A), is_trunc n (x = y)) ≃ (is_trunc (n .+1) A) :=
begin begin
fapply equiv.mk, fapply equiv.mk,
intro H, apply is_trunc_succ, exact H, intro H, apply is_trunc_succ,
fapply is_equiv.adjointify, 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 H, apply (is_trunc.rec_on H), intro Hint, apply idp,
intro P, intro P,
exact sorry, exact sorry,
@ -56,9 +56,6 @@ namespace truncation
apply trunc_equiv, apply trunc_equiv,
apply equiv.to_is_equiv, apply equiv.to_is_equiv,
apply is_trunc.pi_char, apply is_trunc.pi_char,
apply trunc_pi, intro a,
apply trunc_pi, intro b,
apply (IH (a = b)),
end end
end truncation end truncation

View file

@ -99,7 +99,7 @@ namespace pi
apply (trunc_index.rec_on n), apply (trunc_index.rec_on n),
intros (B, H), intros (B, H),
fapply is_contr.mk, 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 f, apply path_pi,
intro x, apply (contr (f x)), intro x, apply (contr (f x)),
intros (n, IH, B, H), 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) := [H : ∀a, is_trunc n (f a = g a)] : is_trunc n (f = g) :=
begin begin
apply trunc_equiv', apply equiv.symm, apply trunc_equiv', apply equiv.symm,
apply path_equiv_homotopy, apply path_equiv_homotopy
apply trunc_pi, exact H,
end end
end pi end pi

View file

@ -451,14 +451,13 @@ namespace sigma
intros (A, B, HA, HB), intros (A, B, HA, HB),
fapply trunc_equiv', fapply trunc_equiv',
apply equiv.symm, apply equiv.symm,
apply equiv_center_of_contr, apply HA, --should be solved by term synthesis apply equiv_center_of_contr,
apply HB,
intros (n, IH, A, B, HA, HB), intros (n, IH, A, B, HA, HB),
fapply is_trunc_succ, intros (u, v), fapply is_trunc_succ, intros (u, v),
fapply trunc_equiv', fapply trunc_equiv',
apply equiv_sigma_path, apply equiv_sigma_path,
apply IH, apply IH,
apply succ_is_trunc, assumption, apply succ_is_trunc,
intro p, intro p,
show is_trunc n (p ▹ u .2 = v .2), from show is_trunc n (p ▹ u .2 = v .2), from
succ_is_trunc n (p ▹ u.2) (v.2), succ_is_trunc n (p ▹ u.2) (v.2),