(* 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] \ Term" and appl :: "[Term, Term] \ Term" (infixl "`" 60) \ \Application binds tighter than abstraction.\ syntax "_PROD" :: "[idt, Term, Term] \ Term" ("(3\_:_./ _)" 30) "_LAMBDA" :: "[idt, Term, Term] \ Term" ("(1\<^bold>\_:_./ _)" 30) "_PROD_ASCII" :: "[idt, Term, Term] \ Term" ("(3PROD _:_./ _)" 30) "_LAMBDA_ASCII" :: "[idt, Term, Term] \ Term" ("(3%%_:_./ _)" 30) text "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)" 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: "\i A B. \A : U(i); B: A \ U(i)\ \ \x:A. B x : U(i)" and Prod_intro: "\i A B b. \A : U(i); \x. x : A \ b x : B x\ \ \<^bold>\x:A. b x : \x:A. B x" and Prod_elim: "\A B f a. \f : \x:A. B x; a : A\ \ f`a : B a" and Prod_comp: "\A B b a. \\x. x : A \ b x : B x; a : A\ \ (\<^bold>\x:A. b x)`a \ b a" and Prod_uniq: "\A B f. f : \x:A. B x \ \<^bold>\x:A. (f`x) \ f" 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). " text " In addition to the usual type rules, it is a meta-theorem (*PROVE THIS!*) 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: "\i A B. (\x:A. B x : U(i)) \ A : U(i)" and Prod_form_cond2: "\i A B. (\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 lemmas Prod_form_conds [intro (*elim, wellform*)] = Prod_form_cond1 Prod_form_cond2 lemmas Prod_comps [comp] = Prod_comp Prod_uniq section \Unit type\ axiomatization Unit :: Term ("\") and pt :: Term ("\") and indUnit :: "[Typefam, Term, Term] \ Term" ("(1ind\<^sub>\)") where Unit_form: "\ : U(O)" and Unit_intro: "\ : \" and Unit_elim: "\i C c a. \C: \ \ U(i); c : C \; a : \\ \ ind\<^sub>\ C c a : C a" and Unit_comp: "\i C c. \C: \ \ U(i); c : C \\ \ ind\<^sub>\ C c \ \ c" lemmas Unit_rules [intro] = Unit_form Unit_intro Unit_elim Unit_comp lemmas Unit_comps [comp] = Unit_comp end