@@ -512,7 +512,7 @@ begin
512
512
end
513
513
514
514
/-!
515
- `homotopy.mk_inductive` allows us to build a homotopy inductively,
515
+ `homotopy.mk_inductive` allows us to build a homotopy of chain complexes inductively,
516
516
so that as we construct each component, we have available the previous two components,
517
517
and the fact that they satisfy the homotopy condition.
518
518
@@ -645,6 +645,140 @@ end
645
645
646
646
end mk_inductive
647
647
648
+ /-!
649
+ `homotopy.mk_coinductive` allows us to build a homotopy of cochain complexes inductively,
650
+ so that as we construct each component, we have available the previous two components,
651
+ and the fact that they satisfy the homotopy condition.
652
+ -/
653
+ section mk_coinductive
654
+
655
+ variables {P Q : cochain_complex V ℕ}
656
+
657
+ @[simp] lemma d_next_cochain_complex (f : Π i j, P.X i ⟶ Q.X j) (j : ℕ) :
658
+ d_next j f = P.d _ _ ≫ f (j+1 ) j :=
659
+ begin
660
+ dsimp [d_next],
661
+ simp only [cochain_complex.next],
662
+ refl,
663
+ end
664
+
665
+ @[simp] lemma prev_d_succ_cochain_complex (f : Π i j, P.X i ⟶ Q.X j) (i : ℕ) :
666
+ prev_d (i+1 ) f = f (i+1 ) _ ≫ Q.d i (i+1 ) :=
667
+ begin
668
+ dsimp [prev_d],
669
+ simp [cochain_complex.prev_nat_succ],
670
+ refl,
671
+ end
672
+
673
+ @[simp] lemma prev_d_zero_cochain_complex (f : Π i j, P.X i ⟶ Q.X j) :
674
+ prev_d 0 f = 0 :=
675
+ begin
676
+ dsimp [prev_d],
677
+ simp only [cochain_complex.prev_nat_zero],
678
+ refl,
679
+ end
680
+
681
+ variables (e : P ⟶ Q)
682
+ (zero : P.X 1 ⟶ Q.X 0 )
683
+ (comm_zero : e.f 0 = P.d 0 1 ≫ zero)
684
+ (one : P.X 2 ⟶ Q.X 1 )
685
+ (comm_one : e.f 1 = zero ≫ Q.d 0 1 + P.d 1 2 ≫ one)
686
+ (succ : ∀ (n : ℕ)
687
+ (p : Σ' (f : P.X (n+1 ) ⟶ Q.X n) (f' : P.X (n+2 ) ⟶ Q.X (n+1 )),
688
+ e.f (n+1 ) = f ≫ Q.d n (n+1 ) + P.d (n+1 ) (n+2 ) ≫ f'),
689
+ Σ' f'' : P.X (n+3 ) ⟶ Q.X (n+2 ), e.f (n+2 ) = p.2 .1 ≫ Q.d (n+1 ) (n+2 ) + P.d (n+2 ) (n+3 ) ≫ f'')
690
+
691
+ include comm_one comm_zero succ
692
+
693
+ /--
694
+ An auxiliary construction for `mk_coinductive`.
695
+
696
+ Here we build by induction a family of diagrams,
697
+ but don't require at the type level that these successive diagrams actually agree.
698
+ They do in fact agree, and we then capture that at the type level (i.e. by constructing a homotopy)
699
+ in `mk_coinductive`.
700
+
701
+ At this stage, we don't check the homotopy condition in degree 0,
702
+ because it "falls off the end", and is easier to treat using `X_next` and `X_prev`,
703
+ which we do in `mk_inductive_aux₂`.
704
+ -/
705
+ @[simp, nolint unused_arguments]
706
+ def mk_coinductive_aux₁ :
707
+ Π n, Σ' (f : P.X (n+1 ) ⟶ Q.X n) (f' : P.X (n+2 ) ⟶ Q.X (n+1 )),
708
+ e.f (n+1 ) = f ≫ Q.d n (n+1 ) + P.d (n+1 ) (n+2 ) ≫ f'
709
+ | 0 := ⟨zero, one, comm_one⟩
710
+ | 1 := ⟨one, (succ 0 ⟨zero, one, comm_one⟩).1 , (succ 0 ⟨zero, one, comm_one⟩).2 ⟩
711
+ | (n+2 ) :=
712
+ ⟨(mk_coinductive_aux₁ (n+1 )).2 .1 ,
713
+ (succ (n+1 ) (mk_coinductive_aux₁ (n+1 ))).1 ,
714
+ (succ (n+1 ) (mk_coinductive_aux₁ (n+1 ))).2 ⟩
715
+
716
+ section
717
+
718
+ variable [has_zero_object V]
719
+
720
+ /--
721
+ An auxiliary construction for `mk_inductive`.
722
+ -/
723
+ @[simp]
724
+ def mk_coinductive_aux₂ :
725
+ Π n, Σ' (f : P.X n ⟶ Q.X_prev n) (f' : P.X_next n ⟶ Q.X n),
726
+ e.f n = f ≫ Q.d_to n + P.d_from n ≫ f'
727
+ | 0 := ⟨0 , (P.X_next_iso rfl).hom ≫ zero, by simpa using comm_zero⟩
728
+ | (n+1 ) := let I := mk_coinductive_aux₁ e zero comm_zero one comm_one succ n in
729
+ ⟨I.1 ≫ (Q.X_prev_iso rfl).inv, (P.X_next_iso rfl).hom ≫ I.2 .1 , by simpa using I.2 .2 ⟩
730
+
731
+ lemma mk_coinductive_aux₃ (i : ℕ) :
732
+ (mk_coinductive_aux₂ e zero comm_zero one comm_one succ i).2 .1 ≫ (Q.X_prev_iso rfl).inv
733
+ = (P.X_next_iso rfl).hom ≫ (mk_coinductive_aux₂ e zero comm_zero one comm_one succ (i+1 )).1 :=
734
+ by rcases i with (_|_|i); { dsimp, simp, }
735
+
736
+ /--
737
+ A constructor for a `homotopy e 0`, for `e` a chain map between `ℕ`-indexed cochain complexes,
738
+ working by induction.
739
+
740
+ You need to provide the components of the homotopy in degrees 0 and 1,
741
+ show that these satisfy the homotopy condition,
742
+ and then give a construction of each component,
743
+ and the fact that it satisfies the homotopy condition,
744
+ using as an inductive hypothesis the data and homotopy condition for the previous two components.
745
+ -/
746
+ def mk_coinductive : homotopy e 0 :=
747
+ { hom := λ i j, if h : j + 1 = i then
748
+ (P.X_next_iso h).inv ≫ (mk_coinductive_aux₂ e zero comm_zero one comm_one succ j).2 .1
749
+ else
750
+ 0 ,
751
+ zero' := λ i j w, by rwa dif_neg,
752
+ comm := λ i, begin
753
+ dsimp,
754
+ rw [add_zero, add_comm],
755
+ convert (mk_coinductive_aux₂ e zero comm_zero one comm_one succ i).2 .2 using 2 ,
756
+ { rcases i with (_|_|_|i),
757
+ { simp only [mk_coinductive_aux₂, prev_d_zero_cochain_complex, zero_comp] },
758
+ all_goals
759
+ { simp only [prev_d_succ_cochain_complex],
760
+ dsimp,
761
+ simp only [eq_self_iff_true, iso.inv_hom_id_assoc, dite_eq_ite,
762
+ if_true, category.assoc, X_prev_iso_comp_d_to], }, },
763
+ { cases i,
764
+ { dsimp,
765
+ simp only [eq_self_iff_true, d_next_cochain_complex, dif_pos,
766
+ d_from_comp_X_next_iso_assoc, ←comm_zero],
767
+ rw mk_coinductive_aux₂,
768
+ dsimp,
769
+ convert comm_zero.symm,
770
+ simp only [iso.inv_hom_id_assoc], },
771
+ { dsimp,
772
+ simp only [eq_self_iff_true, d_next_cochain_complex, dif_pos, d_from_comp_X_next_iso_assoc],
773
+ rw mk_coinductive_aux₂,
774
+ dsimp only,
775
+ simp only [iso.inv_hom_id_assoc], }, },
776
+ end }
777
+
778
+ end
779
+
780
+ end mk_coinductive
781
+
648
782
end homotopy
649
783
650
784
/--
@@ -771,4 +905,3 @@ def functor.map_homotopy_equiv (F : V ⥤ W) [F.additive] (h : homotopy_equiv C
771
905
end }
772
906
773
907
end category_theory
774
-
0 commit comments