(* Title: HoTT/Sum.thy Author: Josh Chen Date: Jun 2018 Dependent sum type. *) theory Sum imports HoTT_Base begin axiomatization Sum :: "[Term, Typefam] \ Term" and pair :: "[Term, Term] \ Term" ("(1'(_,/ _'))") and indSum :: "[Term, Typefam, Typefam, [Term, Term] \ Term, Term] \ Term" ("(1indSum[_,/ _])") section \Syntax\ 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)" section \Type rules\ axiomatization where Sum_form: "\A B. \A : U; B: A \ U\ \ \x:A. B x : U" and Sum_form_cond1: "\A B. (\x:A. B x : U) \ A : U" and Sum_form_cond2: "\A B. (\x:A. B x : U) \ B: A \ U" and Sum_intro: "\A B a b. \B: A \ U; a : A; b : B a\ \ (a,b) : \x:A. B x" and Sum_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: "\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" lemmas Sum_rules [intro] = Sum_form Sum_intro Sum_elim Sum_comp lemmas Sum_form_conds [elim, wellform] = Sum_form_cond1 Sum_form_cond2 lemmas Sum_comps [comp] = Sum_comp \ \Nondependent pair\ abbreviation Pair :: "[Term, Term] \ Term" (infixr "\" 50) where "A \ B \ \_:A. B" text "The nondependent pair needs its own separate introduction rule." lemma Pair_intro [intro]: "\A B a b. \a : A; b : B\ \ (a,b) : A \ B" proof fix b B assume "b : B" then show "B : U" .. qed end