(* Title: HoTT/Coprod.thy Author: Josh Chen Coproduct type *) theory Coprod imports HoTT_Base begin section \Constants and type rules\ axiomatization Coprod :: "[Term, Term] \ Term" (infixr "+" 50) and inl :: "Term \ Term" and inr :: "Term \ Term" and indCoprod :: "[Term \ Term, Term \ Term, Term] \ Term" ("(1ind\<^sub>+)") where Coprod_form: "\A: U(i); B: U(i)\ \ A + B: U(i)" and Coprod_intro_inl: "\a: A; B: U(i)\ \ inl(a): A + B" and Coprod_intro_inr: "\b: B; A: U(i)\ \ inr(b): A + B" and Coprod_elim: "\ C: A + B \ U(i); \x. x: A \ c(x): C(inl(x)); \y. y: B \ d(y): C(inr(y)); u: A + B \ \ ind\<^sub>+(c)(d)(u) : C(u)" and Coprod_comp_inl: "\ C: A + B \ U(i); \x. x: A \ c(x): C(inl(x)); \y. y: B \ d(y): C(inr(y)); a: A \ \ ind\<^sub>+(c)(d)(inl(a)) \ c(a)" and Coprod_comp_inr: "\ C: A + B \ U(i); \x. x: A \ c(x): C(inl(x)); \y. y: B \ d(y): C(inr(y)); b: B \ \ ind\<^sub>+(c)(d)(inr(b)) \ d(b)" text "Admissible formation inference rules:" axiomatization where Coprod_wellform1: "A + B: U(i) \ A: U(i)" and Coprod_wellform2: "A + B: U(i) \ B: U(i)" text "Rule attribute declarations:" lemmas Coprod_intro = Coprod_intro_inl Coprod_intro_inr lemmas Coprod_comp [comp] = Coprod_comp_inl Coprod_comp_inr lemmas Coprod_wellform [wellform] = Coprod_wellform1 Coprod_wellform2 lemmas Coprod_routine [intro] = Coprod_form Coprod_intro Coprod_elim end