(* Title: HoTT/Prod.thy Author: Josh Chen 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: "\\x. x: A \ b x: B x; A: U i\ \ \<^bold>\x. b x: \x:A. B x" and Prod_elim: "\f: \x:A. B x; a: A\ \ f`a: B a" and Prod_appl: "\\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 \ c x; A: U i\ \ \<^bold>\x. b x \ \<^bold>\x. c 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_wellform1: "(\x:A. B x: U i) \ A: U i" and Prod_wellform2: "(\x:A. B x: U i) \ B: A \ U i" text "Rule attribute declarations---set up various methods to use the type rules." lemmas Prod_comp [comp] = Prod_appl Prod_uniq lemmas Prod_wellform [wellform] = Prod_wellform1 Prod_wellform2 lemmas Prod_routine [intro] = Prod_form Prod_intro Prod_elim section \Function composition\ definition compose :: "[Term, Term] \ Term" (infixr "o" 110) where "g o f \ \<^bold>\x. g`(f`x)" syntax "_COMPOSE" :: "[Term, Term] \ Term" (infixr "\" 110) translations "g \ f" \ "g o f" section \Polymorphic identity function\ abbreviation id :: Term where "id \ \<^bold>\x. x" end