(* Title: HoTT/Proj.thy Author: Josh Chen Date: Jun 2018 Projection functions \fst\ and \snd\ for the dependent sum type. *) theory Proj imports HoTT_Methods Prod Sum begin axiomatization fst :: "Term \ Term" where fst_def: "fst \ \p. ind\<^sub>\(\x y. x)(p)" axiomatization snd :: "Term \ Term" where snd_def: "snd \ \p. ind\<^sub>\(\x y. y)(p)" text "Typing judgments and computation rules for the dependent and non-dependent projection functions." lemma fst_type: assumes "\x:A. B(x): U(i)" and "p: \x:A. B(x)" shows "fst(p): A" unfolding fst_def proof show "A: U(i)" using assms(1) by (rule Sum_wellform) qed (fact assms | assumption)+ lemma fst_comp: assumes "A: U(i)" and "B: A \ U(i)" and "a: A" and "b: B(a)" shows "fst() \ a" unfolding fst_def proof show "\x. x: A \ x: A" . qed (rule assms)+ lemma snd_type: assumes "\x:A. B(x): U(i)" and "p: \x:A. B(x)" shows "snd(p): B(fst p)" unfolding snd_def proof show "\p. p: \x:A. B(x) \ B(fst p): U(i)" proof - have "\p. p: \x:A. B(x) \ fst(p): A" using assms(1) by (rule fst_type) with assms(1) show "\p. p: \x:A. B(x) \ B(fst p): U(i)" by (rule Sum_wellform) qed fix x y assume asm: "x: A" "y: B(x)" show "y: B(fst )" proof (subst fst_comp) show "A: U(i)" using assms(1) by (rule Sum_wellform) show "\x. x: A \ B(x): U(i)" using assms(1) by (rule Sum_wellform) qed (rule asm)+ qed (fact assms) lemma snd_comp: assumes "A: U(i)" and "B: A \ U(i)" and "a: A" and "b: B(a)" shows "snd() \ b" unfolding snd_def proof show "\x y. y: B(x) \ y: B(x)" . show "a: A" by (fact assms) show "b: B(a)" by (fact assms) show *: "B(a): U(i)" using assms(3) by (rule assms(2)) show "B(a): U(i)" by (fact *) qed lemmas Proj_comps [intro] = fst_comp snd_comp end