(* Title: HoTT/Prod.thy Author: Josh Chen Date: Aug 2018 Dependent product (function) type. *) theory Prod imports HoTT_Base begin section \Constants and syntax\ axiomatization Prod :: "[Term, Typefam] \ Term" and lambda :: "(Term \ Term) \ Term" (binder "\<^bold>\" 30) and appl :: "[Term, Term] \ Term" (infixl "`" 60) \ \Application binds tighter than abstraction.\ syntax "_PROD" :: "[idt, Term, Term] \ Term" ("(3\_:_./ _)" 30) "_PROD_ASCII" :: "[idt, Term, Term] \ Term" ("(3PROD _:_./ _)" 30) text "The translations below bind the variable \x\ in the expressions \B\ and \b\." translations "\x:A. B" \ "CONST Prod A (\x. B)" "PROD x:A. B" \ "CONST Prod A (\x. B)" text "Nondependent functions are a special case." abbreviation Function :: "[Term, Term] \ Term" (infixr "\" 40) where "A \ B \ \_: A. B" section \Type rules\ axiomatization where Prod_form: "\A: U(i); B: A \ U(i)\ \ \x:A. B(x): U(i)" and Prod_intro: "\A: U(i); \x. x: A \ b(x): B(x)\ \ \<^bold>\x. b(x): \x:A. B(x)" and Prod_elim: "\f: \x:A. B(x); a: A\ \ f`a: B(a)" and Prod_comp: "\\x. x: A \ b(x): B(x); a: A\ \ (\<^bold>\x. b(x))`a \ b(a)" and Prod_uniq: "f : \x:A. B(x) \ \<^bold>\x. (f`x) \ f" and Prod_eq: "\\x. x: A \ b(x) \ b'(x); A: U(i)\ \ \<^bold>\x. b(x) \ \<^bold>\x. b'(x)" text " The Pure rules for \\\ only let us judge strict syntactic equality of object lambda expressions; Prod_eq is the actual definitional equality rule. Note that the syntax \\<^bold>\\ (bold lambda) used for dependent functions clashes with the proof term syntax (cf. \
2.5.2 of the Isabelle/Isar Implementation). " text " In addition to the usual type rules, it is a meta-theorem that whenever \\x:A. B x: U(i)\ is derivable from some set of premises \, then so are \A: U(i)\ and \B: A \ U(i)\. That is to say, the following inference rules are admissible, and it simplifies proofs greatly to axiomatize them directly. " axiomatization where Prod_form_cond1: "(\x:A. B(x): U(i)) \ A: U(i)" and Prod_form_cond2: "(\x:A. B(x): U(i)) \ B: A \ U(i)" text "Set up the standard reasoner to use the type rules:" lemmas Prod_rules [intro] = Prod_form Prod_intro Prod_elim Prod_comp Prod_uniq Prod_eq lemmas Prod_wellform [wellform] = Prod_form_cond1 Prod_form_cond2 lemmas Prod_comps [comp] = Prod_comp Prod_uniq Prod_eq section \Function composition\ definition compose :: "[Term, Term] \ Term" (infixr "o" 70) where "g o f \ \<^bold>\x. g`(f`x)" syntax "_COMPOSE" :: "[Term, Term] \ Term" (infixr "\" 70) translations "g \ f" \ "g o f" section \Unit type\ axiomatization Unit :: Term ("\") and pt :: Term ("\") and indUnit :: "[Term, Term] \ Term" ("(1ind\<^sub>\)") where Unit_form: "\: U(O)" and Unit_intro: "\: \" and Unit_elim: "\C: \ \ U(i); c: C(\); a: \\ \ ind\<^sub>\(c)(a) : C(a)" and Unit_comp: "\C: \ \ U(i); c: C(\)\ \ ind\<^sub>\(c)(\) \ c" lemmas Unit_rules [intro] = Unit_form Unit_intro Unit_elim Unit_comp lemmas Unit_comps [comp] = Unit_comp end