(* Title: Prod.thy Author: Joshua Chen Date: 2018 Dependent product type *) theory Prod imports HoTT_Base HoTT_Methods begin section \Basic definitions\ axiomatization Prod :: "[t, tf] \ t" and lambda :: "(t \ t) \ t" (binder "\<^bold>\" 30) and appl :: "[t, t] \ t" ("(1_`/_)" [60, 61] 60) \ \Application binds tighter than abstraction.\ syntax "_prod" :: "[idt, t, t] \ t" ("(3\_:_./ _)" 30) "_prod_ascii" :: "[idt, t, t] \ t" ("(3II _:_./ _)" 30) text \The translations below bind the variable @{term x} in the expressions @{term B} and @{term b}.\ translations "\x:A. B" \ "CONST Prod A (\x. B)" "II x:A. B" \ "CONST Prod A (\x. B)" text \Non-dependent functions are a special case.\ abbreviation Fun :: "[t, t] \ t" (infixr "\" 40) where "A \ B \ \_: A. B" axiomatization where \ \Type rules\ 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_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 \ \Congruence rules\ Prod_form_eq: "\A: U i; B: A \ U i; C: A \ U i; \x. x: A \ B x \ C x\ \ \x:A. B x \ \x:A. C x" and Prod_intro_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. The actual definitional equality rule is @{thm Prod_intro_eq}. Note that this is a separate rule from function extensionality. Note that the bold lambda symbol \\<^bold>\\ used for dependent functions clashes with the proof term syntax (cf. \
2.5.2 of the Isabelle/Isar Implementation). \ lemmas Prod_routine [intro] = Prod_form Prod_intro Prod_elim lemmas Prod_comps [comp] = Prod_comp Prod_uniq section \Additional definitions\ definition compose :: "[t, t] \ t" (infixr "o" 110) where "g o f \ \<^bold>\x. g`(f`x)" syntax "_compose" :: "[t, t] \ t" (infixr "\" 110) translations "g \ f" \ "g o f" declare compose_def [comp] lemma compose_assoc: assumes "A: U i" and "f: A \ B" "g: B \ C" "h: \x:C. D x" shows "(h \ g) \ f \ h \ (g \ f)" by (derive lems: assms Prod_intro_eq) lemma compose_comp: assumes "A: U i" and "\x. x: A \ b x: B" and "\x. x: B \ c x: C x" shows "(\<^bold>\x. c x) \ (\<^bold>\x. b x) \ \<^bold>\x. c (b x)" by (derive lems: assms Prod_intro_eq) abbreviation id :: t where "id \ \<^bold>\x. x" \ \Polymorphic identity function\ end