(* Title: HoTT/Proj.thy Author: Josh Chen Projection functions \fst\ and \snd\ for the dependent sum type. *) theory Proj imports HoTT_Methods Prod Sum begin definition fst :: "Term \ Term" where "fst p \ ind\<^sub>\ (\x y. x) p" definition snd :: "Term \ Term" where "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 by (derive lems: assms) 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 compute show "a: A" and "b: B a" by fact+ qed (routine lems: 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" by (derive lems: assms fst_type) fix x y assume asm: "x: A" "y: B x" show "y: B (fst )" proof (subst fst_comp) show "A: U i" by (wellformed lems: assms(1)) show "\x. x: A \ B x: U i" by (wellformed lems: assms(1)) qed fact+ qed fact 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 compute show "\x y. y: B x \ y: B x" . show "a: A" by fact show "b: B a" by fact qed (routine lems: assms) text "Rule attribute declarations:" lemmas Proj_type [intro] = fst_type snd_type lemmas Proj_comp [comp] = fst_comp snd_comp end