(******** Isabelle/HoTT: Foundational definitions and notation Jan 2019 This file is the starting point, and contains the base setup and definitions of the object logic. Among other things, it defines: * Meta-numerals to index the universe hierarchy, and related axioms. * The universe hierarchy and its governing rules. * Named theorems for later use by proof methods. ********) theory HoTT_Base imports Pure begin section \Basic setup\ declare[[eta_contract=false]] (* I'm not actually sure this does anything... *) typedecl t \ \Type of object-logic terms (which includes the types)\ typedecl ord \ \Type of meta-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: "n < (Suc n)" and lt_trans: "\m1 < m2; m2 < m3\ \ m1 < m3" and leq_min: "O \ n" declare lt_Suc [intro!] leq_min [intro!] lt_trans [intro] section \Typing 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 very easily lead to non-termination. Instead use @{method elim}, or instantiate @{thm U_cumulative} suitably. @{thm U_cumulative'} is an alternative rule used by the method \cumulativity\ in @{file HoTT_Methods.thy}. \ section \Named theorems\ named_theorems comp named_theorems form text \ The named theorems above will 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