Skip to content

Commit

Permalink
Merge remote-tracking branch 'origin/master' into coq-8.17
Browse files Browse the repository at this point in the history
  • Loading branch information
jwiegley committed Nov 13, 2023
2 parents 6bb150a + 41ab38d commit 2aaea15
Show file tree
Hide file tree
Showing 11 changed files with 100 additions and 27 deletions.
2 changes: 1 addition & 1 deletion Construction/Comma/Adjunction.v
Original file line number Diff line number Diff line change
Expand Up @@ -56,7 +56,7 @@ Context {G : C ⟶ D}.
Definition πD := comma_proj1.
Definition πC := comma_proj2.

#[local] Notation "⟨πD,πC⟩" := comma_proj (at level 100).
#[local] Notation "⟨πD,πC⟩" := comma_proj (at level 0).

Record lawvere_equiv := {
lawvere_iso : F ↓ Id[C] ≅[Cat] Id[D] ↓ G;
Expand Down
2 changes: 1 addition & 1 deletion Functor/Applicative.v
Original file line number Diff line number Diff line change
Expand Up @@ -36,4 +36,4 @@ Arguments Applicative {_ _} F.
Arguments pure {C _ F _ _ _}.

Notation "pure[ F ]" := (@pure _ _ F _ _ _)
(at level 9, format "pure[ F ]") : morphism_scope.
(at level 0, format "pure[ F ]") : morphism_scope.
6 changes: 3 additions & 3 deletions Functor/Diagonal.v
Original file line number Diff line number Diff line change
Expand Up @@ -26,12 +26,12 @@ Program Instance Diagonal_Product `(C : Category) : C ⟶ C ∏ C.

End Diagonal_Product.

Notation "Δ[ J ]( c )" := (Diagonal J c) (at level 90, format "Δ[ J ]( c )") : functor_scope.
Notation "Δ[ J ]( c )" := (Diagonal J c) (at level 0, format "Δ[ J ]( c )") : functor_scope.

Require Import Category.Instance.One.

Notation "Δ( c )" := (Diagonal _ c) (at level 90, format "Δ( c )") : functor_scope.
Notation "=( c )" := (Diagonal 1 c) (at level 90, format "=( c )") : functor_scope.
Notation "Δ( c )" := (Diagonal _ c) (at level 0, format "Δ( c )") : functor_scope.
Notation "=( c )" := (Diagonal 1 c) (at level 0, format "=( c )") : functor_scope.

Definition Δ {C J : Category} := @Diagonal C J.

Expand Down
2 changes: 1 addition & 1 deletion Functor/Product/Internal.v
Original file line number Diff line number Diff line change
Expand Up @@ -25,4 +25,4 @@ Next Obligation.
Qed.

Notation "×( C )" := (@InternalProductFunctor C _)
(at level 90, format "×( C )") : functor_scope.
(at level 0, format "×( C )") : functor_scope.
2 changes: 1 addition & 1 deletion Functor/Structure/Monoidal/Pure.v
Original file line number Diff line number Diff line change
Expand Up @@ -53,4 +53,4 @@ End Pure.
Arguments pure {C _ F _ _ A}.

Notation "pure[ F ]" := (@pure _ _ F _ _ _)
(at level 9, format "pure[ F ]") : morphism_scope.
(at level 0, format "pure[ F ]") : morphism_scope.
4 changes: 2 additions & 2 deletions Instance/Cones/Limit.v
Original file line number Diff line number Diff line change
Expand Up @@ -11,7 +11,7 @@ Program Definition Limit_Cones `(F : J ⟶ C) `{T : @Terminal (Cones F)} :
Limit F := {|
limit_cone := @terminal_obj _ T;
ump_limits := fun N =>
{| unique_obj := `1 @one _ T N
; unique_property := `2 @one _ T N
{| unique_obj := `1 (@one _ T N)
; unique_property := `2 (@one _ T N)
; uniqueness := fun v H => @one_unique _ T N (@one _ T N) (v; H) |}
|}.
8 changes: 4 additions & 4 deletions Instance/Coq/Applicative.v
Original file line number Diff line number Diff line change
Expand Up @@ -35,11 +35,11 @@ Arguments ap {F _ _ _} _ x.

Coercion Applicative_Functor `{@Applicative F} : Coq ⟶ Coq := F.

Notation "pure[ F ]" := (@pure F _ _) (at level 19, F at next level).
Notation "pure[ F G ]" := (@pure (fun x => F (G x)) _ _) (at level 9).
Notation "pure[ F ]" := (@pure F _ _) (at level 0, F at next level).
Notation "pure[ F G ]" := (@pure (fun x => F (G x)) _ _) (at level 0).

Notation "ap[ F ]" := (@ap F _ _ _) (at level 9).
Notation "ap[ F G ]" := (@ap (fun x => F (G x)) _ _ _) (at level 9).
Notation "ap[ F ]" := (@ap F _ _ _) (at level 0).
Notation "ap[ F G ]" := (@ap (fun x => F (G x)) _ _ _) (at level 0).

Infix "<*>" := ap (at level 29, left associativity).
Notation "x <**> F" := (ap F x) (at level 29, left associativity).
Expand Down
10 changes: 5 additions & 5 deletions Theory/Coq/Applicative.v
Original file line number Diff line number Diff line change
Expand Up @@ -28,16 +28,16 @@ Module ApplicativeNotations.
Export FunctorNotations.

Notation "pure[ M ]" := (@pure M _ _)
(at level 9, format "pure[ M ]").
(at level 0, format "pure[ M ]").
Notation "pure[ M N ]" := (@pure (fun X => M (N X)) _ _)
(at level 9, format "pure[ M N ]").
(at level 0, format "pure[ M N ]").

Notation "ap[ M ]" := (@ap M _ _ _)
(at level 9, format "ap[ M ]").
(at level 0, format "ap[ M ]").
Notation "ap[ M N ]" := (@ap (fun X => M (N X)) _ _ _)
(at level 9, format "ap[ M N ]").
(at level 0, format "ap[ M N ]").
Notation "ap[ M N O ]" := (@ap (fun X => M (N (O X))) _ _ _)
(at level 9, format "ap[ M N O ]").
(at level 0, format "ap[ M N O ]").

Infix "<*>" := ap (at level 29, left associativity).
Infix "<*[ M ]>" := (@ap M _ _ _)
Expand Down
12 changes: 6 additions & 6 deletions Theory/Coq/Functor.v
Original file line number Diff line number Diff line change
Expand Up @@ -32,22 +32,22 @@ Notation "x <&> f" := (fmap f x)
(at level 29, left associativity, only parsing).

Notation "fmap[ M ]" := (@fmap M _ _ _)
(at level 9, format "fmap[ M ]").
(at level 0, format "fmap[ M ]").
Notation "fmap[ M N ]" := (@fmap (λ X, M (N X)) _ _ _)
(at level 9, format "fmap[ M N ]").
(at level 0, format "fmap[ M N ]").
Notation "fmap[ M N O ]" := (@fmap (λ X, M (N (O X))) _ _ _)
(at level 9, format "fmap[ M N O ]").
(at level 0, format "fmap[ M N O ]").

Infix ">$<" := contramap (at level 29, left associativity, only parsing).
Notation "x >&< f" :=
(contramap f x) (at level 29, left associativity, only parsing).

Notation "contramap[ M ]" := (@contramap M _ _ _)
(at level 9, format "contramap[ M ]").
(at level 0, format "contramap[ M ]").
Notation "contramap[ M N ]" := (@contramap (λ X, M (N X)) _ _ _)
(at level 9, format "contramap[ M N ]").
(at level 0, format "contramap[ M N ]").
Notation "contramap[ M N O ]" := (@contramap (λ X, M (N (O X))) _ _ _)
(at level 9, format "contramap[ M N O ]").
(at level 0, format "contramap[ M N O ]").

End FunctorNotations.

Expand Down
6 changes: 3 additions & 3 deletions Theory/Functor.v
Original file line number Diff line number Diff line change
Expand Up @@ -57,9 +57,9 @@ Notation "x <&> f" := (fmap f x)
(at level 29, left associativity, only parsing) : morphism_scope.

Notation "fobj[ F ]" := (@fobj _ _ F%functor)
(at level 9, format "fobj[ F ]") : object_scope.
(at level 0, format "fobj[ F ]") : object_scope.
Notation "fmap[ F ]" := (@fmap _ _ F%functor _ _)
(at level 9, format "fmap[ F ]") : morphism_scope.
(at level 0, format "fmap[ F ]") : morphism_scope.

#[export] Hint Rewrite @fmap_id : categories.

Expand Down Expand Up @@ -168,7 +168,7 @@ Program Definition Id {C : Category} : C ⟶ C := {|

Arguments Id {C} /.

Notation "Id[ C ]" := (@Id C) (at level 9, format "Id[ C ]") : functor_scope.
Notation "Id[ C ]" := (@Id C) (at level 0, format "Id[ C ]") : functor_scope.

Program Definition Compose {C D E : Category}
(F : D ⟶ E) (G : C ⟶ D) : C ⟶ E := {|
Expand Down
73 changes: 73 additions & 0 deletions Theory/Sheaf.v
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,8 @@ Require Import Category.Theory.Category.
Require Import Category.Theory.Functor.
Require Import Category.Construction.Opposite.
Require Import Category.Instance.Fun.
Require Import Category.Instance.Sets.
Require Import Coq.Vectors.Vector.

Generalizable All Variables.

Expand All @@ -16,3 +18,74 @@ Definition Presheaves {U C : Category} := [U^op, C].
Definition Copresheaf (U C : Category) := U ⟶ C.

Definition Copresheaves {U C : Category} := [U, C].

(* Custom data type to express Forall propositions on vectors over Type. *)
Inductive ForallT {A : Type} (P : A → Type) :
∀ {n : nat}, t A n → Type :=
| ForallT_nil : ForallT (nil A)
| ForallT_cons (n : nat) (x : A) (v : t A n) :
P x → ForallT v → ForallT (cons A x n v).

Inductive ExistsT {A : Type} (P : A → Type) :
∀ {n : nat}, t A n → Type :=
| ExistsT_cons_hd (m : nat) (x : A) (v : t A m) :
P x → ExistsT (cons A x m v)
| ExistsT_cons_tl (m : nat) (x : A) (v : t A m) :
ExistsT v → ExistsT (cons A x m v).

(* A coverage on a category C consists of a function assigning to each object
U ∈ C a collection of families of morphisms { fᵢ : Uᵢ → U } (i∈I), called
covering families, such that
if { fᵢ : Uᵢ → U } (i∈I) is a covering family
and g : V → U is a morphism,
then there exists a covering family { hⱼ : Vⱼ → V }
such that each composite g ∘ hⱼ factors through some fᵢ. *)

Class Site (C : Category) := {
covering_family (u : C) := sigT (Vector.t (∃ uᵢ, uᵢ ~> u));

coverage (u : C) : covering_family u;

coverage_condition
(u : C) (fs : covering_family u)
(v : C) (g : v ~> u) :
∃ hs : covering_family v,
ForallT (A := ∃ vⱼ, vⱼ ~> v) (λ '(vⱼ; hⱼ),
ExistsT (A := ∃ uᵢ, uᵢ ~> u) (λ '(uᵢ; fᵢ),
∃ k : vⱼ ~> uᵢ, fᵢ ∘ k ≈ g ∘ hⱼ)
(`2 fs))
(`2 hs)
}.

(* If { fᵢ : Uᵢ → U } (i∈I) is a family of morphisms with codomain U,
a presheaf X : Cᵒᵖ → Set is called a sheaf for this family if:
for any collection of elements xᵢ ∈ X(Uᵢ)
such that,
whenever g : V → Uᵢ and h : V → Uⱼ
are such that fᵢ ∘ g = fⱼ ∘ h,
we have X(g)(xᵢ) = X(h)(xⱼ),
then there exists a unique x ∈ X(U)
such that X(fᵢ)(x) = xᵢ for all i . *)

Class Sheaf `{@Site C} (X : Presheaf C Sets) := {
restriction :
∀ u : C,
let '(i; f) := coverage u in
ForallT
(A := ∃ uᵢ, uᵢ ~> u)
(λ '(uᵢ; fᵢ),
∀ xᵢ : X uᵢ,
(∀ (v : C) (g : v ~> uᵢ),
ForallT
(A := ∃ uⱼ, uⱼ ~> u)
(λ '(uⱼ; fⱼ),
∀ h : v ~> uⱼ,
fᵢ ∘ g ≈ fⱼ ∘ h →
∀ xⱼ : X uⱼ,
fmap[X] g xᵢ = fmap[X] h xⱼ)
f) →
∃! x : X u, fmap[X] fᵢ x = xᵢ)
f
}.

0 comments on commit 2aaea15

Please sign in to comment.