(******** Isabelle/HoTT: Basic logical definitions and notation Feb 2019 This file completes the basic logical and functional setup of the HoTT logic. It defines: * The universe hierarchy and its governing rules. * Some notational abbreviations. * Named theorems for later use by proof methods. ********) theory HoTT_Base imports Pure begin section \Basic setup\ declare[[eta_contract=false]] \ \Do not eta-contract\ typedecl t \ \Type of object-logic terms (which includes the types)\ judgment has_type :: "[t, t] \ prop" ("(2_ :/ _)") section \Universes\ 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] 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} and @{thm U_cumulative'} is unsafe: if applied blindly it will very easily lead to non-termination. Instead use @{method elim}, or instantiate the rules suitably. @{thm U_cumulative'} is an alternative rule used by the method @{theory_text cumulativity} in @{file HoTT_Methods.thy}. \ section \Notation\ abbreviation (input) constraint :: "[t \ t, t, t] \ prop" ("(1_:/ _ \ _)") where "f: A \ B \ (\x. x: A \ f x: B)" text \We use the notation @{prop "B: A \ U i"} to abbreviate type families.\ abbreviation (input) K_combinator :: "'a \ 'b \ 'a" ("^_" [1000]) where "^A \ \_. A" section \Named theorems\ named_theorems form named_theorems comp named_theorems cong text \ The named theorems above will be used by proof methods defined in @{file HoTT_Methods.thy}. @{attribute form} declares type formation rules. These are mainly used by the \cumulativity\ method, which lifts types into higher universes. @{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 cong} declares congruence rules, the definitional equality rules of the type theory. \ (* Todo: Set up the Simplifier! *) end