(* Title: HoTT/Sum.thy Author: Josh Chen Date: Jun 2018 Dependent sum type. *) theory Sum imports HoTT_Base begin section \Constants and syntax\ axiomatization Sum :: "[Term, Typefam] \ Term" and pair :: "[Term, Term] \ Term" ("(1'(_,/ _'))") and indSum :: "[Term, Typefam, Typefam, [Term, Term] \ Term, Term] \ Term" ("(1ind\<^sub>\[_,/ _])") 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)" text "Nondependent pair." abbreviation Pair :: "[Term, Term] \ Term" (infixr "\" 50) where "A \ B \ \_:A. B" section \Type rules\ axiomatization where Sum_form: "\i A B. \A : U(i); B: A \ U(i)\ \ \x:A. B x : U(i)" and Sum_intro: "\i A B a b. \B: A \ U(i); a : A; b : B a\ \ (a,b) : \x:A. B x" and Sum_elim: "\i A B C f p. \ C: \x:A. B x \ U(i); \x y. \x : A; y : B x\ \ f x y : C (x,y); p : \x:A. B x \ \ ind\<^sub>\[A,B] C f p : C p" and Sum_comp: "\i A B C f a b. \ C: \x:A. B x \ U(i); \x y. \x : A; y : B x\ \ f x y : C (x,y); a : A; b : B a \ \ ind\<^sub>\[A,B] C f (a,b) \ f a b" text "Admissible inference rules for sum formation:" axiomatization where Sum_form_cond1: "\i A B. (\x:A. B x : U(i)) \ A : U(i)" and Sum_form_cond2: "\i A B. (\x:A. B x : U(i)) \ B: A \ U(i)" lemmas Sum_rules [intro] = Sum_form Sum_intro Sum_elim Sum_comp lemmas Sum_form_conds [intro (*elim, wellform*)] = Sum_form_cond1 Sum_form_cond2 lemmas Sum_comps [comp] = Sum_comp section \Null type\ axiomatization Null :: Term ("\") and indNull :: "[Typefam, Term] \ Term" ("(1ind\<^sub>\)") where Null_form: "\ : U(O)" and Null_elim: "\i C z. \C: \ \ U(i); z : \\ \ ind\<^sub>\ C z : C z" lemmas Null_rules [intro] = Null_form Null_elim end