fix(hott): notation spacing and markdown files
This commit is contained in:
parent
cd48114c47
commit
115dedbd1c
26 changed files with 97 additions and 87 deletions
|
@ -47,7 +47,7 @@ namespace binary
|
|||
variable {f : A → A → A}
|
||||
variable H_comm : commutative f
|
||||
variable H_assoc : associative f
|
||||
local infixl `*` := f
|
||||
local infixl * := f
|
||||
theorem left_comm : left_commutative f :=
|
||||
take a b c, calc
|
||||
a*(b*c) = (a*b)*c : H_assoc
|
||||
|
@ -71,7 +71,7 @@ namespace binary
|
|||
variable {A : Type}
|
||||
variable {f : A → A → A}
|
||||
variable H_assoc : associative f
|
||||
local infixl `*` := f
|
||||
local infixl * := f
|
||||
theorem assoc4helper (a b c d) : (a*b)*(c*d) = a*((b*c)*d) :=
|
||||
calc
|
||||
(a*b)*(c*d) = a*(b*(c*d)) : H_assoc
|
||||
|
|
|
@ -33,7 +33,7 @@ namespace category
|
|||
(is_iso_counit : is_iso ε)
|
||||
|
||||
abbreviation inverse := @is_equivalence.G
|
||||
postfix `⁻¹` := inverse
|
||||
postfix ⁻¹ := inverse
|
||||
--a second notation for the inverse, which is not overloaded
|
||||
postfix [parsing-only] `⁻¹F`:std.prec.max_plus := inverse
|
||||
|
||||
|
|
|
@ -11,5 +11,7 @@ Development of Category Theory. The following files are in this folder (sorted s
|
|||
* [nat_trans](nat_trans.hlean) : Natural transformations
|
||||
* [strict](strict.hlean) : Strict categories
|
||||
* [constructions](constructions/constructions.md) (subfolder) : basic constructions on categories and examples of categories
|
||||
* [adjoint](adjoint.hlean) : Adjoint functors and Equivalences (TODO)
|
||||
* [yoneda](yoneda.hlean) : Yoneda Embedding (TODO)
|
||||
* [adjoint](adjoint.hlean) : Adjoint functors and Equivalences (WIP)
|
||||
* [yoneda](yoneda.hlean) : Yoneda Embedding (WIP)
|
||||
* [limits](limits.hlean) : Limits in a category, defined as terminal object in the cone category
|
||||
* [colimits](colimits.hlean) : Colimits in a category, defined as the limit of the opposite functor
|
|
@ -3,7 +3,19 @@ algebra.category.constructions
|
|||
|
||||
Common categories and constructions on categories. The following files are in this folder.
|
||||
|
||||
* [opposite](opposite.hlean) : Opposite category
|
||||
* [product](product.hlean) : Product category
|
||||
* [hset](hset.hlean) : Category of sets
|
||||
* [functor](functor.hlean) : Functor category
|
||||
* [opposite](opposite.hlean) : Opposite category
|
||||
* [hset](hset.hlean) : Category of sets
|
||||
* [sum](sum.hlean) : Sum category
|
||||
* [product](product.hlean) : Product category
|
||||
* [comma](comma.hlean) : Comma category
|
||||
* [cone](cone.hlean) : Cone category
|
||||
|
||||
|
||||
Discrete, indiscrete or finite categories:
|
||||
|
||||
* [finite_cats](finite_cats.hlean) : Some finite categories, which are diagrams of common limits (the diagram for the pullback or the equalizer). Also contains a general construction of categories where you give some generators for the morphisms, with the condition that you cannot compose two of thosex
|
||||
* [discrete](discrete.hlean)
|
||||
* [indiscrete](indiscrete.hlean)
|
||||
* [terminal](terminal.hlean)
|
||||
* [initial](initial.hlean)
|
||||
|
|
|
@ -1,7 +1,7 @@
|
|||
/-
|
||||
Copyright (c) 2014 Jakob von Raumer. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Author: Jakob von Raumer
|
||||
Authors: Jakob von Raumer, Floris van Doorn
|
||||
|
||||
Ported from Coq HoTT
|
||||
-/
|
||||
|
|
|
@ -31,7 +31,7 @@ namespace iso
|
|||
abbreviation inverse [unfold 6] := @is_iso.inverse
|
||||
abbreviation left_inverse [unfold 6] := @is_iso.left_inverse
|
||||
abbreviation right_inverse [unfold 6] := @is_iso.right_inverse
|
||||
postfix `⁻¹` := inverse
|
||||
postfix ⁻¹ := inverse
|
||||
--a second notation for the inverse, which is not overloaded
|
||||
postfix [parsing-only] `⁻¹ʰ`:std.prec.max_plus := inverse -- input using \-1h
|
||||
|
||||
|
|
|
@ -33,7 +33,7 @@ namespace category
|
|||
attribute precategory [multiple-instances]
|
||||
attribute precategory.is_hset_hom [instance]
|
||||
|
||||
infixr `∘` := precategory.comp
|
||||
infixr ∘ := precategory.comp
|
||||
-- input ⟶ using \--> (this is a different arrow than \-> (→))
|
||||
infixl [parsing-only] ` ⟶ `:25 := precategory.hom
|
||||
namespace hom
|
||||
|
@ -86,7 +86,7 @@ namespace category
|
|||
section squares
|
||||
parameters {ob : Type} [C : precategory ob]
|
||||
local infixl ` ⟶ `:25 := @precategory.hom ob C
|
||||
local infixr `∘` := @precategory.comp ob C _ _ _
|
||||
local infixr ∘ := @precategory.comp ob C _ _ _
|
||||
definition compose_squares {xa xb xc ya yb yc : ob}
|
||||
{xg : xb ⟶ xc} {xf : xa ⟶ xb} {yg : yb ⟶ yc} {yf : ya ⟶ yb}
|
||||
{wa : xa ⟶ ya} {wb : xb ⟶ yb} {wc : xc ⟶ yc}
|
||||
|
@ -152,7 +152,6 @@ namespace category
|
|||
notation g ` ∘[`:60 C:0 `] `:0 f:60 :=
|
||||
@comp (Precategory.carrier C) (Precategory.struct C) _ _ _ g f
|
||||
-- TODO: make this left associative
|
||||
-- TODO: change this notation?
|
||||
|
||||
definition Precategory.eta (C : Precategory) : Precategory.mk C C = C :=
|
||||
Precategory.rec (λob c, idp) C
|
||||
|
|
|
@ -26,7 +26,7 @@ section division_ring
|
|||
include s
|
||||
|
||||
definition divide (a b : A) : A := a * b⁻¹
|
||||
infix `/` := divide
|
||||
infix / := divide
|
||||
|
||||
-- only in this file
|
||||
local attribute divide [reducible]
|
||||
|
|
|
@ -34,10 +34,10 @@ structure has_inv [class] (A : Type) :=
|
|||
structure has_neg [class] (A : Type) :=
|
||||
(neg : A → A)
|
||||
|
||||
infixl `*` := has_mul.mul
|
||||
infixl `+` := has_add.add
|
||||
postfix `⁻¹` := has_inv.inv
|
||||
prefix `-` := has_neg.neg
|
||||
infixl * := has_mul.mul
|
||||
infixl + := has_add.add
|
||||
postfix ⁻¹ := has_inv.inv
|
||||
prefix - := has_neg.neg
|
||||
notation 1 := !has_one.one
|
||||
notation 0 := !has_zero.zero
|
||||
|
||||
|
@ -387,7 +387,7 @@ section add_group
|
|||
-- TODO: derive corresponding facts for div in a field
|
||||
definition sub [reducible] (a b : A) : A := a + -b
|
||||
|
||||
infix `-` := sub
|
||||
infix - := sub
|
||||
|
||||
theorem sub_eq_add_neg (a b : A) : a - b = a + -b := rfl
|
||||
|
||||
|
|
|
@ -29,9 +29,9 @@ structure has_le.{l} [class] (A : Type.{l}) : Type.{l+1} :=
|
|||
structure has_lt [class] (A : Type) :=
|
||||
(lt : A → A → Type₀)
|
||||
|
||||
infixl `<=` := has_le.le
|
||||
infixl `≤` := has_le.le
|
||||
infixl `<` := has_lt.lt
|
||||
infixl <= := has_le.le
|
||||
infixl ≤ := has_le.le
|
||||
infixl < := has_lt.lt
|
||||
|
||||
definition has_le.ge [reducible] {A : Type} [s : has_le A] (a b : A) := b ≤ a
|
||||
notation a ≥ b := has_le.ge a b
|
||||
|
|
|
@ -16,12 +16,12 @@ namespace bool
|
|||
definition bor (a b : bool) :=
|
||||
bool.rec_on a (bool.rec_on b ff tt) tt
|
||||
|
||||
notation a || b := bor a b
|
||||
infix || := bor
|
||||
|
||||
definition band (a b : bool) :=
|
||||
bool.rec_on a ff (bool.rec_on b ff tt)
|
||||
|
||||
notation a && b := band a b
|
||||
infix && := band
|
||||
|
||||
definition bnot (a : bool) :=
|
||||
bool.rec_on a tt ff
|
||||
|
|
|
@ -30,7 +30,7 @@ structure equiv (A B : Type) :=
|
|||
|
||||
namespace is_equiv
|
||||
/- Some instances and closure properties of equivalences -/
|
||||
postfix `⁻¹` := inv
|
||||
postfix ⁻¹ := inv
|
||||
/- a second notation for the inverse, which is not overloaded -/
|
||||
postfix [parsing-only] `⁻¹ᶠ`:std.prec.max_plus := inv
|
||||
|
||||
|
@ -358,7 +358,7 @@ namespace equiv
|
|||
|
||||
|
||||
namespace ops
|
||||
postfix `⁻¹` := equiv.symm -- overloaded notation for inverse
|
||||
postfix ⁻¹ := equiv.symm -- overloaded notation for inverse
|
||||
end ops
|
||||
end equiv
|
||||
|
||||
|
|
|
@ -53,15 +53,11 @@ definition curry [reducible] [unfold-full] : (A × B → C) → A → B → C :=
|
|||
definition uncurry [reducible] [unfold 5] : (A → B → C) → (A × B → C) :=
|
||||
λ f p, match p with (a, b) := f a b end
|
||||
|
||||
precedence `∘'`:60
|
||||
precedence `on`:1
|
||||
precedence `$`:1
|
||||
|
||||
|
||||
infixr ∘ := compose
|
||||
infixr ∘' := dcompose
|
||||
infixl on := on_fun
|
||||
infixr $ := app
|
||||
infixr ` ∘ ` := compose
|
||||
infixr ` ∘' `:60 := dcompose
|
||||
infixl ` on `:1 := on_fun
|
||||
infixr ` $ `:1 := app
|
||||
notation f ` -[` op `]- ` g := combine f op g
|
||||
|
||||
end function
|
||||
|
|
|
@ -10,7 +10,7 @@ import init.reserved_notation
|
|||
/- not -/
|
||||
|
||||
definition not [reducible] (a : Type) := a → empty
|
||||
prefix `¬` := not
|
||||
prefix ¬ := not
|
||||
|
||||
definition absurd {a b : Type} (H₁ : a) (H₂ : ¬a) : b :=
|
||||
empty.rec (λ e, b) (H₂ H₁)
|
||||
|
@ -36,7 +36,7 @@ assume Hb : b, absurd (assume Ha : a, Hb) H
|
|||
|
||||
/- eq -/
|
||||
|
||||
notation a = b := eq a b
|
||||
infix = := eq
|
||||
definition rfl {A : Type} {a : A} := eq.refl a
|
||||
|
||||
namespace eq
|
||||
|
@ -52,9 +52,9 @@ namespace eq
|
|||
subst H (refl a)
|
||||
|
||||
namespace ops
|
||||
notation H `⁻¹` := symm H --input with \sy or \-1 or \inv
|
||||
notation H1 ⬝ H2 := trans H1 H2
|
||||
notation H1 ▸ H2 := subst H1 H2
|
||||
postfix ⁻¹ := symm --input with \sy or \-1 or \inv
|
||||
infixl ⬝ := trans
|
||||
infixr ▸ := subst
|
||||
end ops
|
||||
end eq
|
||||
|
||||
|
@ -94,7 +94,7 @@ end lift
|
|||
/- ne -/
|
||||
|
||||
definition ne {A : Type} (a b : A) := ¬(a = b)
|
||||
notation a ≠ b := ne a b
|
||||
infix ≠ := ne
|
||||
|
||||
namespace ne
|
||||
open eq.ops
|
||||
|
@ -132,8 +132,8 @@ end
|
|||
|
||||
definition iff (a b : Type) := prod (a → b) (b → a)
|
||||
|
||||
notation a <-> b := iff a b
|
||||
notation a ↔ b := iff a b
|
||||
infix <-> := iff
|
||||
infix ↔ := iff
|
||||
|
||||
namespace iff
|
||||
variables {a b c : Type}
|
||||
|
|
|
@ -584,11 +584,11 @@ namespace eq
|
|||
whisker_right h idp = h :=
|
||||
eq.rec_on h (eq.rec_on p idp)
|
||||
|
||||
definition whisker_right_idp_left (p : x = y) (q : y = z) :
|
||||
definition whisker_right_idp_left [unfold-full] (p : x = y) (q : y = z) :
|
||||
whisker_right idp q = idp :> (p ⬝ q = p ⬝ q) :=
|
||||
idp
|
||||
|
||||
definition whisker_left_idp_right (p : x = y) (q : y = z) :
|
||||
definition whisker_left_idp_right [unfold-full] (p : x = y) (q : y = z) :
|
||||
whisker_left p idp = idp :> (p ⬝ q = p ⬝ q) :=
|
||||
idp
|
||||
|
||||
|
@ -596,11 +596,11 @@ namespace eq
|
|||
(idp_con p) ⁻¹ ⬝ whisker_left idp h ⬝ idp_con q = h :=
|
||||
eq.rec_on h (eq.rec_on p idp)
|
||||
|
||||
definition con2_idp {p q : x = y} (h : p = q) :
|
||||
definition con2_idp [unfold-full] {p q : x = y} (h : p = q) :
|
||||
h ◾ idp = whisker_right h idp :> (p ⬝ idp = q ⬝ idp) :=
|
||||
idp
|
||||
|
||||
definition idp_con2 {p q : x = y} (h : p = q) :
|
||||
definition idp_con2 [unfold-full] {p q : x = y} (h : p = q) :
|
||||
idp ◾ h = whisker_left idp h :> (idp ⬝ p = idp ⬝ q) :=
|
||||
idp
|
||||
|
||||
|
|
|
@ -12,14 +12,16 @@ Types (not necessarily HoTT-related):
|
|||
* [sum](sum.hlean)
|
||||
* [pi](pi.hlean)
|
||||
* [arrow](arrow.hlean)
|
||||
* [W](W.hlean): W-types (not loaded by default)
|
||||
* [arrow_2](arrow_2.hlean): alternative development of properties of arrows
|
||||
* [W](W.hlean): W-types (not loaded by default)
|
||||
* [lift](lift.hlean)
|
||||
|
||||
HoTT types
|
||||
|
||||
* [eq](eq.hlean): show that functions related to the identity type are equivalences
|
||||
* [pointed](pointed.hlean)
|
||||
* [pointed](pointed.hlean): pointed types, maps, homotopies, and equivalences
|
||||
* [fiber](fiber.hlean)
|
||||
* [equiv](equiv.hlean)
|
||||
* [trunc](trunc.hlean): truncation levels, n-Types, truncation
|
||||
|
||||
* [pullback](pullback.hlean)
|
||||
* [univ](univ.hlean)
|
|
@ -47,7 +47,6 @@ namespace univ
|
|||
assume H : is_hset Type,
|
||||
absurd (is_trunc_is_embedding_closed lift star) not_is_hset_type0
|
||||
|
||||
--set_option pp.notation false
|
||||
definition not_double_negation_elimination0 : ¬Π(A : Type₀), ¬¬A → A :=
|
||||
begin
|
||||
intro f,
|
||||
|
|
Loading…
Reference in a new issue