blob: 7254104bcec06bc0e7e5100204c6d7f55f6c2af3 (
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
66
|
(* Title: HoTT/Equal.thy
Author: Josh Chen
Equality type
*)
theory Equal
imports HoTT_Base
begin
section \<open>Constants and syntax\<close>
axiomatization
Equal :: "[Term, Term, Term] \<Rightarrow> Term" and
refl :: "Term \<Rightarrow> Term" and
indEqual :: "[Term \<Rightarrow> Term, Term] \<Rightarrow> Term" ("(1ind\<^sub>=)")
syntax
"_EQUAL" :: "[Term, Term, Term] \<Rightarrow> Term" ("(3_ =\<^sub>_/ _)" [101, 0, 101] 100)
"_EQUAL_ASCII" :: "[Term, Term, Term] \<Rightarrow> Term" ("(3_ =[_]/ _)" [101, 0, 101] 100)
translations
"a =[A] b" \<rightleftharpoons> "CONST Equal A a b"
"a =\<^sub>A b" \<rightharpoonup> "CONST Equal A a b"
section \<open>Type rules\<close>
axiomatization where
Equal_form: "\<lbrakk>A: U i; a: A; b: A\<rbrakk> \<Longrightarrow> a =\<^sub>A b : U i"
and
Equal_intro: "a : A \<Longrightarrow> (refl a): a =\<^sub>A a"
and
Equal_elim: "\<lbrakk>
x: A;
y: A;
p: x =\<^sub>A y;
\<And>x. x: A \<Longrightarrow> f x: C x x (refl x);
\<And>x y. \<lbrakk>x: A; y: A\<rbrakk> \<Longrightarrow> C x y: x =\<^sub>A y \<longrightarrow> U i
\<rbrakk> \<Longrightarrow> ind\<^sub>= f p : C x y p"
and
Equal_comp: "\<lbrakk>
a: A;
\<And>x. x: A \<Longrightarrow> f x: C x x (refl x);
\<And>x y. \<lbrakk>x: A; y: A\<rbrakk> \<Longrightarrow> C x y: x =\<^sub>A y \<longrightarrow> U i
\<rbrakk> \<Longrightarrow> ind\<^sub>= f (refl a) \<equiv> f a"
text "Admissible inference rules for equality type formation:"
axiomatization where
Equal_wellform1: "a =\<^sub>A b: U i \<Longrightarrow> A: U i"
and
Equal_wellform2: "a =\<^sub>A b: U i \<Longrightarrow> a: A"
and
Equal_wellform3: "a =\<^sub>A b: U i \<Longrightarrow> b: A"
text "Rule attribute declarations:"
lemmas Equal_comp [comp]
lemmas Equal_wellform [wellform] = Equal_wellform1 Equal_wellform2 Equal_wellform3
lemmas Equal_routine [intro] = Equal_form Equal_intro Equal_elim
end
|