(* 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 consts fst :: "[Term, 'a] \ Term" ("(1fst[/_,/ _])") snd :: "[Term, 'a] \ Term" ("(1snd[/_,/ _])") section \Overloaded syntax for dependent and nondependent pairs\ overloading fst_dep \ fst fst_nondep \ fst begin definition fst_dep :: "[Term, Typefam] \ Term" where "fst_dep A B \ \<^bold>\p: (\x:A. B x). indSum[A,B] (\_. A) (\x y. x) p" definition fst_nondep :: "[Term, Term] \ Term" where "fst_nondep A B \ \<^bold>\p: A \ B. indSum[A, \_. B] (\_. A) (\x y. x) p" end overloading snd_dep \ snd snd_nondep \ snd begin definition snd_dep :: "[Term, Typefam] \ Term" where "snd_dep A B \ \<^bold>\p: (\x:A. B x). indSum[A,B] (\q. B (fst[A,B]`q)) (\x y. y) p" definition snd_nondep :: "[Term, Term] \ Term" where "snd_nondep A B \ \<^bold>\p: A \ B. indSum[A, \_. B] (\_. B) (\x y. y) p" end section \Properties\ text "Typing judgments and computation rules for the dependent and non-dependent projection functions." lemma fst_dep_type: assumes "p : \x:A. B x" shows "fst[A,B]`p : A" by (derive lems: fst_dep_def assms) lemma fst_dep_comp: assumes "B: A \ U" and "a : A" and "b : B a" shows "fst[A,B]`(a,b) \ a" proof - have "fst[A,B]`(a,b) \ indSum[A,B] (\_. A) (\x y. x) (a,b)" by (derive lems: fst_dep_def assms) also have "indSum[A,B] (\_. A) (\x y. x) (a,b) \ a" by (derive lems: assms) finally show "fst[A,B]`(a,b) \ a" . qed text "In proving results about the second dependent projection function we often use the following two lemmas." lemma lemma1: assumes "p : \x:A. B x" shows "B (fst[A,B]`p) : U" by (derive lems: assms fst_dep_def) lemma lemma2: assumes "B: A \ U" and "x : A" and "y : B x" shows "y : B (fst[A,B]`(x,y))" proof - have "fst[A,B]`(x,y) \ x" using assms by (rule fst_dep_comp) then show "y : B (fst[A,B]`(x,y))" using assms by simp qed lemma snd_dep_type: assumes "p : \x:A. B x" shows "snd[A,B]`p : B (fst[A,B]`p)" proof (derive lems: assms snd_dep_def) show "snd[A, B] : \p:(\x:A. B x). B (fst[A,B]`p)" proof (unfold snd_dep_def, standard) show "\p. p : \x:A. B x \ indSum[A,B] (\q. B (fst[A,B]`q)) (\x y. y) p : B (fst[A,B]`p)" proof (standard, elim lemma1) fix p assume *: "p : \x:A. B x" have **: "B: A \ U" by (wellformed jdgmt: *) fix x y assume "x : A" and "y : B x" with ** show "y : B (fst[A,B]`(x,y))" by (rule lemma2) qed qed (wellformed jdgmt: assms) qed (rule assms) lemma snd_dep_comp: assumes "B: A \ U" and "a : A" and "b : B a" shows "snd[A,B]`(a,b) \ b" proof - have "snd[A,B]`(a,b) \ indSum[A, B] (\q. B (fst[A,B]`q)) (\x y. y) (a,b)" proof (unfold snd_dep_def, standard) show "(a,b) : \x:A. B x" using assms .. fix p assume *: "p : \x:A. B x" show "indSum[A,B] (\q. B (fst[A,B]`q)) (\x y. y) p : B (fst[A,B]`p)" proof (standard, elim lemma1) fix x y assume "x : A" and "y : B x" with assms(1) show "y : B (fst[A,B]`(x,y))" by (rule lemma2) qed (rule *) qed also have "indSum[A, B] (\q. B (fst[A,B]`q)) (\x y. y) (a,b) \ b" proof (standard, elim lemma1) fix x y assume "x : A" and "y : B x" with assms(1) show "y : B (fst[A,B]`(x,y))" by (rule lemma2) qed (rule assms)+ finally show "snd[A,B]`(a,b) \ b" . qed text "For non-dependent projection functions:" lemma fst_nondep_type: assumes "p : A \ B" shows "fst[A,B]`p : A" proof show "fst[A,B] : A \ B \ A" proof (unfold fst_nondep_def, standard) fix q assume *: "q : A \ B" show "indSum[A, \_. B] (\_. A) (\x y. x) q : A" proof show "A : U" by (wellformed jdgmt: assms) qed (assumption, rule *) qed (wellformed jdgmt: assms) qed (rule assms) lemma fst_nondep_comp: assumes "a : A" and "b : B" shows "fst[A,B]`(a,b) \ a" proof - have "fst[A,B]`(a,b) \ indSum[A, \_. B] (\_. A) (\x y. x) (a,b)" proof (unfold fst_nondep_def, standard) show "(a,b) : A \ B" using assms .. show "\p. p : A \ B \ indSum[A, \_. B] (\_. A) (\x y. x) p : A" proof show "A : U" by (wellformed jdgmt: assms(1)) qed qed also have "indSum[A, \_. B] (\_. A) (\x y. x) (a,b) \ a" proof show "A : U" by (wellformed jdgmt: assms(1)) qed (assumption, (rule assms)+) finally show "fst[A,B]`(a,b) \ a" . qed lemma snd_nondep_type: assumes "p : A \ B" shows "snd[A,B]`p : B" proof show "snd[A,B] : A \ B \ B" proof (unfold snd_nondep_def, standard) fix q assume *: "q : A \ B" show "indSum[A, \_. B] (\_. B) (\x y. y) q : B" proof have **: "\_. B: A \ U" by (wellformed jdgmt: assms) have "fst[A,B]`p : A" using assms by (rule fst_nondep_type) then show "B : U" by (rule **) qed (assumption, rule *) qed (wellformed jdgmt: assms) qed (rule assms) lemma snd_nondep_comp: assumes "a : A" and "b : B" shows "snd[A,B]`(a,b) \ b" proof - have "snd[A,B]`(a,b) \ indSum[A, \_. B] (\_. B) (\x y. y) (a,b)" proof (unfold snd_nondep_def, standard) show "(a,b) : A \ B" by (simple conds: assms) show "\p. p : A \ B \ indSum[A, \_. B] (\_. B) (\x y. y) p : B" proof show "B : U" by (wellformed jdgmt: assms(2)) qed qed also have "indSum[A, \_. B] (\_. B) (\x y. y) (a,b) \ b" proof show "B : U" by (wellformed jdgmt: assms(2)) qed (assumption, (rule assms)+) finally show "snd[A,B]`(a,b) \ b" . qed end