(* Title: HoTT/HoTT_Base.thy Author: Josh Chen Date: Aug 2018 Basic setup and definitions of a homotopy type theory object logic with a cumulative universe hierarchy à la Russell. *) theory HoTT_Base imports Pure begin section \Foundational definitions\ text "Meta syntactic type for object-logic types and terms." typedecl Term section \Judgments\ text " Formalize the typing judgment \a: A\. For judgmental/definitional equality we use the existing Pure equality \\\ and hence do not need to define a separate judgment for it. " consts has_type :: "[Term, Term] \ prop" ("(1_: _)" [0, 0] 1000) section \Universe hierarchy\ text "Finite meta-ordinals to index the universes." typedecl Ord axiomatization O :: Ord and S :: "Ord \ Ord" ("S_" [0] 1000) and lt :: "[Ord, Ord] \ prop" (infix "<-" 999) where Ord_min: "\n. O <- S(n)" and Ord_monotone: "\m n. m <- n \ S(m) <- S(n)" lemmas Ord_rules [intro] = Ord_min Ord_monotone \ \Enables \standard\ to automatically solve inequalities.\ text "Define the universe types." axiomatization U :: "Ord \ Term" where U_hierarchy: "\i j. i <- j \ U(i): U(j)" and U_cumulative: "\A i j. \A: U(i); i <- j\ \ A: U(j)" \ \WARNING: \rule Universe_cumulative\ can result in an infinite rewrite loop!\ section \Type families\ text " The following abbreviation constrains the output type of a meta lambda expression when given input of certain type. " abbreviation (input) constrained :: "[Term \ Term, Term, Term] \ prop" ("(1_: _ \ _)") where "f: A \ B \ (\x. x : A \ f(x): B)" text " The above is used to define type families, which are constrained meta-lambdas \P: A \ B\ where \A\ and \B\ are small types. " type_synonym Typefam = "Term \ Term" section \Named theorems\ text " Named theorems to be used by proof methods later (see HoTT_Methods.thy). \wellform\ declares necessary wellformedness conditions for type and inhabitation judgments, while \comp\ declares computation rules, which are usually passed to invocations of the method \subst\ to perform equational rewriting. " named_theorems wellform named_theorems comp end