Require Import Coq.Logic.FunctionalExtensionality.
Require Import Program.Tactics.

Tactic Notation "nointr" tactic(t) :=
  let m := fresh "marker" in
  pose (m := tt);
  t; revert_until m; clear m.

Ltac fext := nointr repeat (
  match goal with
    [ |- ?x = ?y ]
    (refine (@functional_extensionality_dep _ _ _ _ _) ||
     refine (@forall_extensionality _ _ _ _) ||
     refine (@forall_extensionalityP _ _ _ _) ||
     refine (@forall_extensionalityS _ _ _ _)); intro
  end).

Axiom pext : P Q : Prop, (P Q) (P = Q).

Lemma pi {P : Prop} (p q : P) : p = q.
Proof.
  assert (P = True) by (apply pext; tauto). subst. now destruct p,q.
Qed.



Definition funcomp {X Y Z} (g : Y Z) (f : X Y) :=
   x g (f x).

Require Export Vector.

Definition vec_ext {A B n} {f g : A B} :
  ( x, f x = g x) xs : Vector.t A n, Vector.map f xs = Vector.map g xs.
Proof.
  intros H. induction xs. reflexivity.
  cbn. f_equal. apply H. apply IHxs.
Defined.


Definition vec_id {A n} { f : A A} :
  ( x, f x = x) xs : Vector.t A n, Vector.map f xs = xs.
Proof.
  intros H. induction xs. reflexivity.
  cbn. rewrite H. rewrite IHxs; eauto.
Defined.


Definition vec_comp {A B C n} {f: A B} {g: B C} {h} :
  ( x, (funcomp g f) x = h x) xs : Vector.t A n, map g (map f xs) = map h xs.
Proof.
  induction xs. reflexivity.
  cbn. rewrite H. f_equal. apply IHxs.
Defined.


Require Export List.

Definition list_ext {A B} {f g : A B} :
  ( x, f x = g x) xs, map f xs = map g xs.
Proof.
  intros H. induction xs. reflexivity.
  cbn. f_equal. apply H. apply IHxs.
Defined.


Definition list_id {A} { f : A A} :
  ( x, f x = x) xs, List.map f xs = xs.
Proof.
  intros H. induction xs. reflexivity.
  cbn. rewrite H. rewrite IHxs; eauto.
Defined.


Definition list_comp {A B C} {f: A B} {g: B C} {h} :
  ( x, (funcomp g f) x = h x) xs, map g (map f xs) = map h xs.
Proof.
  induction xs. reflexivity.
  cbn. rewrite H. f_equal. apply IHxs.
Defined.



Definition prod_map {A B C D} (f : A C) (g : B D) (p : A * B) :
  C * D.
Proof.
  destruct p. split. auto. auto.
Defined.


Definition prod_id {A B} {f : A A} {g : B B} :
  ( x, f x = x) ( x, g x = x) p, prod_map f g p = p.
Proof.
  intros. destruct p. cbn. f_equal; auto.
Defined.


Definition prod_ext {A B C D} {f f' : A C} {g g': B D} :
  ( x, f x = f' x) ( x, g x = g' x) p, prod_map f g p = prod_map f' g' p.
Proof.
  intros. destruct p. cbn. f_equal; auto.
Defined.


Definition prod_comp {A B C D E F} {f1 : A C} {g1 : C E} { h1} {f2: B D} {g2: D F} {h2}:
  ( x, (funcomp ) x = x) ( x, (funcomp ) x = x) p, prod_map (prod_map p) = prod_map p.
Proof.
  intros. destruct p. cbn. f_equal; auto.
  now rewrite H. now rewrite .
Defined.



Definition cod X A: Type := X A.

Definition cod_map {X} {A B} (f : A B) (p : X A) :
  X B.
Proof. eauto. Defined.

Definition cod_id {X} {A} {f : A A} :
  ( x, f x = x) (p: X A), cod_map f p = p.
Proof. intros H p. unfold cod_map. fext. congruence. Defined.

Definition cod_ext {X} {A B} {f f' : A B} :
  ( x, f x = f' x) (p: X A), cod_map f p = cod_map f' p.
Proof. intros H p. unfold cod_map. fext. congruence. Defined.

Definition cod_ext' {X} {A B} {f f' : A B} (p : X A) :
  ( x, f (p x) = f' (p x)) cod_map f p = cod_map f' p.
Proof. intros H. unfold cod_map. fext. congruence. Defined.

Definition cod_comp {X} {A B C} {f : A B} {g : B C} {h} :
  ( x, (funcomp g f) x = h x) (p: X _), cod_map g (cod_map f p) = cod_map h p.
Proof. intros H p. unfold cod_map. fext. intros x. now rewrite H. Defined.

Hint Rewrite in_map_iff : FunctorInstances.