(* Title: HoTT_Base.thy Author: Joshua Chen Date: 2018 Basic setup of a homotopy type theory object logic with a cumulative Russell-style universe hierarchy. *) theory HoTT_Base imports Pure begin section \Basic setup\ typedecl t \ \Type of object types and terms\ typedecl ord \ \Type of meta-level numerals\ axiomatization O :: ord and Suc :: "ord \ ord" and lt :: "[ord, ord] \ prop" (infix "<" 999) and leq :: "[ord, ord] \ prop" (infix "\" 999) where lt_Suc [intro]: "n < (Suc n)" and lt_trans [intro]: "\m1 < m2; m2 < m3\ \ m1 < m3" and leq_min [intro]: "O \ n" section \Judgment\ judgment hastype :: "[t, t] \ prop" ("(3_:/ _)") section \Universes\ axiomatization U :: "ord \ t" where U_hierarchy: "i < j \ U i: U j" and U_cumulative: "\A: U i; i < j\ \ A: U j" and U_cumulative': "\A: U i; i \ j\ \ A: U j" text \ Using method @{method rule} with @{thm U_cumulative} is unsafe, if applied blindly it will typically lead to non-termination. One should instead use @{method elim}, or instantiate @{thm U_cumulative} suitably. @{thm U_cumulative'} is an alternative rule used by the method \lift\ in @{file HoTT_Methods.thy}. \ section \Type families\ abbreviation (input) constrained :: "[t \ t, t, t] \ prop" ("(1_:/ _ \ _)") where "f: A \ B \ (\x. x : A \ f x: B)" text \ The abbreviation @{abbrev constrained} is used to define type families, which are constrained expressions @{term "P: A \ B"}, where @{term "A::t"} and @{term "B::t"} are small types. \ type_synonym tf = "t \ t" \ \Type of type families\ section \Named theorems\ named_theorems comp named_theorems form text \ Declare named theorems to be used by proof methods defined in @{file HoTT_Methods.thy}. @{attribute comp} declares computation rules, which are used by the \compute\ method, and may also be passed to invocations of the method \subst\ to perform equational rewriting. @{attribute form} declares type formation rules. These are mainly used by the \cumulativity\ method, which lifts types into higher universes. \ (* Todo: Set up the Simplifier! *) end