chore(hott): update default files and some markdown files

This commit is contained in:
Floris van Doorn 2016-05-13 15:17:50 -04:00 committed by Leonardo de Moura
parent ac2afb6d82
commit 9f13527c25
12 changed files with 33 additions and 22 deletions

View file

@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
-/
import .functor .set .opposite .product .comma .sum .discrete .indiscrete .terminal .initial
import .functor .set .opposite .product .comma .sum .discrete .indiscrete .terminal .initial .order

View file

@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
-/
import .category .strict .groupoid .constructions
import .category .strict .groupoid .constructions .limits .functor

View file

@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
-/
import .basic
import .exponential_laws

View file

@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
-/
import .limits .colimits
import .set .functor .adjoint .functor_preserve

View file

@ -6,7 +6,4 @@ Authors: Floris van Doorn
The core of the HoTT library
-/
import types
import cubical
import homotopy.circle
import algebra.hott
import types cubical homotopy hit choice

8
hott/hit/default.hlean Normal file
View file

@ -0,0 +1,8 @@
/-
Copyright (c) 2016 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
-/
import .two_quotient .colimit .coeq .refl_quotient

View file

@ -10,14 +10,6 @@ open eq is_trunc is_equiv nat equiv trunc prod pushout sigma sphere_index unit
-- where should this be?
definition family : Type := ΣX, X → Type
-- this should be in init!
namespace nat
definition cases {M : → Type} (mz : M zero) (ms : Πn, M (succ n)) : Πn, M n :=
nat.rec mz (λn dummy, ms n)
end nat
namespace cellcomplex
/-

View file

@ -0,0 +1,8 @@
/-
Copyright (c) 2016 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
-/
import .sphere2 .EM .torus .red_susp .quaternionic_hopf .smash .cellcomplex .interval .cylinder

View file

@ -3,11 +3,12 @@ The Lean Homotopy Type Theory Library
The Lean Homotopy Type Theory library consists of the following directories:
* [init](init/init.md) : constants and theorems needed for low-level system operations
* [init](init/init.md) : basic definitions and theorems. These are imported in each ".hlean" file by default, unless the `prelude` command is used.
* [types](types/types.md) : concrete datatypes and type constructors
* [hit](hit/hit.md): higher inductive types
* [algebra](algebra/algebra.md) : algebraic structures
* [cubical](cubical/cubical.md): cubical types
* [algebra](algebra/algebra.md) : algebraic structures, group theory and category theory.
* [cubical](cubical/cubical.md) : cubical types (e.g. squares in a types, squareovers)
* [hit](hit/hit.md): higher inductive types. Some higher inductive types which are mostly relevant in homotopy theory are in the [homotopy](homotopy/homotopy.md) folder.
* [homotopy](homotopy/homotopy.md) : synthetic homotopy theory.
The following files don't fit in any of the subfolders:
* [prop_trunc](prop_trunc.hlean): in this file we prove that `is_trunc n A` is a mere proposition. We separate this from [types.trunc](types/trunc.hlean) to avoid circularity in imports.

View file

@ -16,6 +16,10 @@ namespace nat
{C : → Type} (n : ) (H₁ : C 0) (H₂ : Π (a : ), C a → C (succ a)) : C n :=
nat.rec H₁ H₂ n
protected definition cases [reducible] [unfold 4] {M : → Type} (mz : M zero)
(ms : Πn, M (succ n)) : Πn, M n :=
nat.rec mz (λn dummy, ms n)
protected definition cases_on [reducible] [recursor] [unfold 2]
{C : → Type} (n : ) (H₁ : C 0) (H₂ : Π (a : ), C (succ a)) : C n :=
nat.rec H₁ (λ a ih, H₂ a) n

View file

@ -6,4 +6,4 @@ Authors: Floris van Doorn
import .bool .unit .prod .sigma .pi .arrow .pointed .fiber
import .nat .int
import .eq .equiv .trunc
import .eq .equiv .trunc .univ .W .type_functor

View file

@ -29,4 +29,5 @@ Types in HoTT:
* [equiv](equiv.hlean)
* [trunc](trunc.hlean): truncation levels, n-types, truncation
* [pullback](pullback.hlean)
* [univ](univ.hlean)
* [univ](univ.hlean)
* [type_functor](type_functor.hlean)