(******** Isabelle/HoTT: Foundational definitions and notation Feb 2019 This file is the starting point of the Isabelle/HoTT object logic, and contains the basic setup and definitions. Among other things, it: * Calls the setup of object-level type inference functionality. * Defines the universe hierarchy and its governing rules. * Defines named theorems for later use by proof methods. ********) theory HoTT_Base imports Pure keywords "assume_type" "assume_types" :: thy_decl begin section \Terms and typing\ typedecl t \ \Type of object-logic terms (which includes the types)\ judgment has_type :: "[t, t] \ prop" ("(3_:/ _)") section \Setup non-core logic functionality\ declare[[eta_contract=false]] \ \Do not eta-contract\ ML \val trace = Attrib.setup_config_bool @{binding "trace"} (K false)\ \ \Context attribute for tracing; use declare[[trace=true]] at any point in a theory to turn on.\ text \The following file sets up type assumption and inference functionality.\ ML_file "typing.ML" ML \ val _ = let fun store_typing ctxt = Typing.put_declared_typing o (Syntax.read_prop ctxt) in Outer_Syntax.local_theory @{command_keyword "assume_type"} "Declare typing assumptions" (Parse.prop >> (fn prop => fn lthy => Local_Theory.background_theory (store_typing lthy prop) lthy)) end val _ = let val parser = Parse.and_list (Parse.prop) fun store_typings ctxt = fold (Typing.put_declared_typing o (Syntax.read_prop ctxt)) in Outer_Syntax.local_theory @{command_keyword "assume_types"} "Declare typing assumptions" (parser >> (fn props => fn lthy => Local_Theory.background_theory (store_typings lthy props) lthy)) end \ 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 \Type families\ 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.\ 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