(* Title: HoTT/Sum.thy Author: Josh Chen Date: Jun 2018 Dependent sum type. *) theory Sum imports Prod begin axiomatization Sum :: "[Term, Typefam] \ Term" and pair :: "[Term, Term] \ Term" ("(1'(_,/ _'))") and indSum :: "[Term, Typefam, Typefam, [Term, Term] \ Term] \ Term" ("(1indSum[_,/ _])") syntax "_SUM" :: "[idt, Term, Term] \ Term" ("(3\_:_./ _)" 20) "_SUM_ASCII" :: "[idt, Term, Term] \ Term" ("(3SUM _:_./ _)" 20) translations "\x:A. B" \ "CONST Sum A (\x. B)" "SUM x:A. B" \ "CONST Sum A (\x. B)" axiomatization where Sum_form [intro]: "\A B. \A : U; B: A \ U\ \ \x:A. B x : U" and Sum_intro [intro]: "\A B a b. \B: A \ U; a : A; b : B a\ \ (a,b) : \x:A. B x" and Sum_elim [elim]: "\A B C f p. \ C: \x:A. B x \ U; \x y. \x : A; y : B x\ \ f x y : C (x,y); p : \x:A. B x \ \ (indSum[A,B] C f)`p : C p" and Sum_comp [simp]: "\A B C f a b. \ C: \x:A. B x \ U; \x y. \x : A; y : B x\ \ f x y : C (x,y); a : A; b : B a \ \ (indSum[A,B] C f)`(a,b) \ f a b" \ \Nondependent pair\ abbreviation Pair :: "[Term, Term] \ Term" (infixr "\" 50) where "A \ B \ \_:A. B" section \Projections\ consts fst :: "[Term, 'a] \ Term" ("(1fst[/_,/ _])") snd :: "[Term, 'a] \ Term" ("(1snd[/_,/ _])") overloading fst_dep \ fst fst_nondep \ fst begin definition fst_dep :: "[Term, Typefam] \ Term" where "fst_dep A B \ indSum[A,B] (\_. A) (\x y. x)" definition fst_nondep :: "[Term, Term] \ Term" where "fst_nondep A B \ indSum[A, \_. B] (\_. A) (\x y. x)" end overloading snd_dep \ snd snd_nondep \ snd begin definition snd_dep :: "[Term, Typefam] \ Term" where "snd_dep A B \ indSum[A,B] (\p. B fst[A,B]`p) (\x y. y)" definition snd_nondep :: "[Term, Term] \ Term" where "snd_nondep A B \ indSum[A, \_. B] (\_. B) (\x y. y)" end text "Properties of projections:" lemma fst_dep_type: assumes "p : \x:A. B x" shows "fst[A,B]`p : A" proof - have "\x:A. B x : U" using assms .. then have "A : U" by (rule Sum_intro) unfolding fst_dep_def using assms by (rule Sum_comp) lemma fst_dep_comp: assumes "a : A" and "b : B a" shows "fst[A,B]`(a,b) \ a" proof - have "A : U" using assms(1) .. then have "\_. A: \x:A. B x \ U" . moreover have "\x y. x : A \ (\x y. x) x y : A" . ultimately show "fst[A,B]`(a,b) \ a" unfolding fst_dep_def using assms by (rule Sum_comp) qed lemma snd_dep_comp: assumes "a : A" and "b : B a" shows "snd[A,B]`(a,b) \ b" proof - have "\p. B fst[A,B]`p: \x:A. B x \ U" ultimately show "snd[A,B]`(a,b) \ b" unfolding snd_dep_def by (rule Sum_comp) qed lemma fst_nondep_comp: assumes "a : A" and "b : B" shows "fst[A,B]`(a,b) \ a" proof - have "A : U" using assms(1) .. then show "fst[A,B]`(a,b) \ a" unfolding fst_nondep_def by simp qed lemma snd_nondep_comp: "\a : A; b : B\ \ snd[A,B]`(a,b) \ b" proof - assume "a : A" and "b : B" then have "(a, b) : A \ B" .. then show "snd[A,B]`(a,b) \ b" unfolding snd_nondep_def by simp qed lemmas proj_simps [simp] = fst_dep_comp snd_dep_comp fst_nondep_comp snd_nondep_comp end