blob: 19e3939f1907004a3dde4d63f4205f76f94225de (
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
|
(*
Title: Equal.thy
Author: Joshua Chen
Date: 2018
Equality type
*)
theory Equal
imports HoTT_Base
begin
section \<open>Basic definitions\<close>
axiomatization
Equal :: "[t, t, t] \<Rightarrow> t" and
refl :: "t \<Rightarrow> t" and
indEqual :: "[t \<Rightarrow> t, t] \<Rightarrow> t" ("(1ind\<^sub>=)")
syntax
"_equal" :: "[t, t, t] \<Rightarrow> t" ("(3_ =\<^sub>_/ _)" [101, 0, 101] 100)
"_equal_ascii" :: "[t, t, t] \<Rightarrow> t" ("(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"
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>
p: x =\<^sub>A y;
x: A;
y: 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>= (\<lambda>x. f x) 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>= (\<lambda>x. f x) (refl a) \<equiv> f a"
lemmas Equal_form [form]
lemmas Equal_routine [intro] = Equal_form Equal_intro Equal_elim
lemmas Equal_comp [comp]
end
|