Library HoTT.Categories.ExponentialLaws.Law0
Require Import Category.Core Functor.Core FunctorCategory.Core Functor.Identity.
Require Import InitialTerminalCategory.Core InitialTerminalCategory.Functors InitialTerminalCategory.NaturalTransformations.
Require Import HoTT.Basics HoTT.Types.
Set Universe Polymorphism.
Set Implicit Arguments.
Generalizable All Variables.
Set Asymmetric Patterns.
Local Open Scope functor_scope.
Require Import InitialTerminalCategory.Core InitialTerminalCategory.Functors InitialTerminalCategory.NaturalTransformations.
Require Import HoTT.Basics HoTT.Types.
Set Universe Polymorphism.
Set Implicit Arguments.
Generalizable All Variables.
Set Asymmetric Patterns.
Local Open Scope functor_scope.
Section law0.
Context `{Funext}.
Context `{IsInitialCategory zero}.
Context `{IsTerminalCategory one}.
Local Notation "0" := zero : category_scope.
Local Notation "1" := one : category_scope.
Variable C : PreCategory.
#[export] Instance IsTerminalCategory_functors_from_initial
: IsTerminalCategory (0 → C) := {}.
Context `{Funext}.
Context `{IsInitialCategory zero}.
Context `{IsTerminalCategory one}.
Local Notation "0" := zero : category_scope.
Local Notation "1" := one : category_scope.
Variable C : PreCategory.
#[export] Instance IsTerminalCategory_functors_from_initial
: IsTerminalCategory (0 → C) := {}.
There is only one functor to the terminal category 1. 
We have already proven in InitialTerminalCategory.v that 0 → C is a terminal category, so there is only one functor to it. 
Since the objects and morphisms in terminal categories are contractible, functors to a terminal category are also contractible, by trunc_functor. 
0ˣ ≅ 0 if x ≠ 0
Section law0'.
Context `{Funext}.
Context `{IsInitialCategory zero}.
Context `{IsTerminalCategory one}.
Local Notation "0" := zero : category_scope.
Local Notation "1" := one : category_scope.
Variable C : PreCategory.
Variable c : C.
Local Instance IsInitialCategory_functors_to_initial_from_inhabited
: IsInitialCategory (C → 0)
:= fun P F ⇒ @Functors.to_initial_category_empty C _ _ F P c.
Context `{Funext}.
Context `{IsInitialCategory zero}.
Context `{IsTerminalCategory one}.
Local Notation "0" := zero : category_scope.
Local Notation "1" := one : category_scope.
Variable C : PreCategory.
Variable c : C.
Local Instance IsInitialCategory_functors_to_initial_from_inhabited
: IsInitialCategory (C → 0)
:= fun P F ⇒ @Functors.to_initial_category_empty C _ _ F P c.
There is exactly one functor from an initial category, and we proved above that if C is inhabited, then C → 0 is initial. 
There is exactly one functor from the initial category 0. 
Since objects and morphisms in an initial category are -1-truncated, so are functors to an initial category. 
Section law00.
Context `{Funext}.
Context `{IsInitialCategory zero}.
Context `{IsInitialCategory zero'}.
Context `{IsTerminalCategory one}.
Local Notation "0" := zero : category_scope.
Local Notation "00" := zero' : category_scope.
Local Notation "1" := one : category_scope.
Context `{Funext}.
Context `{IsInitialCategory zero}.
Context `{IsInitialCategory zero'}.
Context `{IsTerminalCategory one}.
Local Notation "0" := zero : category_scope.
Local Notation "00" := zero' : category_scope.
Local Notation "1" := one : category_scope.
This is just a special case of the first law above.