aboutsummaryrefslogtreecommitdiff
path: root/Coprod.thy
blob: b87958ae5b159ac33003f0cf8af9bcc4f3bc8bf7 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
(*  Title:  HoTT/Coprod.thy
    Author: Josh Chen

Coproduct type
*)

theory Coprod
  imports HoTT_Base
begin


section \<open>Constants and type rules\<close>

axiomatization
  Coprod :: "[Term, Term] \<Rightarrow> Term"  (infixr "+" 50) and
  inl :: "Term \<Rightarrow> Term" and
  inr :: "Term \<Rightarrow> Term" and
  indCoprod :: "[Term \<Rightarrow> Term, Term \<Rightarrow> Term, Term] \<Rightarrow> Term"  ("(1ind\<^sub>+)")
where
  Coprod_form: "\<lbrakk>A: U i; B: U i\<rbrakk> \<Longrightarrow> A + B: U i"
and
  Coprod_intro_inl: "\<lbrakk>a: A; B: U i\<rbrakk> \<Longrightarrow> inl a: A + B"
and
  Coprod_intro_inr: "\<lbrakk>b: B; A: U i\<rbrakk> \<Longrightarrow> inr b: A + B"
and
  Coprod_elim: "\<lbrakk>
    C: A + B \<longrightarrow> U i;
    \<And>x. x: A \<Longrightarrow> c x: C (inl x);
    \<And>y. y: B \<Longrightarrow> d y: C (inr y);
    u: A + B
    \<rbrakk> \<Longrightarrow> ind\<^sub>+ c d u : C u"
and
  Coprod_comp_inl: "\<lbrakk>
    C: A + B \<longrightarrow> U i;
    \<And>x. x: A \<Longrightarrow> c x: C (inl x);
    \<And>y. y: B \<Longrightarrow> d y: C (inr y);
    a: A
    \<rbrakk> \<Longrightarrow> ind\<^sub>+ c d (inl a) \<equiv> c a"
and
  Coprod_comp_inr: "\<lbrakk>
    C: A + B \<longrightarrow> U i;
    \<And>x. x: A \<Longrightarrow> c x: C (inl x);
    \<And>y. y: B \<Longrightarrow> d y: C (inr y);
    b: B
    \<rbrakk> \<Longrightarrow> ind\<^sub>+ c d (inr b) \<equiv> d b"


text "Admissible formation inference rules:"

axiomatization where
  Coprod_wellform1: "A + B: U i \<Longrightarrow> A: U i"
and
  Coprod_wellform2: "A + B: U i \<Longrightarrow> 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