(* Title: HoTT/Prod.thy Author: Josh Chen Date: Jun 2018 Dependent product (function) type for the HoTT logic. *) theory Prod imports HoTT_Base begin axiomatization Prod :: "[Term, Typefam] \ Term" and lambda :: "[Term, Term \ Term] \ Term" and \ \Application binds tighter than abstraction.\ appl :: "[Term, Term] \ Term" (infixl "`" 60) syntax "_PROD" :: "[idt, Term, Term] \ Term" ("(3\_:_./ _)" 30) "_LAMBDA" :: "[idt, Term, Term] \ Term" ("(3\<^bold>\_:_./ _)" 30) "_PROD_ASCII" :: "[idt, Term, Term] \ Term" ("(3PROD _:_./ _)" 30) "_LAMBDA_ASCII" :: "[idt, Term, Term] \ Term" ("(3%%_:_./ _)" 30) \ \The translations below bind the variable \x\ in the expressions \B\ and \b\.\ translations "\x:A. B" \ "CONST Prod A (\x. B)" "\<^bold>\x:A. b" \ "CONST lambda A (\x. b)" "PROD x:A. B" \ "CONST Prod A (\x. B)" "%%x:A. b" \ "CONST lambda A (\x. b)" \ \Type rules\ axiomatization where Prod_form [intro]: "\A B. \A : U; B : A \ U\ \ \x:A. B x : U" and Prod_intro [intro]: "\A B b. \A : U; \x. x : A \ b x : B x\ \ \<^bold>\x:A. b x : \x:A. B x" and Prod_elim [elim]: "\A B f a. \f : \x:A. B x; a : A\ \ f`a : B a" and Prod_comp [simp]: "\A B b a. \\x. x : A \ b x : B x; a : A\ \ (\<^bold>\x:A. b x)`a \ b a" and Prod_uniq [simp]: "\A B f. f : \x:A. B x \ \<^bold>\x:A. (f`x) \ f" \ \The funny thing about the first premises of the computation and uniqueness rules is that they introduce a variable B that doesn't actually explicitly appear in the statement of the conclusion. In a sense, they say something like "if this condition holds for some type family B... (then we can apply the rule)". This forces the theorem prover to search for a suitable B. Is this additional overhead necessary? It *is* a safety check for well-formedness...\ text "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)." \ \Nondependent functions are a special case.\ abbreviation Function :: "[Term, Term] \ Term" (infixr "\" 40) where "A \ B \ \_:A. B" end